Skip to content

Commit 1e752b0

Browse files
authored
chore: cleanup simp lemmas, following the simpNF linter (#8481)
1 parent 11f7d6d commit 1e752b0

37 files changed

+221
-259
lines changed

src/Init/Data/Array/Attach.lean

Lines changed: 6 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -69,11 +69,11 @@ well-founded recursion mechanism to prove that the function terminates.
6969
simp [pmap]
7070

7171
@[simp] theorem toList_attachWith {xs : Array α} {P : α → Prop} {H : ∀ x ∈ xs, P x} :
72-
(xs.attachWith P H).toList = xs.toList.attachWith P (by simpa [mem_toList] using H) := by
72+
(xs.attachWith P H).toList = xs.toList.attachWith P (by simpa [mem_toList_iff] using H) := by
7373
simp [attachWith]
7474

7575
@[simp] theorem toList_attach {xs : Array α} :
76-
xs.attach.toList = xs.toList.attachWith (· ∈ xs) (by simp [mem_toList]) := by
76+
xs.attach.toList = xs.toList.attachWith (· ∈ xs) (by simp [mem_toList_iff]) := by
7777
simp [attach]
7878

7979
@[simp] theorem toList_pmap {xs : Array α} {P : α → Prop} {f : ∀ a, P a → β} {H : ∀ a ∈ xs, P a} :
@@ -574,9 +574,12 @@ state, the right approach is usually the tactic `simp [Array.unattach, -Array.ma
574574
-/
575575
def unattach {α : Type _} {p : α → Prop} (xs : Array { x // p x }) : Array α := xs.map (·.val)
576576

577-
@[simp] theorem unattach_nil {p : α → Prop} : (#[] : Array { x // p x }).unattach = #[] := by
577+
@[simp] theorem unattach_empty {p : α → Prop} : (#[] : Array { x // p x }).unattach = #[] := by
578578
simp [unattach]
579579

580+
@[deprecated unattach_empty (since := "2025-05-26")]
581+
abbrev unattach_nil := @unattach_empty
582+
580583
@[simp] theorem unattach_push {p : α → Prop} {a : { x // p x }} {xs : Array { x // p x }} :
581584
(xs.push a).unattach = xs.unattach.push a.1 := by
582585
simp only [unattach, Array.map_push]

src/Init/Data/Array/Bootstrap.lean

Lines changed: 5 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -89,9 +89,13 @@ theorem foldrM_eq_reverse_foldlM_toList [Monad m] {f : α → β → m β} {init
8989
xs.toList.foldr f init = xs.foldr f init :=
9090
List.foldr_eq_foldrM .. ▸ foldrM_toList ..
9191

92-
@[simp, grind =] theorem push_toList {xs : Array α} {a : α} : (xs.push a).toList = xs.toList ++ [a] := by
92+
@[simp, grind =] theorem toList_push {xs : Array α} {x : α} : (xs.push x).toList = xs.toList ++ [x] := by
93+
rcases xs with ⟨xs⟩
9394
simp [push, List.concat_eq_append]
9495

96+
@[deprecated toList_push (since := "2025-05-26")]
97+
abbrev push_toList := @toList_push
98+
9599
@[simp, grind =] theorem toListAppend_eq {xs : Array α} {l : List α} : xs.toListAppend l = xs.toList ++ l := by
96100
simp [toListAppend, ← foldr_toList]
97101

src/Init/Data/Array/Count.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -52,8 +52,8 @@ theorem countP_push {a : α} {xs : Array α} : countP p (xs.push a) = countP p x
5252
rcases xs with ⟨xs⟩
5353
simp_all
5454

55-
@[simp] theorem countP_singleton {a : α} : countP p #[a] = if p a then 1 else 0 := by
56-
simp [countP_push]
55+
theorem countP_singleton {a : α} : countP p #[a] = if p a then 1 else 0 := by
56+
simp
5757

5858
theorem size_eq_countP_add_countP {xs : Array α} : xs.size = countP p xs + countP (fun a => ¬p a) xs := by
5959
rcases xs with ⟨xs⟩

src/Init/Data/Array/Erase.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -24,7 +24,7 @@ open Nat
2424

2525
/-! ### eraseP -/
2626

27-
@[simp] theorem eraseP_empty : #[].eraseP p = #[] := by simp
27+
theorem eraseP_empty : #[].eraseP p = #[] := by simp
2828

2929
theorem eraseP_of_forall_mem_not {xs : Array α} (h : ∀ a, a ∈ xs → ¬p a) : xs.eraseP p = xs := by
3030
rcases xs with ⟨xs⟩

src/Init/Data/Array/Extract.lean

Lines changed: 2 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -238,11 +238,9 @@ theorem extract_append_left {as bs : Array α} :
238238
(as ++ bs).extract 0 as.size = as.extract 0 as.size := by
239239
simp
240240

241-
@[simp] theorem extract_append_right {as bs : Array α} :
241+
theorem extract_append_right {as bs : Array α} :
242242
(as ++ bs).extract as.size (as.size + i) = bs.extract 0 i := by
243-
simp only [extract_append, extract_size_left, Nat.sub_self, empty_append]
244-
congr 1
245-
omega
243+
simp
246244

247245
@[simp] theorem map_extract {as : Array α} {i j : Nat} :
248246
(as.extract i j).map f = (as.map f).extract i j := by

src/Init/Data/Array/Find.lean

Lines changed: 8 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -142,9 +142,9 @@ abbrev findSome?_mkArray_of_isNone := @findSome?_replicate_of_isNone
142142

143143
@[simp] theorem find?_empty : find? p #[] = none := rfl
144144

145-
@[simp] theorem find?_singleton {a : α} {p : α → Bool} :
145+
theorem find?_singleton {a : α} {p : α → Bool} :
146146
#[a].find? p = if p a then some a else none := by
147-
simp [singleton_eq_toArray_singleton]
147+
simp
148148

149149
@[simp] theorem findRev?_push_of_pos {xs : Array α} (h : p a) :
150150
findRev? p (xs.push a) = some a := by
@@ -347,7 +347,8 @@ theorem find?_eq_some_iff_getElem {xs : Array α} {p : α → Bool} {b : α} :
347347

348348
/-! ### findIdx -/
349349

350-
@[simp] theorem findIdx_empty : findIdx p #[] = 0 := rfl
350+
theorem findIdx_empty : findIdx p #[] = 0 := rfl
351+
351352
theorem findIdx_singleton {a : α} {p : α → Bool} :
352353
#[a].findIdx p = if p a then 0 else 1 := by
353354
simp
@@ -600,7 +601,8 @@ theorem findIdx?_eq_some_le_of_findIdx?_eq_some {xs : Array α} {p q : α → Bo
600601

601602
/-! ### findFinIdx? -/
602603

603-
@[simp] theorem findFinIdx?_empty {p : α → Bool} : findFinIdx? p #[] = none := by simp
604+
theorem findFinIdx?_empty {p : α → Bool} : findFinIdx? p #[] = none := by simp
605+
604606
theorem findFinIdx?_singleton {a : α} {p : α → Bool} :
605607
#[a].findFinIdx? p = if p a then some ⟨0, by simp⟩ else none := by
606608
simp
@@ -699,7 +701,7 @@ The verification API for `idxOf?` is still incomplete.
699701
The lemmas below should be made consistent with those for `findIdx?` (and proved using them).
700702
-/
701703

702-
@[simp] theorem idxOf?_empty [BEq α] : (#[] : Array α).idxOf? a = none := by simp
704+
theorem idxOf?_empty [BEq α] : (#[] : Array α).idxOf? a = none := by simp
703705

704706
@[simp] theorem idxOf?_eq_none_iff [BEq α] [LawfulBEq α] {xs : Array α} {a : α} :
705707
xs.idxOf? a = none ↔ a ∉ xs := by
@@ -712,14 +714,10 @@ theorem isSome_idxOf? [BEq α] [LawfulBEq α] {xs : Array α} {a : α} :
712714
rcases xs with ⟨xs⟩
713715
simp
714716

715-
@[simp]
716717
theorem isNone_idxOf? [BEq α] [LawfulBEq α] {xs : Array α} {a : α} :
717718
(xs.idxOf? a).isNone = ¬ a ∈ xs := by
718-
rcases xs with ⟨xs⟩
719719
simp
720720

721-
722-
723721
/-! ### finIdxOf?
724722
725723
The verification API for `finIdxOf?` is still incomplete.
@@ -730,7 +728,7 @@ theorem idxOf?_eq_map_finIdxOf?_val [BEq α] {xs : Array α} {a : α} :
730728
xs.idxOf? a = (xs.finIdxOf? a).map (·.val) := by
731729
simp [idxOf?, finIdxOf?, findIdx?_eq_map_findFinIdx?_val]
732730

733-
@[simp] theorem finIdxOf?_empty [BEq α] : (#[] : Array α).finIdxOf? a = none := by simp
731+
theorem finIdxOf?_empty [BEq α] : (#[] : Array α).finIdxOf? a = none := by simp
734732

735733
@[simp] theorem finIdxOf?_eq_none_iff [BEq α] [LawfulBEq α] {xs : Array α} {a : α} :
736734
xs.finIdxOf? a = none ↔ a ∉ xs := by
@@ -748,10 +746,8 @@ theorem isSome_finIdxOf? [BEq α] [LawfulBEq α] {xs : Array α} {a : α} :
748746
rcases xs with ⟨xs⟩
749747
simp
750748

751-
@[simp]
752749
theorem isNone_finIdxOf? [BEq α] [LawfulBEq α] {xs : Array α} {a : α} :
753750
(xs.finIdxOf? a).isNone = ¬ a ∈ xs := by
754-
rcases xs with ⟨xs⟩
755751
simp
756752

757753
end Array

src/Init/Data/Array/Lemmas.lean

Lines changed: 43 additions & 56 deletions
Original file line numberDiff line numberDiff line change
@@ -140,13 +140,13 @@ theorem getElem_of_getElem? {xs : Array α} : xs[i]? = some a → ∃ h : i < xs
140140
theorem some_eq_getElem?_iff {xs : Array α} : some b = xs[i]? ↔ ∃ h : i < xs.size, xs[i] = b := by
141141
rw [eq_comm, getElem?_eq_some_iff]
142142

143-
@[simp] theorem some_getElem_eq_getElem?_iff (xs : Array α) (i : Nat) (h : i < xs.size) :
143+
theorem some_getElem_eq_getElem?_iff (xs : Array α) (i : Nat) (h : i < xs.size) :
144144
(some xs[i] = xs[i]?) ↔ True := by
145-
simp [h]
145+
simp
146146

147-
@[simp] theorem getElem?_eq_some_getElem_iff (xs : Array α) (i : Nat) (h : i < xs.size) :
147+
theorem getElem?_eq_some_getElem_iff (xs : Array α) (i : Nat) (h : i < xs.size) :
148148
(xs[i]? = some xs[i]) ↔ True := by
149-
simp [h]
149+
simp
150150

151151
theorem getElem_eq_iff {xs : Array α} {i : Nat} {h : i < xs.size} : xs[i] = x ↔ xs[i]? = some x := by
152152
simp only [getElem?_eq_some_iff]
@@ -186,12 +186,11 @@ theorem getElem?_push {xs : Array α} {x} : (xs.push x)[i]? = if i = xs.size the
186186
simp [getElem?_def, getElem_push]
187187
(repeat' split) <;> first | rfl | omega
188188

189-
@[simp] theorem getElem?_push_size {xs : Array α} {x} : (xs.push x)[xs.size]? = some x := by
190-
simp [getElem?_push]
189+
theorem getElem?_push_size {xs : Array α} {x} : (xs.push x)[xs.size]? = some x := by
190+
simp
191191

192-
@[simp] theorem getElem_singleton {a : α} {i : Nat} (h : i < 1) : #[a][i] = a :=
193-
match i, h with
194-
| 0, _ => rfl
192+
theorem getElem_singleton {a : α} {i : Nat} (h : i < 1) : #[a][i] = a := by
193+
simp
195194

196195
@[grind]
197196
theorem getElem?_singleton {a : α} {i : Nat} : #[a][i]? = if i = 0 then some a else none := by
@@ -240,10 +239,6 @@ theorem back?_pop {xs : Array α} :
240239

241240
@[simp] theorem push_empty : #[].push x = #[x] := rfl
242241

243-
@[simp] theorem toList_push {xs : Array α} {x : α} : (xs.push x).toList = xs.toList ++ [x] := by
244-
rcases xs with ⟨xs⟩
245-
simp
246-
247242
@[simp] theorem push_ne_empty {a : α} {xs : Array α} : xs.push a ≠ #[] := by
248243
cases xs
249244
simp
@@ -423,8 +418,7 @@ theorem eq_empty_iff_forall_not_mem {xs : Array α} : xs = #[] ↔ ∀ a, a ∉
423418
theorem eq_of_mem_singleton (h : a ∈ #[b]) : a = b := by
424419
simpa using h
425420

426-
@[simp] theorem mem_singleton {a b : α} : a ∈ #[b] ↔ a = b :=
427-
⟨eq_of_mem_singleton, (by simp [·])⟩
421+
theorem mem_singleton {a b : α} : a ∈ #[b] ↔ a = b := by simp
428422

429423
theorem forall_mem_push {p : α → Prop} {xs : Array α} {a : α} :
430424
(∀ x, x ∈ xs.push a → p x) ↔ p a ∧ ∀ x, x ∈ xs → p x := by
@@ -868,8 +862,8 @@ theorem elem_eq_mem [BEq α] [LawfulBEq α] {a : α} {xs : Array α} :
868862
@[simp, grind] theorem contains_eq_mem [BEq α] [LawfulBEq α] {a : α} {xs : Array α} :
869863
xs.contains a = decide (a ∈ xs) := by rw [← elem_eq_contains, elem_eq_mem]
870864

871-
@[simp, grind] theorem any_empty [BEq α] {p : α → Bool} : (#[] : Array α).any p = false := by simp
872-
@[simp, grind] theorem all_empty [BEq α] {p : α → Bool} : (#[] : Array α).all p = true := by simp
865+
@[grind] theorem any_empty [BEq α] {p : α → Bool} : (#[] : Array α).any p = false := by simp
866+
@[grind] theorem all_empty [BEq α] {p : α → Bool} : (#[] : Array α).all p = true := by simp
873867

874868
/-- Variant of `any_push` with a side condition on `stop`. -/
875869
@[simp, grind] theorem any_push' [BEq α] {xs : Array α} {a : α} {p : α → Bool} (h : stop = xs.size + 1) :
@@ -1228,7 +1222,7 @@ where
12281222
@[simp] theorem mapM_empty [Monad m] (f : α → m β) : mapM f #[] = pure #[] := by
12291223
rw [mapM, mapM.map]; rfl
12301224

1231-
@[simp, grind] theorem map_empty {f : α → β} : map f #[] = #[] := mapM_empty f
1225+
@[grind] theorem map_empty {f : α → β} : map f #[] = #[] := by simp
12321226

12331227
@[simp, grind] theorem map_push {f : α → β} {as : Array α} {x : α} :
12341228
(as.push x).map f = (as.map f).push (f x) := by
@@ -1813,7 +1807,8 @@ theorem toArray_append {xs : List α} {ys : Array α} :
18131807

18141808
theorem singleton_eq_toArray_singleton {a : α} : #[a] = [a].toArray := rfl
18151809

1816-
@[simp] theorem empty_append_fun : ((#[] : Array α) ++ ·) = id := by
1810+
@[deprecated empty_append (since := "2025-05-26")]
1811+
theorem empty_append_fun : ((#[] : Array α) ++ ·) = id := by
18171812
funext ⟨l⟩
18181813
simp
18191814

@@ -1964,8 +1959,8 @@ theorem append_left_inj {xs₁ xs₂ : Array α} (ys) : xs₁ ++ ys = xs₂ ++ y
19641959
theorem eq_empty_of_append_eq_empty {xs ys : Array α} (h : xs ++ ys = #[]) : xs = #[] ∧ ys = #[] :=
19651960
append_eq_empty_iff.mp h
19661961

1967-
@[simp] theorem empty_eq_append_iff {xs ys : Array α} : #[] = xs ++ ys ↔ xs = #[] ∧ ys = #[] := by
1968-
rw [eq_comm, append_eq_empty_iff]
1962+
theorem empty_eq_append_iff {xs ys : Array α} : #[] = xs ++ ys ↔ xs = #[] ∧ ys = #[] := by
1963+
simp
19691964

19701965
theorem append_ne_empty_of_left_ne_empty {xs ys : Array α} (h : xs ≠ #[]) : xs ++ ys ≠ #[] := by
19711966
simp_all
@@ -2033,7 +2028,7 @@ theorem append_eq_append_iff {ws xs ys zs : Array α} :
20332028
simp only [List.append_toArray, List.set_toArray, List.set_append]
20342029
split <;> simp
20352030

2036-
@[simp] theorem set_append_left {xs ys : Array α} {i : Nat} {x : α} (h : i < xs.size) :
2031+
@[simp] theorem set_append_left {xs ys : Array α} {i : Nat} {x : α} (h : i < xs.size) :
20372032
(xs ++ ys).set i x (by simp; omega) = xs.set i x ++ ys := by
20382033
simp [set_append, h]
20392034

@@ -2108,14 +2103,13 @@ theorem append_eq_map_iff {f : α → β} :
21082103
| nil => simp
21092104
| cons as => induction as.toList <;> simp [*]
21102105

2111-
@[simp] theorem flatten_map_toArray {L : List (List α)} :
2112-
(L.toArray.map List.toArray).flatten = L.flatten.toArray := by
2106+
@[simp] theorem flatten_toArray_map {L : List (List α)} :
2107+
(L.map List.toArray).toArray.flatten = L.flatten.toArray := by
21132108
apply ext'
21142109
simp [Function.comp_def]
21152110

2116-
@[simp] theorem flatten_toArray_map {L : List (List α)} :
2117-
(L.map List.toArray).toArray.flatten = L.flatten.toArray := by
2118-
rw [← flatten_map_toArray]
2111+
theorem flatten_map_toArray {L : List (List α)} :
2112+
(L.toArray.map List.toArray).flatten = L.flatten.toArray := by
21192113
simp
21202114

21212115
-- We set this to lower priority so that `flatten_toArray_map` is applied first when relevant.
@@ -2143,8 +2137,8 @@ theorem mem_flatten : ∀ {xss : Array (Array α)}, a ∈ xss.flatten ↔ ∃ xs
21432137
induction xss using array₂_induction
21442138
simp
21452139

2146-
@[simp] theorem empty_eq_flatten_iff {xss : Array (Array α)} : #[] = xss.flatten ↔ ∀ xs ∈ xss, xs = #[] := by
2147-
rw [eq_comm, flatten_eq_empty_iff]
2140+
theorem empty_eq_flatten_iff {xss : Array (Array α)} : #[] = xss.flatten ↔ ∀ xs ∈ xss, xs = #[] := by
2141+
simp
21482142

21492143
theorem flatten_ne_empty_iff {xss : Array (Array α)} : xss.flatten ≠ #[] ↔ ∃ xs, xs ∈ xss ∧ xs ≠ #[] := by
21502144
simp
@@ -2284,15 +2278,9 @@ theorem eq_iff_flatten_eq {xss₁ xss₂ : Array (Array α)} :
22842278
rw [List.map_inj_right]
22852279
simp +contextual
22862280

2287-
@[simp] theorem flatten_toArray_map_toArray {xss : List (List α)} :
2281+
theorem flatten_toArray_map_toArray {xss : List (List α)} :
22882282
(xss.map List.toArray).toArray.flatten = xss.flatten.toArray := by
2289-
simp [flatten]
2290-
suffices ∀ as, List.foldl (fun acc bs => acc ++ bs) as (List.map List.toArray xss) = as ++ xss.flatten.toArray by
2291-
simpa using this #[]
2292-
intro as
2293-
induction xss generalizing as with
2294-
| nil => simp
2295-
| cons xs xss ih => simp [ih]
2283+
simp
22962284

22972285
/-! ### flatMap -/
22982286

@@ -2322,13 +2310,9 @@ theorem flatMap_toArray_cons {β} {f : α → Array β} {a : α} {as : List α}
23222310
intro cs
23232311
induction as generalizing cs <;> simp_all
23242312

2325-
@[simp, grind =] theorem flatMap_toArray {β} {f : α → Array β} {as : List α} :
2313+
theorem flatMap_toArray {β} {f : α → Array β} {as : List α} :
23262314
as.toArray.flatMap f = (as.flatMap (fun a => (f a).toList)).toArray := by
2327-
induction as with
2328-
| nil => simp
2329-
| cons a as ih =>
2330-
apply ext'
2331-
simp [ih, flatMap_toArray_cons]
2315+
simp
23322316

23332317
@[simp] theorem flatMap_id {xss : Array (Array α)} : xss.flatMap id = xss.flatten := by simp [flatMap_def]
23342318

@@ -3030,19 +3014,21 @@ theorem take_size {xs : Array α} : xs.take xs.size = xs := by
30303014
| succ n ih =>
30313015
simp [shrink.loop, ih]
30323016

3033-
@[simp] theorem size_shrink {xs : Array α} {i : Nat} : (xs.shrink i).size = min i xs.size := by
3017+
-- This doesn't need to be a simp lemma, as shortly we will simplify `shrink` to `take`.
3018+
theorem size_shrink {xs : Array α} {i : Nat} : (xs.shrink i).size = min i xs.size := by
30343019
simp [shrink]
30353020
omega
30363021

3037-
@[simp] theorem getElem_shrink {xs : Array α} {i j : Nat} (h : j < (xs.shrink i).size) :
3038-
(xs.shrink i)[j] = xs[j]'(by simp at h; omega) := by
3022+
-- This doesn't need to be a simp lemma, as shortly we will simplify `shrink` to `take`.
3023+
theorem getElem_shrink {xs : Array α} {i j : Nat} (h : j < (xs.shrink i).size) :
3024+
(xs.shrink i)[j] = xs[j]'(by simp [size_shrink] at h; omega) := by
30393025
simp [shrink]
30403026

3041-
@[simp] theorem toList_shrink {xs : Array α} {i : Nat} : (xs.shrink i).toList = xs.toList.take i := by
3042-
apply List.ext_getElem <;> simp
3043-
30443027
@[simp] theorem shrink_eq_take {xs : Array α} {i : Nat} : xs.shrink i = xs.take i := by
3045-
ext <;> simp
3028+
ext <;> simp [size_shrink, getElem_shrink]
3029+
3030+
theorem toList_shrink {xs : Array α} {i : Nat} : (xs.shrink i).toList = xs.toList.take i := by
3031+
simp
30463032

30473033
/-! ### foldlM and foldrM -/
30483034

@@ -3249,7 +3235,7 @@ theorem foldrM_reverse [Monad m] {xs : Array α} {f : α → β → m β} {b} :
32493235

32503236
theorem foldrM_push [Monad m] {f : α → β → m β} {init : β} {xs : Array α} {a : α} :
32513237
(xs.push a).foldrM f init = f a init >>= xs.foldrM f := by
3252-
simp only [foldrM_eq_reverse_foldlM_toList, push_toList, List.reverse_append, List.reverse_cons,
3238+
simp only [foldrM_eq_reverse_foldlM_toList, toList_push, List.reverse_append, List.reverse_cons,
32533239
List.reverse_nil, List.nil_append, List.singleton_append, List.foldlM_cons, List.foldlM_reverse]
32543240

32553241
/--
@@ -3654,7 +3640,7 @@ theorem foldr_rel {xs : Array α} {f g : α → β → β} {a b : β} {r : β
36543640
theorem back?_eq_some_iff {xs : Array α} {a : α} :
36553641
xs.back? = some a ↔ ∃ ys : Array α, xs = ys.push a := by
36563642
rcases xs with ⟨xs⟩
3657-
simp only [List.back?_toArray, List.getLast?_eq_some_iff, toArray_eq, push_toList]
3643+
simp only [List.back?_toArray, List.getLast?_eq_some_iff, toArray_eq, toList_push]
36583644
constructor
36593645
· rintro ⟨ys, rfl⟩
36603646
exact ⟨ys.toArray, by simp⟩
@@ -4425,7 +4411,8 @@ theorem getElem!_eq_getD [Inhabited α] {xs : Array α} {i} : xs[i]! = xs.getD i
44254411

44264412
/-! # mem -/
44274413

4428-
@[simp, grind =] theorem mem_toList {a : α} {xs : Array α} : a ∈ xs.toList ↔ a ∈ xs := mem_def.symm
4414+
@[deprecated mem_toList_iff (since := "2025-05-26")]
4415+
theorem mem_toList {a : α} {xs : Array α} : a ∈ xs.toList ↔ a ∈ xs := mem_def.symm
44294416

44304417
@[deprecated not_mem_empty (since := "2025-03-25")]
44314418
theorem not_mem_nil (a : α) : ¬ a ∈ #[] := nofun
@@ -4474,7 +4461,7 @@ theorem getElem?_push_eq {xs : Array α} {x : α} : (xs.push x)[xs.size]? = some
44744461
simp
44754462

44764463
@[simp, grind =] theorem forIn'_toList [Monad m] {xs : Array α} {b : β} {f : (a : α) → a ∈ xs.toList → β → m (ForInStep β)} :
4477-
forIn' xs.toList b f = forIn' xs b (fun a m b => f a (mem_toList.mpr m) b) := by
4464+
forIn' xs.toList b f = forIn' xs b (fun a m b => f a (mem_toList_iff.mpr m) b) := by
44784465
cases xs
44794466
simp
44804467

@@ -4571,8 +4558,8 @@ Our goal is to have `simp` "pull `List.toArray` outwards" as much as possible.
45714558

45724559
theorem toListRev_toArray {l : List α} : l.toArray.toListRev = l.reverse := by simp
45734560

4574-
@[simp, grind =] theorem take_toArray {l : List α} {i : Nat} : l.toArray.take i = (l.take i).toArray := by
4575-
apply Array.ext <;> simp
4561+
@[grind =] theorem take_toArray {l : List α} {i : Nat} : l.toArray.take i = (l.take i).toArray := by
4562+
simp
45764563

45774564
@[simp, grind =] theorem mapM_toArray [Monad m] [LawfulMonad m] {f : α → m β} {l : List α} :
45784565
l.toArray.mapM f = List.toArray <$> l.mapM f := by

0 commit comments

Comments
 (0)