aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <mattam@mattam.org>2015-10-07 16:51:53 +0200
committerGravatar Matthieu Sozeau <mattam@mattam.org>2015-10-07 16:51:53 +0200
commite26b4dbedd29acbfb9cbf2320193cc68afa60cf3 (patch)
tree397f85c4b33d5b7339027371df67f35e09272c6d
parent08a0c44e3525d1f0c7303d189e826e25c3e3d914 (diff)
Fix bug #4069: f_equal regression.
-rw-r--r--plugins/cc/cctac.ml17
-rw-r--r--test-suite/bugs/closed/4069.v51
2 files changed, 63 insertions, 5 deletions
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index 6439f58d2..cbd95eaea 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -482,6 +482,15 @@ let congruence_tac depth l =
This isn't particularly related with congruence, apart from
the fact that congruence is called internally.
*)
+
+let new_app_global_check f args k =
+ new_app_global f args
+ (fun c ->
+ Proofview.Goal.enter
+ begin fun gl ->
+ let evm, _ = Tacmach.New.pf_apply type_of gl c in
+ Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) (k c)
+ end)
let f_equal =
Proofview.Goal.nf_enter begin fun gl ->
@@ -490,11 +499,9 @@ let f_equal =
let cut_eq c1 c2 =
try (* type_of can raise an exception *)
let ty = (* Termops.refresh_universes *) (type_of c1) in
- if eq_constr_nounivs c1 c2 then Proofview.tclUNIT ()
- else
- Tacticals.New.tclTRY (Tacticals.New.tclTHEN
- ((new_app_global _eq [|ty; c1; c2|]) Tactics.cut)
- (Tacticals.New.tclTRY ((new_app_global _refl_equal [||]) apply)))
+ Tacticals.New.tclTHEN
+ ((new_app_global_check _eq [|ty; c1; c2|]) Tactics.cut)
+ (Tacticals.New.tclTRY ((new_app_global _refl_equal [||]) apply))
with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e
in
Proofview.tclORELSE
diff --git a/test-suite/bugs/closed/4069.v b/test-suite/bugs/closed/4069.v
new file mode 100644
index 000000000..21b03ce54
--- /dev/null
+++ b/test-suite/bugs/closed/4069.v
@@ -0,0 +1,51 @@
+
+Lemma test1 :
+forall (v : nat) (f g : nat -> nat),
+f v = g v.
+intros. f_equal.
+(*
+Goal in v8.5: f v = g v
+Goal in v8.4: v = v -> f v = g v
+Expected: f = g
+*)
+Admitted.
+
+Lemma test2 :
+forall (v u : nat) (f g : nat -> nat),
+f v = g u.
+intros. f_equal.
+(*
+In both v8.4 And v8.5
+Goal 1: v = u -> f v = g u
+Goal 2: v = u
+
+Expected Goal 1: f = g
+Expected Goal 2: v = u
+*)
+Admitted.
+
+Lemma test3 :
+forall (v : nat) (u : list nat) (f : nat -> nat) (g : list nat -> nat),
+f v = g u.
+intros. f_equal.
+(*
+In both v8.4 And v8.5, the goal is unchanged.
+*)
+Admitted.
+
+Require Import List.
+Lemma foo n (l k : list nat) : k ++ skipn n l = skipn n l.
+Proof. f_equal.
+(*
+ 8.4: leaves the goal unchanged, i.e. k ++ skipn n l = skipn n l
+ 8.5: 2 goals, skipn n l = l -> k ++ skipn n l = skipn n l
+ and skipn n l = l
+*)
+Require Import List.
+Fixpoint replicate {A} (n : nat) (x : A) : list A :=
+ match n with 0 => nil | S n => x :: replicate n x end.
+Lemma bar {A} n m (x : A) :
+ skipn n (replicate m x) = replicate (m - n) x ->
+ skipn n (replicate m x) = replicate (m - n) x.
+Proof. intros. f_equal.
+(* 8.5: one goal, n = m - n *)