[ new ] Add Bifoldable and Bitraversable interfaces to Prelude (#1265)

This commit is contained in:
Stefan Höck 2021-04-08 18:25:37 +02:00 committed by GitHub
parent c4efa0d29f
commit 61c43e5337
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
7 changed files with 230 additions and 10 deletions

View File

@ -0,0 +1,128 @@
||| Additional utility functions for the `Bifoldable` interface.
module Data.Bifoldable
||| Left associative monadic bifold over a structure.
public export
bifoldlM : (Bifoldable p, Monad m)
=> (f: a -> b -> m a)
-> (g: a -> c -> m a)
-> (init: a)
-> (input: p b c) -> m a
bifoldlM f g a0 = bifoldl (\ma,b => ma >>= flip f b)
(\ma,c => ma >>= flip g c)
(pure a0)
||| Combines the elements of a structure,
||| given ways of mapping them to a common monoid.
public export
bifoldMap : (Bifoldable p, Monoid m) => (a -> m) -> (b -> m) -> p a b -> m
bifoldMap f g = bifoldr ((<+>) . f) ((<+>) . g) neutral
||| Combines the elements of a structure using a monoid.
public export
biconcat : (Bifoldable p, Monoid m) => p m m -> m
biconcat = bifoldr (<+>) (<+>) neutral
||| Combines the elements of a structure,
||| given ways of mapping them to a common monoid.
public export
biconcatMap : (Bifoldable p, Monoid m) => (a -> m) -> (b -> m) -> p a b -> m
biconcatMap f g = bifoldr ((<+>) . f) ((<+>) . g) neutral
||| The conjunction of all elements of a structure containing lazy boolean
||| values. `biand` short-circuits from left to right, evaluating until either an
||| element is `False` or no elements remain.
public export
biand : Bifoldable p => p (Lazy Bool) (Lazy Bool) -> Bool
biand = bifoldl (&&) (&&) True
||| The disjunction of all elements of a structure containing lazy boolean
||| values. `bior` short-circuits from left to right, evaluating either until an
||| element is `True` or no elements remain.
public export
bior : Bifoldable p => p (Lazy Bool) (Lazy Bool) -> Bool
bior = bifoldl (||) (||) False
||| The disjunction of the collective results of applying a predicate to all
||| elements of a structure. `biany` short-circuits from left to right.
public export
biany : Bifoldable p => (a -> Bool) -> (b -> Bool) -> p a b -> Bool
biany f g = bifoldl (\x,y => x || f y) (\x,y => x || g y) False
||| The disjunction of the collective results of applying a predicate to all
||| elements of a structure. `biall` short-circuits from left to right.
public export
biall : Bifoldable p => (a -> Bool) -> (b -> Bool) -> p a b -> Bool
biall f g = bifoldl (\x,y => x && f y) (\x,y => x && g y) True
||| Add together all the elements of a structure.
public export
bisum : (Bifoldable p, Num a) => p a a -> a
bisum = bifoldr (+) (+) 0
||| Add together all the elements of a structure.
||| Same as `bisum` but tail recursive.
export
bisum' : (Bifoldable p, Num a) => p a a -> a
bisum' = bifoldl (+) (+) 0
||| Multiply together all elements of a structure.
public export
biproduct : (Bifoldable p, Num a) => p a a -> a
biproduct = bifoldr (*) (*) 1
||| Multiply together all elements of a structure.
||| Same as `product` but tail recursive.
export
biproduct' : (Bifoldable p, Num a) => p a a -> a
biproduct' = bifoldl (*) (*) 1
||| Map each element of a structure to a computation, evaluate those
||| computations and discard the results.
public export
bitraverse_ : (Bifoldable p, Applicative f)
=> (a -> f x)
-> (b -> f y)
-> p a b
-> f ()
bitraverse_ f g = bifoldr ((*>) . f) ((*>) . g) (pure ())
||| Evaluate each computation in a structure and discard the results.
public export
bisequence_ : (Bifoldable p, Applicative f) => p (f a) (f b) -> f ()
bisequence_ = bifoldr (*>) (*>) (pure ())
||| Like `bitraverse_` but with the arguments flipped.
public export
bifor_ : (Bifoldable p, Applicative f)
=> p a b
-> (a -> f x)
-> (b -> f y)
-> f ()
bifor_ p f g = bitraverse_ f g p
||| Bifold using Alternative.
|||
||| If you have a left-biased alternative operator `<|>`, then `choice` performs
||| left-biased choice from a list of alternatives, which means that it
||| evaluates to the left-most non-`empty` alternative.
public export
bichoice : (Bifoldable p, Alternative f) => p (Lazy (f a)) (Lazy (f a)) -> f a
bichoice t = bifoldr {a = Lazy (f a)} {b = Lazy (f a)} {acc = Lazy (f a)}
(\ x, xs => x <|> xs)
(\ x, xs => x <|> xs)
empty
t
||| A fused version of `bichoice` and `bimap`.
public export
bichoiceMap : (Bifoldable p, Alternative f)
=> (a -> f x)
-> (b -> f x)
-> p a b ->
f x
bichoiceMap fa fb t = bifoldr {a} {b} {acc = Lazy (f x)}
(\e, fx => fa e <|> fx)
(\e, fx => fb e <|> fx)
empty
t

View File

@ -40,6 +40,26 @@ Bifunctor These where
bimap f g (That b) = That (g b) bimap f g (That b) = That (g b)
bimap f g (Both a b) = Both (f a) (g b) bimap f g (Both a b) = Both (f a) (g b)
%inline
public export
Bifoldable These where
bifoldr f _ acc (This a) = f a acc
bifoldr _ g acc (That b) = g b acc
bifoldr f g acc (Both a b) = f a (g b acc)
bifoldl f _ acc (This a) = f acc a
bifoldl _ g acc (That b) = g acc b
bifoldl f g acc (Both a b) = g (f acc a) b
binull _ = False
%inline
public export
Bitraversable These where
bitraverse f _ (This a) = This <$> f a
bitraverse _ g (That b) = That <$> g b
bitraverse f g (Both a b) = [| Both (f a) (g b) |]
%inline %inline
public export public export
Functor (These a) where Functor (These a) where
@ -50,9 +70,3 @@ bifold : Monoid m => These m m -> m
bifold (This a) = a bifold (This a) = a
bifold (That b) = b bifold (That b) = b
bifold (Both a b) = a <+> b bifold (Both a b) = a <+> b
public export
bitraverse : Applicative f => (a -> f c) -> (b -> f d) -> These a b -> f (These c d)
bitraverse f g (This a) = [| This (f a) |]
bitraverse f g (That b) = [| That (g b) |]
bitraverse f g (Both a b) = [| Both (f a) (g b) |]

View File

@ -28,6 +28,7 @@ modules = Control.App,
Control.Monad.Writer.Interface, Control.Monad.Writer.Interface,
Control.WellFounded, Control.WellFounded,
Data.Bifoldable,
Data.Bits, Data.Bits,
Data.Bool, Data.Bool,
Data.Bool.Xor, Data.Bool.Xor,

View File

@ -33,6 +33,21 @@ Bifunctor Validated where
bimap _ s $ Valid x = Valid $ s x bimap _ s $ Valid x = Valid $ s x
bimap f _ $ Invalid e = Invalid $ f e bimap f _ $ Invalid e = Invalid $ f e
public export
Bifoldable Validated where
bifoldr _ g acc (Valid a) = g a acc
bifoldr f _ acc (Invalid e) = f e acc
bifoldl _ g acc (Valid a) = g acc a
bifoldl f _ acc (Invalid e) = f acc e
binull _ = False
public export
Bitraversable Validated where
bitraverse _ g (Valid a) = Valid <$> g a
bitraverse f _ (Invalid e) = Invalid <$> f e
||| Applicative composition preserves invalidity sequentially accumulating all errors. ||| Applicative composition preserves invalidity sequentially accumulating all errors.
public export public export
Semigroup e => Applicative (Validated e) where Semigroup e => Applicative (Validated e) where

View File

@ -372,6 +372,19 @@ namespace Foldable
foldl = foldl . foldl foldl = foldl . foldl
null tf = null tf || all (force . null) tf null tf = null tf || all (force . null) tf
||| `Bifoldable` identifies foldable structures with two different varieties
||| of elements (as opposed to `Foldable`, which has one variety of element).
||| Common examples are `Either` and `Pair`.
public export
interface Bifoldable p where
bifoldr : (a -> acc -> acc) -> (b -> acc -> acc) -> acc -> p a b -> acc
bifoldl : (acc -> a -> acc) -> (acc -> b -> acc) -> acc -> p a b -> acc
bifoldl f g z t = bifoldr (flip (.) . flip f) (flip (.) . flip g) id t z
binull : p a b -> Lazy Bool
binull = bifoldr {acc = Lazy Bool} (\ _,_ => False) (\ _,_ => False) True
public export public export
interface (Functor t, Foldable t) => Traversable t where interface (Functor t, Foldable t) => Traversable t where
||| Map each element of a structure to a computation, evaluate those ||| Map each element of a structure to a computation, evaluate those
@ -388,6 +401,26 @@ public export
for : (Traversable t, Applicative f) => t a -> (a -> f b) -> f (t b) for : (Traversable t, Applicative f) => t a -> (a -> f b) -> f (t b)
for = flip traverse for = flip traverse
public export
interface (Bifunctor p, Bifoldable p) => Bitraversable p where
||| Map each element of a structure to a computation, evaluate those
||| computations and combine the results.
bitraverse : Applicative f => (a -> f c) -> (b -> f d) -> p a b -> f (p c d)
||| Evaluate each computation in a structure and collect the results.
public export
bisequence : (Bitraversable p, Applicative f) => p (f a) (f b) -> f (p a b)
bisequence = bitraverse id id
||| Like `bitraverse` but with the arguments flipped.
public export
bifor : (Bitraversable p, Applicative f)
=> p a b
-> (a -> f c)
-> (b -> f d)
-> f (p c d)
bifor t f g = bitraverse f g t
namespace Traversable namespace Traversable
||| Composition of traversables is traversable. ||| Composition of traversables is traversable.
public export public export

View File

@ -102,6 +102,18 @@ public export
Bifunctor Pair where Bifunctor Pair where
bimap f g (x, y) = (f x, g y) bimap f g (x, y) = (f x, g y)
%inline
public export
Bifoldable Pair where
bifoldr f g acc (x, y) = f x (g y acc)
bifoldl f g acc (x, y) = g (f acc x) y
binull _ = False
%inline
public export
Bitraversable Pair where
bitraverse f g (a,b) = [| (,) (f a) (g b) |]
%inline %inline
public export public export
Functor (Pair a) where Functor (Pair a) where
@ -272,6 +284,23 @@ Bifunctor Either where
bimap f _ (Left x) = Left (f x) bimap f _ (Left x) = Left (f x)
bimap _ g (Right y) = Right (g y) bimap _ g (Right y) = Right (g y)
%inline
public export
Bifoldable Either where
bifoldr f _ acc (Left a) = f a acc
bifoldr _ g acc (Right b) = g b acc
bifoldl f _ acc (Left a) = f acc a
bifoldl _ g acc (Right b) = g acc b
binull _ = False
%inline
public export
Bitraversable Either where
bitraverse f _ (Left a) = Left <$> f a
bitraverse _ g (Right b) = Right <$> g b
%inline %inline
public export public export
Applicative (Either e) where Applicative (Either e) where

View File

@ -45,12 +45,12 @@ Prelude.Types.rangeFromThenTo = [{arg:0}, {arg:1}, {arg:2}, {arg:3}, {arg:4}]: (
Prelude.Types.null = [{arg:0}, {arg:1}]: (%case !{arg:1} [(%concase Prelude.Types.Nil Just 0 [] (%delay Lazy 0)), (%concase Prelude.Types.:: Just 1 [{e:2}, {e:3}] (%delay Lazy 1))] Nothing) Prelude.Types.null = [{arg:0}, {arg:1}]: (%case !{arg:1} [(%concase Prelude.Types.Nil Just 0 [] (%delay Lazy 0)), (%concase Prelude.Types.:: Just 1 [{e:2}, {e:3}] (%delay Lazy 1))] Nothing)
Prelude.Types.foldr = [{arg:0}, {arg:1}, {arg:2}, {arg:3}, {arg:4}]: (%case !{arg:4} [(%concase Prelude.Types.Nil Just 0 [] !{arg:3}), (%concase Prelude.Types.:: Just 1 [{e:2}, {e:3}] ((!{arg:2} [!{e:2}]) [(Prelude.Types.foldr [___, ___, !{arg:2}, !{arg:3}, !{e:3}])]))] Nothing) Prelude.Types.foldr = [{arg:0}, {arg:1}, {arg:2}, {arg:3}, {arg:4}]: (%case !{arg:4} [(%concase Prelude.Types.Nil Just 0 [] !{arg:3}), (%concase Prelude.Types.:: Just 1 [{e:2}, {e:3}] ((!{arg:2} [!{e:2}]) [(Prelude.Types.foldr [___, ___, !{arg:2}, !{arg:3}, !{e:3}])]))] Nothing)
Prelude.Types.foldl = [{arg:0}, {arg:1}, {arg:2}, {arg:3}, {arg:4}]: (%case !{arg:4} [(%concase Prelude.Types.Nil Just 0 [] !{arg:3}), (%concase Prelude.Types.:: Just 1 [{e:2}, {e:3}] (Prelude.Types.foldl [___, ___, !{arg:2}, ((!{arg:2} [!{arg:3}]) [!{e:2}]), !{e:3}]))] Nothing) Prelude.Types.foldl = [{arg:0}, {arg:1}, {arg:2}, {arg:3}, {arg:4}]: (%case !{arg:4} [(%concase Prelude.Types.Nil Just 0 [] !{arg:3}), (%concase Prelude.Types.:: Just 1 [{e:2}, {e:3}] (Prelude.Types.foldl [___, ___, !{arg:2}, ((!{arg:2} [!{arg:3}]) [!{e:2}]), !{e:3}]))] Nothing)
Prelude.Types.Range at Prelude/Types.idr:719:1--725:37 = Constructor tag Just 0 arity 4 Prelude.Types.Range at Prelude/Types.idr:748:1--754:37 = Constructor tag Just 0 arity 4
Prelude.Types.Range implementation at Prelude/Types.idr:751:1--773:48 = [{arg:0}, {arg:1}]: (%con Prelude.Types.Range at Prelude/Types.idr:719:1--725:37 Just 0 [(%lam {arg:4770} (%lam {arg:4771} (Prelude.Types.rangeFromTo [___, !{arg:1}, !{arg:4770}, !{arg:4771}]))), (%lam {arg:4772} (%lam {arg:4773} (%lam {arg:4774} (Prelude.Types.rangeFromThenTo [___, !{arg:1}, !{arg:4772}, !{arg:4773}, !{arg:4774}])))), (%lam {arg:4775} (Prelude.Types.rangeFrom [___, !{arg:1}, !{arg:4775}])), (%lam {arg:4776} (%lam {arg:4777} (Prelude.Types.rangeFromThen [___, !{arg:1}, !{arg:4776}, !{arg:4777}])))]) Prelude.Types.Range implementation at Prelude/Types.idr:780:1--802:48 = [{arg:0}, {arg:1}]: (%con Prelude.Types.Range at Prelude/Types.idr:748:1--754:37 Just 0 [(%lam {arg:5514} (%lam {arg:5515} (Prelude.Types.rangeFromTo [___, !{arg:1}, !{arg:5514}, !{arg:5515}]))), (%lam {arg:5516} (%lam {arg:5517} (%lam {arg:5518} (Prelude.Types.rangeFromThenTo [___, !{arg:1}, !{arg:5516}, !{arg:5517}, !{arg:5518}])))), (%lam {arg:5519} (Prelude.Types.rangeFrom [___, !{arg:1}, !{arg:5519}])), (%lam {arg:5520} (%lam {arg:5521} (Prelude.Types.rangeFromThen [___, !{arg:1}, !{arg:5520}, !{arg:5521}])))])
Prelude.Types.Foldable implementation at Prelude/Types.idr:356:1--365:22 = []: (%con Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:219:1--244:55 Just 0 [(%lam acc (%lam elem (%lam func (%lam init (%lam input (Prelude.Types.foldr [___, ___, !func, !init, !input])))))), (%lam elem (%lam acc (%lam func (%lam init (%lam input (Prelude.Types.foldl [___, ___, !func, !init, !input])))))), (%lam elem (%lam {arg:1123} (Prelude.Types.null [___, !{arg:1123}])))]) Prelude.Types.Foldable implementation at Prelude/Types.idr:385:1--394:22 = []: (%con Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:219:1--244:55 Just 0 [(%lam acc (%lam elem (%lam func (%lam init (%lam input (Prelude.Types.foldr [___, ___, !func, !init, !input])))))), (%lam elem (%lam acc (%lam func (%lam init (%lam input (Prelude.Types.foldl [___, ___, !func, !init, !input])))))), (%lam elem (%lam {arg:1123} (Prelude.Types.null [___, !{arg:1123}])))])
Prelude.Types.takeUntil = [{arg:0}, {arg:1}, {arg:2}]: (%case !{arg:2} [(%concase Prelude.Types.Stream.:: Just 0 [{e:1}, {e:2}] (Prelude.Types.case block in takeUntil [___, !{e:1}, !{e:2}, !{arg:1}, (!{arg:1} [!{e:1}])]))] Nothing) Prelude.Types.takeUntil = [{arg:0}, {arg:1}, {arg:2}]: (%case !{arg:2} [(%concase Prelude.Types.Stream.:: Just 0 [{e:1}, {e:2}] (Prelude.Types.case block in takeUntil [___, !{e:1}, !{e:2}, !{arg:1}, (!{arg:1} [!{e:1}])]))] Nothing)
Prelude.Types.takeBefore = [{arg:0}, {arg:1}, {arg:2}]: (%case !{arg:2} [(%concase Prelude.Types.Stream.:: Just 0 [{e:1}, {e:2}] (Prelude.Types.case block in takeBefore [___, !{e:1}, !{e:2}, !{arg:1}, (!{arg:1} [!{e:1}])]))] Nothing) Prelude.Types.takeBefore = [{arg:0}, {arg:1}, {arg:2}]: (%case !{arg:2} [(%concase Prelude.Types.Stream.:: Just 0 [{e:1}, {e:2}] (Prelude.Types.case block in takeBefore [___, !{e:1}, !{e:2}, !{arg:1}, (!{arg:1} [!{e:1}])]))] Nothing)
Prelude.Types.rangeFromTo = [{arg:0}, {arg:1}]: (%case !{arg:1} [(%concase Prelude.Types.Range at Prelude/Types.idr:719:1--725:37 Just 0 [{e:1}, {e:2}, {e:3}, {e:4}] (%lam {arg:2} (%lam {arg:3} ((!{e:1} [!{arg:2}]) [!{arg:3}]))))] Nothing) Prelude.Types.rangeFromTo = [{arg:0}, {arg:1}]: (%case !{arg:1} [(%concase Prelude.Types.Range at Prelude/Types.idr:748:1--754:37 Just 0 [{e:1}, {e:2}, {e:3}, {e:4}] (%lam {arg:2} (%lam {arg:3} ((!{e:1} [!{arg:2}]) [!{arg:3}]))))] Nothing)
Prelude.Types.prim__integerToNat = [{arg:0}]: (Prelude.Types.case block in prim__integerToNat [!{arg:0}, (%case (<=Integer [0, !{arg:0}]) [(%constcase 0 1)] Just 0)]) Prelude.Types.prim__integerToNat = [{arg:0}]: (Prelude.Types.case block in prim__integerToNat [!{arg:0}, (%case (<=Integer [0, !{arg:0}]) [(%constcase 0 1)] Just 0)])
Prelude.Types.countFrom = [{arg:0}, {arg:1}, {arg:2}]: (%con Prelude.Types.Stream.:: Just 0 [!{arg:1}, (%delay Inf (Prelude.Types.countFrom [___, (!{arg:2} [!{arg:1}]), !{arg:2}]))]) Prelude.Types.countFrom = [{arg:0}, {arg:1}, {arg:2}]: (%con Prelude.Types.Stream.:: Just 0 [!{arg:1}, (%delay Inf (Prelude.Types.countFrom [___, (!{arg:2} [!{arg:1}]), !{arg:2}]))])
Prelude.Types.Z = Constructor tag Just 0 arity 0 Prelude.Types.Z = Constructor tag Just 0 arity 0