feat: Make MulActionTensor
This commit is contained in:
parent
99f4e85839
commit
a65fb06605
11 changed files with 177 additions and 55 deletions
|
@ -73,11 +73,14 @@ import HepLean.SpaceTime.LorentzGroup.Rotations
|
|||
import HepLean.SpaceTime.LorentzTensor.Basic
|
||||
import HepLean.SpaceTime.LorentzTensor.Contractions
|
||||
import HepLean.SpaceTime.LorentzTensor.Fin
|
||||
import HepLean.SpaceTime.LorentzTensor.LorentzTensorStruct
|
||||
import HepLean.SpaceTime.LorentzTensor.MulActionTensor
|
||||
import HepLean.SpaceTime.LorentzTensor.Notation
|
||||
import HepLean.SpaceTime.LorentzTensor.RisingLowering
|
||||
import HepLean.SpaceTime.LorentzTensor.Real.Basic
|
||||
import HepLean.SpaceTime.LorentzVector.AsSelfAdjointMatrix
|
||||
import HepLean.SpaceTime.LorentzVector.Basic
|
||||
import HepLean.SpaceTime.LorentzVector.Contraction
|
||||
import HepLean.SpaceTime.LorentzVector.Covariant
|
||||
import HepLean.SpaceTime.LorentzVector.LorentzAction
|
||||
import HepLean.SpaceTime.LorentzVector.NormOne
|
||||
import HepLean.SpaceTime.MinkowskiMetric
|
||||
import HepLean.SpaceTime.SL2C.Basic
|
||||
|
|
|
@ -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 ?_ ?_
|
||||
|
|
|
@ -46,6 +46,7 @@ def contrDualMidAux {V1 V2 V3 V4 : Type} [AddCommMonoid V1] [AddCommMonoid V2] [
|
|||
(V4 ⊗[R] V1) ⊗[R] (V2 ⊗[R] V3) →ₗ[R] V4 ⊗[R] V3 :=
|
||||
(TensorProduct.map (LinearEquiv.refl R V4).toLinearMap (contrDualLeftAux f)) ∘ₗ
|
||||
(TensorProduct.assoc R _ _ _).toLinearMap
|
||||
|
||||
/-- An initial structure specifying a tensor system (e.g. a system in which you can
|
||||
define real Lorentz tensors or Einstein notation convention). -/
|
||||
structure TensorStructure (R : Type) [CommSemiring R] where
|
||||
|
@ -72,9 +73,9 @@ structure TensorStructure (R : Type) [CommSemiring R] where
|
|||
/-- The unit is a right identity. -/
|
||||
unit_lid : ∀ μ (x : ColorModule μ),
|
||||
TensorProduct.rid R _
|
||||
(TensorProduct.map (LinearEquiv.refl R (ColorModule μ)).toLinearMap
|
||||
(contrDual μ ∘ₗ (TensorProduct.comm R _ _).toLinearMap )
|
||||
((TensorProduct.assoc R _ _ _) (unit μ ⊗ₜ[R] x ))) = x
|
||||
(TensorProduct.map (LinearEquiv.refl R (ColorModule μ)).toLinearMap
|
||||
(contrDual μ ∘ₗ (TensorProduct.comm R _ _).toLinearMap)
|
||||
((TensorProduct.assoc R _ _ _) (unit μ ⊗ₜ[R] x))) = x
|
||||
/-- The metric for a given color. -/
|
||||
metric : (μ : Color) → ColorModule μ ⊗[R] ColorModule μ
|
||||
/-- The metric contracted with its dual is the unit. -/
|
||||
|
|
|
@ -7,7 +7,7 @@ import HepLean.SpaceTime.LorentzTensor.Contractions
|
|||
import Mathlib.RepresentationTheory.Basic
|
||||
/-!
|
||||
|
||||
# Structure to form Lorentz-style Tensor
|
||||
# Group actions on tensor structures
|
||||
|
||||
-/
|
||||
|
||||
|
@ -18,29 +18,65 @@ open TensorProduct
|
|||
variable {R : Type} [CommSemiring R]
|
||||
|
||||
/-! TODO: Add preservation of the unit, and the metric. -/
|
||||
/-- A `DualizeTensorStructure` with a group action. -/
|
||||
structure LorentzTensorStructure (R : Type) [CommSemiring R]
|
||||
(G : Type) [Group G] extends TensorStructure R where
|
||||
|
||||
class MulActionTensor (G : Type) [Monoid G] (𝓣 : TensorStructure R) where
|
||||
/-- For each color `μ` a representation of `G` on `ColorModule μ`. -/
|
||||
repColorModule : (μ : Color) → Representation R G (ColorModule μ)
|
||||
repColorModule : (μ : 𝓣.Color) → Representation R G (𝓣.ColorModule μ)
|
||||
/-- The contraction of a vector with its dual is invariant under the group action. -/
|
||||
contrDual_inv : ∀ μ g, contrDual μ ∘ₗ
|
||||
TensorProduct.map (repColorModule μ g) (repColorModule (τ μ) g) = contrDual μ
|
||||
contrDual_inv : ∀ μ g, 𝓣.contrDual μ ∘ₗ
|
||||
TensorProduct.map (repColorModule μ g) (repColorModule (𝓣.τ μ) g) = 𝓣.contrDual μ
|
||||
|
||||
namespace LorentzTensorStructure
|
||||
open TensorStructure
|
||||
namespace MulActionTensor
|
||||
|
||||
variable {G : Type} [Group G]
|
||||
variable {G H : Type} [Group G] [Group H]
|
||||
|
||||
variable (𝓣 : LorentzTensorStructure R G)
|
||||
variable (𝓣 : TensorStructure R) [MulActionTensor G 𝓣]
|
||||
|
||||
variable {d : ℕ} {X Y Y' Z : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z]
|
||||
{cX cX2 : X → 𝓣.Color} {cY : Y → 𝓣.Color} {cZ : Z → 𝓣.Color} {cY' : Y' → 𝓣.Color} {μ ν: 𝓣.Color}
|
||||
|
||||
/-!
|
||||
|
||||
# Instances of `MulActionTensor`
|
||||
|
||||
-/
|
||||
|
||||
def compHom (f : H →* G) : MulActionTensor H 𝓣 where
|
||||
repColorModule μ := MonoidHom.comp (repColorModule μ) f
|
||||
contrDual_inv μ h := by
|
||||
simp only [MonoidHom.coe_comp, Function.comp_apply]
|
||||
rw [contrDual_inv]
|
||||
|
||||
def trivial : MulActionTensor G 𝓣 where
|
||||
repColorModule μ := Representation.trivial R
|
||||
contrDual_inv μ g := by
|
||||
simp only [Representation.trivial, MonoidHom.one_apply, TensorProduct.map_one]
|
||||
rfl
|
||||
|
||||
end MulActionTensor
|
||||
|
||||
namespace TensorStructure
|
||||
open TensorStructure
|
||||
open MulActionTensor
|
||||
|
||||
variable {G : Type} [Group G]
|
||||
|
||||
variable (𝓣 : TensorStructure R) [MulActionTensor G 𝓣]
|
||||
|
||||
variable {d : ℕ} {X Y Y' Z : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z]
|
||||
{cX cX2 : X → 𝓣.Color} {cY : Y → 𝓣.Color} {cZ : Z → 𝓣.Color} {cY' : Y' → 𝓣.Color} {μ ν: 𝓣.Color}
|
||||
|
||||
/-!
|
||||
|
||||
## Representation of tensor products
|
||||
|
||||
-/
|
||||
|
||||
/-- The representation of the group `G` on the vector space of tensors. -/
|
||||
def rep : Representation R G (𝓣.Tensor cX) where
|
||||
toFun g := PiTensorProduct.map (fun x => 𝓣.repColorModule (cX x) g)
|
||||
toFun g := PiTensorProduct.map (fun x => repColorModule (cX x) g)
|
||||
map_one' := by
|
||||
simp_all only [_root_.map_one, PiTensorProduct.map_one]
|
||||
map_mul' g g' := by
|
||||
|
@ -50,15 +86,15 @@ def rep : Representation R G (𝓣.Tensor cX) where
|
|||
local infixl:78 " • " => 𝓣.rep
|
||||
|
||||
lemma repColorModule_colorModuleCast_apply (h : μ = ν) (g : G) (x : 𝓣.ColorModule μ) :
|
||||
(𝓣.repColorModule ν g) (𝓣.colorModuleCast h x) =
|
||||
(𝓣.colorModuleCast h) (𝓣.repColorModule μ g x) := by
|
||||
(repColorModule ν g) (𝓣.colorModuleCast h x) =
|
||||
(𝓣.colorModuleCast h) (repColorModule μ g x) := by
|
||||
subst h
|
||||
simp [colorModuleCast]
|
||||
|
||||
@[simp]
|
||||
lemma repColorModule_colorModuleCast (h : μ = ν) (g : G) :
|
||||
(𝓣.repColorModule ν g) ∘ₗ (𝓣.colorModuleCast h).toLinearMap =
|
||||
(𝓣.colorModuleCast h).toLinearMap ∘ₗ (𝓣.repColorModule μ g) := by
|
||||
(repColorModule ν g) ∘ₗ (𝓣.colorModuleCast h).toLinearMap =
|
||||
(𝓣.colorModuleCast h).toLinearMap ∘ₗ (repColorModule μ g) := by
|
||||
apply LinearMap.ext
|
||||
intro x
|
||||
simp [repColorModule_colorModuleCast_apply]
|
||||
|
@ -73,7 +109,7 @@ lemma rep_mapIso (e : X ≃ Y) (h : cX = cY ∘ e) (g : G) :
|
|||
Function.comp_apply]
|
||||
erw [mapIso_tprod]
|
||||
simp [rep, repColorModule_colorModuleCast_apply]
|
||||
change (PiTensorProduct.map fun x => (𝓣.repColorModule (cY x)) g)
|
||||
change (PiTensorProduct.map fun x => (repColorModule (cY x)) g)
|
||||
((PiTensorProduct.tprod R) fun i => (𝓣.colorModuleCast _) (x (e.symm i))) =
|
||||
(𝓣.mapIso e h) ((PiTensorProduct.map _) ((PiTensorProduct.tprod R) x))
|
||||
rw [PiTensorProduct.map_tprod, PiTensorProduct.map_tprod, mapIso_tprod]
|
||||
|
@ -92,7 +128,7 @@ lemma rep_mapIso_apply (e : X ≃ Y) (h : cX = cY ∘ e) (g : G) (x : 𝓣.Tenso
|
|||
@[simp]
|
||||
lemma rep_tprod (g : G) (f : (i : X) → 𝓣.ColorModule (cX i)) :
|
||||
g • (PiTensorProduct.tprod R f) = PiTensorProduct.tprod R (fun x =>
|
||||
𝓣.repColorModule (cX x) g (f x)) := by
|
||||
repColorModule (cX x) g (f x)) := by
|
||||
simp [rep]
|
||||
change (PiTensorProduct.map _) ((PiTensorProduct.tprod R) f) = _
|
||||
rw [PiTensorProduct.map_tprod]
|
||||
|
@ -163,7 +199,7 @@ lemma contrAll_rep {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X ≃ Y) (
|
|||
apply congrArg
|
||||
funext x
|
||||
rw [← repColorModule_colorModuleCast_apply]
|
||||
nth_rewrite 2 [← 𝓣.contrDual_inv (c x) g]
|
||||
nth_rewrite 2 [← contrDual_inv (c x) g]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
|
@ -177,9 +213,9 @@ lemma contrAll_rep_apply {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X
|
|||
lemma contrAll_rep_tmul {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X ≃ Y) (h : c = 𝓣.τ ∘ d ∘ e)
|
||||
(g : G) (x : 𝓣.Tensor c) (y : 𝓣.Tensor d) :
|
||||
𝓣.contrAll e h ((g • x) ⊗ₜ[R] (g • y)) = 𝓣.contrAll e h (x ⊗ₜ[R] y) := by
|
||||
nth_rewrite 2 [← contrAll_rep_apply]
|
||||
nth_rewrite 2 [← @contrAll_rep_apply R _ G]
|
||||
rfl
|
||||
|
||||
end LorentzTensorStructure
|
||||
end TensorStructure
|
||||
|
||||
end
|
|
@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.SpaceTime.LorentzVector.Contraction
|
||||
import HepLean.SpaceTime.LorentzTensor.Basic
|
||||
import HepLean.SpaceTime.LorentzVector.LorentzAction
|
||||
import HepLean.SpaceTime.LorentzTensor.MulActionTensor
|
||||
/-!
|
||||
|
||||
# Real Lorentz tensors
|
||||
|
@ -85,4 +86,15 @@ def realLorentzTensor (d : ℕ) : TensorStructure ℝ where
|
|||
| realTensor.ColorType.up => asProdLorentzVector_contr_asCovariantProdLorentzVector
|
||||
| realTensor.ColorType.down => asProdCovariantLorentzVector_contr_asProdLorentzVector
|
||||
|
||||
/-- The action of the Lorentz group on real Lorentz tensors. -/
|
||||
instance : MulActionTensor (LorentzGroup d) (realLorentzTensor d) where
|
||||
repColorModule μ :=
|
||||
match μ with
|
||||
| .up => LorentzVector.rep
|
||||
| .down => CovariantLorentzVector.rep
|
||||
contrDual_inv μ _ :=
|
||||
match μ with
|
||||
| .up => TensorProduct.ext' (fun _ _ => LorentzVector.contrUpDown_invariant_lorentzAction)
|
||||
| .down => TensorProduct.ext' (fun _ _ => LorentzVector.contrDownUp_invariant_lorentzAction)
|
||||
|
||||
end
|
||||
|
|
|
@ -86,7 +86,7 @@ lemma decomp_stdBasis (v : LorentzVector d) : ∑ i, v i • e i = v := by
|
|||
|
||||
@[simp]
|
||||
lemma decomp_stdBasis' (v : LorentzVector d) :
|
||||
v (Sum.inl 0) • e (Sum.inl 0) + ∑ a₂ : Fin d, v (Sum.inr a₂) • e (Sum.inr a₂) = v := by
|
||||
v (Sum.inl 0) • e (Sum.inl 0) + ∑ a₂ : Fin d, v (Sum.inr a₂) • e (Sum.inr a₂) = v := by
|
||||
trans ∑ i, v i • e i
|
||||
simp only [Fin.isValue, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero,
|
||||
Finset.sum_singleton]
|
||||
|
|
|
@ -3,7 +3,7 @@ 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.SpaceTime.LorentzVector.Basic
|
||||
import HepLean.SpaceTime.LorentzVector.LorentzAction
|
||||
import HepLean.SpaceTime.LorentzVector.Covariant
|
||||
import HepLean.SpaceTime.LorentzTensor.Basic
|
||||
/-!
|
||||
|
@ -24,9 +24,9 @@ open TensorProduct
|
|||
|
||||
namespace LorentzVector
|
||||
|
||||
open Matrix
|
||||
variable {d : ℕ} (v : LorentzVector d)
|
||||
|
||||
|
||||
def contrUpDownBi : LorentzVector d →ₗ[ℝ] CovariantLorentzVector d →ₗ[ℝ] ℝ where
|
||||
toFun v := {
|
||||
toFun := fun w => ∑ i, v i * w i,
|
||||
|
@ -60,6 +60,10 @@ def contrUpDownBi : LorentzVector d →ₗ[ℝ] CovariantLorentzVector d →ₗ[
|
|||
def contrUpDown : LorentzVector d ⊗[ℝ] CovariantLorentzVector d →ₗ[ℝ] ℝ :=
|
||||
TensorProduct.lift contrUpDownBi
|
||||
|
||||
lemma contrUpDown_tmul_eq_dotProduct {x : LorentzVector d} {y : CovariantLorentzVector d} :
|
||||
contrUpDown (x ⊗ₜ[ℝ] y) = x ⬝ᵥ y := by
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma contrUpDown_stdBasis_left (x : LorentzVector d) (i : Fin 1 ⊕ Fin d) :
|
||||
contrUpDown (x ⊗ₜ[ℝ] (CovariantLorentzVector.stdBasis i)) = x i := by
|
||||
|
@ -92,16 +96,26 @@ lemma contrUpDown_stdBasis_right (x : CovariantLorentzVector d) (i : Fin 1 ⊕ F
|
|||
def contrDownUp : CovariantLorentzVector d ⊗[ℝ] LorentzVector d →ₗ[ℝ] ℝ :=
|
||||
contrUpDown ∘ₗ (TensorProduct.comm ℝ _ _).toLinearMap
|
||||
|
||||
lemma contrDownUp_tmul_eq_dotProduct {x : CovariantLorentzVector d} {y : LorentzVector d} :
|
||||
contrDownUp (x ⊗ₜ[ℝ] y) = x ⬝ᵥ y := by
|
||||
rw [dotProduct_comm x y]
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## The unit of the contraction
|
||||
|
||||
-/
|
||||
|
||||
def unitUp : LorentzVector d ⊗[ℝ] CovariantLorentzVector d :=
|
||||
∑ i, LorentzVector.stdBasis i ⊗ₜ[ℝ] CovariantLorentzVector.stdBasis i
|
||||
|
||||
lemma unitUp_lid (x : LorentzVector d) :
|
||||
TensorProduct.rid ℝ _
|
||||
(TensorProduct.map (LinearEquiv.refl ℝ (_)).toLinearMap
|
||||
(contrUpDown ∘ₗ (TensorProduct.comm ℝ _ _).toLinearMap )
|
||||
((TensorProduct.assoc ℝ _ _ _) (unitUp ⊗ₜ[ℝ] x ))) = x := by
|
||||
simp only [LinearEquiv.refl_toLinearMap, unitUp]
|
||||
(TensorProduct.map (LinearEquiv.refl ℝ _).toLinearMap
|
||||
(contrUpDown ∘ₗ (TensorProduct.comm ℝ _ _).toLinearMap)
|
||||
((TensorProduct.assoc ℝ _ _ _) (unitUp ⊗ₜ[ℝ] x))) = x := by
|
||||
simp only [LinearEquiv.refl_toLinearMap, unitUp]
|
||||
rw [sum_tmul]
|
||||
simp only [Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
||||
Finset.sum_singleton, map_add, assoc_tmul, map_sum, map_tmul, LinearMap.id_coe, id_eq,
|
||||
|
@ -113,8 +127,8 @@ def unitDown : CovariantLorentzVector d ⊗[ℝ] LorentzVector d :=
|
|||
|
||||
lemma unitDown_lid (x : CovariantLorentzVector d) :
|
||||
TensorProduct.rid ℝ _
|
||||
(TensorProduct.map (LinearEquiv.refl ℝ (_)).toLinearMap
|
||||
(contrDownUp ∘ₗ (TensorProduct.comm ℝ _ _).toLinearMap )
|
||||
(TensorProduct.map (LinearEquiv.refl ℝ _).toLinearMap
|
||||
(contrDownUp ∘ₗ (TensorProduct.comm ℝ _ _).toLinearMap)
|
||||
((TensorProduct.assoc ℝ _ _ _) (unitDown ⊗ₜ[ℝ] x))) = x := by
|
||||
simp only [LinearEquiv.refl_toLinearMap, unitDown]
|
||||
rw [sum_tmul]
|
||||
|
@ -123,6 +137,29 @@ lemma unitDown_lid (x : CovariantLorentzVector d) :
|
|||
id_eq, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, comm_tmul,
|
||||
contrUpDown_stdBasis_right, rid_tmul, CovariantLorentzVector.decomp_stdBasis']
|
||||
|
||||
/-!
|
||||
|
||||
# Contractions and the Lorentz actions
|
||||
|
||||
-/
|
||||
open Matrix
|
||||
|
||||
@[simp]
|
||||
lemma contrUpDown_invariant_lorentzAction : @contrUpDown d ((LorentzVector.rep g) x ⊗ₜ[ℝ]
|
||||
(CovariantLorentzVector.rep g) y) = contrUpDown (x ⊗ₜ[ℝ] y) := by
|
||||
rw [contrUpDown_tmul_eq_dotProduct, contrUpDown_tmul_eq_dotProduct]
|
||||
simp only [rep_apply, CovariantLorentzVector.rep_apply]
|
||||
rw [Matrix.dotProduct_mulVec, vecMul_transpose, mulVec_mulVec]
|
||||
simp only [LorentzGroup.subtype_inv_mul, one_mulVec]
|
||||
|
||||
@[simp]
|
||||
lemma contrDownUp_invariant_lorentzAction : @contrDownUp d ((CovariantLorentzVector.rep g) x ⊗ₜ[ℝ]
|
||||
(LorentzVector.rep g) y) = contrDownUp (x ⊗ₜ[ℝ] y) := by
|
||||
rw [contrDownUp_tmul_eq_dotProduct, contrDownUp_tmul_eq_dotProduct]
|
||||
rw [dotProduct_comm, dotProduct_comm x y]
|
||||
simp only [rep_apply, CovariantLorentzVector.rep_apply]
|
||||
rw [Matrix.dotProduct_mulVec, vecMul_transpose, mulVec_mulVec]
|
||||
simp only [LorentzGroup.subtype_inv_mul, one_mulVec]
|
||||
|
||||
end LorentzVector
|
||||
|
||||
|
@ -132,7 +169,7 @@ open scoped minkowskiMatrix
|
|||
variable {d : ℕ}
|
||||
|
||||
def asProdLorentzVector : LorentzVector d ⊗[ℝ] LorentzVector d :=
|
||||
∑ μ, η μ μ • (LorentzVector.stdBasis μ ⊗ₜ[ℝ] LorentzVector.stdBasis μ)
|
||||
∑ μ, η μ μ • (LorentzVector.stdBasis μ ⊗ₜ[ℝ] LorentzVector.stdBasis μ)
|
||||
|
||||
def asProdCovariantLorentzVector : CovariantLorentzVector d ⊗[ℝ] CovariantLorentzVector d :=
|
||||
∑ μ, η μ μ • (CovariantLorentzVector.stdBasis μ ⊗ₜ[ℝ] CovariantLorentzVector.stdBasis μ)
|
||||
|
@ -173,16 +210,15 @@ lemma asProdLorentzVector_contr_asCovariantProdLorentzVector :
|
|||
refine Finset.sum_congr rfl (fun μ _ => ?_)
|
||||
rw [← tmul_smul, TensorProduct.assoc_tmul]
|
||||
simp only [map_tmul, LinearMap.id_coe, id_eq, contrLeft_asProdCovariantLorentzVector]
|
||||
rw [tmul_sum, Finset.sum_eq_single_of_mem μ]
|
||||
rw [tmul_smul]
|
||||
change (η μ μ * (η μ μ * e μ μ)) • e μ ⊗ₜ[ℝ] CovariantLorentzVector.stdBasis μ = _
|
||||
rw [tmul_sum, Finset.sum_eq_single_of_mem μ, tmul_smul]
|
||||
change (η μ μ * (η μ μ * e μ μ)) • e μ ⊗ₜ[ℝ] CovariantLorentzVector.stdBasis μ = _
|
||||
rw [LorentzVector.stdBasis]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp
|
||||
simp only [LinearMap.stdBasis_same, mul_one, η_apply_mul_η_apply_diag, one_smul]
|
||||
exact Finset.mem_univ μ
|
||||
intro ν _ hμν
|
||||
rw [tmul_smul]
|
||||
change (η ν ν * (η μ μ * e μ ν)) • (e μ ⊗ₜ[ℝ] CovariantLorentzVector.stdBasis ν) = _
|
||||
change (η ν ν * (η μ μ * e μ ν)) • (e μ ⊗ₜ[ℝ] CovariantLorentzVector.stdBasis ν) = _
|
||||
rw [LorentzVector.stdBasis]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp only [LinearMap.stdBasis_apply', mul_ite, mul_one, mul_zero, ite_smul, zero_smul,
|
||||
|
@ -197,15 +233,12 @@ lemma asProdCovariantLorentzVector_contr_asProdLorentzVector :
|
|||
LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply]
|
||||
rw [sum_tmul, map_sum, map_sum, unitDown]
|
||||
refine Finset.sum_congr rfl (fun μ _ => ?_)
|
||||
rw [smul_tmul,]
|
||||
rw [TensorProduct.assoc_tmul]
|
||||
rw [smul_tmul, TensorProduct.assoc_tmul]
|
||||
simp only [tmul_smul, LinearMapClass.map_smul, map_tmul, LinearMap.id_coe, id_eq,
|
||||
contrLeft_asProdLorentzVector]
|
||||
rw [tmul_sum, Finset.sum_eq_single_of_mem μ]
|
||||
rw [tmul_smul, smul_smul]
|
||||
rw [LorentzVector.stdBasis]
|
||||
rw [tmul_sum, Finset.sum_eq_single_of_mem μ, tmul_smul, smul_smul, LorentzVector.stdBasis]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp
|
||||
simp only [LinearMap.stdBasis_same, mul_one, η_apply_mul_η_apply_diag, one_smul]
|
||||
exact Finset.mem_univ μ
|
||||
intro ν _ hμν
|
||||
rw [tmul_smul]
|
||||
|
@ -217,5 +250,4 @@ lemma asProdCovariantLorentzVector_contr_asProdLorentzVector :
|
|||
|
||||
end minkowskiMatrix
|
||||
|
||||
|
||||
end
|
||||
|
|
|
@ -61,7 +61,7 @@ lemma decomp_stdBasis (v : CovariantLorentzVector d) : ∑ i, v i • stdBasis i
|
|||
|
||||
@[simp]
|
||||
lemma decomp_stdBasis' (v : CovariantLorentzVector d) :
|
||||
v (Sum.inl 0) • stdBasis (Sum.inl 0) + ∑ a₂ : Fin d, v (Sum.inr a₂) • stdBasis (Sum.inr a₂) = v := by
|
||||
v (Sum.inl 0) • stdBasis (Sum.inl 0) + ∑ a₂ : Fin d, v (Sum.inr a₂) • stdBasis (Sum.inr a₂) = v := by
|
||||
trans ∑ i, v i • stdBasis i
|
||||
simp only [Fin.isValue, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero,
|
||||
Finset.sum_singleton]
|
||||
|
@ -82,6 +82,17 @@ def rep : Representation ℝ (LorentzGroup d) (CovariantLorentzVector d) where
|
|||
simp only [mul_inv_rev, lorentzGroupIsGroup_inv, LorentzGroup.transpose_mul,
|
||||
lorentzGroupIsGroup_mul_coe, map_mul]
|
||||
|
||||
open Matrix in
|
||||
@[simp]
|
||||
lemma rep_apply (g : LorentzGroup d) : rep g v = (g.1⁻¹)ᵀ *ᵥ v := by
|
||||
trans (LorentzGroup.transpose g⁻¹) *ᵥ v
|
||||
rfl
|
||||
apply congrFun
|
||||
apply congrArg
|
||||
simp only [LorentzGroup.transpose, lorentzGroupIsGroup_inv, transpose_inj]
|
||||
rw [← LorentzGroup.coe_inv]
|
||||
rfl
|
||||
|
||||
end CovariantLorentzVector
|
||||
|
||||
end
|
||||
|
|
|
@ -25,6 +25,10 @@ def rep : Representation ℝ (LorentzGroup d) (LorentzVector d) where
|
|||
map_mul' x y := by
|
||||
simp only [lorentzGroupIsGroup_mul_coe, map_mul]
|
||||
|
||||
open Matrix in
|
||||
@[simp]
|
||||
lemma rep_apply (g : LorentzGroup d) : rep g v = g *ᵥ v := rfl
|
||||
|
||||
end LorentzVector
|
||||
|
||||
end
|
||||
|
|
|
@ -68,10 +68,17 @@ def checkMissingImports (modData : ModuleData) (reqImports : Array Name) :
|
|||
IO Bool := do
|
||||
let names : HashSet Name := HashSet.ofArray (modData.imports.map (·.module))
|
||||
let mut warned := false
|
||||
for req in reqImports do
|
||||
if !names.contains req then
|
||||
IO.print s!"File {req} is not imported. \n"
|
||||
warned := true
|
||||
let nameArray := reqImports.filterMap (
|
||||
fun req => if !names.contains req then
|
||||
some req
|
||||
else
|
||||
none)
|
||||
if nameArray.size ≠ 0 then
|
||||
let nameArraySort := nameArray.qsort (·.toString < ·.toString)
|
||||
IO.print s!"Files are not imported add the following to the `HepLean` file: \n"
|
||||
for name in nameArraySort do
|
||||
IO.print s!"import {name}\n"
|
||||
warned := true
|
||||
pure warned
|
||||
|
||||
def main (_ : List String) : IO UInt32 := do
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue