Skip to content

Commit

Permalink
chore: restore simpler code in foldrRecOn and foldlRecOn (#8508)
Browse files Browse the repository at this point in the history
#1720 made `List.rec` computable and therefore made it possible to restore the simpler versions of
`foldrRecOn` and `foldlRecOn`.
  • Loading branch information
dwrensha committed Nov 19, 2023
1 parent 35f9ee3 commit 778f663
Showing 1 changed file with 4 additions and 30 deletions.
34 changes: 4 additions & 30 deletions Mathlib/Data/List/Basic.lean
Expand Up @@ -2441,56 +2441,30 @@ theorem injective_foldl_comp {α : Type*} {l : List (α → α)} {f : α → α}
apply hl _ (List.mem_cons_self _ _)
#align list.injective_foldl_comp List.injective_foldl_comp

/- Porting note: couldn't do induction proof because "code generator does not support recursor
'List.rec' yet". Earlier proof:
induction l with
| nil => exact hb
| cons hd tl IH =>
refine' hl _ _ hd (mem_cons_self hd tl)
refine' IH _
intro y hy x hx
exact hl y hy x (mem_cons_of_mem hd hx)
-/
/-- Induction principle for values produced by a `foldr`: if a property holds
for the seed element `b : β` and for all incremental `op : α → β → β`
performed on the elements `(a : α) ∈ l`. The principle is given for
a `Sort`-valued predicate, i.e., it can also be used to construct data. -/
def foldrRecOn {C : β → Sort*} (l : List α) (op : α → β → β) (b : β) (hb : C b)
(hl : ∀ (b : β) (_ : C b) (a : α) (_ : a ∈ l), C (op a b)) : C (foldr op b l) := by
cases l with
induction l with
| nil => exact hb
| cons hd tl =>
have IH : ((b : β) → C b → (a : α) → a ∈ tl → C (op a b)) → C (foldr op b tl) :=
foldrRecOn _ _ _ hb
| cons hd tl IH =>
refine' hl _ _ hd (mem_cons_self hd tl)
refine' IH _
intro y hy x hx
exact hl y hy x (mem_cons_of_mem hd hx)
#align list.foldr_rec_on List.foldrRecOn

/- Porting note: couldn't do induction proof because "code generator does not support recursor
'List.rec' yet". Earlier proof:
induction l generalizing b with
| nil => exact hb
| cons hd tl IH =>
refine' IH _ _ _
· exact hl b hb hd (mem_cons_self hd tl)
· intro y hy x hx
exact hl y hy x (mem_cons_of_mem hd hx)
-/
/-- Induction principle for values produced by a `foldl`: if a property holds
for the seed element `b : β` and for all incremental `op : β → α → β`
performed on the elements `(a : α) ∈ l`. The principle is given for
a `Sort`-valued predicate, i.e., it can also be used to construct data. -/
def foldlRecOn {C : β → Sort*} (l : List α) (op : β → α → β) (b : β) (hb : C b)
(hl : ∀ (b : β) (_ : C b) (a : α) (_ : a ∈ l), C (op b a)) : C (foldl op b l) := by
cases l with
induction l generalizing b with
| nil => exact hb
| cons hd tl =>
have IH : (b : β) → C b → ((b : β) → C b → (a : α) → a ∈ tl → C (op b a)) → C (foldl op b tl) :=
foldlRecOn _ _
| cons hd tl IH =>
refine' IH _ _ _
· exact hl b hb hd (mem_cons_self hd tl)
· intro y hy x hx
Expand Down

0 comments on commit 778f663

Please sign in to comment.