mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-20 10:02:03 +03:00
1617d95961
* add `strerror` function * move `getErrno` to `System.Errno` * use `strerror` in `Show FileError` * on node there's no access to `strerror`, so `strerror` just converts the number to string
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
|