2020-07-12 17:54:10 +03:00
|
|
|
module Prelude.IO
|
|
|
|
|
|
|
|
import Builtin
|
|
|
|
import PrimIO
|
|
|
|
import Prelude.Basics
|
|
|
|
import Prelude.Interfaces
|
|
|
|
import Prelude.Show
|
|
|
|
|
|
|
|
%default total
|
|
|
|
|
|
|
|
--------
|
|
|
|
-- IO --
|
|
|
|
--------
|
|
|
|
|
|
|
|
public export
|
|
|
|
Functor IO where
|
2021-06-16 17:22:30 +03:00
|
|
|
map f io = io_bind io $ io_pure . f
|
2020-07-12 17:54:10 +03:00
|
|
|
|
|
|
|
%inline
|
|
|
|
public export
|
|
|
|
Applicative IO where
|
|
|
|
pure x = io_pure x
|
|
|
|
f <*> a
|
|
|
|
= io_bind f (\f' =>
|
|
|
|
io_bind a (\a' =>
|
|
|
|
io_pure (f' a')))
|
|
|
|
|
|
|
|
%inline
|
|
|
|
public export
|
|
|
|
Monad IO where
|
|
|
|
b >>= k = io_bind b k
|
|
|
|
|
|
|
|
public export
|
|
|
|
interface Monad io => HasIO io where
|
2021-05-06 18:32:51 +03:00
|
|
|
constructor MkHasIO
|
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 22:58:35 +03:00
|
|
|
liftIO : IO a -> io a
|
2020-07-12 17:54:10 +03:00
|
|
|
|
2021-01-11 14:24:43 +03:00
|
|
|
public export
|
|
|
|
interface Monad io => HasLinearIO io where
|
2021-05-06 18:32:51 +03:00
|
|
|
constructor MkHasLinearIO
|
2021-01-11 14:24:43 +03:00
|
|
|
liftIO1 : (1 _ : IO a) -> io a
|
|
|
|
|
|
|
|
public export %inline
|
|
|
|
HasLinearIO IO where
|
|
|
|
liftIO1 x = x
|
|
|
|
|
2020-07-12 17:54:10 +03:00
|
|
|
public export %inline
|
2021-01-11 14:24:43 +03:00
|
|
|
HasLinearIO io => HasIO io where
|
|
|
|
liftIO x = liftIO1 x
|
2020-07-12 17:54:10 +03:00
|
|
|
|
|
|
|
export %inline
|
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 22:58:35 +03:00
|
|
|
primIO : HasIO io => (fn : (1 x : %World) -> IORes a) -> io a
|
2020-07-12 17:54:10 +03:00
|
|
|
primIO op = liftIO (fromPrim op)
|
|
|
|
|
2021-01-11 14:24:43 +03:00
|
|
|
export %inline
|
|
|
|
primIO1 : HasLinearIO io => (1 fn : (1 x : %World) -> IORes a) -> io a
|
|
|
|
primIO1 op = liftIO1 (fromPrim op)
|
|
|
|
|
2020-07-12 17:54:10 +03:00
|
|
|
%extern
|
|
|
|
prim__onCollectAny : AnyPtr -> (AnyPtr -> PrimIO ()) -> PrimIO GCAnyPtr
|
|
|
|
%extern
|
|
|
|
prim__onCollect : Ptr t -> (Ptr t -> PrimIO ()) -> PrimIO (GCPtr t)
|
|
|
|
|
|
|
|
export
|
2021-07-16 23:05:52 +03:00
|
|
|
onCollectAny : HasIO io => AnyPtr -> (AnyPtr -> IO ()) -> io GCAnyPtr
|
|
|
|
onCollectAny ptr c = primIO (prim__onCollectAny ptr (\x => toPrim (c x)))
|
2020-07-12 17:54:10 +03:00
|
|
|
|
|
|
|
export
|
2021-07-16 23:05:52 +03:00
|
|
|
onCollect : HasIO io => Ptr t -> (Ptr t -> IO ()) -> io (GCPtr t)
|
|
|
|
onCollect ptr c = primIO (prim__onCollect ptr (\x => toPrim (c x)))
|
2020-07-12 17:54:10 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idris2_getString, libidris2_support, idris_support.h"
|
2020-07-12 17:54:10 +03:00
|
|
|
"javascript:lambda:x=>x"
|
|
|
|
export
|
|
|
|
prim__getString : Ptr String -> String
|
|
|
|
|
|
|
|
%foreign "C:putchar,libc 6"
|
2023-02-13 18:48:07 +03:00
|
|
|
"node:lambda:x=>process.stdout.write(x)"
|
|
|
|
"browser:lambda:x=>console.log(x)"
|
2020-07-12 17:54:10 +03:00
|
|
|
prim__putChar : Char -> (1 x : %World) -> IORes ()
|
2023-02-13 18:48:07 +03:00
|
|
|
|
2020-07-12 17:54:10 +03:00
|
|
|
%foreign "C:getchar,libc 6"
|
2023-02-13 18:48:07 +03:00
|
|
|
"node:support:getChar,support_system_file"
|
2020-07-12 17:54:10 +03:00
|
|
|
%extern prim__getChar : (1 x : %World) -> IORes Char
|
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idris2_getStr, libidris2_support, idris_support.h"
|
2020-07-12 17:54:10 +03:00
|
|
|
"node:support:getStr,support_system_file"
|
|
|
|
prim__getStr : PrimIO String
|
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idris2_putStr, libidris2_support, idris_support.h"
|
2020-07-12 17:54:10 +03:00
|
|
|
"node:lambda:x=>process.stdout.write(x)"
|
2021-08-24 06:44:57 +03:00
|
|
|
"browser:lambda:x=>console.log(x)"
|
2020-07-12 17:54:10 +03:00
|
|
|
prim__putStr : String -> PrimIO ()
|
|
|
|
|
|
|
|
||| Output a string to stdout without a trailing newline.
|
2022-12-01 18:31:00 +03:00
|
|
|
%inline export
|
2020-07-12 17:54:10 +03:00
|
|
|
putStr : HasIO io => String -> io ()
|
|
|
|
putStr str = primIO (prim__putStr str)
|
|
|
|
|
|
|
|
||| Output a string to stdout with a trailing newline.
|
|
|
|
export
|
2022-12-01 18:31:00 +03:00
|
|
|
%inline putStrLn : HasIO io => String -> io ()
|
2020-07-12 17:54:10 +03:00
|
|
|
putStrLn str = putStr (prim__strAppend str "\n")
|
|
|
|
|
|
|
|
||| Read one line of input from stdin, without the trailing newline.
|
2022-12-01 18:31:00 +03:00
|
|
|
%inline export
|
2020-07-12 17:54:10 +03:00
|
|
|
getLine : HasIO io => io String
|
|
|
|
getLine = primIO prim__getStr
|
|
|
|
|
2020-08-08 01:43:51 +03:00
|
|
|
||| Write one single-byte character to stdout.
|
2022-12-01 18:31:00 +03:00
|
|
|
%inline export
|
2020-07-12 17:54:10 +03:00
|
|
|
putChar : HasIO io => Char -> io ()
|
|
|
|
putChar c = primIO (prim__putChar c)
|
|
|
|
|
2020-08-08 01:43:51 +03:00
|
|
|
||| Write one multi-byte character to stdout, with a trailing newline.
|
2022-12-01 18:31:00 +03:00
|
|
|
%inline export
|
2020-07-12 17:54:10 +03:00
|
|
|
putCharLn : HasIO io => Char -> io ()
|
|
|
|
putCharLn c = putStrLn (prim__cast_CharString c)
|
|
|
|
|
2020-08-08 01:43:51 +03:00
|
|
|
||| Read one single-byte character from stdin.
|
2022-12-01 18:31:00 +03:00
|
|
|
%inline export
|
2020-07-12 17:54:10 +03:00
|
|
|
getChar : HasIO io => io Char
|
|
|
|
getChar = primIO prim__getChar
|
|
|
|
|
2020-08-24 19:38:29 +03:00
|
|
|
%foreign "scheme:blodwen-thread"
|
2022-04-27 15:59:32 +03:00
|
|
|
"C:refc_fork"
|
2020-08-25 14:30:57 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__fork : (1 prog : PrimIO ()) -> PrimIO ThreadID
|
2020-07-12 17:54:10 +03:00
|
|
|
|
|
|
|
export
|
|
|
|
fork : (1 prog : IO ()) -> IO ThreadID
|
2020-08-24 19:38:29 +03:00
|
|
|
fork act = fromPrim (prim__fork (toPrim act))
|
2020-07-12 17:54:10 +03:00
|
|
|
|
2021-02-05 19:16:20 +03:00
|
|
|
%foreign "scheme:blodwen-thread-wait"
|
|
|
|
export
|
|
|
|
prim__threadWait : (1 threadID : ThreadID) -> PrimIO ()
|
|
|
|
|
|
|
|
export
|
|
|
|
threadWait : (1 threadID : ThreadID) -> IO ()
|
|
|
|
threadWait threadID = fromPrim (prim__threadWait threadID)
|
|
|
|
|
2020-07-12 17:54:10 +03:00
|
|
|
||| Output something showable to stdout, without a trailing newline.
|
2022-12-01 18:31:00 +03:00
|
|
|
%inline export
|
|
|
|
print : HasIO io => Show a => a -> io ()
|
2022-08-10 23:22:58 +03:00
|
|
|
print = putStr . show
|
2020-07-12 17:54:10 +03:00
|
|
|
|
|
|
|
||| Output something showable to stdout, with a trailing newline.
|
2022-12-01 18:31:00 +03:00
|
|
|
%inline export
|
|
|
|
printLn : HasIO io => Show a => a -> io ()
|
2022-08-10 23:22:58 +03:00
|
|
|
printLn = putStrLn . show
|