Idris2/libs/linear/Data/Linear/Notation.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

27 lines
503 B
Idris

module Data.Linear.Notation
%default total
export infixr 0 -@
||| Infix notation for linear implication
public export
(-@) : Type -> Type -> Type
a -@ b = (1 _ : a) -> b
||| 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)
export prefix 5 !*
||| Prefix notation for the linear unrestricted modality
public export
record (!*) (a : Type) where
constructor MkBang
unrestricted : a