Skip to content

Commit 781e3c6

Browse files
authored
chore: remove unhelpful grind annotations (#10435)
This PR removes some `grind` annotations for `Array.attach` and related functions. These lemmas introduce lambda on the right hand side which `grind` can't do much with. I've added a test file that verifies that the theorems with removed annotations can actually be proved already by grind. Removing the annotations will help with excessive instantiation.
1 parent b73b8a7 commit 781e3c6

File tree

6 files changed

+121
-94
lines changed

6 files changed

+121
-94
lines changed

src/Init/Data/Array/Attach.lean

Lines changed: 18 additions & 28 deletions
Original file line numberDiff line numberDiff line change
@@ -95,16 +95,16 @@ well-founded recursion mechanism to prove that the function terminates.
9595
intro a m h₁ h₂
9696
congr
9797

98-
@[simp, grind =] theorem pmap_empty {P : α → Prop} (f : ∀ a, P a → β) : pmap f #[] (by simp) = #[] := rfl
98+
@[simp] theorem pmap_empty {P : α → Prop} (f : ∀ a, P a → β) : pmap f #[] (by simp) = #[] := rfl
9999

100-
@[simp, grind =] theorem pmap_push {P : α → Prop} (f : ∀ a, P a → β) (a : α) (xs : Array α) (h : ∀ b ∈ xs.push a, P b) :
100+
@[simp] theorem pmap_push {P : α → Prop} (f : ∀ a, P a → β) (a : α) (xs : Array α) (h : ∀ b ∈ xs.push a, P b) :
101101
pmap f (xs.push a) h =
102102
(pmap f xs (fun a m => by simp at h; exact h a (.inl m))).push (f a (h a (by simp))) := by
103103
simp [pmap]
104104

105-
@[simp, grind =] theorem attach_empty : (#[] : Array α).attach = #[] := rfl
105+
@[simp] theorem attach_empty : (#[] : Array α).attach = #[] := rfl
106106

107-
@[simp, grind =] theorem attachWith_empty {P : α → Prop} (H : ∀ x ∈ #[], P x) : (#[] : Array α).attachWith P H = #[] := rfl
107+
@[simp] theorem attachWith_empty {P : α → Prop} (H : ∀ x ∈ #[], P x) : (#[] : Array α).attachWith P H = #[] := rfl
108108

109109
@[simp] theorem _root_.List.attachWith_mem_toArray {l : List α} :
110110
l.attachWith (fun x => x ∈ l.toArray) (fun x h => by simpa using h) =
@@ -125,13 +125,11 @@ theorem pmap_congr_left {p q : α → Prop} {f : ∀ a, p a → β} {g : ∀ a,
125125
simp only [List.pmap_toArray, mk.injEq]
126126
rw [List.pmap_congr_left _ h]
127127

128-
@[grind =]
129128
theorem map_pmap {p : α → Prop} {g : β → γ} {f : ∀ a, p a → β} {xs : Array α} (H) :
130129
map g (pmap f xs H) = pmap (fun a h => g (f a h)) xs H := by
131130
cases xs
132131
simp [List.map_pmap]
133132

134-
@[grind =]
135133
theorem pmap_map {p : β → Prop} {g : ∀ b, p b → γ} {f : α → β} {xs : Array α} (H) :
136134
pmap g (map f xs) H = pmap (fun a h => g (f a) h) xs fun _ h => H _ (mem_map_of_mem h) := by
137135
cases xs
@@ -147,14 +145,14 @@ theorem attachWith_congr {xs ys : Array α} (w : xs = ys) {P : α → Prop} {H :
147145
subst w
148146
simp
149147

150-
@[simp, grind =] theorem attach_push {a : α} {xs : Array α} :
148+
@[simp] theorem attach_push {a : α} {xs : Array α} :
151149
(xs.push a).attach =
152150
(xs.attach.map (fun ⟨x, h⟩ => ⟨x, mem_push_of_mem a h⟩)).push ⟨a, by simp⟩ := by
153151
cases xs
154152
rw [attach_congr (List.push_toArray _ _)]
155153
simp [Function.comp_def]
156154

157-
@[simp, grind =] theorem attachWith_push {a : α} {xs : Array α} {P : α → Prop} {H : ∀ x ∈ xs.push a, P x} :
155+
@[simp] theorem attachWith_push {a : α} {xs : Array α} {P : α → Prop} {H : ∀ x ∈ xs.push a, P x} :
158156
(xs.push a).attachWith P H =
159157
(xs.attachWith P (fun x h => by simp at H; exact H x (.inl h))).push ⟨a, H a (by simp)⟩ := by
160158
cases xs
@@ -288,25 +286,23 @@ theorem getElem_attach {xs : Array α} {i : Nat} (h : i < xs.attach.size) :
288286
xs.attach[i] = ⟨xs[i]'(by simpa using h), getElem_mem (by simpa using h)⟩ :=
289287
getElem_attachWith h
290288

291-
@[simp, grind =] theorem pmap_attach {xs : Array α} {p : {x // x ∈ xs} → Prop} {f : ∀ a, p a → β} (H) :
289+
@[simp] theorem pmap_attach {xs : Array α} {p : {x // x ∈ xs} → Prop} {f : ∀ a, p a → β} (H) :
292290
pmap f xs.attach H =
293291
xs.pmap (P := fun a => ∃ h : a ∈ xs, p ⟨a, h⟩)
294292
(fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨h, H ⟨a, h⟩ (by simp)⟩) := by
295293
ext <;> simp
296294

297-
@[simp, grind =] theorem pmap_attachWith {xs : Array α} {p : {x // q x} → Prop} {f : ∀ a, p a → β} (H₁ H₂) :
295+
@[simp] theorem pmap_attachWith {xs : Array α} {p : {x // q x} → Prop} {f : ∀ a, p a → β} (H₁ H₂) :
298296
pmap f (xs.attachWith q H₁) H₂ =
299297
xs.pmap (P := fun a => ∃ h : q a, p ⟨a, h⟩)
300298
(fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨H₁ _ h, H₂ ⟨a, H₁ _ h⟩ (by simpa)⟩) := by
301299
ext <;> simp
302300

303-
@[grind =]
304301
theorem foldl_pmap {xs : Array α} {P : α → Prop} {f : (a : α) → P a → β}
305302
(H : ∀ (a : α), a ∈ xs → P a) (g : γ → β → γ) (x : γ) :
306303
(xs.pmap f H).foldl g x = xs.attach.foldl (fun acc a => g acc (f a.1 (H _ a.2))) x := by
307304
rw [pmap_eq_map_attach, foldl_map]
308305

309-
@[grind =]
310306
theorem foldr_pmap {xs : Array α} {P : α → Prop} {f : (a : α) → P a → β}
311307
(H : ∀ (a : α), a ∈ xs → P a) (g : β → γ → γ) (x : γ) :
312308
(xs.pmap f H).foldr g x = xs.attach.foldr (fun a acc => g (f a.1 (H _ a.2)) acc) x := by
@@ -364,20 +360,18 @@ theorem foldr_attach {xs : Array α} {f : α → β → β} {b : β} :
364360
ext
365361
simpa using fun a => List.mem_of_getElem? a
366362

367-
@[grind =]
368363
theorem attach_map {xs : Array α} {f : α → β} :
369364
(xs.map f).attach = xs.attach.map (fun ⟨x, h⟩ => ⟨f x, mem_map_of_mem h⟩) := by
370365
cases xs
371366
ext <;> simp
372367

373-
@[grind =]
374368
theorem attachWith_map {xs : Array α} {f : α → β} {P : β → Prop} (H : ∀ (b : β), b ∈ xs.map f → P b) :
375369
(xs.map f).attachWith P H = (xs.attachWith (P ∘ f) (fun _ h => H _ (mem_map_of_mem h))).map
376370
fun ⟨x, h⟩ => ⟨f x, h⟩ := by
377371
cases xs
378372
simp [List.attachWith_map]
379373

380-
@[simp, grind =] theorem map_attachWith {xs : Array α} {P : α → Prop} {H : ∀ (a : α), a ∈ xs → P a}
374+
@[simp] theorem map_attachWith {xs : Array α} {P : α → Prop} {H : ∀ (a : α), a ∈ xs → P a}
381375
{f : { x // P x } → β} :
382376
(xs.attachWith P H).map f = xs.attach.map fun ⟨x, h⟩ => f ⟨x, H _ h⟩ := by
383377
cases xs <;> simp_all
@@ -430,15 +424,14 @@ theorem filter_attachWith {q : α → Prop} {xs : Array α} {p : {x // q x} →
430424
cases xs
431425
simp [Function.comp_def, List.filter_map]
432426

433-
@[grind =]
434427
theorem pmap_pmap {p : α → Prop} {q : β → Prop} {g : ∀ a, p a → β} {f : ∀ b, q b → γ} {xs} (H₁ H₂) :
435428
pmap f (pmap g xs H₁) H₂ =
436429
pmap (α := { x // x ∈ xs }) (fun a h => f (g a h) (H₂ (g a h) (mem_pmap_of_mem a.2))) xs.attach
437430
(fun a _ => H₁ a a.2) := by
438431
cases xs
439432
simp [List.pmap_pmap, List.pmap_map]
440433

441-
@[simp, grind =] theorem pmap_append {p : ι → Prop} {f : ∀ a : ι, p a → α} {xs ys : Array ι}
434+
@[simp] theorem pmap_append {p : ι → Prop} {f : ∀ a : ι, p a → α} {xs ys : Array ι}
442435
(h : ∀ a ∈ xs ++ ys, p a) :
443436
(xs ++ ys).pmap f h =
444437
(xs.pmap f fun a ha => h a (mem_append_left ys ha)) ++
@@ -453,70 +446,67 @@ theorem pmap_append' {p : α → Prop} {f : ∀ a : α, p a → β} {xs ys : Arr
453446
xs.pmap f h₁ ++ ys.pmap f h₂ :=
454447
pmap_append _
455448

456-
@[simp, grind =] theorem attach_append {xs ys : Array α} :
449+
@[simp] theorem attach_append {xs ys : Array α} :
457450
(xs ++ ys).attach = xs.attach.map (fun ⟨x, h⟩ => ⟨x, mem_append_left ys h⟩) ++
458451
ys.attach.map fun ⟨x, h⟩ => ⟨x, mem_append_right xs h⟩ := by
459452
cases xs
460453
cases ys
461454
rw [attach_congr (List.append_toArray _ _)]
462455
simp [List.attach_append, Function.comp_def]
463456

464-
@[simp, grind =] theorem attachWith_append {P : α → Prop} {xs ys : Array α}
457+
@[simp] theorem attachWith_append {P : α → Prop} {xs ys : Array α}
465458
{H : ∀ (a : α), a ∈ xs ++ ys → P a} :
466459
(xs ++ ys).attachWith P H = xs.attachWith P (fun a h => H a (mem_append_left ys h)) ++
467460
ys.attachWith P (fun a h => H a (mem_append_right xs h)) := by
468461
simp [attachWith]
469462

470-
@[simp, grind =] theorem pmap_reverse {P : α → Prop} {f : (a : α) → P a → β} {xs : Array α}
463+
@[simp] theorem pmap_reverse {P : α → Prop} {f : (a : α) → P a → β} {xs : Array α}
471464
(H : ∀ (a : α), a ∈ xs.reverse → P a) :
472465
xs.reverse.pmap f H = (xs.pmap f (fun a h => H a (by simpa using h))).reverse := by
473466
induction xs <;> simp_all
474467

475-
@[grind =]
476468
theorem reverse_pmap {P : α → Prop} {f : (a : α) → P a → β} {xs : Array α}
477469
(H : ∀ (a : α), a ∈ xs → P a) :
478470
(xs.pmap f H).reverse = xs.reverse.pmap f (fun a h => H a (by simpa using h)) := by
479471
rw [pmap_reverse]
480472

481-
@[simp, grind =] theorem attachWith_reverse {P : α → Prop} {xs : Array α}
473+
@[simp] theorem attachWith_reverse {P : α → Prop} {xs : Array α}
482474
{H : ∀ (a : α), a ∈ xs.reverse → P a} :
483475
xs.reverse.attachWith P H =
484476
(xs.attachWith P (fun a h => H a (by simpa using h))).reverse := by
485477
cases xs
486478
simp
487479

488-
@[grind =]
489480
theorem reverse_attachWith {P : α → Prop} {xs : Array α}
490481
{H : ∀ (a : α), a ∈ xs → P a} :
491482
(xs.attachWith P H).reverse = (xs.reverse.attachWith P (fun a h => H a (by simpa using h))) := by
492483
cases xs
493484
simp
494485

495-
@[simp, grind =] theorem attach_reverse {xs : Array α} :
486+
@[simp] theorem attach_reverse {xs : Array α} :
496487
xs.reverse.attach = xs.attach.reverse.map fun ⟨x, h⟩ => ⟨x, by simpa using h⟩ := by
497488
cases xs
498489
rw [attach_congr List.reverse_toArray]
499490
simp
500491

501-
@[grind =]
502492
theorem reverse_attach {xs : Array α} :
503493
xs.attach.reverse = xs.reverse.attach.map fun ⟨x, h⟩ => ⟨x, by simpa using h⟩ := by
504494
cases xs
505495
simp
506496

507-
@[simp, grind =] theorem back?_pmap {P : α → Prop} {f : (a : α) → P a → β} {xs : Array α}
497+
@[simp] theorem back?_pmap {P : α → Prop} {f : (a : α) → P a → β} {xs : Array α}
508498
(H : ∀ (a : α), a ∈ xs → P a) :
509499
(xs.pmap f H).back? = xs.attach.back?.map fun ⟨a, m⟩ => f a (H a m) := by
510500
cases xs
511501
simp
512502

513-
@[simp, grind =] theorem back?_attachWith {P : α → Prop} {xs : Array α}
503+
@[simp] theorem back?_attachWith {P : α → Prop} {xs : Array α}
514504
{H : ∀ (a : α), a ∈ xs → P a} :
515505
(xs.attachWith P H).back? = xs.back?.pbind (fun a h => some ⟨a, H _ (mem_of_back? h)⟩) := by
516506
cases xs
517507
simp
518508

519-
@[simp, grind =]
509+
@[simp]
520510
theorem back?_attach {xs : Array α} :
521511
xs.attach.back? = xs.back?.pbind fun a h => some ⟨a, mem_of_back? h⟩ := by
522512
cases xs

src/Init/Data/List/Attach.lean

Lines changed: 14 additions & 26 deletions
Original file line numberDiff line numberDiff line change
@@ -95,14 +95,12 @@ theorem pmap_congr_left {p q : α → Prop} {f : ∀ a, p a → β} {g : ∀ a,
9595
| cons x l ih =>
9696
rw [pmap, pmap, h _ mem_cons_self, ih fun a ha => h a (mem_cons_of_mem _ ha)]
9797

98-
@[grind =]
9998
theorem map_pmap {p : α → Prop} {g : β → γ} {f : ∀ a, p a → β} {l : List α} (H) :
10099
map g (pmap f l H) = pmap (fun a h => g (f a h)) l H := by
101100
induction l
102101
· rfl
103102
· simp only [*, pmap, map]
104103

105-
@[grind =]
106104
theorem pmap_map {p : β → Prop} {g : ∀ b, p b → γ} {f : α → β} {l : List α} (H) :
107105
pmap g (map f l) H = pmap (fun a h => g (f a) h) l fun _ h => H _ (mem_map_of_mem h) := by
108106
induction l
@@ -285,13 +283,13 @@ theorem getElem_attach {xs : List α} {i : Nat} (h : i < xs.attach.length) :
285283
xs.attach[i] = ⟨xs[i]'(by simpa using h), getElem_mem (by simpa using h)⟩ :=
286284
getElem_attachWith h
287285

288-
@[simp, grind =] theorem pmap_attach {l : List α} {p : {x // x ∈ l} → Prop} {f : ∀ a, p a → β} (H) :
286+
@[simp] theorem pmap_attach {l : List α} {p : {x // x ∈ l} → Prop} {f : ∀ a, p a → β} (H) :
289287
pmap f l.attach H =
290288
l.pmap (P := fun a => ∃ h : a ∈ l, p ⟨a, h⟩)
291289
(fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨h, H ⟨a, h⟩ (by simp)⟩) := by
292290
apply ext_getElem <;> simp
293291

294-
@[simp, grind =] theorem pmap_attachWith {l : List α} {p : {x // q x} → Prop} {f : ∀ a, p a → β} (H₁ H₂) :
292+
@[simp] theorem pmap_attachWith {l : List α} {p : {x // q x} → Prop} {f : ∀ a, p a → β} (H₁ H₂) :
295293
pmap f (l.attachWith q H₁) H₂ =
296294
l.pmap (P := fun a => ∃ h : q a, p ⟨a, h⟩)
297295
(fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨H₁ _ h, H₂ ⟨a, H₁ _ h⟩ (by simpa)⟩) := by
@@ -349,26 +347,24 @@ theorem getElem_attach {xs : List α} {i : Nat} (h : i < xs.attach.length) :
349347
xs.attach.tail = xs.tail.attach.map (fun ⟨x, h⟩ => ⟨x, mem_of_mem_tail h⟩) := by
350348
cases xs <;> simp
351349

352-
@[grind =]
353350
theorem foldl_pmap {l : List α} {P : α → Prop} {f : (a : α) → P a → β}
354351
(H : ∀ (a : α), a ∈ l → P a) (g : γ → β → γ) (x : γ) :
355352
(l.pmap f H).foldl g x = l.attach.foldl (fun acc a => g acc (f a.1 (H _ a.2))) x := by
356353
rw [pmap_eq_map_attach, foldl_map]
357354

358-
@[grind =]
359355
theorem foldr_pmap {l : List α} {P : α → Prop} {f : (a : α) → P a → β}
360356
(H : ∀ (a : α), a ∈ l → P a) (g : β → γ → γ) (x : γ) :
361357
(l.pmap f H).foldr g x = l.attach.foldr (fun a acc => g (f a.1 (H _ a.2)) acc) x := by
362358
rw [pmap_eq_map_attach, foldr_map]
363359

364-
@[simp, grind =] theorem foldl_attachWith
360+
@[simp] theorem foldl_attachWith
365361
{l : List α} {q : α → Prop} (H : ∀ a, a ∈ l → q a) {f : β → { x // q x } → β} {b} :
366362
(l.attachWith q H).foldl f b = l.attach.foldl (fun b ⟨a, h⟩ => f b ⟨a, H _ h⟩) b := by
367363
induction l generalizing b with
368364
| nil => simp
369365
| cons a l ih => simp [ih, foldl_map]
370366

371-
@[simp, grind =] theorem foldr_attachWith
367+
@[simp] theorem foldr_attachWith
372368
{l : List α} {q : α → Prop} (H : ∀ a, a ∈ l → q a) {f : { x // q x } → β → β} {b} :
373369
(l.attachWith q H).foldr f b = l.attach.foldr (fun a acc => f ⟨a.1, H _ a.2⟩ acc) b := by
374370
induction l generalizing b with
@@ -407,18 +403,16 @@ theorem foldr_attach {l : List α} {f : α → β → β} {b : β} :
407403
| nil => simp
408404
| cons a l ih => rw [foldr_cons, attach_cons, foldr_cons, foldr_map, ih]
409405

410-
@[grind =]
411406
theorem attach_map {l : List α} {f : α → β} :
412407
(l.map f).attach = l.attach.map (fun ⟨x, h⟩ => ⟨f x, mem_map_of_mem h⟩) := by
413408
induction l <;> simp [*]
414409

415-
@[grind =]
416410
theorem attachWith_map {l : List α} {f : α → β} {P : β → Prop} (H : ∀ (b : β), b ∈ l.map f → P b) :
417411
(l.map f).attachWith P H = (l.attachWith (P ∘ f) (fun _ h => H _ (mem_map_of_mem h))).map
418412
fun ⟨x, h⟩ => ⟨f x, h⟩ := by
419413
induction l <;> simp [*]
420414

421-
@[simp, grind =] theorem map_attachWith {l : List α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a}
415+
@[simp] theorem map_attachWith {l : List α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a}
422416
{f : { x // P x } → β} :
423417
(l.attachWith P H).map f = l.attach.map fun ⟨x, h⟩ => f ⟨x, H _ h⟩ := by
424418
induction l <;> simp_all
@@ -444,7 +438,6 @@ theorem map_attach_eq_pmap {l : List α} {f : { x // x ∈ l } → β} :
444438
apply pmap_congr_left
445439
simp
446440

447-
@[grind =]
448441
theorem attach_filterMap {l : List α} {f : α → Option β} :
449442
(l.filterMap f).attach = l.attach.filterMap
450443
fun ⟨x, h⟩ => (f x).pbind (fun b m => some ⟨b, mem_filterMap.mpr ⟨x, h, m⟩⟩) := by
@@ -475,7 +468,6 @@ theorem attach_filterMap {l : List α} {f : α → Option β} :
475468
ext
476469
simp
477470

478-
@[grind =]
479471
theorem attach_filter {l : List α} (p : α → Bool) :
480472
(l.filter p).attach = l.attach.filterMap
481473
fun x => if w : p x.1 then some ⟨x.1, mem_filter.mpr ⟨x.2, w⟩⟩ else none := by
@@ -487,7 +479,7 @@ theorem attach_filter {l : List α} (p : α → Bool) :
487479

488480
-- We are still missing here `attachWith_filterMap` and `attachWith_filter`.
489481

490-
@[simp, grind =]
482+
@[simp]
491483
theorem filterMap_attachWith {q : α → Prop} {l : List α} {f : {x // q x} → Option β} (H) :
492484
(l.attachWith q H).filterMap f = l.attach.filterMap (fun ⟨x, h⟩ => f ⟨x, H _ h⟩) := by
493485
induction l with
@@ -496,7 +488,7 @@ theorem filterMap_attachWith {q : α → Prop} {l : List α} {f : {x // q x} →
496488
simp only [attachWith_cons, filterMap_cons]
497489
split <;> simp_all [Function.comp_def]
498490

499-
@[simp, grind =]
491+
@[simp]
500492
theorem filter_attachWith {q : α → Prop} {l : List α} {p : {x // q x} → Bool} (H) :
501493
(l.attachWith q H).filter p =
502494
(l.attach.filter (fun ⟨x, h⟩ => p ⟨x, H _ h⟩)).map (fun ⟨x, h⟩ => ⟨x, H _ h⟩) := by
@@ -506,14 +498,13 @@ theorem filter_attachWith {q : α → Prop} {l : List α} {p : {x // q x} → Bo
506498
simp only [attachWith_cons, filter_cons]
507499
split <;> simp_all [Function.comp_def, filter_map]
508500

509-
@[grind =]
510501
theorem pmap_pmap {p : α → Prop} {q : β → Prop} {g : ∀ a, p a → β} {f : ∀ b, q b → γ} {l} (H₁ H₂) :
511502
pmap f (pmap g l H₁) H₂ =
512503
pmap (α := { x // x ∈ l }) (fun a h => f (g a h) (H₂ (g a h) (mem_pmap_of_mem a.2))) l.attach
513504
(fun a _ => H₁ a a.2) := by
514505
simp [pmap_eq_map_attach, attach_map]
515506

516-
@[simp, grind =] theorem pmap_append {p : ι → Prop} {f : ∀ a : ι, p a → α} {l₁ l₂ : List ι}
507+
@[simp] theorem pmap_append {p : ι → Prop} {f : ∀ a : ι, p a → α} {l₁ l₂ : List ι}
517508
(h : ∀ a ∈ l₁ ++ l₂, p a) :
518509
(l₁ ++ l₂).pmap f h =
519510
(l₁.pmap f fun a ha => h a (mem_append_left l₂ ha)) ++
@@ -530,50 +521,47 @@ theorem pmap_append' {p : α → Prop} {f : ∀ a : α, p a → β} {l₁ l₂ :
530521
l₁.pmap f h₁ ++ l₂.pmap f h₂ :=
531522
pmap_append _
532523

533-
@[simp, grind =] theorem attach_append {xs ys : List α} :
524+
@[simp] theorem attach_append {xs ys : List α} :
534525
(xs ++ ys).attach = xs.attach.map (fun ⟨x, h⟩ => ⟨x, mem_append_left ys h⟩) ++
535526
ys.attach.map fun ⟨x, h⟩ => ⟨x, mem_append_right xs h⟩ := by
536527
simp only [attach, attachWith, map_pmap, pmap_append]
537528
congr 1 <;>
538529
exact pmap_congr_left _ fun _ _ _ _ => rfl
539530

540-
@[simp, grind =] theorem attachWith_append {P : α → Prop} {xs ys : List α}
531+
@[simp] theorem attachWith_append {P : α → Prop} {xs ys : List α}
541532
{H : ∀ (a : α), a ∈ xs ++ ys → P a} :
542533
(xs ++ ys).attachWith P H = xs.attachWith P (fun a h => H a (mem_append_left ys h)) ++
543534
ys.attachWith P (fun a h => H a (mem_append_right xs h)) := by
544535
simp only [attachWith, pmap_append]
545536

546-
@[simp, grind =] theorem pmap_reverse {P : α → Prop} {f : (a : α) → P a → β} {xs : List α}
537+
@[simp] theorem pmap_reverse {P : α → Prop} {f : (a : α) → P a → β} {xs : List α}
547538
(H : ∀ (a : α), a ∈ xs.reverse → P a) :
548539
xs.reverse.pmap f H = (xs.pmap f (fun a h => H a (by simpa using h))).reverse := by
549540
induction xs <;> simp_all
550541

551-
@[grind =]
552542
theorem reverse_pmap {P : α → Prop} {f : (a : α) → P a → β} {xs : List α}
553543
(H : ∀ (a : α), a ∈ xs → P a) :
554544
(xs.pmap f H).reverse = xs.reverse.pmap f (fun a h => H a (by simpa using h)) := by
555545
rw [pmap_reverse]
556546

557-
@[simp, grind =] theorem attachWith_reverse {P : α → Prop} {xs : List α}
547+
@[simp] theorem attachWith_reverse {P : α → Prop} {xs : List α}
558548
{H : ∀ (a : α), a ∈ xs.reverse → P a} :
559549
xs.reverse.attachWith P H =
560550
(xs.attachWith P (fun a h => H a (by simpa using h))).reverse :=
561551
pmap_reverse ..
562552

563-
@[grind =]
564553
theorem reverse_attachWith {P : α → Prop} {xs : List α}
565554
{H : ∀ (a : α), a ∈ xs → P a} :
566555
(xs.attachWith P H).reverse = (xs.reverse.attachWith P (fun a h => H a (by simpa using h))) :=
567556
reverse_pmap ..
568557

569-
@[simp, grind =] theorem attach_reverse {xs : List α} :
558+
@[simp] theorem attach_reverse {xs : List α} :
570559
xs.reverse.attach = xs.attach.reverse.map fun ⟨x, h⟩ => ⟨x, by simpa using h⟩ := by
571560
simp only [attach, attachWith, reverse_pmap, map_pmap]
572561
apply pmap_congr_left
573562
intros
574563
rfl
575564

576-
@[grind =]
577565
theorem reverse_attach {xs : List α} :
578566
xs.attach.reverse = xs.reverse.attach.map fun ⟨x, h⟩ => ⟨x, by simpa using h⟩ := by
579567
simp only [attach, attachWith, reverse_pmap, map_pmap]
@@ -627,7 +615,7 @@ theorem countP_attachWith {p : α → Prop} {q : α → Bool} {l : List α} (H :
627615
(l.attachWith p H).countP (fun a : {x // p x} => q a) = l.countP q := by
628616
simp only [← Function.comp_apply (g := Subtype.val), ← countP_map, attachWith_map_subtype_val]
629617

630-
@[simp]
618+
@[simp, grind =]
631619
theorem count_attach [BEq α] {l : List α} {a : {x // x ∈ l}} :
632620
l.attach.count a = l.count ↑a :=
633621
Eq.trans (countP_congr fun _ _ => by simp) <| countP_attach

0 commit comments

Comments
 (0)