PhysLean/HepLean/Meta/Notes/HTMLNote.lean

72 lines
2.6 KiB
Text
Raw Normal View History

2024-12-04 13:37:23 +00:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
2024-12-04 16:42:59 +00:00
import HepLean.Meta.Notes.NoteFile
import HepLean.Meta.Informal.Post
2024-12-04 13:37:23 +00:00
/-!
## Turns a delaration into a html note structure.
-/
namespace HepLean
open Lean System Meta
2024-12-04 16:24:23 +00:00
/-- A `HTMLNote` is a structure containing the html information from
2025-01-14 00:11:25 +01:00
individual contributions (commands, informal commands, note ..) etc. to a note file. -/
2024-12-04 13:37:23 +00:00
structure HTMLNote where
2024-12-04 16:24:23 +00:00
/-- The name of the file the contribution came from. -/
2024-12-04 13:37:23 +00:00
fileName : Name
2024-12-04 16:24:23 +00:00
/-- The html contribution of the content. -/
2024-12-04 13:37:23 +00:00
content : String
2024-12-04 16:24:23 +00:00
/-- The line in the file where the contribution came from. -/
2024-12-04 13:37:23 +00:00
line : Nat
2024-12-04 16:24:23 +00:00
/-- Converts a note infor into a HTMLNote. -/
def HTMLNote.ofNodeInfo (ni : NoteInfo) : MetaM HTMLNote := do
2024-12-04 13:37:23 +00:00
let line := ni.line
2024-12-04 16:24:23 +00:00
let decl := ni.content
2024-12-04 13:37:23 +00:00
let fileName := ni.fileName
pure { content := decl, fileName := fileName, line := line }
2024-12-04 16:24:23 +00:00
/-- An formal definition or lemma to html for a note. -/
def HTMLNote.ofFormal (name : Name) : MetaM HTMLNote := do
2024-12-04 13:37:23 +00:00
let line ← Name.lineNumber name
let decl ← Name.getDeclString name
let fileName ← Name.fileName name
2024-12-05 16:35:00 +00:00
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>"
2024-12-04 13:37:23 +00:00
pure { content := content, fileName := fileName, line := line }
2024-12-04 16:24:23 +00:00
/-- An informal definition or lemma to html for a note. -/
unsafe def HTMLNote.ofInformal (name : Name) : MetaM HTMLNote := do
2024-12-04 13:37:23 +00:00
let line ← Name.lineNumber name
let fileName ← Name.fileName name
let constInfo ← getConstInfo name
2024-12-05 07:12:56 +00:00
let webAddress : String ← Name.toGitHubLink fileName line
2024-12-04 13:37:23 +00:00
let mut content := ""
if Informal.isInformalDef constInfo then
let X ← Informal.constantInfoToInformalDefinition constInfo
content := "<div class=\"informal-def\">"
2024-12-05 07:12:56 +00:00
++ "<a href=\"" ++ webAddress ++ "\" class=\"button\">Improve/Formalize</a>"
2024-12-04 13:37:23 +00:00
++"<b>Informal definition:</b> " ++ name.toString ++ "<br>"
++ X.math.replace "\n" "<br>"
++ "</div>"
else if Informal.isInformalLemma constInfo then
let X ← Informal.constantInfoToInformalLemma constInfo
2024-12-09 12:54:12 +00:00
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>"
2024-12-04 13:37:23 +00:00
pure { content := content, fileName := fileName, line := line }
end HepLean