feat: make informal_definition and informal_lemma commands (#300)
* make informal_definition and informal_lemma commands * drop the fields "math", "physics", and "proof" from InformalDefinition/InformalLemma and use docstrings instead * render informal docstring in dependency graph
This commit is contained in:
parent
6aab0ba3cd
commit
f8f94979ab
33 changed files with 666 additions and 1089 deletions
|
@ -3,8 +3,9 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|||
Released under Apache 2.0 license.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.Meta.Notes.NoteFile
|
||||
-- import DocGen4.Output.DocString
|
||||
import HepLean.Meta.Informal.Post
|
||||
import HepLean.Meta.Notes.NoteFile
|
||||
/-!
|
||||
|
||||
## Turns a delaration into a html note structure.
|
||||
|
@ -12,7 +13,7 @@ import HepLean.Meta.Informal.Post
|
|||
-/
|
||||
|
||||
namespace HepLean
|
||||
open Lean System Meta
|
||||
open Lean
|
||||
|
||||
/-- A `HTMLNote` is a structure containing the html information from
|
||||
individual contributions (commands, informal commands, note ..) etc. to a note file. -/
|
||||
|
@ -25,47 +26,52 @@ structure HTMLNote where
|
|||
line : Nat
|
||||
|
||||
/-- Converts a note infor into a HTMLNote. -/
|
||||
def HTMLNote.ofNodeInfo (ni : NoteInfo) : MetaM HTMLNote := do
|
||||
let line := ni.line
|
||||
let decl := ni.content
|
||||
let fileName := ni.fileName
|
||||
pure { content := decl, fileName := fileName, line := line }
|
||||
def HTMLNote.ofNodeInfo (ni : NoteInfo) : HTMLNote :=
|
||||
{ ni with }
|
||||
|
||||
/-- An formal definition or lemma to html for a note. -/
|
||||
def HTMLNote.ofFormal (name : Name) : MetaM HTMLNote := do
|
||||
let line ← Name.lineNumber name
|
||||
let decl ← Name.getDeclString name
|
||||
let fileName ← Name.fileName name
|
||||
let webAddress : String ← Name.toGitHubLink fileName line
|
||||
let content :=
|
||||
"<div class=\"code-block-container\">"
|
||||
++ "<a href=\"" ++ webAddress ++ "\" class=\"code-button\">View/Improve</a>"
|
||||
++"<pre><code>"
|
||||
++ decl ++
|
||||
"</code></pre></div>"
|
||||
pure { content := content, fileName := fileName, line := line }
|
||||
def HTMLNote.ofFormal (name : Name) : CoreM HTMLNote := do
|
||||
let line ← name.lineNumber
|
||||
let fileName ← name.fileName
|
||||
return {
|
||||
fileName, line,
|
||||
content := s!"
|
||||
<div class=\"code-block-container\">
|
||||
<a href=\"{fileName.toGitHubLink line}\" class=\"code-button\">View/Improve</a>
|
||||
<pre><code>{← name.getDeclString}</code></pre>
|
||||
</div>"
|
||||
}
|
||||
|
||||
/-- An informal definition or lemma to html for a note. -/
|
||||
unsafe def HTMLNote.ofInformal (name : Name) : MetaM HTMLNote := do
|
||||
let line ← Name.lineNumber name
|
||||
let fileName ← Name.fileName name
|
||||
let constInfo ← getConstInfo name
|
||||
let webAddress : String ← Name.toGitHubLink fileName line
|
||||
unsafe def HTMLNote.ofInformal (name : Name) : CoreM HTMLNote := do
|
||||
let line ← name.lineNumber
|
||||
let fileName ← name.fileName
|
||||
let constInfo ← getConstInfoDefn name
|
||||
let webAddress := fileName.toGitHubLink line
|
||||
let mut content := ""
|
||||
if Informal.isInformalDef constInfo then
|
||||
let X ← Informal.constantInfoToInformalDefinition constInfo
|
||||
content := "<div class=\"informal-def\">"
|
||||
++ "<a href=\"" ++ webAddress ++ "\" class=\"button\">Improve/Formalize</a>"
|
||||
++"<b>Informal definition:</b> " ++ name.toString ++ "<br>"
|
||||
++ X.math.replace "\n" "<br>"
|
||||
++ "</div>"
|
||||
else if Informal.isInformalLemma constInfo then
|
||||
let X ← Informal.constantInfoToInformalLemma constInfo
|
||||
content := "<div class=\"informal-def\">"
|
||||
++ "<a href=\"" ++ webAddress ++ "\" class=\"button\">Improve/Formalize</a>"
|
||||
++"<b>Informal lemma:</b> " ++ name.toString ++ "<br>"
|
||||
++ X.math.replace "\n" "<br>"
|
||||
++ "</div>"
|
||||
pure { content := content, fileName := fileName, line := line }
|
||||
if constInfo.type.isConstOf ``InformalDefinition then
|
||||
let doc ← name.getDocString
|
||||
-- let html ← DocGen4.Output.docStringToHtml doc name.toString
|
||||
-- let X ← Informal.constantInfoToInformalDefinition constInfo
|
||||
-- let fragment := X.math.replace "\n" "<br>"
|
||||
let fragment := doc.replace "\n" "<br>"
|
||||
content := s!"
|
||||
<div class=\"informal-def\">
|
||||
<a href=\"{webAddress}\" class=\"button\">Improve/Formalize</a>
|
||||
<b>Informal definition:</b> {name}<br>
|
||||
{fragment}
|
||||
</div>"
|
||||
else if constInfo.type.isConstOf ``InformalLemma then
|
||||
-- let X ← Informal.constantInfoToInformalLemma constInfo
|
||||
-- let fragment := X.math.replace "\n" "<br>"
|
||||
let doc ← name.getDocString
|
||||
let fragment := doc.replace "\n" "<br>"
|
||||
content := s!"
|
||||
<div class=\"informal-def\">
|
||||
<a href=\"{webAddress}\" class=\"button\">Improve/Formalize</a>
|
||||
<b>Informal lemma:</b> {name}<br>
|
||||
{fragment}
|
||||
</div>"
|
||||
return { content, fileName, line }
|
||||
|
||||
end HepLean
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue