refactor: Slight adjustments to doc-strings

This commit is contained in:
jstoobysmith 2025-02-12 06:14:11 +00:00
parent 3fc88eaca8
commit ea29b15e4a
15 changed files with 64 additions and 30 deletions

View file

@ -353,7 +353,7 @@ lemma orderedInsert_eraseIdx_orderedInsertPos_le {I : Type} (le1 : I → I → P
omega omega
/-- The equivalence between `Fin (r0 :: r).length` and `Fin (List.orderedInsert le1 r0 r).length` /-- The equivalence between `Fin (r0 :: r).length` and `Fin (List.orderedInsert le1 r0 r).length`
according to where the elements map. I.e. `0` is taken to `orderedInsertPos le1 r r0`. -/ according to where the elements map, i.e. `0` is taken to `orderedInsertPos le1 r r0`. -/
def orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I) (r0 : I) : def orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I) (r0 : I) :
Fin (r0 :: r).length ≃ Fin (List.orderedInsert le1 r0 r).length := by Fin (r0 :: r).length ≃ Fin (List.orderedInsert le1 r0 r).length := by
let e2 : Fin (List.orderedInsert le1 r0 r).length ≃ Fin (r0 :: r).length := let e2 : Fin (List.orderedInsert le1 r0 r).length ≃ Fin (r0 :: r).length :=

View file

