Documentation

Mathlib.Data.List.Basic

Basic properties of lists #

instance List.uniqueOfIsEmpty {α : Type u} [IsEmpty α] :

There is only one list of an empty type

Equations
@[simp]
theorem List.cons_injective {α : Type u} {a : α} :
theorem List.singleton_injective {α : Type u} :
Function.Injective fun (a : α) => [a]
theorem List.set_of_mem_cons {α : Type u} (l : List α) (a : α) :
{x : α | x a :: l} = insert a {x : α | x l}

mem #

theorem Decidable.List.eq_or_ne_mem_of_mem {α : Type u} [DecidableEq α] {a b : α} {l : List α} (h : a b :: l) :
a = b a b a l
theorem List.mem_pair {α : Type u} {a b c : α} :
a [b, c] a = b a = c
@[simp]
theorem List.mem_map_of_injective {α : Type u} {β : Type v} {f : αβ} (H : Function.Injective f) {a : α} {l : List α} :
f a map f l a l
@[simp]
theorem Function.Involutive.exists_mem_and_apply_eq_iff {α : Type u} {f : αα} (hf : Involutive f) (x : α) (l : List α) :
( (y : α), y l f y = x) f x l
theorem List.mem_map_of_involutive {α : Type u} {f : αα} (hf : Function.Involutive f) {a : α} {l : List α} :
a map f l f a l

length #

theorem List.length_pos_of_ne_nil {α : Type u_1} {l : List α} :
l []0 < l.length

Alias of the reverse direction of List.length_pos.

theorem List.length_pos_iff_ne_nil {α : Type u} {l : List α} :
0 < l.length l []
theorem List.exists_of_length_succ {α : Type u} {n : } (l : List α) :
l.length = n + 1 (h : α), (t : List α), l = h :: t
theorem List.length_eq_two {α : Type u} {l : List α} :
l.length = 2 (a : α), (b : α), l = [a, b]
theorem List.length_eq_three {α : Type u} {l : List α} :
l.length = 3 (a : α), (b : α), (c : α), l = [a, b, c]

set-theoretic notation of lists #

instance List.instSingletonList {α : Type u} :
Singleton α (List α)
Equations
theorem List.singleton_eq {α : Type u} (x : α) :
{x} = [x]
theorem List.insert_neg {α : Type u} [DecidableEq α] {x : α} {l : List α} (h : ¬x l) :
insert x l = x :: l
theorem List.insert_pos {α : Type u} [DecidableEq α] {x : α} {l : List α} (h : x l) :
insert x l = l
theorem List.doubleton_eq {α : Type u} [DecidableEq α] {x y : α} (h : x y) :
{x, y} = [x, y]

bounded quantifiers over lists #

theorem List.forall_mem_of_forall_mem_cons {α : Type u} {p : αProp} {a : α} {l : List α} (h : ∀ (x : α), x a :: lp x) (x : α) :
x lp x
theorem List.exists_mem_cons_of {α : Type u} {p : αProp} {a : α} (l : List α) (h : p a) :
(x : α), x a :: l p x
theorem List.exists_mem_cons_of_exists {α : Type u} {p : αProp} {a : α} {l : List α} :
( (x : α), x l p x) → (x : α), x a :: l p x
theorem List.or_exists_of_exists_mem_cons {α : Type u} {p : αProp} {a : α} {l : List α} :
( (x : α), x a :: l p x) → p a (x : α), x l p x
theorem List.exists_mem_cons_iff {α : Type u} (p : αProp) (a : α) (l : List α) :
( (x : α), x a :: l p x) p a (x : α), x l p x

list subset #

theorem List.cons_subset_of_subset_of_mem {α : Type u} {a : α} {l m : List α} (ainm : a m) (lsubm : l m) :
a :: l m
theorem List.append_subset_of_subset_of_subset {α : Type u} {l₁ l₂ l : List α} (l₁subl : l₁ l) (l₂subl : l₂ l) :
l₁ ++ l₂ l
theorem List.map_subset_iff {α : Type u} {β : Type v} {l₁ l₂ : List α} (f : αβ) (h : Function.Injective f) :
map f l₁ map f l₂ l₁ l₂

append #

theorem List.append_eq_has_append {α : Type u} {L₁ L₂ : List α} :
L₁.append L₂ = L₁ ++ L₂
theorem List.append_right_injective {α : Type u} (s : List α) :
Function.Injective fun (t : List α) => s ++ t
theorem List.append_left_injective {α : Type u} (t : List α) :
Function.Injective fun (s : List α) => s ++ t

replicate #

theorem List.eq_replicate_length {α : Type u} {a : α} {l : List α} :
l = replicate l.length a ∀ (b : α), b lb = a
theorem List.replicate_add {α : Type u} (m n : ) (a : α) :
replicate (m + n) a = replicate m a ++ replicate n a
theorem List.replicate_subset_singleton {α : Type u} (n : ) (a : α) :
theorem List.subset_singleton_iff {α : Type u} {a : α} {L : List α} :
theorem List.replicate_right_inj {α : Type u} {a b : α} {n : } (hn : n 0) :
replicate n a = replicate n b a = b
theorem List.replicate_right_inj' {α : Type u} {a b : α} {n : } :
replicate n a = replicate n b n = 0 a = b
theorem List.replicate_left_injective {α : Type u} (a : α) :
theorem List.replicate_left_inj {α : Type u} {a : α} {n m : } :
replicate n a = replicate m a n = m

pure #

theorem List.mem_pure {α : Type u} (x y : α) :
x pure y x = y

bind #

