refactor: Major refactor of Lorentz vecs
This commit is contained in:
parent
3eb5da875f
commit
5cc188146f
20 changed files with 494 additions and 1005 deletions
|
@ -35,21 +35,19 @@ variable {d : ℕ}
|
|||
|
||||
/-- An auxillary linear map used in the definition of a generalised boost. -/
|
||||
def genBoostAux₁ (u v : FuturePointing d) : ContrMod d →ₗ[ℝ] ContrMod d where
|
||||
toFun x := (2 * toField d ⟪x, u.val.val⟫ₘ) • v.1.1
|
||||
toFun x := (2 * ⟪x, u.val.val⟫ₘ) • v.1.1
|
||||
map_add' x y := by
|
||||
simp [map_add, LinearMap.add_apply, tmul_add, add_tmul, mul_add, add_smul]
|
||||
map_smul' c x := by
|
||||
simp only [toField, Action.instMonoidalCategory_tensorObj_V,
|
||||
simp only [ Action.instMonoidalCategory_tensorObj_V,
|
||||
Action.instMonoidalCategory_tensorUnit_V, CategoryTheory.Equivalence.symm_inverse,
|
||||
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
|
||||
smul_tmul, tmul_smul, map_smul, smul_eq_mul, RingHom.id_apply]
|
||||
erw [LinearMap.id_apply]
|
||||
rw [← mul_assoc, mul_comm 2 c, mul_assoc, mul_smul]
|
||||
rfl
|
||||
|
||||
/-- An auxillary linear map used in the definition of a genearlised boost. -/
|
||||
def genBoostAux₂ (u v : FuturePointing d) : ContrMod d →ₗ[ℝ] ContrMod d where
|
||||
toFun x := - (toField d ⟪x, u.1.1 + v.1.1⟫ₘ / (1 + toField d ⟪u.1.1, v.1.1⟫ₘ)) • (u.1.1 + v.1.1)
|
||||
toFun x := - (⟪x, u.1.1 + v.1.1⟫ₘ / (1 + ⟪u.1.1, v.1.1⟫ₘ)) • (u.1.1 + v.1.1)
|
||||
map_add' x y := by
|
||||
simp only
|
||||
rw [← add_smul]
|
||||
|
@ -62,7 +60,7 @@ def genBoostAux₂ (u v : FuturePointing d) : ContrMod d →ₗ[ℝ] ContrMod d
|
|||
rw [map_smul]
|
||||
conv =>
|
||||
lhs; lhs; rhs; lhs
|
||||
change (c * toField d (contrContrContract.hom (x ⊗ₜ[ℝ] (u.val.val + v.val.val))))
|
||||
change (c * (contrContrContractField (x ⊗ₜ[ℝ] (u.val.val + v.val.val))))
|
||||
rw [mul_div_assoc, neg_mul_eq_mul_neg, smul_smul]
|
||||
rfl
|
||||
|
||||
|
@ -75,21 +73,47 @@ lemma genBoostAux₂_self (u : FuturePointing d) : genBoostAux₂ u u = - genBoo
|
|||
rw [smul_smul]
|
||||
congr 1
|
||||
rw [u.1.2]
|
||||
simp [toField]
|
||||
conv => lhs; lhs; rhs; rhs; change 1
|
||||
rw [show 1 + (1 : ℝ) = (2 : ℝ ) by ring]
|
||||
simp only [isUnit_iff_ne_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
|
||||
IsUnit.div_mul_cancel]
|
||||
rw [← (two_smul ℝ u.val.val)]
|
||||
simp only [tmul_smul, map_smul, smul_eq_mul]
|
||||
rfl
|
||||
|
||||
|
||||
/-- An generalised boost. This is a Lorentz transformation which takes the four velocity `u`
|
||||
to `v`. -/
|
||||
def genBoost (u v : FuturePointing d) : ContrMod d →ₗ[ℝ] ContrMod d :=
|
||||
LinearMap.id + genBoostAux₁ u v + genBoostAux₂ u v
|
||||
|
||||
lemma genBoost_mul_one_plus_contr (u v : FuturePointing d) (x : Contr d) :
|
||||
(1 + ⟪u.1.1, v.1.1⟫ₘ) • genBoost u v x =
|
||||
(1 + ⟪u.1.1, v.1.1⟫ₘ) • x +
|
||||
(2 * ⟪x, u.val.val⟫ₘ * (1 + ⟪u.1.1, v.1.1⟫ₘ)) • v.1.1
|
||||
- (⟪x, u.1.1 + v.1.1⟫ₘ) • (u.1.1 + v.1.1) := by
|
||||
simp only [genBoost, LinearMap.add_apply, LinearMap.id_apply, id_eq]
|
||||
rw [smul_add, smul_add]
|
||||
trans (1 + ⟪u.1.1, v.1.1⟫ₘ) • x +
|
||||
(2 * ⟪x, u.val.val⟫ₘ * (1 + ⟪u.1.1, v.1.1⟫ₘ)) • v.1.1
|
||||
+ (- ⟪x, u.1.1 + v.1.1⟫ₘ) • (u.1.1 + v.1.1)
|
||||
· congr 1
|
||||
· congr 1
|
||||
rw [genBoostAux₁]
|
||||
simp
|
||||
rw [smul_smul]
|
||||
congr 1
|
||||
ring
|
||||
· rw [genBoostAux₂]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
Action.FunctorCategoryEquivalence.functor_obj_obj, neg_smul, LinearMap.coe_mk,
|
||||
AddHom.coe_mk, smul_neg]
|
||||
rw [smul_smul]
|
||||
congr
|
||||
have h1 := FuturePointing.one_add_metric_non_zero u v
|
||||
field_simp
|
||||
· rw [neg_smul]
|
||||
rfl
|
||||
|
||||
namespace genBoost
|
||||
|
||||
/--
|
||||
|
@ -110,7 +134,7 @@ lemma toMatrix_mulVec (u v : FuturePointing d) (x : Contr d) :
|
|||
(toMatrix u v) *ᵥ x = genBoost u v x :=
|
||||
ContrMod.ext (LinearMap.toMatrix_mulVec_repr ContrMod.stdBasis ContrMod.stdBasis (genBoost u v) x)
|
||||
|
||||
open minkowskiMatrix LorentzVector in
|
||||
open minkowskiMatrix in
|
||||
@[simp]
|
||||
lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) :
|
||||
(toMatrix u v) μ ν = η μ μ * (toField d ⟪ContrMod.stdBasis μ, ContrMod.stdBasis ν⟫ₘ + 2 *
|
||||
|
@ -118,31 +142,39 @@ lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) :
|
|||
- toField d ⟪ContrMod.stdBasis μ, u.val.val + v.val.val⟫ₘ *
|
||||
toField d ⟪ContrMod.stdBasis ν, u.val.val + v.val.val⟫ₘ /
|
||||
(1 + toField d ⟪u.val.val, v.val.val⟫ₘ)) := by
|
||||
sorry
|
||||
rw [matrix_apply_stdBasis (toMatrix u v) μ ν, toMatrix_mulVec]
|
||||
rw [contrContrContractField.matrix_apply_stdBasis (Λ := toMatrix u v) μ ν, toMatrix_mulVec]
|
||||
simp only [genBoost, genBoostAux₁, genBoostAux₂, map_add, smul_add, neg_smul, LinearMap.add_apply,
|
||||
LinearMap.id_apply, LinearMap.coe_mk, AddHom.coe_mk, basis_left, map_smul, smul_eq_mul, map_neg,
|
||||
mul_eq_mul_left_iff]
|
||||
LinearMap.id_apply, LinearMap.coe_mk, AddHom.coe_mk, contrContrContractField.basis_left, map_smul, smul_eq_mul, map_neg,
|
||||
mul_eq_mul_left_iff, toField_apply]
|
||||
ring_nf
|
||||
exact (true_or (η μ μ = 0)).mpr trivial
|
||||
simp only [Pi.add_apply, Action.instMonoidalCategory_tensorObj_V,
|
||||
Action.instMonoidalCategory_tensorUnit_V, CategoryTheory.Equivalence.symm_inverse,
|
||||
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
|
||||
Pi.smul_apply, smul_eq_mul, Pi.sub_apply, Pi.neg_apply]
|
||||
apply Or.inl
|
||||
ring
|
||||
|
||||
open minkowskiMatrix LorentzVector in
|
||||
open minkowskiMatrix in
|
||||
lemma toMatrix_continuous (u : FuturePointing d) : Continuous (toMatrix u) := by
|
||||
refine continuous_matrix ?_
|
||||
intro i j
|
||||
simp only [toMatrix_apply]
|
||||
refine (continuous_mul_left (η i i)).comp' ?_
|
||||
refine Continuous.sub ?_ ?_
|
||||
· refine .comp' (continuous_add_left (minkowskiMetric (e i) (e j))) ?_
|
||||
refine .comp' (continuous_mul_left (2 * ⟪e j, u⟫ₘ)) ?_
|
||||
· refine .comp' (continuous_add_left _) ?_
|
||||
refine .comp' (continuous_mul_left _) ?_
|
||||
exact FuturePointing.metric_continuous _
|
||||
refine .mul ?_ ?_
|
||||
· refine .mul ?_ ?_
|
||||
· simp only [(minkowskiMetric _).map_add]
|
||||
· simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
Action.FunctorCategoryEquivalence.functor_obj_obj, tmul_add, map_add]
|
||||
refine .comp' ?_ ?_
|
||||
· exact continuous_add_left ((minkowskiMetric (stdBasis i)) ↑u)
|
||||
· exact continuous_add_left _
|
||||
· exact FuturePointing.metric_continuous _
|
||||
· simp only [(minkowskiMetric _).map_add]
|
||||
· simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
Action.FunctorCategoryEquivalence.functor_obj_obj, tmul_add, map_add]
|
||||
refine .comp' ?_ ?_
|
||||
· exact continuous_add_left _
|
||||
· exact FuturePointing.metric_continuous _
|
||||
|
@ -150,17 +182,28 @@ lemma toMatrix_continuous (u : FuturePointing d) : Continuous (toMatrix u) := by
|
|||
· refine .comp' (continuous_add_left 1) (FuturePointing.metric_continuous _)
|
||||
exact fun x => FuturePointing.one_add_metric_non_zero u x
|
||||
|
||||
lemma toMatrix_in_lorentzGroup (u v : FuturePointing d) : (toMatrix u v) ∈ LorentzGroup d:= by
|
||||
lemma toMatrix_in_lorentzGroup (u v : FuturePointing d) : (toMatrix u v) ∈ LorentzGroup d := by
|
||||
rw [LorentzGroup.mem_iff_invariant]
|
||||
intro x y
|
||||
rw [toMatrix_mulVec, toMatrix_mulVec, genBoost, genBoostAux₁, genBoostAux₂]
|
||||
have h1 : (1 + (minkowskiMetric ↑u) ↑v.1.1) ≠ 0 := FuturePointing.one_add_metric_non_zero u v
|
||||
simp only [map_add, smul_add, neg_smul, LinearMap.add_apply, LinearMap.id_coe,
|
||||
id_eq, LinearMap.coe_mk, AddHom.coe_mk, LinearMapClass.map_smul, map_neg, LinearMap.smul_apply,
|
||||
smul_eq_mul, LinearMap.neg_apply]
|
||||
field_simp
|
||||
rw [u.1.2, v.1.2, minkowskiMetric.symm v.1.1 u, minkowskiMetric.symm u y,
|
||||
minkowskiMetric.symm v y]
|
||||
ring
|
||||
rw [toMatrix_mulVec, toMatrix_mulVec]
|
||||
have h1 : (((1 + ⟪u.1.1, v.1.1⟫ₘ)) * (1 + ⟪u.1.1, v.1.1⟫ₘ)) •
|
||||
(contrContrContractField ((genBoost u v) y ⊗ₜ[ℝ] (genBoost u v) x))
|
||||
= (((1 + ⟪u.1.1, v.1.1⟫ₘ)) * (1 + ⟪u.1.1, v.1.1⟫ₘ)) • ⟪y, x⟫ₘ := by
|
||||
conv_lhs =>
|
||||
erw [← map_smul]
|
||||
rw [← smul_smul]
|
||||
rw [← tmul_smul, ← smul_tmul, ← tmul_smul, genBoost_mul_one_plus_contr,
|
||||
genBoost_mul_one_plus_contr]
|
||||
simp only [add_smul, one_smul, tmul_add, map_add, smul_add, tmul_sub, sub_tmul, add_tmul,
|
||||
smul_tmul, tmul_smul, map_sub, map_smul, smul_eq_mul, contr_self, mul_one]
|
||||
rw [contrContrContractField.symm v.1.1 u, contrContrContractField.symm v.1.1 x,
|
||||
contrContrContractField.symm u.1.1 x]
|
||||
simp
|
||||
ring
|
||||
have hn (a : ℝ ) {b c : ℝ} (h : a ≠ 0) (hab : a * b = a * c ) : b = c := by
|
||||
simp_all only [smul_eq_mul, ne_eq, mul_eq_mul_left_iff, or_false]
|
||||
refine hn _ ?_ h1
|
||||
simpa using (FuturePointing.one_add_metric_non_zero u v)
|
||||
|
||||
/-- A generalised boost as an element of the Lorentz Group. -/
|
||||
def toLorentz (u v : FuturePointing d) : LorentzGroup d :=
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue