aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2016-09-30 17:25:58 +0200
committerGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2016-09-30 18:19:31 +0200
commit064de6f6839c4ef963b83018812c5d4113eb2bb9 (patch)
treeb375d90699c0b7304b5820f4b32d87bfb84c0c8c
parent024cf5ae087024399cc894b121437d72cd11b480 (diff)
Fix bug #5045: [generalize] creates ill-typed terms in 8.6.
-rw-r--r--tactics/tactics.ml2
-rw-r--r--test-suite/bugs/closed/5045.v3
2 files changed, 5 insertions, 0 deletions
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 2fe8e0bc3..c550df101 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -2821,10 +2821,12 @@ let generalize_dep ?(with_let = false) c =
(** *)
let generalize_gen_let lconstr = Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ let env = Proofview.Goal.env gl in
let newcl, evd =
List.fold_right_i (Tacmach.New.of_old generalize_goal gl) 0 lconstr
(Tacmach.New.pf_concl gl,Tacmach.New.project gl)
in
+ let (evd, _) = Typing.type_of env evd newcl in
let map ((_, c, b),_) = if Option.is_empty b then Some c else None in
let tac = apply_type newcl (List.map_filter map lconstr) in
Sigma.Unsafe.of_pair (tac, evd)
diff --git a/test-suite/bugs/closed/5045.v b/test-suite/bugs/closed/5045.v
new file mode 100644
index 000000000..dc38738d8
--- /dev/null
+++ b/test-suite/bugs/closed/5045.v
@@ -0,0 +1,3 @@
+Axiom silly : 1 = 1 -> nat -> nat.
+Goal forall pf : 1 = 1, silly pf 0 = 0 -> True.
+ Fail generalize (@eq nat).