Skip to content

Commit cdb994b

Browse files
authored
chore: remove @[grind =] from List.countP_eq_length_filter (#11542)
This PR removes `@[grind =]` from `List.countP_eq_length_filter` and `Array.countP_eq_size_filter`, as users [reported](https://leanprover.zulipchat.com/#narrow/channel/270676-lean4/topic/.60countP_eq_length_filter.60.20grind.20attribute/near/561386848)[#lean4 > `countP_eq_length_filter` grind attribute @ 💬](https://leanprover.zulipchat.com/#narrow/channel/270676-lean4/topic/.60countP_eq_length_filter.60.20grind.20attribute/near/561386848) this was problematic.
1 parent b7ac624 commit cdb994b

File tree

4 files changed

+17
-17
lines changed

4 files changed

+17
-17
lines changed

src/Init/Data/Array/Count.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -62,12 +62,12 @@ theorem size_eq_countP_add_countP {xs : Array α} : xs.size = countP p xs + coun
6262
rcases xs with ⟨xs⟩
6363
simp [List.length_eq_countP_add_countP (p := p)]
6464

65-
@[grind =]
6665
theorem countP_eq_size_filter {xs : Array α} : countP p xs = (filter p xs).size := by
6766
rcases xs with ⟨xs⟩
6867
simp [List.countP_eq_length_filter]
6968

70-
@[grind =]
69+
grind_pattern countP_eq_size_filter => xs.countP p, xs.filter p
70+
7171
theorem countP_eq_size_filter' : countP p = size ∘ filter p := by
7272
funext xs
7373
apply countP_eq_size_filter

src/Init/Data/List/Count.lean

Lines changed: 4 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -29,7 +29,7 @@ section countP
2929

3030
variable {p q : α → Bool}
3131

32-
@[simp] theorem countP_nil : countP p [] = 0 := rfl
32+
@[simp, grind =] theorem countP_nil : countP p [] = 0 := rfl
3333

3434
protected theorem countP_go_eq_add {l} : countP.go p l n = n + countP.go p l 0 := by
3535
induction l generalizing n with
@@ -47,6 +47,7 @@ protected theorem countP_go_eq_add {l} : countP.go p l n = n + countP.go p l 0 :
4747
@[simp] theorem countP_cons_of_neg {l} (pa : ¬p a) : countP p (a :: l) = countP p l := by
4848
simp [countP, countP.go, pa]
4949

50+
@[grind =]
5051
theorem countP_cons {a : α} {l : List α} : countP p (a :: l) = countP p l + if p a then 1 else 0 := by
5152
by_cases h : p a <;> simp [h]
5253

@@ -66,7 +67,6 @@ theorem length_eq_countP_add_countP (p : α → Bool) {l : List α} : length l =
6667
· rfl
6768
· simp [h]
6869

69-
@[grind =] -- This to quite aggressive, as it introduces `filter` based reasoning whenever we see `countP`.
7070
theorem countP_eq_length_filter {l : List α} : countP p l = (filter p l).length := by
7171
induction l with
7272
| nil => rfl
@@ -75,7 +75,8 @@ theorem countP_eq_length_filter {l : List α} : countP p l = (filter p l).length
7575
then rw [countP_cons_of_pos h, ih, filter_cons_of_pos h, length]
7676
else rw [countP_cons_of_neg h, ih, filter_cons_of_neg h]
7777

78-
@[grind =]
78+
grind_pattern countP_eq_length_filter => l.countP p, l.filter p
79+
7980
theorem countP_eq_length_filter' : countP p = length ∘ filter p := by
8081
funext l
8182
apply countP_eq_length_filter

tests/lean/run/grind_11081.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -75,5 +75,6 @@ theorem countP_diff (hl : Subperm l₂ l₁) (p : α → Bool) :
7575
List.Perm.subperm_left,
7676
List.Sublist.subperm,
7777
List.Subperm.perm_of_length_le,
78-
List.Perm.countP_congr
78+
List.Perm.countP_congr,
79+
List.countP_eq_length_filter
7980
]

tests/lean/run/grind_list_count.lean

Lines changed: 9 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -29,7 +29,7 @@ theorem countP_eq_length_filter' : countP p = length ∘ filter p := by
2929
grind
3030

3131
theorem countP_le_length : countP p l ≤ l.length := by
32-
grind
32+
induction l with grind
3333

3434
theorem countP_append {l₁ l₂ : List α} : countP p (l₁ ++ l₂) = countP p l₁ + countP p l₂ := by
3535
grind
@@ -44,11 +44,11 @@ theorem countP_eq_zero {p} : countP p l = 0 ↔ ∀ a ∈ l, ¬p a := by
4444
induction l with grind
4545

4646
theorem countP_eq_length {p} : countP p l = l.length ↔ ∀ a ∈ l, p a := by
47-
induction l with grind
47+
induction l with grind [countP_eq_length_filter]
4848

4949
theorem countP_replicate {p : α → Bool} {a : α} {n : Nat} :
5050
countP p (replicate n a) = if p a then n else 0 := by
51-
grind
51+
induction n with grind
5252

5353
theorem boole_getElem_le_countP {p : α → Bool} {l : List α} {i : Nat} (h : i < l.length) :
5454
(if p l[i] then 1 else 0) ≤ l.countP p := by
@@ -66,11 +66,9 @@ theorem countP_tail_le (l) : countP p l.tail ≤ countP p l := by grind
6666

6767
-- See `Init.Data.List.Nat.Count` for `le_countP_tail : countP p l - 1 ≤ countP p l.tail`.
6868

69-
-- TODO Should we have `@[grind]` for `filter_filter`?
70-
7169
theorem countP_filter {l : List α} :
7270
countP p (filter q l) = countP (fun a => p a && q a) l := by
73-
grind [List.filter_filter]
71+
induction l with grind
7472

7573
theorem countP_true : (countP fun (_ : α) => true) = length := by
7674
funext l
@@ -81,23 +79,23 @@ theorem countP_false : (countP fun (_ : α) => false) = Function.const _ 0 := by
8179
induction l with grind
8280

8381
theorem countP_map {p : β → Bool} {f : α → β} {l} : countP p (map f l) = countP (p ∘ f) l := by
84-
grind
82+
induction l with grind
8583

8684
theorem length_filterMap_eq_countP {f : α → Option β} {l : List α} :
8785
(filterMap f l).length = countP (fun a => (f a).isSome) l := by
88-
induction l with grind -- TODO
86+
induction l with grind
8987

9088
theorem countP_filterMap {p : β → Bool} {f : α → Option β} {l : List α} :
9189
countP p (filterMap f l) = countP (fun a => ((f a).map p).getD false) l := by
92-
induction l with grind -- TODO
90+
induction l with grind
9391

9492
theorem countP_flatten {l : List (List α)} :
9593
countP p l.flatten = (l.map (countP p)).sum := by
96-
grind
94+
induction l with grind
9795

9896
theorem countP_flatMap {p : β → Bool} {l : List α} {f : α → List β} :
9997
countP p (l.flatMap f) = sum (map (countP p ∘ f) l) := by
100-
grind
98+
induction l with grind
10199

102100
theorem countP_reverse {l : List α} : countP p l.reverse = countP p l := by
103101
grind

0 commit comments

Comments
 (0)