Documentation

Init.Data.Iterators.Lemmas.Consumers.Loop

theorem Std.Iterators.Iter.forIn'_eq {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [hl : LawfulIteratorLoop α Id m] {γ : Type x} {it : Iter β} {init : γ} {f : (b : β) → it.IsPlausibleIndirectOutput bγm (ForInStep γ)} :
forIn' it init f = IterM.DefaultConsumers.forIn' (fun (x : Type w) (x_1 : Type x) (f : xm x_1) (x_2 : Id x) => f x_2.run) γ (fun (x : β) (x_1 : γ) (x_2 : ForInStep γ) => True) it.toIterM init it.toIterM.IsPlausibleIndirectOutput fun (out : β) (h : it.toIterM.IsPlausibleIndirectOutput out) (acc : γ) => (fun (x : ForInStep γ) => x, True.intro) <$> f out acc
theorem Std.Iterators.Iter.forIn_eq {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [hl : LawfulIteratorLoop α Id m] {γ : Type x} {it : Iter β} {init : γ} {f : βγm (ForInStep γ)} :
forIn it init f = IterM.DefaultConsumers.forIn' (fun (x : Type w) (x_1 : Type x) (f : xm x_1) (c : Id x) => f c.run) γ (fun (x : β) (x_1 : γ) (x_2 : ForInStep γ) => True) it.toIterM init it.toIterM.IsPlausibleIndirectOutput fun (out : β) (x : it.toIterM.IsPlausibleIndirectOutput out) (acc : γ) => (fun (x : ForInStep γ) => x, True.intro) <$> f out acc
theorem Std.Iterators.Iter.forIn'_congr {γ α β : Type w} {m : Type w → Type w'} [Monad m] [Iterator α Id β] [Finite α Id] [IteratorLoop α Id m] {ita itb : Iter β} (w : ita = itb) {b b' : γ} (hb : b = b') {f : (a' : β) → a' itaγm (ForInStep γ)} {g : (a' : β) → a' itbγm (ForInStep γ)} (h : ∀ (a : β) (m_1 : a itb) (b : γ), f a b = g a m_1 b) :
forIn' ita b f = forIn' itb b' g
theorem Std.Iterators.Iter.forIn_congr {γ α β : Type w} {m : Type w → Type w'} [Monad m] [Iterator α Id β] [Finite α Id] [IteratorLoop α Id m] {ita itb : Iter β} (w : ita = itb) {b b' : γ} (hb : b = b') {f g : βγm (ForInStep γ)} (h : ∀ (a : β) (b : γ), f a b = g a b) :
forIn ita b f = forIn itb b' g
theorem Std.Iterators.Iter.forIn'_eq_forIn'_toIterM {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type w → Type w''} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] {γ : Type w} {it : Iter β} {init : γ} {f : (out : β) → out itγm (ForInStep γ)} :
forIn' it init f = forIn' it.toIterM init fun (out : β) (h : out it.toIterM) (acc : γ) => f out acc
theorem Std.Iterators.Iter.forIn_eq_forIn_toIterM {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type w → Type w''} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] {γ : Type w} {it : Iter β} {init : γ} {f : βγm (ForInStep γ)} :
forIn it init f = forIn it.toIterM init f
theorem Std.Iterators.Iter.forIn'_eq_match_step {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type x → Type x''} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] {γ : Type x} {it : Iter β} {init : γ} {f : (out : β) → out itγm (ForInStep γ)} :
forIn' it init f = match it.step with | IterStep.yield it' out, h => do let __do_liftf out init match __do_lift with | ForInStep.yield c => forIn' it' c fun (out_1 : β) (h'' : out_1 it') (acc : γ) => f out_1 acc | ForInStep.done c => pure c | IterStep.skip it', h => forIn' it' init fun (out : β) (h' : out it') (acc : γ) => f out acc | IterStep.done, property => pure init
theorem Std.Iterators.Iter.forIn_eq_match_step {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] {γ : Type x} {it : Iter β} {init : γ} {f : βγm (ForInStep γ)} :
forIn it init f = match it.step with | IterStep.yield it' out, property => do let __do_liftf out init match __do_lift with | ForInStep.yield c => forIn it' c f | ForInStep.done c => pure c | IterStep.skip it', property => forIn it' init f | IterStep.done, property => pure init
theorem Std.Iterators.Iter.forIn'_toList {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] [LawfulDeterministicIterator α Id] {γ : Type x} {it : Iter β} {init : γ} {f : (out : β) → out it.toListγm (ForInStep γ)} :
forIn' it.toList init f = forIn' it init fun (out : β) (h : out it) (acc : γ) => f out acc
theorem Std.Iterators.Iter.forIn'_toArray {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] [LawfulDeterministicIterator α Id] {γ : Type x} {it : Iter β} {init : γ} {f : (out : β) → out it.toArrayγm (ForInStep γ)} :
forIn' it.toArray init f = forIn' it init fun (out : β) (h : out it) (acc : γ) => f out acc
theorem Std.Iterators.Iter.forIn'_eq_forIn'_toList {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] [LawfulDeterministicIterator α Id] {γ : Type x} {it : Iter β} {init : γ} {f : (out : β) → out itγm (ForInStep γ)} :
forIn' it init f = forIn' it.toList init fun (out : β) (h : out it.toList) (acc : γ) => f out acc
theorem Std.Iterators.Iter.forIn'_eq_forIn'_toArray {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] [LawfulDeterministicIterator α Id] {γ : Type x} {it : Iter β} {init : γ} {f : (out : β) → out itγm (ForInStep γ)} :
forIn' it init f = forIn' it.toArray init fun (out : β) (h : out it.toArray) (acc : γ) => f out acc
theorem Std.Iterators.Iter.forIn_toList {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {γ : Type x} {it : Iter β} {init : γ} {f : βγm (ForInStep γ)} :
forIn it.toList init f = forIn it init f
theorem Std.Iterators.Iter.forIn_toArray {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {γ : Type x} {it : Iter β} {init : γ} {f : βγm (ForInStep γ)} :
forIn it.toArray init f = forIn it init f
theorem Std.Iterators.Iter.foldM_eq_forIn {α β : Type w} {γ : Type x} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [IteratorLoop α Id m] {f : γβm γ} {init : γ} {it : Iter β} :
foldM f init it = forIn it init fun (x : β) (acc : γ) => ForInStep.yield <$> f acc x
theorem Std.Iterators.Iter.foldM_eq_foldM_toIterM {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type w → Type w''} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] {γ : Type w} {it : Iter β} {init : γ} {f : γβm γ} :
foldM f init it = IterM.foldM f init it.toIterM
theorem Std.Iterators.Iter.forIn_yield_eq_foldM {α β : Type w} {γ δ : Type x} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] {f : βγm δ} {g : βγδγ} {init : γ} {it : Iter β} :
(forIn it init fun (c : β) (b : γ) => (fun (d : δ) => ForInStep.yield (g c b d)) <$> f c b) = foldM (fun (b : γ) (c : β) => g c b <$> f c b) init it
theorem Std.Iterators.Iter.foldM_eq_match_step {α β : Type w} {γ : Type x} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] {f : γβm γ} {init : γ} {it : Iter β} :
foldM f init it = match it.step with | IterStep.yield it' out, property => do let __do_liftf init out foldM f __do_lift it' | IterStep.skip it', property => foldM f init it' | IterStep.done, property => pure init
theorem Std.Iterators.Iter.foldlM_toList {α β : Type w} {γ : Type x} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {f : γβm γ} {init : γ} {it : Iter β} :
List.foldlM f init it.toList = foldM f init it
theorem Std.Iterators.Iter.foldlM_toArray {α β : Type w} {γ : Type x} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {f : γβm γ} {init : γ} {it : Iter β} :
Array.foldlM f init it.toArray = foldM f init it
theorem Std.Iterators.IterM.forIn_eq_foldM {α β : Type w} [Iterator α Id β] [Finite α Id] {m : Type x → Type x'} [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {γ : Type x} {it : Iter β} {init : γ} {f : βγm (ForInStep γ)} :
forIn it init f = ForInStep.value <$> Iter.foldM (fun (c : ForInStep γ) (b : β) => match c with | ForInStep.yield c => f b c | ForInStep.done c => pure (ForInStep.done c)) (ForInStep.yield init) it
theorem Std.Iterators.Iter.fold_eq_forIn {α β : Type w} {γ : Type x} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] {f : γβγ} {init : γ} {it : Iter β} :
fold f init it = (forIn it init fun (x : β) (acc : γ) => pure (ForInStep.yield (f acc x))).run
theorem Std.Iterators.Iter.fold_eq_foldM {α β : Type w} {γ : Type x} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] {f : γβγ} {init : γ} {it : Iter β} :
fold f init it = (foldM (fun (x1 : γ) (x2 : β) => pure (f x1 x2)) init it).run
theorem Std.Iterators.Iter.fold_eq_fold_toIterM {α β γ : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] {f : γβγ} {init : γ} {it : Iter β} :
fold f init it = (IterM.fold f init it.toIterM).run
@[simp]
theorem Std.Iterators.Iter.forIn_pure_yield_eq_fold {α β : Type w} {γ : Type x} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {f : βγγ} {init : γ} {it : Iter β} :
(forIn it init fun (c : β) (b : γ) => pure (ForInStep.yield (f c b))) = pure (fold (fun (b : γ) (c : β) => f c b) init it)
theorem Std.Iterators.Iter.fold_eq_match_step {α β : Type w} {γ : Type x} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {f : γβγ} {init : γ} {it : Iter β} :
fold f init it = match it.step with | IterStep.yield it' out, property => fold f (f init out) it' | IterStep.skip it', property => fold f init it' | IterStep.done, property => init
theorem Std.Iterators.Iter.fold_hom {α β : Type u_1} {γ₁ : Type x₁} {γ₂ : Type x₂} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} (f : γ₁γ₂) {g₁ : γ₁βγ₁} {g₂ : γ₂βγ₂} {init : γ₁} (H : ∀ (x : γ₁) (y : β), g₂ (f x) y = f (g₁ x y)) :
fold g₂ (f init) it = f (fold g₁ init it)
theorem Std.Iterators.Iter.toList_eq_fold {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {it : Iter β} :
it.toList = fold (fun (l : List β) (out : β) => l ++ [out]) [] it
theorem Std.Iterators.Iter.toArray_eq_fold {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {it : Iter β} :
it.toArray = fold (fun (xs : Array β) (out : β) => xs.push out) #[] it
@[simp]
theorem Std.Iterators.Iter.foldl_toList {α β : Type w} {γ : Type x} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {f : γβγ} {init : γ} {it : Iter β} :
List.foldl f init it.toList = fold f init it
@[simp]
theorem Std.Iterators.Iter.foldl_toArray {α β : Type w} {γ : Type x} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {f : γβγ} {init : γ} {it : Iter β} :
Array.foldl f init it.toArray = fold f init it
theorem Std.Iterators.Iter.count_eq_fold {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} :
it.count = fold (fun (acc : Nat) (x : β) => acc + 1) 0 it
theorem Std.Iterators.Iter.count_eq_forIn {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} :
it.count = (forIn it 0 fun (x : β) (acc : Nat) => pure (ForInStep.yield (acc + 1))).run
theorem Std.Iterators.Iter.count_eq_match_step {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} :
it.count = match it.step.val with | IterStep.yield it' out => it'.count + 1 | IterStep.skip it' => it'.count | IterStep.done => 0
@[deprecated Std.Iterators.Iter.size_toArray_eq_count (since := "2025-10-29")]
Equations
Instances For
    @[deprecated Std.Iterators.Iter.length_toList_eq_count (since := "2025-10-29")]
    Equations
    Instances For
      @[deprecated Std.Iterators.Iter.length_toListRev_eq_count (since := "2025-10-29")]
      Equations
      Instances For
        theorem Std.Iterators.Iter.anyM_eq_forIn {α β : Type w} {m : TypeType w'} [Iterator α Id β] [Finite α Id] [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] {it : Iter β} {p : βm Bool} :
        anyM p it = forIn it false fun (x : β) (x_1 : Bool) => do let __do_liftp x if __do_lift = true then pure (ForInStep.done true) else pure (ForInStep.yield false)
        theorem Std.Iterators.Iter.anyM_eq_match_step {α β : Type w} {m : TypeType w'} [Iterator α Id β] [Finite α Id] [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] {it : Iter β} {p : βm Bool} :
        anyM p it = match it.step.val with | IterStep.yield it' x => do let __do_liftp x if __do_lift = true then pure true else anyM p it' | IterStep.skip it' => anyM p it' | IterStep.done => pure false
        theorem Std.Iterators.Iter.anyM_toList {α β : Type w} {m : TypeType w'} [Iterator α Id β] [Finite α Id] [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {it : Iter β} {p : βm Bool} :
        theorem Std.Iterators.Iter.anyM_toArray {α β : Type w} {m : TypeType w'} [Iterator α Id β] [Finite α Id] [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {it : Iter β} {p : βm Bool} :
        theorem Std.Iterators.Iter.any_eq_anyM {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} {p : βBool} :
        any p it = (anyM (fun (x : β) => pure (p x)) it).run
        theorem Std.Iterators.Iter.anyM_pure {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} {p : βBool} :
        anyM (fun (x : β) => pure (p x)) it = pure (any (fun (x : β) => p x) it)
        theorem Std.Iterators.Iter.any_eq_match_step {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} {p : βBool} :
        any p it = match it.step.val with | IterStep.yield it' x => if p x = true then true else any p it' | IterStep.skip it' => any p it' | IterStep.done => false
        theorem Std.Iterators.Iter.any_eq_forIn {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} {p : βBool} :
        any p it = (forIn it false fun (x : β) (x_1 : Bool) => if p x = true then pure (ForInStep.done true) else pure (ForInStep.yield false)).run
        theorem Std.Iterators.Iter.any_toList {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {it : Iter β} {p : βBool} :
        it.toList.any p = any p it
        theorem Std.Iterators.Iter.any_toArray {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {it : Iter β} {p : βBool} :
        it.toArray.any p = any p it
        theorem Std.Iterators.Iter.allM_eq_forIn {α β : Type w} {m : TypeType w'} [Iterator α Id β] [Finite α Id] [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] {it : Iter β} {p : βm Bool} :
        allM p it = forIn it true fun (x : β) (x_1 : Bool) => do let __do_liftp x if __do_lift = true then pure (ForInStep.yield true) else pure (ForInStep.done false)
        theorem Std.Iterators.Iter.allM_eq_match_step {α β : Type w} {m : TypeType w'} [Iterator α Id β] [Finite α Id] [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] {it : Iter β} {p : βm Bool} :
        allM p it = match it.step.val with | IterStep.yield it' x => do let __do_liftp x if __do_lift = true then allM p it' else pure false | IterStep.skip it' => allM p it' | IterStep.done => pure true
        theorem Std.Iterators.Iter.all_eq_allM {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} {p : βBool} :
        all p it = (allM (fun (x : β) => pure (p x)) it).run
        theorem Std.Iterators.Iter.allM_pure {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} {p : βBool} :
        allM (fun (x : β) => pure (p x)) it = pure (all (fun (x : β) => p x) it)
        theorem Std.Iterators.Iter.all_eq_match_step {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} {p : βBool} :
        all p it = match it.step.val with | IterStep.yield it' x => if p x = true then all p it' else false | IterStep.skip it' => all p it' | IterStep.done => true
        theorem Std.Iterators.Iter.all_eq_forIn {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} {p : βBool} :
        all p it = (forIn it true fun (x : β) (x_1 : Bool) => if p x = true then pure (ForInStep.yield true) else pure (ForInStep.done false)).run
        theorem Std.Iterators.Iter.all_toList {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {it : Iter β} {p : βBool} :
        it.toList.all p = all p it
        theorem Std.Iterators.Iter.all_toArray {α β : Type w} [Iterator α Id β] [Finite α Id] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] [IteratorCollect α Id Id] [LawfulIteratorCollect α Id Id] {it : Iter β} {p : βBool} :
        it.toArray.all p = all p it
        theorem Std.Iterators.Iter.allM_eq_not_anyM_not {α β : Type w} {m : TypeType w'} [Iterator α Id β] [Finite α Id] [Monad m] [LawfulMonad m] [IteratorLoop α Id m] [LawfulIteratorLoop α Id m] {it : Iter β} {p : βm Bool} :
        allM p it = (fun (x : Bool) => !x) <$> anyM (fun (x : β) => (fun (x : Bool) => !x) <$> p x) it
        theorem Std.Iterators.Iter.all_eq_not_any_not {m : Type u_1 → Type u_2} {α β : Type w} [Iterator α Id β] [Finite α Id] [Monad m] [LawfulMonad m] [IteratorLoop α Id Id] [LawfulIteratorLoop α Id Id] {it : Iter β} {p : βBool} :
        all p it = !any (fun (x : β) => !p x) it
        theorem Std.Iterators.Iter.findSomeM?_eq_match_step {α β : Type w} {γ : Type x} {m : Type x → Type w'} [Monad m] [Iterator α Id β] [IteratorLoop α Id m] [LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] {it : Iter β} {f : βm (Option γ)} :
        it.findSomeM? f = match it.step.val with | IterStep.yield it' out => do let __do_liftf out match __do_lift with | none => it'.findSomeM? f | some fx => pure (some fx) | IterStep.skip it' => it'.findSomeM? f | IterStep.done => pure none
        theorem Std.Iterators.Iter.findSomeM?_toList {α β : Type w} {γ : Type x} {m : Type x → Type w'} [Monad m] [Iterator α Id β] [IteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] [LawfulIteratorCollect α Id Id] {it : Iter β} {f : βm (Option γ)} :
        theorem Std.Iterators.Iter.findSome?_eq_findSomeM? {α β : Type w} {γ : Type x} [Iterator α Id β] [IteratorLoop α Id Id] [Finite α Id] {it : Iter β} {f : βOption γ} :
        it.findSome? f = (it.findSomeM? fun (x : β) => pure (f x)).run
        theorem Std.Iterators.Iter.findSome?_eq_findSome?_toIterM {α β γ : Type w} [Iterator α Id β] [IteratorLoop α Id Id] [Finite α Id] {it : Iter β} {f : βOption γ} :
        theorem Std.Iterators.Iter.findSome?_eq_match_step {α β : Type w} {γ : Type x} [Iterator α Id β] [IteratorLoop α Id Id] [Finite α Id] [LawfulIteratorLoop α Id Id] {it : Iter β} {f : βOption γ} :
        it.findSome? f = match it.step.val with | IterStep.yield it' out => match f out with | none => it'.findSome? f | some fx => some fx | IterStep.skip it' => it'.findSome? f | IterStep.done => none
        theorem Std.Iterators.Iter.findSomeM?_pure {α β : Type w} {γ : Type x} {m : Type x → Type w'} [Monad m] [Iterator α Id β] [IteratorLoop α Id m] [IteratorLoop α Id Id] [LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] [LawfulIteratorLoop α Id Id] {it : Iter β} {f : βOption γ} :
        (it.findSomeM? fun (x : β) => pure (f x)) = pure (it.findSome? f)
        theorem Std.Iterators.Iter.findM?_eq_findSomeM? {α β : Type w} {m : Type w → Type w'} [Monad m] [Iterator α Id β] [IteratorLoop α Id m] [Finite α Id] {it : Iter β} {f : βm (ULift Bool)} :
        it.findM? f = it.findSomeM? fun (x : β) => do let __do_liftf x pure (if __do_lift.down = true then some x else none)
        theorem Std.Iterators.Iter.findM?_eq_match_step {α β : Type w} {m : Type w → Type w'} [Monad m] [Iterator α Id β] [IteratorLoop α Id m] [LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] {it : Iter β} {f : βm (ULift Bool)} :
        it.findM? f = match it.step.val with | IterStep.yield it' out => do let __do_liftf out if __do_lift.down = true then pure (some out) else it'.findM? f | IterStep.skip it' => it'.findM? f | IterStep.done => pure none
        theorem Std.Iterators.Iter.findM?_toList {α β : Type} {m : TypeType w'} [Monad m] [Iterator α Id β] [IteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] [LawfulIteratorCollect α Id Id] {it : Iter β} {f : βm Bool} :
        List.findM? f it.toList = it.findM? fun (x : β) => ULift.up <$> f x
        theorem Std.Iterators.Iter.findM?_eq_findM?_toList {α β : Type} {m : TypeType w'} [Monad m] [Iterator α Id β] [IteratorLoop α Id m] [IteratorCollect α Id Id] [LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] [LawfulIteratorCollect α Id Id] {it : Iter β} {f : βm (ULift Bool)} :
        it.findM? f = List.findM? (fun (x : β) => ULift.down <$> f x) it.toList
        theorem Std.Iterators.Iter.find?_eq_findM? {α β : Type w} [Iterator α Id β] [IteratorLoop α Id Id] [Finite α Id] {it : Iter β} {f : βBool} :
        it.find? f = (it.findM? fun (x : β) => pure { down := f x }).run
        theorem Std.Iterators.Iter.find?_eq_find?_toIterM {α β : Type w} [Iterator α Id β] [IteratorLoop α Id Id] [Finite α Id] {it : Iter β} {f : βBool} :
        it.find? f = (it.toIterM.find? f).run
        theorem Std.Iterators.Iter.find?_eq_findSome? {α β : Type w} [Iterator α Id β] [IteratorLoop α Id Id] [Finite α Id] {it : Iter β} {f : βBool} :
        it.find? f = it.findSome? fun (x : β) => if f x = true then some x else none
        theorem Std.Iterators.Iter.find?_eq_match_step {α β : Type w} [Iterator α Id β] [IteratorLoop α Id Id] [Finite α Id] [LawfulIteratorLoop α Id Id] {it : Iter β} {f : βBool} :
        it.find? f = match it.step.val with | IterStep.yield it' out => if f out = true then some out else it'.find? f | IterStep.skip it' => it'.find? f | IterStep.done => none
        theorem Std.Iterators.Iter.findM?_pure {α β : Type w} {m : Type w → Type w'} [Monad m] [Iterator α Id β] [IteratorLoop α Id m] [IteratorLoop α Id Id] [LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] [LawfulIteratorLoop α Id Id] {it : Iter β} {f : βULift Bool} :
        (it.findM? fun (x : β) => pure (f x)) = pure (it.find? fun (x : β) => (f x).down)