bookshelf-doc/DocGen4/Output/Module.lean

190 lines
6.7 KiB
Plaintext
Raw Normal View History

/-
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 DocGen4.Output.Template
import DocGen4.Output.Inductive
import DocGen4.Output.Structure
import DocGen4.Output.Class
import DocGen4.Output.Definition
2022-02-04 21:48:08 +00:00
import DocGen4.Output.Instance
2022-02-06 19:14:36 +00:00
import DocGen4.Output.ClassInductive
2022-02-17 05:47:38 +00:00
import DocGen4.Output.DocString
2022-02-15 18:19:40 +00:00
import Lean.Data.Xml.Parser
namespace DocGen4
namespace Output
open scoped DocGen4.Jsx
open Lean
def argToHtml (arg : Arg) : HtmlM Html := do
let (l, r, implicit) := match arg.binderInfo with
| BinderInfo.default => ("(", ")", false)
| BinderInfo.implicit => ("{", "}", true)
| BinderInfo.strictImplicit => ("⦃", "⦄", true)
| BinderInfo.instImplicit => ("[", "]", true)
-- TODO: Can this ever be reached here? What does it mean?
| BinderInfo.auxDecl => unreachable!
let mut nodes := #[Html.text s!"{l}{arg.name.toString} : "]
nodes := nodes.append (←infoFormatToHtml arg.type)
nodes := nodes.push r
let inner := Html.element "span" true #[("class", "fn")] nodes
let html := Html.element "span" false #[("class", "decl_args")] #[inner]
if implicit then
2022-02-12 14:09:13 +00:00
pure <span «class»="impl_arg">{html}</span>
else
2022-02-12 14:09:13 +00:00
pure html
def structureInfoHeader (s : StructureInfo) : HtmlM (Array Html) := do
let mut nodes := #[]
if s.parents.size > 0 then
nodes := nodes.push <span «class»="decl_extends">extends</span>
let mut parents := #[]
for parent in s.parents do
let link := <a «class»="break_within" href={←declNameToLink parent}>{parent.toString}</a>
let inner := Html.element "span" true #[("class", "fn")] #[link]
let html:= Html.element "span" false #[("class", "decl_parent")] #[inner]
parents := parents.push html
nodes := nodes.append (parents.toList.intersperse (Html.text ", ")).toArray
2022-02-12 14:09:13 +00:00
pure nodes
def docInfoHeader (doc : DocInfo) : HtmlM Html := do
let mut nodes := #[]
nodes := nodes.push <span «class»="decl_kind">{doc.getKindDescription}</span>
2021-12-17 16:20:44 +00:00
nodes := nodes.push
<span «class»="decl_name">
<a «class»="break_within" href={←declNameToLink doc.getName}>
-- TODO: HTMLify the name
{doc.getName.toString}
</a>
</span>
for arg in doc.getArgs do
nodes := nodes.push (←argToHtml arg)
-- TODO: dedup this
match doc with
| DocInfo.structureInfo i => nodes := nodes.append (←structureInfoHeader i)
| DocInfo.classInfo i => nodes := nodes.append (←structureInfoHeader i.toStructureInfo)
| _ => nodes := nodes
nodes := nodes.push <span «class»="decl_args">:</span>
nodes := nodes.push $ Html.element "div" true #[("class", "decl_type")] (←infoFormatToHtml doc.getType)
2022-02-12 14:09:13 +00:00
pure <div «class»="decl_header"> [nodes] </div>
2022-01-09 15:57:19 +00:00
def docInfoToHtml (module : Name) (doc : DocInfo) : HtmlM Html := do
2022-02-18 03:28:44 +00:00
let docInfoHtml ← match doc with
| DocInfo.inductiveInfo i => inductiveToHtml i
| DocInfo.structureInfo i => structureToHtml i
| DocInfo.classInfo i => classToHtml i
| DocInfo.definitionInfo i => definitionToHtml i
2022-02-04 21:48:08 +00:00
| DocInfo.instanceInfo i => instanceToHtml i
2022-02-06 19:14:36 +00:00
| DocInfo.classInductiveInfo i => classInductiveToHtml i
2022-02-18 03:28:44 +00:00
| i => pure #[]
let docStringHtml ← match doc.getDocString with
| some s => docStringToHtml s
| none => pure #[]
2022-02-13 02:32:53 +00:00
let attrs := doc.getAttrs
let attrsHtml :=
if attrs.size > 0 then
let attrStr := "@[" ++ String.intercalate ", " doc.getAttrs.toList ++ "]"
#[Html.element "div" false #[("class", "attributes")] #[attrStr]]
else
#[]
2022-02-12 14:09:13 +00:00
pure
<div «class»="decl" id={doc.getName.toString}>
<div «class»={doc.getKind}>
<div «class»="gh_link">
<a href={←getSourceUrl module doc.getDeclarationRange}>source</a>
</div>
2022-02-13 02:32:53 +00:00
[attrsHtml]
2022-02-12 14:09:13 +00:00
{←docInfoHeader doc}
2022-02-18 03:28:44 +00:00
[docStringHtml]
[docInfoHtml]
</div>
</div>
2022-02-15 17:06:46 +00:00
def modDocToHtml (module : Name) (mdoc : ModuleDoc) : HtmlM Html := do
2022-02-17 05:47:38 +00:00
pure
2022-02-15 17:06:46 +00:00
<div «class»="mod_doc">
2022-02-17 13:26:02 +00:00
[←docStringToHtml mdoc.doc]
2022-02-15 17:06:46 +00:00
</div>
2022-02-18 03:28:44 +00:00
def moduleMemberToHtml (module : Name) (member : ModuleMember) : HtmlM Html := do
2022-02-15 17:06:46 +00:00
match member with
| ModuleMember.docInfo d => docInfoToHtml module d
| ModuleMember.modDoc d => modDocToHtml module d
2022-01-16 13:22:53 +00:00
def declarationToNavLink (module : Name) : Html :=
<div «class»="nav_link">
<a «class»="break_within" href={s!"#{module.toString}"}>{module.toString}</a>
</div>
2022-01-16 13:22:53 +00:00
-- TODO: Similar functions are used all over the place, we should dedup them
def moduleToNavLink (module : Name) : HtmlM Html := do
2022-02-12 14:09:13 +00:00
pure <a href={←moduleNameToLink module}>{module.toString}</a>
2022-01-16 13:22:53 +00:00
def getImports (module : Name) : HtmlM (Array Name) := do
let res ← getResult
2022-02-12 14:09:13 +00:00
let some idx := res.moduleNames.findIdx? (. == module) | unreachable!
2022-01-16 13:22:53 +00:00
let adj := res.importAdj.get! idx
let mut imports := #[]
for i in [:adj.size] do
if adj.get! i then
imports := imports.push (res.moduleNames.get! i)
2022-02-12 14:09:13 +00:00
pure imports
2022-01-16 13:22:53 +00:00
def getImportedBy (module : Name) : HtmlM (Array Name) := do
let res ← getResult
2022-02-12 14:09:13 +00:00
let some idx := res.moduleNames.findIdx? (. == module) | unreachable!
2022-01-16 13:22:53 +00:00
let adj := res.importAdj
let mut impBy := #[]
for i in [:adj.size] do
if adj.get! i |>.get! idx then
impBy := impBy.push (res.moduleNames.get! i)
2022-02-12 14:09:13 +00:00
pure impBy
2022-01-16 13:22:53 +00:00
def importedByHtml (moduleName : Name) : HtmlM (Array Html) := do
let imports := (←getImportedBy moduleName) |>.qsort Name.lt
2022-02-12 14:09:13 +00:00
imports.mapM (λ i => do pure <li>{←moduleToNavLink i}</li>)
2022-01-16 13:22:53 +00:00
def importsHtml (moduleName : Name) : HtmlM (Array Html) := do
let imports := (←getImports moduleName) |>.qsort Name.lt
2022-02-12 14:09:13 +00:00
imports.mapM (λ i => do pure <li>{←moduleToNavLink i}</li>)
2022-01-16 13:22:53 +00:00
2022-01-09 15:57:19 +00:00
def internalNav (members : Array Name) (moduleName : Name) : HtmlM Html := do
2022-02-12 14:09:13 +00:00
pure
<nav «class»="internal_nav">
<h3><a «class»="break_within" href="#top">{moduleName.toString}</a></h3>
<p «class»="gh_nav_link"><a href={←getSourceUrl moduleName none}>source</a></p>
<div «class»="imports">
<details>
<summary>Imports</summary>
<ul>
[←importsHtml moduleName]
</ul>
</details>
<details>
<summary>Imported by</summary>
<ul>
[←importedByHtml moduleName]
</ul>
</details>
</div>
[members.map declarationToNavLink]
</nav>
def moduleToHtml (module : Module) : HtmlM Html := withReader (setCurrentName module.name) do
2022-02-18 04:52:01 +00:00
let memberDocs ← module.members.mapM (λ i => moduleMemberToHtml module.name i)
let memberNames := filterMapDocInfo module.members |>.map DocInfo.getName
2022-02-12 14:09:13 +00:00
templateExtends (baseHtmlArray module.name.toString) $ pure #[
2022-02-18 04:52:01 +00:00
←internalNav memberNames module.name,
Html.element "main" false #[] memberDocs
]
end Output
end DocGen4