[ fix ] Make traverse and friends lazy for LazyList

This commit is contained in:
Denis Buzdalov 2023-09-15 16:49:17 +03:00 committed by G. Allais
parent 6dc06cd67d
commit 3886200d29
7 changed files with 51 additions and 5 deletions

View File

@ -228,6 +228,10 @@
about the modulo operator's upper bound, which can be useful when working with about the modulo operator's upper bound, which can be useful when working with
indices (for example). 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 #### Papers
* In `Control.DivideAndConquer`: a port of the paper * In `Control.DivideAndConquer`: a port of the paper

View File

@ -124,16 +124,16 @@ Monad LazyList where
-- The result of a traversal will be a non-lazy list in general -- The result of a traversal will be a non-lazy list in general
-- (you can't delay the "effect" of an applicative functor). -- (you can't delay the "effect" of an applicative functor).
public export 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 [] = pure []
traverse g (x :: xs) = [| g x :: traverse g xs |] traverse g (x::xs) = pure $ !(g x) :: !(traverse g xs)
public export %inline 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 for = flip traverse
public export %inline public export %inline
sequence : Applicative f => LazyList (f a) -> f (List a) sequence : Monad f => LazyList (f a) -> f (List a)
sequence = traverse id sequence = traverse id
-- Traverse a lazy list with lazy effect lazily -- Traverse a lazy list with lazy effect lazily

View File

@ -125,7 +125,7 @@ idrisTests = MkTestPool "Misc" [] Nothing
"import001", "import002", "import003", "import004", "import005", "import006", "import001", "import002", "import003", "import004", "import005", "import006",
"import007", "import008", "import009", "import007", "import008", "import009",
-- Implicit laziness, lazy evaluation -- Implicit laziness, lazy evaluation
"lazy001", "lazy002", "lazy003", "lazy001", "lazy002", "lazy003", "lazy004",
-- Namespace blocks -- Namespace blocks
"namespace001", "namespace002", "namespace003", "namespace004", "namespace005", "namespace001", "namespace002", "namespace003", "namespace004", "namespace005",
-- Parameters blocks -- Parameters blocks

View File

@ -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'

View File

@ -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!

View File

@ -0,0 +1,2 @@
:exec putStrLn $ show ys
:exec putStrLn $ show ys'

3
tests/idris2/misc/lazy004/run Executable file
View File

@ -0,0 +1,3 @@
. ../../../testutils.sh
idris2 -p contrib LazyFor.idr < input