diff options
| author | Gaëtan Gilbert | 2020-11-04 15:02:11 +0100 |
|---|---|---|
| committer | Gaëtan Gilbert | 2020-11-16 11:24:56 +0100 |
| commit | f09eef2a1c55a1e9a617b7bad0759b4760a8978a (patch) | |
| tree | 35a7d229335f482714dd7053b1158f3eef2fe871 | |
| parent | 41840f92c0ee8c10691411f2a44e4059a4754456 (diff) | |
Only lower inductives to Prop if the type is syntactically an arity.
Fix #13300
| -rw-r--r-- | test-suite/bugs/closed/bug_13300.v | 7 | ||||
| -rw-r--r-- | vernac/comInductive.ml | 9 |
2 files changed, 14 insertions, 2 deletions
diff --git a/test-suite/bugs/closed/bug_13300.v b/test-suite/bugs/closed/bug_13300.v new file mode 100644 index 0000000000..e4fcd6dacc --- /dev/null +++ b/test-suite/bugs/closed/bug_13300.v @@ -0,0 +1,7 @@ +Polymorphic Definition type := Type. + +Inductive bad : type := . + +Fail Check bad : Prop. +Check bad : Set. +(* lowered as much as possible *) diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml index cc3f7bc093..6cf6606a93 100644 --- a/vernac/comInductive.ml +++ b/vernac/comInductive.ml @@ -327,8 +327,13 @@ let inductive_levels env evd arities inds = let duu = Sorts.univ_of_sort du in let template_prop, evd = if not (Univ.is_small_univ duu) && Univ.Universe.equal cu duu then - if is_flexible_sort evd duu && not (Evd.check_leq evd Univ.type0_univ duu) then - true, Evd.set_eq_sort env evd Sorts.prop du + if is_flexible_sort evd duu && not (Evd.check_leq evd Univ.type0_univ duu) + then if Term.isArity arity + (* If not a syntactic arity, the universe may be used in a + polymorphic instance and so cannot be lowered to Prop. + See #13300. *) + then true, Evd.set_eq_sort env evd Sorts.prop du + else false, Evd.set_eq_sort env evd Sorts.set du else false, evd else false, Evd.set_eq_sort env evd (sort_of_univ cu) du in |
