/- Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved. Released under Apache 2.0 license. Authors: Joseph Tooby-Smith -/ import HepLean.Meta.Notes.HTMLNote /-! ## Turns a declaration into a html note structure. -/ namespace HepLean open Lean System Meta namespace NoteFile variable (N : NoteFile) /-- Sorts `NoteInfo`'s by file name then by line number. -/ def sortLE (ni1 ni2 : HTMLNote) : Bool := if N.files.indexOf ni1.fileName ≠ N.files.indexOf ni2.fileName then N.files.indexOf ni1.fileName ≤ N.files.indexOf ni2.fileName else ni1.line ≤ ni2.line /-- Returns a sorted list of NodeInfos for a file system. -/ unsafe def getNodeInfo : CoreM (List HTMLNote) := do let env ← getEnv let allNotes := (noteExtension.getState env) let allDecl := (noteDeclExtension.getState env) let allInformalDecl := noteInformalDeclExtension.getState env let allNoteInfo := allNotes.map HTMLNote.ofNodeInfo ++ (← allDecl.mapM HTMLNote.ofFormal) ++ (← allInformalDecl.mapM HTMLNote.ofInformal) let noteInfo := allNoteInfo.filter (fun x => x.fileName ∈ N.files) let noteInfoSort := noteInfo.toList.mergeSort N.sortLE pure noteInfoSort /-- The HTML code needed to have syntax highlighting. -/ def codeBlockHTML : String := " " /-- The html styles for informal definitions. -/ def informalDefStyle : String := " " /-- The html styles for code button. -/ def codeButton : String := " " /-- HTML allowing the use of mathjax. -/ def mathJaxScript : String := " " /-- The header to the html code. -/ def headerHTML : String := "--- layout: default ---
" ++ codeBlockHTML ++ mathJaxScript ++ informalDefStyle ++ codeButton ++ " " /-- The html code corresponding to the title, abstract and authors. -/ def titleHTML : String := "Note: These are not ordinary notes. They are created using an interactive theorem
prover called Lean.
Lean formally checks definitions, theorems and proofs for correctness.
These notes are part of a much larger project called
HepLean, which aims to digitalize
high energy physics into Lean. Please consider contributing to this project.
Please provide feedback or suggestions for improvements by creating a GitHub issue
here.