feat: Add reps to real modules

This commit is contained in:
jstoobysmith 2024-11-08 06:41:33 +00:00
parent 1350ab732d
commit 6b6f9261ca
4 changed files with 38 additions and 34 deletions

View file

@ -68,18 +68,24 @@ lemma val_smul (r : ) (ψ : ContrModule d) : (r • ψ).val = r • ψ.val
/-- 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
def toFin1dEquiv : ContrModule 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 (ψ : ContrModule d) := toFin1dEquiv ψ
/-- 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]
end ContrModule
/-- The module for covariant (up-index) complex Lorentz vectors. -/
@ -112,18 +118,26 @@ 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
def toFin1dEquiv : CoModule 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 (ψ : CoModule d) := toFin1dEquiv ψ
/-- 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]
end CoModule
end