aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorThéo Zimmermann2020-11-27 16:22:33 +0100
committerThéo Zimmermann2020-11-27 16:22:41 +0100
commit058ac643c5e93da2472e3a0a717b864f49f90b3b (patch)
tree12af50d60c3956ef06375eafa26cda87779e8df8
parentc294664df8e9190a2fbb6153c70c208f58c7db70 (diff)
Revert "Remove deprecated tactic cutrewrite."
This reverts commit f3642ad8bdf6d9aa1b411892e5e6815a6a75e4d5.
-rw-r--r--doc/changelog/04-tactics/12993-remove-cutrewrite.rst4
-rw-r--r--doc/sphinx/proofs/writing-proofs/rewriting.rst7
-rw-r--r--doc/tools/docgram/common.edit_mlg3
-rw-r--r--doc/tools/docgram/fullGrammar2
-rw-r--r--doc/tools/docgram/orderedGrammar1
-rw-r--r--plugins/ltac/extratactics.mlg8
-rw-r--r--tactics/equality.ml11
-rw-r--r--tactics/equality.mli4
8 files changed, 35 insertions, 5 deletions
diff --git a/doc/changelog/04-tactics/12993-remove-cutrewrite.rst b/doc/changelog/04-tactics/12993-remove-cutrewrite.rst
deleted file mode 100644
index b719c5618e..0000000000
--- a/doc/changelog/04-tactics/12993-remove-cutrewrite.rst
+++ /dev/null
@@ -1,4 +0,0 @@
-- **Removed:**
- Deprecated ``cutrewrite`` tactic. Use :tacn:`replace` instead
- (`#12993 <https://github.com/coq/coq/pull/12993>`_,
- by Théo Zimmermann).
diff --git a/doc/sphinx/proofs/writing-proofs/rewriting.rst b/doc/sphinx/proofs/writing-proofs/rewriting.rst
index 2de6b2a18c..b7f2927000 100644
--- a/doc/sphinx/proofs/writing-proofs/rewriting.rst
+++ b/doc/sphinx/proofs/writing-proofs/rewriting.rst
@@ -146,6 +146,13 @@ simply :g:`t=u` dropping the implicit type of :g:`t` and :g:`u`.
only in the conclusion of the goal. The clause argument must not contain
any ``type of`` nor ``value of``.
+ .. tacv:: cutrewrite {? {| <- | -> } } (@term__1 = @term__2) {? in @ident }
+ :name: cutrewrite
+
+ .. deprecated:: 8.5
+
+ Use :tacn:`replace` instead.
+
.. tacn:: subst @ident
:name: subst
diff --git a/doc/tools/docgram/common.edit_mlg b/doc/tools/docgram/common.edit_mlg
index 8efda825de..75b3260166 100644
--- a/doc/tools/docgram/common.edit_mlg
+++ b/doc/tools/docgram/common.edit_mlg
@@ -1021,6 +1021,9 @@ simple_tactic: [
| DELETE "unify" constr constr
| REPLACE "unify" constr constr "with" preident
| WITH "unify" constr constr OPT ( "with" preident )
+| DELETE "cutrewrite" orient constr
+| REPLACE "cutrewrite" orient constr "in" hyp
+| WITH "cutrewrite" orient constr OPT ( "in" hyp )
| DELETE "destauto"
| REPLACE "destauto" "in" hyp
| WITH "destauto" OPT ( "in" hyp )
diff --git a/doc/tools/docgram/fullGrammar b/doc/tools/docgram/fullGrammar
index cf90eea5a1..ccf38d2c15 100644
--- a/doc/tools/docgram/fullGrammar
+++ b/doc/tools/docgram/fullGrammar
@@ -1583,6 +1583,8 @@ simple_tactic: [
| "simple" "injection" destruction_arg
| "dependent" "rewrite" orient constr
| "dependent" "rewrite" orient constr "in" hyp
+| "cutrewrite" orient constr
+| "cutrewrite" orient constr "in" hyp
| "decompose" "sum" constr
| "decompose" "record" constr
| "absurd" constr
diff --git a/doc/tools/docgram/orderedGrammar b/doc/tools/docgram/orderedGrammar
index 7c709baa48..d950b32160 100644
--- a/doc/tools/docgram/orderedGrammar
+++ b/doc/tools/docgram/orderedGrammar
@@ -1685,6 +1685,7 @@ simple_tactic: [
| "einjection" OPT destruction_arg OPT ( "as" LIST0 simple_intropattern )
| "simple" "injection" OPT destruction_arg
| "dependent" "rewrite" OPT [ "->" | "<-" ] one_term OPT ( "in" ident )
+| "cutrewrite" OPT [ "->" | "<-" ] one_term OPT ( "in" ident )
| "decompose" "sum" one_term
| "decompose" "record" one_term
| "absurd" one_term
diff --git a/plugins/ltac/extratactics.mlg b/plugins/ltac/extratactics.mlg
index 6ab82b1253..0b5d36b845 100644
--- a/plugins/ltac/extratactics.mlg
+++ b/plugins/ltac/extratactics.mlg
@@ -41,7 +41,7 @@ DECLARE PLUGIN "ltac_plugin"
(**********************************************************************)
(* replace, discriminate, injection, simplify_eq *)
-(* dependent rewrite *)
+(* cutrewrite, dependent rewrite *)
let with_delayed_uconstr ist c tac =
let flags = {
@@ -201,6 +201,12 @@ TACTIC EXTEND dependent_rewrite
-> { rewriteInHyp b c id }
END
+TACTIC EXTEND cut_rewrite
+| [ "cutrewrite" orient(b) constr(eqn) ] -> { cutRewriteInConcl b eqn }
+| [ "cutrewrite" orient(b) constr(eqn) "in" hyp(id) ]
+ -> { cutRewriteInHyp b eqn id }
+END
+
(**********************************************************************)
(* Decompose *)
diff --git a/tactics/equality.ml b/tactics/equality.ml
index 486575d229..fcdd23a9c1 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -1655,6 +1655,17 @@ let cutSubstClause l2r eqn cls =
| None -> cutSubstInConcl l2r eqn
| Some id -> cutSubstInHyp l2r eqn id
+let warn_deprecated_cutrewrite =
+ CWarnings.create ~name:"deprecated-cutrewrite" ~category:"deprecated"
+ (fun () -> strbrk"\"cutrewrite\" is deprecated. Use \"replace\" instead.")
+
+let cutRewriteClause l2r eqn cls =
+ warn_deprecated_cutrewrite ();
+ try_rewrite (cutSubstClause l2r eqn cls)
+
+let cutRewriteInHyp l2r eqn id = cutRewriteClause l2r eqn (Some id)
+let cutRewriteInConcl l2r eqn = cutRewriteClause l2r eqn None
+
let substClause l2r c cls =
Proofview.Goal.enter begin fun gl ->
let eq = pf_apply get_type_of gl c in
diff --git a/tactics/equality.mli b/tactics/equality.mli
index 5a4fe47cab..fdcbbc0e3c 100644
--- a/tactics/equality.mli
+++ b/tactics/equality.mli
@@ -107,6 +107,10 @@ val dEqThen : keep_proofs:(bool option) -> evars_flag -> (clear_flag -> constr -
val make_iterated_tuple :
env -> evar_map -> constr -> (constr * types) -> evar_map * (constr * constr * constr)
+(* The family cutRewriteIn expect an equality statement *)
+val cutRewriteInHyp : bool -> types -> Id.t -> unit Proofview.tactic
+val cutRewriteInConcl : bool -> constr -> unit Proofview.tactic
+
(* The family rewriteIn expect the proof of an equality *)
val rewriteInHyp : bool -> constr -> Id.t -> unit Proofview.tactic
val rewriteInConcl : bool -> constr -> unit Proofview.tactic