feat: Add metric invariance for Real Lorentz Tensors
This commit is contained in:
parent
8b2c853fd8
commit
7b0b979d51
7 changed files with 175 additions and 24 deletions
|
@ -167,6 +167,27 @@ lemma subtype_mul_inv : (Subtype.val Λ) * (Subtype.val Λ)⁻¹ = 1 := by
|
|||
rw [mul_inv_self Λ]
|
||||
simp only [lorentzGroupIsGroup_one_coe]
|
||||
|
||||
@[simp]
|
||||
lemma mul_minkowskiMatrix_mul_transpose :
|
||||
(Subtype.val Λ) * minkowskiMatrix * (Subtype.val Λ).transpose = minkowskiMatrix := by
|
||||
have h2 := Λ.prop
|
||||
rw [LorentzGroup.mem_iff_self_mul_dual] at h2
|
||||
simp [minkowskiMetric.dual] at h2
|
||||
symm
|
||||
refine right_inv_eq_left_inv minkowskiMatrix.sq ?_
|
||||
rw [← h2]
|
||||
noncomm_ring
|
||||
|
||||
@[simp]
|
||||
lemma transpose_mul_minkowskiMatrix_mul_self :
|
||||
(Subtype.val Λ).transpose * minkowskiMatrix * (Subtype.val Λ) = minkowskiMatrix := by
|
||||
have h2 := Λ.prop
|
||||
rw [LorentzGroup.mem_iff_dual_mul_self] at h2
|
||||
simp [minkowskiMetric.dual] at h2
|
||||
refine right_inv_eq_left_inv ?_ minkowskiMatrix.sq
|
||||
rw [← h2]
|
||||
noncomm_ring
|
||||
|
||||
/-- The transpose of a matrix in the Lorentz group is an element of the Lorentz group. -/
|
||||
def transpose (Λ : LorentzGroup d) : LorentzGroup d :=
|
||||
⟨Λ.1ᵀ, mem_iff_transpose.mp Λ.2⟩
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue