mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-29 03:24:02 +03:00
a972778eab
They don't all pass yet, for minor reasons. Coming shortly... Unfortunately the startup overhead for chez is really noticeable here!
50 lines
1.0 KiB
Idris
50 lines
1.0 KiB
Idris
data Maybe a = Nothing
|
|
| Just a
|
|
|
|
infixl 1 >>=
|
|
|
|
(>>=) : Maybe a -> (a -> Maybe b) -> Maybe b
|
|
(>>=) Nothing k = Nothing
|
|
(>>=) (Just x) k = k x
|
|
|
|
data Nat : Type where
|
|
Z : Nat
|
|
S : Nat -> Nat
|
|
|
|
plus : Nat -> Nat -> Nat
|
|
plus Z y = y
|
|
plus (S k) y = S (plus k y)
|
|
|
|
maybeAdd' : Maybe Nat -> Maybe Nat -> Maybe Nat
|
|
maybeAdd' x y
|
|
= x >>= \x' =>
|
|
y >>= \y' =>
|
|
Just (plus x' y')
|
|
|
|
maybeAdd : Maybe Nat -> Maybe Nat -> Maybe Nat
|
|
maybeAdd x y
|
|
= do x' <- x
|
|
y' <- y
|
|
Just (plus x' y')
|
|
|
|
data Either : Type -> Type -> Type where
|
|
Left : a -> Either a b
|
|
Right : b -> Either a b
|
|
|
|
mEmbed : Maybe a -> Maybe (Either String a)
|
|
mEmbed Nothing = Just (Left "FAIL")
|
|
mEmbed (Just x) = Just (Right x)
|
|
|
|
mPatBind : Maybe Nat -> Maybe Nat -> Maybe Nat
|
|
mPatBind x y
|
|
= do Right res <- mEmbed (maybeAdd x y)
|
|
| Left err => Just Z
|
|
Just res
|
|
|
|
mLetBind : Maybe Nat -> Maybe Nat -> Maybe Nat
|
|
mLetBind x y
|
|
= do let Just res = maybeAdd x y
|
|
| Nothing => Just Z
|
|
Just res
|
|
|