[ fix ] support .lidr.md and .lidr.tex extensions (#3071)

This commit is contained in:
Steve Dunham 2023-09-25 14:25:26 -07:00 committed by GitHub
parent 3886200d29
commit 46483fd120
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
8 changed files with 55 additions and 4 deletions

View File

@ -216,7 +216,7 @@ mbPathToNS wdir sdir fname =
sdir = fromMaybe "" sdir
base = if isAbsolute fname then wdir </> sdir else sdir
in
unsafeFoldModuleIdent . reverse . splitPath . Path.dropExtension
unsafeFoldModuleIdent . reverse . splitPath . Path.dropExtensions
<$> (Path.dropBase `on` cleanPath) base fname
export

View File

@ -54,7 +54,7 @@ findModules start = do
go acc ((path, (root ** iot)) :: iots) = do
t <- liftIO iot
let mods = flip map t.files $ \ entry =>
let fname = fst (splitFileName (fileName entry)) in
let fname = fst (splitExtensions (fileName entry)) in
let mod = unsafeFoldModuleIdent (fname :: path) in
let fp = toFilePath entry in
(mod, fp)

View File

@ -1083,7 +1083,7 @@ process (ImportPackage package) = do
tree <- coreLift $ explore packageDirPath
fentries <- coreLift $ toPaths (toRelative tree)
errs <- for fentries $ \entry => do
let entry' = dropExtension entry
let entry' = dropExtensions entry
let sp = forget $ split (== dirSeparator) entry'
let ns = concat $ intersperse "." sp
let ns' = mkNamespace ns

View File

@ -21,6 +21,7 @@ import System.File
infixl 5 </>, />
infixr 7 <.>
infixr 7 <..>
||| The character that separates directories in the path.
@ -595,11 +596,43 @@ export
show $ setFileName' (stem ++ ext) path'
Nothing => path
||| Appends a extension to the path, replacing multiple extensions.
|||
||| 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, all of the extensions of the path will be dropped;
||| otherwise, the extensions will be replaced.
|||
||| ```idris example
||| "/tmp/Foo.lidr.md" <.> "idr" == "/tmp/Foo.idr"
||| ```
export
(<..>) : String -> (extension : String) -> String
(<..>) path ext =
let
path' = parse path
ext = pack $ dropWhile (\char => char == '.' || isSpace char) (unpack ext)
ext = if ext == "" then "" else "." ++ ext
in
case fileName' path' of
Just name =>
let (stem, _) = splitExtensions name in
show $ setFileName' (stem ++ ext) path'
Nothing => path
||| Drops the extension of the path.
export
dropExtension : String -> String
dropExtension path = path <.> ""
||| Drops all of the extensions of a path.
||| ```idris example
||| "/tmp/Foo.lidr.md" == "/tmp/Foo"
export
dropExtensions : String -> String
dropExtensions path = path <..> ""
||| Looks up an executable from a list of candidate names in the PATH
export
pathLookup : List String -> IO (Maybe String)

View File

@ -47,12 +47,15 @@ styleTeX = MkLitStyle
export
listOfExtensionsLiterate : List String
listOfExtensionsLiterate
= concatMap file_extensions
= do let exts = concatMap file_extensions
[ styleBird
, styleOrg
, styleCMark
, styleTeX
]
pfx <- [ "", ".idr", ".lidr"]
ext <- exts
pure (pfx ++ ext)
||| Are we dealing with a valid literate file name, if so return the base name and used extension.
export

View File

@ -0,0 +1,11 @@
# Markdown test
Idris should allow multiple extensions on literate files with a module declaration.
```idris
module Test
main : IO ()
main = printLn "hello"
```

View File

@ -0,0 +1 @@
1/1: Building Test (Test.lidr.md)

View File

@ -0,0 +1,3 @@
rm -rf build
$1 --no-color --console-width 0 --no-banner --check Test.lidr.md