2022-01-25 15:25:55 +03:00
|
|
|
module Data.Linear.LEither
|
|
|
|
|
2022-02-02 00:22:16 +03:00
|
|
|
import Data.Linear.Bifunctor
|
|
|
|
import Data.Linear.Interface
|
2022-01-25 15:25:55 +03:00
|
|
|
import Data.Linear.Notation
|
|
|
|
|
|
|
|
%default total
|
|
|
|
|
|
|
|
public export
|
|
|
|
data LEither : (a, b : Type) -> Type where
|
|
|
|
Left : a -@ LEither a b
|
|
|
|
Right : b -@ LEither a b
|
|
|
|
|
|
|
|
export
|
|
|
|
(Consumable a, Consumable b) => Consumable (LEither a b) where
|
|
|
|
consume (Left a) = consume a
|
|
|
|
consume (Right b) = consume b
|
2022-02-02 00:22:16 +03:00
|
|
|
|
|
|
|
export
|
|
|
|
(Duplicable a, Duplicable b) => Duplicable (LEither a b) where
|
|
|
|
duplicate (Left a) = Left <$> duplicate a
|
|
|
|
duplicate (Right b) = Right <$> duplicate b
|