/- 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.CrAnAlgebra.SuperCommute import HepLean.PerturbationTheory.Koszul.KoszulSign /-! # Normal Ordering in the CrAnAlgebra 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 `CrAnAlgebra`. We derive properties of this normal ordering. -/ namespace FieldSpecification variable {๐“• : FieldSpecification} open FieldStatistic namespace CrAnAlgebra 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 normalOrder : CrAnAlgebra ๐“• โ†’โ‚—[โ„‚] CrAnAlgebra ๐“• := Basis.constr ofCrAnListBasis โ„‚ fun ฯ†s => normalOrderSign ฯ†s โ€ข ofCrAnList (normalOrderList ฯ†s) lemma normalOrder_ofCrAnList (ฯ†s : List ๐“•.CrAnStates) : normalOrder (ofCrAnList ฯ†s) = normalOrderSign ฯ†s โ€ข ofCrAnList (normalOrderList ฯ†s) := by rw [โ† ofListBasis_eq_ofList] simp only [normalOrder, Basis.constr_basis] lemma ofCrAnList_eq_normalOrder (ฯ†s : List ๐“•.CrAnStates) : ofCrAnList (normalOrderList ฯ†s) = normalOrderSign ฯ†s โ€ข normalOrder (ofCrAnList ฯ†s) := by rw [normalOrder_ofCrAnList, normalOrderList] rw [smul_smul] simp only [normalOrderSign] rw [Wick.koszulSign_mul_self] simp lemma normalOrder_one : normalOrder (๐“• := ๐“•) 1 = 1 := by rw [โ† ofCrAnList_nil, normalOrder_ofCrAnList] simp /-! ## Normal ordering with a creation operator on the left or annihilation on the right -/ lemma normalOrder_ofCrAnList_cons_create (ฯ† : ๐“•.CrAnStates) (hฯ† : ๐“• |>แถœ ฯ† = CreateAnnihilate.create) (ฯ†s : List ๐“•.CrAnStates) : normalOrder (ofCrAnList (ฯ† :: ฯ†s)) = ofCrAnState ฯ† * normalOrder (ofCrAnList ฯ†s) := by rw [normalOrder_ofCrAnList] rw [normalOrderSign_cons_create ฯ† hฯ†, normalOrderList_cons_create ฯ† hฯ† ฯ†s] rw [ofCrAnList_cons, normalOrder_ofCrAnList, mul_smul_comm] lemma normalOrder_create_mul (ฯ† : ๐“•.CrAnStates) (hฯ† : ๐“• |>แถœ ฯ† = CreateAnnihilate.create) (a : CrAnAlgebra ๐“•) : normalOrder (ofCrAnState ฯ† * a) = ofCrAnState ฯ† * normalOrder a := by change (normalOrder โˆ˜โ‚— mulLinearMap (ofCrAnState ฯ†)) a = (mulLinearMap (ofCrAnState ฯ†) โˆ˜โ‚— normalOrder) a refine LinearMap.congr_fun ?h a apply ofCrAnListBasis.ext intro l simp only [mulLinearMap, LinearMap.coe_mk, AddHom.coe_mk, ofListBasis_eq_ofList, LinearMap.coe_comp, Function.comp_apply] rw [โ† ofCrAnList_cons] rw [normalOrder_ofCrAnList_cons_create ฯ† hฯ†] lemma normalOrder_ofCrAnList_append_annihilate (ฯ† : ๐“•.CrAnStates) (hฯ† : ๐“• |>แถœ ฯ† = CreateAnnihilate.annihilate) (ฯ†s : List ๐“•.CrAnStates) : normalOrder (ofCrAnList (ฯ†s ++ [ฯ†])) = normalOrder (ofCrAnList ฯ†s) * ofCrAnState ฯ† := by rw [normalOrder_ofCrAnList] rw [normalOrderSign_append_annihlate ฯ† hฯ† ฯ†s, normalOrderList_append_annihilate ฯ† hฯ† ฯ†s] rw [ofCrAnList_append, ofCrAnList_singleton, normalOrder_ofCrAnList, smul_mul_assoc] lemma normalOrder_mul_annihilate (ฯ† : ๐“•.CrAnStates) (hฯ† : ๐“• |>แถœ ฯ† = CreateAnnihilate.annihilate) (a : CrAnAlgebra ๐“•) : normalOrder (a * ofCrAnState ฯ†) = normalOrder a * ofCrAnState ฯ† := by change (normalOrder โˆ˜โ‚— mulLinearMap.flip (ofCrAnState ฯ†)) a = (mulLinearMap.flip (ofCrAnState ฯ†) โˆ˜โ‚— normalOrder) a refine LinearMap.congr_fun ?h a apply ofCrAnListBasis.ext intro l 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] rw [normalOrder_ofCrAnList_append_annihilate ฯ† hฯ†] lemma normalOrder_crPart_mul (ฯ† : ๐“•.States) (a : CrAnAlgebra ๐“•) : normalOrder (crPart (StateAlgebra.ofState ฯ†) * a) = crPart (StateAlgebra.ofState ฯ†) * normalOrder a := by match ฯ† with | .negAsymp ฯ† => dsimp only [crPart, StateAlgebra.ofState] simp only [FreeAlgebra.lift_ฮน_apply] exact normalOrder_create_mul โŸจStates.negAsymp ฯ†, ()โŸฉ rfl a | .position ฯ† => dsimp only [crPart, StateAlgebra.ofState] simp only [FreeAlgebra.lift_ฮน_apply] refine normalOrder_create_mul _ ?_ _ simp [crAnStatesToCreateAnnihilate] | .posAsymp ฯ† => simp lemma normalOrder_mul_anPart (ฯ† : ๐“•.States) (a : CrAnAlgebra ๐“•) : normalOrder (a * anPart (StateAlgebra.ofState ฯ†)) = normalOrder a * anPart (StateAlgebra.ofState ฯ†) := by match ฯ† with | .negAsymp ฯ† => simp | .position ฯ† => dsimp only [anPart, StateAlgebra.ofState] simp only [FreeAlgebra.lift_ฮน_apply] refine normalOrder_mul_annihilate _ ?_ _ simp [crAnStatesToCreateAnnihilate] | .posAsymp ฯ† => dsimp only [anPart, StateAlgebra.ofState] simp only [FreeAlgebra.lift_ฮน_apply] refine normalOrder_mul_annihilate _ ?_ _ simp [crAnStatesToCreateAnnihilate] /-! ## Normal ordering for an adjacent creation and annihliation state The main result of this section is `normalOrder_superCommute_annihilate_create`. -/ lemma normalOrder_swap_create_annihlate_ofCrAnList_ofCrAnList (ฯ†c ฯ†a : ๐“•.CrAnStates) (hฯ†c : ๐“• |>แถœ ฯ†c = CreateAnnihilate.create) (hฯ†a : ๐“• |>แถœ ฯ†a = CreateAnnihilate.annihilate) (ฯ†s ฯ†s' : List ๐“•.CrAnStates) : normalOrder (ofCrAnList ฯ†s' * ofCrAnState ฯ†c * ofCrAnState ฯ†a * ofCrAnList ฯ†s) = ๐“ข(๐“• |>โ‚› ฯ†c, ๐“• |>โ‚› ฯ†a) โ€ข normalOrder (ofCrAnList ฯ†s' * ofCrAnState ฯ†a * ofCrAnState ฯ†c * ofCrAnList ฯ†s) := by rw [mul_assoc, mul_assoc, โ† ofCrAnList_cons, โ† ofCrAnList_cons, โ† ofCrAnList_append] rw [normalOrder_ofCrAnList, normalOrderSign_swap_create_annihlate ฯ†c ฯ†a hฯ†c hฯ†a] rw [normalOrderList_swap_create_annihlate ฯ†c ฯ†a hฯ†c hฯ†a] rw [โ† smul_smul, โ† normalOrder_ofCrAnList] congr rw [ofCrAnList_append, ofCrAnList_cons, ofCrAnList_cons] noncomm_ring lemma normalOrder_swap_create_annihlate_ofCrAnList (ฯ†c ฯ†a : ๐“•.CrAnStates) (hฯ†c : ๐“• |>แถœ ฯ†c = CreateAnnihilate.create) (hฯ†a : ๐“• |>แถœ ฯ†a = CreateAnnihilate.annihilate) (ฯ†s : List ๐“•.CrAnStates) (a : ๐“•.CrAnAlgebra) : normalOrder (ofCrAnList ฯ†s * ofCrAnState ฯ†c * ofCrAnState ฯ†a * a) = ๐“ข(๐“• |>โ‚› ฯ†c, ๐“• |>โ‚› ฯ†a) โ€ข normalOrder (ofCrAnList ฯ†s * ofCrAnState ฯ†a * ofCrAnState ฯ†c * a) := by change (normalOrder โˆ˜โ‚— mulLinearMap (ofCrAnList ฯ†s * ofCrAnState ฯ†c * ofCrAnState ฯ†a)) a = (smulLinearMap _ โˆ˜โ‚— normalOrder โˆ˜โ‚— mulLinearMap (ofCrAnList ฯ†s * ofCrAnState ฯ†a * ofCrAnState ฯ†c)) a refine LinearMap.congr_fun ?h a apply ofCrAnListBasis.ext intro l simp only [mulLinearMap, LinearMap.coe_mk, AddHom.coe_mk, ofListBasis_eq_ofList, LinearMap.coe_comp, Function.comp_apply, instCommGroup.eq_1] rw [normalOrder_swap_create_annihlate_ofCrAnList_ofCrAnList ฯ†c ฯ†a hฯ†c hฯ†a] rfl lemma normalOrder_swap_create_annihlate (ฯ†c ฯ†a : ๐“•.CrAnStates) (hฯ†c : ๐“• |>แถœ ฯ†c = CreateAnnihilate.create) (hฯ†a : ๐“• |>แถœ ฯ†a = CreateAnnihilate.annihilate) (a b : ๐“•.CrAnAlgebra) : normalOrder (a * ofCrAnState ฯ†c * ofCrAnState ฯ†a * b) = ๐“ข(๐“• |>โ‚› ฯ†c, ๐“• |>โ‚› ฯ†a) โ€ข normalOrder (a * ofCrAnState ฯ†a * ofCrAnState ฯ†c * b) := by rw [mul_assoc, mul_assoc, mul_assoc, mul_assoc] change (normalOrder โˆ˜โ‚— mulLinearMap.flip (ofCrAnState ฯ†c * (ofCrAnState ฯ†a * b))) a = (smulLinearMap (๐“ข(๐“• |>โ‚› ฯ†c, ๐“• |>โ‚› ฯ†a)) โˆ˜โ‚— normalOrder โˆ˜โ‚— mulLinearMap.flip (ofCrAnState ฯ†a * (ofCrAnState ฯ†c * b))) a apply LinearMap.congr_fun apply ofCrAnListBasis.ext intro 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] repeat rw [โ† mul_assoc] rw [normalOrder_swap_create_annihlate_ofCrAnList ฯ†c ฯ†a hฯ†c hฯ†a] rfl lemma normalOrder_superCommute_create_annihilate (ฯ†c ฯ†a : ๐“•.CrAnStates) (hฯ†c : ๐“• |>แถœ ฯ†c = CreateAnnihilate.create) (hฯ†a : ๐“• |>แถœ ฯ†a = CreateAnnihilate.annihilate) (a b : ๐“•.CrAnAlgebra) : normalOrder (a * superCommute (ofCrAnState ฯ†c) (ofCrAnState ฯ†a) * b) = 0 := by rw [superCommute_ofCrAnState_ofCrAnState] simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [mul_sub, sub_mul, map_sub, โ† smul_mul_assoc] rw [โ† mul_assoc, โ† mul_assoc] rw [normalOrder_swap_create_annihlate ฯ†c ฯ†a hฯ†c hฯ†a] simp only [FieldStatistic.instCommGroup.eq_1, Algebra.mul_smul_comm, Algebra.smul_mul_assoc, map_smul, sub_self] lemma normalOrder_superCommute_annihilate_create (ฯ†c ฯ†a : ๐“•.CrAnStates) (hฯ†c : ๐“• |>แถœ ฯ†c = CreateAnnihilate.create) (hฯ†a : ๐“• |>แถœ ฯ†a = CreateAnnihilate.annihilate) (a b : ๐“•.CrAnAlgebra) : normalOrder (a * superCommute (ofCrAnState ฯ†a) (ofCrAnState ฯ†c) * b) = 0 := by rw [superCommute_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] apply Or.inr exact normalOrder_superCommute_create_annihilate ฯ†c ฯ†a hฯ†c hฯ†a a b lemma normalOrder_swap_crPart_anPart (ฯ† ฯ†' : ๐“•.States) (a b : CrAnAlgebra ๐“•) : normalOrder (a * (crPart (StateAlgebra.ofState ฯ†)) * (anPart (StateAlgebra.ofState ฯ†')) * b) = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข normalOrder (a * (anPart (StateAlgebra.ofState ฯ†')) * (crPart (StateAlgebra.ofState ฯ†)) * b) := by match ฯ†, ฯ†' with | _, .negAsymp ฯ†' => simp | .posAsymp ฯ†, _ => simp | .position ฯ†, .position ฯ†' => simp only [crPart_position, anPart_position, instCommGroup.eq_1] rw [normalOrder_swap_create_annihlate] simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod] rfl rfl | .negAsymp ฯ†, .posAsymp ฯ†' => simp only [crPart_negAsymp, anPart_posAsymp, instCommGroup.eq_1] rw [normalOrder_swap_create_annihlate] simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod] rfl rfl | .negAsymp ฯ†, .position ฯ†' => simp only [crPart_negAsymp, anPart_position, instCommGroup.eq_1] rw [normalOrder_swap_create_annihlate] simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod] rfl rfl | .position ฯ†, .posAsymp ฯ†' => simp only [crPart_position, anPart_posAsymp, instCommGroup.eq_1] rw [normalOrder_swap_create_annihlate] simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod] rfl rfl /-! ## Normal ordering for an anPart and crPart Using the results from above. -/ lemma normalOrder_swap_anPart_crPart (ฯ† ฯ†' : ๐“•.States) (a b : CrAnAlgebra ๐“•) : normalOrder (a * (anPart (StateAlgebra.ofState ฯ†)) * (crPart (StateAlgebra.ofState ฯ†')) * b) = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข normalOrder (a * (crPart (StateAlgebra.ofState ฯ†')) * (anPart (StateAlgebra.ofState ฯ†)) * b) := by rw [normalOrder_swap_crPart_anPart] rw [smul_smul, FieldStatistic.exchangeSign_symm, FieldStatistic.exchangeSign_mul_self] simp lemma normalOrder_superCommute_crPart_anPart (ฯ† ฯ†' : ๐“•.States) (a b : CrAnAlgebra ๐“•) : normalOrder (a * superCommute (crPart (StateAlgebra.ofState ฯ†)) (anPart (StateAlgebra.ofState ฯ†')) * b) = 0 := by match ฯ†, ฯ†' with | _, .negAsymp ฯ†' => simp | .posAsymp ฯ†', _ => simp | .position ฯ†, .position ฯ†' => simp only [crPart_position, anPart_position] refine normalOrder_superCommute_create_annihilate _ _ (by rfl) (by rfl) _ _ | .negAsymp ฯ†, .posAsymp ฯ†' => simp only [crPart_negAsymp, anPart_posAsymp] refine normalOrder_superCommute_create_annihilate _ _ (by rfl) (by rfl) _ _ | .negAsymp ฯ†, .position ฯ†' => simp only [crPart_negAsymp, anPart_position] refine normalOrder_superCommute_create_annihilate _ _ (by rfl) (by rfl) _ _ | .position ฯ†, .posAsymp ฯ†' => simp only [crPart_position, anPart_posAsymp] refine normalOrder_superCommute_create_annihilate _ _ (by rfl) (by rfl) _ _ lemma normalOrder_superCommute_anPart_crPart (ฯ† ฯ†' : ๐“•.States) (a b : CrAnAlgebra ๐“•) : normalOrder (a * superCommute (anPart (StateAlgebra.ofState ฯ†)) (crPart (StateAlgebra.ofState ฯ†')) * b) = 0 := by match ฯ†, ฯ†' with | .negAsymp ฯ†', _ => simp | _, .posAsymp ฯ†' => simp | .position ฯ†, .position ฯ†' => simp only [anPart_position, crPart_position] refine normalOrder_superCommute_annihilate_create _ _ (by rfl) (by rfl) _ _ | .posAsymp ฯ†', .negAsymp ฯ† => simp only [anPart_posAsymp, crPart_negAsymp] refine normalOrder_superCommute_annihilate_create _ _ (by rfl) (by rfl) _ _ | .position ฯ†', .negAsymp ฯ† => simp only [anPart_position, crPart_negAsymp] refine normalOrder_superCommute_annihilate_create _ _ (by rfl) (by rfl) _ _ | .posAsymp ฯ†, .position ฯ†' => simp only [anPart_posAsymp, crPart_position] refine normalOrder_superCommute_annihilate_create _ _ (by rfl) (by rfl) _ _ /-! ## The normal ordering of a product of two states -/ @[simp] lemma normalOrder_crPart_mul_crPart (ฯ† ฯ†' : ๐“•.States) : normalOrder (crPart (StateAlgebra.ofState ฯ†) * crPart (StateAlgebra.ofState ฯ†')) = crPart (StateAlgebra.ofState ฯ†) * crPart (StateAlgebra.ofState ฯ†') := by rw [normalOrder_crPart_mul] conv_lhs => rw [โ† mul_one (crPart (StateAlgebra.ofState ฯ†'))] rw [normalOrder_crPart_mul, normalOrder_one] simp @[simp] lemma normalOrder_anPart_mul_anPart (ฯ† ฯ†' : ๐“•.States) : normalOrder (anPart (StateAlgebra.ofState ฯ†) * anPart (StateAlgebra.ofState ฯ†')) = anPart (StateAlgebra.ofState ฯ†) * anPart (StateAlgebra.ofState ฯ†') := by rw [normalOrder_mul_anPart] conv_lhs => rw [โ† one_mul (anPart (StateAlgebra.ofState ฯ†))] rw [normalOrder_mul_anPart, normalOrder_one] simp @[simp] lemma normalOrder_crPart_mul_anPart (ฯ† ฯ†' : ๐“•.States) : normalOrder (crPart (StateAlgebra.ofState ฯ†) * anPart (StateAlgebra.ofState ฯ†')) = crPart (StateAlgebra.ofState ฯ†) * anPart (StateAlgebra.ofState ฯ†') := by rw [normalOrder_crPart_mul] conv_lhs => rw [โ† one_mul (anPart (StateAlgebra.ofState ฯ†'))] rw [normalOrder_mul_anPart, normalOrder_one] simp @[simp] lemma normalOrder_anPart_mul_crPart (ฯ† ฯ†' : ๐“•.States) : normalOrder (anPart (StateAlgebra.ofState ฯ†) * crPart (StateAlgebra.ofState ฯ†')) = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข (crPart (StateAlgebra.ofState ฯ†') * anPart (StateAlgebra.ofState ฯ†)) := by conv_lhs => rw [โ† one_mul (anPart (StateAlgebra.ofState ฯ†) * crPart (StateAlgebra.ofState ฯ†'))] conv_lhs => rw [โ† mul_one (1 * (anPart (StateAlgebra.ofState ฯ†) * crPart (StateAlgebra.ofState ฯ†')))] rw [โ† mul_assoc, normalOrder_swap_anPart_crPart] simp lemma normalOrder_ofState_mul_ofState (ฯ† ฯ†' : ๐“•.States) : normalOrder (ofState ฯ† * ofState ฯ†') = crPart (StateAlgebra.ofState ฯ†) * crPart (StateAlgebra.ofState ฯ†') + ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข (crPart (StateAlgebra.ofState ฯ†') * anPart (StateAlgebra.ofState ฯ†)) + crPart (StateAlgebra.ofState ฯ†) * anPart (StateAlgebra.ofState ฯ†') + anPart (StateAlgebra.ofState ฯ†) * anPart (StateAlgebra.ofState ฯ†') := by rw [ofState_eq_crPart_add_anPart, ofState_eq_crPart_add_anPart] rw [mul_add, add_mul, add_mul] simp only [map_add, normalOrder_crPart_mul_crPart, normalOrder_anPart_mul_crPart, instCommGroup.eq_1, normalOrder_crPart_mul_anPart, normalOrder_anPart_mul_anPart] abel /-! ## Normal order with super commutors -/ /-! TODO: Split the following two lemmas up into smaller parts. -/ lemma normalOrder_superCommute_ofCrAnList_create_create_ofCrAnList (ฯ†c ฯ†c' : ๐“•.CrAnStates) (hฯ†c : ๐“• |>แถœ ฯ†c = CreateAnnihilate.create) (hฯ†c' : ๐“• |>แถœ ฯ†c' = CreateAnnihilate.create) (ฯ†s ฯ†s' : List ๐“•.CrAnStates) : (normalOrder (ofCrAnList ฯ†s * superCommute (ofCrAnState ฯ†c) (ofCrAnState ฯ†c') * ofCrAnList ฯ†s')) = normalOrderSign (ฯ†s ++ ฯ†c' :: ฯ†c :: ฯ†s') โ€ข (ofCrAnList (createFilter ฯ†s) * superCommute (ofCrAnState ฯ†c) (ofCrAnState ฯ†c') * ofCrAnList (createFilter ฯ†s') * ofCrAnList (annihilateFilter (ฯ†s ++ ฯ†s'))) := by rw [superCommute_ofCrAnState_ofCrAnState] rw [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 [normalOrder_ofCrAnList] rw [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] 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 [normalOrder_ofCrAnList] rw [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] congr rw [ofCrAnList_append, ofCrAnList_singleton, ofCrAnList_singleton] rw [ofCrAnList_append, ofCrAnList_singleton, ofCrAnList_singleton, smul_mul_assoc] lemma normalOrder_superCommute_ofCrAnList_annihilate_annihilate_ofCrAnList (ฯ†a ฯ†a' : ๐“•.CrAnStates) (hฯ†a : ๐“• |>แถœ ฯ†a = CreateAnnihilate.annihilate) (hฯ†a' : ๐“• |>แถœ ฯ†a' = CreateAnnihilate.annihilate) (ฯ†s ฯ†s' : List ๐“•.CrAnStates) : (normalOrder (ofCrAnList ฯ†s * superCommute (ofCrAnState ฯ†a) (ofCrAnState ฯ†a') * ofCrAnList ฯ†s')) = normalOrderSign (ฯ†s ++ ฯ†a' :: ฯ†a :: ฯ†s') โ€ข (ofCrAnList (createFilter (ฯ†s ++ ฯ†s')) * ofCrAnList (annihilateFilter ฯ†s) * superCommute (ofCrAnState ฯ†a) (ofCrAnState ฯ†a') * ofCrAnList (annihilateFilter ฯ†s')) := by rw [superCommute_ofCrAnState_ofCrAnState] rw [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 [normalOrder_ofCrAnList] rw [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 [normalOrder_ofCrAnList] rw [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 congr 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_superCommute_normalOrder_ofCrAnList (ฯ†s ฯ†s' : List ๐“•.CrAnStates) : โŸจofCrAnList ฯ†s, normalOrder (ofCrAnList ฯ†s')โŸฉโ‚›ca = ofCrAnList ฯ†s * normalOrder (ofCrAnList ฯ†s') - ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†s') โ€ข normalOrder (ofCrAnList ฯ†s') * ofCrAnList ฯ†s := by simp [normalOrder_ofCrAnList, map_smul, superCommute_ofCrAnList_ofCrAnList, ofCrAnList_append, smul_sub, smul_smul, mul_comm] lemma ofCrAnList_superCommute_normalOrder_ofStateList (ฯ†s : List ๐“•.CrAnStates) (ฯ†s' : List ๐“•.States) : โŸจofCrAnList ฯ†s, normalOrder (ofStateList ฯ†s')โŸฉโ‚›ca = ofCrAnList ฯ†s * normalOrder (ofStateList ฯ†s') - ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†s') โ€ข normalOrder (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_superCommute_normalOrder_ofCrAnList, CrAnSection.statistics_eq_state_statistics] /-! ## Multiplications with normal order written in terms of super commute. -/ lemma ofCrAnList_mul_normalOrder_ofStateList_eq_superCommute (ฯ†s : List ๐“•.CrAnStates) (ฯ†s' : List ๐“•.States) : ofCrAnList ฯ†s * normalOrder (ofStateList ฯ†s') = ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†s') โ€ข normalOrder (ofStateList ฯ†s') * ofCrAnList ฯ†s + โŸจofCrAnList ฯ†s, normalOrder (ofStateList ฯ†s')โŸฉโ‚›ca := by rw [ofCrAnList_superCommute_normalOrder_ofStateList] simp lemma ofCrAnState_mul_normalOrder_ofStateList_eq_superCommute (ฯ† : ๐“•.CrAnStates) (ฯ†s' : List ๐“•.States) : ofCrAnState ฯ† * normalOrder (ofStateList ฯ†s') = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†s') โ€ข normalOrder (ofStateList ฯ†s') * ofCrAnState ฯ† + โŸจofCrAnState ฯ†, normalOrder (ofStateList ฯ†s')โŸฉโ‚›ca := by rw [โ† ofCrAnList_singleton, ofCrAnList_mul_normalOrder_ofStateList_eq_superCommute] simp [ofCrAnList_singleton] lemma anPart_mul_normalOrder_ofStateList_eq_superCommute (ฯ† : ๐“•.States) (ฯ†s' : List ๐“•.States) : anPart (StateAlgebra.ofState ฯ†) * normalOrder (ofStateList ฯ†s') = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†s') โ€ข normalOrder (ofStateList ฯ†s' * anPart (StateAlgebra.ofState ฯ†)) + โŸจanPart (StateAlgebra.ofState ฯ†), normalOrder (ofStateList ฯ†s')โŸฉโ‚›ca := by rw [normalOrder_mul_anPart] match ฯ† with | .negAsymp ฯ† => simp | .position ฯ† => simp only [anPart_position, instCommGroup.eq_1] rw [ofCrAnState_mul_normalOrder_ofStateList_eq_superCommute] simp [crAnStatistics] | .posAsymp ฯ† => simp only [anPart_posAsymp, instCommGroup.eq_1] rw [ofCrAnState_mul_normalOrder_ofStateList_eq_superCommute] simp [crAnStatistics] end end CrAnAlgebra end FieldSpecification