Idris2/libs/base/Data/These.idr

108 lines
2.3 KiB
Idris
Raw Normal View History

2020-05-17 17:56:45 +03:00
module Data.These
import Control.Function
2020-05-17 17:56:45 +03:00
%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
fromThis : These a b -> Maybe a
fromThis (This a) = Just a
fromThis (That _) = Nothing
fromThis (Both a _) = Just a
public export
fromThat : These a b -> Maybe b
fromThat (This _) = Nothing
fromThat (That b) = Just b
fromThat (Both _ b) = Just b
2020-05-17 17:56:45 +03:00
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
2021-10-17 16:07:20 +03:00
public export
swap : These a b -> These b a
swap (This a) = That a
swap (That b) = This b
swap (Both a b) = Both b a
export
Injective This where
injective Refl = Refl
export
Injective That where
injective Refl = Refl
export
{x : _} -> Injective (Both x) where
injective Refl = Refl
export
{y : _} -> Injective (`Both` y) where
injective Refl = Refl
export
Biinjective Both where
biinjective Refl = (Refl, Refl)
2020-05-17 17:56:45 +03:00
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
Eq a => Eq b => Eq (These a b) where
This x == This x' = x == x'
That x == That x' = x == x'
Both x y == Both x' y' = x == x' && y == y'
_ == _ = False
2020-09-30 12:51:07 +03:00
%inline
2020-05-17 17:56:45 +03:00
public export
2020-09-30 12:51:07 +03:00
Bifunctor These where
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)
2020-05-17 17:56:45 +03:00
%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) |]
2020-09-30 12:51:07 +03:00
%inline
2020-05-17 17:56:45 +03:00
public export
2020-09-30 12:51:07 +03:00
Functor (These a) where
map = mapSnd
2020-05-17 17:56:45 +03:00
public export
bifold : Semigroup m => These m m -> m
2020-05-17 17:56:45 +03:00
bifold (This a) = a
bifold (That b) = b
bifold (Both a b) = a <+> b