feat: Contractions and involutions

This commit is contained in:
jstoobysmith 2025-01-03 15:13:16 +00:00
parent bab9f10763
commit 7d1f15e18a
5 changed files with 662 additions and 99 deletions

View file

@ -33,13 +33,13 @@ lemma static_wick_nil {A : Type} [Semiring A] [Algebra A]
simp [ofListLift_empty]
lemma static_wick_cons [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) × f i) le]
{A : Type} [Semiring A] [Algebra A] (r : List 𝓕) (a : 𝓕)
{A : Type} [Semiring A] [Algebra A] (φs : List 𝓕) (φ : 𝓕)
(F : FreeAlgebra (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le F]
(S : Contractions.Splitting f le)
(ih : F (ofListLift f r 1) =
∑ c : Contractions r, c.toCenterTerm f q le F S * F (koszulOrder (fun i => q i.fst) le
(ih : F (ofListLift f φs 1) =
∑ c : Contractions φs, c.toCenterTerm f q le F S * F (koszulOrder (fun i => q i.fst) le
(ofListLift f c.normalize 1))) :
F (ofListLift f (a :: r) 1) = ∑ c : Contractions (a :: r),
F (ofListLift f (φ :: φs) 1) = ∑ c : Contractions (φ :: φs),
c.toCenterTerm f q le F S *
F (koszulOrder (fun i => q i.fst) le (ofListLift f c.normalize 1)) := by
rw [ofListLift_cons_eq_ofListLift, map_mul, ih, Finset.mul_sum,
@ -47,7 +47,7 @@ lemma static_wick_cons [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) ×
erw [Finset.sum_sigma]
congr
funext c
have hb := S.h𝓑 a
have hb := S.h𝓑 φ
rw [← mul_assoc]
have hi := c.toCenterTerm_center f q le F S
rw [Subalgebra.mem_center_iff] at hi
@ -80,16 +80,16 @@ lemma static_wick_cons [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) ×
funext n
rw [← mul_assoc]
rfl
exact S.h𝓑p a
exact S.h𝓑n a
exact S.h𝓑p φ
exact S.h𝓑n φ
theorem static_wick_theorem [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) × f i) le]
{A : Type} [Semiring A] [Algebra A] (r : List 𝓕)
{A : Type} [Semiring A] [Algebra A] (φs : List 𝓕)
(F : FreeAlgebra (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le F]
(S : Contractions.Splitting f le) :
F (ofListLift f r 1) = ∑ c : Contractions r, c.toCenterTerm f q le F S *
F (ofListLift f φs 1) = ∑ c : Contractions φs, c.toCenterTerm f q le F S *
F (koszulOrder (fun i => q i.fst) le (ofListLift f c.normalize 1)) := by
induction r with
induction φs with
| nil => exact static_wick_nil q le F S
| cons a r ih => exact static_wick_cons q le r a F S ih