Library UniMath.Combinatorics.Lists
Require Import UniMath.MoreFoundations.Tactics.
Require Import UniMath.Combinatorics.StandardFiniteSets.
Section preamble.
Definition iterprod (n : nat) (A : UU) : UU.
Show proof.
induction n as [|n IHn].
- apply unit.
- apply (A × IHn).
- apply unit.
- apply (A × IHn).
End preamble.
Section lists.
Context {A : UU}.
Context {A : UU}.
The type of lists
Definition list : UU := ∑ n, iterprod n A.
The empty list
Definition nil : list := (0,,tt).
List cons
Definition cons (x : A) (xs : list) : list :=
(S (pr1 xs),, (x,, pr2 xs)).
Local Notation "[]" := nil (at level 0, format "[]").
Local Infix "::" := cons.
Lemma list_ind : ∏ (P : list -> UU),
P nil
-> (∏ (x : A) (xs : list), P xs -> P (x :: xs))
-> ∏ xs, P xs.
Show proof.
Lemma list_ind_compute_2
(P : list -> UU)
(p0 : P nil)
(ind : ∏ (x : A) (xs : list), P xs -> P (x :: xs))
(x : A) (xs : list)
(f := list_ind P p0 ind) :
f (x::xs) = ind x xs (f xs).
Show proof.
Definition foldr {B : UU} (f : A -> B -> B) (b : B) : list -> B :=
list_ind (λ _, B) b (λ a _ b', f a b').
Definition length : list -> nat := pr1.
(S (pr1 xs),, (x,, pr2 xs)).
Local Notation "[]" := nil (at level 0, format "[]").
Local Infix "::" := cons.
Lemma list_ind : ∏ (P : list -> UU),
P nil
-> (∏ (x : A) (xs : list), P xs -> P (x :: xs))
-> ∏ xs, P xs.
Show proof.
intros P Hnil Hcons xs.
induction xs as [n xs].
induction n as [|n IHn].
- induction xs.
apply Hnil.
- simpl in xs.
induction xs as [x xs].
apply (Hcons x (n,,xs) (IHn xs)).
induction xs as [n xs].
induction n as [|n IHn].
- induction xs.
apply Hnil.
- simpl in xs.
induction xs as [x xs].
apply (Hcons x (n,,xs) (IHn xs)).
Lemma list_ind_compute_2
(P : list -> UU)
(p0 : P nil)
(ind : ∏ (x : A) (xs : list), P xs -> P (x :: xs))
(x : A) (xs : list)
(f := list_ind P p0 ind) :
f (x::xs) = ind x xs (f xs).
Show proof.
reflexivity.
Definition foldr {B : UU} (f : A -> B -> B) (b : B) : list -> B :=
list_ind (λ _, B) b (λ a _ b', f a b').
Definition length : list -> nat := pr1.
Variation of foldr that returns a for the empty list and folds the
rest with the first element as new default value
Definition foldr1 (f : A -> A -> A) (a : A) : list → A.
Show proof.
Show proof.
apply list_ind.
- exact a.
- intros a' l fl. revert l. apply list_ind.
+ exact a'.
+ intros _ _ _. exact (f a' fl).
- exact a.
- intros a' l fl. revert l. apply list_ind.
+ exact a'.
+ intros _ _ _. exact (f a' fl).
The n-th element of a list
Fixpoint nth'' n i : i < n -> iterprod n A -> A
:= match n, i with
| 0, _ => λ r x, fromempty (nopathsfalsetotrue r)
| S n, 0 => λ r x, pr1 x
| S n, S i => λ r x, nth'' n i r (pr2 x)
end.
Definition nth' n : iterprod n A -> stn n -> A.
Show proof.
intros x i.
exact (nth'' n (pr1 i) (pr2 i) x).
exact (nth'' n (pr1 i) (pr2 i) x).
Lemma nth'_step n (x:iterprod (S n) A) i (I:i<n) :
nth' (S n) x (stnpair (S n) (S i) I) = nth' n (pr2 x) (stnpair n i I).
Show proof.
reflexivity.
Definition nth x : stn(length x) -> A.
Show proof.
intros i. exact (nth' (length x) (pr2 x) i).
Definition functionToList' n : (stn n -> A) -> iterprod n A.
Show proof.
intros f.
induction n as [|n I].
- exact tt.
- exists (f (●0))%stn.
exact (I(f ∘ dni (●0)))%stn.
induction n as [|n I].
- exact tt.
- exists (f (●0))%stn.
exact (I(f ∘ dni (●0)))%stn.
Definition functionToList n : (stn n -> A) -> list.
Show proof.
intros f.
exact (n ,, functionToList' n f).
exact (n ,, functionToList' n f).
Section Test.
Local Open Scope stn.
Context {a b c d:A}.
Let x := a::b::c::d::[].
Goal nth x (●0) = a. reflexivity. Qed.
Goal nth x (●1) = b. reflexivity. Qed.
Goal nth x (●2) = c. reflexivity. Qed.
Goal nth x (●3) = d. reflexivity. Qed.
Goal functionToList _ (nth x) = x. reflexivity. Qed.
End Test.
End lists.
Make the type not implicit for list
Arguments list : clear implicits.
Section more_lists.
Definition map {A B : UU} (f : A -> B) : list A -> list B :=
foldr (λ a l, cons (f a) l) nil.
Lemma mapStep {A B : UU} (f : A -> B) (a:A) (x:list A) : map f (cons a x) = cons (f a) (map f x).
Show proof.
Section more_lists.
Definition map {A B : UU} (f : A -> B) : list A -> list B :=
foldr (λ a l, cons (f a) l) nil.
Lemma mapStep {A B : UU} (f : A -> B) (a:A) (x:list A) : map f (cons a x) = cons (f a) (map f x).
Show proof.
reflexivity.
Various unfolding lemmas
Lemma foldr_cons {A B : UU} (f : A -> B -> B) (b : B) (x : A) (xs : list A) :
foldr f b (cons x xs) = f x (foldr f b xs).
Show proof.
Lemma map_nil {A B : UU} (f : A -> B) : map f nil = nil.
Show proof.
Lemma map_cons {A B : UU} (f : A -> B) (x : A) (xs : list A) :
map f (cons x xs) = cons (f x) (map f xs).
Show proof.
Lemma map_compose {A B C : UU} (f : A → B) (g : B → C) (xs : list A) :
map (g ∘ f) xs = map g (map f xs).
Show proof.
Lemma map_idfun {A : UU} (xs : list A) :
map (idfun A) xs = xs.
Show proof.
Lemma map_homot {A B : UU} {f g : A → B} (h : f ~ g) (xs : list A) :
map f xs = map g xs.
Show proof.
Lemma foldr1_cons_nil {A : UU} (f : A -> A -> A) (a : A) (x : A) :
foldr1 f a (cons x nil) = x.
Show proof.
Lemma foldr1_cons {A : UU} (f : A -> A -> A) (a : A) (x y : A) (xs : list A) :
foldr1 f a (cons x (cons y xs)) = f x (foldr1 f a (cons y xs)).
Show proof.
End more_lists.
Local Notation "[]" := nil (at level 0, format "[]").
Local Infix "::" := cons.
foldr f b (cons x xs) = f x (foldr f b xs).
Show proof.
reflexivity.
Lemma map_nil {A B : UU} (f : A -> B) : map f nil = nil.
Show proof.
apply idpath.
Lemma map_cons {A B : UU} (f : A -> B) (x : A) (xs : list A) :
map f (cons x xs) = cons (f x) (map f xs).
Show proof.
apply idpath.
Lemma map_compose {A B C : UU} (f : A → B) (g : B → C) (xs : list A) :
map (g ∘ f) xs = map g (map f xs).
Show proof.
revert xs. apply list_ind.
- reflexivity.
- intros x xs IH. now rewrite !map_cons, IH.
- reflexivity.
- intros x xs IH. now rewrite !map_cons, IH.
Lemma map_idfun {A : UU} (xs : list A) :
map (idfun A) xs = xs.
Show proof.
revert xs. apply list_ind.
- reflexivity.
- intros x xs IH. now rewrite !map_cons, IH.
- reflexivity.
- intros x xs IH. now rewrite !map_cons, IH.
Lemma map_homot {A B : UU} {f g : A → B} (h : f ~ g) (xs : list A) :
map f xs = map g xs.
Show proof.
revert xs. apply list_ind.
- reflexivity.
- intros x xs IH. now rewrite !map_cons, h, IH.
- reflexivity.
- intros x xs IH. now rewrite !map_cons, h, IH.
Lemma foldr1_cons_nil {A : UU} (f : A -> A -> A) (a : A) (x : A) :
foldr1 f a (cons x nil) = x.
Show proof.
apply idpath.
Lemma foldr1_cons {A : UU} (f : A -> A -> A) (a : A) (x y : A) (xs : list A) :
foldr1 f a (cons x (cons y xs)) = f x (foldr1 f a (cons y xs)).
Show proof.
apply idpath.
End more_lists.
Local Notation "[]" := nil (at level 0, format "[]").
Local Infix "::" := cons.
concatenate two lists
Definition concatenate {X} : list X -> list X -> list X
:= λ r s, foldr cons s r.
Local Infix "++" := concatenate.
Lemma concatenateStep {X} (x:X) (r s:list X) :
(x::r) ++ s = x :: (r ++ s).
Show proof.
reflexivity.
Lemma nil_concatenate {X} (r : list X) : nil ++ r = r.
Show proof.
reflexivity.
Lemma concatenate_nil {X} (r : list X) : r ++ nil = r.
Show proof.
revert r. apply list_ind. reflexivity. intros x xs p. exact (maponpaths (cons x) p).
Lemma assoc_concatenate {X} (r s t : list X) : (r ++ s) ++ t = r ++ (s ++ t).
Show proof.
revert r. apply list_ind.
- reflexivity.
- intros x xs p. now rewrite !concatenateStep, p.
- reflexivity.
- intros x xs p. now rewrite !concatenateStep, p.
Lemma map_concatenate {X Y} (f : X → Y) (r s : list X) : map f (r ++ s) = map f r ++ map f s.
Show proof.
revert r. apply list_ind.
- reflexivity.
- intros x xs p. now rewrite mapStep, !concatenateStep, mapStep, p.
- reflexivity.
- intros x xs p. now rewrite mapStep, !concatenateStep, mapStep, p.
Lemma foldr_concatenate {X Y : UU} (f : X → Y) (l : list X) :
foldr concatenate [] (map (λ x, f x::[]) l) = map f l.
Show proof.
revert l. apply list_ind.
- reflexivity.
- intros x l IH. now rewrite !map_cons, foldr_cons, IH.
- reflexivity.
- intros x l IH. now rewrite !map_cons, foldr_cons, IH.
Lemma foldr1_concatenate {X Y : UU} (f : X → Y) (l : list X) :
map f l = foldr1 concatenate [] (map (λ x, f x::[]) l).
Show proof.
revert l. apply list_ind.
- reflexivity.
- intros x. refine (list_ind _ _ _).
+ reflexivity.
+ intros x' l _ IH. exact (maponpaths (cons (f x)) IH).
- reflexivity.
- intros x. refine (list_ind _ _ _).
+ reflexivity.
+ intros x' l _ IH. exact (maponpaths (cons (f x)) IH).
Append a single element to a list
Definition append {X} (x : X) (l : list X) : list X :=
l ++ x::[].
Lemma appendStep {X} (x y : X) (l : list X) : append x (y::l) = y::append x l.
Show proof.
reflexivity.
Lemma append_concatenate {X} (x : X) (l s : list X) : append x (l ++ s) = l ++ append x s.
Show proof.
apply assoc_concatenate.
Lemma map_append {X Y} (f : X → Y) (x : X) (r : list X) : map f (append x r) = append (f x) (map f r).
Show proof.
exact (map_concatenate _ _ _).
Reverse a list
Definition reverse {X} : list X → list X :=
foldr append [].
Lemma reverse_nil (X : Type) : reverse (@nil X) = [].
Show proof.
reflexivity.
Lemma reverseStep {X} (x : X) (r : list X) : reverse (x::r) = append x (reverse r).
Show proof.
reflexivity.
Lemma map_reverse {X Y} (f : X → Y) (r : list X) : map f (reverse r) = reverse (map f r).
Show proof.
revert r. apply list_ind.
- reflexivity.
- intros x xs p. now rewrite mapStep, !reverseStep, map_append, p.
- reflexivity.
- intros x xs p. now rewrite mapStep, !reverseStep, map_append, p.
Lemma reverse_concatenate {X} (l s : list X) : reverse (l ++ s) = reverse s ++ reverse l.
Show proof.
revert l. apply list_ind.
- symmetry. apply concatenate_nil.
- intros x xs p. now rewrite concatenateStep, !reverseStep, p, append_concatenate.
- symmetry. apply concatenate_nil.
- intros x xs p. now rewrite concatenateStep, !reverseStep, p, append_concatenate.
Lemma reverse_append {X} (x : X) (l : list X) : reverse (append x l) = x :: reverse l.
Show proof.
unfold append. now rewrite reverse_concatenate, reverseStep, reverse_nil.
Lemma reverse_reverse {X} (r : list X) : reverse (reverse r) = r.
Show proof.
revert r. apply list_ind.
- reflexivity.
- intros x xs p. now rewrite !reverseStep, reverse_append, p.
- reflexivity.
- intros x xs p. now rewrite !reverseStep, reverse_append, p.
flatten lists of lists
Definition flatten {X} : list (list X) → list X.
Show proof.
apply list_ind.
+ exact [].
+ intros s _ f. exact (concatenate s f).
+ exact [].
+ intros s _ f. exact (concatenate s f).
Lemma flattenStep {X} (x:list X) (m : list(list X)) : flatten (x::m) = concatenate x (flatten m).
Show proof.
unfold flatten.
rewrite list_ind_compute_2.
reflexivity.
rewrite list_ind_compute_2.
reflexivity.
between ≃ lists and functions
Lemma isweqlistfun {A} n : isweq (@nth' A n).
Show proof.
simple refine (isweq_iso _ (functionToList' _) _ _).
- intros. induction n as [|n N].
+ apply isapropunit.
+ simpl in x. induction x as [a x]. apply dirprodeq.
* simpl. reflexivity.
* simpl. apply N.
- intros. apply funextfun; intro i.
induction n as [|n N].
+ apply fromempty. now apply negstn0.
+ induction i as [i I].
induction i as [|i IHi].
* unfold nth', functionToList'; simpl. apply maponpaths. now apply subtypeEquality_prop.
* change (hProptoType (i<n)) in I.
exact (nth'_step _ (functionToList' _ _) _ _ @ N _ _).
- intros. induction n as [|n N].
+ apply isapropunit.
+ simpl in x. induction x as [a x]. apply dirprodeq.
* simpl. reflexivity.
* simpl. apply N.
- intros. apply funextfun; intro i.
induction n as [|n N].
+ apply fromempty. now apply negstn0.
+ induction i as [i I].
induction i as [|i IHi].
* unfold nth', functionToList'; simpl. apply maponpaths. now apply subtypeEquality_prop.
* change (hProptoType (i<n)) in I.
exact (nth'_step _ (functionToList' _ _) _ _ @ N _ _).
Corollary weqlistfun {A} n : (iterprod n A) ≃ (stn n -> A).
Show proof.
exact (weqpair _ (isweqlistfun _)).
Lemma isofhleveliterprod (n : nat) (k : nat) {X : UU} (is1 : isofhlevel n X) : isofhlevel n (iterprod k X).
Show proof.
induction k as [|k IH].
- apply isofhlevelcontr, iscontrunit.
- apply isofhleveldirprod.
+ apply is1.
+ apply IH.
- apply isofhlevelcontr, iscontrunit.
- apply isofhleveldirprod.
+ apply is1.
+ apply IH.
Lemma isofhlevellist (n : nat) {X : UU} (is1 : isofhlevel (S (S n)) X) : isofhlevel (S (S n)) (list X).
Show proof.
use isofhleveltotal2.
- intros m k. apply isofhlevelsnprop, isasetnat.
- intro m. apply isofhleveliterprod, is1.
- intros m k. apply isofhlevelsnprop, isasetnat.
- intro m. apply isofhleveliterprod, is1.