HITs-Examples/FiniteSets/fsets/properties_cons_repr.v

65 lines
1.7 KiB
Coq
Raw Normal View History

2017-08-02 11:40:03 +02:00
(* Properties of the operations on [FSetC A] *)
Require Import HoTT HitTactics.
Require Import representations.cons_repr.
From fsets Require Import operations_cons_repr.
Section properties.
Context {A : Type}.
2017-08-08 15:29:50 +02:00
Definition append_nl : forall (x: FSetC A), x = x
:= fun _ => 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 set_path2).
- reflexivity.
2017-08-08 15:29:50 +02:00
- intros. apply (ap (fun y => a;;y) X).
2017-08-02 11:40:03 +02:00
Defined.
Lemma append_assoc {H: Funext}:
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-08-08 15:29:50 +02:00
hinduction
; try (intros ; apply path_forall ; intro
; apply path_forall ; intro ; apply set_path2).
2017-08-02 11:40:03 +02:00
- reflexivity.
- intros a x HR y z.
2017-08-02 11:40:03 +02:00
specialize (HR y z).
apply (ap (fun y => a;;y) HR).
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 set_path2).
- reflexivity.
- intros b x HR. refine (_ @ _).
+ apply comm.
+ apply (ap (fun y => b ;; y) HR).
Defined.
Lemma append_comm {H: Funext}:
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.
2017-08-08 15:29:50 +02:00
hinduction ; try (intros ; apply path_forall ; intro ; apply set_path2).
- intros. symmetry. apply append_nr.
2017-08-02 11:40:03 +02:00
- intros a x1 HR x2.
etransitivity.
apply (ap (fun y => a;;y) (HR x2)).
2017-08-08 15:29:50 +02:00
transitivity ((x2 x1) (a;;)).
+ apply append_singleton.
2017-08-02 11:40:03 +02:00
+ etransitivity.
* symmetry. apply append_assoc.
2017-08-08 15:29:50 +02:00
* simple 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
unfold singleton.
intro.
apply dupl.
2017-08-02 11:40:03 +02:00
Defined.
End properties.