Add System.File fRead function

This commit is contained in:
Robert Wright 2021-11-03 15:56:29 +00:00
parent c1fc487bec
commit dc47df688c
6 changed files with 35 additions and 0 deletions

View File

@ -3,6 +3,7 @@ module System.File.ReadWrite
import public Data.Fuel
import Data.List
import Data.SnocList
import System.File.Handle
import public System.File.Error
@ -113,6 +114,21 @@ fEOF (FHandle f)
= do res <- primIO (prim__eof f)
pure (res /= 0)
||| Read all the remaining contents of a file handle
export
covering
fRead : HasIO io => (h : File) -> io (Either FileError String)
fRead h = fRead' h [<]
where
fRead' : HasIO io' => (h : File) -> (acc : SnocList String) -> io' (Either FileError String)
fRead' h acc = do
if !(fEOF h)
then pure $ Right $ concat acc
else do
Right line <- fGetLine h
| Left err => pure $ Left err
fRead' h $ acc :< line
||| Delete the file at the given name.
|||
||| @ fname the file to delete

View File

@ -0,0 +1,9 @@
import System.File
main : IO ()
main = do
Right f <- openFile "sampleFile.txt" Read
| Left err => printLn err
Right contents <- fRead f
| Left err => printLn err
printLn contents

View File

@ -0,0 +1,3 @@
1/1: Building ReadFile (ReadFile.idr)
Main> "Hello, world\nLorem ipsum\n"
Main> Bye for now!

View File

@ -0,0 +1,2 @@
:exec main
:q

View File

@ -0,0 +1,3 @@
rm -rf build
$1 --no-color --console-width 0 --no-banner ReadFile.idr < input

View File

@ -0,0 +1,2 @@
Hello, world
Lorem ipsum