Idris2/tests/idris2/interface006/Bimonad.idr
Edwin Brady a972778eab Add test script
They don't all pass yet, for minor reasons. Coming shortly...
Unfortunately the startup overhead for chez is really noticeable here!
2020-05-19 18:25:18 +01:00

25 lines
754 B
Idris

module Bimonad
import Bifunctor
import Biapplicative
infixl 4 >>==
||| Bimonads
||| @p the action of the first Bifunctor component on pairs of objects
||| @q the action of the second Bifunctor component on pairs of objects
interface (Biapplicative p, Biapplicative q) =>
Bimonad (p : Type -> Type -> Type) (q : Type -> Type -> Type) where
bijoin : (p (p a b) (q a b), q (p a b) (q a b)) -> (p a b, q a b)
bijoin p = p >>== (id, id)
(>>==) : (p a b, q a b) -> ((a -> p c d), (b -> q c d)) -> (p c d, q c d)
(pab, qab) >>== (f, g) = bijoin $ (bimap f g, bimap f g) <<*>> (pab, qab)
biunit : (Bimonad p q) => a -> b -> (p a b, q a b)
biunit a b = (bipure a b, bipure a b)
implementation Bimonad Pair Pair where
bijoin = bimap fst snd