Idris2/tests/idris2/basic/basic028/Do.idr
André Videla 75032a7164
Emit warning for fixities with no export modifiers (#3234)
* Emit warning for fixities with no export modifiers

This is to help update all the existing code to program with explicit
fixity export directives in preparation for the behavioral change where
they will become private by default.
2024-04-03 15:41:57 +01:00

49 lines
1.0 KiB
Idris

data Maybe a = Nothing
| Just a
private 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