Skip to content

Commit 70b4b2b

Browse files
authored
feat: polymorphic ranges (#8784)
This PR introduces ranges that are polymorphic, in contrast to the existing `Std.Range` which only supports natural numbers. Breakdown of core changes: * `Lean.Parser.Basic`: Modified the number parser (`Lean.Parser.Basic`) so that it will only consider a *single* dot to be part of a decimal number. `1..` will no longer be parsed as `1.` followed by `.`, but as `1` followed by `..`. * The test `ellipsisProjIssue` ensures that `#check Nat.add ...succ` produces a syntax error. After introducing the new range notation (see below), it returns a different (less nice) error message. I updated the test to reflect the new error message. (The error message will become nicer as soon as a delaborator for the ranges is implemented. This is out of scope for this PR.) Breakdown of standard library changes: Modified modules: `Init.Data.Range.Polymorphic` (added), `Init.Data.Iterators`, `Std.Data.Iterators` * Introduced the type `Std.PRange` that is parameterized over the type in which the range operates and the shapes of the lower and upper bound. * Introduced a new notation for ranges. Examples for this notation are: `1...*`, `1...=3`, `1...<3`, `1<...=2`, `*...=3`. * Defined lots of typeclasses for different capabilities of ranges, depending on their shape and underlying type. * Introduced `Iter(M).size`. * Introduced the `Iter(M).stepSize n` combinator, which iterates over an iterator with the given step size `n`. It will drop `n - 1` values between every value it emits. * Replaced `LawfulPureIterator` with a new and better typeclass `LawfulDeterministicIterator`. * Simplified some lemma statements in the iterator library such as `IterM.toList_eq_match`, which unnecessarily matched over a `Subtype`, hindering rewrites due to type dependencies. Reasons for the concrete choice of notation: * `lean4-cli` uses `...`-based notation for the `Cmd` notation and it clashes with `...a` range notation. * test `2461` fails when using two-dot-based notation because of the existing `{ a.. }` notation.
1 parent 3695059 commit 70b4b2b

File tree

33 files changed

+2551
-85
lines changed

33 files changed

+2551
-85
lines changed

src/Init/Data.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -47,3 +47,4 @@ import Init.Data.Function
4747
import Init.Data.RArray
4848
import Init.Data.Vector
4949
import Init.Data.Iterators
50+
import Init.Data.Range.Polymorphic

src/Init/Data/Iterators/Basic.lean

Lines changed: 95 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -354,7 +354,7 @@ Makes a single step with the given iterator `it`, potentially emitting a value a
354354
succeeding iterator. If this function is used recursively, termination can sometimes be proved with
355355
the termination measures `it.finitelyManySteps` and `it.finitelyManySkips`.
356356
-/
357-
@[always_inline, inline]
357+
@[always_inline, inline, expose]
358358
def IterM.step {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
359359
(it : IterM (α := α) m β) : m it.Step :=
360360
Iterator.step it
@@ -383,6 +383,24 @@ inductive IterM.IsPlausibleIndirectOutput {α β : Type w} {m : Type w → Type
383383
| indirect {it it' : IterM (α := α) m β} {out : β} : it'.IsPlausibleSuccessorOf it →
384384
it'.IsPlausibleIndirectOutput out → it.IsPlausibleIndirectOutput out
385385

386+
/--
387+
Asserts that an iterator `it'` could plausibly produce `it'` as a successor iterator after
388+
finitely many steps. This relation is reflexive.
389+
-/
390+
inductive IterM.IsPlausibleIndirectSuccessorOf {α β : Type w} {m : Type w → Type w'}
391+
[Iterator α m β] : IterM (α := α) m β → IterM (α := α) m β → Prop where
392+
| refl (it : IterM (α := α) m β) : it.IsPlausibleIndirectSuccessorOf it
393+
| cons_right {it'' it' it : IterM (α := α) m β} (h' : it''.IsPlausibleIndirectSuccessorOf it')
394+
(h : it'.IsPlausibleSuccessorOf it) : it''.IsPlausibleIndirectSuccessorOf it
395+
396+
theorem IterM.IsPlausibleIndirectSuccessorOf.trans {α β : Type w} {m : Type w → Type w'}
397+
[Iterator α m β] {it'' it' it : IterM (α := α) m β}
398+
(h' : it''.IsPlausibleIndirectSuccessorOf it') (h : it'.IsPlausibleIndirectSuccessorOf it) :
399+
it''.IsPlausibleIndirectSuccessorOf it := by
400+
induction h
401+
case refl => exact h'
402+
case cons_right ih => exact IsPlausibleIndirectSuccessorOf.cons_right ih ‹_›
403+
386404
/--
387405
The type of the step object returned by `Iter.step`, containing an `IterStep`
388406
and a proof that this is a plausible step for the given iterator.
@@ -431,6 +449,16 @@ def Iter.IsPlausibleOutput {α : Type w} {β : Type w} [Iterator α Id β]
431449
(it : Iter (α := α) β) (out : β) : Prop :=
432450
it.toIterM.IsPlausibleOutput out
433451

452+
theorem Iter.isPlausibleOutput_iff_exists {α : Type w} {β : Type w} [Iterator α Id β]
453+
{it : Iter (α := α) β} {out : β} :
454+
it.IsPlausibleOutput out ↔ ∃ it', it.IsPlausibleStep (.yield it' out) := by
455+
simp only [IsPlausibleOutput, IterM.IsPlausibleOutput]
456+
constructor
457+
· rintro ⟨it', h⟩
458+
exact ⟨it'.toIter, h⟩
459+
· rintro ⟨it', h⟩
460+
exact ⟨it'.toIterM, h⟩
461+
434462
/--
435463
Asserts that a certain iterator `it'` could plausibly be the directly succeeding iterator of another
436464
given iterator `it`.
@@ -440,6 +468,18 @@ def Iter.IsPlausibleSuccessorOf {α : Type w} {β : Type w} [Iterator α Id β]
440468
(it' it : Iter (α := α) β) : Prop :=
441469
it'.toIterM.IsPlausibleSuccessorOf it.toIterM
442470

471+
theorem Iter.isPlausibleSuccessorOf_iff_exists {α : Type w} {β : Type w} [Iterator α Id β]
472+
{it' it : Iter (α := α) β} :
473+
it'.IsPlausibleSuccessorOf it ↔ ∃ step, step.successor = some it' ∧ it.IsPlausibleStep step := by
474+
simp only [IsPlausibleSuccessorOf, IterM.IsPlausibleSuccessorOf]
475+
constructor
476+
· rintro ⟨step, h₁, h₂⟩
477+
exact ⟨step.mapIterator IterM.toIter,
478+
by cases step <;> simp_all [IterStep.successor, Iter.IsPlausibleStep]⟩
479+
· rintro ⟨step, h₁, h₂⟩
480+
exact ⟨step.mapIterator Iter.toIterM,
481+
by cases step <;> simp_all [IterStep.successor, Iter.IsPlausibleStep]⟩
482+
443483
/--
444484
Asserts that a certain iterator `it` could plausibly yield the value `out` after an arbitrary
445485
number of steps.
@@ -472,6 +512,45 @@ theorem Iter.isPlausibleIndirectOutput_iff_isPlausibleIndirectOutput_toIterM {α
472512
replace h : it'.toIter.IsPlausibleSuccessorOf it.toIter := h
473513
exact .indirect (α := α) h ih
474514

515+
/--
516+
Asserts that an iterator `it'` could plausibly produce `it'` as a successor iterator after
517+
finitely many steps. This relation is reflexive.
518+
-/
519+
inductive Iter.IsPlausibleIndirectSuccessorOf {α : Type w} {β : Type w} [Iterator α Id β] :
520+
Iter (α := α) β → Iter (α := α) β → Prop where
521+
| refl (it : Iter (α := α) β) : IsPlausibleIndirectSuccessorOf it it
522+
| cons_right {it'' it' it : Iter (α := α) β} (h' : it''.IsPlausibleIndirectSuccessorOf it')
523+
(h : it'.IsPlausibleSuccessorOf it) : it''.IsPlausibleIndirectSuccessorOf it
524+
525+
theorem Iter.isPlausibleIndirectSuccessor_iff_isPlausibleIndirectSuccessor_toIterM {α β : Type w}
526+
[Iterator α Id β] {it' it : Iter (α := α) β} :
527+
it'.IsPlausibleIndirectSuccessorOf it ↔ it'.toIterM.IsPlausibleIndirectSuccessorOf it.toIterM := by
528+
constructor
529+
· intro h
530+
induction h with
531+
| refl => exact .refl _
532+
| cons_right _ h ih => exact .cons_right ih h
533+
· intro h
534+
rw [← Iter.toIter_toIterM (it := it), ← Iter.toIter_toIterM (it := it')]
535+
generalize it.toIterM = it at ⊢ h
536+
induction h with
537+
| refl => exact .refl _
538+
| cons_right _ h ih => exact .cons_right ih h
539+
540+
theorem Iter.IsPlausibleIndirectSuccessorOf.trans {α : Type w} {β : Type w} [Iterator α Id β]
541+
{it'' it' it : Iter (α := α) β} (h' : it''.IsPlausibleIndirectSuccessorOf it')
542+
(h : it'.IsPlausibleIndirectSuccessorOf it) : it''.IsPlausibleIndirectSuccessorOf it := by
543+
induction h
544+
case refl => exact h'
545+
case cons_right ih => exact IsPlausibleIndirectSuccessorOf.cons_right ih ‹_›
546+
547+
theorem Iter.IsPlausibleIndirectOutput.trans {α : Type w} {β : Type w} [Iterator α Id β]
548+
{it' it : Iter (α := α) β} {out : β} (h : it'.IsPlausibleIndirectSuccessorOf it)
549+
(h' : it'.IsPlausibleIndirectOutput out) : it.IsPlausibleIndirectOutput out := by
550+
induction h
551+
case refl => exact h'
552+
case cons_right ih => exact IsPlausibleIndirectOutput.indirect ‹_› ih
553+
475554
/--
476555
Asserts that a certain iterator `it'` could plausibly be the directly succeeding iterator of another
477556
given iterator `it` while no value is emitted (see `IterStep.skip`).
@@ -687,6 +766,21 @@ instance [Iterator α m β] [Finite α m] : Productive α m where
687766

688767
end Productive
689768

769+
/--
770+
This typeclass characterizes iterators that have deterministic return values. This typeclass does
771+
*not* guarantee that there are no monadic side effects such as exceptions.
772+
773+
General monadic iterators can be nondeterministic, so that `it.IsPlausibleStep step` will be true
774+
for no or more than one choice of `step`. This typeclass ensures that there is exactly one such
775+
choice.
776+
777+
This is an experimental instance and it should not be explicitly used downstream of the standard
778+
library.
779+
-/
780+
class LawfulDeterministicIterator (α : Type w) (m : Type w → Type w') [Iterator α m β]
781+
where
782+
isPlausibleStep_eq_eq : ∀ it : IterM (α := α) m β, ∃ step, it.IsPlausibleStep = (· = step)
783+
690784
end Iterators
691785

692786
export Iterators (Iter IterM)

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

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -7,6 +7,8 @@ module
77

88
prelude
99
import Init.Data.Iterators.Consumers.Partial
10+
import Init.Data.Iterators.Consumers.Loop
11+
import Init.Data.Iterators.Consumers.Monadic.Access
1012

1113
namespace Std.Iterators
1214

@@ -51,4 +53,12 @@ partial def Iter.Partial.atIdxSlow? {α β} [Iterator α Id β] [Monad Id]
5153
| .skip it' _ => (⟨it'⟩ : Iter.Partial (α := α) β).atIdxSlow? n
5254
| .done _ => none
5355

56+
@[always_inline, inline, inherit_doc IterM.atIdx?]
57+
def Iter.atIdx? {α β} [Iterator α Id β] [Productive α Id] [IteratorAccess α Id]
58+
(n : Nat) (it : Iter (α := α) β) : Option β :=
59+
match (IteratorAccess.nextAtIdx? it.toIterM n).run.val with
60+
| .yield _ out => some out
61+
| .skip _ => none
62+
| .done => none
63+
5464
end Std.Iterators

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

Lines changed: 0 additions & 14 deletions
Original file line numberDiff line numberDiff line change
@@ -56,18 +56,4 @@ def Iter.Partial.toList {α : Type w} {β : Type w}
5656
[Iterator α Id β] [IteratorCollectPartial α Id Id] (it : Iter.Partial (α := α) β) : List β :=
5757
it.it.toIterM.allowNontermination.toList.run
5858

59-
/--
60-
This class charaterizes how the plausibility behavior (`IsPlausibleStep`) and the actual iteration
61-
behavior (`it.step`) should relate to each other for pure iterators. Intuitively, a step should
62-
only be plausible if it is possible. For simplicity's sake, the actual definition is weaker but
63-
presupposes that the iterator is finite.
64-
65-
This is an experimental instance and it should not be explicitly used downstream of the standard
66-
library.
67-
-/
68-
class LawfulPureIterator (α : Type w) [Iterator α Id β]
69-
[Finite α Id] [IteratorCollect α Id Id] where
70-
mem_toList_iff_isPlausibleIndirectOutput {it : Iter (α := α) β} {out : β} :
71-
out ∈ it.toList ↔ it.IsPlausibleIndirectOutput out
72-
7359
end Std.Iterators

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

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Paul Reichert
66
module
77

88
prelude
9+
import Init.Data.Iterators.Consumers.Collect
910
import Init.Data.Iterators.Consumers.Monadic.Loop
1011
import Init.Data.Iterators.Consumers.Partial
1112

@@ -29,6 +30,7 @@ A `ForIn'` instance for iterators. Its generic membership relation is not easy t
2930
so this is not marked as `instance`. This way, more convenient instances can be built on top of it
3031
or future library improvements will make it more comfortable.
3132
-/
33+
@[always_inline, inline]
3234
def Iter.instForIn' {α : Type w} {β : Type w} {n : Type w → Type w'} [Monad n]
3335
[Iterator α Id β] [Finite α Id] [IteratorLoop α Id n] :
3436
ForIn' n (Iter (α := α) β) β ⟨fun it out => it.IsPlausibleIndirectOutput out⟩ where
@@ -135,4 +137,18 @@ def Iter.Partial.size {α : Type w} {β : Type w} [Iterator α Id β] [IteratorS
135137
(it : Iter (α := α) β) : Nat :=
136138
(IteratorSizePartial.size it.toIterM).run.down
137139

140+
/--
141+
`LawfulIteratorSize α m` ensures that the `size` function of an iterator behaves as if it
142+
iterated over the whole iterator, counting its elements and causing all the monadic side effects
143+
of the iterations. This is a fairly strong condition for monadic iterators and it will be false
144+
for many efficient implementations of `size` that compute the size without actually iterating.
145+
146+
This class is experimental and users of the iterator API should not explicitly depend on it.
147+
-/
148+
class LawfulIteratorSize (α : Type w) {β : Type w} [Iterator α Id β] [Finite α Id]
149+
[IteratorSize α Id] where
150+
size_eq_size_toArray {it : Iter (α := α) β} : it.size =
151+
haveI : IteratorCollect α Id Id := .defaultImplementation
152+
it.toArray.size
153+
138154
end Std.Iterators

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

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Paul Reichert
66
module
77

88
prelude
9+
import Init.Data.Iterators.Consumers.Monadic.Access
910
import Init.Data.Iterators.Consumers.Monadic.Collect
1011
import Init.Data.Iterators.Consumers.Monadic.Loop
1112
import Init.Data.Iterators.Consumers.Monadic.Partial
Lines changed: 95 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,95 @@
1+
/-
2+
Copyright (c) 2025 Lean FRO, LLC. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Paul Reichert
5+
-/
6+
module
7+
8+
prelude
9+
import Init.Data.Iterators.Basic
10+
11+
namespace Std.Iterators
12+
13+
/--
14+
`it.IsPlausibleNthOutputStep n step` is the proposition that according to the
15+
`IsPlausibleStep` relation, it is plausible that `step` returns the step in which the `n`-th value
16+
of `it` is emitted, or `.done` if `it` can plausibly terminate before emitting `n` values.
17+
-/
18+
inductive IterM.IsPlausibleNthOutputStep {α β : Type w} {m : Type w → Type w'} [Iterator α m β] :
19+
Nat → IterM (α := α) m β → IterStep (IterM (α := α) m β) β → Prop where
20+
/-- If `it` plausibly yields in its immediate next step, this step is a plausible `0`-th output step. -/
21+
| zero_yield {it : IterM (α := α) m β} : it.IsPlausibleStep (.yield it' out) →
22+
it.IsPlausibleNthOutputStep 0 (.yield it' out)
23+
/--
24+
If `it` plausibly terminates in its immediate next step (`.done`), then `.done` is a plausible
25+
`n`-th output step for arbitrary `n`.
26+
-/
27+
| done {it : IterM (α := α) m β} : it.IsPlausibleStep .done →
28+
it.IsPlausibleNthOutputStep n .done
29+
/--
30+
If `it` plausibly yields in its immediate next step, the successor iterator being `it'`, and
31+
if `step` is a plausible `n`-th output step of `it'`, then `step` is a plausible `n + 1`-th
32+
output step of `it`.
33+
-/
34+
| yield {it it' : IterM (α := α) m β} {out step} : it.IsPlausibleStep (.yield it' out) →
35+
it'.IsPlausibleNthOutputStep n step → it.IsPlausibleNthOutputStep (n + 1) step
36+
/--
37+
If `it` plausibly skips in its immediate next step, the successor iterator being `it'`, and
38+
if `step` is a plausible `n`-th output step of `it'`, then `step` is also a plausible `n`-th
39+
output step of `it`.
40+
-/
41+
| skip {it it' : IterM (α := α) m β} {step} : it.IsPlausibleStep (.skip it') →
42+
it'.IsPlausibleNthOutputStep n step → it.IsPlausibleNthOutputStep n step
43+
44+
/--
45+
`IteratorAccess α m` provides efficient implementations for random access or iterators that support
46+
it. `it.nextAtIdx? n` either returns the step in which the `n`-th value of `it` is emitted
47+
(necessarily of the form `.yield _ _`) or `.done` if `it` terminates before emitting the `n`-th
48+
value.
49+
50+
For monadic iterators, the monadic effects of this operation may differ from manually iterating
51+
to the `n`-th value because `nextAtIdx?` can take shortcuts. By the signature, the return value
52+
is guaranteed to plausible in the sense of `IterM.IsPlausibleNthOutputStep`.
53+
54+
This class is experimental and users of the iterator API should not explicitly depend on it.
55+
-/
56+
class IteratorAccess (α : Type w) (m : Type w → Type w') {β : Type w} [Iterator α m β] where
57+
nextAtIdx? (it : IterM (α := α) m β) (n : Nat) :
58+
m (PlausibleIterStep (it.IsPlausibleNthOutputStep n))
59+
60+
/--
61+
Returns the step in which `it` yields its `n`-th element, or `.done` if it terminates earlier.
62+
In contrast to `step`, this function will always return either `.yield` or `.done` but never a
63+
`.skip` step.
64+
65+
For monadic iterators, the monadic effects of this operation may differ from manually iterating
66+
to the `n`-th value because `nextAtIdx?` can take shortcuts. By the signature, the return value
67+
is guaranteed to plausible in the sense of `IterM.IsPlausibleNthOutputStep`.
68+
69+
This function is only available for iterators that explicitly support it by implementing
70+
the `IteratorAccess` typeclass.
71+
-/
72+
@[always_inline, inline]
73+
def IterM.nextAtIdx? [Iterator α m β] [IteratorAccess α m] (it : IterM (α := α) m β)
74+
(n : Nat) : m (PlausibleIterStep (it.IsPlausibleNthOutputStep n)) :=
75+
IteratorAccess.nextAtIdx? it n
76+
77+
/--
78+
Returns the `n`-th value emitted by `it`, or `none` if `it` terminates earlier.
79+
80+
For monadic iterators, the monadic effects of this operation may differ from manually iterating
81+
to the `n`-th value because `atIdx?` can take shortcuts. By the signature, the return value
82+
is guaranteed to plausible in the sense of `IterM.IsPlausibleNthOutputStep`.
83+
84+
This function is only available for iterators that explicitly support it by implementing
85+
the `IteratorAccess` typeclass.
86+
-/
87+
@[always_inline, inline]
88+
def IterM.atIdx? [Iterator α m β] [IteratorAccess α m] [Monad m] (it : IterM (α := α) m β)
89+
(n : Nat) : m (Option β) := do
90+
match (← IteratorAccess.nextAtIdx? it n).val with
91+
| .yield _ out => return some out
92+
| .skip _ => return none
93+
| .done => return none
94+
95+
end Std.Iterators

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

Lines changed: 13 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -105,12 +105,14 @@ class IteratorSizePartial (α : Type w) (m : Type w → Type w') {β : Type w} [
105105

106106
end Typeclasses
107107

108-
private def IteratorLoop.WFRel {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
108+
/-- Internal implementation detail of the iterator library. -/
109+
def IteratorLoop.WFRel {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
109110
{γ : Type x} {plausible_forInStep : β → γ → ForInStep γ → Prop}
110111
(_wf : WellFounded α m plausible_forInStep) :=
111112
IterM (α := α) m β × γ
112113

113-
private def IteratorLoop.WFRel.mk {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
114+
/-- Internal implementation detail of the iterator library. -/
115+
def IteratorLoop.WFRel.mk {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
114116
{γ : Type x} {plausible_forInStep : β → γ → ForInStep γ → Prop}
115117
(wf : WellFounded α m plausible_forInStep) (it : IterM (α := α) m β) (c : γ) :
116118
IteratorLoop.WFRel wf :=
@@ -134,20 +136,21 @@ def IterM.DefaultConsumers.forIn' {m : Type w → Type w'} {α : Type w} {β : T
134136
(plausible_forInStep : β → γ → ForInStep γ → Prop)
135137
(wf : IteratorLoop.WellFounded α m plausible_forInStep)
136138
(it : IterM (α := α) m β) (init : γ)
137-
(f : (b : β) → it.IsPlausibleIndirectOutput b → (c : γ) → n (Subtype (plausible_forInStep b c))) : n γ :=
139+
(P : β → Prop) (hP : ∀ b, it.IsPlausibleIndirectOutput b → P b)
140+
(f : (b : β) → P b → (c : γ) → n (Subtype (plausible_forInStep b c))) : n γ :=
138141
haveI : WellFounded _ := wf
139142
letI : MonadLift m n := ⟨fun {γ} => lift γ⟩
140143
do
141144
match ← it.step with
142145
| .yield it' out h =>
143-
match ← f out (.direct ⟨_, h⟩) init with
146+
match ← f out (hP _ <| .direct ⟨_, h⟩) init with
144147
| ⟨.yield c, _⟩ =>
145-
IterM.DefaultConsumers.forIn' lift _ plausible_forInStep wf it' c
146-
(fun out h' acc => f out (.indirect ⟨_, rfl, h⟩ h') acc)
148+
IterM.DefaultConsumers.forIn' lift _ plausible_forInStep wf it' c P
149+
(fun _ h' => hP _ <| .indirect ⟨_, rfl, h⟩ h') f
147150
| ⟨.done c, _⟩ => return c
148151
| .skip it' h =>
149-
IterM.DefaultConsumers.forIn' lift _ plausible_forInStep wf it' init
150-
(fun out h' acc => f out (.indirect ⟨_, rfl, h⟩ h') acc)
152+
IterM.DefaultConsumers.forIn' lift _ plausible_forInStep wf it' init P
153+
(fun _ h' => hP _ <| .indirect ⟨_, rfl, h⟩ h') f
151154
| .done _ => return init
152155
termination_by IteratorLoop.WFRel.mk wf it init
153156
decreasing_by
@@ -163,7 +166,7 @@ implementations are possible and should be used instead.
163166
def IteratorLoop.defaultImplementation {α : Type w} {m : Type w → Type w'} {n : Type w → Type w''}
164167
[Monad n] [Iterator α m β] :
165168
IteratorLoop α m n where
166-
forIn lift := IterM.DefaultConsumers.forIn' lift
169+
forIn lift γ Pl wf it init := IterM.DefaultConsumers.forIn' lift γ Pl wf it init _ (fun _ => id)
167170

168171
/--
169172
Asserts that a given `IteratorLoop` instance is equal to `IteratorLoop.defaultImplementation`.
@@ -246,6 +249,7 @@ A `ForIn'` instance for iterators. Its generic membership relation is not easy t
246249
so this is not marked as `instance`. This way, more convenient instances can be built on top of it
247250
or future library improvements will make it more comfortable.
248251
-/
252+
@[always_inline, inline]
249253
def IterM.instForIn' {m : Type w → Type w'} {n : Type w → Type w''}
250254
{α : Type w} {β : Type w} [Iterator α m β] [Finite α m] [IteratorLoop α m n]
251255
[MonadLiftT m n] :

0 commit comments

Comments
 (0)