2020-05-18 15:59:07 +03:00
|
|
|
module Data.Either
|
|
|
|
|
2021-11-26 13:55:17 +03:00
|
|
|
import public Control.Function
|
2020-09-22 17:07:40 +03:00
|
|
|
import Data.List1
|
|
|
|
|
2020-06-30 01:35:34 +03:00
|
|
|
%default total
|
|
|
|
|
2020-09-22 17:07:40 +03:00
|
|
|
--------------------------------------------------------------------------------
|
|
|
|
-- Checking for a specific constructor
|
|
|
|
|
|
|
|
||| Extract the Left value, if possible
|
|
|
|
public export
|
|
|
|
getLeft : Either a b -> Maybe a
|
|
|
|
getLeft (Left a) = Just a
|
|
|
|
getLeft _ = Nothing
|
|
|
|
|
|
|
|
||| Extract the Right value, if possible
|
|
|
|
public export
|
|
|
|
getRight : Either a b -> Maybe b
|
|
|
|
getRight (Right b) = Just b
|
|
|
|
getRight _ = Nothing
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
||| True if the argument is Left, False otherwise
|
|
|
|
public export
|
|
|
|
isLeft : Either a b -> Bool
|
2020-06-30 01:35:44 +03:00
|
|
|
isLeft (Left _) = True
|
|
|
|
isLeft (Right _) = False
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| True if the argument is Right, False otherwise
|
|
|
|
public export
|
|
|
|
isRight : Either a b -> Bool
|
2020-06-30 01:35:44 +03:00
|
|
|
isRight (Left _) = False
|
|
|
|
isRight (Right _) = True
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 22:41:23 +03:00
|
|
|
||| Proof that an `Either` is actually a Right value
|
|
|
|
public export
|
|
|
|
data IsRight : Either a b -> Type where
|
|
|
|
ItIsRight : IsRight (Right x)
|
|
|
|
|
|
|
|
export
|
|
|
|
Uninhabited (IsRight (Left x)) where
|
|
|
|
uninhabited ItIsRight impossible
|
|
|
|
|
|
|
|
||| Proof that an `Either` is actually a Left value
|
|
|
|
public export
|
|
|
|
data IsLeft : Either a b -> Type where
|
|
|
|
ItIsLeft : IsLeft (Left x)
|
|
|
|
|
|
|
|
export
|
|
|
|
Uninhabited (IsLeft (Right x)) where
|
|
|
|
uninhabited ItIsLeft impossible
|
|
|
|
|
2020-09-22 17:07:40 +03:00
|
|
|
--------------------------------------------------------------------------------
|
|
|
|
-- Grouping values
|
|
|
|
|
2021-08-09 17:36:51 +03:00
|
|
|
||| Compress the list of Lefts and Rights by accumulating
|
|
|
|
||| all of the lefts and rights into non-empty blocks.
|
|
|
|
export
|
|
|
|
compress : List (Either a b) -> List (Either (List1 a) (List1 b))
|
|
|
|
compress [] = []
|
|
|
|
compress (Left x :: abs) = compressLefts (singleton x) abs where
|
2020-09-22 17:07:40 +03:00
|
|
|
compressLefts : List1 a -> List (Either a b) -> List (Either (List1 a) (List1 b))
|
|
|
|
compressLefts acc (Left a :: abs) = compressLefts (cons a acc) abs
|
|
|
|
compressLefts acc abs = Left (reverse acc) :: compress abs
|
2021-08-09 17:36:51 +03:00
|
|
|
compress (Right y :: abs) = compressRights (singleton y) abs where
|
2020-09-22 17:07:40 +03:00
|
|
|
compressRights : List1 b -> List (Either a b) -> List (Either (List1 a) (List1 b))
|
|
|
|
compressRights acc (Right b :: abs) = compressRights (cons b acc) abs
|
|
|
|
compressRights acc abs = Right (reverse acc) :: compress abs
|
|
|
|
|
|
|
|
||| Decompress a compressed list. This is the left inverse of `compress` but not its
|
|
|
|
||| right inverse because nothing forces the input to be maximally compressed!
|
|
|
|
export
|
|
|
|
decompress : List (Either (List1 a) (List1 b)) -> List (Either a b)
|
2021-08-09 17:36:51 +03:00
|
|
|
decompress = concatMap $ \case
|
|
|
|
Left as => map Left $ forget as
|
2020-09-22 17:07:40 +03:00
|
|
|
Right bs => map Right $ forget bs
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
||| Keep the payloads of all Left constructors in a list of Eithers
|
|
|
|
public export
|
|
|
|
lefts : List (Either a b) -> List a
|
2020-06-30 01:35:44 +03:00
|
|
|
lefts [] = []
|
|
|
|
lefts (Left l :: xs) = l :: lefts xs
|
|
|
|
lefts (Right _ :: xs) = lefts xs
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Keep the payloads of all Right constructors in a list of Eithers
|
|
|
|
public export
|
|
|
|
rights : List (Either a b) -> List b
|
2020-06-30 01:35:44 +03:00
|
|
|
rights [] = []
|
|
|
|
rights (Left _ :: xs) = rights xs
|
|
|
|
rights (Right r :: xs) = r :: rights xs
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Split a list of Eithers into a list of the left elements and a list of the right elements
|
|
|
|
public export
|
|
|
|
partitionEithers : List (Either a b) -> (List a, List b)
|
|
|
|
partitionEithers l = (lefts l, rights l)
|
|
|
|
|
|
|
|
||| Remove a "useless" Either by collapsing the case distinction
|
|
|
|
public export
|
|
|
|
fromEither : Either a a -> a
|
|
|
|
fromEither (Left l) = l
|
|
|
|
fromEither (Right r) = r
|
|
|
|
|
|
|
|
||| Right becomes left and left becomes right
|
|
|
|
public export
|
|
|
|
mirror : Either a b -> Either b a
|
|
|
|
mirror (Left x) = Right x
|
|
|
|
mirror (Right x) = Left x
|
|
|
|
|
2023-10-16 20:21:00 +03:00
|
|
|
public export
|
|
|
|
Zippable (Either a) where
|
|
|
|
zipWith f x y = [| f x y |]
|
|
|
|
zipWith3 f x y z = [| f x y z |]
|
|
|
|
|
|
|
|
unzipWith f (Left e) = (Left e, Left e)
|
|
|
|
unzipWith f (Right xy) = let (x, y) = f xy in (Right x, Right y)
|
|
|
|
|
|
|
|
unzipWith3 f (Left e) = (Left e, Left e, Left e)
|
|
|
|
unzipWith3 f (Right xyz) = let (x, y, z) = f xyz in (Right x, Right y, Right z)
|
|
|
|
|
2020-09-22 17:07:40 +03:00
|
|
|
--------------------------------------------------------------------------------
|
|
|
|
-- Bifunctor
|
|
|
|
|
|
|
|
export
|
|
|
|
pushInto : c -> Either a b -> Either (c, a) (c, b)
|
|
|
|
pushInto c = bimap (\ a => (c, a)) (\ b => (c, b))
|
|
|
|
-- ^ not using sections to keep it backwards compatible
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
--------------------------------------------------------------------------------
|
|
|
|
-- Conversions
|
|
|
|
--------------------------------------------------------------------------------
|
|
|
|
|
|
|
|
||| Convert a Maybe to an Either by using a default value in case of Nothing
|
|
|
|
||| @ e the default value
|
|
|
|
public export
|
|
|
|
maybeToEither : (def : Lazy e) -> Maybe a -> Either e a
|
|
|
|
maybeToEither def (Just j) = Right j
|
|
|
|
maybeToEither def Nothing = Left def
|
|
|
|
|
|
|
|
||| Convert an Either to a Maybe from Right injection
|
|
|
|
public export
|
|
|
|
eitherToMaybe : Either e a -> Maybe a
|
|
|
|
eitherToMaybe (Left _) = Nothing
|
|
|
|
eitherToMaybe (Right x) = Just x
|
|
|
|
|
2021-10-18 01:35:28 +03:00
|
|
|
export
|
|
|
|
eitherMapFusion : (f : _) -> (g : _) -> (p : _) -> (e : Either a b) -> either f g (map p e) = either f (g . p) e
|
|
|
|
eitherMapFusion f g p $ Left x = Refl
|
|
|
|
eitherMapFusion f g p $ Right x = Refl
|
|
|
|
|
|
|
|
export
|
|
|
|
eitherBimapFusion : (f : _) -> (g : _) -> (p : _) -> (q : _) -> (e : _) -> either f g (bimap p q e) = either (f . p) (g . q) e
|
|
|
|
eitherBimapFusion f g p q $ Left z = Refl
|
|
|
|
eitherBimapFusion f g p q $ Right z = Refl
|
2021-11-26 13:55:17 +03:00
|
|
|
|
|
|
|
-- Injectivity of constructors
|
|
|
|
|
|
|
|
||| Left is injective
|
|
|
|
export
|
|
|
|
Injective Left where
|
|
|
|
injective Refl = Refl
|
|
|
|
|
|
|
|
||| Right is injective
|
|
|
|
export
|
|
|
|
Injective Right where
|
|
|
|
injective Refl = Refl
|