Idris2Doc : Data.Either

Data.Either(source)

Reexports

importpublic Control.Function

Definitions

getLeft : Eitherab->Maybea
  Extract the Left value, if possible

Totality: total
Visibility: public export
getRight : Eitherab->Maybeb
  Extract the Right value, if possible

Totality: total
Visibility: public export
isLeft : Eitherab->Bool
  True if the argument is Left, False otherwise

Totality: total
Visibility: public export
isRight : Eitherab->Bool
  True if the argument is Right, False otherwise

Totality: total
Visibility: public export
dataIsRight : Eitherab->Type
  Proof that an `Either` is actually a Right value

Totality: total
Visibility: public export
Constructor: 
ItIsRight : IsRight (Rightx)

Hint: 
Uninhabited (IsRight (Leftx))
fromRight : (e : Eitherlr) -> {auto0_ : IsRighte} ->r
  Returns the `r` value of an `Either l r` which is proved `Right`.

Totality: total
Visibility: public export
dataIsLeft : Eitherab->Type
  Proof that an `Either` is actually a Left value

Totality: total
Visibility: public export
Constructor: 
ItIsLeft : IsLeft (Leftx)

Hint: 
Uninhabited (IsLeft (Rightx))
fromLeft : (e : Eitherlr) -> {auto0_ : IsLefte} ->l
  Returns the `l` value of an `Either l r` which is proved `Left`.

Totality: total
Visibility: public export
compress : List (Eitherab) ->List (Either (List1a) (List1b))
  Compress the list of Lefts and Rights by accumulating
all of the lefts and rights into non-empty blocks.

Totality: total
Visibility: export
decompress : List (Either (List1a) (List1b)) ->List (Eitherab)
  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!

Totality: total
Visibility: export
lefts : List (Eitherab) ->Lista
  Keep the payloads of all Left constructors in a list of Eithers

Totality: total
Visibility: public export
rights : List (Eitherab) ->Listb
  Keep the payloads of all Right constructors in a list of Eithers

Totality: total
Visibility: public export
partitionEithers : List (Eitherab) -> (Lista, Listb)
  Split a list of Eithers into a list of the left elements and a list of the right elements

Totality: total
Visibility: public export
fromEither : Eitheraa->a
  Remove a "useless" Either by collapsing the case distinction

Totality: total
Visibility: public export
mirror : Eitherab->Eitherba
  Right becomes left and left becomes right

Totality: total
Visibility: public export
pushInto : c->Eitherab->Either (c, a) (c, b)
Totality: total
Visibility: export
maybeToEither : Lazy e->Maybea->Eitherea
  Convert a Maybe to an Either by using a default value in case of Nothing
@ e the default value

Totality: total
Visibility: public export
eitherToMaybe : Eitherea->Maybea
  Convert an Either to a Maybe from Right injection

Totality: total
Visibility: public export
eitherMapFusion : (f : (a->{a:5010})) -> (g : ({b:5008}->{a:5010})) -> (p : (b->{b:5008})) -> (e : Eitherab) ->either (Delay f) (Delay g) (mappe) =either (Delay f) (Delay (g.p)) e
Totality: total
Visibility: export
eitherBimapFusion : (f : ({a:5108}->{a:5114})) -> (g : ({b:5109}->{a:5114})) -> (p : ({a:5106}->{a:5108})) -> (q : ({b:5107}->{b:5109})) -> (e : Either{a:5106}{b:5107}) ->either (Delay f) (Delay g) (bimappqe) =either (Delay (f.p)) (Delay (g.q)) e
Totality: total
Visibility: export