Skip to content

Commit 365203e

Browse files
committed
feat: grind golf in Mathlib.Data.List (#29492)
Much of this comes from locations identified by the new `linter.tacticAnalysis.terminalToGrind`, and then cleaning up in flagged files.
1 parent 21fd553 commit 365203e

File tree

10 files changed

+74
-212
lines changed

10 files changed

+74
-212
lines changed

Mathlib/Data/List/Basic.lean

Lines changed: 18 additions & 72 deletions
Original file line numberDiff line numberDiff line change
@@ -157,15 +157,10 @@ theorem exists_mem_cons_of_exists {p : α → Prop} {a : α} {l : List α} : (
157157
fun ⟨x, xl, px⟩ => ⟨x, mem_cons_of_mem _ xl, px⟩
158158

159159
theorem or_exists_of_exists_mem_cons {p : α → Prop} {a : α} {l : List α} : (∃ x ∈ a :: l, p x) →
160-
p a ∨ ∃ x ∈ l, p x :=
161-
fun ⟨x, xal, px⟩ =>
162-
Or.elim (eq_or_mem_of_mem_cons xal) (fun h : x = a => by rw [← h]; left; exact px)
163-
fun h : x ∈ l => Or.inr ⟨x, h, px⟩
160+
p a ∨ ∃ x ∈ l, p x := by grind
164161

165162
theorem exists_mem_cons_iff (p : α → Prop) (a : α) (l : List α) :
166-
(∃ x ∈ a :: l, p x) ↔ p a ∨ ∃ x ∈ l, p x :=
167-
Iff.intro or_exists_of_exists_mem_cons fun h =>
168-
Or.elim h (exists_mem_cons_of l) exists_mem_cons_of_exists
163+
(∃ x ∈ a :: l, p x) ↔ p a ∨ ∃ x ∈ l, p x := by grind
169164

170165
/-! ### list subset -/
171166

@@ -303,9 +298,7 @@ theorem getLast_append_singleton {a : α} (l : List α) :
303298

304299
theorem getLast_append_of_right_ne_nil (l₁ l₂ : List α) (h : l₂ ≠ []) :
305300
getLast (l₁ ++ l₂) (append_ne_nil_of_right_ne_nil l₁ h) = getLast l₂ h := by
306-
induction l₁ with
307-
| nil => simp
308-
| cons _ _ ih => simp only [cons_append]; rw [List.getLast_cons]; exact ih
301+
induction l₁ with grind
309302

310303
theorem getLast_concat' {a : α} (l : List α) : getLast (concat l a) (by simp) = a := by
311304
simp
@@ -332,15 +325,9 @@ theorem getLast_replicate_succ (m : ℕ) (a : α) :
332325
/-! ### getLast? -/
333326

334327
theorem mem_getLast?_eq_getLast : ∀ {l : List α} {x : α}, x ∈ l.getLast? → ∃ h, x = getLast l h
335-
| [], x, hx => False.elim <| by simp at hx
336-
| [a], x, hx =>
337-
have : a = x := by simpa using hx
338-
this ▸ ⟨cons_ne_nil a [], rfl⟩
339-
| a :: b :: l, x, hx => by
340-
rw [getLast?_cons_cons] at hx
341-
rcases mem_getLast?_eq_getLast hx with ⟨_, h₂⟩
342-
use cons_ne_nil _ _
343-
assumption
328+
| [], x, hx
329+
| [a], x, hx
330+
| a :: b :: l, x, hx => by grind
344331

345332
theorem getLast?_eq_getLast_of_ne_nil : ∀ {l : List α} (h : l ≠ []), l.getLast? = some (l.getLast h)
346333
| [], h => (h rfl).elim
@@ -378,11 +365,7 @@ theorem getLast?_append_of_ne_nil (l₁ : List α) :
378365
| b :: l₂, _ => getLast?_append_cons l₁ b l₂
379366

380367
theorem mem_getLast?_append_of_mem_getLast? {l₁ l₂ : List α} {x : α} (h : x ∈ l₂.getLast?) :
381-
x ∈ (l₁ ++ l₂).getLast? := by
382-
cases l₂
383-
· contradiction
384-
· rw [List.getLast?_append_cons]
385-
exact h
368+
x ∈ (l₁ ++ l₂).getLast? := by grind
386369

387370
/-! ### head(!?) and tail -/
388371

@@ -433,10 +416,7 @@ theorem head?_append_of_ne_nil :
433416
| _ :: _, _, _ => rfl
434417

435418
theorem tail_append_singleton_of_ne_nil {a : α} {l : List α} (h : l ≠ nil) :
436-
tail (l ++ [a]) = tail l ++ [a] := by
437-
induction l
438-
· contradiction
439-
· rw [tail, cons_append, tail]
419+
tail (l ++ [a]) = tail l ++ [a] := by grind
440420

441421
theorem cons_head?_tail : ∀ {l : List α} {a : α}, a ∈ head? l → a :: tail l = l
442422
| [], a, h => by contradiction
@@ -518,12 +498,10 @@ theorem idxOf_of_notMem {l : List α} {a : α} : a ∉ l → idxOf a l = length
518498

519499
@[deprecated (since := "2025-05-23")] alias idxOf_of_not_mem := idxOf_of_notMem
520500

521-
theorem idxOf_append_of_mem {a : α} (h : a ∈ l₁) : idxOf a (l₁ ++ l₂) = idxOf a l₁ := by
522-
grind
501+
theorem idxOf_append_of_mem {a : α} (h : a ∈ l₁) : idxOf a (l₁ ++ l₂) = idxOf a l₁ := by grind
523502

524503
theorem idxOf_append_of_notMem {a : α} (h : a ∉ l₁) :
525-
idxOf a (l₁ ++ l₂) = l₁.length + idxOf a l₂ := by
526-
grind
504+
idxOf a (l₁ ++ l₂) = l₁.length + idxOf a l₂ := by grind
527505

528506
@[deprecated (since := "2025-05-23")] alias idxOf_append_of_not_mem := idxOf_append_of_notMem
529507

@@ -551,10 +529,7 @@ theorem take_one_drop_eq_of_lt_length {l : List α} {n : ℕ} (h : n < l.length)
551529
theorem ext_getElem?' {l₁ l₂ : List α} (h' : ∀ n < max l₁.length l₂.length, l₁[n]? = l₂[n]?) :
552530
l₁ = l₂ := by
553531
apply ext_getElem?
554-
intro n
555-
rcases Nat.lt_or_ge n <| max l₁.length l₂.length with hn | hn
556-
· exact h' n hn
557-
· simp_all [Nat.max_le]
532+
grind
558533

559534
theorem ext_get_iff {l₁ l₂ : List α} :
560535
l₁ = l₂ ↔ l₁.length = l₂.length ∧ ∀ n h₁ h₂, get l₁ ⟨n, h₁⟩ = get l₂ ⟨n, h₂⟩ := by
@@ -604,11 +579,8 @@ theorem get_reverse' (l : List α) (n) (hn') :
604579
l.reverse.get n = l.get ⟨l.length - 1 - n, hn'⟩ := by
605580
simp
606581

