HITs-Examples/FiniteSets/list_representation/properties.v

61 lines
1.5 KiB
Coq
Raw Permalink Normal View History

(** Properties of the operations on [FSetC A]. These are needed to prove that the
representations are isomorphic. *)
2017-08-02 11:40:03 +02:00
Require Import HoTT HitTactics.
2017-09-07 15:19:48 +02:00
Require Import list_representation list_representation.operations.
2017-08-02 11:40:03 +02:00
Section properties.
Context {A : Type}.
2017-09-19 17:22:15 +02:00
Definition append_nl (x: FSetC A) : x = x
:= idpath.
2017-08-02 11:40:03 +02:00
2017-08-08 15:29:50 +02:00
Lemma append_nr : forall (x: FSetC A), x = x.
2017-08-02 11:40:03 +02:00
Proof.
hinduction; try (intros; apply path_ishprop).
2017-09-07 15:19:48 +02:00
- reflexivity.
- intros. apply (ap (fun y => a;;y) X).
2017-08-02 11:40:03 +02:00
Defined.
2017-09-19 17:22:15 +02:00
Lemma append_assoc :
2017-08-08 15:29:50 +02:00
forall (x y z: FSetC A), x (y z) = (x y) z.
2017-08-02 11:40:03 +02:00
Proof.
2017-09-19 17:22:15 +02:00
intros x y z.
hinduction x ; try (intros ; apply path_ishprop).
- reflexivity.
2017-09-19 17:22:15 +02:00
- intros.
cbn.
f_ap.
2017-08-02 11:40:03 +02:00
Defined.
Lemma append_singleton: forall (a: A) (x: FSetC A),
2017-08-08 15:29:50 +02:00
a ;; x = x (a ;; ).
2017-08-02 11:40:03 +02:00
Proof.
intro a. hinduction; try (intros; apply path_ishprop).
2017-08-02 11:40:03 +02:00
- reflexivity.
- intros b x HR.
refine (comm_s _ _ _ @ ap (fun y => b ;; y) HR).
2017-08-02 11:40:03 +02:00
Defined.
Lemma append_comm :
2017-08-08 15:29:50 +02:00
forall (x1 x2: FSetC A), x1 x2 = x2 x1.
2017-08-02 11:40:03 +02:00
Proof.
intros x1 x2.
hinduction x1 ; try (intros ; apply path_ishprop).
2017-09-07 15:19:48 +02:00
- intros.
apply (append_nr _)^.
- intros a x HR.
refine (ap (fun y => a;;y) HR @ _).
2017-09-07 15:19:48 +02:00
refine (append_singleton _ _ @ _).
refine ((append_assoc _ _ _)^ @ _).
refine (ap (x2 ) (append_singleton _ _)^).
2017-08-02 11:40:03 +02:00
Defined.
Lemma singleton_idem: forall (a: A),
2017-08-08 15:29:50 +02:00
{|a|} {|a|} = {|a|}.
2017-08-02 11:40:03 +02:00
Proof.
2017-08-08 15:29:50 +02:00
intro.
apply dupl.
2017-08-02 11:40:03 +02:00
Defined.
End properties.