PhysLean/HepLean/SpaceTime/LorentzVector/Real/Modules.lean
2024-11-08 06:13:03 +00:00

130 lines
4.1 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
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.
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 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!]
def toFin1dEquiv : ContrModule d ≃ₗ[] (Fin 1 ⊕ Fin d → ) where
toFun := toFin1dFun
map_add' := fun _ _ => rfl
map_smul' := fun _ _ => rfl
invFun := toFin1dFun.symm
left_inv := fun _ => rfl
right_inv := fun _ => rfl
/-- The underlying element of `Fin 1 ⊕ Fin d → ` of a element in `ContrModule` defined
through the linear equivalence `toFin1dEquiv`. -/
abbrev toFin1d (ψ : ContrModule d) := toFin1dEquiv ψ
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!]
def toFin1dEquiv : CoModule d ≃ₗ[] (Fin 1 ⊕ Fin d → ) where
toFun := toFin1dFun
map_add' := fun _ _ => rfl
map_smul' := fun _ _ => rfl
invFun := toFin1dFun.symm
left_inv := fun _ => rfl
right_inv := fun _ => rfl
/-- The underlying element of `Fin 1 ⊕ Fin d → ` of a element in `CoModule` defined
through the linear equivalence `toFin1dEquiv`. -/
abbrev toFin1d (ψ : CoModule d) := toFin1dEquiv ψ
end CoModule
end
end Lorentz