Idris2/libs/prelude/PrimIO.idr
Edwin Brady ad632d825d Remove linearity subtyping
It's disappointing to have to do this, but I think necessary because
various issue reports have shown it to be unsound (at least as far as
inference goes) and, at the very least, confusing. This patch brings us
back to the basic rules of QTT.

On the one hand, this makes the 1 multiplicity less useful, because it
means we can't flag arguments as being used exactly once which would be
useful for optimisation purposes as well as precision in the type. On
the other hand, it removes some complexity (and a hack) from
unification, and has the advantage of being correct! Also, I still
consider the 1 multiplicity an experiment.

We can still do interesting things like protocol state tracking, which
is my primary motivation at least.

Ideally, if the 1 multiplicity is going to be more generall useful,
we'll need some kind of way of doing multiplicity polymorphism in the
future. I don't think subtyping is the way (I've pretty much always come
to regret adding some form of subtyping).

Fixes #73 (and maybe some others).
2020-12-27 19:58:35 +00:00

105 lines
2.7 KiB
Idris

module PrimIO
import Builtin
public export
data IORes : Type -> Type where
MkIORes : (result : a) -> (1 x : %World) -> IORes a
||| Idris's primitive IO, for building abstractions on top of.
public export
PrimIO : Type -> Type
PrimIO a = (1 x : %World) -> IORes a
export
data IO : Type -> Type where
MkIO : (1 fn : PrimIO a) -> IO a
export
prim__io_pure : a -> PrimIO a
prim__io_pure x = \w => MkIORes x w
%inline
export
io_pure : a -> IO a
io_pure x = MkIO (\w => MkIORes x w)
export
prim__io_bind : (1 act : PrimIO a) -> (1 k : a -> PrimIO b) -> PrimIO b
prim__io_bind fn k w
= let MkIORes x' w' = fn w in k x' w'
-- There's a special case for inlining this is Compiler.Inline, because
-- the inliner is cautious about case blocks at the moment. Once it's less
-- cautious, add an explicit %inline directive and take out the special case.
-- See also dealing with the newtype optimisation via %World in
-- Compiler.CompileExpr
export
io_bind : (1 act : IO a) -> (1 k : a -> IO b) -> IO b
io_bind (MkIO fn) k
= MkIO (\w => let MkIORes x' w' = fn w
MkIO res = k x' in
res w')
-- A pointer representing a given parameter type
-- The parameter is a phantom type, to help differentiate between
-- different pointer types
public export
data Ptr : Type -> Type where [external]
-- A pointer to any type (representing a void* in foreign calls)
public export
data AnyPtr : Type where [external]
-- As Ptr, but associated with a finaliser that is run on garbage collection
public export
data GCPtr : Type -> Type where [external]
-- As AnyPtr, but associated with a finaliser that is run on garbage collection
public export
data GCAnyPtr : Type where [external]
public export
data ThreadID : Type where [external]
export %inline
fromPrim : (1 fn : (1 x : %World) -> IORes a) -> IO a
fromPrim op = MkIO op
export %inline
toPrim : (1 act : IO a) -> PrimIO a
toPrim (MkIO fn) = fn
%foreign "C:idris2_isNull, libidris2_support"
"javascript:lambda:x=>x===undefined||x===null?1n:0n"
export
prim__nullAnyPtr : AnyPtr -> Int
%foreign "C:idris2_getNull,libidris2_support"
export
prim__getNullAnyPtr : AnyPtr
export
prim__castPtr : AnyPtr -> Ptr t
prim__castPtr = believe_me
export
prim__forgetPtr : Ptr t -> AnyPtr
prim__forgetPtr = believe_me
export %inline
prim__nullPtr : Ptr t -> Int -- can't pass 'type' to a foreign function
prim__nullPtr p = prim__nullAnyPtr (prim__forgetPtr p)
unsafeCreateWorld : (1 f : (1 x : %World) -> a) -> a
unsafeCreateWorld f = f %MkWorld
unsafeDestroyWorld : (1 x : %World) -> a -> a
unsafeDestroyWorld %MkWorld x = x
export
unsafePerformIO : IO a -> a
unsafePerformIO (MkIO f)
= unsafeCreateWorld (\w => case f w of
MkIORes res w' => unsafeDestroyWorld w' res)