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

Generalize `gset_to_propset` to `set_to_propset` for any SemiSet.

This closes issue #25.
parent ff302761
No related branches found
No related tags found
No related merge requests found
......@@ -237,11 +237,6 @@ Section gset.
Global Instance gset_dom {A} : Dom (gmap K A) (gset K) := mapset_dom.
Global Instance gset_dom_spec : FinMapDom K (gmap K) (gset K) := mapset_dom_spec.
Definition gset_to_propset (X : gset K) : propset K :=
{[ x | x X ]}.
Lemma elem_of_gset_to_propset (X : gset K) x : x gset_to_propset X x X.
Proof. done. Qed.
Definition gset_to_gmap {A} (x : A) (X : gset K) : gmap K A :=
(λ _, x) <$> mapset_car X.
......
......@@ -34,6 +34,12 @@ Lemma top_subseteq {A} (X : propset A) : X ⊆ ⊤.
Proof. done. Qed.
Hint Resolve top_subseteq : core.
Definition set_to_propset `{ElemOf A C} (X : C) : propset A :=
{[ x | x X ]}.
Lemma elem_of_set_to_propset `{SemiSet A C} x (X : C) :
x set_to_propset X x X.
Proof. done. Qed.
Instance propset_ret : MRet propset := λ A (x : A), {[ x ]}.
Instance propset_bind : MBind propset := λ A B (f : A propset B) (X : propset A),
PropSet (λ b, a, b f a a X).
......
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