Merge pull request #823 from andylokandy/path

Changes to System.Path

- Rename `startWith` to `isBaseOf`
- Rename `stripPrefix` to `dropBase`
- Add `dropExtension`
- Add `splitPath`
This commit is contained in:
André Videla 2021-01-14 22:09:04 +00:00 committed by GitHub
commit bea840418a
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
3 changed files with 584 additions and 425 deletions

View File

@ -6,13 +6,13 @@ import Data.Nat
import Data.Strings import Data.Strings
import Data.String.Extra import Data.String.Extra
import System.Info
import Text.Token import Text.Token
import Text.Lexer import Text.Lexer
import Text.Parser import Text.Parser
import Text.Quantity import Text.Quantity
import System.Info
infixr 5 </> infixr 5 </>
infixr 7 <.> infixr 7 <.>
@ -27,44 +27,48 @@ export
pathSeparator : Char pathSeparator : Char
pathSeparator = if isWindows then ';' else ':' pathSeparator = if isWindows then ';' else ':'
||| Windows' path prefixes of path component. ||| Windows path prefix.
public export public export
data Volume data Volume
= =
||| Windows' Uniform Naming Convention, e.g., a network sharing ||| Windows Uniform Naming Convention, consisting of server name and share
||| directory: `\\host\c$\Windows\System32` ||| directory.
|||
||| Example: `\\localhost\share`
UNC String String | UNC String String |
||| The drive, e.g., "C:". The disk character is in upper case ||| The drive.
|||
||| Example: `C:`
Disk Char Disk Char
||| A single body of path component. ||| A single body in path.
public export public export
data Body data Body
= =
||| Represents "." ||| Represents ".".
CurDir | CurDir |
||| Represents ".." ||| Represents "..".
ParentDir | ParentDir |
||| Common directory or file ||| Directory or file.
Normal String Normal String
||| A parsed cross-platform file system path. ||| A parsed cross-platform file system path.
||| |||
||| The function `parse` constructs a Path component from String, ||| Use the function `parse` to constructs a Path from String, and use the
||| and the function `show` converts in reverse. ||| function `show` to convert in reverse.
||| |||
||| Trailing separator is only used for display and is ignored while ||| Trailing separator is only used for display and is ignored when comparing
||| comparing paths. ||| paths.
public export public export
record Path where record Path where
constructor MkPath constructor MkPath
||| Windows' path prefix (only on Windows) ||| Windows path prefix (only for Windows).
volume : Maybe Volume volume : Maybe Volume
||| Whether the path contains a root ||| Whether the path contains root.
hasRoot : Bool hasRoot : Bool
||| Path bodies ||| Path bodies.
body : List Body body : List Body
||| Whether the path terminates with a separator ||| Whether the path terminates with a separator.
hasTrailSep : Bool hasTrailSep : Bool
export export
@ -82,11 +86,10 @@ Eq Body where
export export
Eq Path where Eq Path where
(==) (MkPath l1 l2 l3 _) (MkPath r1 r2 r3 _) = l1 == r1 (==) (MkPath l1 l2 l3 _) (MkPath r1 r2 r3 _) =
&& l2 == r2 l1 == r1 && l2 == r2 && l3 == r3
&& l3 == r3
||| An empty path that represents "". ||| An empty path, which represents "".
public export public export
emptyPath : Path emptyPath : Path
emptyPath = MkPath Nothing False [] False emptyPath = MkPath Nothing False [] False
@ -99,22 +102,25 @@ export
Show Body where Show Body where
show CurDir = "." show CurDir = "."
show ParentDir = ".." show ParentDir = ".."
show (Normal s) = s show (Normal normal) = normal
export export
Show Volume where Show Volume where
show (UNC server share) = "\\\\" ++ server ++ "\\" ++ share show (UNC server share) = "\\\\" ++ server ++ "\\" ++ share
show (Disk disk) = singleton disk ++ ":" show (Disk disk) = Strings.singleton disk ++ ":"
||| Display the path in the format on the platform. ||| Displays the path in the format of this platform.
export export
Show Path where Show Path where
show p = let sep = singleton dirSeparator show path =
volStr = fromMaybe "" (map show p.volume) let
rootStr = if p.hasRoot then sep else "" sep = Strings.singleton dirSeparator
bodyStr = join sep $ map show p.body showVol = maybe "" show path.volume
trailStr = if p.hasTrailSep then sep else "" in showRoot = if path.hasRoot then sep else ""
volStr ++ rootStr ++ bodyStr ++ trailStr showBody = join sep $ map show path.body
showTrail = if path.hasTrailSep then sep else ""
in
showVol ++ showRoot ++ showBody ++ showTrail
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
-- Parser -- Parser
@ -147,203 +153,244 @@ pathTokenMap = toTokenMap $
] ]
lexPath : String -> List PathToken lexPath : String -> List PathToken
lexPath str = let (tokens, _, _, _) = lex pathTokenMap str in lexPath str =
map TokenData.tok tokens let
(tokens, _, _, _) = lex pathTokenMap str
in
map TokenData.tok tokens
-- match both '/' and '\\' regardless of the platform. -- match both '/' and '\\' regardless of the platform.
bodySeparator : Grammar PathToken True () bodySeparator : Grammar PathToken True ()
bodySeparator = (match $ PTPunct '\\') <|> (match $ PTPunct '/') bodySeparator = (match $ PTPunct '\\') <|> (match $ PTPunct '/')
-- Windows will automatically translate '/' to '\\'. And the verbatim prefix,
-- i.e., `\\?\`, can be used to disable the translation.
-- However, we just parse it and ignore it.
--
-- Example: \\?\ -- Example: \\?\
-- Windows can automatically translate '/' to '\\'. The verbatim prefix,
-- i.e., `\\?\`, disables the translation.
-- Here, we simply parse and then ignore it.
verbatim : Grammar PathToken True () verbatim : Grammar PathToken True ()
verbatim = do count (exactly 2) $ match $ PTPunct '\\' verbatim =
match $ PTPunct '?' do
match $ PTPunct '\\' count (exactly 2) $ match $ PTPunct '\\'
pure () match $ PTPunct '?'
match $ PTPunct '\\'
pure ()
-- Example: \\server\share -- Example: \\server\share
unc : Grammar PathToken True Volume unc : Grammar PathToken True Volume
unc = do count (exactly 2) $ match $ PTPunct '\\' unc =
server <- match PTText do
bodySeparator count (exactly 2) $ match $ PTPunct '\\'
share <- match PTText server <- match PTText
Core.pure $ UNC server share bodySeparator
share <- match PTText
pure $ UNC server share
-- Example: \\?\server\share -- Example: \\?\server\share
verbatimUnc : Grammar PathToken True Volume verbatimUnc : Grammar PathToken True Volume
verbatimUnc = do verbatim verbatimUnc =
server <- match PTText do
bodySeparator verbatim
share <- match PTText server <- match PTText
Core.pure $ UNC server share bodySeparator
share <- match PTText
pure $ UNC server share
-- Example: C: -- Example: C:
disk : Grammar PathToken True Volume disk : Grammar PathToken True Volume
disk = do text <- match PTText disk =
disk <- case unpack text of do
(disk :: xs) => pure disk text <- match PTText
[] => fail "Expect Disk" disk <- case unpack text of
match $ PTPunct ':' (disk :: xs) => pure disk
pure $ Disk (toUpper disk) [] => fail "Expects disk"
match $ PTPunct ':'
pure $ Disk (toUpper disk)
-- Example: \\?\C: -- Example: \\?\C:
verbatimDisk : Grammar PathToken True Volume verbatimDisk : Grammar PathToken True Volume
verbatimDisk = do verbatim verbatimDisk =
d <- disk do
pure d verbatim
disk <- disk
pure disk
parseVolume : Grammar PathToken True Volume parseVolume : Grammar PathToken True Volume
parseVolume = verbatimUnc parseVolume =
<|> verbatimDisk verbatimUnc
<|> unc <|> verbatimDisk
<|> disk <|> unc
<|> disk
parseBody : Grammar PathToken True Body parseBody : Grammar PathToken True Body
parseBody = do text <- match PTText parseBody =
the (Grammar _ False _) $ do
case text of text <- match PTText
".." => pure ParentDir the (Grammar _ False _) $
"." => pure CurDir case text of
s => pure (Normal s) ".." => pure ParentDir
"." => pure CurDir
normal => pure (Normal normal)
parsePath : Grammar PathToken False Path parsePath : Grammar PathToken False Path
parsePath = do vol <- optional parseVolume parsePath =
root <- optional (some bodySeparator) do
body <- sepBy (some bodySeparator) parseBody vol <- optional parseVolume
trailSep <- optional (some bodySeparator) root <- optional (some bodySeparator)
let body = filter (\case Normal s => ltrim s /= "" body <- sepBy (some bodySeparator) parseBody
_ => True) body trailSep <- optional (some bodySeparator)
let body = case body of let body = filter (\case Normal s => ltrim s /= ""; _ => True) body
[] => [] let body = case body of
(x::xs) => x :: delete CurDir xs [] => []
pure $ MkPath vol (isJust root) body (isJust trailSep) (x::xs) => x :: delete CurDir xs
pure $ MkPath vol (isJust root) body (isJust trailSep)
||| Parse a String into Path component. ||| Parses a String into Path.
||| |||
||| Returns the path parsed as much as possible from left to right, the ||| The string is parsed as much as possible from left to right, and the invalid
||| invalid parts on the right end is ignored. ||| parts on the right is ignored.
||| |||
||| Some kind of invalid path is accepted. Relaxing rules: ||| Some kind of invalid paths are accepted. The relax rules:
||| |||
||| - Both slash('/') and backslash('\\') are parsed as valid directory ||| - Both slash('/') and backslash('\\') are parsed as valid directory separator,
||| separator, regardless of the platform; ||| regardless of the platform;
||| - Any characters in path body in allowed, e.g., glob like "/root/*"; ||| - Any characters in the body in allowed, e.g., glob like "/root/*";
||| - Verbatim prefix(`\\?\`) that disables the forward ||| - Verbatim prefix(`\\?\`) that disables the forward slash is ignored (for
||| slash (Windows only) is ignored. ||| Windows only).
||| - Repeated separators are ignored, so "a/b" and "a//b" both have "a" ||| - Repeated separators are ignored, therefore, "a/b" and "a//b" both have "a"
||| and "b" as bodies. ||| and "b" as bodies.
||| - Occurrences of "." are normalized away, except if they are at the ||| - "." in the body are removed, unless they are at the beginning of the path.
||| beginning of the path. For example, "a/./b", "a/b/", "a/b/". and ||| For example, "a/./b", "a/b/", "a/b/." and "a/b" will have "a" and "b" as
||| "a/b" all have "a" and "b" as bodies, but "./a/b" starts with an ||| bodies, and "./a/b" will starts with `CurDir`.
||| additional `CurDir` body.
||| |||
||| ```idris example ||| ```idris example
||| parse "C:\\Windows/System32" ||| parse "C:\\Windows/System32"
||| ```
||| ```idris example
||| parse "/usr/local/etc/*" ||| parse "/usr/local/etc/*"
||| ``` ||| ```
export export
parse : String -> Path parse : String -> Path
parse str = case parse parsePath (lexPath str) of parse str =
Right (p, _) => p case parse parsePath (lexPath str) of
_ => emptyPath Right (path, _) => path
_ => emptyPath
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
-- Utils -- Utils
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
isAbsolute' : Path -> Bool isAbsolute' : Path -> Bool
isAbsolute' p = if isWindows isAbsolute' path =
then case p.volume of if isWindows then
Just (UNC _ _) => True case path.volume of
Just (Disk _) => p.hasRoot Just (UNC _ _) => True
Nothing => False Just (Disk _) => path.hasRoot
else p.hasRoot Nothing => False
else
path.hasRoot
append' : (left : Path) -> (right : Path) -> Path append' : (left : Path) -> (right : Path) -> Path
append' l r = if isAbsolute' r || isJust r.volume append' left right =
then r if isAbsolute' right || isJust right.volume then
else if hasRoot r right
then record { volume = l.volume } r else if hasRoot right then
else record { body = l.body ++ r.body, record { volume = left.volume } right
hasTrailSep = r.hasTrailSep } l else
record { body = left.body ++ right.body, hasTrailSep = right.hasTrailSep } left
splitPath' : Path -> List Path
splitPath' path =
case splitRoot path of
(Just root, other) => root :: iterateBody (path.body) (path.hasTrailSep)
(Nothing, other) => iterateBody (path.body) (path.hasTrailSep)
where
splitRoot : Path -> (Maybe Path, Path)
splitRoot path@(MkPath Nothing False _ _) = (Nothing, path)
splitRoot (MkPath vol root xs trailSep) =
(Just $ MkPath vol root [] False, MkPath Nothing False xs trailSep)
iterateBody : List Body -> (trailSep : Bool) -> List Path
iterateBody [] _ = []
iterateBody [x] trailSep = [MkPath Nothing False [x] trailSep]
iterateBody (x::y::xs) trailSep =
(MkPath Nothing False [x] False) :: iterateBody (y::xs) trailSep
splitParent' : Path -> Maybe (Path, Path) splitParent' : Path -> Maybe (Path, Path)
splitParent' p splitParent' path =
= case p.body of case path.body of
[] => Nothing [] => Nothing
(x::xs) => let parentPath = record { body = init (x::xs), (x::xs) =>
hasTrailSep = False } p let
lastPath = MkPath Nothing False [last (x::xs)] p.hasTrailSep in parent = record { body = init (x::xs), hasTrailSep = False } path
Just (parentPath, lastPath) child = MkPath Nothing False [last (x::xs)] path.hasTrailSep
in
Just (parent, child)
parent' : Path -> Maybe Path parent' : Path -> Maybe Path
parent' p = map fst (splitParent' p) parent' = map fst . splitParent'
fileName' : Path -> Maybe String fileName' : Path -> Maybe String
fileName' p = findNormal (reverse p.body) fileName' path =
where findNormal $ reverse path.body
findNormal : List Body -> Maybe String where
findNormal ((Normal s)::xs) = Just s findNormal : List Body -> Maybe String
findNormal (CurDir::xs) = findNormal xs findNormal ((Normal normal)::xs) = Just normal
findNormal _ = Nothing findNormal (CurDir::xs) = findNormal xs
findNormal _ = Nothing
setFileName' : (name : String) -> Path -> Path setFileName' : (name : String) -> Path -> Path
setFileName' name p = if isJust (fileName' p) setFileName' name path =
then append' (fromMaybe emptyPath $ parent' p) (parse name) if isJust (fileName' path) then
else append' p (parse name) append' (fromMaybe emptyPath $ parent' path) (parse name)
else
append' path (parse name)
splitFileName : String -> (String, String) splitFileName : String -> (String, String)
splitFileName name splitFileName name =
= case break (== '.') $ reverse $ unpack name of case break (== '.') $ reverse $ unpack name of
(_, []) => (name, "") (_, []) => (name, "")
(_, ['.']) => (name, "") (_, ['.']) => (name, "")
(revExt, (dot :: revStem)) (revExt, (dot :: revStem)) =>
=> ((pack $ reverse revStem), (pack $ reverse revExt)) ((pack $ reverse revStem), (pack $ reverse revExt))
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
-- Manipulations -- Methods
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
||| Returns true if the path is absolute. ||| Returns true if the path is absolute.
||| |||
||| - On Unix, a path is absolute if it starts with the root, ||| - On Unix, a path is absolute if it starts with the root, so `isAbsolute` and
||| so isAbsolute and hasRoot are equivalent. ||| `hasRoot` are equivalent.
||| |||
||| - On Windows, a path is absolute if it has a volume and starts ||| - On Windows, a path is absolute if it starts with a disk and has root or
||| with the root. e.g., `c:\\windows` is absolute, while `c:temp` ||| starts with UNC. For example, `c:\\windows` is absolute, while `c:temp`
||| and `\temp` are not. In addition, a path with UNC volume is absolute. ||| and `\temp` are not.
export export
isAbsolute : String -> Bool isAbsolute : String -> Bool
isAbsolute p = isAbsolute' (parse p) isAbsolute = isAbsolute' . parse
||| Returns true if the path is relative, i.e., not absolute. ||| Returns true if the path is relative.
export export
isRelative : String -> Bool isRelative : String -> Bool
isRelative = not . isAbsolute isRelative = not . isAbsolute
||| Appends the right path to the left one. ||| Appends the right path to the left path.
||| |||
||| If the path on the right is absolute, it replaces the left path. ||| If the path on the right is absolute, it replaces the left path.
||| |||
||| On Windows: ||| On Windows:
||| |||
||| - If the right path has a root but no volume (e.g., `\windows`), it ||| - If the right path has a root but no volume (e.g., `\windows`), it replaces
||| replaces everything except for the volume (if any) of left. ||| everything except for the volume (if any) of left.
||| - If the right path has a volume but no root, it replaces left. ||| - If the right path has a volume but no root, it replaces left.
||| |||
||| ```idris example ||| ```idris example
||| "/usr" </> "local/etc" ||| "/usr" </> "local/etc" == "/usr/local/etc"
||| ``` ||| ```
export export
(</>) : (left : String) -> (right : String) -> String (</>) : (left : String) -> (right : String) -> String
(</>) l r = show $ append' (parse l) (parse r) (</>) left right = show $ append' (parse left) (parse right)
||| Join path elements together. ||| Joins path components into one.
||| |||
||| ```idris example ||| ```idris example
||| joinPath ["/usr", "local/etc"] == "/usr/local/etc" ||| joinPath ["/usr", "local/etc"] == "/usr/local/etc"
@ -352,115 +399,147 @@ export
joinPath : List String -> String joinPath : List String -> String
joinPath xs = foldl (</>) "" xs joinPath xs = foldl (</>) "" xs
||| Returns the parent and child. ||| Splits path into components.
|||
||| ```idris example
||| splitPath "/usr/local/etc" == ["/", "usr", "local", "etc"]
||| splitPath "tmp/Foo.idr" == ["tmp", "Foo.idr"]
||| ```
export
splitPath : String -> List String
splitPath = map show . splitPath' . parse
||| Splits the path into parent and child.
||| |||
||| ```idris example ||| ```idris example
||| splitParent "/usr/local/etc" == Just ("/usr/local", "etc") ||| splitParent "/usr/local/etc" == Just ("/usr/local", "etc")
||| ``` ||| ```
export export
splitParent : String -> Maybe (String, String) splitParent : String -> Maybe (String, String)
splitParent p = do (a, b) <- splitParent' (parse p) splitParent path =
pure $ (show a, show b) do
(parent, child) <- splitParent' (parse path)
pure (show parent, show child)
||| Returns the path without its final component, if there is one. ||| Returns the path without its final component, if there is one.
||| |||
||| Returns Nothing if the path terminates in a root or volume. ||| Returns Nothing if the path terminates by a root or volume.
export export
parent : String -> Maybe String parent : String -> Maybe String
parent p = map show $ parent' (parse p) parent = map show . parent' . parse
||| Returns a list of all the parents of the path, longest first, ||| Returns the list of all parents of the path, longest first, self included.
||| self included.
||| |||
||| ```idris example ||| ```idris example
||| parents "/etc/kernel" == ["/etc/kernel", "/etc", "/"] ||| parents "/etc/kernel" == ["/etc/kernel", "/etc", "/"]
||| ``` ||| ```
export export
parents : String -> List String parents : String -> List String
parents p = map show $ iterate parent' (parse p) parents = map show . List.iterate parent' . parse
||| Determines whether base is either one of the parents of full. ||| Determines whether the base is one of the parents of target.
||| |||
||| Trailing separator is ignored. ||| Trailing separator is ignored.
export
startWith : (base : String) -> (full : String) -> Bool
startWith base full = (parse base) `elem` (iterate parent' (parse full))
||| Returns a path that, when appended onto base, yields full.
||| |||
||| If base is not a prefix of full (i.e., startWith returns false), ||| ```idris example
||| returns Nothing. ||| "/etc" `isBaseOf` "/etc/kernel"
||| ```
export export
stripPrefix : (base : String) -> (full : String) -> Maybe String isBaseOf : (base : String) -> (target : String) -> Bool
stripPrefix base full isBaseOf base target =
= do let MkPath vol1 root1 body1 _ = parse base let
let MkPath vol2 root2 body2 trialSep = parse full MkPath baseVol baseRoot baseBody _ = parse base
if vol1 == vol2 && root1 == root2 then Just () else Nothing MkPath targetVol targetRoot targetBody _ = parse target
body <- stripBody body1 body2 in
pure $ show $ MkPath Nothing False body trialSep baseVol == targetVol
&& baseRoot == targetRoot
&& (baseBody `isPrefixOf` targetBody)
||| Returns a path that, when appended to base, yields target.
|||
||| Returns Nothing if the base is not a prefix of the target.
export
dropBase : (base : String) -> (target : String) -> Maybe String
dropBase base target =
do
let MkPath baseVol baseRoot baseBody _ = parse base
let MkPath targetVol targetRoot targetBody targetTrialSep = parse target
if baseVol == targetVol && baseRoot == targetRoot then Just () else Nothing
body <- dropBody baseBody targetBody
pure $ show $ MkPath Nothing False body targetTrialSep
where where
stripBody : (base : List Body) -> (full : List Body) -> Maybe (List Body) dropBody : (base : List Body) -> (target : List Body) -> Maybe (List Body)
stripBody [] ys = Just ys dropBody [] ys = Just ys
stripBody xs [] = Nothing dropBody xs [] = Nothing
stripBody (x::xs) (y::ys) = if x == y then stripBody xs ys else Nothing dropBody (x::xs) (y::ys) = if x == y then dropBody xs ys else Nothing
||| Returns the final body of the path, if there is one. ||| Returns the last body of the path.
||| |||
||| If the path is a normal file, this is the file name. If it's the ||| If the last body is a file, this is the file name;
||| path of a directory, this is the directory name. ||| if it's a directory, this is the directory name;
||| ||| if it's ".", it recurses on the previous body;
||| Returns Nothing if the final body is "..". ||| if it's "..", returns Nothing.
export export
fileName : String -> Maybe String fileName : String -> Maybe String
fileName p = fileName' (parse p) fileName = fileName' . parse
||| Extracts the stem (non-extension) portion of the file name of path. ||| Extracts the file name in the path without extension.
||| |||
||| The stem is: ||| The stem is:
||| |||
||| - Nothing, if there is no file name; ||| - Nothing, if there is no file name;
||| - The entire file name if there is no embedded "."; ||| - The entire file name if there is no embedded ".";
||| - The entire file name if the file name begins with "." and has ||| - The entire file name if the file name begins with a "." and has no other ".";
||| no other "."s within; ||| - Otherwise, the portion of the file name before the last ".".
||| - Otherwise, the portion of the file name before the final "."
export export
fileStem : String -> Maybe String fileStem : String -> Maybe String
fileStem p = pure $ fst $ splitFileName !(fileName p) fileStem path = pure $ fst $ splitFileName !(fileName path)
||| Extracts the extension of the file name of path. ||| Extracts the extension of the file name in the path.
||| |||
||| The extension is: ||| The extension is:
||| |||
||| - Nothing, if there is no file name; ||| - Nothing, if there is no file name;
||| - Nothing, if there is no embedded "."; ||| - Nothing, if there is no embedded ".";
||| - Nothing, if the file name begins with "." and has no other "."s within; ||| - Nothing, if the file name begins with a "." and has no other ".";
||| - Otherwise, the portion of the file name after the final "." ||| - Otherwise, the portion of the file name after the last ".".
export export
extension : String -> Maybe String extension : String -> Maybe String
extension p = pure $ snd $ splitFileName !(fileName p) extension path = pure $ snd $ splitFileName !(fileName path)
||| Updates the file name of the path. ||| Updates the file name in the path.
||| |||
||| If no file name, this is equivalent to appending the name; ||| If there is no file name, it appends the name to the path;
||| Otherwise it is equivalent to appending the name to the parent. ||| otherwise, it appends the name to the parent of the path.
export export
setFileName : (name : String) -> String -> String setFileName : (name : String) -> String -> String
setFileName name p = show $ setFileName' name (parse p) setFileName name path = show $ setFileName' name (parse path)
||| Append a extension to the path. ||| Appends a extension to the path.
||| |||
||| Returns the path as it is if no file name. ||| If there is no file name, the path will not change;
||| if the path has no extension, the extension will be appended;
||| if the given extension is empty, the extension of the path will be dropped;
||| otherwise, the extension will be replaced.
||| |||
||| If `extension` of the path is Nothing, the extension is added; otherwise ||| ```idris example
||| it is replaced. ||| "/tmp/Foo" <.> "idr" == "/tmp/Foo.idr"
||| ||| ```
||| If the ext is empty, the extension is dropped.
export export
(<.>) : String -> (ext : String) -> String (<.>) : String -> (extension : String) -> String
(<.>) p ext = let p' = parse p (<.>) path ext =
ext = pack $ dropWhile (== '.') (unpack ext) let
ext = if ltrim ext == "" then "" else "." ++ ext in path' = parse path
case fileName' p' of ext = pack $ dropWhile (\char => char == '.' || isSpace char) (unpack ext)
Just name => let (stem, _) = splitFileName name in ext = if ext == "" then "" else "." ++ ext
show $ setFileName' (stem ++ ext) p' in
Nothing => p case fileName' path' of
Just name =>
let (stem, _) = splitFileName name in
show $ setFileName' (stem ++ ext) path'
Nothing => path
||| Drops the extension of the path.
export
dropExtension : String -> String
dropExtension path = path <.> ""

View File

@ -88,15 +88,19 @@ nsToSource loc ns
-- namespace for it -- namespace for it
export export
pathToNS : String -> Maybe String -> String -> Core ModuleIdent pathToNS : String -> Maybe String -> String -> Core ModuleIdent
pathToNS wdir sdir fname pathToNS wdir sdir fname =
= let sdir = fromMaybe "" sdir let
base = if isAbsolute fname then wdir </> sdir else sdir sdir = fromMaybe "" sdir
in base = if isAbsolute fname then wdir </> sdir else sdir
case stripPrefix base fname of in
Nothing => throw (UserError ("Source file " ++ show fname case Path.dropBase base fname of
++ " is not in the source directory " Nothing => throw (UserError (
++ show (wdir </> sdir))) "Source file "
Just p => pure $ unsafeFoldModuleIdent $ map show $ reverse $ (parse (p <.> "")).body ++ show fname
++ " is not in the source directory "
++ show (wdir </> sdir)))
Just relPath =>
pure $ unsafeFoldModuleIdent $ reverse $ splitPath $ Path.dropExtension relPath
dirExists : String -> IO Bool dirExists : String -> IO Bool
dirExists dir = do Right d <- openDir dir dirExists dir = do Right d <- openDir dir

View File

@ -6,13 +6,13 @@ import Data.Nat
import Data.Strings import Data.Strings
import Data.String.Extra import Data.String.Extra
import System.Info
import Text.Token import Text.Token
import Text.Lexer import Text.Lexer
import Text.Parser import Text.Parser
import Text.Quantity import Text.Quantity
import System.Info
infixr 5 </> infixr 5 </>
infixr 7 <.> infixr 7 <.>
@ -27,44 +27,48 @@ export
pathSeparator : Char pathSeparator : Char
pathSeparator = if isWindows then ';' else ':' pathSeparator = if isWindows then ';' else ':'
||| Windows' path prefixes of path component. ||| Windows path prefix.
public export public export
data Volume data Volume
= =
||| Windows' Uniform Naming Convention, e.g., a network sharing ||| Windows Uniform Naming Convention, consisting of server name and share
||| directory: `\\host\c$\Windows\System32` ||| directory.
|||
||| Example: `\\localhost\share`
UNC String String | UNC String String |
||| The drive, e.g., "C:". The disk character is in upper case ||| The drive.
|||
||| Example: `C:`
Disk Char Disk Char
||| A single body of path component. ||| A single body in path.
public export public export
data Body data Body
= =
||| Represents "." ||| Represents ".".
CurDir | CurDir |
||| Represents ".." ||| Represents "..".
ParentDir | ParentDir |
||| Common directory or file ||| Directory or file.
Normal String Normal String
||| A parsed cross-platform file system path. ||| A parsed cross-platform file system path.
||| |||
||| The function `parse` constructs a Path component from String, ||| Use the function `parse` to constructs a Path from String, and use the
||| and the function `show` converts in reverse. ||| function `show` to convert in reverse.
||| |||
||| Trailing separator is only used for display and is ignored while ||| Trailing separator is only used for display and is ignored when comparing
||| comparing paths. ||| paths.
public export public export
record Path where record Path where
constructor MkPath constructor MkPath
||| Windows' path prefix (only on Windows) ||| Windows path prefix (only for Windows).
volume : Maybe Volume volume : Maybe Volume
||| Whether the path contains a root ||| Whether the path contains root.
hasRoot : Bool hasRoot : Bool
||| Path bodies ||| Path bodies.
body : List Body body : List Body
||| Whether the path terminates with a separator ||| Whether the path terminates with a separator.
hasTrailSep : Bool hasTrailSep : Bool
export export
@ -82,11 +86,10 @@ Eq Body where
export export
Eq Path where Eq Path where
(==) (MkPath l1 l2 l3 _) (MkPath r1 r2 r3 _) = l1 == r1 (==) (MkPath l1 l2 l3 _) (MkPath r1 r2 r3 _) =
&& l2 == r2 l1 == r1 && l2 == r2 && l3 == r3
&& l3 == r3
||| An empty path that represents "". ||| An empty path, which represents "".
public export public export
emptyPath : Path emptyPath : Path
emptyPath = MkPath Nothing False [] False emptyPath = MkPath Nothing False [] False
@ -99,22 +102,25 @@ export
Show Body where Show Body where
show CurDir = "." show CurDir = "."
show ParentDir = ".." show ParentDir = ".."
show (Normal s) = s show (Normal normal) = normal
export export
Show Volume where Show Volume where
show (UNC server share) = "\\\\" ++ server ++ "\\" ++ share show (UNC server share) = "\\\\" ++ server ++ "\\" ++ share
show (Disk disk) = singleton disk ++ ":" show (Disk disk) = Strings.singleton disk ++ ":"
||| Display the path in the format on the platform. ||| Displays the path in the format of this platform.
export export
Show Path where Show Path where
show p = let sep = singleton dirSeparator show path =
volStr = fromMaybe "" (map show p.volume) let
rootStr = if p.hasRoot then sep else "" sep = Strings.singleton dirSeparator
bodyStr = join sep $ map show p.body showVol = maybe "" show path.volume
trailStr = if p.hasTrailSep then sep else "" in showRoot = if path.hasRoot then sep else ""
volStr ++ rootStr ++ bodyStr ++ trailStr showBody = join sep $ map show path.body
showTrail = if path.hasTrailSep then sep else ""
in
showVol ++ showRoot ++ showBody ++ showTrail
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
-- Parser -- Parser
@ -153,196 +159,234 @@ lexPath str = let (tokens, _, _, _) = lex pathTokenMap str in tokens
bodySeparator : Grammar PathToken True () bodySeparator : Grammar PathToken True ()
bodySeparator = (match $ PTPunct '\\') <|> (match $ PTPunct '/') bodySeparator = (match $ PTPunct '\\') <|> (match $ PTPunct '/')
-- Windows will automatically translate '/' to '\\'. And the verbatim prefix,
-- i.e., `\\?\`, can be used to disable the translation.
-- However, we just parse it and ignore it.
--
-- Example: \\?\ -- Example: \\?\
-- Windows can automatically translate '/' to '\\'. The verbatim prefix,
-- i.e., `\\?\`, disables the translation.
-- Here, we simply parse and then ignore it.
verbatim : Grammar PathToken True () verbatim : Grammar PathToken True ()
verbatim = do count (exactly 2) $ match $ PTPunct '\\' verbatim =
match $ PTPunct '?' do
match $ PTPunct '\\' count (exactly 2) $ match $ PTPunct '\\'
pure () match $ PTPunct '?'
match $ PTPunct '\\'
pure ()
-- Example: \\server\share -- Example: \\server\share
unc : Grammar PathToken True Volume unc : Grammar PathToken True Volume
unc = do count (exactly 2) $ match $ PTPunct '\\' unc =
server <- match PTText do
bodySeparator count (exactly 2) $ match $ PTPunct '\\'
share <- match PTText server <- match PTText
Core.pure $ UNC server share bodySeparator
share <- match PTText
pure $ UNC server share
-- Example: \\?\server\share -- Example: \\?\server\share
verbatimUnc : Grammar PathToken True Volume verbatimUnc : Grammar PathToken True Volume
verbatimUnc = do verbatim verbatimUnc =
server <- match PTText do
bodySeparator verbatim
share <- match PTText server <- match PTText
Core.pure $ UNC server share bodySeparator
share <- match PTText
pure $ UNC server share
-- Example: C: -- Example: C:
disk : Grammar PathToken True Volume disk : Grammar PathToken True Volume
disk = do text <- match PTText disk =
disk <- case unpack text of do
(disk :: xs) => pure disk text <- match PTText
[] => fail "Expect Disk" disk <- case unpack text of
match $ PTPunct ':' (disk :: xs) => pure disk
pure $ Disk (toUpper disk) [] => fail "Expects disk"
match $ PTPunct ':'
pure $ Disk (toUpper disk)
-- Example: \\?\C: -- Example: \\?\C:
verbatimDisk : Grammar PathToken True Volume verbatimDisk : Grammar PathToken True Volume
verbatimDisk = do verbatim verbatimDisk =
d <- disk do
pure d verbatim
disk <- disk
pure disk
parseVolume : Grammar PathToken True Volume parseVolume : Grammar PathToken True Volume
parseVolume = verbatimUnc parseVolume =
<|> verbatimDisk verbatimUnc
<|> unc <|> verbatimDisk
<|> disk <|> unc
<|> disk
parseBody : Grammar PathToken True Body parseBody : Grammar PathToken True Body
parseBody = do text <- match PTText parseBody =
the (Grammar _ False _) $ do
case text of text <- match PTText
".." => pure ParentDir the (Grammar _ False _) $
"." => pure CurDir case text of
s => pure (Normal s) ".." => pure ParentDir
"." => pure CurDir
normal => pure (Normal normal)
parsePath : Grammar PathToken False Path parsePath : Grammar PathToken False Path
parsePath = do vol <- optional parseVolume parsePath =
root <- optional (some bodySeparator) do
body <- sepBy (some bodySeparator) parseBody vol <- optional parseVolume
trailSep <- optional (some bodySeparator) root <- optional (some bodySeparator)
let body = filter (\case Normal s => ltrim s /= "" body <- sepBy (some bodySeparator) parseBody
_ => True) body trailSep <- optional (some bodySeparator)
let body = case body of let body = filter (\case Normal s => ltrim s /= ""; _ => True) body
[] => [] let body = case body of
(x::xs) => x :: delete CurDir xs [] => []
pure $ MkPath vol (isJust root) body (isJust trailSep) (x::xs) => x :: delete CurDir xs
pure $ MkPath vol (isJust root) body (isJust trailSep)
||| Parse a String into Path component. ||| Parses a String into Path.
||| |||
||| Returns the path parsed as much as possible from left to right, the ||| The string is parsed as much as possible from left to right, and the invalid
||| invalid parts on the right end is ignored. ||| parts on the right is ignored.
||| |||
||| Some kind of invalid path is accepted. Relaxing rules: ||| Some kind of invalid paths are accepted. The relax rules:
||| |||
||| - Both slash('/') and backslash('\\') are parsed as valid directory ||| - Both slash('/') and backslash('\\') are parsed as valid directory separator,
||| separator, regardless of the platform; ||| regardless of the platform;
||| - Any characters in path body in allowed, e.g., glob like "/root/*"; ||| - Any characters in the body in allowed, e.g., glob like "/root/*";
||| - Verbatim prefix(`\\?\`) that disables the forward ||| - Verbatim prefix(`\\?\`) that disables the forward slash is ignored (for
||| slash (Windows only) is ignored. ||| Windows only).
||| - Repeated separators are ignored, so "a/b" and "a//b" both have "a" ||| - Repeated separators are ignored, therefore, "a/b" and "a//b" both have "a"
||| and "b" as bodies. ||| and "b" as bodies.
||| - Occurrences of "." are normalized away, except if they are at the ||| - "." in the body are removed, unless they are at the beginning of the path.
||| beginning of the path. For example, "a/./b", "a/b/", "a/b/". and ||| For example, "a/./b", "a/b/", "a/b/." and "a/b" will have "a" and "b" as
||| "a/b" all have "a" and "b" as bodies, but "./a/b" starts with an ||| bodies, and "./a/b" will starts with `CurDir`.
||| additional `CurDir` body.
||| |||
||| ```idris example ||| ```idris example
||| parse "C:\\Windows/System32" ||| parse "C:\\Windows/System32"
||| ```
||| ```idris example
||| parse "/usr/local/etc/*" ||| parse "/usr/local/etc/*"
||| ``` ||| ```
export export
parse : String -> Path parse : String -> Path
parse str = case parse parsePath (lexPath str) of parse str =
Right (p, _) => p case parse parsePath (lexPath str) of
_ => emptyPath Right (path, _) => path
_ => emptyPath
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
-- Utils -- Utils
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
isAbsolute' : Path -> Bool isAbsolute' : Path -> Bool
isAbsolute' p = if isWindows isAbsolute' path =
then case p.volume of if isWindows then
Just (UNC _ _) => True case path.volume of
Just (Disk _) => p.hasRoot Just (UNC _ _) => True
Nothing => False Just (Disk _) => path.hasRoot
else p.hasRoot Nothing => False
else
path.hasRoot
append' : (left : Path) -> (right : Path) -> Path append' : (left : Path) -> (right : Path) -> Path
append' l r = if isAbsolute' r || isJust r.volume append' left right =
then r if isAbsolute' right || isJust right.volume then
else if hasRoot r right
then record { volume = l.volume } r else if hasRoot right then
else record { body = l.body ++ r.body, record { volume = left.volume } right
hasTrailSep = r.hasTrailSep } l else
record { body = left.body ++ right.body, hasTrailSep = right.hasTrailSep } left
splitPath' : Path -> List Path
splitPath' path =
case splitRoot path of
(Just root, other) => root :: iterateBody (path.body) (path.hasTrailSep)
(Nothing, other) => iterateBody (path.body) (path.hasTrailSep)
where
splitRoot : Path -> (Maybe Path, Path)
splitRoot path@(MkPath Nothing False _ _) = (Nothing, path)
splitRoot (MkPath vol root xs trailSep) =
(Just $ MkPath vol root [] False, MkPath Nothing False xs trailSep)
iterateBody : List Body -> (trailSep : Bool) -> List Path
iterateBody [] _ = []
iterateBody [x] trailSep = [MkPath Nothing False [x] trailSep]
iterateBody (x::y::xs) trailSep =
(MkPath Nothing False [x] False) :: iterateBody (y::xs) trailSep
splitParent' : Path -> Maybe (Path, Path) splitParent' : Path -> Maybe (Path, Path)
splitParent' p splitParent' path =
= case p.body of case path.body of
[] => Nothing [] => Nothing
(x::xs) => let parentPath = record { body = init (x::xs), (x::xs) =>
hasTrailSep = False } p let
lastPath = MkPath Nothing False [last (x::xs)] p.hasTrailSep in parent = record { body = init (x::xs), hasTrailSep = False } path
Just (parentPath, lastPath) child = MkPath Nothing False [last (x::xs)] path.hasTrailSep
in
Just (parent, child)
parent' : Path -> Maybe Path parent' : Path -> Maybe Path
parent' p = map fst (splitParent' p) parent' = map fst . splitParent'
fileName' : Path -> Maybe String fileName' : Path -> Maybe String
fileName' p = findNormal (reverse p.body) fileName' path =
where findNormal $ reverse path.body
findNormal : List Body -> Maybe String where
findNormal ((Normal s)::xs) = Just s findNormal : List Body -> Maybe String
findNormal (CurDir::xs) = findNormal xs findNormal ((Normal normal)::xs) = Just normal
findNormal _ = Nothing findNormal (CurDir::xs) = findNormal xs
findNormal _ = Nothing
setFileName' : (name : String) -> Path -> Path setFileName' : (name : String) -> Path -> Path
setFileName' name p = if isJust (fileName' p) setFileName' name path =
then append' (fromMaybe emptyPath $ parent' p) (parse name) if isJust (fileName' path) then
else append' p (parse name) append' (fromMaybe emptyPath $ parent' path) (parse name)
else
append' path (parse name)
splitFileName : String -> (String, String) splitFileName : String -> (String, String)
splitFileName name splitFileName name =
= case break (== '.') $ reverse $ unpack name of case break (== '.') $ reverse $ unpack name of
(_, []) => (name, "") (_, []) => (name, "")
(_, ['.']) => (name, "") (_, ['.']) => (name, "")
(revExt, (dot :: revStem)) (revExt, (dot :: revStem)) =>
=> ((pack $ reverse revStem), (pack $ reverse revExt)) ((pack $ reverse revStem), (pack $ reverse revExt))
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
-- Manipulations -- Methods
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
||| Returns true if the path is absolute. ||| Returns true if the path is absolute.
||| |||
||| - On Unix, a path is absolute if it starts with the root, ||| - On Unix, a path is absolute if it starts with the root, so `isAbsolute` and
||| so isAbsolute and hasRoot are equivalent. ||| `hasRoot` are equivalent.
||| |||
||| - On Windows, a path is absolute if it has a volume and starts ||| - On Windows, a path is absolute if it starts with a disk and has root or
||| with the root. e.g., `c:\\windows` is absolute, while `c:temp` ||| starts with UNC. For example, `c:\\windows` is absolute, while `c:temp`
||| and `\temp` are not. In addition, a path with UNC volume is absolute. ||| and `\temp` are not.
export export
isAbsolute : String -> Bool isAbsolute : String -> Bool
isAbsolute p = isAbsolute' (parse p) isAbsolute = isAbsolute' . parse
||| Returns true if the path is relative, i.e., not absolute. ||| Returns true if the path is relative.
export export
isRelative : String -> Bool isRelative : String -> Bool
isRelative = not . isAbsolute isRelative = not . isAbsolute
||| Appends the right path to the left one. ||| Appends the right path to the left path.
||| |||
||| If the path on the right is absolute, it replaces the left path. ||| If the path on the right is absolute, it replaces the left path.
||| |||
||| On Windows: ||| On Windows:
||| |||
||| - If the right path has a root but no volume (e.g., `\windows`), it ||| - If the right path has a root but no volume (e.g., `\windows`), it replaces
||| replaces everything except for the volume (if any) of left. ||| everything except for the volume (if any) of left.
||| - If the right path has a volume but no root, it replaces left. ||| - If the right path has a volume but no root, it replaces left.
||| |||
||| ```idris example ||| ```idris example
||| "/usr" </> "local/etc" ||| "/usr" </> "local/etc" == "/usr/local/etc"
||| ``` ||| ```
export export
(</>) : (left : String) -> (right : String) -> String (</>) : (left : String) -> (right : String) -> String
(</>) l r = show $ append' (parse l) (parse r) (</>) left right = show $ append' (parse left) (parse right)
||| Join path elements together. ||| Joins path components into one.
||| |||
||| ```idris example ||| ```idris example
||| joinPath ["/usr", "local/etc"] == "/usr/local/etc" ||| joinPath ["/usr", "local/etc"] == "/usr/local/etc"
@ -351,115 +395,147 @@ export
joinPath : List String -> String joinPath : List String -> String
joinPath xs = foldl (</>) "" xs joinPath xs = foldl (</>) "" xs
||| Returns the parent and child. ||| Splits path into components.
|||
||| ```idris example
||| splitPath "/usr/local/etc" == ["/", "usr", "local", "etc"]
||| splitPath "tmp/Foo.idr" == ["tmp", "Foo.idr"]
||| ```
export
splitPath : String -> List String
splitPath = map show . splitPath' . parse
||| Splits the path into parent and child.
||| |||
||| ```idris example ||| ```idris example
||| splitParent "/usr/local/etc" == Just ("/usr/local", "etc") ||| splitParent "/usr/local/etc" == Just ("/usr/local", "etc")
||| ``` ||| ```
export export
splitParent : String -> Maybe (String, String) splitParent : String -> Maybe (String, String)
splitParent p = do (a, b) <- splitParent' (parse p) splitParent path =
pure $ (show a, show b) do
(parent, child) <- splitParent' (parse path)
pure (show parent, show child)
||| Returns the path without its final component, if there is one. ||| Returns the path without its final component, if there is one.
||| |||
||| Returns Nothing if the path terminates in a root or volume. ||| Returns Nothing if the path terminates by a root or volume.
export export
parent : String -> Maybe String parent : String -> Maybe String
parent p = map show $ parent' (parse p) parent = map show . parent' . parse
||| Returns a list of all the parents of the path, longest first, ||| Returns the list of all parents of the path, longest first, self included.
||| self included.
||| |||
||| ```idris example ||| ```idris example
||| parents "/etc/kernel" == ["/etc/kernel", "/etc", "/"] ||| parents "/etc/kernel" == ["/etc/kernel", "/etc", "/"]
||| ``` ||| ```
export export
parents : String -> List String parents : String -> List String
parents p = map show $ iterate parent' (parse p) parents = map show . List.iterate parent' . parse
||| Determines whether base is either one of the parents of full. ||| Determines whether the base is one of the parents of target.
||| |||
||| Trailing separator is ignored. ||| Trailing separator is ignored.
export
startWith : (base : String) -> (full : String) -> Bool
startWith base full = (parse base) `elem` (iterate parent' (parse full))
||| Returns a path that, when appended onto base, yields full.
||| |||
||| If base is not a prefix of full (i.e., startWith returns false), ||| ```idris example
||| returns Nothing. ||| "/etc" `isBaseOf` "/etc/kernel"
||| ```
export export
stripPrefix : (base : String) -> (full : String) -> Maybe String isBaseOf : (base : String) -> (target : String) -> Bool
stripPrefix base full isBaseOf base target =
= do let MkPath vol1 root1 body1 _ = parse base let
let MkPath vol2 root2 body2 trialSep = parse full MkPath baseVol baseRoot baseBody _ = parse base
if vol1 == vol2 && root1 == root2 then Just () else Nothing MkPath targetVol targetRoot targetBody _ = parse target
body <- stripBody body1 body2 in
pure $ show $ MkPath Nothing False body trialSep baseVol == targetVol
&& baseRoot == targetRoot
&& (baseBody `isPrefixOf` targetBody)
||| Returns a path that, when appended to base, yields target.
|||
||| Returns Nothing if the base is not a prefix of the target.
export
dropBase : (base : String) -> (target : String) -> Maybe String
dropBase base target =
do
let MkPath baseVol baseRoot baseBody _ = parse base
let MkPath targetVol targetRoot targetBody targetTrialSep = parse target
if baseVol == targetVol && baseRoot == targetRoot then Just () else Nothing
body <- dropBody baseBody targetBody
pure $ show $ MkPath Nothing False body targetTrialSep
where where
stripBody : (base : List Body) -> (full : List Body) -> Maybe (List Body) dropBody : (base : List Body) -> (target : List Body) -> Maybe (List Body)
stripBody [] ys = Just ys dropBody [] ys = Just ys
stripBody xs [] = Nothing dropBody xs [] = Nothing
stripBody (x::xs) (y::ys) = if x == y then stripBody xs ys else Nothing dropBody (x::xs) (y::ys) = if x == y then dropBody xs ys else Nothing
||| Returns the final body of the path, if there is one. ||| Returns the last body of the path.
||| |||
||| If the path is a normal file, this is the file name. If it's the ||| If the last body is a file, this is the file name;
||| path of a directory, this is the directory name. ||| if it's a directory, this is the directory name;
||| ||| if it's ".", it recurses on the previous body;
||| Returns Nothing if the final body is "..". ||| if it's "..", returns Nothing.
export export
fileName : String -> Maybe String fileName : String -> Maybe String
fileName p = fileName' (parse p) fileName = fileName' . parse
||| Extracts the stem (non-extension) portion of the file name of path. ||| Extracts the file name in the path without extension.
||| |||
||| The stem is: ||| The stem is:
||| |||
||| - Nothing, if there is no file name; ||| - Nothing, if there is no file name;
||| - The entire file name if there is no embedded "."; ||| - The entire file name if there is no embedded ".";
||| - The entire file name if the file name begins with "." and has ||| - The entire file name if the file name begins with a "." and has no other ".";
||| no other "."s within; ||| - Otherwise, the portion of the file name before the last ".".
||| - Otherwise, the portion of the file name before the final "."
export export
fileStem : String -> Maybe String fileStem : String -> Maybe String
fileStem p = pure $ fst $ splitFileName !(fileName p) fileStem path = pure $ fst $ splitFileName !(fileName path)
||| Extracts the extension of the file name of path. ||| Extracts the extension of the file name in the path.
||| |||
||| The extension is: ||| The extension is:
||| |||
||| - Nothing, if there is no file name; ||| - Nothing, if there is no file name;
||| - Nothing, if there is no embedded "."; ||| - Nothing, if there is no embedded ".";
||| - Nothing, if the file name begins with "." and has no other "."s within; ||| - Nothing, if the file name begins with a "." and has no other ".";
||| - Otherwise, the portion of the file name after the final "." ||| - Otherwise, the portion of the file name after the last ".".
export export
extension : String -> Maybe String extension : String -> Maybe String
extension p = pure $ snd $ splitFileName !(fileName p) extension path = pure $ snd $ splitFileName !(fileName path)
||| Updates the file name of the path. ||| Updates the file name in the path.
||| |||
||| If no file name, this is equivalent to appending the name; ||| If there is no file name, it appends the name to the path;
||| Otherwise it is equivalent to appending the name to the parent. ||| otherwise, it appends the name to the parent of the path.
export export
setFileName : (name : String) -> String -> String setFileName : (name : String) -> String -> String
setFileName name p = show $ setFileName' name (parse p) setFileName name path = show $ setFileName' name (parse path)
||| Append a extension to the path. ||| Appends a extension to the path.
||| |||
||| Returns the path as it is if no file name. ||| If there is no file name, the path will not change;
||| if the path has no extension, the extension will be appended;
||| if the given extension is empty, the extension of the path will be dropped;
||| otherwise, the extension will be replaced.
||| |||
||| If `extension` of the path is Nothing, the extension is added; otherwise ||| ```idris example
||| it is replaced. ||| "/tmp/Foo" <.> "idr" == "/tmp/Foo.idr"
||| ||| ```
||| If the ext is empty, the extension is dropped.
export export
(<.>) : String -> (ext : String) -> String (<.>) : String -> (extension : String) -> String
(<.>) p ext = let p' = parse p (<.>) path ext =
ext = pack $ dropWhile (== '.') (unpack ext) let
ext = if ltrim ext == "" then "" else "." ++ ext in path' = parse path
case fileName' p' of ext = pack $ dropWhile (\char => char == '.' || isSpace char) (unpack ext)
Just name => let (stem, _) = splitFileName name in ext = if ext == "" then "" else "." ++ ext
show $ setFileName' (stem ++ ext) p' in
Nothing => p case fileName' path' of
Just name =>
let (stem, _) = splitFileName name in
show $ setFileName' (stem ++ ext) path'
Nothing => path
||| Drops the extension of the path.
export
dropExtension : String -> String
dropExtension path = path <.> ""