@[simp]
theorem List.bind_eq_flatMap {α β : Type u_2} (f : αList β) (l : List α) :
l >>= f = flatMap f l
@[deprecated List.bind_eq_flatMap (since := "2024-10-16")]
theorem List.bind_eq_bind {α β : Type u_2} (f : αList β) (l : List α) :
l >>= f = flatMap f l

Alias of List.bind_eq_flatMap.

concat #

reverse #

theorem List.reverse_cons' {α : Type u} (a : α) (l : List α) :
theorem List.reverse_concat' {α : Type u} (l : List α) (a : α) :
(l ++ [a]).reverse = a :: l.reverse
@[simp]
theorem List.reverse_singleton {α : Type u} (a : α) :
theorem List.concat_eq_reverse_cons {α : Type u} (a : α) (l : List α) :
theorem List.map_reverseAux {α : Type u} {β : Type v} (f : αβ) (l₁ l₂ : List α) :
map f (l₁.reverseAux l₂) = (map f l₁).reverseAux (map f l₂)

getLast #

theorem List.getLast_append_singleton {α : Type u} {a : α} (l : List α) :
(l ++ [a]).getLast = a
theorem List.getLast_append_of_right_ne_nil {α : Type u} (l₁ l₂ : List α) (h : l₂ []) :
(l₁ ++ l₂).getLast = l₂.getLast h
@[deprecated List.getLast_append_of_right_ne_nil (since := "2025-02-06")]
theorem List.getLast_append' {α : Type u} (l₁ l₂ : List α) (h : l₂ []) :
(l₁ ++ l₂).getLast = l₂.getLast h

Alias of List.getLast_append_of_right_ne_nil.

theorem List.getLast_concat' {α : Type u} {a : α} (l : List α) :
(l.concat a).getLast = a
@[simp]
theorem List.getLast_singleton' {α : Type u} (a : α) :
[a].getLast = a
@[simp]
theorem List.getLast_cons_cons {α : Type u} (a₁ a₂ : α) (l : List α) :
(a₁ :: a₂ :: l).getLast = (a₂ :: l).getLast
theorem List.dropLast_append_getLast {α : Type u} {l : List α} (h : l []) :
theorem List.getLast_congr {α : Type u} {l₁ l₂ : List α} (h₁ : l₁ []) (h₂ : l₂ []) (h₃ : l₁ = l₂) :
l₁.getLast h₁ = l₂.getLast h₂
theorem List.getLast_replicate_succ {α : Type u} (m : ) (a : α) :
(replicate (m + 1) a).getLast = a
@[deprecated List.getLast_filter_of_pos (since := "2025-02-07")]
theorem List.getLast_filter' {α : Type u_1} {p : αBool} {l : List α} (w : l []) (h : p (l.getLast w) = true) :
(filter p l).getLast = l.getLast w

Alias of List.getLast_filter_of_pos.

getLast? #

@[deprecated List.getLast?_eq_none_iff (since := "2024-09-06")]
theorem List.getLast?_eq_none {α : Type u_1} {xs : List α} :

Alias of List.getLast?_eq_none_iff.

theorem List.mem_getLast?_eq_getLast {α : Type u} {l : List α} {x : α} :
x l.getLast? (h : l []), x = l.getLast h
theorem List.getLast?_eq_getLast_of_ne_nil {α : Type u} {l : List α} (h : l []) :
theorem List.mem_getLast?_cons {α : Type u} {x y : α} {l : List α} :
x l.getLast?x (y :: l).getLast?
theorem List.dropLast_append_getLast? {α : Type u} {l : List α} (a : α) :
a l.getLast?l.dropLast ++ [a] = l
theorem List.getLast?_append_cons {α : Type u} (l₁ : List α) (a : α) (l₂ : List α) :
(l₁ ++ a :: l₂).getLast? = (a :: l₂).getLast?
theorem List.getLast?_append_of_ne_nil {α : Type u} (l₁ : List α) {l₂ : List α} :
l₂ [](l₁ ++ l₂).getLast? = l₂.getLast?
theorem List.mem_getLast?_append_of_mem_getLast? {α : Type u} {l₁ l₂ : List α} {x : α} (h : x l₂.getLast?) :
x (l₁ ++ l₂).getLast?

head(!?) and tail #

@[simp]
theorem List.head!_nil {α : Type u} [Inhabited α] :
@[simp]
theorem List.head_cons_tail {α : Type u} (x : List α) (h : x []) :
x.head h :: x.tail = x
theorem List.head_eq_getElem_zero {α : Type u} {l : List α} (hl : l []) :
l.head hl = l[0]
theorem List.head!_eq_head? {α : Type u} [Inhabited α] (l : List α) :
theorem List.eq_cons_of_mem_head? {α : Type u} {x : α} {l : List α} :
x l.head?l = x :: l.tail
@[simp]
theorem List.head!_cons {α : Type u} [Inhabited α] (a : α) (l : List α) :
(a :: l).head! = a
@[simp]
theorem List.head!_append {α : Type u} [Inhabited α] (t : List α) {s : List α} (h : s []) :
(s ++ t).head! = s.head!
theorem List.mem_head?_append_of_mem_head? {α : Type u} {s t : List α} {x : α} (h : x s.head?) :
x (s ++ t).head?
theorem List.head?_append_of_ne_nil {α : Type u} (l₁ : List α) {l₂ : List α} :
l₁ [](l₁ ++ l₂).head? = l₁.head?
theorem List.tail_append_singleton_of_ne_nil {α : Type u} {a : α} {l : List α} (h : l []) :
(l ++ [a]).tail = l.tail ++ [a]
theorem List.cons_head?_tail {α : Type u} {l : List α} {a : α} :
a l.head?a :: l.tail = l
theorem List.head!_mem_head? {α : Type u} [Inhabited α] {l : List α} :
l []l.head! l.head?
theorem List.cons_head!_tail {α : Type u} [Inhabited α] {l : List α} (h : l []) :
l.head! :: l.tail = l
theorem List.head!_mem_self {α : Type u} [Inhabited α] {l : List α} (h : l []) :
theorem List.get_eq_getElem? {α : Type u} (l : List α) (i : Fin l.length) :
l.get i = l[i]?.get
@[deprecated List.get_eq_getElem? (since := "2025-02-15")]
theorem List.get_eq_get? {α : Type u} (l : List α) (i : Fin l.length) :
l.get i = l[i]?.get

