aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--tactics/tactics.ml9
-rw-r--r--test-suite/success/intros.v8
2 files changed, 16 insertions, 1 deletions
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index d0724804b..ad7ff14e6 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -2209,6 +2209,11 @@ and intro_pattern_action loc b style pat thin destopt tac id = match pat with
| IntroApplyOn (f,(loc,pat)) ->
let naming,tac_ipat =
prepare_intros_loc loc (IntroIdentifier id) destopt pat in
+ let doclear =
+ if naming = NamingMustBe (loc,id) then
+ Proofview.tclUNIT () (* apply_in_once do a replacement *)
+ else
+ Proofview.V82.tactic (clear [id]) in
Proofview.Goal.enter begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
@@ -2217,7 +2222,9 @@ and intro_pattern_action loc b style pat thin destopt tac id = match pat with
(Tacticals.New.tclTHENFIRST
(* Skip the side conditions of the apply *)
(apply_in_once false true true true naming id
- (None,(sigma,(c,NoBindings))) tac_ipat) (tac ((dloc,id)::thin) None []))
+ (None,(sigma,(c,NoBindings)))
+ (fun id -> Tacticals.New.tclTHEN doclear (tac_ipat id)))
+ (tac thin None []))
sigma
end
diff --git a/test-suite/success/intros.v b/test-suite/success/intros.v
index ae1694c58..35ba94fb6 100644
--- a/test-suite/success/intros.v
+++ b/test-suite/success/intros.v
@@ -61,3 +61,11 @@ Goal forall n, n = S n -> 0=0.
intros n H/n_Sn.
destruct H.
Qed.
+
+(* Another check about generated names and cleared hypotheses with
+ pat/c patterns *)
+Goal (True -> 0=0 /\ 1=1) -> True -> 0=0.
+intros H (H1,?)/H.
+change (1=1) in H0.
+exact H1.
+Qed.