feat: Make MulActionTensor
This commit is contained in:
parent
99f4e85839
commit
a65fb06605
11 changed files with 177 additions and 55 deletions
|
@ -151,6 +151,22 @@ lemma coe_inv : (Λ⁻¹).1 = Λ.1⁻¹:= by
|
|||
refine (inv_eq_left_inv ?h).symm
|
||||
exact mem_iff_dual_mul_self.mp Λ.2
|
||||
|
||||
@[simp]
|
||||
lemma subtype_inv_mul : (Subtype.val Λ)⁻¹ * (Subtype.val Λ) = 1 := by
|
||||
trans Subtype.val (Λ⁻¹ * Λ)
|
||||
rw [← coe_inv]
|
||||
simp only [lorentzGroupIsGroup_inv, lorentzGroupIsGroup_mul_coe]
|
||||
rw [inv_mul_self Λ]
|
||||
simp only [lorentzGroupIsGroup_one_coe]
|
||||
|
||||
@[simp]
|
||||
lemma subtype_mul_inv : (Subtype.val Λ) * (Subtype.val Λ)⁻¹ = 1 := by
|
||||
trans Subtype.val (Λ * Λ⁻¹)
|
||||
rw [← coe_inv]
|
||||
simp only [lorentzGroupIsGroup_inv, lorentzGroupIsGroup_mul_coe]
|
||||
rw [mul_inv_self Λ]
|
||||
simp only [lorentzGroupIsGroup_one_coe]
|
||||
|
||||
/-- 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⟩
|
||||
|
|
|
@ -118,7 +118,7 @@ lemma toMatrix_continuous (u : FuturePointing d) : Continuous (toMatrix u) := by
|
|||
simp only [toMatrix_apply]
|
||||
refine Continuous.comp' (continuous_mul_left (η i i)) ?hf
|
||||
refine Continuous.sub ?_ ?_
|
||||
refine Continuous.comp' (continuous_add_left ⟪e i, e j⟫ₘ) ?_
|
||||
refine Continuous.comp' (continuous_add_left (minkowskiMetric (e i) (e j))) ?_
|
||||
refine Continuous.comp' (continuous_mul_left (2 * ⟪e j, u⟫ₘ)) ?_
|
||||
exact FuturePointing.metric_continuous _
|
||||
refine Continuous.mul ?_ ?_
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue