HITs-Examples/FiniteSets/implementations/lists.v

106 lines
2.6 KiB
Coq
Raw Permalink Normal View History

2017-08-02 11:40:03 +02:00
(* Implementation of [FSet A] using lists *)
2017-06-20 17:33:35 +02:00
Require Import HoTT HitTactics.
2017-08-07 12:20:43 +02:00
Require Import FSets implementations.interface.
2017-06-20 17:33:35 +02:00
Section Operations.
2017-08-07 12:20:43 +02:00
Context `{Univalence}.
2017-06-20 17:33:35 +02:00
Global Instance list_empty A : hasEmpty (list A) := nil.
2017-06-20 17:33:35 +02:00
Global Instance list_single A: hasSingleton (list A) A := fun a => cons a nil.
Global Instance list_union A : hasUnion (list A).
2017-08-07 12:20:43 +02:00
Proof.
intros l1 l2.
2017-08-07 12:20:43 +02:00
induction l1.
* apply l2.
* apply (cons a IHl1).
Defined.
2017-06-20 17:33:35 +02:00
Global Instance list_membership A : hasMembership (list A) A.
2017-08-07 12:20:43 +02:00
Proof.
intros a l.
induction l as [ | b l IHl].
- apply False_hp.
- apply (hor (a = b) IHl).
Defined.
2017-06-20 17:33:35 +02:00
Global Instance list_comprehension A: hasComprehension (list A) A.
2017-08-07 12:20:43 +02:00
Proof.
intros ϕ l.
2017-08-07 12:20:43 +02:00
induction l as [ | b l IHl].
- apply nil.
- apply (if ϕ b then cons b IHl else IHl).
Defined.
2017-06-20 17:33:35 +02:00
2017-08-07 12:20:43 +02:00
Fixpoint list_to_set A (l : list A) : FSet A :=
2017-06-20 17:33:35 +02:00
match l with
| nil =>
| cons a l => {|a|} (list_to_set A l)
2017-06-20 17:33:35 +02:00
end.
End Operations.
Section ListToSet.
Variable A : Type.
2017-08-07 12:20:43 +02:00
Context `{Univalence}.
2017-06-20 17:33:35 +02:00
2017-08-07 12:20:43 +02:00
Lemma member_isIn (a : A) (l : list A) :
2017-08-07 16:49:46 +02:00
member a l = a (list_to_set A l).
2017-06-20 17:33:35 +02:00
Proof.
2017-08-07 12:20:43 +02:00
induction l ; unfold member in * ; simpl in *.
2017-06-20 17:33:35 +02:00
- reflexivity.
2017-08-07 12:20:43 +02:00
- rewrite IHl.
unfold hor, merely, lor.
apply path_iff_hprop ; intros z ; strip_truncations ; destruct z as [z1 | z2].
* apply (tr (inl (tr z1))).
* apply (tr (inr z2)).
* strip_truncations ; apply (tr (inl z1)).
* apply (tr (inr z2)).
2017-06-20 17:33:35 +02:00
Defined.
Definition empty_empty : list_to_set A = := idpath.
2017-06-20 17:33:35 +02:00
2017-08-07 12:20:43 +02:00
Lemma filter_comprehension (ϕ : A -> Bool) (l : list A) :
list_to_set A (filter ϕ l) = {| list_to_set A l & ϕ |}.
2017-06-20 17:33:35 +02:00
Proof.
induction l ; cbn in *.
- reflexivity.
- destruct (ϕ a) ; cbn in * ; unfold list_to_set in IHl.
2017-08-07 16:49:46 +02:00
* refine (ap (fun y => {|a|} y) _).
2017-06-20 17:33:35 +02:00
apply IHl.
* rewrite nl.
apply IHl.
Defined.
2017-08-07 16:49:46 +02:00
Definition singleton_single (a : A) : list_to_set A (singleton a) = {|a|} :=
nr {|a|}.
2017-06-20 17:33:35 +02:00
2017-08-07 12:20:43 +02:00
Lemma append_union (l1 l2 : list A) :
2017-08-07 16:49:46 +02:00
list_to_set A (union l1 l2) = (list_to_set A l1) (list_to_set A l2).
2017-06-20 17:33:35 +02:00
Proof.
2017-08-07 12:20:43 +02:00
induction l1 ; induction l2 ; cbn.
- apply (nl _)^.
2017-08-07 12:20:43 +02:00
- apply (nl _)^.
- rewrite IHl1.
apply assoc.
- rewrite IHl1.
cbn.
apply assoc.
2017-06-20 17:33:35 +02:00
Defined.
2017-08-07 12:20:43 +02:00
End ListToSet.
2017-06-20 17:33:35 +02:00
2017-08-07 12:20:43 +02:00
Section lists_are_sets.
Context `{Univalence}.
2017-06-20 17:33:35 +02:00
2017-08-07 12:20:43 +02:00
Instance lists_sets : sets list list_to_set.
Proof.
split ; intros.
- apply empty_empty.
- apply singleton_single.
- apply append_union.
- apply filter_comprehension.
- apply member_isIn.
Defined.
End lists_are_sets.