2019-06-13 15:23:21 +03:00
|
|
|
module PrimIO
|
|
|
|
|
|
|
|
import Builtin
|
|
|
|
|
|
|
|
public export
|
|
|
|
data IORes : Type -> Type where
|
|
|
|
MkIORes : (result : a) -> (1 x : %World) -> IORes a
|
|
|
|
|
|
|
|
export
|
|
|
|
data IO : Type -> Type where
|
|
|
|
MkIO : (1 fn : (1 x : %World) -> IORes a) -> IO a
|
|
|
|
|
|
|
|
export
|
|
|
|
io_pure : a -> IO a
|
|
|
|
io_pure x = MkIO (\w => MkIORes x w)
|
|
|
|
|
|
|
|
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')
|
|
|
|
|
Experimenting with a new FFI
Functions can be declared as %foreign with a list of calling
conventions, which a backend will work through until it finds one it can
understand. Currently implemented only in Chez backend. If this works
out, I'll implement it for Racket too, and remove the old primitive
functions.
There's a bit more boiler plate here than before, but it has the benefit
of being more extensible and portable between different back ends.
Some examples, pending proper documentation:
%foreign "C:puts,libc" "scheme:display"
putline : String -> PrimIO ()
%foreign "C:exp, libm.so.6, math.h"
fexp : Double -> Double
%foreign "C:initscr, ncurses_glue.so, ncurses.h"
prim_initscr : PrimIO ()
2019-09-02 19:10:48 +03:00
|
|
|
public export
|
|
|
|
PrimIO : Type -> Type
|
|
|
|
PrimIO a = (1 x : %World) -> IORes a
|
|
|
|
|
2019-06-13 15:23:21 +03:00
|
|
|
%extern prim__putStr : String -> (1 x : %World) -> IORes ()
|
|
|
|
%extern prim__getStr : (1 x : %World) -> IORes String
|
|
|
|
|
2019-09-04 12:25:45 +03:00
|
|
|
-- A pointer representing a given parameter type
|
|
|
|
-- The parameter is a phantom type, to help differentiate between
|
|
|
|
-- different pointer types
|
2019-06-13 15:23:21 +03:00
|
|
|
public export
|
2019-09-04 12:25:45 +03:00
|
|
|
data Ptr : Type -> Type where
|
|
|
|
|
|
|
|
-- A pointer to any type (representing a void* in foreign calls)
|
|
|
|
public export
|
|
|
|
data AnyPtr : Type where
|
2019-06-13 15:23:21 +03:00
|
|
|
|
|
|
|
public export
|
|
|
|
data ThreadID : Type where
|
|
|
|
|
|
|
|
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
|
|
|
|
|
2019-09-29 19:25:26 +03:00
|
|
|
export %inline
|
|
|
|
toPrim : IO a -> PrimIO a
|
|
|
|
toPrim (MkIO fn) = fn
|
|
|
|
|
2019-06-13 15:23:21 +03:00
|
|
|
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)
|
|
|
|
|
|
|
|
export
|
|
|
|
putStr : String -> IO ()
|
|
|
|
putStr str = primIO (prim__putStr str)
|
|
|
|
|
|
|
|
export
|
|
|
|
putStrLn : String -> IO ()
|
|
|
|
putStrLn str = putStr (prim__strAppend str "\n")
|
|
|
|
|
|
|
|
export
|
|
|
|
getLine : IO String
|
|
|
|
getLine = primIO prim__getStr
|
|
|
|
|
|
|
|
export
|
|
|
|
fork : (1 prog : IO ()) -> IO ThreadID
|
|
|
|
fork (MkIO act) = schemeCall ThreadID "blodwen-thread" [act]
|
|
|
|
|
|
|
|
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)
|