Merge pull request #376 from andrevidela/kleisli-arrow

Add Kleisli arrow operators to contrib
This commit is contained in:
Niklas Larsson 2020-06-28 00:18:04 +02:00 committed by GitHub
commit cd777a00ce
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
2 changed files with 22 additions and 0 deletions

View File

@ -0,0 +1,21 @@
module Control.Monad.Syntax
%default total
infixr 1 =<<, <=<, >=>
||| Left-to-right Kleisli composition of monads.
public export
(>=>) : Monad m => (a -> m b) -> (b -> m c) -> (a -> m c)
(>=>) f g = \x => f x >>= g
public export
||| Right-to-left Kleisli composition of monads, flipped version of `>=>`.
(<=<) : Monad m => (b -> m c) -> (a -> m b) -> (a -> m c)
(<=<) = flip (>=>)
public export
||| Right-to-left monadic bind, flipped version of `>>=`.
(=<<) : Monad m => (a -> m b) -> m a -> m b
(=<<) = flip (>>=)

View File

@ -2,6 +2,7 @@ package contrib
modules = Control.Delayed,
Control.Linear.LIO,
Control.Monad.Syntax,
Data.Linear.Array,