PhysLean/HepLean/PerturbationTheory/Wick/Contraction.lean

183 lines
7 KiB
Text
Raw Normal View History

2024-12-15 12:42:50 +00:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
2024-12-19 14:25:09 +00:00
import HepLean.PerturbationTheory.Wick.OperatorMap
2024-12-15 12:42:50 +00:00
/-!
# Koszul signs and ordering for lists and algebras
-/
namespace Wick
noncomputable section
2024-12-17 07:15:47 +00:00
open HepLean.List
2024-12-15 12:42:50 +00:00
2024-12-19 15:40:04 +00:00
/-- Given a list of fields `l`, the type of pairwise-contractions associated with `l`
which have the list `aux` uncontracted. -/
2024-12-15 12:42:50 +00:00
inductive ContractionsAux {I : Type} : (l : List I) → (aux : List I) → Type
| nil : ContractionsAux [] []
2024-12-17 07:15:47 +00:00
| cons {l : List I} {aux : List I} {a : I} (i : Option (Fin aux.length)) :
ContractionsAux l aux → ContractionsAux (a :: l) (optionEraseZ aux a i)
2024-12-15 12:42:50 +00:00
2024-12-19 15:40:04 +00:00
/-- Given a list of fields `l`, the type of pairwise-contractions associated with `l`. -/
2024-12-15 12:42:50 +00:00
def Contractions {I : Type} (l : List I) : Type := Σ aux, ContractionsAux l aux
namespace Contractions
variable {I : Type} {l : List I} (c : Contractions l)
2024-12-19 15:40:04 +00:00
/-- The list of uncontracted fields. -/
2024-12-15 12:42:50 +00:00
def normalize : List I := c.1
lemma contractions_nil (a : Contractions ([] : List I)) : a = ⟨[], ContractionsAux.nil⟩ := by
cases a
rename_i aux c
cases c
rfl
2024-12-17 07:15:47 +00:00
lemma contractions_single {i : I} (a : Contractions [i]) : a =
2024-12-19 12:59:14 +00:00
⟨[i], ContractionsAux.cons none ContractionsAux.nil⟩ := by
2024-12-15 12:42:50 +00:00
cases a
rename_i aux c
cases c
2024-12-17 07:15:47 +00:00
rename_i aux' c'
cases c'
cases aux'
2024-12-19 12:59:14 +00:00
simp only [List.length_nil, optionEraseZ]
2024-12-17 07:15:47 +00:00
rename_i x
exact Fin.elim0 x
2024-12-19 15:40:04 +00:00
/-- For the nil list of fields there is only one contraction. -/
2024-12-17 07:15:47 +00:00
def nilEquiv : Contractions ([] : List I) ≃ Unit where
toFun _ := ()
invFun _ := ⟨[], ContractionsAux.nil⟩
2024-12-19 12:59:14 +00:00
left_inv a := Eq.symm (contractions_nil a)
right_inv _ := rfl
2024-12-17 07:15:47 +00:00
2024-12-19 15:40:04 +00:00
/-- The equivalence between contractions of `a :: l` and contractions of
`Contractions l` paired with an optional element of `Fin (c.normalize).length` specifying
what `a` contracts with if any. -/
2024-12-17 07:15:47 +00:00
def consEquiv {a : I} {l : List I} :
Contractions (a :: l) ≃ (c : Contractions l) × Option (Fin (c.normalize).length) where
2024-12-15 12:42:50 +00:00
toFun c :=
match c with
| ⟨aux, c⟩ =>
match c with
| ContractionsAux.cons (aux := aux') i c => ⟨⟨aux', c⟩, i⟩
invFun ci :=
2024-12-17 07:15:47 +00:00
⟨(optionEraseZ (ci.fst.normalize) a ci.2), ContractionsAux.cons (a := a) ci.2 ci.1.2⟩
2024-12-15 12:42:50 +00:00
left_inv c := by
match c with
| ⟨aux, c⟩ =>
match c with
| ContractionsAux.cons (aux := aux') i c => rfl
right_inv ci := by rfl
2024-12-19 15:40:04 +00:00
/-- The type of contractions is decidable. -/
2024-12-15 12:42:50 +00:00
instance decidable : (l : List I) → DecidableEq (Contractions l)
| [] => fun a b =>
match a, b with
| ⟨_, a⟩, ⟨_, b⟩ =>
match a, b with
2024-12-19 12:59:14 +00:00
| ContractionsAux.nil, ContractionsAux.nil => isTrue rfl
| _ :: l =>
2024-12-17 07:15:47 +00:00
haveI : DecidableEq (Contractions l) := decidable l
haveI : DecidableEq ((c : Contractions l) × Option (Fin (c.normalize).length)) :=
2024-12-15 12:42:50 +00:00
Sigma.instDecidableEqSigma
2024-12-17 07:15:47 +00:00
Equiv.decidableEq consEquiv
2024-12-15 12:42:50 +00:00
2024-12-19 15:40:04 +00:00
/-- The type of contractions is finite. -/
2024-12-19 12:59:14 +00:00
instance fintype : (l : List I) → Fintype (Contractions l)
2024-12-15 12:42:50 +00:00
| [] => {
elems := {⟨[], ContractionsAux.nil⟩}
complete := by
intro a
rw [Finset.mem_singleton]
exact contractions_nil a}
2024-12-19 12:59:14 +00:00
| a :: l =>
2024-12-17 07:15:47 +00:00
haveI : Fintype (Contractions l) := fintype l
haveI : Fintype ((c : Contractions l) × Option (Fin (c.normalize).length)) :=
2024-12-15 12:42:50 +00:00
Sigma.instFintype
2024-12-17 07:15:47 +00:00
Fintype.ofEquiv _ consEquiv.symm
2024-12-15 12:42:50 +00:00
2024-12-19 15:40:04 +00:00
/-- A structure specifying when a type `I` and a map `f :I → Type` corresponds to
the splitting of a fields `φ` into a creation `n` and annihlation part `p`. -/
2024-12-17 07:15:47 +00:00
structure Splitting {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1] where
2024-12-19 15:40:04 +00:00
/-- The creation part of the fields. The label `n` corresponds to the fact that
in normal ordering these feilds get pushed to the negative (left) direction. -/
2024-12-19 12:59:14 +00:00
𝓑n : I → (Σ i, f i)
2024-12-19 15:40:04 +00:00
/-- The annhilation part of the fields. The label `p` corresponds to the fact that
in normal ordering these feilds get pushed to the positive (right) direction. -/
2024-12-19 12:59:14 +00:00
𝓑p : I → (Σ i, f i)
2024-12-19 15:40:04 +00:00
/-- The complex coefficent of creation part of a field `i`. This is usually `0` or `1`. -/
2024-12-19 12:59:14 +00:00
𝓧n : I →
2024-12-19 15:40:04 +00:00
/-- The complex coefficent of annhilation part of a field `i`. This is usually `0` or `1`. -/
2024-12-19 12:59:14 +00:00
𝓧p : I →
2024-12-19 14:25:09 +00:00
h𝓑 : ∀ i, ofListLift f [i] 1 = ofList [𝓑n i] (𝓧n i) + ofList [𝓑p i] (𝓧p i)
2024-12-17 07:15:47 +00:00
h𝓑n : ∀ i j, le1 (𝓑n i) j
h𝓑p : ∀ i j, le1 j (𝓑p i)
2024-12-15 12:42:50 +00:00
2024-12-19 15:40:04 +00:00
/-- In the static wick's theorem, the term associated with a contraction `c` containing
the contractions. -/
2024-12-17 07:15:47 +00:00
def toCenterTerm {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
2024-12-19 11:23:49 +00:00
(q : I → Fin 2)
2024-12-17 07:15:47 +00:00
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
2024-12-15 12:42:50 +00:00
{A : Type} [Semiring A] [Algebra A]
2024-12-19 15:40:04 +00:00
(F : FreeAlgebra (Σ i, f i) →ₐ[] A) :
2024-12-19 12:59:14 +00:00
{r : List I} → (c : Contractions r) → (S : Splitting f le1) → A
2024-12-19 11:23:49 +00:00
| [], ⟨[], .nil⟩, _ => 1
| _ :: _, ⟨_, .cons (aux := aux') none c⟩, S => toCenterTerm f q le1 F ⟨aux', c⟩ S
2024-12-19 12:59:14 +00:00
| a :: _, ⟨_, .cons (aux := aux') (some n) c⟩, S => toCenterTerm f q le1 F ⟨aux', c⟩ S *
2024-12-17 07:15:47 +00:00
superCommuteCoef q [aux'.get n] (List.take (↑n) aux') •
2024-12-19 14:25:09 +00:00
F (((superCommute fun i => q i.fst) (ofList [S.𝓑p a] (S.𝓧p a))) (ofListLift f [aux'.get n] 1))
2024-12-17 07:15:47 +00:00
lemma toCenterTerm_none {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
(q : I → Fin 2) {r : List I}
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
{A : Type} [Semiring A] [Algebra A]
2024-12-19 15:40:04 +00:00
(F : FreeAlgebra (Σ i, f i) →ₐ A)
2024-12-19 12:59:14 +00:00
(S : Splitting f le1) (a : I) (c : Contractions r) :
toCenterTerm (r := a :: r) f q le1 F (Contractions.consEquiv.symm ⟨c, none⟩) S =
toCenterTerm f q le1 F c S := by
2024-12-17 07:15:47 +00:00
rw [consEquiv]
2024-12-19 12:59:14 +00:00
simp only [Equiv.coe_fn_symm_mk]
2024-12-17 07:15:47 +00:00
dsimp [toCenterTerm]
rfl
2024-12-15 12:42:50 +00:00
2024-12-17 07:15:47 +00:00
lemma toCenterTerm_center {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
2024-12-19 11:23:49 +00:00
(q : I → Fin 2)
2024-12-15 12:42:50 +00:00
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
2024-12-17 07:15:47 +00:00
{A : Type} [Semiring A] [Algebra A]
2024-12-19 12:59:14 +00:00
(F : FreeAlgebra (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le1 F] :
{r : List I} → (c : Contractions r) → (S : Splitting f le1) →
2024-12-19 11:23:49 +00:00
(c.toCenterTerm f q le1 F S) ∈ Subalgebra.center A
| [], ⟨[], .nil⟩, _ => by
dsimp [toCenterTerm]
exact Subalgebra.one_mem (Subalgebra.center A)
| _ :: _, ⟨_, .cons (aux := aux') none c⟩, S => by
dsimp [toCenterTerm]
exact toCenterTerm_center f q le1 F ⟨aux', c⟩ S
2024-12-19 12:59:14 +00:00
| a :: _, ⟨_, .cons (aux := aux') (some n) c⟩, S => by
2024-12-19 11:23:49 +00:00
dsimp [toCenterTerm]
refine Subalgebra.mul_mem (Subalgebra.center A) ?hx ?hy
exact toCenterTerm_center f q le1 F ⟨aux', c⟩ S
apply Subalgebra.smul_mem
2024-12-19 14:25:09 +00:00
rw [ofListLift_expand]
2024-12-19 11:23:49 +00:00
rw [map_sum, map_sum]
refine Subalgebra.sum_mem (Subalgebra.center A) ?hy.hx.h
intro x _
2024-12-19 14:25:09 +00:00
simp only [CreateAnnilateSect.toList]
2024-12-19 11:23:49 +00:00
rw [ofList_singleton]
2024-12-19 12:59:14 +00:00
exact OperatorMap.superCommute_ofList_singleton_ι_center (q := fun i => q i.1)
(le1 := le1) F (S.𝓑p a) ⟨aux'[↑n], x.head⟩
2024-12-15 12:42:50 +00:00
end Contractions
end
end Wick