Add a (>>) operator, implemented in terms of (>>=).

This mirrors the (>>) found in Haskell, which is the same as (>>=), except the
second computation (on the right hand side) takes no arguments, and the result
of the first computation is discarded. This is a trivial implementation written
in terms of (>>=).
This commit is contained in:
Molly Miller 2020-10-01 12:12:14 +01:00 committed by G. Allais
parent 28cf2a3083
commit 7d5ec53b53
2 changed files with 5 additions and 1 deletions

View File

@ -176,6 +176,10 @@ interface Applicative m => Monad m where
%allow_overloads (>>=)
public export
(>>) : (Monad m) => m a -> m b -> m b
a >> b = a >>= \_ => b
||| `guard a` is `pure ()` if `a` is `True` and `empty` if `a` is `False`.
public export
guard : Alternative f => Bool -> f ()

View File

@ -14,7 +14,7 @@ infixr 4 ||
infixr 7 ::, ++
-- Functor/Applicative/Monad/Algebra operators
infixl 1 >>=
infixl 1 >>=, >>
infixr 2 <|>
infixl 3 <*>, *>, <*
infixr 4 <$>, $>, <$