mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-24 06:52:19 +03:00
[ new ] Applicative and Monad for Pair (#1188)
This commit is contained in:
parent
f4a790ded4
commit
1784593abb
@ -107,6 +107,17 @@ public export
|
||||
Functor (Pair a) where
|
||||
map = mapSnd
|
||||
|
||||
%inline
|
||||
public export
|
||||
Monoid a => Applicative (Pair a) where
|
||||
pure = (neutral,)
|
||||
(a1,f) <*> (a2,v) = (a1 <+> a2, f v)
|
||||
|
||||
%inline
|
||||
public export
|
||||
Monoid a => Monad (Pair a) where
|
||||
(a1,a) >>= f = let (a2,b) = f a in (a1 <+> a2, b)
|
||||
|
||||
-----------
|
||||
-- MAYBE --
|
||||
-----------
|
||||
|
@ -76,6 +76,7 @@ join : m (m a) -> m a
|
||||
|
||||
Implementations:
|
||||
Monad IO
|
||||
Monoid a => Monad (Pair a)
|
||||
Monad Maybe
|
||||
Monad (Either e)
|
||||
Monad List
|
||||
|
Loading…
Reference in New Issue
Block a user