refactor: Lint

This commit is contained in:
jstoobysmith 2025-02-06 13:31:59 +00:00
parent ee2134e448
commit e4c6da1cd6
12 changed files with 30 additions and 38 deletions

View file

@ -29,7 +29,7 @@ def staticWickTerm {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length
/-- The static Wick term for the empty contraction of the empty list is `1`. -/
@[simp]
lemma staticWickTerm_empty_nil :
lemma staticWickTerm_empty_nil :
staticWickTerm (empty (n := ([] : List 𝓕.FieldOp).length)) = 1 := by
rw [staticWickTerm, uncontractedListGet, nil_zero_uncontractedList]
simp [sign, empty, staticContract]
@ -51,7 +51,7 @@ lemma staticWickTerm_insert_zero_none (φ : 𝓕.FieldOp) (φs : List 𝓕.Field
simp only [staticContract_insert_none, insertAndContract_uncontractedList_none_zero,
Algebra.smul_mul_assoc]
/-- Let `φsΛ` be a Wick contraction for `φs = φ₀φ₁…φₙ`. Then`(φsΛ ↩Λ φ 0 (some k)).wickTerm`
/-- Let `φsΛ` be a Wick contraction for `φs = φ₀φ₁…φₙ`. Then`(φsΛ ↩Λ φ 0 (some k)).wickTerm`
is equal the product of
- the sign `𝓢(φ, φ₀…φᵢ₋₁) `
- the sign `φsΛ.sign`
@ -62,7 +62,7 @@ is equal the product of
The proof of this result relies on
- `staticContract_insert_some_of_lt` to rewrite static
contractions.
contractions.
- `normalOrder_uncontracted_some` to rewrite normal orderings.
- `sign_insert_some_zero` to rewrite signs.
-/
@ -105,14 +105,13 @@ lemma staticWickTerm_insert_zero_some (φ : 𝓕.FieldOp) (φs : List 𝓕.Field
rw [h1]
simp
/--
Let `φsΛ` be a Wick contraction for `φs = φ₀φ₁…φₙ`. Then
Let `φsΛ` be a Wick contraction for `φs = φ₀φ₁…φₙ`. Then
`φ * φsΛ.staticWickTerm = ∑ k, (φsΛ ↩Λ φ i k).wickTerm`
where the sum is over all `k` in `Option φsΛ.uncontracted` (so either `none` or `some k`).
The proof of proceeds as follows:
- `ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum` is used to expand 𝓝([φsΛ]ᵘᶜ)` as
- `ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum` is used to expand `φ 𝓝([φsΛ]ᵘᶜ)` as
a sum over `k` in `Option φsΛ.uncontracted` of terms involving `[φ, φs[k]]` etc.
- Then `staticWickTerm_insert_zero_none` and `staticWickTerm_insert_zero_some` are
used to equate terms.

View file

@ -10,8 +10,6 @@ import HepLean.PerturbationTheory.FieldOpAlgebra.StaticWickTerm
-/
namespace FieldSpecification
variable {𝓕 : FieldSpecification}
open FieldOpFreeAlgebra

View file

@ -19,11 +19,11 @@ namespace FieldOpAlgebra
variable {𝓕 : FieldSpecification}
/-- For a field specification, `𝓕`, given an algebra `A` and a function `f : 𝓕.CrAnFieldOp → A`
such that the lift of `f` to `FreeAlgebra.lift f : FreeAlgebra 𝓕.CrAnFieldOp → A` is
such that the lift of `f` to `FreeAlgebra.lift f : FreeAlgebra 𝓕.CrAnFieldOp → A` is
zero on the ideal defining `𝓕.FieldOpAlgebra`, the corresponding map `𝓕.FieldOpAlgebra → A`.
-/
def universalLiftMap {A : Type} [Semiring A] [Algebra A] (f : 𝓕.CrAnFieldOp → A)
(h1 : ∀ a ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet, FreeAlgebra.lift f a = 0):
(h1 : ∀ a ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet, FreeAlgebra.lift f a = 0) :
FieldOpAlgebra 𝓕 → A :=
Quotient.lift (FreeAlgebra.lift f) (by
intro a b h
@ -38,9 +38,8 @@ lemma universalLiftMap_ι {A : Type} [Semiring A] [Algebra A] (f : 𝓕.CrAn
(h1 : ∀ a ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet, FreeAlgebra.lift f a = 0) :
universalLiftMap f h1 (ι a) = FreeAlgebra.lift f a := by rfl
/-- For a field specification, `𝓕`, given an algebra `A` and a function `f : 𝓕.CrAnFieldOp → A`
such that the lift of `f` to `FreeAlgebra.lift f : FreeAlgebra 𝓕.CrAnFieldOp → A` is
such that the lift of `f` to `FreeAlgebra.lift f : FreeAlgebra 𝓕.CrAnFieldOp → A` is
zero on the ideal defining `𝓕.FieldOpAlgebra`, the corresponding algebra map
`𝓕.FieldOpAlgebra → A`.
-/

View file

@ -31,7 +31,7 @@ def wickTerm {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) :
/-- The Wick term for the empty contraction of the empty list is `1`. -/
@[simp]
lemma wickTerm_empty_nil :
lemma wickTerm_empty_nil :
wickTerm (empty (n := ([] : List 𝓕.FieldOp).length)) = 1 := by
rw [wickTerm]
simp [sign_empty]
@ -52,7 +52,7 @@ lemma wickTerm_insert_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
• (φsΛ.sign • φsΛ.timeContract * 𝓝(ofFieldOpList (φ :: [φsΛ]ᵘᶜ))) := by
rw [wickTerm]
by_cases hg : GradingCompliant φs φsΛ
· rw [normalOrder_uncontracted_none, sign_insert_none _ _ _ _ hg]
· rw [normalOrder_uncontracted_none, sign_insert_none _ _ _ _ hg]
simp only [Nat.succ_eq_add_one, timeContract_insert_none, instCommGroup.eq_1,
Algebra.mul_smul_comm, Algebra.smul_mul_assoc, smul_smul]
congr 1
@ -99,8 +99,8 @@ is equal the product of
The proof of this result relies on
- `timeContract_insert_some_of_not_lt`
and `timeContract_insert_some_of_lt` to rewrite time
contractions.
and `timeContract_insert_some_of_lt` to rewrite time
contractions.
- `normalOrder_uncontracted_some` to rewrite normal orderings.
- `sign_insert_some_of_not_lt` and `sign_insert_some_of_lt` to rewrite signs.
-/
@ -173,7 +173,7 @@ all files in `φ₀…φᵢ₋₁` have time strictly less then `φ`. Then
where the sum is over all `k` in `Option φsΛ.uncontracted` (so either `none` or `some k`).
The proof of proceeds as follows:
- `ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum` is used to expand 𝓝([φsΛ]ᵘᶜ)` as
- `ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum` is used to expand `φ 𝓝([φsΛ]ᵘᶜ)` as
a sum over `k` in `Option φsΛ.uncontracted` of terms involving `[φ, φs[k]]` etc.
- Then `wickTerm_insert_none` and `wickTerm_insert_some` are used to equate terms.
-/

View file

@ -40,7 +40,6 @@ abbrev FieldOpFreeAlgebra (𝓕 : FieldSpecification) : Type := FreeAlgebra
namespace FieldOpFreeAlgebra
remark naming_convention := "
For mathematicial objects defined in relation to `FieldOpFreeAlgebra` we will often postfix
their names with an `F` to indicate that they are related to the free algebra.
@ -62,7 +61,7 @@ The algebra `𝓕.FieldOpFreeAlgebra` satisfies the universal property that for
The unique `g` is given by `FreeAlgebra.lift f`.
-/
lemma universality {A : Type} [Semiring A] [Algebra A] (f : 𝓕.CrAnFieldOp → A) :
∃! g : FieldOpFreeAlgebra 𝓕 →ₐ[] A, g ∘ ofCrAnOpF = f := by
∃! g : FieldOpFreeAlgebra 𝓕 →ₐ[] A, g ∘ ofCrAnOpF = f := by
use FreeAlgebra.lift f
apply And.intro
· funext x

View file

@ -28,7 +28,7 @@ open FieldStatistic
which on the lists `φs` and `φs'` of `𝓕.CrAnFieldOp` gives
`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`. -/
noncomputable def superCommuteF : 𝓕.FieldOpFreeAlgebra →ₗ[] 𝓕.FieldOpFreeAlgebra →ₗ[]
𝓕.FieldOpFreeAlgebra :=
Basis.constr ofCrAnListFBasis fun φs =>

View file

@ -87,7 +87,6 @@ inductive FieldOp (𝓕 : FieldSpecification) where
| position : (Σ f, 𝓕.PositionLabel f) × SpaceTime → 𝓕.FieldOp
| outAsymp : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → ) → 𝓕.FieldOp
/-- The bool on `FieldOp` which is true only for position field operator. -/
def statesIsPosition : 𝓕.FieldOp → Bool
| FieldOp.position _ => true
@ -105,12 +104,12 @@ def fieldOpToField : 𝓕.FieldOp → 𝓕.Field
the field underlying `φ`.
The following notation is used in relation to `fieldOpStatistic`:
- For `φ` an element of `𝓕.FieldOp`, `𝓕 |>ₛ φ` is `fieldOpStatistic φ`.
- For `φs` a list of `𝓕.FieldOp`, `𝓕 |>ₛ φs` is the product of `fieldOpStatistic φ` over
- For `φ` an element of `𝓕.FieldOp`, `𝓕 |>ₛ φ` is `fieldOpStatistic φ`.
- For `φs` a list of `𝓕.FieldOp`, `𝓕 |>ₛ φs` is the product of `fieldOpStatistic φ` over
the list `φs`.
- For a function `f : Fin n → 𝓕.FieldOp` and a finset `a` of `Fin n`, `𝓕 |>ₛ ⟨f, a⟩` is the
- For a function `f : Fin n → 𝓕.FieldOp` and a finset `a` of `Fin n`, `𝓕 |>ₛ ⟨f, a⟩` is the
product of `fieldOpStatistic (f i)` for all `i ∈ a`. -/
def fieldOpStatistic : 𝓕.FieldOp → FieldStatistic := 𝓕.statistic ∘ 𝓕.fieldOpToField
def fieldOpStatistic : 𝓕.FieldOp → FieldStatistic := 𝓕.statistic ∘ 𝓕.fieldOpToField
@[inherit_doc fieldOpStatistic]
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => fieldOpStatistic 𝓕 φ

View file

@ -95,8 +95,8 @@ def crAnFieldOpToCreateAnnihilate : 𝓕.CrAnFieldOp → CreateAnnihilate
(or equivalently `𝓕.FieldOp`) underlying `φ`.
The following notation is used in relation to `crAnStatistics`:
- For `φ` an element of `𝓕.CrAnFieldOp`, `𝓕 |>ₛ φ` is `crAnStatistics φ`.
- For `φs` a list of `𝓕.CrAnFieldOp`, `𝓕 |>ₛ φs` is the product of `crAnStatistics φ` over
- For `φ` an element of `𝓕.CrAnFieldOp`, `𝓕 |>ₛ φ` is `crAnStatistics φ`.
- For `φs` a list of `𝓕.CrAnFieldOp`, `𝓕 |>ₛ φs` is the product of `crAnStatistics φ` over
the list `φs`.
-/
def crAnStatistics : 𝓕.CrAnFieldOp → FieldStatistic :=
@ -124,7 +124,7 @@ some notation within doc-strings and in code. The main notation used is:
`φᵃ` to indicate the annihilation part of an operator.
Some examples:
- `𝓢(φ, φs)` corresponds to `𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs)`
- `𝓢(φ, φs)` corresponds to `𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs)`
- `φ₀…φᵢ₋₁φᵢ₊₁…φₙ` corresponds to a (given) list `φs = φ₀…φₙ` with the element at the `i`th position
removed.
"

View file

@ -191,7 +191,7 @@ lemma timeOrderList_eq_maxTimeField_timeOrderList (φ : 𝓕.FieldOp) (φs : Lis
-/
/-- For a field specification `𝓕`, `𝓕.crAnTimeOrderRel` is a relation on
/-- For a field specification `𝓕`, `𝓕.crAnTimeOrderRel` is a relation on
`𝓕.CrAnFieldOp` representing time ordering.
It is defined as such that `𝓕.crAnTimeOrderRel φ₀ φ₁` is true if and only if one of the following
holds
@ -201,7 +201,7 @@ lemma timeOrderList_eq_maxTimeField_timeOrderList (φ : 𝓕.FieldOp) (φs : Lis
the `SpaceTime` point of `φ₀` has a time *greater* then or equal to that of `φ₁`.
Thus, colloquially `𝓕.crAnTimeOrderRel φ₀ φ₁` if `φ₀` has time *greater* then or equal to `φ₁`.
-/
-/
def crAnTimeOrderRel (a b : 𝓕.CrAnFieldOp) : Prop := 𝓕.timeOrderRel a.1 b.1
/-- The relation `crAnTimeOrderRel` is decidable, but not computablly so due to

View file

@ -32,7 +32,7 @@ open HepLean.Fin
elements and contracting `φ` optionally with `j`.
The notation `φsΛ ↩Λ φ i j` is used to denote `φsΛ.insertAndContract φ i j`. Thus,
`φsΛ ↩Λ φ i none` indicates the case when we insert `φ` into `φs` but do not contract it. -/
`φsΛ ↩Λ φ i none` indicates the case when we insert `φ` into `φs` but do not contract it. -/
def insertAndContract {φs : List 𝓕.FieldOp} (φ : 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
(i : Fin φs.length.succ) (j : Option φsΛ.uncontracted) :
WickContraction (φs.insertIdx i φ).length :=

View file

@ -856,7 +856,7 @@ lemma signInsertSome_mul_filter_contracted_of_not_lt (φ : 𝓕.FieldOp) (φs :
/--
For `k < i`, the sign of `φsΛ ↩Λ φ i (some k)` is equal to the product of
- the sign associated with moving `φ` through the `φsΛ`-uncontracted fields in `φ₀…φₖ`,
- the sign associated with moving `φ` through the `φsΛ`-uncontracted fields in `φ₀…φₖ`,
- the sign associated with moving `φ` through the fields in `φ₀…φᵢ₋₁`,
- the sign of `φsΛ`.
@ -877,10 +877,9 @@ lemma sign_insert_some_of_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
rw [mul_comm, ← mul_assoc]
simp
/--
For `i ≤ k`, the sign of `φsΛ ↩Λ φ i (some k)` is equal to the product of
- the sign associated with moving `φ` through the `φsΛ`-uncontracted fields in `φ₀…φₖ₋₁`,
- the sign associated with moving `φ` through the `φsΛ`-uncontracted fields in `φ₀…φₖ₋₁`,
- the sign associated with moving `φ` through the fields in `φ₀…φᵢ₋₁`,
- the sign of `φsΛ`.
@ -903,13 +902,12 @@ lemma sign_insert_some_of_not_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
lemma sign_insert_some_zero (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (k : φsΛ.uncontracted)
(hn : GradingCompliant φs φsΛ ∧ (𝓕|>ₛφ) = 𝓕|>ₛφs[k.1]):
(φsΛ ↩Λ φ 0 k).sign = 𝓢(𝓕|>ₛφ, 𝓕 |>ₛ ⟨φs.get, (φsΛ.uncontracted.filter (fun x => x < ↑k))⟩) *
(hn : GradingCompliant φs φsΛ ∧ (𝓕|>ₛφ) = 𝓕|>ₛφs[k.1]) :
(φsΛ ↩Λ φ 0 k).sign = 𝓢(𝓕|>ₛφ, 𝓕 |>ₛ ⟨φs.get, (φsΛ.uncontracted.filter (fun x => x < ↑k))⟩) *
φsΛ.sign := by
rw [sign_insert_some_of_not_lt]
· simp
· simp
· exact hn
end WickContraction

View file

@ -418,7 +418,7 @@ lemma join_sign_induction {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs
exact join_uncontractedListGet (singleton hij) φsucΛ'
/-- Let `φsΛ` be a grading compliant Wick contraction for `φs` and
`φsucΛ` a Wick contraction for `[φsΛ]ᵘᶜ`. Then `(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign`.
`φsucΛ` a Wick contraction for `[φsΛ]ᵘᶜ`. Then `(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign`.
This lemma manifests the fact that it does not matter which order contracted pairs are brought
together when defining the sign of a contraction. -/