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

`TopSet` instance for `boolset`.

parent e09f7ce3
No related branches found
No related tags found
No related merge requests found
...@@ -19,15 +19,16 @@ Instance boolset_intersection {A} : Intersection (boolset A) := λ X1 X2, ...@@ -19,15 +19,16 @@ Instance boolset_intersection {A} : Intersection (boolset A) := λ X1 X2,
BoolSet (λ x, boolset_car X1 x && boolset_car X2 x). BoolSet (λ x, boolset_car X1 x && boolset_car X2 x).
Instance boolset_difference {A} : Difference (boolset A) := λ X1 X2, Instance boolset_difference {A} : Difference (boolset A) := λ X1 X2,
BoolSet (λ x, boolset_car X1 x && negb (boolset_car X2 x)). BoolSet (λ x, boolset_car X1 x && negb (boolset_car X2 x)).
Instance boolset_set `{EqDecision A} : Set_ A (boolset A). Instance boolset_top_set `{EqDecision A} : TopSet A (boolset A).
Proof. Proof.
split; [split| |]. split; [split; [split| |]|].
- by intros x ?. - by intros x ?.
- by intros x y; rewrite <-(bool_decide_spec (x = y)). - by intros x y; rewrite <-(bool_decide_spec (x = y)).
- split. apply orb_prop_elim. apply orb_prop_intro. - split. apply orb_prop_elim. apply orb_prop_intro.
- split. apply andb_prop_elim. apply andb_prop_intro. - split. apply andb_prop_elim. apply andb_prop_intro.
- intros X Y x; unfold elem_of, boolset_elem_of; simpl. - intros X Y x; unfold elem_of, boolset_elem_of; simpl.
destruct (boolset_car X x), (boolset_car Y x); simpl; tauto. destruct (boolset_car X x), (boolset_car Y x); simpl; tauto.
- done.
Qed. Qed.
Instance boolset_elem_of_dec {A} : RelDecision (∈@{boolset A}). Instance boolset_elem_of_dec {A} : RelDecision (∈@{boolset A}).
Proof. refine (λ x X, cast_if (decide (boolset_car X x))); done. Defined. Proof. refine (λ x X, cast_if (decide (boolset_car X x))); done. Defined.
......
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