refactor: Rename CrAnAlgebra
This commit is contained in:
parent
9a5676e134
commit
b0735a1e13
16 changed files with 214 additions and 214 deletions
|
@ -0,0 +1,225 @@
|
|||
/-
|
||||
Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.PerturbationTheory.FieldSpecification.CrAnStates
|
||||
import HepLean.PerturbationTheory.FieldSpecification.CrAnSection
|
||||
/-!
|
||||
|
||||
# Creation and annihlation free-algebra
|
||||
|
||||
This module defines the creation and annihilation algebra for a field structure.
|
||||
|
||||
The creation and annihilation algebra extends from the state algebra by adding information about
|
||||
whether a state is a creation or annihilation operator.
|
||||
|
||||
The algebra is spanned by lists of creation/annihilation states.
|
||||
|
||||
The main structures defined in this module are:
|
||||
|
||||
* `FieldOpFreeAlgebra` - The creation and annihilation algebra
|
||||
* `ofCrAnState` - Maps a creation/annihilation state to the algebra
|
||||
* `ofCrAnList` - Maps a list of creation/annihilation states to the algebra
|
||||
* `ofState` - Maps a state to a sum of creation and annihilation operators
|
||||
* `crPartF` - The creation part of a state in the algebra
|
||||
* `anPartF` - The annihilation part of a state in the algebra
|
||||
* `superCommuteF` - The super commutator on the algebra
|
||||
|
||||
The key lemmas show how these operators interact, particularly focusing on the
|
||||
super commutation relations between creation and annihilation operators.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldSpecification
|
||||
variable {𝓕 : FieldSpecification}
|
||||
|
||||
/-- The creation and annihlation free-algebra.
|
||||
The free algebra generated by `CrAnStates`,
|
||||
that is a position based states or assymptotic states with a specification of
|
||||
whether the state is a creation or annihlation state.
|
||||
As a module `FieldOpFreeAlgebra` is spanned by lists of `CrAnStates`. -/
|
||||
abbrev FieldOpFreeAlgebra (𝓕 : FieldSpecification) : Type := FreeAlgebra ℂ 𝓕.CrAnStates
|
||||
|
||||
namespace FieldOpFreeAlgebra
|
||||
|
||||
/-- Maps a creation and annihlation state to the creation and annihlation free-algebra. -/
|
||||
def ofCrAnState (φ : 𝓕.CrAnStates) : FieldOpFreeAlgebra 𝓕 :=
|
||||
FreeAlgebra.ι ℂ φ
|
||||
|
||||
/-- Maps a list creation and annihlation state to the creation and annihlation free-algebra
|
||||
by taking their product. -/
|
||||
def ofCrAnList (φs : List 𝓕.CrAnStates) : FieldOpFreeAlgebra 𝓕 := (List.map ofCrAnState φs).prod
|
||||
|
||||
@[simp]
|
||||
lemma ofCrAnList_nil : ofCrAnList ([] : List 𝓕.CrAnStates) = 1 := rfl
|
||||
|
||||
lemma ofCrAnList_cons (φ : 𝓕.CrAnStates) (φs : List 𝓕.CrAnStates) :
|
||||
ofCrAnList (φ :: φs) = ofCrAnState φ * ofCrAnList φs := rfl
|
||||
|
||||
lemma ofCrAnList_append (φs φs' : List 𝓕.CrAnStates) :
|
||||
ofCrAnList (φs ++ φs') = ofCrAnList φs * ofCrAnList φs' := by
|
||||
simp [ofCrAnList, List.map_append]
|
||||
|
||||
lemma ofCrAnList_singleton (φ : 𝓕.CrAnStates) :
|
||||
ofCrAnList [φ] = ofCrAnState φ := by simp [ofCrAnList]
|
||||
|
||||
/-- Maps a state to the sum of creation and annihilation operators in
|
||||
creation and annihilation free-algebra. -/
|
||||
def ofState (φ : 𝓕.States) : FieldOpFreeAlgebra 𝓕 :=
|
||||
∑ (i : 𝓕.statesToCrAnType φ), ofCrAnState ⟨φ, i⟩
|
||||
|
||||
/-- Maps a list of states to the creation and annihilation free-algebra by taking
|
||||
the product of their sums of creation and annihlation operators.
|
||||
Roughly `[φ1, φ2]` gets sent to `(φ1ᶜ+ φ1ᵃ) * (φ2ᶜ+ φ2ᵃ)` etc. -/
|
||||
def ofStateList (φs : List 𝓕.States) : FieldOpFreeAlgebra 𝓕 := (List.map ofState φs).prod
|
||||
|
||||
/-- Coercion from `List 𝓕.States` to `FieldOpFreeAlgebra 𝓕` through `ofStateList`. -/
|
||||
instance : Coe (List 𝓕.States) (FieldOpFreeAlgebra 𝓕) := ⟨ofStateList⟩
|
||||
|
||||
@[simp]
|
||||
lemma ofStateList_nil : ofStateList ([] : List 𝓕.States) = 1 := rfl
|
||||
|
||||
lemma ofStateList_cons (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
ofStateList (φ :: φs) = ofState φ * ofStateList φs := rfl
|
||||
|
||||
lemma ofStateList_singleton (φ : 𝓕.States) :
|
||||
ofStateList [φ] = ofState φ := by simp [ofStateList]
|
||||
|
||||
lemma ofStateList_append (φs φs' : List 𝓕.States) :
|
||||
ofStateList (φs ++ φs') = ofStateList φs * ofStateList φs' := by
|
||||
dsimp only [ofStateList]
|
||||
rw [List.map_append, List.prod_append]
|
||||
|
||||
lemma ofStateList_sum (φs : List 𝓕.States) :
|
||||
ofStateList φs = ∑ (s : CrAnSection φs), ofCrAnList s.1 := by
|
||||
induction φs with
|
||||
| nil => simp
|
||||
| cons φ φs ih =>
|
||||
rw [CrAnSection.sum_cons]
|
||||
dsimp only [CrAnSection.cons, ofCrAnList_cons]
|
||||
conv_rhs =>
|
||||
enter [2, x]
|
||||
rw [← Finset.mul_sum]
|
||||
rw [← Finset.sum_mul, ofStateList_cons, ← ih]
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Creation and annihilation parts of a state
|
||||
|
||||
-/
|
||||
|
||||
/-- The algebra map taking an element of the free-state algbra to
|
||||
the part of it in the creation and annihlation free algebra
|
||||
spanned by creation operators. -/
|
||||
def crPartF : 𝓕.States → 𝓕.FieldOpFreeAlgebra := fun φ =>
|
||||
match φ with
|
||||
| States.inAsymp φ => ofCrAnState ⟨States.inAsymp φ, ()⟩
|
||||
| States.position φ => ofCrAnState ⟨States.position φ, CreateAnnihilate.create⟩
|
||||
| States.outAsymp _ => 0
|
||||
|
||||
@[simp]
|
||||
lemma crPartF_negAsymp (φ : 𝓕.IncomingAsymptotic) :
|
||||
crPartF (States.inAsymp φ) = ofCrAnState ⟨States.inAsymp φ, ()⟩ := by
|
||||
simp [crPartF]
|
||||
|
||||
@[simp]
|
||||
lemma crPartF_position (φ : 𝓕.PositionStates) :
|
||||
crPartF (States.position φ) =
|
||||
ofCrAnState ⟨States.position φ, CreateAnnihilate.create⟩ := by
|
||||
simp [crPartF]
|
||||
|
||||
@[simp]
|
||||
lemma crPartF_posAsymp (φ : 𝓕.OutgoingAsymptotic) :
|
||||
crPartF (States.outAsymp φ) = 0 := by
|
||||
simp [crPartF]
|
||||
|
||||
/-- The algebra map taking an element of the free-state algbra to
|
||||
the part of it in the creation and annihilation free algebra
|
||||
spanned by annihilation operators. -/
|
||||
def anPartF : 𝓕.States → 𝓕.FieldOpFreeAlgebra := fun φ =>
|
||||
match φ with
|
||||
| States.inAsymp _ => 0
|
||||
| States.position φ => ofCrAnState ⟨States.position φ, CreateAnnihilate.annihilate⟩
|
||||
| States.outAsymp φ => ofCrAnState ⟨States.outAsymp φ, ()⟩
|
||||
|
||||
@[simp]
|
||||
lemma anPartF_negAsymp (φ : 𝓕.IncomingAsymptotic) :
|
||||
anPartF (States.inAsymp φ) = 0 := by
|
||||
simp [anPartF]
|
||||
|
||||
@[simp]
|
||||
lemma anPartF_position (φ : 𝓕.PositionStates) :
|
||||
anPartF (States.position φ) =
|
||||
ofCrAnState ⟨States.position φ, CreateAnnihilate.annihilate⟩ := by
|
||||
simp [anPartF]
|
||||
|
||||
@[simp]
|
||||
lemma anPartF_posAsymp (φ : 𝓕.OutgoingAsymptotic) :
|
||||
anPartF (States.outAsymp φ) = ofCrAnState ⟨States.outAsymp φ, ()⟩ := by
|
||||
simp [anPartF]
|
||||
|
||||
lemma ofState_eq_crPartF_add_anPartF (φ : 𝓕.States) :
|
||||
ofState φ = crPartF φ + anPartF φ := by
|
||||
rw [ofState]
|
||||
cases φ with
|
||||
| inAsymp φ => simp [statesToCrAnType]
|
||||
| position φ => simp [statesToCrAnType, CreateAnnihilate.sum_eq]
|
||||
| outAsymp φ => simp [statesToCrAnType]
|
||||
|
||||
/-!
|
||||
|
||||
## The basis of the creation and annihlation free-algebra.
|
||||
|
||||
-/
|
||||
|
||||
/-- The basis of the free creation and annihilation algebra formed by lists of CrAnStates. -/
|
||||
noncomputable def ofCrAnListBasis : Basis (List 𝓕.CrAnStates) ℂ (FieldOpFreeAlgebra 𝓕) where
|
||||
repr := FreeAlgebra.equivMonoidAlgebraFreeMonoid.toLinearEquiv
|
||||
|
||||
@[simp]
|
||||
lemma ofListBasis_eq_ofList (φs : List 𝓕.CrAnStates) :
|
||||
ofCrAnListBasis φs = ofCrAnList φs := by
|
||||
simp only [ofCrAnListBasis, FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
Basis.coe_ofRepr, AlgEquiv.toLinearEquiv_symm, AlgEquiv.toLinearEquiv_apply,
|
||||
AlgEquiv.ofAlgHom_symm_apply, ofCrAnList]
|
||||
erw [MonoidAlgebra.lift_apply]
|
||||
simp only [zero_smul, Finsupp.sum_single_index, one_smul]
|
||||
rw [@FreeMonoid.lift_apply]
|
||||
match φs with
|
||||
| [] => rfl
|
||||
| φ :: φs => erw [List.map_cons]
|
||||
|
||||
/-!
|
||||
|
||||
## Some useful multi-linear maps.
|
||||
|
||||
-/
|
||||
|
||||
/-- The bi-linear map associated with multiplication in `FieldOpFreeAlgebra`. -/
|
||||
noncomputable def mulLinearMap : FieldOpFreeAlgebra 𝓕 →ₗ[ℂ] FieldOpFreeAlgebra 𝓕 →ₗ[ℂ] FieldOpFreeAlgebra 𝓕 where
|
||||
toFun a := {
|
||||
toFun := fun b => a * b,
|
||||
map_add' := fun c d => by simp [mul_add]
|
||||
map_smul' := by simp}
|
||||
map_add' := fun a b => by
|
||||
ext c
|
||||
simp [add_mul]
|
||||
map_smul' := by
|
||||
intros
|
||||
ext c
|
||||
simp [smul_mul']
|
||||
|
||||
lemma mulLinearMap_apply (a b : FieldOpFreeAlgebra 𝓕) :
|
||||
mulLinearMap a b = a * b := rfl
|
||||
|
||||
/-- The linear map associated with scalar-multiplication in `FieldOpFreeAlgebra`. -/
|
||||
noncomputable def smulLinearMap (c : ℂ) : FieldOpFreeAlgebra 𝓕 →ₗ[ℂ] FieldOpFreeAlgebra 𝓕 where
|
||||
toFun a := c • a
|
||||
map_add' := by simp
|
||||
map_smul' m x := by simp [smul_smul, RingHom.id_apply, NonUnitalNormedCommRing.mul_comm]
|
||||
|
||||
end FieldOpFreeAlgebra
|
||||
|
||||
end FieldSpecification
|
|
@ -0,0 +1,412 @@
|
|||
/-
|
||||
Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpFreeAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
import Mathlib.RingTheory.GradedAlgebra.Basic
|
||||
/-!
|
||||
|
||||
# Grading on the FieldOpFreeAlgebra
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldSpecification
|
||||
variable {𝓕 : FieldSpecification}
|
||||
open FieldStatistic
|
||||
|
||||
namespace FieldOpFreeAlgebra
|
||||
|
||||
noncomputable section
|
||||
|
||||
/-- The submodule of `FieldOpFreeAlgebra` spanned by lists of field statistic `f`. -/
|
||||
def statisticSubmodule (f : FieldStatistic) : Submodule ℂ 𝓕.FieldOpFreeAlgebra :=
|
||||
Submodule.span ℂ {a | ∃ φs, a = ofCrAnList φs ∧ (𝓕 |>ₛ φs) = f}
|
||||
|
||||
lemma ofCrAnList_mem_statisticSubmodule_of (φs : List 𝓕.CrAnStates) (f : FieldStatistic)
|
||||
(h : (𝓕 |>ₛ φs) = f) :
|
||||
ofCrAnList φs ∈ statisticSubmodule f := by
|
||||
refine Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl, h⟩⟩
|
||||
|
||||
lemma ofCrAnList_bosonic_or_fermionic (φs : List 𝓕.CrAnStates) :
|
||||
ofCrAnList φs ∈ statisticSubmodule bosonic ∨ ofCrAnList φs ∈ statisticSubmodule fermionic := by
|
||||
by_cases h : (𝓕 |>ₛ φs) = bosonic
|
||||
· left
|
||||
exact ofCrAnList_mem_statisticSubmodule_of φs bosonic h
|
||||
· right
|
||||
exact ofCrAnList_mem_statisticSubmodule_of φs fermionic (by simpa using h)
|
||||
|
||||
lemma ofCrAnState_bosonic_or_fermionic (φ : 𝓕.CrAnStates) :
|
||||
ofCrAnState φ ∈ statisticSubmodule bosonic ∨ ofCrAnState φ ∈ statisticSubmodule fermionic := by
|
||||
rw [← ofCrAnList_singleton]
|
||||
exact ofCrAnList_bosonic_or_fermionic [φ]
|
||||
|
||||
/-- The projection of an element of `FieldOpFreeAlgebra` onto it's bosonic part. -/
|
||||
def bosonicProj : 𝓕.FieldOpFreeAlgebra →ₗ[ℂ] statisticSubmodule (𝓕 := 𝓕) bosonic :=
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs =>
|
||||
if h : (𝓕 |>ₛ φs) = bosonic then
|
||||
⟨ofCrAnList φs, Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl, h⟩⟩⟩
|
||||
else
|
||||
0
|
||||
|
||||
lemma bosonicProj_ofCrAnList (φs : List 𝓕.CrAnStates) :
|
||||
bosonicProj (ofCrAnList φs) = if h : (𝓕 |>ₛ φs) = bosonic then
|
||||
⟨ofCrAnList φs, Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl, h⟩⟩⟩ else 0 := by
|
||||
conv_lhs =>
|
||||
rw [← ofListBasis_eq_ofList, bosonicProj, Basis.constr_basis]
|
||||
|
||||
lemma bosonicProj_of_mem_bosonic (a : 𝓕.FieldOpFreeAlgebra) (h : a ∈ statisticSubmodule bosonic) :
|
||||
bosonicProj a = ⟨a, h⟩ := by
|
||||
let p (a : 𝓕.FieldOpFreeAlgebra) (hx : a ∈ statisticSubmodule bosonic) : Prop :=
|
||||
bosonicProj a = ⟨a, hx⟩
|
||||
change p a h
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
simp only [Set.mem_setOf_eq] at hx
|
||||
obtain ⟨φs, rfl, h⟩ := hx
|
||||
simp [p, bosonicProj_ofCrAnList, h]
|
||||
· simp only [map_zero, p]
|
||||
rfl
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all [p]
|
||||
· intro a x hx hy
|
||||
simp_all [p]
|
||||
|
||||
lemma bosonicProj_of_mem_fermionic (a : 𝓕.FieldOpFreeAlgebra) (h : a ∈ statisticSubmodule fermionic) :
|
||||
bosonicProj a = 0 := by
|
||||
let p (a : 𝓕.FieldOpFreeAlgebra) (hx : a ∈ statisticSubmodule fermionic) : Prop :=
|
||||
bosonicProj a = 0
|
||||
change p a h
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
simp only [Set.mem_setOf_eq] at hx
|
||||
obtain ⟨φs, rfl, h⟩ := hx
|
||||
simp [p, bosonicProj_ofCrAnList, h]
|
||||
· simp [p]
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all [p]
|
||||
· intro a x hx hy
|
||||
simp_all [p]
|
||||
|
||||
@[simp]
|
||||
lemma bosonicProj_of_bonosic_part
|
||||
(a : DirectSum FieldStatistic (fun i => (statisticSubmodule (𝓕 := 𝓕) i))) :
|
||||
bosonicProj (a bosonic) = a bosonic := by
|
||||
apply bosonicProj_of_mem_bosonic
|
||||
|
||||
@[simp]
|
||||
lemma bosonicProj_of_fermionic_part
|
||||
(a : DirectSum FieldStatistic (fun i => (statisticSubmodule (𝓕 := 𝓕) i))) :
|
||||
bosonicProj (a fermionic).1 = 0 := by
|
||||
apply bosonicProj_of_mem_fermionic
|
||||
exact Submodule.coe_mem (a.toFun fermionic)
|
||||
|
||||
/-- The projection of an element of `FieldOpFreeAlgebra` onto it's fermionic part. -/
|
||||
def fermionicProj : 𝓕.FieldOpFreeAlgebra →ₗ[ℂ] statisticSubmodule (𝓕 := 𝓕) fermionic :=
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs =>
|
||||
if h : (𝓕 |>ₛ φs) = fermionic then
|
||||
⟨ofCrAnList φs, Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl, h⟩⟩⟩
|
||||
else
|
||||
0
|
||||
|
||||
lemma fermionicProj_ofCrAnList (φs : List 𝓕.CrAnStates) :
|
||||
fermionicProj (ofCrAnList φs) = if h : (𝓕 |>ₛ φs) = fermionic then
|
||||
⟨ofCrAnList φs, Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl, h⟩⟩⟩ else 0 := by
|
||||
conv_lhs =>
|
||||
rw [← ofListBasis_eq_ofList, fermionicProj, Basis.constr_basis]
|
||||
|
||||
lemma fermionicProj_ofCrAnList_if_bosonic (φs : List 𝓕.CrAnStates) :
|
||||
fermionicProj (ofCrAnList φs) = if h : (𝓕 |>ₛ φs) = bosonic then
|
||||
0 else ⟨ofCrAnList φs, Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl,
|
||||
by simpa using h⟩⟩⟩ := by
|
||||
rw [fermionicProj_ofCrAnList]
|
||||
by_cases h1 : (𝓕 |>ₛ φs) = fermionic
|
||||
· simp [h1]
|
||||
· simp only [h1, ↓reduceDIte]
|
||||
simp only [neq_fermionic_iff_eq_bosonic] at h1
|
||||
simp [h1]
|
||||
|
||||
lemma fermionicProj_of_mem_fermionic (a : 𝓕.FieldOpFreeAlgebra) (h : a ∈ statisticSubmodule fermionic) :
|
||||
fermionicProj a = ⟨a, h⟩ := by
|
||||
let p (a : 𝓕.FieldOpFreeAlgebra) (hx : a ∈ statisticSubmodule fermionic) : Prop :=
|
||||
fermionicProj a = ⟨a, hx⟩
|
||||
change p a h
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
simp only [Set.mem_setOf_eq] at hx
|
||||
obtain ⟨φs, rfl, h⟩ := hx
|
||||
simp [p, fermionicProj_ofCrAnList, h]
|
||||
· simp only [map_zero, p]
|
||||
rfl
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all [p]
|
||||
· intro a x hx hy
|
||||
simp_all [p]
|
||||
|
||||
lemma fermionicProj_of_mem_bosonic (a : 𝓕.FieldOpFreeAlgebra) (h : a ∈ statisticSubmodule bosonic) :
|
||||
fermionicProj a = 0 := by
|
||||
let p (a : 𝓕.FieldOpFreeAlgebra) (hx : a ∈ statisticSubmodule bosonic) : Prop :=
|
||||
fermionicProj a = 0
|
||||
change p a h
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
simp only [Set.mem_setOf_eq] at hx
|
||||
obtain ⟨φs, rfl, h⟩ := hx
|
||||
simp [p, fermionicProj_ofCrAnList, h]
|
||||
· simp [p]
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all [p]
|
||||
· intro a x hx hy
|
||||
simp_all [p]
|
||||
|
||||
@[simp]
|
||||
lemma fermionicProj_of_bosonic_part
|
||||
(a : DirectSum FieldStatistic (fun i => (statisticSubmodule (𝓕 := 𝓕) i))) :
|
||||
fermionicProj (a bosonic).1 = 0 := by
|
||||
apply fermionicProj_of_mem_bosonic
|
||||
exact Submodule.coe_mem (a.toFun bosonic)
|
||||
|
||||
@[simp]
|
||||
lemma fermionicProj_of_fermionic_part
|
||||
(a : DirectSum FieldStatistic (fun i => (statisticSubmodule (𝓕 := 𝓕) i))) :
|
||||
fermionicProj (a fermionic) = a fermionic := by
|
||||
apply fermionicProj_of_mem_fermionic
|
||||
|
||||
lemma bosonicProj_add_fermionicProj (a : 𝓕.FieldOpFreeAlgebra) :
|
||||
a.bosonicProj + (a.fermionicProj).1 = a := by
|
||||
let f1 :𝓕.FieldOpFreeAlgebra →ₗ[ℂ] 𝓕.FieldOpFreeAlgebra :=
|
||||
(statisticSubmodule bosonic).subtype ∘ₗ bosonicProj
|
||||
let f2 :𝓕.FieldOpFreeAlgebra →ₗ[ℂ] 𝓕.FieldOpFreeAlgebra :=
|
||||
(statisticSubmodule fermionic).subtype ∘ₗ fermionicProj
|
||||
change (f1 + f2) a = LinearMap.id (R := ℂ) a
|
||||
refine LinearMap.congr_fun (ofCrAnListBasis.ext fun φs ↦ ?_) a
|
||||
simp only [ofListBasis_eq_ofList, LinearMap.add_apply, LinearMap.coe_comp, Submodule.coe_subtype,
|
||||
Function.comp_apply, LinearMap.id_coe, id_eq, f1, f2]
|
||||
rw [bosonicProj_ofCrAnList, fermionicProj_ofCrAnList_if_bosonic]
|
||||
by_cases h : (𝓕 |>ₛ φs) = bosonic
|
||||
· simp [h]
|
||||
· simp [h]
|
||||
|
||||
lemma coeAddMonoidHom_apply_eq_bosonic_plus_fermionic
|
||||
(a : DirectSum FieldStatistic (fun i => (statisticSubmodule (𝓕 := 𝓕) i))) :
|
||||
DirectSum.coeAddMonoidHom statisticSubmodule a = a.1 bosonic + a.1 fermionic := by
|
||||
let C : DirectSum FieldStatistic (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) → Prop :=
|
||||
fun a => DirectSum.coeAddMonoidHom statisticSubmodule a = a.1 bosonic + a.1 fermionic
|
||||
change C a
|
||||
apply DirectSum.induction_on
|
||||
· simp [C]
|
||||
· intro i x
|
||||
simp only [DFinsupp.toFun_eq_coe, DirectSum.coeAddMonoidHom_of, C]
|
||||
rw [DirectSum.of_apply, DirectSum.of_apply]
|
||||
match i with
|
||||
| bosonic => simp
|
||||
| fermionic => simp
|
||||
· intro x y hx hy
|
||||
simp_all only [C, DFinsupp.toFun_eq_coe, map_add, DirectSum.add_apply, Submodule.coe_add]
|
||||
abel
|
||||
|
||||
lemma directSum_eq_bosonic_plus_fermionic
|
||||
(a : DirectSum FieldStatistic (fun i => (statisticSubmodule (𝓕 := 𝓕) i))) :
|
||||
a = (DirectSum.of (fun i => ↥(statisticSubmodule i)) bosonic) (a bosonic) +
|
||||
(DirectSum.of (fun i => ↥(statisticSubmodule i)) fermionic) (a fermionic) := by
|
||||
let C : DirectSum FieldStatistic (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) → Prop :=
|
||||
fun a => a = (DirectSum.of (fun i => ↥(statisticSubmodule i)) bosonic) (a bosonic) +
|
||||
(DirectSum.of (fun i => ↥(statisticSubmodule i)) fermionic) (a fermionic)
|
||||
change C a
|
||||
apply DirectSum.induction_on
|
||||
· simp [C]
|
||||
· intro i x
|
||||
simp only [C]
|
||||
match i with
|
||||
| bosonic =>
|
||||
simp only [DirectSum.of_eq_same, self_eq_add_right]
|
||||
rw [DirectSum.of_eq_of_ne]
|
||||
simp only [map_zero]
|
||||
simp
|
||||
| fermionic =>
|
||||
simp only [DirectSum.of_eq_same, add_zero]
|
||||
rw [DirectSum.of_eq_of_ne]
|
||||
simp only [map_zero, zero_add]
|
||||
simp
|
||||
· intro x y hx hy
|
||||
simp only [DirectSum.add_apply, map_add, C] at hx hy ⊢
|
||||
conv_lhs => rw [hx, hy]
|
||||
abel
|
||||
|
||||
/-- The instance of a graded algebra on `FieldOpFreeAlgebra`. -/
|
||||
instance fieldOpFreeAlgebraGrade : GradedAlgebra (A := 𝓕.FieldOpFreeAlgebra) statisticSubmodule where
|
||||
one_mem := by
|
||||
simp only [statisticSubmodule]
|
||||
refine Submodule.mem_span.mpr fun p a => a ?_
|
||||
simp only [Set.mem_setOf_eq]
|
||||
use []
|
||||
simp only [ofCrAnList_nil, ofList_empty, true_and]
|
||||
rfl
|
||||
mul_mem f1 f2 a1 a2 h1 h2 := by
|
||||
let p (a2 : 𝓕.FieldOpFreeAlgebra) (hx : a2 ∈ statisticSubmodule f2) : Prop :=
|
||||
a1 * a2 ∈ statisticSubmodule (f1 + f2)
|
||||
change p a2 h2
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
simp only [Set.mem_setOf_eq] at hx
|
||||
obtain ⟨φs, rfl, h⟩ := hx
|
||||
simp only [p]
|
||||
let p (a1 : 𝓕.FieldOpFreeAlgebra) (hx : a1 ∈ statisticSubmodule f1) : Prop :=
|
||||
a1 * ofCrAnList φs ∈ statisticSubmodule (f1 + f2)
|
||||
change p a1 h1
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro y hy
|
||||
obtain ⟨φs', rfl, h'⟩ := hy
|
||||
simp only [p]
|
||||
rw [← ofCrAnList_append]
|
||||
refine Submodule.mem_span.mpr fun p a => a ?_
|
||||
simp only [Set.mem_setOf_eq]
|
||||
use φs' ++ φs
|
||||
simp only [ofList_append, h', h, true_and]
|
||||
cases f1 <;> cases f2 <;> rfl
|
||||
· simp [p]
|
||||
· intro x y hx hy hx1 hx2
|
||||
simp only [add_mul, p]
|
||||
exact Submodule.add_mem _ hx1 hx2
|
||||
· intro c a hx h1
|
||||
simp only [Algebra.smul_mul_assoc, p]
|
||||
exact Submodule.smul_mem _ _ h1
|
||||
· exact h1
|
||||
· simp [p]
|
||||
· intro x y hx hy hx1 hx2
|
||||
simp only [mul_add, p]
|
||||
exact Submodule.add_mem _ hx1 hx2
|
||||
· intro c a hx h1
|
||||
simp only [Algebra.mul_smul_comm, p]
|
||||
exact Submodule.smul_mem _ _ h1
|
||||
· exact h2
|
||||
decompose' a := DirectSum.of (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) bosonic (bosonicProj a)
|
||||
+ DirectSum.of (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) fermionic (fermionicProj a)
|
||||
left_inv a := by
|
||||
trans a.bosonicProj + fermionicProj a
|
||||
· simp
|
||||
· exact bosonicProj_add_fermionicProj a
|
||||
right_inv a := by
|
||||
rw [coeAddMonoidHom_apply_eq_bosonic_plus_fermionic]
|
||||
simp only [DFinsupp.toFun_eq_coe, map_add, bosonicProj_of_bonosic_part,
|
||||
bosonicProj_of_fermionic_part, add_zero, fermionicProj_of_bosonic_part,
|
||||
fermionicProj_of_fermionic_part, zero_add]
|
||||
conv_rhs => rw [directSum_eq_bosonic_plus_fermionic a]
|
||||
|
||||
lemma eq_zero_of_bosonic_and_fermionic {a : 𝓕.FieldOpFreeAlgebra}
|
||||
(hb : a ∈ statisticSubmodule bosonic) (hf : a ∈ statisticSubmodule fermionic) : a = 0 := by
|
||||
have ha := bosonicProj_of_mem_bosonic a hb
|
||||
have hb := fermionicProj_of_mem_fermionic a hf
|
||||
have hc := (bosonicProj_add_fermionicProj a)
|
||||
rw [ha, hb] at hc
|
||||
simpa using hc
|
||||
|
||||
lemma bosonicProj_mul (a b : 𝓕.FieldOpFreeAlgebra) :
|
||||
(a * b).bosonicProj.1 = a.bosonicProj.1 * b.bosonicProj.1
|
||||
+ a.fermionicProj.1 * b.fermionicProj.1 := by
|
||||
conv_lhs =>
|
||||
rw [← bosonicProj_add_fermionicProj a]
|
||||
rw [← bosonicProj_add_fermionicProj b]
|
||||
simp only [mul_add, add_mul, map_add, Submodule.coe_add]
|
||||
rw [bosonicProj_of_mem_bosonic]
|
||||
conv_lhs =>
|
||||
left
|
||||
right
|
||||
rw [bosonicProj_of_mem_fermionic _
|
||||
(by
|
||||
have h1 : fermionic = fermionic + bosonic := by simp
|
||||
conv_lhs => rw [h1]
|
||||
apply fieldOpFreeAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
conv_lhs =>
|
||||
right
|
||||
left
|
||||
rw [bosonicProj_of_mem_fermionic _
|
||||
(by
|
||||
have h1 : fermionic = bosonic + fermionic := by simp
|
||||
conv_lhs => rw [h1]
|
||||
apply fieldOpFreeAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
conv_lhs =>
|
||||
right
|
||||
right
|
||||
rw [bosonicProj_of_mem_bosonic _
|
||||
(by
|
||||
have h1 : bosonic = fermionic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
conv_lhs => rw [h1]
|
||||
apply fieldOpFreeAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
simp only [ZeroMemClass.coe_zero, add_zero, zero_add]
|
||||
· have h1 : bosonic = bosonic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
conv_lhs => rw [h1]
|
||||
apply fieldOpFreeAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp
|
||||
|
||||
lemma fermionicProj_mul (a b : 𝓕.FieldOpFreeAlgebra) :
|
||||
(a * b).fermionicProj.1 = a.bosonicProj.1 * b.fermionicProj.1
|
||||
+ a.fermionicProj.1 * b.bosonicProj.1 := by
|
||||
conv_lhs =>
|
||||
rw [← bosonicProj_add_fermionicProj a]
|
||||
rw [← bosonicProj_add_fermionicProj b]
|
||||
simp only [mul_add, add_mul, map_add, Submodule.coe_add]
|
||||
conv_lhs =>
|
||||
left
|
||||
left
|
||||
rw [fermionicProj_of_mem_bosonic _
|
||||
(by
|
||||
have h1 : bosonic = bosonic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
conv_lhs => rw [h1]
|
||||
apply fieldOpFreeAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
conv_lhs =>
|
||||
left
|
||||
right
|
||||
rw [fermionicProj_of_mem_fermionic _
|
||||
(by
|
||||
have h1 : fermionic = fermionic + bosonic := by simp
|
||||
conv_lhs => rw [h1]
|
||||
apply fieldOpFreeAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
conv_lhs =>
|
||||
right
|
||||
left
|
||||
rw [fermionicProj_of_mem_fermionic _
|
||||
(by
|
||||
have h1 : fermionic = bosonic + fermionic := by simp
|
||||
conv_lhs => rw [h1]
|
||||
apply fieldOpFreeAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
conv_lhs =>
|
||||
right
|
||||
right
|
||||
rw [fermionicProj_of_mem_bosonic _
|
||||
(by
|
||||
have h1 : bosonic = fermionic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
conv_lhs => rw [h1]
|
||||
apply fieldOpFreeAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
simp only [ZeroMemClass.coe_zero, zero_add, add_zero]
|
||||
abel
|
||||
|
||||
end
|
||||
|
||||
end FieldOpFreeAlgebra
|
||||
|
||||
end FieldSpecification
|
|
@ -0,0 +1,47 @@
|
|||
/-
|
||||
Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.PerturbationTheory.FieldSpecification.TimeOrder
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpFreeAlgebra.SuperCommute
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
/-!
|
||||
|
||||
# Norm-time Ordering in the FieldOpFreeAlgebra
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldSpecification
|
||||
variable {𝓕 : FieldSpecification}
|
||||
open FieldStatistic
|
||||
|
||||
namespace FieldOpFreeAlgebra
|
||||
|
||||
noncomputable section
|
||||
open HepLean.List
|
||||
|
||||
/-!
|
||||
|
||||
## Norm-time order
|
||||
|
||||
-/
|
||||
|
||||
/-- The normal-time ordering on `FieldOpFreeAlgebra`. -/
|
||||
def normTimeOrder : FieldOpFreeAlgebra 𝓕 →ₗ[ℂ] FieldOpFreeAlgebra 𝓕 :=
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs =>
|
||||
normTimeOrderSign φs • ofCrAnList (normTimeOrderList φs)
|
||||
|
||||
@[inherit_doc normTimeOrder]
|
||||
scoped[FieldSpecification.FieldOpFreeAlgebra] notation "𝓣𝓝ᶠ(" a ")" => normTimeOrder a
|
||||
|
||||
lemma normTimeOrder_ofCrAnList (φs : List 𝓕.CrAnStates) :
|
||||
𝓣𝓝ᶠ(ofCrAnList φs) = normTimeOrderSign φs • ofCrAnList (normTimeOrderList φs) := by
|
||||
rw [← ofListBasis_eq_ofList]
|
||||
simp only [normTimeOrder, Basis.constr_basis]
|
||||
|
||||
end
|
||||
|
||||
end FieldOpFreeAlgebra
|
||||
|
||||
end FieldSpecification
|
|
@ -0,0 +1,575 @@
|
|||
/-
|
||||
Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.PerturbationTheory.FieldSpecification.NormalOrder
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpFreeAlgebra.SuperCommute
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
/-!
|
||||
|
||||
# Normal Ordering in the FieldOpFreeAlgebra
|
||||
|
||||
In the module
|
||||
`HepLean.PerturbationTheory.FieldSpecification.NormalOrder`
|
||||
we defined the normal ordering of a list of `CrAnStates`.
|
||||
In this module we extend the normal ordering to a linear map on `FieldOpFreeAlgebra`.
|
||||
|
||||
We derive properties of this normal ordering.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldSpecification
|
||||
variable {𝓕 : FieldSpecification}
|
||||
open FieldStatistic
|
||||
|
||||
namespace FieldOpFreeAlgebra
|
||||
|
||||
noncomputable section
|
||||
|
||||
/-- The linear map on the free creation and annihlation
|
||||
algebra defined as the map taking
|
||||
a list of CrAnStates to the normal-ordered list of states multiplied by
|
||||
the sign corresponding to the number of fermionic-fermionic
|
||||
exchanges done in ordering. -/
|
||||
def normalOrderF : FieldOpFreeAlgebra 𝓕 →ₗ[ℂ] FieldOpFreeAlgebra 𝓕 :=
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs =>
|
||||
normalOrderSign φs • ofCrAnList (normalOrderList φs)
|
||||
|
||||
@[inherit_doc normalOrderF]
|
||||
scoped[FieldSpecification.FieldOpFreeAlgebra] notation "𝓝ᶠ(" a ")" => normalOrderF a
|
||||
|
||||
lemma normalOrderF_ofCrAnList (φs : List 𝓕.CrAnStates) :
|
||||
𝓝ᶠ(ofCrAnList φs) = normalOrderSign φs • ofCrAnList (normalOrderList φs) := by
|
||||
rw [← ofListBasis_eq_ofList, normalOrderF, Basis.constr_basis]
|
||||
|
||||
lemma ofCrAnList_eq_normalOrderF (φs : List 𝓕.CrAnStates) :
|
||||
ofCrAnList (normalOrderList φs) = normalOrderSign φs • 𝓝ᶠ(ofCrAnList φs) := by
|
||||
rw [normalOrderF_ofCrAnList, normalOrderList, smul_smul, normalOrderSign,
|
||||
Wick.koszulSign_mul_self, one_smul]
|
||||
|
||||
lemma normalOrderF_one : normalOrderF (𝓕 := 𝓕) 1 = 1 := by
|
||||
rw [← ofCrAnList_nil, normalOrderF_ofCrAnList, normalOrderSign_nil, normalOrderList_nil,
|
||||
ofCrAnList_nil, one_smul]
|
||||
|
||||
lemma normalOrderF_normalOrderF_mid (a b c : 𝓕.FieldOpFreeAlgebra) :
|
||||
𝓝ᶠ(a * b * c) = 𝓝ᶠ(a * 𝓝ᶠ(b) * c) := by
|
||||
let pc (c : 𝓕.FieldOpFreeAlgebra) (hc : c ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓝ᶠ(a * b * c) = 𝓝ᶠ(a * 𝓝ᶠ(b) * c)
|
||||
change pc c (Basis.mem_span _ c)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pc]
|
||||
let pb (b : 𝓕.FieldOpFreeAlgebra) (hb : b ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓝ᶠ(a * b * ofCrAnList φs) = 𝓝ᶠ(a * 𝓝ᶠ(b) * ofCrAnList φs)
|
||||
change pb b (Basis.mem_span _ b)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pb]
|
||||
let pa (a : 𝓕.FieldOpFreeAlgebra) (ha : a ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓝ᶠ(a * ofCrAnList φs' * ofCrAnList φs) = 𝓝ᶠ(a * 𝓝ᶠ(ofCrAnList φs') * ofCrAnList φs)
|
||||
change pa a (Basis.mem_span _ a)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs'', rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pa]
|
||||
rw [normalOrderF_ofCrAnList]
|
||||
simp only [← ofCrAnList_append, Algebra.mul_smul_comm,
|
||||
Algebra.smul_mul_assoc, map_smul]
|
||||
rw [normalOrderF_ofCrAnList, normalOrderF_ofCrAnList, smul_smul]
|
||||
congr 1
|
||||
· simp only [normalOrderSign, normalOrderList]
|
||||
rw [Wick.koszulSign_of_append_eq_insertionSort, mul_comm]
|
||||
· congr 1
|
||||
simp only [normalOrderList]
|
||||
rw [HepLean.List.insertionSort_append_insertionSort_append]
|
||||
· simp [pa]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pa, add_mul]
|
||||
· intro x hx h
|
||||
simp_all [pa]
|
||||
· simp [pb]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pb, mul_add, add_mul]
|
||||
· intro x hx h
|
||||
simp_all [pb]
|
||||
· simp [pc]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pc, mul_add]
|
||||
· intro x hx h hp
|
||||
simp_all [pc]
|
||||
|
||||
lemma normalOrderF_normalOrderF_right (a b : 𝓕.FieldOpFreeAlgebra) : 𝓝ᶠ(a * b) = 𝓝ᶠ(a * 𝓝ᶠ(b)) := by
|
||||
trans 𝓝ᶠ(a * b * 1)
|
||||
· simp
|
||||
· rw [normalOrderF_normalOrderF_mid]
|
||||
simp
|
||||
|
||||
lemma normalOrderF_normalOrderF_left (a b : 𝓕.FieldOpFreeAlgebra) : 𝓝ᶠ(a * b) = 𝓝ᶠ(𝓝ᶠ(a) * b) := by
|
||||
trans 𝓝ᶠ(1 * a * b)
|
||||
· simp
|
||||
· rw [normalOrderF_normalOrderF_mid]
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## Normal ordering with a creation operator on the left or annihilation on the right
|
||||
|
||||
-/
|
||||
|
||||
lemma normalOrderF_ofCrAnList_cons_create (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create) (φs : List 𝓕.CrAnStates) :
|
||||
𝓝ᶠ(ofCrAnList (φ :: φs)) = ofCrAnState φ * 𝓝ᶠ(ofCrAnList φs) := by
|
||||
rw [normalOrderF_ofCrAnList, normalOrderSign_cons_create φ hφ,
|
||||
normalOrderList_cons_create φ hφ φs]
|
||||
rw [ofCrAnList_cons, normalOrderF_ofCrAnList, mul_smul_comm]
|
||||
|
||||
lemma normalOrderF_create_mul (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create) (a : FieldOpFreeAlgebra 𝓕) :
|
||||
𝓝ᶠ(ofCrAnState φ * a) = ofCrAnState φ * 𝓝ᶠ(a) := by
|
||||
change (normalOrderF ∘ₗ mulLinearMap (ofCrAnState φ)) a =
|
||||
(mulLinearMap (ofCrAnState φ) ∘ₗ normalOrderF) a
|
||||
refine LinearMap.congr_fun (ofCrAnListBasis.ext fun l ↦ ?_) a
|
||||
simp only [mulLinearMap, LinearMap.coe_mk, AddHom.coe_mk, ofListBasis_eq_ofList,
|
||||
LinearMap.coe_comp, Function.comp_apply]
|
||||
rw [← ofCrAnList_cons, normalOrderF_ofCrAnList_cons_create φ hφ]
|
||||
|
||||
lemma normalOrderF_ofCrAnList_append_annihilate (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) (φs : List 𝓕.CrAnStates) :
|
||||
𝓝ᶠ(ofCrAnList (φs ++ [φ])) = 𝓝ᶠ(ofCrAnList φs) * ofCrAnState φ := by
|
||||
rw [normalOrderF_ofCrAnList, normalOrderSign_append_annihlate φ hφ φs,
|
||||
normalOrderList_append_annihilate φ hφ φs, ofCrAnList_append, ofCrAnList_singleton,
|
||||
normalOrderF_ofCrAnList, smul_mul_assoc]
|
||||
|
||||
lemma normalOrderF_mul_annihilate (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate)
|
||||
(a : FieldOpFreeAlgebra 𝓕) : 𝓝ᶠ(a * ofCrAnState φ) = 𝓝ᶠ(a) * ofCrAnState φ := by
|
||||
change (normalOrderF ∘ₗ mulLinearMap.flip (ofCrAnState φ)) a =
|
||||
(mulLinearMap.flip (ofCrAnState φ) ∘ₗ normalOrderF) a
|
||||
refine LinearMap.congr_fun (ofCrAnListBasis.ext fun l ↦ ?_) a
|
||||
simp only [mulLinearMap, ofListBasis_eq_ofList, LinearMap.coe_comp, Function.comp_apply,
|
||||
LinearMap.flip_apply, LinearMap.coe_mk, AddHom.coe_mk]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_append, ofCrAnList_singleton,
|
||||
normalOrderF_ofCrAnList_append_annihilate φ hφ]
|
||||
|
||||
lemma normalOrderF_crPartF_mul (φ : 𝓕.States) (a : FieldOpFreeAlgebra 𝓕) :
|
||||
𝓝ᶠ(crPartF φ * a) =
|
||||
crPartF φ * 𝓝ᶠ(a) := by
|
||||
match φ with
|
||||
| .inAsymp φ =>
|
||||
rw [crPartF]
|
||||
exact normalOrderF_create_mul ⟨States.inAsymp φ, ()⟩ rfl a
|
||||
| .position φ =>
|
||||
rw [crPartF]
|
||||
exact normalOrderF_create_mul _ rfl _
|
||||
| .outAsymp φ => simp
|
||||
|
||||
lemma normalOrderF_mul_anPartF (φ : 𝓕.States) (a : FieldOpFreeAlgebra 𝓕) :
|
||||
𝓝ᶠ(a * anPartF φ) =
|
||||
𝓝ᶠ(a) * anPartF φ := by
|
||||
match φ with
|
||||
| .inAsymp φ => simp
|
||||
| .position φ =>
|
||||
rw [anPartF]
|
||||
exact normalOrderF_mul_annihilate _ rfl _
|
||||
| .outAsymp φ =>
|
||||
rw [anPartF]
|
||||
refine normalOrderF_mul_annihilate _ rfl _
|
||||
|
||||
/-!
|
||||
|
||||
## Normal ordering for an adjacent creation and annihliation state
|
||||
|
||||
The main result of this section is `normalOrderF_superCommuteF_annihilate_create`.
|
||||
-/
|
||||
|
||||
lemma normalOrderF_swap_create_annihlate_ofCrAnList_ofCrAnList (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(φs φs' : List 𝓕.CrAnStates) :
|
||||
𝓝ᶠ(ofCrAnList φs' * ofCrAnState φc * ofCrAnState φa * ofCrAnList φs) = 𝓢(𝓕 |>ₛ φc, 𝓕 |>ₛ φa) •
|
||||
𝓝ᶠ(ofCrAnList φs' * ofCrAnState φa * ofCrAnState φc * ofCrAnList φs) := by
|
||||
rw [mul_assoc, mul_assoc, ← ofCrAnList_cons, ← ofCrAnList_cons, ← ofCrAnList_append]
|
||||
rw [normalOrderF_ofCrAnList, normalOrderSign_swap_create_annihlate φc φa hφc hφa]
|
||||
rw [normalOrderList_swap_create_annihlate φc φa hφc hφa, ← smul_smul, ← normalOrderF_ofCrAnList]
|
||||
rw [ofCrAnList_append, ofCrAnList_cons, ofCrAnList_cons]
|
||||
noncomm_ring
|
||||
|
||||
lemma normalOrderF_swap_create_annihlate_ofCrAnList (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(φs : List 𝓕.CrAnStates) (a : 𝓕.FieldOpFreeAlgebra) :
|
||||
𝓝ᶠ(ofCrAnList φs * ofCrAnState φc * ofCrAnState φa * a) = 𝓢(𝓕 |>ₛ φc, 𝓕 |>ₛ φa) •
|
||||
𝓝ᶠ(ofCrAnList φs * ofCrAnState φa * ofCrAnState φc * a) := by
|
||||
change (normalOrderF ∘ₗ mulLinearMap (ofCrAnList φs * ofCrAnState φc * ofCrAnState φa)) a =
|
||||
(smulLinearMap _ ∘ₗ normalOrderF ∘ₗ
|
||||
mulLinearMap (ofCrAnList φs * ofCrAnState φa * ofCrAnState φc)) a
|
||||
refine LinearMap.congr_fun (ofCrAnListBasis.ext fun l ↦ ?_) a
|
||||
simp only [mulLinearMap, LinearMap.coe_mk, AddHom.coe_mk, ofListBasis_eq_ofList,
|
||||
LinearMap.coe_comp, Function.comp_apply, instCommGroup.eq_1]
|
||||
rw [normalOrderF_swap_create_annihlate_ofCrAnList_ofCrAnList φc φa hφc hφa]
|
||||
rfl
|
||||
|
||||
lemma normalOrderF_swap_create_annihlate (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(a b : 𝓕.FieldOpFreeAlgebra) :
|
||||
𝓝ᶠ(a * ofCrAnState φc * ofCrAnState φa * b) = 𝓢(𝓕 |>ₛ φc, 𝓕 |>ₛ φa) •
|
||||
𝓝ᶠ(a * ofCrAnState φa * ofCrAnState φc * b) := by
|
||||
rw [mul_assoc, mul_assoc, mul_assoc, mul_assoc]
|
||||
change (normalOrderF ∘ₗ mulLinearMap.flip (ofCrAnState φc * (ofCrAnState φa * b))) a =
|
||||
(smulLinearMap (𝓢(𝓕 |>ₛ φc, 𝓕 |>ₛ φa)) ∘ₗ
|
||||
normalOrderF ∘ₗ mulLinearMap.flip (ofCrAnState φa * (ofCrAnState φc * b))) a
|
||||
refine LinearMap.congr_fun (ofCrAnListBasis.ext fun l ↦ ?_) _
|
||||
simp only [mulLinearMap, ofListBasis_eq_ofList, LinearMap.coe_comp, Function.comp_apply,
|
||||
LinearMap.flip_apply, LinearMap.coe_mk, AddHom.coe_mk, instCommGroup.eq_1, ← mul_assoc,
|
||||
normalOrderF_swap_create_annihlate_ofCrAnList φc φa hφc hφa]
|
||||
rfl
|
||||
|
||||
lemma normalOrderF_superCommuteF_create_annihilate (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(a b : 𝓕.FieldOpFreeAlgebra) :
|
||||
𝓝ᶠ(a * [ofCrAnState φc, ofCrAnState φa]ₛca * b) = 0 := by
|
||||
simp only [superCommuteF_ofCrAnState_ofCrAnState, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [mul_sub, sub_mul, map_sub, ← smul_mul_assoc, ← mul_assoc, ← mul_assoc,
|
||||
normalOrderF_swap_create_annihlate φc φa hφc hφa]
|
||||
simp
|
||||
|
||||
lemma normalOrderF_superCommuteF_annihilate_create (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(a b : 𝓕.FieldOpFreeAlgebra) :
|
||||
𝓝ᶠ(a * [ofCrAnState φa, ofCrAnState φc]ₛca * b) = 0 := by
|
||||
rw [superCommuteF_ofCrAnState_ofCrAnState_symm]
|
||||
simp only [instCommGroup.eq_1, neg_smul, mul_neg, Algebra.mul_smul_comm, neg_mul,
|
||||
Algebra.smul_mul_assoc, map_neg, map_smul, neg_eq_zero, smul_eq_zero]
|
||||
exact Or.inr (normalOrderF_superCommuteF_create_annihilate φc φa hφc hφa ..)
|
||||
|
||||
lemma normalOrderF_swap_crPartF_anPartF (φ φ' : 𝓕.States) (a b : FieldOpFreeAlgebra 𝓕) :
|
||||
𝓝ᶠ(a * (crPartF φ) * (anPartF φ') * b) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
𝓝ᶠ(a * (anPartF φ') * (crPartF φ) * b) := by
|
||||
match φ, φ' with
|
||||
| _, .inAsymp φ' => simp
|
||||
| .outAsymp φ, _ => simp
|
||||
| .position φ, .position φ' =>
|
||||
simp only [crPartF_position, anPartF_position, instCommGroup.eq_1]
|
||||
rw [normalOrderF_swap_create_annihlate]
|
||||
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod]
|
||||
rfl; rfl
|
||||
| .inAsymp φ, .outAsymp φ' =>
|
||||
simp only [crPartF_negAsymp, anPartF_posAsymp, instCommGroup.eq_1]
|
||||
rw [normalOrderF_swap_create_annihlate]
|
||||
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod]
|
||||
rfl; rfl
|
||||
| .inAsymp φ, .position φ' =>
|
||||
simp only [crPartF_negAsymp, anPartF_position, instCommGroup.eq_1]
|
||||
rw [normalOrderF_swap_create_annihlate]
|
||||
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod]
|
||||
rfl; rfl
|
||||
| .position φ, .outAsymp φ' =>
|
||||
simp only [crPartF_position, anPartF_posAsymp, instCommGroup.eq_1]
|
||||
rw [normalOrderF_swap_create_annihlate]
|
||||
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod]
|
||||
rfl; rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Normal ordering for an anPartF and crPartF
|
||||
|
||||
Using the results from above.
|
||||
|
||||
-/
|
||||
|
||||
lemma normalOrderF_swap_anPartF_crPartF (φ φ' : 𝓕.States) (a b : FieldOpFreeAlgebra 𝓕) :
|
||||
𝓝ᶠ(a * (anPartF φ) * (crPartF φ') * b) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • 𝓝ᶠ(a * (crPartF φ') *
|
||||
(anPartF φ) * b) := by
|
||||
simp [normalOrderF_swap_crPartF_anPartF, smul_smul]
|
||||
|
||||
lemma normalOrderF_superCommuteF_crPartF_anPartF (φ φ' : 𝓕.States) (a b : FieldOpFreeAlgebra 𝓕) :
|
||||
𝓝ᶠ(a * superCommuteF
|
||||
(crPartF φ) (anPartF φ') * b) = 0 := by
|
||||
match φ, φ' with
|
||||
| _, .inAsymp φ' => simp
|
||||
| .outAsymp φ', _ => simp
|
||||
| .position φ, .position φ' =>
|
||||
rw [crPartF_position, anPartF_position]
|
||||
exact normalOrderF_superCommuteF_create_annihilate _ _ rfl rfl ..
|
||||
| .inAsymp φ, .outAsymp φ' =>
|
||||
rw [crPartF_negAsymp, anPartF_posAsymp]
|
||||
exact normalOrderF_superCommuteF_create_annihilate _ _ rfl rfl ..
|
||||
| .inAsymp φ, .position φ' =>
|
||||
rw [crPartF_negAsymp, anPartF_position]
|
||||
exact normalOrderF_superCommuteF_create_annihilate _ _ rfl rfl ..
|
||||
| .position φ, .outAsymp φ' =>
|
||||
rw [crPartF_position, anPartF_posAsymp]
|
||||
exact normalOrderF_superCommuteF_create_annihilate _ _ rfl rfl ..
|
||||
|
||||
lemma normalOrderF_superCommuteF_anPartF_crPartF (φ φ' : 𝓕.States) (a b : FieldOpFreeAlgebra 𝓕) :
|
||||
𝓝ᶠ(a * superCommuteF
|
||||
(anPartF φ) (crPartF φ') * b) = 0 := by
|
||||
match φ, φ' with
|
||||
| .inAsymp φ', _ => simp
|
||||
| _, .outAsymp φ' => simp
|
||||
| .position φ, .position φ' =>
|
||||
rw [anPartF_position, crPartF_position]
|
||||
exact normalOrderF_superCommuteF_annihilate_create _ _ rfl rfl ..
|
||||
| .outAsymp φ', .inAsymp φ =>
|
||||
simp only [anPartF_posAsymp, crPartF_negAsymp]
|
||||
exact normalOrderF_superCommuteF_annihilate_create _ _ rfl rfl ..
|
||||
| .position φ', .inAsymp φ =>
|
||||
simp only [anPartF_position, crPartF_negAsymp]
|
||||
exact normalOrderF_superCommuteF_annihilate_create _ _ rfl rfl ..
|
||||
| .outAsymp φ, .position φ' =>
|
||||
simp only [anPartF_posAsymp, crPartF_position]
|
||||
exact normalOrderF_superCommuteF_annihilate_create _ _ rfl rfl ..
|
||||
|
||||
/-!
|
||||
|
||||
## The normal ordering of a product of two states
|
||||
|
||||
-/
|
||||
|
||||
@[simp]
|
||||
lemma normalOrderF_crPartF_mul_crPartF (φ φ' : 𝓕.States) :
|
||||
𝓝ᶠ(crPartF φ * crPartF φ') =
|
||||
crPartF φ * crPartF φ' := by
|
||||
rw [normalOrderF_crPartF_mul]
|
||||
conv_lhs => rw [← mul_one (crPartF φ')]
|
||||
rw [normalOrderF_crPartF_mul, normalOrderF_one]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma normalOrderF_anPartF_mul_anPartF (φ φ' : 𝓕.States) :
|
||||
𝓝ᶠ(anPartF φ * anPartF φ') =
|
||||
anPartF φ * anPartF φ' := by
|
||||
rw [normalOrderF_mul_anPartF]
|
||||
conv_lhs => rw [← one_mul (anPartF φ)]
|
||||
rw [normalOrderF_mul_anPartF, normalOrderF_one]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma normalOrderF_crPartF_mul_anPartF (φ φ' : 𝓕.States) :
|
||||
𝓝ᶠ(crPartF φ * anPartF φ') =
|
||||
crPartF φ * anPartF φ' := by
|
||||
rw [normalOrderF_crPartF_mul]
|
||||
conv_lhs => rw [← one_mul (anPartF φ')]
|
||||
rw [normalOrderF_mul_anPartF, normalOrderF_one]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma normalOrderF_anPartF_mul_crPartF (φ φ' : 𝓕.States) :
|
||||
𝓝ᶠ(anPartF φ * crPartF φ') =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
(crPartF φ' * anPartF φ) := by
|
||||
conv_lhs => rw [← one_mul (anPartF φ * crPartF φ')]
|
||||
conv_lhs => rw [← mul_one (1 * (anPartF φ *
|
||||
crPartF φ'))]
|
||||
rw [← mul_assoc, normalOrderF_swap_anPartF_crPartF]
|
||||
simp
|
||||
|
||||
lemma normalOrderF_ofState_mul_ofState (φ φ' : 𝓕.States) :
|
||||
𝓝ᶠ(ofState φ * ofState φ') =
|
||||
crPartF φ * crPartF φ' +
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
(crPartF φ' * anPartF φ) +
|
||||
crPartF φ * anPartF φ' +
|
||||
anPartF φ * anPartF φ' := by
|
||||
rw [ofState_eq_crPartF_add_anPartF, ofState_eq_crPartF_add_anPartF, mul_add, add_mul, add_mul]
|
||||
simp only [map_add, normalOrderF_crPartF_mul_crPartF, normalOrderF_anPartF_mul_crPartF,
|
||||
instCommGroup.eq_1, normalOrderF_crPartF_mul_anPartF, normalOrderF_anPartF_mul_anPartF]
|
||||
abel
|
||||
|
||||
/-!
|
||||
|
||||
## Normal order with super commutors
|
||||
|
||||
-/
|
||||
|
||||
TODO "Split the following two lemmas up into smaller parts."
|
||||
|
||||
lemma normalOrderF_superCommuteF_ofCrAnList_create_create_ofCrAnList
|
||||
(φc φc' : 𝓕.CrAnStates) (hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφc' : 𝓕 |>ᶜ φc' = CreateAnnihilate.create) (φs φs' : List 𝓕.CrAnStates) :
|
||||
(𝓝ᶠ(ofCrAnList φs * [ofCrAnState φc, ofCrAnState φc']ₛca * ofCrAnList φs')) =
|
||||
normalOrderSign (φs ++ φc' :: φc :: φs') •
|
||||
(ofCrAnList (createFilter φs) * [ofCrAnState φc, ofCrAnState φc']ₛca *
|
||||
ofCrAnList (createFilter φs') * ofCrAnList (annihilateFilter (φs ++ φs'))) := by
|
||||
rw [superCommuteF_ofCrAnState_ofCrAnState, mul_sub, sub_mul, map_sub]
|
||||
conv_lhs =>
|
||||
lhs; rhs
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_append, ← ofCrAnList_append,
|
||||
← ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
rw [normalOrderF_ofCrAnList, normalOrderList_eq_createFilter_append_annihilateFilter]
|
||||
rw [createFilter_append, createFilter_append, createFilter_append,
|
||||
createFilter_singleton_create _ hφc, createFilter_singleton_create _ hφc']
|
||||
rw [annihilateFilter_append, annihilateFilter_append, annihilateFilter_append,
|
||||
annihilateFilter_singleton_create _ hφc, annihilateFilter_singleton_create _ hφc']
|
||||
enter [2, 1, 2]
|
||||
simp only [List.singleton_append, List.append_assoc, List.cons_append, List.append_nil,
|
||||
instCommGroup.eq_1, Algebra.smul_mul_assoc, Algebra.mul_smul_comm, map_smul]
|
||||
rw [← annihilateFilter_append]
|
||||
conv_lhs =>
|
||||
rhs; rhs
|
||||
rw [smul_mul_assoc, Algebra.mul_smul_comm, smul_mul_assoc]
|
||||
rhs
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_append, ← ofCrAnList_append,
|
||||
← ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rw [map_smul]
|
||||
rhs
|
||||
rw [normalOrderF_ofCrAnList, normalOrderList_eq_createFilter_append_annihilateFilter]
|
||||
rw [createFilter_append, createFilter_append, createFilter_append,
|
||||
createFilter_singleton_create _ hφc, createFilter_singleton_create _ hφc']
|
||||
rw [annihilateFilter_append, annihilateFilter_append, annihilateFilter_append,
|
||||
annihilateFilter_singleton_create _ hφc, annihilateFilter_singleton_create _ hφc']
|
||||
enter [2, 1, 2]
|
||||
simp only [List.singleton_append, List.append_assoc, List.cons_append, instCommGroup.eq_1,
|
||||
List.append_nil, Algebra.smul_mul_assoc]
|
||||
rw [← annihilateFilter_append]
|
||||
conv_lhs =>
|
||||
lhs; lhs
|
||||
simp
|
||||
conv_lhs =>
|
||||
rhs; rhs; lhs
|
||||
simp
|
||||
rw [normalOrderSign_swap_create_create φc φc' hφc hφc']
|
||||
rw [smul_smul, mul_comm, ← smul_smul]
|
||||
rw [← smul_sub, ofCrAnList_append, ofCrAnList_append, ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
rhs; rhs
|
||||
rw [ofCrAnList_append, ofCrAnList_append, ofCrAnList_append]
|
||||
rw [← smul_mul_assoc, ← smul_mul_assoc, ← Algebra.mul_smul_comm]
|
||||
rw [← sub_mul, ← sub_mul, ← mul_sub, ofCrAnList_append, ofCrAnList_singleton,
|
||||
ofCrAnList_singleton]
|
||||
rw [ofCrAnList_append, ofCrAnList_singleton, ofCrAnList_singleton, smul_mul_assoc]
|
||||
|
||||
lemma normalOrderF_superCommuteF_ofCrAnList_annihilate_annihilate_ofCrAnList
|
||||
(φa φa' : 𝓕.CrAnStates)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(hφa' : 𝓕 |>ᶜ φa' = CreateAnnihilate.annihilate)
|
||||
(φs φs' : List 𝓕.CrAnStates) :
|
||||
𝓝ᶠ(ofCrAnList φs * [ofCrAnState φa, ofCrAnState φa']ₛca * ofCrAnList φs') =
|
||||
normalOrderSign (φs ++ φa' :: φa :: φs') •
|
||||
(ofCrAnList (createFilter (φs ++ φs'))
|
||||
* ofCrAnList (annihilateFilter φs) * [ofCrAnState φa, ofCrAnState φa']ₛca
|
||||
* ofCrAnList (annihilateFilter φs')) := by
|
||||
rw [superCommuteF_ofCrAnState_ofCrAnState, mul_sub, sub_mul, map_sub]
|
||||
conv_lhs =>
|
||||
lhs; rhs
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_append, ← ofCrAnList_append,
|
||||
← ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
rw [normalOrderF_ofCrAnList, normalOrderList_eq_createFilter_append_annihilateFilter]
|
||||
rw [createFilter_append, createFilter_append, createFilter_append,
|
||||
createFilter_singleton_annihilate _ hφa, createFilter_singleton_annihilate _ hφa']
|
||||
rw [annihilateFilter_append, annihilateFilter_append, annihilateFilter_append,
|
||||
annihilateFilter_singleton_annihilate _ hφa, annihilateFilter_singleton_annihilate _ hφa']
|
||||
enter [2, 1, 1]
|
||||
simp only [List.singleton_append, List.append_assoc, List.cons_append, List.append_nil,
|
||||
instCommGroup.eq_1, Algebra.smul_mul_assoc, Algebra.mul_smul_comm, map_smul]
|
||||
rw [← createFilter_append]
|
||||
conv_lhs =>
|
||||
rhs; rhs
|
||||
rw [smul_mul_assoc]
|
||||
rw [Algebra.mul_smul_comm, smul_mul_assoc]
|
||||
rhs
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_append, ← ofCrAnList_append,
|
||||
← ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rw [map_smul]
|
||||
rhs
|
||||
rw [normalOrderF_ofCrAnList, normalOrderList_eq_createFilter_append_annihilateFilter]
|
||||
rw [createFilter_append, createFilter_append, createFilter_append,
|
||||
createFilter_singleton_annihilate _ hφa, createFilter_singleton_annihilate _ hφa']
|
||||
rw [annihilateFilter_append, annihilateFilter_append, annihilateFilter_append,
|
||||
annihilateFilter_singleton_annihilate _ hφa, annihilateFilter_singleton_annihilate _ hφa']
|
||||
enter [2, 1, 1]
|
||||
simp only [List.singleton_append, List.append_assoc, List.cons_append, instCommGroup.eq_1,
|
||||
List.append_nil, Algebra.smul_mul_assoc]
|
||||
rw [← createFilter_append]
|
||||
conv_lhs =>
|
||||
lhs; lhs
|
||||
simp
|
||||
conv_lhs =>
|
||||
rhs; rhs; lhs
|
||||
simp
|
||||
rw [normalOrderSign_swap_annihilate_annihilate φa φa' hφa hφa']
|
||||
rw [smul_smul, mul_comm, ← smul_smul]
|
||||
rw [← smul_sub, ofCrAnList_append, ofCrAnList_append, ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
rhs; rhs
|
||||
rw [ofCrAnList_append, ofCrAnList_append, ofCrAnList_append]
|
||||
rw [← Algebra.mul_smul_comm, ← smul_mul_assoc, ← Algebra.mul_smul_comm]
|
||||
rw [← mul_sub, ← sub_mul, ← mul_sub]
|
||||
apply congrArg
|
||||
conv_rhs => rw [mul_assoc, mul_assoc]
|
||||
apply congrArg
|
||||
rw [mul_assoc]
|
||||
apply congrArg
|
||||
rw [ofCrAnList_append, ofCrAnList_singleton, ofCrAnList_singleton]
|
||||
rw [ofCrAnList_append, ofCrAnList_singleton, ofCrAnList_singleton, smul_mul_assoc]
|
||||
|
||||
/-!
|
||||
|
||||
## Super commututators involving a normal order.
|
||||
|
||||
-/
|
||||
|
||||
lemma ofCrAnList_superCommuteF_normalOrderF_ofCrAnList (φs φs' : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs, 𝓝ᶠ(ofCrAnList φs')]ₛca =
|
||||
ofCrAnList φs * 𝓝ᶠ(ofCrAnList φs') -
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • 𝓝ᶠ(ofCrAnList φs') * ofCrAnList φs := by
|
||||
simp [normalOrderF_ofCrAnList, map_smul, superCommuteF_ofCrAnList_ofCrAnList, ofCrAnList_append,
|
||||
smul_sub, smul_smul, mul_comm]
|
||||
|
||||
lemma ofCrAnList_superCommuteF_normalOrderF_ofStateList (φs : List 𝓕.CrAnStates)
|
||||
(φs' : List 𝓕.States) : [ofCrAnList φs, 𝓝ᶠ(ofStateList φs')]ₛca =
|
||||
ofCrAnList φs * 𝓝ᶠ(ofStateList φs') -
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • 𝓝ᶠ(ofStateList φs') * ofCrAnList φs := by
|
||||
rw [ofStateList_sum, map_sum, Finset.mul_sum, Finset.smul_sum, Finset.sum_mul,
|
||||
← Finset.sum_sub_distrib, map_sum]
|
||||
congr
|
||||
funext n
|
||||
rw [ofCrAnList_superCommuteF_normalOrderF_ofCrAnList,
|
||||
CrAnSection.statistics_eq_state_statistics]
|
||||
|
||||
/-!
|
||||
|
||||
## Multiplications with normal order written in terms of super commute.
|
||||
|
||||
-/
|
||||
|
||||
lemma ofCrAnList_mul_normalOrderF_ofStateList_eq_superCommuteF (φs : List 𝓕.CrAnStates)
|
||||
(φs' : List 𝓕.States) :
|
||||
ofCrAnList φs * 𝓝ᶠ(ofStateList φs') =
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • 𝓝ᶠ(ofStateList φs') * ofCrAnList φs
|
||||
+ [ofCrAnList φs, 𝓝ᶠ(ofStateList φs')]ₛca := by
|
||||
simp [ofCrAnList_superCommuteF_normalOrderF_ofStateList]
|
||||
|
||||
lemma ofCrAnState_mul_normalOrderF_ofStateList_eq_superCommuteF (φ : 𝓕.CrAnStates)
|
||||
(φs' : List 𝓕.States) : ofCrAnState φ * 𝓝ᶠ(ofStateList φs') =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • 𝓝ᶠ(ofStateList φs') * ofCrAnState φ
|
||||
+ [ofCrAnState φ, 𝓝ᶠ(ofStateList φs')]ₛca := by
|
||||
simp [← ofCrAnList_singleton, ofCrAnList_mul_normalOrderF_ofStateList_eq_superCommuteF]
|
||||
|
||||
lemma anPartF_mul_normalOrderF_ofStateList_eq_superCommuteF (φ : 𝓕.States)
|
||||
(φs' : List 𝓕.States) :
|
||||
anPartF φ * 𝓝ᶠ(ofStateList φs') =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • 𝓝ᶠ(ofStateList φs' * anPartF φ)
|
||||
+ [anPartF φ, 𝓝ᶠ(ofStateList φs')]ₛca := by
|
||||
rw [normalOrderF_mul_anPartF]
|
||||
match φ with
|
||||
| .inAsymp φ => simp
|
||||
| .position φ => simp [ofCrAnState_mul_normalOrderF_ofStateList_eq_superCommuteF, crAnStatistics]
|
||||
| .outAsymp φ => simp [ofCrAnState_mul_normalOrderF_ofStateList_eq_superCommuteF, crAnStatistics]
|
||||
|
||||
end
|
||||
|
||||
end FieldOpFreeAlgebra
|
||||
|
||||
end FieldSpecification
|
|
@ -0,0 +1,875 @@
|
|||
/-
|
||||
Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpFreeAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpFreeAlgebra.Grading
|
||||
/-!
|
||||
|
||||
# Super Commute
|
||||
-/
|
||||
|
||||
namespace FieldSpecification
|
||||
variable {𝓕 : FieldSpecification}
|
||||
|
||||
namespace FieldOpFreeAlgebra
|
||||
|
||||
/-!
|
||||
|
||||
## The super commutor on the FieldOpFreeAlgebra.
|
||||
|
||||
-/
|
||||
|
||||
open FieldStatistic
|
||||
|
||||
/-- The super commutor on the creation and annihlation algebra. For two bosonic operators
|
||||
or a bosonic and fermionic operator this corresponds to the usual commutator
|
||||
whilst for two fermionic operators this corresponds to the anti-commutator. -/
|
||||
noncomputable def superCommuteF : 𝓕.FieldOpFreeAlgebra →ₗ[ℂ] 𝓕.FieldOpFreeAlgebra →ₗ[ℂ] 𝓕.FieldOpFreeAlgebra :=
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs =>
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs' =>
|
||||
ofCrAnList (φs ++ φs') - 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofCrAnList (φs' ++ φs)
|
||||
|
||||
/-- The super commutor on the creation and annihlation algebra. For two bosonic operators
|
||||
or a bosonic and fermionic operator this corresponds to the usual commutator
|
||||
whilst for two fermionic operators this corresponds to the anti-commutator. -/
|
||||
scoped[FieldSpecification.FieldOpFreeAlgebra] notation "[" φs "," φs' "]ₛca" => superCommuteF φs φs'
|
||||
|
||||
/-!
|
||||
|
||||
## The super commutor of different types of elements
|
||||
|
||||
-/
|
||||
|
||||
lemma superCommuteF_ofCrAnList_ofCrAnList (φs φs' : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs, ofCrAnList φs']ₛca =
|
||||
ofCrAnList (φs ++ φs') - 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofCrAnList (φs' ++ φs) := by
|
||||
rw [← ofListBasis_eq_ofList, ← ofListBasis_eq_ofList]
|
||||
simp only [superCommuteF, Basis.constr_basis]
|
||||
|
||||
lemma superCommuteF_ofCrAnState_ofCrAnState (φ φ' : 𝓕.CrAnStates) :
|
||||
[ofCrAnState φ, ofCrAnState φ']ₛca =
|
||||
ofCrAnState φ * ofCrAnState φ' - 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • ofCrAnState φ' * ofCrAnState φ := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList, ofCrAnList_append]
|
||||
congr
|
||||
rw [ofCrAnList_append]
|
||||
rw [FieldStatistic.ofList_singleton, FieldStatistic.ofList_singleton, smul_mul_assoc]
|
||||
|
||||
lemma superCommuteF_ofCrAnList_ofStatesList (φcas : List 𝓕.CrAnStates) (φs : List 𝓕.States) :
|
||||
[ofCrAnList φcas, ofStateList φs]ₛca = ofCrAnList φcas * ofStateList φs -
|
||||
𝓢(𝓕 |>ₛ φcas, 𝓕 |>ₛ φs) • ofStateList φs * ofCrAnList φcas := by
|
||||
conv_lhs => rw [ofStateList_sum]
|
||||
rw [map_sum]
|
||||
conv_lhs =>
|
||||
enter [2, x]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList, CrAnSection.statistics_eq_state_statistics,
|
||||
ofCrAnList_append, ofCrAnList_append]
|
||||
rw [Finset.sum_sub_distrib, ← Finset.mul_sum, ← Finset.smul_sum,
|
||||
← Finset.sum_mul, ← ofStateList_sum]
|
||||
simp
|
||||
|
||||
lemma superCommuteF_ofStateList_ofStatesList (φ : List 𝓕.States) (φs : List 𝓕.States) :
|
||||
[ofStateList φ, ofStateList φs]ₛca = ofStateList φ * ofStateList φs -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) • ofStateList φs * ofStateList φ := by
|
||||
conv_lhs => rw [ofStateList_sum]
|
||||
simp only [map_sum, LinearMap.coeFn_sum, Finset.sum_apply, instCommGroup.eq_1,
|
||||
Algebra.smul_mul_assoc]
|
||||
conv_lhs =>
|
||||
enter [2, x]
|
||||
rw [superCommuteF_ofCrAnList_ofStatesList]
|
||||
simp only [instCommGroup.eq_1, CrAnSection.statistics_eq_state_statistics,
|
||||
Algebra.smul_mul_assoc, Finset.sum_sub_distrib]
|
||||
rw [← Finset.sum_mul, ← Finset.smul_sum, ← Finset.mul_sum, ← ofStateList_sum]
|
||||
|
||||
lemma superCommuteF_ofState_ofStatesList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
[ofState φ, ofStateList φs]ₛca = ofState φ * ofStateList φs -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) • ofStateList φs * ofState φ := by
|
||||
rw [← ofStateList_singleton, superCommuteF_ofStateList_ofStatesList, ofStateList_singleton]
|
||||
simp
|
||||
|
||||
lemma superCommuteF_ofStateList_ofState (φs : List 𝓕.States) (φ : 𝓕.States) :
|
||||
[ofStateList φs, ofState φ]ₛca = ofStateList φs * ofState φ -
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ) • ofState φ * ofStateList φs := by
|
||||
rw [← ofStateList_singleton, superCommuteF_ofStateList_ofStatesList, ofStateList_singleton]
|
||||
simp
|
||||
|
||||
lemma superCommuteF_anPartF_crPartF (φ φ' : 𝓕.States) :
|
||||
[anPartF φ, crPartF φ']ₛca = anPartF φ * crPartF φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • crPartF φ' * anPartF φ := by
|
||||
match φ, φ' with
|
||||
| States.inAsymp φ, _ =>
|
||||
simp
|
||||
| _, States.outAsymp φ =>
|
||||
simp only [crPartF_posAsymp, map_zero, mul_zero, instCommGroup.eq_1, smul_zero, zero_mul,
|
||||
sub_self]
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [anPartF_position, crPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.outAsymp φ, States.position φ' =>
|
||||
simp only [anPartF_posAsymp, crPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.inAsymp φ' =>
|
||||
simp only [anPartF_position, crPartF_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp only [List.singleton_append, instCommGroup.eq_1, crAnStatistics,
|
||||
FieldStatistic.ofList_singleton, Function.comp_apply, crAnStatesToStates_prod, ←
|
||||
ofCrAnList_append]
|
||||
| States.outAsymp φ, States.inAsymp φ' =>
|
||||
simp only [anPartF_posAsymp, crPartF_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma superCommuteF_crPartF_anPartF (φ φ' : 𝓕.States) :
|
||||
[crPartF φ, anPartF φ']ₛca = crPartF φ * anPartF φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • anPartF φ' * crPartF φ := by
|
||||
match φ, φ' with
|
||||
| States.outAsymp φ, _ =>
|
||||
simp only [crPartF_posAsymp, map_zero, LinearMap.zero_apply, zero_mul, instCommGroup.eq_1,
|
||||
mul_zero, sub_self]
|
||||
| _, States.inAsymp φ =>
|
||||
simp only [anPartF_negAsymp, map_zero, mul_zero, instCommGroup.eq_1, smul_zero, zero_mul,
|
||||
sub_self]
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [crPartF_position, anPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.outAsymp φ' =>
|
||||
simp only [crPartF_position, anPartF_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.inAsymp φ, States.position φ' =>
|
||||
simp only [crPartF_negAsymp, anPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.inAsymp φ, States.outAsymp φ' =>
|
||||
simp only [crPartF_negAsymp, anPartF_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma superCommuteF_crPartF_crPartF (φ φ' : 𝓕.States) :
|
||||
[crPartF φ, crPartF φ']ₛca = crPartF φ * crPartF φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • crPartF φ' * crPartF φ := by
|
||||
match φ, φ' with
|
||||
| States.outAsymp φ, _ =>
|
||||
simp only [crPartF_posAsymp, map_zero, LinearMap.zero_apply, zero_mul, instCommGroup.eq_1,
|
||||
mul_zero, sub_self]
|
||||
| _, States.outAsymp φ =>
|
||||
simp only [crPartF_posAsymp, map_zero, mul_zero, instCommGroup.eq_1, smul_zero, zero_mul,
|
||||
sub_self]
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [crPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.inAsymp φ' =>
|
||||
simp only [crPartF_position, crPartF_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.inAsymp φ, States.position φ' =>
|
||||
simp only [crPartF_negAsymp, crPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.inAsymp φ, States.inAsymp φ' =>
|
||||
simp only [crPartF_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma superCommuteF_anPartF_anPartF (φ φ' : 𝓕.States) :
|
||||
[anPartF φ, anPartF φ']ₛca =
|
||||
anPartF φ * anPartF φ' - 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • anPartF φ' * anPartF φ := by
|
||||
match φ, φ' with
|
||||
| States.inAsymp φ, _ =>
|
||||
simp
|
||||
| _, States.inAsymp φ =>
|
||||
simp
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [anPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.outAsymp φ' =>
|
||||
simp only [anPartF_position, anPartF_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.outAsymp φ, States.position φ' =>
|
||||
simp only [anPartF_posAsymp, anPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.outAsymp φ, States.outAsymp φ' =>
|
||||
simp only [anPartF_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma superCommuteF_crPartF_ofStateList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
[crPartF φ, ofStateList φs]ₛca =
|
||||
crPartF φ * ofStateList φs - 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) • ofStateList φs *
|
||||
crPartF φ := by
|
||||
match φ with
|
||||
| States.inAsymp φ =>
|
||||
simp only [crPartF_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
| States.position φ =>
|
||||
simp only [crPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
| States.outAsymp φ =>
|
||||
simp
|
||||
|
||||
lemma superCommuteF_anPartF_ofStateList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
[anPartF φ, ofStateList φs]ₛca =
|
||||
anPartF φ * ofStateList φs - 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) •
|
||||
ofStateList φs * anPartF φ := by
|
||||
match φ with
|
||||
| States.inAsymp φ =>
|
||||
simp
|
||||
| States.position φ =>
|
||||
simp only [anPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
| States.outAsymp φ =>
|
||||
simp only [anPartF_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
|
||||
lemma superCommuteF_crPartF_ofState (φ φ' : 𝓕.States) :
|
||||
[crPartF φ, ofState φ']ₛca =
|
||||
crPartF φ * ofState φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • ofState φ' * crPartF φ := by
|
||||
rw [← ofStateList_singleton, superCommuteF_crPartF_ofStateList]
|
||||
simp
|
||||
|
||||
lemma superCommuteF_anPartF_ofState (φ φ' : 𝓕.States) :
|
||||
[anPartF φ, ofState φ']ₛca =
|
||||
anPartF φ * ofState φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • ofState φ' * anPartF φ := by
|
||||
rw [← ofStateList_singleton, superCommuteF_anPartF_ofStateList]
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## Mul equal superCommuteF
|
||||
|
||||
Lemmas which rewrite a multiplication of two elements of the algebra as their commuted
|
||||
multiplication with a sign plus the super commutor.
|
||||
|
||||
-/
|
||||
lemma ofCrAnList_mul_ofCrAnList_eq_superCommuteF (φs φs' : List 𝓕.CrAnStates) :
|
||||
ofCrAnList φs * ofCrAnList φs' = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofCrAnList φs' * ofCrAnList φs
|
||||
+ [ofCrAnList φs, ofCrAnList φs']ₛca := by
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [ofCrAnList_append]
|
||||
|
||||
lemma ofCrAnState_mul_ofCrAnList_eq_superCommuteF (φ : 𝓕.CrAnStates) (φs' : List 𝓕.CrAnStates) :
|
||||
ofCrAnState φ * ofCrAnList φs' = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • ofCrAnList φs' * ofCrAnState φ
|
||||
+ [ofCrAnState φ, ofCrAnList φs']ₛca := by
|
||||
rw [← ofCrAnList_singleton, ofCrAnList_mul_ofCrAnList_eq_superCommuteF]
|
||||
simp
|
||||
|
||||
lemma ofStateList_mul_ofStateList_eq_superCommuteF (φs φs' : List 𝓕.States) :
|
||||
ofStateList φs * ofStateList φs' = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofStateList φs' * ofStateList φs
|
||||
+ [ofStateList φs, ofStateList φs']ₛca := by
|
||||
rw [superCommuteF_ofStateList_ofStatesList]
|
||||
simp
|
||||
|
||||
lemma ofState_mul_ofStateList_eq_superCommuteF (φ : 𝓕.States) (φs' : List 𝓕.States) :
|
||||
ofState φ * ofStateList φs' = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • ofStateList φs' * ofState φ
|
||||
+ [ofState φ, ofStateList φs']ₛca := by
|
||||
rw [superCommuteF_ofState_ofStatesList]
|
||||
simp
|
||||
|
||||
lemma ofStateList_mul_ofState_eq_superCommuteF (φs : List 𝓕.States) (φ : 𝓕.States) :
|
||||
ofStateList φs * ofState φ = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ) • ofState φ * ofStateList φs
|
||||
+ [ofStateList φs, ofState φ]ₛca := by
|
||||
rw [superCommuteF_ofStateList_ofState]
|
||||
simp
|
||||
|
||||
lemma crPartF_mul_anPartF_eq_superCommuteF (φ φ' : 𝓕.States) :
|
||||
crPartF φ * anPartF φ' =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • anPartF φ' * crPartF φ +
|
||||
[crPartF φ, anPartF φ']ₛca := by
|
||||
rw [superCommuteF_crPartF_anPartF]
|
||||
simp
|
||||
|
||||
lemma anPartF_mul_crPartF_eq_superCommuteF (φ φ' : 𝓕.States) :
|
||||
anPartF φ * crPartF φ' =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
crPartF φ' * anPartF φ +
|
||||
[anPartF φ, crPartF φ']ₛca := by
|
||||
rw [superCommuteF_anPartF_crPartF]
|
||||
simp
|
||||
|
||||
lemma crPartF_mul_crPartF_eq_superCommuteF (φ φ' : 𝓕.States) :
|
||||
crPartF φ * crPartF φ' =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • crPartF φ' * crPartF φ +
|
||||
[crPartF φ, crPartF φ']ₛca := by
|
||||
rw [superCommuteF_crPartF_crPartF]
|
||||
simp
|
||||
|
||||
lemma anPartF_mul_anPartF_eq_superCommuteF (φ φ' : 𝓕.States) :
|
||||
anPartF φ * anPartF φ' = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • anPartF φ' * anPartF φ +
|
||||
[anPartF φ, anPartF φ']ₛca := by
|
||||
rw [superCommuteF_anPartF_anPartF]
|
||||
simp
|
||||
|
||||
lemma ofCrAnList_mul_ofStateList_eq_superCommuteF (φs : List 𝓕.CrAnStates) (φs' : List 𝓕.States) :
|
||||
ofCrAnList φs * ofStateList φs' = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofStateList φs' * ofCrAnList φs
|
||||
+ [ofCrAnList φs, ofStateList φs']ₛca := by
|
||||
rw [superCommuteF_ofCrAnList_ofStatesList]
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## Symmetry of the super commutor.
|
||||
|
||||
-/
|
||||
|
||||
lemma superCommuteF_ofCrAnList_ofCrAnList_symm (φs φs' : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs, ofCrAnList φs']ₛca =
|
||||
(- 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs')) • [ofCrAnList φs', ofCrAnList φs]ₛca := by
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList, superCommuteF_ofCrAnList_ofCrAnList, smul_sub]
|
||||
simp only [instCommGroup.eq_1, neg_smul, sub_neg_eq_add]
|
||||
rw [smul_smul]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [exchangeSign_symm, exchangeSign_mul_self]
|
||||
simp only [one_smul]
|
||||
abel
|
||||
|
||||
lemma superCommuteF_ofCrAnState_ofCrAnState_symm (φ φ' : 𝓕.CrAnStates) :
|
||||
[ofCrAnState φ, ofCrAnState φ']ₛca =
|
||||
(- 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ')) • [ofCrAnState φ', ofCrAnState φ]ₛca := by
|
||||
rw [superCommuteF_ofCrAnState_ofCrAnState, superCommuteF_ofCrAnState_ofCrAnState]
|
||||
rw [smul_sub]
|
||||
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc, neg_smul, sub_neg_eq_add]
|
||||
rw [smul_smul]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [exchangeSign_symm, exchangeSign_mul_self]
|
||||
simp only [one_smul]
|
||||
abel
|
||||
|
||||
/-!
|
||||
|
||||
## Splitting the super commutor on lists into sums.
|
||||
|
||||
-/
|
||||
|
||||
lemma superCommuteF_ofCrAnList_ofCrAnList_cons (φ : 𝓕.CrAnStates) (φs φs' : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs, ofCrAnList (φ :: φs')]ₛca =
|
||||
[ofCrAnList φs, ofCrAnState φ]ₛca * ofCrAnList φs' +
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ)
|
||||
• ofCrAnState φ * [ofCrAnList φs, ofCrAnList φs']ₛca := by
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
conv_rhs =>
|
||||
lhs
|
||||
rw [← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList, sub_mul, ← ofCrAnList_append]
|
||||
rhs
|
||||
rw [FieldStatistic.ofList_singleton, ofCrAnList_append, ofCrAnList_singleton, smul_mul_assoc,
|
||||
mul_assoc, ← ofCrAnList_append]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList, mul_sub, smul_mul_assoc]
|
||||
simp only [instCommGroup.eq_1, List.cons_append, List.append_assoc, List.nil_append,
|
||||
Algebra.mul_smul_comm, Algebra.smul_mul_assoc, sub_add_sub_cancel, sub_right_inj]
|
||||
rw [← ofCrAnList_cons, smul_smul, FieldStatistic.ofList_cons_eq_mul]
|
||||
simp only [instCommGroup, map_mul, mul_comm]
|
||||
|
||||
lemma superCommuteF_ofCrAnList_ofStateList_cons (φ : 𝓕.States) (φs : List 𝓕.CrAnStates)
|
||||
(φs' : List 𝓕.States) : [ofCrAnList φs, ofStateList (φ :: φs')]ₛca =
|
||||
[ofCrAnList φs, ofState φ]ₛca * ofStateList φs' +
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ) • ofState φ * [ofCrAnList φs, ofStateList φs']ₛca := by
|
||||
rw [superCommuteF_ofCrAnList_ofStatesList]
|
||||
conv_rhs =>
|
||||
lhs
|
||||
rw [← ofStateList_singleton, superCommuteF_ofCrAnList_ofStatesList, sub_mul, mul_assoc,
|
||||
← ofStateList_append]
|
||||
rhs
|
||||
rw [FieldStatistic.ofList_singleton, ofStateList_singleton, smul_mul_assoc,
|
||||
smul_mul_assoc, mul_assoc]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [superCommuteF_ofCrAnList_ofStatesList, mul_sub, smul_mul_assoc]
|
||||
simp only [instCommGroup, Algebra.smul_mul_assoc, List.singleton_append, Algebra.mul_smul_comm,
|
||||
sub_add_sub_cancel, sub_right_inj]
|
||||
rw [ofStateList_cons, mul_assoc, smul_smul, FieldStatistic.ofList_cons_eq_mul]
|
||||
simp [mul_comm]
|
||||
|
||||
/--
|
||||
Within the creation and annihilation algebra, we have that
|
||||
`[φᶜᵃs, φᶜᵃ₀ … φᶜᵃₙ]ₛca = ∑ i, sᵢ • φᶜᵃs₀ … φᶜᵃᵢ₋₁ * [φᶜᵃs, φᶜᵃᵢ]ₛca * φᶜᵃᵢ₊₁ … φᶜᵃₙ`
|
||||
where `sᵢ` is the exchange sign for `φᶜᵃs` and `φᶜᵃs₀ … φᶜᵃᵢ₋₁`.
|
||||
-/
|
||||
lemma superCommuteF_ofCrAnList_ofCrAnList_eq_sum (φs : List 𝓕.CrAnStates) :
|
||||
(φs' : List 𝓕.CrAnStates) → [ofCrAnList φs, ofCrAnList φs']ₛca =
|
||||
∑ (n : Fin φs'.length), 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs'.take n) •
|
||||
ofCrAnList (φs'.take n) * [ofCrAnList φs, ofCrAnState (φs'.get n)]ₛca *
|
||||
ofCrAnList (φs'.drop (n + 1))
|
||||
| [] => by
|
||||
simp [← ofCrAnList_nil, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
| φ :: φs' => by
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList_cons, superCommuteF_ofCrAnList_ofCrAnList_eq_sum φs φs']
|
||||
conv_rhs => erw [Fin.sum_univ_succ]
|
||||
congr 1
|
||||
· simp
|
||||
· simp [Finset.mul_sum, smul_smul, ofCrAnList_cons, mul_assoc,
|
||||
FieldStatistic.ofList_cons_eq_mul, mul_comm]
|
||||
|
||||
lemma superCommuteF_ofCrAnList_ofStateList_eq_sum (φs : List 𝓕.CrAnStates) : (φs' : List 𝓕.States) →
|
||||
[ofCrAnList φs, ofStateList φs']ₛca =
|
||||
∑ (n : Fin φs'.length), 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs'.take n) •
|
||||
ofStateList (φs'.take n) * [ofCrAnList φs, ofState (φs'.get n)]ₛca *
|
||||
ofStateList (φs'.drop (n + 1))
|
||||
| [] => by
|
||||
simp only [superCommuteF_ofCrAnList_ofStatesList, instCommGroup, ofList_empty,
|
||||
exchangeSign_bosonic, one_smul, List.length_nil, Finset.univ_eq_empty, List.take_nil,
|
||||
List.get_eq_getElem, List.drop_nil, Finset.sum_empty]
|
||||
simp
|
||||
| φ :: φs' => by
|
||||
rw [superCommuteF_ofCrAnList_ofStateList_cons,
|
||||
superCommuteF_ofCrAnList_ofStateList_eq_sum φs φs']
|
||||
conv_rhs => erw [Fin.sum_univ_succ]
|
||||
congr 1
|
||||
· simp
|
||||
· simp [Finset.mul_sum, smul_smul, ofStateList_cons, mul_assoc,
|
||||
FieldStatistic.ofList_cons_eq_mul, mul_comm]
|
||||
|
||||
lemma summerCommute_jacobi_ofCrAnList (φs1 φs2 φs3 : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs1, [ofCrAnList φs2, ofCrAnList φs3]ₛca]ₛca =
|
||||
𝓢(𝓕 |>ₛ φs1, 𝓕 |>ₛ φs3) •
|
||||
(- 𝓢(𝓕 |>ₛ φs2, 𝓕 |>ₛ φs3) • [ofCrAnList φs3, [ofCrAnList φs1, ofCrAnList φs2]ₛca]ₛca -
|
||||
𝓢(𝓕 |>ₛ φs1, 𝓕 |>ₛ φs2) • [ofCrAnList φs2, [ofCrAnList φs3, ofCrAnList φs1]ₛca]ₛca) := by
|
||||
repeat rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp only [instCommGroup, map_sub, map_smul, neg_smul]
|
||||
repeat rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp only [instCommGroup.eq_1, ofList_append_eq_mul, List.append_assoc]
|
||||
by_cases h1 : (𝓕 |>ₛ φs1) = bosonic <;>
|
||||
by_cases h2 : (𝓕 |>ₛ φs2) = bosonic <;>
|
||||
by_cases h3 : (𝓕 |>ₛ φs3) = bosonic
|
||||
· simp only [h1, h2, h3, mul_self, bosonic_exchangeSign, one_smul, exchangeSign_bosonic, neg_sub]
|
||||
abel
|
||||
· simp only [h1, h2, bosonic_exchangeSign, one_smul, mul_bosonic, mul_self, map_one,
|
||||
exchangeSign_bosonic, neg_sub]
|
||||
abel
|
||||
· simp only [h1, h3, mul_bosonic, bosonic_exchangeSign, one_smul, exchangeSign_bosonic, neg_sub,
|
||||
mul_self, map_one]
|
||||
abel
|
||||
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
|
||||
simp only [h1, h2, h3, mul_self, bosonic_exchangeSign, one_smul,
|
||||
fermionic_exchangeSign_fermionic, neg_smul, neg_sub, bosonic_mul_fermionic, sub_neg_eq_add,
|
||||
mul_bosonic, smul_add, exchangeSign_bosonic]
|
||||
abel
|
||||
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
|
||||
simp only [h1, h2, h3, mul_self, map_one, one_smul, exchangeSign_bosonic, mul_bosonic,
|
||||
bosonic_exchangeSign, bosonic_mul_fermionic, neg_sub]
|
||||
abel
|
||||
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
|
||||
simp only [h1, h2, h3, bosonic_mul_fermionic, fermionic_exchangeSign_fermionic, neg_smul,
|
||||
one_smul, sub_neg_eq_add, bosonic_exchangeSign, mul_bosonic, smul_add, exchangeSign_bosonic,
|
||||
neg_sub, mul_self]
|
||||
abel
|
||||
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
|
||||
simp only [h1, h2, h3, mul_bosonic, fermionic_exchangeSign_fermionic, neg_smul, one_smul,
|
||||
sub_neg_eq_add, exchangeSign_bosonic, bosonic_mul_fermionic, smul_add, mul_self,
|
||||
bosonic_exchangeSign, neg_sub]
|
||||
abel
|
||||
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
|
||||
simp only [h1, h2, h3, mul_self, map_one, one_smul, fermionic_exchangeSign_fermionic, neg_smul,
|
||||
neg_sub]
|
||||
abel
|
||||
|
||||
/-!
|
||||
|
||||
## Interaction with grading.
|
||||
|
||||
-/
|
||||
|
||||
lemma superCommuteF_grade {a b : 𝓕.FieldOpFreeAlgebra} {f1 f2 : FieldStatistic}
|
||||
(ha : a ∈ statisticSubmodule f1) (hb : b ∈ statisticSubmodule f2) :
|
||||
[a, b]ₛca ∈ statisticSubmodule (f1 + f2) := by
|
||||
let p (a2 : 𝓕.FieldOpFreeAlgebra) (hx : a2 ∈ statisticSubmodule f2) : Prop :=
|
||||
[a, a2]ₛca ∈ statisticSubmodule (f1 + f2)
|
||||
change p b hb
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl, hφs⟩ := hx
|
||||
simp only [add_eq_mul, instCommGroup, p]
|
||||
let p (a2 : 𝓕.FieldOpFreeAlgebra) (hx : a2 ∈ statisticSubmodule f1) : Prop :=
|
||||
[a2, ofCrAnList φs]ₛca ∈ statisticSubmodule (f1 + f2)
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [add_eq_mul, instCommGroup, p]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
apply Submodule.sub_mem _
|
||||
· apply ofCrAnList_mem_statisticSubmodule_of
|
||||
rw [ofList_append_eq_mul, hφs, hφs']
|
||||
· apply Submodule.smul_mem
|
||||
apply ofCrAnList_mem_statisticSubmodule_of
|
||||
rw [ofList_append_eq_mul, hφs, hφs']
|
||||
rw [mul_comm]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp only [add_eq_mul, instCommGroup, map_add, LinearMap.add_apply, p]
|
||||
exact Submodule.add_mem _ hp1 hp2
|
||||
· intro c x hx hp1
|
||||
simp only [add_eq_mul, instCommGroup, map_smul, LinearMap.smul_apply, p]
|
||||
exact Submodule.smul_mem _ c hp1
|
||||
· exact ha
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp only [add_eq_mul, instCommGroup, map_add, p]
|
||||
exact Submodule.add_mem _ hp1 hp2
|
||||
· intro c x hx hp1
|
||||
simp only [add_eq_mul, instCommGroup, map_smul, p]
|
||||
exact Submodule.smul_mem _ c hp1
|
||||
· exact hb
|
||||
|
||||
lemma superCommuteF_bosonic_bosonic {a b : 𝓕.FieldOpFreeAlgebra}
|
||||
(ha : a ∈ statisticSubmodule bosonic) (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
let p (a2 : 𝓕.FieldOpFreeAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop :=
|
||||
[a, a2]ₛca = a * a2 - a2 * a
|
||||
change p b hb
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl, hφs⟩ := hx
|
||||
let p (a2 : 𝓕.FieldOpFreeAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop :=
|
||||
[a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact ha
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, mul_add, add_mul]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommuteF_bosonic_fermionic {a b : 𝓕.FieldOpFreeAlgebra}
|
||||
(ha : a ∈ statisticSubmodule bosonic) (hb : b ∈ statisticSubmodule fermionic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
let p (a2 : 𝓕.FieldOpFreeAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop :=
|
||||
[a, a2]ₛca = a * a2 - a2 * a
|
||||
change p b hb
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl, hφs⟩ := hx
|
||||
let p (a2 : 𝓕.FieldOpFreeAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop :=
|
||||
[a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, hφs', ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact ha
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, mul_add, add_mul]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommuteF_fermionic_bonsonic {a b : 𝓕.FieldOpFreeAlgebra}
|
||||
(ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
let p (a2 : 𝓕.FieldOpFreeAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop :=
|
||||
[a, a2]ₛca = a * a2 - a2 * a
|
||||
change p b hb
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl, hφs⟩ := hx
|
||||
let p (a2 : 𝓕.FieldOpFreeAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop :=
|
||||
[a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, hφs', ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact ha
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [map_add, mul_add, add_mul, p]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommuteF_bonsonic {a b : 𝓕.FieldOpFreeAlgebra} (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
rw [← bosonicProj_add_fermionicProj a]
|
||||
simp only [map_add, LinearMap.add_apply]
|
||||
rw [superCommuteF_bosonic_bosonic (by simp) hb, superCommuteF_fermionic_bonsonic (by simp) hb]
|
||||
simp only [add_mul, mul_add]
|
||||
abel
|
||||
|
||||
lemma bosonic_superCommuteF {a b : 𝓕.FieldOpFreeAlgebra} (ha : a ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
rw [← bosonicProj_add_fermionicProj b]
|
||||
simp only [map_add, LinearMap.add_apply]
|
||||
rw [superCommuteF_bosonic_bosonic ha (by simp), superCommuteF_bosonic_fermionic ha (by simp)]
|
||||
simp only [add_mul, mul_add]
|
||||
abel
|
||||
|
||||
lemma superCommuteF_bonsonic_symm {a b : 𝓕.FieldOpFreeAlgebra} (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = - [b, a]ₛca := by
|
||||
rw [bosonic_superCommuteF hb, superCommuteF_bonsonic hb]
|
||||
simp
|
||||
|
||||
lemma bonsonic_superCommuteF_symm {a b : 𝓕.FieldOpFreeAlgebra} (ha : a ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = - [b, a]ₛca := by
|
||||
rw [bosonic_superCommuteF ha, superCommuteF_bonsonic ha]
|
||||
simp
|
||||
|
||||
lemma superCommuteF_fermionic_fermionic {a b : 𝓕.FieldOpFreeAlgebra}
|
||||
(ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule fermionic) :
|
||||
[a, b]ₛca = a * b + b * a := by
|
||||
let p (a2 : 𝓕.FieldOpFreeAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop :=
|
||||
[a, a2]ₛca = a * a2 + a2 * a
|
||||
change p b hb
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl, hφs⟩ := hx
|
||||
let p (a2 : 𝓕.FieldOpFreeAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop :=
|
||||
[a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs + ofCrAnList φs * a2
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, hφs', ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact ha
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [map_add, mul_add, add_mul, p]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommuteF_fermionic_fermionic_symm {a b : 𝓕.FieldOpFreeAlgebra}
|
||||
(ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule fermionic) :
|
||||
[a, b]ₛca = [b, a]ₛca := by
|
||||
rw [superCommuteF_fermionic_fermionic ha hb]
|
||||
rw [superCommuteF_fermionic_fermionic hb ha]
|
||||
abel
|
||||
|
||||
lemma superCommuteF_expand_bosonicProj_fermionicProj (a b : 𝓕.FieldOpFreeAlgebra) :
|
||||
[a, b]ₛca = bosonicProj a * bosonicProj b - bosonicProj b * bosonicProj a +
|
||||
bosonicProj a * fermionicProj b - fermionicProj b * bosonicProj a +
|
||||
fermionicProj a * bosonicProj b - bosonicProj b * fermionicProj a +
|
||||
fermionicProj a * fermionicProj b + fermionicProj b * fermionicProj a := by
|
||||
conv_lhs => rw [← bosonicProj_add_fermionicProj a, ← bosonicProj_add_fermionicProj b]
|
||||
simp only [map_add, LinearMap.add_apply]
|
||||
rw [superCommuteF_bonsonic (by simp),
|
||||
superCommuteF_fermionic_bonsonic (by simp) (by simp),
|
||||
superCommuteF_bosonic_fermionic (by simp) (by simp),
|
||||
superCommuteF_fermionic_fermionic (by simp) (by simp)]
|
||||
abel
|
||||
|
||||
lemma superCommuteF_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule bosonic ∨
|
||||
[ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule fermionic := by
|
||||
by_cases h1 : (𝓕 |>ₛ φs) = bosonic <;> by_cases h2 : (𝓕 |>ₛ φs') = bosonic
|
||||
· left
|
||||
have h : bosonic = bosonic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h1
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h2
|
||||
· right
|
||||
have h : fermionic = bosonic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h1
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h2)
|
||||
· right
|
||||
have h : fermionic = fermionic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h1)
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h2
|
||||
· left
|
||||
have h : bosonic = fermionic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h1)
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h2)
|
||||
|
||||
lemma superCommuteF_ofCrAnState_ofCrAnState_bosonic_or_fermionic (φ φ' : 𝓕.CrAnStates) :
|
||||
[ofCrAnState φ, ofCrAnState φ']ₛca ∈ statisticSubmodule bosonic ∨
|
||||
[ofCrAnState φ, ofCrAnState φ']ₛca ∈ statisticSubmodule fermionic := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
exact superCommuteF_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ] [φ']
|
||||
|
||||
lemma superCommuteF_superCommuteF_ofCrAnState_bosonic_or_fermionic (φ1 φ2 φ3 : 𝓕.CrAnStates) :
|
||||
[ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca ∈ statisticSubmodule bosonic ∨
|
||||
[ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca ∈ statisticSubmodule fermionic := by
|
||||
rcases superCommuteF_ofCrAnState_ofCrAnState_bosonic_or_fermionic φ2 φ3 with hs | hs
|
||||
<;> rcases ofCrAnState_bosonic_or_fermionic φ1 with h1 | h1
|
||||
· left
|
||||
have h : bosonic = bosonic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommuteF_grade h1 hs
|
||||
· right
|
||||
have h : fermionic = fermionic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommuteF_grade h1 hs
|
||||
· right
|
||||
have h : fermionic = bosonic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommuteF_grade h1 hs
|
||||
· left
|
||||
have h : bosonic = fermionic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommuteF_grade h1 hs
|
||||
|
||||
lemma superCommuteF_bosonic_ofCrAnList_eq_sum (a : 𝓕.FieldOpFreeAlgebra) (φs : List 𝓕.CrAnStates)
|
||||
(ha : a ∈ statisticSubmodule bosonic) :
|
||||
[a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length),
|
||||
ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca *
|
||||
ofCrAnList (φs.drop (n + 1)) := by
|
||||
let p (a : 𝓕.FieldOpFreeAlgebra) (ha : a ∈ statisticSubmodule bosonic) : Prop :=
|
||||
[a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length),
|
||||
ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca *
|
||||
ofCrAnList (φs.drop (n + 1))
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro a ha
|
||||
obtain ⟨φs, rfl, hφs⟩ := ha
|
||||
simp only [List.get_eq_getElem, p]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList_eq_sum]
|
||||
congr
|
||||
funext n
|
||||
simp [hφs]
|
||||
· simp [p]
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all only [List.get_eq_getElem, map_add, LinearMap.add_apply, p]
|
||||
rw [← Finset.sum_add_distrib]
|
||||
congr
|
||||
funext n
|
||||
simp [mul_add, add_mul]
|
||||
· intro c x hx hpx
|
||||
simp_all [p, Finset.smul_sum]
|
||||
· exact ha
|
||||
|
||||
lemma superCommuteF_fermionic_ofCrAnList_eq_sum (a : 𝓕.FieldOpFreeAlgebra) (φs : List 𝓕.CrAnStates)
|
||||
(ha : a ∈ statisticSubmodule fermionic) :
|
||||
[a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), 𝓢(fermionic, 𝓕 |>ₛ φs.take n) •
|
||||
ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca *
|
||||
ofCrAnList (φs.drop (n + 1)) := by
|
||||
let p (a : 𝓕.FieldOpFreeAlgebra) (ha : a ∈ statisticSubmodule fermionic) : Prop :=
|
||||
[a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), 𝓢(fermionic, 𝓕 |>ₛ φs.take n) •
|
||||
ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca *
|
||||
ofCrAnList (φs.drop (n + 1))
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro a ha
|
||||
obtain ⟨φs, rfl, hφs⟩ := ha
|
||||
simp only [instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, p]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList_eq_sum]
|
||||
congr
|
||||
funext n
|
||||
simp [hφs]
|
||||
· simp [p]
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all only [p, instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, map_add,
|
||||
LinearMap.add_apply]
|
||||
rw [← Finset.sum_add_distrib]
|
||||
congr
|
||||
funext n
|
||||
simp [mul_add, add_mul]
|
||||
· intro c x hx hpx
|
||||
simp_all only [p, instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, map_smul,
|
||||
LinearMap.smul_apply, Finset.smul_sum, Algebra.mul_smul_comm]
|
||||
congr
|
||||
funext x
|
||||
simp [smul_smul, mul_comm]
|
||||
· exact ha
|
||||
|
||||
lemma statistic_neq_of_superCommuteF_fermionic {φs φs' : List 𝓕.CrAnStates}
|
||||
(h : [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule fermionic) :
|
||||
(𝓕 |>ₛ φs) ≠ (𝓕 |>ₛ φs') ∨ [ofCrAnList φs, ofCrAnList φs']ₛca = 0 := by
|
||||
by_cases h0 : [ofCrAnList φs, ofCrAnList φs']ₛca = 0
|
||||
· simp [h0]
|
||||
simp only [ne_eq, h0, or_false]
|
||||
by_contra hn
|
||||
refine h0 (eq_zero_of_bosonic_and_fermionic ?_ h)
|
||||
by_cases hc : (𝓕 |>ₛ φs) = bosonic
|
||||
· have h1 : bosonic = bosonic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h1]
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ hc
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _
|
||||
rw [← hn, hc]
|
||||
· have h1 : bosonic = fermionic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h1]
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _
|
||||
simpa using hc
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _
|
||||
rw [← hn]
|
||||
simpa using hc
|
||||
|
||||
end FieldOpFreeAlgebra
|
||||
|
||||
end FieldSpecification
|
|
@ -0,0 +1,366 @@
|
|||
/-
|
||||
Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.PerturbationTheory.FieldSpecification.TimeOrder
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpFreeAlgebra.SuperCommute
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
/-!
|
||||
|
||||
# Time Ordering in the FieldOpFreeAlgebra
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldSpecification
|
||||
variable {𝓕 : FieldSpecification}
|
||||
open FieldStatistic
|
||||
|
||||
namespace FieldOpFreeAlgebra
|
||||
|
||||
noncomputable section
|
||||
open HepLean.List
|
||||
/-!
|
||||
|
||||
## Time order
|
||||
|
||||
-/
|
||||
|
||||
/-- Time ordering for the `FieldOpFreeAlgebra`. -/
|
||||
def timeOrderF : FieldOpFreeAlgebra 𝓕 →ₗ[ℂ] FieldOpFreeAlgebra 𝓕 :=
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs =>
|
||||
crAnTimeOrderSign φs • ofCrAnList (crAnTimeOrderList φs)
|
||||
|
||||
@[inherit_doc timeOrderF]
|
||||
scoped[FieldSpecification.FieldOpFreeAlgebra] notation "𝓣ᶠ(" a ")" => timeOrderF a
|
||||
|
||||
lemma timeOrderF_ofCrAnList (φs : List 𝓕.CrAnStates) :
|
||||
𝓣ᶠ(ofCrAnList φs) = crAnTimeOrderSign φs • ofCrAnList (crAnTimeOrderList φs) := by
|
||||
rw [← ofListBasis_eq_ofList]
|
||||
simp only [timeOrderF, Basis.constr_basis]
|
||||
|
||||
lemma timeOrderF_timeOrderF_mid (a b c : 𝓕.FieldOpFreeAlgebra) : 𝓣ᶠ(a * b * c) = 𝓣ᶠ(a * 𝓣ᶠ(b) * c) := by
|
||||
let pc (c : 𝓕.FieldOpFreeAlgebra) (hc : c ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓣ᶠ(a * b * c) = 𝓣ᶠ(a * 𝓣ᶠ(b) * c)
|
||||
change pc c (Basis.mem_span _ c)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pc]
|
||||
let pb (b : 𝓕.FieldOpFreeAlgebra) (hb : b ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓣ᶠ(a * b * ofCrAnList φs) = 𝓣ᶠ(a * 𝓣ᶠ(b) * ofCrAnList φs)
|
||||
change pb b (Basis.mem_span _ b)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pb]
|
||||
let pa (a : 𝓕.FieldOpFreeAlgebra) (ha : a ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓣ᶠ(a * ofCrAnList φs' * ofCrAnList φs) = 𝓣ᶠ(a * 𝓣ᶠ(ofCrAnList φs') * ofCrAnList φs)
|
||||
change pa a (Basis.mem_span _ a)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs'', rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pa]
|
||||
rw [timeOrderF_ofCrAnList]
|
||||
simp only [← ofCrAnList_append, Algebra.mul_smul_comm,
|
||||
Algebra.smul_mul_assoc, map_smul]
|
||||
rw [timeOrderF_ofCrAnList, timeOrderF_ofCrAnList, smul_smul]
|
||||
congr 1
|
||||
· simp only [crAnTimeOrderSign, crAnTimeOrderList]
|
||||
rw [Wick.koszulSign_of_append_eq_insertionSort, mul_comm]
|
||||
· congr 1
|
||||
simp only [crAnTimeOrderList]
|
||||
rw [insertionSort_append_insertionSort_append]
|
||||
· simp [pa]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pa, add_mul]
|
||||
· intro x hx h
|
||||
simp_all [pa]
|
||||
· simp [pb]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pb, mul_add, add_mul]
|
||||
· intro x hx h
|
||||
simp_all [pb]
|
||||
· simp [pc]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pc, mul_add]
|
||||
· intro x hx h hp
|
||||
simp_all [pc]
|
||||
|
||||
lemma timeOrderF_timeOrderF_right (a b : 𝓕.FieldOpFreeAlgebra) : 𝓣ᶠ(a * b) = 𝓣ᶠ(a * 𝓣ᶠ(b)) := by
|
||||
trans 𝓣ᶠ(a * b * 1)
|
||||
· simp
|
||||
· rw [timeOrderF_timeOrderF_mid]
|
||||
simp
|
||||
|
||||
lemma timeOrderF_timeOrderF_left (a b : 𝓕.FieldOpFreeAlgebra) : 𝓣ᶠ(a * b) = 𝓣ᶠ(𝓣ᶠ(a) * b) := by
|
||||
trans 𝓣ᶠ(1 * a * b)
|
||||
· simp
|
||||
· rw [timeOrderF_timeOrderF_mid]
|
||||
simp
|
||||
|
||||
lemma timeOrderF_ofStateList (φs : List 𝓕.States) :
|
||||
𝓣ᶠ(ofStateList φs) = timeOrderSign φs • ofStateList (timeOrderList φs) := by
|
||||
conv_lhs =>
|
||||
rw [ofStateList_sum, map_sum]
|
||||
enter [2, x]
|
||||
rw [timeOrderF_ofCrAnList]
|
||||
simp only [crAnTimeOrderSign_crAnSection]
|
||||
rw [← Finset.smul_sum]
|
||||
congr
|
||||
rw [ofStateList_sum, sum_crAnSections_timeOrder]
|
||||
rfl
|
||||
|
||||
lemma timeOrderF_ofStateList_nil : timeOrderF (𝓕 := 𝓕) (ofStateList []) = 1 := by
|
||||
rw [timeOrderF_ofStateList]
|
||||
simp [timeOrderSign, Wick.koszulSign, timeOrderList]
|
||||
|
||||
@[simp]
|
||||
lemma timeOrderF_ofStateList_singleton (φ : 𝓕.States) : 𝓣ᶠ(ofStateList [φ]) = ofStateList [φ] := by
|
||||
simp [timeOrderF_ofStateList, timeOrderSign, timeOrderList]
|
||||
|
||||
lemma timeOrderF_ofState_ofState_ordered {φ ψ : 𝓕.States} (h : timeOrderRel φ ψ) :
|
||||
𝓣ᶠ(ofState φ * ofState ψ) = ofState φ * ofState ψ := by
|
||||
rw [← ofStateList_singleton, ← ofStateList_singleton, ← ofStateList_append,
|
||||
timeOrderF_ofStateList]
|
||||
simp only [List.singleton_append]
|
||||
rw [timeOrderSign_pair_ordered h, timeOrderList_pair_ordered h]
|
||||
simp
|
||||
|
||||
lemma timeOrderF_ofState_ofState_not_ordered {φ ψ : 𝓕.States} (h : ¬ timeOrderRel φ ψ) :
|
||||
𝓣ᶠ(ofState φ * ofState ψ) = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ψ) • ofState ψ * ofState φ := by
|
||||
rw [← ofStateList_singleton, ← ofStateList_singleton,
|
||||
← ofStateList_append, timeOrderF_ofStateList]
|
||||
simp only [List.singleton_append, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [timeOrderSign_pair_not_ordered h, timeOrderList_pair_not_ordered h]
|
||||
simp [← ofStateList_append]
|
||||
|
||||
lemma timeOrderF_ofState_ofState_not_ordered_eq_timeOrderF {φ ψ : 𝓕.States}
|
||||
(h : ¬ timeOrderRel φ ψ) :
|
||||
𝓣ᶠ(ofState φ * ofState ψ) = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ψ) • 𝓣ᶠ(ofState ψ * ofState φ) := by
|
||||
rw [timeOrderF_ofState_ofState_not_ordered h]
|
||||
rw [timeOrderF_ofState_ofState_ordered]
|
||||
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
have hx := IsTotal.total (r := timeOrderRel) ψ φ
|
||||
simp_all
|
||||
|
||||
lemma timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel
|
||||
{φ ψ : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ ψ) :
|
||||
𝓣ᶠ([ofCrAnState φ, ofCrAnState ψ]ₛca) = 0 := by
|
||||
rw [superCommuteF_ofCrAnState_ofCrAnState]
|
||||
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc, map_sub, map_smul]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton,
|
||||
← ofCrAnList_append, ← ofCrAnList_append, timeOrderF_ofCrAnList, timeOrderF_ofCrAnList]
|
||||
simp only [List.singleton_append]
|
||||
rw [crAnTimeOrderSign_pair_not_ordered h, crAnTimeOrderList_pair_not_ordered h]
|
||||
rw [sub_eq_zero, smul_smul]
|
||||
have h1 := IsTotal.total (r := crAnTimeOrderRel) φ ψ
|
||||
congr
|
||||
· rw [crAnTimeOrderSign_pair_ordered, exchangeSign_symm]
|
||||
simp only [instCommGroup.eq_1, mul_one]
|
||||
simp_all
|
||||
· rw [crAnTimeOrderList_pair_ordered]
|
||||
simp_all
|
||||
|
||||
lemma timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right
|
||||
{φ ψ : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ ψ) (a : 𝓕.FieldOpFreeAlgebra) :
|
||||
𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca) = 0 := by
|
||||
rw [timeOrderF_timeOrderF_right,
|
||||
timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel h]
|
||||
simp
|
||||
|
||||
lemma timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left
|
||||
{φ ψ : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ ψ) (a : 𝓕.FieldOpFreeAlgebra) :
|
||||
𝓣ᶠ([ofCrAnState φ, ofCrAnState ψ]ₛca * a) = 0 := by
|
||||
rw [timeOrderF_timeOrderF_left,
|
||||
timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel h]
|
||||
simp
|
||||
|
||||
lemma timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_mid
|
||||
{φ ψ : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ ψ) (a b : 𝓕.FieldOpFreeAlgebra) :
|
||||
𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = 0 := by
|
||||
rw [timeOrderF_timeOrderF_mid,
|
||||
timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel h]
|
||||
simp
|
||||
|
||||
lemma timeOrderF_superCommuteF_superCommuteF_ofCrAnState_not_crAnTimeOrderRel
|
||||
{φ1 φ2 : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ1 φ2) (a : 𝓕.FieldOpFreeAlgebra) :
|
||||
𝓣ᶠ([a, [ofCrAnState φ1, ofCrAnState φ2]ₛca]ₛca) = 0 := by
|
||||
rw [← bosonicProj_add_fermionicProj a]
|
||||
simp only [map_add, LinearMap.add_apply]
|
||||
rw [bosonic_superCommuteF (Submodule.coe_mem (bosonicProj a))]
|
||||
simp only [map_sub]
|
||||
rw [timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left h]
|
||||
rw [timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right h]
|
||||
simp only [sub_self, zero_add]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rcases superCommuteF_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ1] [φ2] with h' | h'
|
||||
· rw [superCommuteF_bonsonic h']
|
||||
simp only [ofCrAnList_singleton, map_sub]
|
||||
rw [timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left h]
|
||||
rw [timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right h]
|
||||
simp
|
||||
· rw [superCommuteF_fermionic_fermionic (Submodule.coe_mem (fermionicProj a)) h']
|
||||
simp only [ofCrAnList_singleton, map_add]
|
||||
rw [timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left h]
|
||||
rw [timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right h]
|
||||
simp
|
||||
|
||||
lemma timeOrderF_superCommuteF_ofCrAnState_superCommuteF_not_crAnTimeOrderRel
|
||||
{φ1 φ2 φ3 : 𝓕.CrAnStates} (h12 : ¬ crAnTimeOrderRel φ1 φ2)
|
||||
(h13 : ¬ crAnTimeOrderRel φ1 φ3) :
|
||||
𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rw [summerCommute_jacobi_ofCrAnList]
|
||||
simp only [instCommGroup.eq_1, ofList_singleton, ofCrAnList_singleton, neg_smul, map_smul,
|
||||
map_sub, map_neg, smul_eq_zero]
|
||||
right
|
||||
rw [timeOrderF_superCommuteF_superCommuteF_ofCrAnState_not_crAnTimeOrderRel h12]
|
||||
rw [superCommuteF_ofCrAnState_ofCrAnState_symm φ3]
|
||||
simp only [smul_zero, neg_zero, instCommGroup.eq_1, neg_smul, map_neg, map_smul, smul_neg,
|
||||
sub_neg_eq_add, zero_add, smul_eq_zero]
|
||||
rw [timeOrderF_superCommuteF_superCommuteF_ofCrAnState_not_crAnTimeOrderRel h13]
|
||||
simp
|
||||
|
||||
lemma timeOrderF_superCommuteF_ofCrAnState_superCommuteF_not_crAnTimeOrderRel'
|
||||
{φ1 φ2 φ3 : 𝓕.CrAnStates} (h12 : ¬ crAnTimeOrderRel φ2 φ1)
|
||||
(h13 : ¬ crAnTimeOrderRel φ3 φ1) :
|
||||
𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rw [summerCommute_jacobi_ofCrAnList]
|
||||
simp only [instCommGroup.eq_1, ofList_singleton, ofCrAnList_singleton, neg_smul, map_smul,
|
||||
map_sub, map_neg, smul_eq_zero]
|
||||
right
|
||||
rw [superCommuteF_ofCrAnState_ofCrAnState_symm φ1]
|
||||
simp only [instCommGroup.eq_1, neg_smul, map_neg, map_smul, smul_neg, neg_neg]
|
||||
rw [timeOrderF_superCommuteF_superCommuteF_ofCrAnState_not_crAnTimeOrderRel h12]
|
||||
simp only [smul_zero, zero_sub, neg_eq_zero, smul_eq_zero]
|
||||
rw [timeOrderF_superCommuteF_superCommuteF_ofCrAnState_not_crAnTimeOrderRel h13]
|
||||
simp
|
||||
|
||||
lemma timeOrderF_superCommuteF_ofCrAnState_superCommuteF_all_not_crAnTimeOrderRel
|
||||
(φ1 φ2 φ3 : 𝓕.CrAnStates) (h : ¬
|
||||
(crAnTimeOrderRel φ1 φ2 ∧ crAnTimeOrderRel φ1 φ3 ∧
|
||||
crAnTimeOrderRel φ2 φ1 ∧ crAnTimeOrderRel φ2 φ3 ∧
|
||||
crAnTimeOrderRel φ3 φ1 ∧ crAnTimeOrderRel φ3 φ2)) :
|
||||
𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by
|
||||
simp only [not_and] at h
|
||||
by_cases h23 : ¬ crAnTimeOrderRel φ2 φ3
|
||||
· simp_all only [IsEmpty.forall_iff, implies_true]
|
||||
rw [timeOrderF_superCommuteF_superCommuteF_ofCrAnState_not_crAnTimeOrderRel h23]
|
||||
simp_all only [Decidable.not_not, forall_const]
|
||||
by_cases h32 : ¬ crAnTimeOrderRel φ3 φ2
|
||||
· simp_all only [not_false_eq_true, implies_true]
|
||||
rw [superCommuteF_ofCrAnState_ofCrAnState_symm]
|
||||
simp only [instCommGroup.eq_1, neg_smul, map_neg, map_smul, neg_eq_zero, smul_eq_zero]
|
||||
rw [timeOrderF_superCommuteF_superCommuteF_ofCrAnState_not_crAnTimeOrderRel h32]
|
||||
simp
|
||||
simp_all only [imp_false, Decidable.not_not]
|
||||
by_cases h12 : ¬ crAnTimeOrderRel φ1 φ2
|
||||
· have h13 : ¬ crAnTimeOrderRel φ1 φ3 := by
|
||||
intro h13
|
||||
apply h12
|
||||
exact IsTrans.trans φ1 φ3 φ2 h13 h32
|
||||
rw [timeOrderF_superCommuteF_ofCrAnState_superCommuteF_not_crAnTimeOrderRel h12 h13]
|
||||
simp_all only [Decidable.not_not, forall_const]
|
||||
have h13 : crAnTimeOrderRel φ1 φ3 := IsTrans.trans φ1 φ2 φ3 h12 h23
|
||||
simp_all only [forall_const]
|
||||
by_cases h21 : ¬ crAnTimeOrderRel φ2 φ1
|
||||
· simp_all only [IsEmpty.forall_iff]
|
||||
have h31 : ¬ crAnTimeOrderRel φ3 φ1 := by
|
||||
intro h31
|
||||
apply h21
|
||||
exact IsTrans.trans φ2 φ3 φ1 h23 h31
|
||||
rw [timeOrderF_superCommuteF_ofCrAnState_superCommuteF_not_crAnTimeOrderRel' h21 h31]
|
||||
simp_all only [Decidable.not_not, forall_const]
|
||||
refine False.elim (h ?_)
|
||||
exact IsTrans.trans φ3 φ2 φ1 h32 h21
|
||||
|
||||
lemma timeOrderF_superCommuteF_ofCrAnState_ofCrAnState_eq_time
|
||||
{φ ψ : 𝓕.CrAnStates} (h1 : crAnTimeOrderRel φ ψ) (h2 : crAnTimeOrderRel ψ φ) :
|
||||
𝓣ᶠ([ofCrAnState φ, ofCrAnState ψ]ₛca) = [ofCrAnState φ, ofCrAnState ψ]ₛca := by
|
||||
rw [superCommuteF_ofCrAnState_ofCrAnState]
|
||||
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc, map_sub, map_smul]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton,
|
||||
← ofCrAnList_append, ← ofCrAnList_append, timeOrderF_ofCrAnList, timeOrderF_ofCrAnList]
|
||||
simp only [List.singleton_append]
|
||||
rw [crAnTimeOrderSign_pair_ordered h1, crAnTimeOrderList_pair_ordered h1,
|
||||
crAnTimeOrderSign_pair_ordered h2, crAnTimeOrderList_pair_ordered h2]
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## Interaction with maxTimeField
|
||||
|
||||
-/
|
||||
|
||||
/-- In the state algebra time, ordering obeys `T(φ₀φ₁…φₙ) = s * φᵢ * T(φ₀φ₁…φᵢ₋₁φᵢ₊₁…φₙ)`
|
||||
where `φᵢ` is the state
|
||||
which has maximum time and `s` is the exchange sign of `φᵢ` and `φ₀φ₁…φᵢ₋₁`. -/
|
||||
lemma timeOrderF_eq_maxTimeField_mul (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
𝓣ᶠ(ofStateList (φ :: φs)) =
|
||||
𝓢(𝓕 |>ₛ maxTimeField φ φs, 𝓕 |>ₛ (φ :: φs).take (maxTimeFieldPos φ φs)) •
|
||||
ofState (maxTimeField φ φs) * 𝓣ᶠ(ofStateList (eraseMaxTimeField φ φs)) := by
|
||||
rw [timeOrderF_ofStateList, timeOrderList_eq_maxTimeField_timeOrderList]
|
||||
rw [ofStateList_cons, timeOrderF_ofStateList]
|
||||
simp only [instCommGroup.eq_1, Algebra.mul_smul_comm, Algebra.smul_mul_assoc, smul_smul]
|
||||
congr
|
||||
rw [timerOrderSign_of_eraseMaxTimeField, mul_assoc]
|
||||
simp
|
||||
|
||||
/-- In the state algebra time, ordering obeys `T(φ₀φ₁…φₙ) = s * φᵢ * T(φ₀φ₁…φᵢ₋₁φᵢ₊₁…φₙ)`
|
||||
where `φᵢ` is the state
|
||||
which has maximum time and `s` is the exchange sign of `φᵢ` and `φ₀φ₁…φᵢ₋₁`.
|
||||
Here `s` is written using finite sets. -/
|
||||
lemma timeOrderF_eq_maxTimeField_mul_finset (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
𝓣ᶠ(ofStateList (φ :: φs)) = 𝓢(𝓕 |>ₛ maxTimeField φ φs, 𝓕 |>ₛ ⟨(eraseMaxTimeField φ φs).get,
|
||||
(Finset.filter (fun x =>
|
||||
(maxTimeFieldPosFin φ φs).succAbove x < maxTimeFieldPosFin φ φs) Finset.univ)⟩) •
|
||||
ofState (maxTimeField φ φs) * 𝓣ᶠ(ofStateList (eraseMaxTimeField φ φs)) := by
|
||||
rw [timeOrderF_eq_maxTimeField_mul]
|
||||
congr 3
|
||||
apply FieldStatistic.ofList_perm
|
||||
nth_rewrite 1 [← List.finRange_map_get (φ :: φs)]
|
||||
simp only [List.length_cons, eraseMaxTimeField, insertionSortDropMinPos]
|
||||
rw [eraseIdx_get, ← List.map_take, ← List.map_map]
|
||||
refine List.Perm.map (φ :: φs).get ?_
|
||||
apply (List.perm_ext_iff_of_nodup _ _).mpr
|
||||
· intro i
|
||||
simp only [List.length_cons, maxTimeFieldPos, mem_take_finrange, Fin.val_fin_lt, List.mem_map,
|
||||
Finset.mem_sort, Finset.mem_filter, Finset.mem_univ, true_and, Function.comp_apply]
|
||||
refine Iff.intro (fun hi => ?_) (fun h => ?_)
|
||||
· have h2 := (maxTimeFieldPosFin φ φs).2
|
||||
simp only [eraseMaxTimeField, insertionSortDropMinPos, List.length_cons, Nat.succ_eq_add_one,
|
||||
maxTimeFieldPosFin, insertionSortMinPosFin] at h2
|
||||
use ⟨i, by omega⟩
|
||||
apply And.intro
|
||||
· simp only [Fin.succAbove, List.length_cons, Fin.castSucc_mk, maxTimeFieldPosFin,
|
||||
insertionSortMinPosFin, Nat.succ_eq_add_one, Fin.mk_lt_mk, Fin.val_fin_lt, Fin.succ_mk]
|
||||
rw [Fin.lt_def]
|
||||
split
|
||||
· simp only [Fin.val_fin_lt]
|
||||
omega
|
||||
· omega
|
||||
· simp only [Fin.succAbove, List.length_cons, Fin.castSucc_mk, Fin.succ_mk, Fin.ext_iff,
|
||||
Fin.coe_cast]
|
||||
split
|
||||
· simp
|
||||
· simp_all [Fin.lt_def]
|
||||
· obtain ⟨j, h1, h2⟩ := h
|
||||
subst h2
|
||||
simp only [Fin.lt_def, Fin.coe_cast]
|
||||
exact h1
|
||||
· exact List.Sublist.nodup (List.take_sublist _ _) <|
|
||||
List.nodup_finRange (φs.length + 1)
|
||||
· refine List.Nodup.map ?_ ?_
|
||||
· refine Function.Injective.comp ?hf.hg Fin.succAbove_right_injective
|
||||
exact Fin.cast_injective (eraseIdx_length (φ :: φs) (insertionSortMinPos timeOrderRel φ φs))
|
||||
· exact Finset.sort_nodup (fun x1 x2 => x1 ≤ x2)
|
||||
(Finset.filter (fun x => (maxTimeFieldPosFin φ φs).succAbove x < maxTimeFieldPosFin φ φs)
|
||||
Finset.univ)
|
||||
|
||||
end
|
||||
|
||||
end FieldOpFreeAlgebra
|
||||
|
||||
end FieldSpecification
|
Loading…
Add table
Add a link
Reference in a new issue