aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Maxime Dénès <mail@maximedenes.fr>2018-03-09 15:09:30 +0100
committerGravatar Maxime Dénès <mail@maximedenes.fr>2018-03-09 15:09:30 +0100
commitd04f262f4c27a086b9ddfef1931586ef1062614d (patch)
tree14c13f9879e6c92561627e6308ba6801ea2ed2b6
parent2069c0aec44a50031e33bcb9f7d86c1535ef6b84 (diff)
parentaa95d82692684a416d7a1c25fce38b4eca1e49c9 (diff)
Merge PR #6820: Tacticals assert_fails and assert_succeeds
-rw-r--r--CHANGES3
-rw-r--r--doc/refman/RefMan-ltac.tex18
-rw-r--r--theories/Init/Tactics.v13
3 files changed, 34 insertions, 0 deletions
diff --git a/CHANGES b/CHANGES
index 1c7c53f29..a514f5934 100644
--- a/CHANGES
+++ b/CHANGES
@@ -57,6 +57,9 @@ Tactics
- A bug fixed in "rewrite H in *" and "rewrite H in * |-" may cause a
few rare incompatibilities (it was unintendedly recursively
rewriting in the side conditions generated by H).
+- Added tactics "assert_succeeds tac" and "assert_fails tac" to ensure
+ properties of the executation of a tactic without keeping the effect
+ of the execution.
Focusing
diff --git a/doc/refman/RefMan-ltac.tex b/doc/refman/RefMan-ltac.tex
index 0a4d0ef9a..3ed697d8b 100644
--- a/doc/refman/RefMan-ltac.tex
+++ b/doc/refman/RefMan-ltac.tex
@@ -156,6 +156,8 @@ is understood as
& | & {\tt type\_term} {\term}\\
& | & {\tt numgoals} \\
& | & {\tt guard} {\it test}\\
+& | & {\tt assert\_fails} {\tacexprpref}\\
+& | & {\tt assert\_succeds} {\tacexprpref}\\
& | & \atomictac\\
& | & {\qualid} \nelist{\tacarg}{}\\
& | & {\atom}
@@ -598,6 +600,22 @@ The experimental status of this tactic pertains to the fact if $v$ performs side
\ErrMsg \errindex{This tactic has more than one success}
+\subsubsection[Checking the failure]{Checking the failure\tacindex{assert\_fails}\index{Tacticals!assert\_fails@{\tt assert\_fails}}}
+
+Coq provides a derived tactic to check that a tactic \emph{fails}:
+\begin{quote}
+{\tt assert\_fails} {\tacexpr}
+\end{quote}
+This behaves like {\tt tryif {\tacexpr} then fail 0 tac "succeeds" else idtac}.
+
+\subsubsection[Checking the success]{Checking the success\tacindex{assert\_succeeds}\index{Tacticals!assert\_succeeds@{\tt assert\_succeeds}}}
+
+Coq provides a derived tactic to check that a tactic has \emph{at least one} success:
+\begin{quote}
+{\tt assert\_succeeds} {\tacexpr}
+\end{quote}
+This behaves like {\tt tryif (assert\_fails tac) then fail 0 tac "fails" else idtac}.
+
\subsubsection[Solving]{Solving\tacindex{solve}
\index{Tacticals!solve@{\tt solve}}}
diff --git a/theories/Init/Tactics.v b/theories/Init/Tactics.v
index 9e6c26b10..8df533e74 100644
--- a/theories/Init/Tactics.v
+++ b/theories/Init/Tactics.v
@@ -255,6 +255,7 @@ Tactic Notation "dependent" "induction" ident(H) :=
writing a version of [inversion] / [dependent destruction] which
does not lose information, i.e., does not turn a goal which is
provable into one which requires axiom K / UIP. *)
+
Ltac simpl_proj_exist_in H :=
repeat match type of H with
| context G[proj1_sig (exist _ ?x ?p)]
@@ -310,8 +311,20 @@ Ltac inversion_sigma_step :=
Ltac inversion_sigma := repeat inversion_sigma_step.
(** A version of [time] that works for constrs *)
+
Ltac time_constr tac :=
let eval_early := match goal with _ => restart_timer end in
let ret := tac () in
let eval_early := match goal with _ => finish_timing ( "Tactic evaluation" ) end in
ret.
+
+(** Useful combinators *)
+
+Ltac assert_fails tac :=
+ tryif tac then fail 0 tac "succeeds" else idtac.
+Ltac assert_succeeds tac :=
+ tryif (assert_fails tac) then fail 0 tac "fails" else idtac.
+Tactic Notation "assert_succeeds" tactic3(tac) :=
+ assert_succeeds tac.
+Tactic Notation "assert_fails" tactic3(tac) :=
+ assert_fails tac.