PhysLean/HepLean/Meta/Informal/Basic.lean

60 lines
2.2 KiB
Text
Raw Normal View History

2024-09-15 10:14:34 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import Lean.Parser.Term
2024-09-15 10:14:34 -04:00
/-!
## Informal definitions and lemmas
This file contains the necessary structures that must be imported into a file for it to contain
informal definitions and lemmas.
Everything else about informal definitions and lemmas are in the `Informal.Post` module.
2024-09-15 10:14:34 -04:00
-/
2024-09-15 10:18:05 -04:00
2025-01-14 00:11:25 +01:00
/-- The structure representing an informal definition. -/
2024-09-15 10:14:34 -04:00
structure InformalDefinition where
2024-09-16 05:32:40 -04:00
/-- The names of top-level commands we expect this definition to depend on. -/
deps : List Lean.Name
2024-09-15 10:14:34 -04:00
/-- The structure representing an informal lemma. -/
2024-09-15 10:14:34 -04:00
structure InformalLemma where
2024-09-16 05:32:40 -04:00
/-- The names of top-level commands we expect this lemma to depend on. -/
deps : List Lean.Name
2024-09-15 10:14:34 -04:00
namespace Informal
2024-12-04 13:37:23 +00:00
/-!
## Syntax
Using macros for syntax rewriting works better with the language server compared to
`Lake.DSL.LeanLibDecl`. Hovering over any definition of `informal_definition` or `informal_lemma`
gives a proper type hint like any proper definition using `def` whereas definitions of `lake_lib`
and `lake_exe` don't show docstrings and infer the type `Lean.Name`.
2024-09-15 10:14:34 -04:00
-/
2024-09-15 10:14:34 -04:00
open Lean.Parser.Term
2024-12-04 13:37:23 +00:00
/-- A placeholder for definitions to be formalized in the future. Docstrings of informal definitions
should outline its mathematical or physical content and specify useful references. Use the attribute
`note_attr_informal` from `HepLean.Meta.Notes.Basic` to mark the informal definition as a note.
-/
macro (name := informalDefinitionDecl)
doc?:(docComment)? attrs?:(attributes)? "informal_definition " name:ident body:declVal : command =>
`($[$doc?]? $[$attrs?]? def $name : InformalDefinition $body:declVal)
2024-09-16 07:40:15 -04:00
/-- A placeholder for lemmas to be formalized in the future. Docstrings of informal lemmas
should outline its mathematical or physical content and specify useful references. Use the attribute
`note_attr_informal` from `HepLean.Meta.Notes.Basic` to mark the informal definition as a note.
-/
macro (name := informalLemmaDecl)
doc?:(docComment)? attrs?:(attributes)? "informal_lemma " name:ident body:declVal : command =>
`($[$doc?]? $[$attrs?]? def $name : InformalLemma $body:declVal)
2024-09-15 10:14:34 -04:00
end Informal