PhysLean/HepLean/SpaceTime/LorentzVector/Real/Modules.lean

145 lines
5.1 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
import HepLean.SpaceTime.SL2C.Basic
import Mathlib.RepresentationTheory.Rep
import Mathlib.Logic.Equiv.TransferInstance
/-!
## 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. -/
structure ContrModule (d : ) where
/-- The underlying value as a vector `Fin 1 ⊕ Fin d → `. -/
val : Fin 1 ⊕ Fin d →
namespace ContrModule
variable {d : }
/-- The equivalence between `ContrModule` and `Fin 1 ⊕ Fin d → `. -/
def toFin1dFun : ContrModule 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 (ContrModule d) := Equiv.addCommMonoid toFin1dFun
/-- The instance of `AddCommGroup` on `ContrModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
instance : AddCommGroup (ContrModule d) := Equiv.addCommGroup toFin1dFun
/-- The instance of `Module` on `ContrModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
instance : Module (ContrModule d) := Equiv.module toFin1dFun
@[ext]
lemma ext (ψ ψ' : ContrModule d) (h : ψ.val = ψ'.val) : ψ = ψ' := by
cases ψ
cases ψ'
subst h
rfl
@[simp]
lemma val_add (ψ ψ' : ContrModule d) : (ψ + ψ').val = ψ.val + ψ'.val := rfl
@[simp]
lemma val_smul (r : ) (ψ : ContrModule d) : (r • ψ).val = r • ψ.val := rfl
/-- The linear equivalence between `ContrModule` and `(Fin 1 ⊕ Fin d → )`. -/
@[simps!]
2024-11-08 06:41:33 +00:00
def toFin1dEquiv : ContrModule d ≃ₗ[] (Fin 1 ⊕ Fin d → ) :=
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`. -/
abbrev toFin1d (ψ : ContrModule d) := toFin1dEquiv ψ
2024-11-08 06:41:33 +00:00
/-- The standard basis of `ContrModule` indexed by `Fin 1 ⊕ Fin d`. -/
@[simps!]
def stdBasis : Basis (Fin 1 ⊕ Fin d) (ContrModule d) := Basis.ofEquivFun toFin1dEquiv
/-- The representation of the Lorentz group acting on `ContrModule d`. -/
def rep : Representation (LorentzGroup d) (ContrModule 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]
2024-11-08 06:07:18 +00:00
end ContrModule
/-- The module for covariant (up-index) complex Lorentz vectors. -/
structure CoModule (d : ) where
/-- The underlying value as a vector `Fin 1 ⊕ Fin d → `. -/
val : Fin 1 ⊕ Fin d →
namespace CoModule
variable {d : }
/-- The equivalence between `CoModule` and `Fin 1 ⊕ Fin d → `. -/
def toFin1dFun : CoModule 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 (CoModule d) := Equiv.addCommMonoid toFin1dFun
/-- The instance of `AddCommGroup` on `CoModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
instance : AddCommGroup (CoModule d) := Equiv.addCommGroup toFin1dFun
/-- The instance of `Module` on `CoModule` defined via its equivalence
with `Fin 1 ⊕ Fin d → `. -/
instance : Module (CoModule d) := Equiv.module toFin1dFun
/-- The linear equivalence between `CoModule` and `(Fin 1 ⊕ Fin d → )`. -/
@[simps!]
2024-11-08 06:41:33 +00:00
def toFin1dEquiv : CoModule d ≃ₗ[] (Fin 1 ⊕ Fin d → ) :=
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 06:13:03 +00:00
abbrev toFin1d (ψ : CoModule d) := toFin1dEquiv ψ
2024-11-08 06:07:18 +00:00
2024-11-08 06:41:33 +00:00
/-- The standard basis of `CoModule` indexed by `Fin 1 ⊕ Fin d`. -/
@[simps!]
def stdBasis : Basis (Fin 1 ⊕ Fin d) (CoModule d) := Basis.ofEquivFun toFin1dEquiv
/-- The representation of the Lorentz group acting on `CoModule d`. -/
def rep : Representation (LorentzGroup d) (CoModule 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]
2024-11-08 06:07:18 +00:00
end CoModule
end
end Lorentz