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

A negligible change in the structure

This commit is contained in:
Niels
2017-09-07 15:19:48 +02:00
parent 4ab70ae1fe
commit 474c9324ca
29 changed files with 2082 additions and 1459 deletions

View File

@@ -0,0 +1,162 @@
(* 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}.
Context `{Univalence}.
Definition dupl' (a : A) (X : FSet A) :
{|a|} {|a|} X = {|a|} X := assoc _ _ _ @ ap ( X) (idem _).
Definition comm' (a b : A) (X : FSet A) :
{|a|} {|b|} X = {|b|} {|a|} X :=
assoc _ _ _ @ ap ( X) (comm _ _) @ (assoc _ _ _)^.
Definition FSetC_to_FSet: FSetC A -> FSet A.
Proof.
hrecursion.
- apply E.
- intros a x.
apply ({|a|} x).
- intros. cbn.
apply dupl'.
- intros. cbn.
apply comm'.
Defined.
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.
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 _)).
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.
- intros x y p q.
refine (append_union _ _ @ _).
refine (ap ( _) p @ _).
apply (ap (_ ) q).
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.
Global Instance: IsEquiv FSet_to_FSetC.
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.
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.
Theorem fset_fsetc : FSet A = FSetC A.
Proof.
apply (equiv_path _ _)^-1.
exact repr_iso.
Defined.
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.
- intros a Y PY.
refine (_ @ (Pdupl _ _ _)).
etransitivity.
{ apply (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)). }
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 )
(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.
Proof.
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). *)
End Iso.

View File

@@ -0,0 +1,84 @@
(** Definition of Finite Sets as via lists. *)
Require Import HoTT HitTactics.
Require Export set_names.
Module Export FSetC.
Section FSetC.
Private Inductive FSetC (A : Type) : Type :=
| Nil : FSetC A
| Cns : A -> FSetC A -> FSetC A.
Global Instance fset_empty : forall A,hasEmpty (FSetC A) := Nil.
Variable A : Type.
Arguments Cns {_} _ _.
Infix ";;" := Cns (at level 8, right associativity).
Axiom dupl : forall (a : A) (x : FSetC A),
a ;; a ;; x = a ;; x.
Axiom comm : forall (a b : A) (x : FSetC A),
a ;; b ;; x = b ;; a ;; x.
Axiom trunc : IsHSet (FSetC A).
End FSetC.
Arguments Cns {_} _ _.
Arguments dupl {_} _ _.
Arguments comm {_} _ _ _.
Infix ";;" := Cns (at level 8, right associativity).
Section FSetC_induction.
Variable (A : Type)
(P : FSetC A -> Type)
(H : forall x : FSetC A, IsHSet (P x))
(eP : P )
(cnsP : forall (a:A) (x: FSetC A), P x -> P (a ;; x))
(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) =
cnsP b (a;;x) (cnsP a x px)).
(* Induction principle *)
Fixpoint FSetC_ind
(x : FSetC A)
{struct x}
: P x
:= (match x return _ -> _ -> _ -> P x with
| Nil => fun _ => fun _ => fun _ => eP
| a ;; y => fun _ => fun _ => fun _ => cnsP a y (FSetC_ind y)
end) H duplP commP.
End FSetC_induction.
Section FSetC_recursion.
Variable (A : Type)
(P : Type)
(H : IsHSet P)
(nil : P)
(cns : A -> P -> P)
(duplP : forall (a: A) (x: P), cns a (cns a x) = (cns a x))
(commP : forall (a b: A) (x: P), cns a (cns b x) = cns b (cns a x)).
(* Recursion principle *)
Definition FSetC_rec : FSetC A -> P.
Proof.
simple refine (FSetC_ind _ _ _ _ _ _ _ );
try (intros; simple refine ((transport_const _ _) @ _ )); cbn.
- apply nil.
- apply (fun a => fun _ => cns a).
- apply duplP.
- apply commP.
Defined.
End FSetC_recursion.
Instance FSetC_recursion A : HitRecursion (FSetC A) :=
{
indTy := _; recTy := _;
H_inductor := FSetC_ind A; H_recursor := FSetC_rec A
}.
End FSetC.
Infix ";;" := Cns (at level 8, right associativity).

View File

@@ -0,0 +1,18 @@
(** Operations on [FSetC A] *)
Require Import HoTT HitTactics.
Require Import list_representation.
Section operations.
Global Instance fsetc_union : forall A, hasUnion (FSetC A).
Proof.
intros A x y.
hinduction x.
- apply y.
- apply Cns.
- apply dupl.
- apply comm.
Defined.
Global Instance fsetc_singleton : forall A, hasSingleton (FSetC A) A := fun A a => a;;.
End operations.

View File

@@ -0,0 +1,60 @@
(** Properties of the operations on [FSetC A] *)
Require Import HoTT HitTactics.
Require Import list_representation list_representation.operations.
Section properties.
Context {A : Type}.
Definition append_nl : forall (x: FSetC A), x = x
:= fun _ => idpath.
Lemma append_nr : forall (x: FSetC A), x = x.
Proof.
hinduction; try (intros; apply set_path2).
- reflexivity.
- intros. apply (ap (fun y => a;;y) X).
Defined.
Lemma append_assoc {H: Funext}:
forall (x y z: FSetC A), x (y z) = (x y) z.
Proof.
hinduction
; try (intros ; apply path_forall ; intro
; apply path_forall ; intro ; apply set_path2).
- reflexivity.
- intros a x HR y z.
specialize (HR y z).
apply (ap (fun y => a;;y) HR).
Defined.
Lemma append_singleton: forall (a: A) (x: FSetC A),
a ;; x = x (a ;; ).
Proof.
intro a. hinduction; try (intros; apply set_path2).
- reflexivity.
- intros b x HR. refine (_ @ _).
+ apply comm.
+ apply (ap (fun y => b ;; y) HR).
Defined.
Lemma append_comm {H: Funext}:
forall (x1 x2: FSetC A), x1 x2 = x2 x1.
Proof.
hinduction ; try (intros ; apply path_forall ; intro ; apply set_path2).
- intros.
apply (append_nr _)^.
- intros a x1 HR x2.
refine (ap (fun y => a;;y) (HR x2) @ _).
refine (append_singleton _ _ @ _).
refine ((append_assoc _ _ _)^ @ _).
refine (ap (x2 ) (append_singleton _ _)^).
Defined.
Lemma singleton_idem: forall (a: A),
{|a|} {|a|} = {|a|}.
Proof.
intro.
apply dupl.
Defined.
End properties.