Skip to content

Commit

Permalink
feat: List.SatisfiesM_foldlM
Browse files Browse the repository at this point in the history
  • Loading branch information
kim-em committed Nov 11, 2024
1 parent 325d06c commit 0d3360b
Show file tree
Hide file tree
Showing 2 changed files with 30 additions and 0 deletions.
1 change: 1 addition & 0 deletions Batteries/Data/List.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ import Batteries.Data.List.Count
import Batteries.Data.List.EraseIdx
import Batteries.Data.List.Init.Lemmas
import Batteries.Data.List.Lemmas
import Batteries.Data.List.Monadic
import Batteries.Data.List.OfFn
import Batteries.Data.List.Pairwise
import Batteries.Data.List.Perm
29 changes: 29 additions & 0 deletions Batteries/Data/List/Monadic.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,29 @@
/-
Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Kim Morrison
-/
import Batteries.Classes.SatisfiesM

/-!
# Results about monadic operations on `List`, in terms of `SatisfiesM`.
-/

namespace List

theorem SatisfiesM_foldlM [Monad m] [LawfulMonad m] {f : β → α → m β} (h₀ : motive b)
(h₁ : ∀ (b) (_ : motive b) (a : α) (_ : a ∈ l), SatisfiesM motive (f b a)) :
SatisfiesM motive (List.foldlM f b l) := by
have g b hb a am := Classical.indefiniteDescription _ (h₁ b hb a am)
clear h₁
induction l generalizing b with
| nil => exact SatisfiesM.pure h₀
| cons hd tl ih =>
simp only [foldlM_cons]
apply SatisfiesM.bind_pre
let ⟨q, qh⟩ := g b h₀ hd (List.mem_cons_self hd tl)
exact ⟨(fun ⟨b, bh⟩ => ⟨b, ih bh (fun b bh a am => g b bh a (mem_cons_of_mem hd am))⟩) <$> q,
by simpa using qh⟩

end List

0 comments on commit 0d3360b

Please sign in to comment.