aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorPierre-Marie Pédrot2020-03-05 12:40:17 +0100
committerPierre-Marie Pédrot2020-03-05 12:40:17 +0100
commitdca89bdb43c1fe557f1cf681da273f6a8993c338 (patch)
tree068e758c7eb550bb35596f2f253e1ca4e6dc328f
parenta3d1646b59b4233b87b902b627583cf9f028311d (diff)
parent7c1733040a19846359cdf7725b20c93398b50d3b (diff)
Merge PR #11522: Adding an alias `pose proof (x:=t)` for `pose proof t as x` following the model of `pose (x:=t)`.
Reviewed-by: Zimmi48 Reviewed-by: ppedrot
-rw-r--r--doc/changelog/04-tactics/11522-master+pose-proof-wo-as-syntax.rst6
-rw-r--r--doc/sphinx/proof-engine/tactics.rst13
-rw-r--r--plugins/ltac/g_tactic.mlg10
-rw-r--r--test-suite/success/pose.v9
4 files changed, 38 insertions, 0 deletions
diff --git a/doc/changelog/04-tactics/11522-master+pose-proof-wo-as-syntax.rst b/doc/changelog/04-tactics/11522-master+pose-proof-wo-as-syntax.rst
new file mode 100644
index 0000000000..3dd103b115
--- /dev/null
+++ b/doc/changelog/04-tactics/11522-master+pose-proof-wo-as-syntax.rst
@@ -0,0 +1,6 @@
+- **Added:**
+ Syntax :n:`pose proof (@ident:=@term)` as an
+ alternative to :n:`pose proof @term as @ident`, following the model of
+ :n:`pose (@ident:=@term)`. See documentation of :tacn:`pose proof`
+ (`#11522 <https://github.com/coq/coq/pull/11522>`_,
+ by Hugo Herbelin).
diff --git a/doc/sphinx/proof-engine/tactics.rst b/doc/sphinx/proof-engine/tactics.rst
index 6a0ce20c79..206342fef1 100644
--- a/doc/sphinx/proof-engine/tactics.rst
+++ b/doc/sphinx/proof-engine/tactics.rst
@@ -1452,6 +1452,19 @@ Controlling the proof flow
While :tacn:`pose proof` expects that no existential variables are generated by
the tactic, :tacn:`epose proof` removes this constraint.
+.. tacv:: pose proof (@ident := @term)
+
+ This is an alternative syntax for :n:`assert (@ident := @term)` and
+ :n:`pose proof @term as @ident`, following the model of :n:`pose
+ (@ident := @term)` but dropping the value of :token:`ident`.
+
+.. tacv:: epose proof (@ident := @term)
+
+ This is an alternative syntax for :n:`eassert (@ident := @term)`
+ and :n:`epose proof @term as @ident`, following the model of
+ :n:`epose (@ident := @term)` but dropping the value of
+ :token:`ident`.
+
.. tacv:: enough (@ident : @type)
:name: enough
diff --git a/plugins/ltac/g_tactic.mlg b/plugins/ltac/g_tactic.mlg
index 3e4c7ba782..8e1e5559af 100644
--- a/plugins/ltac/g_tactic.mlg
+++ b/plugins/ltac/g_tactic.mlg
@@ -589,6 +589,16 @@ GRAMMAR EXTEND Gram
{ TacAtom (CAst.make ~loc @@ TacAssert (false,true,Some tac,ipat,c)) }
| IDENT "eassert"; c = constr; ipat = as_ipat; tac = by_tactic ->
{ TacAtom (CAst.make ~loc @@ TacAssert (true,true,Some tac,ipat,c)) }
+
+ (* Alternative syntax for "pose proof c as id by tac" *)
+ | IDENT "pose"; IDENT "proof"; test_lpar_id_coloneq; "("; lid = identref; ":=";
+ c = lconstr; ")" ->
+ { let { CAst.loc = loc; v = id } = lid in
+ TacAtom (CAst.make ?loc @@ TacAssert (false,true,None,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) }
+ | IDENT "epose"; IDENT "proof"; test_lpar_id_coloneq; "("; lid = identref; ":=";
+ c = lconstr; ")" ->
+ { let { CAst.loc = loc; v = id } = lid in
+ TacAtom (CAst.make ?loc @@ TacAssert (true,true,None,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) }
| IDENT "pose"; IDENT "proof"; c = lconstr; ipat = as_ipat ->
{ TacAtom (CAst.make ~loc @@ TacAssert (false,true,None,ipat,c)) }
| IDENT "epose"; IDENT "proof"; c = lconstr; ipat = as_ipat ->
diff --git a/test-suite/success/pose.v b/test-suite/success/pose.v
new file mode 100644
index 0000000000..17007915fe
--- /dev/null
+++ b/test-suite/success/pose.v
@@ -0,0 +1,9 @@
+(* Test syntax *)
+
+Goal 0=0.
+pose proof (a := I).
+Fail clearbody a.
+epose proof (b := fun _ => eq_refl).
+Fail clearbody b.
+exact (b a).
+Qed.