refactor: Remove ProtoOperatorAlgebra
This commit is contained in:
parent
c18b4850e5
commit
fc20099282
15 changed files with 615 additions and 792 deletions
|
@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.PerturbationTheory.WickContraction.Sign
|
||||
import HepLean.PerturbationTheory.Algebras.ProtoOperatorAlgebra.TimeContraction
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.TimeContraction
|
||||
/-!
|
||||
|
||||
# Time contractions
|
||||
|
@ -17,16 +17,16 @@ variable {𝓕 : FieldSpecification}
|
|||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
|
||||
open FieldOpAlgebra
|
||||
/-- Given a Wick contraction `φsΛ` associated with a list `φs`, the
|
||||
product of all time-contractions of pairs of contracted elements in `φs`,
|
||||
as a member of the center of `𝓞.A`. -/
|
||||
noncomputable def timeContract (𝓞 : 𝓕.ProtoOperatorAlgebra) {φs : List 𝓕.States}
|
||||
noncomputable def timeContract {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length) :
|
||||
Subalgebra.center ℂ 𝓞.A :=
|
||||
∏ (a : φsΛ.1), ⟨𝓞.timeContract
|
||||
Subalgebra.center ℂ 𝓕.FieldOpAlgebra :=
|
||||
∏ (a : φsΛ.1), ⟨FieldOpAlgebra.timeContract
|
||||
(φs.get (φsΛ.fstFieldOfContract a)) (φs.get (φsΛ.sndFieldOfContract a)),
|
||||
𝓞.timeContract_mem_center _ _⟩
|
||||
timeContract_mem_center _ _⟩
|
||||
|
||||
/-- For `φsΛ` a Wick contraction for `φs`, the time contraction `(φsΛ ↩Λ φ i none).timeContract 𝓞`
|
||||
is equal to `φsΛ.timeContract 𝓞`.
|
||||
|
@ -34,10 +34,10 @@ noncomputable def timeContract (𝓞 : 𝓕.ProtoOperatorAlgebra) {φs : List
|
|||
This result follows from the fact that `timeContract` only depends on contracted pairs,
|
||||
and `(φsΛ ↩Λ φ i none)` has the 'same' contracted pairs as `φsΛ`. -/
|
||||
@[simp]
|
||||
lemma timeContract_insertAndContract_none (𝓞 : 𝓕.ProtoOperatorAlgebra)
|
||||
lemma timeContract_insertAndContract_none
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) :
|
||||
(φsΛ ↩Λ φ i none).timeContract 𝓞 = φsΛ.timeContract 𝓞 := by
|
||||
(φsΛ ↩Λ φ i none).timeContract = φsΛ.timeContract := by
|
||||
rw [timeContract, insertAndContract_none_prod_contractions]
|
||||
congr
|
||||
ext a
|
||||
|
@ -51,13 +51,13 @@ lemma timeContract_insertAndContract_none (𝓞 : 𝓕.ProtoOperatorAlgebra)
|
|||
This follows from the fact that `(φsΛ ↩Λ φ i (some j))` has one more contracted pair than `φsΛ`,
|
||||
corresponding to `φ` contracted with `φⱼ`. The order depends on whether we insert `φ` before
|
||||
or after `φⱼ`. -/
|
||||
lemma timeConract_insertAndContract_some (𝓞 : 𝓕.ProtoOperatorAlgebra)
|
||||
lemma timeConract_insertAndContract_some
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
(φsΛ ↩Λ φ i (some j)).timeContract 𝓞 =
|
||||
(φsΛ ↩Λ φ i (some j)).timeContract =
|
||||
(if i < i.succAbove j then
|
||||
⟨𝓞.timeContract φ φs[j.1], 𝓞.timeContract_mem_center _ _⟩
|
||||
else ⟨𝓞.timeContract φs[j.1] φ, 𝓞.timeContract_mem_center _ _⟩) * φsΛ.timeContract 𝓞 := by
|
||||
⟨FieldOpAlgebra.timeContract φ φs[j.1], timeContract_mem_center _ _⟩
|
||||
else ⟨FieldOpAlgebra.timeContract φs[j.1] φ, timeContract_mem_center _ _⟩) * φsΛ.timeContract := by
|
||||
rw [timeContract, insertAndContract_some_prod_contractions]
|
||||
congr 1
|
||||
· simp only [Nat.succ_eq_add_one, insertAndContract_fstFieldOfContract_some_incl, finCongr_apply,
|
||||
|
@ -72,27 +72,25 @@ lemma timeConract_insertAndContract_some (𝓞 : 𝓕.ProtoOperatorAlgebra)
|
|||
open FieldStatistic
|
||||
|
||||
lemma timeConract_insertAndContract_some_eq_mul_contractStateAtIndex_lt
|
||||
(𝓞 : 𝓕.ProtoOperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
||||
(ht : 𝓕.timeOrderRel φ φs[k.1]) (hik : i < i.succAbove k) :
|
||||
(φsΛ ↩Λ φ i (some k)).timeContract 𝓞 =
|
||||
(φsΛ ↩Λ φ i (some k)).timeContract =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (φsΛ.uncontracted.filter (fun x => x < k))⟩)
|
||||
• (𝓞.contractStateAtIndex φ [φsΛ]ᵘᶜ ((uncontractedStatesEquiv φs φsΛ) (some k)) *
|
||||
φsΛ.timeContract 𝓞) := by
|
||||
• (contractStateAtIndex φ [φsΛ]ᵘᶜ ((uncontractedStatesEquiv φs φsΛ) (some k)) *
|
||||
φsΛ.timeContract ) := by
|
||||
rw [timeConract_insertAndContract_some]
|
||||
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
|
||||
ProtoOperatorAlgebra.contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
|
||||
contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
|
||||
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
|
||||
List.getElem_map, uncontractedList_getElem_uncontractedIndexEquiv_symm, List.get_eq_getElem,
|
||||
Algebra.smul_mul_assoc, uncontractedListGet]
|
||||
· simp only [hik, ↓reduceIte, MulMemClass.coe_mul]
|
||||
rw [𝓞.timeContract_of_timeOrderRel]
|
||||
trans (1 : ℂ) • (𝓞.crAnF ((CrAnAlgebra.superCommuteF
|
||||
(CrAnAlgebra.anPartF φ)) (CrAnAlgebra.ofState φs[k.1])) *
|
||||
↑(timeContract 𝓞 φsΛ))
|
||||
rw [timeContract_of_timeOrderRel]
|
||||
trans (1 : ℂ) • ((superCommute (anPart φ)) (ofFieldOp φs[k.1]) * ↑φsΛ.timeContract )
|
||||
· simp
|
||||
simp only [smul_smul]
|
||||
congr
|
||||
congr 1
|
||||
have h1 : ofList 𝓕.statesStatistic (List.take (↑(φsΛ.uncontractedIndexEquiv.symm k))
|
||||
(List.map φs.get φsΛ.uncontractedList))
|
||||
= (𝓕 |>ₛ ⟨φs.get, (Finset.filter (fun x => x < k) φsΛ.uncontracted)⟩) := by
|
||||
|
@ -107,21 +105,21 @@ lemma timeConract_insertAndContract_some_eq_mul_contractStateAtIndex_lt
|
|||
· exact ht
|
||||
|
||||
lemma timeConract_insertAndContract_some_eq_mul_contractStateAtIndex_not_lt
|
||||
(𝓞 : 𝓕.ProtoOperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
||||
(ht : ¬ 𝓕.timeOrderRel φs[k.1] φ) (hik : ¬ i < i.succAbove k) :
|
||||
(φsΛ ↩Λ φ i (some k)).timeContract 𝓞 =
|
||||
(φsΛ ↩Λ φ i (some k)).timeContract =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (φsΛ.uncontracted.filter (fun x => x ≤ k))⟩)
|
||||
• (𝓞.contractStateAtIndex φ [φsΛ]ᵘᶜ
|
||||
((uncontractedStatesEquiv φs φsΛ) (some k)) * φsΛ.timeContract 𝓞) := by
|
||||
• (contractStateAtIndex φ [φsΛ]ᵘᶜ
|
||||
((uncontractedStatesEquiv φs φsΛ) (some k)) * φsΛ.timeContract) := by
|
||||
rw [timeConract_insertAndContract_some]
|
||||
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
|
||||
ProtoOperatorAlgebra.contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
|
||||
contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
|
||||
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
|
||||
List.getElem_map, uncontractedList_getElem_uncontractedIndexEquiv_symm, List.get_eq_getElem,
|
||||
Algebra.smul_mul_assoc, uncontractedListGet]
|
||||
simp only [hik, ↓reduceIte, MulMemClass.coe_mul]
|
||||
rw [𝓞.timeContract_of_not_timeOrderRel, 𝓞.timeContract_of_timeOrderRel]
|
||||
rw [timeContract_of_not_timeOrderRel, timeContract_of_timeOrderRel]
|
||||
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc, smul_smul]
|
||||
congr
|
||||
have h1 : ofList 𝓕.statesStatistic (List.take (↑(φsΛ.uncontractedIndexEquiv.symm k))
|
||||
|
@ -158,9 +156,9 @@ lemma timeConract_insertAndContract_some_eq_mul_contractStateAtIndex_not_lt
|
|||
simp_all only [Fin.getElem_fin, Nat.succ_eq_add_one, not_lt, false_or]
|
||||
exact ht
|
||||
|
||||
lemma timeContract_of_not_gradingCompliant (𝓞 : 𝓕.ProtoOperatorAlgebra) (φs : List 𝓕.States)
|
||||
lemma timeContract_of_not_gradingCompliant (φs : List 𝓕.States)
|
||||
(φsΛ : WickContraction φs.length) (h : ¬ GradingCompliant φs φsΛ) :
|
||||
φsΛ.timeContract 𝓞 = 0 := by
|
||||
φsΛ.timeContract = 0 := by
|
||||
rw [timeContract]
|
||||
simp only [GradingCompliant, Fin.getElem_fin, Subtype.forall, not_forall] at h
|
||||
obtain ⟨a, ha⟩ := h
|
||||
|
@ -169,7 +167,7 @@ lemma timeContract_of_not_gradingCompliant (𝓞 : 𝓕.ProtoOperatorAlgebra) (
|
|||
simp only [Finset.univ_eq_attach, Finset.mem_attach]
|
||||
apply Subtype.eq
|
||||
simp only [List.get_eq_getElem, ZeroMemClass.coe_zero]
|
||||
rw [ProtoOperatorAlgebra.timeContract_zero_of_diff_grade]
|
||||
rw [timeContract_zero_of_diff_grade]
|
||||
simp [ha2]
|
||||
|
||||
end WickContraction
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue