diff --git a/src/Init/Data/List/Lemmas.lean b/src/Init/Data/List/Lemmas.lean index 18a043a36cc2..7003a515fd45 100644 --- a/src/Init/Data/List/Lemmas.lean +++ b/src/Init/Data/List/Lemmas.lean @@ -745,6 +745,58 @@ theorem foldr_hom (f : β₁ → β₂) (g₁ : α → β₁ → β₁) (g₂ : (H : ∀ x y, g₂ x (f y) = f (g₁ x y)) : l.foldr g₂ (f init) = f (l.foldr g₁ init) := by induction l <;> simp [*, H] +/-- +Prove a proposition about the result of `List.foldl`, +by proving it for the initial data, +and the implication that the operation applied to any element of the list preserves the property. + +The motive can take values in `Sort _`, so this may be used to construct data, +as well as to prove propositions. +-/ +def foldlRecOn {motive : β → Sort _} : ∀ (l : List α) (op : β → α → β) (b : β) (_ : motive b) + (_ : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ l), motive (op b a)), motive (List.foldl op b l) + | [], _, _, hb, _ => hb + | hd :: tl, op, b, hb, hl => + foldlRecOn tl op (op b hd) (hl b hb hd (mem_cons_self hd tl)) + fun y hy x hx => hl y hy x (mem_cons_of_mem hd hx) + +@[simp] theorem foldlRecOn_nil {motive : β → Sort _} (hb : motive b) + (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ []), motive (op b a)) : + foldlRecOn [] op b hb hl = hb := rfl + +@[simp] theorem foldlRecOn_cons {motive : β → Sort _} (hb : motive b) + (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ x :: l), motive (op b a)) : + foldlRecOn (x :: l) op b hb hl = + foldlRecOn l op (op b x) (hl b hb x (mem_cons_self x l)) + (fun b c a m => hl b c a (mem_cons_of_mem x m)) := + rfl + +/-- +Prove a proposition about the result of `List.foldr`, +by proving it for the initial data, +and the implication that the operation applied to any element of the list preserves the property. + +The motive can take values in `Sort _`, so this may be used to construct data, +as well as to prove propositions. +-/ +def foldrRecOn {motive : β → Sort _} : ∀ (l : List α) (op : α → β → β) (b : β) (_ : motive b) + (_ : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ l), motive (op a b)), motive (List.foldr op b l) + | nil, _, _, hb, _ => hb + | x :: l, op, b, hb, hl => + hl (foldr op b l) + (foldrRecOn l op b hb fun b c a m => hl b c a (mem_cons_of_mem x m)) x (mem_cons_self x l) + +@[simp] theorem foldrRecOn_nil {motive : β → Sort _} (hb : motive b) + (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ []), motive (op a b)) : + foldrRecOn [] op b hb hl = hb := rfl + +@[simp] theorem foldrRecOn_cons {motive : β → Sort _} (hb : motive b) + (hl : ∀ (b : β) (_ : motive b) (a : α) (_ : a ∈ x :: l), motive (op a b)) : + foldrRecOn (x :: l) op b hb hl = + hl _ (foldrRecOn l op b hb fun b c a m => hl b c a (mem_cons_of_mem x m)) + x (mem_cons_self x l) := + rfl + /-! ### getLast -/ theorem getLast_eq_getElem : ∀ (l : List α) (h : l ≠ []),