/- 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 -/ import HepLean.PerturbationTheory.Wick.SuperCommute /-! # Operator map -/ namespace Wick noncomputable section open FieldStatistic variable {๐“• : Type} /-- A map from the free algebra of fields `FreeAlgebra โ„‚ ๐“•` to an algebra `A`, to be thought of as the operator algebra is said to be an operator map if all super commutors of fields land in the center of `A`, if two fields are of a different grade then their super commutor lands on zero, and the `koszulOrder` (normal order) of any term with a super commutor of two fields present is zero. This can be thought as as a condition on the operator algebra `A` as much as it can on `F`. -/ class OperatorMap {A : Type} [Semiring A] [Algebra โ„‚ A] (q : ๐“• โ†’ FieldStatistic) (le : ๐“• โ†’ ๐“• โ†’ Prop) [DecidableRel le] (F : FreeAlgebra โ„‚ ๐“• โ†’โ‚[โ„‚] A) : Prop where superCommute_mem_center : โˆ€ i j, F (superCommute q (FreeAlgebra.ฮน โ„‚ i) (FreeAlgebra.ฮน โ„‚ j)) โˆˆ Subalgebra.center โ„‚ A superCommute_diff_grade_zero : โˆ€ i j, q i โ‰  q j โ†’ F (superCommute q (FreeAlgebra.ฮน โ„‚ i) (FreeAlgebra.ฮน โ„‚ j)) = 0 superCommute_ordered_zero : โˆ€ i j, โˆ€ a b, F (koszulOrder q le (a * superCommute q (FreeAlgebra.ฮน โ„‚ i) (FreeAlgebra.ฮน โ„‚ j) * b)) = 0 namespace OperatorMap variable {A : Type} [Semiring A] [Algebra โ„‚ A] {q : ๐“• โ†’ FieldStatistic} {le : ๐“• โ†’ ๐“• โ†’ Prop} [DecidableRel le] (F : FreeAlgebra โ„‚ ๐“• โ†’โ‚[โ„‚] A) lemma superCommute_ofList_singleton_ฮน_center [OperatorMap q le F] (i j : ๐“•) : F (superCommute q (ofList [i] xa) (FreeAlgebra.ฮน โ„‚ j)) โˆˆ Subalgebra.center โ„‚ A := by have h1 : F (superCommute q (ofList [i] xa) (FreeAlgebra.ฮน โ„‚ j)) = xa โ€ข F (superCommute q (FreeAlgebra.ฮน โ„‚ i) (FreeAlgebra.ฮน โ„‚ j)) := by rw [โ† map_smul] congr rw [ofList_eq_smul_one, ofList_singleton] rw [map_smul] rfl rw [h1] refine Subalgebra.smul_mem (Subalgebra.center โ„‚ A) ?_ xa exact superCommute_mem_center (le := le) i j end OperatorMap variable {๐“• : Type} (q : ๐“• โ†’ FieldStatistic) (le : ๐“• โ†’ ๐“• โ†’ Prop) [DecidableRel le] lemma superCommuteSplit_operatorMap (lb : List ๐“•) (xa xb : โ„‚) (n : โ„•) (hn : n < lb.length) {A : Type} [Semiring A] [Algebra โ„‚ A] (f : FreeAlgebra โ„‚ ๐“• โ†’โ‚[โ„‚] A) [OperatorMap q le f] (i : ๐“•) : f (superCommuteSplit q [i] lb xa xb n hn) = f (superCommute q (ofList [i] xa) (FreeAlgebra.ฮน โ„‚ (lb.get โŸจn, hnโŸฉ))) * (superCommuteCoef q [i] (List.take n lb) โ€ข f (ofList (List.eraseIdx lb n) xb)) := by have hn : f ((superCommute q) (ofList [i] xa) (FreeAlgebra.ฮน โ„‚ (lb.get โŸจn, hnโŸฉ))) โˆˆ Subalgebra.center โ„‚ A := OperatorMap.superCommute_ofList_singleton_ฮน_center (le := le) f i (lb.get โŸจn, hnโŸฉ) rw [Subalgebra.mem_center_iff] at hn rw [superCommuteSplit, map_mul, map_mul, map_smul, hn, mul_assoc, smul_mul_assoc, โ† map_mul, โ† ofList_pair] congr ยท exact Eq.symm (List.eraseIdx_eq_take_drop_succ lb n) ยท exact one_mul xb lemma superCommuteLiftSplit_operatorMap {f : ๐“• โ†’ Type} [โˆ€ i, Fintype (f i)] (c : (ฮฃ i, f i)) (r : List ๐“•) (x y : โ„‚) (n : โ„•) (hn : n < r.length) (le : (ฮฃ i, f i) โ†’ (ฮฃ i, f i) โ†’ Prop) [DecidableRel le] {A : Type} [Semiring A] [Algebra โ„‚ A] (F : FreeAlgebra โ„‚ (ฮฃ i, f i) โ†’โ‚[โ„‚] A) [OperatorMap (fun i => q i.1) le F] : F (superCommuteLiftSplit q [c] r x y n hn) = superCommuteLiftCoef q [c] (List.take n r) โ€ข (F (superCommute (fun i => q i.1) (ofList [c] x) (sumFiber f (FreeAlgebra.ฮน โ„‚ (r.get โŸจn, hnโŸฉ)))) * F (ofListLift f (List.eraseIdx r n) y)) := by rw [superCommuteLiftSplit] rw [map_smul] congr rw [map_mul, map_mul] have h1 : F ((superCommute fun i => q i.fst) (ofList [c] x) ((sumFiber f) (FreeAlgebra.ฮน โ„‚ (r.get โŸจn, hnโŸฉ)))) โˆˆ Subalgebra.center โ„‚ A := by rw [sumFiber_ฮน] rw [map_sum, map_sum] refine Subalgebra.sum_mem _ ?_ intro n exact fun a => OperatorMap.superCommute_ofList_singleton_ฮน_center (le := le) F c _ rw [Subalgebra.mem_center_iff] at h1 rw [h1, mul_assoc, โ† map_mul] congr rw [ofListLift, ofListLift, ofListLift, โ† map_mul] congr rw [โ† ofList_pair, one_mul] congr exact Eq.symm (List.eraseIdx_eq_take_drop_succ r n) lemma superCommute_koszulOrder_le_ofList [IsTotal ๐“• le] [IsTrans ๐“• le] (r : List ๐“•) (x : โ„‚) (i : ๐“•) {A : Type} [Semiring A] [Algebra โ„‚ A] (F : FreeAlgebra โ„‚ ๐“• โ†’โ‚ A) [OperatorMap q le F] : F ((superCommute q (FreeAlgebra.ฮน โ„‚ i) (koszulOrder q le (ofList r x)))) = โˆ‘ n : Fin r.length, (superCommuteCoef q [r.get n] (r.take n)) โ€ข (F (((superCommute q) (ofList [i] 1)) (FreeAlgebra.ฮน โ„‚ (r.get n))) * F ((koszulOrder q le) (ofList (r.eraseIdx โ†‘n) x))) := by rw [koszulOrder_ofList, map_smul, map_smul, โ† ofList_singleton, superCommute_ofList_sum] rw [map_sum, โ† (HepLean.List.insertionSortEquiv le r).sum_comp] conv_lhs => enter [2, 2] intro n rw [superCommuteSplit_operatorMap (le := le)] enter [1, 2, 2, 2] change ((List.insertionSort le r).get โˆ˜ (HepLean.List.insertionSortEquiv le r)) n rw [HepLean.List.insertionSort_get_comp_insertionSortEquiv] conv_lhs => enter [2, 2] intro n rw [HepLean.List.eraseIdx_insertionSort_fin le r n] rw [ofList_insertionSort_eq_koszulOrder q le] rw [Finset.smul_sum] conv_lhs => rhs intro n rw [map_smul, smul_smul, Algebra.mul_smul_comm, smul_smul] congr funext n by_cases hq : q i โ‰  q (r.get n) ยท have hn := OperatorMap.superCommute_diff_grade_zero (q := q) (F := F) le i (r.get n) hq conv_lhs => enter [2, 1] rw [ofList_singleton, hn] conv_rhs => enter [2, 1] rw [ofList_singleton, hn] simp ยท congr 1 trans staticWickCoef q le r i n ยท rw [staticWickCoef, mul_assoc] refine staticWickCoef_eq_get q le r i n ?_ simpa using hq lemma koszulOrder_of_le_all_ofList (r : List ๐“•) (x : โ„‚) (i : ๐“•) {A : Type} [Semiring A] [Algebra โ„‚ A] (F : FreeAlgebra โ„‚ ๐“• โ†’โ‚ A) [OperatorMap q le F] : F (koszulOrder q le (ofList r x * FreeAlgebra.ฮน โ„‚ i)) = superCommuteCoef q [i] r โ€ข F (koszulOrder q le (FreeAlgebra.ฮน โ„‚ i * ofList r x)) := by conv_lhs => enter [2, 2] rw [โ† ofList_singleton] rw [ofListLift_ofList_superCommute' q] rw [map_sub] rw [sub_eq_add_neg] rw [map_add] conv_lhs => enter [2, 2] rw [map_smul] rw [โ† neg_smul] rw [map_smul, map_smul, map_smul] conv_lhs => rhs rhs rw [superCommute_ofList_sum] rw [map_sum, map_sum] dsimp [superCommuteSplit] rw [ofList_singleton] rhs intro n rw [Algebra.smul_mul_assoc, Algebra.smul_mul_assoc] rw [map_smul, map_smul] rw [OperatorMap.superCommute_ordered_zero] simp only [smul_zero, Finset.sum_const_zero, add_zero] rw [ofList_singleton] lemma le_all_mul_koszulOrder_ofList (r : List ๐“•) (x : โ„‚) (i : ๐“•) (hi : โˆ€ (j : ๐“•), le j i) {A : Type} [Semiring A] [Algebra โ„‚ A] (F : FreeAlgebra โ„‚ ๐“• โ†’โ‚ A) [OperatorMap q le F] : F (FreeAlgebra.ฮน โ„‚ i * koszulOrder q le (ofList r x)) = F ((koszulOrder q le) (FreeAlgebra.ฮน โ„‚ i * ofList r x)) + F (((superCommute q) (ofList [i] 1)) ((koszulOrder q le) (ofList r x))) := by rw [koszulOrder_ofList, Algebra.mul_smul_comm, map_smul, โ† ofList_singleton, ofList_ofList_superCommute q, map_add, smul_add, โ† map_smul] conv_lhs => enter [1, 2] rw [โ† Algebra.smul_mul_assoc, smul_smul, mul_comm, โ† smul_smul, โ† koszulOrder_ofList, Algebra.smul_mul_assoc, ofList_singleton] rw [koszulOrder_mul_ge, map_smul] congr ยท rw [koszulOrder_of_le_all_ofList] rw [superCommuteCoef_perm_snd q [i] (List.insertionSort le r) r (List.perm_insertionSort le r)] rw [smul_smul] rw [superCommuteCoef_mul_self] simp [ofList_singleton] ยท rw [map_smul, map_smul] ยท exact fun j => hi j /-- In the expansions of `F (FreeAlgebra.ฮน โ„‚ i * koszulOrder q le (ofList r x))` the ter multiplying `F ((koszulOrder q le) (ofList (optionEraseZ r i n) x))`. -/ def superCommuteCenterOrder (r : List ๐“•) (i : ๐“•) {A : Type} [Semiring A] [Algebra โ„‚ A] (F : FreeAlgebra โ„‚ ๐“• โ†’โ‚[โ„‚] A) (n : Option (Fin r.length)) : A := match n with | none => 1 | some n => superCommuteCoef q [r.get n] (r.take n) โ€ข F (((superCommute q) (ofList [i] 1)) (FreeAlgebra.ฮน โ„‚ (r.get n))) @[simp] lemma superCommuteCenterOrder_none (r : List ๐“•) (i : ๐“•) {A : Type} [Semiring A] [Algebra โ„‚ A] (F : FreeAlgebra โ„‚ ๐“• โ†’โ‚[โ„‚] A) : superCommuteCenterOrder q r i F none = 1 := by simp [superCommuteCenterOrder] open HepLean.List lemma le_all_mul_koszulOrder_ofList_expand [IsTotal ๐“• le] [IsTrans ๐“• le] (r : List ๐“•) (x : โ„‚) (i : ๐“•) (hi : โˆ€ (j : ๐“•), le j i) {A : Type} [Semiring A] [Algebra โ„‚ A] (F : FreeAlgebra โ„‚ ๐“• โ†’โ‚[โ„‚] A) [OperatorMap q le F] : F (FreeAlgebra.ฮน โ„‚ i * koszulOrder q le (ofList r x)) = โˆ‘ n, superCommuteCenterOrder q r i F n * F ((koszulOrder q le) (ofList (optionEraseZ r i n) x)) := by rw [le_all_mul_koszulOrder_ofList] conv_lhs => rhs rw [ofList_singleton] rw [superCommute_koszulOrder_le_ofList] simp only [List.get_eq_getElem, Fintype.sum_option, superCommuteCenterOrder_none, one_mul] congr ยท rw [โ† ofList_singleton, โ† ofList_pair] simp only [List.singleton_append, one_mul] rfl ยท funext n simp only [superCommuteCenterOrder, List.get_eq_getElem, Algebra.smul_mul_assoc] rfl exact fun j => hi j lemma le_all_mul_koszulOrder_ofListLift_expand {f : ๐“• โ†’ Type} [โˆ€ i, Fintype (f i)] (r : List ๐“•) (x : โ„‚) (le : (ฮฃ i, f i) โ†’ (ฮฃ i, f i) โ†’ Prop) [DecidableRel le] [IsTotal (ฮฃ i, f i) le] [IsTrans (ฮฃ i, f i) le] (i : (ฮฃ i, f i)) (hi : โˆ€ (j : (ฮฃ i, f i)), le j i) {A : Type} [Semiring A] [Algebra โ„‚ A] (F : FreeAlgebra โ„‚ (ฮฃ i, f i) โ†’โ‚ A) [OperatorMap (fun i => q i.1) le F] : F (ofList [i] 1 * koszulOrder (fun i => q i.1) le (ofListLift f r x)) = F ((koszulOrder (fun i => q i.fst) le) (ofList [i] 1 * ofListLift f r x)) + โˆ‘ n : (Fin r.length), superCommuteCoef q [r.get n] (List.take (โ†‘n) r) โ€ข F (((superCommute fun i => q i.fst) (ofList [i] 1)) (ofListLift f [r.get n] 1)) * F ((koszulOrder (fun i => q i.fst) le) (ofListLift f (r.eraseIdx โ†‘n) x)) := by match r with | [] => simp only [map_mul, List.length_nil, Finset.univ_eq_empty, List.get_eq_getElem, List.take_nil, List.eraseIdx_nil, Algebra.smul_mul_assoc, Finset.sum_empty, add_zero] rw [ofListLift_empty_smul] simp only [map_smul, koszulOrder_one, map_one, Algebra.mul_smul_comm, mul_one] rw [ofList_singleton, koszulOrder_ฮน] | r0 :: r => rw [ofListLift_expand, map_sum, Finset.mul_sum, map_sum] let e1 (a : CreateAnnihilateSect f (r0 :: r)) : Option (Fin a.toList.length) โ‰ƒ Option (Fin (r0 :: r).length) := Equiv.optionCongr (Fin.castOrderIso (CreateAnnihilateSect.toList_length a)).toEquiv conv_lhs => rhs intro a rw [ofList_singleton, le_all_mul_koszulOrder_ofList_expand _ _ _ _ _ hi] rw [โ† (e1 a).symm.sum_comp] rhs intro n rw [Finset.sum_comm] simp only [Fintype.sum_option] congr 1 ยท simp only [List.length_cons, List.get_eq_getElem, superCommuteCenterOrder, Equiv.optionCongr_symm, OrderIso.toEquiv_symm, Fin.symm_castOrderIso, Equiv.optionCongr_apply, RelIso.coe_fn_toEquiv, Option.map_none', optionEraseZ, one_mul, e1] rw [โ† map_sum, Finset.mul_sum, โ† map_sum] apply congrArg apply congrArg congr funext x rw [ofList_cons_eq_ofList] ยท congr funext n rw [โ† (CreateAnnihilateSect.extractEquiv n).symm.sum_comp] simp only [List.get_eq_getElem, List.length_cons, Equiv.optionCongr_symm, OrderIso.toEquiv_symm, Fin.symm_castOrderIso, Equiv.optionCongr_apply, RelIso.coe_fn_toEquiv, Option.map_some', Fin.castOrderIso_apply, Algebra.smul_mul_assoc, e1] erw [Finset.sum_product] have h1 (a0 : f (r0 :: r)[โ†‘n]) (a : CreateAnnihilateSect f ((r0 :: r).eraseIdx โ†‘n)) : superCommuteCenterOrder (fun i => q i.fst) ((CreateAnnihilateSect.extractEquiv n).symm (a0, a)).toList i F (some (Fin.cast (by simp) n)) = superCommuteCoef q [(r0 :: r).get n] (List.take (โ†‘n) (r0 :: r)) โ€ข F (((superCommute fun i => q i.fst) (ofList [i] 1)) (FreeAlgebra.ฮน โ„‚ โŸจ(r0 :: r).get n, a0โŸฉ)) := by simp only [superCommuteCenterOrder, List.get_eq_getElem, List.length_cons, Fin.coe_cast] erw [CreateAnnihilateSect.extractEquiv_symm_toList_get_same] have hsc : superCommuteCoef (fun i => q i.fst) [โŸจ(r0 :: r).get n, a0โŸฉ] (List.take (โ†‘n) ((CreateAnnihilateSect.extractEquiv n).symm (a0, a)).toList) = superCommuteCoef q [(r0 :: r).get n] (List.take (โ†‘n) ((r0 :: r))) := by simp only [superCommuteCoef, List.get_eq_getElem, List.length_cons, Fin.isValue, CreateAnnihilateSect.toList_grade_take] rfl erw [hsc] rfl conv_lhs => rhs intro a0 rhs intro a erw [h1] conv_lhs => rhs intro a0 rw [โ† Finset.mul_sum] conv_lhs => rhs intro a0 enter [2, 2] intro a simp [optionEraseZ] enter [2, 2, 1] rw [โ† CreateAnnihilateSect.eraseIdx_toList] erw [CreateAnnihilateSect.extractEquiv_symm_eraseIdx] rw [โ† Finset.sum_mul] conv_lhs => lhs rw [โ† Finset.smul_sum] erw [โ† map_sum, โ† map_sum, โ† ofListLift_singleton_one] conv_lhs => rhs rw [โ† map_sum, โ† map_sum] simp only [List.get_eq_getElem, List.length_cons, Equiv.symm_apply_apply, Algebra.smul_mul_assoc] erw [โ† ofListLift_expand] simp only [List.get_eq_getElem, List.length_cons, Algebra.smul_mul_assoc] end end Wick