Idris2/libs/linear/Data/Linear/Notation.idr

27 lines
489 B
Idris
Raw Normal View History

module Data.Linear.Notation
%default total
infixr 0 -@
||| Infix notation for linear implication
public export
(-@) : Type -> Type -> Type
a -@ b = (1 _ : a) -> b
2022-02-11 12:28:15 +03:00
||| Linear identity function
public export
id : a -@ a
id x = x
||| Linear function composition
public export
(.) : (b -@ c) -@ (a -@ b) -@ (a -@ c)
(.) f g v = f (g v)
prefix 5 !*
||| Prefix notation for the linear unrestricted modality
public export
record (!*) (a : Type) where
constructor MkBang
unrestricted : a