@ -119,7 +119,7 @@ lemma crPart_mul_normalOrder (φ : 𝓕.FieldOp) (a : 𝓕.FieldOpAlgebra) :
-/ -/
/-- For a field specification `𝓕`, and `a` and `b` in `𝓕.FieldOpAlgebra` the normal ordering /-- For a field specification `𝓕`, and `a` and `b` in `𝓕.FieldOpAlgebra` the normal ordering
of the super commutator of `a` and `b` vanishes. I.e. `𝓝([a,b]ₛ) = 0`. -/ of the super commutator of `a` and `b` vanishes, i.e. `𝓝([a,b]ₛ) = 0`. -/
@[simp] @[simp]
lemma normalOrder_superCommute_eq_zero (a b : 𝓕.FieldOpAlgebra) : lemma normalOrder_superCommute_eq_zero (a b : 𝓕.FieldOpAlgebra) :
𝓝([a, b]ₛ) = 0 := by 𝓝([a, b]ₛ) = 0 := by
@ -346,15 +346,24 @@ noncomputable def contractStateAtIndex (φ : 𝓕.FieldOp) (φs : List 𝓕.Fiel
/-- /--
For a field specification `𝓕`, a `φ` in `𝓕.FieldOp` and a list `φs` of `𝓕.FieldOp` For a field specification `𝓕`, a `φ` in `𝓕.FieldOp` and a list `φs` of `𝓕.FieldOp`
the following relation holds in the algebra `𝓕.FieldOpAlgebra`, then `φ * 𝓝(φ₀φ₁…φₙ)` is equal to
`φ * 𝓝(φ₀φ₁…φₙ) = 𝓝(φφ₀φ₁…φₙ) + ∑ i, (𝓢(φ,φ₀φ₁…φᵢ₋₁) • [anPart φ, φᵢ]ₛ) * 𝓝(φ₀φ₁…φᵢ₋₁φᵢ₊₁…φₙ)`.
`𝓝(φφ₀φ₁…φₙ) + ∑ i, (𝓢(φ,φ₀φ₁…φᵢ₋₁) • [anPart φ, φᵢ]ₛ) * 𝓝(φ₀…φᵢ₋₁φᵢ₊₁…φₙ)`.
The proof of ultimately goes as follows: The proof of ultimately goes as follows:
- `ofFieldOp_eq_crPart_add_anPart` is used to split `φ` into its creation and annihilation parts. - `ofFieldOp_eq_crPart_add_anPart` is used to split `φ` into its creation and annihilation parts.
- The fact that `crPart φ * 𝓝(φ₀φ₁…φₙ) = 𝓝(crPart φ * φ₀φ₁…φₙ)` is used. - The following relation is then used
- The fact that `anPart φ * 𝓝(φ₀φ₁…φₙ)` is
`𝓢(φ, φ₀φ₁…φₙ) 𝓝(φ₀φ₁…φₙ) * anPart φ + [anPart φ, 𝓝(φ₀φ₁…φₙ)]` is used `crPart φ * 𝓝(φ₀φ₁…φₙ) = 𝓝(crPart φ * φ₀φ₁…φₙ)`.
- The fact that `𝓢(φ, φ₀φ₁…φₙ) 𝓝(φ₀φ₁…φₙ) * anPart φ = 𝓝(anPart φ * φ₀φ₁…φₙ)`
- It used that `anPart φ * 𝓝(φ₀φ₁…φₙ)` is equal to
`𝓢(φ, φ₀φ₁…φₙ) 𝓝(φ₀φ₁…φₙ) * anPart φ + [anPart φ, 𝓝(φ₀φ₁…φₙ)]`
- Then it is used that
`𝓢(φ, φ₀φ₁…φₙ) 𝓝(φ₀φ₁…φₙ) * anPart φ = 𝓝(anPart φ * φ₀φ₁…φₙ)`
- The result `ofCrAnOp_superCommute_normalOrder_ofCrAnList_sum` is used - The result `ofCrAnOp_superCommute_normalOrder_ofCrAnList_sum` is used
to expand `[anPart φ, 𝓝(φ₀φ₁…φₙ)]` as a sum. to expand `[anPart φ, 𝓝(φ₀φ₁…φₙ)]` as a sum.
-/ -/

View file

@ -41,9 +41,9 @@ lemma staticWickTerm_empty_nil :
/-- /--
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, and an element `φ` of For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, and an element `φ` of
`𝓕.FieldOp`, the following relation holds `𝓕.FieldOp`, then `(φsΛ ↩Λ φ 0 none).staticWickTerm` is equal to
`(φsΛ ↩Λ φ 0 none).staticWickTerm = φsΛ.sign • φsΛ.staticWickTerm * 𝓝(φ :: [φsΛ]ᵘᶜ)` `φsΛ.sign • φsΛ.staticWickTerm * 𝓝(φ :: [φsΛ]ᵘᶜ)`
The proof of this result relies on The proof of this result relies on
- `staticContract_insert_none` to rewrite the static contract. - `staticContract_insert_none` to rewrite the static contract.

View file

@ -89,7 +89,9 @@ lemma superCommuteRight_eq_of_equiv (a1 a2 : 𝓕.FieldOpFreeAlgebra) (h : a1
simp simp
/-- For a field specification `𝓕`, `superCommute` is the linear map /-- For a field specification `𝓕`, `superCommute` is the linear map
`FieldOpAlgebra 𝓕 →ₗ[] FieldOpAlgebra 𝓕 →ₗ[] FieldOpAlgebra 𝓕` `FieldOpAlgebra 𝓕 →ₗ[] FieldOpAlgebra 𝓕 →ₗ[] FieldOpAlgebra 𝓕`
defined as the decent of `ι ∘ superCommuteF` in both arguments. defined as the decent of `ι ∘ superCommuteF` in both arguments.
In particular for `φs` and `φs'` lists of `𝓕.CrAnFieldOp` in `FieldOpAlgebra 𝓕` the following In particular for `φs` and `φs'` lists of `𝓕.CrAnFieldOp` in `FieldOpAlgebra 𝓕` the following
relation holds: relation holds:

View file

@ -43,9 +43,9 @@ lemma wickTerm_empty_nil :
/-- /--
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
`𝓕.FieldOp`, and `i ≤ φs.length` the following relation holds `𝓕.FieldOp`, and `i ≤ φs.length`, then `(φsΛ ↩Λ φ i none).wickTerm` is equal to
`(φsΛ ↩Λ φ i none).wickTerm = 𝓢(φ, φ₀…φᵢ₋₁) φsΛ.sign • φsΛ.timeContract * 𝓝(φ :: [φsΛ]ᵘᶜ)` `𝓢(φ, φ₀…φᵢ₋₁) φsΛ.sign • φsΛ.timeContract * 𝓝(φ :: [φsΛ]ᵘᶜ)`
The proof of this result relies on The proof of this result relies on
- `normalOrder_uncontracted_none` to rewrite normal orderings. - `normalOrder_uncontracted_none` to rewrite normal orderings.

View file

@ -23,7 +23,7 @@ open EqTimeOnly
/-- /--
For a list `φs` of `𝓕.FieldOp`, then For a list `φs` of `𝓕.FieldOp`, then
`𝓣(φs) = ∑ φsΛ, φsΛ.1.wickTerm • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))` `𝓣(φs) = ∑ φsΛ, φsΛ.sign • φsΛ.timeContract * 𝓣(𝓝([φsΛ]ᵘᶜ))`
where the sum is over all Wick contraction `φsΛ` which only have equal time contractions. where the sum is over all Wick contraction `φsΛ` which only have equal time contractions.
@ -92,7 +92,7 @@ lemma timeOrder_ofFieldOpList_eq_eqTimeOnly_empty (φs : List 𝓕.FieldOp) :
/-- /--
For a list `φs` of `𝓕.FieldOp`, then For a list `φs` of `𝓕.FieldOp`, then
`𝓣(𝓝(φs)) = 𝓣(φs) - ∑ φsΛ, φsΛ.1.wickTerm • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))` `𝓣(𝓝(φs)) = 𝓣(φs) - ∑ φsΛ, φsΛ.sign • φsΛ.timeContract.1 * 𝓣(𝓝([φsΛ]ᵘᶜ))`
where the sum is over all *non-empty* Wick contraction `φsΛ` which only where the sum is over all *non-empty* Wick contraction `φsΛ` which only
have equal time contractions. have equal time contractions.
@ -222,7 +222,7 @@ lemma wicks_theorem_normal_order_empty : 𝓣(𝓝(ofFieldOpList [])) =
where the sum is over all Wick contraction `φsΛ` in which no two contracted elements where the sum is over all Wick contraction `φsΛ` in which no two contracted elements
have the same time. have the same time.
The proof of proceeds by induction on `φs`, with the base case `[]` holding by following The proof proceeds by induction on `φs`, with the base case `[]` holding by following
through definitions. and the inductive case holding as a result of through definitions. and the inductive case holding as a result of
- `timeOrder_haveEqTime_split` - `timeOrder_haveEqTime_split`
- `normalOrder_timeOrder_ofFieldOpList_eq_eqTimeOnly_empty` - `normalOrder_timeOrder_ofFieldOpList_eq_eqTimeOnly_empty`

View file

@ -28,9 +28,13 @@ namespace FieldOpFreeAlgebra
noncomputable section noncomputable section
/-- For a field specification `𝓕`, `normalOrderF` is the linear map /-- For a field specification `𝓕`, `normalOrderF` is the linear map
`FieldOpFreeAlgebra 𝓕 →ₗ[] FieldOpFreeAlgebra 𝓕` `FieldOpFreeAlgebra 𝓕 →ₗ[] FieldOpFreeAlgebra 𝓕`
defined by its action on the basis `ofCrAnListF φs`, taking `ofCrAnListF φs` to defined by its action on the basis `ofCrAnListF φs`, taking `ofCrAnListF φs` to
`normalOrderSign φs • ofCrAnListF (normalOrderList φs)`. `normalOrderSign φs • ofCrAnListF (normalOrderList φs)`.
That is, `normalOrderF` normal-orders the field operators and multiplies by the sign of the That is, `normalOrderF` normal-orders the field operators and multiplies by the sign of the
normal order. normal order.

View file

@ -26,6 +26,7 @@ open FieldStatistic
/-- For a field specification `𝓕`, the super commutator `superCommuteF` is defined as the linear /-- For a field specification `𝓕`, the super commutator `superCommuteF` is defined as the linear
map `𝓕.FieldOpFreeAlgebra →ₗ[] 𝓕.FieldOpFreeAlgebra →ₗ[] 𝓕.FieldOpFreeAlgebra` map `𝓕.FieldOpFreeAlgebra →ₗ[] 𝓕.FieldOpFreeAlgebra →ₗ[] 𝓕.FieldOpFreeAlgebra`
which on the lists `φs` and `φs'` of `𝓕.CrAnFieldOp` gives which on the lists `φs` and `φs'` of `𝓕.CrAnFieldOp` gives
`superCommuteF φs φs' = φs * φs' - 𝓢(φs, φs') • φs' * φs`. `superCommuteF φs φs' = φs * φs' - 𝓢(φs, φs') • φs' * φs`.
The notation `[a, b]ₛca` can be used for `superCommuteF a b`. -/ The notation `[a, b]ₛca` can be used for `superCommuteF a b`. -/

View file

@ -27,9 +27,14 @@ open HepLean.List
-/ -/
/-- For a field specification `𝓕`, `timeOrderF` is the linear map /-- For a field specification `𝓕`, `timeOrderF` is the linear map
`FieldOpFreeAlgebra 𝓕 →ₗ[] FieldOpFreeAlgebra 𝓕` `FieldOpFreeAlgebra 𝓕 →ₗ[] FieldOpFreeAlgebra 𝓕`
defined by its action on the basis `ofCrAnListF φs`, taking defined by its action on the basis `ofCrAnListF φs`, taking
`ofCrAnListF φs` to `crAnTimeOrderSign φs • ofCrAnListF (crAnTimeOrderList φs)`. `ofCrAnListF φs` to
`crAnTimeOrderSign φs • ofCrAnListF (crAnTimeOrderList φs)`.
That is, `timeOrderF` time-orders the field operators and multiplies by the sign of the That is, `timeOrderF` time-orders the field operators and multiplies by the sign of the
time order. time order.

View file

@ -85,12 +85,14 @@ As some intuition, if `f` corresponds to a Weyl-fermion field, then
- For `inAsymp f e p`, `e` would correspond to a spin `s`, and `inAsymp f e p` would, once - For `inAsymp f e p`, `e` would correspond to a spin `s`, and `inAsymp f e p` would, once
represented in the operator algebra, represented in the operator algebra,
be proportional to the creation operator `a(p, s)`. be proportional to the creation operator `a(p, s)`.
- `position f e x`, `e` would correspond to a Lorentz index `α`, and `position f e x` would, - `position f e x`, `e` would correspond to a Lorentz index `a`, and `position f e x` would,
once represented in the operator algebra, be proportional to the operator once represented in the operator algebra, be proportional to the operator
`∑ s, ∫ d^3p/(…) (x_α(p,s) a(p, s) e^{-i p x} + y_α(p,s) a^†(p, s) e^{-i p x})`.
`∑ s, ∫ d³p/(…) (xₐ(p,s) a(p, s) e ^ (-i p x) + yₐ(p,s) a†(p, s) e ^ (-i p x))`.
- `outAsymp f e p`, `e` would correspond to a spin `s`, and `outAsymp f e p` would, - `outAsymp f e p`, `e` would correspond to a spin `s`, and `outAsymp f e p` would,
once represented in the operator algebra, be proportional to the once represented in the operator algebra, be proportional to the
annihilation operator `a^†(p, s)`. annihilation operator `a†(p, s)`.
This type contains all operators which are related to a field. This type contains all operators which are related to a field.
-/ -/

View file

@ -81,12 +81,14 @@ As some intuition, if `f` corresponds to a Weyl-fermion field, it would contribu
the following elements to `𝓕.CrAnFieldOp` the following elements to `𝓕.CrAnFieldOp`
- an element corresponding to incoming asymptotic operators for each spin `s`: `a(p, s)`. - an element corresponding to incoming asymptotic operators for each spin `s`: `a(p, s)`.
- an element corresponding to the creation parts of position operators for each each Lorentz - an element corresponding to the creation parts of position operators for each each Lorentz
index `α`: index `a`:
`∑ s, ∫ d^3p/(…) (x_α(p,s) a(p, s) e^{-i p x})`.
`∑ s, ∫ d³p/(…) (xₐ (p,s) a(p, s) e ^ (-i p x))`.
- an element corresponding to annihilation parts of position operator, - an element corresponding to annihilation parts of position operator,
for each each Lorentz index `α`: for each each Lorentz index `a`:
`∑ s, ∫ d^3p/(…) (y_α(p,s) a^†(p, s) e^{-i p x})`.
- an element corresponding to outgoing asymptotic operators for each spin `s`: `a^†(p, s)`. `∑ s, ∫ d³p/(…) (yₐ(p,s) a†(p, s) e ^ (-i p x))`.
- an element corresponding to outgoing asymptotic operators for each spin `s`: `a†(p, s)`.
-/ -/
def CrAnFieldOp : Type := Σ (s : 𝓕.FieldOp), 𝓕.fieldOpToCrAnType s def CrAnFieldOp : Type := Σ (s : 𝓕.FieldOp), 𝓕.fieldOpToCrAnType s

View file

@ -25,7 +25,9 @@ namespace FieldStatistic
variable {𝓕 : Type} variable {𝓕 : Type}
/-- The exchange sign, `exchangeSign`, is defined as the group homomorphism /-- The exchange sign, `exchangeSign`, is defined as the group homomorphism
`FieldStatistic →* FieldStatistic →* `, `FieldStatistic →* FieldStatistic →* `,
for which `exchangeSign a b` is `-1` if both `a` and `b` are `fermionic` and `1` otherwise. for which `exchangeSign a b` is `-1` if both `a` and `b` are `fermionic` and `1` otherwise.
The exchange sign is the sign one picks up on exchanging an operator or field `φ₁` of statistic The exchange sign is the sign one picks up on exchanging an operator or field `φ₁` of statistic
`a` with an operator or field `φ₂` of statistic `b`, i.e. `φ₁φ₂ → φ₂φ₁`. `a` with an operator or field `φ₂` of statistic `b`, i.e. `φ₁φ₂ → φ₂φ₁`.

View file

@ -291,9 +291,15 @@ lemma insert_fin_eq_self (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp}
Wick contractions of `φs` with `φ` inserted at `i` is equal to the sum over Wick contractions Wick contractions of `φs` with `φ` inserted at `i` is equal to the sum over Wick contractions
`φsΛ` of just `φs` and the sum over optional uncontracted elements of the `φsΛ`. `φsΛ` of just `φs` and the sum over optional uncontracted elements of the `φsΛ`.
In other words, `∑ (φsΛ : WickContraction (φs.insertIdx i φ).length), f φsΛ` is equal to In other words,
`∑ (φsΛ : WickContraction φs.length), ∑ (k : Option φsΛ.uncontracted), f (φsΛ ↩Λ φ i k) `.
where `(φs.insertIdx i φ)` is `φs` with `φ` inserted at position `i`. -/ `∑ (φsΛ : WickContraction (φs.insertIdx i φ).length), f φsΛ`
where `(φs.insertIdx i φ)` is `φs` with `φ` inserted at position `i`. is equal to
`∑ (φsΛ : WickContraction φs.length), ∑ k, f (φsΛ ↩Λ φ i k) `.
where the sum over `k` is over all `k` in `Option φsΛ.uncontracted`. -/
lemma insertLift_sum (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp} lemma insertLift_sum (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp}
(i : Fin φs.length.succ) [AddCommMonoid M] (f : WickContraction (φs.insertIdx i φ).length → M) : (i : Fin φs.length.succ) [AddCommMonoid M] (f : WickContraction (φs.insertIdx i φ).length → M) :
∑ c, f c = ∑ c, f c =

View file

@ -22,7 +22,8 @@ open FieldStatistic
/-- Given a Wick contraction `c : WickContraction n` and `i1 i2 : Fin n` the finite set /-- Given a Wick contraction `c : WickContraction n` and `i1 i2 : Fin n` the finite set
of elements of `Fin n` between `i1` and `i2` which are either uncontracted of elements of `Fin n` between `i1` and `i2` which are either uncontracted
or are contracted but are contracted with an element occurring after `i1`. or are contracted but are contracted with an element occurring after `i1`.
I.e. the elements of `Fin n` between `i1` and `i2` which are not contracted with before `i1`. In other words, the elements of `Fin n` between `i1` and `i2` which are not
contracted with before `i1`.
One should assume `i1 < i2` otherwise this finite set is empty. -/ One should assume `i1 < i2` otherwise this finite set is empty. -/
def signFinset (c : WickContraction n) (i1 i2 : Fin n) : Finset (Fin n) := def signFinset (c : WickContraction n) (i1 i2 : Fin n) : Finset (Fin n) :=
Finset.univ.filter (fun i => i1 < i ∧ i < i2 ∧ Finset.univ.filter (fun i => i1 < i ∧ i < i2 ∧

View file

@ -39,8 +39,8 @@ import HepLean.Lorentz.ComplexTensor.Basic
## Comments ## Comments
- In all of theses expressions `μ`, `ν` etc are free. It does not matter what they are called, - In all of theses expressions `μ`, `ν` etc are free. It does not matter what they are called,
Lean will elaborate them in the same way. I.e. `{T | μ ν ⊗ T3 | μ ν }ᵀ` is exactly the same Lean will elaborate them in the same way. In other words, `{T | μ ν ⊗ T3 | μ ν }ᵀ` is exactly
to Lean as `{T | α β ⊗ T3 | α β }ᵀ`. the same to Lean as `{T | α β ⊗ T3 | α β }ᵀ`.
- Note that compared to ordinary index notation, we do not rise or lower the indices. - Note that compared to ordinary index notation, we do not rise or lower the indices.
This is for two reasons: 1) It is difficult to make this general for all tensor species, This is for two reasons: 1) It is difficult to make this general for all tensor species,
2) It is a redundancy in ordinary index notation, since the tensor `T` itself already tells you 2) It is a redundancy in ordinary index notation, since the tensor `T` itself already tells you