aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorMatthieu Sozeau2018-09-10 15:49:11 +0200
committerMatthieu Sozeau2018-09-10 15:49:11 +0200
commitea33ac1992bffc6d603760de2a46e70607db3ea0 (patch)
treeab4fc3af7da8ee5e7c60e039321c429fd871b6cd
parent2e931a4537aadfa4c15359c7b89ce4cb0291dad3 (diff)
parent6c1783de49e0b11f65cc511a94e9ce37f7210562 (diff)
Merge PR #8290: Fix #8288: cumulativity inferance ignores args to bound variables
-rw-r--r--pretyping/inferCumulativity.ml2
-rw-r--r--test-suite/bugs/closed/8288.v7
2 files changed, 8 insertions, 1 deletions
diff --git a/pretyping/inferCumulativity.ml b/pretyping/inferCumulativity.ml
index eb283a0220..be79b8b07d 100644
--- a/pretyping/inferCumulativity.ml
+++ b/pretyping/inferCumulativity.ml
@@ -99,7 +99,7 @@ let rec infer_fterm cv_pb infos variances hd stk =
| FEvar ((_,args),e) ->
let variances = infer_stack infos variances stk in
infer_vect infos variances (Array.map (mk_clos e) args)
- | FRel _ -> variances
+ | FRel _ -> infer_stack infos variances stk
| FFlex fl ->
let variances = infer_table_key infos variances fl in
infer_stack infos variances stk
diff --git a/test-suite/bugs/closed/8288.v b/test-suite/bugs/closed/8288.v
new file mode 100644
index 0000000000..0350be9c06
--- /dev/null
+++ b/test-suite/bugs/closed/8288.v
@@ -0,0 +1,7 @@
+Set Universe Polymorphism.
+Set Printing Universes.
+
+Set Polymorphic Inductive Cumulativity.
+
+Inductive foo := C : (forall A : Type -> Type, A Type) -> foo.
+(* anomaly invalid subtyping relation *)