Skip to content

Commit 86eded3

Browse files
authored
refactor: partially move iterators to Init (#8767)
This PR moves parts of the iterator library from `Std` to `Init`. The reason is that the polymorphic range API must be in `Init` and it depends on the iterators.
1 parent f0fdab8 commit 86eded3

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

62 files changed

+979
-802
lines changed

src/Init/Data.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -46,3 +46,4 @@ import Init.Data.NeZero
4646
import Init.Data.Function
4747
import Init.Data.RArray
4848
import Init.Data.Vector
49+
import Init.Data.Iterators

src/Init/Data/Iterators.lean

Lines changed: 19 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,19 @@
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+
import Init.Data.Iterators.PostconditionMonad
11+
import Init.Data.Iterators.Consumers
12+
import Init.Data.Iterators.Lemmas
13+
import Init.Data.Iterators.Internal
14+
15+
/-!
16+
# Iterators
17+
18+
See `Std.Data.Iterators` for an overview over the iterator API.
19+
-/
Lines changed: 42 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -3,6 +3,8 @@ Copyright (c) 2025 Lean FRO, LLC. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Paul Reichert
55
-/
6+
module
7+
68
prelude
79
import Init.Core
810
import Init.Classical
@@ -31,7 +33,7 @@ See `Std.Data.Iterators.Producers` for ways to iterate over common data structur
3133
By convention, the monadic iterator associated with an object can be obtained via dot notation.
3234
For example, `List.iterM IO` creates an iterator over a list in the monad `IO`.
3335
34-
See `Std.Data.Iterators.Consumers` for ways to use an iterator. For example, `it.toList` will
36+
See `Init.Data.Iterators.Consumers` for ways to use an iterator. For example, `it.toList` will
3537
convert a provably finite iterator `it` into a list and `it.allowNontermination.toList` will
3638
do so even if finiteness cannot be proved. It is also always possible to manually iterate using
3739
`it.step`, relying on the termination measures `it.finitelyManySteps` and `it.finitelyManySkips`.
@@ -75,7 +77,7 @@ See `Std.Data.Iterators.Producers` for ways to iterate over common data structur
7577
By convention, the monadic iterator associated with an object can be obtained via dot notation.
7678
For example, `List.iterM IO` creates an iterator over a list in the monad `IO`.
7779
78-
See `Std.Data.Iterators.Consumers` for ways to use an iterator. For example, `it.toList` will
80+
See `Init.Data.Iterators.Consumers` for ways to use an iterator. For example, `it.toList` will
7981
convert a provably finite iterator `it` into a list and `it.allowNontermination.toList` will
8082
do so even if finiteness cannot be proved. It is also always possible to manually iterate using
8183
`it.step`, relying on the termination measures `it.finitelyManySteps` and `it.finitelyManySkips`.
@@ -111,12 +113,14 @@ structure Iter {α : Type w} (β : Type w) where
111113
Converts a pure iterator (`Iter β`) into a monadic iterator (`IterM Id β`) in the
112114
identity monad `Id`.
113115
-/
116+
@[expose]
114117
def Iter.toIterM {α : Type w} {β : Type w} (it : Iter (α := α) β) : IterM (α := α) Id β :=
115118
⟨it.internalState⟩
116119

117120
/--
118121
Converts a monadic iterator (`IterM Id β`) over `Id` into a pure iterator (`Iter β`).
119122
-/
123+
@[expose]
120124
def IterM.toIter {α : Type w} {β : Type w} (it : IterM (α := α) Id β) : Iter (α := α) β :=
121125
⟨it.internalState⟩
122126

@@ -170,6 +174,7 @@ inductive IterStep (α β) where
170174
Returns the succeeding iterator stored in an iterator step or `none` if the step is `.done`
171175
and the iterator has finished.
172176
-/
177+
@[expose]
173178
def IterStep.successor : IterStep α β → Option α
174179
| .yield it _ => some it
175180
| .skip it => some it
@@ -179,7 +184,7 @@ def IterStep.successor : IterStep α β → Option α
179184
If present, applies `f` to the iterator of an `IterStep` and replaces the iterator
180185
with the result of the application of `f`.
181186
-/
182-
@[always_inline, inline]
187+
@[always_inline, inline, expose]
183188
def IterStep.mapIterator {α' : Type u'} (f : α → α') : IterStep α β → IterStep α' β
184189
| .yield it out => .yield (f it) out
185190
| .skip it => .skip (f it)
@@ -224,12 +229,13 @@ of another state. Having this proof bundled up with the step is important for te
224229
225230
See `IterM.Step` and `Iter.Step` for the concrete choice of the plausibility predicate.
226231
-/
232+
@[expose]
227233
def PlausibleIterStep (IsPlausibleStep : IterStep α β → Prop) := Subtype IsPlausibleStep
228234

229235
/--
230236
Match pattern for the `yield` case. See also `IterStep.yield`.
231237
-/
232-
@[match_pattern, simp]
238+
@[match_pattern, simp, expose]
233239
def PlausibleIterStep.yield {IsPlausibleStep : IterStep α β → Prop}
234240
(it' : α) (out : β) (h : IsPlausibleStep (.yield it' out)) :
235241
PlausibleIterStep IsPlausibleStep :=
@@ -238,15 +244,15 @@ def PlausibleIterStep.yield {IsPlausibleStep : IterStep α β → Prop}
238244
/--
239245
Match pattern for the `skip` case. See also `IterStep.skip`.
240246
-/
241-
@[match_pattern, simp]
247+
@[match_pattern, simp, expose]
242248
def PlausibleIterStep.skip {IsPlausibleStep : IterStep α β → Prop}
243249
(it' : α) (h : IsPlausibleStep (.skip it')) : PlausibleIterStep IsPlausibleStep :=
244250
⟨.skip it', h⟩
245251

