Skip to content
Snippets Groups Projects
Commit c2f64d9d authored by Thibaut Pérami's avatar Thibaut Pérami
Browse files

Allow patterns and type annotations in propset notation

parent 32edabf7
No related branches found
No related tags found
No related merge requests found
Pipeline #91426 passed
...@@ -6,8 +6,12 @@ Record propset (A : Type) : Type := PropSet { propset_car : A → Prop }. ...@@ -6,8 +6,12 @@ Record propset (A : Type) : Type := PropSet { propset_car : A → Prop }.
Add Printing Constructor propset. Add Printing Constructor propset.
Global Arguments PropSet {_} _ : assert. Global Arguments PropSet {_} _ : assert.
Global Arguments propset_car {_} _ _ : 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 *)
Notation "{[ x | P ]}" := (PropSet (λ x, P)) 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. 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 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.
Finish editing this message first!
Please register or to comment