Skip to content

Commit

Permalink
Merge PR #13317: [ssr] intro pattern extensions for dup, swap and apply
Browse files Browse the repository at this point in the history
Reviewed-by: gares
Ack-by: Zimmi48
  • Loading branch information
coqbot-app[bot] committed Nov 12, 2020
2 parents af42e1b + 7bd93b0 commit 9dfc627
Show file tree
Hide file tree
Showing 6 changed files with 103 additions and 1 deletion.
4 changes: 4 additions & 0 deletions doc/changelog/06-ssreflect/13317-ssr_dup_swap_apply_ipat.rst
Original file line number Diff line number Diff line change
@@ -0,0 +1,4 @@
- **Added:**
SSReflect intro pattern ltac views ``/[dup]``, ``/[swap]`` and ``/[apply]``
(`#13317 <https://github.com/coq/coq/pull/13317>`_,
by Cyril Cohen).
5 changes: 4 additions & 1 deletion doc/sphinx/proof-engine/ssreflect-proof-language.rst
Original file line number Diff line number Diff line change
Expand Up @@ -1647,7 +1647,10 @@ Notations can be used to name tactics, for example
Notation "'myop'" := (ltac:(my ltac code)) : ssripat_scope.

lets one write just ``/myop`` in the intro pattern. Note the scope
annotation: views are interpreted opening the ``ssripat`` scope.
annotation: views are interpreted opening the ``ssripat`` scope. We
provide the following ltac views: ``/[dup]`` to duplicate the top of
the stack, ``/[swap]`` to swap the two first elements and ``/[apply]``
to apply the top of the stack to the next.

Intro patterns
``````````````
Expand Down
13 changes: 13 additions & 0 deletions test-suite/ssr/ipat_apply.v
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
Require Import ssreflect.

Section Apply.

Variable P : nat -> Prop.
Lemma test_apply A B : forall (f : A -> B) (a : A), B.

Proof.
move=> /[apply] b.
exact.
Qed.

End Apply.
13 changes: 13 additions & 0 deletions test-suite/ssr/ipat_dup.v
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
Require Import ssreflect.

Section Dup.

Variable P : nat -> Prop.

Lemma test_dup1 : forall n : nat, P n.
Proof. move=> /[dup] m n; suff: P n by []. Abort.

Lemma test_dup2 : let n := 1 in False.
Proof. move=> /[dup] m n; have : m = n := eq_refl. Abort.

End Dup.
13 changes: 13 additions & 0 deletions test-suite/ssr/ipat_swap.v
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
Require Import ssreflect.

Section Swap.

Definition P n := match n with 1 => true | _ => false end.

Lemma test_swap1 : forall (n : nat) (b : bool), P n = b.
Proof. move=> /[swap] b n; suff: P n = b by []. Abort.

Lemma test_swap1 : let n := 1 in let b := true in False.
Proof. move=> /[swap] b n; have : P n = b := eq_refl. Abort.

End Swap.
56 changes: 56 additions & 0 deletions theories/ssr/ssreflect.v
Original file line number Diff line number Diff line change
Expand Up @@ -59,6 +59,15 @@ Declare ML Module "ssreflect_plugin".
Canonical foo_unlockable := #[#unlockable fun foo#]#.
This minimizes the comparison overhead for foo, while still allowing
rewrite unlock to expose big_foo_expression.
Additionally we provide default intro pattern ltac views:
- top of the stack actions:
=> /[apply] := => hyp {}/hyp
=> /[swap] := => x y; move: y x
(also swap and perserves let bindings)
=> /[dup] := => x; have copy := x; move: copy x
(also copies and preserves let bindings)
More information about these definitions and their use can be found in the
ssreflect manual, and in specific comments below. **)

Expand Down Expand Up @@ -654,3 +663,50 @@ End Exports.

End NonPropType.
Export NonPropType.Exports.

Module Export ipat.

Notation "'[' 'apply' ']'" := (ltac:(let f := fresh "_top_" in move=> f {}/f))
(at level 0, only parsing) : ssripat_scope.

(** We try to preserve the naming by matching the names from the goal.
We do 'move' to perform a hnf before trying to match. **)
Notation "'[' 'swap' ']'" := (ltac:(move;
lazymatch goal with
| |- forall (x : _), _ => let x := fresh x in move=> x; move;
lazymatch goal with
| |- forall (y : _), _ => let y := fresh y in move=> y; move: y x
| |- let y := _ in _ => let y := fresh y in move=> y; move: @y x
| _ => let y := fresh "_top_" in move=> y; move: y x
end
| |- let x := _ in _ => let x := fresh x in move => x; move;
lazymatch goal with
| |- forall (y : _), _ => let y := fresh y in move=> y; move: y @x
| |- let y := _ in _ => let y := fresh y in move=> y; move: @y @x
| _ => let y := fresh "_top_" in move=> y; move: y x
end
| _ => let x := fresh "_top_" in let x := fresh x in move=> x; move;
lazymatch goal with
| |- forall (y : _), _ => let y := fresh y in move=> y; move: y @x
| |- let y := _ in _ => let y := fresh y in move=> y; move: @y @x
| _ => let y := fresh "_top_" in move=> y; move: y x
end
end))
(at level 0, only parsing) : ssripat_scope.

Notation "'[' 'dup' ']'" := (ltac:(move;
lazymatch goal with
| |- forall (x : _), _ =>
let x := fresh x in move=> x;
let copy := fresh x in have copy := x; move: copy x
| |- let x := _ in _ =>
let x := fresh x in move=> x;
let copy := fresh x in pose copy := x;
do [unfold x in (value of copy)]; move: @copy @x
| |- _ =>
let x := fresh "_top_" in move=> x;
let copy := fresh "_top" in have copy := x; move: copy x
end))
(at level 0, only parsing) : ssripat_scope.

End ipat.

0 comments on commit 9dfc627

Please sign in to comment.