Idris2-boot/libs/prelude/PrimIO.idr
2020-04-22 01:40:40 +06:00

132 lines
3.5 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
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'
%inline
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')
%extern prim__putStr : String -> (1 x : %World) -> IORes ()
%extern prim__getStr : (1 x : %World) -> IORes String
%extern prim__putChar : Char -> (1 x : %World) -> IORes ()
%extern prim__getChar : (1 x : %World) -> IORes Char
-- 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]
public export
data ThreadID : Type where [external]
public export
data FArgList : Type where
Nil : FArgList
(::) : {a : Type} -> (1 arg : a) -> (1 args : FArgList) -> FArgList
%extern prim__cCall : (ret : Type) -> String -> (1 args : FArgList) ->
(1 x : %World) -> IORes ret
%extern prim__schemeCall : (ret : Type) -> String -> (1 args : FArgList) ->
(1 x : %World) -> IORes ret
export %inline
primIO : (1 fn : (1 x : %World) -> IORes a) -> IO a
primIO op = MkIO op
export %inline
toPrim : (1 act : IO a) -> PrimIO a
toPrim (MkIO fn) = fn
export %inline
schemeCall : (ret : Type) -> String -> (1 args : FArgList) -> IO ret
schemeCall ret fn args = primIO (prim__schemeCall ret fn args)
export %inline
cCall : (ret : Type) -> String -> FArgList -> IO ret
cCall ret fn args = primIO (prim__cCall ret fn args)
||| Output a string to stdout without a trailing newline.
export
putStr : String -> IO ()
putStr str = primIO (prim__putStr str)
||| Output a string to stdout with a trailing newline.
export
putStrLn : String -> IO ()
putStrLn str = putStr (prim__strAppend str "\n")
||| Read one line of input from stdin, without the trailing newline.
export
getLine : IO String
getLine = primIO prim__getStr
||| Write a single character to stdout.
export
putChar : Char -> IO ()
putChar c = primIO (prim__putChar c)
||| Write a single character to stdout, with a trailing newline.
export
putCharLn : Char -> IO ()
putCharLn c = putStrLn (prim__cast_CharString c)
||| Read a single character from stdin.
export
getChar : IO Char
getChar = primIO prim__getChar
export
fork : (1 prog : IO ()) -> IO ThreadID
fork (MkIO act) = schemeCall ThreadID "blodwen-thread" [act]
export
prim_fork : (1 prog : PrimIO ()) -> PrimIO ThreadID
prim_fork act w = prim__schemeCall ThreadID "blodwen-thread" [act] w
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)