aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs/closed/4726.v
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <matthieu.sozeau@inria.fr>2016-05-30 23:52:58 +0200
committerGravatar Matthieu Sozeau <mattam@mattam.org>2016-06-29 11:52:52 +0200
commitc200c2b41e88dd7d4a1b9e90e0c35a7ed047309c (patch)
tree1ff729200d9e587ace74d4b83611af4a69e6afdf /test-suite/bugs/closed/4726.v
parentd4cdb7e41844e1bb77bac9a7b9df423364b996e2 (diff)
Univs: Fix bug #4726
When using Record and an explicit sort constraint, the universe was wrongly made flexible and minimized.
Diffstat (limited to 'test-suite/bugs/closed/4726.v')
-rw-r--r--test-suite/bugs/closed/4726.v19
1 files changed, 19 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/4726.v b/test-suite/bugs/closed/4726.v
new file mode 100644
index 000000000..0037b6fde
--- /dev/null
+++ b/test-suite/bugs/closed/4726.v
@@ -0,0 +1,19 @@
+Set Universe Polymorphism.
+
+Definition le@{i j} : Type@{j} :=
+ (fun A : Type@{j} => A)
+ (unit : Type@{i}).
+Definition eq@{i j} : Type@{j} := let x := le@{i j} in le@{j i}.
+
+Record Inj@{i j} (A : Type@{i}) (B : Type@{j}) : Type@{j} :=
+ { inj : A }.
+
+Monomorphic Universe u1.
+Let ty1 : Type@{u1} := Set.
+Check Inj@{Set u1}.
+(* Would fail with univ inconsistency if the universe was minimized *)
+
+Record Inj'@{i j} (A : Type@{i}) (B : Type@{j}) : Type@{j} :=
+ { inj' : A; foo : Type@{j} := eq@{i j} }.
+Fail Check Inj'@{Set u1}. (* Do not drop constraint i = j *)
+Check Inj'@{Set Set}.