/- Copyright (c) 2021 Henrik Böving. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Henrik Böving -/ import Lean import Std.Data.HashMap import DocGen4.Process import DocGen4.ToHtmlFormat import DocGen4.IncludeStr namespace DocGen4 open Lean Std open scoped DocGen4.Jsx open IO System structure SiteContext where root : String result : AnalyzerResult abbrev HtmlM := Reader SiteContext def getRoot : HtmlM String := do (←read).root def getResult : HtmlM AnalyzerResult := do (←read).result def templateExtends {α β : Type} (base : α → HtmlM β) (new : HtmlM α) : HtmlM β := new >>= base def nameToUrl (n : Name) : String := (parts.intersperse "/").foldl (· ++ ·) "" ++ ".html" where parts := n.components.map Name.toString def nameToDirectory (basePath : FilePath) (n : Name) : FilePath := basePath / parts.foldl (λ acc p => acc / FilePath.mk p) (FilePath.mk ".") where parts := n.components.dropLast.map Name.toString def moduleListFile (file : Name) : HtmlM Html := do
{file.toString}
partial def moduleListDir (h : Hierarchy) : HtmlM Html := do let children := Array.mk (h.getChildren.toList.map Prod.snd) let dirs := children.filter (λ c => c.getChildren.toList.length != 0) let files := children.filter Hierarchy.isFile |>.map Hierarchy.getName return
{h.getName.toString} [(←(dirs.mapM moduleListDir))] [(←(files.mapM moduleListFile))]
def moduleList : HtmlM (Array Html) := do let hierarchy := (←getResult).hierarchy let mut list := Array.empty for (n, cs) in hierarchy.getChildren do list := list.push

{n.toString}

list := list.push $ ←moduleListDir cs list def navbar : HtmlM Html := do def baseHtml (title : String) (site : Html) : HtmlM Html := do {title}

Documentation

{title}

-- TODO: Replace this form with our own search
{site} {←navbar} -- TODO Add the js stuff def notFound : HtmlM Html := do templateExtends (baseHtml "404") $

404 Not Found

Unfortunately, the page you were looking for is no longer here.

def index : HtmlM Html := do templateExtends (baseHtml "Index") $

Welcome to the documentation page

What is up?
def styleCss : String := include_str "./static/style.css" def moduleToHtml (module : Module) : HtmlM Html := do templateExtends (baseHtml module.name.toString) $

This is the page of {module.name.toString}

def htmlOutput (result : AnalyzerResult) : IO Unit := do -- TODO: parameterize this let config := { root := "/", result := result } let basePath := FilePath.mk "./build/doc/" let indexHtml := ReaderT.run index config let notFoundHtml := ReaderT.run notFound config FS.createDirAll basePath FS.writeFile (basePath / "index.html") indexHtml.toString FS.writeFile (basePath / "style.css") styleCss FS.writeFile (basePath / "404.html") notFoundHtml.toString for (module, content) in result.modules.toArray do let moduleHtml := ReaderT.run (moduleToHtml content) config let path := basePath / (nameToUrl module) FS.createDirAll $ nameToDirectory basePath module FS.writeFile path moduleHtml.toString end DocGen4