PhysLean/HepLean/Lorentz/RealVector/Modules.lean

424 lines
16 KiB
Text
Raw Normal View History

2024-11-08 06:07:18 +00:00
/-
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.Basic
2024-11-09 17:43:48 +00:00
import HepLean.Lorentz.PauliMatrices.SelfAdjoint
2024-11-08 06:07:18 +00:00
/-!
## Modules associated with Real Lorentz vectors
We define the modules underlying real Lorentz vectors.
2024-11-08 06:13:03 +00:00
These definitions are preludes to the definitions of
`Lorentz.contr` and `Lorentz.co`.
2024-11-08 06:07:18 +00:00
-/
namespace Lorentz
noncomputable section
open Matrix
open MatrixGroups
open Complex
/-- The module for contravariant (up-index) real Lorentz vectors. -/
2024-11-08 07:11:57 +00:00
structure ContrMod (d : ) where
2024-11-08 06:07:18 +00:00
/-- The underlying value as a vector `Fin 1 ⊕ Fin d → `. -/
val : Fin 1 ⊕ Fin d →
2024-11-08 07:11:57 +00:00
namespace ContrMod
2024-11-08 06:07:18 +00:00
variable {d : }
2024-11-08 13:20:00 +00:00
@[ext]
lemma ext {ψ ψ' : ContrMod d} (h : ψ.val = ψ'.val) : ψ = ψ' := by
cases ψ
cases ψ'
subst h
rfl
2024-11-08 06:07:18 +00:00
/-- The equivalence between `ContrModule` and `Fin 1 ⊕ Fin d → `. -/
2024-11-08 07:11:57 +00:00
def toFin1dFun : ContrMod d ≃ (Fin 1 ⊕ Fin d → ) where
2024-11-08 06:07:18 +00:00
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 → `. -/
2024-11-08 07:11:57 +00:00
instance : AddCommMonoid (ContrMod d) := Equiv.addCommMonoid toFin1dFun
2024-11-08 06:07:18 +00:00
/-- The instance of `AddCommGroup` on `ContrModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
2024-11-08 07:11:57 +00:00
instance : AddCommGroup (ContrMod d) := Equiv.addCommGroup toFin1dFun
2024-11-08 06:07:18 +00:00
/-- The instance of `Module` on `ContrModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
2024-11-08 07:11:57 +00:00
instance : Module (ContrMod d) := Equiv.module toFin1dFun
2024-11-08 06:07:18 +00:00
@[simp]
2024-11-08 07:11:57 +00:00
lemma val_add (ψ ψ' : ContrMod d) : (ψ + ψ').val = ψ.val + ψ'.val := rfl
2024-11-08 06:07:18 +00:00
@[simp]
2024-11-08 07:11:57 +00:00
lemma val_smul (r : ) (ψ : ContrMod d) : (r • ψ).val = r • ψ.val := rfl
2024-11-08 06:07:18 +00:00
/-- The linear equivalence between `ContrModule` and `(Fin 1 ⊕ Fin d → )`. -/
2024-11-08 07:11:57 +00:00
def toFin1dEquiv : ContrMod d ≃ₗ[] (Fin 1 ⊕ Fin d → ) :=
2024-11-08 06:41:33 +00:00
Equiv.linearEquiv toFin1dFun
2024-11-08 06:07:18 +00:00
/-- The underlying element of `Fin 1 ⊕ Fin d → ` of a element in `ContrModule` defined
through the linear equivalence `toFin1dEquiv`. -/
2024-11-08 07:11:57 +00:00
abbrev toFin1d (ψ : ContrMod d) := toFin1dEquiv ψ
2024-11-08 06:07:18 +00:00
lemma toFin1d_eq_val (ψ : ContrMod d) : ψ.toFin1d = ψ.val := by rfl
2024-11-08 09:27:54 +00:00
/-!
## The standard basis.
-/
2024-11-08 06:41:33 +00:00
/-- The standard basis of `ContrModule` indexed by `Fin 1 ⊕ Fin d`. -/
@[simps!]
2024-11-08 07:11:57 +00:00
def stdBasis : Basis (Fin 1 ⊕ Fin d) (ContrMod d) := Basis.ofEquivFun toFin1dEquiv
2024-11-08 06:41:33 +00:00
2024-11-08 09:27:54 +00:00
@[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 μ
2024-11-08 09:27:54 +00:00
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
2024-11-09 18:06:48 +00:00
simp only [stdBasis, Basis.coe_ofEquivFun]
change Pi.single μ 1 ν = _
2024-11-09 18:06:48 +00:00
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)
2024-11-08 09:27:54 +00:00
/-- 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
2024-11-09 18:12:05 +00:00
simp only [map_sum, _root_.map_smul]
2024-11-08 09:27:54 +00:00
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]
/-!
2024-11-08 11:01:54 +00:00
## mulVec
-/
2024-11-10 06:57:41 +00:00
/-- Multiplication of a matrix with a vector in `ContrMod`. -/
2024-11-08 11:01:54 +00:00
abbrev mulVec (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (v : ContrMod d) :
ContrMod d := Matrix.toLinAlgEquiv stdBasis M v
2024-11-10 06:57:41 +00:00
/-- Multiplication of a matrix with a vector in `ContrMod`. -/
2024-11-08 13:20:00 +00:00
scoped[Lorentz] infixr:73 " *ᵥ " => ContrMod.mulVec
2024-11-08 11:01:54 +00:00
@[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
2024-11-08 13:20:00 +00:00
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]
2024-11-08 11:01:54 +00:00
/-!
## The norm
(Not the Minkowski norm, but the norm of a vector in `ContrModule d`.)
-/
2024-11-10 06:57:41 +00:00
/-- 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)
2024-11-10 06:57:41 +00:00
/-- 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
/-!
2024-11-08 09:27:54 +00:00
## The representation.
-/
2024-11-08 06:41:33 +00:00
/-- The representation of the Lorentz group acting on `ContrModule d`. -/
2024-11-08 07:11:57 +00:00
def rep : Representation (LorentzGroup d) (ContrMod d) where
2024-11-08 06:41:33 +00:00
toFun g := Matrix.toLinAlgEquiv stdBasis g
2024-12-10 13:44:39 +00:00
map_one' := EmbeddingLike.map_eq_one_iff.mpr rfl
2024-11-08 06:41:33 +00:00
map_mul' x y := by
simp only [lorentzGroupIsGroup_mul_coe, _root_.map_mul]
2024-11-08 09:27:54 +00:00
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]
2024-11-10 07:00:48 +00:00
/-!
## Topology
-/
2024-11-11 16:55:15 +00:00
/-- The type `ContrMod d` carries an instance of a topological group, induced by
it's equivalence to `Fin 1 ⊕ Fin d → `. -/
2024-11-10 07:00:48 +00:00
instance : TopologicalSpace (ContrMod d) := TopologicalSpace.induced
ContrMod.toFin1dEquiv (Pi.topologicalSpace)
2024-12-10 13:44:39 +00:00
open Topology
2024-11-10 07:00:48 +00:00
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
2024-11-08 07:11:57 +00:00
end ContrMod
2024-11-08 06:07:18 +00:00
/-- The module for covariant (up-index) complex Lorentz vectors. -/
2024-11-08 07:11:57 +00:00
structure CoMod (d : ) where
2024-11-08 06:07:18 +00:00
/-- The underlying value as a vector `Fin 1 ⊕ Fin d → `. -/
val : Fin 1 ⊕ Fin d →
2024-11-08 07:11:57 +00:00
namespace CoMod
2024-11-08 06:07:18 +00:00
variable {d : }
2024-11-08 13:20:00 +00:00
@[ext]
lemma ext {ψ ψ' : CoMod d} (h : ψ.val = ψ'.val) : ψ = ψ' := by
cases ψ
cases ψ'
subst h
rfl
2024-11-08 06:07:18 +00:00
/-- The equivalence between `CoModule` and `Fin 1 ⊕ Fin d → `. -/
2024-11-08 07:11:57 +00:00
def toFin1dFun : CoMod d ≃ (Fin 1 ⊕ Fin d → ) where
2024-11-08 06:07:18 +00:00
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 → `. -/
2024-11-08 07:11:57 +00:00
instance : AddCommMonoid (CoMod d) := Equiv.addCommMonoid toFin1dFun
2024-11-08 06:07:18 +00:00
/-- The instance of `AddCommGroup` on `CoModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
2024-11-08 07:11:57 +00:00
instance : AddCommGroup (CoMod d) := Equiv.addCommGroup toFin1dFun
2024-11-08 06:07:18 +00:00
/-- The instance of `Module` on `CoModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
2024-11-08 07:11:57 +00:00
instance : Module (CoMod d) := Equiv.module toFin1dFun
2024-11-08 06:07:18 +00:00
/-- The linear equivalence between `CoModule` and `(Fin 1 ⊕ Fin d → )`. -/
2024-11-08 07:11:57 +00:00
def toFin1dEquiv : CoMod d ≃ₗ[] (Fin 1 ⊕ Fin d → ) :=
2024-11-08 06:41:33 +00:00
Equiv.linearEquiv toFin1dFun
2024-11-08 06:07:18 +00:00
/-- The underlying element of `Fin 1 ⊕ Fin d → ` of a element in `CoModule` defined
through the linear equivalence `toFin1dEquiv`. -/
2024-11-08 07:11:57 +00:00
abbrev toFin1d (ψ : CoMod d) := toFin1dEquiv ψ
2024-11-08 06:07:18 +00:00
2024-11-08 09:27:54 +00:00
/-!
## The standard basis.
-/
2024-11-08 06:41:33 +00:00
/-- The standard basis of `CoModule` indexed by `Fin 1 ⊕ Fin d`. -/
@[simps!]
2024-11-08 07:11:57 +00:00
def stdBasis : Basis (Fin 1 ⊕ Fin d) (CoMod d) := Basis.ofEquivFun toFin1dEquiv
2024-11-08 06:41:33 +00:00
@[simp]
2024-11-08 09:27:54 +00:00
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 : μ ≠ ν) :
2024-11-08 09:27:54 +00:00
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
2024-11-09 18:06:48 +00:00
simp only [stdBasis, Basis.coe_ofEquivFun]
change Pi.single μ 1 ν = _
2024-11-09 18:06:48 +00:00
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
2024-11-09 18:12:05 +00:00
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 μ)]
2024-11-08 09:27:54 +00:00
· 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]
2024-11-08 09:27:54 +00:00
/-!
2024-11-08 11:01:54 +00:00
## mulVec
-/
2024-11-10 06:57:41 +00:00
/-- Multiplication of a matrix with a vector in `CoMod`. -/
2024-11-08 11:01:54 +00:00
abbrev mulVec (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (v : CoMod d) :
CoMod d := Matrix.toLinAlgEquiv stdBasis M v
2024-11-10 06:57:41 +00:00
/-- Multiplication of a matrix with a vector in `CoMod`. -/
2024-11-08 13:20:00 +00:00
scoped[Lorentz] infixr:73 " *ᵥ " => CoMod.mulVec
2024-11-08 11:01:54 +00:00
@[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
/-!
2024-11-08 09:27:54 +00:00
## The representation
-/
2024-11-08 06:41:33 +00:00
/-- The representation of the Lorentz group acting on `CoModule d`. -/
2024-11-08 07:11:57 +00:00
def rep : Representation (LorentzGroup d) (CoMod d) where
2024-11-08 06:41:33 +00:00
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]
2024-11-08 07:11:57 +00:00
end CoMod
2024-11-08 06:07:18 +00:00
end
end Lorentz