Skip to content
Merged
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
12 changes: 1 addition & 11 deletions src/Init/Control/Lawful/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -170,6 +170,7 @@ theorem bind_pure_unit [Monad m] [LawfulMonad m] {x : m PUnit} : (x >>= fun _ =>
theorem map_congr [Functor m] {x : m α} {f g : α → β} (h : ∀ a, f a = g a) : (f <$> x : m β) = g <$> x := by
simp [funext h]

@[deprecated seq_eq_bind_map (since := "2025-10-26")]
theorem seq_eq_bind {α β : Type u} [Monad m] [LawfulMonad m] (mf : m (α → β)) (x : m α) : mf <*> x = mf >>= fun f => f <$> x := by
rw [bind_map]

Expand Down Expand Up @@ -256,14 +257,3 @@ instance : LawfulMonad Id := by
@[simp] theorem run_seq (f : Id (α → β)) (x : Id α) : (f <*> x).run = f.run x.run := rfl

end Id

/-! # Option -/

instance : LawfulMonad Option := LawfulMonad.mk'
(id_map := fun x => by cases x <;> rfl)
(pure_bind := fun _ _ => rfl)
(bind_assoc := fun x _ _ => by cases x <;> rfl)
(bind_pure_comp := fun _ x => by cases x <;> rfl)

instance : LawfulApplicative Option := inferInstance
instance : LawfulFunctor Option := inferInstance
6 changes: 3 additions & 3 deletions src/Init/Control/Lawful/Instances.lean
Original file line number Diff line number Diff line change
Expand Up @@ -189,12 +189,12 @@ instance [Monad m] [LawfulMonad m] : LawfulMonad (OptionT m) where

@[simp] theorem run_seq [Monad m] [LawfulMonad m] (f : OptionT m (α → β)) (x : OptionT m α) :
(f <*> x).run = Option.elimM f.run (pure none) (fun f => Option.map f <$> x.run) := by
simp [seq_eq_bind, Option.elimM, Option.elim]
simp [seq_eq_bind_map, Option.elimM, Option.elim]

@[simp] theorem run_seqLeft [Monad m] [LawfulMonad m] (x : OptionT m α) (y : OptionT m β) :
(x <* y).run = Option.elimM x.run (pure none)
(fun x => Option.map (Function.const β x) <$> y.run) := by
simp [seqLeft_eq, seq_eq_bind, Option.elimM, OptionT.run_bind]
simp [seqLeft_eq, seq_eq_bind_map, Option.elimM, OptionT.run_bind]

@[simp] theorem run_seqRight [Monad m] [LawfulMonad m] (x : OptionT m α) (y : OptionT m β) :
(x *> y).run = Option.elimM x.run (pure none) (Function.const α y.run) := by
Expand All @@ -219,7 +219,7 @@ instance : LawfulMonad Option := LawfulMonad.mk'
(id_map := fun x => by cases x <;> rfl)
(pure_bind := fun _ _ => by rfl)
(bind_assoc := fun a _ _ => by cases a <;> rfl)
(bind_pure_comp := bind_pure_comp)
(bind_pure_comp := fun _ x => by cases x <;> rfl)

instance : LawfulApplicative Option := inferInstance
instance : LawfulFunctor Option := inferInstance
Expand Down
2 changes: 1 addition & 1 deletion src/Init/Control/Lawful/MonadLift/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -23,7 +23,7 @@ theorem monadLift_map [LawfulMonad m] [LawfulMonad n] (f : α → β) (ma : m α

theorem monadLift_seq [LawfulMonad m] [LawfulMonad n] (mf : m (α → β)) (ma : m α) :
monadLift (mf <*> ma) = monadLift mf <*> (monadLift ma : n α) := by
simp only [seq_eq_bind, monadLift_map, monadLift_bind]
simp only [seq_eq_bind_map, monadLift_map, monadLift_bind]

theorem monadLift_seqLeft [LawfulMonad m] [LawfulMonad n] (x : m α) (y : m β) :
monadLift (x <* y) = (monadLift x : n α) <* (monadLift y : n β) := by
Expand Down
20 changes: 14 additions & 6 deletions src/Init/Core.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1348,8 +1348,12 @@ namespace Subtype
theorem exists_of_subtype {α : Type u} {p : α → Prop} : { x // p x } → Exists (fun x => p x)
| ⟨a, h⟩ => ⟨a, h⟩

variable {α : Type u} {p : α → Prop}
variable {α : Sort u} {p : α → Prop}

protected theorem ext : ∀ {a1 a2 : {x // p x}}, val a1 = val a2 → a1 = a2
| ⟨_, _⟩, ⟨_, _⟩, rfl => rfl

@[deprecated Subtype.ext (since := "2025-10-26")]
protected theorem eq : ∀ {a1 a2 : {x // p x}}, val a1 = val a2 → a1 = a2
| ⟨_, _⟩, ⟨_, _⟩, rfl => rfl

Expand All @@ -1364,9 +1368,9 @@ instance {α : Type u} {p : α → Prop} [BEq α] [ReflBEq α] : ReflBEq {x : α
rfl {x} := BEq.refl x.1

instance {α : Type u} {p : α → Prop} [BEq α] [LawfulBEq α] : LawfulBEq {x : α // p x} where
eq_of_beq h := Subtype.eq (eq_of_beq h)
eq_of_beq h := Subtype.ext (eq_of_beq h)

instance {α : Type u} {p : α → Prop} [DecidableEq α] : DecidableEq {x : α // p x} :=
instance {α : Sort u} {p : α → Prop} [DecidableEq α] : DecidableEq {x : α // p x} :=
fun ⟨a, h₁⟩ ⟨b, h₂⟩ =>
if h : a = b then isTrue (by subst h; exact rfl)
else isFalse (fun h' => Subtype.noConfusion h' (fun h' => absurd h' h))
Expand Down Expand Up @@ -1483,20 +1487,24 @@ protected theorem PSigma.eta {α : Sort u} {β : α → Sort v} {a₁ a₂ : α}

/-! # Universe polymorphic unit -/

theorem PUnit.ext (a b : PUnit) : a = b := by
cases a; cases b; exact rfl

@[deprecated PUnit.ext (since := "2025-10-26")]
theorem PUnit.subsingleton (a b : PUnit) : a = b := by
cases a; cases b; exact rfl

theorem PUnit.eq_punit (a : PUnit) : a = ⟨⟩ :=
PUnit.subsingleton a ⟨⟩
PUnit.ext a ⟨⟩

instance : Subsingleton PUnit :=
Subsingleton.intro PUnit.subsingleton
Subsingleton.intro PUnit.ext

instance : Inhabited PUnit where
default := ⟨⟩

instance : DecidableEq PUnit :=
fun a b => isTrue (PUnit.subsingleton a b)
fun a b => isTrue (PUnit.ext a b)

/-! # Setoid -/

Expand Down
2 changes: 1 addition & 1 deletion src/Init/Data/Array/Bootstrap.lean
Original file line number Diff line number Diff line change
Expand Up @@ -31,7 +31,7 @@ theorem foldlM_toList.aux [Monad m]
· cases Nat.not_le_of_gt ‹_› (Nat.zero_add _ ▸ H)
· rename_i i; rw [Nat.succ_add] at H
simp [foldlM_toList.aux (j := j+1) H]
rw (occs := [2]) [← List.getElem_cons_drop_succ_eq_drop ‹_›]
rw (occs := [2]) [← List.getElem_cons_drop ‹_›]
simp
· rw [List.drop_of_length_le (Nat.ge_of_not_lt ‹_›)]; simp

Expand Down
2 changes: 1 addition & 1 deletion src/Init/Data/Array/GetLit.lean
Original file line number Diff line number Diff line change
Expand Up @@ -50,6 +50,6 @@ where
getLit_eq (xs : Array α) (i : Nat) (h₁ : xs.size = n) (h₂ : i < n) : xs.getLit i h₁ h₂ = getElem xs.toList i ((id (α := xs.toList.length = n) h₁) ▸ h₂) :=
rfl
go (i : Nat) (hi : i ≤ xs.size) : toListLitAux xs n hsz i hi (xs.toList.drop i) = xs.toList := by
induction i <;> simp only [List.drop, toListLitAux, getLit_eq, List.getElem_cons_drop_succ_eq_drop, *]
induction i <;> simp only [List.drop, toListLitAux, getLit_eq, List.getElem_cons_drop, *]

end Array
84 changes: 48 additions & 36 deletions src/Init/Data/Array/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -265,8 +265,6 @@ theorem push_eq_push {a b : α} {xs ys : Array α} : xs.push a = ys.push b ↔ a
· rintro ⟨rfl, rfl⟩
rfl

theorem push_eq_append_singleton {as : Array α} {x : α} : as.push x = as ++ #[x] := rfl

theorem exists_push_of_ne_empty {xs : Array α} (h : xs ≠ #[]) :
∃ (ys : Array α) (a : α), xs = ys.push a := by
rcases xs with ⟨xs⟩
Expand Down Expand Up @@ -817,6 +815,11 @@ theorem contains_eq_true_of_mem [BEq α] [ReflBEq α] {a : α} {as : Array α} (
theorem elem_iff [BEq α] [LawfulBEq α] {a : α} {xs : Array α} :
elem a xs = true ↔ a ∈ xs := ⟨mem_of_contains_eq_true, contains_eq_true_of_mem⟩

@[grind =]
theorem contains_iff_mem [BEq α] [LawfulBEq α] {a : α} {xs : Array α} :
xs.contains a = true ↔ a ∈ xs := ⟨mem_of_contains_eq_true, contains_eq_true_of_mem⟩

@[deprecated contains_iff_mem (since := "2025-10-26")]
theorem contains_iff [BEq α] [LawfulBEq α] {a : α} {xs : Array α} :
xs.contains a = true ↔ a ∈ xs := ⟨mem_of_contains_eq_true, contains_eq_true_of_mem⟩

Expand Down Expand Up @@ -1139,7 +1142,7 @@ where
aux (i bs) :
mapM.map f xs i bs = (xs.toList.drop i).foldlM (fun bs a => bs.push <$> f a) bs := by
unfold mapM.map; split
· rw [← List.getElem_cons_drop_succ_eq_drop ‹_›]
· rw [← List.getElem_cons_drop ‹_›]
simp only [aux (i + 1), map_eq_pure_bind, List.foldlM_cons, bind_assoc,
pure_bind]
rfl
Expand Down Expand Up @@ -1853,6 +1856,9 @@ theorem getElem_of_append {xs ys zs : Array α} (eq : xs = ys.push a ++ zs) (h :

theorem push_eq_append {a : α} {as : Array α} : as.push a = as ++ #[a] := rfl

@[deprecated push_eq_append (since := "2025-10-26")]
theorem push_eq_append_singleton {as : Array α} {x : α} : as.push x = as ++ #[x] := rfl

theorem append_inj {xs₁ xs₂ ys₁ ys₂ : Array α} (h : xs₁ ++ ys₁ = xs₂ ++ ys₂) (hl : xs₁.size = xs₂.size) :
xs₁ = xs₂ ∧ ys₁ = ys₂ := by
rcases xs₁ with ⟨s₁⟩
Expand Down Expand Up @@ -2053,11 +2059,22 @@ theorem append_eq_map_iff {f : α → β} :
| nil => simp
| cons as => induction as.toList <;> simp [*]

@[simp] theorem flatten_toArray_map {L : List (List α)} :
@[simp] theorem flatten_toArray_map_toArray {L : List (List α)} :
(L.map List.toArray).toArray.flatten = L.flatten.toArray := by
apply ext'
simp [Function.comp_def]

@[deprecated flatten_toArray_map_toArray (since := "2025-10-26")]
theorem flatten_toArray_map {L : List (List α)} :
(L.map List.toArray).toArray.flatten = L.flatten.toArray := by
simp

@[grind =]
theorem flatten_map_toArray_toArray {L : List (List α)} :
(L.toArray.map List.toArray).flatten = L.flatten.toArray := by
simp

@[deprecated flatten_map_toArray_toArray (since := "2025-10-26")]
theorem flatten_map_toArray {L : List (List α)} :
(L.toArray.map List.toArray).flatten = L.flatten.toArray := by
simp
Expand Down Expand Up @@ -2104,32 +2121,33 @@ theorem forall_mem_flatten {p : α → Prop} {xss : Array (Array α)} :

theorem flatten_eq_flatMap {xss : Array (Array α)} : flatten xss = xss.flatMap id := by
induction xss using array₂_induction
rw [flatten_toArray_map, List.flatten_eq_flatMap]
rw [flatten_toArray_map_toArray, List.flatten_eq_flatMap]
simp [List.flatMap_map]

@[simp, grind _=_] theorem map_flatten {f : α → β} {xss : Array (Array α)} :
(flatten xss).map f = (map (map f) xss).flatten := by
induction xss using array₂_induction with
| of xss =>
simp only [flatten_toArray_map, List.map_toArray, List.map_flatten, List.map_map,
simp only [flatten_toArray_map_toArray, List.map_toArray, List.map_flatten, List.map_map,
Function.comp_def]
rw [← Function.comp_def, ← List.map_map, flatten_toArray_map]
rw [← Function.comp_def, ← List.map_map, flatten_toArray_map_toArray]

@[simp, grind =] theorem filterMap_flatten {f : α → Option β} {xss : Array (Array α)} {stop : Nat} (w : stop = xss.flatten.size) :
filterMap f (flatten xss) 0 stop = flatten (map (filterMap f) xss) := by
subst w
induction xss using array₂_induction
simp only [flatten_toArray_map, List.size_toArray, List.length_flatten, List.filterMap_toArray',
List.filterMap_flatten, List.map_toArray, List.map_map, Function.comp_def]
rw [← Function.comp_def, ← List.map_map, flatten_toArray_map]
simp only [flatten_toArray_map_toArray, List.size_toArray, List.length_flatten,
List.filterMap_toArray', List.filterMap_flatten, List.map_toArray, List.map_map,
Function.comp_def]
rw [← Function.comp_def, ← List.map_map, flatten_toArray_map_toArray]

@[simp, grind =] theorem filter_flatten {p : α → Bool} {xss : Array (Array α)} {stop : Nat} (w : stop = xss.flatten.size) :
filter p (flatten xss) 0 stop = flatten (map (filter p) xss) := by
subst w
induction xss using array₂_induction
simp only [flatten_toArray_map, List.size_toArray, List.length_flatten, List.filter_toArray',
List.filter_flatten, List.map_toArray, List.map_map, Function.comp_def]
rw [← Function.comp_def, ← List.map_map, flatten_toArray_map]
simp only [flatten_toArray_map_toArray, List.size_toArray, List.length_flatten,
List.filter_toArray', List.filter_flatten, List.map_toArray, List.map_map, Function.comp_def]
rw [← Function.comp_def, ← List.map_map, flatten_toArray_map_toArray]

theorem flatten_filter_not_isEmpty {xss : Array (Array α)} :
flatten (xss.filter fun xs => !xs.isEmpty) = xss.flatten := by
Expand All @@ -2152,23 +2170,23 @@ theorem flatten_filter_ne_empty [DecidablePred fun xs : Array α => xs ≠ #[]]
induction xss using array₂_induction
rcases xs with ⟨l⟩
have this : [l.toArray] = [l].map List.toArray := by simp
simp only [List.push_toArray, flatten_toArray_map, List.append_toArray]
rw [this, ← List.map_append, flatten_toArray_map]
simp only [List.push_toArray, flatten_toArray_map_toArray, List.append_toArray]
rw [this, ← List.map_append, flatten_toArray_map_toArray]
simp

theorem flatten_flatten {xss : Array (Array (Array α))} : flatten (flatten xss) = flatten (map flatten xss) := by
induction xss using array₃_induction with
| of xss =>
rw [flatten_toArray_map]
rw [flatten_toArray_map_toArray]
have : (xss.map (fun xs => xs.map List.toArray)).flatten = xss.flatten.map List.toArray := by
induction xss with
| nil => simp
| cons xs xss ih =>
simp only [List.map_cons, List.flatten_cons, ih, List.map_append]
rw [this, flatten_toArray_map, List.flatten_flatten, ← List.map_toArray, Array.map_map,
rw [this, flatten_toArray_map_toArray, List.flatten_flatten, ← List.map_toArray, Array.map_map,
← List.map_toArray, map_map, Function.comp_def]
simp only [Function.comp_apply, flatten_toArray_map]
rw [List.map_toArray, ← Function.comp_def, ← List.map_map, flatten_toArray_map]
simp only [Function.comp_apply, flatten_toArray_map_toArray]
rw [List.map_toArray, ← Function.comp_def, ← List.map_map, flatten_toArray_map_toArray]

theorem flatten_eq_push_iff {xss : Array (Array α)} {ys : Array α} {y : α} :
xss.flatten = ys.push y ↔
Expand All @@ -2177,13 +2195,13 @@ theorem flatten_eq_push_iff {xss : Array (Array α)} {ys : Array α} {y : α} :
induction xss using array₂_induction with
| of xs =>
rcases ys with ⟨ys⟩
rw [flatten_toArray_map, List.push_toArray, mk.injEq, List.flatten_eq_append_iff]
rw [flatten_toArray_map_toArray, List.push_toArray, mk.injEq, List.flatten_eq_append_iff]
constructor
· rintro (⟨as, bs, rfl, rfl, h⟩ | ⟨as, bs, c, cs, ds, rfl, rfl, h⟩)
· rw [List.singleton_eq_flatten_iff] at h
obtain ⟨xs, ys, rfl, h₁, h₂⟩ := h
exact ⟨((as ++ xs).map List.toArray).toArray, #[], (ys.map List.toArray).toArray, by simp,
by simpa using h₂, by rw [flatten_toArray_map]; simpa⟩
by simpa using h₂, by rw [flatten_toArray_map_toArray]; simpa⟩
· rw [List.singleton_eq_append_iff] at h
obtain (⟨h₁, h₂⟩ | ⟨h₁, h₂⟩) := h
· simp at h₁
Expand Down Expand Up @@ -2228,10 +2246,6 @@ theorem eq_iff_flatten_eq {xss₁ xss₂ : Array (Array α)} :
rw [List.map_inj_right]
simp +contextual

theorem flatten_toArray_map_toArray {xss : List (List α)} :
(xss.map List.toArray).toArray.flatten = xss.flatten.toArray := by
simp

/-! ### flatMap -/

theorem flatMap_def {xs : Array α} {f : α → Array β} : xs.flatMap f = flatten (map f xs) := by
Expand Down Expand Up @@ -2535,8 +2549,8 @@ theorem getElem?_swap {xs : Array α} {i j : Nat} (hi hj) {k : Nat} : (xs.swap i
split <;> rename_i h₃
· simp only [← h₃, Nat.not_le.2 (Nat.lt_succ_self _), Nat.le_refl, false_and]
exact (List.getElem?_reverse' (Eq.trans (by simp +arith) h)).symm
simp only [Nat.succ_le, Nat.lt_iff_le_and_ne.trans (and_iff_left h₃),
Nat.lt_succ.symm.trans (Nat.lt_iff_le_and_ne.trans (and_iff_left (Ne.symm h₂)))]
simp only [Nat.succ_le_iff, Nat.lt_iff_le_and_ne.trans (and_iff_left h₃),
Nat.lt_succ_iff.symm.trans (Nat.lt_iff_le_and_ne.trans (and_iff_left (Ne.symm h₂)))]
· rw [H]; split <;> rename_i h₂
· cases Nat.le_antisymm (Nat.not_lt.1 h₁) (Nat.le_trans h₂.1 h₂.2)
cases Nat.le_antisymm h₂.1 h₂.2
Expand All @@ -2551,7 +2565,7 @@ theorem getElem?_swap {xs : Array α} {i j : Nat} (hi hj) {k : Nat} : (xs.swap i
split
· rfl
· rename_i h
simp only [← show k < _ + 1 ↔ _ from Nat.lt_succ (n := xs.size - 1), this, Nat.zero_le,
simp only [← show k < _ + 1 ↔ _ from Nat.lt_succ_iff (n := xs.size - 1), this, Nat.zero_le,
true_and, Nat.not_lt] at h
rw [List.getElem?_eq_none_iff.2 ‹_›, List.getElem?_eq_none_iff.2 (xs.toList.length_reverse ▸ ‹_›)]

Expand Down Expand Up @@ -2699,8 +2713,8 @@ theorem extract_loop_eq_aux {xs ys : Array α} {size start : Nat} :
| zero => rw [extract_loop_zero, extract_loop_zero, append_empty]
| succ size ih =>
if h : start < xs.size then
rw [extract_loop_succ (h := h), ih, push_eq_append_singleton]
rw [extract_loop_succ (h := h), ih (ys := #[].push _), push_eq_append_singleton, empty_append]
rw [extract_loop_succ (h := h), ih, push_eq_append]
rw [extract_loop_succ (h := h), ih (ys := #[].push _), push_eq_append, empty_append]
rw [append_assoc]
else
rw [extract_loop_of_ge (h := Nat.le_of_not_lt h)]
Expand Down Expand Up @@ -3622,11 +3636,6 @@ theorem contains_iff_exists_mem_beq [BEq α] {xs : Array α} {a : α} :
-- With `LawfulBEq α`, it would be better to use `contains_iff_mem` directly.
grind_pattern contains_iff_exists_mem_beq => xs.contains a

@[grind =]
theorem contains_iff_mem [BEq α] [LawfulBEq α] {xs : Array α} {a : α} :
xs.contains a ↔ a ∈ xs := by
simp

@[simp, grind =]
theorem contains_toList [BEq α] {xs : Array α} {x : α} :
xs.toList.contains x = xs.contains x := by
Expand Down Expand Up @@ -4276,6 +4285,7 @@ theorem getElem_fin_eq_getElem_toList {xs : Array α} {i : Fin xs.size} : xs[i]
@[simp] theorem ugetElem_eq_getElem {xs : Array α} {i : USize} (h : i.toNat < xs.size) :
xs[i] = xs[i.toNat] := rfl

@[deprecated getElem?_eq_none (since := "2025-10-26")]
theorem getElem?_size_le {xs : Array α} {i : Nat} (h : xs.size ≤ i) : xs[i]? = none := by
simp [getElem?_neg, h]

Expand All @@ -4295,6 +4305,7 @@ theorem getElem?_push_lt {xs : Array α} {x : α} {i : Nat} (h : i < xs.size) :
(xs.push x)[i]? = some xs[i] := by
rw [getElem?_pos (xs.push x) i (size_push _ ▸ Nat.lt_succ_of_lt h), getElem_push_lt]

@[deprecated getElem?_push_size (since := "2025-10-26")]
theorem getElem?_push_eq {xs : Array α} {x : α} : (xs.push x)[xs.size]? = some x := by
rw [getElem?_pos (xs.push x) xs.size (size_push _ ▸ Nat.lt_succ_self xs.size), getElem_push_eq]

Expand Down Expand Up @@ -4426,7 +4437,8 @@ theorem uset_toArray {l : List α} {i : USize} {a : α} {h : i.toNat < l.toArray
apply ext'
simp

@[simp, grind =] theorem flatten_toArray {L : List (List α)} :
@[deprecated Array.flatten_map_toArray_toArray (since := "2025-10-26")]
theorem flatten_toArray {L : List (List α)} :
(L.toArray.map List.toArray).flatten = L.flatten.toArray := by
apply ext'
simp
Expand Down
17 changes: 11 additions & 6 deletions src/Init/Data/Array/Lex/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -34,7 +34,18 @@ grind_pattern _root_.List.le_toArray => l₁.toArray ≤ l₂.toArray
grind_pattern lt_toList => xs.toList < ys.toList
grind_pattern le_toList => xs.toList ≤ ys.toList

@[simp]
protected theorem not_lt [LT α] {xs ys : Array α} : ¬ xs < ys ↔ ys ≤ xs := Iff.rfl

@[deprecated Array.not_lt (since := "2025-10-26")]
protected theorem not_lt_iff_ge [LT α] {xs ys : Array α} : ¬ xs < ys ↔ ys ≤ xs := Iff.rfl

@[simp]
protected theorem not_le [LT α] {xs ys : Array α} :
¬ xs ≤ ys ↔ ys < xs :=
Classical.not_not

@[deprecated Array.not_le (since := "2025-10-26")]
protected theorem not_le_iff_gt [LT α] {xs ys : Array α} :
¬ xs ≤ ys ↔ ys < xs :=
Classical.not_not
Expand Down Expand Up @@ -178,12 +189,6 @@ protected theorem le_total [LT α]
[i : Std.Asymm (· < · : α → α → Prop)] (xs ys : Array α) : xs ≤ ys ∨ ys ≤ xs :=
List.le_total xs.toList ys.toList

@[simp] protected theorem not_lt [LT α]
{xs ys : Array α} : ¬ xs < ys ↔ ys ≤ xs := Iff.rfl

@[simp] protected theorem not_le [LT α]
{xs ys : Array α} : ¬ ys ≤ xs ↔ xs < ys := Classical.not_not

protected theorem le_of_lt [LT α]
[i : Std.Asymm (· < · : α → α → Prop)]
{xs ys : Array α} (h : xs < ys) : xs ≤ ys :=
Expand Down
Loading
Loading