refactor: Rename superCommute for CrAnAlgebra
This commit is contained in:
parent
32aefb7eb7
commit
289050c829
13 changed files with 428 additions and 425 deletions
|
@ -26,7 +26,7 @@ open FieldStatistic
|
|||
/-- The super commutor on the creation and annihlation algebra. For two bosonic operators
|
||||
or a bosonic and fermionic operator this corresponds to the usual commutator
|
||||
whilst for two fermionic operators this corresponds to the anti-commutator. -/
|
||||
noncomputable def superCommute : 𝓕.CrAnAlgebra →ₗ[ℂ] 𝓕.CrAnAlgebra →ₗ[ℂ] 𝓕.CrAnAlgebra :=
|
||||
noncomputable def superCommuteF : 𝓕.CrAnAlgebra →ₗ[ℂ] 𝓕.CrAnAlgebra →ₗ[ℂ] 𝓕.CrAnAlgebra :=
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs =>
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs' =>
|
||||
ofCrAnList (φs ++ φs') - 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofCrAnList (φs' ++ φs)
|
||||
|
@ -34,7 +34,7 @@ noncomputable def superCommute : 𝓕.CrAnAlgebra →ₗ[ℂ] 𝓕.CrAnAlgebra
|
|||
/-- The super commutor on the creation and annihlation algebra. For two bosonic operators
|
||||
or a bosonic and fermionic operator this corresponds to the usual commutator
|
||||
whilst for two fermionic operators this corresponds to the anti-commutator. -/
|
||||
scoped[FieldSpecification.CrAnAlgebra] notation "[" φs "," φs' "]ₛca" => superCommute φs φs'
|
||||
scoped[FieldSpecification.CrAnAlgebra] notation "[" φs "," φs' "]ₛca" => superCommuteF φs φs'
|
||||
|
||||
/-!
|
||||
|
||||
|
@ -42,35 +42,35 @@ scoped[FieldSpecification.CrAnAlgebra] notation "[" φs "," φs' "]ₛca" => sup
|
|||
|
||||
-/
|
||||
|
||||
lemma superCommute_ofCrAnList_ofCrAnList (φs φs' : List 𝓕.CrAnStates) :
|
||||
lemma superCommuteF_ofCrAnList_ofCrAnList (φs φs' : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs, ofCrAnList φs']ₛca =
|
||||
ofCrAnList (φs ++ φs') - 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofCrAnList (φs' ++ φs) := by
|
||||
rw [← ofListBasis_eq_ofList, ← ofListBasis_eq_ofList]
|
||||
simp only [superCommute, Basis.constr_basis]
|
||||
simp only [superCommuteF, Basis.constr_basis]
|
||||
|
||||
lemma superCommute_ofCrAnState_ofCrAnState (φ φ' : 𝓕.CrAnStates) :
|
||||
lemma superCommuteF_ofCrAnState_ofCrAnState (φ φ' : 𝓕.CrAnStates) :
|
||||
[ofCrAnState φ, ofCrAnState φ']ₛca =
|
||||
ofCrAnState φ * ofCrAnState φ' - 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • ofCrAnState φ' * ofCrAnState φ := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList, ofCrAnList_append]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList, ofCrAnList_append]
|
||||
congr
|
||||
rw [ofCrAnList_append]
|
||||
rw [FieldStatistic.ofList_singleton, FieldStatistic.ofList_singleton, smul_mul_assoc]
|
||||
|
||||
lemma superCommute_ofCrAnList_ofStatesList (φcas : List 𝓕.CrAnStates) (φs : List 𝓕.States) :
|
||||
lemma superCommuteF_ofCrAnList_ofStatesList (φcas : List 𝓕.CrAnStates) (φs : List 𝓕.States) :
|
||||
[ofCrAnList φcas, ofStateList φs]ₛca = ofCrAnList φcas * ofStateList φs -
|
||||
𝓢(𝓕 |>ₛ φcas, 𝓕 |>ₛ φs) • ofStateList φs * ofCrAnList φcas := by
|
||||
conv_lhs => rw [ofStateList_sum]
|
||||
rw [map_sum]
|
||||
conv_lhs =>
|
||||
enter [2, x]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList, CrAnSection.statistics_eq_state_statistics,
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList, CrAnSection.statistics_eq_state_statistics,
|
||||
ofCrAnList_append, ofCrAnList_append]
|
||||
rw [Finset.sum_sub_distrib, ← Finset.mul_sum, ← Finset.smul_sum,
|
||||
← Finset.sum_mul, ← ofStateList_sum]
|
||||
simp
|
||||
|
||||
lemma superCommute_ofStateList_ofStatesList (φ : List 𝓕.States) (φs : List 𝓕.States) :
|
||||
lemma superCommuteF_ofStateList_ofStatesList (φ : List 𝓕.States) (φs : List 𝓕.States) :
|
||||
[ofStateList φ, ofStateList φs]ₛca = ofStateList φ * ofStateList φs -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) • ofStateList φs * ofStateList φ := by
|
||||
conv_lhs => rw [ofStateList_sum]
|
||||
|
@ -78,24 +78,24 @@ lemma superCommute_ofStateList_ofStatesList (φ : List 𝓕.States) (φs : List
|
|||
Algebra.smul_mul_assoc]
|
||||
conv_lhs =>
|
||||
enter [2, x]
|
||||
rw [superCommute_ofCrAnList_ofStatesList]
|
||||
rw [superCommuteF_ofCrAnList_ofStatesList]
|
||||
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, ← ofStateList_sum]
|
||||
|
||||
lemma superCommute_ofState_ofStatesList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
lemma superCommuteF_ofState_ofStatesList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
[ofState φ, ofStateList φs]ₛca = ofState φ * ofStateList φs -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) • ofStateList φs * ofState φ := by
|
||||
rw [← ofStateList_singleton, superCommute_ofStateList_ofStatesList, ofStateList_singleton]
|
||||
rw [← ofStateList_singleton, superCommuteF_ofStateList_ofStatesList, ofStateList_singleton]
|
||||
simp
|
||||
|
||||
lemma superCommute_ofStateList_ofState (φs : List 𝓕.States) (φ : 𝓕.States) :
|
||||
lemma superCommuteF_ofStateList_ofState (φs : List 𝓕.States) (φ : 𝓕.States) :
|
||||
[ofStateList φs, ofState φ]ₛca = ofStateList φs * ofState φ -
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ) • ofState φ * ofStateList φs := by
|
||||
rw [← ofStateList_singleton, superCommute_ofStateList_ofStatesList, ofStateList_singleton]
|
||||
rw [← ofStateList_singleton, superCommuteF_ofStateList_ofStatesList, ofStateList_singleton]
|
||||
simp
|
||||
|
||||
lemma superCommute_anPart_crPart (φ φ' : 𝓕.States) :
|
||||
lemma superCommuteF_anPart_crPart (φ φ' : 𝓕.States) :
|
||||
[anPart φ, crPart φ']ₛca =
|
||||
anPart φ * crPart φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • crPart φ' * anPart φ := by
|
||||
|
@ -107,24 +107,24 @@ lemma superCommute_anPart_crPart (φ φ' : 𝓕.States) :
|
|||
sub_self]
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [anPart_position, crPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.outAsymp φ, States.position φ' =>
|
||||
simp only [anPart_posAsymp, crPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.inAsymp φ' =>
|
||||
simp only [anPart_position, crPart_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp only [List.singleton_append, instCommGroup.eq_1, crAnStatistics,
|
||||
FieldStatistic.ofList_singleton, Function.comp_apply, crAnStatesToStates_prod, ←
|
||||
ofCrAnList_append]
|
||||
| States.outAsymp φ, States.inAsymp φ' =>
|
||||
simp only [anPart_posAsymp, crPart_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma superCommute_crPart_anPart (φ φ' : 𝓕.States) :
|
||||
lemma superCommuteF_crPart_anPart (φ φ' : 𝓕.States) :
|
||||
[crPart φ, anPart φ']ₛca =
|
||||
crPart φ * anPart φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
|
@ -138,22 +138,22 @@ lemma superCommute_crPart_anPart (φ φ' : 𝓕.States) :
|
|||
sub_self]
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [crPart_position, anPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.outAsymp φ' =>
|
||||
simp only [crPart_position, anPart_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.inAsymp φ, States.position φ' =>
|
||||
simp only [crPart_negAsymp, anPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.inAsymp φ, States.outAsymp φ' =>
|
||||
simp only [crPart_negAsymp, anPart_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma superCommute_crPart_crPart (φ φ' : 𝓕.States) :
|
||||
lemma superCommuteF_crPart_crPart (φ φ' : 𝓕.States) :
|
||||
[crPart φ, crPart φ']ₛca =
|
||||
crPart φ * crPart φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
|
@ -166,22 +166,22 @@ lemma superCommute_crPart_crPart (φ φ' : 𝓕.States) :
|
|||
simp only [crPart_posAsymp, map_zero, mul_zero, instCommGroup.eq_1, smul_zero, zero_mul, sub_self]
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [crPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.inAsymp φ' =>
|
||||
simp only [crPart_position, crPart_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.inAsymp φ, States.position φ' =>
|
||||
simp only [crPart_negAsymp, crPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.inAsymp φ, States.inAsymp φ' =>
|
||||
simp only [crPart_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma superCommute_anPart_anPart (φ φ' : 𝓕.States) :
|
||||
lemma superCommuteF_anPart_anPart (φ φ' : 𝓕.States) :
|
||||
[anPart φ, anPart φ']ₛca =
|
||||
anPart φ * anPart φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
|
@ -193,38 +193,38 @@ lemma superCommute_anPart_anPart (φ φ' : 𝓕.States) :
|
|||
simp
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [anPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.outAsymp φ' =>
|
||||
simp only [anPart_position, anPart_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.outAsymp φ, States.position φ' =>
|
||||
simp only [anPart_posAsymp, anPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.outAsymp φ, States.outAsymp φ' =>
|
||||
simp only [anPart_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma superCommute_crPart_ofStateList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
lemma superCommuteF_crPart_ofStateList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
[crPart φ, ofStateList φs]ₛca =
|
||||
crPart φ * ofStateList φs - 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) • ofStateList φs *
|
||||
crPart φ := by
|
||||
match φ with
|
||||
| States.inAsymp φ =>
|
||||
simp only [crPart_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofStatesList]
|
||||
rw [← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
| States.position φ =>
|
||||
simp only [crPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofStatesList]
|
||||
rw [← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
| States.outAsymp φ =>
|
||||
simp
|
||||
|
||||
lemma superCommute_anPart_ofStateList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
lemma superCommuteF_anPart_ofStateList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
[anPart φ, ofStateList φs]ₛca =
|
||||
anPart φ * ofStateList φs - 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) •
|
||||
ofStateList φs * anPart φ := by
|
||||
|
@ -233,97 +233,97 @@ lemma superCommute_anPart_ofStateList (φ : 𝓕.States) (φs : List 𝓕.States
|
|||
simp
|
||||
| States.position φ =>
|
||||
simp only [anPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofStatesList]
|
||||
rw [← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
| States.outAsymp φ =>
|
||||
simp only [anPart_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofStatesList]
|
||||
rw [← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
|
||||
lemma superCommute_crPart_ofState (φ φ' : 𝓕.States) :
|
||||
lemma superCommuteF_crPart_ofState (φ φ' : 𝓕.States) :
|
||||
[crPart φ, ofState φ']ₛca =
|
||||
crPart φ * ofState φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • ofState φ' * crPart φ := by
|
||||
rw [← ofStateList_singleton, superCommute_crPart_ofStateList]
|
||||
rw [← ofStateList_singleton, superCommuteF_crPart_ofStateList]
|
||||
simp
|
||||
|
||||
lemma superCommute_anPart_ofState (φ φ' : 𝓕.States) :
|
||||
lemma superCommuteF_anPart_ofState (φ φ' : 𝓕.States) :
|
||||
[anPart φ, ofState φ']ₛca =
|
||||
anPart φ * ofState φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • ofState φ' * anPart φ := by
|
||||
rw [← ofStateList_singleton, superCommute_anPart_ofStateList]
|
||||
rw [← ofStateList_singleton, superCommuteF_anPart_ofStateList]
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## Mul equal superCommute
|
||||
## Mul equal superCommuteF
|
||||
|
||||
Lemmas which rewrite a multiplication of two elements of the algebra as their commuted
|
||||
multiplication with a sign plus the super commutor.
|
||||
|
||||
-/
|
||||
lemma ofCrAnList_mul_ofCrAnList_eq_superCommute (φs φs' : List 𝓕.CrAnStates) :
|
||||
lemma ofCrAnList_mul_ofCrAnList_eq_superCommuteF (φs φs' : List 𝓕.CrAnStates) :
|
||||
ofCrAnList φs * ofCrAnList φs' = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofCrAnList φs' * ofCrAnList φs
|
||||
+ [ofCrAnList φs, ofCrAnList φs']ₛca := by
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [ofCrAnList_append]
|
||||
|
||||
lemma ofCrAnState_mul_ofCrAnList_eq_superCommute (φ : 𝓕.CrAnStates) (φs' : List 𝓕.CrAnStates) :
|
||||
lemma ofCrAnState_mul_ofCrAnList_eq_superCommuteF (φ : 𝓕.CrAnStates) (φs' : List 𝓕.CrAnStates) :
|
||||
ofCrAnState φ * ofCrAnList φs' = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • ofCrAnList φs' * ofCrAnState φ
|
||||
+ [ofCrAnState φ, ofCrAnList φs']ₛca := by
|
||||
rw [← ofCrAnList_singleton, ofCrAnList_mul_ofCrAnList_eq_superCommute]
|
||||
rw [← ofCrAnList_singleton, ofCrAnList_mul_ofCrAnList_eq_superCommuteF]
|
||||
simp
|
||||
|
||||
lemma ofStateList_mul_ofStateList_eq_superCommute (φs φs' : List 𝓕.States) :
|
||||
lemma ofStateList_mul_ofStateList_eq_superCommuteF (φs φs' : List 𝓕.States) :
|
||||
ofStateList φs * ofStateList φs' = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofStateList φs' * ofStateList φs
|
||||
+ [ofStateList φs, ofStateList φs']ₛca := by
|
||||
rw [superCommute_ofStateList_ofStatesList]
|
||||
rw [superCommuteF_ofStateList_ofStatesList]
|
||||
simp
|
||||
|
||||
lemma ofState_mul_ofStateList_eq_superCommute (φ : 𝓕.States) (φs' : List 𝓕.States) :
|
||||
lemma ofState_mul_ofStateList_eq_superCommuteF (φ : 𝓕.States) (φs' : List 𝓕.States) :
|
||||
ofState φ * ofStateList φs' = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • ofStateList φs' * ofState φ
|
||||
+ [ofState φ, ofStateList φs']ₛca := by
|
||||
rw [superCommute_ofState_ofStatesList]
|
||||
rw [superCommuteF_ofState_ofStatesList]
|
||||
simp
|
||||
|
||||
lemma ofStateList_mul_ofState_eq_superCommute (φs : List 𝓕.States) (φ : 𝓕.States) :
|
||||
lemma ofStateList_mul_ofState_eq_superCommuteF (φs : List 𝓕.States) (φ : 𝓕.States) :
|
||||
ofStateList φs * ofState φ = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ) • ofState φ * ofStateList φs
|
||||
+ [ofStateList φs, ofState φ]ₛca := by
|
||||
rw [superCommute_ofStateList_ofState]
|
||||
rw [superCommuteF_ofStateList_ofState]
|
||||
simp
|
||||
|
||||
lemma crPart_mul_anPart_eq_superCommute (φ φ' : 𝓕.States) :
|
||||
lemma crPart_mul_anPart_eq_superCommuteF (φ φ' : 𝓕.States) :
|
||||
crPart φ * anPart φ' =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • anPart φ' * crPart φ +
|
||||
[crPart φ, anPart φ']ₛca := by
|
||||
rw [superCommute_crPart_anPart]
|
||||
rw [superCommuteF_crPart_anPart]
|
||||
simp
|
||||
|
||||
lemma anPart_mul_crPart_eq_superCommute (φ φ' : 𝓕.States) :
|
||||
lemma anPart_mul_crPart_eq_superCommuteF (φ φ' : 𝓕.States) :
|
||||
anPart φ * crPart φ' =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
crPart φ' * anPart φ +
|
||||
[anPart φ, crPart φ']ₛca := by
|
||||
rw [superCommute_anPart_crPart]
|
||||
rw [superCommuteF_anPart_crPart]
|
||||
simp
|
||||
|
||||
lemma crPart_mul_crPart_eq_superCommute (φ φ' : 𝓕.States) :
|
||||
lemma crPart_mul_crPart_eq_superCommuteF (φ φ' : 𝓕.States) :
|
||||
crPart φ * crPart φ' =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • crPart φ' * crPart φ +
|
||||
[crPart φ, crPart φ']ₛca := by
|
||||
rw [superCommute_crPart_crPart]
|
||||
rw [superCommuteF_crPart_crPart]
|
||||
simp
|
||||
|
||||
lemma anPart_mul_anPart_eq_superCommute (φ φ' : 𝓕.States) :
|
||||
lemma anPart_mul_anPart_eq_superCommuteF (φ φ' : 𝓕.States) :
|
||||
anPart φ * anPart φ' = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • anPart φ' * anPart φ +
|
||||
[anPart φ, anPart φ']ₛca := by
|
||||
rw [superCommute_anPart_anPart]
|
||||
rw [superCommuteF_anPart_anPart]
|
||||
simp
|
||||
|
||||
lemma ofCrAnList_mul_ofStateList_eq_superCommute (φs : List 𝓕.CrAnStates) (φs' : List 𝓕.States) :
|
||||
lemma ofCrAnList_mul_ofStateList_eq_superCommuteF (φs : List 𝓕.CrAnStates) (φs' : List 𝓕.States) :
|
||||
ofCrAnList φs * ofStateList φs' = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofStateList φs' * ofCrAnList φs
|
||||
+ [ofCrAnList φs, ofStateList φs']ₛca := by
|
||||
rw [superCommute_ofCrAnList_ofStatesList]
|
||||
rw [superCommuteF_ofCrAnList_ofStatesList]
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
@ -332,10 +332,10 @@ lemma ofCrAnList_mul_ofStateList_eq_superCommute (φs : List 𝓕.CrAnStates) (
|
|||
|
||||
-/
|
||||
|
||||
lemma superCommute_ofCrAnList_ofCrAnList_symm (φs φs' : List 𝓕.CrAnStates) :
|
||||
lemma superCommuteF_ofCrAnList_ofCrAnList_symm (φs φs' : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs, ofCrAnList φs']ₛca =
|
||||
(- 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs')) • [ofCrAnList φs', ofCrAnList φs]ₛca := by
|
||||
rw [superCommute_ofCrAnList_ofCrAnList, superCommute_ofCrAnList_ofCrAnList, smul_sub]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList, superCommuteF_ofCrAnList_ofCrAnList, smul_sub]
|
||||
simp only [instCommGroup.eq_1, neg_smul, sub_neg_eq_add]
|
||||
rw [smul_smul]
|
||||
conv_rhs =>
|
||||
|
@ -344,10 +344,10 @@ lemma superCommute_ofCrAnList_ofCrAnList_symm (φs φs' : List 𝓕.CrAnStates)
|
|||
simp only [one_smul]
|
||||
abel
|
||||
|
||||
lemma superCommute_ofCrAnState_ofCrAnState_symm (φ φ' : 𝓕.CrAnStates) :
|
||||
lemma superCommuteF_ofCrAnState_ofCrAnState_symm (φ φ' : 𝓕.CrAnStates) :
|
||||
[ofCrAnState φ, ofCrAnState φ']ₛca =
|
||||
(- 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ')) • [ofCrAnState φ', ofCrAnState φ]ₛca := by
|
||||
rw [superCommute_ofCrAnState_ofCrAnState, superCommute_ofCrAnState_ofCrAnState]
|
||||
rw [superCommuteF_ofCrAnState_ofCrAnState, superCommuteF_ofCrAnState_ofCrAnState]
|
||||
rw [smul_sub]
|
||||
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc, neg_smul, sub_neg_eq_add]
|
||||
rw [smul_smul]
|
||||
|
@ -362,41 +362,41 @@ lemma superCommute_ofCrAnState_ofCrAnState_symm (φ φ' : 𝓕.CrAnStates) :
|
|||
## Splitting the super commutor on lists into sums.
|
||||
|
||||
-/
|
||||
lemma superCommute_ofCrAnList_ofCrAnList_cons (φ : 𝓕.CrAnStates) (φs φs' : List 𝓕.CrAnStates) :
|
||||
lemma superCommuteF_ofCrAnList_ofCrAnList_cons (φ : 𝓕.CrAnStates) (φs φs' : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs, ofCrAnList (φ :: φs')]ₛca =
|
||||
[ofCrAnList φs, ofCrAnState φ]ₛca * ofCrAnList φs' +
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ)
|
||||
• ofCrAnState φ * [ofCrAnList φs, ofCrAnList φs']ₛca := by
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
conv_rhs =>
|
||||
lhs
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList, sub_mul, ← ofCrAnList_append]
|
||||
rw [← ofCrAnList_singleton, superCommuteF_ofCrAnList_ofCrAnList, sub_mul, ← ofCrAnList_append]
|
||||
rhs
|
||||
rw [FieldStatistic.ofList_singleton, ofCrAnList_append, ofCrAnList_singleton, smul_mul_assoc,
|
||||
mul_assoc, ← ofCrAnList_append]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [superCommute_ofCrAnList_ofCrAnList, mul_sub, smul_mul_assoc]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList, 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 [← ofCrAnList_cons, smul_smul, FieldStatistic.ofList_cons_eq_mul]
|
||||
simp only [instCommGroup, map_mul, mul_comm]
|
||||
|
||||
lemma superCommute_ofCrAnList_ofStateList_cons (φ : 𝓕.States) (φs : List 𝓕.CrAnStates)
|
||||
lemma superCommuteF_ofCrAnList_ofStateList_cons (φ : 𝓕.States) (φs : List 𝓕.CrAnStates)
|
||||
(φs' : List 𝓕.States) : [ofCrAnList φs, ofStateList (φ :: φs')]ₛca =
|
||||
[ofCrAnList φs, ofState φ]ₛca * ofStateList φs' +
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ) • ofState φ * [ofCrAnList φs, ofStateList φs']ₛca := by
|
||||
rw [superCommute_ofCrAnList_ofStatesList]
|
||||
rw [superCommuteF_ofCrAnList_ofStatesList]
|
||||
conv_rhs =>
|
||||
lhs
|
||||
rw [← ofStateList_singleton, superCommute_ofCrAnList_ofStatesList, sub_mul, mul_assoc,
|
||||
rw [← ofStateList_singleton, superCommuteF_ofCrAnList_ofStatesList, sub_mul, mul_assoc,
|
||||
← ofStateList_append]
|
||||
rhs
|
||||
rw [FieldStatistic.ofList_singleton, ofStateList_singleton, smul_mul_assoc,
|
||||
smul_mul_assoc, mul_assoc]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [superCommute_ofCrAnList_ofStatesList, mul_sub, smul_mul_assoc]
|
||||
rw [superCommuteF_ofCrAnList_ofStatesList, 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 [ofStateList_cons, mul_assoc, smul_smul, FieldStatistic.ofList_cons_eq_mul]
|
||||
|
@ -407,33 +407,33 @@ Within the creation and annihilation algebra, we have that
|
|||
`[φᶜᵃs, φᶜᵃ₀ … φᶜᵃₙ]ₛca = ∑ i, sᵢ • φᶜᵃs₀ … φᶜᵃᵢ₋₁ * [φᶜᵃs, φᶜᵃᵢ]ₛca * φᶜᵃᵢ₊₁ … φᶜᵃₙ`
|
||||
where `sᵢ` is the exchange sign for `φᶜᵃs` and `φᶜᵃs₀ … φᶜᵃᵢ₋₁`.
|
||||
-/
|
||||
lemma superCommute_ofCrAnList_ofCrAnList_eq_sum (φs : List 𝓕.CrAnStates) :
|
||||
lemma superCommuteF_ofCrAnList_ofCrAnList_eq_sum (φs : List 𝓕.CrAnStates) :
|
||||
(φs' : List 𝓕.CrAnStates) → [ofCrAnList φs, ofCrAnList φs']ₛca =
|
||||
∑ (n : Fin φs'.length), 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs'.take n) •
|
||||
ofCrAnList (φs'.take n) * [ofCrAnList φs, ofCrAnState (φs'.get n)]ₛca *
|
||||
ofCrAnList (φs'.drop (n + 1))
|
||||
| [] => by
|
||||
simp [← ofCrAnList_nil, superCommute_ofCrAnList_ofCrAnList]
|
||||
simp [← ofCrAnList_nil, superCommuteF_ofCrAnList_ofCrAnList]
|
||||
| φ :: φs' => by
|
||||
rw [superCommute_ofCrAnList_ofCrAnList_cons, superCommute_ofCrAnList_ofCrAnList_eq_sum φs φs']
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList_cons, superCommuteF_ofCrAnList_ofCrAnList_eq_sum φs φs']
|
||||
conv_rhs => erw [Fin.sum_univ_succ]
|
||||
congr 1
|
||||
· simp
|
||||
· simp [Finset.mul_sum, smul_smul, ofCrAnList_cons, mul_assoc,
|
||||
FieldStatistic.ofList_cons_eq_mul, mul_comm]
|
||||
|
||||
lemma superCommute_ofCrAnList_ofStateList_eq_sum (φs : List 𝓕.CrAnStates) : (φs' : List 𝓕.States) →
|
||||
lemma superCommuteF_ofCrAnList_ofStateList_eq_sum (φs : List 𝓕.CrAnStates) : (φs' : List 𝓕.States) →
|
||||
[ofCrAnList φs, ofStateList φs']ₛca =
|
||||
∑ (n : Fin φs'.length), 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs'.take n) •
|
||||
ofStateList (φs'.take n) * [ofCrAnList φs, ofState (φs'.get n)]ₛca *
|
||||
ofStateList (φs'.drop (n + 1))
|
||||
| [] => by
|
||||
simp only [superCommute_ofCrAnList_ofStatesList, instCommGroup, ofList_empty,
|
||||
simp only [superCommuteF_ofCrAnList_ofStatesList, 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 [superCommute_ofCrAnList_ofStateList_cons, superCommute_ofCrAnList_ofStateList_eq_sum φs φs']
|
||||
rw [superCommuteF_ofCrAnList_ofStateList_cons, superCommuteF_ofCrAnList_ofStateList_eq_sum φs φs']
|
||||
conv_rhs => erw [Fin.sum_univ_succ]
|
||||
congr 1
|
||||
· simp
|
||||
|
@ -445,9 +445,9 @@ lemma summerCommute_jacobi_ofCrAnList (φs1 φs2 φs3 : List 𝓕.CrAnStates) :
|
|||
𝓢(𝓕 |>ₛ φs1, 𝓕 |>ₛ φs3) •
|
||||
(- 𝓢(𝓕 |>ₛ φs2, 𝓕 |>ₛ φs3) • [ofCrAnList φs3, [ofCrAnList φs1, ofCrAnList φs2]ₛca]ₛca -
|
||||
𝓢(𝓕 |>ₛ φs1, 𝓕 |>ₛ φs2) • [ofCrAnList φs2, [ofCrAnList φs3, ofCrAnList φs1]ₛca]ₛca) := by
|
||||
repeat rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
repeat rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp only [instCommGroup, map_sub, map_smul, neg_smul]
|
||||
repeat rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
repeat rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp only [instCommGroup.eq_1, ofList_append_eq_mul, List.append_assoc]
|
||||
by_cases h1 : (𝓕 |>ₛ φs1) = bosonic <;>
|
||||
by_cases h2 : (𝓕 |>ₛ φs2) = bosonic <;>
|
||||
|
@ -489,7 +489,7 @@ lemma summerCommute_jacobi_ofCrAnList (φs1 φs2 φs3 : List 𝓕.CrAnStates) :
|
|||
|
||||
-/
|
||||
|
||||
lemma superCommute_grade {a b : 𝓕.CrAnAlgebra} {f1 f2 : FieldStatistic}
|
||||
lemma superCommuteF_grade {a b : 𝓕.CrAnAlgebra} {f1 f2 : FieldStatistic}
|
||||
(ha : a ∈ statisticSubmodule f1) (hb : b ∈ statisticSubmodule f2) :
|
||||
[a, b]ₛca ∈ statisticSubmodule (f1 + f2) := by
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule f2) : Prop :=
|
||||
|
@ -506,7 +506,7 @@ lemma superCommute_grade {a b : 𝓕.CrAnAlgebra} {f1 f2 : FieldStatistic}
|
|||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [add_eq_mul, instCommGroup, p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
apply Submodule.sub_mem _
|
||||
· apply ofCrAnList_mem_statisticSubmodule_of
|
||||
rw [ofList_append_eq_mul, hφs, hφs']
|
||||
|
@ -531,7 +531,7 @@ lemma superCommute_grade {a b : 𝓕.CrAnAlgebra} {f1 f2 : FieldStatistic}
|
|||
exact Submodule.smul_mem _ c hp1
|
||||
· exact hb
|
||||
|
||||
lemma superCommute_bosonic_bosonic {a b : 𝓕.CrAnAlgebra}
|
||||
lemma superCommuteF_bosonic_bosonic {a b : 𝓕.CrAnAlgebra}
|
||||
(ha : a ∈ statisticSubmodule bosonic) (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop :=
|
||||
|
@ -547,7 +547,7 @@ lemma superCommute_bosonic_bosonic {a b : 𝓕.CrAnAlgebra}
|
|||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
|
@ -564,7 +564,7 @@ lemma superCommute_bosonic_bosonic {a b : 𝓕.CrAnAlgebra}
|
|||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommute_bosonic_fermionic {a b : 𝓕.CrAnAlgebra}
|
||||
lemma superCommuteF_bosonic_fermionic {a b : 𝓕.CrAnAlgebra}
|
||||
(ha : a ∈ statisticSubmodule bosonic) (hb : b ∈ statisticSubmodule fermionic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop :=
|
||||
|
@ -580,7 +580,7 @@ lemma superCommute_bosonic_fermionic {a b : 𝓕.CrAnAlgebra}
|
|||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, hφs', ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
|
@ -597,7 +597,7 @@ lemma superCommute_bosonic_fermionic {a b : 𝓕.CrAnAlgebra}
|
|||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommute_fermionic_bonsonic {a b : 𝓕.CrAnAlgebra}
|
||||
lemma superCommuteF_fermionic_bonsonic {a b : 𝓕.CrAnAlgebra}
|
||||
(ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop :=
|
||||
|
@ -613,7 +613,7 @@ lemma superCommute_fermionic_bonsonic {a b : 𝓕.CrAnAlgebra}
|
|||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, hφs', ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
|
@ -630,33 +630,33 @@ lemma superCommute_fermionic_bonsonic {a b : 𝓕.CrAnAlgebra}
|
|||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommute_bonsonic {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) :
|
||||
lemma superCommuteF_bonsonic {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
rw [← bosonicProj_add_fermionicProj a]
|
||||
simp only [map_add, LinearMap.add_apply]
|
||||
rw [superCommute_bosonic_bosonic (by simp) hb, superCommute_fermionic_bonsonic (by simp) hb]
|
||||
rw [superCommuteF_bosonic_bosonic (by simp) hb, superCommuteF_fermionic_bonsonic (by simp) hb]
|
||||
simp only [add_mul, mul_add]
|
||||
abel
|
||||
|
||||
lemma bosonic_superCommute {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) :
|
||||
lemma bosonic_superCommuteF {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
rw [← bosonicProj_add_fermionicProj b]
|
||||
simp only [map_add, LinearMap.add_apply]
|
||||
rw [superCommute_bosonic_bosonic ha (by simp), superCommute_bosonic_fermionic ha (by simp)]
|
||||
rw [superCommuteF_bosonic_bosonic ha (by simp), superCommuteF_bosonic_fermionic ha (by simp)]
|
||||
simp only [add_mul, mul_add]
|
||||
abel
|
||||
|
||||
lemma superCommute_bonsonic_symm {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) :
|
||||
lemma superCommuteF_bonsonic_symm {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = - [b, a]ₛca := by
|
||||
rw [bosonic_superCommute hb, superCommute_bonsonic hb]
|
||||
rw [bosonic_superCommuteF hb, superCommuteF_bonsonic hb]
|
||||
simp
|
||||
|
||||
lemma bonsonic_superCommute_symm {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) :
|
||||
lemma bonsonic_superCommuteF_symm {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = - [b, a]ₛca := by
|
||||
rw [bosonic_superCommute ha, superCommute_bonsonic ha]
|
||||
rw [bosonic_superCommuteF ha, superCommuteF_bonsonic ha]
|
||||
simp
|
||||
|
||||
lemma superCommute_fermionic_fermionic {a b : 𝓕.CrAnAlgebra}
|
||||
lemma superCommuteF_fermionic_fermionic {a b : 𝓕.CrAnAlgebra}
|
||||
(ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule fermionic) :
|
||||
[a, b]ₛca = a * b + b * a := by
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop :=
|
||||
|
@ -672,7 +672,7 @@ lemma superCommute_fermionic_fermionic {a b : 𝓕.CrAnAlgebra}
|
|||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, hφs', ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
|
@ -689,27 +689,27 @@ lemma superCommute_fermionic_fermionic {a b : 𝓕.CrAnAlgebra}
|
|||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommute_fermionic_fermionic_symm {a b : 𝓕.CrAnAlgebra}
|
||||
lemma superCommuteF_fermionic_fermionic_symm {a b : 𝓕.CrAnAlgebra}
|
||||
(ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule fermionic) :
|
||||
[a, b]ₛca = [b, a]ₛca := by
|
||||
rw [superCommute_fermionic_fermionic ha hb]
|
||||
rw [superCommute_fermionic_fermionic hb ha]
|
||||
rw [superCommuteF_fermionic_fermionic ha hb]
|
||||
rw [superCommuteF_fermionic_fermionic hb ha]
|
||||
abel
|
||||
|
||||
lemma superCommute_expand_bosonicProj_fermionicProj (a b : 𝓕.CrAnAlgebra) :
|
||||
lemma superCommuteF_expand_bosonicProj_fermionicProj (a b : 𝓕.CrAnAlgebra) :
|
||||
[a, b]ₛca = bosonicProj a * bosonicProj b - bosonicProj b * bosonicProj a +
|
||||
bosonicProj a * fermionicProj b - fermionicProj b * bosonicProj a +
|
||||
fermionicProj a * bosonicProj b - bosonicProj b * fermionicProj a +
|
||||
fermionicProj a * fermionicProj b + fermionicProj b * fermionicProj a := by
|
||||
conv_lhs => rw [← bosonicProj_add_fermionicProj a, ← bosonicProj_add_fermionicProj b]
|
||||
simp only [map_add, LinearMap.add_apply]
|
||||
rw [superCommute_bonsonic (by simp),
|
||||
superCommute_fermionic_bonsonic (by simp) (by simp),
|
||||
superCommute_bosonic_fermionic (by simp) (by simp),
|
||||
superCommute_fermionic_fermionic (by simp) (by simp)]
|
||||
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 superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List 𝓕.CrAnStates) :
|
||||
lemma superCommuteF_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule bosonic ∨
|
||||
[ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule fermionic := by
|
||||
by_cases h1 : (𝓕 |>ₛ φs) = bosonic <;> by_cases h2 : (𝓕 |>ₛ φs') = bosonic
|
||||
|
@ -718,7 +718,7 @@ lemma superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List
|
|||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h1
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h2
|
||||
· right
|
||||
|
@ -726,7 +726,7 @@ lemma superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List
|
|||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h1
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h2)
|
||||
· right
|
||||
|
@ -734,7 +734,7 @@ lemma superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List
|
|||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h1)
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h2
|
||||
· left
|
||||
|
@ -742,47 +742,47 @@ lemma superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List
|
|||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h1)
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h2)
|
||||
|
||||
lemma superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic (φ φ' : 𝓕.CrAnStates) :
|
||||
lemma superCommuteF_ofCrAnState_ofCrAnState_bosonic_or_fermionic (φ φ' : 𝓕.CrAnStates) :
|
||||
[ofCrAnState φ, ofCrAnState φ']ₛca ∈ statisticSubmodule bosonic ∨
|
||||
[ofCrAnState φ, ofCrAnState φ']ₛca ∈ statisticSubmodule fermionic := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
exact superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ] [φ']
|
||||
exact superCommuteF_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ] [φ']
|
||||
|
||||
lemma superCommute_superCommute_ofCrAnState_bosonic_or_fermionic (φ1 φ2 φ3 : 𝓕.CrAnStates) :
|
||||
lemma superCommuteF_superCommuteF_ofCrAnState_bosonic_or_fermionic (φ1 φ2 φ3 : 𝓕.CrAnStates) :
|
||||
[ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca ∈ statisticSubmodule bosonic ∨
|
||||
[ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca ∈ statisticSubmodule fermionic := by
|
||||
rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φ2 φ3 with hs | hs
|
||||
rcases superCommuteF_ofCrAnState_ofCrAnState_bosonic_or_fermionic φ2 φ3 with hs | hs
|
||||
<;> rcases ofCrAnState_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 superCommute_grade h1 hs
|
||||
apply superCommuteF_grade h1 hs
|
||||
· right
|
||||
have h : fermionic = fermionic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade h1 hs
|
||||
apply superCommuteF_grade h1 hs
|
||||
· right
|
||||
have h : fermionic = bosonic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade h1 hs
|
||||
apply superCommuteF_grade h1 hs
|
||||
· left
|
||||
have h : bosonic = fermionic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade h1 hs
|
||||
apply superCommuteF_grade h1 hs
|
||||
|
||||
lemma superCommute_bosonic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List 𝓕.CrAnStates)
|
||||
lemma superCommuteF_bosonic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List 𝓕.CrAnStates)
|
||||
(ha : a ∈ statisticSubmodule bosonic) :
|
||||
[a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length),
|
||||
ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca *
|
||||
|
@ -796,7 +796,7 @@ lemma superCommute_bosonic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List
|
|||
· intro a ha
|
||||
obtain ⟨φs, rfl, hφs⟩ := ha
|
||||
simp only [List.get_eq_getElem, p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList_eq_sum]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList_eq_sum]
|
||||
congr
|
||||
funext n
|
||||
simp [hφs]
|
||||
|
@ -811,7 +811,7 @@ lemma superCommute_bosonic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List
|
|||
simp_all [p, Finset.smul_sum]
|
||||
· exact ha
|
||||
|
||||
lemma superCommute_fermionic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List 𝓕.CrAnStates)
|
||||
lemma superCommuteF_fermionic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List 𝓕.CrAnStates)
|
||||
(ha : a ∈ statisticSubmodule fermionic) :
|
||||
[a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), 𝓢(fermionic, 𝓕 |>ₛ φs.take n) •
|
||||
ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca *
|
||||
|
@ -825,7 +825,7 @@ lemma superCommute_fermionic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : Lis
|
|||
· intro a ha
|
||||
obtain ⟨φs, rfl, hφs⟩ := ha
|
||||
simp only [instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList_eq_sum]
|
||||
rw [superCommuteF_ofCrAnList_ofCrAnList_eq_sum]
|
||||
congr
|
||||
funext n
|
||||
simp [hφs]
|
||||
|
@ -845,7 +845,7 @@ lemma superCommute_fermionic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : Lis
|
|||
simp [smul_smul, mul_comm]
|
||||
· exact ha
|
||||
|
||||
lemma statistic_neq_of_superCommute_fermionic {φs φs' : List 𝓕.CrAnStates}
|
||||
lemma statistic_neq_of_superCommuteF_fermionic {φs φs' : List 𝓕.CrAnStates}
|
||||
(h : [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule fermionic) :
|
||||
(𝓕 |>ₛ φs) ≠ (𝓕 |>ₛ φs') ∨ [ofCrAnList φs, ofCrAnList φs']ₛca = 0 := by
|
||||
by_cases h0 : [ofCrAnList φs, ofCrAnList φs']ₛca = 0
|
||||
|
@ -858,7 +858,7 @@ lemma statistic_neq_of_superCommute_fermionic {φs φs' : List 𝓕.CrAnStates}
|
|||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h1]
|
||||
apply superCommute_grade
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ hc
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _
|
||||
rw [← hn, hc]
|
||||
|
@ -866,7 +866,7 @@ lemma statistic_neq_of_superCommute_fermionic {φs φs' : List 𝓕.CrAnStates}
|
|||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h1]
|
||||
apply superCommute_grade
|
||||
apply superCommuteF_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _
|
||||
simpa using hc
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue