Skip to content

Commit

Permalink
[ fix ] Make traverse and friends lazy for LazyList
Browse files Browse the repository at this point in the history
  • Loading branch information
buzden committed Sep 15, 2023
1 parent dc79c6d commit 4856cde
Show file tree
Hide file tree
Showing 7 changed files with 51 additions and 5 deletions.
4 changes: 4 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -226,6 +226,10 @@
about the modulo operator's upper bound, which can be useful when working with
indices (for example).

* Existing specialised variants of functions from the `Traversable` for `LazyList`
were made to be indeed lazy by the effect, but their requirements were strengthened
from `Applicative` to `Monad`.

#### Papers

* In `Control.DivideAndConquer`: a port of the paper
Expand Down
8 changes: 4 additions & 4 deletions libs/contrib/Data/List/Lazy.idr
Original file line number Diff line number Diff line change
Expand Up @@ -124,16 +124,16 @@ Monad LazyList where
-- The result of a traversal will be a non-lazy list in general
-- (you can't delay the "effect" of an applicative functor).
public export
traverse : Applicative f => (a -> f b) -> LazyList a -> f (List b)
traverse : Monad f => (a -> f b) -> LazyList a -> f (List b)
traverse g [] = pure []
traverse g (x :: xs) = [| g x :: traverse g xs |]
traverse g (x::xs) = pure $ !(g x) :: !(traverse g xs)

public export %inline
for : Applicative f => LazyList a -> (a -> f b) -> f (List b)
for : Monad f => LazyList a -> (a -> f b) -> f (List b)
for = flip traverse

public export %inline
sequence : Applicative f => LazyList (f a) -> f (List a)
sequence : Monad f => LazyList (f a) -> f (List a)
sequence = traverse id

-- Traverse a lazy list with lazy effect lazily
Expand Down
2 changes: 1 addition & 1 deletion tests/Main.idr
Original file line number Diff line number Diff line change
Expand Up @@ -125,7 +125,7 @@ idrisTests = MkTestPool "Misc" [] Nothing
"import001", "import002", "import003", "import004", "import005", "import006",
"import007", "import008", "import009",
-- Implicit laziness, lazy evaluation
"lazy001", "lazy002", "lazy003",
"lazy001", "lazy002", "lazy003", "lazy004",
-- Namespace blocks
"namespace001", "namespace002", "namespace003", "namespace004", "namespace005",
-- Parameters blocks
Expand Down
26 changes: 26 additions & 0 deletions tests/idris2/misc/lazy004/LazyFor.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,26 @@
import Data.List.Lazy

import Debug.Trace

%default total

xs : LazyList Nat
xs = iterateN 5 (traceValBy (("xs: " ++) . show) . S) Z

ys : Maybe (List Nat)
ys = for xs $ \n => if n >= 3 then Nothing else Just (10 * n)

xs' : LazyList Nat
xs' = iterateN 5 (traceValBy (("xs': " ++) . show) . S) Z

for' : Monad f => LazyList a -> (a -> f b) -> f $ List b
for' [] _ = pure []
for' (x::xs) g = pure $ !(g x) :: !(for' xs g)

ys' : Maybe (List Nat)
ys' = for' xs' $ \n => if n >= 3 then Nothing else Just (10 * n)

main : IO ()
main = do
putStrLn $ show ys
putStrLn $ show ys'
11 changes: 11 additions & 0 deletions tests/idris2/misc/lazy004/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,11 @@
1/1: Building LazyFor (LazyFor.idr)
Main> xs: 1
xs: 2
xs: 3
Nothing
Main> xs': 1
xs': 2
xs': 3
Nothing
Main>
Bye for now!
2 changes: 2 additions & 0 deletions tests/idris2/misc/lazy004/input
Original file line number Diff line number Diff line change
@@ -0,0 +1,2 @@
:exec putStrLn $ show ys
:exec putStrLn $ show ys'
3 changes: 3 additions & 0 deletions tests/idris2/misc/lazy004/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
. ../../../testutils.sh

idris2 -p contrib LazyFor.idr < input

0 comments on commit 4856cde

Please sign in to comment.