2020-05-18 15:59:07 +03:00
|
|
|
module System.File
|
|
|
|
|
|
|
|
import Data.List
|
|
|
|
import Data.Strings
|
2020-06-11 16:53:54 +03:00
|
|
|
import System.Info
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
public export
|
|
|
|
data Mode = Read | WriteTruncate | Append | ReadWrite | ReadWriteTruncate | ReadAppend
|
|
|
|
|
|
|
|
public export
|
|
|
|
FilePtr : Type
|
|
|
|
FilePtr = AnyPtr
|
|
|
|
|
|
|
|
support : String -> String
|
|
|
|
support fn = "C:" ++ fn ++ ", libidris2_support"
|
|
|
|
|
|
|
|
libc : String -> String
|
|
|
|
libc fn = "C:" ++ fn ++ ", libc 6"
|
|
|
|
|
|
|
|
%foreign support "idris2_openFile"
|
2020-07-06 18:58:02 +03:00
|
|
|
"node:support:openFile,support_system_file"
|
2020-07-07 13:44:37 +03:00
|
|
|
prim__open : String -> String -> PrimIO FilePtr
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_closeFile"
|
2020-07-06 18:58:02 +03:00
|
|
|
"node:lambdaRequire:fs:(fp) => __require_fs.closeSync(fp.fd)"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim__close : FilePtr -> PrimIO ()
|
|
|
|
|
|
|
|
%foreign support "idris2_fileError"
|
2020-06-18 01:29:54 +03:00
|
|
|
"node:lambda:x=>(x===1n?BigInt(1):BigInt(0))"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim_error : FilePtr -> PrimIO Int
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_fileErrno"
|
2020-06-21 19:54:50 +03:00
|
|
|
"node:lambda:()=>-BigInt(process.__lasterr.errno)"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim_fileErrno : PrimIO Int
|
|
|
|
|
|
|
|
%foreign support "idris2_readLine"
|
2020-07-06 18:58:02 +03:00
|
|
|
"node:support:readLine,support_system_file"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim__readLine : FilePtr -> PrimIO (Ptr String)
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_readChars"
|
|
|
|
prim__readChars : Int -> FilePtr -> PrimIO (Ptr String)
|
2020-06-21 08:24:45 +03:00
|
|
|
%foreign "C:fgetc,libc 6"
|
2020-05-24 00:19:10 +03:00
|
|
|
prim__readChar : FilePtr -> PrimIO Int
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_writeLine"
|
2020-06-21 19:54:50 +03:00
|
|
|
"node:lambdaRequire:fs:(filePtr, line) => __require_fs.writeSync(filePtr.fd, line, undefined, 'utf-8')"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim__writeLine : FilePtr -> String -> PrimIO Int
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_eof"
|
2020-06-21 19:54:50 +03:00
|
|
|
"node:lambda:x=>(x.eof?1n:0n)"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim__eof : FilePtr -> PrimIO Int
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign "C:fflush,libc 6"
|
|
|
|
prim__flush : FilePtr -> PrimIO Int
|
2020-05-23 19:42:05 +03:00
|
|
|
%foreign support "idris2_popen"
|
|
|
|
prim__popen : String -> String -> PrimIO FilePtr
|
2020-05-28 17:12:16 +03:00
|
|
|
%foreign support "idris2_pclose"
|
2020-05-23 19:42:05 +03:00
|
|
|
prim__pclose : FilePtr -> PrimIO ()
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2020-05-21 15:32:35 +03:00
|
|
|
%foreign support "idris2_removeFile"
|
|
|
|
prim__removeFile : String -> PrimIO Int
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_fileSize"
|
2020-06-18 01:29:54 +03:00
|
|
|
"node:lambdaRequire:fs:fp=>__require_fs.fstatSync(fp.fd, {bigint: true}).size"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim__fileSize : FilePtr -> PrimIO Int
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_fileSize"
|
|
|
|
prim__fPoll : FilePtr -> PrimIO Int
|
|
|
|
|
|
|
|
%foreign support "idris2_fileAccessTime"
|
|
|
|
prim__fileAccessTime : FilePtr -> PrimIO Int
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_fileModifiedTime"
|
2020-06-18 01:29:54 +03:00
|
|
|
"node:lambdaRequire:fs:fp=>__require_fs.fstatSync(fp.fd, {bigint: true}).mtimeMs / 1000n"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim__fileModifiedTime : FilePtr -> PrimIO Int
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_fileStatusTime"
|
|
|
|
prim__fileStatusTime : FilePtr -> PrimIO Int
|
|
|
|
|
|
|
|
%foreign support "idris2_stdin"
|
2020-06-18 01:29:54 +03:00
|
|
|
"node:lambda:x=>({fd:0, buffer: Buffer.alloc(0), name:'<stdin>', eof: false})"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim__stdin : FilePtr
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_stdout"
|
2020-06-18 01:29:54 +03:00
|
|
|
"node:lambda:x=>({fd:1, buffer: Buffer.alloc(0), name:'<stdout>', eof: false})"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim__stdout : FilePtr
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_stderr"
|
2020-06-18 01:29:54 +03:00
|
|
|
"node:lambda:x=>({fd:2, buffer: Buffer.alloc(0), name:'<stderr>', eof: false})"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim__stderr : FilePtr
|
|
|
|
|
|
|
|
%foreign libc "chmod"
|
2020-07-06 18:58:02 +03:00
|
|
|
"node:support:chmod,support_system_file"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim__chmod : String -> Int -> PrimIO Int
|
|
|
|
|
|
|
|
modeStr : Mode -> String
|
2020-06-11 16:53:54 +03:00
|
|
|
modeStr Read = if isWindows then "rb" else "r"
|
|
|
|
modeStr WriteTruncate = if isWindows then "wb" else "w"
|
|
|
|
modeStr Append = if isWindows then "ab" else "a"
|
|
|
|
modeStr ReadWrite = if isWindows then "rb+" else "r+"
|
|
|
|
modeStr ReadWriteTruncate = if isWindows then "wb+" else "w+"
|
|
|
|
modeStr ReadAppend = if isWindows then "ab+" else "a+"
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
public export
|
|
|
|
data FileError = GenericFileError Int -- errno
|
|
|
|
| FileReadError
|
|
|
|
| FileWriteError
|
|
|
|
| FileNotFound
|
|
|
|
| PermissionDenied
|
|
|
|
| FileExists
|
|
|
|
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
returnError : HasIO io => io (Either FileError a)
|
2020-05-18 15:59:07 +03:00
|
|
|
returnError
|
|
|
|
= do err <- primIO prim_fileErrno
|
|
|
|
case err of
|
|
|
|
0 => pure $ Left FileReadError
|
|
|
|
1 => pure $ Left FileWriteError
|
|
|
|
2 => pure $ Left FileNotFound
|
|
|
|
3 => pure $ Left PermissionDenied
|
|
|
|
4 => pure $ Left FileExists
|
|
|
|
_ => pure $ Left (GenericFileError (err-5))
|
|
|
|
|
|
|
|
export
|
|
|
|
Show FileError where
|
|
|
|
show (GenericFileError errno) = "File error: " ++ show errno
|
|
|
|
show FileReadError = "File Read Error"
|
|
|
|
show FileWriteError = "File Write Error"
|
|
|
|
show FileNotFound = "File Not Found"
|
|
|
|
show PermissionDenied = "Permission Denied"
|
|
|
|
show FileExists = "File Exists"
|
|
|
|
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
ok : HasIO io => a -> io (Either FileError a)
|
2020-05-18 15:59:07 +03:00
|
|
|
ok x = pure (Right x)
|
|
|
|
|
|
|
|
public export
|
|
|
|
data File : Type where
|
|
|
|
FHandle : FilePtr -> File
|
|
|
|
|
|
|
|
export
|
|
|
|
stdin : File
|
|
|
|
stdin = FHandle prim__stdin
|
|
|
|
|
|
|
|
export
|
|
|
|
stdout : File
|
|
|
|
stdout = FHandle prim__stdout
|
|
|
|
|
|
|
|
export
|
|
|
|
stderr : File
|
|
|
|
stderr = FHandle prim__stderr
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
openFile : HasIO io => String -> Mode -> io (Either FileError File)
|
2020-05-18 15:59:07 +03:00
|
|
|
openFile f m
|
2020-07-07 13:44:37 +03:00
|
|
|
= do res <- primIO (prim__open f (modeStr m))
|
2020-05-18 15:59:07 +03:00
|
|
|
if prim__nullAnyPtr res /= 0
|
|
|
|
then returnError
|
|
|
|
else ok (FHandle res)
|
|
|
|
|
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
closeFile : HasIO io => File -> io ()
|
2020-05-18 15:59:07 +03:00
|
|
|
closeFile (FHandle f) = primIO (prim__close f)
|
|
|
|
|
2020-05-19 19:03:18 +03:00
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fileError : HasIO io => File -> io Bool
|
2020-05-19 19:03:18 +03:00
|
|
|
fileError (FHandle f)
|
|
|
|
= do x <- primIO $ prim_error f
|
|
|
|
pure (x /= 0)
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fGetLine : HasIO io => (h : File) -> io (Either FileError String)
|
2020-05-18 15:59:07 +03:00
|
|
|
fGetLine (FHandle f)
|
|
|
|
= do res <- primIO (prim__readLine f)
|
|
|
|
if prim__nullPtr res /= 0
|
|
|
|
then returnError
|
|
|
|
else ok (prim__getString res)
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fGetChars : HasIO io => (h : File) -> Int -> io (Either FileError String)
|
2020-05-18 15:59:07 +03:00
|
|
|
fGetChars (FHandle f) max
|
|
|
|
= do res <- primIO (prim__readChars max f)
|
|
|
|
if prim__nullPtr res /= 0
|
|
|
|
then returnError
|
|
|
|
else ok (prim__getString res)
|
|
|
|
|
2020-05-19 19:03:18 +03:00
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fGetChar : HasIO io => (h : File) -> io (Either FileError Char)
|
2020-05-19 19:03:18 +03:00
|
|
|
fGetChar (FHandle h)
|
|
|
|
= do c <- primIO (prim__readChar h)
|
|
|
|
ferr <- primIO (prim_error h)
|
|
|
|
if (ferr /= 0)
|
|
|
|
then returnError
|
2020-05-24 00:19:10 +03:00
|
|
|
else ok (cast c)
|
2020-05-19 19:03:18 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fPutStr : HasIO io => (h : File) -> String -> io (Either FileError ())
|
2020-05-18 15:59:07 +03:00
|
|
|
fPutStr (FHandle f) str
|
|
|
|
= do res <- primIO (prim__writeLine f str)
|
|
|
|
if res == 0
|
|
|
|
then returnError
|
|
|
|
else ok ()
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fPutStrLn : HasIO io => (h : File) -> String -> io (Either FileError ())
|
2020-05-18 15:59:07 +03:00
|
|
|
fPutStrLn f str = fPutStr f (str ++ "\n")
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fEOF : HasIO io => (h : File) -> io Bool
|
2020-05-18 15:59:07 +03:00
|
|
|
fEOF (FHandle f)
|
|
|
|
= do res <- primIO (prim__eof f)
|
|
|
|
pure (res /= 0)
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fflush : HasIO io => (h : File) -> io ()
|
2020-05-18 15:59:07 +03:00
|
|
|
fflush (FHandle f)
|
|
|
|
= do primIO (prim__flush f)
|
|
|
|
pure ()
|
|
|
|
|
2020-05-23 19:42:05 +03:00
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
popen : HasIO io => String -> Mode -> io (Either FileError File)
|
2020-05-23 19:42:05 +03:00
|
|
|
popen f m = do
|
|
|
|
ptr <- primIO (prim__popen f (modeStr m))
|
|
|
|
if prim__nullAnyPtr ptr /= 0
|
|
|
|
then returnError
|
|
|
|
else pure (Right (FHandle ptr))
|
|
|
|
|
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
pclose : HasIO io => File -> io ()
|
2020-05-23 19:42:05 +03:00
|
|
|
pclose (FHandle h) = primIO (prim__pclose h)
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fileAccessTime : HasIO io => (h : File) -> io (Either FileError Int)
|
2020-05-18 15:59:07 +03:00
|
|
|
fileAccessTime (FHandle f)
|
|
|
|
= do res <- primIO (prim__fileAccessTime f)
|
|
|
|
if res > 0
|
|
|
|
then ok res
|
|
|
|
else returnError
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fileModifiedTime : HasIO io => (h : File) -> io (Either FileError Int)
|
2020-05-18 15:59:07 +03:00
|
|
|
fileModifiedTime (FHandle f)
|
|
|
|
= do res <- primIO (prim__fileModifiedTime f)
|
|
|
|
if res > 0
|
|
|
|
then ok res
|
|
|
|
else returnError
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fileStatusTime : HasIO io => (h : File) -> io (Either FileError Int)
|
2020-05-18 15:59:07 +03:00
|
|
|
fileStatusTime (FHandle f)
|
|
|
|
= do res <- primIO (prim__fileStatusTime f)
|
|
|
|
if res > 0
|
|
|
|
then ok res
|
|
|
|
else returnError
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
removeFile : HasIO io => String -> io (Either FileError ())
|
2020-05-21 15:32:35 +03:00
|
|
|
removeFile fname
|
|
|
|
= do res <- primIO (prim__removeFile fname)
|
2020-05-18 15:59:07 +03:00
|
|
|
if res == 0
|
|
|
|
then ok ()
|
|
|
|
else returnError
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fileSize : HasIO io => (h : File) -> io (Either FileError Int)
|
2020-05-18 15:59:07 +03:00
|
|
|
fileSize (FHandle f)
|
|
|
|
= do res <- primIO (prim__fileSize f)
|
|
|
|
if res >= 0
|
|
|
|
then ok res
|
|
|
|
else returnError
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
fPoll : HasIO io => File -> io Bool
|
2020-05-18 15:59:07 +03:00
|
|
|
fPoll (FHandle f)
|
|
|
|
= do p <- primIO (prim__fPoll f)
|
|
|
|
pure (p > 0)
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
readFile : HasIO io => String -> io (Either FileError String)
|
2020-05-18 15:59:07 +03:00
|
|
|
readFile file
|
|
|
|
= do Right h <- openFile file Read
|
|
|
|
| Left err => returnError
|
|
|
|
Right content <- read [] h
|
|
|
|
| Left err => do closeFile h
|
|
|
|
returnError
|
|
|
|
closeFile h
|
|
|
|
pure (Right (fastAppend content))
|
|
|
|
where
|
2020-06-21 03:18:43 +03:00
|
|
|
read : List String -> File -> io (Either FileError (List String))
|
2020-05-18 15:59:07 +03:00
|
|
|
read acc h
|
|
|
|
= do eof <- fEOF h
|
|
|
|
if eof
|
|
|
|
then pure (Right (reverse acc))
|
|
|
|
else
|
|
|
|
do Right str <- fGetLine h
|
|
|
|
| Left err => returnError
|
|
|
|
read (str :: acc) h
|
|
|
|
|
|
|
|
||| Write a string to a file
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
writeFile : HasIO io =>
|
2020-06-21 03:18:43 +03:00
|
|
|
(filepath : String) -> (contents : String) ->
|
|
|
|
io (Either FileError ())
|
2020-05-18 15:59:07 +03:00
|
|
|
writeFile fn contents = do
|
|
|
|
Right h <- openFile fn WriteTruncate
|
|
|
|
| Left err => pure (Left err)
|
|
|
|
Right () <- fPutStr h contents
|
|
|
|
| Left err => do closeFile h
|
|
|
|
pure (Left err)
|
|
|
|
closeFile h
|
|
|
|
pure (Right ())
|
|
|
|
|
2020-05-21 15:32:35 +03:00
|
|
|
namespace FileMode
|
2020-05-18 15:59:07 +03:00
|
|
|
public export
|
|
|
|
data FileMode = Read | Write | Execute
|
|
|
|
|
|
|
|
public export
|
|
|
|
record Permissions where
|
|
|
|
constructor MkPermissions
|
|
|
|
user : List FileMode
|
|
|
|
group : List FileMode
|
|
|
|
others : List FileMode
|
|
|
|
|
|
|
|
mkMode : Permissions -> Int
|
|
|
|
mkMode p
|
|
|
|
= getMs (user p) * 64 + getMs (group p) * 8 + getMs (others p)
|
|
|
|
where
|
|
|
|
getM : FileMode -> Int
|
|
|
|
getM Read = 4
|
|
|
|
getM Write = 2
|
|
|
|
getM Execute = 1
|
|
|
|
|
|
|
|
getMs : List FileMode -> Int
|
|
|
|
getMs = sum . map getM
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
chmodRaw : HasIO io => String -> Int -> io (Either FileError ())
|
2020-05-18 15:59:07 +03:00
|
|
|
chmodRaw fname p
|
|
|
|
= do ok <- primIO $ prim__chmod fname p
|
|
|
|
if ok == 0
|
|
|
|
then pure (Right ())
|
|
|
|
else returnError
|
|
|
|
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
chmod : HasIO io => String -> Permissions -> io (Either FileError ())
|
2020-05-18 15:59:07 +03:00
|
|
|
chmod fname p = chmodRaw fname (mkMode p)
|