2017-08-01 15:12:59 +02:00
|
|
|
|
(* Operations on the [FSet A] for an arbitrary [A] *)
|
2017-05-24 13:54:00 +02:00
|
|
|
|
Require Import HoTT HitTactics.
|
2017-08-01 15:41:53 +02:00
|
|
|
|
Require Import representations.definition disjunction lattice.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
|
2017-08-01 15:12:59 +02:00
|
|
|
|
Section operations.
|
2017-08-07 16:49:46 +02:00
|
|
|
|
Context `{Univalence}.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Global Instance fset_member : forall A, hasMembership (FSet A) A.
|
2017-08-07 16:49:46 +02:00
|
|
|
|
Proof.
|
2017-08-08 17:00:30 +02:00
|
|
|
|
intros A a.
|
2017-08-07 16:49:46 +02:00
|
|
|
|
hrecursion.
|
|
|
|
|
- exists Empty.
|
|
|
|
|
exact _.
|
|
|
|
|
- intro a'.
|
|
|
|
|
exists (Trunc (-1) (a = a')).
|
|
|
|
|
exact _.
|
|
|
|
|
- apply lor.
|
|
|
|
|
- intros ; symmetry ; apply lor_assoc.
|
|
|
|
|
- apply lor_commutative.
|
|
|
|
|
- apply lor_nl.
|
|
|
|
|
- apply lor_nr.
|
|
|
|
|
- intros ; apply lor_idem.
|
|
|
|
|
Defined.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Global Instance fset_comprehension : forall A, hasComprehension (FSet A) A.
|
2017-08-07 16:49:46 +02:00
|
|
|
|
Proof.
|
2017-08-08 17:00:30 +02:00
|
|
|
|
intros A P.
|
2017-08-07 16:49:46 +02:00
|
|
|
|
hrecursion.
|
|
|
|
|
- apply ∅.
|
|
|
|
|
- intro a.
|
|
|
|
|
refine (if (P a) then {|a|} else ∅).
|
2017-08-08 15:29:50 +02:00
|
|
|
|
- apply (∪).
|
2017-08-07 16:49:46 +02:00
|
|
|
|
- apply assoc.
|
|
|
|
|
- apply comm.
|
|
|
|
|
- apply nl.
|
|
|
|
|
- apply nr.
|
2017-08-09 18:05:58 +02:00
|
|
|
|
- intros; simpl.
|
2017-08-07 16:49:46 +02:00
|
|
|
|
destruct (P x).
|
|
|
|
|
+ apply idem.
|
|
|
|
|
+ apply nl.
|
|
|
|
|
Defined.
|
2017-05-23 16:30:31 +02:00
|
|
|
|
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Definition isEmpty (A : Type) :
|
2017-08-07 16:49:46 +02:00
|
|
|
|
FSet A -> Bool.
|
|
|
|
|
Proof.
|
|
|
|
|
simple refine (FSet_rec _ _ _ true (fun _ => false) andb _ _ _ _ _)
|
|
|
|
|
; try eauto with bool_lattice_hints typeclass_instances.
|
2017-08-07 22:13:42 +02:00
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Definition single_product {A B : Type} (a : A) : FSet B -> FSet (A * B).
|
2017-08-07 23:15:25 +02:00
|
|
|
|
Proof.
|
|
|
|
|
hrecursion.
|
|
|
|
|
- apply ∅.
|
|
|
|
|
- intro b.
|
|
|
|
|
apply {|(a, b)|}.
|
2017-08-08 15:29:50 +02:00
|
|
|
|
- apply (∪).
|
2017-08-07 23:15:25 +02:00
|
|
|
|
- intros X Y Z ; apply assoc.
|
|
|
|
|
- intros X Y ; apply comm.
|
|
|
|
|
- intros ; apply nl.
|
|
|
|
|
- intros ; apply nr.
|
|
|
|
|
- intros ; apply idem.
|
|
|
|
|
Defined.
|
2017-08-09 18:05:58 +02:00
|
|
|
|
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Definition product {A B : Type} : FSet A -> FSet B -> FSet (A * B).
|
2017-08-07 22:13:42 +02:00
|
|
|
|
Proof.
|
|
|
|
|
intros X Y.
|
|
|
|
|
hrecursion X.
|
|
|
|
|
- apply ∅.
|
|
|
|
|
- intro a.
|
2017-08-07 23:15:25 +02:00
|
|
|
|
apply (single_product a Y).
|
2017-08-08 15:29:50 +02:00
|
|
|
|
- apply (∪).
|
2017-08-07 22:13:42 +02:00
|
|
|
|
- intros ; apply assoc.
|
|
|
|
|
- intros ; apply comm.
|
|
|
|
|
- intros ; apply nl.
|
|
|
|
|
- intros ; apply nr.
|
2017-08-08 13:35:28 +02:00
|
|
|
|
- intros ; apply union_idem.
|
2017-08-07 22:13:42 +02:00
|
|
|
|
Defined.
|
2017-08-09 18:05:58 +02:00
|
|
|
|
|
2017-08-08 17:00:30 +02:00
|
|
|
|
Global Instance fset_subset : forall A, hasSubset (FSet A).
|
2017-08-08 15:29:50 +02:00
|
|
|
|
Proof.
|
2017-08-08 17:00:30 +02:00
|
|
|
|
intros A X Y.
|
2017-08-08 15:29:50 +02:00
|
|
|
|
hrecursion X.
|
|
|
|
|
- exists Unit.
|
|
|
|
|
exact _.
|
|
|
|
|
- intros a.
|
|
|
|
|
apply (a ∈ Y).
|
|
|
|
|
- intros X1 X2.
|
|
|
|
|
exists (prod X1 X2).
|
|
|
|
|
exact _.
|
|
|
|
|
- intros.
|
|
|
|
|
apply path_trunctype ; apply equiv_prod_assoc.
|
|
|
|
|
- intros.
|
|
|
|
|
apply path_trunctype ; apply equiv_prod_symm.
|
|
|
|
|
- intros.
|
|
|
|
|
apply path_trunctype ; apply prod_unit_l.
|
|
|
|
|
- intros.
|
|
|
|
|
apply path_trunctype ; apply prod_unit_r.
|
|
|
|
|
- intros a'.
|
|
|
|
|
apply path_iff_hprop ; cbn.
|
|
|
|
|
* intros [p1 p2]. apply p1.
|
|
|
|
|
* intros p.
|
|
|
|
|
split ; apply p.
|
|
|
|
|
Defined.
|
2017-08-09 17:03:51 +02:00
|
|
|
|
|
|
|
|
|
Local Ltac remove_transport
|
|
|
|
|
:= intros ; apply path_forall ; intro Z ; rewrite transport_arrow
|
|
|
|
|
; rewrite transport_const ; simpl.
|
|
|
|
|
Local Ltac pointwise
|
|
|
|
|
:= repeat (f_ap) ; try (apply path_forall ; intro Z2) ;
|
|
|
|
|
rewrite transport_arrow, transport_const, transport_sigma
|
|
|
|
|
; f_ap ; hott_simpl ; simple refine (path_sigma _ _ _ _ _)
|
|
|
|
|
; try (apply transport_const) ; try (apply path_ishprop).
|
|
|
|
|
|
|
|
|
|
Lemma separation (A B : Type) : forall (X : FSet A) (f : {a | a ∈ X} -> B), FSet B.
|
|
|
|
|
Proof.
|
|
|
|
|
hinduction.
|
|
|
|
|
- intros f.
|
|
|
|
|
apply ∅.
|
|
|
|
|
- intros a f.
|
|
|
|
|
apply {|f (a; tr idpath)|}.
|
|
|
|
|
- intros X1 X2 HX1 HX2 f.
|
|
|
|
|
pose (fX1 := fun Z : {a : A & a ∈ X1} => f (pr1 Z;tr (inl (pr2 Z)))).
|
|
|
|
|
pose (fX2 := fun Z : {a : A & a ∈ X2} => f (pr1 Z;tr (inr (pr2 Z)))).
|
|
|
|
|
specialize (HX1 fX1).
|
|
|
|
|
specialize (HX2 fX2).
|
|
|
|
|
apply (HX1 ∪ HX2).
|
|
|
|
|
- remove_transport.
|
|
|
|
|
rewrite assoc.
|
|
|
|
|
pointwise.
|
|
|
|
|
- remove_transport.
|
|
|
|
|
rewrite comm.
|
|
|
|
|
pointwise.
|
|
|
|
|
- remove_transport.
|
|
|
|
|
rewrite nl.
|
|
|
|
|
pointwise.
|
|
|
|
|
- remove_transport.
|
|
|
|
|
rewrite nr.
|
|
|
|
|
pointwise.
|
|
|
|
|
- remove_transport.
|
|
|
|
|
rewrite <- (idem (Z (x; tr 1%path))).
|
|
|
|
|
pointwise.
|
2017-08-09 18:05:58 +02:00
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
End operations.
|