Merge pull request #152 from HEPLean/informal_defs
feat: Add informal_definition and informal_lemma
This commit is contained in:
commit
9071ee8354
9 changed files with 447 additions and 15 deletions
28
.github/workflows/docs.yml
vendored
28
.github/workflows/docs.yml
vendored
|
@ -1,18 +1,18 @@
|
|||
on:
|
||||
push:
|
||||
branches:
|
||||
on:
|
||||
push:
|
||||
branches:
|
||||
- master
|
||||
|
||||
name: Build and deploy documentation
|
||||
|
||||
# borrowed from https://github.com/teorth/pfr/blob/master/.github/workflows/push.yml
|
||||
permissions:
|
||||
permissions:
|
||||
contents: read
|
||||
pages: write
|
||||
id-token: write
|
||||
|
||||
jobs:
|
||||
build_project:
|
||||
jobs:
|
||||
build_project:
|
||||
runs-on: ubuntu-latest
|
||||
name: Build project
|
||||
steps:
|
||||
|
@ -20,7 +20,7 @@ jobs:
|
|||
uses: actions/checkout@v4
|
||||
with:
|
||||
fetch-depth: 0
|
||||
|
||||
|
||||
##################
|
||||
# Remove this section if you don't want docs to be made
|
||||
- name: Install elan
|
||||
|
@ -53,13 +53,14 @@ jobs:
|
|||
key: MathlibDoc-${{ hashFiles('lake-manifest.json') }}
|
||||
restore-keys: |
|
||||
MathlibDoc-
|
||||
|
||||
|
||||
- name: Build documentation
|
||||
run: lake -Kenv=dev build HepLean:docs
|
||||
|
||||
- name: make TODO list
|
||||
run : lake exe find_TODOs mkFile
|
||||
|
||||
- name: make TODO list
|
||||
run : lake exe find_TODOs mkFile
|
||||
- name: make list of informal proofs and lemmas
|
||||
run : lake exe informal mkFile
|
||||
- name: Copy documentation to `docs/docs`
|
||||
run: |
|
||||
mv .lake/build/doc docs/docs
|
||||
|
@ -71,7 +72,7 @@ jobs:
|
|||
working-directory: docs
|
||||
ruby-version: "3.1" # Not needed with a .ruby-version file
|
||||
bundler-cache: true # runs 'bundle install' and caches installed gems automatically
|
||||
|
||||
|
||||
|
||||
- name: Bundle website
|
||||
working-directory: docs
|
||||
|
@ -85,6 +86,3 @@ jobs:
|
|||
- name: Deploy to GitHub Pages
|
||||
id: deployment
|
||||
uses: actions/deploy-pages@v1
|
||||
|
||||
|
||||
|
||||
|
|
|
@ -61,6 +61,7 @@ import HepLean.FlavorPhysics.CKMMatrix.StandardParameterization.StandardParamete
|
|||
import HepLean.Mathematics.LinearMaps
|
||||
import HepLean.Mathematics.SO3.Basic
|
||||
import HepLean.Meta.AllFilePaths
|
||||
import HepLean.Meta.Informal
|
||||
import HepLean.Meta.TransverseTactics
|
||||
import HepLean.SpaceTime.Basic
|
||||
import HepLean.SpaceTime.CliffordAlgebra
|
||||
|
@ -82,6 +83,7 @@ import HepLean.SpaceTime.LorentzVector.LorentzAction
|
|||
import HepLean.SpaceTime.LorentzVector.NormOne
|
||||
import HepLean.SpaceTime.MinkowskiMetric
|
||||
import HepLean.SpaceTime.SL2C.Basic
|
||||
import HepLean.SpaceTime.WeylFermion.Basic
|
||||
import HepLean.StandardModel.Basic
|
||||
import HepLean.StandardModel.HiggsBoson.Basic
|
||||
import HepLean.StandardModel.HiggsBoson.GaugeAction
|
||||
|
|
157
HepLean/Meta/Informal.lean
Normal file
157
HepLean/Meta/Informal.lean
Normal file
|
@ -0,0 +1,157 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import Lean
|
||||
/-!
|
||||
|
||||
## Informal definitions and lemmas
|
||||
|
||||
-/
|
||||
open Lean Elab System
|
||||
|
||||
/-! TODO: Derive a means to extract informal definitions and informal lemmas. -/
|
||||
|
||||
/-! TODO: Can likely make this a bona-fide command. -/
|
||||
|
||||
/-- The structure representating an informal definition. -/
|
||||
structure InformalDefinition where
|
||||
/-- The name of the informal definition. This is autogenerated. -/
|
||||
name : Name
|
||||
/-- The mathematical description of the definition. -/
|
||||
math : String
|
||||
/-- The physics description of the definition. -/
|
||||
physics : String
|
||||
/-- References. -/
|
||||
ref : String
|
||||
/-- The names of top-level commands we expect this definition to depend on. -/
|
||||
dependencies : List Name
|
||||
|
||||
/-- The structure representating an informal proof. -/
|
||||
structure InformalLemma where
|
||||
/-- The name of the informal lemma. This is autogenerated. -/
|
||||
name : Name
|
||||
/-- The mathematical description of the lemma. -/
|
||||
math : String
|
||||
/-- The physics description of the lemma. -/
|
||||
physics : String
|
||||
/-- A description of the proof. -/
|
||||
proof : String
|
||||
/-- References. -/
|
||||
ref : String
|
||||
/-- The names of top-level commands we expect this lemma to depend on. -/
|
||||
dependencies : List Name
|
||||
|
||||
namespace Informal
|
||||
|
||||
/-- The Parser.Category we will use for assignments. -/
|
||||
declare_syntax_cat informalAssignment
|
||||
|
||||
/-- The syntax describing an informal assignment of `ident` to a string. -/
|
||||
syntax (name := informalAssignment) ident ":≈" str : informalAssignment
|
||||
|
||||
/-- The syntax describing an informal assignment of `ident` to a list. -/
|
||||
syntax (name := informalAssignmentDeps) ident ":≈" "[" sepBy(term, ",") "]" : informalAssignment
|
||||
|
||||
/-- The syntax for the command informal_definition. -/
|
||||
syntax (name := informal_definition) "informal_definition " ident
|
||||
" where " (colGt informalAssignment)* : command
|
||||
|
||||
/-- The macro turning `informal_definition ... where ...` into a definition. -/
|
||||
macro "informal_definition " name:ident " where " assignments:informalAssignment* : command => do
|
||||
let mut math_def? : Option (TSyntax `term) := none
|
||||
let mut physics_def? : Option (TSyntax `term) := none
|
||||
let mut ref_def? : Option (TSyntax `term) := none
|
||||
let mut dep_def? : Option (TSyntax `term) := none
|
||||
for assignment in assignments do
|
||||
match assignment with
|
||||
| `(informalAssignment| physics :≈ $val:str) =>
|
||||
let some strVal := val.raw.isStrLit? | Macro.throwError "Expected string literal for physics"
|
||||
physics_def? := some (← `($(Lean.quote strVal)))
|
||||
| `(informalAssignment| math :≈ $val:str) =>
|
||||
let some strVal := val.raw.isStrLit? | Macro.throwError "Expected string literal for math"
|
||||
math_def? := some (← `($(Lean.quote strVal)))
|
||||
| `(informalAssignment| ref :≈ $val:str) =>
|
||||
let some strVal := val.raw.isStrLit? | Macro.throwError "Expected string literal for ref"
|
||||
ref_def? := some (← `($(Lean.quote strVal)))
|
||||
| `(informalAssignmentDeps| deps :≈ [ $deps,* ]) =>
|
||||
dep_def? := some (← `([ $deps,* ]))
|
||||
| _ => Macro.throwError "invalid assignment syntax in informal_definition"
|
||||
unless math_def?.isSome do
|
||||
Macro.throwError "A 'math' assignments is required"
|
||||
`(
|
||||
/-- An informal definition. -/
|
||||
def $name : InformalDefinition := {
|
||||
name := $(Lean.quote name.getId),
|
||||
physics := $(physics_def?.getD (← `("No physics description provided"))),
|
||||
math := $(math_def?.getD (panic! "math not assigned")),
|
||||
ref := $(ref_def?.getD (← `("No references provided"))),
|
||||
dependencies := $(dep_def?.getD (← `([])))
|
||||
})
|
||||
|
||||
/-- The syntax for the command `informal_lemma`. -/
|
||||
syntax (name := informal_lemma) "informal_lemma " ident " where "
|
||||
(colGt informalAssignment)* : command
|
||||
|
||||
/-- The macro turning `informal_lemma ... where ...` into a definition. -/
|
||||
macro "informal_lemma " name:ident " where " assignments:informalAssignment* : command => do
|
||||
let mut math_def? : Option (TSyntax `term) := none
|
||||
let mut physics_def? : Option (TSyntax `term) := none
|
||||
let mut proof_def? : Option (TSyntax `term) := none
|
||||
let mut ref_def? : Option (TSyntax `term) := none
|
||||
let mut dep_def? : Option (TSyntax `term) := none
|
||||
for assignment in assignments do
|
||||
match assignment with
|
||||
| `(informalAssignment| physics :≈ $val:str) =>
|
||||
let some strVal := val.raw.isStrLit? | Macro.throwError "Expected string literal for physics"
|
||||
physics_def? := some (← `($(Lean.quote strVal)))
|
||||
| `(informalAssignment| math :≈ $val:str) =>
|
||||
let some strVal := val.raw.isStrLit? | Macro.throwError "Expected string literal for math"
|
||||
math_def? := some (← `($(Lean.quote strVal)))
|
||||
| `(informalAssignment| proof :≈ $val:str) =>
|
||||
let some strVal := val.raw.isStrLit? | Macro.throwError "Expected string literal for proof"
|
||||
proof_def? := some (← `($(Lean.quote strVal)))
|
||||
| `(informalAssignment| ref :≈ $val:str) =>
|
||||
let some strVal := val.raw.isStrLit? | Macro.throwError "Expected string literal for ref"
|
||||
ref_def? := some (← `($(Lean.quote strVal)))
|
||||
| `(informalAssignmentDeps| deps :≈ [ $deps,* ]) =>
|
||||
dep_def? := some (← `([ $deps,* ]))
|
||||
| _ => Macro.throwError "invalid assignment syntax"
|
||||
unless math_def?.isSome do
|
||||
Macro.throwError "A 'math' assignments is required"
|
||||
`(
|
||||
|
||||
/-- An informal lemma. -/
|
||||
def $name : InformalLemma := {
|
||||
name := $(Lean.quote name.getId),
|
||||
physics := $(physics_def?.getD (← `("No physics description provided"))),
|
||||
math := $(math_def?.getD (panic! "math not assigned")),
|
||||
proof := $(proof_def?.getD (← `("No proof description provided"))),
|
||||
ref := $(ref_def?.getD (← `("No references provided"))),
|
||||
dependencies := $(dep_def?.getD (← `([])))
|
||||
})
|
||||
|
||||
/-- Is true if and only if a `ConstantInfo` corresponds to an `InformalLemma` or a
|
||||
`InformalDefinition`. -/
|
||||
def isInformal (c : ConstantInfo) : Bool :=
|
||||
match c with
|
||||
| ConstantInfo.defnInfo c =>
|
||||
if c.type.isAppOf ``InformalDefinition ∨ c.type.isAppOf ``InformalLemma then true else false
|
||||
| _ => false
|
||||
|
||||
/-- Is true if and only if a `ConstantInfo` corresponds to an `InformalLemma`. -/
|
||||
def isInformalLemma (c : ConstantInfo) : Bool :=
|
||||
match c with
|
||||
| ConstantInfo.defnInfo c =>
|
||||
if c.type.isAppOf ``InformalLemma then true else false
|
||||
| _ => false
|
||||
|
||||
/-- Is true if and only if a `ConstantInfo` corresponds to an `InformalDefinition`. -/
|
||||
def isInformalDef (c : ConstantInfo) : Bool :=
|
||||
match c with
|
||||
| ConstantInfo.defnInfo c =>
|
||||
if c.type.isAppOf ``InformalDefinition then true else false
|
||||
| _ => false
|
||||
|
||||
end Informal
|
61
HepLean/SpaceTime/WeylFermion/Basic.lean
Normal file
61
HepLean/SpaceTime/WeylFermion/Basic.lean
Normal file
|
@ -0,0 +1,61 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.Meta.Informal
|
||||
/-!
|
||||
|
||||
# Weyl fermions
|
||||
|
||||
-/
|
||||
|
||||
/-!
|
||||
|
||||
## The definition of Weyl fermion vector spaces.
|
||||
|
||||
-/
|
||||
|
||||
informal_definition leftHandedWeylFermion where
|
||||
math :≈ "The vector space ℂ^2 carrying the fundamental representation of SL(2,C)."
|
||||
ref :≈ "https://particle.physics.ucdavis.edu/modernsusy/slides/slideimages/spinorfeynrules.pdf"
|
||||
|
||||
informal_definition rightHandedWeylFermion where
|
||||
math :≈ "The vector space ℂ^2 carrying the conjguate representation of SL(2,C)."
|
||||
ref :≈ "https://particle.physics.ucdavis.edu/modernsusy/slides/slideimages/spinorfeynrules.pdf"
|
||||
|
||||
informal_definition altLeftHandedWeylFermion where
|
||||
math :≈ "The vector space ℂ^2 carrying the representation of SL(2,C) given by
|
||||
M → (M⁻¹)ᵀ."
|
||||
ref :≈ "https://particle.physics.ucdavis.edu/modernsusy/slides/slideimages/spinorfeynrules.pdf"
|
||||
|
||||
informal_definition altRightHandedWeylFermion where
|
||||
math :≈ "The vector space ℂ^2 carrying the representation of SL(2,C) given by
|
||||
M → (M⁻¹)^†."
|
||||
ref :≈ "https://particle.physics.ucdavis.edu/modernsusy/slides/slideimages/spinorfeynrules.pdf"
|
||||
|
||||
/-!
|
||||
|
||||
## Equivalences between Weyl fermion vector spaces.
|
||||
|
||||
-/
|
||||
|
||||
informal_definition leftHandedWeylFermionAltEquiv where
|
||||
math :≈ "The linear equiv between leftHandedWeylFermion and altLeftHandedWeylFermion given
|
||||
by multiplying an element of rightHandedWeylFermion by the matrix `εᵃ⁰ᵃ¹ = !![0, 1; -1, 0]]`."
|
||||
deps :≈ [`leftHandedWeylFermion, `altLeftHandedWeylFermion]
|
||||
|
||||
informal_lemma leftHandedWeylFermionAltEquiv_equivariant where
|
||||
math :≈ "The linear equiv leftHandedWeylFermionAltEquiv is equivariant with respect to the
|
||||
action of SL(2,C) on leftHandedWeylFermion and altLeftHandedWeylFermion."
|
||||
deps :≈ [`leftHandedWeylFermionAltEquiv]
|
||||
|
||||
informal_definition rightHandedWeylFermionAltEquiv where
|
||||
math :≈ "The linear equiv between rightHandedWeylFermion and altRightHandedWeylFermion given
|
||||
by multiplying an element of rightHandedWeylFermion by the matrix `εᵃ⁰ᵃ¹ = !![0, 1; -1, 0]]`"
|
||||
deps :≈ [`rightHandedWeylFermion, `altRightHandedWeylFermion]
|
||||
|
||||
informal_lemma rightHandedWeylFermionAltEquiv_equivariant where
|
||||
math :≈ "The linear equiv rightHandedWeylFermionAltEquiv is equivariant with respect to the
|
||||
action of SL(2,C) on rightHandedWeylFermion and altRightHandedWeylFermion."
|
||||
deps :≈ [`rightHandedWeylFermionAltEquiv]
|
|
@ -11,6 +11,7 @@ import Mathlib.Geometry.Manifold.VectorBundle.SmoothSection
|
|||
import Mathlib.Geometry.Manifold.Instances.Real
|
||||
import Mathlib.Analysis.InnerProductSpace.Basic
|
||||
import Mathlib.Geometry.Manifold.ContMDiff.Product
|
||||
import HepLean.Meta.Informal
|
||||
/-!
|
||||
|
||||
# The Higgs field
|
||||
|
@ -173,6 +174,12 @@ def ofReal (a : ℝ) : HiggsField := (HiggsVec.ofReal a).toField
|
|||
/-- The higgs field which is all zero. -/
|
||||
def zero : HiggsField := ofReal 0
|
||||
|
||||
informal_lemma zero_is_zero_section where
|
||||
physics :≈ "The zero Higgs field is the zero section of the Higgs bundle."
|
||||
math :≈ "The HiggsField `zero` defined by `ofReal 0`
|
||||
is the constant zero-section of the bundle `HiggsBundle`."
|
||||
deps :≈ [`StandardModel.HiggsField.zero]
|
||||
|
||||
end HiggsField
|
||||
|
||||
end
|
||||
|
|
|
@ -190,6 +190,14 @@ theorem rotate_fst_real_snd_zero (φ : HiggsVec) :
|
|||
· simp only [Fin.mk_one, Fin.isValue, Pi.smul_apply, Function.comp_apply, cons_val_one, head_cons,
|
||||
tail_cons, smul_zero]
|
||||
|
||||
informal_lemma stablity_group where
|
||||
physics :≈ "The Higgs boson breaks electroweak symmetry down to the electromagnetic force."
|
||||
math :≈ "The stablity group of the action of `rep` on `![0, Complex.ofReal ‖φ‖]`,
|
||||
for non-zero `‖φ‖` is the `SU(3) x U(1)` subgroup of
|
||||
`gaugeGroup := SU(3) x SU(2) x U(1)` with the embedding given by
|
||||
`(g, e^{i θ}) ↦ (g, diag (e ^ {3 * i θ}, e ^ {- 3 * i θ}), e^{i θ})`."
|
||||
deps :≈ [`StandardModel.HiggsVec, `StandardModel.HiggsVec.rep]
|
||||
|
||||
end HiggsVec
|
||||
|
||||
/-! TODO: Define the global gauge action on HiggsField. -/
|
||||
|
|
|
@ -5,6 +5,7 @@ Authors: Joseph Tooby-Smith
|
|||
-/
|
||||
import Mathlib.Algebra.QuadraticDiscriminant
|
||||
import HepLean.StandardModel.HiggsBoson.PointwiseInnerProd
|
||||
import HepLean.Meta.Informal
|
||||
/-!
|
||||
# The potential of the Higgs field
|
||||
|
||||
|
@ -314,6 +315,12 @@ lemma isBounded_of_𝓵_pos (h : 0 < P.𝓵) : P.IsBounded := by
|
|||
have h2' := h2 φ x
|
||||
linarith
|
||||
|
||||
informal_lemma isBounded_iff_of_𝓵_zero where
|
||||
physics :≈ "When there is no quartic coupling, the potential is bounded iff the mass squared is
|
||||
non-positive."
|
||||
math :≈ "For `P : Potential` then P.IsBounded if and only if P.μ2 ≤ 0.
|
||||
That is to say `- P.μ2 * ‖φ‖_H ^ 2 x` is bounded below if and only if `P.μ2 ≤ 0`."
|
||||
|
||||
/-!
|
||||
|
||||
## Minimum and maximum
|
||||
|
|
|
@ -64,6 +64,10 @@ srcDir = "scripts"
|
|||
name = "free_simps"
|
||||
srcDir = "scripts/MetaPrograms"
|
||||
|
||||
[[lean_exe]]
|
||||
name = "informal"
|
||||
srcDir = "scripts/MetaPrograms"
|
||||
|
||||
# -- Optional inclusion of openAI_doc_check. Needs `llm` above.
|
||||
#[[lean_exe]]
|
||||
#name = "openAI_doc_check"
|
||||
|
|
188
scripts/MetaPrograms/informal.lean
Normal file
188
scripts/MetaPrograms/informal.lean
Normal file
|
@ -0,0 +1,188 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import Batteries.Lean.HashSet
|
||||
import Lean
|
||||
import Mathlib.Lean.Expr.Basic
|
||||
import Mathlib.Lean.CoreM
|
||||
import HepLean.Meta.Informal
|
||||
import ImportGraph.RequiredModules
|
||||
/-!
|
||||
|
||||
# Extracting information about informal definitions and lemmas
|
||||
|
||||
-/
|
||||
|
||||
open Lean System Meta
|
||||
|
||||
def getConst (imp : Import) : IO (Array (Import × ConstantInfo)) := do
|
||||
let mFile ← findOLean imp.module
|
||||
let (modData, _) ← readModuleData mFile
|
||||
pure (modData.constants.map (fun c => (imp, c)))
|
||||
|
||||
def getLineNumber (c : Name) : MetaM Nat := do
|
||||
match ← findDeclarationRanges? c with
|
||||
| some decl => pure decl.range.pos.line
|
||||
| none => panic! s!"{c} is a declaration without position"
|
||||
|
||||
def getModule (c : Name) : MetaM Name := do
|
||||
match Lean.Environment.getModuleFor? (← getEnv) c with
|
||||
| some mod => pure mod
|
||||
| none => panic! s!"{c} is a declaration without position"
|
||||
|
||||
def depToString (d : Name) : MetaM String := do
|
||||
let lineNo ← getLineNumber d
|
||||
let mod ← getModule d
|
||||
pure s!" * {d}: ./{mod.toString.replace "." "/" ++ ".lean"}:{lineNo}"
|
||||
|
||||
def depToWebString (d : Name) : MetaM String := do
|
||||
let lineNo ← getLineNumber d
|
||||
let mod ← getModule d
|
||||
let webPath := "https://github.com/HEPLean/HepLean/blob/master/"++
|
||||
(mod.toString.replace "." "/") ++ ".lean"
|
||||
pure s!" * [{d}]({webPath}#L{lineNo})"
|
||||
|
||||
unsafe def informalLemmaToString (c : Import × ConstantInfo) : MetaM String := do
|
||||
let lineNo ← getLineNumber c.2.name
|
||||
let informalLemma ←
|
||||
match c.2 with
|
||||
| ConstantInfo.defnInfo c =>
|
||||
Lean.Meta.evalExpr' InformalLemma ``InformalLemma c.value
|
||||
| _ => panic! "This should not happen"
|
||||
let dep ← informalLemma.dependencies.mapM fun d => depToString d
|
||||
pure s!"
|
||||
Informal lemma: {informalLemma.name}
|
||||
- ./{c.1.module.toString.replace "." "/" ++ ".lean"}:{lineNo}
|
||||
- Math description: {informalLemma.math}
|
||||
- Physics description: {informalLemma.physics}
|
||||
- Proof description: {informalLemma.proof}
|
||||
- References: {informalLemma.ref}
|
||||
- Dependencies:\n{String.intercalate "\n" dep}"
|
||||
|
||||
unsafe def informalLemmaToWebString (c : Import × ConstantInfo) : MetaM String := do
|
||||
let lineNo ← getLineNumber c.2.name
|
||||
let informalLemma ←
|
||||
match c.2 with
|
||||
| ConstantInfo.defnInfo c =>
|
||||
Lean.Meta.evalExpr' InformalLemma ``InformalLemma c.value
|
||||
| _ => panic! "This should not happen"
|
||||
let dep ← informalLemma.dependencies.mapM fun d => depToWebString d
|
||||
let webPath := "https://github.com/HEPLean/HepLean/blob/master/"++
|
||||
(c.1.module.toString.replace "." "/") ++ ".lean"
|
||||
pure s!"
|
||||
**Informal lemma**: [{informalLemma.name}]({webPath}#L{lineNo}) :=
|
||||
*{informalLemma.math}*
|
||||
- Physics description: {informalLemma.physics}
|
||||
- Proof description: {informalLemma.proof}
|
||||
- References: {informalLemma.ref}
|
||||
- Dependencies:\n{String.intercalate "\n" dep}"
|
||||
|
||||
unsafe def informalDefToString (c : Import × ConstantInfo) : MetaM String := do
|
||||
let lineNo ← getLineNumber c.2.name
|
||||
let informalDef ←
|
||||
match c.2 with
|
||||
| ConstantInfo.defnInfo c =>
|
||||
Lean.Meta.evalExpr' InformalDefinition ``InformalDefinition c.value
|
||||
| _ => panic! "This should not happen"
|
||||
let dep ← informalDef.dependencies.mapM fun d => depToString d
|
||||
pure s!"
|
||||
Informal def: {informalDef.name}
|
||||
- ./{c.1.module.toString.replace "." "/" ++ ".lean"}:{lineNo}
|
||||
- Math description: {informalDef.math}
|
||||
- Physics description: {informalDef.physics}
|
||||
- References: {informalDef.ref}
|
||||
- Dependencies:\n{String.intercalate "\n" dep}"
|
||||
|
||||
unsafe def informalDefToWebString (c : Import × ConstantInfo) : MetaM String := do
|
||||
let lineNo ← getLineNumber c.2.name
|
||||
let informalDef ←
|
||||
match c.2 with
|
||||
| ConstantInfo.defnInfo c =>
|
||||
Lean.Meta.evalExpr' InformalDefinition ``InformalDefinition c.value
|
||||
| _ => panic! "This should not happen"
|
||||
let dep ← informalDef.dependencies.mapM fun d => depToWebString d
|
||||
let webPath := "https://github.com/HEPLean/HepLean/blob/master/"++
|
||||
(c.1.module.toString.replace "." "/") ++ ".lean"
|
||||
pure s!"
|
||||
**Informal def**: [{informalDef.name}]({webPath}#L{lineNo}) :=
|
||||
*{informalDef.math}*
|
||||
- Physics description: {informalDef.physics}
|
||||
- References: {informalDef.ref}
|
||||
- Dependencies:\n{String.intercalate "\n" dep}"
|
||||
|
||||
unsafe def informalToString (c : Import × ConstantInfo) : MetaM String := do
|
||||
if Informal.isInformalLemma c.2 then
|
||||
informalLemmaToString c
|
||||
else if Informal.isInformalDef c.2 then
|
||||
informalDefToString c
|
||||
else
|
||||
pure ""
|
||||
|
||||
unsafe def informalToWebString (c : Import × ConstantInfo) : MetaM String := do
|
||||
if Informal.isInformalLemma c.2 then
|
||||
informalLemmaToWebString c
|
||||
else if Informal.isInformalDef c.2 then
|
||||
informalDefToWebString c
|
||||
else
|
||||
pure ""
|
||||
|
||||
def informalFileHeader : String := s!"
|
||||
# Informal definitions and lemmas
|
||||
|
||||
This file is automatically generated using `informal_lemma` and `informal_definition` commands
|
||||
appearing in the raw Lean code of HepLean.
|
||||
|
||||
There is an implicit invitation to the reader to contribute to the formalization of
|
||||
informal definitions and lemmas. You may also wish to contribute to HepLean by including
|
||||
informal definitions and lemmas in the raw Lean code, especially if you do not have a
|
||||
background in Lean.
|
||||
|
||||
"
|
||||
|
||||
unsafe def importToString (i : Import) : MetaM String := do
|
||||
let constants ← getConst i
|
||||
let constants := constants.filter (fun c => ¬ Lean.Name.isInternalDetail c.2.name)
|
||||
let informalConst := constants.filter fun c => Informal.isInformal c.2
|
||||
let informalConstLineNo ← informalConst.mapM fun c => getLineNumber c.2.name
|
||||
let informalConstWithLineNo := informalConst.zip informalConstLineNo
|
||||
let informalConstWithLineNoSort := informalConstWithLineNo.qsort (fun a b => a.2 < b.2)
|
||||
let informalConst := informalConstWithLineNoSort.map (fun c => c.1)
|
||||
let informalPrint ← (informalConst.mapM informalToString).run'
|
||||
if informalPrint.isEmpty then
|
||||
pure ""
|
||||
else
|
||||
pure ("\n\n" ++ i.module.toString ++ "\n" ++ String.intercalate "\n\n" informalPrint.toList)
|
||||
|
||||
unsafe def importToWebString (i : Import) : MetaM String := do
|
||||
let constants ← getConst i
|
||||
let constants := constants.filter (fun c => ¬ Lean.Name.isInternalDetail c.2.name)
|
||||
let informalConst := constants.filter fun c => Informal.isInformal c.2
|
||||
let informalConstLineNo ← informalConst.mapM fun c => getLineNumber c.2.name
|
||||
let informalConstWithLineNo := informalConst.zip informalConstLineNo
|
||||
let informalConstWithLineNoSort := informalConstWithLineNo.qsort (fun a b => a.2 < b.2)
|
||||
let informalConst := informalConstWithLineNoSort.map (fun c => c.1)
|
||||
let informalPrint ← (informalConst.mapM informalToWebString).run'
|
||||
if informalPrint.isEmpty then
|
||||
pure ""
|
||||
else
|
||||
pure ("\n\n## " ++ i.module.toString ++ "\n" ++ String.intercalate "\n\n" informalPrint.toList)
|
||||
unsafe def main (args : List String) : IO Unit := do
|
||||
initSearchPath (← findSysroot)
|
||||
let mods : Name := `HepLean
|
||||
let imp : Import := {module := mods}
|
||||
let mFile ← findOLean imp.module
|
||||
unless (← mFile.pathExists) do
|
||||
throw <| IO.userError s!"object file '{mFile}' of module {imp.module} does not exist"
|
||||
let (hepLeanMod, _) ← readModuleData mFile
|
||||
let imports := hepLeanMod.imports.filter (fun c => c.module ≠ `Init)
|
||||
let importString ← CoreM.withImportModules #[`HepLean] (imports.mapM importToString).run'
|
||||
IO.println (String.intercalate "" importString.toList)
|
||||
/- Writing out informal file. -/
|
||||
let fileOut : System.FilePath := {toString := "./docs/Informal.md"}
|
||||
if "mkFile" ∈ args then
|
||||
let importWebString ← CoreM.withImportModules #[`HepLean] (imports.mapM importToWebString).run'
|
||||
let out := String.intercalate "\n" importWebString.toList
|
||||
IO.println (s!"Informal file made.")
|
||||
IO.FS.writeFile fileOut (informalFileHeader ++ out)
|
Loading…
Add table
Add a link
Reference in a new issue