Idris2/libs/base/Control/App/FileIO.idr

83 lines
2.2 KiB
Idris
Raw Normal View History

2020-05-18 15:59:07 +03:00
module Control.App.FileIO
import Control.App
import Data.List
import Data.Strings
import System.File
toFileEx : FileError -> FileEx
toFileEx (GenericFileError i) = GenericFileEx i
toFileEx FileReadError = FileReadError
toFileEx FileWriteError = FileWriteError
toFileEx FileNotFound = FileNotFound
toFileEx PermissionDenied = PermissionDenied
toFileEx FileExists = FileExists
public export
interface Has [Exception IOError] e => FileIO e where
2020-06-11 22:35:18 +03:00
withFile : String -> Mode ->
2020-05-18 15:59:07 +03:00
(onError : IOError -> App e a) ->
2020-06-11 22:35:18 +03:00
(onOpen : File -> App e a) ->
2020-05-18 15:59:07 +03:00
App e a
fGetStr : File -> App e String
2020-06-11 22:35:18 +03:00
fGetChars : File -> Int -> App e String
fGetChar : File -> App e Char
2020-05-18 15:59:07 +03:00
fPutStr : File -> String -> App e ()
2020-06-11 22:35:18 +03:00
fPutStrLn : File -> String -> App e ()
2020-06-12 18:30:31 +03:00
fflush : File -> App e ()
2020-05-18 15:59:07 +03:00
fEOF : File -> App e Bool
-- TODO : Add Binary File IO with buffers
export
readFile : FileIO e => String -> App e String
readFile f
= withFile f Read throw $ \h =>
do content <- read [] h
pure (fastAppend content)
where
read : List String -> File -> App e (List String)
read acc h
= do eof <- FileIO.fEOF h
if eof
then pure (reverse acc)
else do str <- fGetStr h
read (str :: acc) h
2020-06-11 22:50:03 +03:00
fileOp : IO (Either FileError a) -> Has [PrimIO, Exception IOError] e => App e a
fileOp fileRes
= do Right res <- primIO $ fileRes
| Left err => throw (FileErr (toFileEx err))
pure res
2020-05-18 15:59:07 +03:00
export
Has [PrimIO, Exception IOError] e => FileIO e where
withFile fname m onError proc
= do Right h <- primIO $ openFile fname m
| Left err => onError (FileErr (toFileEx err))
res <- catch (proc h) onError
2020-06-11 22:35:18 +03:00
primIO $ closeFile h
2020-05-18 15:59:07 +03:00
pure res
2020-06-11 22:50:03 +03:00
fGetStr f = fileOp (fGetLine f)
2020-05-18 15:59:07 +03:00
2020-06-11 22:50:03 +03:00
fGetChars f n = fileOp (fGetChars f n)
2020-06-11 22:35:18 +03:00
2020-06-11 22:50:03 +03:00
fGetChar f = fileOp (fGetChar f)
2020-06-11 22:35:18 +03:00
2020-06-11 22:50:03 +03:00
fPutStr f str = fileOp (fPutStr f str)
2020-05-18 15:59:07 +03:00
2020-06-11 22:50:03 +03:00
fPutStrLn f str = fileOp (File.fPutStrLn f str)
2020-06-11 22:35:18 +03:00
2020-06-12 18:30:31 +03:00
fflush f = primIO $ fflush f
2020-05-18 15:59:07 +03:00
fEOF f = primIO $ fEOF f
export
withFileIO : Has [PrimIO] e =>
App (IOError :: e) a ->
(ok : a -> App e b) ->
(err : IOError -> App e b) -> App e b
withFileIO prog ok err = handle prog ok err