246252
/--
247253
Match pattern for the `done` case. See also `IterStep.done`.
248254
-/
249-
@[match_pattern, simp]
255+
@[match_pattern, simp, expose]
250256
def PlausibleIterStep.done {IsPlausibleStep : IterStep α β → Prop}
251257
(h : IsPlausibleStep .done) : PlausibleIterStep IsPlausibleStep :=
252258
⟨.done, h⟩
@@ -283,7 +289,7 @@ section Monadic
283289
/--
284290
Converts wraps the state of an iterator into an `IterM` object.
285291
-/
286-
@[always_inline, inline]
292+
@[always_inline, inline, expose]
287293
def toIterM {α : Type w} (it : α) (m : Type w → Type w') (β : Type w) :
288294
IterM (α := α) m β :=
289295
⟨it⟩
@@ -302,6 +308,7 @@ theorem internalState_toIterM {α m β} (it : α) :
302308
Asserts that certain step is plausibly the successor of a given iterator. What "plausible" means
303309
is up to the `Iterator` instance but it should be strong enough to allow termination proofs.
304310
-/
311+
@[expose]
305312
abbrev IterM.IsPlausibleStep {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β] :
306313
IterM (α := α) m β → IterStep (IterM (α := α) m β) β → Prop :=
307314
Iterator.IsPlausibleStep (α := α) (m := m)
@@ -310,6 +317,7 @@ abbrev IterM.IsPlausibleStep {α : Type w} {m : Type w → Type w'} {β : Type w
310317
The type of the step object returned by `IterM.step`, containing an `IterStep`
311318
and a proof that this is a plausible step for the given iterator.
312319
-/
320+
@[expose]
313321
abbrev IterM.Step {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
314322
(it : IterM (α := α) m β) :=
315323
PlausibleIterStep it.IsPlausibleStep
@@ -318,6 +326,7 @@ abbrev IterM.Step {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator
318326
Asserts that a certain output value could plausibly be emitted by the given iterator in its next
319327
step.
320328
-/
329+
@[expose]
321330
def IterM.IsPlausibleOutput {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
322331
(it : IterM (α := α) m β) (out : β) : Prop :=
323332
∃ it', it.IsPlausibleStep (.yield it' out)
@@ -326,6 +335,7 @@ def IterM.IsPlausibleOutput {α : Type w} {m : Type w → Type w'} {β : Type w}
326335
Asserts that a certain iterator `it'` could plausibly be the directly succeeding iterator of another
327336
given iterator `it`.
328337
-/
338+
@[expose]
329339
def IterM.IsPlausibleSuccessorOf {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
330340
(it' it : IterM (α := α) m β) : Prop :=
331341
∃ step, step.successor = some it' ∧ it.IsPlausibleStep step
@@ -334,6 +344,7 @@ def IterM.IsPlausibleSuccessorOf {α : Type w} {m : Type w → Type w'} {β : Ty
334344
Asserts that a certain iterator `it'` could plausibly be the directly succeeding iterator of another
335345
given iterator `it` while no value is emitted (see `IterStep.skip`).
336346
-/
347+
@[expose]
337348
def IterM.IsPlausibleSkipSuccessorOf {α : Type w} {m : Type w → Type w'} {β : Type w}
338349
[Iterator α m β] (it' it : IterM (α := α) m β) : Prop :=
339350
it.IsPlausibleStep (.skip it')
@@ -356,6 +367,7 @@ section Pure
356367
Asserts that certain step is plausibly the successor of a given iterator. What "plausible" means
357368
is up to the `Iterator` instance but it should be strong enough to allow termination proofs.
358369
-/
370+
@[expose]
359371
def Iter.IsPlausibleStep {α : Type w} {β : Type w} [Iterator α Id β]
360372
(it : Iter (α := α) β) (step : IterStep (Iter (α := α) β) β) : Prop :=
361373
it.toIterM.IsPlausibleStep (step.mapIterator Iter.toIterM)
@@ -364,6 +376,7 @@ def Iter.IsPlausibleStep {α : Type w} {β : Type w} [Iterator α Id β]
364376
The type of the step object returned by `Iter.step`, containing an `IterStep`
365377
and a proof that this is a plausible step for the given iterator.
366378
-/
379+
@[expose]
367380
def Iter.Step {α : Type w} {β : Type w} [Iterator α Id β] (it : Iter (α := α) β) :=
368381
PlausibleIterStep (Iter.IsPlausibleStep it)
369382

@@ -378,7 +391,7 @@ def Iter.Step.toMonadic {α : Type w} {β : Type w} [Iterator α Id β] {it : It
378391
/--
379392
Converts an `IterM.Step` into an `Iter.Step`.
380393
-/
381-
@[always_inline, inline]
394+
@[always_inline, inline, expose]
382395
def IterM.Step.toPure {α : Type w} {β : Type w} [Iterator α Id β] {it : IterM (α := α) Id β}
383396
(step : it.Step) : it.toIter.Step :=
384397
⟨step.val.mapIterator IterM.toIter, (by simp [Iter.IsPlausibleStep, step.property])⟩
@@ -402,6 +415,7 @@ theorem IterM.Step.toPure_done {α β : Type w} [Iterator α Id β] {it : IterM
402415
Asserts that a certain output value could plausibly be emitted by the given iterator in its next
403416
step.
404417
-/
418+
@[expose]
405419
def Iter.IsPlausibleOutput {α : Type w} {β : Type w} [Iterator α Id β]
406420
(it : Iter (α := α) β) (out : β) : Prop :=
407421
it.toIterM.IsPlausibleOutput out
@@ -410,6 +424,7 @@ def Iter.IsPlausibleOutput {α : Type w} {β : Type w} [Iterator α Id β]
410424
Asserts that a certain iterator `it'` could plausibly be the directly succeeding iterator of another
411425
given iterator `it`.
412426
-/
427+
@[expose]
413428
def Iter.IsPlausibleSuccessorOf {α : Type w} {β : Type w} [Iterator α Id β]
414429
(it' it : Iter (α := α) β) : Prop :=
415430
it'.toIterM.IsPlausibleSuccessorOf it.toIterM
@@ -427,7 +442,7 @@ Makes a single step with the given iterator `it`, potentially emitting a value a
427442
succeeding iterator. If this function is used recursively, termination can sometimes be proved with
428443
the termination measures `it.finitelyManySteps` and `it.finitelyManySkips`.
429444
-/
430-
@[always_inline, inline]
445+
@[always_inline, inline, expose]
431446
def Iter.step {α β : Type w} [Iterator α Id β] (it : Iter (α := α) β) : it.Step :=
432447
it.toIterM.step.run.toPure
433448

@@ -456,6 +471,7 @@ structure IterM.TerminationMeasures.Finite
456471
The relation of plausible successors on `IterM.TerminationMeasures.Finite`. It is well-founded
457472
if there is a `Finite` instance.
458473
-/
474+
@[expose]
459475
def IterM.TerminationMeasures.Finite.Rel
460476
{α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β] :
461477
TerminationMeasures.Finite α m → TerminationMeasures.Finite α m → Prop :=
@@ -464,12 +480,13 @@ def IterM.TerminationMeasures.Finite.Rel
464480
instance {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
465481
[Finite α m] : WellFoundedRelation (IterM.TerminationMeasures.Finite α m) where
466482
rel := IterM.TerminationMeasures.Finite.Rel
467-
wf := (InvImage.wf _ Finite.wf).transGen
483+
wf := by exact (InvImage.wf _ Finite.wf).transGen
468484

469485
/--
470486
Termination measure to be used in well-founded recursive functions recursing over a finite iterator
471487
(see also `Finite`).
472488
-/
489+
@[expose]
473490
def IterM.finitelyManySteps {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
474491
[Finite α m] (it : IterM (α := α) m β) : IterM.TerminationMeasures.Finite α m :=
475492
⟨it⟩
@@ -494,9 +511,10 @@ theorem IterM.TerminationMeasures.Finite.rel_of_skip
494511
macro_rules | `(tactic| decreasing_trivial) => `(tactic|
495512
first
496513
| exact IterM.TerminationMeasures.Finite.rel_of_yield ‹_›
497-
| exact IterM.TerminationMeasures.Finite.rel_of_skip ‹_›)
514+
| exact IterM.TerminationMeasures.Finite.rel_of_skip ‹_›
515+
| fail)
498516

499-
@[inherit_doc IterM.finitelyManySteps]
517+
@[inherit_doc IterM.finitelyManySteps, expose]
500518
def Iter.finitelyManySteps {α : Type w} {β : Type w} [Iterator α Id β] [Finite α Id]
501519
(it : Iter (α := α) β) : IterM.TerminationMeasures.Finite α Id :=
502520
it.toIterM.finitelyManySteps
@@ -521,7 +539,8 @@ theorem Iter.TerminationMeasures.Finite.rel_of_skip
521539
macro_rules | `(tactic| decreasing_trivial) => `(tactic|
522540
first
523541
| exact Iter.TerminationMeasures.Finite.rel_of_yield ‹_›
524-
| exact Iter.TerminationMeasures.Finite.rel_of_skip ‹_›)
542+
| exact Iter.TerminationMeasures.Finite.rel_of_skip ‹_›
543+
| fail)
525544

526545
theorem IterM.isPlausibleSuccessorOf_of_yield
527546
{α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
@@ -561,6 +580,7 @@ structure IterM.TerminationMeasures.Productive
561580
The relation of plausible successors while skipping on `IterM.TerminationMeasures.Productive`.
562581
It is well-founded if there is a `Productive` instance.
563582
-/
583+
@[expose]
564584
def IterM.TerminationMeasures.Productive.Rel
565585
{α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β] :
566586
TerminationMeasures.Productive α m → TerminationMeasures.Productive α m → Prop :=
@@ -569,12 +589,13 @@ def IterM.TerminationMeasures.Productive.Rel
569589
instance {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
570590
[Productive α m] : WellFoundedRelation (IterM.TerminationMeasures.Productive α m) where
571591
rel := IterM.TerminationMeasures.Productive.Rel
572-
wf := (InvImage.wf _ Productive.wf).transGen
592+
wf := by exact (InvImage.wf _ Productive.wf).transGen
573593

574594
/--
575595
Termination measure to be used in well-founded recursive functions recursing over a productive
576596
iterator (see also `Productive`).
577597
-/
598+
@[expose]
578599
def IterM.finitelyManySkips {α : Type w} {m : Type w → Type w'} {β : Type w} [Iterator α m β]
579600
[Productive α m] (it : IterM (α := α) m β) : IterM.TerminationMeasures.Productive α m :=
580601
⟨it⟩
@@ -590,9 +611,11 @@ theorem IterM.TerminationMeasures.Productive.rel_of_skip
590611
.single h
591612

592613
macro_rules | `(tactic| decreasing_trivial) => `(tactic|
593-
exact IterM.TerminationMeasures.Productive.rel_of_skip ‹_›)
614+
first
615+
| exact IterM.TerminationMeasures.Productive.rel_of_skip ‹_›
616+
| fail)
594617

595-
@[inherit_doc IterM.finitelyManySkips]
618+
@[inherit_doc IterM.finitelyManySkips, expose]
596619
def Iter.finitelyManySkips {α : Type w} {β : Type w} [Iterator α Id β] [Productive α Id]
597620
(it : Iter (α := α) β) : IterM.TerminationMeasures.Productive α Id :=
598621
it.toIterM.finitelyManySkips
@@ -608,7 +631,9 @@ theorem Iter.TerminationMeasures.Productive.rel_of_skip
608631
IterM.TerminationMeasures.Productive.rel_of_skip h
609632

610633
macro_rules | `(tactic| decreasing_trivial) => `(tactic|
611-
exact Iter.TerminationMeasures.Productive.rel_of_skip ‹_›)
634+
first
635+
| exact Iter.TerminationMeasures.Productive.rel_of_skip ‹_›
636+
| fail)
612637

613638
instance [Iterator α m β] [Finite α m] : Productive α m where
614639
wf := by
Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,13 @@
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.Consumers.Monadic
10+
import Init.Data.Iterators.Consumers.Access
11+
import Init.Data.Iterators.Consumers.Collect
12+
import Init.Data.Iterators.Consumers.Loop
13+
import Init.Data.Iterators.Consumers.Partial

src/Std/Data/Iterators/Consumers/Access.lean renamed to src/Init/Data/Iterators/Consumers/Access.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,8 +3,10 @@ Copyright (c) 2025 Lean FRO, LLC. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Paul Reichert
55
-/
6+
module
7+
68
prelude
7-
import Std.Data.Iterators.Consumers.Partial
9+
import Init.Data.Iterators.Consumers.Partial
810

911
namespace Std.Iterators
1012

src/Std/Data/Iterators/Consumers/Collect.lean renamed to src/Init/Data/Iterators/Consumers/Collect.lean

Lines changed: 5 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -3,10 +3,12 @@ Copyright (c) 2025 Lean FRO, LLC. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Paul Reichert
55
-/
6+
module
7+
68
prelude
7-
import Std.Data.Iterators.Basic
8-
import Std.Data.Iterators.Consumers.Partial
9-
import Std.Data.Iterators.Consumers.Monadic.Collect
9+
import Init.Data.Iterators.Basic
10+
import Init.Data.Iterators.Consumers.Partial
11+
import Init.Data.Iterators.Consumers.Monadic.Collect
1012

1113
/-!
1214
# Collectors

src/Std/Data/Iterators/Consumers/Loop.lean renamed to src/Init/Data/Iterators/Consumers/Loop.lean

Lines changed: 4 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -3,9 +3,11 @@ Copyright (c) 2025 Lean FRO, LLC. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Paul Reichert
55
-/
6+
module
7+
68
prelude
7-
import Std.Data.Iterators.Consumers.Monadic.Loop
8-
import Std.Data.Iterators.Consumers.Partial
9+
import Init.Data.Iterators.Consumers.Monadic.Loop
10+
import Init.Data.Iterators.Consumers.Partial
911

1012
/-!
1113
# Loop consumers
Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,11 @@
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.Consumers.Monadic.Collect
10+
import Init.Data.Iterators.Consumers.Monadic.Loop
11+
import Init.Data.Iterators.Consumers.Monadic.Partial

src/Std/Data/Iterators/Consumers/Monadic/Collect.lean renamed to src/Init/Data/Iterators/Consumers/Monadic/Collect.lean

Lines changed: 4 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -3,9 +3,11 @@ Copyright (c) 2025 Lean FRO, LLC. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Paul Reichert
55
-/
6+
module
7+
68
prelude
7-
import Std.Data.Iterators.Consumers.Monadic.Partial
8-
import Std.Data.Internal.LawfulMonadLiftFunction
9+
import Init.Data.Iterators.Consumers.Monadic.Partial
10+
import Init.Data.Iterators.Internal.LawfulMonadLiftFunction
911

1012
/-!
1113
# Collectors

0 commit comments

Comments
 (0)