Merge pull request #231 from HEPLean/RealLorentzTensors

feat: Properties of Real Lorentz Tensors
This commit is contained in:
Joseph Tooby-Smith 2024-11-10 07:20:49 +00:00 committed by GitHub
commit df210420ba
No known key found for this signature in database
GPG key ID: B5690EEEBB952194
52 changed files with 2204 additions and 1738 deletions

View file

@ -62,6 +62,47 @@ import HepLean.FlavorPhysics.CKMMatrix.Relations
import HepLean.FlavorPhysics.CKMMatrix.Rows
import HepLean.FlavorPhysics.CKMMatrix.StandardParameterization.Basic
import HepLean.FlavorPhysics.CKMMatrix.StandardParameterization.StandardParameters
import HepLean.Lorentz.Algebra.Basic
import HepLean.Lorentz.Algebra.Basis
import HepLean.Lorentz.ComplexTensor.Basic
import HepLean.Lorentz.ComplexTensor.Basis
import HepLean.Lorentz.ComplexTensor.Bispinors.Basic
import HepLean.Lorentz.ComplexTensor.Lemmas
import HepLean.Lorentz.ComplexTensor.Metrics.Basic
import HepLean.Lorentz.ComplexTensor.Metrics.Basis
import HepLean.Lorentz.ComplexTensor.Metrics.Lemmas
import HepLean.Lorentz.ComplexTensor.PauliMatrices.Basic
import HepLean.Lorentz.ComplexTensor.PauliMatrices.Basis
import HepLean.Lorentz.ComplexTensor.PauliMatrices.CoContractContr
import HepLean.Lorentz.ComplexTensor.Units.Basic
import HepLean.Lorentz.ComplexTensor.Units.Symm
import HepLean.Lorentz.ComplexVector.Basic
import HepLean.Lorentz.ComplexVector.Contraction
import HepLean.Lorentz.ComplexVector.Metric
import HepLean.Lorentz.ComplexVector.Modules
import HepLean.Lorentz.ComplexVector.Two
import HepLean.Lorentz.ComplexVector.Unit
import HepLean.Lorentz.Group.Basic
import HepLean.Lorentz.Group.Boosts
import HepLean.Lorentz.Group.Orthochronous
import HepLean.Lorentz.Group.Proper
import HepLean.Lorentz.Group.Restricted
import HepLean.Lorentz.Group.Rotations
import HepLean.Lorentz.MinkowskiMatrix
import HepLean.Lorentz.PauliMatrices.AsTensor
import HepLean.Lorentz.PauliMatrices.Basic
import HepLean.Lorentz.PauliMatrices.SelfAdjoint
import HepLean.Lorentz.RealVector.Basic
import HepLean.Lorentz.RealVector.Contraction
import HepLean.Lorentz.RealVector.Modules
import HepLean.Lorentz.RealVector.NormOne
import HepLean.Lorentz.SL2C.Basic
import HepLean.Lorentz.Weyl.Basic
import HepLean.Lorentz.Weyl.Contraction
import HepLean.Lorentz.Weyl.Metric
import HepLean.Lorentz.Weyl.Modules
import HepLean.Lorentz.Weyl.Two
import HepLean.Lorentz.Weyl.Unit
import HepLean.Mathematics.Fin
import HepLean.Mathematics.LinearMaps
import HepLean.Mathematics.PiTensorProduct
@ -71,54 +112,12 @@ import HepLean.Meta.Informal
import HepLean.Meta.TransverseTactics
import HepLean.SpaceTime.Basic
import HepLean.SpaceTime.CliffordAlgebra
import HepLean.SpaceTime.LorentzAlgebra.Basic
import HepLean.SpaceTime.LorentzAlgebra.Basis
import HepLean.SpaceTime.LorentzGroup.Basic
import HepLean.SpaceTime.LorentzGroup.Boosts
import HepLean.SpaceTime.LorentzGroup.Orthochronous
import HepLean.SpaceTime.LorentzGroup.Proper
import HepLean.SpaceTime.LorentzGroup.Restricted
import HepLean.SpaceTime.LorentzGroup.Rotations
import HepLean.SpaceTime.LorentzVector.AsSelfAdjointMatrix
import HepLean.SpaceTime.LorentzVector.Basic
import HepLean.SpaceTime.LorentzVector.Complex.Basic
import HepLean.SpaceTime.LorentzVector.Complex.Contraction
import HepLean.SpaceTime.LorentzVector.Complex.Metric
import HepLean.SpaceTime.LorentzVector.Complex.Modules
import HepLean.SpaceTime.LorentzVector.Complex.Two
import HepLean.SpaceTime.LorentzVector.Complex.Unit
import HepLean.SpaceTime.LorentzVector.Covariant
import HepLean.SpaceTime.LorentzVector.LorentzAction
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
import HepLean.SpaceTime.WeylFermion.Metric
import HepLean.SpaceTime.WeylFermion.Modules
import HepLean.SpaceTime.WeylFermion.Two
import HepLean.SpaceTime.WeylFermion.Unit
import HepLean.StandardModel.Basic
import HepLean.StandardModel.HiggsBoson.Basic
import HepLean.StandardModel.HiggsBoson.GaugeAction
import HepLean.StandardModel.HiggsBoson.PointwiseInnerProd
import HepLean.StandardModel.HiggsBoson.Potential
import HepLean.StandardModel.Representations
import HepLean.Tensors.ComplexLorentz.Basic
import HepLean.Tensors.ComplexLorentz.Basis
import HepLean.Tensors.ComplexLorentz.Bispinors.Basic
import HepLean.Tensors.ComplexLorentz.Lemmas
import HepLean.Tensors.ComplexLorentz.Metrics.Basic
import HepLean.Tensors.ComplexLorentz.Metrics.Basis
import HepLean.Tensors.ComplexLorentz.Metrics.Lemmas
import HepLean.Tensors.ComplexLorentz.PauliMatrices.Basic
import HepLean.Tensors.ComplexLorentz.PauliMatrices.Basis
import HepLean.Tensors.ComplexLorentz.PauliMatrices.CoContractContr
import HepLean.Tensors.ComplexLorentz.Units.Basic
import HepLean.Tensors.ComplexLorentz.Units.Symm
import HepLean.Tensors.OverColor.Basic
import HepLean.Tensors.OverColor.Discrete
import HepLean.Tensors.OverColor.Functors

View file

@ -3,8 +3,9 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.MinkowskiMetric
import HepLean.Lorentz.MinkowskiMatrix
import Mathlib.Algebra.Lie.Classical
import HepLean.Lorentz.RealVector.Basic
/-!
# The Lorentz Algebra
@ -17,7 +18,6 @@ We define
-/
namespace SpaceTime
open Matrix
open TensorProduct
@ -83,23 +83,3 @@ lemma space_comps (Λ : lorentzAlgebra) (i j : Fin 3) :
(congrArg (fun M ↦ M (Sum.inr i) (Sum.inr j)) $ mem_iff.mp Λ.2).symm
end lorentzAlgebra
@[simps!]
instance lorentzVectorAsLieRingModule : LieRingModule lorentzAlgebra (LorentzVector 3) where
bracket Λ x := Λ.1.mulVec x
add_lie Λ1 Λ2 x := by
simp [add_mulVec]
lie_add Λ x1 x2 := by
simp only
exact mulVec_add _ _ _
leibniz_lie Λ1 Λ2 x := by
simp [mulVec_add, Bracket.bracket, sub_mulVec]
@[simps!]
instance spaceTimeAsLieModule : LieModule lorentzAlgebra (LorentzVector 3) where
smul_lie r Λ x := smul_mulVec_assoc r Λ.1 x
lie_smul r Λ x := by
simp only [Bracket.bracket]
rw [mulVec_smul]
end SpaceTime

View file

@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzAlgebra.Basic
import HepLean.Lorentz.Algebra.Basic
/-!
# Basis of the Lorentz Algebra

View file

@ -5,14 +5,14 @@ Authors: Joseph Tooby-Smith
-/
import HepLean.Tensors.OverColor.Basic
import HepLean.Tensors.Tree.Dot
import HepLean.SpaceTime.WeylFermion.Contraction
import HepLean.SpaceTime.WeylFermion.Metric
import HepLean.SpaceTime.WeylFermion.Unit
import HepLean.SpaceTime.LorentzVector.Complex.Contraction
import HepLean.SpaceTime.LorentzVector.Complex.Metric
import HepLean.SpaceTime.LorentzVector.Complex.Unit
import HepLean.Lorentz.Weyl.Contraction
import HepLean.Lorentz.Weyl.Metric
import HepLean.Lorentz.Weyl.Unit
import HepLean.Lorentz.ComplexVector.Contraction
import HepLean.Lorentz.ComplexVector.Metric
import HepLean.Lorentz.ComplexVector.Unit
import HepLean.Mathematics.PiTensorProduct
import HepLean.SpaceTime.PauliMatrices.AsTensor
import HepLean.Lorentz.PauliMatrices.AsTensor
/-!
## Complex Lorentz tensors

View file

@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Tensors.Tree.Elab
import HepLean.Tensors.ComplexLorentz.Basic
import HepLean.Lorentz.ComplexTensor.Basic
import Mathlib.LinearAlgebra.TensorProduct.Basis
import HepLean.Tensors.Tree.NodeIdentities.Basic
import HepLean.Tensors.Tree.NodeIdentities.PermProd

View file

@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Tensors.ComplexLorentz.PauliMatrices.Basic
import HepLean.Lorentz.ComplexTensor.PauliMatrices.Basic
import HepLean.Tensors.Tree.NodeIdentities.ProdContr
import HepLean.Tensors.Tree.NodeIdentities.PermContr
import HepLean.Tensors.Tree.NodeIdentities.PermProd

View file

@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Tensors.ComplexLorentz.Basis
import HepLean.Lorentz.ComplexTensor.Basis
import HepLean.Tensors.Tree.NodeIdentities.PermProd
/-!

View file

@ -3,8 +3,8 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Tensors.ComplexLorentz.Metrics.Basic
import HepLean.Tensors.ComplexLorentz.Basis
import HepLean.Lorentz.ComplexTensor.Metrics.Basic
import HepLean.Lorentz.ComplexTensor.Basis
/-!
## Metrics and basis expansions

View file

@ -3,9 +3,9 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Tensors.ComplexLorentz.Metrics.Basis
import HepLean.Tensors.ComplexLorentz.Units.Basic
import HepLean.Tensors.ComplexLorentz.Basis
import HepLean.Lorentz.ComplexTensor.Metrics.Basis
import HepLean.Lorentz.ComplexTensor.Units.Basic
import HepLean.Lorentz.ComplexTensor.Basis
/-!
## Basic lemmas regarding metrics

View file

@ -10,7 +10,7 @@ import HepLean.Tensors.Tree.NodeIdentities.ContrContr
import HepLean.Tensors.Tree.NodeIdentities.ContrSwap
import HepLean.Tensors.Tree.NodeIdentities.PermContr
import HepLean.Tensors.Tree.NodeIdentities.Congr
import HepLean.Tensors.ComplexLorentz.Metrics.Lemmas
import HepLean.Lorentz.ComplexTensor.Metrics.Lemmas
/-!
## Pauli matrices as complex Lorentz tensors

View file

@ -3,8 +3,8 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Tensors.ComplexLorentz.PauliMatrices.Basic
import HepLean.Tensors.ComplexLorentz.Basis
import HepLean.Lorentz.ComplexTensor.PauliMatrices.Basic
import HepLean.Lorentz.ComplexTensor.Basis
/-!
## Pauli matrices and the basis of complex Lorentz tensors

View file

@ -3,8 +3,8 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Tensors.ComplexLorentz.PauliMatrices.Basis
import HepLean.Tensors.ComplexLorentz.Lemmas
import HepLean.Lorentz.ComplexTensor.PauliMatrices.Basis
import HepLean.Lorentz.ComplexTensor.Lemmas
/-!
## Contractiong of indices of Pauli matrix.

View file

@ -3,9 +3,9 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Tensors.ComplexLorentz.Metrics.Basis
import HepLean.Tensors.ComplexLorentz.Units.Basic
import HepLean.Tensors.ComplexLorentz.Basis
import HepLean.Lorentz.ComplexTensor.Metrics.Basis
import HepLean.Lorentz.ComplexTensor.Units.Basic
import HepLean.Lorentz.ComplexTensor.Basis
/-!
## Symmetry lemmas relating to units

View file

