Library prosa.util.list
Require Import prosa.util.ssromega prosa.util.tactics.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq fintype bigop.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq fintype bigop.
We define a few simple operations on
lists. Namely first, last, and max.
Additional lemmas about last0.
Let xs be a non-empty sequence and x be an arbitrary element,
then we prove that last0 (x::xs) = last0 xs.
Similarly, let xs_r be a non-empty sequence and xs_l be an sequence,
we prove that last0 (xs_l ++ xs_r) = last0 xs_r.
We prove that for any non-empty sequence xs there is a sequence xsh
such that xsh ++ [::last0 x] = [xs].
We prove that if x is the last element of a sequence xs and
x satisfies a predicate, then x remains the last element in
the filtered sequence.
Lemma last0_filter:
∀ x xs (P : nat → bool),
xs ≠ [::] →
last0 xs = x →
P x →
last0 [seq x <- xs | P x] = x.
End Last0.
∀ x xs (P : nat → bool),
xs ≠ [::] →
last0 xs = x →
P x →
last0 [seq x <- xs | P x] = x.
End Last0.
Additional lemmas about max0.
Next we prove that one can remove duplicating element from the
head of a sequence.
We prove that one can remove the first element of a sequence if
it is not greater than the second element of this sequence.
Note that the last element is at most the max element.
Let's introduce the notion of the nth element of a sequence.
If any element of a sequence xs is less-than-or-equal-to
the corresponding element of a sequence ys, then max of
xs is less-than-or-equal-to max of ys.
Lemma max_of_dominating_seq:
∀ (xs ys : seq nat),
(∀ n, xs[|n|] ≤ ys[|n|]) →
max0 xs ≤ max0 ys.
End Max0.
(* Additional lemmas about rem for lists. *)
Section RemList.
(* We prove that if x lies in xs excluding y, then x also lies in xs. *)
Lemma rem_in:
∀ (T: eqType) (x y: T) (xs: seq T),
x \in rem y xs → x \in xs.
(* We prove that if we remove an element x for which P x from
a filter, then the size of the filter decreases by 1. *)
Lemma filter_size_rem:
∀ (T: eqType) (x:T) (xs: seq T) (P: T → bool),
(x \in xs) →
P x →
size [seq y <- xs | P y] = size [seq y <- rem x xs | P y] + 1.
End RemList.
(* Additional lemmas about sequences. *)
Section AdditionalLemmas.
(* First, we show that if n > 0, then nth (x::xs) n = nth xs (n-1). *)
Lemma nth0_cons:
∀ x xs n,
n > 0 →
nth 0 (x :: xs) n = nth 0 xs n.-1.
(* We prove that a sequence xs of size n.+1 can be destructed
into a sequence xs_l of size n and an element x such that
x = xs ++ [::x]. *)
Lemma seq_elim_last:
∀ {X : Type} (n : nat) (xs : seq X),
size xs = n.+1 →
∃ x xs__c, xs = xs__c ++ [:: x] ∧ size xs__c = n.
(* Next, we prove that x ∈ xs implies that xs can be split
into two parts such that xs = xsl ++ [::x] ++ [xsr]. *)
Lemma in_cat:
∀ {X : eqType} (x : X) (xs : list X),
x \in xs → ∃ xsl xsr, xs = xsl ++ [::x] ++ xsr.
(* We prove that for any two sequences xs and ys the fact that xs is a sub-sequence
of ys implies that the size of xs is at most the size of ys. *)
Lemma subseq_leq_size:
∀ {X : eqType} (xs ys: seq X),
uniq xs →
(∀ x, x \in xs → x \in ys) →
size xs ≤ size ys.
(* We prove that if no element of a sequence xs satisfies
a predicate P, then filter P xs is equal to an empty
sequence. *)
Lemma filter_in_pred0:
∀ {X : eqType} (xs : seq X) P,
(∀ x, x \in xs → ~~ P x) →
filter P xs = [::].
End AdditionalLemmas.
∀ (xs ys : seq nat),
(∀ n, xs[|n|] ≤ ys[|n|]) →
max0 xs ≤ max0 ys.
End Max0.
(* Additional lemmas about rem for lists. *)
Section RemList.
(* We prove that if x lies in xs excluding y, then x also lies in xs. *)
Lemma rem_in:
∀ (T: eqType) (x y: T) (xs: seq T),
x \in rem y xs → x \in xs.
(* We prove that if we remove an element x for which P x from
a filter, then the size of the filter decreases by 1. *)
Lemma filter_size_rem:
∀ (T: eqType) (x:T) (xs: seq T) (P: T → bool),
(x \in xs) →
P x →
size [seq y <- xs | P y] = size [seq y <- rem x xs | P y] + 1.
End RemList.
(* Additional lemmas about sequences. *)
Section AdditionalLemmas.
(* First, we show that if n > 0, then nth (x::xs) n = nth xs (n-1). *)
Lemma nth0_cons:
∀ x xs n,
n > 0 →
nth 0 (x :: xs) n = nth 0 xs n.-1.
(* We prove that a sequence xs of size n.+1 can be destructed
into a sequence xs_l of size n and an element x such that
x = xs ++ [::x]. *)
Lemma seq_elim_last:
∀ {X : Type} (n : nat) (xs : seq X),
size xs = n.+1 →
∃ x xs__c, xs = xs__c ++ [:: x] ∧ size xs__c = n.
(* Next, we prove that x ∈ xs implies that xs can be split
into two parts such that xs = xsl ++ [::x] ++ [xsr]. *)
Lemma in_cat:
∀ {X : eqType} (x : X) (xs : list X),
x \in xs → ∃ xsl xsr, xs = xsl ++ [::x] ++ xsr.
(* We prove that for any two sequences xs and ys the fact that xs is a sub-sequence
of ys implies that the size of xs is at most the size of ys. *)
Lemma subseq_leq_size:
∀ {X : eqType} (xs ys: seq X),
uniq xs →
(∀ x, x \in xs → x \in ys) →
size xs ≤ size ys.
(* We prove that if no element of a sequence xs satisfies
a predicate P, then filter P xs is equal to an empty
sequence. *)
Lemma filter_in_pred0:
∀ {X : eqType} (xs : seq X) P,
(∀ x, x \in xs → ~~ P x) →
filter P xs = [::].
End AdditionalLemmas.
Function rem from ssreflect removes only the first occurrence of
an element in a sequence. We define function rem_all which
removes all occurrences of an element in a sequence.
Fixpoint rem_all {X : eqType} (x : X) (xs : seq X) :=
match xs with
| [::] ⇒ [::]
| a :: xs ⇒
if a == x then rem_all x xs else a :: rem_all x xs
end.
match xs with
| [::] ⇒ [::]
| a :: xs ⇒
if a == x then rem_all x xs else a :: rem_all x xs
end.
Additional lemmas about rem_all for lists.
To have a more intuitive naming, we introduce the definition of
range a b which is equal to index_iota a b.+1.
Additional lemmas about index_iota and range for lists.
We prove that one can remove duplicating element from the
head of a sequence by which range is filtered.
Lemma range_filter_2cons:
∀ x xs k,
[seq ρ <- range 0 k | ρ \in x :: x :: xs] =
[seq ρ <- range 0 k | ρ \in x :: xs].
∀ x xs k,
[seq ρ <- range 0 k | ρ \in x :: x :: xs] =
[seq ρ <- range 0 k | ρ \in x :: xs].
Consider a, b, and x s.t. a ≤ x < b,
then filter of iota_index a b with predicate
(? == x) yields ::x.
Corollary index_iota_filter_singl:
∀ x a b,
a ≤ x < b →
[seq ρ <- index_iota a b | ρ \in [:: x]] = [::x].
∀ x a b,
a ≤ x < b →
[seq ρ <- index_iota a b | ρ \in [:: x]] = [::x].
Next we prove that if an element x is not in a sequence xs,
then iota_index a b filtered with predicate (_ ∈ xs) is
equal to iota_index a b filtered with predicate (_ ∈ rem_all
x xs).
Lemma index_iota_filter_inxs:
∀ a b x xs,
x < a →
[seq ρ <- index_iota a b | ρ \in xs] =
[seq ρ <- index_iota a b | ρ \in rem_all x xs].
∀ a b x xs,
x < a →
[seq ρ <- index_iota a b | ρ \in xs] =
[seq ρ <- index_iota a b | ρ \in rem_all x xs].
We prove that if an element x is a min of a sequence xs,
then iota_index a b filtered with predicate (_ ∈ x::xs) is
equal to x :: iota_index a b filtered with predicate (_ ∈
rem_all x xs).
Lemma index_iota_filter_step:
∀ x xs a b,
a ≤ x < b →
(∀ y, y \in xs → x ≤ y) →
[seq ρ <- index_iota a b | ρ \in x :: xs] =
x :: [seq ρ <- index_iota a b | ρ \in rem_all x xs].
∀ x xs a b,
a ≤ x < b →
(∀ y, y \in xs → x ≤ y) →
[seq ρ <- index_iota a b | ρ \in x :: xs] =
x :: [seq ρ <- index_iota a b | ρ \in rem_all x xs].