607-
theorem eq_cons_of_length_one {l : List α} (h : l.length = 1) : l = [l.get ⟨0, by cutsat⟩] := by
608-
refine ext_get (by convert h) fun n h₁ h₂ => ?_
609-
simp
610-
congr
611-
cutsat
582+
theorem eq_cons_of_length_one {l : List α} (h : l.length = 1) : l = [l.get ⟨0, by omega⟩] := by
583+
refine ext_get (by convert h) (by grind)
612584

613585
end deprecated
614586

@@ -987,18 +959,7 @@ variable (p)
987959

988960
theorem monotone_filter_right (l : List α) ⦃p q : α → Bool⦄
989961
(h : ∀ a, p a → q a) : l.filter p <+ l.filter q := by
990-
induction l with
991-
| nil => rfl
992-
| cons hd tl IH =>
993-
by_cases hp : p hd
994-
· rw [filter_cons_of_pos hp, filter_cons_of_pos (h _ hp)]
995-
exact IH.cons_cons hd
996-
· rw [filter_cons_of_neg hp]
997-
by_cases hq : q hd
998-
· rw [filter_cons_of_pos hq]
999-
exact sublist_cons_of_sublist hd IH
1000-
· rw [filter_cons_of_neg hq]
1001-
exact IH
962+
induction l with grind
1002963

1003964
lemma map_filter {f : α → β} (hf : Injective f) (l : List α)
1004965
[DecidablePred fun b => ∃ a, p a ∧ f a = b] :
@@ -1040,8 +1001,7 @@ variable {p : α → Bool}
10401001

