mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 20:51:43 +03:00
System.Directory.listDir
This function is what users want 99.9% of time.
This commit is contained in:
parent
c7629e20fe
commit
4eff6ac916
@ -118,3 +118,20 @@ dirEntry d = do r <- nextDirEntry d
|
|||||||
Left e => Left e
|
Left e => Left e
|
||||||
Right (Just n) => Right n
|
Right (Just n) => Right n
|
||||||
Right Nothing => Left FileNotFound
|
Right Nothing => Left FileNotFound
|
||||||
|
|
||||||
|
collectDir : HasIO io => Directory -> io (Either FileError (List String))
|
||||||
|
collectDir d
|
||||||
|
= liftIO $ do let (>>=) : (IO . Either e) a -> (a -> (IO . Either e) b) -> (IO . Either e) b
|
||||||
|
(>>=) = Prelude.(>>=) @{Monad.Compose {m = IO} {t = Either e}}
|
||||||
|
Just n <- nextDirEntry d
|
||||||
|
| Nothing => pure $ Right []
|
||||||
|
ns <- assert_total $ collectDir d
|
||||||
|
pure $ Right (n :: ns)
|
||||||
|
|
||||||
|
export
|
||||||
|
listDir : HasIO io => String -> io (Either FileError (List String))
|
||||||
|
listDir name = do Right d <- openDir name
|
||||||
|
| Left e => pure $ Left e
|
||||||
|
ns <- collectDir d
|
||||||
|
ignore <- closeDir d
|
||||||
|
pure $ ns
|
||||||
|
@ -5,16 +5,8 @@ panic : String -> IO a
|
|||||||
panic s = do putStrLn s
|
panic s = do putStrLn s
|
||||||
exitFailure
|
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
|
|
||||||
pure (n :: ns)
|
|
||||||
|
|
||||||
main : IO ()
|
main : IO ()
|
||||||
main = do Right d <- openDir "dir"
|
main = do Right ["a"] <- listDir "dir"
|
||||||
| Left e => panic (show e)
|
| Left e => panic (show e)
|
||||||
["a"] <- collectEntries d
|
| Right x => panic ("wrong entries: " ++ (show x))
|
||||||
| x => panic ("wrong entries: " ++ (show x))
|
|
||||||
pure ()
|
pure ()
|
||||||
|
Loading…
Reference in New Issue
Block a user