/- 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.FieldOpFreeAlgebra.Basic import HepLean.PerturbationTheory.FieldOpFreeAlgebra.Grading /-! # Super Commute -/ namespace FieldSpecification variable {๐“• : FieldSpecification} namespace FieldOpFreeAlgebra /-! ## The super commutator on the FieldOpFreeAlgebra. -/ open FieldStatistic /-- For a field specification `๐“•`, the super commutator `superCommuteF` is defined as the linear map `๐“•.FieldOpFreeAlgebra โ†’โ‚—[โ„‚] ๐“•.FieldOpFreeAlgebra โ†’โ‚—[โ„‚] ๐“•.FieldOpFreeAlgebra` which on the lists `ฯ†s` and `ฯ†s'` of `๐“•.CrAnFieldOp` gives `superCommuteF ฯ†s ฯ†s' = ฯ†s * ฯ†s' - ๐“ข(ฯ†s, ฯ†s') โ€ข ฯ†s' * ฯ†s`. The notation `[a, b]โ‚›ca` can be used for `superCommuteF a b`. -/ noncomputable def superCommuteF : ๐“•.FieldOpFreeAlgebra โ†’โ‚—[โ„‚] ๐“•.FieldOpFreeAlgebra โ†’โ‚—[โ„‚] ๐“•.FieldOpFreeAlgebra := Basis.constr ofCrAnListFBasis โ„‚ fun ฯ†s => Basis.constr ofCrAnListFBasis โ„‚ fun ฯ†s' => ofCrAnListF (ฯ†s ++ ฯ†s') - ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†s') โ€ข ofCrAnListF (ฯ†s' ++ ฯ†s) @[inherit_doc superCommuteF] scoped[FieldSpecification.FieldOpFreeAlgebra] notation "[" ฯ†s "," ฯ†s' "]โ‚›ca" => superCommuteF ฯ†s ฯ†s' /-! ## The super commutator of different types of elements -/ lemma superCommuteF_ofCrAnListF_ofCrAnListF (ฯ†s ฯ†s' : List ๐“•.CrAnFieldOp) : [ofCrAnListF ฯ†s, ofCrAnListF ฯ†s']โ‚›ca = ofCrAnListF (ฯ†s ++ ฯ†s') - ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†s') โ€ข ofCrAnListF (ฯ†s' ++ ฯ†s) := by rw [โ† ofListBasis_eq_ofList, โ† ofListBasis_eq_ofList] simp only [superCommuteF, Basis.constr_basis] lemma superCommuteF_ofCrAnOpF_ofCrAnOpF (ฯ† ฯ†' : ๐“•.CrAnFieldOp) : [ofCrAnOpF ฯ†, ofCrAnOpF ฯ†']โ‚›ca = ofCrAnOpF ฯ† * ofCrAnOpF ฯ†' - ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข ofCrAnOpF ฯ†' * ofCrAnOpF ฯ† := by rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton] rw [superCommuteF_ofCrAnListF_ofCrAnListF, ofCrAnListF_append] congr rw [ofCrAnListF_append] rw [FieldStatistic.ofList_singleton, FieldStatistic.ofList_singleton, smul_mul_assoc] lemma superCommuteF_ofCrAnListF_ofFieldOpFsList (ฯ†cas : List ๐“•.CrAnFieldOp) (ฯ†s : List ๐“•.FieldOp) : [ofCrAnListF ฯ†cas, ofFieldOpListF ฯ†s]โ‚›ca = ofCrAnListF ฯ†cas * ofFieldOpListF ฯ†s - ๐“ข(๐“• |>โ‚› ฯ†cas, ๐“• |>โ‚› ฯ†s) โ€ข ofFieldOpListF ฯ†s * ofCrAnListF ฯ†cas := by conv_lhs => rw [ofFieldOpListF_sum] rw [map_sum] conv_lhs => enter [2, x] rw [superCommuteF_ofCrAnListF_ofCrAnListF, CrAnSection.statistics_eq_state_statistics, ofCrAnListF_append, ofCrAnListF_append] rw [Finset.sum_sub_distrib, โ† Finset.mul_sum, โ† Finset.smul_sum, โ† Finset.sum_mul, โ† ofFieldOpListF_sum] simp lemma superCommuteF_ofFieldOpListF_ofFieldOpFsList (ฯ† : List ๐“•.FieldOp) (ฯ†s : List ๐“•.FieldOp) : [ofFieldOpListF ฯ†, ofFieldOpListF ฯ†s]โ‚›ca = ofFieldOpListF ฯ† * ofFieldOpListF ฯ†s - ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†s) โ€ข ofFieldOpListF ฯ†s * ofFieldOpListF ฯ† := by conv_lhs => rw [ofFieldOpListF_sum] simp only [map_sum, LinearMap.coeFn_sum, Finset.sum_apply, instCommGroup.eq_1, Algebra.smul_mul_assoc] conv_lhs => enter [2, x] rw [superCommuteF_ofCrAnListF_ofFieldOpFsList] simp only [instCommGroup.eq_1, CrAnSection.statistics_eq_state_statistics, Algebra.smul_mul_assoc, Finset.sum_sub_distrib] rw [โ† Finset.sum_mul, โ† Finset.smul_sum, โ† Finset.mul_sum, โ† ofFieldOpListF_sum] lemma superCommuteF_ofFieldOpF_ofFieldOpFsList (ฯ† : ๐“•.FieldOp) (ฯ†s : List ๐“•.FieldOp) : [ofFieldOpF ฯ†, ofFieldOpListF ฯ†s]โ‚›ca = ofFieldOpF ฯ† * ofFieldOpListF ฯ†s - ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†s) โ€ข ofFieldOpListF ฯ†s * ofFieldOpF ฯ† := by rw [โ† ofFieldOpListF_singleton, superCommuteF_ofFieldOpListF_ofFieldOpFsList, ofFieldOpListF_singleton] simp lemma superCommuteF_ofFieldOpListF_ofFieldOpF (ฯ†s : List ๐“•.FieldOp) (ฯ† : ๐“•.FieldOp) : [ofFieldOpListF ฯ†s, ofFieldOpF ฯ†]โ‚›ca = ofFieldOpListF ฯ†s * ofFieldOpF ฯ† - ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†) โ€ข ofFieldOpF ฯ† * ofFieldOpListF ฯ†s := by rw [โ† ofFieldOpListF_singleton, superCommuteF_ofFieldOpListF_ofFieldOpFsList, ofFieldOpListF_singleton] simp lemma superCommuteF_anPartF_crPartF (ฯ† ฯ†' : ๐“•.FieldOp) : [anPartF ฯ†, crPartF ฯ†']โ‚›ca = anPartF ฯ† * crPartF ฯ†' - ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข crPartF ฯ†' * anPartF ฯ† := by match ฯ†, ฯ†' with | FieldOp.inAsymp ฯ†, _ => simp | _, FieldOp.outAsymp ฯ† => simp only [crPartF_posAsymp, map_zero, mul_zero, instCommGroup.eq_1, smul_zero, zero_mul, sub_self] | FieldOp.position ฯ†, FieldOp.position ฯ†' => simp only [anPartF_position, crPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] | FieldOp.outAsymp ฯ†, FieldOp.position ฯ†' => simp only [anPartF_posAsymp, crPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] | FieldOp.position ฯ†, FieldOp.inAsymp ฯ†' => simp only [anPartF_position, crPartF_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp only [List.singleton_append, instCommGroup.eq_1, crAnStatistics, FieldStatistic.ofList_singleton, Function.comp_apply, crAnFieldOpToFieldOp_prod, โ† ofCrAnListF_append] | FieldOp.outAsymp ฯ†, FieldOp.inAsymp ฯ†' => simp only [anPartF_posAsymp, crPartF_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] lemma superCommuteF_crPartF_anPartF (ฯ† ฯ†' : ๐“•.FieldOp) : [crPartF ฯ†, anPartF ฯ†']โ‚›ca = crPartF ฯ† * anPartF ฯ†' - ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข anPartF ฯ†' * crPartF ฯ† := by match ฯ†, ฯ†' with | FieldOp.outAsymp ฯ†, _ => simp only [crPartF_posAsymp, map_zero, LinearMap.zero_apply, zero_mul, instCommGroup.eq_1, mul_zero, sub_self] | _, FieldOp.inAsymp ฯ† => simp only [anPartF_negAsymp, map_zero, mul_zero, instCommGroup.eq_1, smul_zero, zero_mul, sub_self] | FieldOp.position ฯ†, FieldOp.position ฯ†' => simp only [crPartF_position, anPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] | FieldOp.position ฯ†, FieldOp.outAsymp ฯ†' => simp only [crPartF_position, anPartF_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] | FieldOp.inAsymp ฯ†, FieldOp.position ฯ†' => simp only [crPartF_negAsymp, anPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] | FieldOp.inAsymp ฯ†, FieldOp.outAsymp ฯ†' => simp only [crPartF_negAsymp, anPartF_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] lemma superCommuteF_crPartF_crPartF (ฯ† ฯ†' : ๐“•.FieldOp) : [crPartF ฯ†, crPartF ฯ†']โ‚›ca = crPartF ฯ† * crPartF ฯ†' - ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข crPartF ฯ†' * crPartF ฯ† := by match ฯ†, ฯ†' with | FieldOp.outAsymp ฯ†, _ => simp only [crPartF_posAsymp, map_zero, LinearMap.zero_apply, zero_mul, instCommGroup.eq_1, mul_zero, sub_self] | _, FieldOp.outAsymp ฯ† => simp only [crPartF_posAsymp, map_zero, mul_zero, instCommGroup.eq_1, smul_zero, zero_mul, sub_self] | FieldOp.position ฯ†, FieldOp.position ฯ†' => simp only [crPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] | FieldOp.position ฯ†, FieldOp.inAsymp ฯ†' => simp only [crPartF_position, crPartF_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] | FieldOp.inAsymp ฯ†, FieldOp.position ฯ†' => simp only [crPartF_negAsymp, crPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] | FieldOp.inAsymp ฯ†, FieldOp.inAsymp ฯ†' => simp only [crPartF_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] lemma superCommuteF_anPartF_anPartF (ฯ† ฯ†' : ๐“•.FieldOp) : [anPartF ฯ†, anPartF ฯ†']โ‚›ca = anPartF ฯ† * anPartF ฯ†' - ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข anPartF ฯ†' * anPartF ฯ† := by match ฯ†, ฯ†' with | FieldOp.inAsymp ฯ†, _ => simp | _, FieldOp.inAsymp ฯ† => simp | FieldOp.position ฯ†, FieldOp.position ฯ†' => simp only [anPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] | FieldOp.position ฯ†, FieldOp.outAsymp ฯ†' => simp only [anPartF_position, anPartF_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] | FieldOp.outAsymp ฯ†, FieldOp.position ฯ†' => simp only [anPartF_posAsymp, anPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] | FieldOp.outAsymp ฯ†, FieldOp.outAsymp ฯ†' => simp only [anPartF_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF] simp [crAnStatistics, โ† ofCrAnListF_append] lemma superCommuteF_crPartF_ofFieldOpListF (ฯ† : ๐“•.FieldOp) (ฯ†s : List ๐“•.FieldOp) : [crPartF ฯ†, ofFieldOpListF ฯ†s]โ‚›ca = crPartF ฯ† * ofFieldOpListF ฯ†s - ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†s) โ€ข ofFieldOpListF ฯ†s * crPartF ฯ† := by match ฯ† with | FieldOp.inAsymp ฯ† => simp only [crPartF_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofFieldOpFsList] simp [crAnStatistics] | FieldOp.position ฯ† => simp only [crPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofFieldOpFsList] simp [crAnStatistics] | FieldOp.outAsymp ฯ† => simp lemma superCommuteF_anPartF_ofFieldOpListF (ฯ† : ๐“•.FieldOp) (ฯ†s : List ๐“•.FieldOp) : [anPartF ฯ†, ofFieldOpListF ฯ†s]โ‚›ca = anPartF ฯ† * ofFieldOpListF ฯ†s - ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†s) โ€ข ofFieldOpListF ฯ†s * anPartF ฯ† := by match ฯ† with | FieldOp.inAsymp ฯ† => simp | FieldOp.position ฯ† => simp only [anPartF_position, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofFieldOpFsList] simp [crAnStatistics] | FieldOp.outAsymp ฯ† => simp only [anPartF_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc] rw [โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofFieldOpFsList] simp [crAnStatistics] lemma superCommuteF_crPartF_ofFieldOpF (ฯ† ฯ†' : ๐“•.FieldOp) : [crPartF ฯ†, ofFieldOpF ฯ†']โ‚›ca = crPartF ฯ† * ofFieldOpF ฯ†' - ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข ofFieldOpF ฯ†' * crPartF ฯ† := by rw [โ† ofFieldOpListF_singleton, superCommuteF_crPartF_ofFieldOpListF] simp lemma superCommuteF_anPartF_ofFieldOpF (ฯ† ฯ†' : ๐“•.FieldOp) : [anPartF ฯ†, ofFieldOpF ฯ†']โ‚›ca = anPartF ฯ† * ofFieldOpF ฯ†' - ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข ofFieldOpF ฯ†' * anPartF ฯ† := by rw [โ† ofFieldOpListF_singleton, superCommuteF_anPartF_ofFieldOpListF] simp /-! ## Mul equal superCommuteF Lemmas which rewrite a multiplication of two elements of the algebra as their commuted multiplication with a sign plus the super commutator. -/ lemma ofCrAnListF_mul_ofCrAnListF_eq_superCommuteF (ฯ†s ฯ†s' : List ๐“•.CrAnFieldOp) : ofCrAnListF ฯ†s * ofCrAnListF ฯ†s' = ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†s') โ€ข ofCrAnListF ฯ†s' * ofCrAnListF ฯ†s + [ofCrAnListF ฯ†s, ofCrAnListF ฯ†s']โ‚›ca := by rw [superCommuteF_ofCrAnListF_ofCrAnListF] simp [ofCrAnListF_append] lemma ofCrAnOpF_mul_ofCrAnListF_eq_superCommuteF (ฯ† : ๐“•.CrAnFieldOp) (ฯ†s' : List ๐“•.CrAnFieldOp) : ofCrAnOpF ฯ† * ofCrAnListF ฯ†s' = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†s') โ€ข ofCrAnListF ฯ†s' * ofCrAnOpF ฯ† + [ofCrAnOpF ฯ†, ofCrAnListF ฯ†s']โ‚›ca := by rw [โ† ofCrAnListF_singleton, ofCrAnListF_mul_ofCrAnListF_eq_superCommuteF] simp lemma ofFieldOpListF_mul_ofFieldOpListF_eq_superCommuteF (ฯ†s ฯ†s' : List ๐“•.FieldOp) : ofFieldOpListF ฯ†s * ofFieldOpListF ฯ†s' = ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†s') โ€ข ofFieldOpListF ฯ†s' * ofFieldOpListF ฯ†s + [ofFieldOpListF ฯ†s, ofFieldOpListF ฯ†s']โ‚›ca := by rw [superCommuteF_ofFieldOpListF_ofFieldOpFsList] simp lemma ofFieldOpF_mul_ofFieldOpListF_eq_superCommuteF (ฯ† : ๐“•.FieldOp) (ฯ†s' : List ๐“•.FieldOp) : ofFieldOpF ฯ† * ofFieldOpListF ฯ†s' = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†s') โ€ข ofFieldOpListF ฯ†s' * ofFieldOpF ฯ† + [ofFieldOpF ฯ†, ofFieldOpListF ฯ†s']โ‚›ca := by rw [superCommuteF_ofFieldOpF_ofFieldOpFsList] simp lemma ofFieldOpListF_mul_ofFieldOpF_eq_superCommuteF (ฯ†s : List ๐“•.FieldOp) (ฯ† : ๐“•.FieldOp) : ofFieldOpListF ฯ†s * ofFieldOpF ฯ† = ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†) โ€ข ofFieldOpF ฯ† * ofFieldOpListF ฯ†s + [ofFieldOpListF ฯ†s, ofFieldOpF ฯ†]โ‚›ca := by rw [superCommuteF_ofFieldOpListF_ofFieldOpF] simp lemma crPartF_mul_anPartF_eq_superCommuteF (ฯ† ฯ†' : ๐“•.FieldOp) : crPartF ฯ† * anPartF ฯ†' = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข anPartF ฯ†' * crPartF ฯ† + [crPartF ฯ†, anPartF ฯ†']โ‚›ca := by rw [superCommuteF_crPartF_anPartF] simp lemma anPartF_mul_crPartF_eq_superCommuteF (ฯ† ฯ†' : ๐“•.FieldOp) : anPartF ฯ† * crPartF ฯ†' = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข crPartF ฯ†' * anPartF ฯ† + [anPartF ฯ†, crPartF ฯ†']โ‚›ca := by rw [superCommuteF_anPartF_crPartF] simp lemma crPartF_mul_crPartF_eq_superCommuteF (ฯ† ฯ†' : ๐“•.FieldOp) : crPartF ฯ† * crPartF ฯ†' = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข crPartF ฯ†' * crPartF ฯ† + [crPartF ฯ†, crPartF ฯ†']โ‚›ca := by rw [superCommuteF_crPartF_crPartF] simp lemma anPartF_mul_anPartF_eq_superCommuteF (ฯ† ฯ†' : ๐“•.FieldOp) : anPartF ฯ† * anPartF ฯ†' = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†') โ€ข anPartF ฯ†' * anPartF ฯ† + [anPartF ฯ†, anPartF ฯ†']โ‚›ca := by rw [superCommuteF_anPartF_anPartF] simp lemma ofCrAnListF_mul_ofFieldOpListF_eq_superCommuteF (ฯ†s : List ๐“•.CrAnFieldOp) (ฯ†s' : List ๐“•.FieldOp) : ofCrAnListF ฯ†s * ofFieldOpListF ฯ†s' = ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†s') โ€ข ofFieldOpListF ฯ†s' * ofCrAnListF ฯ†s + [ofCrAnListF ฯ†s, ofFieldOpListF ฯ†s']โ‚›ca := by rw [superCommuteF_ofCrAnListF_ofFieldOpFsList] simp /-! ## Symmetry of the super commutator. -/ lemma superCommuteF_ofCrAnListF_ofCrAnListF_symm (ฯ†s ฯ†s' : List ๐“•.CrAnFieldOp) : [ofCrAnListF ฯ†s, ofCrAnListF ฯ†s']โ‚›ca = (- ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†s')) โ€ข [ofCrAnListF ฯ†s', ofCrAnListF ฯ†s]โ‚›ca := by rw [superCommuteF_ofCrAnListF_ofCrAnListF, superCommuteF_ofCrAnListF_ofCrAnListF, smul_sub] simp only [instCommGroup.eq_1, neg_smul, sub_neg_eq_add] rw [smul_smul] conv_rhs => rhs rw [exchangeSign_symm, exchangeSign_mul_self] simp only [one_smul] abel lemma superCommuteF_ofCrAnOpF_ofCrAnOpF_symm (ฯ† ฯ†' : ๐“•.CrAnFieldOp) : [ofCrAnOpF ฯ†, ofCrAnOpF ฯ†']โ‚›ca = (- ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯ†')) โ€ข [ofCrAnOpF ฯ†', ofCrAnOpF ฯ†]โ‚›ca := by rw [superCommuteF_ofCrAnOpF_ofCrAnOpF, superCommuteF_ofCrAnOpF_ofCrAnOpF] rw [smul_sub] simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc, neg_smul, sub_neg_eq_add] rw [smul_smul] conv_rhs => rhs rw [exchangeSign_symm, exchangeSign_mul_self] simp only [one_smul] abel /-! ## Splitting the super commutator on lists into sums. -/ lemma superCommuteF_ofCrAnListF_ofCrAnListF_cons (ฯ† : ๐“•.CrAnFieldOp) (ฯ†s ฯ†s' : List ๐“•.CrAnFieldOp) : [ofCrAnListF ฯ†s, ofCrAnListF (ฯ† :: ฯ†s')]โ‚›ca = [ofCrAnListF ฯ†s, ofCrAnOpF ฯ†]โ‚›ca * ofCrAnListF ฯ†s' + ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†) โ€ข ofCrAnOpF ฯ† * [ofCrAnListF ฯ†s, ofCrAnListF ฯ†s']โ‚›ca := by rw [superCommuteF_ofCrAnListF_ofCrAnListF] conv_rhs => lhs rw [โ† ofCrAnListF_singleton, superCommuteF_ofCrAnListF_ofCrAnListF, sub_mul, โ† ofCrAnListF_append] rhs rw [FieldStatistic.ofList_singleton, ofCrAnListF_append, ofCrAnListF_singleton, smul_mul_assoc, mul_assoc, โ† ofCrAnListF_append] conv_rhs => rhs rw [superCommuteF_ofCrAnListF_ofCrAnListF, mul_sub, smul_mul_assoc] simp only [instCommGroup.eq_1, List.cons_append, List.append_assoc, List.nil_append, Algebra.mul_smul_comm, Algebra.smul_mul_assoc, sub_add_sub_cancel, sub_right_inj] rw [โ† ofCrAnListF_cons, smul_smul, FieldStatistic.ofList_cons_eq_mul] simp only [instCommGroup, map_mul, mul_comm] lemma superCommuteF_ofCrAnListF_ofFieldOpListF_cons (ฯ† : ๐“•.FieldOp) (ฯ†s : List ๐“•.CrAnFieldOp) (ฯ†s' : List ๐“•.FieldOp) : [ofCrAnListF ฯ†s, ofFieldOpListF (ฯ† :: ฯ†s')]โ‚›ca = [ofCrAnListF ฯ†s, ofFieldOpF ฯ†]โ‚›ca * ofFieldOpListF ฯ†s' + ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†) โ€ข ofFieldOpF ฯ† * [ofCrAnListF ฯ†s, ofFieldOpListF ฯ†s']โ‚›ca := by rw [superCommuteF_ofCrAnListF_ofFieldOpFsList] conv_rhs => lhs rw [โ† ofFieldOpListF_singleton, superCommuteF_ofCrAnListF_ofFieldOpFsList, sub_mul, mul_assoc, โ† ofFieldOpListF_append] rhs rw [FieldStatistic.ofList_singleton, ofFieldOpListF_singleton, smul_mul_assoc, smul_mul_assoc, mul_assoc] conv_rhs => rhs rw [superCommuteF_ofCrAnListF_ofFieldOpFsList, mul_sub, smul_mul_assoc] simp only [instCommGroup, Algebra.smul_mul_assoc, List.singleton_append, Algebra.mul_smul_comm, sub_add_sub_cancel, sub_right_inj] rw [ofFieldOpListF_cons, mul_assoc, smul_smul, FieldStatistic.ofList_cons_eq_mul] simp [mul_comm] /-- For a field specification `๐“•`, and two lists `ฯ†s = ฯ†โ‚€โ€ฆฯ†โ‚™` and `ฯ†s'` of `๐“•.CrAnFieldOp` the following super commutation relation holds: `[ฯ†s', ฯ†โ‚€โ€ฆฯ†โ‚™]โ‚›ca = โˆ‘ i, ๐“ข(ฯ†s', ฯ†โ‚€โ€ฆฯ†แตขโ‚‹โ‚) โ€ข ฯ†โ‚€โ€ฆฯ†แตขโ‚‹โ‚ * [ฯ†s', ฯ†แตข]โ‚›ca * ฯ†แตขโ‚Šโ‚ โ€ฆ ฯ†โ‚™` The proof of this relation is via induction on the length of `ฯ†s`. -/ lemma superCommuteF_ofCrAnListF_ofCrAnListF_eq_sum (ฯ†s : List ๐“•.CrAnFieldOp) : (ฯ†s' : List ๐“•.CrAnFieldOp) โ†’ [ofCrAnListF ฯ†s, ofCrAnListF ฯ†s']โ‚›ca = โˆ‘ (n : Fin ฯ†s'.length), ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†s'.take n) โ€ข ofCrAnListF (ฯ†s'.take n) * [ofCrAnListF ฯ†s, ofCrAnOpF (ฯ†s'.get n)]โ‚›ca * ofCrAnListF (ฯ†s'.drop (n + 1)) | [] => by simp [โ† ofCrAnListF_nil, superCommuteF_ofCrAnListF_ofCrAnListF] | ฯ† :: ฯ†s' => by rw [superCommuteF_ofCrAnListF_ofCrAnListF_cons, superCommuteF_ofCrAnListF_ofCrAnListF_eq_sum ฯ†s ฯ†s'] conv_rhs => erw [Fin.sum_univ_succ] congr 1 ยท simp ยท simp [Finset.mul_sum, smul_smul, ofCrAnListF_cons, mul_assoc, FieldStatistic.ofList_cons_eq_mul, mul_comm] lemma superCommuteF_ofCrAnListF_ofFieldOpListF_eq_sum (ฯ†s : List ๐“•.CrAnFieldOp) : (ฯ†s' : List ๐“•.FieldOp) โ†’ [ofCrAnListF ฯ†s, ofFieldOpListF ฯ†s']โ‚›ca = โˆ‘ (n : Fin ฯ†s'.length), ๐“ข(๐“• |>โ‚› ฯ†s, ๐“• |>โ‚› ฯ†s'.take n) โ€ข ofFieldOpListF (ฯ†s'.take n) * [ofCrAnListF ฯ†s, ofFieldOpF (ฯ†s'.get n)]โ‚›ca * ofFieldOpListF (ฯ†s'.drop (n + 1)) | [] => by simp only [superCommuteF_ofCrAnListF_ofFieldOpFsList, instCommGroup, ofList_empty, exchangeSign_bosonic, one_smul, List.length_nil, Finset.univ_eq_empty, List.take_nil, List.get_eq_getElem, List.drop_nil, Finset.sum_empty] simp | ฯ† :: ฯ†s' => by rw [superCommuteF_ofCrAnListF_ofFieldOpListF_cons, superCommuteF_ofCrAnListF_ofFieldOpListF_eq_sum ฯ†s ฯ†s'] conv_rhs => erw [Fin.sum_univ_succ] congr 1 ยท simp ยท simp [Finset.mul_sum, smul_smul, ofFieldOpListF_cons, mul_assoc, FieldStatistic.ofList_cons_eq_mul, mul_comm] lemma summerCommute_jacobi_ofCrAnListF (ฯ†s1 ฯ†s2 ฯ†s3 : List ๐“•.CrAnFieldOp) : [ofCrAnListF ฯ†s1, [ofCrAnListF ฯ†s2, ofCrAnListF ฯ†s3]โ‚›ca]โ‚›ca = ๐“ข(๐“• |>โ‚› ฯ†s1, ๐“• |>โ‚› ฯ†s3) โ€ข (- ๐“ข(๐“• |>โ‚› ฯ†s2, ๐“• |>โ‚› ฯ†s3) โ€ข [ofCrAnListF ฯ†s3, [ofCrAnListF ฯ†s1, ofCrAnListF ฯ†s2]โ‚›ca]โ‚›ca - ๐“ข(๐“• |>โ‚› ฯ†s1, ๐“• |>โ‚› ฯ†s2) โ€ข [ofCrAnListF ฯ†s2, [ofCrAnListF ฯ†s3, ofCrAnListF ฯ†s1]โ‚›ca]โ‚›ca) := by repeat rw [superCommuteF_ofCrAnListF_ofCrAnListF] simp only [instCommGroup, map_sub, map_smul, neg_smul] repeat rw [superCommuteF_ofCrAnListF_ofCrAnListF] simp only [instCommGroup.eq_1, ofList_append_eq_mul, List.append_assoc] by_cases h1 : (๐“• |>โ‚› ฯ†s1) = bosonic <;> by_cases h2 : (๐“• |>โ‚› ฯ†s2) = bosonic <;> by_cases h3 : (๐“• |>โ‚› ฯ†s3) = bosonic ยท simp only [h1, h2, h3, mul_self, bosonic_exchangeSign, one_smul, exchangeSign_bosonic, neg_sub] abel ยท simp only [h1, h2, bosonic_exchangeSign, one_smul, mul_bosonic, mul_self, map_one, exchangeSign_bosonic, neg_sub] abel ยท simp only [h1, h3, mul_bosonic, bosonic_exchangeSign, one_smul, exchangeSign_bosonic, neg_sub, mul_self, map_one] abel ยท simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3 simp only [h1, h2, h3, mul_self, bosonic_exchangeSign, one_smul, fermionic_exchangeSign_fermionic, neg_smul, neg_sub, bosonic_mul_fermionic, sub_neg_eq_add, mul_bosonic, smul_add, exchangeSign_bosonic] abel ยท simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3 simp only [h1, h2, h3, mul_self, map_one, one_smul, exchangeSign_bosonic, mul_bosonic, bosonic_exchangeSign, bosonic_mul_fermionic, neg_sub] abel ยท simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3 simp only [h1, h2, h3, bosonic_mul_fermionic, fermionic_exchangeSign_fermionic, neg_smul, one_smul, sub_neg_eq_add, bosonic_exchangeSign, mul_bosonic, smul_add, exchangeSign_bosonic, neg_sub, mul_self] abel ยท simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3 simp only [h1, h2, h3, mul_bosonic, fermionic_exchangeSign_fermionic, neg_smul, one_smul, sub_neg_eq_add, exchangeSign_bosonic, bosonic_mul_fermionic, smul_add, mul_self, bosonic_exchangeSign, neg_sub] abel ยท simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3 simp only [h1, h2, h3, mul_self, map_one, one_smul, fermionic_exchangeSign_fermionic, neg_smul, neg_sub] abel /-! ## Interaction with grading. -/ lemma superCommuteF_grade {a b : ๐“•.FieldOpFreeAlgebra} {f1 f2 : FieldStatistic} (ha : a โˆˆ statisticSubmodule f1) (hb : b โˆˆ statisticSubmodule f2) : [a, b]โ‚›ca โˆˆ statisticSubmodule (f1 + f2) := by let p (a2 : ๐“•.FieldOpFreeAlgebra) (hx : a2 โˆˆ statisticSubmodule f2) : Prop := [a, a2]โ‚›ca โˆˆ statisticSubmodule (f1 + f2) change p b hb apply Submodule.span_induction (p := p) ยท intro x hx obtain โŸจฯ†s, rfl, hฯ†sโŸฉ := hx simp only [add_eq_mul, instCommGroup, p] let p (a2 : ๐“•.FieldOpFreeAlgebra) (hx : a2 โˆˆ statisticSubmodule f1) : Prop := [a2, ofCrAnListF ฯ†s]โ‚›ca โˆˆ statisticSubmodule (f1 + f2) change p a ha apply Submodule.span_induction (p := p) ยท intro x hx obtain โŸจฯ†s', rfl, hฯ†s'โŸฉ := hx simp only [add_eq_mul, instCommGroup, p] rw [superCommuteF_ofCrAnListF_ofCrAnListF] apply Submodule.sub_mem _ ยท apply ofCrAnListF_mem_statisticSubmodule_of rw [ofList_append_eq_mul, hฯ†s, hฯ†s'] ยท apply Submodule.smul_mem apply ofCrAnListF_mem_statisticSubmodule_of rw [ofList_append_eq_mul, hฯ†s, hฯ†s'] rw [mul_comm] ยท simp [p] ยท intro x y hx hy hp1 hp2 simp only [add_eq_mul, instCommGroup, map_add, LinearMap.add_apply, p] exact Submodule.add_mem _ hp1 hp2 ยท intro c x hx hp1 simp only [add_eq_mul, instCommGroup, map_smul, LinearMap.smul_apply, p] exact Submodule.smul_mem _ c hp1 ยท exact ha ยท simp [p] ยท intro x y hx hy hp1 hp2 simp only [add_eq_mul, instCommGroup, map_add, p] exact Submodule.add_mem _ hp1 hp2 ยท intro c x hx hp1 simp only [add_eq_mul, instCommGroup, map_smul, p] exact Submodule.smul_mem _ c hp1 ยท exact hb lemma superCommuteF_bosonic_bosonic {a b : ๐“•.FieldOpFreeAlgebra} (ha : a โˆˆ statisticSubmodule bosonic) (hb : b โˆˆ statisticSubmodule bosonic) : [a, b]โ‚›ca = a * b - b * a := by let p (a2 : ๐“•.FieldOpFreeAlgebra) (hx : a2 โˆˆ statisticSubmodule bosonic) : Prop := [a, a2]โ‚›ca = a * a2 - a2 * a change p b hb apply Submodule.span_induction (p := p) ยท intro x hx obtain โŸจฯ†s, rfl, hฯ†sโŸฉ := hx let p (a2 : ๐“•.FieldOpFreeAlgebra) (hx : a2 โˆˆ statisticSubmodule bosonic) : Prop := [a2, ofCrAnListF ฯ†s]โ‚›ca = a2 * ofCrAnListF ฯ†s - ofCrAnListF ฯ†s * a2 change p a ha apply Submodule.span_induction (p := p) ยท intro x hx obtain โŸจฯ†s', rfl, hฯ†s'โŸฉ := hx simp only [p] rw [superCommuteF_ofCrAnListF_ofCrAnListF] simp [hฯ†s, ofCrAnListF_append] ยท simp [p] ยท intro x y hx hy hp1 hp2 simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add] abel ยท intro c x hx hp1 simp_all [p, smul_sub] ยท exact ha ยท simp [p] ยท intro x y hx hy hp1 hp2 simp_all only [p, map_add, mul_add, add_mul] abel ยท intro c x hx hp1 simp_all [p, smul_sub] ยท exact hb lemma superCommuteF_bosonic_fermionic {a b : ๐“•.FieldOpFreeAlgebra} (ha : a โˆˆ statisticSubmodule bosonic) (hb : b โˆˆ statisticSubmodule fermionic) : [a, b]โ‚›ca = a * b - b * a := by let p (a2 : ๐“•.FieldOpFreeAlgebra) (hx : a2 โˆˆ statisticSubmodule fermionic) : Prop := [a, a2]โ‚›ca = a * a2 - a2 * a change p b hb apply Submodule.span_induction (p := p) ยท intro x hx obtain โŸจฯ†s, rfl, hฯ†sโŸฉ := hx let p (a2 : ๐“•.FieldOpFreeAlgebra) (hx : a2 โˆˆ statisticSubmodule bosonic) : Prop := [a2, ofCrAnListF ฯ†s]โ‚›ca = a2 * ofCrAnListF ฯ†s - ofCrAnListF ฯ†s * a2 change p a ha apply Submodule.span_induction (p := p) ยท intro x hx obtain โŸจฯ†s', rfl, hฯ†s'โŸฉ := hx simp only [p] rw [superCommuteF_ofCrAnListF_ofCrAnListF] simp [hฯ†s, hฯ†s', ofCrAnListF_append] ยท simp [p] ยท intro x y hx hy hp1 hp2 simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add] abel ยท intro c x hx hp1 simp_all [p, smul_sub] ยท exact ha ยท simp [p] ยท intro x y hx hy hp1 hp2 simp_all only [p, map_add, mul_add, add_mul] abel ยท intro c x hx hp1 simp_all [p, smul_sub] ยท exact hb lemma superCommuteF_fermionic_bonsonic {a b : ๐“•.FieldOpFreeAlgebra} (ha : a โˆˆ statisticSubmodule fermionic) (hb : b โˆˆ statisticSubmodule bosonic) : [a, b]โ‚›ca = a * b - b * a := by let p (a2 : ๐“•.FieldOpFreeAlgebra) (hx : a2 โˆˆ statisticSubmodule bosonic) : Prop := [a, a2]โ‚›ca = a * a2 - a2 * a change p b hb apply Submodule.span_induction (p := p) ยท intro x hx obtain โŸจฯ†s, rfl, hฯ†sโŸฉ := hx let p (a2 : ๐“•.FieldOpFreeAlgebra) (hx : a2 โˆˆ statisticSubmodule fermionic) : Prop := [a2, ofCrAnListF ฯ†s]โ‚›ca = a2 * ofCrAnListF ฯ†s - ofCrAnListF ฯ†s * a2 change p a ha apply Submodule.span_induction (p := p) ยท intro x hx obtain โŸจฯ†s', rfl, hฯ†s'โŸฉ := hx simp only [p] rw [superCommuteF_ofCrAnListF_ofCrAnListF] simp [hฯ†s, hฯ†s', ofCrAnListF_append] ยท simp [p] ยท intro x y hx hy hp1 hp2 simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add] abel ยท intro c x hx hp1 simp_all [p, smul_sub] ยท exact ha ยท simp [p] ยท intro x y hx hy hp1 hp2 simp_all only [map_add, mul_add, add_mul, p] abel ยท intro c x hx hp1 simp_all [p, smul_sub] ยท exact hb lemma superCommuteF_bonsonic {a b : ๐“•.FieldOpFreeAlgebra} (hb : b โˆˆ statisticSubmodule bosonic) : [a, b]โ‚›ca = a * b - b * a := by rw [โ† bosonicProjF_add_fermionicProjF a] simp only [map_add, LinearMap.add_apply] rw [superCommuteF_bosonic_bosonic (by simp) hb, superCommuteF_fermionic_bonsonic (by simp) hb] simp only [add_mul, mul_add] abel lemma bosonic_superCommuteF {a b : ๐“•.FieldOpFreeAlgebra} (ha : a โˆˆ statisticSubmodule bosonic) : [a, b]โ‚›ca = a * b - b * a := by rw [โ† bosonicProjF_add_fermionicProjF b] simp only [map_add, LinearMap.add_apply] rw [superCommuteF_bosonic_bosonic ha (by simp), superCommuteF_bosonic_fermionic ha (by simp)] simp only [add_mul, mul_add] abel lemma superCommuteF_bonsonic_symm {a b : ๐“•.FieldOpFreeAlgebra} (hb : b โˆˆ statisticSubmodule bosonic) : [a, b]โ‚›ca = - [b, a]โ‚›ca := by rw [bosonic_superCommuteF hb, superCommuteF_bonsonic hb] simp lemma bonsonic_superCommuteF_symm {a b : ๐“•.FieldOpFreeAlgebra} (ha : a โˆˆ statisticSubmodule bosonic) : [a, b]โ‚›ca = - [b, a]โ‚›ca := by rw [bosonic_superCommuteF ha, superCommuteF_bonsonic ha] simp lemma superCommuteF_fermionic_fermionic {a b : ๐“•.FieldOpFreeAlgebra} (ha : a โˆˆ statisticSubmodule fermionic) (hb : b โˆˆ statisticSubmodule fermionic) : [a, b]โ‚›ca = a * b + b * a := by let p (a2 : ๐“•.FieldOpFreeAlgebra) (hx : a2 โˆˆ statisticSubmodule fermionic) : Prop := [a, a2]โ‚›ca = a * a2 + a2 * a change p b hb apply Submodule.span_induction (p := p) ยท intro x hx obtain โŸจฯ†s, rfl, hฯ†sโŸฉ := hx let p (a2 : ๐“•.FieldOpFreeAlgebra) (hx : a2 โˆˆ statisticSubmodule fermionic) : Prop := [a2, ofCrAnListF ฯ†s]โ‚›ca = a2 * ofCrAnListF ฯ†s + ofCrAnListF ฯ†s * a2 change p a ha apply Submodule.span_induction (p := p) ยท intro x hx obtain โŸจฯ†s', rfl, hฯ†s'โŸฉ := hx simp only [p] rw [superCommuteF_ofCrAnListF_ofCrAnListF] simp [hฯ†s, hฯ†s', ofCrAnListF_append] ยท simp [p] ยท intro x y hx hy hp1 hp2 simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add] abel ยท intro c x hx hp1 simp_all [p, smul_sub] ยท exact ha ยท simp [p] ยท intro x y hx hy hp1 hp2 simp_all only [map_add, mul_add, add_mul, p] abel ยท intro c x hx hp1 simp_all [p, smul_sub] ยท exact hb lemma superCommuteF_fermionic_fermionic_symm {a b : ๐“•.FieldOpFreeAlgebra} (ha : a โˆˆ statisticSubmodule fermionic) (hb : b โˆˆ statisticSubmodule fermionic) : [a, b]โ‚›ca = [b, a]โ‚›ca := by rw [superCommuteF_fermionic_fermionic ha hb] rw [superCommuteF_fermionic_fermionic hb ha] abel lemma superCommuteF_expand_bosonicProjF_fermionicProjF (a b : ๐“•.FieldOpFreeAlgebra) : [a, b]โ‚›ca = bosonicProjF a * bosonicProjF b - bosonicProjF b * bosonicProjF a + bosonicProjF a * fermionicProjF b - fermionicProjF b * bosonicProjF a + fermionicProjF a * bosonicProjF b - bosonicProjF b * fermionicProjF a + fermionicProjF a * fermionicProjF b + fermionicProjF b * fermionicProjF a := by conv_lhs => rw [โ† bosonicProjF_add_fermionicProjF a, โ† bosonicProjF_add_fermionicProjF b] simp only [map_add, LinearMap.add_apply] rw [superCommuteF_bonsonic (by simp), superCommuteF_fermionic_bonsonic (by simp) (by simp), superCommuteF_bosonic_fermionic (by simp) (by simp), superCommuteF_fermionic_fermionic (by simp) (by simp)] abel lemma superCommuteF_ofCrAnListF_ofCrAnListF_bosonic_or_fermionic (ฯ†s ฯ†s' : List ๐“•.CrAnFieldOp) : [ofCrAnListF ฯ†s, ofCrAnListF ฯ†s']โ‚›ca โˆˆ statisticSubmodule bosonic โˆจ [ofCrAnListF ฯ†s, ofCrAnListF ฯ†s']โ‚›ca โˆˆ statisticSubmodule fermionic := by by_cases h1 : (๐“• |>โ‚› ฯ†s) = bosonic <;> by_cases h2 : (๐“• |>โ‚› ฯ†s') = bosonic ยท left have h : bosonic = bosonic + bosonic := by simp only [add_eq_mul, instCommGroup, mul_self] rfl rw [h] apply superCommuteF_grade apply ofCrAnListF_mem_statisticSubmodule_of _ _ h1 apply ofCrAnListF_mem_statisticSubmodule_of _ _ h2 ยท right have h : fermionic = bosonic + fermionic := by simp only [add_eq_mul, instCommGroup, mul_self] rfl rw [h] apply superCommuteF_grade apply ofCrAnListF_mem_statisticSubmodule_of _ _ h1 apply ofCrAnListF_mem_statisticSubmodule_of _ _ (by simpa using h2) ยท right have h : fermionic = fermionic + bosonic := by simp only [add_eq_mul, instCommGroup, mul_self] rfl rw [h] apply superCommuteF_grade apply ofCrAnListF_mem_statisticSubmodule_of _ _ (by simpa using h1) apply ofCrAnListF_mem_statisticSubmodule_of _ _ h2 ยท left have h : bosonic = fermionic + fermionic := by simp only [add_eq_mul, instCommGroup, mul_self] rfl rw [h] apply superCommuteF_grade apply ofCrAnListF_mem_statisticSubmodule_of _ _ (by simpa using h1) apply ofCrAnListF_mem_statisticSubmodule_of _ _ (by simpa using h2) lemma superCommuteF_ofCrAnOpF_ofCrAnOpF_bosonic_or_fermionic (ฯ† ฯ†' : ๐“•.CrAnFieldOp) : [ofCrAnOpF ฯ†, ofCrAnOpF ฯ†']โ‚›ca โˆˆ statisticSubmodule bosonic โˆจ [ofCrAnOpF ฯ†, ofCrAnOpF ฯ†']โ‚›ca โˆˆ statisticSubmodule fermionic := by rw [โ† ofCrAnListF_singleton, โ† ofCrAnListF_singleton] exact superCommuteF_ofCrAnListF_ofCrAnListF_bosonic_or_fermionic [ฯ†] [ฯ†'] lemma superCommuteF_superCommuteF_ofCrAnOpF_bosonic_or_fermionic (ฯ†1 ฯ†2 ฯ†3 : ๐“•.CrAnFieldOp) : [ofCrAnOpF ฯ†1, [ofCrAnOpF ฯ†2, ofCrAnOpF ฯ†3]โ‚›ca]โ‚›ca โˆˆ statisticSubmodule bosonic โˆจ [ofCrAnOpF ฯ†1, [ofCrAnOpF ฯ†2, ofCrAnOpF ฯ†3]โ‚›ca]โ‚›ca โˆˆ statisticSubmodule fermionic := by rcases superCommuteF_ofCrAnOpF_ofCrAnOpF_bosonic_or_fermionic ฯ†2 ฯ†3 with hs | hs <;> rcases ofCrAnOpF_bosonic_or_fermionic ฯ†1 with h1 | h1 ยท left have h : bosonic = bosonic + bosonic := by simp only [add_eq_mul, instCommGroup, mul_self] rfl rw [h] apply superCommuteF_grade h1 hs ยท right have h : fermionic = fermionic + bosonic := by simp only [add_eq_mul, instCommGroup, mul_self] rfl rw [h] apply superCommuteF_grade h1 hs ยท right have h : fermionic = bosonic + fermionic := by simp only [add_eq_mul, instCommGroup, mul_self] rfl rw [h] apply superCommuteF_grade h1 hs ยท left have h : bosonic = fermionic + fermionic := by simp only [add_eq_mul, instCommGroup, mul_self] rfl rw [h] apply superCommuteF_grade h1 hs lemma superCommuteF_bosonic_ofCrAnListF_eq_sum (a : ๐“•.FieldOpFreeAlgebra) (ฯ†s : List ๐“•.CrAnFieldOp) (ha : a โˆˆ statisticSubmodule bosonic) : [a, ofCrAnListF ฯ†s]โ‚›ca = โˆ‘ (n : Fin ฯ†s.length), ofCrAnListF (ฯ†s.take n) * [a, ofCrAnOpF (ฯ†s.get n)]โ‚›ca * ofCrAnListF (ฯ†s.drop (n + 1)) := by let p (a : ๐“•.FieldOpFreeAlgebra) (ha : a โˆˆ statisticSubmodule bosonic) : Prop := [a, ofCrAnListF ฯ†s]โ‚›ca = โˆ‘ (n : Fin ฯ†s.length), ofCrAnListF (ฯ†s.take n) * [a, ofCrAnOpF (ฯ†s.get n)]โ‚›ca * ofCrAnListF (ฯ†s.drop (n + 1)) change p a ha apply Submodule.span_induction (p := p) ยท intro a ha obtain โŸจฯ†s, rfl, hฯ†sโŸฉ := ha simp only [List.get_eq_getElem, p] rw [superCommuteF_ofCrAnListF_ofCrAnListF_eq_sum] congr funext n simp [hฯ†s] ยท simp [p] ยท intro x y hx hy hpx hpy simp_all only [List.get_eq_getElem, map_add, LinearMap.add_apply, p] rw [โ† Finset.sum_add_distrib] congr funext n simp [mul_add, add_mul] ยท intro c x hx hpx simp_all [p, Finset.smul_sum] ยท exact ha lemma superCommuteF_fermionic_ofCrAnListF_eq_sum (a : ๐“•.FieldOpFreeAlgebra) (ฯ†s : List ๐“•.CrAnFieldOp) (ha : a โˆˆ statisticSubmodule fermionic) : [a, ofCrAnListF ฯ†s]โ‚›ca = โˆ‘ (n : Fin ฯ†s.length), ๐“ข(fermionic, ๐“• |>โ‚› ฯ†s.take n) โ€ข ofCrAnListF (ฯ†s.take n) * [a, ofCrAnOpF (ฯ†s.get n)]โ‚›ca * ofCrAnListF (ฯ†s.drop (n + 1)) := by let p (a : ๐“•.FieldOpFreeAlgebra) (ha : a โˆˆ statisticSubmodule fermionic) : Prop := [a, ofCrAnListF ฯ†s]โ‚›ca = โˆ‘ (n : Fin ฯ†s.length), ๐“ข(fermionic, ๐“• |>โ‚› ฯ†s.take n) โ€ข ofCrAnListF (ฯ†s.take n) * [a, ofCrAnOpF (ฯ†s.get n)]โ‚›ca * ofCrAnListF (ฯ†s.drop (n + 1)) change p a ha apply Submodule.span_induction (p := p) ยท intro a ha obtain โŸจฯ†s, rfl, hฯ†sโŸฉ := ha simp only [instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, p] rw [superCommuteF_ofCrAnListF_ofCrAnListF_eq_sum] congr funext n simp [hฯ†s] ยท simp [p] ยท intro x y hx hy hpx hpy simp_all only [p, instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, map_add, LinearMap.add_apply] rw [โ† Finset.sum_add_distrib] congr funext n simp [mul_add, add_mul] ยท intro c x hx hpx simp_all only [p, instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, map_smul, LinearMap.smul_apply, Finset.smul_sum, Algebra.mul_smul_comm] congr funext x simp [smul_smul, mul_comm] ยท exact ha lemma statistic_neq_of_superCommuteF_fermionic {ฯ†s ฯ†s' : List ๐“•.CrAnFieldOp} (h : [ofCrAnListF ฯ†s, ofCrAnListF ฯ†s']โ‚›ca โˆˆ statisticSubmodule fermionic) : (๐“• |>โ‚› ฯ†s) โ‰  (๐“• |>โ‚› ฯ†s') โˆจ [ofCrAnListF ฯ†s, ofCrAnListF ฯ†s']โ‚›ca = 0 := by by_cases h0 : [ofCrAnListF ฯ†s, ofCrAnListF ฯ†s']โ‚›ca = 0 ยท simp [h0] simp only [ne_eq, h0, or_false] by_contra hn refine h0 (eq_zero_of_bosonic_and_fermionic ?_ h) by_cases hc : (๐“• |>โ‚› ฯ†s) = bosonic ยท have h1 : bosonic = bosonic + bosonic := by simp only [add_eq_mul, instCommGroup, mul_self] rfl rw [h1] apply superCommuteF_grade apply ofCrAnListF_mem_statisticSubmodule_of _ _ hc apply ofCrAnListF_mem_statisticSubmodule_of _ _ rw [โ† hn, hc] ยท have h1 : bosonic = fermionic + fermionic := by simp only [add_eq_mul, instCommGroup, mul_self] rfl rw [h1] apply superCommuteF_grade apply ofCrAnListF_mem_statisticSubmodule_of _ _ simpa using hc apply ofCrAnListF_mem_statisticSubmodule_of _ _ rw [โ† hn] simpa using hc end FieldOpFreeAlgebra end FieldSpecification