Skip to content

Commit 2fa3f94

Browse files
committed
feat: grind annotations for List/Array/Vector.zip functions
1 parent c2876a1 commit 2fa3f94

File tree

4 files changed

+96
-28
lines changed

4 files changed

+96
-28
lines changed

src/Init/Data/Array/Zip.lean

Lines changed: 29 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -45,6 +45,7 @@ theorem zipWith_self {f : α → α → δ} {xs : Array α} : zipWith f xs xs =
4545
See also `getElem?_zipWith'` for a variant
4646
using `Option.map` and `Option.bind` rather than a `match`.
4747
-/
48+
@[grind =]
4849
theorem getElem?_zipWith {f : α → β → γ} {i : Nat} :
4950
(zipWith f as bs)[i]? = match as[i]?, bs[i]? with
5051
| some a, some b => some (f a b) | _, _ => none := by
@@ -76,31 +77,35 @@ theorem getElem?_zip_eq_some {as : Array α} {bs : Array β} {z : α × β} {i :
7677
· rintro ⟨h₀, h₁⟩
7778
exact ⟨_, _, h₀, h₁, rfl⟩
7879

79-
@[simp]
80+
@[simp, grind =]
8081
theorem zipWith_map {μ} {f : γ → δ → μ} {g : α → γ} {h : β → δ} {as : Array α} {bs : Array β} :
8182
zipWith f (as.map g) (bs.map h) = zipWith (fun a b => f (g a) (h b)) as bs := by
8283
cases as
8384
cases bs
8485
simp [List.zipWith_map]
8586

87+
@[grind =]
8688
theorem zipWith_map_left {as : Array α} {bs : Array β} {f : α → α'} {g : α' → β → γ} :
8789
zipWith g (as.map f) bs = zipWith (fun a b => g (f a) b) as bs := by
8890
cases as
8991
cases bs
9092
simp [List.zipWith_map_left]
9193

94+
@[grind =]
9295
theorem zipWith_map_right {as : Array α} {bs : Array β} {f : β → β'} {g : α → β' → γ} :
9396
zipWith g as (bs.map f) = zipWith (fun a b => g a (f b)) as bs := by
9497
cases as
9598
cases bs
9699
simp [List.zipWith_map_right]
97100

101+
@[grind =]
98102
theorem zipWith_foldr_eq_zip_foldr {f : α → β → γ} {i : δ} :
99103
(zipWith f as bs).foldr g i = (zip as bs).foldr (fun p r => g (f p.1 p.2) r) i := by
100104
cases as
101105
cases bs
102106
simp [List.zipWith_foldr_eq_zip_foldr]
103107

108+
@[grind =]
104109
theorem zipWith_foldl_eq_zip_foldl {f : α → β → γ} {i : δ} :
105110
(zipWith f as bs).foldl g i = (zip as bs).foldl (fun r p => g r (f p.1 p.2)) i := by
106111
cases as
@@ -111,22 +116,26 @@ theorem zipWith_foldl_eq_zip_foldl {f : α → β → γ} {i : δ} :
111116
theorem zipWith_eq_empty_iff {f : α → β → γ} {as : Array α} {bs : Array β} : zipWith f as bs = #[] ↔ as = #[] ∨ bs = #[] := by
112117
cases as <;> cases bs <;> simp
113118

119+
@[grind =]
114120
theorem map_zipWith {δ : Type _} {f : α → β} {g : γ → δ → α} {cs : Array γ} {ds : Array δ} :
115121
map f (zipWith g cs ds) = zipWith (fun x y => f (g x y)) cs ds := by
116122
cases cs
117123
cases ds
118124
simp [List.map_zipWith]
119125

126+
@[grind =]
120127
theorem take_zipWith : (zipWith f as bs).take i = zipWith f (as.take i) (bs.take i) := by
121128
cases as
122129
cases bs
123130
simp [List.take_zipWith]
124131

132+
@[grind =]
125133
theorem extract_zipWith : (zipWith f as bs).extract i j = zipWith f (as.extract i j) (bs.extract i j) := by
126134
cases as
127135
cases bs
128136
simp [List.drop_zipWith, List.take_zipWith]
129137

138+
@[grind =]
130139
theorem zipWith_append {f : α → β → γ} {as as' : Array α} {bs bs' : Array β}
131140
(h : as.size = bs.size) :
132141
zipWith f (as ++ as') (bs ++ bs') = zipWith f as bs ++ zipWith f as' bs' := by
@@ -152,7 +161,7 @@ theorem zipWith_eq_append_iff {f : α → β → γ} {as : Array α} {bs : Array
152161
· rintro ⟨⟨ws⟩, ⟨xs⟩, ⟨ys⟩, ⟨zs⟩, h, rfl, rfl, h₁, h₂⟩
153162
exact ⟨ws, xs, ys, zs, by simp_all⟩
154163

155-
@[simp] theorem zipWith_replicate {a : α} {b : β} {m n : Nat} :
164+
@[simp, grind =] theorem zipWith_replicate {a : α} {b : β} {m n : Nat} :
156165
zipWith f (replicate m a) (replicate n b) = replicate (min m n) (f a b) := by
157166
simp [← List.toArray_replicate]
158167

@@ -184,6 +193,7 @@ theorem zipWith_eq_zipWith_take_min (as : Array α) (bs : Array β) :
184193
simp
185194
rw [List.zipWith_eq_zipWith_take_min]
186195

196+
@[grind =]
187197
theorem reverse_zipWith (h : as.size = bs.size) :
188198
(zipWith f as bs).reverse = zipWith f as.reverse bs.reverse := by
189199
cases as
@@ -200,7 +210,7 @@ theorem lt_size_right_of_zip {i : Nat} {as : Array α} {bs : Array β} (h : i <
200210
i < bs.size :=
201211
lt_size_right_of_zipWith h
202212

203-
@[simp]
213+
@[simp, grind =]
204214
theorem getElem_zip {as : Array α} {bs : Array β} {i : Nat} {h : i < (zip as bs).size} :
205215
(zip as bs)[i] =
206216
(as[i]'(lt_size_left_of_zip h), bs[i]'(lt_size_right_of_zip h)) :=
@@ -211,18 +221,22 @@ theorem zip_eq_zipWith {as : Array α} {bs : Array β} : zip as bs = zipWith Pro
211221
cases bs
212222
simp [List.zip_eq_zipWith]
213223

224+
@[grind _=_]
214225
theorem zip_map {f : α → γ} {g : β → δ} {as : Array α} {bs : Array β} :
215226
zip (as.map f) (bs.map g) = (zip as bs).map (Prod.map f g) := by
216227
cases as
217228
cases bs
218229
simp [List.zip_map]
219230

231+
@[grind _=_]
220232
theorem zip_map_left {f : α → γ} {as : Array α} {bs : Array β} :
221233
zip (as.map f) bs = (zip as bs).map (Prod.map f id) := by rw [← zip_map, map_id]
222234

235+
@[grind _=_]
223236
theorem zip_map_right {f : β → γ} {as : Array α} {bs : Array β} :
224237
zip as (bs.map f) = (zip as bs).map (Prod.map id f) := by rw [← zip_map, map_id]
225238

239+
@[grind =]
226240
theorem zip_append {as bs : Array α} {cs ds : Array β} (_h : as.size = cs.size) :
227241
zip (as ++ bs) (cs ++ ds) = zip as cs ++ zip bs ds := by
228242
cases as
@@ -231,6 +245,7 @@ theorem zip_append {as bs : Array α} {cs ds : Array β} (_h : as.size = cs.size
231245
cases ds
232246
simp_all [List.zip_append]
233247

248+
@[grind =]
234249
theorem zip_map' {f : α → β} {g : α → γ} {xs : Array α} :
235250
zip (xs.map f) (xs.map g) = xs.map fun a => (f a, g a) := by
236251
cases xs
@@ -276,7 +291,7 @@ theorem zip_eq_append_iff {as : Array α} {bs : Array β} :
276291
∃ as₁ as₂ bs₁ bs₂, as₁.size = bs₁.size ∧ as = as₁ ++ as₂ ∧ bs = bs₁ ++ bs₂ ∧ xs = zip as₁ bs₁ ∧ ys = zip as₂ bs₂ := by
277292
simp [zip_eq_zipWith, zipWith_eq_append_iff]
278293

279-
@[simp] theorem zip_replicate {a : α} {b : β} {m n : Nat} :
294+
@[simp, grind =] theorem zip_replicate {a : α} {b : β} {m n : Nat} :
280295
zip (replicate m a) (replicate n b) = replicate (min m n) (a, b) := by
281296
simp [← List.toArray_replicate]
282297

@@ -293,6 +308,7 @@ theorem zip_eq_zip_take_min {as : Array α} {bs : Array β} :
293308

294309
/-! ### zipWithAll -/
295310

311+
@[grind =]
296312
theorem getElem?_zipWithAll {f : Option α → Option β → γ} {i : Nat} :
297313
(zipWithAll f as bs)[i]? = match as[i]?, bs[i]? with
298314
| none, none => .none | a?, b? => some (f a? b?) := by
@@ -301,31 +317,35 @@ theorem getElem?_zipWithAll {f : Option α → Option β → γ} {i : Nat} :
301317
simp [List.getElem?_zipWithAll]
302318
rfl
303319

320+
@[grind =]
304321
theorem zipWithAll_map {μ} {f : Option γ → Option δ → μ} {g : α → γ} {h : β → δ} {as : Array α} {bs : Array β} :
305322
zipWithAll f (as.map g) (bs.map h) = zipWithAll (fun a b => f (g <$> a) (h <$> b)) as bs := by
306323
cases as
307324
cases bs
308325
simp [List.zipWithAll_map]
309326

327+
@[grind =]
310328
theorem zipWithAll_map_left {as : Array α} {bs : Array β} {f : α → α'} {g : Option α' → Option β → γ} :
311329
zipWithAll g (as.map f) bs = zipWithAll (fun a b => g (f <$> a) b) as bs := by
312330
cases as
313331
cases bs
314332
simp [List.zipWithAll_map_left]
315333

334+
@[grind =]
316335
theorem zipWithAll_map_right {as : Array α} {bs : Array β} {f : β → β'} {g : Option α → Option β' → γ} :
317336
zipWithAll g as (bs.map f) = zipWithAll (fun a b => g a (f <$> b)) as bs := by
318337
cases as
319338
cases bs
320339
simp [List.zipWithAll_map_right]
321340

341+
@[grind =]
322342
theorem map_zipWithAll {δ : Type _} {f : α → β} {g : Option γ → Option δ → α} {cs : Array γ} {ds : Array δ} :
323343
map f (zipWithAll g cs ds) = zipWithAll (fun x y => f (g x y)) cs ds := by
324344
cases cs
325345
cases ds
326346
simp [List.map_zipWithAll]
327347

328-
@[simp] theorem zipWithAll_replicate {a : α} {b : β} {n : Nat} :
348+
@[simp, grind =] theorem zipWithAll_replicate {a : α} {b : β} {n : Nat} :
329349
zipWithAll f (replicate n a) (replicate n b) = replicate n (f (some a) (some b)) := by
330350
simp [← List.toArray_replicate]
331351

@@ -342,6 +362,7 @@ theorem unzip_fst : (unzip l).fst = l.map Prod.fst := by
342362
theorem unzip_snd : (unzip l).snd = l.map Prod.snd := by
343363
simp
344364

365+
@[grind =]
345366
theorem unzip_eq_map {xs : Array (α × β)} : unzip xs = (xs.map Prod.fst, xs.map Prod.snd) := by
346367
cases xs
347368
simp [List.unzip_eq_map]
@@ -375,9 +396,11 @@ theorem zip_of_prod {as : Array α} {bs : Array β} {xs : Array (α × β)} (hl
375396
(hr : xs.map Prod.snd = bs) : xs = as.zip bs := by
376397
rw [← hl, ← hr, ← zip_unzip xs, ← fst_unzip, ← snd_unzip, zip_unzip, zip_unzip]
377398

378-
@[simp] theorem unzip_replicate {n : Nat} {a : α} {b : β} :
399+
@[simp, grind =] theorem unzip_replicate {n : Nat} {a : α} {b : β} :
379400
unzip (replicate n (a, b)) = (replicate n a, replicate n b) := by
380401
ext1 <;> simp
381402

382403
@[deprecated unzip_replicate (since := "2025-03-18")]
383404
abbrev unzip_mkArray := @unzip_replicate
405+
406+
end Array

src/Init/Data/List/Nat/TakeDrop.lean

Lines changed: 7 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -538,7 +538,7 @@ theorem dropWhile_eq_drop_findIdx_not {xs : List α} {p : α → Bool} :
538538

539539
/-! ### zipWith -/
540540

541-
@[simp] theorem length_zipWith {f : α → β → γ} {l₁ : List α} {l₂ : List β} :
541+
@[simp, grind =] theorem length_zipWith {f : α → β → γ} {l₁ : List α} {l₂ : List β} :
542542
length (zipWith f l₁ l₂) = min (length l₁) (length l₂) := by
543543
induction l₁ generalizing l₂ <;> cases l₂ <;>
544544
simp_all [succ_min_succ, Nat.zero_min, Nat.min_zero]
@@ -549,7 +549,7 @@ theorem lt_length_left_of_zipWith {f : α → β → γ} {i : Nat} {l : List α}
549549
theorem lt_length_right_of_zipWith {f : α → β → γ} {i : Nat} {l : List α} {l' : List β}
550550
(h : i < (zipWith f l l').length) : i < l'.length := by rw [length_zipWith] at h; omega
551551

552-
@[simp]
552+
@[simp, grind =]
553553
theorem getElem_zipWith {f : α → β → γ} {l : List α} {l' : List β}
554554
{i : Nat} {h : i < (zipWith f l l').length} :
555555
(zipWith f l l')[i] =
@@ -566,6 +566,7 @@ theorem zipWith_eq_zipWith_take_min : ∀ {l₁ : List α} {l₂ : List β},
566566
| _, [] => by simp
567567
| a :: l₁, b :: l₂ => by simp [succ_min_succ, zipWith_eq_zipWith_take_min (l₁ := l₁) (l₂ := l₂)]
568568

569+
@[grind =]
569570
theorem reverse_zipWith (h : l.length = l'.length) :
570571
(zipWith f l l').reverse = zipWith f l.reverse l'.reverse := by
571572
induction l generalizing l' with
@@ -578,14 +579,14 @@ theorem reverse_zipWith (h : l.length = l'.length) :
578579
have : tl.reverse.length = tl'.reverse.length := by simp [h]
579580
simp [hl h, zipWith_append this]
580581

581-
@[simp] theorem zipWith_replicate {a : α} {b : β} {m n : Nat} :
582+
@[simp, grind =] theorem zipWith_replicate {a : α} {b : β} {m n : Nat} :
582583
zipWith f (replicate m a) (replicate n b) = replicate (min m n) (f a b) := by
583584
rw [zipWith_eq_zipWith_take_min]
584585
simp
585586

586587
/-! ### zip -/
587588

588-
@[simp] theorem length_zip {l₁ : List α} {l₂ : List β} :
589+
@[simp, grind =] theorem length_zip {l₁ : List α} {l₂ : List β} :
589590
length (zip l₁ l₂) = min (length l₁) (length l₂) := by
590591
simp [zip]
591592

@@ -597,7 +598,7 @@ theorem lt_length_right_of_zip {i : Nat} {l : List α} {l' : List β} (h : i < (
597598
i < l'.length :=
598599
lt_length_right_of_zipWith h
599600

600-
@[simp]
601+
@[simp, grind =]
601602
theorem getElem_zip {l : List α} {l' : List β} {i : Nat} {h : i < (zip l l').length} :
602603
(zip l l')[i] =
603604
(l[i]'(lt_length_left_of_zip h), l'[i]'(lt_length_right_of_zip h)) :=
@@ -609,7 +610,7 @@ theorem zip_eq_zip_take_min : ∀ {l₁ : List α} {l₂ : List β},
609610
| _, [] => by simp
610611
| a :: l₁, b :: l₂ => by simp [succ_min_succ, zip_eq_zip_take_min (l₁ := l₁) (l₂ := l₂)]
611612

612-
@[simp] theorem zip_replicate {a : α} {b : β} {m n : Nat} :
613+
@[simp, grind =] theorem zip_replicate {a : α} {b : β} {m n : Nat} :
613614
zip (replicate m a) (replicate n b) = replicate (min m n) (a, b) := by
614615
rw [zip_eq_zip_take_min]
615616
simp

0 commit comments

Comments
 (0)