aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--library/lib.ml2
-rw-r--r--test-suite/bugs/closed/4816.v18
2 files changed, 19 insertions, 1 deletions
diff --git a/library/lib.ml b/library/lib.ml
index e4617cafb..f6b4a2458 100644
--- a/library/lib.ml
+++ b/library/lib.ml
@@ -417,7 +417,7 @@ let add_section_variable id impl poly ctx =
match !sectab with
| [] -> () (* because (Co-)Fixpoint temporarily uses local vars *)
| (vars,repl,abs)::sl ->
- check_same_poly poly vars;
+ List.iter (fun tab -> check_same_poly poly (pi1 tab)) !sectab;
sectab := (Variable (id,impl,poly,ctx)::vars,repl,abs)::sl
let add_section_context ctx =
diff --git a/test-suite/bugs/closed/4816.v b/test-suite/bugs/closed/4816.v
new file mode 100644
index 000000000..ef79b9869
--- /dev/null
+++ b/test-suite/bugs/closed/4816.v
@@ -0,0 +1,18 @@
+Section foo.
+Polymorphic Universes A B.
+Constraint A <= B.
+End foo.
+(* gives an anomaly Universe undefined *)
+
+or even, a refinement of #4503:
+Require Coq.Classes.RelationClasses.
+
+Class PreOrder (A : Type) (r : A -> A -> Type) : Type :=
+{ refl : forall x, r x x }.
+
+Section foo.
+ Polymorphic Universes A.
+ Section bar.
+ Context {A : Type@{A}} {rA : A -> A -> Prop} {PO : PreOrder A rA}.
+ End bar.
+End foo. \ No newline at end of file