feat: lemmas relating to index notation
This commit is contained in:
parent
ac11a510cf
commit
672cc1ed8b
11 changed files with 510 additions and 17 deletions
|
@ -5,6 +5,7 @@ Authors: Joseph Tooby-Smith
|
|||
-/
|
||||
import HepLean.SpaceTime.LorentzVector.Complex.Two
|
||||
import HepLean.SpaceTime.MinkowskiMetric
|
||||
import LLMLean
|
||||
/-!
|
||||
|
||||
# Metric for complex Lorentz vectors
|
||||
|
@ -55,6 +56,22 @@ def contrMetric : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ complexContr wh
|
|||
def coMetricVal : (complexCo ⊗ complexCo).V :=
|
||||
coCoToMatrix.symm ((@minkowskiMatrix 3).map ofReal)
|
||||
|
||||
lemma coMetricVal_expand_tmul : coMetricVal =
|
||||
complexCoBasis (Sum.inl 0) ⊗ₜ[ℂ] complexCoBasis (Sum.inl 0)
|
||||
- complexCoBasis (Sum.inr 0) ⊗ₜ[ℂ] complexCoBasis (Sum.inr 0)
|
||||
- complexCoBasis (Sum.inr 1) ⊗ₜ[ℂ] complexCoBasis (Sum.inr 1)
|
||||
- complexCoBasis (Sum.inr 2) ⊗ₜ[ℂ] complexCoBasis (Sum.inr 2) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, coMetricVal, Fin.isValue]
|
||||
erw [coCoToMatrix_symm_expand_tmul]
|
||||
simp only [map_apply, ofReal_eq_coe, coe_smul, Fintype.sum_sum_type, Finset.univ_unique,
|
||||
Fin.default_eq_zero, Fin.isValue, Finset.sum_singleton, Fin.sum_univ_three, ne_eq, reduceCtorEq,
|
||||
not_false_eq_true, minkowskiMatrix.off_diag_zero, zero_smul, add_zero, zero_add, Sum.inr.injEq,
|
||||
zero_ne_one, Fin.reduceEq, one_ne_zero]
|
||||
rw [minkowskiMatrix.inl_0_inl_0, minkowskiMatrix.inr_i_inr_i,
|
||||
minkowskiMatrix.inr_i_inr_i, minkowskiMatrix.inr_i_inr_i]
|
||||
simp only [Fin.isValue, one_smul, neg_smul]
|
||||
rfl
|
||||
|
||||
/-- The metric `ηᵢᵢ` as a morphism `𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexCo ⊗ complexCo`,
|
||||
making its invariance under the action of `SL(2,ℂ)`. -/
|
||||
def coMetric : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexCo ⊗ complexCo where
|
||||
|
@ -84,5 +101,10 @@ def coMetric : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexCo ⊗ complexCo where
|
|||
simp only [lorentzGroupIsGroup_inv, SL2C.toLorentzGroup_apply_coe,
|
||||
LorentzGroup.toComplex_transpose_mul_minkowskiMatrix_mul_self]
|
||||
|
||||
lemma coMetric_apply_one : coMetric.hom (1 : ℂ) = coMetricVal := by
|
||||
change coMetric.hom.toFun (1 : ℂ) = coMetricVal
|
||||
simp [coMetric]
|
||||
|
||||
|
||||
end Lorentz
|
||||
end
|
||||
|
|
|
@ -34,6 +34,19 @@ def coCoToMatrix : (complexCo ⊗ complexCo).V ≃ₗ[ℂ]
|
|||
Finsupp.linearEquivFunOnFinite ℂ ℂ ((Fin 1 ⊕ Fin 3) × (Fin 1 ⊕ Fin 3)) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3)
|
||||
|
||||
/-- Expanding `coCoToMatrix` in terms of the standard basis. -/
|
||||
lemma coCoToMatrix_symm_expand_tmul (M : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℂ) :
|
||||
coCoToMatrix.symm M = ∑ i, ∑ j, M i j • (complexCoBasis i ⊗ₜ[ℂ] complexCoBasis j) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, coCoToMatrix, LinearEquiv.trans_symm,
|
||||
LinearEquiv.trans_apply, Basis.repr_symm_apply]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℂ (s := Finset.univ)]
|
||||
· erw [Finset.sum_product]
|
||||
refine Finset.sum_congr rfl (fun i _ => Finset.sum_congr rfl (fun j _ => ?_))
|
||||
erw [Basis.tensorProduct_apply complexCoBasis complexCoBasis i j]
|
||||
rfl
|
||||
· simp
|
||||
|
||||
|
||||
/-- Equivalence of `complexContr ⊗ complexCo` to `4 x 4` complex matrices. -/
|
||||
def contrCoToMatrix : (complexContr ⊗ complexCo).V ≃ₗ[ℂ]
|
||||
Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℂ :=
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue