Idris2/libs/contrib/Data/List/Lazy.idr

56 lines
1.2 KiB
Idris
Raw Normal View History

2020-09-19 23:46:46 +03:00
module Data.List.Lazy
%default total
-- All functions here are public export
-- because their definitions are pretty much the specification.
public export
data LazyList : Type -> Type where
Nil : LazyList a
2020-09-20 10:56:58 +03:00
(::) : (1 x : a) -> (1 xs : Lazy (LazyList a)) -> LazyList a
2020-09-19 23:46:46 +03:00
public export
Semigroup (LazyList a) where
[] <+> ys = ys
(x :: xs) <+> ys = x :: (xs <+> ys)
public export
Monoid (LazyList a) where
neutral = []
public export
Foldable LazyList where
foldr op nil [] = nil
foldr op nil (x :: xs) = x `op` foldr op nil xs
foldl op acc [] = acc
foldl op acc (x :: xs) = foldl op (acc `op` x) xs
public export
Functor LazyList where
map f [] = []
map f (x :: xs) = f x :: map f xs
public export
Applicative LazyList where
pure x = [x]
fs <*> vs = concatMap (\f => map f vs) fs
public export
Alternative LazyList where
empty = []
(<|>) = (<+>)
public export
Monad LazyList where
m >>= f = concatMap f m
-- There is no Traversable instance for lazy lists.
-- 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 g [] = pure []
traverse g (x :: xs) = [| g x :: traverse g xs |]