Alias of List.get_eq_getElem?.

theorem List.exists_mem_iff_getElem {α : Type u} {l : List α} {p : αProp} :
( (x : α), x l p x) (i : ), (x : i < l.length), p l[i]
theorem List.forall_mem_iff_getElem {α : Type u} {l : List α} {p : αProp} :
(∀ (x : α), x lp x) ∀ (i : ) (x : i < l.length), p l[i]
theorem List.get_tail {α : Type u} (l : List α) (i : ) (h : i < l.tail.length) (h' : i + 1 < l.length := ) :
l.tail.get i, h = l.get i + 1, h'
@[deprecated "No deprecation message was provided." (since := "2024-08-22")]
theorem List.get_cons {α : Type u} {l : List α} {a : α} {n : } (hl : n < (a :: l).length) :
(a :: l).get n, hl = if hn : n = 0 then a else l.get n - 1,

sublists #

theorem List.Sublist.cons_cons {α : Type u} {l₁ l₂ : List α} (a : α) (s : l₁.Sublist l₂) :
(a :: l₁).Sublist (a :: l₂)
theorem List.cons_sublist_cons' {α : Type u} {l₁ l₂ : List α} {a b : α} :
(a :: l₁).Sublist (b :: l₂) (a :: l₁).Sublist l₂ a = b l₁.Sublist l₂
theorem List.sublist_cons_of_sublist {α : Type u} {l₁ l₂ : List α} (a : α) (h : l₁.Sublist l₂) :
l₁.Sublist (a :: l₂)
@[deprecated List.sublist_nil (since := "2025-02-07")]
theorem List.sublist_nil_iff_eq_nil {α : Type u_1} {l : List α} :

Alias of List.sublist_nil.

@[simp]
theorem List.sublist_singleton {α : Type u} {l : List α} {a : α} :
l.Sublist [a] l = [] l = [a]
theorem List.Sublist.antisymm {α : Type u} {l₁ l₂ : List α} (s₁ : l₁.Sublist l₂) (s₂ : l₂.Sublist l₁) :
l₁ = l₂
theorem List.Sublist.of_cons_of_ne {α : Type u} {l₁ l₂ : List α} {a b : α} (h₁ : a b) (h₂ : (a :: l₁).Sublist (b :: l₂)) :
(a :: l₁).Sublist l₂

If the first element of two lists are different, then a sublist relation can be reduced.

indexOf #

theorem List.idxOf_cons_eq {α : Type u} [DecidableEq α] {a b : α} (l : List α) :
b = aidxOf a (b :: l) = 0
@[deprecated List.idxOf_cons_eq (since := "2025-01-30")]
theorem List.indexOf_cons_eq {α : Type u} [DecidableEq α] {a b : α} (l : List α) :
b = aidxOf a (b :: l) = 0

Alias of List.idxOf_cons_eq.

@[simp]
theorem List.idxOf_cons_ne {α : Type u} [DecidableEq α] {a b : α} (l : List α) :
b aidxOf a (b :: l) = (idxOf a l).succ
@[deprecated List.idxOf_cons_ne (since := "2025-01-30")]
theorem List.indexOf_cons_ne {α : Type u} [DecidableEq α] {a b : α} (l : List α) :
b aidxOf a (b :: l) = (idxOf a l).succ

Alias of List.idxOf_cons_ne.

theorem List.idxOf_eq_length_iff {α : Type u} [DecidableEq α] {a : α} {l : List α} :
idxOf a l = l.length ¬a l
@[simp]
theorem List.idxOf_of_not_mem {α : Type u} [DecidableEq α] {l : List α} {a : α} :
¬a lidxOf a l = l.length
@[deprecated List.idxOf_of_not_mem (since := "2025-01-30")]
theorem List.indexOf_of_not_mem {α : Type u} [DecidableEq α] {l : List α} {a : α} :
¬a lidxOf a l = l.length

Alias of List.idxOf_of_not_mem.

theorem List.idxOf_le_length {α : Type u} [DecidableEq α] {a : α} {l : List α} :
@[deprecated List.idxOf_le_length (since := "2025-01-30")]
theorem List.indexOf_le_length {α : Type u} [DecidableEq α] {a : α} {l : List α} :

Alias of List.idxOf_le_length.

theorem List.idxOf_lt_length_iff {α : Type u} [DecidableEq α] {a : α} {l : List α} :
idxOf a l < l.length a l
@[deprecated List.idxOf_lt_length_iff (since := "2025-01-30")]
theorem List.indexOf_lt_length_iff {α : Type u} [DecidableEq α] {a : α} {l : List α} :
idxOf a l < l.length a l

Alias of List.idxOf_lt_length_iff.

theorem List.idxOf_append_of_mem {α : Type u} {l₁ l₂ : List α} [DecidableEq α] {a : α} (h : a l₁) :
idxOf a (l₁ ++ l₂) = idxOf a l₁
@[deprecated List.idxOf_append_of_mem (since := "2025-01-30")]
theorem List.indexOf_append_of_mem {α : Type u} {l₁ l₂ : List α} [DecidableEq α] {a : α} (h : a l₁) :
idxOf a (l₁ ++ l₂) = idxOf a l₁

Alias of List.idxOf_append_of_mem.

theorem List.idxOf_append_of_not_mem {α : Type u} {l₁ l₂ : List α} [DecidableEq α] {a : α} (h : ¬a l₁) :
idxOf a (l₁ ++ l₂) = l₁.length + idxOf a l₂
@[deprecated List.idxOf_append_of_not_mem (since := "2025-01-30")]
theorem List.indexOf_append_of_not_mem {α : Type u} {l₁ l₂ : List α} [DecidableEq α] {a : α} (h : ¬a l₁) :
idxOf a (l₁ ++ l₂) = l₁.length + idxOf a l₂

Alias of List.idxOf_append_of_not_mem.

nth element #

@[simp]
theorem List.getElem?_length {α : Type u} (l : List α) :
theorem List.getElem_map_rev {α : Type u} {β : Type v} (f : αβ) {l : List α} {n : } {h : n < l.length} :
f l[n] = (map f l)[n]

A version of getElem_map that can be used for rewriting.

theorem List.get_length_sub_one {α : Type u} {l : List α} (h : l.length - 1 < l.length) :
l.get l.length - 1, h = l.getLast
theorem List.take_one_drop_eq_of_lt_length {α : Type u} {l : List α} {n : } (h : n < l.length) :
take 1 (drop n l) = [l.get n, h]
theorem List.ext_getElem?' {α : Type u} {l₁ l₂ : List α} (h' : ∀ (n : ), n < l₁.length l₂.lengthl₁[n]? = l₂[n]?) :
l₁ = l₂
@[deprecated List.ext_getElem?' (since := "2025-02-15")]
theorem List.ext_get?' {α : Type u} {l₁ l₂ : List α} (h' : ∀ (n : ), n < l₁.length l₂.lengthl₁[n]? = l₂[n]?) :
l₁ = l₂

