2021-10-29 19:58:29 +03:00
|
|
|
||| Directory access and handling.
|
2020-05-18 15:59:07 +03:00
|
|
|
module System.Directory
|
|
|
|
|
2021-07-16 20:33:48 +03:00
|
|
|
import System.Errno
|
2020-05-18 15:59:07 +03:00
|
|
|
import public System.File
|
|
|
|
|
2021-06-09 01:05:10 +03:00
|
|
|
%default total
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
public export
|
|
|
|
DirPtr : Type
|
|
|
|
DirPtr = AnyPtr
|
|
|
|
|
2022-04-07 12:09:30 +03:00
|
|
|
||| Shorthand for referring to the C support library
|
|
|
|
|||
|
|
|
|
||| @ fn the function name to refer to in the C support library
|
|
|
|
supportC : (fn : String) -> String
|
|
|
|
supportC fn = "C:\{fn}, libidris2_support, idris_directory.h"
|
|
|
|
|
|
|
|
||| Shorthand for referring to the Node system support library
|
|
|
|
|||
|
|
|
|
||| @ fn the function name to refer to in the js/system_support.js file
|
|
|
|
supportNode : (fn : String) -> String
|
|
|
|
supportNode fn = "node:support:\{fn},support_system_directory"
|
2020-05-18 15:59:07 +03:00
|
|
|
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
ok : HasIO io => a -> io (Either FileError a)
|
2020-05-18 15:59:07 +03:00
|
|
|
ok x = pure (Right x)
|
|
|
|
|
2022-04-07 12:09:30 +03:00
|
|
|
%foreign supportC "idris2_currentDirectory"
|
2020-06-18 01:29:54 +03:00
|
|
|
"node:lambda:()=>process.cwd()"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__currentDir : PrimIO (Ptr String)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2022-04-07 12:09:30 +03:00
|
|
|
%foreign supportC "idris2_changeDir"
|
|
|
|
supportNode "changeDir"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__changeDir : String -> PrimIO Int
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2022-04-07 12:09:30 +03:00
|
|
|
%foreign supportC "idris2_createDir"
|
|
|
|
supportNode "createDir"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__createDir : String -> PrimIO Int
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2022-04-07 12:09:30 +03:00
|
|
|
%foreign supportC "idris2_openDir"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__openDir : String -> PrimIO DirPtr
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2022-04-07 12:09:30 +03:00
|
|
|
%foreign supportC "idris2_closeDir"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__closeDir : DirPtr -> PrimIO ()
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2022-04-07 12:09:30 +03:00
|
|
|
%foreign supportC "idris2_removeDir"
|
|
|
|
supportNode "removeDir"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__removeDir : String -> PrimIO ()
|
2020-05-18 20:28:33 +03:00
|
|
|
|
2022-04-07 12:09:30 +03:00
|
|
|
%foreign supportC "idris2_nextDirEntry"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__dirEntry : DirPtr -> PrimIO (Ptr String)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Data structure for managing the pointer to a directory.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
|
|
|
data Directory : Type where
|
|
|
|
MkDir : DirPtr -> Directory
|
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Try to create a directory at the specified path.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
createDir : HasIO io => String -> io (Either FileError ())
|
2020-05-18 15:59:07 +03:00
|
|
|
createDir dir
|
2020-07-21 14:30:33 +03:00
|
|
|
= do res <- primIO (prim__createDir dir)
|
2020-05-18 15:59:07 +03:00
|
|
|
if res == 0
|
|
|
|
then ok ()
|
|
|
|
else returnError
|
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Change the current working directory to the specified path. Returns whether
|
|
|
|
||| the operation succeeded.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
changeDir : HasIO io => String -> io Bool
|
2020-05-18 15:59:07 +03:00
|
|
|
changeDir dir
|
2020-07-21 14:30:33 +03:00
|
|
|
= do ok <- primIO (prim__changeDir dir)
|
2020-05-18 15:59:07 +03:00
|
|
|
pure (ok == 0)
|
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Get the absolute path of the current working directory. Returns `Nothing` if
|
|
|
|
||| an error occurred.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
currentDir : HasIO io => io (Maybe String)
|
2020-05-18 15:59:07 +03:00
|
|
|
currentDir
|
2020-07-21 14:30:33 +03:00
|
|
|
= do res <- primIO prim__currentDir
|
2020-05-18 15:59:07 +03:00
|
|
|
if prim__nullPtr res /= 0
|
|
|
|
then pure Nothing
|
|
|
|
else pure (Just (prim__getString res))
|
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Try to open the directory at the specified path.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
openDir : HasIO io => String -> io (Either FileError Directory)
|
2020-05-21 15:32:35 +03:00
|
|
|
openDir d
|
2020-07-21 14:30:33 +03:00
|
|
|
= do res <- primIO (prim__openDir d)
|
2020-05-18 15:59:07 +03:00
|
|
|
if prim__nullAnyPtr res /= 0
|
|
|
|
then returnError
|
|
|
|
else ok (MkDir res)
|
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Close the given `Directory`.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
closeDir : HasIO io => Directory -> io ()
|
2020-07-21 14:30:33 +03:00
|
|
|
closeDir (MkDir d) = primIO (prim__closeDir d)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Remove the directory at the specified path.
|
2022-04-07 12:09:30 +03:00
|
|
|
||| If the directory is not empty, this operation fails.
|
2020-05-18 20:28:33 +03:00
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
removeDir : HasIO io => String -> io ()
|
2020-07-21 14:30:33 +03:00
|
|
|
removeDir dirName = primIO (prim__removeDir dirName)
|
2020-05-18 20:28:33 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Get the next entry in the `Directory`, omitting the '.' and '..' entries.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2021-07-16 20:33:48 +03:00
|
|
|
nextDirEntry : HasIO io => Directory -> io (Either FileError (Maybe String))
|
|
|
|
nextDirEntry (MkDir d)
|
2020-07-21 14:30:33 +03:00
|
|
|
= do res <- primIO (prim__dirEntry d)
|
2020-05-18 15:59:07 +03:00
|
|
|
if prim__nullPtr res /= 0
|
2021-07-16 20:33:48 +03:00
|
|
|
then if !(getErrno) /= 0
|
|
|
|
then returnError
|
|
|
|
else pure $ Right Nothing
|
2021-07-16 21:34:04 +03:00
|
|
|
else do let n = prim__getString res
|
|
|
|
if n == "." || n == ".."
|
|
|
|
then assert_total $ nextDirEntry (MkDir d)
|
|
|
|
else pure $ Right (Just n)
|
2021-07-16 20:33:48 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Get a list of all the entries in the `Directory`, excluding the '.' and '..'
|
|
|
|
||| entries.
|
2021-07-16 22:13:59 +03:00
|
|
|
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)
|
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Get a list of all the entries in the directory at the specified path,
|
|
|
|
||| excluding the '.' and '..' entries.
|
|
|
|
|||
|
|
|
|
||| @ name the directory to list
|
2021-07-16 22:13:59 +03:00
|
|
|
export
|
2021-10-29 19:58:29 +03:00
|
|
|
listDir : HasIO io => (name : String) -> io (Either FileError (List String))
|
2021-07-16 22:13:59 +03:00
|
|
|
listDir name = do Right d <- openDir name
|
|
|
|
| Left e => pure $ Left e
|
|
|
|
ns <- collectDir d
|
|
|
|
ignore <- closeDir d
|
|
|
|
pure $ ns
|