mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-18 08:42:11 +03:00
49 lines
1.2 KiB
Idris
49 lines
1.2 KiB
Idris
|
module Data.These
|
||
|
|
||
|
%default total
|
||
|
|
||
|
public export
|
||
|
data These a b = This a | That b | Both a b
|
||
|
|
||
|
public export
|
||
|
fromEither : Either a b -> These a b
|
||
|
fromEither = either This That
|
||
|
|
||
|
public export
|
||
|
these : (a -> c) -> (b -> c) -> (a -> b -> c) -> These a b -> c
|
||
|
these l r lr (This a) = l a
|
||
|
these l r lr (That b) = r b
|
||
|
these l r lr (Both a b) = lr a b
|
||
|
|
||
|
public export
|
||
|
bimap : (f : a -> b) -> (g : c -> d) -> These a c -> These b d
|
||
|
bimap f g (This a) = This (f a)
|
||
|
bimap f g (That b) = That (g b)
|
||
|
bimap f g (Both a b) = Both (f a) (g b)
|
||
|
|
||
|
public export
|
||
|
(Show a, Show b) => Show (These a b) where
|
||
|
showPrec d (This x) = showCon d "This" $ showArg x
|
||
|
showPrec d (That x) = showCon d "That" $ showArg x
|
||
|
showPrec d (Both x y) = showCon d "Both" $ showArg x ++ showArg y
|
||
|
|
||
|
public export
|
||
|
Functor (These a) where
|
||
|
map = bimap id
|
||
|
|
||
|
public export
|
||
|
mapFst : (f : a -> b) -> These a c -> These b c
|
||
|
mapFst f = bimap f id
|
||
|
|
||
|
public export
|
||
|
bifold : Monoid m => These m m -> m
|
||
|
bifold (This a) = a
|
||
|
bifold (That b) = 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) |]
|