Skip to content

Commit

Permalink
Merge pull request #4754 from clayrat/patch-15
Browse files Browse the repository at this point in the history
Fix file handle leaks and simplify code
  • Loading branch information
melted authored Aug 14, 2019
2 parents 476234f + 229331c commit 075c0f7
Showing 1 changed file with 20 additions and 25 deletions.
45 changes: 20 additions & 25 deletions libs/prelude/Prelude/File.idr
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ module Prelude.File
import Builtins
import Prelude.List
import Prelude.Maybe
import Prelude.Functor
import Prelude.Monad
import Prelude.Chars
import Prelude.Strings
Expand Down Expand Up @@ -80,7 +81,7 @@ do_ferror h = foreign FFI_C "fileError" (Ptr -> IO Int) h
export
ferror : File -> IO Bool
ferror (FHandle h) = do err <- do_ferror h
pure (not (err == 0))
pure (err /= 0)

||| Call the RTS's file opening function
private
Expand All @@ -95,15 +96,13 @@ export
fopen : (f : String) -> (m : String) -> IO (Either FileError File)
fopen f m = do h <- do_fopen f m
if !(nullPtr h)
then do err <- getFileError
pure (Left err)
then Left <$> getFileError
else pure (Right (FHandle h))

||| Check whether a file handle is actually a null pointer
export
validFile : File -> IO Bool
validFile (FHandle h) = do x <- nullPtr h
pure (not x)
validFile (FHandle h) = not <$> nullPtr h

||| Modes for opening files
data Mode = Read | WriteTruncate | Append | ReadWrite | ReadWriteTruncate | ReadAppend
Expand Down Expand Up @@ -169,38 +168,33 @@ fileSize : File -> IO (Either FileError Int)
fileSize (FHandle h)
= do s <- do_getFileSize h
if (s < 0)
then do err <- getFileError
pure (Left err)
then Left <$> getFileError
else pure (Right s)

export
fileModifiedTime : File -> IO (Either FileError Integer)
fileModifiedTime (FHandle h)
= do s <- do_getFileModifiedTime h
if (s < 0)
then do err <- getFileError
pure (Left err)
then Left <$> getFileError
else pure (Right s)

export
fileAccessTime : File -> IO (Either FileError Integer)
fileAccessTime (FHandle h)
= do s <- do_getFileAccessTime h
if (s < 0)
then do err <- getFileError
pure (Left err)
then Left <$> getFileError
else pure (Right s)

export
fileStatusTime : File -> IO (Either FileError Integer)
fileStatusTime (FHandle h)
= do s <- do_getFileStatusTime h
if (s < 0)
then do err <- getFileError
pure (Left err)
then Left <$> getFileError
else pure (Right s)


private
do_fread : Ptr -> IO' l String
do_fread h = prim_fread h
Expand Down Expand Up @@ -264,8 +258,7 @@ do_fwrite h s = do res <- prim_fwrite h s
then do errno <- getErrno
if errno == 0
then pure (Left FileWriteError)
else do err <- getFileError
pure (Left err)
else Left <$> getFileError
else pure (Right ())

||| Write a line to a file
Expand All @@ -287,7 +280,7 @@ do_feof h = foreign FFI_C "fileEOF" (Ptr -> IO Int) h
export
fEOF : File -> IO Bool
fEOF (FHandle h) = do eof <- do_feof h
pure (not (eof == 0))
pure (eof /= 0)

private
do_fremove : String -> IO Int
Expand Down Expand Up @@ -317,7 +310,8 @@ readFile : (filepath : String) -> IO (Either FileError String)
readFile fn = do Right h <- openFile fn Read
| Left err => pure (Left err)
Right max <- fileSize h
| Left err => pure (Left err)
| Left err => do closeFile h
pure (Left err)
sb <- newStringBuffer (max + 1)
c <- readFile' h max sb
closeFile h
Expand All @@ -340,8 +334,11 @@ export
writeFile : (filepath : String) -> (contents : String) ->
IO (Either FileError ())
writeFile fn contents = do
Right h <- openFile fn WriteTruncate | Left err => pure (Left err)
Right () <- fPutStr h contents | Left err => pure (Left err)
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 ())

Expand All @@ -351,8 +348,7 @@ dirOpen : (d : String) -> IO (Either FileError Directory)
dirOpen d
= do dptr <- foreign FFI_C "idris_dirOpen" (String -> IO Ptr) d
if !(nullPtr dptr)
then do err <- getFileError
pure (Left err)
then Left <$> getFileError
else pure (Right (DHandle dptr))

export
Expand All @@ -363,7 +359,7 @@ export
dirError : Directory -> IO Bool
dirError (DHandle d)
= do err <- foreign FFI_C "idris_dirError" (Ptr -> IO Int) d
pure (not (err == 0))
pure (err /= 0)

export
dirEntry : Directory -> IO (Either FileError String)
Expand All @@ -379,8 +375,7 @@ createDir d
= do ok <- foreign FFI_C "idris_mkdir" (String -> IO Int) d
if (ok == 0)
then pure (Right ())
else do err <- getFileError
pure (Left err)
else Left <$> getFileError

export
changeDir : String -> IO Bool
Expand Down

0 comments on commit 075c0f7

Please sign in to comment.