PhysLean/scripts/MetaPrograms/notes.lean

248 lines
10 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 HepLean.Meta.Basic
2025-01-23 06:31:11 +00:00
import HepLean.Meta.Remark.Properties
2024-12-04 13:37:23 +00:00
import HepLean.Meta.Notes.ToHTML
2025-02-03 12:04:24 +00:00
import Mathlib.Lean.CoreM
2025-02-06 10:06:05 +00:00
import HepLean
2024-12-04 13:37:23 +00:00
/-!
# Extracting notes from Lean files
-/
open Lean System Meta HepLean
2025-02-06 10:06:05 +00:00
inductive NameStatus
| complete : NameStatus
| incomplete : NameStatus
instance : ToString NameStatus where
toString
| NameStatus.complete => "complete"
| NameStatus.incomplete => "incomplete"
2025-01-23 06:31:11 +00:00
inductive NotePart
| h1 : String → NotePart
| h2 : String → NotePart
| p : String → NotePart
2025-02-06 10:06:05 +00:00
| name : Name → NameStatus → NotePart
2025-02-03 15:59:25 +00:00
| warning : String → NotePart
2025-01-22 10:32:39 +00:00
2025-01-23 10:46:50 +00:00
structure DeclInfo where
line : Nat
fileName : Name
name : Name
2025-02-06 10:06:05 +00:00
status : NameStatus
2025-01-23 10:46:50 +00:00
declString : String
docString : String
2025-01-23 06:31:11 +00:00
2025-02-06 10:06:05 +00:00
def DeclInfo.ofName (n : Name) (ns : NameStatus): MetaM DeclInfo := do
2025-01-23 10:46:50 +00:00
let line ← Name.lineNumber n
let fileName ← Name.fileName n
let declString ← Name.getDeclString n
let docString ← Name.getDocString n
pure {
line := line,
fileName := fileName,
name := n,
2025-02-06 10:06:05 +00:00
status := ns,
2025-01-23 10:46:50 +00:00
declString := declString,
docString := docString}
2025-01-23 06:31:11 +00:00
2025-01-23 14:18:02 +00:00
def DeclInfo.toYML (d : DeclInfo) : MetaM String := do
2025-01-23 10:46:50 +00:00
let declStringIndent := d.declString.replace "\n" "\n "
2025-01-23 14:18:02 +00:00
let docStringIndent := d.docString.replace "\n" "\n "
2025-02-03 12:04:24 +00:00
let link := Name.toGitHubLink d.fileName d.line
2025-01-23 14:18:02 +00:00
return s!"
2025-01-23 10:46:50 +00:00
- type: name
name: {d.name}
line: {d.line}
fileName: {d.fileName}
2025-02-06 10:06:05 +00:00
status: \"{d.status}\"
2025-01-23 14:18:02 +00:00
link: \"{link}\"
docString: |
{docStringIndent}
2025-01-23 10:46:50 +00:00
declString: |
{declStringIndent}"
def NotePart.toYMLM : ((List String) × Nat × Nat) → NotePart → MetaM ((List String) × Nat × Nat)
| x, NotePart.h1 s =>
let newString := s!"
2025-01-23 06:31:11 +00:00
- type: h1
2025-01-23 10:46:50 +00:00
sectionNo: {x.2.1.succ}
2025-01-23 06:31:11 +00:00
content: \"{s}\""
2025-01-23 10:46:50 +00:00
return ⟨x.1 ++ [newString], ⟨Nat.succ x.2.1, 0⟩⟩
| x, NotePart.h2 s =>
let newString := s!"
2025-01-23 06:31:11 +00:00
- type: h2
2025-01-23 10:46:50 +00:00
sectionNo: \"{x.2.1}.{x.2.2.succ}\"
2025-01-23 06:31:11 +00:00
content: \"{s}\""
2025-01-23 10:46:50 +00:00
return ⟨x.1 ++ [newString], ⟨x.2.1, Nat.succ x.2.2⟩⟩
| x, NotePart.p s =>
let newString := s!"
2025-01-23 06:31:11 +00:00
- type: p
content: \"{s}\""
2025-01-23 10:46:50 +00:00
return ⟨x.1 ++ [newString], x.2⟩
2025-02-03 15:59:25 +00:00
| x, NotePart.warning s =>
let newString := s!"
- type: warning
content: \"{s}\""
return ⟨x.1 ++ [newString], x.2⟩
2025-02-06 10:06:05 +00:00
| x, NotePart.name n s => do
2025-01-23 06:31:11 +00:00
match (← RemarkInfo.IsRemark n) with
| true =>
let remarkInfo ← RemarkInfo.getRemarkInfo n
let content := remarkInfo.content
let contentIndent := content.replace "\n" "\n "
let shortName := remarkInfo.name.toString
2025-01-23 10:46:50 +00:00
let newString := s!"
2025-01-23 06:31:11 +00:00
- type: remark
name: \"{shortName}\"
2025-02-06 10:06:05 +00:00
status : \"{s}\"
2025-02-03 12:04:24 +00:00
link: \"{Name.toGitHubLink remarkInfo.fileName remarkInfo.line}\"
2025-01-23 06:31:11 +00:00
content: |
{contentIndent}"
2025-01-23 10:46:50 +00:00
return ⟨x.1 ++ [newString], x.2⟩
2025-01-23 06:31:11 +00:00
| false =>
2025-02-06 10:06:05 +00:00
let newString ← (← DeclInfo.ofName n s).toYML
2025-01-23 10:46:50 +00:00
return ⟨x.1 ++ [newString], x.2⟩
2025-01-23 06:31:11 +00:00
structure Note where
title : String
/-- The curators of the note are the people who put the note together.
This may not conicide with the original authors of the material. -/
curators : List String
parts : List NotePart
def Note.toYML : Note → MetaM String
2025-01-23 10:46:50 +00:00
| ⟨title, curators, parts⟩ => do
let parts ← parts.foldlM NotePart.toYMLM ([], ⟨0, 0⟩)
return s!"
2025-01-23 06:31:11 +00:00
title: \"{title}\"
2025-01-23 10:46:50 +00:00
curators: {String.intercalate "," curators}
2025-01-23 06:31:11 +00:00
parts:
2025-01-23 10:46:50 +00:00
{String.intercalate "\n" parts.1}"
2025-01-23 06:31:11 +00:00
def perturbationTheory : Note where
title := "Proof of Wick's theorem"
curators := ["Joseph Tooby-Smith"]
parts := [
.warning "This note is a work in progress and is not finished. Use with caution.
(5th Feb 2025)",
2025-01-23 10:46:50 +00:00
.h1 "Introduction",
2025-02-06 10:06:05 +00:00
.name `FieldSpecification.wicks_theorem_context .incomplete,
2025-01-23 10:46:50 +00:00
.p "In this note we walk through the important parts of the proof of Wick's theorem
for both fermions and bosons,
as it appears in HepLean. We start with some basic definitions.",
2025-02-03 15:59:25 +00:00
.h1 "Field operators",
2025-01-23 10:46:50 +00:00
.h2 "Field statistics",
2025-02-06 10:06:05 +00:00
.name ``FieldStatistic .complete,
.name ``FieldStatistic.instCommGroup .complete,
.name ``FieldStatistic.exchangeSign .complete,
2025-01-23 10:46:50 +00:00
.h2 "Field specifications",
2025-02-06 10:06:05 +00:00
.name ``FieldSpecification .complete,
2025-02-03 15:59:25 +00:00
.h2 "Field operators",
2025-02-06 10:06:05 +00:00
.name ``FieldSpecification.FieldOp .complete,
.name ``FieldSpecification.fieldOpStatistic .complete,
.name ``CreateAnnihilate .complete,
.name ``FieldSpecification.CrAnFieldOp .complete,
.name ``FieldSpecification.crAnStatistics .complete,
.name `FieldSpecification.notation_remark .complete,
2025-02-03 15:59:25 +00:00
.h2 "Field-operator free algebra",
2025-02-06 10:06:05 +00:00
.name ``FieldSpecification.FieldOpFreeAlgebra .incomplete,
.name `FieldSpecification.FieldOpFreeAlgebra.naming_convention .incomplete,
.name ``FieldSpecification.FieldOpFreeAlgebra.ofCrAnOpF .incomplete,
.name ``FieldSpecification.FieldOpFreeAlgebra.ofCrAnListF .incomplete,
.name ``FieldSpecification.FieldOpFreeAlgebra.ofFieldOpF .incomplete,
.name ``FieldSpecification.FieldOpFreeAlgebra.ofFieldOpListF .incomplete,
.name ``FieldSpecification.FieldOpFreeAlgebra.fieldOpFreeAlgebraGrade .incomplete,
.name ``FieldSpecification.FieldOpFreeAlgebra.superCommuteF .incomplete,
.name ``FieldSpecification.FieldOpFreeAlgebra.superCommuteF_ofCrAnListF_ofFieldOpListF_eq_sum .incomplete,
2025-02-03 15:59:25 +00:00
.h2 "Field-operator algebra",
2025-02-06 10:06:05 +00:00
.name ``FieldSpecification.FieldOpAlgebra .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.ofCrAnFieldOp .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.ofCrAnFieldOpList .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.ofFieldOp .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.ofCrAnFieldOpList .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.anPart .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.crPart .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.ofFieldOp_eq_crPart_add_anPart .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.fieldOpAlgebraGrade .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.superCommute .incomplete,
2025-02-03 15:59:25 +00:00
.h1 "Time ordering",
2025-02-06 10:06:05 +00:00
.name ``FieldSpecification.crAnTimeOrderRel .incomplete,
.name ``FieldSpecification.crAnTimeOrderSign .incomplete,
.name ``FieldSpecification.FieldOpFreeAlgebra.timeOrderF .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.timeOrder .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.timeOrder_eq_maxTimeField_mul_finset .incomplete,
2025-02-03 15:59:25 +00:00
.h1 "Normal ordering",
2025-02-06 10:06:05 +00:00
.name ``FieldSpecification.normalOrderRel .incomplete,
.name ``FieldSpecification.normalOrderSign .incomplete,
.name ``FieldSpecification.FieldOpFreeAlgebra.normalOrderF .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.normalOrder .incomplete,
.h2 "Some lemmas",
2025-02-06 10:06:05 +00:00
.name ``FieldSpecification.normalOrderSign_eraseIdx .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.ofCrAnFieldOp_superCommute_normalOrder_ofCrAnFieldOpList_sum .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum .incomplete,
2025-01-23 14:18:02 +00:00
.h1 "Wick Contractions",
2025-02-03 15:59:25 +00:00
.h2 "Definition",
2025-02-06 10:06:05 +00:00
.name ``WickContraction .incomplete,
.name ``WickContraction.GradingCompliant .incomplete,
.h2 "Aside: Cardinality",
2025-02-06 10:06:05 +00:00
.name ``WickContraction.card_eq_cardFun .incomplete,
2025-02-03 15:59:25 +00:00
.h2 "Constructors",
.p "There are a number of ways to construct a Wick contraction from
other Wick contractions or single contractions.",
2025-02-06 10:06:05 +00:00
.name ``WickContraction.insertAndContract .incomplete,
.name ``WickContraction.join .incomplete,
2025-02-03 15:59:25 +00:00
.h2 "Sign",
2025-02-06 10:06:05 +00:00
.name ``WickContraction.sign .incomplete,
.name ``WickContraction.join_sign .incomplete,
.name ``WickContraction.sign_insert_none .incomplete,
.name ``WickContraction.sign_insert_none_zero .incomplete,
.name ``WickContraction.sign_insert_some_of_not_lt .incomplete,
.name ``WickContraction.sign_insert_some_of_lt .incomplete,
.name ``WickContraction.sign_insert_some_zero .incomplete,
.h2 "Normal order",
2025-02-06 10:06:05 +00:00
.name ``FieldSpecification.FieldOpAlgebra.normalOrder_uncontracted_none .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.normalOrder_uncontracted_some .incomplete,
.h1 "Static contractions",
2025-02-06 10:06:05 +00:00
.name ``WickContraction.staticContract .incomplete,
.name ``WickContraction.staticContract_insert_some_of_lt .incomplete,
.name ``WickContraction.staticContract_insert_none .incomplete,
.h1 "Time contractions",
2025-02-06 10:06:05 +00:00
.name ``FieldSpecification.FieldOpAlgebra.timeContract .incomplete,
.name ``WickContraction.timeContract .incomplete,
.name ``WickContraction.timeContract_insert_none .incomplete,
.name ``WickContraction.timeContract_insert_some_of_not_lt .incomplete,
.name ``WickContraction.timeContract_insert_some_of_lt .incomplete,
2025-02-05 10:01:48 +00:00
.h1 "Wick terms",
2025-02-06 10:06:05 +00:00
.name ``WickContraction.wickTerm .incomplete,
.name ``WickContraction.wickTerm_empty_nil .incomplete,
.name ``WickContraction.wickTerm_insert_none .incomplete,
.name ``WickContraction.wickTerm_insert_some .incomplete,
.name ``WickContraction.mul_wickTerm_eq_sum .incomplete,
.h1 "Static wick terms",
2025-02-06 10:06:05 +00:00
.name ``WickContraction.staticWickTerm .incomplete,
.name ``WickContraction.staticWickTerm_empty_nil .incomplete,
.name ``WickContraction.staticWickTerm_insert_zero_none .incomplete,
.name ``WickContraction.staticWickTerm_insert_zero_some .incomplete,
.name ``WickContraction.mul_staticWickTerm_eq_sum .incomplete,
2025-02-04 14:56:38 +00:00
.h1 "The three Wick's theorems",
2025-02-06 10:06:05 +00:00
.name ``FieldSpecification.wicks_theorem .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.static_wick_theorem .incomplete,
.name ``FieldSpecification.FieldOpAlgebra.wicks_theorem_normal_order .incomplete
2025-02-04 14:56:38 +00:00
]
2024-12-04 13:37:23 +00:00
2024-12-05 06:23:10 +00:00
unsafe def main (_ : List String) : IO UInt32 := do
2024-12-04 13:37:23 +00:00
initSearchPath (← findSysroot)
2025-01-23 06:31:11 +00:00
let ymlString ← CoreM.withImportModules #[`HepLean] (perturbationTheory.toYML).run'
let fileOut : System.FilePath := {toString := "./docs/_data/perturbationTheory.yml"}
IO.println (s!"YML file made.")
IO.FS.writeFile fileOut ymlString
2024-12-04 13:37:23 +00:00
pure 0