doc-next-gen

Init.Data.List.Find

Module docstring

{"Lemmas about List.findSome?, List.find?, List.findIdx, List.findIdx?, List.idxOf, and List.lookup. ","### findSome? ","### find? ","### findIdx? (preliminary lemmas) ","### findIdx ","### findIdx? ","### findFinIdx? ","### idxOf

The verification API for idxOf is still incomplete. The lemmas below should be made consistent with those for findIdx (and proved using them). ","### finIdxOf?

The verification API for finIdxOf? is still incomplete. The lemmas below should be made consistent with those for findFinIdx? (and proved using them). ","### idxOf?

The verification API for idxOf? is still incomplete. The lemmas below should be made consistent with those for findIdx? (and proved using them). ","### lookup ","### Deprecations "}

List.findSome?_cons_of_isSome theorem
{l} (h : (f a).isSome) : findSome? f (a :: l) = f a
Full source
@[simp] theorem findSome?_cons_of_isSome {l} (h : (f a).isSome) : findSome? f (a :: l) = f a := by
  simp only [findSome?]
  split <;> simp_all
First Element Yields Result in `findSome?` When Non-None
Informal description
For any function $f : \alpha \to \text{Option } \beta$, list $l : \text{List } \alpha$, and element $a : \alpha$, if $f(a)$ is of the form $\text{some } x$ (i.e., $\text{Option.isSome } (f a) = \text{true}$), then the result of applying $\text{findSome?}$ to $f$ and the list $a :: l$ is equal to $f(a)$.
List.findSome?_cons_of_isNone theorem
{l} (h : (f a).isNone) : findSome? f (a :: l) = findSome? f l
Full source
@[simp] theorem findSome?_cons_of_isNone {l} (h : (f a).isNone) : findSome? f (a :: l) = findSome? f l := by
  simp only [findSome?]
  split <;> simp_all
`findSome?` on Cons List with None Head Reduces to Tail
Informal description
For any function $f : \alpha \to \text{Option } \beta$, list $l : \text{List } \alpha$, and element $a : \alpha$, if $f(a) = \text{none}$, then the result of applying `findSome?` to $f$ and the list $a :: l$ is equal to the result of applying `findSome?` to $f$ and $l$.
List.exists_of_findSome?_eq_some theorem
{l : List α} {f : α → Option β} (w : l.findSome? f = some b) : ∃ a, a ∈ l ∧ f a = b
Full source
theorem exists_of_findSome?_eq_some {l : List α} {f : α → Option β} (w : l.findSome? f = some b) :
    ∃ a, a ∈ l ∧ f a = b := by
  induction l with
  | nil => simp_all
  | cons h l ih =>
    simp_all only [findSome?_cons, mem_cons, exists_eq_or_imp]
    split at w <;> simp_all
Existence of Element Yielding Some Result in `findSome?`
Informal description
For any list $l$ of elements of type $\alpha$ and any function $f : \alpha \to \text{Option } \beta$, if the result of applying `findSome?` to $f$ and $l$ is `some b`, then there exists an element $a \in l$ such that $f(a) = b$.
List.findSome?_eq_none_iff theorem
: findSome? p l = none ↔ ∀ x ∈ l, p x = none
Full source
@[simp] theorem findSome?_eq_none_iff : findSome?findSome? p l = none ↔ ∀ x ∈ l, p x = none := by
  induction l <;> simp [findSome?_cons]; split <;> simp [*]
Characterization of `findSome?` Returning `none`
Informal description
For any function $p : \alpha \to \text{Option } \beta$ and any list $l : \text{List } \alpha$, the result of `findSome? p l` is `none` if and only if for every element $x$ in $l$, the result of $p(x)$ is `none`.
List.findSome?_eq_none abbrev
Full source
@[deprecated findSome?_eq_none_iff (since := "2024-09-05")] abbrev findSome?_eq_none := @findSome?_eq_none_iff
Characterization of `findSome?` Returning `none`
Informal description
For any function $f : \alpha \to \text{Option } \beta$ and any list $l : \text{List } \alpha$, the result of `findSome? f l` is `none` if and only if for every element $x$ in $l$, the result of $f(x)$ is `none$.
List.findSome?_isSome_iff theorem
{f : α → Option β} {l : List α} : (l.findSome? f).isSome ↔ ∃ x, x ∈ l ∧ (f x).isSome
Full source
@[simp] theorem findSome?_isSome_iff {f : α → Option β} {l : List α} :
    (l.findSome? f).isSome ↔ ∃ x, x ∈ l ∧ (f x).isSome := by
  induction l with
  | nil => simp
  | cons x xs ih =>
    simp only [findSome?_cons]
    split <;> simp_all
Existence of Non-empty Result in `findSome?`
Informal description
For any function $f : \alpha \to \text{Option } \beta$ and any list $l : \text{List } \alpha$, the result of `findSome? f l` is non-empty (i.e., `isSome` returns `true`) if and only if there exists an element $x \in l$ such that $f(x)$ is non-empty (i.e., `isSome` returns `true` for $f(x)$).
List.findSome?_eq_some_iff theorem
{f : α → Option β} {l : List α} {b : β} : l.findSome? f = some b ↔ ∃ l₁ a l₂, l = l₁ ++ a :: l₂ ∧ f a = some b ∧ ∀ x ∈ l₁, f x = none
Full source
theorem findSome?_eq_some_iff {f : α → Option β} {l : List α} {b : β} :
    l.findSome? f = some b ↔ ∃ l₁ a l₂, l = l₁ ++ a :: l₂ ∧ f a = some b ∧ ∀ x ∈ l₁, f x = none := by
  induction l with
  | nil => simp
  | cons p l ih =>
    simp only [findSome?_cons]
    split <;> rename_i b' h
    · simp only [Option.some.injEq, exists_and_right]
      constructor
      · rintro rfl
        exact ⟨[], p, ⟨l, rfl⟩, h, by simp⟩
      · rintro ⟨(⟨⟩ | ⟨p', l₁⟩), a, ⟨l₂, h₁⟩, h₂, h₃⟩
        · simp only [nil_append, cons.injEq] at h₁
          apply Option.some.inj
          simp [← h, ← h₂, h₁.1]
        · simp only [cons_append, cons.injEq] at h₁
          obtain ⟨rfl, rfl⟩ := h₁
          specialize h₃ p
          simp_all
    · rw [ih]
      constructor
      · rintro ⟨l₁, a, l₂, rfl, h₁, h₂⟩
        refine ⟨p :: l₁, a, l₂, rfl, h₁, ?_⟩
        intro a w
        simp at w
        rcases w with rfl | w
        · exact h
        · exact h₂ _ w
      · rintro ⟨l₁, a, l₂, h₁, h₂, h₃⟩
        rcases l₁ with (⟨⟩ | ⟨a', l₁⟩)
        · simp_all
        · simp only [cons_append, cons.injEq] at h₁
          obtain ⟨⟨rfl, rfl⟩, rfl⟩ := h₁
          exact ⟨l₁, a, l₂, rfl, h₂, fun a' w => h₃ a' (mem_cons_of_mem p w)⟩
Characterization of `findSome?` Returning `some b` via List Splitting
Informal description
For a function $f : \alpha \to \text{Option } \beta$, a list $l : \text{List } \alpha$, and an element $b \in \beta$, the following are equivalent: 1. The result of `findSome? f l` is `some b`. 2. There exist sublists $l_1, l_2$ and an element $a \in \alpha$ such that: - $l = l_1 ++ [a] ++ l_2$ (i.e., $l$ can be split into $l_1$, $a$, and $l_2$), - $f(a) = \text{some } b$, - For all $x \in l_1$, $f(x) = \text{none}$.
List.findSome?_guard theorem
{l : List α} : findSome? (Option.guard fun x => p x) l = find? p l
Full source
@[simp] theorem findSome?_guard {l : List α} : findSome? (Option.guard fun x => p x) l = find? p l := by
  induction l with
  | nil => simp
  | cons x xs ih =>
    simp [guard, findSome?, find?]
    split <;> rename_i h
    · simp only [Option.guard_eq_some] at h
      obtain ⟨rfl, h⟩ := h
      simp [h]
    · simp only [Option.guard_eq_none] at h
      simp [ih, h]
Equivalence of `findSome?` with `Option.guard` and `find?` on Lists
Informal description
For any list $l$ of elements of type $\alpha$, the result of applying `findSome?` with the function `Option.guard p` to $l$ is equal to the result of applying `find? p` to $l$. In other words, `findSome? (Option.guard p) l = find? p l`.
List.find?_eq_findSome?_guard theorem
{l : List α} : find? p l = findSome? (Option.guard fun x => p x) l
Full source
theorem find?_eq_findSome?_guard {l : List α} : find? p l = findSome? (Option.guard fun x => p x) l :=
  findSome?_guard.symm
Equivalence of `find?` and `findSome?` with `Option.guard`
Informal description
For any list $l$ of elements of type $\alpha$, the result of applying `find? p` to $l$ is equal to the result of applying `findSome?` with the function `Option.guard p` to $l$. In other words, $\text{find?}\ p\ l = \text{findSome?}\ (\text{Option.guard}\ p)\ l$.
List.head?_filterMap theorem
{f : α → Option β} {l : List α} : (l.filterMap f).head? = l.findSome? f
Full source
@[simp] theorem head?_filterMap {f : α → Option β} {l : List α} : (l.filterMap f).head? = l.findSome? f := by
  induction l with
  | nil => simp
  | cons x xs ih =>
    simp only [filterMap_cons, findSome?_cons]
    split <;> simp [*]
First Element of Filter-Mapped List Equals First Non-None Result
Informal description
For any function $f : \alpha \to \text{Option } \beta$ and any list $l$ of elements of type $\alpha$, the first element of the list obtained by filtering and mapping $l$ with $f$ (if it exists) is equal to the first non-`none` result of applying $f$ to elements of $l$.
List.head_filterMap theorem
{f : α → Option β} {l : List α} (h) : (l.filterMap f).head h = (l.findSome? f).get (by simp_all [Option.isSome_iff_ne_none])
Full source
@[simp] theorem head_filterMap {f : α → Option β} {l : List α} (h) :
    (l.filterMap f).head h = (l.findSome? f).get (by simp_all [Option.isSome_iff_ne_none]) := by
  simp [head_eq_iff_head?_eq_some]
Head of Filter-Mapped List Equals First Non-None Result's Value
Informal description
For any function $f : \alpha \to \text{Option } \beta$ and any list $l$ of elements of type $\alpha$, if the filtered and mapped list $l.\text{filterMap}\ f$ is non-empty (with proof $h$), then the head of this filtered list equals the value obtained from the first non-`none` result of applying $f$ to elements of $l$. In other words, when $l.\text{filterMap}\ f$ is non-empty: $$\text{head}(l.\text{filterMap}\ f) = \text{get}(l.\text{findSome?}\ f)$$ where $\text{get}$ extracts the value from a $\text{some}$ constructor.
List.getLast?_filterMap theorem
{f : α → Option β} {l : List α} : (l.filterMap f).getLast? = l.reverse.findSome? f
Full source
@[simp] theorem getLast?_filterMap {f : α → Option β} {l : List α} : (l.filterMap f).getLast? = l.reverse.findSome? f := by
  rw [getLast?_eq_head?_reverse]
  simp [← filterMap_reverse]
Last Element of Filter-Mapped List Equals First Non-None Result in Reversed List
Informal description
For any function $f : \alpha \to \text{Option } \beta$ and any list $l$ of elements of type $\alpha$, the last element of the list obtained by filtering and mapping $l$ with $f$ (if it exists) is equal to the first non-`none` result of applying $f$ to elements of the reversed list $l^{\text{rev}}$. In other words: $$\text{getLast?}(l.\text{filterMap}\ f) = \text{findSome?}(l^{\text{rev}})\ f$$
List.getLast_filterMap theorem
{f : α → Option β} {l : List α} (h) : (l.filterMap f).getLast h = (l.reverse.findSome? f).get (by simp_all [Option.isSome_iff_ne_none])
Full source
@[simp] theorem getLast_filterMap {f : α → Option β} {l : List α} (h) :
    (l.filterMap f).getLast h = (l.reverse.findSome? f).get (by simp_all [Option.isSome_iff_ne_none]) := by
  simp [getLast_eq_iff_getLast?_eq_some]
Last Element of Filter-Mapped List Equals First Non-None Result in Reversed List
Informal description
For any function $f : \alpha \to \text{Option } \beta$ and any list $l$ of elements of type $\alpha$, if the filtered and mapped list $l.\text{filterMap}\ f$ is non-empty (with proof $h$), then the last element of this filtered list equals the value obtained from the first non-`none` result of applying $f$ to elements of the reversed list $l^{\text{rev}}$. In other words, when $l.\text{filterMap}\ f$ is non-empty: $$\text{getLast}(l.\text{filterMap}\ f) = \text{get}(\text{findSome?}(l^{\text{rev}})\ f)$$ where $\text{get}$ extracts the value from a $\text{some}$ constructor.
List.map_findSome? theorem
{f : α → Option β} {g : β → γ} {l : List α} : (l.findSome? f).map g = l.findSome? (Option.map g ∘ f)
Full source
@[simp] theorem map_findSome? {f : α → Option β} {g : β → γ} {l : List α} :
    (l.findSome? f).map g = l.findSome? (Option.mapOption.map g ∘ f) := by
  induction l <;> simp [findSome?_cons]; split <;> simp [*]
Commutativity of Mapping with First Non-None Result: $(l.\text{findSome? } f).\text{map } g = l.\text{findSome? } (g \circ f)$
Informal description
For any function $f : \alpha \to \text{Option } \beta$, any function $g : \beta \to \gamma$, and any list $l$ of elements of type $\alpha$, the following equality holds: $$(l.\text{findSome? } f).\text{map } g = l.\text{findSome? } (g \circ f)$$ where $\circ$ denotes function composition and $\text{findSome?}$ returns the first non-$\text{none}$ result of applying $f$ to each element of $l$, and $\text{map}$ applies $g$ to the value inside an $\text{Option}$ if it exists.
List.findSome?_map theorem
{f : β → γ} {l : List β} : findSome? p (l.map f) = l.findSome? (p ∘ f)
Full source
theorem findSome?_map {f : β → γ} {l : List β} : findSome? p (l.map f) = l.findSome? (p ∘ f) := by
  induction l with
  | nil => simp
  | cons x xs ih =>
    simp only [map_cons, findSome?]
    split <;> simp_all
Commutativity of `findSome?` with List Mapping: $\text{findSome?}\ p \circ \text{map}\ f = \text{findSome?}\ (p \circ f)$
Informal description
For any function $f : \beta \to \gamma$ and list $l$ of type $\text{List}\ \beta$, the result of applying `findSome? p` to the mapped list $l.map\ f$ is equal to applying `findSome? (p \circ f)` directly to $l$. In other words, the following equality holds: $$\text{findSome?}\ p\ (l.map\ f) = l.\text{findSome?}\ (p \circ f)$$
List.findSome?_append theorem
{l₁ l₂ : List α} : (l₁ ++ l₂).findSome? f = (l₁.findSome? f).or (l₂.findSome? f)
Full source
theorem findSome?_append {l₁ l₂ : List α} : (l₁ ++ l₂).findSome? f = (l₁.findSome? f).or (l₂.findSome? f) := by
  induction l₁ with
  | nil => simp
  | cons x xs ih =>
    simp only [cons_append, findSome?]
    split <;> simp_all
Concatenation Preserves First Non-None Result: $\text{findSome? } f (l_1 ++ l_2) = (\text{findSome? } f l_1) \text{ or } (\text{findSome? } f l_2)$
Informal description
For any function $f : \alpha \to \text{Option } \beta$ and any two lists $l_1$ and $l_2$ of type $\text{List } \alpha$, the result of applying `findSome? f` to the concatenated list $l_1 ++ l_2$ is equal to the first non-`none` result obtained by either applying `findSome? f` to $l_1$ or applying `findSome? f` to $l_2$. In other words: $$\text{findSome? } f (l_1 ++ l_2) = (\text{findSome? } f l_1) \text{ or } (\text{findSome? } f l_2)$$ where $\text{or}$ is the non-short-circuiting choice operation on optional values.
List.head_flatten theorem
{L : List (List α)} (h : ∃ l, l ∈ L ∧ l ≠ []) : (flatten L).head (by simpa using h) = (L.findSome? fun l => l.head?).get (by simpa using h)
Full source
theorem head_flatten {L : List (List α)} (h : ∃ l, l ∈ L ∧ l ≠ []) :
    (flatten L).head (by simpa using h) = (L.findSome? fun l => l.head?).get (by simpa using h) := by
  simp [head_eq_iff_head?_eq_some, head?_flatten]
Head of Flattened List Equals First Non-Empty Sublist's Head
Informal description
For any list of lists $L$ of elements of type $\alpha$, if there exists a non-empty sublist $l \in L$, then the head of the flattened list $\text{flatten}(L)$ is equal to the value obtained from the first non-empty sublist's head via $\text{findSome?}$. More precisely, if there exists $l \in L$ such that $l \neq []$, then: $$\text{head}(\text{flatten}(L)) = \text{get}(\text{findSome?} (\lambda l, \text{head?}(l)) L)$$ where $\text{get}$ extracts the value from a $\text{some}$ constructor.
List.getLast_flatten theorem
{L : List (List α)} (h : ∃ l, l ∈ L ∧ l ≠ []) : (flatten L).getLast (by simpa using h) = (L.reverse.findSome? fun l => l.getLast?).get (by simpa using h)
Full source
theorem getLast_flatten {L : List (List α)} (h : ∃ l, l ∈ L ∧ l ≠ []) :
    (flatten L).getLast (by simpa using h) =
      (L.reverse.findSome? fun l => l.getLast?).get (by simpa using h) := by
  simp [getLast_eq_iff_getLast?_eq_some, getLast?_flatten]
Last Element of Flattened List Equals First Non-Empty Last Element in Reversed List
Informal description
For any list of lists $L$ of type $\text{List}(\text{List } \alpha)$, if there exists a non-empty list $l \in L$, then the last element of the flattened list $\text{flatten } L$ is equal to the value obtained by finding the first non-empty list in the reverse of $L$ and taking its last element. More precisely, if $h$ is a proof that there exists a non-empty list $l \in L$, then: $$\text{getLast}(\text{flatten } L, h) = \text{get}(\text{findSome? } (\lambda l \mapsto \text{getLast?}(l)) (\text{reverse } L), h')$$ where $h'$ is derived from $h$ and ensures the $\text{findSome?}$ operation succeeds.
List.findSome?_replicate theorem
: findSome? f (replicate n a) = if n = 0 then none else f a
Full source
theorem findSome?_replicate : findSome? f (replicate n a) = if n = 0 then none else f a := by
  cases n with
  | zero => simp
  | succ n =>
    simp only [replicate_succ, findSome?_cons]
    split <;> simp_all
Behavior of `findSome?` on Replicated Lists: $\text{findSome? } f (\text{replicate } n a) = \text{if } n = 0 \text{ then none else } f(a)$
Informal description
For any function $f : \alpha \to \text{Option } \beta$, element $a \in \alpha$, and natural number $n$, the result of applying `findSome?` to $f$ and a list of $n$ copies of $a$ is equal to `none` if $n = 0$, and equal to $f(a)$ otherwise. That is: $$\text{findSome? } f (\text{replicate } n a) = \begin{cases} \text{none} & \text{if } n = 0 \\ f(a) & \text{otherwise} \end{cases}$$
List.findSome?_replicate_of_pos theorem
(h : 0 < n) : findSome? f (replicate n a) = f a
Full source
@[simp] theorem findSome?_replicate_of_pos (h : 0 < n) : findSome? f (replicate n a) = f a := by
  simp [findSome?_replicate, Nat.ne_of_gt h]
`findSome?` on Non-Empty Replicated List Equals Function Application
Informal description
For any function $f : \alpha \to \text{Option } \beta$, element $a \in \alpha$, and natural number $n > 0$, the result of applying `findSome?` to $f$ and a list of $n$ copies of $a$ is equal to $f(a)$. That is: $$\text{findSome? } f (\text{replicate } n a) = f(a) \quad \text{when } n > 0$$
List.findSome?_replicate_of_isSome theorem
(_ : (f a).isSome) : findSome? f (replicate n a) = if n = 0 then none else f a
Full source
@[simp] theorem findSome?_replicate_of_isSome (_ : (f a).isSome) : findSome? f (replicate n a) = if n = 0 then none else f a := by
  simp [findSome?_replicate]
Behavior of `findSome?` on Replicated Lists When $f(a)$ is Defined: $\text{findSome? } f (\text{replicate } n a) = \text{if } n = 0 \text{ then none else } f(a)$
Informal description
For any function $f : \alpha \to \text{Option } \beta$, element $a \in \alpha$, and natural number $n$, if $f(a)$ is of the form $\text{some } x$ (i.e., $(f a).\text{isSome}$ holds), then the result of applying `findSome?` to $f$ and a list of $n$ copies of $a$ is equal to $\text{none}$ if $n = 0$, and equal to $f(a)$ otherwise. That is: $$\text{findSome? } f (\text{replicate } n a) = \begin{cases} \text{none} & \text{if } n = 0 \\ f(a) & \text{otherwise} \end{cases}$$
List.findSome?_replicate_of_isNone theorem
(h : (f a).isNone) : findSome? f (replicate n a) = none
Full source
@[simp] theorem findSome?_replicate_of_isNone (h : (f a).isNone) : findSome? f (replicate n a) = none := by
  rw [Option.isNone_iff_eq_none] at h
  simp [findSome?_replicate, h]
`findSome?` on Replicated List Yields `none` When Function Does
Informal description
For any function $f : \alpha \to \text{Option } \beta$, element $a \in \alpha$, and natural number $n$, if $f(a)$ is `none`, then applying `findSome?` to $f$ and a list of $n$ copies of $a$ yields `none`. That is: $$(f(a) = \text{none}) \implies \text{findSome? } f (\text{replicate } n a) = \text{none}$$
List.Sublist.findSome?_isSome theorem
{l₁ l₂ : List α} (h : l₁ <+ l₂) : (l₁.findSome? f).isSome → (l₂.findSome? f).isSome
Full source
theorem Sublist.findSome?_isSome {l₁ l₂ : List α} (h : l₁ <+ l₂) :
    (l₁.findSome? f).isSome → (l₂.findSome? f).isSome := by
  induction h with
  | slnil => simp
  | cons a h ih
  | cons₂ a h ih =>
    simp only [findSome?]
    split
    · simp_all
    · exact ih
Sublist preserves non-`none` results in `findSome?`
Informal description
For any two lists $l_1$ and $l_2$ of type $\alpha$, if $l_1$ is a sublist of $l_2$ (denoted $l_1 <+ l_2$) and the function `findSome? f` applied to $l_1$ returns a `some` value (i.e., `(l₁.findSome? f).isSome` holds), then `findSome? f` applied to $l_2$ also returns a `some` value (i.e., `(l₂.findSome? f).isSome` holds).
List.Sublist.findSome?_eq_none theorem
{l₁ l₂ : List α} (h : l₁ <+ l₂) : l₂.findSome? f = none → l₁.findSome? f = none
Full source
theorem Sublist.findSome?_eq_none {l₁ l₂ : List α} (h : l₁ <+ l₂) :
    l₂.findSome? f = none → l₁.findSome? f = none := by
  simp only [List.findSome?_eq_none_iff, Bool.not_eq_true]
  exact fun w x m => w x (Sublist.mem m h)
Sublist preserves `none` results in `findSome?`
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a sublist of $l_2$ (denoted $l_1 <+ l_2$) and the function `findSome? f` applied to $l_2$ returns `none`, then `findSome? f` applied to $l_1$ also returns `none$.
List.IsPrefix.findSome?_eq_some theorem
{l₁ l₂ : List α} {f : α → Option β} (h : l₁ <+: l₂) : List.findSome? f l₁ = some b → List.findSome? f l₂ = some b
Full source
theorem IsPrefix.findSome?_eq_some {l₁ l₂ : List α} {f : α → Option β} (h : l₁ <+: l₂) :
    List.findSome? f l₁ = some b → List.findSome? f l₂ = some b := by
  rw [IsPrefix] at h
  obtain ⟨t, rfl⟩ := h
  simp +contextual [findSome?_append]
Prefix Preservation of First Non-None Result: $\text{findSome? } f l_1 = \text{some } b \implies \text{findSome? } f l_2 = \text{some } b$ when $l_1$ is prefix of $l_2$
Informal description
For any two lists $l_1$ and $l_2$ of type $\alpha$, if $l_1$ is a prefix of $l_2$ (denoted $l_1 <+: l_2$) and applying the function `findSome? f` to $l_1$ yields `some b`, then applying `findSome? f` to $l_2$ also yields `some b$. Here, `findSome? f l` is the first non-`none` result of applying $f : \alpha \to \text{Option } \beta$ to elements of list $l$ in order.
List.IsPrefix.findSome?_eq_none theorem
{l₁ l₂ : List α} {f : α → Option β} (h : l₁ <+: l₂) : List.findSome? f l₂ = none → List.findSome? f l₁ = none
Full source
theorem IsPrefix.findSome?_eq_none {l₁ l₂ : List α} {f : α → Option β} (h : l₁ <+: l₂) :
    List.findSome? f l₂ = noneList.findSome? f l₁ = none :=
  h.sublist.findSome?_eq_none
Prefix Preservation of `none` Results in `findSome?`
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a prefix of $l_2$ (denoted $l_1 <+: l_2$) and applying the function $\text{findSome? } f$ to $l_2$ yields $\text{none}$, then applying $\text{findSome? } f$ to $l_1$ also yields $\text{none}$. Here, $\text{findSome? } f l$ denotes the first non-$\text{none}$ result obtained by applying the function $f : \alpha \to \text{Option } \beta$ to each element of the list $l$ in order.
List.IsSuffix.findSome?_eq_none theorem
{l₁ l₂ : List α} {f : α → Option β} (h : l₁ <:+ l₂) : List.findSome? f l₂ = none → List.findSome? f l₁ = none
Full source
theorem IsSuffix.findSome?_eq_none {l₁ l₂ : List α} {f : α → Option β} (h : l₁ <:+ l₂) :
    List.findSome? f l₂ = noneList.findSome? f l₁ = none :=
  h.sublist.findSome?_eq_none
Suffix Preservation of `none` in `findSome?`
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a suffix of $l_2$ (denoted $l_1 <:+ l_2$) and the function `findSome? f` applied to $l_2$ returns `none`, then `findSome? f` applied to $l_1$ also returns `none$. Here, `findSome? f l` is the first non-`none` result of applying $f : \alpha \to \text{Option } \beta$ to elements of list $l$ in order.
List.IsInfix.findSome?_eq_none theorem
{l₁ l₂ : List α} {f : α → Option β} (h : l₁ <:+: l₂) : List.findSome? f l₂ = none → List.findSome? f l₁ = none
Full source
theorem IsInfix.findSome?_eq_none {l₁ l₂ : List α} {f : α → Option β} (h : l₁ <:+: l₂) :
    List.findSome? f l₂ = noneList.findSome? f l₁ = none :=
  h.sublist.findSome?_eq_none
Infix Preservation of `none` in `findSome?`
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is an infix of $l_2$ (i.e., $l_1$ appears as a contiguous subsequence in $l_2$) and the function `findSome? f` applied to $l_2$ returns `none`, then `findSome? f` applied to $l_1$ also returns `none$.
List.find?_singleton theorem
{a : α} {p : α → Bool} : [a].find? p = if p a then some a else none
Full source
@[simp] theorem find?_singleton {a : α} {p : α → Bool} : [a].find? p = if p a then some a else none := by
  simp only [find?]
  split <;> simp_all
Finding in Singleton List: $\text{find? } p [a] = \text{if } p(a) \text{ then some } a \text{ else none}$
Informal description
For any element $a$ of type $\alpha$ and any predicate $p : \alpha \to \text{Bool}$, the result of applying the `find?` function to the singleton list $[a]$ with predicate $p$ is equal to $\text{some } a$ if $p(a)$ holds, and $\text{none}$ otherwise. In other words, $\text{find? } p [a] = \begin{cases} \text{some } a & \text{if } p(a) \\ \text{none} & \text{otherwise} \end{cases}$.
List.find?_cons_of_pos theorem
{l} (h : p a) : find? p (a :: l) = some a
Full source
@[simp] theorem find?_cons_of_pos {l} (h : p a) : find? p (a :: l) = some a := by
  simp [find?, h]
First Satisfying Element in Cons List When Head Satisfies Predicate
Informal description
For any predicate $p : \alpha \to \text{Bool}$ and any list $l : \text{List } \alpha$, if $p(a)$ holds for an element $a$, then the first element satisfying $p$ in the list $a :: l$ is $\text{some } a$.
List.find?_cons_of_neg theorem
{l} (h : ¬p a) : find? p (a :: l) = find? p l
Full source
@[simp] theorem find?_cons_of_neg {l} (h : ¬p a) : find? p (a :: l) = find? p l := by
  simp [find?, h]
First Element Satisfying Predicate is Unchanged When Head Fails: $\text{find? } p (a :: l) = \text{find? } p l$ given $\neg p(a)$
Informal description
For any predicate $p : \alpha \to \text{Bool}$, element $a : \alpha$, and list $l : \text{List } \alpha$, if $p(a)$ is false, then the first element satisfying $p$ in the list $a :: l$ is equal to the first element satisfying $p$ in the list $l$. In other words, $\text{find? } p (a :: l) = \text{find? } p l$ when $\neg p(a)$.
List.find?_eq_none theorem
: find? p l = none ↔ ∀ x ∈ l, ¬p x
Full source
@[simp] theorem find?_eq_none : find?find? p l = none ↔ ∀ x ∈ l, ¬ p x := by
  induction l <;> simp [find?_cons]; split <;> simp [*]
Characterization of `find?` Returning `none` as Absence of Satisfying Elements
Informal description
For any predicate $p : \alpha \to \text{Bool}$ and list $l : \text{List } \alpha$, the function `find? p l` returns `none` if and only if no element $x$ in $l$ satisfies $p(x)$. In other words, $\text{find? } p l = \text{none} \leftrightarrow \forall x \in l, \neg p(x)$.
List.find?_eq_some_iff_append theorem
: xs.find? p = some b ↔ p b ∧ ∃ as bs, xs = as ++ b :: bs ∧ ∀ a ∈ as, !p a
Full source
theorem find?_eq_some_iff_append :
    xs.find? p = some b ↔ p b ∧ ∃ as bs, xs = as ++ b :: bs ∧ ∀ a ∈ as, !p a := by
  induction xs with
  | nil => simp
  | cons x xs ih =>
    simp only [find?_cons, exists_and_right]
    split <;> rename_i h
    · simp only [Option.some.injEq]
      constructor
      · rintro rfl
        exact ⟨h, [], ⟨xs, rfl⟩, by simp⟩
      · rintro ⟨-, ⟨as, ⟨⟨bs, h₁⟩, h₂⟩⟩⟩
        cases as with
        | nil => simp_all
        | cons a as =>
          specialize h₂ a mem_cons_self
          simp only [cons_append] at h₁
          obtain ⟨rfl, -⟩ := h₁
          simp_all
    · simp only [ih, Bool.not_eq_eq_eq_not, Bool.not_true, exists_and_right, and_congr_right_iff]
      intro pb
      constructor
      · rintro ⟨as, ⟨⟨bs, rfl⟩, h₁⟩⟩
        refine ⟨x :: as, ⟨⟨bs, rfl⟩, ?_⟩⟩
        intro a m
        simp at m
        obtain (rfl|m) := m
        · exact h
        · exact h₁ a m
      · rintro ⟨as, ⟨bs, h₁⟩, h₂⟩
        cases as with
        | nil => simp_all
        | cons a as =>
          refine ⟨as, ⟨⟨bs, ?_⟩, fun a m => h₂ a (mem_cons_of_mem _ m)⟩⟩
          cases h₁
          simp
Characterization of `find?` Result via List Splitting: `find? p xs = some b ↔ p b ∧ ∃ as bs, xs = as ++ b :: bs ∧ ∀ a ∈ as, ¬p a`
Informal description
For a list `xs`, a predicate `p : α → Bool`, and an element `b : α`, the following are equivalent: 1. The function `find? p xs` returns `some b`. 2. The predicate `p` holds at `b`, and there exist sublists `as` and `bs` such that `xs = as ++ b :: bs` and for every element `a` in `as`, the predicate `p` does not hold at `a`. In other words, `find? p xs` returns `some b` if and only if `b` is the first element in `xs` satisfying `p`, and `xs` can be split into a prefix `as` (where no element satisfies `p`), followed by `b`, followed by a suffix `bs`.
List.find?_eq_some abbrev
Full source
@[deprecated find?_eq_some_iff_append (since := "2024-11-06")]
abbrev find?_eq_some := @find?_eq_some_iff_append
Characterization of `find?` Result via First Satisfying Element: `find? p xs = some b ↔ b ∈ xs ∧ p b ∧ ∀ a ∈ xs before b, ¬p a`
Informal description
For a list `xs : List α`, a predicate `p : α → Bool`, and an element `b : α`, the following are equivalent: 1. The function `find? p xs` returns `some b`. 2. The element `b` is in `xs`, the predicate `p` holds at `b`, and for every element `a` that appears before `b` in `xs`, the predicate `p` does not hold at `a`. In other words, `find? p xs = some b` if and only if `b` is the first element in `xs` that satisfies `p`.
List.find?_cons_eq_some theorem
: (a :: xs).find? p = some b ↔ (p a ∧ a = b) ∨ (!p a ∧ xs.find? p = some b)
Full source
@[simp]
theorem find?_cons_eq_some : (a :: xs).find? p = some b ↔ (p a ∧ a = b) ∨ (!p a ∧ xs.find? p = some b) := by
  rw [find?_cons]
  split <;> simp_all
Characterization of List.find? on Cons with Some Result
Informal description
For a list of the form $a :: xs$ and a predicate $p : \alpha \to \text{Bool}$, the following equivalence holds: $(a :: xs).\text{find?}\ p = \text{some}\ b \leftrightarrow (p(a) \land a = b) \lor (\neg p(a) \land xs.\text{find?}\ p = \text{some}\ b)$.
List.find?_isSome theorem
{xs : List α} {p : α → Bool} : (xs.find? p).isSome ↔ ∃ x, x ∈ xs ∧ p x
Full source
@[simp] theorem find?_isSome {xs : List α} {p : α → Bool} : (xs.find? p).isSome ↔ ∃ x, x ∈ xs ∧ p x := by
  induction xs with
  | nil => simp
  | cons x xs ih =>
    simp only [find?_cons, mem_cons, exists_eq_or_imp]
    split <;> simp_all
Existence of Satisfying Element in List via `find?`
Informal description
For any list `xs` of elements of type `α` and any predicate `p : α → Bool`, the function `find? p xs` returns a `some` value if and only if there exists an element `x` in `xs` such that `p x` holds.
List.find?_some theorem
: ∀ {l}, find? p l = some a → p a
Full source
theorem find?_some : ∀ {l}, find? p l = some a → p a
  | b :: l, H => by
    by_cases h : p b <;> simp [find?, h] at H
    · exact H ▸ h
    · exact find?_some H
Predicate Holds for Found Element in List
Informal description
For any list $l$ and predicate $p : \alpha \to \text{Bool}$, if $\text{find?}\ p\ l$ returns $\text{some}\ a$, then $p(a)$ holds.
List.mem_of_find?_eq_some theorem
: ∀ {l}, find? p l = some a → a ∈ l
Full source
theorem mem_of_find?_eq_some : ∀ {l}, find? p l = some a → a ∈ l
  | b :: l, H => by
    by_cases h : p b <;> simp [find?, h] at H
    · exact H ▸ .head _
    · exact .tail _ (mem_of_find?_eq_some H)
Membership from Successful `find?` Operation
Informal description
For any list $l$ of elements of type $\alpha$ and any predicate $p : \alpha \to \text{Bool}$, if the `find?` function applied to $p$ and $l$ returns `some a`, then the element $a$ is a member of the list $l$.
List.get_find?_mem theorem
{xs : List α} {p : α → Bool} (h) : (xs.find? p).get h ∈ xs
Full source
theorem get_find?_mem {xs : List α} {p : α → Bool} (h) : (xs.find? p).get h ∈ xs := by
  induction xs with
  | nil => simp at h
  | cons x xs ih =>
    simp only [find?_cons]
    by_cases h : p x
    · simp [h]
    · simp only [h]
      right
      apply ih
Membership of Extracted Element from `find?` in Original List
Informal description
For any list `xs` of elements of type `α` and any predicate `p : α → Bool`, if `h` is a proof that `xs.find? p` is of the form `some x`, then the element `x` obtained via `Option.get` is a member of the list `xs`. In other words, the element found by `find?` and extracted with `get` is guaranteed to be in the original list.
List.find?_filter theorem
{xs : List α} {p : α → Bool} {q : α → Bool} : (xs.filter p).find? q = xs.find? (fun a => p a ∧ q a)
Full source
@[simp] theorem find?_filter {xs : List α} {p : α → Bool} {q : α → Bool} :
    (xs.filter p).find? q = xs.find? (fun a => p a ∧ q a) := by
  induction xs with
  | nil => simp
  | cons x xs ih =>
    simp only [filter_cons]
    split <;>
    · simp only [find?_cons]
      split <;> simp_all
Equivalence of Filtered Find and Conjunctive Find
Informal description
For any list $xs$ of elements of type $\alpha$ and predicates $p, q : \alpha \to \text{Bool}$, the first element in the filtered list $xs.filter\ p$ that satisfies $q$ is equal to the first element in $xs$ that satisfies both $p$ and $q$ simultaneously. That is, $$ \text{find?}\ q\ (xs.filter\ p) = \text{find?}\ (\lambda a, p\ a \land q\ a)\ xs. $$
List.head?_filter theorem
{p : α → Bool} {l : List α} : (l.filter p).head? = l.find? p
Full source
@[simp] theorem head?_filter {p : α → Bool} {l : List α} : (l.filter p).head? = l.find? p := by
  rw [← filterMap_eq_filter, head?_filterMap, findSome?_guard]
First Element of Filtered List Equals First Satisfying Element
Informal description
For any predicate $p : \alpha \to \text{Bool}$ and any list $l$ of elements of type $\alpha$, the first element of the filtered list $l.\text{filter}\ p$ (if it exists) is equal to the first element of $l$ that satisfies $p$.
List.head_filter theorem
{p : α → Bool} {l : List α} (h) : (l.filter p).head h = (l.find? p).get (by simp_all [Option.isSome_iff_ne_none])
Full source
@[simp] theorem head_filter {p : α → Bool} {l : List α} (h) :
    (l.filter p).head h = (l.find? p).get (by simp_all [Option.isSome_iff_ne_none]) := by
  simp [head_eq_iff_head?_eq_some]
First Element of Non-Empty Filtered List Equals First Satisfying Element
Informal description
For any predicate $p : \alpha \to \text{Bool}$ and any list $l$ of elements of type $\alpha$, if the filtered list $l.\text{filter}\ p$ is non-empty, then its first element is equal to the first element of $l$ that satisfies $p$ (which must exist under this condition).
List.getLast?_filter theorem
{p : α → Bool} {l : List α} : (l.filter p).getLast? = l.reverse.find? p
Full source
@[simp] theorem getLast?_filter {p : α → Bool} {l : List α} : (l.filter p).getLast? = l.reverse.find? p := by
  rw [getLast?_eq_head?_reverse]
  simp [← filter_reverse]
Last Element of Filtered List Equals First Satisfying Element in Reversed List
Informal description
For any predicate $p : \alpha \to \text{Bool}$ and any list $l$ of elements of type $\alpha$, the last element of the filtered list $l.\text{filter}\ p$ (if it exists) is equal to the first element of the reversed list $l^{\text{rev}}$ that satisfies $p$.
List.getLast_filter theorem
{p : α → Bool} {l : List α} (h) : (l.filter p).getLast h = (l.reverse.find? p).get (by simp_all [Option.isSome_iff_ne_none])
Full source
@[simp] theorem getLast_filter {p : α → Bool} {l : List α} (h) :
    (l.filter p).getLast h = (l.reverse.find? p).get (by simp_all [Option.isSome_iff_ne_none]) := by
  simp [getLast_eq_iff_getLast?_eq_some]
Last Element of Non-Empty Filtered List Equals First Satisfying Element in Reversed List
Informal description
For any predicate $p : \alpha \to \text{Bool}$ and any non-empty list $l$ of elements of type $\alpha$, the last element of the filtered list $l.\text{filter}\ p$ is equal to the first element of the reversed list $l^{\text{rev}}$ that satisfies $p$.
List.find?_filterMap theorem
{xs : List α} {f : α → Option β} {p : β → Bool} : (xs.filterMap f).find? p = (xs.find? (fun a => (f a).any p)).bind f
Full source
@[simp] theorem find?_filterMap {xs : List α} {f : α → Option β} {p : β → Bool} :
    (xs.filterMap f).find? p = (xs.find? (fun a => (f a).any p)).bind f := by
  induction xs with
  | nil => simp
  | cons x xs ih =>
    simp only [filterMap_cons]
    split <;>
    · simp only [find?_cons]
      split <;> simp_all
Commutativity of `find?` and `filterMap` with Predicate Filtering
Informal description
For any list `xs` of elements of type `α`, any function `f : α → Option β`, and any predicate `p : β → Bool`, the following equality holds: \[ \text{find? } p \ (\text{filterMap } f \ xs) = \text{bind } (\text{find? } (\lambda a, \text{any } p \ (f \ a)) \ xs) \ f \] Here, `filterMap f xs` applies `f` to each element of `xs` and collects the `some` results, `find? p l` finds the first element in list `l` satisfying `p`, and `bind` is the monadic bind operation for `Option`.
List.find?_map theorem
{f : β → α} {l : List β} : find? p (l.map f) = (l.find? (p ∘ f)).map f
Full source
@[simp] theorem find?_map {f : β → α} {l : List β} : find? p (l.map f) = (l.find? (p ∘ f)).map f := by
  induction l with
  | nil => simp
  | cons x xs ih =>
    simp only [map_cons, find?]
    by_cases h : p (f x) <;> simp [h, ih]
Commutativity of `find?` and `map` with Predicate Composition
Informal description
For any function $f : \beta \to \alpha$, any list $l$ of elements of type $\beta$, and any predicate $p : \alpha \to \text{Bool}$, the first element satisfying $p$ in the mapped list $\text{map } f \ l$ is equal to the result of applying $f$ to the first element in $l$ that satisfies $p \circ f$. Formally: \[ \text{find? } p \ (\text{map } f \ l) = \text{map } f \ (\text{find? } (p \circ f) \ l) \]
List.find?_append theorem
{l₁ l₂ : List α} : (l₁ ++ l₂).find? p = (l₁.find? p).or (l₂.find? p)
Full source
@[simp] theorem find?_append {l₁ l₂ : List α} : (l₁ ++ l₂).find? p = (l₁.find? p).or (l₂.find? p) := by
  induction l₁ with
  | nil => simp
  | cons x xs ih =>
    simp only [cons_append, find?]
    by_cases h : p x <;> simp [h, ih]
First Satisfying Element in Concatenated List Equals First in Either List
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, and any predicate $p : \alpha \to \text{Bool}$, the first element satisfying $p$ in the concatenated list $l_1 ++ l_2$ is equal to the first element satisfying $p$ in $l_1$ if it exists, otherwise it's the first element satisfying $p$ in $l_2$. Formally: \[ \text{find? } p \ (l_1 ++ l_2) = (\text{find? } p \ l_1) \text{ or } (\text{find? } p \ l_2) \]
List.find?_flatten theorem
{xss : List (List α)} {p : α → Bool} : xss.flatten.find? p = xss.findSome? (·.find? p)
Full source
@[simp] theorem find?_flatten {xss : List (List α)} {p : α → Bool} :
    xss.flatten.find? p = xss.findSome? (·.find? p) := by
  induction xss with
  | nil => simp
  | cons _ _ ih =>
    simp only [flatten_cons, find?_append, findSome?_cons, ih]
    split <;> simp [*]
First Satisfying Element in Flattened List Equals First Non-None Result of Sublist Searches
Informal description
For any list of lists `xss : List (List α)` and any predicate `p : α → Bool`, the first element satisfying `p` in the flattened list `xss.flatten` is equal to the first non-`none` result obtained by applying `find? p` to each sublist in `xss`. Formally: \[ \text{find? } p \ (\text{flatten } xss) = \text{findSome? } (\lambda ys \Rightarrow \text{find? } p \ ys) \ xss \]
List.find?_flatten_eq_none_iff theorem
{xs : List (List α)} {p : α → Bool} : xs.flatten.find? p = none ↔ ∀ ys ∈ xs, ∀ x ∈ ys, !p x
Full source
theorem find?_flatten_eq_none_iff {xs : List (List α)} {p : α → Bool} :
    xs.flatten.find? p = none ↔ ∀ ys ∈ xs, ∀ x ∈ ys, !p x := by
  simp
Characterization of No Satisfying Elements in Flattened List
Informal description
For any list of lists `xs : List (List α)` and any predicate `p : α → Bool`, the flattened list `xs.flatten` has no elements satisfying `p` if and only if for every sublist `ys` in `xs` and every element `x` in `ys`, the predicate `p x` evaluates to `false`. Formally: \[ \text{find? } p \ (\text{flatten } xs) = \text{none} \leftrightarrow \forall ys \in xs, \forall x \in ys, \neg p(x) \]
List.find?_flatten_eq_none abbrev
Full source
@[deprecated find?_flatten_eq_none_iff (since := "2025-02-03")]
abbrev find?_flatten_eq_none := @find?_flatten_eq_none_iff
Characterization of No Satisfying Elements in Flattened List
Informal description
For any list of lists `xs : List (List α)` and any predicate `p : α → Bool`, the flattened list `xs.flatten` has no elements satisfying `p` if and only if for every sublist `ys` in `xs` and every element `x` in `ys`, the predicate `p x` evaluates to `false`. Formally: \[ \text{find? } p \ (\text{flatten } xs) = \text{none} \leftrightarrow \forall ys \in xs, \forall x \in ys, \neg p(x) \]
List.find?_flatten_eq_some_iff theorem
{xs : List (List α)} {p : α → Bool} {a : α} : xs.flatten.find? p = some a ↔ p a ∧ ∃ as ys zs bs, xs = as ++ (ys ++ a :: zs) :: bs ∧ (∀ l ∈ as, ∀ x ∈ l, !p x) ∧ (∀ x ∈ ys, !p x)
Full source
/--
If `find? p` returns `some a` from `xs.flatten`, then `p a` holds, and
some list in `xs` contains `a`, and no earlier element of that list satisfies `p`.
Moreover, no earlier list in `xs` has an element satisfying `p`.
-/
theorem find?_flatten_eq_some_iff {xs : List (List α)} {p : α → Bool} {a : α} :
    xs.flatten.find? p = some a ↔
      p a ∧ ∃ as ys zs bs, xs = as ++ (ys ++ a :: zs) :: bs ∧
        (∀ l ∈ as, ∀ x ∈ l, !p x) ∧ (∀ x ∈ ys, !p x) := by
  rw [find?_eq_some_iff_append]
  constructor
  · rintro ⟨h, ⟨ys, zs, h₁, h₂⟩⟩
    refine ⟨h, ?_⟩
    rw [flatten_eq_append_iff] at h₁
    obtain (⟨as, bs, rfl, rfl, h₁⟩ | ⟨as, bs, c, cs, ds, rfl, rfl, h₁⟩) := h₁
    · replace h₁ := h₁.symm
      rw [flatten_eq_cons_iff] at h₁
      obtain ⟨bs, cs, ds, rfl, h₁, rfl⟩ := h₁
      refine ⟨as ++ bs, [], cs, ds, by simp, ?_⟩
      simp
      rintro l (ma | mb) x m
      · simpa using h₂ x (by simpa using ⟨l, ma, m⟩)
      · specialize h₁ _ mb
        simp_all
    · simp [h₁]
      refine ⟨as, bs, ?_⟩
      refine ⟨?_, ?_, ?_⟩
      · simp_all
      · intro l ml a m
        simpa using h₂ a (by simpa using .inl ⟨l, ml, m⟩)
      · intro x m
        simpa using h₂ x (by simpa using .inr m)
  · rintro ⟨h, ⟨as, ys, zs, bs, rfl, h₁, h₂⟩⟩
    refine ⟨h, as.flatten ++ ys, zs ++ bs.flatten, by simp, ?_⟩
    intro a m
    simp at m
    obtain ⟨l, ml, m⟩ | m := m
    · exact h₁ l ml a m
    · exact h₂ a m
Characterization of First Satisfying Element in Flattened List: `find? p (flatten xs) = some a`
Informal description
For a list of lists `xs : List (List α)`, a predicate `p : α → Bool`, and an element `a : α`, the following are equivalent: 1. The function `find? p` applied to the flattened list `xs.flatten` returns `some a`. 2. The predicate `p a` holds, and there exist sublists `as`, `ys`, `zs`, `bs` such that: - `xs` can be written as `as ++ (ys ++ a :: zs) :: bs`, - For every list `l` in `as` and every element `x` in `l`, `¬p x` holds, - For every element `x` in `ys`, `¬p x` holds.
List.find?_flatten_eq_some abbrev
Full source
@[deprecated find?_flatten_eq_some_iff (since := "2025-02-03")]
abbrev find?_flatten_eq_some := @find?_flatten_eq_some_iff
Implication of `find? p (flatten xs) = some a` for List Structure and Predicate Satisfaction
Informal description
For a list of lists `xs : List (List α)`, a predicate `p : α → Bool`, and an element `a : α`, if the function `find? p` applied to the flattened list `xs.flatten` returns `some a`, then: 1. The predicate `p a` holds, and 2. There exist sublists `as`, `ys`, `zs`, `bs` such that: - `xs` can be written as `as ++ (ys ++ a :: zs) :: bs`, - For every list `l` in `as` and every element `x` in `l`, `¬p x` holds, - For every element `x` in `ys`, `¬p x` holds.
List.find?_flatMap theorem
{xs : List α} {f : α → List β} {p : β → Bool} : (xs.flatMap f).find? p = xs.findSome? (fun x => (f x).find? p)
Full source
@[simp] theorem find?_flatMap {xs : List α} {f : α → List β} {p : β → Bool} :
    (xs.flatMap f).find? p = xs.findSome? (fun x => (f x).find? p) := by
  simp [flatMap_def, findSome?_map]; rfl
First Satisfying Element in FlatMap Equals First Non-None Result of Sublist Searches
Informal description
For any list `xs : List α`, function `f : α → List β`, and predicate `p : β → Bool`, the first element satisfying `p` in the flattened list obtained by applying `f` to each element of `xs` is equal to the first non-`none` result of applying `find? p` to each sublist `f x` for `x ∈ xs`. Formally: \[ \text{find? } p \ (\text{flatMap } f \ xs) = \text{findSome? } (\lambda x \Rightarrow \text{find? } p \ (f x)) \ xs \]
List.find?_bind abbrev
Full source
@[deprecated find?_flatMap (since := "2024-10-16")] abbrev find?_bind := @find?_flatMap
First Satisfying Element in Bind Equals First Non-None Result of Sublist Searches
Informal description
Given a list `xs : List α`, a function `f : α → List β`, and a predicate `p : β → Bool`, the first element satisfying `p` in the list obtained by applying `f` to each element of `xs` and concatenating the results is equal to the first non-`none` result of applying `find? p` to each sublist `f x` for `x ∈ xs`. Formally: \[ \text{find? } p \ (\text{bind } f \ xs) = \text{findSome? } (\lambda x \Rightarrow \text{find? } p \ (f x)) \ xs \]
List.find?_flatMap_eq_none_iff theorem
{xs : List α} {f : α → List β} {p : β → Bool} : (xs.flatMap f).find? p = none ↔ ∀ x ∈ xs, ∀ y ∈ f x, !p y
Full source
theorem find?_flatMap_eq_none_iff {xs : List α} {f : α → List β} {p : β → Bool} :
    (xs.flatMap f).find? p = none ↔ ∀ x ∈ xs, ∀ y ∈ f x, !p y := by
  simp
Equivalence of FlatMap Find Failure to Universal Predicate Falsity
Informal description
For any list `xs` of type `List α`, function `f : α → List β`, and predicate `p : β → Bool`, the following equivalence holds: \[ \text{find? } p \ (\text{flatMap } f \ xs) = \text{none} \leftrightarrow \forall x \in xs, \forall y \in f x, \neg p(y) \] This means that the `find?` operation on the flattened list returns `none` if and only if for every element `x` in `xs` and every element `y` in `f x`, the predicate `p(y)` evaluates to `false`.
List.find?_flatMap_eq_none abbrev
Full source
@[deprecated find?_flatMap_eq_none_iff (since := "2024-10-16")]
abbrev find?_flatMap_eq_none := @find?_flatMap_eq_none_iff
Failure of `find?` on FlatMapped List Equivalent to Universal Predicate Falsity
Informal description
For any list `xs` of type `List α`, function `f : α → List β`, and predicate `p : β → Bool`, the `find?` operation on the flattened list `flatMap f xs` returns `none` if and only if for every element `x` in `xs` and every element `y` in `f x`, the predicate `p(y)` evaluates to `false`. That is: \[ \text{find? } p \ (\text{flatMap } f \ xs) = \text{none} \leftrightarrow \forall x \in xs, \forall y \in f x, \neg p(y) \]
List.find?_bind_eq_none abbrev
Full source
@[deprecated find?_flatMap_eq_none (since := "2024-10-16")] abbrev find?_bind_eq_none := @find?_flatMap_eq_none_iff
Equivalence of Bind Find Failure to Universal Predicate Falsity
Informal description
For any list `xs` of type `List α`, function `f : α → Option β`, and predicate `p : β → Bool`, the following equivalence holds: \[ \text{find? } p \ (\text{bind } f \ xs) = \text{none} \leftrightarrow \forall x \in xs, \forall y \in f x, \neg p(y) \] where `bind f xs` is the monadic bind operation that applies `f` to each element of `xs` and collects the `some` results.
List.find?_replicate theorem
: find? p (replicate n a) = if n = 0 then none else if p a then some a else none
Full source
theorem find?_replicate : find? p (replicate n a) = if n = 0 then none else if p a then some a else none := by
  cases n
  · simp
  · by_cases p a <;> simp_all [replicate_succ]
Behavior of `find?` on Replicated Lists: $\text{find? } p (\text{replicate}(n, a)) = \text{if } n = 0 \text{ then none else if } p(a) \text{ then some } a \text{ else none}$
Informal description
For any predicate $p : \alpha \to \text{Bool}$, natural number $n$, and element $a : \alpha$, the first element satisfying $p$ in the list $\text{replicate}(n, a)$ (a list containing $n$ copies of $a$) is: - $\text{none}$ if $n = 0$, - $\text{some } a$ if $n > 0$ and $p(a)$ holds, - $\text{none}$ otherwise.
List.find?_replicate_of_length_pos theorem
(h : 0 < n) : find? p (replicate n a) = if p a then some a else none
Full source
@[simp] theorem find?_replicate_of_length_pos (h : 0 < n) : find? p (replicate n a) = if p a then some a else none := by
  simp [find?_replicate, Nat.ne_of_gt h]
Behavior of `find?` on Non-Empty Replicated Lists: $\text{find? } p (\text{replicate}(n, a)) = \text{if } p(a) \text{ then some } a \text{ else none}$ for $n > 0$
Informal description
For any predicate $p : \alpha \to \text{Bool}$, natural number $n > 0$, and element $a : \alpha$, the first element satisfying $p$ in the list $\text{replicate}(n, a)$ (a list containing $n$ copies of $a$) is: - $\text{some } a$ if $p(a)$ holds, - $\text{none}$ otherwise.
List.find?_replicate_of_pos theorem
(h : p a) : find? p (replicate n a) = if n = 0 then none else some a
Full source
@[simp] theorem find?_replicate_of_pos (h : p a) : find? p (replicate n a) = if n = 0 then none else some a := by
  simp [find?_replicate, h]
Behavior of `find?` on Replicated Lists when Predicate Holds: $\text{find? } p (\text{replicate}(n, a)) = \text{if } n = 0 \text{ then none else some } a$
Informal description
For any predicate $p : \alpha \to \text{Bool}$, natural number $n$, and element $a : \alpha$, if $p(a)$ holds, then the first element satisfying $p$ in the list $\text{replicate}(n, a)$ (a list containing $n$ copies of $a$) is: - $\text{none}$ if $n = 0$, - $\text{some } a$ otherwise.
List.find?_replicate_of_neg theorem
(h : ¬p a) : find? p (replicate n a) = none
Full source
@[simp] theorem find?_replicate_of_neg (h : ¬ p a) : find? p (replicate n a) = none := by
  simp [find?_replicate, h]
`find?` on Replicated List Returns `none` When Predicate Fails
Informal description
For any predicate $p : \alpha \to \text{Bool}$, natural number $n$, and element $a : \alpha$, if $\neg p(a)$ holds, then the first element satisfying $p$ in the list $\text{replicate}(n, a)$ is $\text{none}$.
List.find?_replicate_eq_none_iff theorem
{n : Nat} {a : α} {p : α → Bool} : (replicate n a).find? p = none ↔ n = 0 ∨ !p a
Full source
theorem find?_replicate_eq_none_iff {n : Nat} {a : α} {p : α → Bool} :
    (replicate n a).find? p = none ↔ n = 0 ∨ !p a := by
  simp [Classical.or_iff_not_imp_left]
Condition for `find?` on Replicated List to Return `none`: $n = 0 \lor \neg p(a)$
Informal description
For any natural number $n$, element $a$ of type $\alpha$, and predicate $p : \alpha \to \text{Bool}$, the function `find?` applied to the list `replicate n a` returns `none` if and only if either $n = 0$ or the negation of $p(a)$ holds.
List.find?_replicate_eq_none abbrev
Full source
@[deprecated find?_replicate_eq_none_iff (since := "2025-02-03")]
abbrev find?_replicate_eq_none := @find?_replicate_eq_none_iff
`find?` on Replicated List Returns `none` iff Empty or Predicate Fails
Informal description
For any natural number $n$, element $a$ of type $\alpha$, and predicate $p : \alpha \to \text{Bool}$, the function `find?` applied to the list `replicate n a` returns `none` if and only if either $n = 0$ or the negation of $p(a)$ holds.
List.find?_replicate_eq_some_iff theorem
{n : Nat} {a b : α} {p : α → Bool} : (replicate n a).find? p = some b ↔ n ≠ 0 ∧ p a ∧ a = b
Full source
@[simp] theorem find?_replicate_eq_some_iff {n : Nat} {a b : α} {p : α → Bool} :
    (replicate n a).find? p = some b ↔ n ≠ 0 ∧ p a ∧ a = b := by
  cases n <;> simp
Characterization of `find?` on Replicated Lists: $\text{find? } p (\text{replicate } n a) = \text{some } b \leftrightarrow n \neq 0 \land p(a) \land a = b$
Informal description
For any natural number $n$, elements $a, b$ of type $\alpha$, and predicate $p : \alpha \to \text{Bool}$, the function `find?` applied to the list `replicate n a` (a list containing $n$ copies of $a$) returns `some b` if and only if: 1. $n \neq 0$, 2. $p(a)$ holds, and 3. $a = b$. In other words, $\text{find? } p (\text{replicate } n a) = \text{some } b$ if and only if $n$ is positive, $a$ satisfies $p$, and $b$ equals $a$.
List.find?_replicate_eq_some abbrev
Full source
@[deprecated find?_replicate_eq_some_iff (since := "2025-02-03")]
abbrev find?_replicate_eq_some := @find?_replicate_eq_some_iff
Sufficient Conditions for `find?` on Replicated List to Return `some b`
Informal description
For any natural number $n$, elements $a, b$ of type $\alpha$, and predicate $p : \alpha \to \text{Bool}$, if the list `replicate n a` (a list containing $n$ copies of $a$) satisfies $\text{find? } p (\text{replicate } n a) = \text{some } b$, then: 1. $n \neq 0$, 2. $p(a)$ holds, and 3. $a = b$.
List.get_find?_replicate theorem
{n : Nat} {a : α} {p : α → Bool} (h) : ((replicate n a).find? p).get h = a
Full source
@[simp] theorem get_find?_replicate {n : Nat} {a : α} {p : α → Bool} (h) : ((replicate n a).find? p).get h = a := by
  cases n with
  | zero => simp at h
  | succ n => simp
Extracted Value from Replicated List Find Operation Equals Original Element
Informal description
For any natural number $n$, element $a$ of type $\alpha$, and predicate $p : \alpha \to \text{Bool}$, if the result of finding an element in the replicated list $\text{replicate}(n, a)$ that satisfies $p$ is of the form $\text{some } x$ (as witnessed by the proof $h$), then the extracted value $x$ is equal to $a$.
List.Sublist.find?_isSome theorem
{l₁ l₂ : List α} (h : l₁ <+ l₂) : (l₁.find? p).isSome → (l₂.find? p).isSome
Full source
theorem Sublist.find?_isSome {l₁ l₂ : List α} (h : l₁ <+ l₂) : (l₁.find? p).isSome → (l₂.find? p).isSome := by
  induction h with
  | slnil => simp
  | cons a h ih
  | cons₂ a h ih =>
    simp only [find?]
    split
    · simp
    · simpa using ih
Sublist Preservation of Predicate Satisfaction
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a sublist of $l_2$ (denoted $l_1 <+ l_2$) and there exists an element in $l_1$ that satisfies the predicate $p$ (i.e., $(l_1.\text{find?}\ p).\text{isSome}$ holds), then there also exists an element in $l_2$ that satisfies $p$ (i.e., $(l_2.\text{find?}\ p).\text{isSome}$ holds).
List.Sublist.find?_eq_none theorem
{l₁ l₂ : List α} (h : l₁ <+ l₂) : l₂.find? p = none → l₁.find? p = none
Full source
theorem Sublist.find?_eq_none {l₁ l₂ : List α} (h : l₁ <+ l₂) : l₂.find? p = none → l₁.find? p = none := by
  simp only [List.find?_eq_none, Bool.not_eq_true]
  exact fun w x m => w x (Sublist.mem m h)
Sublist Preservation of Predicate Failure: $l_1 <+ l_2$ and $\text{find?}\ p\ l_2 = \text{none}$ implies $\text{find?}\ p\ l_1 = \text{none}$
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a sublist of $l_2$ (denoted $l_1 <+ l_2$) and the predicate $p$ fails to find any element in $l_2$ (i.e., $l_2.\text{find?}\ p = \text{none}$), then $p$ also fails to find any element in $l_1$ (i.e., $l_1.\text{find?}\ p = \text{none}$).
List.IsPrefix.find?_eq_some theorem
{l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂) : List.find? p l₁ = some b → List.find? p l₂ = some b
Full source
theorem IsPrefix.find?_eq_some {l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂) :
    List.find? p l₁ = some b → List.find? p l₂ = some b := by
  rw [IsPrefix] at h
  obtain ⟨t, rfl⟩ := h
  simp +contextual [find?_append]
Prefix Preservation of First Satisfying Element: $\text{find?}\ p\ l_1 = \text{some}\ b$ implies $\text{find?}\ p\ l_2 = \text{some}\ b$ when $l_1$ is prefix of $l_2$
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a prefix of $l_2$ (i.e., there exists a list $t$ such that $l_2 = l_1 ++ t$) and the first element in $l_1$ satisfying predicate $p$ is $b$ (i.e., $\text{find?}\ p\ l_1 = \text{some}\ b$), then the first element in $l_2$ satisfying $p$ is also $b$ (i.e., $\text{find?}\ p\ l_2 = \text{some}\ b$).
List.IsPrefix.find?_eq_none theorem
{l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂) : List.find? p l₂ = none → List.find? p l₁ = none
Full source
theorem IsPrefix.find?_eq_none {l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂) :
    List.find? p l₂ = noneList.find? p l₁ = none :=
  h.sublist.find?_eq_none
Prefix Preservation of Predicate Failure in List Search
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a prefix of $l_2$ (denoted $l_1 <+: l_2$) and the predicate $p$ fails to find any element in $l_2$ (i.e., $\text{find?}\ p\ l_2 = \text{none}$), then $p$ also fails to find any element in $l_1$ (i.e., $\text{find?}\ p\ l_1 = \text{none}$).
List.IsSuffix.find?_eq_none theorem
{l₁ l₂ : List α} {p : α → Bool} (h : l₁ <:+ l₂) : List.find? p l₂ = none → List.find? p l₁ = none
Full source
theorem IsSuffix.find?_eq_none {l₁ l₂ : List α} {p : α → Bool} (h : l₁ <:+ l₂) :
    List.find? p l₂ = noneList.find? p l₁ = none :=
  h.sublist.find?_eq_none
Suffix Preservation of Predicate Failure: $\text{find?}\ p\ l_2 = \text{none}$ implies $\text{find?}\ p\ l_1 = \text{none}$ when $l_1$ is suffix of $l_2$
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a suffix of $l_2$ (denoted $l_1 <:+ l_2$) and the predicate $p$ fails to find any element in $l_2$ (i.e., $\text{find?}\ p\ l_2 = \text{none}$), then $p$ also fails to find any element in $l_1$ (i.e., $\text{find?}\ p\ l_1 = \text{none}$).
List.IsInfix.find?_eq_none theorem
{l₁ l₂ : List α} {p : α → Bool} (h : l₁ <:+: l₂) : List.find? p l₂ = none → List.find? p l₁ = none
Full source
theorem IsInfix.find?_eq_none {l₁ l₂ : List α} {p : α → Bool} (h : l₁ <:+: l₂) :
    List.find? p l₂ = noneList.find? p l₁ = none :=
  h.sublist.find?_eq_none
Infix Preservation of Predicate Failure: $l_1 \sqsubseteq l_2$ and $\text{find?}\ p\ l_2 = \text{none}$ implies $\text{find?}\ p\ l_1 = \text{none}$
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is an infix of $l_2$ (denoted $l_1 \sqsubseteq l_2$) and the predicate $p$ fails to find any element in $l_2$ (i.e., $\text{find?}\ p\ l_2 = \text{none}$), then $p$ also fails to find any element in $l_1$ (i.e., $\text{find?}\ p\ l_1 = \text{none}$).
List.find?_pmap theorem
{P : α → Prop} {f : (a : α) → P a → β} {xs : List α} (H : ∀ (a : α), a ∈ xs → P a) {p : β → Bool} : (xs.pmap f H).find? p = (xs.attach.find? (fun ⟨a, m⟩ => p (f a (H a m)))).map fun ⟨a, m⟩ => f a (H a m)
Full source
theorem find?_pmap {P : α → Prop} {f : (a : α) → P a → β} {xs : List α}
    (H : ∀ (a : α), a ∈ xs → P a) {p : β → Bool} :
    (xs.pmap f H).find? p = (xs.attach.find? (fun ⟨a, m⟩ => p (f a (H a m)))).map fun ⟨a, m⟩ => f a (H a m) := by
  simp only [pmap_eq_map_attach, find?_map]
  rfl
Commutativity of `find?` and `pmap` with Predicate Composition
Informal description
Let $P$ be a predicate on a type $\alpha$, $f : (a : \alpha) \to P(a) \to \beta$ a function, and $xs$ a list of elements of type $\alpha$ such that every element $a \in xs$ satisfies $P(a)$. For any predicate $p : \beta \to \text{Bool}$, the first element satisfying $p$ in the mapped list $\text{pmap}\ f\ H\ xs$ is equal to the result of applying $f$ to the first element in the attached list $\text{attach}\ xs$ that satisfies $p \circ f$ under the membership proof. Formally: \[ \text{find?}\ p\ (\text{pmap}\ f\ H\ xs) = \text{map}\ (\lambda \langle a, m \rangle, f\ a\ (H\ a\ m))\ (\text{find?}\ (\lambda \langle a, m \rangle, p\ (f\ a\ (H\ a\ m)))\ (\text{attach}\ xs)) \]
List.findIdx?_nil theorem
: ([] : List α).findIdx? p = none
Full source
@[simp] theorem findIdx?_nil : ([] : List α).findIdx? p = none := rfl
Empty List Has No Satisfying Index
Informal description
For any predicate $p : \alpha \to \text{Bool}$, the function `findIdx?` applied to the empty list returns `none`, i.e., $\text{findIdx?}_p([]) = \text{none}$.
List.findIdx?_cons theorem
: (x :: xs).findIdx? p = if p x then some 0 else (xs.findIdx? p).map fun i => i + 1
Full source
@[simp] theorem findIdx?_cons :
    (x :: xs).findIdx? p = if p x then some 0 else (xs.findIdx? p).map fun i => i + 1 := by
  simp [findIdx?, findIdx?_go_eq]
Recursive computation of first satisfying index in a list: $\text{findIdx?}_p(x :: xs) = \text{if } p(x) \text{ then some } 0 \text{ else map } (\lambda i, i + 1) (\text{findIdx?}_p xs)$
Informal description
For any predicate $p : \alpha \to \text{Bool}$, element $x : \alpha$, and list $xs : \text{List } \alpha$, the index of the first element in the list $x :: xs$ satisfying $p$ is equal to $\text{some } 0$ if $p(x)$ holds, otherwise it is equal to $\text{map } (\lambda i, i + 1) (\text{findIdx?}_p xs)$.
List.findIdx_cons theorem
{p : α → Bool} {b : α} {l : List α} : (b :: l).findIdx p = bif p b then 0 else (l.findIdx p) + 1
Full source
theorem findIdx_cons {p : α → Bool} {b : α} {l : List α} :
    (b :: l).findIdx p = bif p b then 0 else (l.findIdx p) + 1 := by
  cases H : p b with
  | true => simp [H, findIdx, findIdx.go]
  | false => simp [H, findIdx, findIdx.go, findIdx_go_succ]
where
  findIdx_go_succ (p : α → Bool) (l : List α) (n : Nat) :
      List.findIdx.go p l (n + 1) = (findIdx.go p l n) + 1 := by
    cases l with
    | nil => unfold findIdx.go; exact Nat.succ_eq_add_one n
    | cons hd tl =>
      unfold findIdx.go
      cases p hd <;> simp only [cond_false, cond_true]
      exact findIdx_go_succ p tl (n + 1)
Recursive Computation of First Satisfying Index in a List: $\text{findIdx}_p(b :: l) = \text{if } p(b) \text{ then } 0 \text{ else } (\text{findIdx}_p l) + 1$
Informal description
For any predicate $p : \alpha \to \text{Bool}$, element $b : \alpha$, and list $l : \text{List } \alpha$, the index of the first element in the list $b :: l$ satisfying $p$ is equal to $0$ if $p(b)$ holds, otherwise it is equal to $1$ plus the index of the first satisfying element in $l$.
List.findIdx_of_getElem?_eq_some theorem
{xs : List α} (w : xs[xs.findIdx p]? = some y) : p y
Full source
theorem findIdx_of_getElem?_eq_some {xs : List α} (w : xs[xs.findIdx p]? = some y) : p y := by
  induction xs with
  | nil => simp_all
  | cons x xs ih => by_cases h : p x <;> simp_all [findIdx_cons]
Predicate Holds at Found Index: $p(y)$ when $xs[\text{findIdx}_p\,xs]? = \text{some}\,y$
Informal description
For any list `xs` of type `List α` and predicate `p : α → Bool`, if the optional indexing operation `xs[xs.findIdx p]?` returns `some y` for some element `y`, then `p y` holds.
List.findIdx_getElem theorem
{xs : List α} {w : xs.findIdx p < xs.length} : p xs[xs.findIdx p]
Full source
theorem findIdx_getElem {xs : List α} {w : xs.findIdx p < xs.length} :
    p xs[xs.findIdx p] :=
  xs.findIdx_of_getElem?_eq_some (getElem?_eq_getElem w)
Predicate Holds at Found Index: $p(xs[\text{findIdx}_p\,xs])$ when $\text{findIdx}_p\,xs < \text{length}\,xs$
Informal description
For any list `xs` of elements of type `α` and predicate `p : α → Bool`, if the index `xs.findIdx p` is a valid index for `xs` (i.e., `xs.findIdx p < xs.length`), then the predicate `p` holds for the element at that index, i.e., `p(xs[xs.findIdx p])`.
List.findIdx_lt_length_of_exists theorem
{xs : List α} (h : ∃ x ∈ xs, p x) : xs.findIdx p < xs.length
Full source
theorem findIdx_lt_length_of_exists {xs : List α} (h : ∃ x ∈ xs, p x) :
    xs.findIdx p < xs.length := by
  induction xs with
  | nil => simp_all
  | cons x xs ih =>
    by_cases p x
    · simp_all only [forall_exists_index, and_imp, mem_cons, exists_eq_or_imp, true_or,
        findIdx_cons, cond_true, length_cons]
      apply Nat.succ_pos
    · simp_all [findIdx_cons, Nat.succ_lt_succ_iff]
      obtain ⟨x', m', h'⟩ := h
      exact ih x' m' h'
Existence of Satisfying Element Implies Valid Index: $\text{findIdx}_p(xs) < \text{length}(xs)$
Informal description
For any list `xs` of elements of type `α` and any predicate `p : α → Bool`, if there exists an element `x ∈ xs` such that `p x` holds, then the index of the first element in `xs` satisfying `p` is strictly less than the length of `xs`.
List.findIdx_getElem?_eq_getElem_of_exists theorem
{xs : List α} (h : ∃ x ∈ xs, p x) : xs[xs.findIdx p]? = some (xs[xs.findIdx p]'(xs.findIdx_lt_length_of_exists h))
Full source
theorem findIdx_getElem?_eq_getElem_of_exists {xs : List α} (h : ∃ x ∈ xs, p x) :
    xs[xs.findIdx p]? = some (xs[xs.findIdx p]'(xs.findIdx_lt_length_of_exists h)) :=
  getElem?_eq_getElem (findIdx_lt_length_of_exists h)
Optional Indexing at First Satisfying Index Yields Element
Informal description
For any list $xs$ of elements of type $\alpha$ and predicate $p : \alpha \to \text{Bool}$, if there exists an element $x \in xs$ such that $p(x)$ holds, then the optional indexing operation at the first index where $p$ holds returns the corresponding element, i.e., $xs[\text{findIdx}_p(xs)]? = \text{some}(xs[\text{findIdx}_p(xs)])$.
List.findIdx_eq_length theorem
{p : α → Bool} {xs : List α} : xs.findIdx p = xs.length ↔ ∀ x ∈ xs, p x = false
Full source
@[simp]
theorem findIdx_eq_length {p : α → Bool} {xs : List α} :
    xs.findIdx p = xs.length ↔ ∀ x ∈ xs, p x = false := by
  induction xs with
  | nil => simp_all
  | cons x xs ih =>
    rw [findIdx_cons, length_cons]
    simp only [cond_eq_if]
    split <;> simp_all [Nat.succ.injEq]
First Satisfying Index Equals List Length if and only if All Elements Fail Predicate
Informal description
For a predicate $p : \alpha \to \text{Bool}$ and a list $xs$ of elements of type $\alpha$, the index of the first element in $xs$ satisfying $p$ equals the length of $xs$ if and only if no element in $xs$ satisfies $p$ (i.e., $p(x) = \text{false}$ for all $x \in xs$).
List.findIdx_eq_length_of_false theorem
{p : α → Bool} {xs : List α} (h : ∀ x ∈ xs, p x = false) : xs.findIdx p = xs.length
Full source
theorem findIdx_eq_length_of_false {p : α → Bool} {xs : List α} (h : ∀ x ∈ xs, p x = false) :
    xs.findIdx p = xs.length := by
  rw [findIdx_eq_length]
  exact h
Index of First Satisfying Element Equals List Length When All Elements Fail Predicate
Informal description
For any predicate $p : \alpha \to \text{Bool}$ and list $xs$ of elements of type $\alpha$, if $p(x) = \text{false}$ for all $x \in xs$, then the index of the first element in $xs$ satisfying $p$ equals the length of $xs$.
List.findIdx_le_length theorem
{p : α → Bool} {xs : List α} : xs.findIdx p ≤ xs.length
Full source
theorem findIdx_le_length {p : α → Bool} {xs : List α} : xs.findIdx p ≤ xs.length := by
  by_cases e : ∃ x ∈ xs, p x
  · exact Nat.le_of_lt (findIdx_lt_length_of_exists e)
  · simp at e
    exact Nat.le_of_eq (findIdx_eq_length.mpr e)
Bounded Index Property for List Search: $\text{findIdx}_p(xs) \leq \text{length}(xs)$
Informal description
For any predicate $p : \alpha \to \text{Bool}$ and any list $xs$ of elements of type $\alpha$, the index of the first element in $xs$ satisfying $p$ is less than or equal to the length of $xs$.
List.findIdx_lt_length theorem
{p : α → Bool} {xs : List α} : xs.findIdx p < xs.length ↔ ∃ x ∈ xs, p x
Full source
@[simp]
theorem findIdx_lt_length {p : α → Bool} {xs : List α} :
    xs.findIdx p < xs.length ↔ ∃ x ∈ xs, p x := by
  rw [← Decidable.not_iff_not, Nat.not_lt]
  have := @Nat.le_antisymm_iff (xs.findIdx p) xs.length
  simp only [findIdx_le_length, true_and] at this
  rw [← this, findIdx_eq_length, not_exists]
  simp only [Bool.not_eq_true, not_and]
First Satisfying Index is Within Bounds if and only if Predicate Holds for Some Element
Informal description
For any predicate $p : \alpha \to \text{Bool}$ and any list $xs$ of elements of type $\alpha$, the index of the first element in $xs$ satisfying $p$ is less than the length of $xs$ if and only if there exists an element $x \in xs$ such that $p(x)$ holds.
List.not_of_lt_findIdx theorem
{p : α → Bool} {xs : List α} {i : Nat} (h : i < xs.findIdx p) : p (xs[i]'(Nat.le_trans h findIdx_le_length)) = false
Full source
/-- `p` does not hold for elements with indices less than `xs.findIdx p`. -/
theorem not_of_lt_findIdx {p : α → Bool} {xs : List α} {i : Nat} (h : i < xs.findIdx p) :
    p (xs[i]'(Nat.le_trans h findIdx_le_length)) = false := by
  revert i
  induction xs with
  | nil => intro i h; rw [findIdx_nil] at h; simp at h
  | cons x xs ih =>
    intro i h
    have ho := h
    rw [findIdx_cons] at h
    have npx : p x = false := by
      apply eq_false_of_ne_true
      intro y
      rw [y, cond_true] at h
      simp at h
    simp [npx, cond_false] at h
    cases i.eq_zero_or_pos with
    | inl e => simpa [e, Fin.zero_eta, get_cons_zero]
    | inr e =>
      have ipm := Nat.succ_pred_eq_of_pos e
      simp +singlePass only [← ipm, getElem_cons_succ]
      rw [← ipm, Nat.succ_lt_succ_iff] at h
      simpa using ih h
Predicate Fails Before First Satisfying Index in List
Informal description
For any predicate $p : \alpha \to \text{Bool}$, list $xs : \text{List } \alpha$, and natural number index $i$, if $i$ is less than the index of the first element in $xs$ satisfying $p$, then $p$ evaluates to false on the $i$-th element of $xs$.
List.le_findIdx_of_not theorem
{p : α → Bool} {xs : List α} {i : Nat} (h : i < xs.length) (h2 : ∀ j (hji : j < i), p (xs[j]'(Nat.lt_trans hji h)) = false) : i ≤ xs.findIdx p
Full source
/-- If `¬ p xs[j]` for all `j < i`, then `i ≤ xs.findIdx p`. -/
theorem le_findIdx_of_not {p : α → Bool} {xs : List α} {i : Nat} (h : i < xs.length)
    (h2 : ∀ j (hji : j < i), p (xs[j]'(Nat.lt_trans hji h)) = false) : i ≤ xs.findIdx p := by
  apply Decidable.byContradiction
  intro f
  simp only [Nat.not_le] at f
  exact absurd (@findIdx_getElem _ p xs (Nat.lt_trans f h)) (by simpa using h2 (xs.findIdx p) f)
Lower bound on first satisfying index in list: $i \leq \text{findIdx}_p\,xs$ when $p$ fails on all $j < i$
Informal description
For any predicate $p : \alpha \to \text{Bool}$, list $xs$ of elements of type $\alpha$, and natural number index $i$ such that $i < \text{length}(xs)$, if for all indices $j < i$ the predicate $p$ evaluates to false on the $j$-th element of $xs$, then $i$ is less than or equal to the index of the first element in $xs$ satisfying $p$. In other words, if no element before position $i$ satisfies $p$, then the first satisfying element (if any) must be at position $i$ or later.
List.lt_findIdx_of_not theorem
{p : α → Bool} {xs : List α} {i : Nat} (h : i < xs.length) (h2 : ∀ j (hji : j ≤ i), ¬p (xs[j]'(Nat.lt_of_le_of_lt hji h))) : i < xs.findIdx p
Full source
/-- If `¬ p xs[j]` for all `j ≤ i`, then `i < xs.findIdx p`. -/
theorem lt_findIdx_of_not {p : α → Bool} {xs : List α} {i : Nat} (h : i < xs.length)
    (h2 : ∀ j (hji : j ≤ i), ¬p (xs[j]'(Nat.lt_of_le_of_lt hji h))) : i < xs.findIdx p := by
  apply Decidable.byContradiction
  intro f
  simp only [Nat.not_lt] at f
  exact absurd (@findIdx_getElem _ p xs (Nat.lt_of_le_of_lt f h)) (h2 (xs.findIdx p) f)
Index Lower Bound for First Satisfying Element in List: $i < \text{findIdx}_p\,xs$ when $p$ fails on all $j \leq i$
Informal description
For any predicate $p : \alpha \to \text{Bool}$, list $xs : \text{List } \alpha$, and natural number index $i$ such that $i < \text{length}(xs)$, if for all $j \leq i$ the predicate $p$ does not hold on the $j$-th element of $xs$, then $i$ is strictly less than the index of the first element in $xs$ satisfying $p$.
List.findIdx_eq theorem
{p : α → Bool} {xs : List α} {i : Nat} (h : i < xs.length) : xs.findIdx p = i ↔ p xs[i] ∧ ∀ j (hji : j < i), p (xs[j]'(Nat.lt_trans hji h)) = false
Full source
/-- `xs.findIdx p = i` iff `p xs[i]` and `¬ p xs [j]` for all `j < i`. -/
theorem findIdx_eq {p : α → Bool} {xs : List α} {i : Nat} (h : i < xs.length) :
    xs.findIdx p = i ↔ p xs[i] ∧ ∀ j (hji : j < i), p (xs[j]'(Nat.lt_trans hji h)) = false := by
  refine ⟨fun f ↦ ⟨f ▸ (@findIdx_getElem _ p xs (f ▸ h)), fun _ hji ↦ not_of_lt_findIdx (f ▸ hji)⟩,
    fun ⟨_, h2⟩ ↦ ?_⟩
  apply Nat.le_antisymm _ (le_findIdx_of_not h h2)
  apply Decidable.byContradiction
  intro h3
  simp at h3
  simp_all [not_of_lt_findIdx h3]
Characterization of First Satisfying Index in List: $\text{findIdx}_p\,xs = i \leftrightarrow p(xs[i]) \land (\forall j < i, \neg p(xs[j]))$
Informal description
For a predicate $p : \alpha \to \text{Bool}$, a list $xs$ of elements of type $\alpha$, and a natural number index $i$ such that $i < \text{length}(xs)$, the index of the first element in $xs$ satisfying $p$ equals $i$ if and only if: 1. The predicate $p$ holds for the $i$-th element of $xs$, i.e., $p(xs[i])$ is true, and 2. For all indices $j < i$, the predicate $p$ does not hold for the $j$-th element of $xs$, i.e., $p(xs[j])$ is false.
List.findIdx_append theorem
{p : α → Bool} {l₁ l₂ : List α} : (l₁ ++ l₂).findIdx p = if l₁.findIdx p < l₁.length then l₁.findIdx p else l₂.findIdx p + l₁.length
Full source
theorem findIdx_append {p : α → Bool} {l₁ l₂ : List α} :
    (l₁ ++ l₂).findIdx p =
      if l₁.findIdx p < l₁.length then l₁.findIdx p else l₂.findIdx p + l₁.length := by
  induction l₁ with
  | nil => simp
  | cons x xs ih =>
    simp only [findIdx_cons, length_cons, cons_append]
    by_cases h : p x
    · simp [h]
    · simp only [h, ih, cond_eq_if, Bool.false_eq_true, ↓reduceIte, add_one_lt_add_one_iff]
      split <;> simp [Nat.add_assoc]
Index of First Satisfying Element in Concatenated Lists: $\text{findIdx}_p(l_1 \mathbin{+\kern-1.5ex+} l_2) = \text{if } \text{findIdx}_p(l_1) < \text{length}(l_1) \text{ then } \text{findIdx}_p(l_1) \text{ else } \text{findIdx}_p(l_2) + \text{length}(l_1)$
Informal description
For any predicate $p : \alpha \to \text{Bool}$ and lists $l_1, l_2 : \text{List } \alpha$, the index of the first element in $l_1 \mathbin{+\kern-1.5ex+} l_2$ satisfying $p$ is equal to the index of the first satisfying element in $l_1$ if such an element exists in $l_1$, otherwise it is equal to the length of $l_1$ plus the index of the first satisfying element in $l_2$.
List.IsPrefix.findIdx_le theorem
{l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂) : l₁.findIdx p ≤ l₂.findIdx p
Full source
theorem IsPrefix.findIdx_le {l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂) :
    l₁.findIdx p ≤ l₂.findIdx p := by
  rw [IsPrefix] at h
  obtain ⟨t, rfl⟩ := h
  simp only [findIdx_append, findIdx_lt_length]
  split
  · exact Nat.le_refl ..
  · simp_all [findIdx_eq_length_of_false]
Monotonicity of First Satisfying Index under List Prefix: $l_1 \prec l_2 \Rightarrow \text{findIdx}_p(l_1) \leq \text{findIdx}_p(l_2)$
Informal description
For any lists $l_1$ and $l_2$ of elements of type $\alpha$ and any predicate $p : \alpha \to \text{Bool}$, if $l_1$ is a prefix of $l_2$, then the index of the first element in $l_1$ satisfying $p$ is less than or equal to the index of the first element in $l_2$ satisfying $p$.
List.IsPrefix.findIdx_eq_of_findIdx_lt_length theorem
{l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂) (lt : l₁.findIdx p < l₁.length) : l₂.findIdx p = l₁.findIdx p
Full source
theorem IsPrefix.findIdx_eq_of_findIdx_lt_length {l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂)
    (lt : l₁.findIdx p < l₁.length) : l₂.findIdx p = l₁.findIdx p := by
  rw [IsPrefix] at h
  obtain ⟨t, rfl⟩ := h
  simp only [findIdx_append, findIdx_lt_length]
  split
  · rfl
  · simp_all
Prefix Preservation of First Satisfying Index: $\text{findIdx}_p(l_2) = \text{findIdx}_p(l_1)$ when $l_1$ is prefix and $\text{findIdx}_p(l_1) < \text{length}(l_1)$
Informal description
For any lists $l_1$ and $l_2$ of elements of type $\alpha$, and any predicate $p : \alpha \to \text{Bool}$, if $l_1$ is a prefix of $l_2$ and the index of the first element in $l_1$ satisfying $p$ is less than the length of $l_1$, then the index of the first satisfying element in $l_2$ equals that in $l_1$.
List.findIdx_le_findIdx theorem
{l : List α} {p q : α → Bool} (h : ∀ x ∈ l, p x → q x) : l.findIdx q ≤ l.findIdx p
Full source
theorem findIdx_le_findIdx {l : List α} {p q : α → Bool} (h : ∀ x ∈ l, p x → q x) : l.findIdx q ≤ l.findIdx p := by
  induction l with
  | nil => simp
  | cons x xs ih =>
    simp only [findIdx_cons, cond_eq_if]
    split
    · simp
    · split
      · simp_all
      · simp only [Nat.add_le_add_iff_right]
        exact ih fun _ m w => h _ (mem_cons_of_mem x m) w
Monotonicity of First Satisfying Index: $(\forall x \in l, p(x) \to q(x)) \to \text{findIdx}_q(l) \leq \text{findIdx}_p(l)$
Informal description
For any list $l$ of elements of type $\alpha$ and any two predicates $p, q : \alpha \to \text{Bool}$, if for every element $x$ in $l$ the implication $p(x) \to q(x)$ holds, then the index of the first element in $l$ satisfying $q$ is less than or equal to the index of the first element in $l$ satisfying $p$.
List.findIdx_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → Bool} {g : α → Bool} (hf : ∀ x h, f ⟨x, h⟩ = g x) : l.findIdx f = l.unattach.findIdx g
Full source
@[simp] theorem findIdx_subtype {p : α → Prop} {l : List { x // p x }}
    {f : { x // p x }Bool} {g : α → Bool} (hf : ∀ x h, f ⟨x, h⟩ = g x) :
    l.findIdx f = l.unattach.findIdx g := by
  unfold unattach
  induction l with
  | nil => simp
  | cons a l ih =>
    simp [ih, hf, findIdx_cons]
Equality of Indices for Predicates on Subtypes and Their Base Types
Informal description
Let $p : \alpha \to \text{Prop}$ be a predicate, $l$ be a list of elements of type $\{x : \alpha \mid p x\}$, and $f : \{x : \alpha \mid p x\} \to \text{Bool}$ and $g : \alpha \to \text{Bool}$ be predicates such that for any $x$ and proof $h$ of $p(x)$, we have $f(\langle x, h \rangle) = g(x)$. Then the index of the first element in $l$ satisfying $f$ is equal to the index of the first element in the unattached list $l$ (viewed as a list of $\alpha$) satisfying $g$.
List.findIdx?_eq_none_iff theorem
{xs : List α} {p : α → Bool} : xs.findIdx? p = none ↔ ∀ x, x ∈ xs → p x = false
Full source
@[simp]
theorem findIdx?_eq_none_iff {xs : List α} {p : α → Bool} :
    xs.findIdx? p = none ↔ ∀ x, x ∈ xs → p x = false := by
  induction xs with
  | nil => simp_all
  | cons x xs ih =>
    simp only [findIdx?_cons]
    split <;> simp_all [cond_eq_if]
Nonexistence of Satisfying Index in List: $\text{findIdx?}_p(xs) = \text{none} \leftrightarrow \forall x \in xs, p(x) = \text{false}$
Informal description
For a list `xs` of elements of type `α` and a predicate `p : α → Bool`, the function `findIdx? p xs` returns `none` if and only if for every element `x` in `xs`, the predicate `p x` evaluates to `false`.
List.findIdx?_isSome theorem
{xs : List α} {p : α → Bool} : (xs.findIdx? p).isSome = xs.any p
Full source
theorem findIdx?_isSome {xs : List α} {p : α → Bool} :
    (xs.findIdx? p).isSome = xs.any p := by
  induction xs with
  | nil => simp
  | cons x xs ih =>
    simp only [findIdx?_cons]
    split <;> simp_all
`findIdx?` Returns `some` Index if and only if `any` Element Satisfies Predicate
Informal description
For any list `xs` of elements of type `α` and any predicate `p : α → Bool`, the boolean check whether `findIdx? p xs` returns `some` index is equal to the result of checking if any element in `xs` satisfies `p`. In other words, `(findIdx? p xs).isSome = any p xs`.
List.findIdx?_isNone theorem
{xs : List α} {p : α → Bool} : (xs.findIdx? p).isNone = xs.all (¬p ·)
Full source
theorem findIdx?_isNone {xs : List α} {p : α → Bool} :
    (xs.findIdx? p).isNone = xs.all (¬p ·) := by
  induction xs with
  | nil => simp
  | cons x xs ih =>
    simp only [findIdx?_cons]
    split <;> simp_all
Equivalence of `findIdx?` Being `none` and All Elements Failing the Predicate
Informal description
For any list `xs` of elements of type `α` and any predicate `p : α → Bool`, the result of checking whether `findIdx? p xs` is `none` is equal to checking whether all elements in `xs` satisfy `¬p x`. In other words, `(xs.findIdx? p).isNone = xs.all (λ x, ¬p x)`.
List.findIdx?_eq_some_iff_findIdx_eq theorem
{xs : List α} {p : α → Bool} {i : Nat} : xs.findIdx? p = some i ↔ i < xs.length ∧ xs.findIdx p = i
Full source
theorem findIdx?_eq_some_iff_findIdx_eq {xs : List α} {p : α → Bool} {i : Nat} :
    xs.findIdx? p = some i ↔ i < xs.length ∧ xs.findIdx p = i := by
  induction xs generalizing i with
  | nil => simp_all
  | cons x xs ih =>
    simp only [findIdx?_cons, findIdx_cons]
    split
    · simp_all [cond_eq_if]
      rintro rfl
      exact zero_lt_succ xs.length
    · simp_all [cond_eq_if, and_assoc]
      constructor
      · rintro ⟨a, lt, rfl, rfl⟩
        simp_all [Nat.succ_lt_succ_iff]
      · rintro ⟨h, rfl⟩
        exact ⟨_, by simp_all [Nat.succ_lt_succ_iff], rfl, rfl⟩
Equivalence of `findIdx?` and `findIdx` for Valid Indices
Informal description
For any list `xs` of elements of type `alpha`, any predicate `p : alpha → Bool`, and any natural number `i`, the following are equivalent: 1. The function `findIdx?` returns `some i`, indicating the first element satisfying `p` is at index `i`. 2. The index `i` is within the bounds of the list (`i < xs.length`) and the function `findIdx` returns `i` for the same predicate. In other words, `xs.findIdx? p = some i` if and only if `i < xs.length` and `xs.findIdx p = i`.
List.findIdx?_eq_some_of_exists theorem
{xs : List α} {p : α → Bool} (h : ∃ x, x ∈ xs ∧ p x) : xs.findIdx? p = some (xs.findIdx p)
Full source
theorem findIdx?_eq_some_of_exists {xs : List α} {p : α → Bool} (h : ∃ x, x ∈ xs ∧ p x) :
    xs.findIdx? p = some (xs.findIdx p) := by
  rw [findIdx?_eq_some_iff_findIdx_eq]
  exact ⟨findIdx_lt_length_of_exists h, rfl⟩
Existence of Satisfying Element Implies `findIdx?` Returns Some Index
Informal description
For any list `xs` of elements of type `α` and any predicate `p : α → Bool`, if there exists an element `x ∈ xs` such that `p x` holds, then the function `findIdx?` returns `some` applied to the index of the first such element as computed by `findIdx`. In other words, `xs.findIdx? p = some (xs.findIdx p)` under the given existence condition.
List.findIdx?_eq_none_iff_findIdx_eq theorem
{xs : List α} {p : α → Bool} : xs.findIdx? p = none ↔ xs.findIdx p = xs.length
Full source
theorem findIdx?_eq_none_iff_findIdx_eq {xs : List α} {p : α → Bool} :
    xs.findIdx? p = none ↔ xs.findIdx p = xs.length := by
  simp
Equivalence of `findIdx?` Returning `none` and `findIdx` Returning List Length
Informal description
For any list `xs` of elements of type `α` and any predicate `p : α → Bool`, the following are equivalent: 1. The function `findIdx?` returns `none`, indicating no element in `xs` satisfies `p`. 2. The function `findIdx` returns the length of `xs`, indicating no element in `xs` satisfies `p`. In other words, `xs.findIdx? p = none` if and only if `xs.findIdx p = xs.length`.
List.findIdx?_eq_guard_findIdx_lt theorem
{xs : List α} {p : α → Bool} : xs.findIdx? p = Option.guard (fun i => i < xs.length) (xs.findIdx p)
Full source
theorem findIdx?_eq_guard_findIdx_lt {xs : List α} {p : α → Bool} :
    xs.findIdx? p = Option.guard (fun i => i < xs.length) (xs.findIdx p) := by
  match h : xs.findIdx? p with
  | none =>
    simp only [findIdx?_eq_none_iff] at h
    simp [findIdx_eq_length_of_false h, Option.guard]
  | somesome i =>
    simp only [findIdx?_eq_some_iff_findIdx_eq] at h
    simp [h]
$\text{findIdx?}\ p\ xs$ as a Guarded Version of $\text{findIdx}\ p\ xs$
Informal description
For any list $xs$ of elements of type $\alpha$ and any predicate $p : \alpha \to \text{Bool}$, the function $\text{findIdx?}\ p\ xs$ is equal to $\text{Option.guard}\ (\lambda i, i < \text{length}(xs))\ (\text{findIdx}\ p\ xs)$. That is, $\text{findIdx?}\ p\ xs$ returns $\text{some}\ i$ if and only if $i = \text{findIdx}\ p\ xs$ and $i < \text{length}(xs)$, and $\text{none}$ otherwise.
List.findIdx?_eq_some_iff_getElem theorem
{xs : List α} {p : α → Bool} {i : Nat} : xs.findIdx? p = some i ↔ ∃ h : i < xs.length, p xs[i] ∧ ∀ j (hji : j < i), ¬p (xs[j]'(Nat.lt_trans hji h))
Full source
theorem findIdx?_eq_some_iff_getElem {xs : List α} {p : α → Bool} {i : Nat} :
    xs.findIdx? p = some i ↔
      ∃ h : i < xs.length, p xs[i] ∧ ∀ j (hji : j < i), ¬p (xs[j]'(Nat.lt_trans hji h)) := by
  induction xs generalizing i with
  | nil => simp
  | cons x xs ih =>
    simp only [findIdx?_cons, Nat.zero_add]
    split
    · simp only [Option.some.injEq, Bool.not_eq_true, length_cons]
      cases i with
      | zero => simp_all
      | succ i =>
        simp only [Bool.not_eq_true, zero_ne_add_one, getElem_cons_succ, false_iff, not_exists,
          not_and, Classical.not_forall, Bool.not_eq_false]
        intros
        refine ⟨0, zero_lt_succ i, ‹_›⟩
    · simp only [Option.map_eq_some', ih, Bool.not_eq_true, length_cons]
      constructor
      · rintro ⟨a, ⟨⟨h, h₁, h₂⟩, rfl⟩⟩
        refine ⟨Nat.succ_lt_succ_iff.mpr h, by simpa, fun j hj => ?_⟩
        cases j with
        | zero => simp_all
        | succ j =>
          apply h₂
          simp_all [Nat.succ_lt_succ_iff]
      · rintro ⟨h, h₁, h₂⟩
        cases i with
        | zero => simp_all
        | succ i =>
          refine ⟨i, ⟨Nat.succ_lt_succ_iff.mp h, by simpa, fun j hj => ?_⟩, rfl⟩
          simpa using h₂ (j + 1) (Nat.succ_lt_succ_iff.mpr hj)
Characterization of `findIdx?` Result as First Satisfying Index
Informal description
For a list `xs` of elements of type `alpha`, a predicate `p : alpha → Bool`, and a natural number `i`, the following are equivalent: 1. The function `findIdx? p xs` returns `some i`. 2. There exists a proof `h` that `i` is a valid index of `xs` (i.e., `i < xs.length`), the element `xs[i]` satisfies `p`, and for all indices `j < i`, the element `xs[j]` does not satisfy `p`. In symbols: $$\text{findIdx?}_p(xs) = \text{some } i \leftrightarrow \exists h : i < \text{length}(xs), p(xs[i]) \land \forall j < i, \neg p(xs[j])$$
List.of_findIdx?_eq_some theorem
{xs : List α} {p : α → Bool} (w : xs.findIdx? p = some i) : match xs[i]? with | some a => p a | none => false
Full source
theorem of_findIdx?_eq_some {xs : List α} {p : α → Bool} (w : xs.findIdx? p = some i) :
    match xs[i]? with | some a => p a | none => false := by
  induction xs generalizing i with
  | nil => simp_all
  | cons x xs ih =>
    simp_all only [findIdx?_cons, Nat.zero_add]
    split at w <;> cases i <;> simp_all [succ_inj']
Characterization of Successful Index Search in List: `findIdx? p xs = some i` Implies `p (xs[i])`
Informal description
For any list `xs` of type `List α` and predicate `p : α → Bool`, if `xs.findIdx? p` returns `some i`, then either: 1. The optional indexing `xs[i]?` returns `some a` and `p a` holds, or 2. `xs[i]?` returns `none` (which is impossible since `i` must be a valid index by the definition of `findIdx?`).
List.findIdx?_of_eq_some abbrev
Full source
@[deprecated of_findIdx?_eq_some (since := "2025-02-02")]
abbrev findIdx?_of_eq_some := @of_findIdx?_eq_some
Implication of Successful Index Search in List: `findIdx? p xs = some i` Implies First Satisfying Index
Informal description
For any list `xs` of elements of type `α`, predicate `p : α → Bool`, and natural number `i`, if `xs.findIdx? p = some i`, then there exists a proof `h` that `i` is a valid index of `xs` (i.e., `i < xs.length`), the element `xs[i]` satisfies `p`, and for all indices `j < i`, the element `xs[j]` does not satisfy `p`. In symbols: $$\text{findIdx?}_p(xs) = \text{some } i \rightarrow \exists h : i < \text{length}(xs), p(xs[i]) \land \forall j < i, \neg p(xs[j])$$
List.of_findIdx?_eq_none theorem
{xs : List α} {p : α → Bool} (w : xs.findIdx? p = none) : ∀ i : Nat, match xs[i]? with | some a => ¬p a | none => true
Full source
theorem of_findIdx?_eq_none {xs : List α} {p : α → Bool} (w : xs.findIdx? p = none) :
    ∀ i : Nat, match xs[i]? with | some a => ¬ p a | none => true := by
  intro i
  induction xs generalizing i with
  | nil => simp_all
  | cons x xs ih =>
    simp_all only [Bool.not_eq_true, findIdx?_cons, Nat.zero_add]
    cases i with
    | zero =>
      split at w <;> simp_all
    | succ i =>
      simp only [getElem?_cons_succ]
      apply ih
      split at w <;> simp_all
No Satisfying Index Implies All Elements Fail Predicate or Are Out of Bounds
Informal description
For any list `xs` of type `List α` and predicate `p : α → Bool`, if `findIdx? p xs = none`, then for every natural number index `i`, either the `i`-th element of `xs` does not satisfy `p` (i.e., `¬p a` when `xs[i]? = some a`), or the index `i` is out of bounds (i.e., `xs[i]? = none`).
List.findIdx?_of_eq_none abbrev
Full source
@[deprecated of_findIdx?_eq_none (since := "2025-02-02")]
abbrev findIdx?_of_eq_none := @of_findIdx?_eq_none
No Satisfying Index Implies All Elements Fail Predicate or Are Out of Bounds
Informal description
For any list `xs` of type `List α` and predicate `p : α → Bool`, if `findIdx? p xs = none`, then for every natural number index `i`, either the `i`-th element of `xs` does not satisfy `p` (i.e., `¬p a` when `xs[i]? = some a`), or the index `i` is out of bounds (i.e., `xs[i]? = none`).
List.findIdx?_map theorem
{f : β → α} {l : List β} : findIdx? p (l.map f) = l.findIdx? (p ∘ f)
Full source
@[simp] theorem findIdx?_map {f : β → α} {l : List β} : findIdx? p (l.map f) = l.findIdx? (p ∘ f) := by
  induction l with
  | nil => simp
  | cons x xs ih =>
    simp only [map_cons, findIdx?_cons]
    split <;> simp_all
Index Preservation under List Mapping: $\text{findIdx?}_p (\text{map } f\ l) = \text{findIdx?}_{p \circ f}\ l$
Informal description
For any function $f : \beta \to \alpha$, list $l : \text{List } \beta$, and predicate $p : \alpha \to \text{Bool}$, the index of the first element in the mapped list $\text{map } f\ l$ that satisfies $p$ is equal to the index of the first element in $l$ that satisfies $p \circ f$.
List.findIdx?_append theorem
: (xs ++ ys : List α).findIdx? p = (xs.findIdx? p).or ((ys.findIdx? p).map fun i => i + xs.length)
Full source
@[simp] theorem findIdx?_append :
    (xs ++ ys : List α).findIdx? p =
      (xs.findIdx? p).or ((ys.findIdx? p).map fun i => i + xs.length) := by
  induction xs with simp
  | cons _ _ _ => split <;> simp_all [Option.map_or', Option.map_map]; rfl
Index of First Satisfying Element in Concatenated Lists: $\text{findIdx?}_p(xs \mathbin{+\kern-1.5ex+} ys) = \text{findIdx?}_p(xs) \text{ or } \text{map } (+ \text{length}(xs)) (\text{findIdx?}_p(ys))$
Informal description
For any predicate $p : \alpha \to \text{Bool}$ and lists $xs, ys : \text{List } \alpha$, the index of the first element in the concatenated list $xs \mathbin{+\kern-1.5ex+} ys$ satisfying $p$ is equal to: - $\text{findIdx?}_p(xs)$ if it exists (i.e., is $\text{some } i$), otherwise - $\text{map } (\lambda i, i + \text{length}(xs)) (\text{findIdx?}_p(ys))$ if $\text{findIdx?}_p(ys)$ exists. In other words: \[ \text{findIdx?}_p(xs \mathbin{+\kern-1.5ex+} ys) = \text{findIdx?}_p(xs) \text{ or } \text{map } (\lambda i, i + \text{length}(xs)) (\text{findIdx?}_p(ys)) \] where $\text{or}$ is the non-short-circuiting choice operation on optional values.
List.findIdx?_flatten theorem
{l : List (List α)} {p : α → Bool} : l.flatten.findIdx? p = (l.findIdx? (·.any p)).map fun i => ((l.take i).map List.length).sum + (l[i]?.map fun xs => xs.findIdx p).getD 0
Full source
theorem findIdx?_flatten {l : List (List α)} {p : α → Bool} :
    l.flatten.findIdx? p =
      (l.findIdx? (·.any p)).map
        fun i => ((l.take i).map List.length).sum +
          (l[i]?.map fun xs => xs.findIdx p).getD 0 := by
  induction l with
  | nil => simp
  | cons xs l ih =>
    rw [flatten_cons, findIdx?_append, ih, findIdx?_cons]
    split <;> rename_i h
    · simp only [any_eq_true] at h
      rw [Option.or_of_isSome (by simp_all [findIdx?_isSome])]
      simp_all [findIdx?_eq_some_of_exists]
    · rw [Option.or_of_isNone (by simp_all [findIdx?_isNone])]
      simp [Function.comp_def, Nat.add_comm, Nat.add_assoc]
Index of First Satisfying Element in Flattened List of Lists
Informal description
For any list of lists $l : \text{List}(\text{List}\,\alpha)$ and predicate $p : \alpha \to \text{Bool}$, the index of the first element in the flattened list $\text{flatten}\,l$ that satisfies $p$ is equal to: - $\text{some}\left(\sum_{j < i} \text{length}(l_j) + k\right)$, where $i$ is the index of the first sublist in $l$ containing an element satisfying $p$, and $k$ is the index of the first such element within that sublist, if such a sublist exists; - $\text{none}$ otherwise. More precisely: \[ \text{findIdx?}_p(\text{flatten}\,l) = \begin{cases} \text{some}\left(\sum_{j < i} \text{length}(l_j) + k\right) & \text{if } \exists i, \text{any}\,p\,(l_i) \text{ and } k = \text{findIdx}_p(l_i), \\ \text{none} & \text{otherwise.} \end{cases} \]
List.findIdx?_replicate theorem
: (replicate n a).findIdx? p = if 0 < n ∧ p a then some 0 else none
Full source
@[simp] theorem findIdx?_replicate :
    (replicate n a).findIdx? p = if 0 < n ∧ p a then some 0 else none := by
  cases n with
  | zero => simp
  | succ n =>
    simp only [replicate, findIdx?_cons, Nat.zero_add, zero_lt_succ, true_and]
    split <;> simp_all
First Satisfying Index in Replicated List: $\text{findIdx?}_p(\text{replicate}(n, a)) = \text{if } (0 < n \land p(a)) \text{ then some } 0 \text{ else none}$
Informal description
For any natural number $n$, element $a$ of type $\alpha$, and predicate $p : \alpha \to \text{Bool}$, the index of the first element in the list $\text{replicate}(n, a)$ that satisfies $p$ is $\text{some } 0$ if $n > 0$ and $p(a)$ holds, otherwise it is $\text{none}$.
List.findIdx?_eq_findSome?_zipIdx theorem
{xs : List α} {p : α → Bool} : xs.findIdx? p = xs.zipIdx.findSome? fun ⟨a, i⟩ => if p a then some i else none
Full source
theorem findIdx?_eq_findSome?_zipIdx {xs : List α} {p : α → Bool} :
    xs.findIdx? p = xs.zipIdx.findSome? fun ⟨a, i⟩ => if p a then some i else none := by
  induction xs with
  | nil => simp
  | cons x xs ih =>
    simp only [findIdx?_cons, Nat.zero_add, zipIdx]
    split
    · simp_all
    · simp_all only [zipIdx_cons, ite_false, Option.isNone_none, findSome?_cons_of_isNone, reduceCtorEq]
      rw [← map_snd_add_zipIdx_eq_zipIdx (n := 1) (k := 0)]
      simp [Function.comp_def, findSome?_map]
Equivalence of List Index Search via Zipped Indices
Informal description
For any list `xs` of elements of type `alpha` and any predicate `p : alpha → Bool`, the index of the first element in `xs` satisfying `p` (computed by `xs.findIdx? p`) is equal to the first non-`none` result obtained by applying the function `fun ⟨a, i⟩ => if p a then some i else none` to the list `xs.zipIdx` (which pairs each element of `xs` with its index). In other words, the following equality holds: $$\text{findIdx?}\ p\ xs = \text{findSome?}\ (\lambda \langle a, i \rangle, \text{if } p(a) \text{ then some } i \text{ else none})\ (\text{zipIdx}\ xs)$$
List.findIdx?_eq_fst_find?_zipIdx theorem
{xs : List α} {p : α → Bool} : xs.findIdx? p = (xs.zipIdx.find? fun ⟨x, _⟩ => p x).map (·.2)
Full source
theorem findIdx?_eq_fst_find?_zipIdx {xs : List α} {p : α → Bool} :
    xs.findIdx? p = (xs.zipIdx.find? fun ⟨x, _⟩ => p x).map (·.2) := by
  induction xs with
  | nil => simp
  | cons x xs ih =>
    simp only [findIdx?_cons, Nat.zero_add, zipIdx_cons]
    split
    · simp_all
    · rw [ih, ← map_snd_add_zipIdx_eq_zipIdx (n := 1) (k := 0)]
      simp [Function.comp_def, *]
Equivalence of Index Search via Zipped List: $\text{findIdx? } p \ xs = \text{map } \pi_2 (\text{find? } (p \circ \pi_1) \ (\text{zipIdx } xs))$
Informal description
For any list $xs$ of elements of type $\alpha$ and any predicate $p : \alpha \to \text{Bool}$, the index of the first element in $xs$ satisfying $p$ (computed by $\text{findIdx? } p \ xs$) is equal to the second component of the first pair $(x, i)$ in the list $\text{zipIdx } xs$ (which pairs each element with its index) for which $p(x)$ holds, if such a pair exists (otherwise both sides evaluate to $\text{none}$). Formally: \[ \text{findIdx? } p \ xs = \text{map } (\lambda (x, i), i) \ (\text{find? } (\lambda (x, \_), p x) \ (\text{zipIdx } xs)) \]
List.findIdx?_eq_none_of_findIdx?_eq_none theorem
{xs : List α} {p q : α → Bool} (w : ∀ x ∈ xs, p x → q x) : xs.findIdx? q = none → xs.findIdx? p = none
Full source
theorem findIdx?_eq_none_of_findIdx?_eq_none {xs : List α} {p q : α → Bool} (w : ∀ x ∈ xs, p x → q x) :
    xs.findIdx? q = none → xs.findIdx? p = none := by
  simp only [findIdx?_eq_none_iff]
  intro h x m
  cases z : p x
  · rfl
  · exfalso
    specialize w x m z
    specialize h x m
    simp_all
Implication of Predicate Absence in List Index Search
Informal description
For any list `xs` of elements of type `α` and any predicates `p, q : α → Bool`, if for every element `x ∈ xs` we have that `p x` implies `q x`, then the absence of an index satisfying `q` (i.e., `findIdx? q xs = none`) implies the absence of an index satisfying `p` (i.e., `findIdx? p xs = none`).
List.Sublist.findIdx?_isSome theorem
{l₁ l₂ : List α} (h : l₁ <+ l₂) : (l₁.findIdx? p).isSome → (l₂.findIdx? p).isSome
Full source
theorem Sublist.findIdx?_isSome {l₁ l₂ : List α} (h : l₁ <+ l₂) :
    (l₁.findIdx? p).isSome → (l₂.findIdx? p).isSome := by
  simp only [List.findIdx?_isSome, any_eq_true]
  rintro ⟨w, m, q⟩
  exact ⟨w, h.mem m, q⟩
Sublist Preserves Existence of Satisfying Index: $(l_1 <+ l_2) \to (l_1.\text{findIdx?}\ p \neq \text{none}) \to (l_2.\text{findIdx?}\ p \neq \text{none})$
Informal description
For any lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a sublist of $l_2$ (denoted $l_1 <+ l_2$) and there exists an index in $l_1$ satisfying the predicate $p$ (i.e., $(l_1.\text{findIdx?}\ p).\text{isSome}$ holds), then there also exists an index in $l_2$ satisfying $p$ (i.e., $(l_2.\text{findIdx?}\ p).\text{isSome}$ holds).
List.Sublist.findIdx?_eq_none theorem
{l₁ l₂ : List α} (h : l₁ <+ l₂) : l₂.findIdx? p = none → l₁.findIdx? p = none
Full source
theorem Sublist.findIdx?_eq_none {l₁ l₂ : List α} (h : l₁ <+ l₂) :
    l₂.findIdx? p = none → l₁.findIdx? p = none := by
  simp only [findIdx?_eq_none_iff]
  exact fun w x m => w x (h.mem m)
Sublist Preservation of Nonexistence in `findIdx?`
Informal description
For any lists $l₁$ and $l₂$ of elements of type $\alpha$, if $l₁$ is a sublist of $l₂$ and the predicate $p$ does not hold for any element in $l₂$ (i.e., $l₂.\text{findIdx?}\ p = \text{none}$), then $p$ also does not hold for any element in $l₁$ (i.e., $l₁.\text{findIdx?}\ p = \text{none}$).
List.IsPrefix.findIdx?_eq_some theorem
{l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂) : List.findIdx? p l₁ = some i → List.findIdx? p l₂ = some i
Full source
theorem IsPrefix.findIdx?_eq_some {l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂) :
    List.findIdx? p l₁ = some i → List.findIdx? p l₂ = some i := by
  rw [IsPrefix] at h
  obtain ⟨t, rfl⟩ := h
  intro h
  simp [findIdx?_append, h]
Prefix Preservation of First Satisfying Index in `findIdx?`
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$ and any predicate $p : \alpha \to \text{Bool}$, if $l_1$ is a prefix of $l_2$ and the first index where $p$ holds in $l_1$ is $i$ (i.e., $\text{findIdx?}\ p\ l_1 = \text{some}\ i$), then the first index where $p$ holds in $l_2$ is also $i$ (i.e., $\text{findIdx?}\ p\ l_2 = \text{some}\ i$).
List.IsPrefix.findIdx?_eq_none theorem
{l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂) : List.findIdx? p l₂ = none → List.findIdx? p l₁ = none
Full source
theorem IsPrefix.findIdx?_eq_none {l₁ l₂ : List α} {p : α → Bool} (h : l₁ <+: l₂) :
    List.findIdx? p l₂ = noneList.findIdx? p l₁ = none :=
  h.sublist.findIdx?_eq_none
Prefix Preservation of Nonexistence in `findIdx?`
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a prefix of $l_2$ and the predicate $p$ does not hold for any element in $l_2$ (i.e., $\text{findIdx?}\ p\ l_2 = \text{none}$), then $p$ also does not hold for any element in $l_1$ (i.e., $\text{findIdx?}\ p\ l_1 = \text{none}$).
List.IsSuffix.findIdx?_eq_none theorem
{l₁ l₂ : List α} {p : α → Bool} (h : l₁ <:+ l₂) : List.findIdx? p l₂ = none → List.findIdx? p l₁ = none
Full source
theorem IsSuffix.findIdx?_eq_none {l₁ l₂ : List α} {p : α → Bool} (h : l₁ <:+ l₂) :
    List.findIdx? p l₂ = noneList.findIdx? p l₁ = none :=
  h.sublist.findIdx?_eq_none
Suffix Preservation of Nonexistence in `findIdx?`
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$ and any predicate $p : \alpha \to \text{Bool}$, if $l_1$ is a suffix of $l_2$ and no element in $l_2$ satisfies $p$ (i.e., $\text{findIdx?}\ p\ l_2 = \text{none}$), then no element in $l_1$ satisfies $p$ either (i.e., $\text{findIdx?}\ p\ l_1 = \text{none}$).
List.IsInfix.findIdx?_eq_none theorem
{l₁ l₂ : List α} {p : α → Bool} (h : l₁ <:+: l₂) : List.findIdx? p l₂ = none → List.findIdx? p l₁ = none
Full source
theorem IsInfix.findIdx?_eq_none {l₁ l₂ : List α} {p : α → Bool} (h : l₁ <:+: l₂) :
    List.findIdx? p l₂ = noneList.findIdx? p l₁ = none :=
  h.sublist.findIdx?_eq_none
Nonexistence Preservation of Predicate in Infix Lists for `findIdx?`
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is an infix of $l_2$ (i.e., $l_1$ appears as a contiguous subsequence in $l_2$) and the predicate $p$ does not hold for any element in $l_2$ (i.e., $\text{findIdx?}\ p\ l_2 = \text{none}$), then $p$ also does not hold for any element in $l_1$ (i.e., $\text{findIdx?}\ p\ l_1 = \text{none}$).
List.findIdx_eq_getD_findIdx? theorem
{xs : List α} {p : α → Bool} : xs.findIdx p = (xs.findIdx? p).getD xs.length
Full source
theorem findIdx_eq_getD_findIdx? {xs : List α} {p : α → Bool} :
    xs.findIdx p = (xs.findIdx? p).getD xs.length := by
  induction xs with
  | nil => simp
  | cons x xs ih =>
    simp only [findIdx_cons, findIdx?_cons]
    split <;> simp_all [ih]
Equivalence of `findIdx` and `findIdx?` with Default Length
Informal description
For any list `xs` of elements of type `α` and any predicate `p : α → Bool`, the index of the first element in `xs` satisfying `p` (or the length of `xs` if no such element exists) is equal to the default value obtained from `xs.findIdx? p` with default value `xs.length`. In other words, `findIdx p xs = (findIdx? p xs).getD (length xs)`.
List.findIdx?_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → Bool} {g : α → Bool} (hf : ∀ x h, f ⟨x, h⟩ = g x) : l.findIdx? f = l.unattach.findIdx? g
Full source
@[simp] theorem findIdx?_subtype {p : α → Prop} {l : List { x // p x }}
    {f : { x // p x }Bool} {g : α → Bool} (hf : ∀ x h, f ⟨x, h⟩ = g x) :
    l.findIdx? f = l.unattach.findIdx? g := by
  unfold unattach
  induction l with
  | nil => simp
  | cons a l ih =>
    simp [hf, findIdx?_cons]
    split <;> simp [ih, Function.comp_def]
Equivalence of Index Search Between Subtype and Unattached List
Informal description
Let $p : \alpha \to \text{Prop}$ be a predicate, $l$ a list of elements of the subtype $\{x : \alpha \mid p x\}$, and $f : \{x : \alpha \mid p x\} \to \text{Bool}$ and $g : \alpha \to \text{Bool}$ two predicates such that for any $x$ and proof $h$ of $p x$, we have $f \langle x, h \rangle = g x$. Then the index of the first element in $l$ satisfying $f$ is equal to the index of the first element in the unattached list $l$ (viewed as a list of $\alpha$) satisfying $g$.
List.findFinIdx?_nil theorem
{p : α → Bool} : findFinIdx? p [] = none
Full source
@[simp] theorem findFinIdx?_nil {p : α → Bool} : findFinIdx? p [] = none := rfl
`findFinIdx?` on empty list is `none`
Informal description
For any Boolean predicate $p$ on elements of type $\alpha$, the function `findFinIdx?` applied to $p$ and the empty list `[]` returns `none`.
List.findIdx?_go_eq_map_findFinIdx?_go_val theorem
{xs : List α} {p : α → Bool} {i : Nat} {h} : List.findIdx?.go p xs i = (List.findFinIdx?.go p l xs i h).map (·.val)
Full source
theorem findIdx?_go_eq_map_findFinIdx?_go_val {xs : List α} {p : α → Bool} {i : Nat} {h} :
    List.findIdx?.go p xs i =
      (List.findFinIdx?.go p l xs i h).map (·.val) := by
  unfold findIdx?.go
  unfold findFinIdx?.go
  split
  · simp_all
  · simp only
    split
    · simp
    · rw [findIdx?_go_eq_map_findFinIdx?_go_val]
Equivalence of `findIdx?.go` and `map`-composed `findFinIdx?.go` auxiliary functions
Informal description
For any list `xs` of elements of type `α`, any predicate `p : α → Bool`, any natural number index `i`, and any proof `h`, the auxiliary function `findIdx?.go` applied to `p`, `xs`, and `i` is equal to the composition of the `map` function with the value projection applied to the auxiliary function `findFinIdx?.go` with the same arguments and proof `h`. In mathematical notation: Let $xs$ be a list of elements of type $\alpha$, $p : \alpha \to \text{Bool}$ a predicate, $i \in \mathbb{N}$ an index, and $h$ a proof. Then: $$\text{findIdx?.go}(p, xs, i) = \text{map}(\cdot.\text{val}, \text{findFinIdx?.go}(p, l, xs, i, h))$$
List.findIdx?_eq_map_findFinIdx?_val theorem
{xs : List α} {p : α → Bool} : xs.findIdx? p = (xs.findFinIdx? p).map (·.val)
Full source
theorem findIdx?_eq_map_findFinIdx?_val {xs : List α} {p : α → Bool} :
    xs.findIdx? p = (xs.findFinIdx? p).map (·.val) := by
  simp [findIdx?, findFinIdx?]
  rw [findIdx?_go_eq_map_findFinIdx?_go_val]
Equivalence of Unbounded and Bounded First Index Search: $\text{findIdx?} = \text{map}(\cdot.\text{val}) \circ \text{findFinIdx?}$
Informal description
For any list $xs$ of elements of type $\alpha$ and any predicate $p : \alpha \to \text{Bool}$, the result of finding the first index where $p$ holds in $xs$ (as an unbounded natural number) is equal to the result of mapping the value projection function over the result of finding the first index where $p$ holds in $xs$ (as a bounded natural number in $\text{Fin}(xs.\text{length})$). In mathematical notation: $$\text{findIdx?}(p, xs) = \text{map}(\cdot.\text{val}, \text{findFinIdx?}(p, xs))$$
List.findFinIdx?_eq_pmap_findIdx? theorem
{xs : List α} {p : α → Bool} : xs.findFinIdx? p = (xs.findIdx? p).pmap (fun i m => by simp [findIdx?_eq_some_iff_getElem] at m; exact ⟨i, m.choose⟩) (fun i h => h)
Full source
theorem findFinIdx?_eq_pmap_findIdx? {xs : List α} {p : α → Bool} :
    xs.findFinIdx? p =
      (xs.findIdx? p).pmap
        (fun i m => by simp [findIdx?_eq_some_iff_getElem] at m; exact ⟨i, m.choose⟩)
        (fun i h => h) := by
  simp [findIdx?_eq_map_findFinIdx?_val, Option.pmap_map]
Equivalence of Bounded and Unbounded Index Search via Partial Map
Informal description
For any list `xs` of elements of type `α` and any predicate `p : α → Bool`, the bounded index search function `findFinIdx?` is equal to a partial map applied to the unbounded index search function `findIdx?`. Specifically, when `findIdx? p xs` returns `some i`, the partial map verifies that `i` is a valid index (i.e., `i < xs.length`) and returns it as a bounded natural number in `Fin (xs.length)`. In mathematical notation: $$\text{findFinIdx?}(p, xs) = \text{pmap}\left(\lambda i\ m, \langle i, \text{proof}\rangle\right) (\text{findIdx?}(p, xs))$$ where the proof verifies that `i` is a valid index when `findIdx?` returns `some i`.
List.findFinIdx?_cons theorem
{p : α → Bool} {x : α} {xs : List α} : findFinIdx? p (x :: xs) = if p x then some 0 else (findFinIdx? p xs).map Fin.succ
Full source
@[simp] theorem findFinIdx?_cons {p : α → Bool} {x : α} {xs : List α} :
    findFinIdx? p (x :: xs) = if p x then some 0 else (findFinIdx? p xs).map Fin.succ := by
  rw [← Option.map_inj_right (f := Fin.val) (fun a b => Fin.eq_of_val_eq)]
  rw [← findIdx?_eq_map_findFinIdx?_val]
  rw [findIdx?_cons]
  split
  · simp
  · rw [findIdx?_eq_map_findFinIdx?_val]
    simp [Function.comp_def]
Recursive computation of first satisfying index in a list: $\text{findFinIdx?}_p(x :: xs) = \text{if } p(x) \text{ then some } 0 \text{ else map Fin.succ } (\text{findFinIdx?}_p xs)$
Informal description
For any predicate $p : \alpha \to \text{Bool}$, element $x : \alpha$, and list $xs : \text{List } \alpha$, the first index (as a bounded natural number) where $p$ holds in the list $x :: xs$ is equal to $\text{some } 0$ if $p(x)$ holds, otherwise it is equal to mapping the successor function over the first index where $p$ holds in $xs$. In mathematical notation: $$\text{findFinIdx?}(p, x :: xs) = \begin{cases} \text{some } 0 & \text{if } p(x) \\ \text{map Fin.succ } (\text{findFinIdx?}(p, xs)) & \text{otherwise} \end{cases}$$
List.findFinIdx?_eq_none_iff theorem
{l : List α} {p : α → Bool} : l.findFinIdx? p = none ↔ ∀ x ∈ l, ¬p x
Full source
@[simp] theorem findFinIdx?_eq_none_iff {l : List α} {p : α → Bool} :
    l.findFinIdx? p = none ↔ ∀ x ∈ l, ¬ p x := by
  simp [findFinIdx?_eq_pmap_findIdx?]
Nonexistence of Satisfying Elements in List for Predicate $p$
Informal description
For any list $l$ of elements of type $\alpha$ and any predicate $p : \alpha \to \text{Bool}$, the bounded index search function $\text{findFinIdx?}$ returns $\text{none}$ if and only if no element $x$ in $l$ satisfies $p(x)$. In symbols: $$\text{findFinIdx?}(p, l) = \text{none} \leftrightarrow \forall x \in l, \neg p(x)$$
List.findFinIdx?_eq_some_iff theorem
{xs : List α} {p : α → Bool} {i : Fin xs.length} : xs.findFinIdx? p = some i ↔ p xs[i] ∧ ∀ j (hji : j < i), ¬p (xs[j]'(Nat.lt_trans hji i.2))
Full source
@[simp]
theorem findFinIdx?_eq_some_iff {xs : List α} {p : α → Bool} {i : Fin xs.length} :
    xs.findFinIdx? p = some i ↔
      p xs[i] ∧ ∀ j (hji : j < i), ¬p (xs[j]'(Nat.lt_trans hji i.2)) := by
  simp only [findFinIdx?_eq_pmap_findIdx?, Option.pmap_eq_some_iff, findIdx?_eq_some_iff_getElem,
    Bool.not_eq_true, Option.mem_def, exists_and_left, and_exists_self, Fin.getElem_fin]
  constructor
  · rintro ⟨a, ⟨h, w₁, w₂⟩, rfl⟩
    exact ⟨w₁, fun j hji => by simpa using w₂ j hji⟩
  · rintro ⟨h, w⟩
    exact ⟨i, ⟨i.2, h, fun j hji => w ⟨j, by omega⟩ hji⟩, rfl⟩
Characterization of First Satisfying Index in List: $\text{findFinIdx?}(p, xs) = \text{some } i \leftrightarrow p(xs[i]) \land (\forall j < i, \neg p(xs[j]))$
Informal description
For a list `xs` of elements of type `α`, a predicate `p : α → Bool`, and an index `i` (as a bounded natural number with `i < xs.length`), the bounded index search function `findFinIdx?` returns `some i` if and only if: 1. The predicate `p` holds for the element `xs[i]`, and 2. For all indices `j < i`, the predicate `p` does not hold for the element `xs[j]`. In mathematical notation: $$\text{findFinIdx?}(p, xs) = \text{some } i \leftrightarrow p(xs[i]) \land (\forall j < i, \neg p(xs[j]))$$
List.findFinIdx?_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → Bool} {g : α → Bool} (hf : ∀ x h, f ⟨x, h⟩ = g x) : l.findFinIdx? f = (l.unattach.findFinIdx? g).map (fun i => i.cast (by simp))
Full source
@[simp] theorem findFinIdx?_subtype {p : α → Prop} {l : List { x // p x }}
    {f : { x // p x }Bool} {g : α → Bool} (hf : ∀ x h, f ⟨x, h⟩ = g x) :
    l.findFinIdx? f = (l.unattach.findFinIdx? g).map (fun i => i.cast (by simp)) := by
  unfold unattach
  induction l with
  | nil => simp
  | cons a l ih =>
    simp [hf, findFinIdx?_cons]
    split <;> simp [ih, Function.comp_def]
First Satisfying Index Preservation under Subtype and Predicate Transformation
Informal description
Let $p : \alpha \to \text{Prop}$ be a predicate, $l$ a list of elements of subtype $\{x \mid p x\}$, and $f : \{x \mid p x\} \to \text{Bool}$ and $g : \alpha \to \text{Bool}$ boolean predicates such that $f(\langle x, h\rangle) = g(x)$ for all $x$ and proof $h$ of $p(x)$. Then the first index (as a bounded natural number) where $f$ holds in $l$ is equal to the first index where $g$ holds in the unattached list $l$, composed with a cast adjusting the bound.
List.idxOf_cons theorem
[BEq α] : (x :: xs : List α).idxOf y = bif x == y then 0 else xs.idxOf y + 1
Full source
theorem idxOf_cons [BEq α] :
    (x :: xs : List α).idxOf y = bif x == y then 0 else xs.idxOf y + 1 := by
  dsimp [idxOf]
  simp [findIdx_cons]
Index of First Occurrence in a Cons List: $\text{idxOf}_y(x :: xs) = \text{if } x == y \text{ then } 0 \text{ else } (\text{idxOf}_y xs) + 1$
Informal description
For any type $\alpha$ with a boolean equality relation `==`, and for any elements $x, y \in \alpha$ and list $xs : \text{List } \alpha$, the index of the first occurrence of $y$ in the list $x :: xs$ is equal to $0$ if $x == y$ is true, otherwise it is equal to $1$ plus the index of the first occurrence of $y$ in $xs$.
List.indexOf_cons abbrev
Full source
@[deprecated idxOf_cons (since := "2025-01-29")]
abbrev indexOf_cons := @idxOf_cons
Index of First Occurrence in a Cons List: $\text{indexOf}_y(x :: xs) = \text{if } x == y \text{ then } 0 \text{ else } (\text{indexOf}_y xs) + 1$
Informal description
For any type $\alpha$ with a boolean equality relation `==`, and for any elements $x, y \in \alpha$ and list $xs : \text{List } \alpha$, the index of the first occurrence of $y$ in the list $x :: xs$ is equal to $0$ if $x == y$ is true, otherwise it is equal to $1$ plus the index of the first occurrence of $y$ in $xs$.
List.idxOf_cons_self theorem
[BEq α] [ReflBEq α] {l : List α} : (a :: l).idxOf a = 0
Full source
@[simp] theorem idxOf_cons_self [BEq α] [ReflBEq α] {l : List α} : (a :: l).idxOf a = 0 := by
  simp [idxOf_cons]
Index of First Occurrence in a Cons List with Reflexive Equality: $\text{idxOf}_a(a :: l) = 0$
Informal description
For any type $\alpha$ with a reflexive boolean equality relation `==`, and for any element $a \in \alpha$ and list $l : \text{List } \alpha$, the index of the first occurrence of $a$ in the list $a :: l$ is $0$.
List.indexOf_cons_self abbrev
Full source
@[deprecated idxOf_cons_self (since := "2025-01-29")]
abbrev indexOf_cons_self := @idxOf_cons_self
First Occurrence Index in Cons List with Reflexive Equality: $\text{indexOf}_a(a :: l) = 0$
Informal description
For any type $\alpha$ with a reflexive boolean equality relation `==`, and for any element $a \in \alpha$ and list $l : \text{List } \alpha$, the index of the first occurrence of $a$ in the list $a :: l$ is $0$.
List.idxOf_append theorem
[BEq α] [LawfulBEq α] {l₁ l₂ : List α} {a : α} : (l₁ ++ l₂).idxOf a = if a ∈ l₁ then l₁.idxOf a else l₂.idxOf a + l₁.length
Full source
theorem idxOf_append [BEq α] [LawfulBEq α] {l₁ l₂ : List α} {a : α} :
    (l₁ ++ l₂).idxOf a = if a ∈ l₁ then l₁.idxOf a else l₂.idxOf a + l₁.length := by
  rw [idxOf, findIdx_append]
  split <;> rename_i h
  · rw [if_pos]
    simpa using h
  · rw [if_neg]
    simpa using h
Index of First Occurrence in Concatenated Lists: $\text{idxOf}_a(l_1 \mathbin{+\kern-1.5ex+} l_2) = \text{if } a \in l_1 \text{ then } \text{idxOf}_a(l_1) \text{ else } \text{idxOf}_a(l_2) + \text{length}(l_1)$
Informal description
For any type $\alpha$ with a lawful boolean equality relation, and for any lists $l_1, l_2$ of elements of type $\alpha$, the index of the first occurrence of an element $a$ in the concatenated list $l_1 \mathbin{+\kern-1.5ex+} l_2$ is equal to the index of $a$ in $l_1$ if $a$ appears in $l_1$, otherwise it is equal to the length of $l_1$ plus the index of $a$ in $l_2$.
List.indexOf_append abbrev
Full source
@[deprecated idxOf_append (since := "2025-01-29")]
abbrev indexOf_append := @idxOf_append
Index of First Occurrence in Concatenated Lists: $\text{idxOf}_a(l_1 \mathbin{+\kern-1.5ex+} l_2) = \text{if } a \in l_1 \text{ then } \text{idxOf}_a(l_1) \text{ else } \text{idxOf}_a(l_2) + \text{length}(l_1)$
Informal description
For any type $\alpha$ with a lawful boolean equality relation, and for any lists $l_1, l_2$ of elements of type $\alpha$, the index of the first occurrence of an element $a$ in the concatenated list $l_1 \mathbin{+\kern-1.5ex+} l_2$ is equal to the index of $a$ in $l_1$ if $a$ appears in $l_1$, otherwise it is equal to the length of $l_1$ plus the index of $a$ in $l_2$.
List.idxOf_eq_length theorem
[BEq α] [LawfulBEq α] {l : List α} (h : a ∉ l) : l.idxOf a = l.length
Full source
theorem idxOf_eq_length [BEq α] [LawfulBEq α] {l : List α} (h : a ∉ l) : l.idxOf a = l.length := by
  induction l with
  | nil => rfl
  | cons x xs ih =>
    simp only [mem_cons, not_or] at h
    simp only [idxOf_cons, cond_eq_if, beq_iff_eq]
    split <;> simp_all
Index of Absent Element Equals List Length
Informal description
For any type $\alpha$ with a lawful boolean equality relation, and for any list $l$ of elements of type $\alpha$, if an element $a$ is not in $l$, then the index of the first occurrence of $a$ in $l$ is equal to the length of $l$.
List.indexOf_eq_length abbrev
Full source
@[deprecated idxOf_eq_length (since := "2025-01-29")]
abbrev indexOf_eq_length := @idxOf_eq_length
Index of Absent Element Equals List Length
Informal description
For any type $\alpha$ with a lawful boolean equality relation, and for any list $l$ of elements of type $\alpha$, if an element $a$ is not in $l$, then the index of the first occurrence of $a$ in $l$ is equal to the length of $l$.
List.idxOf_lt_length theorem
[BEq α] [LawfulBEq α] {l : List α} (h : a ∈ l) : l.idxOf a < l.length
Full source
theorem idxOf_lt_length [BEq α] [LawfulBEq α] {l : List α} (h : a ∈ l) : l.idxOf a < l.length := by
  induction l with
  | nil => simp at h
  | cons x xs ih =>
    simp only [mem_cons] at h
    obtain rfl | h := h
    · simp
    · simp only [idxOf_cons, cond_eq_if, beq_iff_eq, length_cons]
      specialize ih h
      split
      · exact zero_lt_succ xs.length
      · exact Nat.add_lt_add_right ih 1
First Occurrence Index is Less Than List Length
Informal description
For any type $\alpha$ with a lawful boolean equality relation, and for any list $l$ of elements of type $\alpha$, if an element $a$ is in $l$, then the index of the first occurrence of $a$ in $l$ is strictly less than the length of $l$.
List.indexOf_lt_length abbrev
Full source
@[deprecated idxOf_lt_length (since := "2025-01-29")]
abbrev indexOf_lt_length := @idxOf_lt_length
First Occurrence Index is Less Than List Length
Informal description
For any type $\alpha$ with a lawful boolean equality relation, and for any list $l$ of elements of type $\alpha$, if an element $a$ is in $l$, then the index of the first occurrence of $a$ in $l$ is strictly less than the length of $l$.
List.idxOf?_eq_map_finIdxOf?_val theorem
[BEq α] {xs : List α} {a : α} : xs.idxOf? a = (xs.finIdxOf? a).map (·.val)
Full source
theorem idxOf?_eq_map_finIdxOf?_val [BEq α] {xs : List α} {a : α} :
    xs.idxOf? a = (xs.finIdxOf? a).map (·.val) := by
  simp [idxOf?, finIdxOf?, findIdx?_eq_map_findFinIdx?_val]
Equivalence of Unbounded and Bounded First Index Search: $\text{idxOf?} = \text{map}(\cdot.\text{val}) \circ \text{finIdxOf?}$
Informal description
For any list $xs$ of elements of type $\alpha$ with a boolean equality relation `==`, and any element $a \in \alpha$, the result of finding the first index of $a$ in $xs$ (as an unbounded natural number) is equal to the result of mapping the value projection function over the result of finding the first index of $a$ in $xs$ (as a bounded natural number in $\text{Fin}(xs.\text{length})$). In mathematical notation: $$\text{idxOf?}(a, xs) = \text{map}(\cdot.\text{val}, \text{finIdxOf?}(a, xs))$$
List.finIdxOf?_nil theorem
[BEq α] : ([] : List α).finIdxOf? a = none
Full source
@[simp] theorem finIdxOf?_nil [BEq α] : ([] : List α).finIdxOf? a = none := rfl
`finIdxOf?` returns `none` on empty list
Informal description
For any type $\alpha$ with a boolean equality relation and any element $a \in \alpha$, the function `finIdxOf?` applied to the empty list returns `none`. That is, $\text{finIdxOf?} \, a \, [] = \text{none}$.
List.finIdxOf?_cons theorem
[BEq α] {a : α} {xs : List α} : (a :: xs).finIdxOf? b = if a == b then some ⟨0, by simp⟩ else (xs.finIdxOf? b).map (·.succ)
Full source
@[simp] theorem finIdxOf?_cons [BEq α] {a : α} {xs : List α} :
    (a :: xs).finIdxOf? b =
      if a == b then some ⟨0, by simp⟩ else (xs.finIdxOf? b).map (·.succ) := by
  simp [finIdxOf?]
First Occurrence Index in Cons List: $\text{finIdxOf?}(b, a :: xs) = \text{if } a == b \text{ then some } 0 \text{ else map Fin.succ } (\text{finIdxOf?}(b, xs))$
Informal description
For any type $\alpha$ with a boolean equality relation `==`, any element $a \in \alpha$, and any list $xs$ of elements of type $\alpha$, the first occurrence of an element $b$ in the list $a :: xs$ is given by: $$\text{finIdxOf?}(b, a :: xs) = \begin{cases} \text{some } 0 & \text{if } a == b \\ \text{map Fin.succ } (\text{finIdxOf?}(b, xs)) & \text{otherwise} \end{cases}$$ where $\text{Fin.succ}$ increments the index while preserving the bound.
List.finIdxOf?_eq_none_iff theorem
[BEq α] [LawfulBEq α] {l : List α} {a : α} : l.finIdxOf? a = none ↔ a ∉ l
Full source
@[simp] theorem finIdxOf?_eq_none_iff [BEq α] [LawfulBEq α] {l : List α} {a : α} :
    l.finIdxOf? a = none ↔ a ∉ l := by
  simp only [finIdxOf?, findFinIdx?_eq_none_iff, beq_iff_eq]
  constructor
  · intro w m
    exact w a m rfl
  · rintro h a m rfl
    exact h m
Characterization of Absence in List via `finIdxOf?`
Informal description
For a list $l$ of elements of type $\alpha$ with a lawful boolean equality relation, and an element $a$ of type $\alpha$, the function `finIdxOf?` returns `none` if and only if $a$ does not appear in $l$.
List.finIdxOf?_eq_some_iff theorem
[BEq α] [LawfulBEq α] {l : List α} {a : α} {i : Fin l.length} : l.finIdxOf? a = some i ↔ l[i] = a ∧ ∀ j (_ : j < i), ¬l[j] = a
Full source
@[simp] theorem finIdxOf?_eq_some_iff [BEq α] [LawfulBEq α] {l : List α} {a : α} {i : Fin l.length} :
    l.finIdxOf? a = some i ↔ l[i] = a ∧ ∀ j (_ : j < i), ¬l[j] = a := by
  simp only [finIdxOf?, findFinIdx?_eq_some_iff, beq_iff_eq]
Characterization of First Occurrence in List via $\mathrm{finIdxOf?}$
Informal description
For a list $l$ of elements of type $\alpha$ with a lawful boolean equality relation, an element $a$ of type $\alpha$, and a finite index $i$ bounded by the length of $l$, the function $\mathrm{finIdxOf?}$ returns $\mathrm{some}\,i$ if and only if the element at index $i$ in $l$ equals $a$ and no element at any smaller index equals $a$.
List.idxOf?_nil theorem
[BEq α] : ([] : List α).idxOf? a = none
Full source
@[simp] theorem idxOf?_nil [BEq α] : ([] : List α).idxOf? a = none := rfl
Empty List Has No Index for Any Element
Informal description
For any type $\alpha$ with a boolean equality relation and any element $a$ of type $\alpha$, the first occurrence index of $a$ in the empty list is `none`.
List.idxOf?_cons theorem
[BEq α] {a : α} {xs : List α} {b : α} : (a :: xs).idxOf? b = if a == b then some 0 else (xs.idxOf? b).map (· + 1)
Full source
theorem idxOf?_cons [BEq α] {a : α} {xs : List α} {b : α} :
    (a :: xs).idxOf? b = if a == b then some 0 else (xs.idxOf? b).map (· + 1) := by
  simp [idxOf?]
Recursive computation of first index in a list: $\text{idxOf?}_b(a :: xs) = \text{if } a == b \text{ then some } 0 \text{ else map } (\lambda i, i + 1) (\text{idxOf?}_b xs)$
Informal description
For any type $\alpha$ with a boolean equality relation `==`, an element $b : \alpha$, and a list $a :: xs$ of type $\text{List } \alpha$, the index of the first occurrence of $b$ in $a :: xs$ is equal to $\text{some } 0$ if $a == b$ holds, otherwise it is equal to $\text{map } (\lambda i, i + 1) (\text{idxOf? } b \text{ } xs)$.
List.idxOf?_eq_none_iff theorem
[BEq α] [LawfulBEq α] {l : List α} {a : α} : l.idxOf? a = none ↔ a ∉ l
Full source
@[simp] theorem idxOf?_eq_none_iff [BEq α] [LawfulBEq α] {l : List α} {a : α} :
    l.idxOf? a = none ↔ a ∉ l := by
  simp only [idxOf?, findIdx?_eq_none_iff, beq_eq_false_iff_ne, ne_eq]
  constructor
  · intro w h
    specialize w _ h
    simp at w
  · rintro w x h rfl
    contradiction
Nonexistence of Element in List via `idxOf?`
Informal description
For a type $\alpha$ with a lawful boolean equality relation `==`, a list $l$ of elements of type $\alpha$, and an element $a$ of type $\alpha$, the function `idxOf?` returns `none` if and only if $a$ does not appear in $l$. In other words, $l.\text{idxOf?}\ a = \text{none} \leftrightarrow a \notin l$.
List.indexOf?_eq_none_iff abbrev
Full source
@[deprecated idxOf?_eq_none_iff (since := "2025-01-29")]
abbrev indexOf?_eq_none_iff := @idxOf?_eq_none_iff
Nonexistence of Element in List via `indexOf?`
Informal description
For a type $\alpha$ with a lawful boolean equality relation `==`, a list $l$ of elements of type $\alpha$, and an element $a$ of type $\alpha$, the function `indexOf?` returns `none` if and only if $a$ does not appear in $l$. In other words, $l.\text{indexOf?}\ a = \text{none} \leftrightarrow a \notin l$.
List.lookup_cons_self theorem
{k : α} : ((k, b) :: es).lookup k = some b
Full source
@[simp] theorem lookup_cons_self  {k : α} : ((k,b)(k,b) :: es).lookup k = some b := by
  simp [lookup_cons]
Lookup of Key in Cons Pair Returns Associated Value
Informal description
For any key $k$ of type $\alpha$ and any list $es$ of key-value pairs, the lookup of $k$ in the list $(k, b) :: es$ returns the value $b$ wrapped in `some`.
List.lookup_eq_findSome? theorem
{l : List (α × β)} {k : α} : l.lookup k = l.findSome? fun p => if k == p.1 then some p.2 else none
Full source
theorem lookup_eq_findSome? {l : List (α × β)} {k : α} :
    l.lookup k = l.findSome? fun p => if k == p.1 then some p.2 else none := by
  induction l with
  | nil => rfl
  | cons p l ih =>
    match p with
    | (k', v) =>
      simp only [lookup_cons, findSome?_cons]
      split <;> simp_all
Equivalence of List Lookup and First Non-None Result via Key Matching
Informal description
For any list $l$ of key-value pairs $(a_i, b_i)$ and any key $k$ of type $\alpha$, the lookup of $k$ in $l$ is equal to the first non-`none` result obtained by applying the function $\lambda p \mapsto \text{if } k == p.1 \text{ then some } p.2 \text{ else none}$ to each element of $l$.
List.lookup_eq_none_iff theorem
{l : List (α × β)} {k : α} : l.lookup k = none ↔ ∀ p ∈ l, k != p.1
Full source
@[simp] theorem lookup_eq_none_iff {l : List (α × β)} {k : α} :
    l.lookup k = none ↔ ∀ p ∈ l, k != p.1 := by
  simp [lookup_eq_findSome?]
Characterization of Failed Lookup in Association List
Informal description
For any list $l$ of key-value pairs $(a_i, b_i)$ and any key $k$ of type $\alpha$, the lookup of $k$ in $l$ returns `none` if and only if for every pair $(a_i, b_i)$ in $l$, the key $k$ is not equal to $a_i$ (using the boolean inequality `!=`).
List.lookup_isSome_iff theorem
{l : List (α × β)} {k : α} : (l.lookup k).isSome ↔ ∃ p ∈ l, k == p.1
Full source
@[simp] theorem lookup_isSome_iff {l : List (α × β)} {k : α} :
    (l.lookup k).isSome ↔ ∃ p ∈ l, k == p.1 := by
  simp [lookup_eq_findSome?]
Existence of Key in List Lookup: $(l.\text{lookup}\ k).\text{isSome} \leftrightarrow \exists (k', v) \in l,\ k == k'$
Informal description
For any list $l$ of key-value pairs $(a_i, b_i)$ and any key $k$ of type $\alpha$, the lookup of $k$ in $l$ returns a value (i.e., `isSome` holds) if and only if there exists a pair $(a_j, b_j) \in l$ such that $k$ is equal to $a_j$ (using the boolean equality `==`).
List.lookup_eq_some_iff theorem
{l : List (α × β)} {k : α} {b : β} : l.lookup k = some b ↔ ∃ l₁ l₂, l = l₁ ++ (k, b) :: l₂ ∧ ∀ p ∈ l₁, k != p.1
Full source
theorem lookup_eq_some_iff {l : List (α × β)} {k : α} {b : β} :
    l.lookup k = some b ↔ ∃ l₁ l₂, l = l₁ ++ (k, b) :: l₂ ∧ ∀ p ∈ l₁, k != p.1 := by
  simp only [lookup_eq_findSome?, findSome?_eq_some_iff]
  constructor
  · rintro ⟨l₁, a, l₂, rfl, h₁, h₂⟩
    simp only [beq_iff_eq, Option.ite_none_right_eq_some, Option.some.injEq] at h₁
    obtain ⟨rfl, rfl⟩ := h₁
    simp at h₂
    exact ⟨l₁, l₂, rfl, by simpa using h₂⟩
  · rintro ⟨l₁, l₂, rfl, h⟩
    exact ⟨l₁, (k, b), l₂, rfl, by simp, by simpa using h⟩
Characterization of Successful Lookup in Association Lists: $l.\text{lookup}\ k = \text{some}\ b$ iff $(k, b)$ is the first occurrence of $k$ in $l$
Informal description
For a list $l$ of key-value pairs $(a_i, b_i)$, a key $k$, and a value $b$, the lookup of $k$ in $l$ returns `some $b$` if and only if there exist sublists $l_1$ and $l_2$ such that: 1. $l = l_1 \mathbin{+\!\!+} [(k, b)] \mathbin{+\!\!+} l_2$ (i.e., the pair $(k, b)$ appears in $l$ after $l_1$ and before $l_2$), and 2. For every pair $(a', b') \in l_1$, the key $a'$ does not equal $k$ (i.e., $k$ does not appear in $l_1$).
List.lookup_append theorem
{l₁ l₂ : List (α × β)} {k : α} : (l₁ ++ l₂).lookup k = (l₁.lookup k).or (l₂.lookup k)
Full source
theorem lookup_append {l₁ l₂ : List (α × β)} {k : α} :
    (l₁ ++ l₂).lookup k = (l₁.lookup k).or (l₂.lookup k) := by
  simp [lookup_eq_findSome?, findSome?_append]
Concatenation Preserves Lookup: $\text{lookup } k (l_1 ++ l_2) = (\text{lookup } k l_1) \text{ or } (\text{lookup } k l_2)$
Informal description
For any two lists $l_1$ and $l_2$ of key-value pairs $(a_i, b_i)$ where $a_i \in \alpha$ and $b_i \in \beta$, and for any key $k \in \alpha$, the lookup of $k$ in the concatenated list $l_1 ++ l_2$ is equal to the first non-`none` result obtained by either looking up $k$ in $l_1$ or looking up $k$ in $l_2$. In other words: $$\text{lookup } k (l_1 ++ l_2) = (\text{lookup } k l_1) \text{ or } (\text{lookup } k l_2)$$ where $\text{or}$ is the non-short-circuiting choice operation on optional values.
List.lookup_replicate theorem
{k : α} : (replicate n (a, b)).lookup k = if n = 0 then none else if k == a then some b else none
Full source
theorem lookup_replicate {k : α} :
    (replicate n (a,b)).lookup k = if n = 0 then none else if k == a then some b else none := by
  induction n with
  | zero => simp
  | succ n ih =>
    simp only [replicate_succ, lookup_cons]
    split <;> simp_all
Lookup in Replicated Association List: $\text{lookup } k \ (\text{replicate } n \ (a, b)) = \text{if } n = 0 \text{ then none else if } k = a \text{ then some } b \text{ else none}$
Informal description
For any key $k$ of type $\alpha$, the lookup operation on a list consisting of $n$ copies of the pair $(a, b)$ returns $\text{some } b$ if $n > 0$ and $k = a$, and returns $\text{none}$ otherwise. More precisely: \[ \text{lookup } k \ (\text{replicate } n \ (a, b)) = \begin{cases} \text{none} & \text{if } n = 0, \\ \text{some } b & \text{if } n > 0 \text{ and } k = a, \\ \text{none} & \text{otherwise.} \end{cases} \]
List.lookup_replicate_of_pos theorem
{k : α} (h : 0 < n) : (replicate n (a, b)).lookup k = if k == a then some b else none
Full source
theorem lookup_replicate_of_pos {k : α} (h : 0 < n) :
    (replicate n (a, b)).lookup k = if k == a then some b else none := by
  simp [lookup_replicate, Nat.ne_of_gt h]
Lookup in Replicated Association List for Positive Length: $\text{lookup } k \ (\text{replicate } n \ (a, b)) = \text{if } k = a \text{ then some } b \text{ else none}$ when $n > 0$
Informal description
For any key $k$ of type $\alpha$ and any positive natural number $n > 0$, the lookup operation on a list consisting of $n$ copies of the pair $(a, b)$ returns $\text{some } b$ if $k = a$, and $\text{none}$ otherwise. That is: \[ \text{lookup } k \ (\text{replicate } n \ (a, b)) = \begin{cases} \text{some } b & \text{if } k = a, \\ \text{none} & \text{otherwise.} \end{cases} \]
List.lookup_replicate_self theorem
{a : α} : (replicate n (a, b)).lookup a = if n = 0 then none else some b
Full source
theorem lookup_replicate_self {a : α} :
    (replicate n (a, b)).lookup a = if n = 0 then none else some b := by
  simp [lookup_replicate]
Lookup of Self in Replicated Association List: $\text{lookup } a \ (\text{replicate } n \ (a, b)) = \text{if } n = 0 \text{ then none else some } b$
Informal description
For any key $a$ of type $\alpha$, the lookup operation on a list consisting of $n$ copies of the pair $(a, b)$ returns $\text{some } b$ if $n > 0$, and returns $\text{none}$ otherwise. More precisely: \[ \text{lookup } a \ (\text{replicate } n \ (a, b)) = \begin{cases} \text{none} & \text{if } n = 0, \\ \text{some } b & \text{if } n > 0. \end{cases} \]
List.lookup_replicate_self_of_pos theorem
{a : α} (h : 0 < n) : (replicate n (a, b)).lookup a = some b
Full source
@[simp] theorem lookup_replicate_self_of_pos {a : α} (h : 0 < n) :
    (replicate n (a, b)).lookup a = some b := by
  simp [lookup_replicate_self, Nat.ne_of_gt h]
Lookup of Self in Non-Empty Replicated Association List Yields Value
Informal description
For any key $a$ of type $\alpha$ and any positive natural number $n$, the lookup operation on a list consisting of $n$ copies of the pair $(a, b)$ returns $\text{some } b$. That is, if $n > 0$, then $\text{lookup } a \ (\text{replicate } n \ (a, b)) = \text{some } b$.
List.lookup_replicate_ne theorem
{k : α} (h : !k == a) : (replicate n (a, b)).lookup k = none
Full source
@[simp] theorem lookup_replicate_ne {k : α} (h : !k == a) :
    (replicate n (a, b)).lookup k = none := by
  simp_all [lookup_replicate]
Lookup in Replicated List Yields `none` for Non-Matching Keys
Informal description
For any key $k$ of type $\alpha$ such that $k \neq a$ (using boolean equality), the lookup operation on a list consisting of $n$ copies of the pair $(a, b)$ returns `none`. That is, $\text{lookup } k \ (\text{replicate } n \ (a, b)) = \text{none}$ when $k \neq a$.
List.Sublist.lookup_isSome theorem
{l₁ l₂ : List (α × β)} (h : l₁ <+ l₂) : (l₁.lookup k).isSome → (l₂.lookup k).isSome
Full source
theorem Sublist.lookup_isSome {l₁ l₂ : List (α × β)} (h : l₁ <+ l₂) :
    (l₁.lookup k).isSome → (l₂.lookup k).isSome := by
  simp only [lookup_eq_findSome?]
  exact h.findSome?_isSome
Sublist preserves non-`none` results in key lookup
Informal description
For any two lists $l_1$ and $l_2$ of key-value pairs of types $\alpha \times \beta$, if $l_1$ is a sublist of $l_2$ (denoted $l_1 <+ l_2$) and the lookup of key $k$ in $l_1$ returns a `some` value (i.e., $(l_1.\text{lookup}\ k).\text{isSome}$ holds), then the lookup of $k$ in $l_2$ also returns a `some` value (i.e., $(l_2.\text{lookup}\ k).\text{isSome}$ holds).
List.Sublist.lookup_eq_none theorem
{l₁ l₂ : List (α × β)} (h : l₁ <+ l₂) : l₂.lookup k = none → l₁.lookup k = none
Full source
theorem Sublist.lookup_eq_none {l₁ l₂ : List (α × β)} (h : l₁ <+ l₂) :
    l₂.lookup k = none → l₁.lookup k = none := by
  simp only [lookup_eq_findSome?]
  exact h.findSome?_eq_none
Sublist Property for Lookup: $\text{lookup}(k, l_2) = \text{none} \rightarrow \text{lookup}(k, l_1) = \text{none}$ when $l_1 <+ l_2$
Informal description
For any two lists $l_1$ and $l_2$ of key-value pairs of types $\alpha \times \beta$, if $l_1$ is a sublist of $l_2$ and the lookup of key $k$ in $l_2$ returns `none`, then the lookup of $k$ in $l_1$ also returns `none$.
List.IsPrefix.lookup_eq_some theorem
{l₁ l₂ : List (α × β)} (h : l₁ <+: l₂) : List.lookup k l₁ = some b → List.lookup k l₂ = some b
Full source
theorem IsPrefix.lookup_eq_some {l₁ l₂ : List (α × β)} (h : l₁ <+: l₂) :
    List.lookup k l₁ = some b → List.lookup k l₂ = some b := by
  simp only [lookup_eq_findSome?]
  exact h.findSome?_eq_some
Prefix Preservation of Lookup Results: $l_1 \text{ prefix of } l_2 \land \text{lookup } k l_1 = \text{some } b \implies \text{lookup } k l_2 = \text{some } b$
Informal description
For any two lists $l_1$ and $l_2$ of key-value pairs $(a_i, b_i)$, if $l_1$ is a prefix of $l_2$ and the lookup of key $k$ in $l_1$ returns `some b`, then the lookup of $k$ in $l_2$ also returns `some b$.
List.IsPrefix.lookup_eq_none theorem
{l₁ l₂ : List (α × β)} (h : l₁ <+: l₂) : List.lookup k l₂ = none → List.lookup k l₁ = none
Full source
theorem IsPrefix.lookup_eq_none {l₁ l₂ : List (α × β)} (h : l₁ <+: l₂) :
    List.lookup k l₂ = noneList.lookup k l₁ = none :=
  h.sublist.lookup_eq_none
Prefix Lookup None Preservation: $\text{lookup}(k, l_2) = \text{none} \rightarrow \text{lookup}(k, l_1) = \text{none}$ when $l_1$ is prefix of $l_2$
Informal description
For any two lists $l_1$ and $l_2$ of key-value pairs of type $\alpha \times \beta$, if $l_1$ is a prefix of $l_2$ and the lookup of key $k$ in $l_2$ returns `none`, then the lookup of $k$ in $l_1$ also returns `none$.
List.IsSuffix.lookup_eq_none theorem
{l₁ l₂ : List (α × β)} (h : l₁ <:+ l₂) : List.lookup k l₂ = none → List.lookup k l₁ = none
Full source
theorem IsSuffix.lookup_eq_none {l₁ l₂ : List (α × β)} (h : l₁ <:+ l₂) :
    List.lookup k l₂ = noneList.lookup k l₁ = none :=
  h.sublist.lookup_eq_none
Suffix Property for Lookup: $\text{lookup}(k, l_2) = \text{none} \rightarrow \text{lookup}(k, l_1) = \text{none}$ when $l_1$ is a suffix of $l_2$
Informal description
For any two lists $l_1$ and $l_2$ of key-value pairs of type $\alpha \times \beta$, if $l_1$ is a suffix of $l_2$ and the lookup of key $k$ in $l_2$ returns `none`, then the lookup of $k$ in $l_1$ also returns `none$.
List.IsInfix.lookup_eq_none theorem
{l₁ l₂ : List (α × β)} (h : l₁ <:+: l₂) : List.lookup k l₂ = none → List.lookup k l₁ = none
Full source
theorem IsInfix.lookup_eq_none {l₁ l₂ : List (α × β)} (h : l₁ <:+: l₂) :
    List.lookup k l₂ = noneList.lookup k l₁ = none :=
  h.sublist.lookup_eq_none
Infix Property for Lookup: $\text{lookup}(k, l_2) = \text{none} \rightarrow \text{lookup}(k, l_1) = \text{none}$ when $l_1$ is an infix of $l_2$
Informal description
For any two lists $l_1$ and $l_2$ of key-value pairs of types $\alpha \times \beta$, if $l_1$ is an infix of $l_2$ and the lookup of key $k$ in $l_2$ returns `none`, then the lookup of $k$ in $l_1$ also returns `none$.
List.head_join abbrev
Full source
@[deprecated head_flatten (since := "2024-10-14")] abbrev head_join := @head_flatten
Head of Joined List Equals First Non-Empty Sublist's Head
Informal description
For any list of lists $L$ of elements of type $\alpha$, if there exists a non-empty sublist $l \in L$, then the head of the joined list $\text{join}(L)$ is equal to the value obtained from the first non-empty sublist's head via $\text{findSome?}$. More precisely, if there exists $l \in L$ such that $l \neq []$, then: $$\text{head}(\text{join}(L)) = \text{get}(\text{findSome?} (\lambda l, \text{head?}(l)) L)$$ where $\text{get}$ extracts the value from a $\text{some}$ constructor.
List.getLast_join abbrev
Full source
@[deprecated getLast_flatten (since := "2024-10-14")] abbrev getLast_join := @getLast_flatten
Last Element of Concatenated List Equals First Non-Empty Last Element in Reversed List
Informal description
For any list of lists $L$ of type $\text{List}(\text{List } \alpha)$, if there exists a non-empty list $l \in L$, then the last element of the concatenated list $\text{join } L$ is equal to the last element of the first non-empty list in the reversed list $L$ that has a last element. More precisely, if $h$ is a proof that there exists a non-empty list $l \in L$, then: $$\text{getLast}(\text{join } L, h) = \text{get}(\text{findSome? } (\lambda l \mapsto \text{getLast?}(l)) (\text{reverse } L), h')$$ where $h'$ is derived from $h$ and ensures the $\text{findSome?}$ operation succeeds.
List.find?_join abbrev
Full source
@[deprecated find?_flatten (since := "2024-10-14")] abbrev find?_join := @find?_flatten
First Satisfying Element in Concatenated List Equals First Non-None Result of Sublist Searches
Informal description
For any list of lists $xss : \text{List}(\text{List } \alpha)$ and any predicate $p : \alpha \to \text{Bool}$, the first element satisfying $p$ in the concatenated list $\text{join } xss$ is equal to the first non-`none` result obtained by applying $\text{find? } p$ to each sublist in $xss$. Formally: \[ \text{find? } p \ (\text{join } xss) = \text{findSome? } (\lambda ys \Rightarrow \text{find? } p \ ys) \ xss \]
List.find?_join_eq_none abbrev
Full source
@[deprecated find?_flatten_eq_none (since := "2024-10-14")] abbrev find?_join_eq_none := @find?_flatten_eq_none_iff
Characterization of Empty Search Result in Concatenated List
Informal description
For any list of lists `xss : List (List α)` and any predicate `p : α → Bool`, the following are equivalent: 1. The predicate `p` does not hold for any element in the concatenated list `join xss` 2. For every sublist `ys` in `xss` and every element `x` in `ys`, the predicate `p x` evaluates to `false` In mathematical notation: \[ \text{find? } p \ (\text{join } xss) = \text{none} \leftrightarrow \forall ys \in xss, \forall x \in ys, \neg p(x) \]
List.find?_join_eq_some abbrev
Full source
@[deprecated find?_flatten_eq_some (since := "2024-10-14")] abbrev find?_join_eq_some := @find?_flatten_eq_some_iff
Characterization of First Satisfying Element in Concatenated List: $\text{find? } p \ (\text{join } xss) = \text{some } a$
Informal description
For any list of lists $xss : \text{List}(\text{List } \alpha)$, predicate $p : \alpha \to \text{Bool}$, and element $a : \alpha$, the following are equivalent: 1. The function $\text{find? } p$ applied to the concatenated list $\text{join } xss$ returns $\text{some } a$. 2. There exists a sublist $ys$ in $xss$ such that $\text{find? } p \ ys = \text{some } a$, and for all sublists $zs$ appearing before $ys$ in $xss$, $\text{find? } p \ zs = \text{none}$.
List.findIdx?_join abbrev
Full source
@[deprecated findIdx?_flatten (since := "2024-10-14")] abbrev findIdx?_join := @findIdx?_flatten
Index of First Satisfying Element in Flattened List of Lists
Informal description
For any list of lists $l : \text{List}(\text{List}\,\alpha)$ and predicate $p : \alpha \to \text{Bool}$, the index of the first element in the flattened list $\text{flatten}\,l$ that satisfies $p$ is equal to: - $\text{some}\left(\sum_{j < i} \text{length}(l_j) + k\right)$, where $i$ is the index of the first sublist in $l$ containing an element satisfying $p$, and $k$ is the index of the first such element within that sublist, if such a sublist exists; - $\text{none}$ otherwise. More precisely: \[ \text{findIdx?}_p(\text{flatten}\,l) = \begin{cases} \text{some}\left(\sum_{j < i} \text{length}(l_j) + k\right) & \text{if } \exists i, \text{any}\,p\,(l_i) \text{ and } k = \text{findIdx}_p(l_i), \\ \text{none} & \text{otherwise.} \end{cases} \]