Documentation

Init.Data.List.MinMax

Lemmas about List.min? and `List.max?. #

Minima and maxima #

min? #

@[simp]
theorem List.min?_nil {α : Type u_1} [Min α] :
theorem List.min?_cons' {α : Type u_1} {x : α} [Min α] {xs : List α} :
(x :: xs).min? = some (foldl min x xs)
@[simp]
theorem List.min?_cons {α : Type u_1} {x : α} [Min α] [Std.Associative min] {xs : List α} :
(x :: xs).min? = some (xs.min?.elim x (min x))
@[simp]
theorem List.min?_eq_none_iff {α : Type u_1} {xs : List α} [Min α] :
xs.min? = none xs = []
theorem List.isSome_min?_of_mem {α : Type u_1} {l : List α} [Min α] {a : α} (h : a l) :
theorem List.min?_eq_head? {α : Type u} [Min α] {l : List α} (h : Pairwise (fun (a b : α) => min a b = a) l) :
theorem List.min?_mem {α : Type u_1} {a : α} [Min α] (min_eq_or : ∀ (a b : α), min a b = a min a b = b) {xs : List α} :
xs.min? = some aa xs
theorem List.le_min?_iff {α : Type u_1} {a : α} [Min α] [LE α] (le_min_iff : ∀ (a b c : α), a min b c a b a c) {xs : List α} :
xs.min? = some a∀ {x : α}, x a ∀ (b : α), b xsx b
theorem List.min?_eq_some_iff {α : Type u_1} {a : α} [Min α] [LE α] (le_refl : ∀ (a : α), a a) (min_eq_or : ∀ (a b : α), min a b = a min a b = b) (le_min_iff : ∀ (a b c : α), a min b c a b a c) {xs : List α} (anti : ∀ (a b : α), a xsb xsa bb aa = b := by exact fun a b _ _ => Std.Antisymm.antisymm a b) :
xs.min? = some a a xs ∀ (b : α), b xsa b
theorem List.min?_replicate {α : Type u_1} [Min α] {n : Nat} {a : α} (w : min a a = a) :
@[simp]
theorem List.min?_replicate_of_pos {α : Type u_1} [Min α] {n : Nat} {a : α} (w : min a a = a) (h : 0 < n) :
theorem List.foldl_min {α : Type u_1} [Min α] [Std.IdempotentOp min] [Std.Associative min] {l : List α} {a : α} :
foldl min a l = min a (l.min?.getD a)

max? #

@[simp]
theorem List.max?_nil {α : Type u_1} [Max α] :
theorem List.max?_cons' {α : Type u_1} {x : α} [Max α] {xs : List α} :
(x :: xs).max? = some (foldl max x xs)
@[simp]
theorem List.max?_cons {α : Type u_1} {x : α} [Max α] [Std.Associative max] {xs : List α} :
(x :: xs).max? = some (xs.max?.elim x (max x))
@[simp]
theorem List.max?_eq_none_iff {α : Type u_1} {xs : List α} [Max α] :
xs.max? = none xs = []
theorem List.isSome_max?_of_mem {α : Type u_1} {l : List α} [Max α] {a : α} (h : a l) :
theorem List.max?_mem {α : Type u_1} {a : α} [Max α] (min_eq_or : ∀ (a b : α), max a b = a max a b = b) {xs : List α} :
xs.max? = some aa xs
theorem List.max?_le_iff {α : Type u_1} {a : α} [Max α] [LE α] (max_le_iff : ∀ (a b c : α), max b c a b a c a) {xs : List α} :
xs.max? = some a∀ {x : α}, a x ∀ (b : α), b xsb x
theorem List.max?_eq_some_iff {α : Type u_1} {a : α} [Max α] [LE α] [anti : Std.Antisymm fun (x1 x2 : α) => x1 x2] (le_refl : ∀ (a : α), a a) (max_eq_or : ∀ (a b : α), max a b = a max a b = b) (max_le_iff : ∀ (a b c : α), max b c a b a c a) {xs : List α} :
xs.max? = some a a xs ∀ (b : α), b xsb a
theorem List.max?_replicate {α : Type u_1} [Max α] {n : Nat} {a : α} (w : max a a = a) :
@[simp]
theorem List.max?_replicate_of_pos {α : Type u_1} [Max α] {n : Nat} {a : α} (w : max a a = a) (h : 0 < n) :
theorem List.foldl_max {α : Type u_1} [Max α] [Std.IdempotentOp max] [Std.Associative max] {l : List α} {a : α} :
foldl max a l = max a (l.max?.getD a)
@[reducible, inline, deprecated List.min?_nil (since := "2024-09-29")]
abbrev List.minimum?_nil {α : Type u_1} [Min α] :
Equations
@[reducible, inline, deprecated List.min?_cons (since := "2024-09-29")]
abbrev List.minimum?_cons {α : Type u_1} {x : α} [Min α] [Std.Associative min] {xs : List α} :
(x :: xs).min? = some (xs.min?.elim x (min x))
Equations
@[reducible, inline, deprecated List.min?_eq_none_iff (since := "2024-09-29")]
abbrev List.mininmum?_eq_none_iff {α : Type u_1} {xs : List α} [Min α] :
xs.min? = none xs = []
Equations
@[reducible, inline, deprecated List.min?_mem (since := "2024-09-29")]
abbrev List.minimum?_mem {α : Type u_1} {a : α} [Min α] (min_eq_or : ∀ (a b : α), min a b = a min a b = b) {xs : List α} :
xs.min? = some aa xs
Equations
@[reducible, inline, deprecated List.le_min?_iff (since := "2024-09-29")]
abbrev List.le_minimum?_iff {α : Type u_1} {a : α} [Min α] [LE α] (le_min_iff : ∀ (a b c : α), a min b c a b a c) {xs : List α} :
xs.min? = some a∀ {x : α}, x a ∀ (b : α), b xsx b
Equations
@[reducible, inline, deprecated List.min?_eq_some_iff (since := "2024-09-29")]
abbrev List.minimum?_eq_some_iff {α : Type u_1} {a : α} [Min α] [LE α] (le_refl : ∀ (a : α), a a) (min_eq_or : ∀ (a b : α), min a b = a min a b = b) (le_min_iff : ∀ (a b c : α), a min b c a b a c) {xs : List α} (anti : ∀ (a b : α), a xsb xsa bb aa = b := by exact fun a b _ _ => Std.Antisymm.antisymm a b) :
xs.min? = some a a xs ∀ (b : α), b xsa b
Equations
@[reducible, inline, deprecated List.min?_replicate (since := "2024-09-29")]
abbrev List.minimum?_replicate {α : Type u_1} [Min α] {n : Nat} {a : α} (w : min a a = a) :
Equations
@[reducible, inline, deprecated List.min?_replicate_of_pos (since := "2024-09-29")]
abbrev List.minimum?_replicate_of_pos {α : Type u_1} [Min α] {n : Nat} {a : α} (w : min a a = a) (h : 0 < n) :
Equations
@[reducible, inline, deprecated List.max?_nil (since := "2024-09-29")]
abbrev List.maximum?_nil {α : Type u_1} [Max α] :
Equations
@[reducible, inline, deprecated List.max?_cons (since := "2024-09-29")]
abbrev List.maximum?_cons {α : Type u_1} {x : α} [Max α] [Std.Associative max] {xs : List α} :
(x :: xs).max? = some (xs.max?.elim x (max x))
Equations
@[reducible, inline, deprecated List.max?_eq_none_iff (since := "2024-09-29")]
abbrev List.maximum?_eq_none_iff {α : Type u_1} {xs : List α} [Max α] :
xs.max? = none xs = []
Equations
@[reducible, inline, deprecated List.max?_mem (since := "2024-09-29")]
abbrev List.maximum?_mem {α : Type u_1} {a : α} [Max α] (min_eq_or : ∀ (a b : α), max a b = a max a b = b) {xs : List α} :
xs.max? = some aa xs
Equations
@[reducible, inline, deprecated List.max?_le_iff (since := "2024-09-29")]
abbrev List.maximum?_le_iff {α : Type u_1} {a : α} [Max α] [LE α] (max_le_iff : ∀ (a b c : α), max b c a b a c a) {xs : List α} :
xs.max? = some a∀ {x : α}, a x ∀ (b : α), b xsb x
Equations
@[reducible, inline, deprecated List.max?_eq_some_iff (since := "2024-09-29")]
abbrev List.maximum?_eq_some_iff {α : Type u_1} {a : α} [Max α] [LE α] [anti : Std.Antisymm fun (x1 x2 : α) => x1 x2] (le_refl : ∀ (a : α), a a) (max_eq_or : ∀ (a b : α), max a b = a max a b = b) (max_le_iff : ∀ (a b c : α), max b c a b a c a) {xs : List α} :
xs.max? = some a a xs ∀ (b : α), b xsb a
Equations
@[reducible, inline, deprecated List.max?_replicate (since := "2024-09-29")]
abbrev List.maximum?_replicate {α : Type u_1} [Max α] {n : Nat} {a : α} (w : max a a = a) :
Equations
@[reducible, inline, deprecated List.max?_replicate_of_pos (since := "2024-09-29")]
abbrev List.maximum?_replicate_of_pos {α : Type u_1} [Max α] {n : Nat} {a : α} (w : max a a = a) (h : 0 < n) :
Equations