Merge pull request #1 from hargoniX/main

Initial implementation
main
Henrik Böving 2022-01-03 14:29:58 +01:00 committed by GitHub
commit 85d1e4608c
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
20 changed files with 1839 additions and 0 deletions

2
.gitignore vendored Normal file
View File

@ -0,0 +1,2 @@
/build
/lean_packages

11
DocGen4.lean Normal file
View File

@ -0,0 +1,11 @@
/-
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.Hierarchy
import DocGen4.Process
import DocGen4.Load
import DocGen4.ToHtmlFormat
import DocGen4.IncludeStr
import DocGen4.Output

82
DocGen4/Hierarchy.lean Normal file
View File

@ -0,0 +1,82 @@
/-
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
namespace DocGen4
open Lean Std Name
def getNLevels (name : Name) (levels: Nat) : Name :=
(components.drop (components.length - levels)).reverse.foldl (· ++ ·) Name.anonymous
where
components := name.components'
inductive Hierarchy where
| node (name : Name) (isFile : Bool) (children : RBNode Name (λ _ => Hierarchy)) : Hierarchy
instance : Inhabited Hierarchy := ⟨Hierarchy.node Name.anonymous false RBNode.leaf⟩
abbrev HierarchyMap := RBNode Name (λ _ => Hierarchy)
-- Everything in this namespace is adapted from stdlib's RBNode
namespace HierarchyMap
def toList : HierarchyMap → List (Name × Hierarchy)
| t => t.revFold (fun ps k v => (k, v)::ps) []
def hForIn [Monad m] (t : HierarchyMap) (init : σ) (f : (Name × Hierarchy) → σ → m (ForInStep σ)) : m σ :=
t.forIn init (fun a b acc => f (a, b) acc)
instance : ForIn m HierarchyMap (Name × Hierarchy) where
forIn := HierarchyMap.hForIn
end HierarchyMap
namespace Hierarchy
def empty (n : Name) (isFile : Bool) : Hierarchy :=
node n isFile RBNode.leaf
def getName : Hierarchy → Name
| node n _ _ => n
def getChildren : Hierarchy → HierarchyMap
| node _ _ c => c
def isFile : Hierarchy → Bool
| node _ f _ => f
partial def insert! (h : Hierarchy) (n : Name) : Hierarchy := Id.run $ do
let hn := h.getName
let mut cs := h.getChildren
assert! getNumParts hn ≤ getNumParts n
if getNumParts hn + 1 == getNumParts n then
match cs.find Name.cmp n with
| none =>
node hn h.isFile (cs.insert Name.cmp n $ empty n true)
| some (node _ true _) => h
| some hierarchy@(node _ false ccs) =>
cs := cs.erase Name.cmp n
node hn h.isFile (cs.insert Name.cmp n $ node n true ccs)
else
let leveledName := getNLevels n (getNumParts hn + 1)
match cs.find Name.cmp leveledName with
| some nextLevel =>
cs := cs.erase Name.cmp leveledName
-- BUG?
node hn h.isFile $ cs.insert Name.cmp leveledName (nextLevel.insert! n)
| none =>
let child := (insert! (empty leveledName false) n)
node hn h.isFile $ cs.insert Name.cmp leveledName child
partial def fromArray (names : Array Name) : Hierarchy :=
names.foldl insert! (empty anonymous false)
end Hierarchy
end DocGen4

26
DocGen4/IncludeStr.lean Normal file
View File

@ -0,0 +1,26 @@
/-
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
namespace DocGen4
open Lean System IO Lean.Elab.Term
syntax (name := includeStr) "include_str" str : term
@[termElab includeStr] def includeStrImpl : TermElab := λ stx expectedType? => do
let str := stx[1].isStrLit?.get!
let path := FilePath.mk str
if ←path.pathExists then
if ←path.isDir then
throwError s!"{str} is a directory"
else
let content ← FS.readFile path
return mkStrLit content
else
throwError s!"\"{str}\" does not exist as a file"
end DocGen4

51
DocGen4/Load.lean Normal file
View File

