Idris2/libs/base/Data/Maybe.idr

63 lines
1.6 KiB
Idris
Raw Normal View History

2020-05-18 15:59:07 +03:00
module Data.Maybe
public export
isNothing : Maybe a -> Bool
isNothing Nothing = True
2020-06-30 01:35:54 +03:00
isNothing (Just _) = False
2020-05-18 15:59:07 +03:00
public export
isJust : Maybe a -> Bool
isJust Nothing = False
2020-06-30 01:35:54 +03:00
isJust (Just _) = True
2020-05-18 15:59:07 +03:00
||| Proof that some `Maybe` is actually `Just`
public export
data IsJust : Maybe a -> Type where
ItIsJust : IsJust (Just x)
export
Uninhabited (IsJust Nothing) where
uninhabited ItIsJust impossible
||| Decide whether a 'Maybe' is 'Just'
public export
isItJust : (v : Maybe a) -> Dec (IsJust v)
2020-06-30 01:35:54 +03:00
isItJust (Just _) = Yes ItIsJust
2020-05-18 15:59:07 +03:00
isItJust Nothing = No absurd
||| Convert a `Maybe a` value to an `a` value by providing a default `a` value
||| in the case that the `Maybe` value is `Nothing`.
public export
fromMaybe : (Lazy a) -> Maybe a -> a
fromMaybe def Nothing = def
2020-06-30 01:35:54 +03:00
fromMaybe _ (Just j) = j
2020-05-18 15:59:07 +03:00
2020-05-20 21:11:35 +03:00
||| Returns the `a` value of a `Maybe a` which is proved `Just`.
public export
fromJust : (v : Maybe a) -> IsJust v => a
fromJust (Just x) = x
fromJust Nothing impossible
2020-05-18 15:59:07 +03:00
||| Returns `Just` the given value if the conditional is `True`
||| and `Nothing` if the conditional is `False`.
public export
toMaybe : Bool -> Lazy a -> Maybe a
toMaybe True j = Just j
2020-06-30 01:35:54 +03:00
toMaybe False _ = Nothing
2020-05-18 15:59:07 +03:00
export
2020-06-30 01:35:54 +03:00
justInjective : Just x = Just y -> x = y
2020-05-18 15:59:07 +03:00
justInjective Refl = Refl
||| Convert a `Maybe a` value to an `a` value, using `neutral` in the case
||| that the `Maybe` value is `Nothing`.
public export
lowerMaybe : Monoid a => Maybe a -> a
lowerMaybe Nothing = neutral
lowerMaybe (Just x) = x
||| Returns `Nothing` when applied to `neutral`, and `Just` the value otherwise.
export
raiseToMaybe : (Monoid a, Eq a) => a -> Maybe a
raiseToMaybe x = if x == neutral then Nothing else Just x