diff --git a/DocGen4/Output/References.lean b/DocGen4/Output/References.lean
new file mode 100644
index 00000000..f18e58f3
--- /dev/null
+++ b/DocGen4/Output/References.lean
@@ -0,0 +1,97 @@
+import DocGen4.Output.Template
+
+/-!
+
+# Generic functions for references support
+
+This file contains functions for references support,
+independent of actual implementation.
+
+The main function is `DocGen4.preprocessBibFile` which preprocess
+the contents of bib file using user provided `process` function,
+and save the bib file and processed json file to output directory,
+as "references.bib" and "references.json", respectively.
+
+Note that "references.bib" is only for user to download it, the actual file
+used later is "references.json". It contains an array of objects with 4 fields:
+'citekey', 'tag', 'html' and 'plaintext'. For the meaning of these fields,
+see `DocGen4.Output.BibItem`.
+
+-/
+
+open Lean DocGen4 Output
+
+namespace DocGen4
+
+/-- Preprocess (using the user provided `process` function)
+and save the bib file to the output path. -/
+def preprocessBibFile (contents : String) (process : String → IO (Array BibItem)) : IO Unit := do
+ -- create directories
+ IO.FS.createDirAll basePath
+ IO.FS.createDirAll declarationsBasePath
+ -- save the contents to "references.bib" and erase "references.json"
+ IO.FS.writeFile (basePath / "references.bib") contents
+ IO.FS.writeFile (declarationsBasePath / "references.json") "[]"
+ -- if contents is empty, just do nothing
+ if contents.trim.isEmpty then
+ return
+ -- run the user provided process function
+ let items ← process contents
+ -- save the result to "references.json"
+ IO.FS.writeFile (declarationsBasePath / "references.json") (toString (toJson items))
+
+/-- Save the bib json to the output path. -/
+def preprocessBibJson (contents : String) : IO Unit := do
+ -- create directories
+ IO.FS.createDirAll basePath
+ IO.FS.createDirAll declarationsBasePath
+ -- erase "references.bib" (since we can't recover it from json)
+ -- and save the contents to "references.json"
+ IO.FS.writeFile (basePath / "references.bib") ""
+ IO.FS.writeFile (declarationsBasePath / "references.json") contents
+
+/-- Erase the contents of bib file in the output path. -/
+def disableBibFile : IO Unit := do
+ -- create directories
+ IO.FS.createDirAll basePath
+ IO.FS.createDirAll declarationsBasePath
+ -- erase files
+ IO.FS.writeFile (basePath / "references.bib") ""
+ IO.FS.writeFile (declarationsBasePath / "references.json") "[]"
+
+namespace Output
+
+open scoped DocGen4.Jsx
+
+def refItem (ref : BibItem) (backrefs : Array BackrefItem) : BaseHtmlM Html := do
+ let backrefs := backrefs.filter (fun x => x.citekey == ref.citekey)
+ let toHtml (i : Fin backrefs.size) (backref : BackrefItem) : BaseHtmlM (Array Html) := do
+ let href := s!"{moduleNameToFile "" backref.modName}#_backref_{backref.index}"
+ let title := s!"File: {backref.modName}" ++
+ if backref.funName.isEmpty then "" else s!"\nLocation: {backref.funName}"
+ pure #[.raw " ",
{.text s!"[{i.1 + 1}]"}]
+ let backrefHtml : Html ← (do
+ if backrefs.isEmpty then
+ pure (.raw "")
+ else
+ pure
[(← backrefs.mapIdxM toHtml).foldl (· ++ ·) #[]])
+ pure <|
+
+ {.text ref.tag}
+ {.raw " "}{.raw ref.html}{backrefHtml}
+
+
+def references (backrefs : Array BackrefItem) :
+ BaseHtmlM Html := templateLiftExtends (baseHtml "References") do
+ pure <|
+
+
+ References
+
+ [← (← read).refs.mapM (refItem · backrefs)]
+
+
+
+end Output
+
+end DocGen4
diff --git a/DocGen4/Output/Structure.lean b/DocGen4/Output/Structure.lean
index ad019949..b8f0aa45 100644
--- a/DocGen4/Output/Structure.lean
+++ b/DocGen4/Output/Structure.lean
@@ -15,7 +15,7 @@ def fieldToHtml (f : Process.NameInfo) : HtmlM Html := do
let shortName := f.name.componentsRev.head!.toString
let name := f.name.toString
if let some doc := f.doc then
- let renderedDoc ← docStringToHtml doc
+ let renderedDoc ← docStringToHtml doc name
pure
{s!"{shortName} "} : [← infoFormatToHtml f.type]
diff --git a/Main.lean b/Main.lean
index 47017f15..904e4973 100644
--- a/Main.lean
+++ b/Main.lean
@@ -39,6 +39,22 @@ def runDocGenCmd (_p : Parsed) : IO UInt32 := do
IO.println "https://github.com/leanprover/doc-gen4"
return 0
+def runBibPrepassCmd (p : Parsed) : IO UInt32 := do
+ if p.hasFlag "none" then
+ IO.println "INFO: reference page disabled"
+ disableBibFile
+ else
+ match p.variableArgsAs! String with
+ | #[source] =>
+ let contents ← IO.FS.readFile source
+ if p.hasFlag "json" then
+ IO.println "INFO: 'references.json' will be copied to the output path; there will be no 'references.bib'"
+ preprocessBibJson contents
+ else
+ preprocessBibFile contents Bibtex.process
+ | _ => throw <| IO.userError "there should be exactly one source file"
+ return 0
+
def singleCmd := `[Cli|
single VIA runSingleCmd;
"Only generate the documentation for the module it was given, might contain broken links unless all documentation is generated."
@@ -58,6 +74,18 @@ def genCoreCmd := `[Cli|
"Generate documentation for the core Lean modules: Init and Lean since they are not Lake projects"
]
+def bibPrepassCmd := `[Cli|
+ bibPrepass VIA runBibPrepassCmd;
+ "Run the bibliography prepass: copy the bibliography file to output directory. By default it assumes the input is '.bib'."
+
+ FLAGS:
+ n, none; "Disable bibliography in this project."
+ j, json; "The input file is '.json' which contains an array of objects with 4 fields: 'citekey', 'tag', 'html' and 'plaintext'."
+
+ ARGS:
+ ...source : String; "The bibliography file. We only support one file for input. Should be '.bib' or '.json' according to flags."
+]
+
def headerDataCmd := `[Cli|
headerData VIA runHeaderDataCmd;
"Produce `header-data.bmp`, this allows embedding of doc-gen declarations into other pages and more."
@@ -71,6 +99,7 @@ def docGenCmd : Cmd := `[Cli|
singleCmd;
indexCmd;
genCoreCmd;
+ bibPrepassCmd;
headerDataCmd
]
diff --git a/lake-manifest.json b/lake-manifest.json
index d7f8ac3f..7c5f3f2c 100644
--- a/lake-manifest.json
+++ b/lake-manifest.json
@@ -11,11 +11,21 @@
"inputRev": "main",
"inherited": false,
"configFile": "lakefile.lean"},
+ {"url": "https://github.com/dupuisf/BibtexQuery",
+ "type": "git",
+ "subDir": null,
+ "scope": "",
+ "rev": "bd8747df9ee72fca365efa5bd3bd0d8dcd083b9f",
+ "name": "BibtexQuery",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "master",
+ "inherited": false,
+ "configFile": "lakefile.lean"},
{"url": "https://github.com/fgdorais/lean4-unicode-basic",
"type": "git",
"subDir": null,
"scope": "",
- "rev": "c74a052aebee847780e165611099854de050adf7",
+ "rev": "f93115d0209de6db335725dee900d379f40c0317",
"name": "UnicodeBasic",
"manifestFile": "lake-manifest.json",
"inputRev": "main",
diff --git a/lakefile.lean b/lakefile.lean
index 2dbdbffa..f550b52a 100644
--- a/lakefile.lean
+++ b/lakefile.lean
@@ -14,6 +14,9 @@ lean_exe «doc-gen4» {
require MD4Lean from git
"https://github.com/acmepjz/md4lean" @ "main"
+require BibtexQuery from git
+ "https://github.com/dupuisf/BibtexQuery" @ "master"
+
require «UnicodeBasic» from git
"https://github.com/fgdorais/lean4-unicode-basic" @ "main"
@@ -120,8 +123,34 @@ def getSrcUri (mod : Module) : IO String := do
| .some "file" => getFileUri mod
| .some _ => throw <| IO.userError "$DOCGEN_SRC should be github, file, or vscode."
+target bibPrepass : FilePath := do
+ let exeJob ← «doc-gen4».fetch
+ let dataPath := (← getWorkspace).root.buildDir / "doc-data"
+ let inputJsonFile := (← getWorkspace).root.srcDir / "docs" / "references.json"
+ let inputBibFile := (← getWorkspace).root.srcDir / "docs" / "references.bib"
+ let outputFile := dataPath / "references.json"
+ let tryJson : JobM (Array String × BuildTrace) := do
+ let inputTrace ← mixTrace (BuildTrace.ofHash (.ofString "json")) <$> computeTrace inputJsonFile
+ pure (#["--json", inputJsonFile.toString], inputTrace)
+ let tryBib : JobM (Array String × BuildTrace) := do
+ let inputTrace ← mixTrace (BuildTrace.ofHash (.ofString "bib")) <$> computeTrace inputBibFile
+ pure (#[inputBibFile.toString], inputTrace)
+ let tryBibFailed : JobM (Array String × BuildTrace) := do
+ pure (#["--none"], .nil)
+ exeJob.bindSync fun exeFile exeTrace => do
+ let (args, inputTrace) ← tryJson <|> tryBib <|> tryBibFailed
+ let depTrace := exeTrace.mix inputTrace
+ let trace ← buildFileUnlessUpToDate outputFile depTrace do
+ proc {
+ cmd := exeFile.toString
+ args := #["bibPrepass"] ++ args
+ env := ← getAugmentedEnv
+ }
+ return (outputFile, trace)
+
module_facet docs (mod) : FilePath := do
let exeJob ← «doc-gen4».fetch
+ let bibPrepassJob ← bibPrepass.fetch
let modJob ← mod.leanArts.fetch
-- Build all documentation imported modules
let imports ← mod.imports.fetch
@@ -130,30 +159,34 @@ module_facet docs (mod) : FilePath := do
let buildDir := (←getWorkspace).root.buildDir
let docFile := mod.filePath (buildDir / "doc") "html"
depDocJobs.bindAsync fun _ depDocTrace => do
- exeJob.bindAsync fun exeFile exeTrace => do
- modJob.bindSync fun _ modTrace => do
- let depTrace := mixTraceArray #[exeTrace, modTrace, depDocTrace]
- let trace ← buildFileUnlessUpToDate docFile depTrace do
- proc {
- cmd := exeFile.toString
- args := #["single", mod.name.toString, srcUri]
- env := ← getAugmentedEnv
- }
- return (docFile, trace)
+ bibPrepassJob.bindAsync fun _ bibPrepassTrace => do
+ exeJob.bindAsync fun exeFile exeTrace => do
+ modJob.bindSync fun _ modTrace => do
+ let depTrace := mixTraceArray #[exeTrace, modTrace, bibPrepassTrace, depDocTrace]
+ let trace ← buildFileUnlessUpToDate docFile depTrace do
+ proc {
+ cmd := exeFile.toString
+ args := #["single", mod.name.toString, srcUri]
+ env := ← getAugmentedEnv
+ }
+ return (docFile, trace)
-- TODO: technically speaking this target does not show all file dependencies
target coreDocs : FilePath := do
let exeJob ← «doc-gen4».fetch
+ let bibPrepassJob ← bibPrepass.fetch
let dataPath := (← getWorkspace).root.buildDir / "doc-data"
let dataFile := dataPath / "declaration-data-Lean.bmp"
- exeJob.bindSync fun exeFile exeTrace => do
- let trace ← buildFileUnlessUpToDate dataFile exeTrace do
- proc {
- cmd := exeFile.toString
- args := #["genCore"]
- env := ← getAugmentedEnv
- }
- return (dataFile, trace)
+ bibPrepassJob.bindAsync fun _ bibPrepassTrace => do
+ exeJob.bindSync fun exeFile exeTrace => do
+ let depTrace := mixTraceArray #[exeTrace, bibPrepassTrace]
+ let trace ← buildFileUnlessUpToDate dataFile depTrace do
+ proc {
+ cmd := exeFile.toString
+ args := #["genCore"]
+ env := ← getAugmentedEnv
+ }
+ return (dataFile, trace)
library_facet docs (lib) : FilePath := do
let mods ← lib.modules.fetch
diff --git a/lean-toolchain b/lean-toolchain
index d69d1ed2..6a4259fa 100644
--- a/lean-toolchain
+++ b/lean-toolchain
@@ -1 +1 @@
-leanprover/lean4:v4.10.0-rc1
+leanprover/lean4:v4.10.0-rc2