aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--engine/uState.ml3
-rw-r--r--test-suite/bugs/closed/bug_10161.v8
2 files changed, 10 insertions, 1 deletions
diff --git a/engine/uState.ml b/engine/uState.ml
index adea78d4c9..9b27ff7916 100644
--- a/engine/uState.ml
+++ b/engine/uState.ml
@@ -214,7 +214,8 @@ let process_universe_constraints ctx cstrs =
| Inr l, Inl r | Inl r, Inr l ->
let alg = LSet.mem l ctx.uctx_univ_algebraic in
let inst = univ_level_rem l r r in
- if alg then (instantiate_variable l inst vars; local)
+ if alg && not (LSet.mem l (Universe.levels inst)) then
+ (instantiate_variable l inst vars; local)
else
let lu = Universe.make l in
if univ_level_mem l r then
diff --git a/test-suite/bugs/closed/bug_10161.v b/test-suite/bugs/closed/bug_10161.v
new file mode 100644
index 0000000000..3d262b89fe
--- /dev/null
+++ b/test-suite/bugs/closed/bug_10161.v
@@ -0,0 +1,8 @@
+Inductive SwitchT (A : Type) : Type :=
+| switchT : forall T, SwitchT T -> SwitchT A.
+
+Set Printing Universes.
+
+Fail Inductive UseSwitchT :=
+| useSwitchT : SwitchT UseSwitchT -> UseSwitchT.
+(* used to stack overflow, should be univ inconsistency cannot satisfy u = u+1 *)