@ -0,0 +1,51 @@
/-
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 DocGen4.Process
import Std.Data.HashMap
namespace DocGen4
open Lean System Std IO
def getLakePath : IO FilePath := do
match (← IO.getEnv "LAKE") with
| some path => System.FilePath.mk path
| none =>
let lakePath := (←findSysroot?) / "bin" / "lake"
lakePath.withExtension System.FilePath.exeExtension
-- Modified from the LSP Server
def lakeSetupSearchPath (lakePath : System.FilePath) (imports : Array String) : IO Lean.SearchPath := do
let args := #["print-paths"] ++ imports
let cmdStr := " ".intercalate (toString lakePath :: args.toList)
let lakeProc ← Process.spawn {
stdin := Process.Stdio.null
stdout := Process.Stdio.piped
stderr := Process.Stdio.piped
cmd := lakePath.toString
args
}
let stdout := String.trim (← lakeProc.stdout.readToEnd)
let stderr := String.trim (← lakeProc.stderr.readToEnd)
match (← lakeProc.wait) with
| 0 =>
let stdout := stdout.split (· == '\n') |>.getLast!
let Except.ok (paths : LeanPaths) ← pure (Json.parse stdout >>= fromJson?)
| throw $ userError s!"invalid output from `{cmdStr}`:\n{stdout}\nstderr:\n{stderr}"
initSearchPath (← findSysroot?) paths.oleanPath
paths.oleanPath.mapM realPathNormalized
| 2 => pure [] -- no lakefile.lean
| _ => throw $ userError s!"`{cmdStr}` failed:\n{stdout}\nstderr:\n{stderr}"
def load (imports : List Name) : IO AnalyzerResult := do
let env ← importModules (List.map (Import.mk · false) imports) Options.empty
-- TODO parameterize maxHeartbeats
IO.println "Processing modules"
Prod.fst <$> (Meta.MetaM.toIO process { maxHeartbeats := 100000000, options := ⟨[(`pp.tagAppFns, true)]⟩ } { env := env} {} {})
end DocGen4

35
DocGen4/Output.lean Normal file
View File

@ -0,0 +1,35 @@
/-
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 DocGen4.Process
import DocGen4.Output.Base
import DocGen4.Output.Index
import DocGen4.Output.Module
import DocGen4.Output.NotFound
namespace DocGen4
open Lean Std IO System Output
def htmlOutput (result : AnalyzerResult) : IO Unit := do
-- TODO: parameterize this
let config := { root := "/", result := result, currentName := none}
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
FS.writeFile (basePath / "nav.js") navJs
for (module, content) in result.moduleInfo.toArray do
let moduleHtml := ReaderT.run (moduleToHtml content) config
let path := moduleNameToFile basePath module
FS.createDirAll $ moduleNameToDirectory basePath module
FS.writeFile path moduleHtml.toString
end DocGen4

53
DocGen4/Output/Base.lean Normal file
View File

@ -0,0 +1,53 @@
/-
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 DocGen4.Process
import DocGen4.IncludeStr
namespace DocGen4
namespace Output
open Lean System
structure SiteContext where
root : String
result : AnalyzerResult
currentName : Option Name
def setCurrentName (name : Name) (ctx : SiteContext) := {ctx with currentName := some name}
abbrev HtmlT := ReaderT SiteContext
abbrev HtmlM := HtmlT Id
def getRoot : HtmlM String := do (←read).root
def getResult : HtmlM AnalyzerResult := do (←read).result
def getCurrentName : HtmlM (Option Name) := do (←read).currentName
def templateExtends {α β : Type} (base : α → HtmlM β) (new : HtmlM α) : HtmlM β :=
new >>= base
-- TODO: Change this to HtmlM and auto add the root URl
def moduleNameToLink (n : Name) : HtmlM String := do
let parts := n.components.map Name.toString
(←getRoot) ++ (parts.intersperse "/").foldl (· ++ ·) "" ++ ".html"
def moduleNameToFile (basePath : FilePath) (n : Name) : FilePath :=
FilePath.withExtension (basePath / parts.foldl (· / ·) (FilePath.mk ".")) "html"
where
parts := n.components.map Name.toString
def moduleNameToDirectory (basePath : FilePath) (n : Name) : FilePath :=
basePath / parts.foldl (· / ·) (FilePath.mk ".")
where
parts := n.components.dropLast.map Name.toString
section Static
def styleCss : String := include_str "./static/style.css"
def navJs : String := include_str "./static/nav.js"
end Static
end Output
end DocGen4

22
DocGen4/Output/Index.lean Normal file
View File

@ -0,0 +1,22 @@
/-
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.ToHtmlFormat
import DocGen4.Output.Template
namespace DocGen4
namespace Output
open scoped DocGen4.Jsx
def index : HtmlM Html := do templateExtends (baseHtml "Index") $
<main>
<a id="top"></a>
<h1> Welcome to the documentation page </h1>
What is up?
</main>
end Output
end DocGen4

111
DocGen4/Output/Module.lean Normal file
View File

@ -0,0 +1,111 @@
/-
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 Lean.PrettyPrinter
import Lean.Widget.TaggedText
import DocGen4.ToHtmlFormat
import DocGen4.Output.Template
namespace DocGen4
namespace Output
open scoped DocGen4.Jsx
open Lean PrettyPrinter Widget Elab
def declNameToLink (name : Name) : HtmlM String := do
let res ← getResult
let module := res.moduleNames[res.name2ModIdx.find! name]
(←moduleNameToLink module) ++ "#" ++ name.toString
def splitWhitespaces (s : String) : (String × String × String) := Id.run do
let front := "".pushn ' ' (s.find (!Char.isWhitespace ·))
let mut s := s.trimLeft
let back := "".pushn ' ' (s.length - s.offsetOfPos (s.find Char.isWhitespace))
s:= s.trimRight
(front, s, back)
partial def infoFormatToHtml (i : CodeWithInfos) : HtmlM (Array Html) := do
match i with
| TaggedText.text t => return #[t]
| TaggedText.append tt => tt.foldlM (λ acc t => do acc ++ (←infoFormatToHtml t)) #[]
| TaggedText.tag a t =>
match a.info.val.info with
| Info.ofTermInfo i =>
match i.expr.consumeMData with
| Expr.const name _ _ =>
match t with
| TaggedText.text t =>
let (front, t, back) := splitWhitespaces t
let elem := Html.element "a" true #[("href", ←declNameToLink name)] #[t]
#[Html.text front, elem, Html.text back]
| _ =>
-- TODO: Is this ever reachable?
#[Html.element "a" true #[("href", ←declNameToLink name)] (←infoFormatToHtml t)]
| _ =>
#[Html.element "span" true #[("class", "fn")] (←infoFormatToHtml t)]
| _ => #[Html.element "span" true #[("class", "fn")] (←infoFormatToHtml t)]
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
<span «class»="impl_arg">{html}</span>
else
html
def docInfoHeader (doc : DocInfo) : HtmlM Html := do
let mut nodes := #[]
-- TODO: noncomputable, partial
-- TODO: Support all the kinds in CSS
nodes := nodes.push <span «class»="decl_kind">{doc.getKind}</span>
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)
nodes := nodes.push <span «class»="decl_args">:</span>
nodes := nodes.push $ Html.element "div" true #[("class", "decl_type")] (←infoFormatToHtml doc.getType)
-- TODO: The final type of the declaration
return <div «class»="decl_header"> [nodes] </div>
def docInfoToHtml (doc : DocInfo) : HtmlM Html := do
<div «class»="decl" id={doc.getName.toString}>
<div «class»={doc.getKind}>
<div «class»="gh_link">
-- TODO: Put the proper source link
<a href="https://github.com">source</a>
</div>
-- TODO: Attributes
{←docInfoHeader doc}
-- TODO: The actual type information we are here for
</div>
</div>
def moduleToHtml (module : Module) : HtmlM Html := withReader (setCurrentName module.name) do
-- TODO: Probably some sort of ordering by line number would be cool?
-- maybe they should already be ordered in members.
let docInfos ← module.members.mapM docInfoToHtml
-- TODO: This is missing imports, imported by, source link, list of decls
templateExtends (baseHtml module.name.toString) $
Html.element "main" false #[] docInfos
end Output
end DocGen4

View File

@ -0,0 +1,71 @@
/-
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 DocGen4.ToHtmlFormat
import DocGen4.Output.Base
namespace DocGen4
namespace Output
open Lean
open scoped DocGen4.Jsx
def moduleListFile (file : Name) : HtmlM Html := do
let attributes := match ←getCurrentName with
| some name =>
if file == name then
#[("class", "nav_link"), ("visible", "")]
else
#[("class", "nav_link")]
| none => #[("class", "nav_link")]
let nodes := #[<a href={s!"{←moduleNameToLink file}"}>{file.toString}</a>]
return Html.element "div" false attributes nodes
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
let dirNodes ← (dirs.mapM moduleListDir)
let fileNodes ← (files.mapM moduleListFile)
let moduleLink ← moduleNameToLink h.getName
let attributes := match ←getCurrentName with
| some name =>
if h.getName.isPrefixOf name then
#[("class", "nav_sect"), ("data-path", moduleLink), ("open", "")]
else
#[("class", "nav_sect"), ("data-path", moduleLink)]
| none =>
#[("class", "nav_sect"), ("data-path", moduleLink)]
let nodes := #[<summary>{h.getName.toString}</summary>] ++ dirNodes ++ fileNodes
return Html.element "details" false attributes nodes
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 <h4>{n.toString}</h4>
list := list.push $ ←moduleListDir cs
list
def navbar : HtmlM Html := do
<nav «class»="nav">
<h3>General documentation</h3>
<div «class»="nav_link"><a href={s!"{←getRoot}"}>index</a></div>
/-
TODO: Add these in later
<div «class»="nav_link"><a href={s!"{←getRoot}tactics.html"}>tactics</a></div>
<div «class»="nav_link"><a href={s!"{←getRoot}commands.html"}>commands</a></div>
<div «class»="nav_link"><a href={s!"{←getRoot}hole_commands.html"}>hole commands</a></div>
<div «class»="nav_link"><a href={s!"{←getRoot}attributes.html"}>attributes</a></div>
<div «class»="nav_link"><a href={s!"{←getRoot}notes.html"}>notes</a></div>
<div «class»="nav_link"><a href={s!"{←getRoot}references.html"}>references</a></div>
-/
<h3>Library</h3>
[←moduleList]
</nav>
end Output
end DocGen4

View File

@ -0,0 +1,22 @@
/-
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.ToHtmlFormat
import DocGen4.Output.Template
namespace DocGen4
namespace Output
open scoped DocGen4.Jsx
def notFound : HtmlM Html := do templateExtends (baseHtml "404") $
<main>
<h1>404 Not Found</h1>
<p> Unfortunately, the page you were looking for is no longer here. </p>
<div id="howabout"></div>
</main>
end Output
end DocGen4

View File

@ -0,0 +1,57 @@
/-
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.ToHtmlFormat
import DocGen4.Output.Navbar
namespace DocGen4
namespace Output
open scoped DocGen4.Jsx
def baseHtml (title : String) (site : Html) : HtmlM Html := do
<html lang="en">
<head>
<link rel="stylesheet" href={s!"{←getRoot}style.css"}/>
<link rel="stylesheet" href={s!"{←getRoot}pygments.css"}/>
<link rel="shortcut icon" href={s!"{←getRoot}favicon.ico"}/>
<title>{title}</title>
<meta charset="UTF-8"/>
<meta name="viewport" content="width=device-width, initial-scale=1"/>
</head>
<body>
<input id="nav_toggle" type="checkbox"/>
<header>
<h1><label «for»="nav_toggle"></label>Documentation</h1>
<p «class»="header_filename break_within">{title}</p>
-- TODO: Replace this form with our own search
<form action="https://google.com/search" method="get" id="search_form">
<input type="hidden" name="sitesearch" value="https://leanprover-community.github.io/mathlib_docs"/>
<input type="text" name="q" autocomplete="off"/>
<button>Google site search</button>
</form>
</header>
<nav «class»="internal_nav"></nav>
{site}
{←navbar}
-- Lean in JS in HTML in Lean...very meta
<script>
siteRoot = "{←getRoot}";
</script>
-- TODO Add more js stuff
<script src={s!"{←getRoot}nav.js"}></script>
</body>
</html>
end Output
end DocGen4

349
DocGen4/Process.lean Normal file
View File

@ -0,0 +1,349 @@
/-
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 Lean.PrettyPrinter
import Std.Data.HashMap
import Lean.Meta.SynthInstance
import DocGen4.Hierarchy
namespace DocGen4
open Lean Meta PrettyPrinter Std Widget
structure NameInfo where
name : Name
type : CodeWithInfos
deriving Inhabited
structure Arg where
name : Name
type : CodeWithInfos
binderInfo : BinderInfo
structure Info extends NameInfo where
args : Array Arg
doc : Option String
declarationRange : DeclarationRange
deriving Inhabited
structure AxiomInfo extends Info where
isUnsafe : Bool
deriving Inhabited
structure TheoremInfo extends Info
deriving Inhabited
structure OpaqueInfo extends Info where
value : CodeWithInfos
isUnsafe : Bool
deriving Inhabited
structure DefinitionInfo extends Info where
--value : CodeWithInfos
unsafeInformation : DefinitionSafety
hints : ReducibilityHints
deriving Inhabited
abbrev InstanceInfo := DefinitionInfo
structure InductiveInfo extends Info where
numParams : Nat -- Number of parameters
numIndices : Nat -- Number of indices
all : List Name -- List of all (including this one) inductive datatypes in the mutual declaration containing this one
ctors : List NameInfo -- List of all constructors and their type for this inductive datatype
isRec : Bool -- `true` Iff it is recursive
isUnsafe : Bool
isReflexive : Bool
isNested : Bool
deriving Inhabited
structure FieldInfo extends NameInfo where
projFn : Name
subobject? : Option Name
deriving Inhabited
structure StructureInfo extends Info where
fieldInfo : Array FieldInfo
parents : Array Name
ctor : NameInfo
deriving Inhabited
structure ClassInfo extends StructureInfo where
hasOutParam : Bool
instances : Array CodeWithInfos
deriving Inhabited
inductive DocInfo where
| axiomInfo (info : AxiomInfo) : DocInfo
| theoremInfo (info : TheoremInfo) : DocInfo
| opaqueInfo (info : OpaqueInfo) : DocInfo
| definitionInfo (info : DefinitionInfo) : DocInfo
| instanceInfo (info : InstanceInfo) : DocInfo
| inductiveInfo (info : InductiveInfo) : DocInfo
| structureInfo (info : StructureInfo) : DocInfo
| classInfo (info : ClassInfo) : DocInfo
deriving Inhabited
structure Module where
name : Name
doc : Option String
members : Array DocInfo
deriving Inhabited
partial def typeToArgsType (e : Expr) : (Array (Name × Expr × BinderInfo) × Expr) :=
match e.consumeMData with
| Expr.lam name type body data =>
let name := name.eraseMacroScopes
let arg := (name, type, data.binderInfo)
let (args, final) := typeToArgsType (Expr.instantiate1 body (mkFVar ⟨name⟩))
(#[arg] ++ args, final)
| Expr.forallE name type body data =>
let name := name.eraseMacroScopes
let arg := (name, type, data.binderInfo)
let (args, final) := typeToArgsType (Expr.instantiate1 body (mkFVar ⟨name⟩))
(#[arg] ++ args, final)
| _ => (#[], e)
def prettyPrintTerm (expr : Expr) : MetaM CodeWithInfos := do
let (fmt, infos) ← formatInfos expr
let tt := TaggedText.prettyTagged fmt
let ctx := {
env := ← getEnv
mctx := ← getMCtx
options := ← getOptions
currNamespace := ← getCurrNamespace
openDecls := ← getOpenDecls
fileMap := arbitrary
}
tagExprInfos ctx infos tt
def Info.ofConstantVal (v : ConstantVal) : MetaM Info := do
let env ← getEnv
let (args, type) := typeToArgsType v.type
let type ← prettyPrintTerm type
let args ← args.mapM (λ (n, e, b) => do Arg.mk n (←prettyPrintTerm e) b)
let doc ← findDocString? env v.name
match ←findDeclarationRanges? v.name with
-- TODO: Maybe selection range is more relevant? Figure this out in the future
| some range => return Info.mk ⟨v.name, type⟩ args doc range.range
| none => panic! s!"{v.name} is a declaration without position"
def AxiomInfo.ofAxiomVal (v : AxiomVal) : MetaM AxiomInfo := do
let info ← Info.ofConstantVal v.toConstantVal
return AxiomInfo.mk info v.isUnsafe
def TheoremInfo.ofTheoremVal (v : TheoremVal) : MetaM TheoremInfo := do
let info ← Info.ofConstantVal v.toConstantVal
return TheoremInfo.mk info
def OpaqueInfo.ofOpaqueVal (v : OpaqueVal) : MetaM OpaqueInfo := do
let info ← Info.ofConstantVal v.toConstantVal
let t ← prettyPrintTerm v.value
return OpaqueInfo.mk info t v.isUnsafe
def isInstance (declName : Name) : MetaM Bool := do
(instanceExtension.getState (←getEnv)).instanceNames.contains declName
def DefinitionInfo.ofDefinitionVal (v : DefinitionVal) : MetaM DefinitionInfo := do
let info ← Info.ofConstantVal v.toConstantVal
-- Elaborating the value yields weird exceptions
--let value ← prettyPrintTerm v.value
return DefinitionInfo.mk info v.safety v.hints
def getConstructorType (ctor : Name) : MetaM CodeWithInfos := do
let env ← getEnv
match env.find? ctor with
| some (ConstantInfo.ctorInfo i) => ←prettyPrintTerm i.type
| _ => panic! s!"Constructor {ctor} was requested but does not exist"
-- TODO: Obtain parameters that come after the inductive Name
def InductiveInfo.ofInductiveVal (v : InductiveVal) : MetaM InductiveInfo := do
let info ← Info.ofConstantVal v.toConstantVal
let env ← getEnv
let ctors ← v.ctors.mapM (λ name => do NameInfo.mk name (←getConstructorType name))
return InductiveInfo.mk info v.numParams v.numIndices v.all ctors v.isRec v.isUnsafe v.isReflexive v.isNested
def getFieldTypeAux (type : Expr) (vars : List Name) : (Expr × List Name) :=
match type with
| Expr.forallE `self _ b .. => (b, (`self :: vars))
| Expr.forallE n _ b .. => getFieldTypeAux b (n :: vars)
| _ => (type, vars)
def getFieldType (projFn : Name) : MetaM Expr := do
let fn ← mkConstWithFreshMVarLevels projFn
let type ← inferType fn
let (type, vars) := getFieldTypeAux type []
type.instantiate $ vars.toArray.map mkConst
def FieldInfo.ofStructureFieldInfo (i : StructureFieldInfo) : MetaM FieldInfo := do
let type ← getFieldType i.projFn
let ni := NameInfo.mk i.fieldName (←prettyPrintTerm type)
FieldInfo.mk ni i.projFn i.subobject?
def StructureInfo.ofInductiveVal (v : InductiveVal) : MetaM StructureInfo := do
let info ← Info.ofConstantVal v.toConstantVal
let env ← getEnv
let parents := getParentStructures env v.name
let ctor := getStructureCtor env v.name |>.name
let ctorType ← getConstructorType ctor
match getStructureInfo? env v.name with
| some i =>
let fieldInfos ← i.fieldInfo.mapM FieldInfo.ofStructureFieldInfo
return StructureInfo.mk info fieldInfos parents ⟨ctor, ctorType⟩
| none => panic! s!"{v.name} is not a structure"
def ClassInfo.ofInductiveVal (v : InductiveVal) : MetaM ClassInfo := do
let sinfo ← StructureInfo.ofInductiveVal v
let fn ← mkConstWithFreshMVarLevels v.name
let (xs, _, _) ← forallMetaTelescopeReducing (← inferType fn)
let insts ← SynthInstance.getInstances (mkAppN fn xs)
let insts_stx ← insts.mapM prettyPrintTerm
return ClassInfo.mk sinfo (hasOutParams (←getEnv) v.name) insts_stx
namespace DocInfo
def isBlackListed (declName : Name) : MetaM Bool := do
match ←findDeclarationRanges? declName with
| some _ =>
let env ← getEnv
declName.isInternal
<||> isAuxRecursor env declName
<||> isNoConfusion env declName
<||> isRec declName
<||> isMatcher declName
-- TODO: Evaluate whether filtering out declarations without range is sensible
| none => true
-- TODO: Is this actually the best way?
def isProjFn (declName : Name) : MetaM Bool := do
let env ← getEnv
match declName with
| Name.str parent name _ =>
if isStructure env parent then
match getStructureInfo? env parent with
| some i =>
match i.fieldNames.find? (· == name) with
| some _ => true
| none => false
| none => panic! s!"{parent} is not a structure"
else
false
| _ => false
def ofConstant : (Name × ConstantInfo) → MetaM (Option DocInfo) := λ (name, info) => do
if (←isBlackListed name) then
return none
match info with
| ConstantInfo.axiomInfo i => some $ axiomInfo (←AxiomInfo.ofAxiomVal i)
| ConstantInfo.thmInfo i => some $ theoremInfo (←TheoremInfo.ofTheoremVal i)
| ConstantInfo.opaqueInfo i => some $ opaqueInfo (←OpaqueInfo.ofOpaqueVal i)
-- TODO: Find a way to extract equations nicely
| ConstantInfo.defnInfo i =>
if ← (isProjFn i.name) then
none
else
let info ← DefinitionInfo.ofDefinitionVal i
if (←isInstance i.name) then
some $ instanceInfo info
else
some $ definitionInfo info
| ConstantInfo.inductInfo i =>
let env ← getEnv
if isStructure env i.name then
if isClass env i.name then
some $ classInfo (←ClassInfo.ofInductiveVal i)
else
some $ structureInfo (←StructureInfo.ofInductiveVal i)
else
some $ inductiveInfo (←InductiveInfo.ofInductiveVal i)
-- we ignore these for now
| ConstantInfo.ctorInfo i => none
| ConstantInfo.recInfo i => none
| ConstantInfo.quotInfo i => none
def getName : DocInfo → Name
| axiomInfo i => i.name
| theoremInfo i => i.name
| opaqueInfo i => i.name
| definitionInfo i => i.name
| instanceInfo i => i.name
| inductiveInfo i => i.name
| structureInfo i => i.name
| classInfo i => i.name
def getKind : DocInfo → String
| axiomInfo _ => "axiom"
| theoremInfo _ => "theorem"
| opaqueInfo _ => "constant"
| definitionInfo _ => "def"
| instanceInfo _ => "instance" -- TODO: This doesnt exist in CSS yet
| inductiveInfo _ => "inductive"
| structureInfo _ => "structure"
| classInfo _ => "class" -- TODO: This is handled as structure right now
def getType : DocInfo → CodeWithInfos
| axiomInfo i => i.type
| theoremInfo i => i.type
| opaqueInfo i => i.type
| definitionInfo i => i.type
| instanceInfo i => i.type
| inductiveInfo i => i.type
| structureInfo i => i.type
| classInfo i => i.type
def getArgs : DocInfo → Array Arg
| axiomInfo i => i.args
| theoremInfo i => i.args
| opaqueInfo i => i.args
| definitionInfo i => i.args
| instanceInfo i => i.args
| inductiveInfo i => i.args
| structureInfo i => i.args
| classInfo i => i.args
end DocInfo
structure AnalyzerResult where
name2ModIdx : HashMap Name ModuleIdx
moduleNames : Array Name
moduleInfo : HashMap Name Module
hierarchy : Hierarchy
deriving Inhabited
def process : MetaM AnalyzerResult := do
let env ← getEnv
let mut res := mkHashMap env.header.moduleNames.size
for module in env.header.moduleNames do
-- TODO: Check why modules can have multiple doc strings and add that later on
let moduleDoc := match getModuleDoc? env module with
| none => none
| some #[] => none
| some doc => doc.get! 0
res := res.insert module (Module.mk module moduleDoc #[])
for cinfo in env.constants.toList do
let d := ←DocInfo.ofConstant cinfo
match d with
| some dinfo =>
match (env.getModuleIdxFor? cinfo.fst) with
| some modidx =>
-- TODO: Check whether this is still efficient
let moduleName := env.allImportedModuleNames.get! modidx
let module := res.find! moduleName
res := res.insert moduleName {module with members := module.members.push dinfo}
| none => panic! "impossible"
| none => ()
return {
name2ModIdx := env.const2ModIdx,
moduleNames := env.header.moduleNames,
moduleInfo := res,
hierarchy := Hierarchy.fromArray env.header.moduleNames
}
end DocGen4

113
DocGen4/ToHtmlFormat.lean Normal file
View File

@ -0,0 +1,113 @@
/-
Copyright (c) 2021 Wojciech Nawrocki. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Wojciech Nawrocki, Sebastian Ullrich, Henrik Böving
-/
import Lean.Data.Json
import Lean.Parser
/-! This module defines:
- a representation of HTML trees
- together with a JSX-like DSL for writing them
- and widget support for visualizing any type as HTML. -/
namespace DocGen4
open Lean
inductive Html where
-- TODO(WN): it's nameless for shorter JSON; re-add names when we have deriving strategies for From/ToJson
-- element (tag : String) (flatten : Bool) (attrs : Array HtmlAttribute) (children : Array Html)
| element : String → Bool → Array (String × String) → Array Html → Html
| text : String → Html
deriving Repr, BEq, Inhabited, FromJson, ToJson
instance : Coe String Html :=
⟨Html.text⟩
namespace Html
def attributesToString (attrs : Array (String × String)) :String :=
attrs.foldl (λ acc (k, v) => acc ++ " " ++ k ++ "=\"" ++ v ++ "\"") ""
-- TODO: Termination proof
partial def toStringAux : Html → String
| element tag false attrs #[text s] => s!"<{tag}{attributesToString attrs}>{s}</{tag}>\n"
| element tag false attrs #[child] => s!"<{tag}{attributesToString attrs}>\n{child.toStringAux}</{tag}>\n"
| element tag false attrs children => s!"<{tag}{attributesToString attrs}>\n{children.foldl (· ++ toStringAux ·) ""}</{tag}>\n"
| element tag true attrs children => s!"<{tag}{attributesToString attrs}>{children.foldl (· ++ toStringAux ·) ""}</{tag}>"
| text s => s
def toString (html : Html) : String :=
html.toStringAux.trimRight
instance : ToString Html :=
⟨toString⟩
end Html
namespace Jsx
open Parser PrettyPrinter
declare_syntax_cat jsxElement
declare_syntax_cat jsxChild
def jsxAttrVal : Parser := strLit <|> group ("{" >> termParser >> "}")
def jsxAttr : Parser := ident >> "=" >> jsxAttrVal
-- JSXTextCharacter : SourceCharacter but not one of {, <, > or }
def jsxText : Parser :=
withAntiquot (mkAntiquot "jsxText" `jsxText) {
fn := fun c s =>
let startPos := s.pos
let s := takeWhile1Fn (not ∘ "[{<>}]$".contains) "expected JSX text" c s
mkNodeToken `jsxText startPos c s }
@[combinatorFormatter DocGen4.Jsx.jsxText] def jsxText.formatter : Formatter := pure ()
@[combinatorParenthesizer DocGen4.Jsx.jsxText] def jsxText.parenthesizer : Parenthesizer := pure ()
scoped syntax "<" ident jsxAttr* "/>" : jsxElement
scoped syntax "<" ident jsxAttr* ">" jsxChild* "</" ident ">" : jsxElement
scoped syntax jsxText : jsxChild
scoped syntax "{" term "}" : jsxChild
scoped syntax "[" term "]" : jsxChild
scoped syntax jsxElement : jsxChild
scoped syntax:max jsxElement : term
macro_rules
| `(<$n $[$ns = $vs]* />) =>
let ns := ns.map (quote <| toString ·.getId)
let vs := vs.map fun
| `(jsxAttrVal| $s:strLit) => s
| `(jsxAttrVal| { $t:term }) => t
| _ => unreachable!
`(Html.element $(quote <| toString n.getId) false #[ $[($ns, $vs)],* ] #[])
| `(<$n $[$ns = $vs]* >$cs*</$m>) =>
if n.getId == m.getId then do
let ns := ns.map (quote <| toString ·.getId)
let vs := vs.map fun
| `(jsxAttrVal| $s:strLit) => s
| `(jsxAttrVal| { $t:term }) => t
| _ => unreachable!
let cs ← cs.mapM fun
| `(jsxChild|$t:jsxText) => `(#[Html.text $(quote t[0].getAtomVal!)])
-- TODO(WN): elab as list of children if type is `t Html` where `Foldable t`
| `(jsxChild|{$t}) => `(#[$t])
| `(jsxChild|[$t]) => `($t)
| `(jsxChild|$e:jsxElement) => `(#[$e:jsxElement])
| _ => unreachable!
let tag := toString n.getId
`(Html.element $(quote tag) false #[ $[($ns, $vs)],* ] (Array.foldl Array.append #[] #[ $[$cs],* ]))
else Macro.throwError ("expected </" ++ toString n.getId ++ ">")
end Jsx
/-- A type which implements `ToHtmlFormat` will be visualized
as the resulting HTML in editors which support it. -/
class ToHtmlFormat (α : Type u) where
formatHtml : α → Html
end DocGen4

12
Main.lean Normal file
View File

@ -0,0 +1,12 @@
import DocGen4
import Lean
open DocGen4 Lean IO
def main (args : List String) : IO Unit := do
let modules := args
let path ← lakeSetupSearchPath (←getLakePath) modules.toArray
IO.println s!"Loading modules from: {path}"
let doc ← load $ modules.map Name.mkSimple
IO.println "Outputting HTML"
htmlOutput doc

View File

@ -1,2 +1,15 @@
# doc-gen4
Document Generator for Lean 4
## Usage
You can call `doc-gen4` from the top of a Lake project like this:
```sh
$ /path/to/doc-gen4 Module
```
Where `Module` is one or more of the top level modules you want to document.
The tool will then proceed to compile the project using lake (if that hasn't happened yet),
analyze it and put the result in `./build/doc`.
You could e.g. host the files locally with the built-in Python webserver:
```sh
$ cd build/doc && python -m http.server
```

7
lakefile.lean Normal file
View File

@ -0,0 +1,7 @@
import Lake
open Lake DSL
package «doc-gen4» {
-- add configuration options here
supportInterpreter := true
}

1
lean-toolchain Normal file
View File

@ -0,0 +1 @@
lean4

239
static/nav.js Normal file
View File

@ -0,0 +1,239 @@
// Persistent expansion cookie for the file tree
// ---------------------------------------------
let expanded = {};
for (const e of (sessionStorage.getItem('expanded') || '').split(',')) {
if (e !== '') {
expanded[e] = true;
}
}
function saveExpanded() {
sessionStorage.setItem("expanded",
Object.getOwnPropertyNames(expanded).filter((e) => expanded[e]).join(","));
}
for (const elem of document.getElementsByClassName('nav_sect')) {
const id = elem.getAttribute('data-path');
if (!id) continue;
if (expanded[id]) {
elem.open = true;
}
elem.addEventListener('toggle', () => {
expanded[id] = elem.open;
saveExpanded();
});
}
for (const currentFileLink of document.getElementsByClassName('visible')) {
currentFileLink.scrollIntoView({block: 'center'});
}
// Tactic list tag filter
// ----------------------
function filterSelectionClass(tagNames, classname) {
if (tagNames.length == 0) {
for (const elem of document.getElementsByClassName(classname)) {
elem.classList.remove("hide");
}
} else {
// Add the "show" class (display:block) to the filtered elements, and remove the "show" class from the elements that are not selected
for (const elem of document.getElementsByClassName(classname)) {
elem.classList.add("hide");
for (const tagName of tagNames) {
if (elem.classList.contains(tagName)) {
elem.classList.remove("hide");
}
}
}
}
}
function filterSelection(c) {
filterSelectionClass(c, "tactic");
filterSelectionClass(c, "taclink");
}
var filterBoxes = document.getElementsByClassName("tagfilter");
function updateDisplay() {
filterSelection(getSelectValues());
}
function getSelectValues() {
var result = [];
for (const opt of filterBoxes) {
if (opt.checked) {
result.push(opt.value);
}
}
return result;
}
function setSelectVal(val) {
for (const opt of filterBoxes) {
opt.checked = val;
}
}
updateDisplay();
for (const opt of filterBoxes) {
opt.addEventListener('change', updateDisplay);
}
const tse = document.getElementById("tagfilter-selectall")
if (tse != null) {
tse.addEventListener('change', function() {
setSelectVal(this.checked);
updateDisplay();
});
}
// Simple declaration search
// -------------------------
const searchWorkerURL = new URL(`${siteRoot}searchWorker.js`, window.location);
const declSearch = (q) => new Promise((resolve, reject) => {
const worker = new SharedWorker(searchWorkerURL);
worker.port.start();
worker.port.onmessage = ({data}) => resolve(data);
worker.port.onmessageerror = (e) => reject(e);
worker.port.postMessage({q});
});
const srId = 'search_results';
document.getElementById('search_form')
.appendChild(document.createElement('div'))
.id = srId;
function handleSearchCursorUpDown(down) {
const sel = document.querySelector(`#${srId} .selected`);
const sr = document.getElementById(srId);
if (sel) {
sel.classList.remove('selected');
const toSelect = down ?
sel.nextSibling || sr.firstChild:
sel.previousSibling || sr.lastChild;
toSelect && toSelect.classList.add('selected');
} else {
const toSelect = down ? sr.firstChild : sr.lastChild;
toSelect && toSelect.classList.add('selected');
}
}
function handleSearchEnter() {
const sel = document.querySelector(`#${srId} .selected`)
|| document.getElementById(srId).firstChild;
sel.click();
}
const searchInput = document.querySelector('#search_form input[name=q]');
searchInput.addEventListener('keydown', (ev) => {
switch (ev.key) {
case 'Down':
case 'ArrowDown':
ev.preventDefault();
handleSearchCursorUpDown(true);
break;
case 'Up':
case 'ArrowUp':
ev.preventDefault();
handleSearchCursorUpDown(false);
break;
case 'Enter':
ev.preventDefault();
handleSearchEnter();
break;
}
});
searchInput.addEventListener('input', async (ev) => {
const text = ev.target.value;
if (!text) {
const sr = document.getElementById(srId);
sr.removeAttribute('state');
sr.replaceWith(sr.cloneNode(false));
return;
}
document.getElementById(srId).setAttribute('state', 'loading');
const result = await declSearch(text);
if (ev.target.value != text) return;
const oldSR = document.getElementById('search_results');
const sr = oldSR.cloneNode(false);
for (const {decl} of result) {
const d = sr.appendChild(document.createElement('a'));
d.innerText = decl;
d.title = decl;
d.href = `${siteRoot}find/${decl}`;
}
sr.setAttribute('state', 'done');
oldSR.replaceWith(sr);
});
// 404 page goodies
// ----------------
const howabout = document.getElementById('howabout');
if (howabout) {
howabout.innerText = "Please wait a second. I'll try to help you.";
howabout.parentNode
.insertBefore(document.createElement('pre'), howabout)
.appendChild(document.createElement('code'))
.innerText = window.location.href.replace(/[/]/g, '/\u200b');
const query = window.location.href.match(/[/]([^/]+)(?:\.html|[/])?$/)[1];
declSearch(query).then((results) => {
howabout.innerText = 'How about one of these instead:';
const ul = howabout.appendChild(document.createElement('ul'));
for (const {decl} of results) {
const li = ul.appendChild(document.createElement('li'));
const a = li.appendChild(document.createElement('a'));
a.href = `${siteRoot}find/${decl}`;
a.appendChild(document.createElement('code')).innerText = decl;
}
});
}
// Rewrite GitHub links
// --------------------
for (const elem of document.getElementsByClassName('gh_link')) {
const a = elem.firstElementChild;
// commit is set in add_commit.js
for (const [prefix, replacement] of commit) {
if (a.href.startsWith(prefix)) {
a.href = a.href.replace(prefix, replacement);
break;
}
}
}

562
static/style.css Normal file
View File

@ -0,0 +1,562 @@
@import url('https://fonts.googleapis.com/css2?family=Merriweather&family=Open+Sans&family=Source+Code+Pro&family=Source+Code+Pro:wght@600&display=swap');
* {
box-sizing: border-box;
}
body {
font-family: 'Open Sans', sans-serif;
}
h1, h2, h3, h4, h5, h6 {
font-family: 'Merriweather', serif;
}
body { line-height: 1.5; }
nav { line-height: normal; }
:root {
--header-height: 3em;
--header-bg: #f8f8f8;
--fragment-offset: calc(var(--header-height) + 1em);
--content-width: 55vw;
}
@supports (width: min(10px, 5vw)) {
:root {
--content-width: clamp(20em, 55vw, 60em);
}
}
#nav_toggle {
display: none;
}
label[for="nav_toggle"] {
display: none;
}
header {
height: var(--header-height);
float: left;
position: fixed;
width: 100vw;
max-width: 100%;
left: 0;
right: 0;
top: 0;
--header-side-padding: 2em;
padding: 0 var(--header-side-padding);
background: var(--header-bg);
z-index: 1;
display: flex;
align-items: center;
justify-content: space-between;
}
@supports (width: min(10px, 5vw)) {
header {
--header-side-padding: calc(max(2em, (100vw - var(--content-width) - 30em) / 2));
}
}
@media screen and (max-width: 1000px) {
:root {
--content-width: 100vw;
}
.internal_nav {
display: none;
}
body .nav {
width: 100vw;
max-width: 100vw;
margin-left: 1em;
z-index: 1;
}
body main {
width: unset;
max-width: unset;
margin-left: unset;
margin-right: unset;
}
body .decl > div {
overflow-x: unset;
}
#nav_toggle:not(:checked) ~ .nav {
display: none;
}
#nav_toggle:checked ~ main {
visibility: hidden;
}
label[for="nav_toggle"]::before {
content: '≡';
}
label[for="nav_toggle"] {
display: inline-block;
margin-right: 1em;
border: 1px solid #ccc;
padding: 0.5ex 1ex;
cursor: pointer;
background: #eee;
}
#nav_toggle:checked ~ * label[for="nav_toggle"] {
background: white;
}
body header h1 {
font-size: 100%;
}
header {
--header-side-padding: 1ex;
}
}
@media screen and (max-width: 700px) {
header h1 span { display: none; }
:root { --header-side-padding: 1ex; }
#search_form button { display: none; }
#search_form input { width: 100%; }
header #search_results {
left: 1ex;
right: 1ex;
width: inherit;
}
body header > * { margin: 0; }
}
header > * {
display: inline-block;
padding: 0;
margin: 0 1em;
vertical-align: middle;
}
header h1 {
font-weight: normal;
font-size: 160%;
}
header header_filename {
font-size: 150%;
}
@media (max-width: 80em) {
.header .header_filename {
display: none;
}
}
/* inserted by nav.js */
#search_results {
position: absolute;
top: var(--header-height);
right: calc(var(--header-side-padding));
width: calc(20em + 4em);
z-index: 1;
background: var(--header-bg);
border: 1px solid #aaa;
border-top: none;
overflow-x: hidden;
overflow-y: auto;
max-height: calc(100vh - var(--header-height));
}
#search_results:empty {
display: none;
}
#search_results[state="loading"]:empty {
display: block;
cursor: progress;
}
#search_results[state="loading"]:empty::before {
display: block;
content: ' 🐙 🐙 🐙 🐙 🐙 🐙 🐙 🐙 🐙 🐙 🐙 🐙 🐙 🐙 🐙 ';
padding: 1ex;
animation: marquee 10s linear infinite;
}
@keyframes marquee {
0% { transform: translate(100%, 0); }
100% { transform: translate(-100%, 0); }
}
#search_results[state="done"]:empty {
display: block;
text-align: center;
padding: 1ex;
}
#search_results[state="done"]:empty::before {
content: '(no results)';
font-style: italic;
}
#search_results a {
display: block;
color: inherit;
padding: 1ex;
border-left: 0.5ex solid transparent;
padding-left: 0.5ex;
cursor: pointer;
}
#search_results .selected {
background: white;
border-color: #f0a202;
}
main, nav {
margin-top: calc(var(--header-height) + 1em);
}
/* extra space for scrolling things to the top */
main {
margin-bottom: 90vh;
}
main {
max-width: var(--content-width);
/* center it: */
margin-left: auto;
margin-right: auto;
}
nav {
float: left;
height: calc(100vh - var(--header-height) - 1em);
position: fixed;
top: 0;
overflow: auto;
scrollbar-width: thin;
scrollbar-color: transparent transparent;
}
nav:hover {
scrollbar-color: gray transparent;
}
nav {
--column-available-space: calc((100vw - var(--content-width) - 5em)/2);
--column-width: calc(var(--column-available-space) - 1ex);
--dist-to-edge: 1ex;
width: var(--content-width);
max-width: var(--column-width);
}
@supports (width: min(10px, 5vw)) {
.nav { --desired-column-width: 20em; }
.internal_nav { --desired-column-width: 30em; }
nav {
--column-available-space: calc(max(0px, (100vw - var(--content-width) - 5em)/2));
--column-width: calc(clamp(0px, var(--column-available-space) - 1ex, var(--desired-column-width)));
--dist-to-edge: calc(max(1ex, var(--column-available-space) - var(--column-width)));
}
}
.nav { left: var(--dist-to-edge); }
.internal_nav { right: var(--dist-to-edge); }
.internal_nav .nav_link, .taclink {
/* indent everything but first line by 2ex */
text-indent: -2ex; padding-left: 2ex;
}
.internal_nav .imports {
margin-bottom: 1rem;
}
.tagfilter-div {
margin-bottom: 1em;
}
.tagfilter-div > summary {
margin-bottom: 1ex;
}
.nav details > * {
padding-left: 2ex;
}
.nav summary {
cursor: pointer;
padding-left: 0;
}
.nav summary::marker {
font-size: 85%;
}
.nav .nav_file {
display: inline-block;
}
.nav h3 {
margin-block-end: 4px;
}
.nav h4 {
margin-bottom: 1ex;
}
/* People use way too long declaration names. */
.internal_nav, .decl_name {
overflow-wrap: break-word;
}
.decl > div, .mod_doc {
padding-left: 8px;
padding-right: 8px;
}
.decl {
margin-top: 20px;
margin-bottom: 20px;
}
.decl > div {
/* sometimes declarations arguments are way too long
and would continue into the right column,
so put a scroll bar there: */
overflow-x: auto;
}
/* Make `#id` links appear below header. */
.decl, h1[id], h2[id], h3[id], h4[id], h5[id], h6[id] {
scroll-margin-top: var(--fragment-offset);
}
/* don't need as much vertical space for these
inline elements */
a[id], li[id] {
scroll-margin-top: var(--header-height);
}
/* HACK: Safari doesn't support scroll-margin-top for
fragment links (yet?)
https://caniuse.com/mdn-css_properties_scroll-margin-top
https://bugs.webkit.org/show_bug.cgi?id=189265
*/
@supports not (scroll-margin-top: var(--fragment-offset)) {
.decl::before, h1[id]::before, h2[id]::before, h3[id]::before,
h4[id]::before, h5[id]::before, h6[id]::before,
a[id]::before, li[id]::before {
content: "";
display: block;
height: var(--fragment-offset);
margin-top: calc(-1 * var(--fragment-offset));
box-sizing: inherit;
visibility: hidden;
width: 1px;
}
}
/* hide # after markdown headings except on hover */
.markdown-heading:not(:hover) > .hover-link {
visibility: hidden;
}
main h2, main h3, main h4, main h5, main h6 {
margin-top: 2rem;
}
.decl + .mod_doc > h2,
.decl + .mod_doc > h3,
.decl + .mod_doc > h4,
.decl + .mod_doc > h5,
.decl + .mod_doc > h6 {
margin-top: 4rem;
}
.def {
border-left: 10px solid #92dce5;
border-top: 2px solid #92dce5;
}
.theorem {
border-left: 10px solid #8fe388;
border-top: 2px solid #8fe388;
}
.axiom, .constant {
border-left: 10px solid #f44708;
border-top: 2px solid #f44708;
}
.structure, .inductive {
border-left: 10px solid #f0a202;
border-top: 2px solid #f0a202;
}
.fn {
display: inline-block;
/* border: 1px dashed red; */
text-indent: -1ex;
padding-left: 1ex;
white-space: pre-wrap;
vertical-align: top;
}
.fn { --fn: 1; }
.fn .fn { --fn: 2; }
.fn .fn .fn { --fn: 3; }
.fn .fn .fn .fn { --fn: 4; }
.fn .fn .fn .fn .fn { --fn: 5; }
.fn .fn .fn .fn .fn .fn { --fn: 6; }
.fn .fn .fn .fn .fn .fn .fn { --fn: 7; }
.fn .fn .fn .fn .fn .fn .fn .fn { --fn: 8; }
.fn {
transition: background-color 100ms ease-in-out;
}
.def .fn:hover {
background-color: hsla(187, 61%, calc(100% - 5%*var(--fn)));
box-shadow: 0 0 0 1px hsla(187, 61%, calc(100% - 5%*(var(--fn) + 1)));
border-radius: 5px;
}
.theorem .fn:hover {
background-color: hsla(115, 62%, calc(100% - 5%*var(--fn)));
box-shadow: 0 0 0 1px hsla(115, 62%, calc(100% - 5%*(var(--fn) + 1)));
border-radius: 5px;
}
.axiom .fn:hover, .constant .fn:hover {
background-color: hsla(16, 94%, calc(100% - 5%*var(--fn)));
box-shadow: 0 0 0 1px hsla(16, 94%, calc(100% - 5%*(var(--fn) + 1)));
border-radius: 5px;
}
.structure .fn:hover, .inductive .fn:hover {
background-color: hsla(40, 98%, calc(100% - 5%*var(--fn)));
box-shadow: 0 0 0 1px hsla(40, 98%, calc(100% - 5%*(var(--fn) + 1)));
border-radius: 5px;
}
.decl_args, .decl_type {
font-weight: normal;
}
.implicits, .impl_arg {
color: black;
white-space: nowrap;
}
.decl_kind, .decl_name {
font-weight: bold;
}
/* break long declaration names at periods where possible */
.break_within {
word-break: break-all;
}
.break_within .name {
word-break: normal;
}
.decl_header {
/* indent everything but first line twice as much as decl_type */
text-indent: -8ex; padding-left: 8ex;
}
.decl_type {
margin-top: 2px;
margin-left: 4ex; /* extra indentation */
}
.imports li, code, .decl_header, .attributes, .structure_field,
.constructor, .instances li, .equation, #search_results div {
font-family: 'Source Code Pro', monospace;
}
pre {
white-space: break-spaces;
}
code, pre { background: #f3f3f3; }
code, pre { border-radius: 5px; }
code { padding: 1px 3px; }
pre { padding: 1ex; }
pre code { padding: 0 0; }
#howabout code { background: inherit; }
#howabout li { margin-bottom: 0.5ex; }
.structure_fields, .constructors {
display: block;
padding-inline-start: 0;
margin-top: 1ex;
text-indent: -2ex; padding-left: 2ex;
}
.structure_field {
display: block;
}
.structure_field:before {
content: '(';
color: gray;
}
.structure_field:after {
content: ')';
color: gray;
}
.constructor {
display: block;
}
.constructor:before {
content: '| ';
color: gray;
}
/** Don't show underline on types, to prevent the ≤ vs < confusion. **/
a:link, a:visited, a:active {
color:hsl(210, 100%, 30%);
text-decoration: none;
}
/** Show it on hover though. **/
a:hover {
text-decoration: underline;
}
.impl_arg {
font-style: italic;
transition: opacity 300ms ease-in;
}
.decl_header:not(:hover) .impl_arg {
opacity: 30%;
transition: opacity 1000ms ease-out;
}
.gh_link {
float: right;
margin-left: 20px;
}
.docfile h2, .note h2 {
margin-block-start: 3px;
margin-block-end: 0px;
}
.docfile h2 a {
color: black;
}
.tags {
margin-bottom: 1ex;
}
.tags ul {
display: inline;
padding: 0;
}
.tags li {
border: 1px solid #555;
border-radius: 4px;
list-style-type: none;
padding: 1px 3px;
margin-left: 1ex;
display: inline-block;
}
/* used by nav.js */
.hide { display: none; }
.tactic, .note {
border-top: 3px solid #0479c7;
padding-top: 2em;
margin-top: 2em;
margin-bottom: 2em;
}