10411002
-- Cannot be @[simp] because `a` cannot be inferred by `simp`.
10421003
theorem length_eraseP_add_one {l : List α} {a} (al : a ∈ l) (pa : p a) :
1043-
(l.eraseP p).length + 1 = l.length := by
1044-
grind
1004+
(l.eraseP p).length + 1 = l.length := by grind
10451005

10461006
end eraseP
10471007

@@ -1069,19 +1029,10 @@ theorem erase_getElem [DecidableEq ι] {l : List ι} {i : ℕ} (hi : i < l.lengt
10691029
Perm (l.erase l[i]) (l.eraseIdx i) := by
10701030
induction l generalizing i with
10711031
| nil => simp
1072-
| cons a l IH =>
1073-
cases i with
1074-
| zero => simp
1075-
| succ i =>
1076-
have hi' : i < l.length := by simpa using hi
1077-
if ha : a = l[i] then
1078-
simpa [ha] using .trans (perm_cons_erase (getElem_mem _)) (.cons _ (IH hi'))
1079-
else
1080-
simpa [ha] using IH hi'
1032+
| cons a l IH => cases i with grind
10811033

10821034
theorem length_eraseIdx_add_one {l : List ι} {i : ℕ} (h : i < l.length) :
1083-
(l.eraseIdx i).length + 1 = l.length := by
1084-
grind
1035+
(l.eraseIdx i).length + 1 = l.length := by grind
10851036

10861037
end Erase
10871038

@@ -1210,12 +1161,7 @@ variable [BEq α] [LawfulBEq α]
12101161

12111162
lemma lookup_graph (f : α → β) {a : α} {as : List α} (h : a ∈ as) :
12121163
lookup a (as.map fun x => (x, f x)) = some (f a) := by
1213-
induction as with
1214-
| nil => exact (not_mem_nil h).elim
1215-
| cons a' as ih =>
1216-
by_cases ha : a = a'
1217-
· simp [ha]
1218-
· simpa [lookup_cons, beq_false_of_ne ha] using ih (List.mem_of_ne_of_mem ha h)
1164+
induction as with grind
12191165

12201166
end lookup
12211167

Mathlib/Data/List/Flatten.lean

Lines changed: 3 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -23,12 +23,7 @@ namespace List
2323
@[gcongr]
2424
protected theorem Sublist.flatten {l₁ l₂ : List (List α)} (h : l₁ <+ l₂) :
2525
l₁.flatten <+ l₂.flatten := by
26-
induction h with
27-
| slnil => simp
28-
| cons _ _ ih =>
29-
rw [flatten_cons]
30-
exact ih.trans (sublist_append_right _ _)
31-
| cons₂ _ _ ih => simpa
26+
induction h with grind
3227

3328
@[gcongr]
3429
protected theorem Sublist.flatMap {l₁ l₂ : List α} (h : l₁ <+ l₂) (f : α → List β) :
@@ -39,20 +34,12 @@ protected theorem Sublist.flatMap {l₁ l₂ : List α} (h : l₁ <+ l₂) (f :
3934
left with a list of length `1` made of the `i`-th element of the original list. -/
4035
theorem drop_take_succ_eq_cons_getElem (L : List α) (i : Nat) (h : i < L.length) :
4136
(L.take (i + 1)).drop i = [L[i]] := by
42-
induction L generalizing i with
43-
| nil => exact (Nat.not_succ_le_zero i h).elim
44-
| cons head tail ih =>
45-
rcases i with _ | i
46-
· simp
47-
· simpa using ih _ (by simpa using h)
37+
induction L generalizing i with grind
4838

4939
/-- We can rebracket `x ++ (l₁ ++ x) ++ (l₂ ++ x) ++ ... ++ (lₙ ++ x)` to
5040
`(x ++ l₁) ++ (x ++ l₂) ++ ... ++ (x ++ lₙ) ++ x` where `L = [l₁, l₂, ..., lₙ]`. -/
5141
theorem append_flatten_map_append (L : List (List α)) (x : List α) :
5242
x ++ (L.map (· ++ x)).flatten = (L.map (x ++ ·)).flatten ++ x := by
53-
induction L with
54-
| nil => rw [map_nil, flatten, append_nil, map_nil, flatten, nil_append]
55-
| cons _ _ ih =>
56-
rw [map_cons, flatten, map_cons, flatten, append_assoc, ih, append_assoc, append_assoc]
43+
induction L with grind
5744

5845
end List

Mathlib/Data/List/Induction.lean

Lines changed: 2 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -47,14 +47,7 @@ theorem reverseRecOn_concat {motive : List α → Sort*} (x : α) (xs : List α)
4747
exact this _ (reverse_reverse xs)
4848
intro ys hy
4949
conv_lhs => unfold reverseRecOn
50-
split
51-
next h => simp at h
52-
next heq =>
53-
revert heq
54-
simp only [reverse_append, reverse_cons, reverse_nil, nil_append, singleton_append, cons.injEq]
55-
rintro ⟨rfl, rfl⟩
56-
subst ys
57-
rfl
50+
grind
5851

5952
/-- Bidirectional induction principle for lists: if a property holds for the empty list, the
6053
singleton list, and `a :: (l ++ [b])` from `l`, then it holds for all lists. This can be used to
@@ -95,20 +88,7 @@ theorem bidirectionalRec_cons_append {motive : List α → Sort*}
9588
bidirectionalRec nil singleton cons_append (a :: (l ++ [b])) =
9689
cons_append a l b (bidirectionalRec nil singleton cons_append l) := by
9790
conv_lhs => unfold bidirectionalRec
98-
cases l with
99-
| nil => rfl
100-
| cons x xs =>
101-
simp only [List.cons_append]
102-
dsimp only [← List.cons_append]
103-
suffices ∀ (ys init : List α) (hinit : init = ys) (last : α) (hlast : last = b),
104-
(cons_append a init last
105-
(bidirectionalRec nil singleton cons_append init)) =
106-
cast (congr_arg motive <| by simp [hinit, hlast])
107-
(cons_append a ys b (bidirectionalRec nil singleton cons_append ys)) by
108-
rw [this (x :: xs) _ (by rw [dropLast_append_cons, dropLast_singleton, append_nil]) _ (by simp)]
109-
simp
110-
rintro ys init rfl last rfl
111-
rfl
91+
cases l with grind
11292

11393
/-- Like `bidirectionalRec`, but with the list parameter placed first. -/
11494
@[elab_as_elim]

Mathlib/Data/List/Lattice.lean

Lines changed: 25 additions & 52 deletions
Original file line numberDiff line numberDiff line change
@@ -95,7 +95,7 @@ end Union
9595

9696
section Inter
9797

98-
@[simp]
98+
@[simp, grind =]
9999
theorem inter_nil (l : List α) : [] ∩ l = [] :=
100100
rfl
101101

@@ -107,13 +107,16 @@ theorem inter_cons_of_mem (l₁ : List α) (h : a ∈ l₂) : (a :: l₁) ∩ l
107107
theorem inter_cons_of_notMem (l₁ : List α) (h : a ∉ l₂) : (a :: l₁) ∩ l₂ = l₁ ∩ l₂ := by
108108
simp [Inter.inter, List.inter, h]
109109

110+
@[grind =]
111+
theorem inter_cons (l₁ : List α) :
112+
(a :: l₁) ∩ l₂ = if a ∈ l₂ then a :: l₁ ∩ l₂ else l₁ ∩ l₂ := by
113+
split_ifs <;> simp_all
114+
110115
@[deprecated (since := "2025-05-23")] alias inter_cons_of_not_mem := inter_cons_of_notMem
111116

112-
@[simp]
117+
@[simp, grind =]
113118
theorem inter_nil' (l : List α) : l ∩ [] = [] := by
114-
induction l with
115-
| nil => rfl
116-
| cons x xs ih => by_cases x ∈ xs <;> simp [ih]
119+
induction l with grind
117120

118121
theorem mem_of_mem_inter_left : a ∈ l₁ ∩ l₂ → a ∈ l₁ :=
119122
mem_of_mem_filter
@@ -159,68 +162,38 @@ end Inter
159162

160163
section BagInter
161164

162-
@[simp]
165+
@[simp, grind =]
163166
theorem nil_bagInter (l : List α) : [].bagInter l = [] := by cases l <;> rfl
164167

165-
@[simp]
168+
@[simp, grind =]
166169
theorem bagInter_nil (l : List α) : l.bagInter [] = [] := by cases l <;> rfl
167170

168171
@[simp]
169172
theorem cons_bagInter_of_pos (l₁ : List α) (h : a ∈ l₂) :
170173
(a :: l₁).bagInter l₂ = a :: l₁.bagInter (l₂.erase a) := by
171-
cases l₂
172-
· exact if_pos h
173-
· simp only [List.bagInter, if_pos (elem_eq_true_of_mem h)]
174+
cases l₂ with grind [List.bagInter]
174175

175176
@[simp]
176177
theorem cons_bagInter_of_neg (l₁ : List α) (h : a ∉ l₂) :
177178
(a :: l₁).bagInter l₂ = l₁.bagInter l₂ := by
178-
cases l₂; · simp only [bagInter_nil]
179-
simp only [erase_of_not_mem h, List.bagInter, if_neg (mt mem_of_elem_eq_true h)]
179+
cases l₂ with grind [List.bagInter]
180+
181+
@[grind =]
182+
theorem cons_bagInteger :
183+
(a :: l₁).bagInter l₂ = if a ∈ l₂ then a :: l₁.bagInter (l₂.erase a) else l₁.bagInter l₂ := by
184+
split_ifs <;> simp_all
180185

181186
@[simp]
182-
theorem mem_bagInter {a : α} : ∀ {l₁ l₂ : List α}, a ∈ l₁.bagInter l₂ ↔ a ∈ l₁ ∧ a ∈ l₂
183-
| [], l₂ => by simp only [nil_bagInter, not_mem_nil, false_and]
184-
| b :: l₁, l₂ => by
185-
by_cases h : b ∈ l₂
186-
· rw [cons_bagInter_of_pos _ h, mem_cons, mem_cons, mem_bagInter]
187-
by_cases ba : a = b
188-
· simp only [ba, h, true_or, true_and]
189-
· simp only [mem_erase_of_ne ba, ba, false_or]
190-
· rw [cons_bagInter_of_neg _ h, mem_bagInter, mem_cons, or_and_right]
191-
symm
192-
apply or_iff_right_of_imp
193-
rintro ⟨rfl, h'⟩
194-
exact h.elim h'
187+
theorem mem_bagInter {a : α} {l₁ l₂ : List α} : a ∈ l₁.bagInter l₂ ↔ a ∈ l₁ ∧ a ∈ l₂ := by
188+
fun_induction List.bagInter with grind
195189

196190
@[simp]
197-
theorem count_bagInter {a : α} :
198-
∀ {l₁ l₂ : List α}, count a (l₁.bagInter l₂) = min (count a l₁) (count a l₂)
199-
| [], l₂ => by simp
200-
| l₁, [] => by simp
201-
| b :: l₁, l₂ => by
202-
by_cases hb : b ∈ l₂
203-
· rw [cons_bagInter_of_pos _ hb, count_cons, count_cons, count_bagInter, count_erase,
204-
← Nat.add_min_add_right]
205-
by_cases ba : b = a
206-
· simp only [beq_iff_eq]
207-
rw [if_pos ba, Nat.sub_add_cancel]
208-
rwa [succ_le_iff, count_pos_iff, ← ba]
209-
· simp only [beq_iff_eq]
210-
rw [if_neg ba, Nat.sub_zero, Nat.add_zero, Nat.add_zero]
211-
· rw [cons_bagInter_of_neg _ hb, count_bagInter]
212-
by_cases ab : a = b
213-
· rw [← ab] at hb
214-
rw [count_eq_zero.2 hb, Nat.min_zero, Nat.min_zero]
215-
· rw [count_cons_of_ne (Ne.symm ab)]
216-
217-
theorem bagInter_sublist_left : ∀ l₁ l₂ : List α, l₁.bagInter l₂ <+ l₁
218-
| [], l₂ => by simp
219-
| b :: l₁, l₂ => by
220-
by_cases h : b ∈ l₂ <;> simp only [h, cons_bagInter_of_pos, cons_bagInter_of_neg, not_false_iff]
221-
· exact (bagInter_sublist_left _ _).cons_cons _
222-
· apply sublist_cons_of_sublist
223-
apply bagInter_sublist_left
191+
theorem count_bagInter {a : α} {l₁ l₂ : List α} :
192+
count a (l₁.bagInter l₂) = min (count a l₁) (count a l₂) := by
193+
fun_induction List.bagInter with grind [count_pos_iff]
194+
195+
theorem bagInter_sublist_left {l₁ l₂ : List α} : l₁.bagInter l₂ <+ l₁ := by
196+
fun_induction List.bagInter with grind
224197

225198
theorem bagInter_nil_iff_inter_nil : ∀ l₁ l₂ : List α, l₁.bagInter l₂ = [] ↔ l₁ ∩ l₂ = []
226199
| [], l₂ => by simp

0 commit comments

Comments
 (0)