Skip to content
Snippets Groups Projects
Commit 69984fab authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Merge branch 'propsetbinder' into 'master'

Allow pattern and type annotations in propset notation

See merge request iris/stdpp!533
parents 62680a79 fc4a5089
No related branches found
No related tags found
1 merge request!533Allow pattern and type annotations in propset notation
Pipeline #97947 passed
......@@ -28,6 +28,8 @@ Coq 8.19 is newly supported by this version of std++.
to refer to a hypothesis on the goal (`inv 1`).
- Add `prod_swap : A * B → B * A` and some basic theory about it.
- Add lemma `join_app`.
- Allow patterns and type annotations in `propset` notation, e.g.,
`{[ (x, y) : nat * nat | x = y ]}`. (by Thibaut Pérami)
The following `sed` script should perform most of the renaming
(on macOS, replace `sed` by `gsed`, installed via e.g. `brew install gnu-sed`).
......
......@@ -6,8 +6,13 @@ Record propset (A : Type) : Type := PropSet { propset_car : A → Prop }.
Add Printing Constructor propset.
Global Arguments PropSet {_} _ : assert.
Global Arguments propset_car {_} _ _ : assert.
(** Here we are using the notation "at level 200 as pattern" because we want to
be compatible with all the rules that start with [ {[ TERM ] such as
records, singletons, and map singletons. See
https://coq.inria.fr/refman/user-extensions/syntax-extensions.html#binders-bound-in-the-notation-and-parsed-as-terms
and https://gitlab.mpi-sws.org/iris/stdpp/-/merge_requests/533#note_98003 *)
Notation "{[ x | P ]}" := (PropSet (λ x, P))
(at level 1, format "{[ x | P ]}") : stdpp_scope.
(at level 1, x at level 200 as pattern, format "{[ x | P ]}") : stdpp_scope.
Global Instance propset_elem_of {A} : ElemOf A (propset A) := λ x X, propset_car X x.
......
From stdpp Require Import propset.
Lemma diagonal : {[ (a,b) : nat * nat | a = b ]} (λ x, (x,x)) <$> .
Proof. set_unfold. intros []. naive_solver. Qed.
Lemma diagonal2 : {[ (a,b) | a =@{nat} b ]} {[ x | x.1 = x.2 ]}.
Proof. set_unfold. intros []. naive_solver. Qed.
Lemma firstis42 : {[ (x, _) : nat * nat | x = 42 ]} (42,.) <$> .
Proof. set_unfold. intros []. naive_solver. Qed.
Inductive foo := Foo (n : nat).
Definition set_of_positive_foos : {[ Foo x | x 0 ]} Foo <$> (Pos.to_nat <$> ).
Proof.
set_unfold. intros [[]]; naive_solver by (rewrite SuccNat2Pos.id_succ || lia).
Qed.
Lemma simple_pattern_does_not_match : {[ x : nat | x = x]} = PropSet (λ x, x = x).
Proof. exact eq_refl. Qed.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Please register or to comment