mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-23 03:32:09 +03:00
21 lines
533 B
Idris
21 lines
533 B
Idris
|
module System.Errno
|
||
|
|
||
|
%default total
|
||
|
|
||
|
%foreign "C:idris2_getErrno, libidris2_support, idris_support.h"
|
||
|
prim__getErrno : PrimIO Int
|
||
|
|
||
|
%foreign "C:idris2_strerror, libidris2_support, idris_support.h"
|
||
|
"node:lambda:errno=>'Error code: '+errno"
|
||
|
prim__strerror : Int -> PrimIO String
|
||
|
|
||
|
||| Fetch libc `errno` global variable.
|
||
|
export
|
||
|
getErrno : HasIO io => io Int
|
||
|
getErrno = primIO prim__getErrno
|
||
|
|
||
|
||| Convert numeric `errno` to string.
|
||
|
export
|
||
|
strerror : Int -> String
|
||
|
strerror errno = unsafePerformIO $ primIO $ prim__strerror errno
|