PhysLean/HepLean/Meta/Notes/Basic.lean

112 lines
3.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
-/
import Batteries.Lean.HashSet
import Mathlib.Lean.Expr.Basic
import Mathlib.Lean.CoreM
import ImportGraph.RequiredModules
/-!
2024-12-04 16:42:59 +00:00
## Underlying structure for notes
This file contains the necessary structures that must be imported into a file
for it to be turned into a Lean note.
It allows for the
- `note ...` command to be used to add notes to a Lean file
- `note_attr` attribute to be used to display formally verified commands within a note.
- `note_attr_informal` attribute to be used to display informal commands within a note.
Other results relating to notes are in other files.
2024-12-04 13:37:23 +00:00
-/
namespace HepLean
open Lean System Meta
2024-12-04 16:24:23 +00:00
/-- The information from a `note ...` command. To be used in a note file-/
2024-12-04 13:37:23 +00:00
structure NoteInfo where
2024-12-04 16:24:23 +00:00
/-- The content of the note. -/
2024-12-04 13:37:23 +00:00
content : String
2024-12-04 16:24:23 +00:00
/-- The file name where the note came from. -/
2024-12-04 13:37:23 +00:00
fileName : Name
2024-12-04 16:24:23 +00:00
/-- The line from where the note came from. -/
2024-12-04 13:37:23 +00:00
line : Nat
2025-01-14 00:11:25 +01:00
/-- Environment extension to store `note ...`. -/
2024-12-04 13:37:23 +00:00
initialize noteExtension : SimplePersistentEnvExtension NoteInfo (Array NoteInfo) ←
registerSimplePersistentEnvExtension {
2024-12-04 16:24:23 +00:00
name := `noteExtension
addEntryFn := fun arr nodeInfor => arr.push nodeInfor
addImportedFn := fun es => es.foldl (· ++ ·) #[]
2024-12-04 13:37:23 +00:00
}
2024-12-04 16:24:23 +00:00
/-- Syntax for the `note ...` command -/
2024-12-04 13:37:23 +00:00
syntax (name := note_comment) "note " str : command
2024-12-04 16:24:23 +00:00
/-- Elaborator for the `note ...` command -/
2024-12-04 13:37:23 +00:00
@[command_elab note_comment]
def elabNote : Lean.Elab.Command.CommandElab := fun stx =>
match stx with
| `(note $s) => do
let str : String := s.getString
let pos := stx.getPos?
match pos with
| some pos => do
let env ← getEnv
let fileMap ← getFileMap
let filePos := fileMap.toPosition pos
let line := filePos.line
let modName := env.mainModule
let noteInfo : NoteInfo := { content := str, fileName := modName, line := line }
modifyEnv fun env => noteExtension.addEntry env noteInfo
| none => throwError "Invalid syntax for `note` command"
| _ => throwError "Invalid syntax for `note` command"
/-!
## Note attribute
-/
2025-01-14 00:11:25 +01:00
/-- Environment extension to store `note_attr`. -/
2024-12-04 13:37:23 +00:00
initialize noteDeclExtension : SimplePersistentEnvExtension Name (Array Name) ←
registerSimplePersistentEnvExtension {
2024-12-04 16:24:23 +00:00
name := `noteDeclExtension
addEntryFn := fun arr nodeInfor => arr.push nodeInfor
addImportedFn := fun es => es.foldl (· ++ ·) #[]
2024-12-04 13:37:23 +00:00
}
2024-12-04 16:24:23 +00:00
/-- The `note_attr` attribute is used to display formally verified commands within a note. -/
2024-12-04 13:37:23 +00:00
initialize noteAttribute : Unit ←
registerBuiltinAttribute {
2024-12-04 16:24:23 +00:00
name := `note_attr
2024-12-04 13:37:23 +00:00
descr := "Note attribute"
applicationTime := AttributeApplicationTime.afterCompilation
2024-12-04 16:24:23 +00:00
add := fun declName _ _ => do
2024-12-04 13:37:23 +00:00
modifyEnv fun env => noteDeclExtension.addEntry env declName
}
2025-01-14 00:11:25 +01:00
/-- Environment extension to store `note_attr_informal`. -/
2024-12-04 13:37:23 +00:00
initialize noteInformalDeclExtension : SimplePersistentEnvExtension Name (Array Name) ←
registerSimplePersistentEnvExtension {
2024-12-04 16:24:23 +00:00
name := `noteInformalDeclExtension
addEntryFn := fun arr nodeInfor => arr.push nodeInfor
addImportedFn := fun es => es.foldl (· ++ ·) #[]
2024-12-04 13:37:23 +00:00
}
2024-12-04 16:24:23 +00:00
/-- The `note_attr_informal` attribute is used to display informal commands within a note. -/
2024-12-04 13:37:23 +00:00
initialize noteInformalAttribute : Unit ←
registerBuiltinAttribute {
2024-12-04 16:24:23 +00:00
name := `note_attr_informal
2024-12-04 13:37:23 +00:00
descr := "Informal note attribute"
applicationTime := AttributeApplicationTime.afterCompilation
2024-12-04 16:24:23 +00:00
add := fun declName _ _ => do
2024-12-04 13:37:23 +00:00
modifyEnv fun env => noteInformalDeclExtension.addEntry env declName
}
end HepLean