2017-06-20 11:33:13 +02:00
|
|
|
|
(* [FSet] is a (strong and stable) finite powerset monad *)
|
|
|
|
|
Require Import HoTT HitTactics.
|
2017-08-16 16:07:51 +02:00
|
|
|
|
Require Export representations.definition fsets.properties fsets.operations.
|
2017-06-20 11:33:13 +02:00
|
|
|
|
|
2017-06-21 11:07:22 +02:00
|
|
|
|
Definition ffmap {A B : Type} : (A -> B) -> FSet A -> FSet B.
|
2017-06-20 11:33:13 +02:00
|
|
|
|
Proof.
|
|
|
|
|
intro f.
|
|
|
|
|
hrecursion.
|
|
|
|
|
- exact ∅.
|
|
|
|
|
- intro a. exact {| f a |}.
|
2017-08-08 15:29:50 +02:00
|
|
|
|
- intros X Y. apply (X ∪ Y).
|
2017-06-20 11:33:13 +02:00
|
|
|
|
- apply assoc.
|
|
|
|
|
- apply comm.
|
|
|
|
|
- apply nl.
|
|
|
|
|
- apply nr.
|
|
|
|
|
- simpl. intro x. apply idem.
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-06-21 11:07:22 +02:00
|
|
|
|
Lemma ffmap_1 `{Funext} {A : Type} : @ffmap A A idmap = idmap.
|
2017-06-20 11:33:13 +02:00
|
|
|
|
Proof.
|
|
|
|
|
apply path_forall.
|
2017-08-01 15:12:59 +02:00
|
|
|
|
intro x. hinduction x; try (intros; f_ap);
|
2017-06-20 11:33:13 +02:00
|
|
|
|
try (intros; apply set_path2).
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-06-21 11:07:22 +02:00
|
|
|
|
Global Instance fset_functorish `{Funext}: Functorish FSet
|
|
|
|
|
:= { fmap := @ffmap; fmap_idmap := @ffmap_1 _ }.
|
|
|
|
|
|
|
|
|
|
Lemma ffmap_compose {A B C : Type} `{Funext} (f : A -> B) (g : B -> C) :
|
|
|
|
|
fmap FSet (g o f) = fmap _ g o fmap _ f.
|
2017-06-20 11:33:13 +02:00
|
|
|
|
Proof.
|
|
|
|
|
apply path_forall. intro x.
|
2017-08-01 15:12:59 +02:00
|
|
|
|
hrecursion x; try (intros; f_ap);
|
2017-06-20 11:33:13 +02:00
|
|
|
|
try (intros; apply set_path2).
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Definition join {A : Type} : FSet (FSet A) -> FSet A.
|
|
|
|
|
Proof.
|
|
|
|
|
hrecursion.
|
|
|
|
|
- exact ∅.
|
|
|
|
|
- exact idmap.
|
2017-08-08 15:29:50 +02:00
|
|
|
|
- intros X Y. apply (X ∪ Y).
|
2017-06-20 11:33:13 +02:00
|
|
|
|
- apply assoc.
|
|
|
|
|
- apply comm.
|
|
|
|
|
- apply nl.
|
|
|
|
|
- apply nr.
|
|
|
|
|
- simpl. apply union_idem.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Lemma join_assoc {A : Type} (X : FSet (FSet (FSet A))) :
|
2017-06-21 11:07:22 +02:00
|
|
|
|
join (ffmap join X) = join (join X).
|
2017-06-20 11:33:13 +02:00
|
|
|
|
Proof.
|
2017-08-01 15:12:59 +02:00
|
|
|
|
hrecursion X; try (intros; f_ap);
|
2017-06-20 11:33:13 +02:00
|
|
|
|
try (intros; apply set_path2).
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Lemma join_return_1 {A : Type} (X : FSet A) :
|
|
|
|
|
join ({| X |}) = X.
|
|
|
|
|
Proof. reflexivity. Defined.
|
|
|
|
|
|
|
|
|
|
Lemma join_return_fmap {A : Type} (X : FSet A) :
|
2017-06-21 11:07:22 +02:00
|
|
|
|
join ({| X |}) = join (ffmap (fun x => {|x|}) X).
|
2017-06-20 11:33:13 +02:00
|
|
|
|
Proof.
|
2017-08-01 15:12:59 +02:00
|
|
|
|
hrecursion X; try (intros; f_ap);
|
2017-06-20 11:33:13 +02:00
|
|
|
|
try (intros; apply set_path2).
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Lemma join_fmap_return_1 {A : Type} (X : FSet A) :
|
2017-06-21 11:07:22 +02:00
|
|
|
|
join (ffmap (fun x => {|x|}) X) = X.
|
2017-06-20 11:33:13 +02:00
|
|
|
|
Proof. refine ((join_return_fmap _)^ @ join_return_1 _). Defined.
|
2017-08-16 16:07:51 +02:00
|
|
|
|
|
|
|
|
|
Lemma fmap_isIn `{Univalence} {A B : Type} (f : A -> B) (a : A) (X : FSet A) :
|
|
|
|
|
a ∈ X -> (f a) ∈ (ffmap f X).
|
|
|
|
|
Proof.
|
|
|
|
|
hinduction X; try (intros; apply path_ishprop).
|
|
|
|
|
- apply idmap.
|
|
|
|
|
- intros b Hab; strip_truncations.
|
|
|
|
|
apply (tr (ap f Hab)).
|
|
|
|
|
- intros X0 X1 HX0 HX1 Ha.
|
|
|
|
|
strip_truncations. apply tr.
|
|
|
|
|
destruct Ha as [Ha | Ha].
|
|
|
|
|
+ left. by apply HX0.
|
|
|
|
|
+ right. by apply HX1.
|
|
|
|
|
Defined.
|