Alias of List.ext_getElem?'.

@[deprecated List.ext_getElem?_iff (since := "2025-02-15")]
theorem List.ext_get?_iff {α : Type u_1} {l₁ l₂ : List α} :
l₁ = l₂ ∀ (i : ), l₁[i]? = l₂[i]?

Alias of List.ext_getElem?_iff.

theorem List.ext_get_iff {α : Type u} {l₁ l₂ : List α} :
l₁ = l₂ l₁.length = l₂.length ∀ (n : ) (h₁ : n < l₁.length) (h₂ : n < l₂.length), l₁.get n, h₁ = l₂.get n, h₂
theorem List.ext_getElem?_iff' {α : Type u} {l₁ l₂ : List α} :
l₁ = l₂ ∀ (n : ), n < l₁.length l₂.lengthl₁[n]? = l₂[n]?
@[deprecated List.ext_getElem?_iff' (since := "2025-02-15")]
theorem List.ext_get?_iff' {α : Type u} {l₁ l₂ : List α} :
l₁ = l₂ ∀ (n : ), n < l₁.length l₂.lengthl₁[n]? = l₂[n]?

Alias of List.ext_getElem?_iff'.

theorem List.ext_getElem! {α : Type u} {l₁ l₂ : List α} [Inhabited α] (hl : l₁.length = l₂.length) (h : ∀ (n : ), l₁[n]! = l₂[n]!) :
l₁ = l₂

If two lists l₁ and l₂ are the same length and l₁[n]! = l₂[n]! for all n, then the lists are equal.

@[simp]
theorem List.getElem_idxOf {α : Type u} [DecidableEq α] {a : α} {l : List α} (h : idxOf a l < l.length) :
l[idxOf a l] = a
@[deprecated List.getElem_idxOf (since := "2025-01-30")]
theorem List.getElem_indexOf {α : Type u} [DecidableEq α] {a : α} {l : List α} (h : idxOf a l < l.length) :
l[idxOf a l] = a

Alias of List.getElem_idxOf.

theorem List.idxOf_get {α : Type u} [DecidableEq α] {a : α} {l : List α} (h : idxOf a l < l.length) :
l.get idxOf a l, h = a
@[deprecated List.idxOf_get (since := "2025-01-30")]
theorem List.indexOf_get {α : Type u} [DecidableEq α] {a : α} {l : List α} (h : idxOf a l < l.length) :
l.get idxOf a l, h = a

Alias of List.idxOf_get.

@[simp]
theorem List.getElem?_idxOf {α : Type u} [DecidableEq α] {a : α} {l : List α} (h : a l) :
l[idxOf a l]? = some a
@[deprecated List.getElem?_idxOf (since := "2025-01-30")]
theorem List.getElem?_indexOf {α : Type u} [DecidableEq α] {a : α} {l : List α} (h : a l) :
l[idxOf a l]? = some a

Alias of List.getElem?_idxOf.

@[deprecated List.getElem?_idxOf (since := "2025-02-15")]
theorem List.idxOf_get? {α : Type u} [DecidableEq α] {a : α} {l : List α} (h : a l) :
l[idxOf a l]? = some a

Alias of List.getElem?_idxOf.

@[deprecated List.getElem?_idxOf (since := "2025-01-30")]
theorem List.indexOf_get? {α : Type u} [DecidableEq α] {a : α} {l : List α} (h : a l) :
l[idxOf a l]? = some a

Alias of List.getElem?_idxOf.

theorem List.idxOf_inj {α : Type u} [DecidableEq α] {l : List α} {x y : α} (hx : x l) (hy : y l) :
idxOf x l = idxOf y l x = y
@[deprecated List.idxOf_inj (since := "2025-01-30")]
theorem List.indexOf_inj {α : Type u} [DecidableEq α] {l : List α} {x y : α} (hx : x l) (hy : y l) :
idxOf x l = idxOf y l x = y

Alias of List.idxOf_inj.

theorem List.get_reverse' {α : Type u} (l : List α) (n : Fin l.reverse.length) (hn' : l.length - 1 - n < l.length) :
l.reverse.get n = l.get l.length - 1 - n, hn'
theorem List.eq_cons_of_length_one {α : Type u} {l : List α} (h : l.length = 1) :
l = [l.get 0, ]
@[deprecated List.modifyTailIdx_modifyTailIdx_le (since := "2024-10-21")]
theorem List.modifyNthTail_modifyNthTail_le {α : Type u_1} {f g : List αList α} (m n : ) (l : List α) (h : n m) :
modifyTailIdx g m (modifyTailIdx f n l) = modifyTailIdx (fun (l : List α) => modifyTailIdx g (m - n) (f l)) n l

