aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorGaëtan Gilbert2019-11-05 15:06:23 +0100
committerGaëtan Gilbert2019-11-05 15:38:43 +0100
commitfea52cee63b678dfda766a0a63435c134d28d77c (patch)
tree95933faa1bf07913b1609ec2af646444d315e033
parent634cb7b8a07a34fc29d074591091f0a6170f7bff (diff)
Fix #11048: uncaught destKO in inductive type
Reduction.whd_all does not commute with to_constr.
-rw-r--r--test-suite/bugs/closed/bug_11048.v5
-rw-r--r--vernac/comInductive.ml5
2 files changed, 6 insertions, 4 deletions
diff --git a/test-suite/bugs/closed/bug_11048.v b/test-suite/bugs/closed/bug_11048.v
new file mode 100644
index 0000000000..d1211587f1
--- /dev/null
+++ b/test-suite/bugs/closed/bug_11048.v
@@ -0,0 +1,5 @@
+
+Inductive foo (HUGE := _) (b : HUGE) A :=
+ bar (X:match _ : HUGE as T return HUGE with T => match (A : T) -> True with _ => T end end)
+ : foo b A.
+(* uncaught destKO *)
diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml
index 36aa7a37a2..80fcb7bc45 100644
--- a/vernac/comInductive.ml
+++ b/vernac/comInductive.ml
@@ -12,7 +12,6 @@ open Pp
open CErrors
open Sorts
open Util
-open Constr
open Context
open Environ
open Names
@@ -164,9 +163,7 @@ let sign_level env evd sign =
match d with
| LocalDef _ -> lev, push_rel d env
| LocalAssum _ ->
- let s = destSort (Reduction.whd_all env
- (EConstr.to_constr evd (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d)))))
- in
+ let s = Retyping.get_sort_of env evd (EConstr.of_constr (RelDecl.get_type d)) in
let u = univ_of_sort s in
(Univ.sup u lev, push_rel d env))
sign (Univ.Universe.sprop,env))