mirror of https://github.com/nmvdw/HITs-Examples
Cleaning in iso
This commit is contained in:
parent
2cd3beec43
commit
8e143c5285
|
@ -5,15 +5,7 @@ Require Import list_representation list_representation.operations
|
|||
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 _ _ _)^.
|
||||
Context {A : Type} `{Univalence}.
|
||||
|
||||
Definition FSetC_to_FSet: FSetC A -> FSet A.
|
||||
Proof.
|
||||
|
@ -21,10 +13,10 @@ Section Iso.
|
|||
- apply E.
|
||||
- intros a x.
|
||||
apply ({|a|} ∪ x).
|
||||
- intros. cbn.
|
||||
apply dupl'.
|
||||
- intros. cbn.
|
||||
apply comm'.
|
||||
- intros a X.
|
||||
apply (assoc _ _ _ @ ap (∪ X) (idem _)).
|
||||
- intros a X Y.
|
||||
apply (assoc _ _ _ @ ap (∪ Y) (comm _ _) @ (assoc _ _ _)^).
|
||||
Defined.
|
||||
|
||||
Definition FSet_to_FSetC: FSet A -> FSetC A.
|
||||
|
|
|
@ -5,8 +5,8 @@ Require Import list_representation list_representation.operations.
|
|||
Section properties.
|
||||
Context {A : Type}.
|
||||
|
||||
Definition append_nl : forall (x: FSetC A), ∅ ∪ x = x
|
||||
:= fun _ => idpath.
|
||||
Definition append_nl (x: FSetC A) : ∅ ∪ x = x
|
||||
:= idpath.
|
||||
|
||||
Lemma append_nr : forall (x: FSetC A), x ∪ ∅ = x.
|
||||
Proof.
|
||||
|
@ -15,16 +15,16 @@ Section properties.
|
|||
- intros. apply (ap (fun y => a;;y) X).
|
||||
Defined.
|
||||
|
||||
Lemma append_assoc {H: Funext}:
|
||||
Lemma append_assoc :
|
||||
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).
|
||||
intros x y z.
|
||||
hinduction x ; try (intros ; apply path_ishprop).
|
||||
- cbn.
|
||||
reflexivity.
|
||||
- intros.
|
||||
cbn.
|
||||
f_ap.
|
||||
Defined.
|
||||
|
||||
Lemma append_singleton: forall (a: A) (x: FSetC A),
|
||||
|
|
Loading…
Reference in New Issue