1
0
mirror of https://github.com/nmvdw/HITs-Examples synced 2025-12-14 06:13:51 +01:00

Improved notatio

This commit is contained in:
Niels
2017-08-08 15:29:50 +02:00
parent de335c3955
commit 2bdec415d9
15 changed files with 227 additions and 284 deletions

View File

@@ -6,38 +6,30 @@ From fsets Require Import operations_cons_repr.
Section properties.
Context {A : Type}.
Lemma append_nl:
forall (x: FSetC A), append x = x.
Proof.
intro. reflexivity.
Defined.
Definition append_nl : forall (x: FSetC A), x = x
:= fun _ => idpath.
Lemma append_nr:
forall (x: FSetC A), append x = x.
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).
- intros. apply (ap (fun y => a;;y) X).
Defined.
Lemma append_assoc {H: Funext}:
forall (x y z: FSetC A), append x (append y z) = append (append x y) z.
forall (x y z: FSetC A), x (y z) = (x y) z.
Proof.
intro x; hinduction x; try (intros; apply set_path2).
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. apply path_forall.
intro. apply path_forall.
intro. apply set_path2.
- intros. apply path_forall.
intro. apply path_forall.
intro. apply set_path2.
apply (ap (fun y => a;;y) HR).
Defined.
Lemma append_singleton: forall (a: A) (x: FSetC A),
a ;; x = append x (a ;; ).
a ;; x = x (a ;; ).
Proof.
intro a. hinduction; try (intros; apply set_path2).
- reflexivity.
@@ -47,29 +39,26 @@ Section properties.
Defined.
Lemma append_comm {H: Funext}:
forall (x1 x2: FSetC A), append x1 x2 = append x2 x1.
forall (x1 x2: FSetC A), x1 x2 = x2 x1.
Proof.
intro x1.
hinduction x1; try (intros; apply set_path2).
hinduction ; try (intros ; apply path_forall ; intro ; apply set_path2).
- intros. symmetry. apply append_nr.
- intros a x1 HR x2.
etransitivity.
apply (ap (fun y => a;;y) (HR x2)).
transitivity (append (append x2 x1) (a;;)).
transitivity ((x2 x1) (a;;)).
+ apply append_singleton.
+ etransitivity.
* symmetry. apply append_assoc.
* simple refine (ap (fun y => append x2 y) (append_singleton _ _)^).
- intros. apply path_forall.
intro. apply set_path2.
- intros. apply path_forall.
intro. apply set_path2.
* simple refine (ap (x2 ) (append_singleton _ _)^).
Defined.
Lemma singleton_idem: forall (a: A),
append (singleton a) (singleton a) = singleton a.
{|a|} {|a|} = {|a|}.
Proof.
unfold singleton. intro. cbn. apply dupl.
unfold singleton.
intro.
apply dupl.
Defined.
End properties.