System.Directory.nextDirEntry

* add `nextDirEntry` which returns `Maybe String`, so `Nothing` on
  the end of directory unlike `dirEntry` which returns unspecified error
  on the end of directory
* `dirEntry` is deprecated now, but not removed because compiler depends on it
* native implementation of `dirEntry` is patched to explicitly reset `errno`
  before the `readdir` call: without it end of directory and error were
  indistinguishable
* test added
This commit is contained in:
Stiopa Koltsov 2021-07-16 18:33:48 +01:00
parent ce44d3b50a
commit 0ecf74e434
8 changed files with 50 additions and 4 deletions

View File

@ -1,5 +1,6 @@
module System.Directory
import System.Errno
import public System.File
%default total
@ -97,9 +98,20 @@ removeDir : HasIO io => String -> io ()
removeDir dirName = primIO (prim__removeDir dirName)
export
dirEntry : HasIO io => Directory -> io (Either FileError String)
dirEntry (MkDir d)
nextDirEntry : HasIO io => Directory -> io (Either FileError (Maybe String))
nextDirEntry (MkDir d)
= do res <- primIO (prim__dirEntry d)
if prim__nullPtr res /= 0
then returnError
else ok (prim__getString res)
then if !(getErrno) /= 0
then returnError
else pure $ Right Nothing
else pure $ Right (Just (prim__getString res))
-- This function is deprecated; to be removed after the next version bump
export
dirEntry : HasIO io => Directory -> io (Either FileError String)
dirEntry d = do r <- nextDirEntry d
pure $ case r of
Left e => Left e
Right (Just n) => Right n
Right Nothing => Left FileNotFound

View File

@ -60,6 +60,10 @@ int idris2_removeDir(char* path) {
char* idris2_nextDirEntry(void* d) {
DirInfo* di = (DirInfo*)d;
// `readdir` keeps `errno` unchanged on end of stream
// so we need to reset `errno` to distinguish between
// end of stream and failure.
errno = 0;
struct dirent* de = readdir(di->dirptr);
if (de == NULL) {

View File

@ -298,6 +298,7 @@ baseLibraryTests = MkTestPool "Base library" [Chez, Node] Nothing
, "data_bits001"
, "data_string_lines001"
, "data_string_unlines001"
, "system_directory"
, "system_errno"
, "system_info001"
, "system_signal001", "system_signal002", "system_signal003", "system_signal004"

View File

@ -0,0 +1,22 @@
import System
import System.Directory
panic : String -> IO a
panic s = do putStrLn s
exitFailure
collectEntries : Directory -> IO (List String)
collectEntries d = do Right (Just n) <- nextDirEntry d
| Right Nothing => pure []
| Left e => panic (show e)
ns <- collectEntries d
if n == "." || n == ".."
then pure ns
else pure (n :: ns)
main : IO ()
main = do Right d <- openDir "dir"
| Left e => panic (show e)
["a"] <- collectEntries d
| x => panic ("wrong entries: " ++ (show x))
pure ()

View File

View File

@ -0,0 +1,2 @@
1/1: Building ReadDir (ReadDir.idr)
Main> 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 ReadDir.idr < input