1
0
mirror of https://github.com/nmvdw/HITs-Examples synced 2025-11-03 23:23:51 +01:00

Uses merely decidable equality, added length.

This commit is contained in:
Niels van der Weide
2017-09-21 14:12:51 +02:00
parent 0def5869cd
commit 39e2ce1c05
15 changed files with 193 additions and 106 deletions

View File

@@ -1,16 +1,16 @@
(* The representations [FSet A] and [FSetC A] are isomorphic for every A *)
(** The representations [FSet A] and [FSetC A] are isomorphic for every A *)
Require Import HoTT HitTactics.
Require Import list_representation list_representation.operations
list_representation.properties.
Require Import kuratowski.kuratowski_sets.
Section Iso.
Context {A : Type} `{Univalence}.
Context {A : Type}.
Definition FSetC_to_FSet: FSetC A -> FSet A.
Proof.
hrecursion.
- apply E.
- apply .
- intros a x.
apply ({|a|} x).
- intros a X.
@@ -23,8 +23,8 @@ Section Iso.
Proof.
hrecursion.
- apply .
- intro a. apply {|a|}.
- intros X Y. apply (X Y).
- apply (fun a => {|a|}).
- apply ().
- apply append_assoc.
- apply append_comm.
- apply append_nl.
@@ -35,19 +35,20 @@ Section Iso.
Lemma append_union: forall (x y: FSetC A),
FSetC_to_FSet (x y) = (FSetC_to_FSet x) (FSetC_to_FSet y).
Proof.
intros x.
hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
- intros. symmetry. apply nl.
- intros a x HR y. unfold union, fsetc_union in *.
refine (_ @ assoc _ _ _).
apply (ap ({|a|} ) (HR _)).
intros x y.
hrecursion x ; try (intros ; apply path_ishprop).
- intros.
apply (nl _)^.
- intros a x HR.
refine (ap ({|a|} ) HR @ assoc _ _ _).
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).
hinduction ; try (intros ; apply path_ishprop).
- reflexivity.
- intro. apply nr.
- intro.
apply nr.
- intros x y p q.
refine (append_union _ _ @ _).
refine (ap ( _) p @ _).
@@ -56,15 +57,16 @@ Section Iso.
Lemma repr_iso_id_r: forall (x: FSetC A), FSet_to_FSetC (FSetC_to_FSet x) = x.
Proof.
hinduction; try (intros; apply set_path2).
hinduction ; try (intros ; apply path_ishprop).
- reflexivity.
- intros a x HR. rewrite HR. reflexivity.
- intros a x HR.
refine (ap ({|a|} ) HR).
Defined.
Global Instance: IsEquiv FSet_to_FSetC.
Proof.
apply isequiv_biinv.
unfold BiInv. split.
split.
exists FSetC_to_FSet.
unfold Sect. apply repr_iso_id_l.
exists FSetC_to_FSet.
@@ -82,12 +84,18 @@ Section Iso.
simple refine (@BuildEquiv (FSet A) (FSetC A) FSet_to_FSetC _ ).
Defined.
Theorem fset_fsetc : FSet A = FSetC A.
Theorem fset_fsetc `{Univalence} : FSet A = FSetC A.
Proof.
apply (equiv_path _ _)^-1.
exact repr_iso.
Defined.
Definition dupl' (a : A) (X : FSet A) : {|a|} {|a|} X = {|a|} X
:= assoc _ _ _ @ ap ( X) (idem a).
Definition comm' (a b : A) (X : FSet A) : {|a|} {|b|} X = {|b|} {|a|} X
:= assoc _ _ _ @ ap ( X) (comm _ _) @ (assoc _ _ _)^.
Theorem FSet_cons_ind (P : FSet A -> Type)
(Pset : forall (X : FSet A), IsHSet (P X))
(Pempt : P )
@@ -95,28 +103,28 @@ Section Iso.
(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.
transport P (comm' a b X) (Pcons a _ (Pcons b X px)) = Pcons b _ (Pcons a X px))
(X : FSet A)
: 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.
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.
- intros a Y HY.
apply (Pcons a _ HY).
- intros a Y PY.
refine (_ @ (Pdupl _ _ _)).
etransitivity.
{ apply (transport_compose _ FSetC_to_FSet (dupl a Y)). }
refine (transport_compose _ FSetC_to_FSet (dupl a Y) _ @ _).
refine (ap (fun z => transport P z _) _).
apply path_ishprop.
- intros a b Y PY. cbn.
refine (_ @ (Pcomm _ _ _ _)).
etransitivity.
{ apply (transport_compose _ FSetC_to_FSet (FSetC.comm a b Y)). }
- intros a b Y PY.
refine (transport_compose _ FSetC_to_FSet (comm_s a b Y) _ @ _ @ (Pcomm _ _ _ _)).
refine (ap (fun z => transport P z _) _).
apply path_ishprop.
Defined.
(*
Theorem FSet_cons_ind_beta_empty (P : FSet A -> Type)
(Pset : forall (X : FSet A), IsHSet (P X))
(Pempt : P ∅)
@@ -130,25 +138,31 @@ Section Iso.
by compute.
Defined.
(* 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). *)
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.
intros.
unfold FSet_cons_ind.
simpl.
rewrite ?transport_pp.
hinduction X ; try(intros ; apply path_ishprop) ; simpl.
- admit.
- intro b.
unfold FSet_cons_ind.
simpl.
admit.
- intros.
unfold FSet_cons_ind.
simpl.
rewrite X.
*)
End Iso.

View File

@@ -18,7 +18,7 @@ Module Export FSetC.
Axiom dupl : forall (a : A) (x : FSetC A),
a ;; a ;; x = a ;; x.
Axiom comm : forall (a b : A) (x : FSetC A),
Axiom comm_s : forall (a b : A) (x : FSetC A),
a ;; b ;; x = b ;; a ;; x.
Axiom trunc : IsHSet (FSetC A).
@@ -26,7 +26,7 @@ Module Export FSetC.
Arguments Cns {_} _ _.
Arguments dupl {_} _ _.
Arguments comm {_} _ _ _.
Arguments comm_s {_} _ _ _.
Infix ";;" := Cns (at level 8, right associativity).
@@ -39,7 +39,7 @@ Module Export FSetC.
(duplP : forall (a: A) (x: FSetC A) (px : P x),
dupl a x # cnsP a (a;;x) (cnsP a x px) = cnsP a x px)
(commP : forall (a b: A) (x: FSetC A) (px: P x),
comm a b x # cnsP a (b;;x) (cnsP b x px) =
comm_s a b x # cnsP a (b;;x) (cnsP b x px) =
cnsP b (a;;x) (cnsP a x px)).
(* Induction principle *)

View File

@@ -10,7 +10,7 @@ Section operations.
- apply y.
- apply Cns.
- apply dupl.
- apply comm.
- apply comm_s.
Defined.
Global Instance fsetc_singleton : forall A, hasSingleton (FSetC A) A := fun A a => a;;.

View File

@@ -1,4 +1,5 @@
(** Properties of the operations on [FSetC A] *)
(** Properties of the operations on [FSetC A]. These are needed to prove that the
representations are isomorphic. *)
Require Import HoTT HitTactics.
Require Import list_representation list_representation.operations.
@@ -10,7 +11,7 @@ Section properties.
Lemma append_nr : forall (x: FSetC A), x = x.
Proof.
hinduction; try (intros; apply set_path2).
hinduction; try (intros; apply path_ishprop).
- reflexivity.
- intros. apply (ap (fun y => a;;y) X).
Defined.
@@ -20,8 +21,7 @@ Section properties.
Proof.
intros x y z.
hinduction x ; try (intros ; apply path_ishprop).
- cbn.
reflexivity.
- reflexivity.
- intros.
cbn.
f_ap.
@@ -30,21 +30,21 @@ Section properties.
Lemma append_singleton: forall (a: A) (x: FSetC A),
a ;; x = x (a ;; ).
Proof.
intro a. hinduction; try (intros; apply set_path2).
intro a. hinduction; try (intros; apply path_ishprop).
- reflexivity.
- intros b x HR. refine (_ @ _).
+ apply comm.
+ apply (ap (fun y => b ;; y) HR).
- intros b x HR.
refine (comm_s _ _ _ @ ap (fun y => b ;; y) HR).
Defined.
Lemma append_comm {H: Funext}:
Lemma append_comm :
forall (x1 x2: FSetC A), x1 x2 = x2 x1.
Proof.
hinduction ; try (intros ; apply path_forall ; intro ; apply set_path2).
intros x1 x2.
hinduction x1 ; try (intros ; apply path_ishprop).
- intros.
apply (append_nr _)^.
- intros a x1 HR x2.
refine (ap (fun y => a;;y) (HR x2) @ _).
- intros a x HR.
refine (ap (fun y => a;;y) HR @ _).
refine (append_singleton _ _ @ _).
refine ((append_assoc _ _ _)^ @ _).
refine (ap (x2 ) (append_singleton _ _)^).