2022-01-25 15:25:55 +03:00
|
|
|
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
|
|
|
|
|
2022-02-02 00:22:16 +03:00
|
|
|
||| Linear function composition
|
2022-01-25 15:25:55 +03:00
|
|
|
public export
|
2022-02-02 00:22:16 +03:00
|
|
|
(.) : (b -@ c) -@ (a -@ b) -@ (a -@ c)
|
|
|
|
(.) f g v = f (g v)
|
2022-01-25 15:25:55 +03:00
|
|
|
|
2022-02-02 00:22:16 +03:00
|
|
|
prefix 5 !*
|
|
|
|
||| Prefix notation for the linear unrestricted modality
|
|
|
|
public export
|
|
|
|
record (!*) (a : Type) where
|
|
|
|
constructor MkBang
|
|
|
|
unrestricted : a
|