@@ -24,7 +24,7 @@ open Nat
2424
2525/-! ### Pairwise -/
2626
27- theorem Pairwise.sublist : l₁ <+ l₂ → l₂.Pairwise R → l₁.Pairwise R
27+ @[grind →] theorem Pairwise.sublist : l₁ <+ l₂ → l₂.Pairwise R → l₁.Pairwise R
2828 | .slnil, h => h
2929 | .cons _ s, .cons _ h₂ => h₂.sublist s
3030 | .cons₂ _ s, .cons h₁ h₂ => (h₂.sublist s).cons fun _ h => h₁ _ (s.subset h)
@@ -37,11 +37,11 @@ theorem Pairwise.imp {α R S} (H : ∀ {a b}, R a b → S a b) :
3737theorem rel_of_pairwise_cons (p : (a :: l).Pairwise R) : ∀ {a'}, a' ∈ l → R a a' :=
3838 (pairwise_cons.1 p).1 _
3939
40- theorem Pairwise.of_cons (p : (a :: l).Pairwise R) : Pairwise R l :=
40+ @[grind →] theorem Pairwise.of_cons (p : (a :: l).Pairwise R) : Pairwise R l :=
4141 (pairwise_cons.1 p).2
4242
4343set_option linter.unusedVariables false in
44- theorem Pairwise.tail : ∀ {l : List α} (h : Pairwise R l), Pairwise R l.tail
44+ @[grind] theorem Pairwise.tail : ∀ {l : List α} (h : Pairwise R l), Pairwise R l.tail
4545 | [], h => h
4646 | _ :: _, h => h.of_cons
4747
@@ -101,11 +101,11 @@ theorem Pairwise.forall_of_forall_of_flip (h₁ : ∀ x ∈ l, R x x) (h₂ : Pa
101101 · exact h₃.1 _ hx
102102 · exact ih (fun x hx => h₁ _ <| mem_cons_of_mem _ hx) h₂.2 h₃.2 hx hy
103103
104- theorem pairwise_singleton (R) (a : α) : Pairwise R [a] := by simp
104+ @[grind] theorem pairwise_singleton (R) (a : α) : Pairwise R [a] := by simp
105105
106- theorem pairwise_pair {a b : α} : Pairwise R [a, b] ↔ R a b := by simp
106+ @[grind =] theorem pairwise_pair {a b : α} : Pairwise R [a, b] ↔ R a b := by simp
107107
108- theorem pairwise_map {l : List α} :
108+ @[grind =] theorem pairwise_map {l : List α} :
109109 (l.map f).Pairwise R ↔ l.Pairwise fun a b => R (f a) (f b) := by
110110 induction l
111111 · simp
@@ -115,11 +115,11 @@ theorem Pairwise.of_map {S : β → β → Prop} (f : α → β) (H : ∀ a b :
115115 (p : Pairwise S (map f l)) : Pairwise R l :=
116116 (pairwise_map.1 p).imp (H _ _)
117117
118- theorem Pairwise.map {S : β → β → Prop } (f : α → β) (H : ∀ a b : α, R a b → S (f a) (f b))
118+ @[grind] theorem Pairwise.map {S : β → β → Prop } (f : α → β) (H : ∀ a b : α, R a b → S (f a) (f b))
119119 (p : Pairwise R l) : Pairwise S (map f l) :=
120120 pairwise_map.2 <| p.imp (H _ _)
121121
122- theorem pairwise_filterMap {f : β → Option α} {l : List β} :
122+ @[grind =] theorem pairwise_filterMap {f : β → Option α} {l : List β} :
123123 Pairwise R (filterMap f l) ↔ Pairwise (fun a a' : β => ∀ b, f a = some b → ∀ b', f a' = some b' → R b b') l := by
124124 let _S (a a' : β) := ∀ b, f a = some b → ∀ b', f a' = some b' → R b b'
125125 induction l with
@@ -134,20 +134,20 @@ theorem pairwise_filterMap {f : β → Option α} {l : List β} :
134134 simpa [IH, e] using fun _ =>
135135 ⟨fun h a ha b hab => h _ _ ha hab, fun h a b ha hab => h _ ha _ hab⟩
136136
137- theorem Pairwise.filterMap {S : β → β → Prop } (f : α → Option β)
137+ @[grind] theorem Pairwise.filterMap {S : β → β → Prop } (f : α → Option β)
138138 (H : ∀ a a' : α, R a a' → ∀ b, f a = some b → ∀ b', f a' = some b' → S b b') {l : List α} (p : Pairwise R l) :
139139 Pairwise S (filterMap f l) :=
140140 pairwise_filterMap.2 <| p.imp (H _ _)
141141
142- theorem pairwise_filter {p : α → Prop } [DecidablePred p] {l : List α} :
142+ @[grind =] theorem pairwise_filter {p : α → Bool} {l : List α} :
143143 Pairwise R (filter p l) ↔ Pairwise (fun x y => p x → p y → R x y) l := by
144144 rw [← filterMap_eq_filter, pairwise_filterMap]
145145 simp
146146
147- theorem Pairwise.filter (p : α → Bool) : Pairwise R l → Pairwise R (filter p l) :=
147+ @[grind] theorem Pairwise.filter (p : α → Bool) : Pairwise R l → Pairwise R (filter p l) :=
148148 Pairwise.sublist filter_sublist
149149
150- theorem pairwise_append {l₁ l₂ : List α} :
150+ @[grind =] theorem pairwise_append {l₁ l₂ : List α} :
151151 (l₁ ++ l₂).Pairwise R ↔ l₁.Pairwise R ∧ l₂.Pairwise R ∧ ∀ a ∈ l₁, ∀ b ∈ l₂, R a b := by
152152 induction l₁ <;> simp [*, or_imp, forall_and, and_assoc, and_left_comm]
153153
@@ -157,13 +157,13 @@ theorem pairwise_append_comm {R : α → α → Prop} (s : ∀ {x y}, R x y →
157157 (x : α) (xm : x ∈ l₂) (y : α) (ym : y ∈ l₁) : R x y := s (H y ym x xm)
158158 simp only [pairwise_append, and_left_comm]; rw [Iff.intro (this l₁ l₂) (this l₂ l₁)]
159159
160- theorem pairwise_middle {R : α → α → Prop } (s : ∀ {x y}, R x y → R y x) {a : α} {l₁ l₂ : List α} :
160+ @[grind =] theorem pairwise_middle {R : α → α → Prop } (s : ∀ {x y}, R x y → R y x) {a : α} {l₁ l₂ : List α} :
161161 Pairwise R (l₁ ++ a :: l₂) ↔ Pairwise R (a :: (l₁ ++ l₂)) := by
162162 show Pairwise R (l₁ ++ ([a] ++ l₂)) ↔ Pairwise R ([a] ++ l₁ ++ l₂)
163163 rw [← append_assoc, pairwise_append, @pairwise_append _ _ ([a] ++ l₁), pairwise_append_comm s]
164164 simp only [mem_append, or_comm]
165165
166- theorem pairwise_flatten {L : List (List α)} :
166+ @[grind =] theorem pairwise_flatten {L : List (List α)} :
167167 Pairwise R (flatten L) ↔
168168 (∀ l ∈ L, Pairwise R l) ∧ Pairwise (fun l₁ l₂ => ∀ x ∈ l₁, ∀ y ∈ l₂, R x y) L := by
169169 induction L with
@@ -174,16 +174,16 @@ theorem pairwise_flatten {L : List (List α)} :
174174 rw [and_comm, and_congr_left_iff]
175175 intros; exact ⟨fun h l' b c d e => h c d e l' b, fun h c d e l' b => h l' b c d e⟩
176176
177- theorem pairwise_flatMap {R : β → β → Prop } {l : List α} {f : α → List β} :
177+ @[grind =] theorem pairwise_flatMap {R : β → β → Prop } {l : List α} {f : α → List β} :
178178 List.Pairwise R (l.flatMap f) ↔
179179 (∀ a ∈ l, Pairwise R (f a)) ∧ Pairwise (fun a₁ a₂ => ∀ x ∈ f a₁, ∀ y ∈ f a₂, R x y) l := by
180180 simp [List.flatMap, pairwise_flatten, pairwise_map]
181181
182- theorem pairwise_reverse {l : List α} :
182+ @[grind =] theorem pairwise_reverse {l : List α} :
183183 l.reverse.Pairwise R ↔ l.Pairwise (fun a b => R b a) := by
184184 induction l <;> simp [*, pairwise_append, and_comm]
185185
186- @[simp] theorem pairwise_replicate {n : Nat} {a : α} :
186+ @[simp, grind = ] theorem pairwise_replicate {n : Nat} {a : α} :
187187 (replicate n a).Pairwise R ↔ n ≤ 1 ∨ R a a := by
188188 induction n with
189189 | zero => simp
@@ -205,10 +205,10 @@ theorem pairwise_reverse {l : List α} :
205205 simp
206206 · exact ⟨fun _ => h, Or.inr h⟩
207207
208- theorem Pairwise.drop {l : List α} {i : Nat} (h : List.Pairwise R l) : List.Pairwise R (l.drop i) :=
208+ @[grind] theorem Pairwise.drop {l : List α} {i : Nat} (h : List.Pairwise R l) : List.Pairwise R (l.drop i) :=
209209 h.sublist (drop_sublist _ _)
210210
211- theorem Pairwise.take {l : List α} {i : Nat} (h : List.Pairwise R l) : List.Pairwise R (l.take i) :=
211+ @[grind] theorem Pairwise.take {l : List α} {i : Nat} (h : List.Pairwise R l) : List.Pairwise R (l.take i) :=
212212 h.sublist (take_sublist _ _)
213213
214214theorem pairwise_iff_forall_sublist : l.Pairwise R ↔ (∀ {a b}, [a,b] <+ l → R a b) := by
@@ -247,7 +247,7 @@ theorem pairwise_of_forall_mem_list {l : List α} {r : α → α → Prop} (h :
247247 intro a b hab
248248 apply h <;> (apply hab.subset; simp)
249249
250- theorem pairwise_pmap {p : β → Prop } {f : ∀ b, p b → α} {l : List β} (h : ∀ x ∈ l, p x) :
250+ @[grind =] theorem pairwise_pmap {p : β → Prop } {f : ∀ b, p b → α} {l : List β} (h : ∀ x ∈ l, p x) :
251251 Pairwise R (l.pmap f h) ↔
252252 Pairwise (fun b₁ b₂ => ∀ (h₁ : p b₁) (h₂ : p b₂), R (f b₁ h₁) (f b₂ h₂)) l := by
253253 induction l with
@@ -259,7 +259,7 @@ theorem pairwise_pmap {p : β → Prop} {f : ∀ b, p b → α} {l : List β} (h
259259 rintro H _ b hb rfl
260260 exact H b hb _ _
261261
262- theorem Pairwise.pmap {l : List α} (hl : Pairwise R l) {p : α → Prop } {f : ∀ a, p a → β}
262+ @[grind] theorem Pairwise.pmap {l : List α} (hl : Pairwise R l) {p : α → Prop } {f : ∀ a, p a → β}
263263 (h : ∀ x ∈ l, p x) {S : β → β → Prop }
264264 (hS : ∀ ⦃x⦄ (hx : p x) ⦃y⦄ (hy : p y), R x y → S (f x hx) (f y hy)) :
265265 Pairwise S (l.pmap f h) := by
@@ -268,15 +268,15 @@ theorem Pairwise.pmap {l : List α} (hl : Pairwise R l) {p : α → Prop} {f :
268268
269269/-! ### Nodup -/
270270
271- @[simp]
271+ @[simp, grind ]
272272theorem nodup_nil : @Nodup α [] :=
273273 Pairwise.nil
274274
275- @[simp]
275+ @[simp, grind = ]
276276theorem nodup_cons {a : α} {l : List α} : Nodup (a :: l) ↔ a ∉ l ∧ Nodup l := by
277277 simp only [Nodup, pairwise_cons, forall_mem_ne]
278278
279- theorem Nodup.sublist : l₁ <+ l₂ → Nodup l₂ → Nodup l₁ :=
279+ @[grind →] theorem Nodup.sublist : l₁ <+ l₂ → Nodup l₂ → Nodup l₁ :=
280280 Pairwise.sublist
281281
282282theorem Sublist.nodup : l₁ <+ l₂ → Nodup l₂ → Nodup l₁ :=
@@ -303,7 +303,7 @@ theorem getElem?_inj {xs : List α}
303303 rw [mem_iff_getElem?]
304304 exact ⟨_, h₂⟩; exact ⟨_ , h₂.symm⟩
305305
306- @[simp] theorem nodup_replicate {n : Nat} {a : α} :
306+ @[simp, grind = ] theorem nodup_replicate {n : Nat} {a : α} :
307307 (replicate n a).Nodup ↔ n ≤ 1 := by simp [Nodup]
308308
309309end List
0 commit comments