2017-05-24 13:54:00 +02:00
|
|
|
Require Import HoTT HitTactics.
|
2017-05-23 16:30:31 +02:00
|
|
|
Require Import definition.
|
|
|
|
|
|
|
|
Section operations.
|
|
|
|
|
2017-05-24 13:54:00 +02:00
|
|
|
Context {A : Type}.
|
|
|
|
Context {A_deceq : DecidablePaths A}.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
|
|
|
Definition isIn : A -> FSet A -> Bool.
|
|
|
|
Proof.
|
|
|
|
intros a.
|
2017-05-24 13:54:00 +02:00
|
|
|
hrecursion.
|
2017-05-23 16:30:31 +02:00
|
|
|
- exact false.
|
2017-05-24 13:54:00 +02:00
|
|
|
- intro a'. destruct (dec (a = a')); [exact true | exact false].
|
2017-05-23 16:30:31 +02:00
|
|
|
- apply orb.
|
2017-05-24 13:54:00 +02:00
|
|
|
- intros x y z. compute. destruct x; reflexivity.
|
|
|
|
- intros x y. compute. destruct x, y; reflexivity.
|
|
|
|
- intros x. compute. reflexivity.
|
|
|
|
- intros x. compute. destruct x; reflexivity.
|
|
|
|
- intros a'. compute. destruct (A_deceq a a'); reflexivity.
|
2017-05-23 16:30:31 +02:00
|
|
|
Defined.
|
|
|
|
|
|
|
|
Infix "∈" := isIn (at level 9, right associativity).
|
|
|
|
|
|
|
|
Definition comprehension :
|
|
|
|
(A -> Bool) -> FSet A -> FSet A.
|
|
|
|
Proof.
|
|
|
|
intros P.
|
2017-05-24 13:54:00 +02:00
|
|
|
hrecursion.
|
2017-05-23 16:30:31 +02:00
|
|
|
- apply E.
|
|
|
|
- intro a.
|
|
|
|
refine (if (P a) then L a else E).
|
|
|
|
- apply U.
|
|
|
|
- intros. cbv. apply assoc.
|
|
|
|
- intros. cbv. apply comm.
|
|
|
|
- intros. cbv. apply nl.
|
|
|
|
- intros. cbv. apply nr.
|
|
|
|
- intros. cbv.
|
|
|
|
destruct (P x); simpl.
|
|
|
|
+ apply idem.
|
|
|
|
+ apply nl.
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
Definition intersection :
|
|
|
|
FSet A -> FSet A -> FSet A.
|
|
|
|
Proof.
|
|
|
|
intros X Y.
|
|
|
|
apply (comprehension (fun (a : A) => isIn a X) Y).
|
|
|
|
Defined.
|
|
|
|
|
2017-05-24 13:54:00 +02:00
|
|
|
End operations.
|