Alias of List.modifyTailIdx_modifyTailIdx_le.

@[deprecated List.modifyTailIdx_modifyTailIdx_self (since := "2024-10-21")]
theorem List.modifyNthTail_modifyNthTail_same {α : Type u_1} {f g : List αList α} (n : ) (l : List α) :

Alias of List.modifyTailIdx_modifyTailIdx_self.

@[deprecated List.modify_eq_set (since := "2024-10-21")]
theorem List.modifyNth_eq_set {α : Type u_1} [Inhabited α] (f : αα) (n : ) (l : List α) :
modify f n l = l.set n (f (l[n]?.getD default))

Alias of List.modify_eq_set.

@[simp]
theorem List.getElem_set_of_ne {α : Type u} {l : List α} {i j : } (h : i j) (a : α) (hj : j < (l.set i a).length) :
(l.set i a)[j] = l[j]

map #

theorem List.flatMap_pure_eq_map {α : Type u} {β : Type v} (f : αβ) (l : List α) :
flatMap (pure f) l = map f l
@[deprecated List.flatMap_pure_eq_map (since := "2024-10-16")]
theorem List.bind_pure_eq_map {α : Type u} {β : Type v} (f : αβ) (l : List α) :
flatMap (pure f) l = map f l

Alias of List.flatMap_pure_eq_map.

theorem List.flatMap_congr {α : Type u} {β : Type v} {l : List α} {f g : αList β} (h : ∀ (x : α), x lf x = g x) :
flatMap f l = flatMap g l
@[deprecated List.flatMap_congr (since := "2024-10-16")]
theorem List.bind_congr {α : Type u} {β : Type v} {l : List α} {f g : αList β} (h : ∀ (x : α), x lf x = g x) :
flatMap f l = flatMap g l

Alias of List.flatMap_congr.

theorem List.infix_flatMap_of_mem {α : Type u} {a : α} {as : List α} (h : a as) (f : αList α) :
f a <:+: flatMap f as
@[deprecated List.infix_flatMap_of_mem (since := "2024-10-16")]
theorem List.infix_bind_of_mem {α : Type u} {a : α} {as : List α} (h : a as) (f : αList α) :
f a <:+: flatMap f as

Alias of List.infix_flatMap_of_mem.

@[simp]
theorem List.map_eq_map {α β : Type u_2} (f : αβ) (l : List α) :
f <$> l = map f l
theorem List.comp_map {α : Type u} {β : Type v} {γ : Type w} (h : βγ) (g : αβ) (l : List α) :
map (h g) l = map h (map g l)

A single List.map of a composition of functions is equal to composing a List.map with another List.map, fully applied. This is the reverse direction of List.map_map.

@[simp]
theorem List.map_comp_map {α : Type u} {β : Type v} {γ : Type w} (g : βγ) (f : αβ) :
map g map f = map (g f)

Composing a List.map with another List.map is equal to a single List.map of composed functions.

theorem Function.LeftInverse.list_map {α : Type u} {β : Type v} {f : αβ} {g : βα} (h : LeftInverse f g) :
theorem Function.RightInverse.list_map {α : Type u} {β : Type v} {f : αβ} {g : βα} (h : RightInverse f g) :
theorem Function.Involutive.list_map {α : Type u} {f : αα} (h : Involutive f) :
@[simp]
theorem List.map_leftInverse_iff {α : Type u} {β : Type v} {f : αβ} {g : βα} :
@[simp]
theorem List.map_rightInverse_iff {α : Type u} {β : Type v} {f : αβ} {g : βα} :
@[simp]
theorem Function.Injective.list_map {α : Type u} {β : Type v} {f : αβ} (h : Injective f) :
@[simp]
theorem List.map_injective_iff {α : Type u} {β : Type v} {f : αβ} :
theorem Function.Surjective.list_map {α : Type u} {β : Type v} {f : αβ} (h : Surjective f) :
@[simp]
theorem List.map_surjective_iff {α : Type u} {β : Type v} {f : αβ} :
theorem Function.Bijective.list_map {α : Type u} {β : Type v} {f : αβ} (h : Bijective f) :
@[simp]
theorem List.map_bijective_iff {α : Type u} {β : Type v} {f : αβ} :
theorem List.eq_of_mem_map_const {α : Type u} {β : Type v} {b₁ b₂ : β} {l : List α} (h : b₁ map (Function.const α b₂) l) :
b₁ = b₂
theorem List.eq_nil_or_concat' {α : Type u} (l : List α) :
l = [] (L : List α), (b : α), l = L ++ [b]

eq_nil_or_concat in simp normal form

foldl, foldr #

