aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--proofs/tacmach.ml3
-rw-r--r--proofs/tacmach.mli1
-rw-r--r--tactics/tactics.ml2
-rw-r--r--test-suite/bugs/closed/3441.v23
4 files changed, 28 insertions, 1 deletions
diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml
index 2b129ad89..330594af5 100644
--- a/proofs/tacmach.ml
+++ b/proofs/tacmach.ml
@@ -171,6 +171,9 @@ module New = struct
let pf_unsafe_type_of gl t =
pf_apply unsafe_type_of gl t
+ let pf_get_type_of gl t =
+ pf_apply (Retyping.get_type_of ~lax:true) gl t
+
let pf_type_of gl t =
pf_apply type_of gl t
diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli
index 727efcf6d..f79fa1d4b 100644
--- a/proofs/tacmach.mli
+++ b/proofs/tacmach.mli
@@ -109,6 +109,7 @@ module New : sig
val pf_concl : ([ `NF ], 'r) Proofview.Goal.t -> types
val pf_unsafe_type_of : ('a, 'r) Proofview.Goal.t -> Term.constr -> Term.types
+ val pf_get_type_of : ('a, 'r) Proofview.Goal.t -> Term.constr -> Term.types
val pf_type_of : ('a, 'r) Proofview.Goal.t -> Term.constr -> evar_map * Term.types
val pf_conv_x : ('a, 'r) Proofview.Goal.t -> Term.constr -> Term.constr -> bool
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 548d2a81f..92cb8a11e 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -2714,7 +2714,7 @@ let forward b usetac ipat c =
match usetac with
| None ->
Proofview.Goal.enter { enter = begin fun gl ->
- let t = Tacmach.New.pf_unsafe_type_of gl c in
+ let t = Tacmach.New.pf_get_type_of gl c in
let hd = head_ident c in
Tacticals.New.tclTHENFIRST (assert_as true hd ipat t) (exact_no_check c)
end }
diff --git a/test-suite/bugs/closed/3441.v b/test-suite/bugs/closed/3441.v
new file mode 100644
index 000000000..50d297807
--- /dev/null
+++ b/test-suite/bugs/closed/3441.v
@@ -0,0 +1,23 @@
+Axiom f : nat -> nat -> nat.
+Fixpoint do_n (n : nat) (k : nat) :=
+ match n with
+ | 0 => k
+ | S n' => do_n n' (f k k)
+ end.
+
+Notation big := (_ = _).
+Axiom k : nat.
+Goal True.
+Timeout 1 let H := fresh "H" in
+ let x := constr:(let n := 17 in do_n n = do_n n) in
+ let y := (eval lazy in x) in
+ pose proof y as H. (* Finished transaction in 1.102 secs (1.084u,0.016s) (successful) *)
+Timeout 1 let H := fresh "H" in
+ let x := constr:(let n := 17 in do_n n = do_n n) in
+ let y := (eval lazy in x) in
+ pose y as H; clearbody H. (* Finished transaction in 0.412 secs (0.412u,0.s) (successful) *)
+
+Timeout 1 Time let H := fresh "H" in
+ let x := constr:(let n := 17 in do_n n = do_n n) in
+ let y := (eval lazy in x) in
+ assert (H := y). (* Finished transaction in 1.19 secs (1.164u,0.024s) (successful) *) \ No newline at end of file