@@ -26,6 +26,10 @@ namespace list
26
26
@[simp] theorem zip_with_nil_right (f : α → β → γ) (l) : zip_with f l [] = [] :=
27
27
by cases l; refl
28
28
29
+ @[simp] lemma zip_with_eq_nil_iff {f : α → β → γ} {l l'} :
30
+ zip_with f l l' = [] ↔ l = [] ∨ l' = [] :=
31
+ by { cases l; cases l'; simp }
32
+
29
33
@[simp] theorem zip_nil_left (l : list α) : zip ([] : list β) l = [] := rfl
30
34
31
35
@[simp] theorem zip_nil_right (l : list α) : zip l ([] : list β) = [] :=
@@ -199,15 +203,14 @@ by rw [← zip_unzip.{u u} (revzip l).reverse, unzip_eq_map]; simp; simp [revzip
199
203
theorem revzip_swap (l : list α) : (revzip l).map prod.swap = revzip l.reverse :=
200
204
by simp [revzip]
201
205
202
- lemma nth_zip_with {α β γ} (f : α → β → γ) (l₁ : list α) (l₂ : list β) (i : ℕ) :
203
- (zip_with f l₁ l₂).nth i = f <$> l₁.nth i <*> l₂.nth i :=
206
+ lemma nth_zip_with (f : α → β → γ) (l₁ : list α) (l₂ : list β) (i : ℕ) :
207
+ (zip_with f l₁ l₂).nth i = (( l₁.nth i).map f).bind (λ g, ( l₂.nth i).map g) :=
204
208
begin
205
209
induction l₁ generalizing l₂ i,
206
210
{ simp [zip_with, (<*>)] },
207
211
{ cases l₂; simp only [zip_with, has_seq.seq, functor.map, nth, option.map_none'],
208
212
{ cases ((l₁_hd :: l₁_tl).nth i); refl },
209
- { cases i; simp only [option.map_some', nth, option.some_bind', *],
210
- refl } },
213
+ { cases i; simp only [option.map_some', nth, option.some_bind', *] } }
211
214
end
212
215
213
216
lemma nth_zip_with_eq_some {α β γ} (f : α → β → γ) (l₁ : list α) (l₂ : list β) (z : γ) (i : ℕ) :
0 commit comments