theorem List.foldl_ext {α : Type u} {β : Type v} (f g : αβα) (a : α) {l : List β} (H : ∀ (a : α) (b : β), b lf a b = g a b) :
foldl f a l = foldl g a l
theorem List.foldr_ext {α : Type u} {β : Type v} (f g : αββ) (b : β) {l : List α} (H : ∀ (a : α), a l∀ (b : β), f a b = g a b) :
foldr f b l = foldr g b l
theorem List.foldl_concat {α : Type u} {β : Type v} (f : βαβ) (b : β) (x : α) (xs : List α) :
foldl f b (xs ++ [x]) = f (foldl f b xs) x
theorem List.foldr_concat {α : Type u} {β : Type v} (f : αββ) (b : β) (x : α) (xs : List α) :
foldr f b (xs ++ [x]) = foldr f (f x b) xs
theorem List.foldl_fixed' {α : Type u} {β : Type v} {f : αβα} {a : α} (hf : ∀ (b : β), f a b = a) (l : List β) :
foldl f a l = a
theorem List.foldr_fixed' {α : Type u} {β : Type v} {f : αββ} {b : β} (hf : ∀ (a : α), f a b = b) (l : List α) :
foldr f b l = b
@[simp]
theorem List.foldl_fixed {α : Type u} {β : Type v} {a : α} (l : List β) :
foldl (fun (a : α) (x : β) => a) a l = a
@[simp]
theorem List.foldr_fixed {α : Type u} {β : Type v} {b : β} (l : List α) :
foldr (fun (x : α) (b : β) => b) b l = b
@[deprecated List.foldr_cons_nil (since := "2025-02-10")]
theorem List.foldr_eta {α : Type u} (l : List α) :
theorem List.reverse_foldl {α : Type u} {l : List α} :
(foldl (fun (t : List α) (h : α) => h :: t) [] l).reverse = l
theorem List.foldl_hom₂ {ι : Type u_1} {α : Type u} {β : Type v} {γ : Type w} (l : List ι) (f : αβγ) (op₁ : αια) (op₂ : βιβ) (op₃ : γιγ) (a : α) (b : β) (h : ∀ (a : α) (b : β) (i : ι), f (op₁ a i) (op₂ b i) = op₃ (f a b) i) :
foldl op₃ (f a b) l = f (foldl op₁ a l) (foldl op₂ b l)
theorem List.foldr_hom₂ {ι : Type u_1} {α : Type u} {β : Type v} {γ : Type w} (l : List ι) (f : αβγ) (op₁ : ιαα) (op₂ : ιββ) (op₃ : ιγγ) (a : α) (b : β) (h : ∀ (a : α) (b : β) (i : ι), f (op₁ i a) (op₂ i b) = op₃ i (f a b)) :
foldr op₃ (f a b) l = f (foldr op₁ a l) (foldr op₂ b l)
theorem List.injective_foldl_comp {α : Type u} {l : List (αα)} {f : αα} (hl : ∀ (f : αα), f lFunction.Injective f) (hf : Function.Injective f) :
theorem List.append_cons_inj_of_not_mem {α : Type u} {x₁ x₂ z₁ z₂ : List α} {a₁ a₂ : α} (notin_x : ¬a₂ x₁) (notin_z : ¬a₂ z₁) :
x₁ ++ a₁ :: z₁ = x₂ ++ a₂ :: z₂ x₁ = x₂ a₁ = a₂ z₁ = z₂

Consider two lists l₁ and l₂ with designated elements a₁ and a₂ somewhere in them: l₁ = x₁ ++ [a₁] ++ z₁ and l₂ = x₂ ++ [a₂] ++ z₂. Assume the designated element a₂ is present in neither x₁ nor z₁. We conclude that the lists are equal (l₁ = l₂) if and only if their respective parts are equal (x₁ = x₂ ∧ a₁ = a₂ ∧ z₁ = z₂).

theorem List.foldl1_eq_foldr1 {α : Type u} {f : ααα} [hassoc : Std.Associative f] (a b : α) (l : List α) :
foldl f a (l ++ [b]) = foldr f b (a :: l)
theorem List.foldl_eq_of_comm_of_assoc {α : Type u} {f : ααα} [hcomm : Std.Commutative f] [hassoc : Std.Associative f] (a b : α) (l : List α) :
foldl f a (b :: l) = f b (foldl f a l)
theorem List.foldl_eq_foldr {α : Type u} {f : ααα} [Std.Commutative f] [Std.Associative f] (a : α) (l : List α) :
foldl f a l = foldr f a l
theorem List.foldl_eq_of_comm' {α : Type u} {β : Type v} {f : αβα} (hf : ∀ (a : α) (b c : β), f (f a b) c = f (f a c) b) (a : α) (b : β) (l : List β) :
foldl f a (b :: l) = f (foldl f a l) b
theorem List.foldl_eq_foldr' {α : Type u} {β : Type v} {f : αβα} (hf : ∀ (a : α) (b c : β), f (f a b) c = f (f a c) b) (a : α) (l : List β) :
foldl f a l = foldr (flip f) a l
theorem List.foldr_eq_of_comm' {α : Type u} {β : Type v} {f : αββ} (hf : ∀ (a b : α) (c : β), f a (f b c) = f b (f a c)) (a : β) (b : α) (l : List α) :
foldr f a (b :: l) = foldr f (f b a) l
theorem List.foldl_op_eq_op_foldr_assoc {α : Type u} {op : ααα} [ha : Std.Associative op] {l : List α} {a₁ a₂ : α} :
op (foldl op a₁ l) a₂ = op a₁ (foldr (fun (x1 x2 : α) => op x1 x2) a₂ l)
theorem List.foldl_assoc_comm_cons {α : Type u} {op : ααα} [ha : Std.Associative op] [hc : Std.Commutative op] {l : List α} {a₁ a₂ : α} :
foldl op a₂ (a₁ :: l) = op a₁ (foldl op a₂ l)

foldlM, foldrM, mapM #

