From f656b979286ba00ea9af17e1263d889f2661b6b5 Mon Sep 17 00:00:00 2001 From: Johann Rudloff Date: Thu, 29 Apr 2021 11:56:47 +0200 Subject: [PATCH] [ refactor ] Move all HTML-related code out of the Idris.Package module --- src/Idris/Doc/HTML.idr | 31 +++++++++++++++++++++++++--- src/Idris/Package.idr | 47 ++++++++++++++---------------------------- 2 files changed, 43 insertions(+), 35 deletions(-) diff --git a/src/Idris/Doc/HTML.idr b/src/Idris/Doc/HTML.idr index 1f6a836c5..1b5c79509 100644 --- a/src/Idris/Doc/HTML.idr +++ b/src/Idris/Doc/HTML.idr @@ -13,6 +13,7 @@ import Libraries.Text.PrettyPrint.Prettyprinter.Render.HTML import Libraries.Text.PrettyPrint.Prettyprinter.SimpleDocTree import Idris.DocString +import Idris.Package.Types import Idris.Pretty import Idris.Version @@ -78,7 +79,6 @@ renderHtml (STAnn ann rest) = do pure $ "" ++ resthtml ++ "" renderHtml (STConcat docs) = pure $ fastConcat !(traverse renderHtml docs) -export docDocToHtml : {auto c : Ref Ctxt Defs} -> Doc IdrisDocAnn -> Core String @@ -86,7 +86,6 @@ docDocToHtml doc = let dt = SimpleDocTree.fromStream $ layoutUnbounded doc in renderHtml dt -export htmlPreamble : String -> String -> String -> String htmlPreamble title root class = "" ++ "" ++ htmlEscape title ++ "" @@ -96,6 +95,32 @@ htmlPreamble title root class = "Index" ++ "
" -export htmlFooter : String htmlFooter = "
" + +export +renderDocIndex : PkgDesc -> String +renderDocIndex pkg = fastConcat $ + [ htmlPreamble (name pkg) "" "index" + , "

Package ", name pkg, " - Namespaces

" + , "" + , htmlFooter + ] + where + moduleLink : (ModuleIdent, String) -> String + moduleLink (mod, filename) = + "
  • " ++ (show mod) ++ "
  • " + +export +renderModuleDoc : {auto c : Ref Ctxt Defs} -> + ModuleIdent -> + Doc IdrisDocAnn -> + Core String +renderModuleDoc mod allModuleDocs = pure $ fastConcat + [ htmlPreamble (show mod) "../" "namespace" + , "

    ", show mod, "

    " + , !(docDocToHtml allModuleDocs) + , htmlFooter + ] diff --git a/src/Idris/Package.idr b/src/Idris/Package.idr index 074ff398e..dab81b4d7 100644 --- a/src/Idris/Package.idr +++ b/src/Idris/Package.idr @@ -426,15 +426,12 @@ makeDoc pkg opts = let build = build_dir (dirs (options defs)) let docBase = build "docs" let docDir = docBase "docs" - ignore $ coreLift $ mkdirAll docDir + Right () <- coreLift $ mkdirAll docDir + | Left err => fileError docDir err u <- newRef UST initUState setPPrint (MkPPOpts False False True) [] <- concat <$> for (modules pkg) (\(mod, filename) => do - let outputFileName = (show mod) ++ ".html" - Right outFile <- coreLift $ openFile (docDir outputFileName) WriteTruncate - | Left err => pure [InternalError $ ("error opening file \"" ++ (docDir outputFileName) ++ "\": " ++ (show err))] - let writeHtml = \s => (coreLift_ $ fPutStrLn outFile s) let ns = miAsNamespace mod addImport (MkImport emptyFC False mod ns) defs <- get Ctxt @@ -442,38 +439,21 @@ makeDoc pkg opts = let allInNamespace = filter (inNS ns) names visibleNames <- filterM (visible defs) allInNamespace - writeHtml $ htmlPreamble (show mod) "../" "namespace" - writeHtml ("

    " ++ show mod ++ "

    ") - writeHtml ("
    ") - ignore $ for (sort visibleNames) (\name => do - doc <- getDocsForName emptyFC name - writeHtml !(docDocToHtml doc) - ) - writeHtml ("
    ") - writeHtml htmlFooter - coreLift $ closeFile outFile + let outputFilePath = docDir (show mod ++ ".html") + allDocs <- annotate Declarations <$> vcat <$> for (sort visibleNames) (getDocsForName emptyFC) + Right () <- coreLift $ writeFile outputFilePath !(renderModuleDoc mod allDocs) + | Left err => fileError (docBase "index.html") err + pure $ the (List Error) [] ) | errs => pure errs - Right outFile <- coreLift $ openFile (docBase "index.html") WriteTruncate - | Left err => pure [InternalError $ ("error opening file \"" ++ (docBase "index.html") ++ "\": " ++ (show err))] - let writeHtml = \s => (coreLift_ $ fPutStrLn outFile s) - writeHtml $ htmlPreamble (name pkg) "" "index" - writeHtml ("

    Package " ++ name pkg ++ " - Namespaces

    ") - writeHtml "" - writeHtml htmlFooter - coreLift_ $ closeFile outFile + Right () <- coreLift $ writeFile (docBase "index.html") $ renderDocIndex pkg + | Left err => fileError (docBase "index.html") err css <- readDataFile "docs/styles.css" - Right outFile <- coreLift $ openFile (docBase "styles.css") WriteTruncate - | Left err => pure [InternalError $ ("error opening file \"" ++ (docBase "styles.css") ++ "\": " ++ (show err))] - coreLift_ $ fPutStr outFile css - coreLift_ $ closeFile outFile + Right () <- coreLift $ writeFile (docBase "styles.css") css + | Left err => fileError (docBase "styles.css") err runScript (postbuild pkg) pure [] @@ -482,7 +462,7 @@ makeDoc pkg opts = visible defs n = do Just def <- lookupCtxtExact n (gamma defs) | Nothing => pure False - -- TODO: if we can find out, wheter a def has been declared as + -- TODO: if we can find out, whether a def has been declared as -- part of an interface, hide it here pure $ case definition def of (DCon _ _ _) => False @@ -497,6 +477,9 @@ makeDoc pkg opts = else full stripNS _ x = x + fileError : String -> FileError -> Core (List Error) + fileError filename err = pure [FileErr filename err] + -- Data.These.bitraverse hand specialised for Core bitraverseC : (a -> Core c) -> (b -> Core d) -> These a b -> Core (These c d) bitraverseC f g (This a) = [| This (f a) |]