Lemmas about List.zip
, List.zipWith
, List.zipWithAll
, and List.unzip
. #
Zippers #
zipWith #
theorem
List.getElem?_zipWith
{α : Type u_1}
{β : Type u_2}
{γ : Type u_3}
{as : List α}
{bs : List β}
{f : α → β → γ}
{i : Nat}
:
See also getElem?_zipWith'
for a variant
using Option.map
and Option.bind
rather than a match
.
theorem
List.getElem?_zipWith'
{α : Type u_1}
{β : Type u_2}
{γ : Type u_3}
{l₁ : List α}
{l₂ : List β}
{f : α → β → γ}
{i : Nat}
:
(zipWith f l₁ l₂)[i]? = (Option.map f l₁[i]?).bind fun (g : β → γ) => Option.map g l₂[i]?
Variant of getElem?_zipWith
using Option.map
and Option.bind
rather than a match
.
@[deprecated List.getElem?_zipWith (since := "2024-06-12")]
@[simp]
theorem
List.zipWith_replicate'
{α : Type u_1}
{β : Type u_2}
{α✝ : Type u_3}
{f : α → β → α✝}
{a : α}
{b : β}
{n : Nat}
:
See also List.zipWith_replicate
in Init.Data.List.TakeDrop
for a generalization with different lengths.
zip #
@[reducible, inline, deprecated List.of_mem_zip (since := "2024-07-28")]
Equations
Instances For
@[simp]
See also List.zip_replicate
in Init.Data.List.TakeDrop
for a generalization with different lengths.
zipWithAll #
@[deprecated List.getElem?_zipWithAll (since := "2024-06-12")]
@[simp]
theorem
List.head_zipWithAll
{α : Type u_1}
{β : Type u_2}
{γ : Type u_3}
{as : List α}
{bs : List β}
{f : Option α → Option β → γ}
(h : zipWithAll f as bs ≠ [])
:
(zipWithAll f as bs).head h = f as.head? bs.head?
@[simp]
theorem
List.tail_zipWithAll
{α : Type u_1}
{β : Type u_2}
{γ : Type u_3}
{as : List α}
{bs : List β}
{f : Option α → Option β → γ}
:
(zipWithAll f as bs).tail = zipWithAll f as.tail bs.tail
theorem
List.zipWithAll_map
{γ : Type u_1}
{δ : Type u_2}
{α : Type u_1}
{β : Type u_2}
{μ : Type u_3}
(f : Option γ → Option δ → μ)
(g : α → γ)
(h : β → δ)
(l₁ : List α)
(l₂ : List β)
:
zipWithAll f (map g l₁) (map h l₂) = zipWithAll (fun (a : Option α) (b : Option β) => f (g <$> a) (h <$> b)) l₁ l₂
unzip #
@[reducible, inline, deprecated List.unzip_fst (since := "2024-07-28")]
Equations
Instances For
@[reducible, inline, deprecated List.unzip_snd (since := "2024-07-28")]