theorem List.foldrM_eq_foldr {α : Type u} {β : Type v} {m : Type v → Type w} [Monad m] [LawfulMonad m] (f : αβm β) (b : β) (l : List α) :
foldrM f b l = foldr (fun (a : α) (mb : m β) => mb >>= f a) (pure b) l
theorem List.foldlM_eq_foldl {α : Type u} {β : Type v} {m : Type v → Type w} [Monad m] [LawfulMonad m] (f : βαm β) (b : β) (l : List α) :
List.foldlM f b l = foldl (fun (mb : m β) (a : α) => do let bmb f b a) (pure b) l

intersperse #

@[deprecated List.intersperse_single (since := "2025-02-07")]
theorem List.intersperse_singleton {α : Type u} {x : α} (sep : α) :

Alias of List.intersperse_single.

@[deprecated List.intersperse_cons₂ (since := "2025-02-07")]
theorem List.intersperse_cons_cons {α : Type u} {x y : α} {zs : List α} (sep : α) :
intersperse sep (x :: y :: zs) = x :: sep :: intersperse sep (y :: zs)

Alias of List.intersperse_cons₂.

map for partial functions #

@[deprecated "Deprecated without replacement." (since := "2025-02-07")]
theorem List.sizeOf_lt_sizeOf_of_mem {α : Type u} [SizeOf α] {x : α} {l : List α} (hx : x l) :

filter #

theorem List.length_eq_length_filter_add {α : Type u} {l : List α} (f : αBool) :
l.length = (filter f l).length + (filter (fun (x : α) => !f x) l).length

filterMap #

theorem List.filterMap_eq_flatMap_toList {α : Type u} {β : Type v} (f : αOption β) (l : List α) :
filterMap f l = flatMap (fun (a : α) => (f a).toList) l
@[deprecated List.filterMap_eq_flatMap_toList (since := "2024-10-16")]
theorem List.filterMap_eq_bind_toList {α : Type u} {β : Type v} (f : αOption β) (l : List α) :
filterMap f l = flatMap (fun (a : α) => (f a).toList) l

Alias of List.filterMap_eq_flatMap_toList.

theorem List.filterMap_congr {α : Type u} {β : Type v} {f g : αOption β} {l : List α} (h : ∀ (x : α), x lf x = g x) :
theorem List.filterMap_eq_map_iff_forall_eq_some {α : Type u} {β : Type v} {f : αOption β} {g : αβ} {l : List α} :
filterMap f l = map g l ∀ (x : α), x lf x = some (g x)

filter #

theorem List.filter_singleton {α : Type u} {p : αBool} {a : α} :
filter p [a] = bif p a then [a] else []
theorem List.filter_eq_foldr {α : Type u} (p : αBool) (l : List α) :
filter p l = foldr (fun (a : α) (out : List α) => bif p a then a :: out else out) [] l
@[simp]
theorem List.filter_subset' {α : Type u} {p : αBool} (l : List α) :
filter p l l
theorem List.of_mem_filter {α : Type u} {p : αBool} {a : α} {l : List α} (h : a filter p l) :
p a = true
theorem List.mem_of_mem_filter {α : Type u} {p : αBool} {a : α} {l : List α} (h : a filter p l) :
a l
theorem List.mem_filter_of_mem {α : Type u} {p : αBool} {a : α} {l : List α} (h₁ : a l) (h₂ : p a = true) :
a filter p l
@[deprecated List.filter_subset (since := "2025-02-07")]
theorem List.monotone_filter_left {α : Type u_1} {l₁ l₂ : List α} (p : αBool) (H : l₁ l₂) :
filter p l₁ filter p l₂

Alias of List.filter_subset.

theorem List.monotone_filter_right {α : Type u} (l : List α) ⦃p q : αBool (h : ∀ (a : α), p a = trueq a = true) :
(filter p l).Sublist (filter q l)
theorem List.map_filter {α : Type u} {β : Type v} (p : αBool) {f : αβ} (hf : Function.Injective f) (l : List α) [DecidablePred fun (b : β) => (a : α), p a = true f a = b] :
map f (filter p l) = filter (fun (b : β) => decide ( (a : α), p a = true f a = b)) (map f l)
@[deprecated List.map_filter (since := "2025-02-07")]
theorem List.map_filter' {α : Type u} {β : Type v} (p : αBool) {f : αβ} (hf : Function.Injective f) (l : List α) [DecidablePred fun (b : β) => (a : α), p a = true f a = b] :
map f (filter p l) = filter (fun (b : β) => decide ( (a : α), p a = true f a = b)) (map f l)

Alias of List.map_filter.

