@@ -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
159159theorem 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
165162theorem 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
304299theorem 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
310303theorem 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
334327theorem 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
345332theorem 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
380367theorem 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
435418theorem 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
441421theorem 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
524503theorem 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)
551529theorem 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
559534theorem 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
613585end deprecated
614586
@@ -987,18 +959,7 @@ variable (p)
987959
988960theorem 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
1003964lemma 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`.
10421003theorem 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
10461006end 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
10821034theorem 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
10861037end Erase
10871038
@@ -1210,12 +1161,7 @@ variable [BEq α] [LawfulBEq α]
12101161
12111162lemma 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
12201166end lookup
12211167
0 commit comments