1 | module Data.Bifoldable
7 | bifoldlM : Monad m => Bifoldable p =>
8 | (f: a -> b -> m a) ->
9 | (g: a -> c -> m a) ->
11 | (input: p b c) -> m a
12 | bifoldlM f g a0 = bifoldl (\ma,b => ma >>= flip f b)
13 | (\ma,c => ma >>= flip g c)
18 | biconcat : Monoid m => Bifoldable p => p m m -> m
19 | biconcat = bifoldr (<+>) (<+>) neutral
24 | biconcatMap : Monoid m => Bifoldable p => (a -> m) -> (b -> m) -> p a b -> m
25 | biconcatMap f g = bifoldr ((<+>) . f) ((<+>) . g) neutral
31 | biand : Bifoldable p => p (Lazy Bool) (Lazy Bool) -> Bool
32 | biand = bifoldl (&&) (&&) True
38 | bior : Bifoldable p => p (Lazy Bool) (Lazy Bool) -> Bool
39 | bior = bifoldl (||) (||) False
44 | biany : Bifoldable p => (a -> Bool) -> (b -> Bool) -> p a b -> Bool
45 | biany f g = bifoldl (\x,y => x || f y) (\x,y => x || g y) False
50 | biall : Bifoldable p => (a -> Bool) -> (b -> Bool) -> p a b -> Bool
51 | biall f g = bifoldl (\x,y => x && f y) (\x,y => x && g y) True
55 | bisum : Num a => Bifoldable p => p a a -> a
56 | bisum = bifoldr (+) (+) 0
61 | bisum' : Num a => Bifoldable p => p a a -> a
62 | bisum' = bifoldl (+) (+) 0
66 | biproduct : Num a => Bifoldable p => p a a -> a
67 | biproduct = bifoldr (*) (*) 1
72 | biproduct' : Num a => Bifoldable p => p a a -> a
73 | biproduct' = bifoldl (*) (*) 1
78 | bitraverse_ : (Bifoldable p, Applicative f)
83 | bitraverse_ f g = bifoldr ((*>) . f) ((*>) . g) (pure ())
87 | bisequence_ : Applicative f => Bifoldable p => p (f a) (f b) -> f ()
88 | bisequence_ = bifoldr (*>) (*>) (pure ())
92 | bifor_ : (Bifoldable p, Applicative f)
97 | bifor_ p f g = bitraverse_ f g p
105 | bichoice : Alternative f => Bifoldable p => p (Lazy (f a)) (Lazy (f a)) -> f a
106 | bichoice t = bifoldr {a = Lazy (f a)} {b = Lazy (f a)} {acc = Lazy (f a)}
107 | (\ x, xs => x <|> xs)
108 | (\ x, xs => x <|> xs)
114 | bichoiceMap : (Bifoldable p, Alternative f)
119 | bichoiceMap fa fb t = bifoldr {a} {b} {acc = Lazy (f x)}
120 | (\e, fx => fa e <|> fx)
121 | (\e, fx => fb e <|> fx)