Skip to content
Snippets Groups Projects
Verified Commit feb655ce authored by Paolo G. Giarrusso's avatar Paolo G. Giarrusso
Browse files

set_bind theory: revise setoid rewriting

- strengthen set_bind_subset to have the same arity
- add missing Params instance to prevent slow failures in setoid rewriting
- `Proper` instances can be proven before `set_bind_ext` via `set_solver`.
- `set_bind_ext` is then provable by set_solver directly.
parent bcc9161a
No related branches found
No related tags found
1 merge request!406Introduce `set_bind` and associated lemmas + set_bind theory: revise setoid rewriting
...@@ -31,6 +31,7 @@ Definition set_bind `{Elements A SA, Empty SB, Union SB} ...@@ -31,6 +31,7 @@ Definition set_bind `{Elements A SA, Empty SB, Union SB}
(f : A SB) (X : SA) : SB := (f : A SB) (X : SA) : SB :=
(f <$> elements X). (f <$> elements X).
Typeclasses Opaque set_bind. Typeclasses Opaque set_bind.
Global Instance: Params (@set_bind) 6 := {}.
Global Instance set_fresh `{Elements A C, Fresh A (list A)} : Fresh A C := Global Instance set_fresh `{Elements A C, Fresh A (list A)} : Fresh A C :=
fresh elements. fresh elements.
...@@ -463,17 +464,15 @@ Section set_bind. ...@@ -463,17 +464,15 @@ Section set_bind.
rewrite elem_of_set_bind. set_solver. rewrite elem_of_set_bind. set_solver.
Qed. Qed.
Lemma set_bind_ext (f g : A SB) (X Y : C) :
( x, x X x Y f x g x) X Y set_bind f X set_bind g Y.
Proof.
intros Hfg HXY b. rewrite !elem_of_set_bind. set_solver.
Qed.
Global Instance set_bind_proper : Proper (pointwise_relation _ () ==> () ==> ()) set_bind. Global Instance set_bind_proper : Proper (pointwise_relation _ () ==> () ==> ()) set_bind.
Proof. intros f1 f2 Hf X1 X2 HX. by apply set_bind_ext. Qed. Proof. unfold pointwise_relation; intros f1 f2 Hf X1 X2 HX. set_solver. Qed.
Global Instance set_bind_subset f : Proper (() ==> ()) (set_bind f). Global Instance set_bind_subset : Proper (pointwise_relation _ () ==> () ==> ()) set_bind.
Proof. intros X Y HXY. set_solver. Qed. Proof. unfold pointwise_relation; intros f1 f2 Hf X1 X2 HX. set_solver. Qed.
Lemma set_bind_ext (f g : A SB) (X Y : C) :
( x, x X x Y f x g x) X Y set_bind f X set_bind g Y.
Proof. set_solver. Qed.
Lemma set_bind_singleton f x : set_bind f {[x]} f x. Lemma set_bind_singleton f x : set_bind f {[x]} f x.
Proof. set_solver. Qed. Proof. set_solver. 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