refactor: Lint
This commit is contained in:
parent
691b7e112e
commit
a1d3616a18
10 changed files with 65 additions and 62 deletions
|
@ -93,6 +93,9 @@ import HepLean.SpaceTime.LorentzVector.LorentzAction
|
|||
import HepLean.SpaceTime.LorentzVector.Modules
|
||||
import HepLean.SpaceTime.LorentzVector.NormOne
|
||||
import HepLean.SpaceTime.MinkowskiMetric
|
||||
import HepLean.SpaceTime.PauliMatrices.AsTensor
|
||||
import HepLean.SpaceTime.PauliMatrices.Basic
|
||||
import HepLean.SpaceTime.PauliMatrices.SelfAdjoint
|
||||
import HepLean.SpaceTime.SL2C.Basic
|
||||
import HepLean.SpaceTime.WeylFermion.Basic
|
||||
import HepLean.SpaceTime.WeylFermion.Contraction
|
||||
|
|
|
@ -7,7 +7,6 @@ import HepLean.SpaceTime.MinkowskiMetric
|
|||
import HepLean.SpaceTime.PauliMatrices.SelfAdjoint
|
||||
import Mathlib.LinearAlgebra.Matrix.SpecialLinearGroup
|
||||
import Mathlib.Tactic.Polyrith
|
||||
import LLMLean
|
||||
/-!
|
||||
# Lorentz vector as a self-adjoint matrix
|
||||
|
||||
|
@ -34,11 +33,11 @@ lemma toSelfAdjointMatrix_apply (x : LorentzVector 3) : toSelfAdjointMatrix x =
|
|||
x (Sum.inl 0) • ⟨PauliMatrix.σ0, PauliMatrix.σ0_selfAdjoint⟩
|
||||
- x (Sum.inr 0) • ⟨PauliMatrix.σ1, PauliMatrix.σ1_selfAdjoint⟩
|
||||
- x (Sum.inr 1) • ⟨PauliMatrix.σ2, PauliMatrix.σ2_selfAdjoint⟩
|
||||
- x (Sum.inr 2) • ⟨PauliMatrix.σ3, PauliMatrix.σ3_selfAdjoint⟩ := by
|
||||
- x (Sum.inr 2) • ⟨PauliMatrix.σ3, PauliMatrix.σ3_selfAdjoint⟩ := by
|
||||
simp only [toSelfAdjointMatrix, PauliMatrix.σSAL, LinearEquiv.trans_apply, Basis.repr_symm_apply,
|
||||
Basis.coe_mk, Fin.isValue]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℝ (s := Finset.univ)]
|
||||
· change (∑ i : Fin 1 ⊕ Fin 3, x i • PauliMatrix.σSAL' i) = _
|
||||
· change (∑ i : Fin 1 ⊕ Fin 3, x i • PauliMatrix.σSAL' i) = _
|
||||
simp [Fin.sum_univ_three, PauliMatrix.σSAL']
|
||||
apply Subtype.ext
|
||||
simp only [Fin.isValue, AddSubgroup.coe_add, selfAdjoint.val_smul, smul_neg,
|
||||
|
@ -50,7 +49,7 @@ lemma toSelfAdjointMatrix_apply_coe (x : LorentzVector 3) : (toSelfAdjointMatrix
|
|||
x (Sum.inl 0) • PauliMatrix.σ0
|
||||
- x (Sum.inr 0) • PauliMatrix.σ1
|
||||
- x (Sum.inr 1) • PauliMatrix.σ2
|
||||
- x (Sum.inr 2) • PauliMatrix.σ3 := by
|
||||
- x (Sum.inr 2) • PauliMatrix.σ3 := by
|
||||
rw [toSelfAdjointMatrix_apply]
|
||||
simp only [Fin.isValue, AddSubgroupClass.coe_sub, selfAdjoint.val_smul]
|
||||
|
||||
|
@ -61,21 +60,21 @@ lemma toSelfAdjointMatrix_stdBasis (i : Fin 1 ⊕ Fin 3) :
|
|||
| Sum.inl 0 =>
|
||||
simp [LorentzVector.stdBasis]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp [PauliMatrix.σSAL, PauliMatrix.σSAL' ]
|
||||
simp [PauliMatrix.σSAL, PauliMatrix.σSAL']
|
||||
| Sum.inr 0 =>
|
||||
simp [LorentzVector.stdBasis]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp [PauliMatrix.σSAL, PauliMatrix.σSAL' ]
|
||||
simp [PauliMatrix.σSAL, PauliMatrix.σSAL']
|
||||
refine Eq.symm (PauliMatrix.selfAdjoint_ext rfl rfl rfl rfl)
|
||||
| Sum.inr 1 =>
|
||||
simp [LorentzVector.stdBasis]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp [PauliMatrix.σSAL, PauliMatrix.σSAL' ]
|
||||
simp [PauliMatrix.σSAL, PauliMatrix.σSAL']
|
||||
refine Eq.symm (PauliMatrix.selfAdjoint_ext rfl rfl rfl rfl)
|
||||
| Sum.inr 2 =>
|
||||
simp [LorentzVector.stdBasis]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp [PauliMatrix.σSAL, PauliMatrix.σSAL' ]
|
||||
simp [PauliMatrix.σSAL, PauliMatrix.σSAL']
|
||||
refine Eq.symm (PauliMatrix.selfAdjoint_ext rfl rfl rfl rfl)
|
||||
|
||||
@[simp]
|
||||
|
|
|
@ -74,6 +74,8 @@ lemma complexCoBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 1 ⊕ Fin 3) :
|
|||
|
||||
-/
|
||||
|
||||
/-- The semilinear map including real Lorentz vectors into complex contravariant
|
||||
lorentz vectors. -/
|
||||
def inclCongrRealLorentz : LorentzVector 3 →ₛₗ[Complex.ofReal] complexContr where
|
||||
toFun v := {val := ofReal ∘ v}
|
||||
map_add' x y := by
|
||||
|
@ -97,9 +99,9 @@ lemma inclCongrRealLorentz_val (v : LorentzVector 3) :
|
|||
lemma complexContrBasis_of_real (i : Fin 1 ⊕ Fin 3) :
|
||||
(complexContrBasis i) = inclCongrRealLorentz (LorentzVector.stdBasis i) := by
|
||||
apply Lorentz.ContrℂModule.ext
|
||||
simp [complexContrBasis, inclCongrRealLorentz, LorentzVector.stdBasis, ]
|
||||
simp [complexContrBasis, inclCongrRealLorentz, LorentzVector.stdBasis]
|
||||
ext j
|
||||
simp
|
||||
simp only [Function.comp_apply, ofReal_eq_coe]
|
||||
erw [Pi.basisFun_apply]
|
||||
change (Pi.single i 1) j = _
|
||||
exact Eq.symm (Pi.apply_single (fun _ => ofReal') (congrFun rfl) i 1 j)
|
||||
|
@ -111,11 +113,10 @@ lemma inclCongrRealLorentz_ρ (M : SL(2, ℂ)) (v : LorentzVector 3) :
|
|||
rw [complexContrBasis_ρ_val, inclCongrRealLorentz_val, inclCongrRealLorentz_val]
|
||||
rw [LorentzGroup.toComplex_mulVec_ofReal]
|
||||
apply congrArg
|
||||
simp
|
||||
simp only [SL2C.toLorentzGroup_apply_coe]
|
||||
rw [SL2C.repLorentzVector_apply_eq_mulVec]
|
||||
rfl
|
||||
|
||||
|
||||
lemma SL2CRep_ρ_basis (M : SL(2, ℂ)) (i : Fin 1 ⊕ Fin 3) :
|
||||
(complexContr.ρ M) (complexContrBasis i) =
|
||||
∑ j, (SL2C.toLorentzGroup M).1 j i •
|
||||
|
|
|
@ -32,7 +32,6 @@ structure ContrℂModule where
|
|||
|
||||
namespace ContrℂModule
|
||||
|
||||
|
||||
/-- The equivalence between `ContrℂModule` and `Fin 1 ⊕ Fin 3 → ℂ`. -/
|
||||
def toFin13ℂFun : ContrℂModule ≃ (Fin 1 ⊕ Fin 3 → ℂ) where
|
||||
toFun v := v.val
|
||||
|
|
|
@ -242,7 +242,6 @@ variable (Λ Λ' : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ)
|
|||
/-- The dual of a matrix with respect to the Minkowski metric. -/
|
||||
def dual : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ := η * Λᵀ * η
|
||||
|
||||
|
||||
@[simp]
|
||||
lemma dual_id : @dual d 1 = 1 := by
|
||||
simpa only [dual, transpose_one, mul_one] using minkowskiMatrix.sq
|
||||
|
@ -279,7 +278,7 @@ lemma det_dual : (dual Λ).det = Λ.det := by
|
|||
simp
|
||||
|
||||
lemma dual_apply (μ ν : Fin 1 ⊕ Fin d) :
|
||||
dual Λ μ ν = η μ μ * Λ ν μ * η ν ν := by
|
||||
dual Λ μ ν = η μ μ * Λ ν μ * η ν ν := by
|
||||
simp only [dual, minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, mul_diagonal,
|
||||
diagonal_mul, transpose_apply, diagonal_apply_eq]
|
||||
|
||||
|
|
|
@ -33,7 +33,7 @@ open SpaceTime
|
|||
/-- The tensor `σ^μ^a^{dot a}` based on the Pauli-matrices as an element of
|
||||
`complexContr ⊗ leftHanded ⊗ rightHanded`. -/
|
||||
def asTensor : (complexContr ⊗ leftHanded ⊗ rightHanded).V :=
|
||||
∑ i, complexContrBasis i ⊗ₜ leftRightToMatrix.symm (σSA i)
|
||||
∑ i, complexContrBasis i ⊗ₜ leftRightToMatrix.symm (σSA i)
|
||||
|
||||
/-- The tensor `σ^μ^a^{dot a}` based on the Pauli-matrices as a morphism,
|
||||
`𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ leftHanded ⊗ rightHanded` manifesting
|
||||
|
@ -56,7 +56,7 @@ def asConsTensor : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ leftHanded ⊗
|
|||
let x' : ℂ := x
|
||||
change x' • asTensor =
|
||||
(TensorProduct.map (complexContr.ρ M) (
|
||||
TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M))) (x' • asTensor)
|
||||
TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M))) (x' • asTensor)
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
|
||||
apply congrArg
|
||||
nth_rewrite 2 [asTensor]
|
||||
|
@ -64,13 +64,13 @@ def asConsTensor : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ leftHanded ⊗
|
|||
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
|
||||
map_sum, map_tmul]
|
||||
symm
|
||||
calc _ = ∑ x, ((complexContr.ρ M) (complexContrBasis x) ⊗ₜ[ℂ]
|
||||
calc _ = ∑ x, ((complexContr.ρ M) (complexContrBasis x) ⊗ₜ[ℂ]
|
||||
leftRightToMatrix.symm (SL2C.repSelfAdjointMatrix M (σSA x))) := by
|
||||
refine Finset.sum_congr rfl (fun x _ => ?_)
|
||||
rw [← leftRightToMatrix_ρ_symm_selfAdjoint]
|
||||
rfl
|
||||
_ = ∑ x, ((∑ i, (SL2C.toLorentzGroup M).1 i x • (complexContrBasis i)) ⊗ₜ[ℂ]
|
||||
∑ j, leftRightToMatrix.symm ((SL2C.toLorentzGroup M⁻¹).1 x j • (σSA j))) := by
|
||||
∑ j, leftRightToMatrix.symm ((SL2C.toLorentzGroup M⁻¹).1 x j • (σSA j))) := by
|
||||
refine Finset.sum_congr rfl (fun x _ => ?_)
|
||||
rw [SL2CRep_ρ_basis, SL2C.repSelfAdjointMatrix_σSA]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, SL2C.toLorentzGroup_apply_coe,
|
||||
|
@ -78,7 +78,7 @@ def asConsTensor : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ leftHanded ⊗
|
|||
Finset.sum_singleton, map_inv, lorentzGroupIsGroup_inv, AddSubgroup.coe_add,
|
||||
selfAdjoint.val_smul, AddSubgroup.val_finset_sum, map_add, map_sum]
|
||||
_ = ∑ x, ∑ i, ∑ j, ((SL2C.toLorentzGroup M).1 i x • (complexContrBasis i)) ⊗ₜ[ℂ]
|
||||
leftRightToMatrix.symm.toLinearMap ((SL2C.toLorentzGroup M⁻¹).1 x j • (σSA j)) := by
|
||||
leftRightToMatrix.symm.toLinearMap ((SL2C.toLorentzGroup M⁻¹).1 x j • (σSA j)) := by
|
||||
refine Finset.sum_congr rfl (fun x _ => ?_)
|
||||
rw [sum_tmul]
|
||||
refine Finset.sum_congr rfl (fun i _ => ?_)
|
||||
|
@ -87,32 +87,32 @@ def asConsTensor : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ leftHanded ⊗
|
|||
_ = ∑ x, ∑ i, ∑ j, ((SL2C.toLorentzGroup M).1 i x • (complexContrBasis i)) ⊗ₜ[ℂ]
|
||||
((SL2C.toLorentzGroup M⁻¹).1 x j • leftRightToMatrix.symm ((σSA j))) := by
|
||||
refine Finset.sum_congr rfl (fun x _ => (Finset.sum_congr rfl (fun i _ =>
|
||||
( Finset.sum_congr rfl (fun j _ => ?_) )) ))
|
||||
(Finset.sum_congr rfl (fun j _ => ?_)))))
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, SL2C.toLorentzGroup_apply_coe,
|
||||
map_inv, lorentzGroupIsGroup_inv, LinearMap.map_smul_of_tower, LinearEquiv.coe_coe,
|
||||
tmul_smul]
|
||||
_ = ∑ x, ∑ i, ∑ j, ((SL2C.toLorentzGroup M).1 i x * (SL2C.toLorentzGroup M⁻¹).1 x j)
|
||||
• ((complexContrBasis i)) ⊗ₜ[ℂ] leftRightToMatrix.symm ((σSA j)) := by
|
||||
refine Finset.sum_congr rfl (fun x _ => (Finset.sum_congr rfl (fun i _ =>
|
||||
( Finset.sum_congr rfl (fun j _ => ?_) )) ))
|
||||
(Finset.sum_congr rfl (fun j _ => ?_)))))
|
||||
rw [smul_tmul, smul_smul, tmul_smul]
|
||||
_ = ∑ i, ∑ x, ∑ j, ((SL2C.toLorentzGroup M).1 i x * (SL2C.toLorentzGroup M⁻¹).1 x j)
|
||||
_ = ∑ i, ∑ x, ∑ j, ((SL2C.toLorentzGroup M).1 i x * (SL2C.toLorentzGroup M⁻¹).1 x j)
|
||||
• ((complexContrBasis i)) ⊗ₜ[ℂ] leftRightToMatrix.symm ((σSA j)) := Finset.sum_comm
|
||||
_ = ∑ i, ∑ j, ∑ x, ((SL2C.toLorentzGroup M).1 i x * (SL2C.toLorentzGroup M⁻¹).1 x j)
|
||||
• ((complexContrBasis i)) ⊗ₜ[ℂ] leftRightToMatrix.symm ((σSA j)) :=
|
||||
Finset.sum_congr rfl (fun x _ => Finset.sum_comm)
|
||||
_ = ∑ i, ∑ j, (∑ x, (SL2C.toLorentzGroup M).1 i x * (SL2C.toLorentzGroup M⁻¹).1 x j)
|
||||
Finset.sum_congr rfl (fun x _ => Finset.sum_comm)
|
||||
_ = ∑ i, ∑ j, (∑ x, (SL2C.toLorentzGroup M).1 i x * (SL2C.toLorentzGroup M⁻¹).1 x j)
|
||||
• ((complexContrBasis i)) ⊗ₜ[ℂ] leftRightToMatrix.symm ((σSA j)) := by
|
||||
refine Finset.sum_congr rfl (fun i _ => (Finset.sum_congr rfl (fun j _ => ?_) ))
|
||||
refine Finset.sum_congr rfl (fun i _ => (Finset.sum_congr rfl (fun j _ => ?_)))
|
||||
rw [Finset.sum_smul]
|
||||
_ = ∑ i, ∑ j, ((1 : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ) i j)
|
||||
_ = ∑ i, ∑ j, ((1 : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ) i j)
|
||||
• ((complexContrBasis i)) ⊗ₜ[ℂ] leftRightToMatrix.symm ((σSA j)) := by
|
||||
refine Finset.sum_congr rfl (fun i _ => (Finset.sum_congr rfl (fun j _ => ?_) ))
|
||||
refine Finset.sum_congr rfl (fun i _ => (Finset.sum_congr rfl (fun j _ => ?_)))
|
||||
congr
|
||||
change ((SL2C.toLorentzGroup M) * (SL2C.toLorentzGroup M⁻¹)).1 i j = _
|
||||
rw [← SL2C.toLorentzGroup.map_mul]
|
||||
simp only [mul_inv_cancel, _root_.map_one, lorentzGroupIsGroup_one_coe]
|
||||
_ = ∑ i, ((1 : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ) i i)
|
||||
_ = ∑ i, ((1 : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ) i i)
|
||||
• ((complexContrBasis i)) ⊗ₜ[ℂ] leftRightToMatrix.symm ((σSA i)) := by
|
||||
refine Finset.sum_congr rfl (fun i _ => ?_)
|
||||
refine Finset.sum_eq_single i (fun b _ hb => ?_) (fun hb => ?_)
|
||||
|
|
|
@ -8,7 +8,6 @@ import Mathlib.RepresentationTheory.Rep
|
|||
import HepLean.Tensors.Basic
|
||||
import Mathlib.Logic.Equiv.TransferInstance
|
||||
import HepLean.SpaceTime.LorentzGroup.Basic
|
||||
import LLMLean
|
||||
/-!
|
||||
|
||||
## Pauli matrices
|
||||
|
@ -69,12 +68,10 @@ lemma σ0_σ0_trace : Matrix.trace (σ0 * σ0) = 2 := by
|
|||
lemma σ0_σ1_trace : Matrix.trace (σ0 * σ1) = 0 := by
|
||||
simp [σ0, σ1]
|
||||
|
||||
|
||||
@[simp]
|
||||
lemma σ0_σ2_trace : Matrix.trace (σ0 * σ2) = 0 := by
|
||||
simp [σ0, σ2]
|
||||
|
||||
|
||||
@[simp]
|
||||
lemma σ0_σ3_trace : Matrix.trace (σ0 * σ3) = 0 := by
|
||||
simp [σ0, σ3]
|
||||
|
|
|
@ -9,7 +9,6 @@ import HepLean.Tensors.Basic
|
|||
import Mathlib.Logic.Equiv.TransferInstance
|
||||
import HepLean.SpaceTime.LorentzGroup.Basic
|
||||
import HepLean.SpaceTime.PauliMatrices.Basic
|
||||
import LLMLean
|
||||
/-!
|
||||
|
||||
## Interaction of Pauli matrices with self-adjoint matrices
|
||||
|
@ -34,7 +33,7 @@ lemma selfAdjoint_trace_σ0_real (A : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ))
|
|||
have hA00 : A.1 0 0 = (A.1 0 0).re := by
|
||||
exact Eq.symm ((fun {z} => Complex.conj_eq_iff_re.mp) h00)
|
||||
rw [hA00]
|
||||
simp
|
||||
simp only [Fin.isValue, Complex.ofReal_re, add_right_inj]
|
||||
have h11 := congrArg (fun f => f 1 1) hA
|
||||
simp only [Fin.isValue, of_apply, cons_val', cons_val_one, head_cons, empty_val',
|
||||
cons_val_fin_one, head_fin_const] at h11
|
||||
|
@ -51,12 +50,11 @@ lemma selfAdjoint_trace_σ1_real (A : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ))
|
|||
rw [eta_fin_two (A.1)ᴴ] at hA
|
||||
simp at hA
|
||||
have h01 := congrArg (fun f => f 0 1) hA
|
||||
have h10 := congrArg (fun f => f 1 0) hA
|
||||
simp at h01 h10
|
||||
simp at h01
|
||||
rw [← h01]
|
||||
simp
|
||||
simp only [Fin.isValue, Complex.conj_re]
|
||||
rw [Complex.add_conj]
|
||||
simp
|
||||
simp only [Fin.isValue, Complex.ofReal_mul, Complex.ofReal_ofNat]
|
||||
ring
|
||||
|
||||
lemma selfAdjoint_trace_σ2_real (A : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)) :
|
||||
|
@ -67,14 +65,17 @@ lemma selfAdjoint_trace_σ2_real (A : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ))
|
|||
Equiv.symm_apply_apply, trace_fin_two_of, Complex.add_re, Complex.ofReal_add]
|
||||
have hA : IsSelfAdjoint A.1 := A.2
|
||||
rw [isSelfAdjoint_iff, star_eq_conjTranspose, eta_fin_two (A.1)ᴴ] at hA
|
||||
simp at hA
|
||||
have h01 := congrArg (fun f => f 0 1) hA
|
||||
simp only [Fin.isValue, conjTranspose_apply, RCLike.star_def] at hA
|
||||
have h10 := congrArg (fun f => f 1 0) hA
|
||||
simp at h01 h10
|
||||
simp only [Fin.isValue, of_apply, cons_val', cons_val_zero, empty_val', cons_val_fin_one,
|
||||
cons_val_one, head_fin_const] at h10
|
||||
rw [← h10]
|
||||
simp
|
||||
simp only [Fin.isValue, neg_smul, smul_cons, smul_eq_mul, smul_empty, neg_cons, neg_empty,
|
||||
trace_fin_two_of, Complex.add_re, Complex.neg_re, Complex.mul_re, Complex.I_re, Complex.conj_re,
|
||||
zero_mul, Complex.I_im, Complex.conj_im, mul_neg, one_mul, sub_neg_eq_add, zero_add, zero_sub,
|
||||
Complex.ofReal_add, Complex.ofReal_neg]
|
||||
trans Complex.I * (A.1 0 1 - (starRingEnd ℂ) (A.1 0 1))
|
||||
· rw [Complex.sub_conj ]
|
||||
· rw [Complex.sub_conj]
|
||||
ring_nf
|
||||
simp
|
||||
· ring
|
||||
|
@ -92,11 +93,11 @@ lemma selfAdjoint_trace_σ3_real (A : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ))
|
|||
have h11 := congrArg (fun f => f 1 1) hA
|
||||
have hA00 : A.1 0 0 = (A.1 0 0).re := Eq.symm ((fun {z} => Complex.conj_eq_iff_re.mp) h00)
|
||||
have hA11 : A.1 1 1 = (A.1 1 1).re := Eq.symm ((fun {z} => Complex.conj_eq_iff_re.mp) h11)
|
||||
simp
|
||||
simp only [Fin.isValue, neg_smul, one_smul, neg_cons, neg_empty, trace_fin_two_of, Complex.add_re,
|
||||
Complex.neg_re, Complex.ofReal_add, Complex.ofReal_neg]
|
||||
rw [hA00, hA11]
|
||||
simp
|
||||
|
||||
|
||||
open Complex
|
||||
|
||||
lemma selfAdjoint_ext_complex {A B : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)}
|
||||
|
@ -114,15 +115,15 @@ lemma selfAdjoint_ext_complex {A B : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)}
|
|||
simp [PauliMatrix.σ3] at h3
|
||||
match i, j with
|
||||
| 0, 0 =>
|
||||
linear_combination (norm := ring_nf) (h0 + h3) / 2
|
||||
linear_combination (norm := ring_nf) (h0 + h3) / 2
|
||||
| 0, 1 =>
|
||||
linear_combination (norm := ring_nf) (h1 - I * h2) / 2
|
||||
linear_combination (norm := ring_nf) (h1 - I * h2) / 2
|
||||
simp
|
||||
| 1, 0 =>
|
||||
linear_combination (norm := ring_nf) (h1 + I * h2) / 2
|
||||
linear_combination (norm := ring_nf) (h1 + I * h2) / 2
|
||||
simp
|
||||
| 1, 1 =>
|
||||
linear_combination (norm := ring_nf) (h0 - h3) / 2
|
||||
linear_combination (norm := ring_nf) (h0 - h3) / 2
|
||||
|
||||
lemma selfAdjoint_ext {A B : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)}
|
||||
(h0 : ((Matrix.trace (PauliMatrix.σ0 * A.1))).re = ((Matrix.trace (PauliMatrix.σ0 * B.1))).re)
|
||||
|
@ -143,6 +144,8 @@ lemma selfAdjoint_ext {A B : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)}
|
|||
|
||||
noncomputable section
|
||||
|
||||
/-- An auxillary function which on `i : Fin 1 ⊕ Fin 3` returns the corresponding
|
||||
Pauli-matrix as a self-adjoint matrix. -/
|
||||
def σSA' (i : Fin 1 ⊕ Fin 3) : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ) :=
|
||||
match i with
|
||||
| Sum.inl 0 => ⟨σ0, σ0_selfAdjoint⟩
|
||||
|
@ -150,7 +153,6 @@ def σSA' (i : Fin 1 ⊕ Fin 3) : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ) :=
|
|||
| Sum.inr 1 => ⟨σ2, σ2_selfAdjoint⟩
|
||||
| Sum.inr 2 => ⟨σ3, σ3_selfAdjoint⟩
|
||||
|
||||
|
||||
lemma σSA_linearly_independent : LinearIndependent ℝ σSA' := by
|
||||
apply Fintype.linearIndependent_iff.mpr
|
||||
intro g hg
|
||||
|
@ -189,17 +191,19 @@ lemma σSA_span : ⊤ ≤ Submodule.span ℝ (Set.range σSA') := by
|
|||
· simp [mul_add, σSA']
|
||||
ring
|
||||
|
||||
/-- The basis of `selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)` formed by Pauli matrices. -/
|
||||
def σSA : Basis (Fin 1 ⊕ Fin 3) ℝ (selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)) :=
|
||||
Basis.mk σSA_linearly_independent σSA_span
|
||||
|
||||
/-- An auxillary function which on `i : Fin 1 ⊕ Fin 3` returns the corresponding
|
||||
Pauli-matrix as a self-adjoint matrix with a minus sign for `Sum.inr _`. -/
|
||||
def σSAL' (i : Fin 1 ⊕ Fin 3) : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ) :=
|
||||
match i with
|
||||
| Sum.inl 0 => ⟨σ0, σ0_selfAdjoint⟩
|
||||
| Sum.inr 0 => ⟨-σ1, by rw [neg_mem_iff]; exact σ1_selfAdjoint⟩
|
||||
| Sum.inr 1 => ⟨-σ2, by rw [neg_mem_iff]; exact σ2_selfAdjoint⟩
|
||||
| Sum.inr 1 => ⟨-σ2, by rw [neg_mem_iff]; exact σ2_selfAdjoint⟩
|
||||
| Sum.inr 2 => ⟨-σ3, by rw [neg_mem_iff]; exact σ3_selfAdjoint⟩
|
||||
|
||||
|
||||
lemma σSAL_linearly_independent : LinearIndependent ℝ σSAL' := by
|
||||
apply Fintype.linearIndependent_iff.mpr
|
||||
intro g hg
|
||||
|
@ -238,6 +242,8 @@ lemma σSAL_span : ⊤ ≤ Submodule.span ℝ (Set.range σSAL') := by
|
|||
· simp [mul_add, σSAL']
|
||||
ring
|
||||
|
||||
/-- The basis of `selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)` formed by Pauli matrices
|
||||
where the `1, 2, 3` pauli matrices are negated. -/
|
||||
def σSAL : Basis (Fin 1 ⊕ Fin 3) ℝ (selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)) :=
|
||||
Basis.mk σSAL_linearly_independent σSAL_span
|
||||
|
||||
|
@ -329,6 +335,5 @@ lemma σSA_minkowskiMetric_σSAL (i : Fin 1 ⊕ Fin 3) :
|
|||
ext i j : 2
|
||||
simp_all only [NegMemClass.coe_neg, neg_apply, neg_neg]
|
||||
|
||||
|
||||
end
|
||||
end PauliMatrix
|
||||
|
|
|
@ -144,9 +144,10 @@ lemma toLorentzGroup_eq_stdBasis (M : SL(2, ℂ)) :
|
|||
(repLorentzVector M) := by rfl
|
||||
|
||||
lemma repLorentzVector_apply_eq_mulVec (v : LorentzVector 3) :
|
||||
SL2C.repLorentzVector M v = (SL2C.toLorentzGroup M).1 *ᵥ v := by
|
||||
SL2C.repLorentzVector M v = (SL2C.toLorentzGroup M).1 *ᵥ v := by
|
||||
simp [toLorentzGroup]
|
||||
have hv : v = (Finsupp.linearEquivFunOnFinite ℝ ℝ (Fin 1 ⊕ Fin 3)) (LorentzVector.stdBasis.repr v) := by rfl
|
||||
have hv : v = (Finsupp.linearEquivFunOnFinite ℝ ℝ (Fin 1 ⊕ Fin 3))
|
||||
(LorentzVector.stdBasis.repr v) := by rfl
|
||||
nth_rewrite 2 [hv]
|
||||
change _ = toLorentzGroup M *ᵥ (LorentzVector.stdBasis.repr v)
|
||||
rw [toLorentzGroup_eq_stdBasis, LinearMap.toMatrix_mulVec_repr]
|
||||
|
@ -159,10 +160,10 @@ lemma repSelfAdjointMatrix_basis (i : Fin 1 ⊕ Fin 3) :
|
|||
rw [toLorentzGroup_eq_σSAL]
|
||||
simp only [LinearMap.toMatrix_apply, Finset.univ_unique,
|
||||
Fin.default_eq_zero, Fin.isValue, Finset.sum_singleton]
|
||||
nth_rewrite 1 [← (Basis.sum_repr PauliMatrix.σSAL ((repSelfAdjointMatrix M) (PauliMatrix.σSAL i)))]
|
||||
nth_rewrite 1 [← (Basis.sum_repr PauliMatrix.σSAL
|
||||
((repSelfAdjointMatrix M) (PauliMatrix.σSAL i)))]
|
||||
congr
|
||||
|
||||
|
||||
lemma repSelfAdjointMatrix_σSA (i : Fin 1 ⊕ Fin 3) :
|
||||
SL2C.repSelfAdjointMatrix M (PauliMatrix.σSA i) =
|
||||
∑ j, (toLorentzGroup M⁻¹).1 i j • PauliMatrix.σSA j := by
|
||||
|
@ -179,7 +180,6 @@ lemma repSelfAdjointMatrix_σSA (i : Fin 1 ⊕ Fin 3) :
|
|||
apply congrArg
|
||||
exact Eq.symm (minkowskiMetric.dual_apply_minkowskiMatrix ((toLorentzGroup M).1) i j)
|
||||
|
||||
|
||||
lemma repLorentzVector_stdBasis (i : Fin 1 ⊕ Fin 3) :
|
||||
SL2C.repLorentzVector M (LorentzVector.stdBasis i) =
|
||||
∑ j, (toLorentzGroup M).1 j i • LorentzVector.stdBasis j := by
|
||||
|
|
|
@ -76,7 +76,7 @@ def altRightRightToMatrix : (altRightHanded ⊗ rightHanded).V ≃ₗ[ℂ] Matri
|
|||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
||||
/-- Equivalence of `altLeftHanded ⊗ altRightHanded` to `2 x 2` complex matrices. -/
|
||||
def altLeftAltRightToMatrix : (altLeftHanded ⊗ altRightHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
def altLeftAltRightToMatrix : (altLeftHanded ⊗ altRightHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
(Basis.tensorProduct altLeftBasis altRightBasis).repr ≪≫ₗ
|
||||
Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
@ -465,7 +465,7 @@ lemma altLeftAltRightToMatrix_ρ (v : (altLeftHanded ⊗ altRightHanded).V) (M :
|
|||
Action.instMonoidalCategory_tensorObj_V]
|
||||
ring
|
||||
|
||||
def leftRightToMatrix_ρ (v : (leftHanded ⊗ rightHanded).V) (M : SL(2,ℂ)) :
|
||||
lemma leftRightToMatrix_ρ (v : (leftHanded ⊗ rightHanded).V) (M : SL(2,ℂ)) :
|
||||
leftRightToMatrix (TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M) v) =
|
||||
M.1 * leftRightToMatrix v * (M.1)ᴴ := by
|
||||
nth_rewrite 1 [leftRightToMatrix]
|
||||
|
@ -490,7 +490,7 @@ def leftRightToMatrix_ρ (v : (leftHanded ⊗ rightHanded).V) (M : SL(2,ℂ)) :
|
|||
erw [Finset.sum_product]
|
||||
simp_rw [kroneckerMap_apply, Matrix.mul_apply]
|
||||
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, M.1 i x1 * leftRightToMatrix v x1 x) * (M.1)ᴴ x j
|
||||
= ∑ x : Fin 2, ∑ x1 : Fin 2, (M.1 i x1 * leftRightToMatrix v x1 x) * (M.1)ᴴ x j := by
|
||||
= ∑ x : Fin 2, ∑ x1 : Fin 2, (M.1 i x1 * leftRightToMatrix v x1 x) * (M.1)ᴴ x j := by
|
||||
congr
|
||||
funext x
|
||||
rw [Finset.sum_mul]
|
||||
|
@ -613,7 +613,7 @@ lemma altLeftAltRightToMatrix_ρ_symm_selfAdjoint (v : Matrix (Fin 2) (Fin 2)
|
|||
rw [← @adjugate_transpose]
|
||||
rfl
|
||||
|
||||
lemma leftRightToMatrix_ρ_symm_selfAdjoint (v : Matrix (Fin 2) (Fin 2) ℂ)
|
||||
lemma leftRightToMatrix_ρ_symm_selfAdjoint (v : Matrix (Fin 2) (Fin 2) ℂ)
|
||||
(hv : IsSelfAdjoint v) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M) (leftRightToMatrix.symm v) =
|
||||
leftRightToMatrix.symm
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue