Skip to content

Commit

Permalink
feat: improvements to simp confluence (#7013)
Browse files Browse the repository at this point in the history
This PR makes improvements to the simp set for List/Array/Vector/Option
to improve confluence, in preparation for `simp_lc`.
  • Loading branch information
kim-em authored Feb 10, 2025
1 parent 2aca375 commit c307e8a
Show file tree
Hide file tree
Showing 11 changed files with 211 additions and 56 deletions.
20 changes: 14 additions & 6 deletions src/Init/Data/Array/Attach.lean
Original file line number Diff line number Diff line change
Expand Up @@ -70,7 +70,7 @@ Unsafe implementation of `attachWith`, taking advantage of the fact that the rep
@[csimp] private theorem pmap_eq_pmapImpl : @pmap = @pmapImpl := by
funext α β p f L h'
cases L
simp only [pmap, pmapImpl, List.attachWith_toArray, List.map_toArray, mk.injEq, List.map_attachWith]
simp only [pmap, pmapImpl, List.attachWith_toArray, List.map_toArray, mk.injEq, List.map_attachWith_eq_pmap]
apply List.pmap_congr_left
intro a m h₁ h₂
congr
Expand Down Expand Up @@ -318,7 +318,7 @@ See however `foldl_subtype` below.
theorem foldl_attach (l : Array α) (f : β → α → β) (b : β) :
l.attach.foldl (fun acc t => f acc t.1) b = l.foldl f b := by
rcases l with ⟨l⟩
simp only [List.attach_toArray, List.attachWith_mem_toArray, List.map_attach, size_toArray,
simp only [List.attach_toArray, List.attachWith_mem_toArray, size_toArray,
List.length_pmap, List.foldl_toArray', mem_toArray, List.foldl_subtype]
congr
ext
Expand All @@ -337,7 +337,7 @@ See however `foldr_subtype` below.
theorem foldr_attach (l : Array α) (f : α → β → β) (b : β) :
l.attach.foldr (fun t acc => f t.1 acc) b = l.foldr f b := by
rcases l with ⟨l⟩
simp only [List.attach_toArray, List.attachWith_mem_toArray, List.map_attach, size_toArray,
simp only [List.attach_toArray, List.attachWith_mem_toArray, size_toArray,
List.length_pmap, List.foldr_toArray', mem_toArray, List.foldr_subtype]
congr
ext
Expand All @@ -354,19 +354,27 @@ theorem attachWith_map {l : Array α} (f : α → β) {P : β → Prop} {H : ∀
cases l
simp [List.attachWith_map]

theorem map_attachWith {l : Array α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a}
@[simp] theorem map_attachWith {l : Array α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a}
(f : { x // P x } → β) :
(l.attachWith P H).map f = l.attach.map fun ⟨x, h⟩ => f ⟨x, H _ h⟩ := by
cases l <;> simp_all

theorem map_attachWith_eq_pmap {l : Array α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a}
(f : { x // P x } → β) :
(l.attachWith P H).map f =
l.pmap (fun a (h : a ∈ l ∧ P a) => f ⟨a, H _ h.1⟩) (fun a h => ⟨h, H a h⟩) := by
cases l
ext <;> simp

/-- See also `pmap_eq_map_attach` for writing `pmap` in terms of `map` and `attach`. -/
theorem map_attach {l : Array α} (f : { x // x ∈ l } → β) :
theorem map_attach_eq_pmap {l : Array α} (f : { x // x ∈ l } → β) :
l.attach.map f = l.pmap (fun a h => f ⟨a, h⟩) (fun _ => id) := by
cases l
ext <;> simp

@[deprecated map_attach_eq_pmap (since := "2025-02-09")]
abbrev map_attach := @map_attach_eq_pmap

theorem attach_filterMap {l : Array α} {f : α → Option β} :
(l.filterMap f).attach = l.attach.filterMap
fun ⟨x, h⟩ => (f x).pbind (fun b m => some ⟨b, mem_filterMap.mpr ⟨x, h, m⟩⟩) := by
Expand Down Expand Up @@ -505,7 +513,7 @@ theorem count_attach [DecidableEq α] (l : Array α) (a : {x // x ∈ l}) :
l.attach.count a = l.count ↑a := by
rcases l with ⟨l⟩
simp only [List.attach_toArray, List.attachWith_mem_toArray, List.count_toArray]
rw [List.map_attach, List.count_eq_countP]
rw [List.map_attach_eq_pmap, List.count_eq_countP]
simp only [Subtype.beq_iff]
rw [List.countP_pmap, List.countP_attach (p := (fun x => x == a.1)), List.count]

Expand Down
2 changes: 1 addition & 1 deletion src/Init/Data/Array/Count.lean
Original file line number Diff line number Diff line change
Expand Up @@ -163,7 +163,7 @@ theorem count_le_size (a : α) (l : Array α) : count a l ≤ l.size := countP_l
theorem count_le_count_push (a b : α) (l : Array α) : count a l ≤ count a (l.push b) := by
simp [count_push]

@[simp] theorem count_singleton (a b : α) : count a #[b] = if b == a then 1 else 0 := by
theorem count_singleton (a b : α) : count a #[b] = if b == a then 1 else 0 := by
simp [count_eq_countP]

@[simp] theorem count_append (a : α) : ∀ l₁ l₂, count a (l₁ ++ l₂) = count a l₁ + count a l₂ :=
Expand Down
121 changes: 100 additions & 21 deletions src/Init/Data/Array/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1350,8 +1350,9 @@ theorem map_filter_eq_foldl (f : α → β) (p : α → Bool) (l : Array α) :
simp only [List.filter_cons, List.foldr_cons]
split <;> simp_all

@[simp] theorem filter_append {p : α → Bool} (l₁ l₂ : Array α) :
filter p (l₁ ++ l₂) = filter p l₁ ++ filter p l₂ := by
@[simp] theorem filter_append {p : α → Bool} (l₁ l₂ : Array α) (w : stop = l₁.size + l₂.size) :
filter p (l₁ ++ l₂) 0 stop = filter p l₁ ++ filter p l₂ := by
subst w
rcases l₁ with ⟨l₁⟩
rcases l₂ with ⟨l₂⟩
simp [List.filter_append]
Expand Down Expand Up @@ -1440,12 +1441,18 @@ theorem _root_.List.filterMap_toArray (f : α → Option β) (l : List α) :
rcases l with ⟨l⟩
simp [h]

@[simp] theorem filterMap_eq_map (f : α → β) (w : stop = as.size ) :
@[simp] theorem filterMap_eq_map (f : α → β) (w : stop = as.size) :
filterMap (some ∘ f) as 0 stop = map f as := by
subst w
cases as
simp

/-- Variant of `filterMap_eq_map` with `some ∘ f` expanded out to a lambda. -/
@[simp]
theorem filterMap_eq_map' (f : α → β) (w : stop = as.size) :
filterMap (fun x => some (f x)) as 0 stop = map f as :=
filterMap_eq_map f w

theorem filterMap_some_fun : filterMap (some : α → Option α) = id := by
funext l
cases l
Expand Down Expand Up @@ -1514,8 +1521,9 @@ theorem forall_mem_filterMap {f : α → Option β} {l : Array α} {P : β → P
intro a
rw [forall_comm]

@[simp] theorem filterMap_append {α β : Type _} (l l' : Array α) (f : α → Option β) :
filterMap f (l ++ l') = filterMap f l ++ filterMap f l' := by
@[simp] theorem filterMap_append {α β : Type _} (l l' : Array α) (f : α → Option β) (w : stop = l.size + l'.size) :
filterMap f (l ++ l') 0 stop = filterMap f l ++ filterMap f l' := by
subst w
cases l
cases l'
simp
Expand Down Expand Up @@ -1557,7 +1565,12 @@ theorem filterMap_eq_push_iff {f : α → Option β} {l : Array α} {l' : Array
@[simp] theorem size_append (as bs : Array α) : (as ++ bs).size = as.size + bs.size := by
simp only [size, toList_append, List.length_append]

@[simp] theorem append_push {as bs : Array α} {a : α} : as ++ bs.push a = (as ++ bs).push a := by
@[simp] theorem push_append {a : α} {xs ys : Array α} : (xs ++ ys).push a = xs ++ ys.push a := by
cases xs
cases ys
simp

theorem append_push {as bs : Array α} {a : α} : as ++ bs.push a = (as ++ bs).push a := by
cases as
cases bs
simp
Expand Down Expand Up @@ -1668,6 +1681,9 @@ theorem getElem_of_append {l l₁ l₂ : Array α} (eq : l = l₁.push a ++ l₂

@[simp] theorem append_singleton {a : α} {as : Array α} : as ++ #[a] = as.push a := rfl

@[simp] theorem append_singleton_assoc {a : α} {as bs : Array α} : as ++ (#[a] ++ bs) = as.push a ++ bs := by
rw [← append_assoc, append_singleton]

theorem push_eq_append {a : α} {as : Array α} : as.push a = as ++ #[a] := rfl

theorem append_inj {s₁ s₂ t₁ t₂ : Array α} (h : s₁ ++ t₁ = s₂ ++ t₂) (hl : s₁.size = s₂.size) :
Expand Down Expand Up @@ -1931,15 +1947,17 @@ theorem flatten_eq_flatMap {L : Array (Array α)} : flatten L = L.flatMap id :=
Function.comp_def]
rw [← Function.comp_def, ← List.map_map, flatten_toArray_map]

@[simp] theorem filterMap_flatten (f : α → Option β) (L : Array (Array α)) :
filterMap f (flatten L) = flatten (map (filterMap f) L) := by
@[simp] theorem filterMap_flatten (f : α → Option β) (L : Array (Array α)) (w : stop = L.flatten.size) :
filterMap f (flatten L) 0 stop = flatten (map (filterMap f) L) := by
subst w
induction L using array₂_induction
simp only [flatten_toArray_map, size_toArray, List.length_flatten, List.filterMap_toArray',
List.filterMap_flatten, List.map_toArray, List.map_map, Function.comp_def]
rw [← Function.comp_def, ← List.map_map, flatten_toArray_map]

@[simp] theorem filter_flatten (p : α → Bool) (L : Array (Array α)) :
filter p (flatten L) = flatten (map (filter p) L) := by
@[simp] theorem filter_flatten (p : α → Bool) (L : Array (Array α)) (w : stop = L.flatten.size) :
filter p (flatten L) 0 stop = flatten (map (filter p) L) := by
subst w
induction L using array₂_induction
simp only [flatten_toArray_map, size_toArray, List.length_flatten, List.filter_toArray',
List.filter_flatten, List.map_toArray, List.map_map, Function.comp_def]
Expand Down Expand Up @@ -2397,11 +2415,25 @@ theorem reverse_eq_iff {as bs : Array α} : as.reverse = bs ↔ as = bs.reverse
@[simp] theorem map_reverse (f : α → β) (l : Array α) : l.reverse.map f = (l.map f).reverse := by
cases l <;> simp [*]

@[simp] theorem filter_reverse (p : α → Bool) (l : Array α) : (l.reverse.filter p) = (l.filter p).reverse := by
/-- Variant of `filter_reverse` with a hypothesis giving the stop condition. -/
@[simp] theorem filter_reverse' (p : α → Bool) (l : Array α) (w : stop = l.size) :
(l.reverse.filter p 0 stop) = (l.filter p).reverse := by
subst w
cases l
simp

@[simp] theorem filterMap_reverse (f : α → Option β) (l : Array α) : (l.reverse.filterMap f) = (l.filterMap f).reverse := by
theorem filter_reverse (p : α → Bool) (l : Array α) : (l.reverse.filter p) = (l.filter p).reverse := by
cases l
simp

/-- Variant of `filterMap_reverse` with a hypothesis giving the stop condition. -/
@[simp] theorem filterMap_reverse' (f : α → Option β) (l : Array α) (w : stop = l.size) :
(l.reverse.filterMap f 0 stop) = (l.filterMap f).reverse := by
subst w
cases l
simp

theorem filterMap_reverse (f : α → Option β) (l : Array α) : (l.reverse.filterMap f) = (l.filterMap f).reverse := by
cases l
simp

Expand Down Expand Up @@ -2877,17 +2909,57 @@ rather than `(arr.push a).size` as the argument.
(h : start = arr.size + 1) : (arr.push a).foldr f init start = arr.foldr f (f a init) :=
foldrM_push' _ _ _ _ h

@[simp] theorem foldl_push_eq_append (l l' : Array α) : l.foldl push l' = l' ++ l := by
cases l
cases l'
@[simp] theorem foldl_push_eq_append {as : Array α} {bs : Array β} {f : α → β} (w : stop = as.size) :
as.foldl (fun b a => Array.push b (f a)) bs 0 stop = bs ++ as.map f := by
subst w
rcases as with ⟨as⟩
rcases bs with ⟨bs⟩
simp only [List.foldl_toArray']
induction as generalizing bs <;> simp [*]

@[simp] theorem foldl_cons_eq_append {as : Array α} {bs : List β} {f : α → β} (w : stop = as.size) :
as.foldl (fun b a => (f a) :: b) bs 0 stop = (as.map f).reverse.toList ++ bs := by
subst w
rcases as with ⟨as⟩
simp

@[simp] theorem foldr_flip_push_eq_append (l l' : Array α) :
l.foldr (fun x y => push y x) l' = l' ++ l.reverse := by
cases l
cases l'
@[simp] theorem foldr_cons_eq_append {as : Array α} {bs : List β} {f : α → β} (w : start = as.size) :
as.foldr (fun a b => (f a) :: b) bs start 0 = (as.map f).toList ++ bs := by
subst w
rcases as with ⟨as⟩
simp

/-- Variant of `foldr_cons_eq_append` specialized to `f = id`. -/
@[simp] theorem foldr_cons_eq_append' {as : Array α} {bs : List α} (w : start = as.size) :
as.foldr List.cons bs start 0 = as.toList ++ bs := by
subst w
rcases as with ⟨as⟩
simp

@[simp] theorem foldr_append_eq_append (l : Array α) (f : α → Array β) (l' : Array β) :
l.foldr (f · ++ ·) l' = (l.map f).flatten ++ l' := by
rcases l with ⟨l⟩
rcases l' with ⟨l'⟩
induction l <;> simp_all [Function.comp_def, flatten_toArray]

@[simp] theorem foldl_append_eq_append (l : Array α) (f : α → Array β) (l' : Array β) :
l.foldl (· ++ f ·) l' = l' ++ (l.map f).flatten := by
rcases l with ⟨l⟩
rcases l' with ⟨l'⟩
induction l generalizing l'<;> simp_all [Function.comp_def, flatten_toArray]

@[simp] theorem foldr_flip_append_eq_append (l : Array α) (f : α → Array β) (l' : Array β) :
l.foldr (fun x y => y ++ f x) l' = l' ++ (l.map f).reverse.flatten := by
rcases l with ⟨l⟩
rcases l' with ⟨l'⟩
induction l generalizing l' <;> simp_all [Function.comp_def, flatten_toArray]

@[simp] theorem foldl_flip_append_eq_append (l : Array α) (f : α → Array β) (l' : Array β) :
l.foldl (fun x y => f y ++ x) l' = (l.map f).reverse.flatten ++ l':= by
rcases l with ⟨l⟩
rcases l' with ⟨l'⟩
induction l generalizing l' <;> simp_all [Function.comp_def, flatten_toArray]

theorem foldl_map' (f : β₁ → β₂) (g : α → β₂ → α) (l : Array β₁) (init : α) (w : stop = l.size) :
(l.map f).foldl g init 0 stop = l.foldl (fun x y => g x (f y)) init := by
subst w
Expand Down Expand Up @@ -3040,6 +3112,13 @@ theorem foldl_eq_foldr_reverse (l : Array α) (f : β → α → β) (b) :
theorem foldr_eq_foldl_reverse (l : Array α) (f : α → β → β) (b) :
l.foldr f b = l.reverse.foldl (fun x y => f y x) b := by simp

@[simp] theorem foldr_push_eq_append {as : Array α} {bs : Array β} {f : α → β} (w : start = as.size) :
as.foldr (fun a b => Array.push b (f a)) bs start 0 = bs ++ (as.map f).reverse := by
subst w
rw [foldr_eq_foldl_reverse, foldl_push_eq_append rfl, map_reverse]

@[deprecated foldr_push_eq_append (since := "2025-02-09")] abbrev foldr_flip_push_eq_append := @foldr_push_eq_append

theorem foldl_assoc {op : α → α → α} [ha : Std.Associative op] {l : Array α} {a₁ a₂} :
l.foldl op (op a₁ a₂) = op a₁ (l.foldl op a₂) := by
rcases l with ⟨l⟩
Expand Down Expand Up @@ -3445,11 +3524,11 @@ theorem map_induction (as : Array α) (f : α → β) (motive : Nat → Prop) (h
(motive := fun i arr => motive i ∧ arr.size = i ∧ ∀ i h2, p i arr[i.1])
(init := #[]) (f := fun r a => r.push (f a)) ?_ ?_
obtain ⟨m, eq, w⟩ := t
· refine ⟨m, by simpa [map_eq_foldl] using eq, ?_⟩
· refine ⟨m, by simp, ?_⟩
intro i h
simp only [eq] at w
specialize w ⟨i, h⟩ h
simpa [map_eq_foldl] using w
simpa using w
· exact ⟨h0, rfl, nofun⟩
· intro i b ⟨m, ⟨eq, w⟩⟩
refine ⟨?_, ?_, ?_⟩
Expand Down
12 changes: 10 additions & 2 deletions src/Init/Data/List/Attach.lean
Original file line number Diff line number Diff line change
Expand Up @@ -416,7 +416,12 @@ theorem attachWith_map {l : List α} (f : α → β) {P : β → Prop} {H : ∀
fun ⟨x, h⟩ => ⟨f x, h⟩ := by
induction l <;> simp [*]

theorem map_attachWith {l : List α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a}
@[simp] theorem map_attachWith {l : List α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a}
(f : { x // P x } → β) :
(l.attachWith P H).map f = l.attach.map fun ⟨x, h⟩ => f ⟨x, H _ h⟩ := by
induction l <;> simp_all

theorem map_attachWith_eq_pmap {l : List α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a}
(f : { x // P x } → β) :
(l.attachWith P H).map f =
l.pmap (fun a (h : a ∈ l ∧ P a) => f ⟨a, H _ h.1⟩) (fun a h => ⟨h, H a h⟩) := by
Expand All @@ -428,7 +433,7 @@ theorem map_attachWith {l : List α} {P : α → Prop} {H : ∀ (a : α), a ∈
simp

/-- See also `pmap_eq_map_attach` for writing `pmap` in terms of `map` and `attach`. -/
theorem map_attach {l : List α} (f : { x // x ∈ l } → β) :
theorem map_attach_eq_pmap {l : List α} (f : { x // x ∈ l } → β) :
l.attach.map f = l.pmap (fun a h => f ⟨a, h⟩) (fun _ => id) := by
induction l with
| nil => rfl
Expand All @@ -437,6 +442,9 @@ theorem map_attach {l : List α} (f : { x // x ∈ l } → β) :
apply pmap_congr_left
simp

@[deprecated map_attach_eq_pmap (since := "2025-02-09")]
abbrev map_attach := @map_attach_eq_pmap

theorem attach_filterMap {l : List α} {f : α → Option β} :
(l.filterMap f).attach = l.attach.filterMap
fun ⟨x, h⟩ => (f x).pbind (fun b m => some ⟨b, mem_filterMap.mpr ⟨x, h, m⟩⟩) := by
Expand Down
21 changes: 17 additions & 4 deletions src/Init/Data/List/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -225,10 +225,23 @@ def lex [BEq α] (l₁ l₂ : List α) (lt : α → α → Bool := by exact (·
| _, [] => false
| a :: as, b :: bs => lt a b || (a == b && lex as bs lt)

@[simp] theorem lex_nil_nil [BEq α] : lex ([] : List α) [] lt = false := rfl
@[simp] theorem lex_nil_cons [BEq α] {b} {bs : List α} : lex [] (b :: bs) lt = true := rfl
@[simp] theorem lex_cons_nil [BEq α] {a} {as : List α} : lex (a :: as) [] lt = false := rfl
@[simp] theorem lex_cons_cons [BEq α] {a b} {as bs : List α} :
theorem nil_lex_nil [BEq α] : lex ([] : List α) [] lt = false := rfl
@[simp] theorem nil_lex_cons [BEq α] {b} {bs : List α} : lex [] (b :: bs) lt = true := rfl
theorem cons_lex_nil [BEq α] {a} {as : List α} : lex (a :: as) [] lt = false := rfl
@[simp] theorem cons_lex_cons [BEq α] {a b} {as bs : List α} :
lex (a :: as) (b :: bs) lt = (lt a b || (a == b && lex as bs lt)) := rfl

@[simp] theorem lex_nil [BEq α] {as : List α} : lex as [] lt = false := by
cases as <;> simp [nil_lex_nil, cons_lex_nil]

@[deprecated nil_lex_nil (since := "2025-02-10")]
theorem lex_nil_nil [BEq α] : lex ([] : List α) [] lt = false := rfl
@[deprecated nil_lex_cons (since := "2025-02-10")]
theorem lex_nil_cons [BEq α] {b} {bs : List α} : lex [] (b :: bs) lt = true := rfl
@[deprecated cons_lex_nil (since := "2025-02-10")]
theorem lex_cons_nil [BEq α] {a} {as : List α} : lex (a :: as) [] lt = false := rfl
@[deprecated cons_lex_cons (since := "2025-02-10")]
theorem lex_cons_cons [BEq α] {a b} {as bs : List α} :
lex (a :: as) (b :: bs) lt = (lt a b || (a == b && lex as bs lt)) := rfl

/-! ## Alternative getters -/
Expand Down
32 changes: 30 additions & 2 deletions src/Init/Data/List/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1344,6 +1344,11 @@ theorem head_filter_of_pos {p : α → Bool} {l : List α} (w : l ≠ []) (h : p
theorem filterMap_eq_map (f : α → β) : filterMap (some ∘ f) = map f := by
funext l; induction l <;> simp [*, filterMap_cons]

/-- Variant of `filterMap_eq_map` with `some ∘ f` expanded out to a lambda. -/
@[simp]
theorem filterMap_eq_map' (f : α → β) : filterMap (fun x => some (f x)) = map f :=
filterMap_eq_map f

@[simp] theorem filterMap_some_fun : filterMap (some : α → Option α) = id := by
funext l
erw [filterMap_eq_map]
Expand Down Expand Up @@ -2516,12 +2521,35 @@ theorem foldr_eq_foldrM (f : α → β → β) (b) (l : List α) :

/-! ### foldl and foldr -/

@[simp] theorem foldr_cons_eq_append (l : List α) : l.foldr cons l' = l ++ l' := by
@[simp] theorem foldr_cons_eq_append (l : List α) (f : α → β) (l' : List β) :
l.foldr (fun x y => f x :: y) l' = l.map f ++ l' := by
induction l <;> simp [*]

/-- Variant of `foldr_cons_eq_append` specalized to `f = id`. -/
@[simp] theorem foldr_cons_eq_append' (l l' : List β) :
l.foldr cons l' = l ++ l' := by
induction l <;> simp [*]

@[deprecated foldr_cons_eq_append (since := "2024-08-22")] abbrev foldr_self_append := @foldr_cons_eq_append

@[simp] theorem foldl_flip_cons_eq_append (l : List α) : l.foldl (fun x y => y :: x) l' = l.reverse ++ l' := by
@[simp] theorem foldl_flip_cons_eq_append (l : List α) (f : α → β) (l' : List β) :
l.foldl (fun x y => f y :: x) l' = (l.map f).reverse ++ l' := by
induction l generalizing l' <;> simp [*]

@[simp] theorem foldr_append_eq_append (l : List α) (f : α → List β) (l' : List β) :
l.foldr (f · ++ ·) l' = (l.map f).flatten ++ l' := by
induction l <;> simp [*]

@[simp] theorem foldl_append_eq_append (l : List α) (f : α → List β) (l' : List β) :
l.foldl (· ++ f ·) l' = l' ++ (l.map f).flatten := by
induction l generalizing l'<;> simp [*]

@[simp] theorem foldr_flip_append_eq_append (l : List α) (f : α → List β) (l' : List β) :
l.foldr (fun x y => y ++ f x) l' = l' ++ (l.map f).reverse.flatten := by
induction l generalizing l' <;> simp [*]

@[simp] theorem foldl_flip_append_eq_append (l : List α) (f : α → List β) (l' : List β) :
l.foldl (fun x y => f y ++ x) l' = (l.map f).reverse.flatten ++ l' := by
induction l generalizing l' <;> simp [*]

theorem foldr_cons_nil (l : List α) : l.foldr cons [] = l := by simp
Expand Down
Loading

0 comments on commit c307e8a

Please sign in to comment.