aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs/closed/4869.v
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <mattam@mattam.org>2016-09-29 15:51:18 +0200
committerGravatar Matthieu Sozeau <mattam@mattam.org>2016-09-29 15:56:10 +0200
commit9615c025a2a09b69f2001d44a66a1fddef74e680 (patch)
tree4643d4f0f38b89402265df46c56119e28a8afe40 /test-suite/bugs/closed/4869.v
parent5348a615a484e379896deac8a6944af1f92b2d4c (diff)
Fix bug #4869, allow Prop, Set, and level names in constraints.
Diffstat (limited to 'test-suite/bugs/closed/4869.v')
-rw-r--r--test-suite/bugs/closed/4869.v18
1 files changed, 18 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/4869.v b/test-suite/bugs/closed/4869.v
new file mode 100644
index 000000000..6d21b66fe
--- /dev/null
+++ b/test-suite/bugs/closed/4869.v
@@ -0,0 +1,18 @@
+Universes i.
+
+Fail Constraint i < Set.
+Fail Constraint i <= Set.
+Fail Constraint i = Set.
+Constraint Set <= i.
+Constraint Set < i.
+Fail Constraint i < j. (* undeclared j *)
+Fail Constraint i < Type. (* anonymous *)
+
+Set Universe Polymorphism.
+
+Section Foo.
+ Universe j.
+ Constraint Set < j.
+
+ Definition foo := Type@{j}.
+End Foo. \ No newline at end of file