Skip to content

Commit 71f1a6c

Browse files
authored
feat: Iterator find? consumer and variants (#10769)
This PR adds a `find?` consumer in analogy to `List.find?` and variants thereof.
1 parent 0d5869b commit 71f1a6c

File tree

5 files changed

+414
-0
lines changed

5 files changed

+414
-0
lines changed

src/Init/Data/Iterators/Consumers/Loop.lean

Lines changed: 53 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -202,6 +202,59 @@ def Iter.all {α β : Type w}
202202
(p : β → Bool) (it : Iter (α := α) β) : Bool :=
203203
(it.allM (fun x => pure (f := Id) (p x))).run
204204

205+
@[inline]
206+
def Iter.findSomeM? {α β : Type w} {γ : Type x} {m : Type x → Type w'} [Monad m] [Iterator α Id β]
207+
[IteratorLoop α Id m] [Finite α Id] (it : Iter (α := α) β) (f : β → m (Option γ)) :
208+
m (Option γ) :=
209+
ForIn.forIn it none (fun x _ => do
210+
match ← f x with
211+
| none => return .yield none
212+
| some fx => return .done (some fx))
213+
214+
@[inline]
215+
def Iter.Partial.findSomeM? {α β : Type w} {γ : Type x} {m : Type x → Type w'} [Monad m]
216+
[Iterator α Id β] [IteratorLoopPartial α Id m] (it : Iter.Partial (α := α) β)
217+
(f : β → m (Option γ)) :
218+
m (Option γ) :=
219+
ForIn.forIn it none (fun x _ => do
220+
match ← f x with
221+
| none => return .yield none
222+
| some fx => return .done (some fx))
223+
224+
@[inline]
225+
def Iter.findSome? {α β : Type w} {γ : Type x} [Iterator α Id β]
226+
[IteratorLoop α Id Id] [Finite α Id] (it : Iter (α := α) β) (f : β → Option γ) :
227+
Option γ :=
228+
Id.run (it.findSomeM? (pure <| f ·))
229+
230+
@[inline]
231+
def Iter.Partial.findSome? {α β : Type w} {γ : Type x} [Iterator α Id β]
232+
[IteratorLoopPartial α Id Id] (it : Iter.Partial (α := α) β) (f : β → Option γ) :
233+
Option γ :=
234+
Id.run (it.findSomeM? (pure <| f ·))
235+
236+
@[inline]
237+
def Iter.findM? {α β : Type w} {m : Type w → Type w'} [Monad m] [Iterator α Id β]
238+
[IteratorLoop α Id m] [Finite α Id] (it : Iter (α := α) β) (f : β → m (ULift Bool)) :
239+
m (Option β) :=
240+
it.findSomeM? (fun x => return if (← f x).down then some x else none)
241+
242+
@[inline]
243+
def Iter.Partial.findM? {α β : Type w} {m : Type w → Type w'} [Monad m] [Iterator α Id β]
244+
[IteratorLoopPartial α Id m] (it : Iter.Partial (α := α) β) (f : β → m (ULift Bool)) :
245+
m (Option β) :=
246+
it.findSomeM? (fun x => return if (← f x).down then some x else none)
247+
248+
@[inline]
249+
def Iter.find? {α β : Type w} [Iterator α Id β] [IteratorLoop α Id Id]
250+
[Finite α Id] (it : Iter (α := α) β) (f : β → Bool) : Option β :=
251+
Id.run (it.findM? (pure <| .up <| f ·))
252+
253+
@[inline]
254+
def Iter.Partial.find? {α β : Type w} [Iterator α Id β] [IteratorLoopPartial α Id Id]
255+
(it : Iter.Partial (α := α) β) (f : β → Bool) : Option β :=
256+
Id.run (it.findM? (pure <| .up <| f ·))
257+
205258
@[always_inline, inline, expose, inherit_doc IterM.size]
206259
def Iter.size {α : Type w} {β : Type w} [Iterator α Id β] [IteratorSize α Id]
207260
(it : Iter (α := α) β) : Nat :=

src/Init/Data/Iterators/Consumers/Monadic/Loop.lean

Lines changed: 54 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -579,6 +579,60 @@ def IterM.Partial.all {α β : Type w} {m : Type w → Type w'} [Monad m]
579579
(p : β → Bool) (it : IterM.Partial (α := α) m β) : m (ULift Bool) := do
580580
it.allM (fun x => pure (.up (p x)))
581581

582+
@[inline]
583+
def IterM.findSomeM? {α β γ : Type w} {m : Type w → Type w'} [Monad m] [Iterator α m β]
584+
[IteratorLoop α m m] [Finite α m] (it : IterM (α := α) m β) (f : β → m (Option γ)) :
585+
m (Option γ) :=
586+
ForIn.forIn it none (fun x _ => do
587+
match ← f x with
588+
| none => return .yield none
589+
| some fx => return .done (some fx))
590+
591+
@[inline]
592+
def IterM.Partial.findSomeM? {α β γ : Type w} {m : Type w → Type w'} [Monad m] [Iterator α m β]
593+
[IteratorLoopPartial α m m] (it : IterM.Partial (α := α) m β) (f : β → m (Option γ)) :
594+
m (Option γ) :=
595+
ForIn.forIn it none (fun x _ => do
596+
match ← f x with
597+
| none => return .yield none
598+
| some fx => return .done (some fx))
599+
600+
@[inline]
601+
def IterM.findSome? {α β γ : Type w} {m : Type w → Type w'} [Monad m] [Iterator α m β]
602+
[IteratorLoop α m m] [Finite α m] (it : IterM (α := α) m β) (f : β → Option γ) :
603+
m (Option γ) :=
604+
it.findSomeM? (pure <| f ·)
605+
606+
@[inline]
607+
def IterM.Partial.findSome? {α β γ : Type w} {m : Type w → Type w'} [Monad m] [Iterator α m β]
608+
[IteratorLoopPartial α m m] (it : IterM.Partial (α := α) m β) (f : β → Option γ) :
609+
m (Option γ) :=
610+
it.findSomeM? (pure <| f ·)
611+
612+
@[inline]
613+
def IterM.findM? {α β : Type w} {m : Type w → Type w'} [Monad m] [Iterator α m β]
614+
[IteratorLoop α m m] [Finite α m] (it : IterM (α := α) m β) (f : β → m (ULift Bool)) :
615+
m (Option β) :=
616+
it.findSomeM? (fun x => return if (← f x).down then some x else none)
617+
618+
@[inline]
619+
def IterM.Partial.findM? {α β : Type w} {m : Type w → Type w'} [Monad m] [Iterator α m β]
620+
[IteratorLoopPartial α m m] (it : IterM.Partial (α := α) m β) (f : β → m (ULift Bool)) :
621+
m (Option β) :=
622+
it.findSomeM? (fun x => return if (← f x).down then some x else none)
623+
624+
@[inline]
625+
def IterM.find? {α β : Type w} {m : Type w → Type w'} [Monad m] [Iterator α m β]
626+
[IteratorLoop α m m] [Finite α m] (it : IterM (α := α) m β) (f : β → Bool) :
627+
m (Option β) :=
628+
it.findM? (pure <| .up <| f ·)
629+
630+
@[inline]
631+
def IterM.Partial.find? {α β : Type w} {m : Type w → Type w'} [Monad m] [Iterator α m β]
632+
[IteratorLoopPartial α m m] (it : IterM.Partial (α := α) m β) (f : β → Bool) :
633+
m (Option β) :=
634+
it.findM? (pure <| .up <| f ·)
635+
582636
section Size
583637

584638
/--

src/Init/Data/Iterators/Lemmas/Consumers/Loop.lean

Lines changed: 166 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -726,4 +726,170 @@ theorem Iter.all_eq_not_any_not {α β : Type w} [Iterator α Id β]
726726
· simp [ihs ‹_›]
727727
· simp
728728

729+
theorem Iter.findSomeM?_eq_match_step {α β : Type w} {γ : Type x} {m : Type x → Type w'} [Monad m]
730+
[Iterator α Id β] [IteratorLoop α Id m] [LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m]
731+
{it : Iter (α := α) β} {f : β → m (Option γ)} :
732+
it.findSomeM? f = (do
733+
match it.step.val with
734+
| .yield it' out =>
735+
match ← f out with
736+
| none => it'.findSomeM? f
737+
| some fx => return (some fx)
738+
| .skip it' => it'.findSomeM? f
739+
| .done => return none) := by
740+
rw [findSomeM?, forIn_eq_match_step]
741+
cases it.step using PlausibleIterStep.casesOn
742+
· simp only [bind_assoc]
743+
apply bind_congr; intro fx
744+
split <;> simp [findSomeM?]
745+
· simp [findSomeM?]
746+
· simp
747+
748+
theorem Iter.findSomeM?_toList {α β : Type w} {γ : Type x} {m : Type x → Type w'} [Monad m]
749+
[Iterator α Id β] [IteratorLoop α Id m] [IteratorCollect α Id Id]
750+
[LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] [LawfulIteratorCollect α Id Id]
751+
{it : Iter (α := α) β} {f : β → m (Option γ)} :
752+
it.toList.findSomeM? f = it.findSomeM? f := by
753+
induction it using Iter.inductSteps with | step it ihy ihs
754+
rw [it.findSomeM?_eq_match_step, it.toList_eq_match_step]
755+
cases it.step using PlausibleIterStep.casesOn
756+
· simp only [List.findSomeM?_cons]
757+
apply bind_congr; intro fx
758+
split <;> simp [ihy ‹_›]
759+
· simp [ihs ‹_›]
760+
· simp
761+
762+
theorem Iter.findSome?_eq_findSomeM? {α β : Type w} {γ : Type x}
763+
[Iterator α Id β] [IteratorLoop α Id Id] [Finite α Id]
764+
{it : Iter (α := α) β} {f : β → Option γ} :
765+
it.findSome? f = Id.run (it.findSomeM? (pure <| f ·)) :=
766+
(rfl)
767+
768+
theorem Iter.findSome?_eq_findSome?_toIterM {α β γ : Type w}
769+
[Iterator α Id β] [IteratorLoop α Id Id.{w}] [Finite α Id]
770+
{it : Iter (α := α) β} {f : β → Option γ} :
771+
it.findSome? f = (it.toIterM.findSome? f).run :=
772+
(rfl)
773+
774+
theorem Iter.findSome?_eq_match_step {α β : Type w} {γ : Type x}
775+
[Iterator α Id β] [IteratorLoop α Id Id] [Finite α Id]
776+
[LawfulIteratorLoop α Id Id] {it : Iter (α := α) β} {f : β → Option γ} :
777+
it.findSome? f = (match it.step.val with
778+
| .yield it' out =>
779+
match f out with
780+
| none => it'.findSome? f
781+
| some fx => some fx
782+
| .skip it' => it'.findSome? f
783+
| .done => none) := by
784+
rw [findSome?_eq_findSomeM?, findSomeM?_eq_match_step]
785+
split
786+
· simp only [pure_bind, findSome?_eq_findSomeM?]
787+
split <;> simp
788+
· simp [findSome?_eq_findSomeM?]
789+
· simp
790+
791+
theorem Iter.findSome?_toList {α β : Type w} {γ : Type x}
792+
[Iterator α Id β] [IteratorLoop α Id Id] [IteratorCollect α Id Id]
793+
[Finite α Id] [LawfulIteratorLoop α Id Id] [LawfulIteratorCollect α Id Id]
794+
{it : Iter (α := α) β} {f : β → Option γ} :
795+
it.toList.findSome? f = it.findSome? f := by
796+
simp [findSome?_eq_findSomeM?, List.findSome?_eq_findSomeM?, findSomeM?_toList]
797+
798+
theorem Iter.findSomeM?_pure {α β : Type w} {γ : Type x} {m : Type x → Type w'} [Monad m]
799+
[Iterator α Id β] [IteratorLoop α Id m] [IteratorLoop α Id Id]
800+
[LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] [LawfulIteratorLoop α Id Id]
801+
{it : Iter (α := α) β} {f : β → Option γ} :
802+
it.findSomeM? (pure <| f ·) = pure (f := m) (it.findSome? f) := by
803+
letI : IteratorCollect α Id Id := .defaultImplementation
804+
simp [← findSomeM?_toList, ← findSome?_toList, List.findSomeM?_pure]
805+
806+
theorem Iter.findM?_eq_findSomeM? {α β : Type w} {m : Type w → Type w'} [Monad m]
807+
[Iterator α Id β] [IteratorLoop α Id m] [Finite α Id]
808+
{it : Iter (α := α) β} {f : β → m (ULift Bool)} :
809+
it.findM? f = it.findSomeM? (fun x => return if (← f x).down then some x else none) :=
810+
(rfl)
811+
812+
theorem Iter.findM?_eq_match_step {α β : Type w} {m : Type w → Type w'} [Monad m]
813+
[Iterator α Id β] [IteratorLoop α Id m] [LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m]
814+
{it : Iter (α := α) β} {f : β → m (ULift Bool)} :
815+
it.findM? f = (do
816+
match it.step.val with
817+
| .yield it' out =>
818+
if (← f out).down then return (some out) else it'.findM? f
819+
| .skip it' => it'.findM? f
820+
| .done => return none) := by
821+
rw [findM?_eq_findSomeM?, findSomeM?_eq_match_step]
822+
split
823+
· simp only [bind_assoc]
824+
apply bind_congr; intro fx
825+
split <;> simp [findM?_eq_findSomeM?]
826+
· simp [findM?_eq_findSomeM?]
827+
· simp
828+
829+
theorem Iter.findM?_toList {α β : Type} {m : TypeType w'} [Monad m]
830+
[Iterator α Id β] [IteratorLoop α Id m] [IteratorCollect α Id Id]
831+
[LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] [LawfulIteratorCollect α Id Id]
832+
{it : Iter (α := α) β} {f : β → m Bool} :
833+
it.toList.findM? f = it.findM? (.up <$> f ·) := by
834+
simp [findM?_eq_findSomeM?, List.findM?_eq_findSomeM?, findSomeM?_toList]
835+
836+
theorem Iter.findM?_eq_findM?_toList {α β : Type} {m : TypeType w'} [Monad m]
837+
[Iterator α Id β] [IteratorLoop α Id m] [IteratorCollect α Id Id]
838+
[LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] [LawfulIteratorCollect α Id Id]
839+
{it : Iter (α := α) β} {f : β → m (ULift Bool)} :
840+
it.findM? f = it.toList.findM? (ULift.down <$> f ·) := by
841+
simp [findM?_toList]
842+
843+
theorem Iter.find?_eq_findM? {α β : Type w} [Iterator α Id β]
844+
[IteratorLoop α Id Id] [Finite α Id] {it : Iter (α := α) β} {f : β → Bool} :
845+
it.find? f = Id.run (it.findM? (pure <| .up <| f ·)) :=
846+
(rfl)
847+
848+
theorem Iter.find?_eq_find?_toIterM {α β : Type w} [Iterator α Id β]
849+
[IteratorLoop α Id Id] [Finite α Id] {it : Iter (α := α) β} {f : β → Bool} :
850+
it.find? f = (it.toIterM.find? f).run :=
851+
(rfl)
852+
853+
theorem Iter.find?_eq_findSome? {α β : Type w} [Iterator α Id β]
854+
[IteratorLoop α Id Id] [Finite α Id] {it : Iter (α := α) β} {f : β → Bool} :
855+
it.find? f = it.findSome? (fun x => if f x then some x else none) := by
856+
simp [find?_eq_findM?, findSome?_eq_findSomeM?, findM?_eq_findSomeM?]
857+
858+
theorem Iter.find?_eq_match_step {α β : Type w}
859+
[Iterator α Id β] [IteratorLoop α Id Id] [Finite α Id] [LawfulIteratorLoop α Id Id]
860+
{it : Iter (α := α) β} {f : β → Bool} :
861+
it.find? f = (match it.step.val with
862+
| .yield it' out =>
863+
if f out then some out else it'.find? f
864+
| .skip it' => it'.find? f
865+
| .done => none) := by
866+
rw [find?_eq_findM?, findM?_eq_match_step]
867+
split
868+
· simp only [pure_bind]
869+
split <;> simp [find?_eq_findM?]
870+
· simp [find?_eq_findM?]
871+
· simp
872+
873+
theorem Iter.find?_toList {α β : Type w}
874+
[Iterator α Id β] [IteratorLoop α Id Id] [IteratorCollect α Id Id]
875+
[Finite α Id] [LawfulIteratorLoop α Id Id] [LawfulIteratorCollect α Id Id]
876+
{it : Iter (α := α) β} {f : β → Bool} :
877+
it.toList.find? f = it.find? f := by
878+
simp [find?_eq_findSome?, List.find?_eq_findSome?_guard, findSome?_toList, Option.guard_def]
879+
880+
theorem Iter.findM?_pure {α β : Type w} {m : Type w → Type w'} [Monad m]
881+
[Iterator α Id β] [IteratorLoop α Id m] [IteratorLoop α Id Id]
882+
[LawfulMonad m] [Finite α Id] [LawfulIteratorLoop α Id m] [LawfulIteratorLoop α Id Id]
883+
{it : Iter (α := α) β} {f : β → ULift Bool} :
884+
it.findM? (pure (f := m) <| f ·) = pure (f := m) (it.find? (ULift.down <| f ·)) := by
885+
induction it using Iter.inductSteps with | step it ihy ihs
886+
rw [findM?_eq_match_step, find?_eq_match_step]
887+
cases it.step using PlausibleIterStep.casesOn
888+
· simp only [pure_bind]
889+
split
890+
· simp
891+
· simp [ihy ‹_›]
892+
· simp [ihs ‹_›]
893+
· simp
894+
729895
end Std.Iterators

0 commit comments

Comments
 (0)