mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-28 02:23:44 +03:00
Add System.File fRead function
This commit is contained in:
parent
c1fc487bec
commit
dc47df688c
@ -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
|
||||
|
9
tests/base/system_file_fRead/ReadFile.idr
Normal file
9
tests/base/system_file_fRead/ReadFile.idr
Normal 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
|
3
tests/base/system_file_fRead/expected
Normal file
3
tests/base/system_file_fRead/expected
Normal file
@ -0,0 +1,3 @@
|
||||
1/1: Building ReadFile (ReadFile.idr)
|
||||
Main> "Hello, world\nLorem ipsum\n"
|
||||
Main> Bye for now!
|
2
tests/base/system_file_fRead/input
Normal file
2
tests/base/system_file_fRead/input
Normal file
@ -0,0 +1,2 @@
|
||||
:exec main
|
||||
:q
|
3
tests/base/system_file_fRead/run
Normal file
3
tests/base/system_file_fRead/run
Normal file
@ -0,0 +1,3 @@
|
||||
rm -rf build
|
||||
|
||||
$1 --no-color --console-width 0 --no-banner ReadFile.idr < input
|
2
tests/base/system_file_fRead/sampleFile.txt
Normal file
2
tests/base/system_file_fRead/sampleFile.txt
Normal file
@ -0,0 +1,2 @@
|
||||
Hello, world
|
||||
Lorem ipsum
|
Loading…
Reference in New Issue
Block a user