2017-08-02 11:40:03 +02:00
|
|
|
|
(* The representations [FSet A] and [FSetC A] are isomorphic for every A *)
|
|
|
|
|
Require Import HoTT HitTactics.
|
2017-09-07 15:19:48 +02:00
|
|
|
|
Require Import list_representation list_representation.operations
|
|
|
|
|
list_representation.properties.
|
|
|
|
|
Require Import kuratowski.kuratowski_sets.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
|
|
|
|
|
Section Iso.
|
|
|
|
|
Context {A : Type}.
|
|
|
|
|
Context `{Univalence}.
|
|
|
|
|
|
2017-08-19 18:55:47 +02:00
|
|
|
|
Definition dupl' (a : A) (X : FSet A) :
|
|
|
|
|
{|a|} ∪ {|a|} ∪ X = {|a|} ∪ X := assoc _ _ _ @ ap (∪ X) (idem _).
|
2017-09-07 15:19:48 +02:00
|
|
|
|
|
2017-08-19 18:55:47 +02:00
|
|
|
|
Definition comm' (a b : A) (X : FSet A) :
|
|
|
|
|
{|a|} ∪ {|b|} ∪ X = {|b|} ∪ {|a|} ∪ X :=
|
|
|
|
|
assoc _ _ _ @ ap (∪ X) (comm _ _) @ (assoc _ _ _)^.
|
|
|
|
|
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Definition FSetC_to_FSet: FSetC A -> FSet A.
|
|
|
|
|
Proof.
|
|
|
|
|
hrecursion.
|
|
|
|
|
- apply E.
|
2017-08-08 15:29:50 +02:00
|
|
|
|
- intros a x.
|
|
|
|
|
apply ({|a|} ∪ x).
|
2017-08-09 18:05:58 +02:00
|
|
|
|
- intros. cbn.
|
2017-08-19 18:55:47 +02:00
|
|
|
|
apply dupl'.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
- intros. cbn.
|
2017-08-19 18:55:47 +02:00
|
|
|
|
apply comm'.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-08 15:29:50 +02:00
|
|
|
|
Definition FSet_to_FSetC: FSet A -> FSetC A.
|
|
|
|
|
Proof.
|
|
|
|
|
hrecursion.
|
|
|
|
|
- apply ∅.
|
|
|
|
|
- intro a. apply {|a|}.
|
|
|
|
|
- intros X Y. apply (X ∪ Y).
|
|
|
|
|
- apply append_assoc.
|
|
|
|
|
- apply append_comm.
|
|
|
|
|
- apply append_nl.
|
|
|
|
|
- apply append_nr.
|
|
|
|
|
- apply singleton_idem.
|
|
|
|
|
Defined.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
|
2017-08-09 18:05:58 +02:00
|
|
|
|
Lemma append_union: forall (x y: FSetC A),
|
2017-08-08 15:29:50 +02:00
|
|
|
|
FSetC_to_FSet (x ∪ y) = (FSetC_to_FSet x) ∪ (FSetC_to_FSet y).
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Proof.
|
2017-08-09 18:05:58 +02:00
|
|
|
|
intros x.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
|
|
|
|
|
- intros. symmetry. apply nl.
|
2017-08-19 18:55:47 +02:00
|
|
|
|
- intros a x HR y. unfold union, fsetc_union in *.
|
|
|
|
|
refine (_ @ assoc _ _ _).
|
|
|
|
|
apply (ap ({|a|} ∪) (HR _)).
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Lemma repr_iso_id_l: forall (x: FSet A), FSetC_to_FSet (FSet_to_FSetC x) = x.
|
|
|
|
|
Proof.
|
|
|
|
|
hinduction; try (intros; apply set_path2).
|
|
|
|
|
- reflexivity.
|
|
|
|
|
- intro. apply nr.
|
2017-08-19 18:55:47 +02:00
|
|
|
|
- intros x y p q.
|
|
|
|
|
refine (append_union _ _ @ _).
|
|
|
|
|
refine (ap (∪ _) p @ _).
|
|
|
|
|
apply (ap (_ ∪) q).
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Lemma repr_iso_id_r: forall (x: FSetC A), FSet_to_FSetC (FSetC_to_FSet x) = x.
|
|
|
|
|
Proof.
|
|
|
|
|
hinduction; try (intros; apply set_path2).
|
|
|
|
|
- reflexivity.
|
|
|
|
|
- intros a x HR. rewrite HR. reflexivity.
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-03 23:21:43 +02:00
|
|
|
|
Global Instance: IsEquiv FSet_to_FSetC.
|
2017-08-02 11:40:03 +02:00
|
|
|
|
Proof.
|
|
|
|
|
apply isequiv_biinv.
|
|
|
|
|
unfold BiInv. split.
|
|
|
|
|
exists FSetC_to_FSet.
|
|
|
|
|
unfold Sect. apply repr_iso_id_l.
|
|
|
|
|
exists FSetC_to_FSet.
|
|
|
|
|
unfold Sect. apply repr_iso_id_r.
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-03 23:21:43 +02:00
|
|
|
|
Global Instance: IsEquiv FSetC_to_FSet.
|
|
|
|
|
Proof.
|
|
|
|
|
change (IsEquiv (FSet_to_FSetC)^-1).
|
|
|
|
|
apply isequiv_inverse.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Theorem repr_iso: FSet A <~> FSetC A.
|
|
|
|
|
Proof.
|
|
|
|
|
simple refine (@BuildEquiv (FSet A) (FSetC A) FSet_to_FSetC _ ).
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-03 14:43:42 +02:00
|
|
|
|
Theorem fset_fsetc : FSet A = FSetC A.
|
|
|
|
|
Proof.
|
|
|
|
|
apply (equiv_path _ _)^-1.
|
|
|
|
|
exact repr_iso.
|
|
|
|
|
Defined.
|
2017-08-19 18:55:47 +02:00
|
|
|
|
|
|
|
|
|
Theorem FSet_cons_ind (P : FSet A -> Type)
|
|
|
|
|
(Pset : forall (X : FSet A), IsHSet (P X))
|
|
|
|
|
(Pempt : P ∅)
|
|
|
|
|
(Pcons : forall (a : A) (X : FSet A), P X -> P ({|a|} ∪ X))
|
|
|
|
|
(Pdupl : forall (a : A) (X : FSet A) (px : P X),
|
|
|
|
|
transport P (dupl' a X) (Pcons a _ (Pcons a X px)) = Pcons a X px)
|
|
|
|
|
(Pcomm : forall (a b : A) (X : FSet A) (px : P X),
|
|
|
|
|
transport P (comm' a b X) (Pcons a _ (Pcons b X px)) = Pcons b _ (Pcons a X px)) :
|
|
|
|
|
forall X, P X.
|
|
|
|
|
Proof.
|
|
|
|
|
intros X.
|
|
|
|
|
refine (transport P (repr_iso_id_l X) _).
|
|
|
|
|
simple refine (FSetC_ind A (fun Z => P (FSetC_to_FSet Z)) _ _ _ _ _ (FSet_to_FSetC X)); simpl.
|
|
|
|
|
- apply Pempt.
|
|
|
|
|
- intros a Y HY. by apply Pcons.
|
2017-09-07 15:19:48 +02:00
|
|
|
|
- intros a Y PY.
|
2017-08-19 18:55:47 +02:00
|
|
|
|
refine (_ @ (Pdupl _ _ _)).
|
|
|
|
|
etransitivity.
|
|
|
|
|
{ apply (transport_compose _ FSetC_to_FSet (dupl a Y)). }
|
|
|
|
|
refine (ap (fun z => transport P z _) _).
|
2017-09-07 15:19:48 +02:00
|
|
|
|
apply path_ishprop.
|
2017-08-19 18:55:47 +02:00
|
|
|
|
- intros a b Y PY. cbn.
|
|
|
|
|
refine (_ @ (Pcomm _ _ _ _)).
|
|
|
|
|
etransitivity.
|
|
|
|
|
{ apply (transport_compose _ FSetC_to_FSet (FSetC.comm a b Y)). }
|
|
|
|
|
refine (ap (fun z => transport P z _) _).
|
2017-09-07 15:19:48 +02:00
|
|
|
|
apply path_ishprop.
|
2017-08-19 18:55:47 +02:00
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Theorem FSet_cons_ind_beta_empty (P : FSet A -> Type)
|
|
|
|
|
(Pset : forall (X : FSet A), IsHSet (P X))
|
|
|
|
|
(Pempt : P ∅)
|
|
|
|
|
(Pcons : forall (a : A) (X : FSet A), P X -> P ({|a|} ∪ X))
|
|
|
|
|
(Pdupl : forall (a : A) (X : FSet A) (px : P X),
|
|
|
|
|
transport P (dupl' a X) (Pcons a _ (Pcons a X px)) = Pcons a X px)
|
|
|
|
|
(Pcomm : forall (a b : A) (X : FSet A) (px : P X),
|
|
|
|
|
transport P (comm' a b X) (Pcons a _ (Pcons b X px)) = Pcons b _ (Pcons a X px)) :
|
|
|
|
|
FSet_cons_ind P Pset Pempt Pcons Pdupl Pcomm ∅ = Pempt.
|
2017-09-07 15:19:48 +02:00
|
|
|
|
Proof.
|
|
|
|
|
by compute.
|
|
|
|
|
Defined.
|
2017-08-19 18:55:47 +02:00
|
|
|
|
|
|
|
|
|
(* TODO *)
|
|
|
|
|
(* Theorem FSet_cons_ind_beta_cons (P : FSet A -> Type) *)
|
|
|
|
|
(* (Pset : forall (X : FSet A), IsHSet (P X)) *)
|
|
|
|
|
(* (Pempt : P ∅) *)
|
|
|
|
|
(* (Pcons : forall (a : A) (X : FSet A), P X -> P ({|a|} ∪ X)) *)
|
|
|
|
|
(* (Pdupl : forall (a : A) (X : FSet A) (px : P X), *)
|
|
|
|
|
(* transport P (dupl' a X) (Pcons a _ (Pcons a X px)) = Pcons a X px) *)
|
|
|
|
|
(* (Pcomm : forall (a b : A) (X : FSet A) (px : P X), *)
|
|
|
|
|
(* transport P (comm' a b X) (Pcons a _ (Pcons b X px)) = Pcons b _ (Pcons a X px)) : *)
|
|
|
|
|
(* forall a X, FSet_cons_ind P Pset Pempt Pcons Pdupl Pcomm ({|a|} ∪ X) = Pcons a X (FSet_cons_ind P Pset Pempt Pcons Pdupl Pcomm X). *)
|
|
|
|
|
(* Proof. *)
|
|
|
|
|
|
|
|
|
|
(* Theorem FSet_cons_ind_beta_dupl (P : FSet A -> Type) *)
|
|
|
|
|
(* (Pset : forall (X : FSet A), IsHSet (P X)) *)
|
|
|
|
|
(* (Pempt : P ∅) *)
|
|
|
|
|
(* (Pcons : forall (a : A) (X : FSet A), P X -> P ({|a|} ∪ X)) *)
|
|
|
|
|
(* (Pdupl : forall (a : A) (X : FSet A) (px : P X), *)
|
|
|
|
|
(* transport P (dupl' a X) (Pcons a _ (Pcons a X px)) = Pcons a X px) *)
|
|
|
|
|
(* (Pcomm : forall (a b : A) (X : FSet A) (px : P X), *)
|
|
|
|
|
(* transport P (comm' a b X) (Pcons a _ (Pcons b X px)) = Pcons b _ (Pcons a X px)) : *)
|
|
|
|
|
(* forall a X, apD (FSet_cons_ind P Pset Pempt Pcons Pdupl Pcomm) (dupl' a X) = Pdupl a X (FSet_cons_ind P Pset Pempt Pcons Pdupl Pcomm X). *)
|
2017-08-03 14:43:42 +02:00
|
|
|
|
End Iso.
|