@ -5,11 +5,11 @@ Authors: Joseph Tooby-Smith
-/
import Mathlib.Data.Complex.Exponential
import Mathlib.Analysis.InnerProductSpace.PiL2
import HepLean.SpaceTime.SL2C.Basic
import HepLean.SpaceTime.LorentzVector.Complex.Modules
import HepLean.Lorentz.SL2C.Basic
import HepLean.Lorentz.ComplexVector.Modules
import HepLean.Meta.Informal
import Mathlib.RepresentationTheory.Rep
import HepLean.SpaceTime.PauliMatrices.SelfAdjoint
import HepLean.Lorentz.PauliMatrices.SelfAdjoint
/-!
# Complex Lorentz vectors
@ -24,7 +24,6 @@ open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
namespace Lorentz
@ -37,8 +36,8 @@ def complexContr : Rep SL(2, ) := Rep.of ContrModule.SL2CRep
def complexCo : Rep SL(2, ) := Rep.of CoModule.SL2CRep
/-- The standard basis of complex contravariant Lorentz vectors. -/
def complexContrBasis : Basis (Fin 1 ⊕ Fin 3) complexContr := Basis.ofEquivFun
(Equiv.linearEquiv ContrModule.toFin13Fun)
def complexContrBasis : Basis (Fin 1 ⊕ Fin 3) complexContr :=
Basis.ofEquivFun ContrModule.toFin13Equiv
@[simp]
lemma complexContrBasis_toFin13 (i :Fin 1 ⊕ Fin 3) :
@ -65,8 +64,8 @@ def complexContrBasisFin4 : Basis (Fin 4) complexContr :=
Basis.reindex complexContrBasis finSumFinEquiv
/-- The standard basis of complex covariant Lorentz vectors. -/
def complexCoBasis : Basis (Fin 1 ⊕ Fin 3) complexCo := Basis.ofEquivFun
(Equiv.linearEquiv CoModule.toFin13Fun)
def complexCoBasis : Basis (Fin 1 ⊕ Fin 3) complexCo :=
Basis.ofEquivFun CoModule.toFin13Equiv
@[simp]
lemma complexCoBasis_toFin13 (i :Fin 1 ⊕ Fin 3) : (complexCoBasis i).toFin13 = Pi.single i 1 := by
@ -94,57 +93,62 @@ def complexCoBasisFin4 : Basis (Fin 4) complexCo :=
/-- The semilinear map including real Lorentz vectors into complex contravariant
lorentz vectors. -/
def inclCongrRealLorentz : LorentzVector 3 →ₛₗ[Complex.ofRealHom] complexContr where
toFun v := {val := ofReal ∘ v}
def inclCongrRealLorentz : ContrMod 3 →ₛₗ[Complex.ofRealHom] complexContr where
toFun v := {val := ofReal ∘ v.toFin1d}
map_add' x y := by
apply Lorentz.ContrModule.ext
rw [Lorentz.ContrModule.val_add]
funext i
simp only [Function.comp_apply, ofRealHom_eq_coe, Pi.add_apply]
change ofReal (x i + y i) = _
simp only [Function.comp_apply, ofRealHom_eq_coe, Pi.add_apply, map_add]
simp only [ofRealHom_eq_coe, ofReal_add]
map_smul' c x := by
apply Lorentz.ContrModule.ext
rw [Lorentz.ContrModule.val_smul]
funext i
simp only [Function.comp_apply, ofRealHom_eq_coe, Pi.smul_apply]
change ofReal (c • x i) = _
simp only [Function.comp_apply, ofRealHom_eq_coe, Pi.smul_apply, _root_.map_smul]
simp only [smul_eq_mul, ofRealHom_eq_coe, ofReal_mul]
lemma inclCongrRealLorentz_val (v : LorentzVector 3) :
(inclCongrRealLorentz v).val = ofRealHom ∘ v := rfl
lemma inclCongrRealLorentz_val (v : ContrMod 3) :
(inclCongrRealLorentz v).val = ofRealHom ∘ v.toFin1d := rfl
lemma complexContrBasis_of_real (i : Fin 1 ⊕ Fin 3) :
(complexContrBasis i) = inclCongrRealLorentz (LorentzVector.stdBasis i) := by
(complexContrBasis i) = inclCongrRealLorentz (ContrMod.stdBasis i) := by
apply Lorentz.ContrModule.ext
simp only [complexContrBasis, Basis.coe_ofEquivFun, inclCongrRealLorentz, LorentzVector.stdBasis,
simp only [complexContrBasis, Basis.coe_ofEquivFun, inclCongrRealLorentz,
LinearMap.coe_mk, AddHom.coe_mk]
ext j
simp only [Function.comp_apply, ofRealHom_eq_coe]
erw [Pi.basisFun_apply]
simp only [Function.comp_apply]
change (Pi.single i 1) j = _
exact Eq.symm (Pi.apply_single (fun _ => ofRealHom) (congrFun rfl) i 1 j)
by_cases h : i = j
· subst h
rw [ContrMod.toFin1d, ContrMod.stdBasis_toFin1dEquiv_apply_same]
simp
· rw [ContrMod.toFin1d, ContrMod.stdBasis_toFin1dEquiv_apply_ne h]
simp [h]
lemma inclCongrRealLorentz_ρ (M : SL(2, )) (v : LorentzVector 3) :
lemma inclCongrRealLorentz_ρ (M : SL(2, )) (v : ContrMod 3) :
(complexContr.ρ M) (inclCongrRealLorentz v) =
inclCongrRealLorentz (SL2C.repLorentzVector M v) := by
inclCongrRealLorentz ((Contr 3).ρ (SL2C.toLorentzGroup M) v) := by
apply Lorentz.ContrModule.ext
rw [complexContrBasis_ρ_val, inclCongrRealLorentz_val, inclCongrRealLorentz_val]
rw [LorentzGroup.toComplex_mulVec_ofReal]
apply congrArg
simp only [SL2C.toLorentzGroup_apply_coe]
rw [SL2C.repLorentzVector_apply_eq_mulVec]
rfl
change _ = ContrMod.toFin1d ((SL2C.toLorentzGroup M) *ᵥ v)
simp only [SL2C.toLorentzGroup_apply_coe, ContrMod.mulVec_toFin1d]
/-! TODO: Rename. -/
lemma SL2CRep_ρ_basis (M : SL(2, )) (i : Fin 1 ⊕ Fin 3) :
(complexContr.ρ M) (complexContrBasis i) =
∑ j, (SL2C.toLorentzGroup M).1 j i •
complexContrBasis j := by
rw [complexContrBasis_of_real, inclCongrRealLorentz_ρ, SL2C.repLorentzVector_stdBasis, map_sum]
rw [complexContrBasis_of_real, inclCongrRealLorentz_ρ]
rw [Contr.ρ_stdBasis, map_sum]
apply congrArg
funext j
simp only [LinearMap.map_smulₛₗ, ofRealHom_eq_coe, coe_smul]
rw [complexContrBasis_of_real]
/-! TODO: Include relation to real Lorentz vectors. -/
end Lorentz
end

View file

@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.Complex.Basic
import HepLean.Lorentz.ComplexVector.Basic
/-!
# Contraction of Lorentz vectors
@ -16,7 +16,6 @@ open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
open CategoryTheory.MonoidalCategory
namespace Lorentz
@ -160,7 +159,7 @@ lemma contrCoContraction_tmul_symm (φ : complexContr) (ψ : complexCo) :
lemma coContrContraction_tmul_symm (φ : complexCo) (ψ : complexContr) :
coContrContraction.hom (φ ⊗ₜ ψ) = contrCoContraction.hom (ψ ⊗ₜ φ) := by
rw [contrCoContraction_hom_tmul, coContrContraction_hom_tmul, dotProduct_comm]
rw [contrCoContraction_tmul_symm]
end Lorentz
end

View file

@ -3,10 +3,10 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.Complex.Two
import HepLean.SpaceTime.MinkowskiMetric
import HepLean.SpaceTime.LorentzVector.Complex.Contraction
import HepLean.SpaceTime.LorentzVector.Complex.Unit
import HepLean.Lorentz.ComplexVector.Two
import HepLean.Lorentz.MinkowskiMatrix
import HepLean.Lorentz.ComplexVector.Contraction
import HepLean.Lorentz.ComplexVector.Unit
/-!
# Metric for complex Lorentz vectors
@ -18,7 +18,6 @@ open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
open CategoryTheory.MonoidalCategory
namespace Lorentz

View file

@ -4,14 +4,16 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Meta.Informal
import HepLean.SpaceTime.SL2C.Basic
import HepLean.Lorentz.SL2C.Basic
import Mathlib.RepresentationTheory.Rep
import Mathlib.Logic.Equiv.TransferInstance
/-!
## Modules associated with complex Lorentz vectors
We define these modules to prevent casting between different types of Lorentz vectors.
We define the modules underlying complex Lorentz vectors.
These definitions are preludes to the definitions of
`Lorentz.complexContr` and `Lorentz.complexCo`.
-/
@ -63,13 +65,8 @@ lemma val_smul (r : ) (ψ : ContrModule) : (r • ψ).val = r • ψ.val :
/-- The linear equivalence between `ContrModule` and `(Fin 1 ⊕ Fin 3 → )`. -/
@[simps!]
def toFin13Equiv : ContrModule ≃ₗ[] (Fin 1 ⊕ Fin 3 → ) where
toFun := toFin13Fun
map_add' := fun _ _ => rfl
map_smul' := fun _ _ => rfl
invFun := toFin13Fun.symm
left_inv := fun _ => rfl
right_inv := fun _ => rfl
def toFin13Equiv : ContrModule ≃ₗ[] (Fin 1 ⊕ Fin 3 → ) :=
Equiv.linearEquiv toFin13Fun
/-- The underlying element of `Fin 1 ⊕ Fin 3 → ` of a element in `ContrModule` defined
through the linear equivalence `toFin13Equiv`. -/
@ -95,7 +92,7 @@ def lorentzGroupRep : Representation (LorentzGroup 3) ContrModule where
/-- The representation of the SL(2, ) on `ContrModule` induced by the representation of the
Lorentz group. -/
def SL2CRep : Representation SL(2, ) ContrModule :=
MonoidHom.comp lorentzGroupRep SpaceTime.SL2C.toLorentzGroup
MonoidHom.comp lorentzGroupRep Lorentz.SL2C.toLorentzGroup
end ContrModule
@ -127,13 +124,8 @@ instance : Module CoModule := Equiv.module toFin13Fun
/-- The linear equivalence between `CoModule` and `(Fin 1 ⊕ Fin 3 → )`. -/
@[simps!]
def toFin13Equiv : CoModule ≃ₗ[] (Fin 1 ⊕ Fin 3 → ) where
toFun := toFin13Fun
map_add' := fun _ _ => rfl
map_smul' := fun _ _ => rfl
invFun := toFin13Fun.symm
left_inv := fun _ => rfl
right_inv := fun _ => rfl
def toFin13Equiv : CoModule ≃ₗ[] (Fin 1 ⊕ Fin 3 → ) :=
Equiv.linearEquiv toFin13Fun
/-- The underlying element of `Fin 1 ⊕ Fin 3 → ` of a element in `CoModule` defined
through the linear equivalence `toFin13Equiv`. -/
@ -164,7 +156,7 @@ def lorentzGroupRep : Representation (LorentzGroup 3) CoModule where
/-- The representation of the SL(2, ) on `ContrModule` induced by the representation of the
Lorentz group. -/
def SL2CRep : Representation SL(2, ) CoModule :=
MonoidHom.comp lorentzGroupRep SpaceTime.SL2C.toLorentzGroup
MonoidHom.comp lorentzGroupRep Lorentz.SL2C.toLorentzGroup
end CoModule

View file

@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.Complex.Basic
import HepLean.Lorentz.ComplexVector.Basic
import Mathlib.LinearAlgebra.TensorProduct.Matrix
/-!
@ -16,7 +16,6 @@ open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
open CategoryTheory.MonoidalCategory
namespace Lorentz

View file

@ -3,8 +3,8 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.Complex.Two
import HepLean.SpaceTime.LorentzVector.Complex.Contraction
import HepLean.Lorentz.ComplexVector.Two
import HepLean.Lorentz.ComplexVector.Contraction
/-!
# Unit for complex Lorentz vectors
@ -16,7 +16,6 @@ open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
open CategoryTheory.MonoidalCategory
namespace Lorentz

View file

@ -3,8 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.MinkowskiMetric
import HepLean.SpaceTime.LorentzVector.NormOne
import HepLean.Lorentz.MinkowskiMatrix
/-!
# The Lorentz Group
@ -32,18 +31,17 @@ These matrices form the Lorentz group, which we will define in the next section
-/
variable {d : }
open minkowskiMetric in
/-- The Lorentz group is the subset of matrices which preserve the minkowski metric. -/
open minkowskiMatrix in
/-- The Lorentz group is the subset of matrices for which
`Λ * dual Λ = 1`. -/
def LorentzGroup (d : ) : Set (Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :=
{Λ : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) |
∀ (x y : LorentzVector d), ⟪Λ *ᵥ x, Λ *ᵥ y⟫ₘ = ⟪x, y⟫ₘ}
{Λ : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) | Λ * dual Λ = 1}
namespace LorentzGroup
/-- Notation for the Lorentz group. -/
scoped[LorentzGroup] notation (name := lorentzGroup_notation) "𝓛" => LorentzGroup
open minkowskiMetric
open minkowskiMatrix
variable {Λ Λ' : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) }
/-!
@ -52,32 +50,11 @@ variable {Λ Λ' : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) }
-/
lemma mem_iff_norm : Λ ∈ LorentzGroup d ↔
∀ (x : LorentzVector d), ⟪Λ *ᵥ x, Λ *ᵥ x⟫ₘ = ⟪x, x⟫ₘ := by
refine Iff.intro (fun h x => h x x) (fun h x y => ?_)
have hp := h (x + y)
have hn := h (x - y)
rw [mulVec_add] at hp
rw [mulVec_sub] at hn
simp only [map_add, LinearMap.add_apply, map_sub, LinearMap.sub_apply] at hp hn
rw [symm (Λ *ᵥ y) (Λ *ᵥ x), symm y x] at hp hn
linear_combination hp / 4 + -1 * hn / 4
lemma mem_iff_on_right : Λ ∈ LorentzGroup d ↔
∀ (x y : LorentzVector d), ⟪x, (dual Λ * Λ) *ᵥ y⟫ₘ = ⟪x, y⟫ₘ := by
refine Iff.intro (fun h x y ↦ ?_) (fun h x y ↦ ?_)
· have h1 := h x y
rw [← dual_mulVec_right, mulVec_mulVec] at h1
exact h1
· rw [← dual_mulVec_right, mulVec_mulVec]
exact h x y
lemma mem_iff_self_mul_dual : Λ ∈ LorentzGroup d ↔ Λ * dual Λ = 1 := by
rfl
lemma mem_iff_dual_mul_self : Λ ∈ LorentzGroup d ↔ dual Λ * Λ = 1 := by
rw [mem_iff_on_right, matrix_eq_id_iff]
exact forall_comm
lemma mem_iff_self_mul_dual : Λ ∈ LorentzGroup d ↔ Λ * dual Λ = 1 := by
rw [mem_iff_dual_mul_self]
rw [mem_iff_self_mul_dual]
exact mul_eq_one_comm
lemma mem_iff_transpose : Λ ∈ LorentzGroup d ↔ Λᵀ ∈ LorentzGroup d := by
@ -123,7 +100,7 @@ instance lorentzGroupIsGroup : Group (LorentzGroup d) where
one := ⟨1, LorentzGroup.one_mem⟩
one_mul A := Subtype.eq (Matrix.one_mul A.1)
mul_one A := Subtype.eq (Matrix.mul_one A.1)
inv A := ⟨minkowskiMetric.dual A.1, LorentzGroup.dual_mem A.2⟩
inv A := ⟨minkowskiMatrix.dual A.1, LorentzGroup.dual_mem A.2⟩
inv_mul_cancel A := Subtype.eq (LorentzGroup.mem_iff_dual_mul_self.mp A.2)
/-- `LorentzGroup` has the subtype topology. -/
@ -131,13 +108,21 @@ instance : TopologicalSpace (LorentzGroup d) := instTopologicalSpaceSubtype
namespace LorentzGroup
open minkowskiMetric
open minkowskiMatrix
variable {Λ Λ' : LorentzGroup d}
lemma coe_inv : (Λ⁻¹).1 = Λ.1⁻¹:= (inv_eq_left_inv (mem_iff_dual_mul_self.mp Λ.2)).symm
@[simp]
instance (M : LorentzGroup d) : Invertible M.1 where
invOf := M⁻¹
invOf_mul_self := by
rw [← coe_inv]
exact (mem_iff_dual_mul_self.mp M.2)
mul_invOf_self := by
rw [← coe_inv]
exact (mem_iff_self_mul_dual.mp M.2)
lemma subtype_inv_mul : (Subtype.val Λ)⁻¹ * (Subtype.val Λ) = 1 := by
trans Subtype.val (Λ⁻¹ * Λ)
· rw [← coe_inv]
@ -145,7 +130,6 @@ lemma subtype_inv_mul : (Subtype.val Λ)⁻¹ * (Subtype.val Λ) = 1 := by
· rw [inv_mul_cancel Λ]
rfl
@[simp]
lemma subtype_mul_inv : (Subtype.val Λ) * (Subtype.val Λ)⁻¹ = 1 := by
trans Subtype.val (Λ * Λ⁻¹)
· rw [← coe_inv]
@ -184,6 +168,32 @@ lemma transpose_one : @transpose d 1 = 1 := Subtype.eq Matrix.transpose_one
lemma transpose_mul : transpose (Λ * Λ') = transpose Λ' * transpose Λ :=
Subtype.eq (Matrix.transpose_mul Λ.1 Λ'.1)
lemma transpose_val : (transpose Λ).1 = Λ.1ᵀ := rfl
lemma transpose_inv : (transpose Λ)⁻¹ = transpose Λ⁻¹ := by
refine Subtype.eq ?_
rw [transpose_val, coe_inv, transpose_val, coe_inv, Matrix.transpose_nonsing_inv]
lemma comm_minkowskiMatrix : Λ.1 * minkowskiMatrix = minkowskiMatrix * (transpose Λ⁻¹).1 := by
conv_rhs => rw [← @mul_minkowskiMatrix_mul_transpose d Λ]
rw [← transpose_inv, coe_inv, transpose_val]
rw [mul_assoc]
have h1 : ((Λ.1)ᵀ * (Λ.1)ᵀ⁻¹) = 1 := by
rw [← transpose_val]
simp only [subtype_mul_inv]
rw [h1]
simp
lemma minkowskiMatrix_comm : minkowskiMatrix * Λ.1 = (transpose Λ⁻¹).1 * minkowskiMatrix := by
conv_rhs => rw [← @transpose_mul_minkowskiMatrix_mul_self d Λ]
rw [← transpose_inv, coe_inv, transpose_val]
rw [← mul_assoc, ← mul_assoc]
have h1 : ((Λ.1)ᵀ⁻¹ * (Λ.1)ᵀ) = 1 := by
rw [← transpose_val]
simp only [subtype_inv_mul]
rw [h1]
simp
/-!
## Lorentz group as a topological group
@ -215,7 +225,7 @@ def toProd : LorentzGroup d →* (Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d)
(Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) )ᵐᵒᵖ :=
MonoidHom.comp (Units.embedProduct _) toGL
lemma toProd_eq_transpose_η : toProd Λ = (Λ.1, MulOpposite.op $ minkowskiMetric.dual Λ.1) := rfl
lemma toProd_eq_transpose_η : toProd Λ = (Λ.1, MulOpposite.op $ minkowskiMatrix.dual Λ.1) := rfl
lemma toProd_injective : Function.Injective (@toProd d) := by
intro A B h
@ -249,43 +259,6 @@ lemma toGL_embedding : IsEmbedding (@toGL d).toFun where
instance : TopologicalGroup (LorentzGroup d) :=
IsInducing.topologicalGroup toGL toGL_embedding.toIsInducing
section
open LorentzVector
/-!
# To a norm one Lorentz vector
-/
/-- The first column of a Lorentz matrix as a `NormOneLorentzVector`. -/
@[simps!]
def toNormOneLorentzVector (Λ : LorentzGroup d) : NormOneLorentzVector d :=
⟨Λ.1 *ᵥ timeVec, by rw [NormOneLorentzVector.mem_iff, Λ.2, minkowskiMetric.on_timeVec]⟩
/-!
# The time like element
-/
/-- The time like element of a Lorentz matrix. -/
@[simp]
def timeComp (Λ : LorentzGroup d) : := Λ.1 (Sum.inl 0) (Sum.inl 0)
lemma timeComp_eq_toNormOneLorentzVector : timeComp Λ = (toNormOneLorentzVector Λ).1.time := by
simp only [time, toNormOneLorentzVector, timeVec, Fin.isValue, timeComp]
erw [Pi.basisFun_apply, Matrix.mulVec_single_one]
rfl
lemma timeComp_mul (Λ Λ' : LorentzGroup d) : timeComp (Λ * Λ') =
⟪toNormOneLorentzVector (transpose Λ), (toNormOneLorentzVector Λ').1.spaceReflection⟫ₘ := by
simp only [timeComp, Fin.isValue, lorentzGroupIsGroup_mul_coe, mul_apply, Fintype.sum_sum_type,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton, toNormOneLorentzVector,
transpose, timeVec, right_spaceReflection, time, space, PiLp.inner_apply, Function.comp_apply,
RCLike.inner_apply, conj_trivial]
erw [Pi.basisFun_apply, Matrix.mulVec_single_one]
simp
/-!
## To Complex matrices
@ -351,6 +324,46 @@ lemma toComplex_mulVec_ofReal (v : Fin 1 ⊕ Fin d → ) (Λ : LorentzGroup d
funext i
rw [← RingHom.map_mulVec]
rfl
/-!
/-!
# To a norm one Lorentz vector
-/
/-- The first column of a Lorentz matrix as a `NormOneLorentzVector`. -/
@[simps!]
def toNormOneLorentzVector (Λ : LorentzGroup d) : NormOneLorentzVector d :=
⟨Λ.1 *ᵥ timeVec, by rw [NormOneLorentzVector.mem_iff, Λ.2, minkowskiMetric.on_timeVec]⟩
/-!
# The time like element
-/
/-- The time like element of a Lorentz matrix. -/
@[simp]
def timeComp (Λ : LorentzGroup d) : := Λ.1 (Sum.inl 0) (Sum.inl 0)
lemma timeComp_eq_toNormOneLorentzVector : timeComp Λ = (toNormOneLorentzVector Λ).1.time := by
simp only [time, toNormOneLorentzVector, timeVec, Fin.isValue, timeComp]
erw [Pi.basisFun_apply, Matrix.mulVec_single_one]
rfl
lemma timeComp_mul (Λ Λ' : LorentzGroup d) : timeComp (Λ * Λ') =
⟪toNormOneLorentzVector (transpose Λ), (toNormOneLorentzVector Λ').1.spaceReflection⟫ₘ := by
simp only [timeComp, Fin.isValue, lorentzGroupIsGroup_mul_coe, mul_apply, Fintype.sum_sum_type,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton, toNormOneLorentzVector,
transpose, timeVec, right_spaceReflection, time, space, PiLp.inner_apply, Function.comp_apply,
RCLike.inner_apply, conj_trivial]
erw [Pi.basisFun_apply, Matrix.mulVec_single_one]
simp
-/
/-- The parity transformation. -/
def parity : LorentzGroup d := ⟨minkowskiMatrix, by
rw [mem_iff_dual_mul_self]
simp only [dual_eta, minkowskiMatrix.sq]⟩
end
end LorentzGroup

View file

@ -0,0 +1,236 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Lorentz.Group.Proper
import Mathlib.Topology.Constructions
import HepLean.Lorentz.RealVector.NormOne
/-!
# Boosts
This file defines those Lorentz which are boosts.
We first define generalised boosts, which are restricted lorentz transformations taking
a four velocity `u` to a four velocity `v`.
A boost is the special case of a generalised boost when `u = stdBasis 0`.
## References
- The main argument follows: Guillem Cobos, The Lorentz Group, 2015:
https://diposit.ub.edu/dspace/bitstream/2445/68763/2/memoria.pdf
-/
/-! TODO: Show that generalised boosts are in the restricted Lorentz group. -/
noncomputable section
namespace LorentzGroup
open Lorentz.Contr.NormOne
open Lorentz
open TensorProduct
variable {d : }
/-- An auxillary linear map used in the definition of a generalised boost. -/
def genBoostAux₁ (u v : FuturePointing d) : ContrMod d →ₗ[] ContrMod d where
toFun x := (2 * ⟪x, u.val.val⟫ₘ) • v.1.1
map_add' x y := by
simp [map_add, LinearMap.add_apply, tmul_add, add_tmul, mul_add, add_smul]
map_smul' c x := by
simp only [Action.instMonoidalCategory_tensorObj_V,
Action.instMonoidalCategory_tensorUnit_V, CategoryTheory.Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
smul_tmul, tmul_smul, map_smul, smul_eq_mul, RingHom.id_apply]
rw [← mul_assoc, mul_comm 2 c, mul_assoc, mul_smul]
/-- An auxillary linear map used in the definition of a genearlised boost. -/
def genBoostAux₂ (u v : FuturePointing d) : ContrMod d →ₗ[] ContrMod d where
toFun x := - (⟪x, u.1.1 + v.1.1⟫ₘ / (1 + ⟪u.1.1, v.1.1⟫ₘ)) • (u.1.1 + v.1.1)
map_add' x y := by
simp only
rw [← add_smul]
apply congrFun (congrArg _ _)
field_simp [add_tmul]
apply congrFun (congrArg _ _)
ring
map_smul' c x := by
simp only [smul_tmul, tmul_smul]
rw [map_smul]
conv =>
lhs; lhs; rhs; lhs
change (c * (contrContrContractField (x ⊗ₜ[] (u.val.val + v.val.val))))
rw [mul_div_assoc, neg_mul_eq_mul_neg, smul_smul]
rfl
lemma genBoostAux₂_self (u : FuturePointing d) : genBoostAux₂ u u = - genBoostAux₁ u u := by
ext1 x
simp only [genBoostAux₂, LinearMap.coe_mk, AddHom.coe_mk, genBoostAux₁, LinearMap.neg_apply]
rw [neg_smul]
apply congrArg
conv => lhs; rhs; rw [← (two_smul u.val.val)]
rw [smul_smul]
congr 1
rw [u.1.2]
conv => lhs; lhs; rhs; rhs; change 1
rw [show 1 + (1 : ) = (2 : ) by ring]
simp only [isUnit_iff_ne_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
IsUnit.div_mul_cancel]
rw [← (two_smul u.val.val)]
simp only [tmul_smul, map_smul, smul_eq_mul]
/-- An generalised boost. This is a Lorentz transformation which takes the four velocity `u`
to `v`. -/
def genBoost (u v : FuturePointing d) : ContrMod d →ₗ[] ContrMod d :=
LinearMap.id + genBoostAux₁ u v + genBoostAux₂ u v
lemma genBoost_mul_one_plus_contr (u v : FuturePointing d) (x : Contr d) :
(1 + ⟪u.1.1, v.1.1⟫ₘ) • genBoost u v x =
(1 + ⟪u.1.1, v.1.1⟫ₘ) • x +
(2 * ⟪x, u.val.val⟫ₘ * (1 + ⟪u.1.1, v.1.1⟫ₘ)) • v.1.1
- (⟪x, u.1.1 + v.1.1⟫ₘ) • (u.1.1 + v.1.1) := by
simp only [genBoost, LinearMap.add_apply, LinearMap.id_apply, id_eq]
rw [smul_add, smul_add]
trans (1 + ⟪u.1.1, v.1.1⟫ₘ) • x +
(2 * ⟪x, u.val.val⟫ₘ * (1 + ⟪u.1.1, v.1.1⟫ₘ)) • v.1.1
+ (- ⟪x, u.1.1 + v.1.1⟫ₘ) • (u.1.1 + v.1.1)
· congr 1
· congr 1
rw [genBoostAux₁]
simp only [LinearMap.coe_mk, AddHom.coe_mk]
rw [smul_smul]
congr 1
ring
· rw [genBoostAux₂]
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, neg_smul, LinearMap.coe_mk,
AddHom.coe_mk, smul_neg]
rw [smul_smul]
congr
have h1 := FuturePointing.one_add_metric_non_zero u v
field_simp
· rw [neg_smul]
rfl
namespace genBoost
/--
This lemma states that for a given four-velocity `u`, the general boost
transformation `genBoost u u` is equal to the identity linear map `LinearMap.id`.
-/
lemma self (u : FuturePointing d) : genBoost u u = LinearMap.id := by
ext x
simp only [genBoost, LinearMap.add_apply, LinearMap.id_coe, id_eq]
rw [genBoostAux₂_self]
simp only [LinearMap.neg_apply, add_neg_cancel_right]
/-- A generalised boost as a matrix. -/
def toMatrix (u v : FuturePointing d) : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) :=
LinearMap.toMatrix ContrMod.stdBasis ContrMod.stdBasis (genBoost u v)
lemma toMatrix_mulVec (u v : FuturePointing d) (x : Contr d) :
(toMatrix u v) *ᵥ x = genBoost u v x :=
ContrMod.ext (LinearMap.toMatrix_mulVec_repr ContrMod.stdBasis ContrMod.stdBasis (genBoost u v) x)
open minkowskiMatrix in
@[simp]
lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) :
(toMatrix u v) μ ν = η μ μ * (⟪ContrMod.stdBasis μ, ContrMod.stdBasis ν⟫ₘ + 2 *
⟪ContrMod.stdBasis ν, u.val.val⟫ₘ * ⟪ContrMod.stdBasis μ, v.val.val⟫ₘ
- ⟪ContrMod.stdBasis μ, u.val.val + v.val.val⟫ₘ *
⟪ContrMod.stdBasis ν, u.val.val + v.val.val⟫ₘ /
(1 + ⟪u.val.val, v.val.val⟫ₘ)) := by
rw [contrContrContractField.matrix_apply_stdBasis (Λ := toMatrix u v) μ ν, toMatrix_mulVec]
simp only [genBoost, genBoostAux₁, genBoostAux₂, smul_add, neg_smul, LinearMap.add_apply,
LinearMap.id_apply, LinearMap.coe_mk, AddHom.coe_mk, contrContrContractField.basis_left,
map_add, map_smul, map_neg, mul_eq_mul_left_iff]
ring_nf
simp only [Pi.add_apply, Action.instMonoidalCategory_tensorObj_V,
Action.instMonoidalCategory_tensorUnit_V, CategoryTheory.Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
Pi.smul_apply, smul_eq_mul, Pi.sub_apply, Pi.neg_apply]
apply Or.inl
ring
open minkowskiMatrix in
lemma toMatrix_continuous (u : FuturePointing d) : Continuous (toMatrix u) := by
refine continuous_matrix ?_
intro i j
simp only [toMatrix_apply]
refine (continuous_mul_left (η i i)).comp' ?_
refine Continuous.sub ?_ ?_
· refine .comp' (continuous_add_left _) ?_
refine .comp' (continuous_mul_left _) ?_
exact FuturePointing.metric_continuous _
refine .mul ?_ ?_
· refine .mul ?_ ?_
· simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, tmul_add, map_add]
refine .comp' ?_ ?_
· exact continuous_add_left _
· exact FuturePointing.metric_continuous _
· simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, tmul_add, map_add]
refine .comp' ?_ ?_
· exact continuous_add_left _
· exact FuturePointing.metric_continuous _
· refine .inv₀ ?_ ?_
· refine .comp' (continuous_add_left 1) (FuturePointing.metric_continuous _)
exact fun x => FuturePointing.one_add_metric_non_zero u x
lemma toMatrix_in_lorentzGroup (u v : FuturePointing d) : (toMatrix u v) ∈ LorentzGroup d := by
rw [LorentzGroup.mem_iff_invariant]
intro x y
rw [toMatrix_mulVec, toMatrix_mulVec]
have h1 : (((1 + ⟪u.1.1, v.1.1⟫ₘ)) * (1 + ⟪u.1.1, v.1.1⟫ₘ)) •
(contrContrContractField ((genBoost u v) y ⊗ₜ[] (genBoost u v) x))
= (((1 + ⟪u.1.1, v.1.1⟫ₘ)) * (1 + ⟪u.1.1, v.1.1⟫ₘ)) • ⟪y, x⟫ₘ := by
conv_lhs =>
erw [← map_smul]
rw [← smul_smul]
rw [← tmul_smul, ← smul_tmul, ← tmul_smul, genBoost_mul_one_plus_contr,
genBoost_mul_one_plus_contr]
simp only [add_smul, one_smul, tmul_add, map_add, smul_add, tmul_sub, sub_tmul, add_tmul,
smul_tmul, tmul_smul, map_sub, map_smul, smul_eq_mul, contr_self, mul_one]
rw [contrContrContractField.symm v.1.1 u, contrContrContractField.symm v.1.1 x,
contrContrContractField.symm u.1.1 x]
simp only [smul_eq_mul]
ring
have hn (a : ) {b c : } (h : a ≠ 0) (hab : a * b = a * c) : b = c := by
simp_all only [smul_eq_mul, ne_eq, mul_eq_mul_left_iff, or_false]
refine hn _ ?_ h1
simpa using (FuturePointing.one_add_metric_non_zero u v)
/-- A generalised boost as an element of the Lorentz Group. -/
def toLorentz (u v : FuturePointing d) : LorentzGroup d :=
⟨toMatrix u v, toMatrix_in_lorentzGroup u v⟩
lemma toLorentz_continuous (u : FuturePointing d) : Continuous (toLorentz u) := by
refine Continuous.subtype_mk ?_ (fun x => toMatrix_in_lorentzGroup u x)
exact toMatrix_continuous u
lemma toLorentz_joined_to_1 (u v : FuturePointing d) : Joined 1 (toLorentz u v) := by
obtain ⟨f, _⟩ := FuturePointing.isPathConnected.joinedIn u trivial v trivial
use ContinuousMap.comp ⟨toLorentz u, toLorentz_continuous u⟩ f
· simp only [ContinuousMap.toFun_eq_coe, ContinuousMap.comp_apply, ContinuousMap.coe_coe,
Path.source, ContinuousMap.coe_mk]
rw [@Subtype.ext_iff, toLorentz]
simp [toMatrix, self u]
· simp
lemma toLorentz_in_connected_component_1 (u v : FuturePointing d) :
toLorentz u v ∈ connectedComponent 1 :=
pathComponent_subset_component _ (toLorentz_joined_to_1 u v)
lemma isProper (u v : FuturePointing d) : IsProper (toLorentz u v) :=
(isProper_on_connected_component (toLorentz_in_connected_component_1 u v)).mp id_IsProper
end genBoost
end LorentzGroup
end

View file

@ -3,8 +3,8 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.NormOne
import HepLean.SpaceTime.LorentzGroup.Proper
import HepLean.Lorentz.RealVector.NormOne
import HepLean.Lorentz.Group.Proper
/-!
# The Orthochronous Lorentz Group
@ -24,46 +24,35 @@ namespace LorentzGroup
variable {d : }
variable (Λ : LorentzGroup d)
open LorentzVector
open minkowskiMetric
open Lorentz.Contr
/-- A Lorentz transformation is `orthochronous` if its `0 0` element is non-negative. -/
def IsOrthochronous : Prop := 0 ≤ timeComp Λ
def IsOrthochronous : Prop := 0 ≤ Λ.1 (Sum.inl 0) (Sum.inl 0)
lemma IsOrthochronous_iff_futurePointing :
IsOrthochronous Λ ↔ (toNormOneLorentzVector Λ) ∈ NormOneLorentzVector.FuturePointing d := by
simp only [IsOrthochronous, timeComp_eq_toNormOneLorentzVector]
rw [NormOneLorentzVector.FuturePointing.mem_iff_time_nonneg]
IsOrthochronous Λ ↔ toNormOne Λ ∈ NormOne.FuturePointing d := by
simp only [IsOrthochronous]
rw [NormOne.FuturePointing.mem_iff_inl_nonneg, toNormOne_inl]
lemma IsOrthochronous_iff_transpose :
IsOrthochronous Λ ↔ IsOrthochronous (transpose Λ) := by rfl
lemma IsOrthochronous_iff_ge_one :
IsOrthochronous Λ ↔ 1 ≤ timeComp Λ := by
rw [IsOrthochronous_iff_futurePointing, NormOneLorentzVector.FuturePointing.mem_iff,
NormOneLorentzVector.time_pos_iff]
simp only [time, toNormOneLorentzVector, timeVec, Fin.isValue]
erw [Pi.basisFun_apply, Matrix.mulVec_single_one]
rfl
IsOrthochronous Λ ↔ 1 ≤ Λ.1 (Sum.inl 0) (Sum.inl 0) := by
rw [IsOrthochronous_iff_futurePointing, NormOne.FuturePointing.mem_iff_inl_one_le_inl,
toNormOne_inl]
lemma not_orthochronous_iff_le_neg_one :
¬ IsOrthochronous Λ ↔ timeComp Λ ≤ -1 := by
rw [timeComp, IsOrthochronous_iff_futurePointing, NormOneLorentzVector.FuturePointing.not_mem_iff,
NormOneLorentzVector.time_nonpos_iff]
simp only [time, toNormOneLorentzVector, timeVec, Fin.isValue]
erw [Pi.basisFun_apply, Matrix.mulVec_single_one]
rfl
¬ IsOrthochronous Λ ↔ Λ.1 (Sum.inl 0) (Sum.inl 0) ≤ -1 := by
rw [IsOrthochronous_iff_futurePointing, NormOne.FuturePointing.not_mem_iff_inl_le_neg_one,
toNormOne_inl]
lemma not_orthochronous_iff_le_zero :
¬ IsOrthochronous Λ ↔ timeComp Λ ≤ 0 := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rw [not_orthochronous_iff_le_neg_one] at h
linarith
· rw [IsOrthochronous_iff_ge_one]
linarith
lemma not_orthochronous_iff_le_zero : ¬ IsOrthochronous Λ ↔ Λ.1 (Sum.inl 0) (Sum.inl 0) ≤ 0 := by
rw [IsOrthochronous_iff_futurePointing, NormOne.FuturePointing.not_mem_iff_inl_le_zero,
toNormOne_inl]
/-- The continuous map taking a Lorentz transformation to its `0 0` element. -/
def timeCompCont : C(LorentzGroup d, ) := ⟨fun Λ => timeComp Λ,
def timeCompCont : C(LorentzGroup d, ) := ⟨fun Λ => Λ.1 (Sum.inl 0) (Sum.inl 0),
Continuous.matrix_elem (continuous_iff_le_induced.mpr fun _ a => a) (Sum.inl 0) (Sum.inl 0)⟩
/-- An auxillary function used in the definition of `orthchroMapReal`. -/
@ -89,7 +78,7 @@ def orthchroMapReal : C(LorentzGroup d, ) := ContinuousMap.comp
lemma orthchroMapReal_on_IsOrthochronous {Λ : LorentzGroup d} (h : IsOrthochronous Λ) :
orthchroMapReal Λ = 1 := by
rw [IsOrthochronous_iff_ge_one, timeComp] at h
rw [IsOrthochronous_iff_ge_one] at h
change stepFunction (Λ.1 _ _) = 1
rw [stepFunction, if_pos h, if_neg]
linarith
@ -97,8 +86,9 @@ lemma orthchroMapReal_on_IsOrthochronous {Λ : LorentzGroup d} (h : IsOrthochron
lemma orthchroMapReal_on_not_IsOrthochronous {Λ : LorentzGroup d} (h : ¬ IsOrthochronous Λ) :
orthchroMapReal Λ = - 1 := by
rw [not_orthochronous_iff_le_neg_one] at h
change stepFunction (timeComp _)= - 1
rw [stepFunction, if_pos h]
change stepFunction _ = - 1
rw [stepFunction, if_pos]
exact h
lemma orthchroMapReal_minus_one_or_one (Λ : LorentzGroup d) :
orthchroMapReal Λ = -1 orthchroMapReal Λ = 1 := by
@ -130,29 +120,29 @@ lemma mul_othchron_of_othchron_othchron {Λ Λ' : LorentzGroup d} (h : IsOrthoch
(h' : IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by
rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_futurePointing] at h h'
rw [IsOrthochronous, timeComp_mul]
exact NormOneLorentzVector.FuturePointing.metric_reflect_mem_mem h h'
rw [IsOrthochronous, LorentzGroup.inl_inl_mul]
exact NormOne.FuturePointing.metric_reflect_mem_mem h h'
lemma mul_othchron_of_not_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h : ¬ IsOrthochronous Λ)
(h' : ¬ IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by
rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_futurePointing] at h h'
rw [IsOrthochronous, timeComp_mul]
exact NormOneLorentzVector.FuturePointing.metric_reflect_not_mem_not_mem h h'
rw [IsOrthochronous, LorentzGroup.inl_inl_mul]
exact NormOne.FuturePointing.metric_reflect_not_mem_not_mem h h'
lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h : IsOrthochronous Λ)
(h' : ¬ IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
rw [not_orthochronous_iff_le_zero, timeComp_mul]
rw [not_orthochronous_iff_le_zero, LorentzGroup.inl_inl_mul]
rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_futurePointing] at h h'
exact NormOneLorentzVector.FuturePointing.metric_reflect_mem_not_mem h h'
exact NormOne.FuturePointing.metric_reflect_mem_not_mem h h'
lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : LorentzGroup d} (h : ¬ IsOrthochronous Λ)
(h' : IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
rw [not_orthochronous_iff_le_zero, timeComp_mul]
rw [not_orthochronous_iff_le_zero, LorentzGroup.inl_inl_mul]
rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_futurePointing] at h h'
exact NormOneLorentzVector.FuturePointing.metric_reflect_not_mem_mem h h'
exact NormOne.FuturePointing.metric_reflect_not_mem_mem h h'
/-- The homomorphism from `LorentzGroup` to `ℤ₂` whose kernel are the Orthochronous elements. -/
def orthchroRep : LorentzGroup d →* ℤ₂ where

View file

@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzGroup.Basic
import HepLean.Lorentz.Group.Basic
/-!
# The Proper Lorentz Group
@ -21,7 +21,7 @@ open ComplexConjugate
namespace LorentzGroup
open minkowskiMetric
open minkowskiMatrix
variable {d : }

View file

@ -3,9 +3,9 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzGroup.Basic
import HepLean.SpaceTime.LorentzGroup.Proper
import HepLean.SpaceTime.LorentzGroup.Orthochronous
import HepLean.Lorentz.Group.Basic
import HepLean.Lorentz.Group.Proper
import HepLean.Lorentz.Group.Orthochronous
import HepLean.Meta.Informal
/-!
# The Restricted Lorentz Group
@ -25,8 +25,6 @@ open ComplexConjugate
namespace LorentzGroup
open minkowskiMetric
informal_definition Restricted where
math :≈ "The subgroup of the Lorentz group consisting of elements which
are proper and orthochronous."

View file

@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzGroup.Basic
import HepLean.Lorentz.Group.Basic
import HepLean.Mathematics.SO3.Basic
import Mathlib.Topology.Constructions
/-!
@ -25,7 +25,7 @@ def SO3ToMatrix (A : SO(3)) : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) :=
lemma SO3ToMatrix_in_LorentzGroup (A : SO(3)) : SO3ToMatrix A ∈ LorentzGroup 3 := by
rw [LorentzGroup.mem_iff_dual_mul_self]
simp only [minkowskiMetric.dual, minkowskiMatrix.as_block, SO3ToMatrix,
simp only [minkowskiMatrix.dual, minkowskiMatrix.as_block, SO3ToMatrix,
Matrix.fromBlocks_transpose, Matrix.transpose_one, Matrix.transpose_zero,
Matrix.fromBlocks_multiply, mul_one, Matrix.mul_zero, add_zero, Matrix.zero_mul, Matrix.mul_one,
neg_mul, one_mul, zero_add, Matrix.mul_neg, neg_zero, mul_neg, neg_neg,
@ -43,8 +43,8 @@ lemma SO3ToMatrix_injective : Function.Injective SO3ToMatrix := by
space-time rotation. -/
def SO3ToLorentz : SO(3) →* LorentzGroup 3 where
toFun A := ⟨SO3ToMatrix A, SO3ToMatrix_in_LorentzGroup A⟩
map_one' := Subtype.eq <|
(minkowskiMetric.matrix_eq_iff_eq_forall _ ↑1).mpr fun w => congrFun rfl
map_one' := Subtype.eq <| by
simp only [SO3ToMatrix, SO3Group_one_coe, Matrix.fromBlocks_one, lorentzGroupIsGroup_one_coe]
map_mul' A B := Subtype.eq <| by
simp only [SO3ToMatrix, SO3Group_mul_coe, lorentzGroupIsGroup_mul_coe,
Matrix.fromBlocks_multiply, mul_one, Matrix.mul_zero, add_zero, Matrix.zero_mul,

View file

@ -0,0 +1,150 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Data.Complex.Exponential
import Mathlib.Analysis.InnerProductSpace.PiL2
import Mathlib.Algebra.Lie.Classical
/-!
# The Minkowski matrix
-/
open Matrix
open InnerProductSpace
/-!
# The definition of the Minkowski Matrix
-/
/-- The `d.succ`-dimensional real matrix of the form `diag(1, -1, -1, -1, ...)`. -/
def minkowskiMatrix {d : } : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) :=
LieAlgebra.Orthogonal.indefiniteDiagonal (Fin 1) (Fin d)
namespace minkowskiMatrix
variable {d : }
/-- Notation for `minkowskiMatrix`. -/
scoped[minkowskiMatrix] notation "η" => minkowskiMatrix
@[simp]
lemma sq : @minkowskiMatrix d * minkowskiMatrix = 1 := by
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_mul_diagonal]
ext1 i j
rcases i with i | i <;> rcases j with j | j
· simp only [diagonal, of_apply, Sum.inl.injEq, Sum.elim_inl, mul_one]
split
· rename_i h
subst h
simp_all only [one_apply_eq]
· simp_all only [ne_eq, Sum.inl.injEq, not_false_eq_true, one_apply_ne]
· rfl
· rfl
· simp only [diagonal, of_apply, Sum.inr.injEq, Sum.elim_inr, mul_neg, mul_one, neg_neg]
split
· rename_i h
subst h
simp_all only [one_apply_eq]
· simp_all only [ne_eq, Sum.inr.injEq, not_false_eq_true, one_apply_ne]
@[simp]
lemma eq_transpose : minkowskiMatrixᵀ = @minkowskiMatrix d := by
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_transpose]
@[simp]
lemma det_eq_neg_one_pow_d : (@minkowskiMatrix d).det = (- 1) ^ d := by
simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
@[simp]
lemma η_apply_mul_η_apply_diag (μ : Fin 1 ⊕ Fin d) : η μ μ * η μ μ = 1 := by
match μ with
| Sum.inl _ => simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
| Sum.inr _ => simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
lemma as_block : @minkowskiMatrix d =
Matrix.fromBlocks (1 : Matrix (Fin 1) (Fin 1) ) 0 0 (-1 : Matrix (Fin d) (Fin d) ) := by
rw [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, ← fromBlocks_diagonal]
refine fromBlocks_inj.mpr ?_
simp only [diagonal_one, true_and]
funext i j
rw [← diagonal_neg]
rfl
@[simp]
lemma off_diag_zero {μ ν : Fin 1 ⊕ Fin d} (h : μ ≠ ν) : η μ ν = 0 := by
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
exact diagonal_apply_ne _ h
lemma inl_0_inl_0 : @minkowskiMatrix d (Sum.inl 0) (Sum.inl 0) = 1 := by
rfl
lemma inr_i_inr_i (i : Fin d) : @minkowskiMatrix d (Sum.inr i) (Sum.inr i) = -1 := by
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
simp_all only [diagonal_apply_eq, Sum.elim_inr]
@[simp]
lemma mulVec_inl_0 (v : (Fin 1 ⊕ Fin d) → ) :
(η *ᵥ v) (Sum.inl 0)= v (Sum.inl 0) := by
simp only [mulVec, minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, mulVec_diagonal]
simp only [Fin.isValue, diagonal_dotProduct, Sum.elim_inl, one_mul]
@[simp]
lemma mulVec_inr_i (v : (Fin 1 ⊕ Fin d) → ) (i : Fin d) :
(η *ᵥ v) (Sum.inr i)= - v (Sum.inr i) := by
simp only [mulVec, minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, mulVec_diagonal]
simp only [diagonal_dotProduct, Sum.elim_inr, neg_mul, one_mul]
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
@[simp]
lemma dual_mul : dual (Λ * Λ') = dual Λ' * dual Λ := by
simp only [dual, transpose_mul]
trans η * Λ'ᵀ * (η * η) * Λᵀ * η
· noncomm_ring [minkowskiMatrix.sq]
· noncomm_ring
@[simp]
lemma dual_dual : dual (dual Λ) = Λ := by
simp only [dual, transpose_mul, transpose_transpose, eq_transpose]
trans (η * η) * Λ * (η * η)
· noncomm_ring
· noncomm_ring [minkowskiMatrix.sq]
@[simp]
lemma dual_eta : @dual d η = η := by
simp only [dual, eq_transpose]
noncomm_ring [minkowskiMatrix.sq]
@[simp]
lemma dual_transpose : dual Λᵀ = (dual Λ)ᵀ := by
simp only [dual, transpose_transpose, transpose_mul, eq_transpose]
noncomm_ring
@[simp]
lemma det_dual : (dual Λ).det = Λ.det := by
simp only [dual, det_mul, minkowskiMatrix.det_eq_neg_one_pow_d, det_transpose]
group
norm_cast
simp
lemma dual_apply (μ ν : Fin 1 ⊕ Fin d) :
dual Λ μ ν = η μ μ * Λ ν μ * η ν ν := by
simp only [dual, minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, mul_diagonal,
diagonal_mul, transpose_apply, diagonal_apply_eq]
lemma dual_apply_minkowskiMatrix (μ ν : Fin 1 ⊕ Fin d) :
dual Λ μ ν * η ν ν = η μ μ * Λ ν μ := by
rw [dual_apply, mul_assoc]
simp
end minkowskiMatrix

View file

@ -5,9 +5,9 @@ Authors: Joseph Tooby-Smith
-/
import HepLean.Tensors.OverColor.Basic
import HepLean.Mathematics.PiTensorProduct
import HepLean.SpaceTime.LorentzVector.Complex.Basic
import HepLean.SpaceTime.WeylFermion.Two
import HepLean.SpaceTime.PauliMatrices.Basic
import HepLean.Lorentz.ComplexVector.Basic
import HepLean.Lorentz.Weyl.Two
import HepLean.Lorentz.PauliMatrices.Basic
/-!
## Pauli matrices
@ -28,7 +28,6 @@ open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
/-- The tensor `σ^μ^a^{dot a}` based on the Pauli-matrices as an element of
`complexContr ⊗ leftHanded ⊗ rightHanded`. -/
@ -131,14 +130,14 @@ def asConsTensor : 𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ leftHanded ⊗
map_sum, map_tmul]
symm
calc _ = ∑ x, ((complexContr.ρ M) (complexContrBasis x) ⊗ₜ[]
leftRightToMatrix.symm (SL2C.repSelfAdjointMatrix M (σSA x))) := by
leftRightToMatrix.symm (SL2C.toLinearMapSelfAdjointMatrix 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
refine Finset.sum_congr rfl (fun x _ => ?_)
rw [SL2CRep_ρ_basis, SL2C.repSelfAdjointMatrix_σSA]
rw [SL2CRep_ρ_basis, SL2C.toLinearMapSelfAdjointMatrix_σSA]
simp only [Action.instMonoidalCategory_tensorObj_V, SL2C.toLorentzGroup_apply_coe,
Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, map_inv, lorentzGroupIsGroup_inv, AddSubgroup.coe_add,

View file

@ -6,7 +6,7 @@ Authors: Joseph Tooby-Smith
import HepLean.Mathematics.PiTensorProduct
import Mathlib.RepresentationTheory.Rep
import Mathlib.Logic.Equiv.TransferInstance
import HepLean.SpaceTime.LorentzGroup.Basic
import HepLean.Lorentz.Group.Basic
/-!
## Pauli matrices
@ -18,16 +18,20 @@ namespace PauliMatrix
open Complex
open Matrix
/-- The zeroth Pauli-matrix as a `2 x 2` complex matrix. -/
/-- The zeroth Pauli-matrix as a `2 x 2` complex matrix.
That is the matrix `!![1, 0; 0, 1]`. -/
def σ0 : Matrix (Fin 2) (Fin 2) := !![1, 0; 0, 1]
/-- The first Pauli-matrix as a `2 x 2` complex matrix. -/
/-- The first Pauli-matrix as a `2 x 2` complex matrix.
That is, the matrix `!![0, 1; 1, 0]`. -/
def σ1 : Matrix (Fin 2) (Fin 2) := !![0, 1; 1, 0]
/-- The second Pauli-matrix as a `2 x 2` complex matrix. -/
/-- The second Pauli-matrix as a `2 x 2` complex matrix.
That is, the matrix `!![0, -I; I, 0]`. -/
def σ2 : Matrix (Fin 2) (Fin 2) := !![0, -I; I, 0]
/-- The third Pauli-matrix as a `2 x 2` complex matrix. -/
/-- The third Pauli-matrix as a `2 x 2` complex matrix.
That is, the matrix `!![1, 0; 0, -1]`. -/
def σ3 : Matrix (Fin 2) (Fin 2) := !![1, 0; 0, -1]
@[simp]

View file

@ -6,8 +6,8 @@ Authors: Joseph Tooby-Smith
import HepLean.Mathematics.PiTensorProduct
import Mathlib.RepresentationTheory.Rep
import Mathlib.Logic.Equiv.TransferInstance
import HepLean.SpaceTime.LorentzGroup.Basic
import HepLean.SpaceTime.PauliMatrices.Basic
import HepLean.Lorentz.Group.Basic
import HepLean.Lorentz.PauliMatrices.Basic
/-!
## Interaction of Pauli matrices with self-adjoint matrices

View file

@ -0,0 +1,136 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Data.Complex.Exponential
import Mathlib.Analysis.InnerProductSpace.PiL2
import HepLean.Lorentz.Group.Basic
import HepLean.Meta.Informal
import Mathlib.RepresentationTheory.Rep
import HepLean.Lorentz.RealVector.Modules
/-!
# Real Lorentz vectors
We define real Lorentz vectors in as representations of the Lorentz group.
-/
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
namespace Lorentz
open minkowskiMatrix
/-- The representation of `LorentzGroup d` on real vectors corresponding to contravariant
Lorentz vectors. In index notation these have an up index `ψⁱ`. -/
def Contr (d : ) : Rep (LorentzGroup d) := Rep.of ContrMod.rep
instance : TopologicalSpace (Contr d) := TopologicalSpace.induced
ContrMod.toFin1dEquiv (Pi.topologicalSpace)
lemma continuous_contr {T : Type} [TopologicalSpace T] (f : T → Contr d)
(h : Continuous (fun i => (f i).toFin1d)) : Continuous f := by
exact continuous_induced_rng.mpr h
lemma contr_continuous {T : Type} [TopologicalSpace T] (f : Contr d → T)
(h : Continuous (f ∘ (@ContrMod.toFin1dEquiv d).symm)) : Continuous f := by
let x := Equiv.toHomeomorphOfIsInducing (@ContrMod.toFin1dEquiv d).toEquiv
ContrMod.toFin1dEquiv_isInducing
rw [← Homeomorph.comp_continuous_iff' x.symm]
exact h
/-- The representation of `LorentzGroup d` on real vectors corresponding to covariant
Lorentz vectors. In index notation these have an up index `ψⁱ`. -/
def Co (d : ) : Rep (LorentzGroup d) := Rep.of CoMod.rep
open CategoryTheory.MonoidalCategory
/-!
## Isomorphism between contravariant and covariant Lorentz vectors
-/
/-- The morphism of representations from `Contr d` to `Co d` defined by multiplication
with the metric. -/
def Contr.toCo (d : ) : Contr d ⟶ Co d where
hom := {
toFun := fun ψ => CoMod.toFin1dEquiv.symm (η *ᵥ ψ.toFin1d),
map_add' := by
intro ψ ψ'
simp only [map_add, mulVec_add]
map_smul' := by
intro r ψ
simp only [_root_.map_smul, mulVec_smul, RingHom.id_apply]}
comm g := by
ext ψ : 2
simp only [ModuleCat.coe_comp, Function.comp_apply]
conv_lhs =>
change CoMod.toFin1dEquiv.symm (η *ᵥ (g.1 *ᵥ ψ.toFin1d))
rw [mulVec_mulVec, LorentzGroup.minkowskiMatrix_comm, ← mulVec_mulVec]
rfl
/-- The morphism of representations from `Co d` to `Contr d` defined by multiplication
with the metric. -/
def Co.toContr (d : ) : Co d ⟶ Contr d where
hom := {
toFun := fun ψ => ContrMod.toFin1dEquiv.symm (η *ᵥ ψ.toFin1d),
map_add' := by
intro ψ ψ'
simp only [map_add, mulVec_add]
map_smul' := by
intro r ψ
simp only [_root_.map_smul, mulVec_smul, RingHom.id_apply]}
comm g := by
ext ψ : 2
simp only [ModuleCat.coe_comp, Function.comp_apply]
conv_lhs =>
change ContrMod.toFin1dEquiv.symm (η *ᵥ ((LorentzGroup.transpose g⁻¹).1 *ᵥ ψ.toFin1d))
rw [mulVec_mulVec, ← LorentzGroup.comm_minkowskiMatrix, ← mulVec_mulVec]
rfl
/-- The isomorphism between `Contr d` and `Co d` induced by multiplication with the
Minkowski metric. -/
def contrIsoCo (d : ) : Contr d ≅ Co d where
hom := Contr.toCo d
inv := Co.toContr d
hom_inv_id := by
ext ψ
simp only [Action.comp_hom, ModuleCat.coe_comp, Function.comp_apply, Action.id_hom,
ModuleCat.id_apply]
conv_lhs => change ContrMod.toFin1dEquiv.symm (η *ᵥ
CoMod.toFin1dEquiv (CoMod.toFin1dEquiv.symm (η *ᵥ ψ.toFin1d)))
rw [LinearEquiv.apply_symm_apply, mulVec_mulVec, minkowskiMatrix.sq]
simp
inv_hom_id := by
ext ψ
simp only [Action.comp_hom, ModuleCat.coe_comp, Function.comp_apply, Action.id_hom,
ModuleCat.id_apply]
conv_lhs => change CoMod.toFin1dEquiv.symm (η *ᵥ
ContrMod.toFin1dEquiv (ContrMod.toFin1dEquiv.symm (η *ᵥ ψ.toFin1d)))
rw [LinearEquiv.apply_symm_apply, mulVec_mulVec, minkowskiMatrix.sq]
simp
/-!
## Other properties
-/
namespace Contr
open Lorentz
lemma ρ_stdBasis (μ : Fin 1 ⊕ Fin 3) (Λ : LorentzGroup 3) :
(Contr 3).ρ Λ (ContrMod.stdBasis μ) = ∑ j, Λ.1 j μ • ContrMod.stdBasis j := by
change Λ *ᵥ ContrMod.stdBasis μ = ∑ j, Λ.1 j μ • ContrMod.stdBasis j
apply ContrMod.ext
simp only [toLinAlgEquiv_self, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero,
Fin.isValue, Finset.sum_singleton, ContrMod.val_add, ContrMod.val_smul]
end Contr
end Lorentz
end

View file

@ -0,0 +1,451 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Lorentz.RealVector.Basic
/-!
# Contraction of Real Lorentz vectors
-/
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open CategoryTheory.MonoidalCategory
open minkowskiMatrix
namespace Lorentz
variable {d : }
/-- The bi-linear map corresponding to contraction of a contravariant Lorentz vector with a
covariant Lorentz vector. -/
def contrModCoModBi (d : ) : ContrMod d →ₗ[] CoMod d →ₗ[] where
toFun ψ := {
toFun := fun φ => ψ.toFin1d ⬝ᵥ φ.toFin1d,
map_add' := by
intro φ φ'
simp only [map_add]
rw [dotProduct_add]
map_smul' := by
intro r φ
simp only [LinearEquiv.map_smul]
rw [dotProduct_smul]
rfl}
map_add' ψ ψ':= by
refine LinearMap.ext (fun φ => ?_)
simp only [map_add, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.add_apply]
rw [add_dotProduct]
map_smul' r ψ := by
refine LinearMap.ext (fun φ => ?_)
simp only [LinearEquiv.map_smul, LinearMap.coe_mk, AddHom.coe_mk]
rw [smul_dotProduct]
rfl
/-- The bi-linear map corresponding to contraction of a covariant Lorentz vector with a
contravariant Lorentz vector. -/
def coModContrModBi (d : ) : CoMod d →ₗ[] ContrMod d →ₗ[] where
toFun φ := {
toFun := fun ψ => φ.toFin1d ⬝ᵥ ψ.toFin1d,
map_add' := by
intro ψ ψ'
simp only [map_add]
rw [dotProduct_add]
map_smul' := by
intro r ψ
simp only [LinearEquiv.map_smul]
rw [dotProduct_smul]
rfl}
map_add' φ φ' := by
refine LinearMap.ext (fun ψ => ?_)
simp only [map_add, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.add_apply]
rw [add_dotProduct]
map_smul' r φ := by
refine LinearMap.ext (fun ψ => ?_)
simp only [LinearEquiv.map_smul, LinearMap.coe_mk, AddHom.coe_mk]
rw [smul_dotProduct]
rfl
/-- The linear map from Contr d ⊗ Co d to given by
summing over components of contravariant Lorentz vector and
covariant Lorentz vector in the
standard basis (i.e. the dot product).
In terms of index notation this is the contraction is ψⁱ φᵢ. -/
def contrCoContract : Contr d ⊗ Co d ⟶ 𝟙_ (Rep (LorentzGroup d)) where
hom := TensorProduct.lift (contrModCoModBi d)
comm M := TensorProduct.ext' fun ψ φ => by
change (M.1 *ᵥ ψ.toFin1d) ⬝ᵥ ((LorentzGroup.transpose M⁻¹).1 *ᵥ φ.toFin1d) = _
rw [dotProduct_mulVec, LorentzGroup.transpose_val,
vecMul_transpose, mulVec_mulVec, LorentzGroup.coe_inv, inv_mul_of_invertible M.1]
simp only [one_mulVec, CategoryTheory.Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
Action.tensorUnit_ρ', CategoryTheory.Category.comp_id, lift.tmul]
rfl
/-- Notation for `contrCoContract` acting on a tmul. -/
scoped[Lorentz] notation "⟪" ψ "," φ "⟫ₘ" => contrCoContract.hom (ψ ⊗ₜ φ)
lemma contrCoContract_hom_tmul (ψ : Contr d) (φ : Co d) : ⟪ψ, φ⟫ₘ = ψ.toFin1d ⬝ᵥ φ.toFin1d := by
rfl
/-- The linear map from Co d ⊗ Contr d to given by
summing over components of contravariant Lorentz vector and
covariant Lorentz vector in the
standard basis (i.e. the dot product).
In terms of index notation this is the contraction is ψⁱ φᵢ. -/
def coContrContract : Co d ⊗ Contr d ⟶ 𝟙_ (Rep (LorentzGroup d)) where
hom := TensorProduct.lift (coModContrModBi d)
comm M := TensorProduct.ext' fun ψ φ => by
change ((LorentzGroup.transpose M⁻¹).1 *ᵥ ψ.toFin1d) ⬝ᵥ (M.1 *ᵥ φ.toFin1d) = _
rw [dotProduct_mulVec, LorentzGroup.transpose_val, mulVec_transpose, vecMul_vecMul,
LorentzGroup.coe_inv, inv_mul_of_invertible M.1]
simp only [vecMul_one, CategoryTheory.Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
Action.tensorUnit_ρ', CategoryTheory.Category.comp_id, lift.tmul]
rfl
/-- Notation for `coContrContract` acting on a tmul. -/
scoped[Lorentz] notation "⟪" φ "," ψ "⟫ₘ" => coContrContract.hom (φ ⊗ₜ ψ)
lemma coContrContract_hom_tmul (φ : Co d) (ψ : Contr d) : ⟪φ, ψ⟫ₘ = φ.toFin1d ⬝ᵥ ψ.toFin1d := by
rfl
/-!
## Symmetry relations
-/
lemma contrCoContract_tmul_symm (φ : Contr d) (ψ : Co d) : ⟪φ, ψ⟫ₘ = ⟪ψ, φ⟫ₘ := by
rw [contrCoContract_hom_tmul, coContrContract_hom_tmul, dotProduct_comm]
lemma coContrContract_tmul_symm (φ : Co d) (ψ : Contr d) : ⟪φ, ψ⟫ₘ = ⟪ψ, φ⟫ₘ := by
rw [contrCoContract_tmul_symm]
/-!
## Contracting contr vectors with contr vectors etc.
-/
open CategoryTheory.MonoidalCategory
open CategoryTheory
/-- The linear map from Contr d ⊗ Contr d to induced by the homomorphism
`Contr.toCo` and the contraction `contrCoContract`. -/
def contrContrContract : Contr d ⊗ Contr d ⟶ 𝟙_ (Rep (LorentzGroup d)) :=
(Contr d ◁ Contr.toCo d) ≫ contrCoContract
/-- The linear map from Contr d ⊗ Contr d to induced by the homomorphism
`Contr.toCo` and the contraction `contrCoContract`. -/
def contrContrContractField : (Contr d).V ⊗[] (Contr d).V →ₗ[] :=
contrContrContract.hom
/-- Notation for `contrContrContractField` acting on a tmul. -/
scoped[Lorentz] notation "⟪" ψ "," φ "⟫ₘ" => contrContrContractField (ψ ⊗ₜ φ)
lemma contrContrContract_hom_tmul (φ : Contr d) (ψ : Contr d) :
⟪φ, ψ⟫ₘ = φ.toFin1d ⬝ᵥ η *ᵥ ψ.toFin1d:= by
simp only [Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
contrContrContractField, Action.comp_hom, Action.instMonoidalCategory_whiskerLeft_hom,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, ModuleCat.coe_comp, Function.comp_apply,
ModuleCat.MonoidalCategory.whiskerLeft_apply]
erw [contrCoContract_hom_tmul]
rfl
/-- The linear map from Co d ⊗ Co d to induced by the homomorphism
`Co.toContr` and the contraction `coContrContract`. -/
def coCoContract : Co d ⊗ Co d ⟶ 𝟙_ (Rep (LorentzGroup d)) :=
(Co d ◁ Co.toContr d) ≫ coContrContract
/-- Notation for `coCoContract` acting on a tmul. -/
scoped[Lorentz] notation "⟪" ψ "," φ "⟫ₘ" => coCoContract.hom (ψ ⊗ₜ φ)
lemma coCoContract_hom_tmul (φ : Co d) (ψ : Co d) :
⟪φ, ψ⟫ₘ = φ.toFin1d ⬝ᵥ η *ᵥ ψ.toFin1d:= by
simp only [Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
contrContrContract, Action.comp_hom, Action.instMonoidalCategory_whiskerLeft_hom,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, ModuleCat.coe_comp, Function.comp_apply,
ModuleCat.MonoidalCategory.whiskerLeft_apply]
erw [coContrContract_hom_tmul]
rfl
/-!
## Lemmas related to contraction.
We derive the lemmas in main for `contrContrContractField`.
-/
namespace contrContrContractField
variable (x y : Contr d)
@[simp]
lemma action_tmul (g : LorentzGroup d) : ⟪(Contr d).ρ g x, (Contr d).ρ g y⟫ₘ = ⟪x, y⟫ₘ := by
conv =>
lhs
change (CategoryTheory.CategoryStruct.comp
((CategoryTheory.MonoidalCategory.tensorObj (Contr d) (Contr d)).ρ g)
contrContrContract.hom) (x ⊗ₜ[] y)
arg 1
apply contrContrContract.comm g
rfl
lemma as_sum : ⟪x, y⟫ₘ = x.val (Sum.inl 0) * y.val (Sum.inl 0) -
∑ i, x.val (Sum.inr i) * y.val (Sum.inr i) := by
rw [contrContrContract_hom_tmul]
simp only [dotProduct, minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, mulVec_diagonal,
Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue, Sum.elim_inl,
one_mul, Finset.sum_singleton, Sum.elim_inr, neg_mul, mul_neg, Finset.sum_neg_distrib]
rfl
open InnerProductSpace
lemma as_sum_toSpace : ⟪x, y⟫ₘ = x.val (Sum.inl 0) * y.val (Sum.inl 0) -
⟪x.toSpace, y.toSpace⟫_ := by
rw [as_sum]
rfl
lemma stdBasis_inl {d : } :
⟪@ContrMod.stdBasis d (Sum.inl 0), ContrMod.stdBasis (Sum.inl 0)⟫ₘ = (1 : ) := by
rw [as_sum]
trans (1 : ) - (0 : )
congr
· rw [ContrMod.stdBasis_apply_same]
simp
· rw [Fintype.sum_eq_zero]
intro a
simp
· ring
lemma symm : ⟪x, y⟫ₘ = ⟪y, x⟫ₘ := by
rw [as_sum, as_sum]
congr 1
rw [mul_comm]
congr
funext i
rw [mul_comm]
lemma dual_mulVec_right : ⟪x, dual Λ *ᵥ y⟫ₘ = ⟪Λ *ᵥ x, y⟫ₘ := by
rw [contrContrContract_hom_tmul, contrContrContract_hom_tmul]
simp only [Action.instMonoidalCategory_tensorUnit_V, ContrMod.mulVec_toFin1d, mulVec_mulVec]
simp only [dual, ← mul_assoc, minkowskiMatrix.sq, one_mul]
rw [← mulVec_mulVec, dotProduct_mulVec, vecMul_transpose]
lemma dual_mulVec_left : ⟪dual Λ *ᵥ x, y⟫ₘ = ⟪x, Λ *ᵥ y⟫ₘ := by
rw [symm, dual_mulVec_right, symm]
lemma right_parity : ⟪x, (Contr d).ρ LorentzGroup.parity y⟫ₘ = ∑ i, x.val i * y.val i := by
rw [as_sum]
simp only [Action.instMonoidalCategory_tensorUnit_V, Fin.isValue, Fintype.sum_sum_type,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
trans x.val (Sum.inl 0) * (((Contr d).ρ LorentzGroup.parity) y).val (Sum.inl 0) +
∑ i : Fin d, - (x.val (Sum.inr i) * (((Contr d).ρ LorentzGroup.parity) y).val (Sum.inr i))
· simp only [Fin.isValue, Finset.sum_neg_distrib]
rfl
congr 1
· change x.val (Sum.inl 0) * (η *ᵥ y.toFin1d) (Sum.inl 0) = _
simp only [Fin.isValue, mulVec_inl_0, mul_eq_mul_left_iff]
exact mul_eq_mul_left_iff.mp rfl
· congr
funext i
change - (x.val (Sum.inr i) * ((η *ᵥ y.toFin1d) (Sum.inr i))) = _
simp only [mulVec_inr_i, mul_neg, neg_neg, mul_eq_mul_left_iff]
exact mul_eq_mul_left_iff.mp rfl
lemma self_parity_eq_zero_iff : ⟪y, (Contr d).ρ LorentzGroup.parity y⟫ₘ = 0 ↔ y = 0 := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rw [right_parity] at h
have hn := Fintype.sum_eq_zero_iff_of_nonneg (f := fun i => y.val i * y.val i) (fun i => by
simpa using mul_self_nonneg (y.val i))
rw [h] at hn
simp only [true_iff] at hn
apply ContrMod.ext
funext i
simpa using congrFun hn i
· rw [h]
simp only [Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, map_zero, tmul_zero]
/-- The metric tensor is non-degenerate. -/
lemma nondegenerate : (∀ (x : Contr d), ⟪x, y⟫ₘ = 0) ↔ y = 0 := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· exact (self_parity_eq_zero_iff _).mp ((symm _ _).trans $ h _)
· simp [h]
lemma matrix_apply_eq_iff_sub : ⟪x, Λ *ᵥ y⟫ₘ = ⟪x, Λ' *ᵥ y⟫ₘ ↔ ⟪x, (Λ - Λ') *ᵥ y⟫ₘ = 0 := by
rw [← sub_eq_zero, ← LinearMap.map_sub, ← tmul_sub, ← ContrMod.sub_mulVec Λ Λ' y]
lemma matrix_eq_iff_eq_forall' : (∀ (v : Contr d), (Λ *ᵥ v) = Λ' *ᵥ v) ↔
∀ (w v : Contr d), ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, Λ' *ᵥ w⟫ₘ := by
refine Iff.intro (fun h ↦ fun w v ↦ ?_) (fun h ↦ fun v ↦ ?_)
· rw [h w]
· simp only [matrix_apply_eq_iff_sub] at h
refine sub_eq_zero.1 ?_
have h1 := h v
rw [nondegenerate] at h1
simp only [ContrMod.sub_mulVec] at h1
exact h1
lemma matrix_eq_iff_eq_forall : Λ = Λ' ↔ ∀ (w v : Contr d), ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, Λ' *ᵥ w⟫ₘ := by
rw [← matrix_eq_iff_eq_forall']
refine Iff.intro (fun h => ?_) (fun h => ?_)
· subst h
exact fun v => rfl
· rw [← (LinearMap.toMatrix ContrMod.stdBasis ContrMod.stdBasis).toEquiv.symm.apply_eq_iff_eq]
ext1 v
exact h v
lemma matrix_eq_id_iff : Λ = 1 ↔ ∀ (w v : Contr d), ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, w⟫ₘ := by
rw [matrix_eq_iff_eq_forall]
simp only [Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, ContrMod.one_mulVec]
lemma _root_.LorentzGroup.mem_iff_invariant : Λ ∈ LorentzGroup d ↔
∀ (w v : Contr d), ⟪Λ *ᵥ v, Λ *ᵥ w⟫ₘ = ⟪v, w⟫ₘ := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· intro x y
rw [← dual_mulVec_right, ContrMod.mulVec_mulVec]
have h1 := LorentzGroup.mem_iff_dual_mul_self.mp h
rw [h1]
simp only [Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, ContrMod.one_mulVec]
· conv at h =>
intro x y
rw [← dual_mulVec_right, ContrMod.mulVec_mulVec]
rw [← matrix_eq_id_iff] at h
exact LorentzGroup.mem_iff_dual_mul_self.mpr h
lemma _root_.LorentzGroup.mem_iff_norm : Λ ∈ LorentzGroup d ↔
∀ (w : Contr d), ⟪Λ *ᵥ w, Λ *ᵥ w⟫ₘ = ⟪w, w⟫ₘ := by
rw [LorentzGroup.mem_iff_invariant]
refine Iff.intro (fun h x => h x x) (fun h x y => ?_)
have hp := h (x + y)
have hn := h (x - y)
rw [ContrMod.mulVec_add, tmul_add, add_tmul, add_tmul, tmul_add, add_tmul, add_tmul] at hp
rw [ContrMod.mulVec_sub, tmul_sub, sub_tmul, sub_tmul, tmul_sub, sub_tmul, sub_tmul] at hn
simp only [map_add, LinearMap.add_apply, map_sub, LinearMap.sub_apply] at hp hn
rw [symm (Λ *ᵥ y) (Λ *ᵥ x), symm y x] at hp hn
let e : 𝟙_ (Rep ↑(LorentzGroup d)) ≃ₗ[] :=
LinearEquiv.refl (CoeSort.coe (𝟙_ (Rep ↑(LorentzGroup d))))
apply e.injective
have hp' := e.injective.eq_iff.mpr hp
have hn' := e.injective.eq_iff.mpr hn
simp only [Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, map_add, map_sub] at hp' hn'
linear_combination (norm := ring_nf) (1 / 4) * hp' + (-1/ 4) * hn'
rw [symm (Λ *ᵥ y) (Λ *ᵥ x), symm y x]
simp only [Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, add_sub_cancel, neg_add_cancel, e]
/-!
## Some equalities and inequalities
-/
lemma inl_sq_eq (v : Contr d) : v.val (Sum.inl 0) ^ 2 =
(⟪v, v⟫ₘ) + ∑ i, v.val (Sum.inr i) ^ 2:= by
rw [as_sum]
apply sub_eq_iff_eq_add.mp
congr
· exact pow_two (v.val (Sum.inl 0))
· funext i
exact pow_two (v.val (Sum.inr i))
lemma le_inl_sq (v : Contr d) : ⟪v, v⟫ₘ ≤ v.val (Sum.inl 0) ^ 2 := by
rw [inl_sq_eq]
apply (le_add_iff_nonneg_right _).mpr
refine Fintype.sum_nonneg ?hf
exact fun i => pow_two_nonneg (v.val (Sum.inr i))
lemma ge_abs_inner_product (v w : Contr d) : v.val (Sum.inl 0) * w.val (Sum.inl 0) -
‖⟪v.toSpace, w.toSpace⟫_‖ ≤ ⟪v, w⟫ₘ := by
rw [as_sum_toSpace, sub_le_sub_iff_left]
exact Real.le_norm_self ⟪v.toSpace, w.toSpace⟫_
lemma ge_sub_norm (v w : Contr d) : v.val (Sum.inl 0) * w.val (Sum.inl 0) -
‖v.toSpace‖ * ‖w.toSpace‖ ≤ ⟪v, w⟫ₘ := by
apply le_trans _ (ge_abs_inner_product v w)
rw [sub_le_sub_iff_left]
exact norm_inner_le_norm v.toSpace w.toSpace
/-!
# The Minkowski metric and the standard basis
-/
@[simp]
lemma basis_left {v : Contr d} (μ : Fin 1 ⊕ Fin d) :
⟪ ContrMod.stdBasis μ, v⟫ₘ = η μ μ * v.toFin1d μ := by
rw [as_sum]
rcases μ with μ | μ
· fin_cases μ
simp only [Fin.zero_eta, Fin.isValue, ContrMod.stdBasis_apply_same, one_mul,
ContrMod.stdBasis_inl_apply_inr, zero_mul, Finset.sum_const_zero, sub_zero, minkowskiMatrix,
LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_apply_eq, Sum.elim_inl]
rfl
· simp only [Action.instMonoidalCategory_tensorUnit_V, Fin.isValue, ContrMod.stdBasis_apply,
reduceCtorEq, ↓reduceIte, zero_mul, Sum.inr.injEq, ite_mul, one_mul, Finset.sum_ite_eq,
Finset.mem_univ, zero_sub, minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal,
diagonal_apply_eq, Sum.elim_inr, neg_mul, neg_inj]
rfl
lemma on_basis_mulVec (μ ν : Fin 1 ⊕ Fin d) :
⟪ContrMod.stdBasis μ, Λ *ᵥ ContrMod.stdBasis ν⟫ₘ = η μ μ * Λ μ ν := by
rw [basis_left, ContrMod.mulVec_toFin1d]
simp [basis_left, mulVec, dotProduct, ContrMod.stdBasis_apply, ContrMod.toFin1d_eq_val]
lemma on_basis (μ ν : Fin 1 ⊕ Fin d) : ⟪ContrMod.stdBasis μ, ContrMod.stdBasis ν⟫ₘ = η μ ν := by
trans ⟪ContrMod.stdBasis μ, 1 *ᵥ ContrMod.stdBasis ν⟫ₘ
· rw [ContrMod.one_mulVec]
rw [on_basis_mulVec]
by_cases h : μ = ν
· subst h
simp
· simp only [Action.instMonoidalCategory_tensorUnit_V, ne_eq, h, not_false_eq_true, one_apply_ne,
mul_zero, off_diag_zero]
lemma matrix_apply_stdBasis (ν μ : Fin 1 ⊕ Fin d) :
Λ ν μ = η ν ν * ⟪ ContrMod.stdBasis ν, Λ *ᵥ ContrMod.stdBasis μ⟫ₘ := by
rw [on_basis_mulVec, ← mul_assoc]
simp [η_apply_mul_η_apply_diag ν]
/-!
## Self-adjoint
-/
lemma same_eq_det_toSelfAdjoint (x : ContrMod 3) :
⟪x, x⟫ₘ = det (ContrMod.toSelfAdjoint x).1 := by
rw [ContrMod.toSelfAdjoint_apply_coe]
simp only [Fin.isValue, as_sum_toSpace,
PiLp.inner_apply, Function.comp_apply, RCLike.inner_apply, conj_trivial, Fin.sum_univ_three,
ofReal_sub, ofReal_mul, ofReal_add]
simp only [Fin.isValue, PauliMatrix.σ0, smul_of, smul_cons, real_smul, mul_one, smul_zero,
smul_empty, PauliMatrix.σ1, of_sub_of, sub_cons, head_cons, sub_zero, tail_cons, zero_sub,
sub_self, zero_empty, PauliMatrix.σ2, smul_neg, sub_neg_eq_add, PauliMatrix.σ3, det_fin_two_of]
ring_nf
simp only [Fin.isValue, ContrMod.toFin1d_eq_val, I_sq, mul_neg, mul_one, ContrMod.toSpace,
Function.comp_apply]
ring
end contrContrContractField
end Lorentz
end

View file

@ -0,0 +1,424 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Meta.Informal
import HepLean.Lorentz.Group.Basic
import Mathlib.RepresentationTheory.Rep
import Mathlib.Logic.Equiv.TransferInstance
import HepLean.Lorentz.PauliMatrices.SelfAdjoint
/-!
## Modules associated with Real Lorentz vectors
We define the modules underlying real Lorentz vectors.
These definitions are preludes to the definitions of
`Lorentz.contr` and `Lorentz.co`.
-/
namespace Lorentz
noncomputable section
open Matrix
open MatrixGroups
open Complex
/-- The module for contravariant (up-index) real Lorentz vectors. -/
structure ContrMod (d : ) where
/-- The underlying value as a vector `Fin 1 ⊕ Fin d → `. -/
val : Fin 1 ⊕ Fin d →
namespace ContrMod
variable {d : }
@[ext]
lemma ext {ψ ψ' : ContrMod d} (h : ψ.val = ψ'.val) : ψ = ψ' := by
cases ψ
cases ψ'
subst h
rfl
/-- The equivalence between `ContrModule` and `Fin 1 ⊕ Fin d → `. -/
def toFin1dFun : ContrMod d ≃ (Fin 1 ⊕ Fin d → ) where
toFun v := v.val
invFun f := ⟨f⟩
left_inv _ := rfl
right_inv _ := rfl
/-- The instance of `AddCommMonoid` on `ContrModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
instance : AddCommMonoid (ContrMod d) := Equiv.addCommMonoid toFin1dFun
/-- The instance of `AddCommGroup` on `ContrModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
instance : AddCommGroup (ContrMod d) := Equiv.addCommGroup toFin1dFun
/-- The instance of `Module` on `ContrModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
instance : Module (ContrMod d) := Equiv.module toFin1dFun
@[simp]
lemma val_add (ψ ψ' : ContrMod d) : (ψ + ψ').val = ψ.val + ψ'.val := rfl
@[simp]
lemma val_smul (r : ) (ψ : ContrMod d) : (r • ψ).val = r • ψ.val := rfl
/-- The linear equivalence between `ContrModule` and `(Fin 1 ⊕ Fin d → )`. -/
def toFin1dEquiv : ContrMod d ≃ₗ[] (Fin 1 ⊕ Fin d → ) :=
Equiv.linearEquiv toFin1dFun
/-- The underlying element of `Fin 1 ⊕ Fin d → ` of a element in `ContrModule` defined
through the linear equivalence `toFin1dEquiv`. -/
abbrev toFin1d (ψ : ContrMod d) := toFin1dEquiv ψ
lemma toFin1d_eq_val (ψ : ContrMod d) : ψ.toFin1d = ψ.val := by rfl
/-!
## The standard basis.
-/
/-- The standard basis of `ContrModule` indexed by `Fin 1 ⊕ Fin d`. -/
@[simps!]
def stdBasis : Basis (Fin 1 ⊕ Fin d) (ContrMod d) := Basis.ofEquivFun toFin1dEquiv
@[simp]
lemma stdBasis_toFin1dEquiv_apply_same (μ : Fin 1 ⊕ Fin d) :
toFin1dEquiv (stdBasis μ) μ = 1 := by
simp only [stdBasis, Basis.ofEquivFun, Basis.coe_ofRepr, LinearEquiv.trans_symm,
LinearEquiv.symm_symm, LinearEquiv.trans_apply, Finsupp.linearEquivFunOnFinite_single]
rw [@LinearEquiv.apply_symm_apply]
exact Pi.single_eq_same μ 1
@[simp]
lemma stdBasis_apply_same (μ : Fin 1 ⊕ Fin d) : (stdBasis μ).val μ = 1 :=
stdBasis_toFin1dEquiv_apply_same μ
lemma stdBasis_toFin1dEquiv_apply_ne {μ ν : Fin 1 ⊕ Fin d} (h : μ ≠ ν) :
toFin1dEquiv (stdBasis μ) ν = 0 := by
simp only [stdBasis, Basis.ofEquivFun, Basis.coe_ofRepr, LinearEquiv.trans_symm,
LinearEquiv.symm_symm, LinearEquiv.trans_apply, Finsupp.linearEquivFunOnFinite_single]
rw [@LinearEquiv.apply_symm_apply]
exact Pi.single_eq_of_ne' h 1
@[simp]
lemma stdBasis_inl_apply_inr (i : Fin d) : (stdBasis (Sum.inl 0)).val (Sum.inr i) = 0 := by
refine stdBasis_toFin1dEquiv_apply_ne ?_
simp
lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : (stdBasis μ).val ν = if μ = ν then 1 else 0 := by
simp only [stdBasis, Basis.coe_ofEquivFun]
change Pi.single μ 1 ν = _
simp only [Pi.single_apply]
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
/-- Decomposition of a contrvariant Lorentz vector into the standard basis. -/
lemma stdBasis_decomp (v : ContrMod d) : v = ∑ i, v.toFin1d i • stdBasis i := by
apply toFin1dEquiv.injective
simp only [map_sum, _root_.map_smul]
funext μ
rw [Fintype.sum_apply μ fun c => toFin1dEquiv v c • toFin1dEquiv (stdBasis c)]
change _ = ∑ x : Fin 1 ⊕ Fin d, toFin1dEquiv v x • (toFin1dEquiv (stdBasis x) μ)
rw [Finset.sum_eq_single_of_mem μ (Finset.mem_univ μ)]
· simp only [stdBasis_toFin1dEquiv_apply_same, smul_eq_mul, mul_one]
· intro b _ hbμ
rw [stdBasis_toFin1dEquiv_apply_ne hbμ]
simp only [smul_eq_mul, mul_zero]
/-!
## mulVec
-/
/-- Multiplication of a matrix with a vector in `ContrMod`. -/
abbrev mulVec (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (v : ContrMod d) :
ContrMod d := Matrix.toLinAlgEquiv stdBasis M v
/-- Multiplication of a matrix with a vector in `ContrMod`. -/
scoped[Lorentz] infixr:73 " *ᵥ " => ContrMod.mulVec
@[simp]
lemma mulVec_toFin1d (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (v : ContrMod d) :
(M *ᵥ v).toFin1d = M *ᵥ v.toFin1d := by
rfl
lemma mulVec_sub (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (v w : ContrMod d) :
M *ᵥ (v - w) = M *ᵥ v - M *ᵥ w := by
simp only [mulVec, LinearMap.map_sub]
lemma sub_mulVec (M N : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (v : ContrMod d) :
(M - N) *ᵥ v = M *ᵥ v - N *ᵥ v := by
simp only [mulVec, map_sub, LinearMap.sub_apply]
lemma mulVec_add (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (v w : ContrMod d) :
M *ᵥ (v + w) = M *ᵥ v + M *ᵥ w := by
simp only [mulVec, LinearMap.map_add]
@[simp]
lemma one_mulVec (v : ContrMod d) : (1 : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) *ᵥ v = v := by
simp only [mulVec, _root_.map_one, LinearMap.one_apply]
lemma mulVec_mulVec (M N : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (v : ContrMod d) :
M *ᵥ (N *ᵥ v) = (M * N) *ᵥ v := by
simp only [mulVec, _root_.map_mul, LinearMap.mul_apply]
/-!
## The norm
(Not the Minkowski norm, but the norm of a vector in `ContrModule d`.)
-/
/-- A `NormedAddCommGroup` structure on `ContrMod`. This is not an instance, as we
don't want it to be applied always. -/
def norm : NormedAddCommGroup (ContrMod d) where
norm v := ‖v.val‖₊
dist_self x := Pi.normedAddCommGroup.dist_self x.val
dist_triangle x y z := Pi.normedAddCommGroup.dist_triangle x.val y.val z.val
dist_comm x y := Pi.normedAddCommGroup.dist_comm x.val y.val
eq_of_dist_eq_zero {x y} := fun h => ext (MetricSpace.eq_of_dist_eq_zero h)
/-- The underlying space part of a `ContrMod` formed by removing the first element.
A better name for this might be `tail`. -/
def toSpace (v : ContrMod d) : EuclideanSpace (Fin d) := v.val ∘ Sum.inr
/-!
## The representation.
-/
/-- The representation of the Lorentz group acting on `ContrModule d`. -/
def rep : Representation (LorentzGroup d) (ContrMod d) where
toFun g := Matrix.toLinAlgEquiv stdBasis g
map_one' := (MulEquivClass.map_eq_one_iff (Matrix.toLinAlgEquiv stdBasis)).mpr rfl
map_mul' x y := by
simp only [lorentzGroupIsGroup_mul_coe, _root_.map_mul]
lemma rep_apply_toFin1d (g : LorentzGroup d) (ψ : ContrMod d) :
(rep g ψ).toFin1d = g.1 *ᵥ ψ.toFin1d := by
rfl
/-!
## To Self-Adjoint Matrix
-/
/-- The linear equivalence between the vector-space `ContrMod 3` and self-adjoint
`2×2`-complex matrices. -/
def toSelfAdjoint : ContrMod 3 ≃ₗ[] selfAdjoint (Matrix (Fin 2) (Fin 2) ) :=
toFin1dEquiv ≪≫ₗ (Finsupp.linearEquivFunOnFinite (Fin 1 ⊕ Fin 3)).symm ≪≫ₗ
PauliMatrix.σSAL.repr.symm
lemma toSelfAdjoint_apply (x : ContrMod 3) : toSelfAdjoint x =
x.toFin1d (Sum.inl 0) • ⟨PauliMatrix.σ0, PauliMatrix.σ0_selfAdjoint⟩
- x.toFin1d (Sum.inr 0) • ⟨PauliMatrix.σ1, PauliMatrix.σ1_selfAdjoint⟩
- x.toFin1d (Sum.inr 1) • ⟨PauliMatrix.σ2, PauliMatrix.σ2_selfAdjoint⟩
- x.toFin1d (Sum.inr 2) • ⟨PauliMatrix.σ3, PauliMatrix.σ3_selfAdjoint⟩ := by
simp only [toSelfAdjoint, 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.toFin1d i • PauliMatrix.σSAL' i) = _
simp only [PauliMatrix.σSAL', Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero,
Fin.isValue, Finset.sum_singleton, Fin.sum_univ_three]
apply Subtype.ext
simp only [Fin.isValue, AddSubgroup.coe_add, selfAdjoint.val_smul, smul_neg,
AddSubgroupClass.coe_sub]
simp only [neg_add, add_assoc, sub_eq_add_neg]
· simp_all only [Finset.coe_univ, Finsupp.supported_univ, Submodule.mem_top]
lemma toSelfAdjoint_apply_coe (x : ContrMod 3) : (toSelfAdjoint x).1 =
x.toFin1d (Sum.inl 0) • PauliMatrix.σ0
- x.toFin1d (Sum.inr 0) • PauliMatrix.σ1
- x.toFin1d (Sum.inr 1) • PauliMatrix.σ2
- x.toFin1d (Sum.inr 2) • PauliMatrix.σ3 := by
rw [toSelfAdjoint_apply]
rfl
lemma toSelfAdjoint_stdBasis (i : Fin 1 ⊕ Fin 3) :
toSelfAdjoint (stdBasis i) = PauliMatrix.σSAL i := by
rw [toSelfAdjoint_apply]
match i with
| Sum.inl 0 =>
simp only [stdBasis, Fin.isValue, Basis.coe_ofEquivFun, LinearEquiv.apply_symm_apply,
Pi.single_eq_same, one_smul, ne_eq, reduceCtorEq, not_false_eq_true, Pi.single_eq_of_ne,
zero_smul, sub_zero, PauliMatrix.σSAL, Basis.coe_mk, PauliMatrix.σSAL']
| Sum.inr 0 =>
simp only [stdBasis, Fin.isValue, Basis.coe_ofEquivFun, LinearEquiv.apply_symm_apply, ne_eq,
reduceCtorEq, not_false_eq_true, Pi.single_eq_of_ne, zero_smul, Pi.single_eq_same, one_smul,
zero_sub, Sum.inr.injEq, one_ne_zero, sub_zero, Fin.reduceEq, PauliMatrix.σSAL, Basis.coe_mk,
PauliMatrix.σSAL']
rfl
| Sum.inr 1 =>
simp only [stdBasis, Fin.isValue, Basis.coe_ofEquivFun, LinearEquiv.apply_symm_apply, ne_eq,
reduceCtorEq, not_false_eq_true, Pi.single_eq_of_ne, zero_smul, Sum.inr.injEq, zero_ne_one,
sub_self, Pi.single_eq_same, one_smul, zero_sub, Fin.reduceEq, sub_zero, PauliMatrix.σSAL,
Basis.coe_mk, PauliMatrix.σSAL']
rfl
| Sum.inr 2 =>
simp only [stdBasis, Fin.isValue, Basis.coe_ofEquivFun, LinearEquiv.apply_symm_apply, ne_eq,
reduceCtorEq, not_false_eq_true, Pi.single_eq_of_ne, zero_smul, Sum.inr.injEq, Fin.reduceEq,
sub_self, Pi.single_eq_same, one_smul, zero_sub, PauliMatrix.σSAL, Basis.coe_mk,
PauliMatrix.σSAL']
rfl
@[simp]
lemma toSelfAdjoint_symm_basis (i : Fin 1 ⊕ Fin 3) :
toSelfAdjoint.symm (PauliMatrix.σSAL i) = (stdBasis i) := by
refine (LinearEquiv.symm_apply_eq toSelfAdjoint).mpr ?_
rw [toSelfAdjoint_stdBasis]
/-!
## Topology
-/
instance : TopologicalSpace (ContrMod d) := TopologicalSpace.induced
ContrMod.toFin1dEquiv (Pi.topologicalSpace)
lemma toFin1dEquiv_isInducing : IsInducing (@ContrMod.toFin1dEquiv d) := by
exact { eq_induced := rfl }
lemma toFin1dEquiv_symm_isInducing : IsInducing ((@ContrMod.toFin1dEquiv d).symm) := by
let x := Equiv.toHomeomorphOfIsInducing (@ContrMod.toFin1dEquiv d).toEquiv
toFin1dEquiv_isInducing
exact Homeomorph.isInducing x.symm
end ContrMod
/-- The module for covariant (up-index) complex Lorentz vectors. -/
structure CoMod (d : ) where
/-- The underlying value as a vector `Fin 1 ⊕ Fin d → `. -/
val : Fin 1 ⊕ Fin d →
namespace CoMod
variable {d : }
@[ext]
lemma ext {ψ ψ' : CoMod d} (h : ψ.val = ψ'.val) : ψ = ψ' := by
cases ψ
cases ψ'
subst h
rfl
/-- The equivalence between `CoModule` and `Fin 1 ⊕ Fin d → `. -/
def toFin1dFun : CoMod d ≃ (Fin 1 ⊕ Fin d → ) where
toFun v := v.val
invFun f := ⟨f⟩
left_inv _ := rfl
right_inv _ := rfl
/-- The instance of `AddCommMonoid` on `CoModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
instance : AddCommMonoid (CoMod d) := Equiv.addCommMonoid toFin1dFun
/-- The instance of `AddCommGroup` on `CoModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
instance : AddCommGroup (CoMod d) := Equiv.addCommGroup toFin1dFun
/-- The instance of `Module` on `CoModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
instance : Module (CoMod d) := Equiv.module toFin1dFun
/-- The linear equivalence between `CoModule` and `(Fin 1 ⊕ Fin d → )`. -/
def toFin1dEquiv : CoMod d ≃ₗ[] (Fin 1 ⊕ Fin d → ) :=
Equiv.linearEquiv toFin1dFun
/-- The underlying element of `Fin 1 ⊕ Fin d → ` of a element in `CoModule` defined
through the linear equivalence `toFin1dEquiv`. -/
abbrev toFin1d (ψ : CoMod d) := toFin1dEquiv ψ
/-!
## The standard basis.
-/
/-- The standard basis of `CoModule` indexed by `Fin 1 ⊕ Fin d`. -/
@[simps!]
def stdBasis : Basis (Fin 1 ⊕ Fin d) (CoMod d) := Basis.ofEquivFun toFin1dEquiv
@[simp]
lemma stdBasis_toFin1dEquiv_apply_same (μ : Fin 1 ⊕ Fin d) :
toFin1dEquiv (stdBasis μ) μ = 1 := by
simp only [stdBasis, Basis.ofEquivFun, Basis.coe_ofRepr, LinearEquiv.trans_symm,
LinearEquiv.symm_symm, LinearEquiv.trans_apply, Finsupp.linearEquivFunOnFinite_single]
rw [@LinearEquiv.apply_symm_apply]
exact Pi.single_eq_same μ 1
@[simp]
lemma stdBasis_apply_same (μ : Fin 1 ⊕ Fin d) : (stdBasis μ).val μ = 1 :=
stdBasis_toFin1dEquiv_apply_same μ
lemma stdBasis_toFin1dEquiv_apply_ne {μ ν : Fin 1 ⊕ Fin d} (h : μ ≠ ν) :
toFin1dEquiv (stdBasis μ) ν = 0 := by
simp only [stdBasis, Basis.ofEquivFun, Basis.coe_ofRepr, LinearEquiv.trans_symm,
LinearEquiv.symm_symm, LinearEquiv.trans_apply, Finsupp.linearEquivFunOnFinite_single]
rw [@LinearEquiv.apply_symm_apply]
exact Pi.single_eq_of_ne' h 1
lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : (stdBasis μ).val ν = if μ = ν then 1 else 0 := by
simp only [stdBasis, Basis.coe_ofEquivFun]
change Pi.single μ 1 ν = _
simp only [Pi.single_apply]
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
/-- Decomposition of a covariant Lorentz vector into the standard basis. -/
lemma stdBasis_decomp (v : CoMod d) : v = ∑ i, v.toFin1d i • stdBasis i := by
apply toFin1dEquiv.injective
simp only [map_sum, _root_.map_smul]
funext μ
rw [Fintype.sum_apply μ fun c => toFin1dEquiv v c • toFin1dEquiv (stdBasis c)]
change _ = ∑ x : Fin 1 ⊕ Fin d, toFin1dEquiv v x • (toFin1dEquiv (stdBasis x) μ)
rw [Finset.sum_eq_single_of_mem μ (Finset.mem_univ μ)]
· simp only [stdBasis_toFin1dEquiv_apply_same, smul_eq_mul, mul_one]
· intro b _ hbμ
rw [stdBasis_toFin1dEquiv_apply_ne hbμ]
simp only [smul_eq_mul, mul_zero]
/-!
## mulVec
-/
/-- Multiplication of a matrix with a vector in `CoMod`. -/
abbrev mulVec (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (v : CoMod d) :
CoMod d := Matrix.toLinAlgEquiv stdBasis M v
/-- Multiplication of a matrix with a vector in `CoMod`. -/
scoped[Lorentz] infixr:73 " *ᵥ " => CoMod.mulVec
@[simp]
lemma mulVec_toFin1d (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (v : CoMod d) :
(M *ᵥ v).toFin1d = M *ᵥ v.toFin1d := by
rfl
/-!
## The representation
-/
/-- The representation of the Lorentz group acting on `CoModule d`. -/
def rep : Representation (LorentzGroup d) (CoMod d) where
toFun g := Matrix.toLinAlgEquiv stdBasis (LorentzGroup.transpose g⁻¹)
map_one' := by
simp only [inv_one, LorentzGroup.transpose_one, lorentzGroupIsGroup_one_coe, _root_.map_one]
map_mul' x y := by
simp only [_root_.mul_inv_rev, lorentzGroupIsGroup_inv, LorentzGroup.transpose_mul,
lorentzGroupIsGroup_mul_coe, _root_.map_mul]
end CoMod
end
end Lorentz

View file

@ -0,0 +1,331 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Lorentz.RealVector.Contraction
import Mathlib.GroupTheory.GroupAction.Blocks
/-!
# Lorentz vectors with norm one
-/
open TensorProduct
namespace Lorentz
namespace Contr
variable {d : }
/-- The set of Lorentz vectors with norm 1. -/
def NormOne (d : ) : Set (Contr d) := fun v => ⟪v, v⟫ₘ = (1 : )
noncomputable section
namespace NormOne
lemma mem_iff {v : Contr d} : v ∈ NormOne d ↔ ⟪v, v⟫ₘ = (1 : ) := by
rfl
@[simp]
lemma contr_self (v : NormOne d) : ⟪v.1, v.1⟫ₘ = (1 : ) := v.2
lemma mem_mulAction (g : LorentzGroup d) (v : Contr d) :
v ∈ NormOne d ↔ (Contr d).ρ g v ∈ NormOne d := by
rw [mem_iff, mem_iff, contrContrContractField.action_tmul]
instance : TopologicalSpace (NormOne d) := instTopologicalSpaceSubtype
variable (v w : NormOne d)
/-- The negative of a `NormOne` as a `NormOne`. -/
def neg : NormOne d := ⟨- v, by
rw [mem_iff]
simp only [Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, tmul_neg, neg_tmul, neg_neg]
exact v.2⟩
/-- The first column of a Lorentz matrix as a `NormOneLorentzVector`. -/
@[simps!]
def _root_.LorentzGroup.toNormOne (Λ : LorentzGroup d) : NormOne d :=
⟨(Contr d).ρ Λ (ContrMod.stdBasis (Sum.inl 0)), by
rw [mem_iff, contrContrContractField.action_tmul, contrContrContractField.stdBasis_inl]⟩
lemma _root_.LorentzGroup.toNormOne_inl (Λ : LorentzGroup d) :
(LorentzGroup.toNormOne Λ).val.val (Sum.inl 0) = Λ.1 (Sum.inl 0) (Sum.inl 0) := by
simp only [Fin.isValue, LorentzGroup.toNormOne_coe_val, Finsupp.single, one_ne_zero, ↓reduceIte,
Finsupp.coe_mk, Matrix.mulVec_single, mul_one]
lemma _root_.LorentzGroup.toNormOne_inr (Λ : LorentzGroup d) (i : Fin d) :
(LorentzGroup.toNormOne Λ).val.val (Sum.inr i) = Λ.1 (Sum.inr i) (Sum.inl 0) := by
simp only [LorentzGroup.toNormOne_coe_val, Finsupp.single, one_ne_zero, ↓reduceIte, Fin.isValue,
Finsupp.coe_mk, Matrix.mulVec_single, mul_one]
lemma _root_.LorentzGroup.inl_inl_mul (Λ Λ' : LorentzGroup d) : (Λ * Λ').1 (Sum.inl 0) (Sum.inl 0) =
⟪(LorentzGroup.toNormOne (LorentzGroup.transpose Λ)).1,
(Contr d).ρ LorentzGroup.parity (LorentzGroup.toNormOne Λ').1⟫ₘ := by
rw [contrContrContractField.right_parity]
simp only [Fin.isValue, lorentzGroupIsGroup_mul_coe, Matrix.mul_apply, Fintype.sum_sum_type,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton,
LorentzGroup.transpose, PiLp.inner_apply, Function.comp_apply,
RCLike.inner_apply, conj_trivial]
congr
· rw [LorentzGroup.toNormOne_inl]
rfl
· rw [LorentzGroup.toNormOne_inl]
· funext x
rw [LorentzGroup.toNormOne_inr, LorentzGroup.toNormOne_inr]
rfl
lemma inl_sq : v.val.val (Sum.inl 0) ^ 2 = 1 + ‖ContrMod.toSpace v.val‖ ^ 2 := by
rw [contrContrContractField.inl_sq_eq, v.2]
congr
rw [← real_inner_self_eq_norm_sq]
simp only [PiLp.inner_apply, RCLike.inner_apply, conj_trivial]
congr
funext x
exact pow_two ((v.val).val (Sum.inr x))
lemma one_le_abs_inl : 1 ≤ |v.val.val (Sum.inl 0)| := by
have h1 := contrContrContractField.le_inl_sq v.val
rw [v.2] at h1
exact (one_le_sq_iff_one_le_abs _).mp h1
lemma inl_le_neg_one_or_one_le_inl : v.val.val (Sum.inl 0) ≤ -1 1 ≤ v.val.val (Sum.inl 0) :=
le_abs'.mp (one_le_abs_inl v)
lemma norm_space_le_abs_inl : ‖v.1.toSpace‖ < |v.val.val (Sum.inl 0)| := by
rw [(abs_norm _).symm, ← @sq_lt_sq, inl_sq]
change ‖ContrMod.toSpace v.val‖ ^ 2 < 1 + ‖ContrMod.toSpace v.val‖ ^ 2
exact lt_one_add (‖(v.1).toSpace‖ ^ 2)
lemma norm_space_leq_abs_inl : ‖v.1.toSpace‖ ≤ |v.val.val (Sum.inl 0)| :=
le_of_lt (norm_space_le_abs_inl v)
lemma inl_abs_sub_space_norm :
0 ≤ |v.val.val (Sum.inl 0)| * |w.val.val (Sum.inl 0)| - ‖v.1.toSpace‖ * ‖w.1.toSpace‖ := by
apply sub_nonneg.mpr
apply mul_le_mul (norm_space_leq_abs_inl v) (norm_space_leq_abs_inl w) ?_ ?_
· exact norm_nonneg _
· exact abs_nonneg _
/-!
# Future pointing norm one Lorentz vectors
-/
/-- The future pointing Lorentz vectors with Norm one. -/
def FuturePointing (d : ) : Set (NormOne d) :=
fun x => 0 < x.val.val (Sum.inl 0)
namespace FuturePointing
lemma mem_iff : v ∈ FuturePointing d ↔ 0 < v.val.val (Sum.inl 0) := by
rfl
lemma mem_iff_inl_nonneg : v ∈ FuturePointing d ↔ 0 ≤ v.val.val (Sum.inl 0) := by
refine Iff.intro (fun h => le_of_lt h) (fun h => ?_)
rw [mem_iff]
rcases inl_le_neg_one_or_one_le_inl v with (h | h)
· linarith
· linarith
lemma mem_iff_inl_one_le_inl : v ∈ FuturePointing d ↔ 1 ≤ v.val.val (Sum.inl 0) := by
rw [mem_iff_inl_nonneg]
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rcases inl_le_neg_one_or_one_le_inl v with (h | h)
· linarith
· linarith
· linarith
lemma mem_iff_parity_mem : v ∈ FuturePointing d ↔ ⟨(Contr d).ρ LorentzGroup.parity v.1,
(NormOne.mem_mulAction _ _).mp v.2⟩ ∈ FuturePointing d := by
rw [mem_iff, mem_iff]
change _ ↔ 0 < (minkowskiMatrix.mulVec v.val.val) (Sum.inl 0)
simp only [Fin.isValue, minkowskiMatrix.mulVec_inl_0]
lemma not_mem_iff_inl_le_zero : v ∉ FuturePointing d ↔ v.val.val (Sum.inl 0) ≤ 0 := by
rw [mem_iff]
simp
lemma not_mem_iff_inl_lt_zero : v ∉ FuturePointing d ↔ v.val.val (Sum.inl 0) < 0 := by
rw [mem_iff_inl_nonneg]
simp
lemma not_mem_iff_inl_le_neg_one : v ∉ FuturePointing d ↔ v.val.val (Sum.inl 0) ≤ -1 := by
rw [not_mem_iff_inl_le_zero]
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rcases inl_le_neg_one_or_one_le_inl v with (h | h)
· linarith
· linarith
· linarith
lemma not_mem_iff_neg : v ∉ FuturePointing d ↔ neg v ∈ FuturePointing d := by
rw [not_mem_iff_inl_le_zero, mem_iff_inl_nonneg]
simp only [Fin.isValue, neg]
change (v).val.val (Sum.inl 0) ≤ 0 ↔ 0 ≤ - (v.val).val (Sum.inl 0)
simp
variable (f f' : FuturePointing d)
lemma inl_nonneg : 0 ≤ f.val.val.val (Sum.inl 0):= le_of_lt f.2
lemma abs_inl : |f.val.val.val (Sum.inl 0)| = f.val.val.val (Sum.inl 0) :=
abs_of_nonneg (inl_nonneg f)
lemma inl_eq_sqrt : f.val.val.val (Sum.inl 0) = √(1 + ‖f.1.1.toSpace‖ ^ 2) := by
symm
rw [Real.sqrt_eq_cases]
apply Or.inl
rw [← inl_sq, sq]
exact ⟨rfl, inl_nonneg f⟩
open InnerProductSpace
lemma metric_nonneg : 0 ≤ ⟪f.1.1, f'.1.1⟫ₘ := by
apply le_trans (inl_abs_sub_space_norm f f'.1)
rw [abs_inl f, abs_inl f']
exact contrContrContractField.ge_sub_norm f.1.1 f'.1.1
lemma one_add_metric_non_zero : 1 + ⟪f.1.1, f'.1.1⟫ₘ ≠ 0 := by
linarith [metric_nonneg f f']
variable {v w : NormOne d}
lemma metric_reflect_mem_mem (h : v ∈ FuturePointing d) (hw : w ∈ FuturePointing d) :
0 ≤ ⟪v.val, (Contr d).ρ LorentzGroup.parity w.1⟫ₘ :=
metric_nonneg ⟨v, h⟩ ⟨⟨(Contr d).ρ LorentzGroup.parity w.1,
(NormOne.mem_mulAction _ _).mp w.2⟩, (mem_iff_parity_mem w).mp hw⟩
lemma metric_reflect_not_mem_not_mem (h : v ∉ FuturePointing d) (hw : w ∉ FuturePointing d) :
0 ≤ ⟪v.val, (Contr d).ρ LorentzGroup.parity w.1⟫ₘ := by
have h1 := metric_reflect_mem_mem ((not_mem_iff_neg v).mp h) ((not_mem_iff_neg w).mp hw)
apply le_of_le_of_eq h1 ?_
simp [neg, neg_tmul, tmul_neg]
lemma metric_reflect_mem_not_mem (h : v ∈ FuturePointing d) (hw : w ∉ FuturePointing d) :
⟪v.val, (Contr d).ρ LorentzGroup.parity w.1⟫ₘ ≤ 0 := by
rw [show (0 : ) = - 0 from zero_eq_neg.mpr rfl, le_neg]
have h1 := metric_reflect_mem_mem h ((not_mem_iff_neg w).mp hw)
apply le_of_le_of_eq h1 ?_
simp [neg, neg_tmul, tmul_neg]
lemma metric_reflect_not_mem_mem (h : v ∉ FuturePointing d) (hw : w ∈ FuturePointing d) :
⟪v.val, (Contr d).ρ LorentzGroup.parity w.1⟫ₘ ≤ 0 := by
rw [show (0 : ) = - 0 from zero_eq_neg.mpr rfl, le_neg]
have h1 := metric_reflect_mem_mem ((not_mem_iff_neg v).mp h) hw
apply le_of_le_of_eq h1 ?_
simp [neg, neg_tmul, tmul_neg]
end FuturePointing
end NormOne
namespace NormOne
namespace FuturePointing
/-!
## Topology
-/
/-- The `FuturePointing d` which has all space components zero. -/
@[simps!]
noncomputable def timeVecNormOneFuture : FuturePointing d := ⟨⟨ContrMod.stdBasis (Sum.inl 0), by
rw [NormOne.mem_iff, contrContrContractField.on_basis]
rfl⟩, by
rw [mem_iff]
simp⟩
/-- A continuous path from `timeVecNormOneFuture` to any other. -/
noncomputable def pathFromTime (u : FuturePointing d) : Path timeVecNormOneFuture u where
toFun t := ⟨
⟨{val := fun i => match i with
| Sum.inl 0 => √(1 + t ^ 2 * ‖u.1.1.toSpace‖ ^ 2)
| Sum.inr i => t * u.1.1.toSpace i},
by
rw [NormOne.mem_iff, contrContrContractField.as_sum_toSpace]
simp only [ContrMod.toSpace, Function.comp_apply, PiLp.inner_apply, RCLike.inner_apply,
map_mul, conj_trivial]
rw [Real.mul_self_sqrt, ← @real_inner_self_eq_norm_sq, @PiLp.inner_apply]
· simp only [Function.comp_apply, RCLike.inner_apply, conj_trivial]
refine Eq.symm (eq_sub_of_add_eq (congrArg (HAdd.hAdd _) ?_))
rw [Finset.mul_sum]
apply Finset.sum_congr rfl
intro i _
ring_nf
· exact Right.add_nonneg (zero_le_one' ) $ mul_nonneg (sq_nonneg _) (sq_nonneg _)⟩,
by
simp only [ContrMod.toSpace, Function.comp_apply, mem_iff_inl_nonneg, Real.sqrt_pos]
exact Real.sqrt_nonneg _⟩
continuous_toFun := by
refine Continuous.subtype_mk ?_ _
refine Continuous.subtype_mk ?_ _
refine continuous_contr _ ?_
apply (continuous_pi_iff).mpr
intro i
match i with
| Sum.inl 0 =>
continuity
| Sum.inr i =>
continuity
source' := by
ext
apply ContrMod.ext
funext i
match i with
| Sum.inl 0 =>
simp only [Set.Icc.coe_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow,
zero_mul, add_zero, Real.sqrt_one, timeVecNormOneFuture, Fin.isValue,
ContrMod.stdBasis_apply_same]
| Sum.inr i =>
simp only [Set.Icc.coe_zero, zero_mul, timeVecNormOneFuture, Fin.isValue,
ContrMod.stdBasis_inl_apply_inr]
target' := by
ext
apply ContrMod.ext
funext i
match i with
| Sum.inl 0 =>
simp only [Set.Icc.coe_one, one_pow, one_mul, Fin.isValue]
exact (inl_eq_sqrt u).symm
| Sum.inr i =>
simp only [Set.Icc.coe_one, one_pow, one_mul, Fin.isValue]
rfl
lemma isPathConnected : IsPathConnected (@Set.univ (FuturePointing d)) := by
use timeVecNormOneFuture
apply And.intro trivial ?_
intro y a
use pathFromTime y
exact fun _ => a
lemma metric_continuous (u : Contr d) :
Continuous (fun (a : FuturePointing d) => ⟪u, a.1.1⟫ₘ) := by
simp only [contrContrContractField.as_sum_toSpace]
refine Continuous.add ?_ ?_
· refine Continuous.comp' (continuous_mul_left _) $ Continuous.comp'
(continuous_apply (Sum.inl 0))
(Continuous.comp' ?_ ?_)
· exact continuous_iff_le_induced.mpr fun U a => a
· exact Continuous.comp' continuous_subtype_val continuous_subtype_val
· refine Continuous.comp' continuous_neg $ Continuous.inner
(Continuous.comp' (?_) continuous_const)
(Continuous.comp' (?_) (Continuous.comp'
continuous_subtype_val continuous_subtype_val))
· apply contr_continuous
exact Pi.continuous_precomp Sum.inr
· apply contr_continuous
exact Pi.continuous_precomp Sum.inr
end FuturePointing
end NormOne
end
end Contr
end Lorentz

View file

@ -0,0 +1,192 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Lorentz.Group.Basic
import HepLean.Lorentz.RealVector.Basic
import Mathlib.RepresentationTheory.Basic
import HepLean.Lorentz.Group.Restricted
import HepLean.Lorentz.PauliMatrices.SelfAdjoint
import HepLean.Meta.Informal
/-!
# The group SL(2, ) and it's relation to the Lorentz group
The aim of this file is to give the relationship between `SL(2, )` and the Lorentz group.
-/
namespace Lorentz
open Matrix
open MatrixGroups
open Complex
namespace SL2C
noncomputable section
/-!
## Some basic properties about SL(2, )
Possibly to be moved to mathlib at some point.
-/
lemma inverse_coe (M : SL(2, )) : M.1⁻¹ = (M⁻¹).1 := by
apply Matrix.inv_inj
simp only [SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq, one_ne_zero, not_false_eq_true,
nonsing_inv_nonsing_inv, SpecialLinearGroup.coe_inv]
have h1 : IsUnit M.1.det := by
simp
rw [Matrix.inv_adjugate M.1 h1]
· simp
· simp
lemma transpose_coe (M : SL(2, )) : M.1ᵀ = (M.transpose).1 := rfl
/-!
## Representation of SL(2, ) on spacetime
Through the correspondence between spacetime and self-adjoint matrices,
we can define a representation a representation of `SL(2, )` on spacetime.
-/
/-- Given an element `M ∈ SL(2, )` the linear map from `selfAdjoint (Matrix (Fin 2) (Fin 2) )` to
itself defined by `A ↦ M * A * Mᴴ`. -/
@[simps!]
def toLinearMapSelfAdjointMatrix (M : SL(2, )) :
selfAdjoint (Matrix (Fin 2) (Fin 2) ) →ₗ[] selfAdjoint (Matrix (Fin 2) (Fin 2) ) where
toFun A := ⟨M.1 * A.1 * Matrix.conjTranspose M,
by
noncomm_ring [selfAdjoint.mem_iff, star_eq_conjTranspose,
conjTranspose_mul, conjTranspose_conjTranspose,
(star_eq_conjTranspose A.1).symm.trans $ selfAdjoint.mem_iff.mp A.2]⟩
map_add' A B := by
simp only [AddSubgroup.coe_add, AddMemClass.mk_add_mk, Subtype.mk.injEq]
noncomm_ring [AddSubmonoid.coe_add, AddSubgroup.coe_toAddSubmonoid, AddSubmonoid.mk_add_mk,
Subtype.mk.injEq]
map_smul' r A := by
noncomm_ring [selfAdjoint.val_smul, Algebra.mul_smul_comm, Algebra.smul_mul_assoc,
RingHom.id_apply]
lemma toLinearMapSelfAdjointMatrix_det (M : SL(2, )) (A : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
det ((toLinearMapSelfAdjointMatrix M) A).1 = det A.1 := by
simp only [LinearMap.coe_mk, AddHom.coe_mk, toLinearMapSelfAdjointMatrix, det_mul,
selfAdjoint.mem_iff, det_conjTranspose, det_mul, det_one, RingHom.id_apply]
simp only [SpecialLinearGroup.det_coe, one_mul, star_one, mul_one]
/-- The monoid homomorphisms from `SL(2, )` to matrices indexed by `Fin 1 ⊕ Fin 3`
formed by the action `M A Mᴴ`. -/
def toMatrix : SL(2, ) →* Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) where
toFun M := LinearMap.toMatrix PauliMatrix.σSAL PauliMatrix.σSAL (toLinearMapSelfAdjointMatrix M)
map_one' := by
simp only [toLinearMapSelfAdjointMatrix, SpecialLinearGroup.coe_one, one_mul, conjTranspose_one,
mul_one, Subtype.coe_eta]
erw [LinearMap.toMatrix_one]
map_mul' M N := by
simp only
rw [← LinearMap.toMatrix_mul]
apply congrArg
ext1 x
simp only [toLinearMapSelfAdjointMatrix, SpecialLinearGroup.coe_mul, conjTranspose_mul,
LinearMap.coe_mk, AddHom.coe_mk, LinearMap.mul_apply, Subtype.mk.injEq]
noncomm_ring
open Lorentz in
lemma toMatrix_apply_contrMod (M : SL(2, )) (v : ContrMod 3) :
(toMatrix M) *ᵥ v = ContrMod.toSelfAdjoint.symm
((toLinearMapSelfAdjointMatrix M) (ContrMod.toSelfAdjoint v)) := by
simp only [ContrMod.mulVec, toMatrix, MonoidHom.coe_mk, OneHom.coe_mk]
obtain ⟨a, ha⟩ := ContrMod.toSelfAdjoint.symm.surjective v
subst ha
rw [LinearEquiv.apply_symm_apply]
simp only [ContrMod.toSelfAdjoint, LinearEquiv.trans_symm, LinearEquiv.symm_symm,
LinearEquiv.trans_apply]
change ContrMod.toFin1dEquiv.symm
((((LinearMap.toMatrix PauliMatrix.σSAL PauliMatrix.σSAL) (toLinearMapSelfAdjointMatrix M)))
*ᵥ (((Finsupp.linearEquivFunOnFinite (Fin 1 ⊕ Fin 3)) (PauliMatrix.σSAL.repr a)))) = _
apply congrArg
erw [LinearMap.toMatrix_mulVec_repr]
rfl
lemma toMatrix_mem_lorentzGroup (M : SL(2, )) : toMatrix M ∈ LorentzGroup 3 := by
rw [LorentzGroup.mem_iff_norm]
intro x
apply ofReal_injective
rw [Lorentz.contrContrContractField.same_eq_det_toSelfAdjoint]
rw [toMatrix_apply_contrMod]
rw [LinearEquiv.apply_symm_apply]
rw [toLinearMapSelfAdjointMatrix_det]
rw [Lorentz.contrContrContractField.same_eq_det_toSelfAdjoint]
/-- The group homomorphism from `SL(2, )` to the Lorentz group `𝓛`. -/
@[simps!]
def toLorentzGroup : SL(2, ) →* LorentzGroup 3 where
toFun M := ⟨toMatrix M, toMatrix_mem_lorentzGroup M⟩
map_one' := by
simp only [_root_.map_one]
rfl
map_mul' M N := by
ext1
simp only [_root_.map_mul, lorentzGroupIsGroup_mul_coe]
lemma toLorentzGroup_eq_σSAL (M : SL(2, )) :
toLorentzGroup M = LinearMap.toMatrix
PauliMatrix.σSAL PauliMatrix.σSAL (toLinearMapSelfAdjointMatrix M) := by
rfl
lemma toLinearMapSelfAdjointMatrix_basis (i : Fin 1 ⊕ Fin 3) :
toLinearMapSelfAdjointMatrix M (PauliMatrix.σSAL i) =
∑ j, (toLorentzGroup M).1 j i •
PauliMatrix.σSAL j := by
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
((toLinearMapSelfAdjointMatrix M) (PauliMatrix.σSAL i)))]
rfl
lemma toLinearMapSelfAdjointMatrix_σSA (i : Fin 1 ⊕ Fin 3) :
toLinearMapSelfAdjointMatrix M (PauliMatrix.σSA i) =
∑ j, (toLorentzGroup M⁻¹).1 i j • PauliMatrix.σSA j := by
have h1 : (toLorentzGroup M⁻¹).1 = minkowskiMatrix.dual (toLorentzGroup M).1 := by
simp
simp only [h1]
rw [PauliMatrix.σSA_minkowskiMetric_σSAL, _root_.map_smul]
rw [toLinearMapSelfAdjointMatrix_basis]
rw [Finset.smul_sum]
apply congrArg
funext j
rw [smul_smul, PauliMatrix.σSA_minkowskiMetric_σSAL, smul_smul]
apply congrFun
apply congrArg
exact Eq.symm (minkowskiMatrix.dual_apply_minkowskiMatrix ((toLorentzGroup M).1) i j)
/-!
## Homomorphism to the restricted Lorentz group
The homomorphism `toLorentzGroup` restricts to a homomorphism to the restricted Lorentz group.
In this section we will define this homomorphism.
-/
informal_lemma toLorentzGroup_det_one where
math :≈ "The determinant of the image of `SL(2, )` in the Lorentz group is one."
deps :≈ [``toLorentzGroup]
informal_lemma toLorentzGroup_inl_inl_nonneg where
math :≈ "The time coponent of the image of `SL(2, )` in the Lorentz group is non-negative."
deps :≈ [``toLorentzGroup]
informal_lemma toRestrictedLorentzGroup where
math :≈ "The homomorphism from `SL(2, )` to the restricted Lorentz group."
deps :≈ [``toLorentzGroup, ``toLorentzGroup_det_one, ``toLorentzGroup_inl_inl_nonneg,
``LorentzGroup.Restricted]
/-! TODO: Define homomorphism from `SL(2, )` to the restricted Lorentz group. -/
end
end SL2C
end Lorentz

View file

@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Meta.Informal
import HepLean.SpaceTime.SL2C.Basic
import HepLean.Lorentz.SL2C.Basic
import Mathlib.RepresentationTheory.Rep
import HepLean.SpaceTime.WeylFermion.Modules
import HepLean.Lorentz.Weyl.Modules
import Mathlib.Logic.Equiv.TransferInstance
/-!
@ -207,7 +207,7 @@ def leftHandedToAlt : leftHanded ⟶ altLeftHanded where
change AltLeftHandedModule.toFin2Equiv.symm (!![0, 1; -1, 0] *ᵥ M.1 *ᵥ ψ.val) =
AltLeftHandedModule.toFin2Equiv.symm ((M.1⁻¹)ᵀ *ᵥ !![0, 1; -1, 0] *ᵥ ψ.val)
apply congrArg
rw [mulVec_mulVec, mulVec_mulVec, SpaceTime.SL2C.inverse_coe, eta_fin_two M.1]
rw [mulVec_mulVec, mulVec_mulVec, Lorentz.SL2C.inverse_coe, eta_fin_two M.1]
refine congrFun (congrArg _ ?_) _
rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two,
Matrix.mul_fin_two, eta_fin_two !![M.1 1 1, -M.1 0 1; -M.1 1 0, M.1 0 0]ᵀ]
@ -238,7 +238,7 @@ def leftHandedAltTo : altLeftHanded ⟶ leftHanded where
refine LinearMap.ext (fun ψ => ?_)
change LeftHandedModule.toFin2Equiv.symm (!![0, -1; 1, 0] *ᵥ (M.1⁻¹)ᵀ *ᵥ ψ.val) =
LeftHandedModule.toFin2Equiv.symm (M.1 *ᵥ !![0, -1; 1, 0] *ᵥ ψ.val)
rw [EquivLike.apply_eq_iff_eq, mulVec_mulVec, mulVec_mulVec, SpaceTime.SL2C.inverse_coe,
rw [EquivLike.apply_eq_iff_eq, mulVec_mulVec, mulVec_mulVec, Lorentz.SL2C.inverse_coe,
eta_fin_two M.1]
refine congrFun (congrArg _ ?_) _
rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two,

View file

@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.WeylFermion.Basic
import HepLean.Lorentz.Weyl.Basic
/-!
# Contraction of Weyl fermions

View file

@ -3,11 +3,11 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.WeylFermion.Basic
import HepLean.SpaceTime.WeylFermion.Contraction
import HepLean.Lorentz.Weyl.Basic
import HepLean.Lorentz.Weyl.Contraction
import Mathlib.LinearAlgebra.TensorProduct.Matrix
import HepLean.SpaceTime.WeylFermion.Two
import HepLean.SpaceTime.WeylFermion.Unit
import HepLean.Lorentz.Weyl.Two
import HepLean.Lorentz.Weyl.Unit
/-!
# Metrics of Weyl fermions
@ -32,7 +32,7 @@ def metricRaw : Matrix (Fin 2) (Fin 2) := !![0, 1; -1, 0]
lemma comm_metricRaw (M : SL(2,)) : M.1 * metricRaw = metricRaw * (M.1⁻¹)ᵀ := by
rw [metricRaw]
rw [SpaceTime.SL2C.inverse_coe, eta_fin_two M.1]
rw [Lorentz.SL2C.inverse_coe, eta_fin_two M.1]
rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two,
Matrix.mul_fin_two, eta_fin_two !![M.1 1 1, -M.1 0 1; -M.1 1 0, M.1 0 0]ᵀ]
simp only [Fin.isValue, mul_zero, mul_neg, mul_one, zero_add, add_zero, transpose_apply, of_apply,
@ -42,7 +42,7 @@ lemma comm_metricRaw (M : SL(2,)) : M.1 * metricRaw = metricRaw * (M.1⁻¹)
lemma metricRaw_comm (M : SL(2,)) : metricRaw * M.1 = (M.1⁻¹)ᵀ * metricRaw := by
rw [metricRaw]
rw [SpaceTime.SL2C.inverse_coe, eta_fin_two M.1]
rw [Lorentz.SL2C.inverse_coe, eta_fin_two M.1]
rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two,
Matrix.mul_fin_two, eta_fin_two !![M.1 1 1, -M.1 0 1; -M.1 1 0, M.1 0 0]ᵀ]
simp only [Fin.isValue, zero_mul, one_mul, zero_add, neg_mul, add_zero, transpose_apply, of_apply,
@ -53,7 +53,7 @@ lemma metricRaw_comm (M : SL(2,)) : metricRaw * M.1 = (M.1⁻¹)ᵀ * metricR
lemma star_comm_metricRaw (M : SL(2,)) : M.1.map star * metricRaw = metricRaw * ((M.1)⁻¹)ᴴ := by
rw [metricRaw]
rw [SpaceTime.SL2C.inverse_coe, eta_fin_two M.1]
rw [Lorentz.SL2C.inverse_coe, eta_fin_two M.1]
rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two,
eta_fin_two !![M.1 1 1, -M.1 0 1; -M.1 1 0, M.1 0 0]ᴴ]
rw [eta_fin_two (!![M.1 0 0, M.1 0 1; M.1 1 0, M.1 1 1].map star)]
@ -61,7 +61,7 @@ lemma star_comm_metricRaw (M : SL(2,)) : M.1.map star * metricRaw = metricRaw
lemma metricRaw_comm_star (M : SL(2,)) : metricRaw * M.1.map star = ((M.1)⁻¹)ᴴ * metricRaw := by
rw [metricRaw]
rw [SpaceTime.SL2C.inverse_coe, eta_fin_two M.1]
rw [Lorentz.SL2C.inverse_coe, eta_fin_two M.1]
rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two,
eta_fin_two !![M.1 1 1, -M.1 0 1; -M.1 1 0, M.1 0 0]ᴴ]
rw [eta_fin_two (!![M.1 0 0, M.1 0 1; M.1 1 0, M.1 1 1].map star)]

View file

@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Meta.Informal
import HepLean.SpaceTime.SL2C.Basic
import HepLean.Lorentz.SL2C.Basic
import Mathlib.RepresentationTheory.Rep
import Mathlib.Logic.Equiv.TransferInstance
/-!

View file

@ -3,8 +3,8 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.WeylFermion.Basic
import HepLean.SpaceTime.WeylFermion.Contraction
import HepLean.Lorentz.Weyl.Basic
import HepLean.Lorentz.Weyl.Contraction
import Mathlib.LinearAlgebra.TensorProduct.Matrix
/-!
@ -713,16 +713,16 @@ lemma leftRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ) (M : SL(2,))
rw [← h1]
simp
open SpaceTime
open Lorentz
lemma altLeftAltRightToMatrix_ρ_symm_selfAdjoint (v : Matrix (Fin 2) (Fin 2) )
(hv : IsSelfAdjoint v) (M : SL(2,)) :
TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M) (altLeftAltRightToMatrix.symm v) =
altLeftAltRightToMatrix.symm
(SL2C.repSelfAdjointMatrix (M.transpose⁻¹) ⟨v, hv⟩) := by
(SL2C.toLinearMapSelfAdjointMatrix (M.transpose⁻¹) ⟨v, hv⟩) := by
rw [altLeftAltRightToMatrix_ρ_symm]
apply congrArg
simp only [SL2C.repSelfAdjointMatrix, MonoidHom.coe_mk, OneHom.coe_mk,
simp only [MonoidHom.coe_mk, OneHom.coe_mk,
SL2C.toLinearMapSelfAdjointMatrix_apply_coe, SpecialLinearGroup.coe_inv,
SpecialLinearGroup.coe_transpose]
congr
@ -738,7 +738,7 @@ 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
(SL2C.repSelfAdjointMatrix M ⟨v, hv⟩) := by
(SL2C.toLinearMapSelfAdjointMatrix M ⟨v, hv⟩) := by
rw [leftRightToMatrix_ρ_symm]
rfl

View file

@ -3,10 +3,10 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.WeylFermion.Basic
import HepLean.SpaceTime.WeylFermion.Contraction
import HepLean.Lorentz.Weyl.Basic
import HepLean.Lorentz.Weyl.Contraction
import Mathlib.LinearAlgebra.TensorProduct.Matrix
import HepLean.SpaceTime.WeylFermion.Two
import HepLean.Lorentz.Weyl.Two
/-!
# Units of Weyl fermions

View file

@ -1,167 +0,0 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzGroup.Proper
import Mathlib.Topology.Constructions
import HepLean.SpaceTime.LorentzVector.NormOne
/-!
# Boosts
This file defines those Lorentz which are boosts.
We first define generalised boosts, which are restricted lorentz transformations taking
a four velocity `u` to a four velocity `v`.
A boost is the special case of a generalised boost when `u = stdBasis 0`.
## References
- The main argument follows: Guillem Cobos, The Lorentz Group, 2015:
https://diposit.ub.edu/dspace/bitstream/2445/68763/2/memoria.pdf
-/
/-! TODO: Show that generalised boosts are in the restricted Lorentz group. -/
noncomputable section
namespace LorentzGroup
open NormOneLorentzVector
open minkowskiMetric
variable {d : }
/-- An auxillary linear map used in the definition of a generalised boost. -/
def genBoostAux₁ (u v : FuturePointing d) : LorentzVector d →ₗ[] LorentzVector d where
toFun x := (2 * ⟪x, u⟫ₘ) • v.1.1
map_add' x y := by
simp only [map_add, LinearMap.add_apply]
rw [mul_add, add_smul]
map_smul' c x := by
simp only [LinearMapClass.map_smul, LinearMap.smul_apply, smul_eq_mul,
RingHom.id_apply]
rw [← mul_assoc, mul_comm 2 c, mul_assoc, mul_smul]
/-- An auxillary linear map used in the definition of a genearlised boost. -/
def genBoostAux₂ (u v : FuturePointing d) : LorentzVector d →ₗ[] LorentzVector d where
toFun x := - (⟪x, u.1.1 + v⟫ₘ / (1 + ⟪u.1.1, v⟫ₘ)) • (u.1.1 + v)
map_add' x y := by
simp only
rw [← add_smul]
apply congrFun (congrArg _ _)
field_simp
apply congrFun (congrArg _ _)
ring
map_smul' c x := by
simp only
rw [map_smul, show (c • minkowskiMetric x) (↑u + ↑v) = c * minkowskiMetric x (u+v) from rfl,
mul_div_assoc, neg_mul_eq_mul_neg, smul_smul]
rfl
/-- An generalised boost. This is a Lorentz transformation which takes the four velocity `u`
to `v`. -/
def genBoost (u v : FuturePointing d) : LorentzVector d →ₗ[] LorentzVector d :=
LinearMap.id + genBoostAux₁ u v + genBoostAux₂ u v
namespace genBoost
/--
This lemma states that for a given four-velocity `u`, the general boost
transformation `genBoost u u` is equal to the identity linear map `LinearMap.id`.
-/
lemma self (u : FuturePointing d) : genBoost u u = LinearMap.id := by
ext x
simp only [genBoost, LinearMap.add_apply, LinearMap.id_coe, id_eq]
rw [add_assoc, add_right_eq_self, add_eq_zero_iff_eq_neg, genBoostAux₁, genBoostAux₂]
simp only [LinearMap.coe_mk, AddHom.coe_mk, map_add, smul_add, neg_smul, neg_add_rev, neg_neg]
rw [← add_smul]
apply congrFun (congrArg _ _)
rw [u.1.2]
ring
/-- A generalised boost as a matrix. -/
def toMatrix (u v : FuturePointing d) : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) :=
LinearMap.toMatrix LorentzVector.stdBasis LorentzVector.stdBasis (genBoost u v)
lemma toMatrix_mulVec (u v : FuturePointing d) (x : LorentzVector d) :
(toMatrix u v).mulVec x = genBoost u v x :=
LinearMap.toMatrix_mulVec_repr LorentzVector.stdBasis LorentzVector.stdBasis (genBoost u v) x
open minkowskiMatrix LorentzVector in
@[simp]
lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) :
(toMatrix u v) μ ν = η μ μ * (⟪e μ, e ν⟫ₘ + 2 * ⟪e ν, u⟫ₘ * ⟪e μ, v⟫ₘ
- ⟪e μ, u + v⟫ₘ * ⟪e ν, u + v⟫ₘ / (1 + ⟪u, v.1.1⟫ₘ)) := by
rw [matrix_apply_stdBasis (toMatrix u v) μ ν, toMatrix_mulVec]
simp only [genBoost, genBoostAux₁, genBoostAux₂, map_add, smul_add, neg_smul, LinearMap.add_apply,
LinearMap.id_apply, LinearMap.coe_mk, AddHom.coe_mk, basis_left, map_smul, smul_eq_mul, map_neg,
mul_eq_mul_left_iff]
ring_nf
exact (true_or (η μ μ = 0)).mpr trivial
open minkowskiMatrix LorentzVector in
lemma toMatrix_continuous (u : FuturePointing d) : Continuous (toMatrix u) := by
refine continuous_matrix ?_
intro i j
simp only [toMatrix_apply]
refine (continuous_mul_left (η i i)).comp' ?_
refine Continuous.sub ?_ ?_
· refine .comp' (continuous_add_left (minkowskiMetric (e i) (e j))) ?_
refine .comp' (continuous_mul_left (2 * ⟪e j, u⟫ₘ)) ?_
exact FuturePointing.metric_continuous _
refine .mul ?_ ?_
· refine .mul ?_ ?_
· simp only [(minkowskiMetric _).map_add]
refine .comp' ?_ ?_
· exact continuous_add_left ((minkowskiMetric (stdBasis i)) ↑u)
· exact FuturePointing.metric_continuous _
· simp only [(minkowskiMetric _).map_add]
refine .comp' ?_ ?_
· exact continuous_add_left _
· exact FuturePointing.metric_continuous _
· refine .inv₀ ?_ ?_
· refine .comp' (continuous_add_left 1) (FuturePointing.metric_continuous _)
exact fun x => FuturePointing.one_add_metric_non_zero u x
lemma toMatrix_in_lorentzGroup (u v : FuturePointing d) : (toMatrix u v) ∈ LorentzGroup d:= by
intro x y
rw [toMatrix_mulVec, toMatrix_mulVec, genBoost, genBoostAux₁, genBoostAux₂]
have h1 : (1 + (minkowskiMetric ↑u) ↑v.1.1) ≠ 0 := FuturePointing.one_add_metric_non_zero u v
simp only [map_add, smul_add, neg_smul, LinearMap.add_apply, LinearMap.id_coe,
id_eq, LinearMap.coe_mk, AddHom.coe_mk, LinearMapClass.map_smul, map_neg, LinearMap.smul_apply,
smul_eq_mul, LinearMap.neg_apply]
field_simp
rw [u.1.2, v.1.2, minkowskiMetric.symm v.1.1 u, minkowskiMetric.symm u y,
minkowskiMetric.symm v y]
ring
/-- A generalised boost as an element of the Lorentz Group. -/
def toLorentz (u v : FuturePointing d) : LorentzGroup d :=
⟨toMatrix u v, toMatrix_in_lorentzGroup u v⟩
lemma toLorentz_continuous (u : FuturePointing d) : Continuous (toLorentz u) := by
refine Continuous.subtype_mk ?_ (fun x => toMatrix_in_lorentzGroup u x)
exact toMatrix_continuous u
lemma toLorentz_joined_to_1 (u v : FuturePointing d) : Joined 1 (toLorentz u v) := by
obtain ⟨f, _⟩ := FuturePointing.isPathConnected.joinedIn u trivial v trivial
use ContinuousMap.comp ⟨toLorentz u, toLorentz_continuous u⟩ f
· simp only [ContinuousMap.toFun_eq_coe, ContinuousMap.comp_apply, ContinuousMap.coe_coe,
Path.source, ContinuousMap.coe_mk]
rw [@Subtype.ext_iff, toLorentz]
simp [toMatrix, self u]
· simp
lemma toLorentz_in_connected_component_1 (u v : FuturePointing d) :
toLorentz u v ∈ connectedComponent 1 :=
pathComponent_subset_component _ (toLorentz_joined_to_1 u v)
lemma isProper (u v : FuturePointing d) : IsProper (toLorentz u v) :=
(isProper_on_connected_component (toLorentz_in_connected_component_1 u v)).mp id_IsProper
end genBoost
end LorentzGroup
end

View file

@ -1,107 +0,0 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.MinkowskiMetric
import HepLean.SpaceTime.PauliMatrices.SelfAdjoint
import Mathlib.LinearAlgebra.Matrix.SpecialLinearGroup
import Mathlib.Tactic.Polyrith
/-!
# Lorentz vector as a self-adjoint matrix
There is a linear equivalence between the vector space of space-time points
and the vector space of 2×2-complex self-adjoint matrices.
In this file we define this linear equivalence in `toSelfAdjointMatrix`.
-/
/-! TODO: Generalize rep of Lorentz vector as a self-adjoint matrix to arbitrary dimension. -/
namespace SpaceTime
open Matrix
open MatrixGroups
open Complex
noncomputable section
/-- The linear equivalence between the vector-space `spaceTime` and self-adjoint
`2×2`-complex matrices. -/
def toSelfAdjointMatrix : LorentzVector 3 ≃ₗ[] selfAdjoint (Matrix (Fin 2) (Fin 2) ) :=
(Finsupp.linearEquivFunOnFinite (Fin 1 ⊕ Fin 3)).symm ≪≫ₗ PauliMatrix.σSAL.repr.symm
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
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) = _
simp only [PauliMatrix.σSAL', Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero,
Fin.isValue, Finset.sum_singleton, Fin.sum_univ_three]
apply Subtype.ext
simp only [Fin.isValue, AddSubgroup.coe_add, selfAdjoint.val_smul, smul_neg,
AddSubgroupClass.coe_sub]
simp only [neg_add, add_assoc, sub_eq_add_neg]
· simp_all only [Finset.coe_univ, Finsupp.supported_univ, Submodule.mem_top]
lemma toSelfAdjointMatrix_apply_coe (x : LorentzVector 3) : (toSelfAdjointMatrix x).1 =
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
rw [toSelfAdjointMatrix_apply]
rfl
lemma toSelfAdjointMatrix_stdBasis (i : Fin 1 ⊕ Fin 3) :
toSelfAdjointMatrix (LorentzVector.stdBasis i) = PauliMatrix.σSAL i := by
rw [toSelfAdjointMatrix_apply]
match i with
| Sum.inl 0 =>
simp only [LorentzVector.stdBasis, Fin.isValue]
erw [Pi.basisFun_apply]
simp [PauliMatrix.σSAL, PauliMatrix.σSAL']
| Sum.inr 0 =>
simp only [LorentzVector.stdBasis, Fin.isValue]
erw [Pi.basisFun_apply]
simp only [Fin.isValue, ne_eq, reduceCtorEq, not_false_eq_true, Pi.single_eq_of_ne, zero_smul,
Pi.single_eq_same, one_smul, zero_sub, Sum.inr.injEq, one_ne_zero, sub_zero, Fin.reduceEq,
PauliMatrix.σSAL, Basis.coe_mk, PauliMatrix.σSAL']
rfl
| Sum.inr 1 =>
simp only [LorentzVector.stdBasis, Fin.isValue]
erw [Pi.basisFun_apply]
simp only [Fin.isValue, ne_eq, reduceCtorEq, not_false_eq_true, Pi.single_eq_of_ne, zero_smul,
Sum.inr.injEq, zero_ne_one, sub_self, Pi.single_eq_same, one_smul, zero_sub, Fin.reduceEq,
sub_zero, PauliMatrix.σSAL, Basis.coe_mk, PauliMatrix.σSAL']
rfl
| Sum.inr 2 =>
simp only [LorentzVector.stdBasis, Fin.isValue]
erw [Pi.basisFun_apply]
simp only [Fin.isValue, ne_eq, reduceCtorEq, not_false_eq_true, Pi.single_eq_of_ne, zero_smul,
Sum.inr.injEq, Fin.reduceEq, sub_self, Pi.single_eq_same, one_smul, zero_sub,
PauliMatrix.σSAL, Basis.coe_mk, PauliMatrix.σSAL']
rfl
@[simp]
lemma toSelfAdjointMatrix_symm_basis (i : Fin 1 ⊕ Fin 3) :
toSelfAdjointMatrix.symm (PauliMatrix.σSAL i) = (LorentzVector.stdBasis i) := by
refine (LinearEquiv.symm_apply_eq toSelfAdjointMatrix).mpr ?_
rw [toSelfAdjointMatrix_stdBasis]
open minkowskiMetric in
lemma det_eq_ηLin (x : LorentzVector 3) : det (toSelfAdjointMatrix x).1 = ⟪x, x⟫ₘ := by
rw [toSelfAdjointMatrix_apply_coe]
simp only [Fin.isValue, eq_time_minus_inner_prod, LorentzVector.time, LorentzVector.space,
PiLp.inner_apply, Function.comp_apply, RCLike.inner_apply, conj_trivial, Fin.sum_univ_three,
ofReal_sub, ofReal_mul, ofReal_add]
simp only [Fin.isValue, PauliMatrix.σ0, smul_of, smul_cons, real_smul, mul_one, smul_zero,
smul_empty, PauliMatrix.σ1, of_sub_of, sub_cons, head_cons, sub_zero, tail_cons, zero_sub,
sub_self, zero_empty, PauliMatrix.σ2, smul_neg, sub_neg_eq_add, PauliMatrix.σ3, det_fin_two_of]
ring_nf
simp only [Fin.isValue, I_sq, mul_neg, mul_one]
ring
end
end SpaceTime

View file

@ -1,136 +0,0 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Data.Complex.Exponential
import Mathlib.Analysis.InnerProductSpace.PiL2
/-!
# Lorentz vectors
(aka 4-vectors)
In this file we define a Lorentz vector (in 4d, this is more often called a 4-vector).
One of the most important example of a Lorentz vector is SpaceTime.
We will define the group action of the Lorentz group on Lorentz vectors in
`HepLean.SpaceTime.LorentzVector.LorentzAction` in such a way that `LorentzVector`
corresponds to contravariant Lorentz tensors.
-/
/- The number of space dimensions . -/
variable (d : )
/-- The type of (contravariant) Lorentz Vectors in `d`-space dimensions. -/
def LorentzVector : Type := (Fin 1 ⊕ Fin d) →
/-- An instance of an additive commutative monoid on `LorentzVector`. -/
instance : AddCommMonoid (LorentzVector d) := Pi.addCommMonoid
/-- An instance of a module on `LorentzVector`. -/
noncomputable instance : Module (LorentzVector d) := Pi.module _ _ _
instance : AddCommGroup (LorentzVector d) := Pi.addCommGroup
/-- The structure of a topological space `LorentzVector d`. -/
instance : TopologicalSpace (LorentzVector d) :=
haveI : NormedAddCommGroup (LorentzVector d) := Pi.normedAddCommGroup
UniformSpace.toTopologicalSpace
namespace LorentzVector
variable {d : } (v : LorentzVector d)
/-- The space components. -/
@[simp]
def space : EuclideanSpace (Fin d) := v ∘ Sum.inr
/-- The time component. -/
@[simp]
def time : := v (Sum.inl 0)
/-!
# The standard basis
-/
/-- The standard basis of `LorentzVector` indexed by `Fin 1 ⊕ Fin (d)`. -/
@[simps!]
noncomputable def stdBasis : Basis (Fin 1 ⊕ Fin (d)) (LorentzVector d) := Pi.basisFun _
/-- Notation for `stdBasis`. -/
scoped[LorentzVector] notation "e" => stdBasis
lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : e μ ν = if μ = ν then 1 else 0 := by
erw [stdBasis, Pi.basisFun_apply, Pi.single_apply]
refine Eq.symm (ite_congr ?h₁ (congrFun rfl) (congrFun rfl))
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
lemma decomp_stdBasis (v : LorentzVector d) : ∑ i, v i • e i = v := by
funext ν
rw [Finset.sum_apply]
rw [Finset.sum_eq_single_of_mem ν]
· simp only [HSMul.hSMul, SMul.smul, stdBasis]
erw [Pi.basisFun_apply]
simp only [Pi.single_eq_same, mul_one]
· exact Finset.mem_univ ν
· intros b _ hbi
simp only [HSMul.hSMul, SMul.smul, stdBasis, mul_eq_zero]
erw [Pi.basisFun_apply]
simp only [Pi.single]
apply Or.inr $ Function.update_noteq (id (Ne.symm hbi)) 1 0
@[simp]
lemma decomp_stdBasis' (v : LorentzVector d) :
v (Sum.inl 0) • e (Sum.inl 0) + ∑ a₂ : Fin d, v (Sum.inr a₂) • e (Sum.inr a₂) = v := by
trans ∑ i, v i • e i
· simp only [Fin.isValue, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero,
Finset.sum_singleton]
· exact decomp_stdBasis v
/-- The standard unit time vector. -/
noncomputable abbrev timeVec : (LorentzVector d) := e (Sum.inl 0)
lemma timeVec_space : (@timeVec d).space = 0 := by
funext i
simp only [space, Function.comp_apply, stdBasis_apply, Fin.isValue, reduceCtorEq, ↓reduceIte,
PiLp.zero_apply]
lemma timeVec_time: (@timeVec d).time = 1 := by
simp only [time, Fin.isValue, stdBasis_apply, ↓reduceIte]
/-!
# Reflection of space
-/
/-- The reflection of space as a linear map. -/
@[simps!]
def spaceReflectionLin : LorentzVector d →ₗ[] LorentzVector d where
toFun x := Sum.elim (x ∘ Sum.inl) (- x ∘ Sum.inr)
map_add' x y := by
funext i
rcases i with i | i
· rfl
· simp only [Sum.elim_inr, Pi.neg_apply]
apply neg_add
map_smul' c x := by
funext i
rcases i with i | i
· rfl
· simp [HSMul.hSMul, SMul.smul]
/-- The reflection of space. -/
@[simp]
def spaceReflection : LorentzVector d := spaceReflectionLin v
lemma spaceReflection_space : v.spaceReflection.space = - v.space := rfl
lemma spaceReflection_time : v.spaceReflection.time = v.time := rfl
end LorentzVector

View file

@ -1,108 +0,0 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.Basic
import HepLean.SpaceTime.LorentzGroup.Basic
import Mathlib.RepresentationTheory.Basic
/-!
# Covariant Lorentz vectors
The type `LorentzVector` corresponds to contravariant Lorentz tensors.
In this section we define covariant Lorentz tensors.
-/
/-! TODO: Define equivariant map between contravariant and covariant lorentz tensors. -/
noncomputable section
/- The number of space dimensions . -/
variable (d : )
/-- The type of covariant Lorentz Vectors in `d`-space dimensions. -/
def CovariantLorentzVector : Type := (Fin 1 ⊕ Fin d) →
/-- An instance of an additive commutative monoid on `LorentzVector`. -/
instance : AddCommMonoid (CovariantLorentzVector d) := Pi.addCommMonoid
/-- An instance of a module on `LorentzVector`. -/
noncomputable instance : Module (CovariantLorentzVector d) := Pi.module _ _ _
instance : AddCommGroup (CovariantLorentzVector d) := Pi.addCommGroup
/-- The structure of a topological space `LorentzVector d`. -/
instance : TopologicalSpace (CovariantLorentzVector d) :=
haveI : NormedAddCommGroup (CovariantLorentzVector d) := Pi.normedAddCommGroup
UniformSpace.toTopologicalSpace
namespace CovariantLorentzVector
variable {d : } (v : CovariantLorentzVector d)
/-- The standard basis of `LorentzVector` indexed by `Fin 1 ⊕ Fin (d)`. -/
@[simps!]
noncomputable def stdBasis : Basis (Fin 1 ⊕ Fin (d)) (CovariantLorentzVector d) := Pi.basisFun _
lemma decomp_stdBasis (v : CovariantLorentzVector d) : ∑ i, v i • stdBasis i = v := by
funext ν
rw [Finset.sum_apply, Finset.sum_eq_single_of_mem ν]
· simp only [HSMul.hSMul, SMul.smul, stdBasis]
erw [Pi.basisFun_apply]
simp only [Pi.single_eq_same, mul_one]
· exact Finset.mem_univ ν
· intros b _ hbi
simp only [HSMul.hSMul, SMul.smul, stdBasis, mul_eq_zero]
erw [Pi.basisFun_apply]
simp only [Pi.single]
apply Or.inr (Function.update_noteq (id (Ne.symm hbi)) 1 0)
@[simp]
lemma decomp_stdBasis' (v : CovariantLorentzVector d) :
v (Sum.inl 0) • stdBasis (Sum.inl 0)
+ ∑ a₂ : Fin d, v (Sum.inr a₂) • stdBasis (Sum.inr a₂) = v := by
trans ∑ i, v i • stdBasis i
· simp only [Fin.isValue, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero,
Finset.sum_singleton]
· exact decomp_stdBasis v
/-!
## Lorentz group action on covariant Lorentz vectors
-/
/-- The representation of the Lorentz group acting on covariant Lorentz vectors. -/
def rep : Representation (LorentzGroup d) (CovariantLorentzVector d) where
toFun g := Matrix.toLinAlgEquiv stdBasis (LorentzGroup.transpose g⁻¹)
map_one' := by
simp only [inv_one, LorentzGroup.transpose_one, lorentzGroupIsGroup_one_coe, map_one]
map_mul' x y := by
simp only [mul_inv_rev, lorentzGroupIsGroup_inv, LorentzGroup.transpose_mul,
lorentzGroupIsGroup_mul_coe, map_mul]
open Matrix in
lemma rep_apply (g : LorentzGroup d) : rep g v = (g.1⁻¹)ᵀ *ᵥ v := by
trans (LorentzGroup.transpose g⁻¹) *ᵥ v
· rfl
· apply congrFun
apply congrArg
simp only [LorentzGroup.transpose, lorentzGroupIsGroup_inv, transpose_inj]
rw [← LorentzGroup.coe_inv]
rfl
lemma rep_apply_stdBasis (g : LorentzGroup d) (μ : Fin 1 ⊕ Fin d) :
rep g (stdBasis μ) = ∑ ν, g⁻¹.1 μ ν • stdBasis ν := by
simp only [rep_apply, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, decomp_stdBasis']
funext ν
simp only [lorentzGroupIsGroup_inv]
erw [Pi.basisFun_apply, Matrix.mulVec_single_one]
rw [← LorentzGroup.coe_inv]
rfl
end CovariantLorentzVector
end

View file

@ -1,41 +0,0 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.Basic
import HepLean.SpaceTime.LorentzGroup.Basic
import Mathlib.RepresentationTheory.Basic
/-!
# Lorentz group action on Lorentz vectors.
-/
noncomputable section
namespace LorentzVector
variable {d : } (v : LorentzVector d)
/-- The contravariant action of the Lorentz group on a Lorentz vector. -/
def rep : Representation (LorentzGroup d) (LorentzVector d) where
toFun g := Matrix.toLinAlgEquiv e g
map_one' := (MulEquivClass.map_eq_one_iff (Matrix.toLinAlgEquiv e)).mpr rfl
map_mul' x y := by
simp only [lorentzGroupIsGroup_mul_coe, map_mul]
open Matrix in
lemma rep_apply (g : LorentzGroup d) : rep g v = g *ᵥ v := rfl
lemma rep_apply_stdBasis (g : LorentzGroup d) (μ : Fin 1 ⊕ Fin d) :
rep g (stdBasis μ) = ∑ ν, g.1.transpose μ ν • stdBasis ν := by
simp only [rep_apply, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, decomp_stdBasis']
funext ν
simp only [stdBasis, Matrix.transpose_apply]
erw [Pi.basisFun_apply, Matrix.mulVec_single_one]
rfl
end LorentzVector
end

View file

@ -1,287 +0,0 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.Basic
import HepLean.SpaceTime.MinkowskiMetric
/-!
# Lorentz vectors with norm one
-/
open minkowskiMetric
/-- The set of Lorentz vectors with norm 1. -/
@[simp]
def NormOneLorentzVector (d : ) : Set (LorentzVector d) :=
fun x => ⟪x, x⟫ₘ = 1
instance : TopologicalSpace (NormOneLorentzVector d) := instTopologicalSpaceSubtype
namespace NormOneLorentzVector
variable {d : }
section
variable (v w : NormOneLorentzVector d)
lemma mem_iff {x : LorentzVector d} : x ∈ NormOneLorentzVector d ↔ ⟪x, x⟫ₘ = 1 := by
rfl
/-- The negative of a `NormOneLorentzVector` as a `NormOneLorentzVector`. -/
def neg : NormOneLorentzVector d := ⟨- v, by
rw [mem_iff]
simp only [map_neg, LinearMap.neg_apply, neg_neg]
exact v.2⟩
lemma time_sq : v.1.time ^ 2 = 1 + ‖v.1.space‖ ^ 2 := by
rw [time_sq_eq_metric_add_space, v.2]
lemma abs_time_ge_one : 1 ≤ |v.1.time| := by
have h1 := leq_time_sq v.1
rw [v.2] at h1
exact (one_le_sq_iff_one_le_abs _).mp h1
lemma norm_space_le_abs_time : ‖v.1.space‖ < |v.1.time| := by
rw [(abs_norm _).symm, ← @sq_lt_sq, time_sq]
exact lt_one_add (‖(v.1).space‖ ^ 2)
lemma norm_space_leq_abs_time : ‖v.1.space‖ ≤ |v.1.time| :=
le_of_lt (norm_space_le_abs_time v)
lemma time_le_minus_one_or_ge_one : v.1.time ≤ -1 1 ≤ v.1.time :=
le_abs'.mp (abs_time_ge_one v)
lemma time_nonpos_iff : v.1.time ≤ 0 ↔ v.1.time ≤ - 1 := by
apply Iff.intro
· intro h
cases' time_le_minus_one_or_ge_one v with h1 h1
· exact h1
· linarith
· intro h
linarith
lemma time_nonneg_iff : 0 ≤ v.1.time ↔ 1 ≤ v.1.time := by
apply Iff.intro
· intro h
cases' time_le_minus_one_or_ge_one v with h1 h1
· linarith
· exact h1
· intro h
linarith
lemma time_pos_iff : 0 < v.1.time ↔ 1 ≤ v.1.time := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· exact (time_nonneg_iff v).mp (le_of_lt h)
· linarith
lemma time_abs_sub_space_norm :
0 ≤ |v.1.time| * |w.1.time| - ‖v.1.space‖ * ‖w.1.space‖ := by
apply sub_nonneg.mpr
apply mul_le_mul (norm_space_leq_abs_time v) (norm_space_leq_abs_time w) ?_ ?_
· exact norm_nonneg w.1.space
· exact abs_nonneg (v.1 _)
/-!
# Future pointing norm one Lorentz vectors
-/
/-- The future pointing Lorentz vectors with Norm one. -/
def FuturePointing (d : ) : Set (NormOneLorentzVector d) :=
fun x => 0 < x.1.time
instance : TopologicalSpace (FuturePointing d) := instTopologicalSpaceSubtype
namespace FuturePointing
section
variable (f f' : FuturePointing d)
lemma mem_iff : v ∈ FuturePointing d ↔ 0 < v.1.time := by
rfl
lemma mem_iff_time_nonneg : v ∈ FuturePointing d ↔ 0 ≤ v.1.time := by
refine Iff.intro (fun h => le_of_lt h) (fun h => ?_)
rw [time_nonneg_iff] at h
rw [mem_iff]
linarith
lemma not_mem_iff : v ∉ FuturePointing d ↔ v.1.time ≤ 0 := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· exact le_of_not_lt ((mem_iff v).mp.mt h)
· have h1 := (mem_iff v).mp.mt
simp only [LorentzVector.time, Fin.isValue, not_lt] at h1
exact h1 h
lemma not_mem_iff_neg : v ∉ FuturePointing d ↔ neg v ∈ FuturePointing d := by
rw [not_mem_iff, mem_iff_time_nonneg]
simp only [Fin.isValue, neg]
change _ ↔ 0 ≤ - v.1 _
exact Iff.symm neg_nonneg
lemma time_nonneg : 0 ≤ f.1.1.time := le_of_lt f.2
lemma abs_time : |f.1.1.time| = f.1.1.time := abs_of_nonneg (time_nonneg f)
lemma time_eq_sqrt : f.1.1.time = √(1 + ‖f.1.1.space‖ ^ 2) := by
symm
rw [Real.sqrt_eq_cases]
apply Or.inl
rw [← time_sq, sq]
exact ⟨rfl, time_nonneg f⟩
/-!
# The sign of ⟪v, w.1⟫ₘ different v and w
-/
lemma metric_nonneg : 0 ≤ ⟪f, f'.1.1⟫ₘ := by
apply le_trans (time_abs_sub_space_norm f f'.1)
rw [abs_time f, abs_time f']
exact ge_sub_norm f.1.1 f'.1.1
lemma one_add_metric_non_zero : 1 + ⟪f, f'.1.1⟫ₘ ≠ 0 := by
linarith [metric_nonneg f f']
/-!
# The sign of ⟪v, w.1.spaceReflection⟫ₘ different v and w
-/
section
variable {v w : NormOneLorentzVector d}
open InnerProductSpace
lemma metric_reflect_mem_mem (h : v ∈ FuturePointing d) (hw : w ∈ FuturePointing d) :
0 ≤ ⟪v.1, w.1.spaceReflection⟫ₘ := by
apply le_trans (time_abs_sub_space_norm v w)
rw [abs_time ⟨v, h⟩, abs_time ⟨w, hw⟩, sub_eq_add_neg, right_spaceReflection]
apply (add_le_add_iff_left _).mpr
rw [neg_le]
apply le_trans (neg_le_abs _ : _ ≤ |⟪(v.1).space, (w.1).space⟫_|) ?_
exact abs_real_inner_le_norm (v.1).space (w.1).space
lemma metric_reflect_not_mem_not_mem (h : v ∉ FuturePointing d) (hw : w ∉ FuturePointing d) :
0 ≤ ⟪v.1, w.1.spaceReflection⟫ₘ := by
have h1 := metric_reflect_mem_mem ((not_mem_iff_neg v).mp h) ((not_mem_iff_neg w).mp hw)
apply le_of_le_of_eq h1 ?_
simp [neg]
lemma metric_reflect_mem_not_mem (h : v ∈ FuturePointing d) (hw : w ∉ FuturePointing d) :
⟪v.1, w.1.spaceReflection⟫ₘ ≤ 0 := by
rw [show (0 : ) = - 0 from zero_eq_neg.mpr rfl, le_neg]
have h1 := metric_reflect_mem_mem h ((not_mem_iff_neg w).mp hw)
apply le_of_le_of_eq h1 ?_
simp [neg]
lemma metric_reflect_not_mem_mem (h : v ∉ FuturePointing d) (hw : w ∈ FuturePointing d) :
⟪v.1, w.1.spaceReflection⟫ₘ ≤ 0 := by
rw [show (0 : ) = - 0 from zero_eq_neg.mpr rfl, le_neg]
have h1 := metric_reflect_mem_mem ((not_mem_iff_neg v).mp h) hw
apply le_of_le_of_eq h1 ?_
simp [neg]
end
end
end FuturePointing
end
namespace FuturePointing
/-!
# Topology
-/
open LorentzVector
/-- The `FuturePointing d` which has all space components zero. -/
@[simps!]
noncomputable def timeVecNormOneFuture : FuturePointing d := ⟨⟨timeVec, by
rw [NormOneLorentzVector.mem_iff, on_timeVec]⟩, by
rw [mem_iff, timeVec_time]; exact Real.zero_lt_one⟩
/-- A continuous path from `timeVecNormOneFuture` to any other. -/
noncomputable def pathFromTime (u : FuturePointing d) : Path timeVecNormOneFuture u where
toFun t := ⟨
⟨fun i => match i with
| Sum.inl 0 => √(1 + t ^ 2 * ‖u.1.1.space‖ ^ 2)
| Sum.inr i => t * u.1.1.space i,
by
rw [NormOneLorentzVector.mem_iff, minkowskiMetric.eq_time_minus_inner_prod]
simp only [time, space, Function.comp_apply, PiLp.inner_apply, RCLike.inner_apply, map_mul,
conj_trivial]
rw [Real.mul_self_sqrt, ← @real_inner_self_eq_norm_sq, @PiLp.inner_apply]
· simp only [Function.comp_apply, RCLike.inner_apply, conj_trivial]
refine Eq.symm (eq_sub_of_add_eq (congrArg (HAdd.hAdd 1) ?_))
rw [Finset.mul_sum]
apply Finset.sum_congr rfl
intro i _
ring
· exact Right.add_nonneg (zero_le_one' ) $ mul_nonneg (sq_nonneg _) (sq_nonneg _)⟩,
by
simp only [space, Function.comp_apply, mem_iff_time_nonneg, time, Real.sqrt_pos]
exact Real.sqrt_nonneg _⟩
continuous_toFun := by
refine Continuous.subtype_mk ?_ _
refine Continuous.subtype_mk ?_ _
apply (continuous_pi_iff).mpr
intro i
match i with
| Sum.inl 0 =>
continuity
| Sum.inr i =>
continuity
source' := by
ext
funext i
match i with
| Sum.inl 0 =>
simp only [Set.Icc.coe_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow, space,
zero_mul, add_zero, Real.sqrt_one, Fin.isValue, timeVecNormOneFuture_coe_coe]
exact Eq.symm timeVec_time
| Sum.inr i =>
simp only [Set.Icc.coe_zero, space, Function.comp_apply, zero_mul,
timeVecNormOneFuture_coe_coe]
change _ = timeVec.space i
rw [timeVec_space]
rfl
target' := by
ext
funext i
match i with
| Sum.inl 0 =>
simp only [Set.Icc.coe_one, one_pow, space, one_mul, Fin.isValue]
exact (time_eq_sqrt u).symm
| Sum.inr i =>
simp only [Set.Icc.coe_one, one_pow, space, one_mul, Fin.isValue]
exact rfl
lemma isPathConnected : IsPathConnected (@Set.univ (FuturePointing d)) := by
use timeVecNormOneFuture
apply And.intro trivial ?_
intro y a
use pathFromTime y
exact fun _ => a
lemma metric_continuous (u : LorentzVector d) :
Continuous (fun (a : FuturePointing d) => ⟪u, a.1.1⟫ₘ) := by
simp only [minkowskiMetric.eq_time_minus_inner_prod]
refine Continuous.add ?_ ?_
· refine Continuous.comp' (continuous_mul_left _) $ Continuous.comp'
(continuous_apply (Sum.inl 0))
(Continuous.comp' continuous_subtype_val continuous_subtype_val)
· refine Continuous.comp' continuous_neg $ Continuous.inner
(Continuous.comp' (Pi.continuous_precomp Sum.inr) continuous_const)
(Continuous.comp' (Pi.continuous_precomp Sum.inr) (Continuous.comp'
continuous_subtype_val continuous_subtype_val))
end FuturePointing
end NormOneLorentzVector

View file

@ -1,362 +0,0 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.Basic
import Mathlib.Algebra.Lie.Classical
/-!
# The Minkowski Metric
This file introduces the Minkowski metric on spacetime in the mainly-minus signature.
We define the Minkowski metric as a bilinear map on the vector space
of Lorentz vectors in d dimensions.
-/
open Matrix
open InnerProductSpace
/-!
# The definition of the Minkowski Matrix
-/
/-- The `d.succ`-dimensional real matrix of the form `diag(1, -1, -1, -1, ...)`. -/
def minkowskiMatrix {d : } : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) :=
LieAlgebra.Orthogonal.indefiniteDiagonal (Fin 1) (Fin d)
namespace minkowskiMatrix
variable {d : }
/-- Notation for `minkowskiMatrix`. -/
scoped[minkowskiMatrix] notation "η" => minkowskiMatrix
@[simp]
lemma sq : @minkowskiMatrix d * minkowskiMatrix = 1 := by
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_mul_diagonal]
ext1 i j
rcases i with i | i <;> rcases j with j | j
· simp only [diagonal, of_apply, Sum.inl.injEq, Sum.elim_inl, mul_one]
split
· rename_i h
subst h
simp_all only [one_apply_eq]
· simp_all only [ne_eq, Sum.inl.injEq, not_false_eq_true, one_apply_ne]
· rfl
· rfl
· simp only [diagonal, of_apply, Sum.inr.injEq, Sum.elim_inr, mul_neg, mul_one, neg_neg]
split
· rename_i h
subst h
simp_all only [one_apply_eq]
· simp_all only [ne_eq, Sum.inr.injEq, not_false_eq_true, one_apply_ne]
@[simp]
lemma eq_transpose : minkowskiMatrixᵀ = @minkowskiMatrix d := by
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_transpose]
@[simp]
lemma det_eq_neg_one_pow_d : (@minkowskiMatrix d).det = (- 1) ^ d := by
simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
@[simp]
lemma η_apply_mul_η_apply_diag (μ : Fin 1 ⊕ Fin d) : η μ μ * η μ μ = 1 := by
match μ with
| Sum.inl _ => simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
| Sum.inr _ => simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
lemma as_block : @minkowskiMatrix d =
Matrix.fromBlocks (1 : Matrix (Fin 1) (Fin 1) ) 0 0 (-1 : Matrix (Fin d) (Fin d) ) := by
rw [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, ← fromBlocks_diagonal]
refine fromBlocks_inj.mpr ?_
simp only [diagonal_one, true_and]
funext i j
rw [← diagonal_neg]
rfl
@[simp]
lemma off_diag_zero {μ ν : Fin 1 ⊕ Fin d} (h : μ ≠ ν) : η μ ν = 0 := by
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
exact diagonal_apply_ne _ h
lemma inl_0_inl_0 : @minkowskiMatrix d (Sum.inl 0) (Sum.inl 0) = 1 := by
rfl
lemma inr_i_inr_i (i : Fin d) : @minkowskiMatrix d (Sum.inr i) (Sum.inr i) = -1 := by
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
simp_all only [diagonal_apply_eq, Sum.elim_inr]
end minkowskiMatrix
/-!
# The definition of the Minkowski Metric
-/
/-- Given a Lorentz vector `v` we define the the linear map `w ↦ v * η * w`. -/
@[simps!]
def minkowskiLinearForm {d : } (v : LorentzVector d) : LorentzVector d →ₗ[] where
toFun w := v ⬝ᵥ (minkowskiMatrix *ᵥ w)
map_add' y z := by
noncomm_ring
rw [mulVec_add, dotProduct_add]
map_smul' c y := by
simp only [RingHom.id_apply, smul_eq_mul]
rw [mulVec_smul, dotProduct_smul]
rfl
/-- The Minkowski metric as a bilinear map. -/
def minkowskiMetric {d : } : LorentzVector d →ₗ[] LorentzVector d →ₗ[] where
toFun v := minkowskiLinearForm v
map_add' y z := by
ext1 x
simp only [minkowskiLinearForm_apply, LinearMap.add_apply]
apply add_dotProduct
map_smul' c y := by
ext1 x
simp only [minkowskiLinearForm_apply, RingHom.id_apply, LinearMap.smul_apply, smul_eq_mul]
rw [smul_dotProduct]
rfl
namespace minkowskiMetric
open minkowskiMatrix
open LorentzVector
variable {d : }
variable (v w : LorentzVector d)
/-- Notation for `minkowskiMetric`. -/
scoped[minkowskiMetric] notation:102 "⟪" v "," w "⟫ₘ" => minkowskiMetric v w
/-!
# Equalitites involving the Minkowski metric
-/
/-- The Minkowski metric expressed as a sum. -/
lemma as_sum :
⟪v, w⟫ₘ = v.time * w.time - ∑ i, v.space i * w.space i := by
simp only [minkowskiMetric, LinearMap.coe_mk, AddHom.coe_mk, minkowskiLinearForm_apply,
dotProduct, LieAlgebra.Orthogonal.indefiniteDiagonal, mulVec_diagonal, Fintype.sum_sum_type,
Finset.univ_unique, Fin.default_eq_zero, Fin.isValue, Sum.elim_inl, one_mul,
Finset.sum_singleton, Sum.elim_inr, neg_mul, mul_neg, Finset.sum_neg_distrib, time, space,
Function.comp_apply, minkowskiMatrix]
rfl
/-- The Minkowski metric expressed as a sum for a single vector. -/
lemma as_sum_self : ⟪v, v⟫ₘ = v.time ^ 2 - ‖v.space‖ ^ 2 := by
rw [← real_inner_self_eq_norm_sq, PiLp.inner_apply, as_sum]
noncomm_ring
lemma eq_time_minus_inner_prod : ⟪v, w⟫ₘ = v.time * w.time - ⟪v.space, w.space⟫_ := by
rw [as_sum, @PiLp.inner_apply]
rfl
lemma self_eq_time_minus_norm : ⟪v, v⟫ₘ = v.time ^ 2 - ‖v.space‖ ^ 2 := by
rw [← real_inner_self_eq_norm_sq, PiLp.inner_apply, as_sum]
noncomm_ring
/-- The Minkowski metric is symmetric in its arguments.. -/
lemma symm : ⟪v, w⟫ₘ = ⟪w, v⟫ₘ := by
simp only [as_sum]
ac_rfl
lemma time_sq_eq_metric_add_space : v.time ^ 2 = ⟪v, v⟫ₘ + ‖v.space‖ ^ 2 := by
rw [self_eq_time_minus_norm]
exact Eq.symm (sub_add_cancel (v.time ^ 2) (‖v.space‖ ^ 2))
/-!
# Minkowski metric and space reflections
-/
lemma right_spaceReflection : ⟪v, w.spaceReflection⟫ₘ = v.time * w.time + ⟪v.space, w.space⟫_ := by
rw [eq_time_minus_inner_prod, spaceReflection_space, spaceReflection_time]
simp only [time, Fin.isValue, space, inner_neg_right, PiLp.inner_apply, Function.comp_apply,
RCLike.inner_apply, conj_trivial, sub_neg_eq_add]
lemma self_spaceReflection_eq_zero_iff : ⟪v, v.spaceReflection⟫ₘ = 0 ↔ v = 0 := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rw [right_spaceReflection] at h
have h2 : 0 ≤ ⟪v.space, v.space⟫_ := real_inner_self_nonneg
have h3 : v.time * v.time = 0 := by linarith [mul_self_nonneg v.time]
have h4 : ⟪v.space, v.space⟫_ = 0 := by linarith
simp only [time, Fin.isValue, mul_eq_zero, or_self] at h3
rw [@inner_self_eq_zero] at h4
funext i
rcases i with i | i
· fin_cases i
exact h3
· simpa using congrFun h4 i
· rw [h]
exact LinearMap.map_zero₂ minkowskiMetric (spaceReflection 0)
/-!
# Non degeneracy of the Minkowski metric
-/
/-- The metric tensor is non-degenerate. -/
lemma nondegenerate : (∀ w, ⟪w, v⟫ₘ = 0) ↔ v = 0 := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· exact (self_spaceReflection_eq_zero_iff _).mp ((symm _ _).trans $ h v.spaceReflection)
· simp [h]
/-!
# Inequalities involving the Minkowski metric
-/
lemma leq_time_sq : ⟪v, v⟫ₘ ≤ v.time ^ 2 := by
rw [time_sq_eq_metric_add_space]
exact (le_add_iff_nonneg_right _).mpr $ sq_nonneg ‖v.space‖
lemma ge_abs_inner_product : v.time * w.time - ‖⟪v.space, w.space⟫_‖ ≤ ⟪v, w⟫ₘ := by
rw [eq_time_minus_inner_prod, sub_le_sub_iff_left]
exact Real.le_norm_self ⟪v.space, w.space⟫_
lemma ge_sub_norm : v.time * w.time - ‖v.space‖ * ‖w.space‖ ≤ ⟪v, w⟫ₘ := by
apply le_trans _ (ge_abs_inner_product v w)
rw [sub_le_sub_iff_left]
exact norm_inner_le_norm v.space w.space
/-!
# The Minkowski metric and matrices
-/
section matrices
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
@[simp]
lemma dual_mul : dual (Λ * Λ') = dual Λ' * dual Λ := by
simp only [dual, transpose_mul]
trans η * Λ'ᵀ * (η * η) * Λᵀ * η
· noncomm_ring [minkowskiMatrix.sq]
· noncomm_ring
@[simp]
lemma dual_dual : dual (dual Λ) = Λ := by
simp only [dual, transpose_mul, transpose_transpose, eq_transpose]
trans (η * η) * Λ * (η * η)
· noncomm_ring
· noncomm_ring [minkowskiMatrix.sq]
@[simp]
lemma dual_eta : @dual d η = η := by
simp only [dual, eq_transpose]
noncomm_ring [minkowskiMatrix.sq]
@[simp]
lemma dual_transpose : dual Λᵀ = (dual Λ)ᵀ := by
simp only [dual, transpose_transpose, transpose_mul, eq_transpose]
noncomm_ring
@[simp]
lemma det_dual : (dual Λ).det = Λ.det := by
simp only [dual, det_mul, minkowskiMatrix.det_eq_neg_one_pow_d, det_transpose]
group
norm_cast
simp
lemma dual_apply (μ ν : Fin 1 ⊕ Fin d) :
dual Λ μ ν = η μ μ * Λ ν μ * η ν ν := by
simp only [dual, minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, mul_diagonal,
diagonal_mul, transpose_apply, diagonal_apply_eq]
lemma dual_apply_minkowskiMatrix (μ ν : Fin 1 ⊕ Fin d) :
dual Λ μ ν * η ν ν = η μ μ * Λ ν μ := by
rw [dual_apply, mul_assoc]
simp
@[simp]
lemma dual_mulVec_right : ⟪x, (dual Λ) *ᵥ y⟫ₘ = ⟪Λ *ᵥ x, y⟫ₘ := by
simp only [minkowskiMetric, LinearMap.coe_mk, AddHom.coe_mk, dual, minkowskiLinearForm_apply,
mulVec_mulVec, ← mul_assoc, minkowskiMatrix.sq, one_mul, (vecMul_transpose Λ x).symm, ←
dotProduct_mulVec]
@[simp]
lemma dual_mulVec_left : ⟪(dual Λ) *ᵥ x, y⟫ₘ = ⟪x, Λ *ᵥ y⟫ₘ := by
rw [symm, dual_mulVec_right, symm]
lemma matrix_apply_eq_iff_sub : ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, Λ' *ᵥ w⟫ₘ ↔ ⟪v, (Λ - Λ') *ᵥ w⟫ₘ = 0 := by
rw [← sub_eq_zero, ← LinearMap.map_sub, sub_mulVec]
lemma matrix_eq_iff_eq_forall' : (∀ v, Λ *ᵥ v= Λ' *ᵥ v) ↔ ∀ w v, ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, Λ' *ᵥ w⟫ₘ := by
refine Iff.intro (fun h ↦ fun w v ↦ ?_) (fun h ↦ fun v ↦ ?_)
· rw [h w]
· simp only [matrix_apply_eq_iff_sub] at h
refine sub_eq_zero.1 ?_
have h1 := h v
rw [nondegenerate] at h1
simp only [sub_mulVec] at h1
exact h1
lemma matrix_eq_iff_eq_forall : Λ = Λ' ↔ ∀ w v, ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, Λ' *ᵥ w⟫ₘ := by
rw [← matrix_eq_iff_eq_forall']
refine Iff.intro (fun h => ?_) (fun h => ?_)
· exact fun _ => congrFun (congrArg mulVec h) _
· rw [← (LinearMap.toMatrix stdBasis stdBasis).toEquiv.symm.apply_eq_iff_eq]
ext x
simp only [LinearEquiv.coe_toEquiv_symm, LinearMap.toMatrix_symm, EquivLike.coe_coe,
toLin_apply, h, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton]
lemma matrix_eq_id_iff : Λ = 1 ↔ ∀ w v, ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, w⟫ₘ := by
rw [matrix_eq_iff_eq_forall]
simp
/-!
# The Minkowski metric and the standard basis
-/
@[simp]
lemma basis_left (μ : Fin 1 ⊕ Fin d) : ⟪e μ, v⟫ₘ = η μ μ * v μ := by
rw [as_sum]
rcases μ with μ | μ
· fin_cases μ
simp [stdBasis_apply, minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
· simp [stdBasis_apply, minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
lemma on_timeVec : ⟪timeVec, @timeVec d⟫ₘ = 1 := by
simp only [timeVec, Fin.isValue, basis_left, minkowskiMatrix,
LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_apply_eq, Sum.elim_inl, stdBasis_apply,
↓reduceIte, mul_one]
lemma on_basis_mulVec (μ ν : Fin 1 ⊕ Fin d) : ⟪e μ, Λ *ᵥ e ν⟫ₘ = η μ μ * Λ μ ν := by
simp [basis_left, mulVec, dotProduct, stdBasis_apply]
lemma on_basis (μ ν : Fin 1 ⊕ Fin d) : ⟪e μ, e ν⟫ₘ = η μ ν := by
rw [basis_left, stdBasis_apply]
by_cases h : μ = ν
· simp [h]
· simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_apply_eq,
mul_ite, mul_one, mul_zero, ne_eq, h, not_false_eq_true, diagonal_apply_ne, ite_eq_right_iff]
exact fun a => False.elim (h (id (Eq.symm a)))
lemma matrix_apply_stdBasis (ν μ : Fin 1 ⊕ Fin d) :
Λ ν μ = η ν ν * ⟪e ν, Λ *ᵥ e μ⟫ₘ := by
rw [on_basis_mulVec, ← mul_assoc]
simp [η_apply_mul_η_apply_diag ν]
end matrices
end minkowskiMetric

View file

@ -1,221 +0,0 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzGroup.Basic
import Mathlib.RepresentationTheory.Basic
import HepLean.SpaceTime.LorentzVector.AsSelfAdjointMatrix
import HepLean.SpaceTime.LorentzGroup.Restricted
import HepLean.Meta.Informal
/-!
# The group SL(2, ) and it's relation to the Lorentz group
The aim of this file is to give the relationship between `SL(2, )` and the Lorentz group.
-/
namespace SpaceTime
open Matrix
open MatrixGroups
open Complex
namespace SL2C
open SpaceTime
noncomputable section
/-!
## Some basic properties about SL(2, )
Possibly to be moved to mathlib at some point.
-/
lemma inverse_coe (M : SL(2, )) : M.1⁻¹ = (M⁻¹).1 := by
apply Matrix.inv_inj
simp only [SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq, one_ne_zero, not_false_eq_true,
nonsing_inv_nonsing_inv, SpecialLinearGroup.coe_inv]
have h1 : IsUnit M.1.det := by
simp
rw [Matrix.inv_adjugate M.1 h1]
· simp
· simp
lemma transpose_coe (M : SL(2, )) : M.1ᵀ = (M.transpose).1 := rfl
/-!
## Representation of SL(2, ) on spacetime
Through the correspondence between spacetime and self-adjoint matrices,
we can define a representation a representation of `SL(2, )` on spacetime.
-/
/-- Given an element `M ∈ SL(2, )` the linear map from `selfAdjoint (Matrix (Fin 2) (Fin 2) )` to
itself defined by `A ↦ M * A * Mᴴ`. -/
@[simps!]
def toLinearMapSelfAdjointMatrix (M : SL(2, )) :
selfAdjoint (Matrix (Fin 2) (Fin 2) ) →ₗ[] selfAdjoint (Matrix (Fin 2) (Fin 2) ) where
toFun A := ⟨M.1 * A.1 * Matrix.conjTranspose M,
by
noncomm_ring [selfAdjoint.mem_iff, star_eq_conjTranspose,
conjTranspose_mul, conjTranspose_conjTranspose,
(star_eq_conjTranspose A.1).symm.trans $ selfAdjoint.mem_iff.mp A.2]⟩
map_add' A B := by
simp only [AddSubgroup.coe_add, AddMemClass.mk_add_mk, Subtype.mk.injEq]
noncomm_ring [AddSubmonoid.coe_add, AddSubgroup.coe_toAddSubmonoid, AddSubmonoid.mk_add_mk,
Subtype.mk.injEq]
map_smul' r A := by
noncomm_ring [selfAdjoint.val_smul, Algebra.mul_smul_comm, Algebra.smul_mul_assoc,
RingHom.id_apply]
/-- The representation of `SL(2, )` on `selfAdjoint (Matrix (Fin 2) (Fin 2) )` given by
`M A ↦ M * A * Mᴴ`. -/
@[simps!]
def repSelfAdjointMatrix : Representation SL(2, ) $ selfAdjoint (Matrix (Fin 2) (Fin 2) ) where
toFun := toLinearMapSelfAdjointMatrix
map_one' := by
noncomm_ring [toLinearMapSelfAdjointMatrix, SpecialLinearGroup.coe_one, one_mul,
conjTranspose_one, mul_one, Subtype.coe_eta]
map_mul' M N := by
ext x i j : 3
noncomm_ring [toLinearMapSelfAdjointMatrix, SpecialLinearGroup.coe_mul, mul_assoc,
conjTranspose_mul, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.mul_apply]
/-- The representation of `SL(2, )` on `spaceTime` obtained from `toSelfAdjointMatrix` and
`repSelfAdjointMatrix`. -/
def repLorentzVector : Representation SL(2, ) (LorentzVector 3) where
toFun M := toSelfAdjointMatrix.symm.comp ((repSelfAdjointMatrix M).comp
toSelfAdjointMatrix.toLinearMap)
map_one' := by
ext
simp
map_mul' M N := by
ext x : 3
simp
/-!
## Homomorphism to the Lorentz group
There is a group homomorphism from `SL(2, )` to the Lorentz group `𝓛`.
The purpose of this section is to define this homomorphism.
In the next section we will restrict this homomorphism to the restricted Lorentz group.
-/
lemma iff_det_selfAdjoint (Λ : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ) : Λ ∈ LorentzGroup 3 ↔
∀ (x : selfAdjoint (Matrix (Fin 2) (Fin 2) )),
det ((toSelfAdjointMatrix ∘
toLin LorentzVector.stdBasis LorentzVector.stdBasis Λ ∘ toSelfAdjointMatrix.symm) x).1
= det x.1 := by
rw [LorentzGroup.mem_iff_norm]
refine Iff.intro (fun h x => ?_) (fun h x => ?_)
· simpa [← det_eq_ηLin] using congrArg ofReal $ h (toSelfAdjointMatrix.symm x)
· simpa [det_eq_ηLin] using h (toSelfAdjointMatrix x)
/-- Given an element `M ∈ SL(2, )` the corresponding element of the Lorentz group. -/
@[simps!]
def toLorentzGroupElem (M : SL(2, )) : LorentzGroup 3 :=
⟨LinearMap.toMatrix LorentzVector.stdBasis LorentzVector.stdBasis (repLorentzVector M),
by simp [repLorentzVector, iff_det_selfAdjoint]⟩
/-- The group homomorphism from ` SL(2, )` to the Lorentz group `𝓛`. -/
@[simps!]
def toLorentzGroup : SL(2, ) →* LorentzGroup 3 where
toFun := toLorentzGroupElem
map_one' := by
simp only [toLorentzGroupElem, _root_.map_one, LinearMap.toMatrix_one]
rfl
map_mul' M N := by
apply Subtype.eq
simp only [toLorentzGroupElem, _root_.map_mul, LinearMap.toMatrix_mul,
lorentzGroupIsGroup_mul_coe]
lemma toLorentzGroup_eq_σSAL (M : SL(2, )) :
toLorentzGroup M = LinearMap.toMatrix
PauliMatrix.σSAL PauliMatrix.σSAL (repSelfAdjointMatrix M) := by
rfl
lemma toLorentzGroup_eq_stdBasis (M : SL(2, )) :
toLorentzGroup M = LinearMap.toMatrix LorentzVector.stdBasis LorentzVector.stdBasis
(repLorentzVector M) := by rfl
lemma repLorentzVector_apply_eq_mulVec (v : LorentzVector 3) :
SL2C.repLorentzVector M v = (SL2C.toLorentzGroup M).1 *ᵥ v := by
simp only [toLorentzGroup, MonoidHom.coe_mk, OneHom.coe_mk, toLorentzGroupElem_coe]
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]
rfl
lemma repSelfAdjointMatrix_basis (i : Fin 1 ⊕ Fin 3) :
SL2C.repSelfAdjointMatrix M (PauliMatrix.σSAL i) =
∑ j, (toLorentzGroup M).1 j i •
PauliMatrix.σSAL j := by
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)))]
rfl
lemma repSelfAdjointMatrix_σSA (i : Fin 1 ⊕ Fin 3) :
SL2C.repSelfAdjointMatrix M (PauliMatrix.σSA i) =
∑ j, (toLorentzGroup M⁻¹).1 i j • PauliMatrix.σSA j := by
have h1 : (toLorentzGroup M⁻¹).1 = minkowskiMetric.dual (toLorentzGroup M).1 := by
simp
simp only [h1]
rw [PauliMatrix.σSA_minkowskiMetric_σSAL, _root_.map_smul]
rw [repSelfAdjointMatrix_basis]
rw [Finset.smul_sum]
apply congrArg
funext j
rw [smul_smul, PauliMatrix.σSA_minkowskiMetric_σSAL, smul_smul]
apply congrFun
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
simp only [repLorentzVector, MonoidHom.coe_mk, OneHom.coe_mk, LinearMap.coe_comp,
LinearEquiv.coe_coe, Function.comp_apply]
rw [toSelfAdjointMatrix_stdBasis]
rw [repSelfAdjointMatrix_basis]
rw [map_sum]
apply congrArg
funext j
simp
/-!
## Homomorphism to the restricted Lorentz group
The homomorphism `toLorentzGroup` restricts to a homomorphism to the restricted Lorentz group.
In this section we will define this homomorphism.
-/
informal_lemma toLorentzGroup_det_one where
math :≈ "The determinant of the image of `SL(2, )` in the Lorentz group is one."
deps :≈ [``toLorentzGroup]
informal_lemma toLorentzGroup_timeComp_nonneg where
math :≈ "The time coponent of the image of `SL(2, )` in the Lorentz group is non-negative."
deps :≈ [``toLorentzGroup, ``LorentzGroup.timeComp]
informal_lemma toRestrictedLorentzGroup where
math :≈ "The homomorphism from `SL(2, )` to the restricted Lorentz group."
deps :≈ [``toLorentzGroup, ``toLorentzGroup_det_one, ``toLorentzGroup_timeComp_nonneg,
``LorentzGroup.Restricted]
/-! TODO: Define homomorphism from `SL(2, )` to the restricted Lorentz group. -/
end
end SL2C
end SpaceTime

View file

@ -9,8 +9,8 @@ import Mathlib.CategoryTheory.Monoidal.Category
import Mathlib.CategoryTheory.Comma.Over
import Mathlib.CategoryTheory.Core
import Mathlib.CategoryTheory.Monoidal.Braided.Basic
import HepLean.SpaceTime.WeylFermion.Basic
import HepLean.SpaceTime.LorentzVector.Complex.Basic
import HepLean.Lorentz.Weyl.Basic
import HepLean.Lorentz.ComplexVector.Basic
/-!
# Over color category.

View file

@ -6,7 +6,7 @@ Authors: Joseph Tooby-Smith
import HepLean.Tensors.Tree.Basic
import Lean.Elab.Term
import HepLean.Tensors.Tree.Dot
import HepLean.Tensors.ComplexLorentz.Basic
import HepLean.Lorentz.ComplexTensor.Basic
/-!
# Elaboration of tensor trees