Idris2/libs/base/Data/Singleton.idr

26 lines
578 B
Idris
Raw Normal View History

2022-01-19 14:10:05 +03:00
module Data.Singleton
||| The type containing only a particular value.
||| This is useful for calculating type-level information at runtime.
public export
data Singleton : a -> Type where
Val : (x : a) -> Singleton x
public export %inline
unVal : Singleton {a} x -> a
unVal $ Val x = x
public export %inline
(.unVal) : Singleton {a} x -> a
(.unVal) = unVal
2022-01-19 14:10:05 +03:00
-- pure and <*> implementations for idiom bracket notation
public export
pure : (x : a) -> Singleton x
pure = Val
public export
(<*>) : Singleton f -> Singleton x -> Singleton (f x)
Val f <*> Val x = Val (f x)