theorem List.filter_attach' {α : Type u} (l : List α) (p : { a : α // a l }Bool) [DecidableEq α] :
filter p l.attach = map (Subtype.map id ) (filter (fun (x : α) => decide ( (h : x l), p x, h = true)) l).attach
theorem List.filter_attach {α : Type u} (l : List α) (p : αBool) :
filter (fun (x : { x : α // x l }) => p x) l.attach = map (Subtype.map id ) (filter p l).attach
theorem List.filter_comm {α : Type u} (p q : αBool) (l : List α) :
filter p (filter q l) = filter q (filter p l)
@[simp]
theorem List.filter_true {α : Type u} (l : List α) :
filter (fun (x : α) => true) l = l
@[simp]
theorem List.filter_false {α : Type u} (l : List α) :
filter (fun (x : α) => false) l = []
@[deprecated List.getElem_cons (since := "2024-08-19")]
theorem List.nthLe_cons {α : Type u_1} {i : } {a : α} {l : List α} (w : i < (a :: l).length) :
(a :: l)[i] = if h : i = 0 then a else l[i - 1]

Alias of List.getElem_cons.

eraseP #

@[simp]
theorem List.length_eraseP_add_one {α : Type u} {p : αBool} {l : List α} {a : α} (al : a l) (pa : p a = true) :
(eraseP p l).length + 1 = l.length

erase #

@[simp]
theorem List.length_erase_add_one {α : Type u} [DecidableEq α] {a : α} {l : List α} (h : a l) :
(l.erase a).length + 1 = l.length
theorem List.map_erase {α : Type u} {β : Type v} [DecidableEq α] [DecidableEq β] {f : αβ} (finj : Function.Injective f) {a : α} (l : List α) :
map f (l.erase a) = (map f l).erase (f a)
theorem List.map_foldl_erase {α : Type u} {β : Type v} [DecidableEq α] [DecidableEq β] {f : αβ} (finj : Function.Injective f) {l₁ l₂ : List α} :
map f (foldl List.erase l₁ l₂) = foldl (fun (l : List β) (a : α) => l.erase (f a)) (map f l₁) l₂
theorem List.erase_getElem {ι : Type u_1} [DecidableEq ι] {l : List ι} {i : } (hi : i < l.length) :
(l.erase l[i]).Perm (l.eraseIdx i)
@[deprecated List.erase_getElem (since := "2024-08-03")]
theorem List.erase_get {ι : Type u_1} [DecidableEq ι] {l : List ι} (i : Fin l.length) :
(l.erase (l.get i)).Perm (l.eraseIdx i)
theorem List.length_eraseIdx_add_one {ι : Type u_1} {l : List ι} {i : } (h : i < l.length) :

diff #

@[simp]
theorem List.map_diff {α : Type u} {β : Type v} [DecidableEq α] [DecidableEq β] {f : αβ} (finj : Function.Injective f) {l₁ l₂ : List α} :
map f (l₁.diff l₂) = (map f l₁).diff (map f l₂)
theorem List.erase_diff_erase_sublist_of_sublist {α : Type u} [DecidableEq α] {a : α} {l₁ l₂ : List α} :
l₁.Sublist l₂((l₂.erase a).diff (l₁.erase a)).Sublist (l₂.diff l₁)
theorem List.choose_spec {α : Type u} (p : αProp) [DecidablePred p] (l : List α) (hp : (a : α), a l p a) :
choose p l hp l p (choose p l hp)
theorem List.choose_mem {α : Type u} (p : αProp) [DecidablePred p] (l : List α) (hp : (a : α), a l p a) :
choose p l hp l
theorem List.choose_property {α : Type u} (p : αProp) [DecidablePred p] (l : List α) (hp : (a : α), a l p a) :
p (choose p l hp)

Forall #

@[simp]
theorem List.forall_cons {α : Type u} (p : αProp) (x : α) (l : List α) :
Forall p (x :: l) p x Forall p l
theorem List.forall_iff_forall_mem {α : Type u} {p : αProp} {l : List α} :
Forall p l ∀ (x : α), x lp x
theorem List.Forall.imp {α : Type u} {p q : αProp} (h : ∀ (x : α), p xq x) {l : List α} :
Forall p lForall q l
@[simp]
theorem List.forall_map_iff {α : Type u} {β : Type v} {l : List α} {p : βProp} (f : αβ) :
Forall p (map f l) Forall (p f) l
Equations

Miscellaneous lemmas #

theorem List.get_attach {α : Type u} (L : List α) (i : Fin L.attach.length) :
(L.attach.get i) = L.get i,
theorem List.disjoint_pmap {α : Type u} {β : Type v} {p : αProp} {f : (a : α) → p aβ} {s t : List α} (hs : ∀ (a : α), a sp a) (ht : ∀ (a : α), a tp a) (hf : ∀ (a a' : α) (ha : p a) (ha' : p a'), f a ha = f a' ha'a = a') (h : s.Disjoint t) :
(pmap f s hs).Disjoint (pmap f t ht)

The images of disjoint lists under a partially defined map are disjoint

theorem List.disjoint_map {α : Type u} {β : Type v} {f : αβ} {s t : List α} (hf : Function.Injective f) (h : s.Disjoint t) :
(map f s).Disjoint (map f t)

The images of disjoint lists under an injective map are disjoint

theorem List.Disjoint.map {α : Type u} {β : Type v} {f : αβ} {s t : List α} (hf : Function.Injective f) (h : s.Disjoint t) :

Alias of List.disjoint_map.


The images of disjoint lists under an injective map are disjoint

theorem List.Disjoint.of_map {α : Type u} {β : Type v} {f : αβ} {s t : List α} (h : (List.map f s).Disjoint (List.map f t)) :
theorem List.Disjoint.map_iff {α : Type u} {β : Type v} {f : αβ} {s t : List α} (hf : Function.Injective f) :
theorem List.Perm.disjoint_left {α : Type u} {l₁ l₂ l : List α} (p : l₁.Perm l₂) :
l₁.Disjoint l l₂.Disjoint l
theorem List.Perm.disjoint_right {α : Type u} {l₁ l₂ l : List α} (p : l₁.Perm l₂) :
l.Disjoint l₁ l.Disjoint l₂
@[simp]
theorem List.disjoint_reverse_left {α : Type u} {l₁ l₂ : List α} :
l₁.reverse.Disjoint l₂ l₁.Disjoint l₂
@[simp]
theorem List.disjoint_reverse_right {α : Type u} {l₁ l₂ : List α} :
l₁.Disjoint l₂.reverse l₁.Disjoint l₂
theorem List.lookup_graph {α : Type u} {β : Type v} [BEq α] [LawfulBEq α] (f : αβ) {a : α} {as : List α} (h : a as) :
lookup a (map (fun (x : α) => (x, f x)) as) = some (f a)