aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--kernel/univ.ml4
-rw-r--r--tactics/tactics.ml1
-rw-r--r--test-suite/bugs/closed/shouldsucceed/2299.v13
3 files changed, 16 insertions, 2 deletions
diff --git a/kernel/univ.ml b/kernel/univ.ml
index ef2024c7a..0a22733d0 100644
--- a/kernel/univ.ml
+++ b/kernel/univ.ml
@@ -56,7 +56,7 @@ let cmp_univ_level u v = match u,v with
else compare dp1 dp2
let string_of_univ_level = function
- | Set -> "0"
+ | Set -> "Set"
| Level (d,n) -> Names.string_of_dirpath d^"."^string_of_int n
module UniverseLMap =
@@ -603,7 +603,7 @@ let pr_constraints c =
in pp_std ++ pr_uni_level u1 ++ str op_str ++
pr_uni_level u2 ++ fnl () ) c (str "")
-(* Dumping constrains to a file *)
+(* Dumping constraints to a file *)
let dump_universes output g =
let dump_arc _ = function
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 51aa86613..41fab4e71 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -2903,6 +2903,7 @@ let apply_induction_with_discharge induct_tac elim indhyps destopt avoid names t
let apply_induction_in_context hyp0 elim indvars names induct_tac gl =
let env = pf_env gl in
let statuslists,lhyp0,indhyps,deps = cook_sign hyp0 indvars env in
+ let deps = List.map (on_pi3 refresh_universes_strict) deps in
let tmpcl = it_mkNamedProd_or_LetIn (pf_concl gl) deps in
let dephyps = List.map (fun (id,_,_) -> id) deps in
let deps_cstr =
diff --git a/test-suite/bugs/closed/shouldsucceed/2299.v b/test-suite/bugs/closed/shouldsucceed/2299.v
new file mode 100644
index 000000000..c0552ca7b
--- /dev/null
+++ b/test-suite/bugs/closed/shouldsucceed/2299.v
@@ -0,0 +1,13 @@
+(* Check that destruct refreshes universes in what it generalizes *)
+
+Section test.
+
+Variable A: Type.
+
+Inductive T: unit -> Type := C: A -> unit -> T tt.
+
+Let unused := T tt.
+
+Goal T tt -> False.
+ intro X.
+ destruct X.