refactor: Lint

This commit is contained in:
jstoobysmith 2024-07-26 14:54:09 -04:00
parent 1f51e718f2
commit 62fdab3ace
9 changed files with 37 additions and 1937 deletions

View file

@ -1,559 +0,0 @@
/-
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 Mathlib.Logic.Function.CompTypeclasses
import Mathlib.Data.Real.Basic
import Mathlib.Data.Fintype.BigOperators
import Mathlib.Logic.Equiv.Fin
import Mathlib.Tactic.FinCases
import Mathlib.Logic.Equiv.Fintype
import Mathlib.Algebra.Module.Pi
import Mathlib.Algebra.Module.Equiv
import Mathlib.Algebra.Module.LinearMap.Basic
import Mathlib.LinearAlgebra.TensorProduct.Basic
import Mathlib.LinearAlgebra.TensorProduct.Basis
import Mathlib.LinearAlgebra.PiTensorProduct
import HepLean.Mathematics.PiTensorProduct
/-!
# Real Lorentz Tensors
In this file we define real Lorentz tensors.
We implicitly follow the definition of a modular operad.
This will relation should be made explicit in the future.
## References
-- For modular operads see: [Raynor][raynor2021graphical]
-/
/-! TODO: Relate the constructions here to `PiTensorProduct`. -/
/-! TODO: Generalize to maps into Lorentz tensors. -/
section PiTensorProductResults
variable {ι ι₂ ι₃ : Type*}
variable {R : Type*} [CommSemiring R]
variable {R₁ R₂ : Type*}
variable {s : ι → Type*} [∀ i, AddCommMonoid (s i)] [∀ i, Module R (s i)]
variable {M : Type*} [AddCommMonoid M] [Module R M]
variable {E : Type*} [AddCommMonoid E] [Module R E]
variable {F : Type*} [AddCommMonoid F]
end PiTensorProductResults
open TensorProduct
noncomputable section
/-- The possible `color` of an index for a RealLorentzTensor.
There are two possiblities, `up` and `down`. -/
inductive RealLorentzTensor.Color where
| up : RealLorentzTensor.Color
| down : RealLorentzTensor.Color
/-- To set of allowed index values associated to a color of index. -/
def RealLorentzTensor.ColorIndex (d : ) (μ : RealLorentzTensor.Color) : Type :=
match μ with
| RealLorentzTensor.Color.up => Fin 1 ⊕ Fin d
| RealLorentzTensor.Color.down => Fin 1 ⊕ Fin d
/-- The instance of a finite type on the set of allowed index values for a given color. -/
instance (d : ) (μ : RealLorentzTensor.Color) : Fintype (RealLorentzTensor.ColorIndex d μ) :=
match μ with
| RealLorentzTensor.Color.up => instFintypeSum (Fin 1) (Fin d)
| RealLorentzTensor.Color.down => instFintypeSum (Fin 1) (Fin d)
/-- The color index set for each color has a decidable equality. -/
instance (d : ) (μ : RealLorentzTensor.Color) : DecidableEq (RealLorentzTensor.ColorIndex d μ) :=
match μ with
| RealLorentzTensor.Color.up => instDecidableEqSum
| RealLorentzTensor.Color.down => instDecidableEqSum
def RealLorentzTensor.ColorModule (d : ) (μ : RealLorentzTensor.Color) : Type :=
RealLorentzTensor.ColorIndex d μ →
instance (d : ) (μ : RealLorentzTensor.Color) :
AddCommMonoid (RealLorentzTensor.ColorModule d μ) :=
Pi.addCommMonoid
instance (d : ) (μ : RealLorentzTensor.Color) : Module (RealLorentzTensor.ColorModule d μ) :=
Pi.module _ _ _
instance (d : ) (μ : RealLorentzTensor.Color) : AddCommGroup (RealLorentzTensor.ColorModule d μ) :=
Pi.addCommGroup
/-- Real Lorentz tensors. -/
def RealLorentzTensor {X : Type} (d : ) (c : X → RealLorentzTensor.Color) : Type :=
⨂[] i : X, RealLorentzTensor.ColorModule d (c i)
instance {X : Type} (d : ) (c : X → RealLorentzTensor.Color) :
AddCommMonoid (RealLorentzTensor d c) :=
PiTensorProduct.instAddCommMonoid fun i => RealLorentzTensor.ColorModule d (c i)
instance {X : Type} (d : ) (c : X → RealLorentzTensor.Color) :
Module (RealLorentzTensor d c) := PiTensorProduct.instModule
instance {X : Type} (d : ) (c : X → RealLorentzTensor.Color) :
AddCommGroup (RealLorentzTensor d c) := PiTensorProduct.instAddCommGroup
namespace RealLorentzTensor
variable {d : } {X Y Z : Type} (c : X → Color)
[Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] [Fintype Z] [DecidableEq Z]
/-- An `IndexValue` is a set of actual values an index can take. e.g. for a
3-tensor (0, 1, 2). -/
def IndexValue {X : Type} (d : ) (c : X → RealLorentzTensor.Color) :
Type 0 := (x : X) → ColorIndex d (c x)
instance [Fintype X] [DecidableEq X] : Fintype (IndexValue d c) := Pi.fintype
instance [Fintype X] : DecidableEq (IndexValue d c) :=
Fintype.decidablePiFintype
def basisColorModule {d : } (μ : Color) :
Basis (ColorIndex d μ) (ColorModule d μ) := Pi.basisFun _ _
def basis (d : ) (c : X → RealLorentzTensor.Color) :
Basis (IndexValue d c) (RealLorentzTensor d c) :=
Basis.piTensorProduct (fun x => basisColorModule (c x))
abbrev PiModule (d : ) (c : X → Color) := IndexValue d c →
def asPi {d : } {c : X → Color} :
RealLorentzTensor d c ≃ₗ[] PiModule d c :=
(basis d c).repr ≪≫ₗ
Finsupp.linearEquivFunOnFinite (IndexValue d c)
/-!
## Colors
-/
/-- The involution acting on colors. -/
def τ : Color → Color
| Color.up => Color.down
| Color.down => Color.up
/-- The map τ is an involution. -/
@[simp]
lemma τ_involutive : Function.Involutive τ := by
intro x
cases x <;> rfl
lemma color_eq_dual_symm {μ ν : Color} (h : μ = τ ν) : ν = τ μ :=
(Function.Involutive.eq_iff τ_involutive).mp h.symm
lemma color_comp_τ_symm {c1 c2 : X → Color} (h : c1 = τ ∘ c2) : c2 = τ ∘ c1 :=
funext (fun x => color_eq_dual_symm (congrFun h x))
/-- An equivalence of `ColorIndex` types given an equality of colors. -/
def colorIndexCast {d : } {μ₁ μ₂ : Color} (h : μ₁ = μ₂) :
ColorIndex d μ₁ ≃ ColorIndex d μ₂ :=
Equiv.cast (congrArg (ColorIndex d) h)
@[simp]
lemma colorIndexCast_symm {d : } {μ₁ μ₂ : Color} (h : μ₁ = μ₂) :
(@colorIndexCast d _ _ h).symm = colorIndexCast h.symm := by
rfl
/-- An equivalence of `ColorIndex` between that of a color and that of its dual.
I.e. the allowed index values for a color and its dual are equivalent. -/
def colorIndexDualCastSelf {d : } {μ : Color}:
ColorIndex d μ ≃ ColorIndex d (τ μ) where
toFun x :=
match μ with
| Color.up => x
| Color.down => x
invFun x :=
match μ with
| Color.up => x
| Color.down => x
left_inv x := by cases μ <;> rfl
right_inv x := by cases μ <;> rfl
/-- An equivalence between the allowed index values of a color and a color dual to it. -/
def colorIndexDualCast {μ ν : Color} (h : μ = τ ν) : ColorIndex d μ ≃ ColorIndex d ν :=
(colorIndexCast h).trans colorIndexDualCastSelf.symm
@[simp]
lemma colorIndexDualCast_symm {μ ν : Color} (h : μ = τ ν) : (colorIndexDualCast h).symm =
@colorIndexDualCast d _ _ ((Function.Involutive.eq_iff τ_involutive).mp h.symm) := by
match μ, ν with
| Color.up, Color.down => rfl
| Color.down, Color.up => rfl
@[simp]
lemma colorIndexDualCast_trans {μ ν ξ : Color} (h : μ = τ ν) (h' : ν = τ ξ) :
(@colorIndexDualCast d _ _ h).trans (colorIndexDualCast h') =
colorIndexCast (by rw [h, h', τ_involutive]) := by
match μ, ν, ξ with
| Color.up, Color.down, Color.up => rfl
| Color.down, Color.up, Color.down => rfl
@[simp]
lemma colorIndexDualCast_trans_colorsIndexCast {μ ν ξ : Color} (h : μ = τ ν) (h' : ν = ξ) :
(@colorIndexDualCast d _ _ h).trans (colorIndexCast h') =
colorIndexDualCast (by rw [h, h']) := by
match μ, ν, ξ with
| Color.down, Color.up, Color.up => rfl
| Color.up, Color.down, Color.down => rfl
@[simp]
lemma colorIndexCast_trans_colorsIndexDualCast {μ ν ξ : Color} (h : μ = ν) (h' : ν = τ ξ) :
(colorIndexCast h).trans (@colorIndexDualCast d _ _ h') =
colorIndexDualCast (by rw [h, h']) := by
match μ, ν, ξ with
| Color.up, Color.up, Color.down => rfl
| Color.down, Color.down, Color.up => rfl
/-!
## Index values
-/
/-!
## Induced isomorphisms between IndexValue sets
-/
/-- An isomorphism of the type of index values given an isomorphism of sets. -/
def indexValueIso (d : ) (f : X ≃ Y) {i : X → Color} {j : Y → Color} (h : i = j ∘ f) :
IndexValue d i ≃ IndexValue d j :=
(Equiv.piCongrRight (fun μ => colorIndexCast (congrFun h μ))).trans $
Equiv.piCongrLeft (fun y => ColorIndex d (j y)) f
@[simp]
lemma indexValueIso_symm_apply' (d : ) (f : X ≃ Y) {i : X → Color} {j : Y → Color}
(h : i = j ∘ f) (y : IndexValue d j) (x : X) :
(indexValueIso d f h).symm y x = (colorIndexCast (congrFun h x)).symm (y (f x)) := by
rfl
@[simp]
lemma indexValueIso_trans (d : ) (f : X ≃ Y) (g : Y ≃ Z) {i : X → Color}
{j : Y → Color} {k : Z → Color} (h : i = j ∘ f) (h' : j = k ∘ g) :
(indexValueIso d f h).trans (indexValueIso d g h') =
indexValueIso d (f.trans g) (by rw [h, h', Function.comp.assoc]; rfl) := by
have h1 : ((indexValueIso d f h).trans (indexValueIso d g h')).symm =
(indexValueIso d (f.trans g) (by rw [h, h', Function.comp.assoc]; rfl)).symm := by
subst h' h
exact Equiv.coe_inj.mp rfl
simpa only [Equiv.symm_symm] using congrArg (fun e => e.symm) h1
@[simp]
lemma indexValueIso_symm (d : ) (f : X ≃ Y) (h : i = j ∘ f) :
(indexValueIso d f h).symm =
indexValueIso d f.symm ((Equiv.eq_comp_symm f j i).mpr h.symm) := by
ext i : 1
rw [← Equiv.symm_apply_eq]
funext y
rw [indexValueIso_symm_apply', indexValueIso_symm_apply']
simp only [Function.comp_apply, colorIndexCast, Equiv.cast_symm, Equiv.cast_apply, cast_cast]
apply cast_eq_iff_heq.mpr
rw [Equiv.apply_symm_apply]
lemma indexValueIso_eq_symm (d : ) (f : X ≃ Y) (h : i = j ∘ f) :
indexValueIso d f h =
(indexValueIso d f.symm ((Equiv.eq_comp_symm f j i).mpr h.symm)).symm := by
rw [indexValueIso_symm]
rfl
@[simp]
lemma indexValueIso_refl (d : ) (i : X → Color) :
indexValueIso d (Equiv.refl X) (rfl : i = i) = Equiv.refl _ := by
rfl
/-!
## Dual isomorphism for index values
-/
/-- The isomorphism between the index values of a color map and its dual. -/
def indexValueDualIso (d : ) {i : X → Color} {j : Y → Color} (e : X ≃ Y)
(h : i = τ ∘ j ∘ e) : IndexValue d i ≃ IndexValue d j :=
(Equiv.piCongrRight (fun μ => colorIndexDualCast (congrFun h μ))).trans $
Equiv.piCongrLeft (fun y => ColorIndex d (j y)) e
lemma indexValueDualIso_symm_apply' (d : ) {i : X → Color} {j : Y → Color} (e : X ≃ Y)
(h : i = τ ∘ j ∘ e) (y : IndexValue d j) (x : X) :
(indexValueDualIso d e h).symm y x = (colorIndexDualCast (congrFun h x)).symm (y (e x)) := by
rfl
lemma indexValueDualIso_cond_symm {i : X → Color} {j : Y → Color} {e : X ≃ Y}
(h : i = τ ∘ j ∘ e) : j = τ ∘ i ∘ e.symm := by
subst h
simp only [Function.comp.assoc, Equiv.self_comp_symm, CompTriple.comp_eq]
rw [← Function.comp.assoc]
funext a
simp only [τ_involutive, Function.Involutive.comp_self, Function.comp_apply, id_eq]
@[simp]
lemma indexValueDualIso_symm {d : } {i : X → Color} {j : Y → Color} (e : X ≃ Y)
(h : i = τ ∘ j ∘ e) : (indexValueDualIso d e h).symm =
indexValueDualIso d e.symm (indexValueDualIso_cond_symm h) := by
ext i : 1
rw [← Equiv.symm_apply_eq]
funext a
rw [indexValueDualIso_symm_apply', indexValueDualIso_symm_apply']
erw [← Equiv.trans_apply, colorIndexDualCast_symm, colorIndexDualCast_symm,
colorIndexDualCast_trans]
simp only [Function.comp_apply, colorIndexCast, Equiv.cast_apply]
apply cast_eq_iff_heq.mpr
rw [Equiv.apply_symm_apply]
lemma indexValueDualIso_eq_symm {d : } {i : X → Color} {j : Y → Color} (e : X ≃ Y)
(h : i = τ ∘ j ∘ e) :
indexValueDualIso d e h = (indexValueDualIso d e.symm (indexValueDualIso_cond_symm h)).symm := by
rw [indexValueDualIso_symm]
simp
lemma indexValueDualIso_cond_trans {i : X → Color} {j : Y → Color} {k : Z → Color}
{e : X ≃ Y} {f : Y ≃ Z} (h : i = τ ∘ j ∘ e) (h' : j = τ ∘ k ∘ f) :
i = k ∘ (e.trans f) := by
rw [h, h']
funext a
simp only [Function.comp_apply, Equiv.coe_trans]
rw [τ_involutive]
@[simp]
lemma indexValueDualIso_trans {d : } {i : X → Color} {j : Y → Color} {k : Z → Color}
(e : X ≃ Y) (f : Y ≃ Z) (h : i = τ ∘ j ∘ e) (h' : j = τ ∘ k ∘ f) :
(indexValueDualIso d e h).trans (indexValueDualIso d f h') =
indexValueIso d (e.trans f) (indexValueDualIso_cond_trans h h') := by
ext i
rw [Equiv.trans_apply]
rw [← Equiv.eq_symm_apply, ← Equiv.eq_symm_apply, indexValueIso_eq_symm]
funext a
rw [indexValueDualIso_symm_apply', indexValueDualIso_symm_apply',
indexValueIso_symm_apply']
erw [← Equiv.trans_apply]
rw [colorIndexDualCast_symm, colorIndexDualCast_symm, colorIndexDualCast_trans]
simp only [Function.comp_apply, colorIndexCast, Equiv.symm_trans_apply, Equiv.cast_symm,
Equiv.cast_apply, cast_cast]
symm
apply cast_eq_iff_heq.mpr
rw [Equiv.symm_apply_apply, Equiv.symm_apply_apply]
lemma indexValueDualIso_cond_trans_indexValueIso {i : X → Color} {j : Y → Color} {k : Z → Color}
{e : X ≃ Y} {f : Y ≃ Z} (h : i = τ ∘ j ∘ e) (h' : j = k ∘ f) :
i = τ ∘ k ∘ (e.trans f) := by
rw [h, h']
funext a
simp only [Function.comp_apply, Equiv.coe_trans]
@[simp]
lemma indexValueDualIso_trans_indexValueIso {d : } {i : X → Color} {j : Y → Color} {k : Z → Color}
(e : X ≃ Y) (f : Y ≃ Z) (h : i = τ ∘ j ∘ e) (h' : j = k ∘ f) :
(indexValueDualIso d e h).trans (indexValueIso d f h') =
indexValueDualIso d (e.trans f) (indexValueDualIso_cond_trans_indexValueIso h h') := by
ext i
rw [Equiv.trans_apply, ← Equiv.eq_symm_apply]
funext a
rw [ indexValueDualIso_eq_symm, indexValueDualIso_symm_apply',
indexValueIso_symm_apply',indexValueDualIso_eq_symm, indexValueDualIso_symm_apply']
rw [Equiv.symm_apply_eq]
erw [← Equiv.trans_apply, ← Equiv.trans_apply]
simp only [Function.comp_apply, Equiv.symm_trans_apply, colorIndexDualCast_symm,
colorIndexCast_symm, colorIndexCast_trans_colorsIndexDualCast, colorIndexDualCast_trans]
simp only [colorIndexCast, Equiv.cast_apply]
symm
apply cast_eq_iff_heq.mpr
rw [Equiv.symm_apply_apply]
lemma indexValueIso_trans_indexValueDualIso_cond {i : X → Color} {j : Y → Color} {k : Z → Color}
{e : X ≃ Y} {f : Y ≃ Z} (h : i = j ∘ e) (h' : j = τ ∘ k ∘ f) :
i = τ ∘ k ∘ (e.trans f) := by
rw [h, h']
funext a
simp only [Function.comp_apply, Equiv.coe_trans]
@[simp]
lemma indexValueIso_trans_indexValueDualIso {d : } {i : X → Color} {j : Y → Color} {k : Z → Color}
(e : X ≃ Y) (f : Y ≃ Z) (h : i = j ∘ e) (h' : j = τ ∘ k ∘ f) :
(indexValueIso d e h).trans (indexValueDualIso d f h') =
indexValueDualIso d (e.trans f) (indexValueIso_trans_indexValueDualIso_cond h h') := by
ext i
rw [Equiv.trans_apply, ← Equiv.eq_symm_apply, ← Equiv.eq_symm_apply]
funext a
rw [indexValueIso_symm_apply', indexValueDualIso_symm_apply',
indexValueDualIso_eq_symm, indexValueDualIso_symm_apply']
erw [← Equiv.trans_apply, ← Equiv.trans_apply]
simp only [Function.comp_apply, Equiv.symm_trans_apply, colorIndexDualCast_symm,
colorIndexCast_symm, colorIndexDualCast_trans_colorsIndexCast, colorIndexDualCast_trans]
simp only [colorIndexCast, Equiv.cast_apply]
symm
apply cast_eq_iff_heq.mpr
rw [Equiv.symm_apply_apply, Equiv.symm_apply_apply]
/-!
## Mapping isomorphisms on fibers
-/
@[simps!]
def mapIso {d : } (f : X ≃ Y) {i : X → Color} {j : Y → Color} (h : i = j ∘ f) :
RealLorentzTensor d i ≃ₗ[] RealLorentzTensor d j where
toFun F := F ∘ (indexValueIso d f h).symm
invFun F := F ∘ indexValueIso d f h
map_add' F G := by rfl
map_smul' a F := by rfl
left_inv F := by
funext x
simp only [Function.comp_apply]
exact congrArg _ $ Equiv.symm_apply_apply (indexValueIso d f h) x
right_inv F := by
funext x
simp only [Function.comp_apply]
exact congrArg _ $ Equiv.apply_symm_apply (indexValueIso d f h) x
@[simp]
lemma mapIso_trans (d : ) (f : X ≃ Y) (g : Y ≃ Z) {i : X → Color}
{j : Y → Color} {k : Z → Color} (h : i = j ∘ f) (h' : j = k ∘ g) :
(@mapIso _ _ d f _ _ h).trans (mapIso g h') =
mapIso (f.trans g) (by rw [h, h', Function.comp.assoc]; rfl) := by
refine LinearEquiv.toEquiv_inj.mp (Equiv.ext (fun x => (funext (fun a => ?_))))
simp only [LinearEquiv.coe_toEquiv, LinearEquiv.trans_apply, mapIso_apply,
indexValueIso_symm, ← Equiv.trans_apply, indexValueIso_trans]
rfl
lemma mapIso_symm (d : ) (f : X ≃ Y) (h : i = j ∘ f) :
(@mapIso _ _ d f _ _ h).symm = mapIso f.symm ((Equiv.eq_comp_symm f j i).mpr h.symm) := by
refine LinearEquiv.toEquiv_inj.mp (Equiv.ext (fun x => (funext (fun a => ?_))))
simp only [LinearEquiv.coe_toEquiv, mapIso_symm_apply, mapIso_apply, indexValueIso_symm,
Equiv.symm_symm]
/-!
## Sums
-/
/-- An equivalence that splits elements of `IndexValue d (Sum.elim TX TY)` into two components. -/
def indexValueTensorator {X Y : Type} {TX : X → Color} {TY : Y → Color} :
IndexValue d (Sum.elim TX TY) ≃ IndexValue d TX × IndexValue d TY where
toFun i := (fun x => i (Sum.inl x), fun x => i (Sum.inr x))
invFun p := fun c => match c with
| Sum.inl x => (p.1 x)
| Sum.inr x => (p.2 x)
left_inv i := by
simp only [IndexValue]
ext1 x
cases x with
| inl val => rfl
| inr val_1 => rfl
right_inv p := rfl
/-! TODO : Move -/
lemma tensorator_mapIso_cond {cX : X → Color} {cY : Y → Color} {cX' : X' → Color}
{cY' : Y' → Color} {eX : X ≃ X'} {eY : Y ≃ Y'} (hX : cX = cX' ∘ eX) (hY : cY = cY' ∘ eY) :
Sum.elim cX cY = Sum.elim cX' cY' ∘ ⇑(eX.sumCongr eY) := by
subst hX hY
ext1 x
simp_all only [Function.comp_apply, Equiv.sumCongr_apply]
cases x with
| inl val => simp_all only [Sum.elim_inl, Function.comp_apply, Sum.map_inl]
| inr val_1 => simp_all only [Sum.elim_inr, Function.comp_apply, Sum.map_inr]
lemma indexValueTensorator_indexValueMapIso {cX : X → Color} {cY : Y → Color} {cX' : X' → Color}
{cY' : Y' → Color} (eX : X ≃ X') (eY : Y ≃ Y') (hX : cX = cX' ∘ eX) (hY : cY = cY' ∘ eY) :
(indexValueIso d (Equiv.sumCongr eX eY) (tensorator_mapIso_cond hX hY)).trans indexValueTensorator =
indexValueTensorator.trans (Equiv.prodCongr (indexValueIso d eX hX) (indexValueIso d eY hY)) := by
apply Equiv.ext
intro i
simp
simp [ indexValueTensorator]
apply And.intro
all_goals
funext x
rw [indexValueIso_eq_symm, indexValueIso_symm_apply']
rw [indexValueIso_eq_symm, indexValueIso_symm_apply']
simp
/-!
## A basis for Lorentz tensors
-/
noncomputable section basis
open TensorProduct
open Set LinearMap Submodule
variable {d : } {X Y Y' X' : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
[Fintype Y'] [DecidableEq Y'] [Fintype X'] [DecidableEq X']
(cX : X → Color) (cY : Y → Color)
def basis : Basis (IndexValue d cX) (RealLorentzTensor d cX) := Pi.basisFun _ _
@[simp]
def basisProd :
Basis (IndexValue d cX × IndexValue d cY) (RealLorentzTensor d cX ⊗[] RealLorentzTensor d cY) :=
(Basis.tensorProduct (basis cX) (basis cY))
lemma mapIsoFiber_basis {cX : X → Color} {cY : Y → Color} (e : X ≃ Y) (h : cX = cY ∘ e)
(i : IndexValue d cX) : (mapIso e h) (basis cX i)
= basis cY (indexValueIso d e h i) := by
funext a
rw [mapIso_apply]
by_cases ha : a = ((indexValueIso d e h) i)
· subst ha
nth_rewrite 2 [indexValueIso_eq_symm]
rw [Equiv.apply_symm_apply]
simp only [basis]
erw [Pi.basisFun_apply, Pi.basisFun_apply]
simp only [stdBasis_same]
· simp only [basis]
erw [Pi.basisFun_apply, Pi.basisFun_apply]
simp
rw [if_neg, if_neg]
exact fun a_1 => ha (_root_.id (Eq.symm a_1))
rw [← Equiv.symm_apply_eq, indexValueIso_symm]
exact fun a_1 => ha (_root_.id (Eq.symm a_1))
end basis
/-! TODO: Following the ethos of modular operads, prove properties of contraction. -/
/-! TODO: Use `contr` to generalize to any pair of marked index. -/
/-!
## Rising and lowering indices
Rising or lowering an index corresponds to changing the color of that index.
-/
/-! TODO: Define the rising and lowering of indices using contraction with the metric. -/
/-!
## Graphical species and Lorentz tensors
-/
/-! TODO: From Lorentz tensors graphical species. -/
/-! TODO: Show that the action of the Lorentz group defines an action on the graphical species. -/
end RealLorentzTensor

View file

@ -1,401 +0,0 @@
/-
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.LorentzTensor.Real.Basic
import HepLean.SpaceTime.LorentzTensor.Real.LorentzAction
import HepLean.SpaceTime.LorentzTensor.Real.Multiplication
/-!
# Constructors for real Lorentz tensors
In this file we will constructors of real Lorentz tensors from real numbers,
vectors and matrices.
We will derive properties of these constructors.
-/
namespace RealLorentzTensor
/-!
# Tensors from and to the reals
An important point that we shall see here is that there is a well defined map
to the real numbers, i.e. which is invariant under transformations of mapIso.
-/
/-- An equivalence from Real tensors on an empty set to ``. -/
@[simps!]
def toReal (d : ) {X : Type} (h : IsEmpty X) (f : X → Color) : RealLorentzTensor d f ≃ where
toFun := fun T => T.coord (IsEmpty.elim h)
invFun := fun r => {
color := fun x => IsEmpty.elim h x,
coord := fun _ => r}
left_inv T := by
refine ext ?_ ?_
funext x
exact IsEmpty.elim h x
funext a
change T.coord (IsEmpty.elim h) = _
apply congrArg
funext x
exact IsEmpty.elim h x
right_inv x := rfl
/-- The real number obtained from a tensor is invariant under `mapIso`. -/
lemma toReal_mapIso (d : ) {X Y : Type} (h : IsEmpty X) (f : X ≃ Y) :
(mapIso d f).trans (toReal d (Equiv.isEmpty f.symm)) = toReal d h := by
apply Equiv.ext
intro x
simp only [Equiv.trans_apply, toReal_apply, mapIso_apply_color, mapIso_apply_coord]
apply congrArg
funext x
exact IsEmpty.elim h x
/-!
# Tensors from reals, vectors and matrices
Note that that these definitions are not equivariant with respect to an
action of the Lorentz group. They are provided for constructive purposes.
-/
/-- A marked 1-tensor with a single up index constructed from a vector.
Note: This is not the same as rising indices on `ofVecDown`. -/
def ofVecUp {d : } (v : Fin 1 ⊕ Fin d → ) :
Marked d Empty 1 where
color := fun _ => Colors.up
coord := fun i => v <| i <| Sum.inr <| 0
/-- A marked 1-tensor with a single down index constructed from a vector.
Note: This is not the same as lowering indices on `ofVecUp`. -/
def ofVecDown {d : } (v : Fin 1 ⊕ Fin d → ) :
Marked d Empty 1 where
color := fun _ => Colors.down
coord := fun i => v <| i <| Sum.inr <| 0
/-- A tensor with two up indices constructed from a matrix.
Note: This is not the same as rising or lowering indices on other `ofMat...`. -/
def ofMatUpUp {d : } (m : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
Marked d Empty 2 where
color := fun _ => Colors.up
coord := fun i => m (i (Sum.inr 0)) (i (Sum.inr 1))
/-- A tensor with two down indices constructed from a matrix.
Note: This is not the same as rising or lowering indices on other `ofMat...`. -/
def ofMatDownDown {d : } (m : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
Marked d Empty 2 where
color := fun _ => Colors.down
coord := fun i => m (i (Sum.inr 0)) (i (Sum.inr 1))
/-- A marked 2-tensor with the first index up and the second index down.
Note: This is not the same as rising or lowering indices on other `ofMat...`. -/
@[simps!]
def ofMatUpDown {d : } (m : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
Marked d Empty 2 where
color := fun i => match i with
| Sum.inr 0 => Colors.up
| Sum.inr 1 => Colors.down
coord := fun i => m (i (Sum.inr 0)) (i (Sum.inr 1))
/-- A marked 2-tensor with the first index down and the second index up.
Note: This is not the same as rising or lowering indices on other `ofMat...`. -/
def ofMatDownUp {d : } (m : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
Marked d Empty 2 where
color := fun i => match i with
| Sum.inr 0 => Colors.down
| Sum.inr 1 => Colors.up
coord := fun i => m (i (Sum.inr 0)) (i (Sum.inr 1))
/-!
## Equivalence of `IndexValue` for constructors
-/
/-- Index values for `ofVecUp v` are equivalent to elements of `Fin 1 ⊕ Fin d`. -/
def ofVecUpIndexValue (v : Fin 1 ⊕ Fin d → ) :
IndexValue d (ofVecUp v).color ≃ (Fin 1 ⊕ Fin d) where
toFun i := i (Sum.inr 0)
invFun x := fun i => match i with
| Sum.inr 0 => x
left_inv i := by
funext y
match y with
| Sum.inr 0 => rfl
right_inv x := rfl
/-- Index values for `ofVecDown v` are equivalent to elements of `Fin 1 ⊕ Fin d`. -/
def ofVecDownIndexValue (v : Fin 1 ⊕ Fin d → ) :
IndexValue d (ofVecDown v).color ≃ (Fin 1 ⊕ Fin d) where
toFun i := i (Sum.inr 0)
invFun x := fun i => match i with
| Sum.inr 0 => x
left_inv i := by
funext y
match y with
| Sum.inr 0 => rfl
right_inv x := rfl
/-- Index values for `ofMatUpUp v` are equivalent to elements of
`(Fin 1 ⊕ Fin d) × (Fin 1 ⊕ Fin d)`. -/
def ofMatUpUpIndexValue (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
IndexValue d (ofMatUpUp M).color ≃ (Fin 1 ⊕ Fin d) × (Fin 1 ⊕ Fin d) where
toFun i := (i (Sum.inr 0), i (Sum.inr 1))
invFun x := fun i => match i with
| Sum.inr 0 => x.1
| Sum.inr 1 => x.2
left_inv i := by
funext y
match y with
| Sum.inr 0 => rfl
| Sum.inr 1 => rfl
right_inv x := rfl
/-- Index values for `ofMatDownDown v` are equivalent to elements of
`(Fin 1 ⊕ Fin d) × (Fin 1 ⊕ Fin d)`. -/
def ofMatDownDownIndexValue (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
IndexValue d (ofMatDownDown M).color ≃ (Fin 1 ⊕ Fin d) × (Fin 1 ⊕ Fin d) where
toFun i := (i (Sum.inr 0), i (Sum.inr 1))
invFun x := fun i => match i with
| Sum.inr 0 => x.1
| Sum.inr 1 => x.2
left_inv i := by
funext y
match y with
| Sum.inr 0 => rfl
| Sum.inr 1 => rfl
right_inv x := rfl
/-- Index values for `ofMatUpDown v` are equivalent to elements of
`(Fin 1 ⊕ Fin d) × (Fin 1 ⊕ Fin d)`. -/
def ofMatUpDownIndexValue (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
IndexValue d (ofMatUpDown M).color ≃ (Fin 1 ⊕ Fin d) × (Fin 1 ⊕ Fin d) where
toFun i := (i (Sum.inr 0), i (Sum.inr 1))
invFun x := fun i => match i with
| Sum.inr 0 => x.1
| Sum.inr 1 => x.2
left_inv i := by
funext y
match y with
| Sum.inr 0 => rfl
| Sum.inr 1 => rfl
right_inv x := rfl
/-- Index values for `ofMatDownUp v` are equivalent to elements of
`(Fin 1 ⊕ Fin d) × (Fin 1 ⊕ Fin d)`. -/
def ofMatDownUpIndexValue (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
IndexValue d (ofMatDownUp M).color ≃ (Fin 1 ⊕ Fin d) × (Fin 1 ⊕ Fin d) where
toFun i := (i (Sum.inr 0), i (Sum.inr 1))
invFun x := fun i => match i with
| Sum.inr 0 => x.1
| Sum.inr 1 => x.2
left_inv i := by
funext y
match y with
| Sum.inr 0 => rfl
| Sum.inr 1 => rfl
right_inv x := rfl
/-!
## Contraction of indices of tensors from matrices
-/
open Matrix
open Marked
/-- Contracting the indices of `ofMatUpDown` returns the trace of the matrix. -/
lemma contr_ofMatUpDown_eq_trace {d : } (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
contr (ofMatUpDown M) (by rfl) = (toReal d instIsEmptyEmpty).symm M.trace := by
refine ext ?_ rfl
· funext i
exact Empty.elim i
/-- Contracting the indices of `ofMatDownUp` returns the trace of the matrix. -/
lemma contr_ofMatDownUp_eq_trace {d : } (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
contr (ofMatDownUp M) (by rfl) = (toReal d instIsEmptyEmpty).symm M.trace := by
refine ext ?_ rfl
· funext i
exact Empty.elim i
/-!
## Multiplication of tensors from vectors and matrices
-/
/-- Multiplying `ofVecUp` with `ofVecDown` gives the dot product. -/
@[simp]
lemma mul_ofVecUp_ofVecDown_eq_dot_prod {d : } (v₁ v₂ : Fin 1 ⊕ Fin d → ) :
multMarked (ofVecUp v₁) (ofVecDown v₂) rfl = (toReal d instIsEmptySum).symm (v₁ ⬝ᵥ v₂) := by
refine ext ?_ rfl
· funext i
exact IsEmpty.elim instIsEmptySum i
/-- Multiplying `ofVecDown` with `ofVecUp` gives the dot product. -/
@[simp]
lemma mul_ofVecDown_ofVecUp_eq_dot_prod {d : } (v₁ v₂ : Fin 1 ⊕ Fin d → ) :
multMarked (ofVecDown v₁) (ofVecUp v₂) rfl = (toReal d instIsEmptySum).symm (v₁ ⬝ᵥ v₂) := by
refine ext ?_ rfl
· funext i
exact IsEmpty.elim instIsEmptySum i
lemma mul_ofMatUpDown_ofVecUp_eq_mulVec {d : } (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) )
(v : Fin 1 ⊕ Fin d → ) :
mapIso d ((Equiv.sumEmpty (Empty ⊕ Fin 1) Empty))
(multMarked (unmarkFirst $ ofMatUpDown M) (ofVecUp v) rfl) = ofVecUp (M *ᵥ v) := by
refine ext ?_ rfl
· funext i
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, mapIso_apply_color, mul_color, Equiv.symm_symm]
fin_cases i
rfl
lemma mul_ofMatDownUp_ofVecDown_eq_mulVec {d : } (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) )
(v : Fin 1 ⊕ Fin d → ) :
mapIso d (Equiv.sumEmpty (Empty ⊕ Fin 1) Empty)
(multMarked (unmarkFirst $ ofMatDownUp M) (ofVecDown v) rfl) = ofVecDown (M *ᵥ v) := by
refine ext ?_ rfl
· funext i
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, mapIso_apply_color, mul_color, Equiv.symm_symm]
fin_cases i
rfl
/-!
## The Lorentz action on constructors
-/
section lorentzAction
variable {d : } {X : Type} [Fintype X] [DecidableEq X] (T : RealLorentzTensor d X) (c : X → Colors)
variable (Λ Λ' : LorentzGroup d)
open Matrix
/-- The action of the Lorentz group on `ofReal v` is trivial. -/
@[simp]
lemma lorentzAction_toReal (h : IsEmpty X) (r : ) :
Λ • (toReal d h).symm r = (toReal d h).symm r :=
lorentzAction_on_isEmpty Λ ((toReal d h).symm r)
/-- The action of the Lorentz group on `ofVecUp v` is by vector multiplication. -/
@[simp]
lemma lorentzAction_ofVecUp (v : Fin 1 ⊕ Fin d → ) :
Λ • ofVecUp v = ofVecUp (Λ *ᵥ v) := by
refine ext rfl ?_
funext i
erw [lorentzAction_smul_coord]
simp only [ofVecUp, IndexValue, Fin.isValue, Fintype.prod_sum_type, Finset.univ_eq_empty,
Finset.prod_empty, one_mul]
rw [mulVec]
simp only [Fin.isValue, dotProduct, Finset.univ_unique, Fin.default_eq_zero,
Finset.sum_singleton]
erw [Finset.sum_equiv (ofVecUpIndexValue v)]
intro i
simp_all only [Finset.mem_univ, Fin.isValue, Equiv.coe_fn_mk]
intro j _
simp_all only [Finset.mem_univ, Fin.isValue, Finset.prod_singleton, indexValueIso_refl]
rfl
/-- The action of the Lorentz group on `ofVecDown v` is
by vector multiplication of the transpose-inverse. -/
@[simp]
lemma lorentzAction_ofVecDown (v : Fin 1 ⊕ Fin d → ) :
Λ • ofVecDown v = ofVecDown ((LorentzGroup.transpose Λ⁻¹) *ᵥ v) := by
refine ext rfl ?_
funext i
erw [lorentzAction_smul_coord]
simp only [ofVecDown, IndexValue, Fin.isValue, Fintype.prod_sum_type, Finset.univ_eq_empty,
Finset.prod_empty, one_mul, lorentzGroupIsGroup_inv]
rw [mulVec]
simp only [Fin.isValue, dotProduct, Finset.univ_unique, Fin.default_eq_zero,
Finset.sum_singleton]
erw [Finset.sum_equiv (ofVecUpIndexValue v)]
intro i
simp_all only [Finset.mem_univ, Fin.isValue, Equiv.coe_fn_mk]
intro j _
simp_all only [Finset.mem_univ, Fin.isValue, Finset.prod_singleton, indexValueIso_refl]
rfl
@[simp]
lemma lorentzAction_ofMatUpUp (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
Λ • ofMatUpUp M = ofMatUpUp (Λ.1 * M * (LorentzGroup.transpose Λ).1) := by
refine ext rfl ?_
funext i
erw [lorentzAction_smul_coord]
erw [← Equiv.sum_comp (ofMatUpUpIndexValue M).symm]
simp only [ofMatUpUp, IndexValue, Fin.isValue, Fintype.prod_sum_type, Finset.univ_eq_empty,
Finset.prod_empty, one_mul, mul_apply]
erw [Finset.sum_product]
rw [Finset.sum_comm]
refine Finset.sum_congr rfl (fun x _ => ?_)
rw [Finset.sum_mul]
refine Finset.sum_congr rfl (fun y _ => ?_)
simp only [Fin.prod_univ_two, Fin.isValue, indexValueIso_refl, IndexValue]
exact mul_right_comm _ _ _
@[simp]
lemma lorentzAction_ofMatDownDown (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
Λ • ofMatDownDown M = ofMatDownDown ((LorentzGroup.transpose Λ⁻¹).1 * M * (Λ⁻¹).1) := by
refine ext rfl ?_
funext i
erw [lorentzAction_smul_coord]
erw [← Equiv.sum_comp (ofMatDownDownIndexValue M).symm]
simp only [ofMatDownDown, IndexValue, Fin.isValue, Fintype.prod_sum_type, Finset.univ_eq_empty,
Finset.prod_empty, one_mul, mul_apply]
erw [Finset.sum_product]
rw [Finset.sum_comm]
refine Finset.sum_congr rfl (fun x _ => ?_)
rw [Finset.sum_mul]
refine Finset.sum_congr rfl (fun y _ => ?_)
simp only [Fin.prod_univ_two, Fin.isValue, indexValueIso_refl, IndexValue]
exact mul_right_comm _ _ _
@[simp]
lemma lorentzAction_ofMatUpDown (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
Λ • ofMatUpDown M = ofMatUpDown (Λ.1 * M * (Λ⁻¹).1) := by
refine ext rfl ?_
funext i
erw [lorentzAction_smul_coord]
erw [← Equiv.sum_comp (ofMatUpDownIndexValue M).symm]
simp only [ofMatUpDown, IndexValue, Fin.isValue, Fintype.prod_sum_type, Finset.univ_eq_empty,
Finset.prod_empty, one_mul, mul_apply]
erw [Finset.sum_product]
rw [Finset.sum_comm]
refine Finset.sum_congr rfl (fun x _ => ?_)
rw [Finset.sum_mul]
refine Finset.sum_congr rfl (fun y _ => ?_)
simp only [Fin.prod_univ_two, Fin.isValue, indexValueIso_refl, IndexValue]
exact mul_right_comm _ _ _
@[simp]
lemma lorentzAction_ofMatDownUp (M : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :
Λ • ofMatDownUp M =
ofMatDownUp ((LorentzGroup.transpose Λ⁻¹).1 * M * (LorentzGroup.transpose Λ).1) := by
refine ext rfl ?_
funext i
erw [lorentzAction_smul_coord]
erw [← Equiv.sum_comp (ofMatDownUpIndexValue M).symm]
simp only [ofMatDownUp, IndexValue, Fin.isValue, Fintype.prod_sum_type, Finset.univ_eq_empty,
Finset.prod_empty, one_mul, mul_apply]
erw [Finset.sum_product]
rw [Finset.sum_comm]
refine Finset.sum_congr rfl (fun x _ => ?_)
rw [Finset.sum_mul]
refine Finset.sum_congr rfl (fun y _ => ?_)
simp only [Fin.prod_univ_two, Fin.isValue, indexValueIso_refl, IndexValue]
exact mul_right_comm _ _ _
end lorentzAction
end RealLorentzTensor

View file

@ -1,119 +0,0 @@
/-
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.LorentzTensor.Real.Basic
import HepLean.SpaceTime.LorentzTensor.Real.TensorProducts
/-!
# Multiplication of Real Lorentz Tensors along an index
We define the multiplication of two singularly marked Lorentz tensors along the
marked index. This is equivalent to contracting two Lorentz tensors.
We prove various results about this multiplication.
-/
/-! TODO: Set up a good notation for the multiplication. -/
namespace RealLorentzTensor
open TensorProduct
open Set LinearMap Submodule
variable {d : } {X Y Y' X' : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
[Fintype Y'] [DecidableEq Y'] [Fintype X'] [DecidableEq X']
(cX : X → Color) (cY : Y → Color)
/-- The contraction of all indices of two tensors with dual index-colors.
This is a bilinear map to . -/
@[simps!]
def contrAll {f1 : X → Color} {f2 : Y → Color} (e : X ≃ Y) (hc : f1 = τ ∘ f2 ∘ e) :
RealLorentzTensor d f1 →ₗ[] RealLorentzTensor d f2 →ₗ[] where
toFun T := {
toFun := fun S => ∑ i, T i * S (indexValueDualIso d e hc i),
map_add' := fun S F => by
trans ∑ i, (T i * S (indexValueDualIso d e hc i) + T i * F (indexValueDualIso d e hc i))
exact Finset.sum_congr rfl (fun i _ => mul_add _ _ _ )
exact Finset.sum_add_distrib,
map_smul' := fun r S => by
trans ∑ i , r * (T i * S (indexValueDualIso d e hc i))
refine Finset.sum_congr rfl (fun x _ => ?_)
ring_nf
rw [mul_assoc]
rfl
rw [← Finset.mul_sum]
rfl}
map_add' := fun T S => by
ext F
trans ∑ i , (T i * F (indexValueDualIso d e hc i) + S i * F (indexValueDualIso d e hc i))
exact Finset.sum_congr rfl (fun x _ => add_mul _ _ _)
exact Finset.sum_add_distrib
map_smul' := fun r T => by
ext S
trans ∑ i , r * (T i * S (indexValueDualIso d e hc i))
refine Finset.sum_congr rfl (fun x _ => mul_assoc _ _ _)
rw [← Finset.mul_sum]
rfl
lemma contrAll_symm {f1 : X → Color} {f2 : Y → Color} (e : X ≃ Y)
(h : f1 = τ ∘ f2 ∘ e) (T : RealLorentzTensor d f1) (S : RealLorentzTensor d f2) :
contrAll e h T S = contrAll e.symm (indexValueDualIso_cond_symm h) S T := by
rw [contrAll_apply_apply, contrAll_apply_apply, ← Equiv.sum_comp (indexValueDualIso d e h)]
refine Finset.sum_congr rfl (fun i _ => ?_)
rw [mul_comm, ← Equiv.trans_apply]
simp
lemma contrAll_mapIsoFiber_right {f1 : X → Color} {f2 : Y → Color}
{g2 : Y' → Color} (e : X ≃ Y) (eY : Y ≃ Y')
(h : f1 = τ ∘ f2 ∘ e) (hY : f2 = g2 ∘ eY) (T : RealLorentzTensor d f1) (S : RealLorentzTensor d f2) :
contrAll e h T S = contrAll (e.trans eY) (indexValueDualIso_cond_trans_indexValueIso h hY)
T (mapIso eY hY S) := by
rw [contrAll_apply_apply, contrAll_apply_apply]
refine Finset.sum_congr rfl (fun i _ => Mathlib.Tactic.Ring.mul_congr rfl ?_ rfl)
rw [mapIso_apply, ← Equiv.trans_apply]
simp only [indexValueDualIso_trans_indexValueIso]
congr
ext1 x
simp only [Equiv.trans_apply, Equiv.symm_apply_apply]
lemma contrAll_mapIsoFiber_left {f1 : X → Color} {f2 : Y → Color}
{g1 : X' → Color} (e : X ≃ Y) (eX : X ≃ X')
(h : f1 = τ ∘ f2 ∘ e) (hX : f1 = g1 ∘ eX) (T : RealLorentzTensor d f1) (S : RealLorentzTensor d f2) :
contrAll e h T S = contrAll (eX.symm.trans e)
(indexValueIso_trans_indexValueDualIso_cond ((Equiv.eq_comp_symm eX g1 f1).mpr hX.symm) h)
(mapIso eX hX T) S := by
rw [contrAll_symm, contrAll_mapIsoFiber_right _ eX _ hX, contrAll_symm]
rfl
lemma contrAll_lorentzAction {f1 : X → Color} {f2 : Y → Color} (e : X ≃ Y)
(hc : f1 = τ ∘ f2 ∘ e) (T : RealLorentzTensor d f1) (S : RealLorentzTensor d f2) (Λ : LorentzGroup d) :
contrAll e hc (Λ • T) (Λ • S) = contrAll e hc T S := by
change ∑ i, (∑ j, toTensorRepMat Λ i j * T j) *
(∑ k, toTensorRepMat Λ (indexValueDualIso d e hc i) k * S k) = _
trans ∑ i, ∑ j, ∑ k, (toTensorRepMat Λ (indexValueDualIso d e hc i) k * toTensorRepMat Λ i j)
* T j * S k
· apply Finset.sum_congr rfl (fun x _ => ?_)
rw [Finset.sum_mul_sum]
apply Finset.sum_congr rfl (fun j _ => ?_)
apply Finset.sum_congr rfl (fun k _ => ?_)
ring
rw [Finset.sum_comm]
trans ∑ j, ∑ k, ∑ i, (toTensorRepMat Λ (indexValueDualIso d e hc i) k
* toTensorRepMat Λ i j) * T j * S k
· apply Finset.sum_congr rfl (fun j _ => ?_)
rw [Finset.sum_comm]
trans ∑ j, ∑ k, (toTensorRepMat 1 (indexValueDualIso d e.symm (indexValueDualIso_cond_symm hc) k) j) * T j * S k
· apply Finset.sum_congr rfl (fun j _ => Finset.sum_congr rfl (fun k _ => ?_))
rw [← Finset.sum_mul, ← Finset.sum_mul]
erw [toTensorRepMat_indexValueDualIso_sum]
rw [Finset.sum_comm]
trans ∑ k, T (indexValueDualIso d e.symm (indexValueDualIso_cond_symm hc) k) * S k
· apply Finset.sum_congr rfl (fun k _ => ?_)
rw [← Finset.sum_mul, ← toTensorRepMat_one_coord_sum' T]
rw [← Equiv.sum_comp (indexValueDualIso d e hc), ← indexValueDualIso_symm e hc]
simp only [Equiv.symm_apply_apply, contrAll_apply_apply]
end RealLorentzTensor

View file

@ -1,308 +0,0 @@
/-
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.LorentzTensor.Real.Basic
import HepLean.SpaceTime.LorentzGroup.Basic
import Mathlib.RepresentationTheory.Basic
/-!
# Lorentz group action on Real Lorentz Tensors
We define the action of the Lorentz group on Real Lorentz Tensors.
The Lorentz action is currently only defined for finite and decidable types `X`.
-/
namespace RealLorentzTensor
variable {d : } {X Y : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
(c : X → Color) (Λ Λ' : LorentzGroup d) {μ : Color}
open LorentzGroup BigOperators
/-- Monoid homomorphism from the Lorentz group to matrices indexed by `ColorsIndex d μ` for a
color `μ`.
This can be thought of as the representation of the Lorentz group for that color index. -/
def colorMatrix (μ : Color) : LorentzGroup d →* Matrix (ColorIndex d μ) (ColorIndex d μ) where
toFun Λ := match μ with
| .up => fun i j => Λ.1 i j
| .down => fun i j => (LorentzGroup.transpose Λ⁻¹).1 i j
map_one' := by
ext i j
match μ with
| .up =>
simp only [lorentzGroupIsGroup_one_coe, OfNat.ofNat, One.one, ColorIndex]
congr
| .down =>
simp only [transpose, inv_one, lorentzGroupIsGroup_one_coe, Matrix.transpose_one]
simp only [OfNat.ofNat, One.one, ColorIndex]
congr
map_mul' Λ Λ' := by
ext i j
match μ with
| .up =>
simp only [lorentzGroupIsGroup_mul_coe]
| .down =>
simp only [transpose, mul_inv_rev, lorentzGroupIsGroup_inv, lorentzGroupIsGroup_mul_coe,
Matrix.transpose_mul, Matrix.transpose_apply]
rfl
lemma colorMatrix_ext {μ : Color} {a b c d : ColorIndex d μ} (hab : a = b) (hcd : c = d) :
(colorMatrix μ) Λ a c = (colorMatrix μ) Λ b d := by
rw [hab, hcd]
lemma colorMatrix_cast {μ ν : Color} (h : μ = ν) (Λ : LorentzGroup d) :
colorMatrix ν Λ =
Matrix.reindex (colorIndexCast h) (colorIndexCast h) (colorMatrix μ Λ) := by
subst h
rfl
lemma colorMatrix_dual_cast {μ ν : Color} (Λ : LorentzGroup d) (h : μ = τ ν) :
colorMatrix ν Λ = Matrix.reindex (colorIndexDualCast h) (colorIndexDualCast h)
(colorMatrix μ (LorentzGroup.transpose Λ⁻¹)) := by
subst h
match ν with
| .up =>
ext i j
simp only [colorMatrix, MonoidHom.coe_mk, OneHom.coe_mk, τ, transpose, lorentzGroupIsGroup_inv,
Matrix.transpose_apply, minkowskiMetric.dual_transpose, minkowskiMetric.dual_dual,
Matrix.reindex_apply, colorIndexDualCast_symm, Matrix.submatrix_apply]
rfl
| .down =>
ext i j
simp only [τ, colorMatrix, MonoidHom.coe_mk, OneHom.coe_mk, colorIndexDualCastSelf, transpose,
lorentzGroupIsGroup_inv, Matrix.transpose_apply, minkowskiMetric.dual_transpose,
minkowskiMetric.dual_dual, Matrix.reindex_apply, Equiv.coe_fn_symm_mk, Matrix.submatrix_apply]
rfl
lemma colorMatrix_transpose {μ : Color} (Λ : LorentzGroup d) :
colorMatrix μ (LorentzGroup.transpose Λ) = (colorMatrix μ Λ).transpose := by
match μ with
| .up => rfl
| .down =>
ext i j
simp only [colorMatrix, transpose, lorentzGroupIsGroup_inv, Matrix.transpose_apply,
MonoidHom.coe_mk, OneHom.coe_mk, minkowskiMetric.dual_transpose]
/-!
## Lorentz group to tensor representation matrices.
-/
/-- The matrix representation of the Lorentz group given a color of index. -/
@[simps!]
def toTensorRepMat {c : X → Color} :
LorentzGroup d →* Matrix (IndexValue d c) (IndexValue d c) where
toFun Λ := fun i j => ∏ x, colorMatrix (c x) Λ (i x) (j x)
map_one' := by
ext i j
by_cases hij : i = j
· subst hij
simp only [map_one, Matrix.one_apply_eq, Finset.prod_const_one]
· obtain ⟨x, hijx⟩ := Function.ne_iff.mp hij
simp only [map_one]
rw [@Finset.prod_eq_zero _ _ _ _ _ x]
exact Eq.symm (Matrix.one_apply_ne' fun a => hij (id (Eq.symm a)))
exact Finset.mem_univ x
exact Matrix.one_apply_ne' (id (Ne.symm hijx))
map_mul' Λ Λ' := by
ext i j
rw [Matrix.mul_apply]
trans ∑ (k : IndexValue d c), ∏ x,
(colorMatrix (c x) Λ (i x) (k x)) * (colorMatrix (c x) Λ' (k x) (j x))
have h1 : ∑ (k : IndexValue d c), ∏ x,
(colorMatrix (c x) Λ (i x) (k x)) * (colorMatrix (c x) Λ' (k x) (j x)) =
∏ x, ∑ y, (colorMatrix (c x) Λ (i x) y) * (colorMatrix (c x) Λ' y (j x)) := by
rw [Finset.prod_sum]
simp only [Finset.prod_attach_univ, Finset.sum_univ_pi]
rfl
rw [h1]
simp only [map_mul]
rfl
refine Finset.sum_congr rfl (fun k _ => ?_)
rw [Finset.prod_mul_distrib]
lemma toTensorRepMat_mul' (i j : IndexValue d c) :
toTensorRepMat (Λ * Λ') i j = ∑ (k : IndexValue d c),
∏ x, colorMatrix (c x) Λ (i x) (k x) * colorMatrix (c x) Λ' (k x) (j x) := by
simp [Matrix.mul_apply, IndexValue]
refine Finset.sum_congr rfl (fun k _ => ?_)
rw [Finset.prod_mul_distrib]
rfl
lemma toTensorRepMat_of_indexValueSumEquiv {cX : X → Color} {cY : Y → Color}
(i j : IndexValue d (Sum.elim cX cY)) :
toTensorRepMat Λ i j = toTensorRepMat Λ (indexValueTensorator i).1 (indexValueTensorator j).1 *
toTensorRepMat Λ (indexValueTensorator i).2 (indexValueTensorator j).2 :=
Fintype.prod_sum_type fun x => (colorMatrix (Sum.elim cX cY x)) Λ (i x) (j x)
lemma toTensorRepMat_of_indexValueSumEquiv' {cX : X → Color} {cY : Y → Color}
(i j : IndexValue d cX) (k l : IndexValue d cY) :
toTensorRepMat Λ i j * toTensorRepMat Λ k l =
toTensorRepMat Λ (indexValueTensorator.symm (i, k)) (indexValueTensorator.symm (j, l)) :=
(Fintype.prod_sum_type fun x => (colorMatrix (Sum.elim cX cY x)) Λ
(indexValueTensorator.symm (i, k) x) (indexValueTensorator.symm (j, l) x)).symm
/-!
## Tensor representation matrices and marked tensors.
-/
lemma toTensorRepMat_indexValueDualIso_left {f1 : X → Color} {f2 : Y → Color}
(e : X ≃ Y) (hc : f1 = τ ∘ f2 ∘ e) (i : IndexValue d f1) (j : IndexValue d f2) :
toTensorRepMat Λ (indexValueDualIso d e hc i) j =
toTensorRepMat Λ⁻¹ (indexValueDualIso d e.symm (indexValueDualIso_cond_symm hc) j) i := by
rw [toTensorRepMat_apply, toTensorRepMat_apply, ← Equiv.prod_comp e]
apply Finset.prod_congr rfl (fun x _ => ?_)
erw [colorMatrix_dual_cast Λ (congrFun hc x)]
rw [Matrix.reindex_apply, colorMatrix_transpose]
simp only [Function.comp_apply, colorIndexDualCast_symm,
Matrix.submatrix_apply, Matrix.transpose_apply]
rw [indexValueDualIso_eq_symm, indexValueDualIso_symm_apply',
indexValueDualIso_eq_symm, indexValueDualIso_symm_apply']
rw [← Equiv.trans_apply, colorIndexDualCast_symm, colorsIndexDualCast_trans]
apply colorMatrix_ext
simp
simp [colorIndexCast]
apply cast_eq_iff_heq.mpr
rw [Equiv.symm_apply_apply]
lemma toTensorRepMat_indexValueDualIso_sum {f1 : X → Color} {f2 : Y → Color}
(e : X ≃ Y) (hc : f1 = τ ∘ f2 ∘ e) (j : IndexValue d f1) (k : IndexValue d f2) :
(∑ i : IndexValue d f1, toTensorRepMat Λ ((indexValueDualIso d e hc) i) k * toTensorRepMat Λ i j) =
toTensorRepMat 1 (indexValueDualIso d e.symm (indexValueDualIso_cond_symm hc) k) j := by
trans ∑ i, toTensorRepMat Λ⁻¹ (indexValueDualIso d e.symm (indexValueDualIso_cond_symm hc) k) i
* toTensorRepMat Λ i j
apply Finset.sum_congr rfl (fun i _ => ?_)
rw [toTensorRepMat_indexValueDualIso_left]
rw [← Matrix.mul_apply, ← toTensorRepMat.map_mul, inv_mul_self Λ]
lemma toTensorRepMat_one_coord_sum' {f1 : X → Color}
(T : RealLorentzTensor d f1) (k : IndexValue d f1) :
∑ j, (toTensorRepMat 1 k j) * T j = T k := by
erw [Finset.sum_eq_single_of_mem k]
simp only [IndexValue, map_one, Matrix.one_apply_eq, one_mul]
exact Finset.mem_univ k
intro j _ hjk
simp only [IndexValue, map_one, mul_eq_zero]
exact Or.inl (Matrix.one_apply_ne' hjk)
/-!
## Definition of the Lorentz group action on Real Lorentz Tensors.
-/
@[simps!]
def lorentzAction {c : X → Color} :
Representation (LorentzGroup d) (RealLorentzTensor d c) where
toFun Λ := {
toFun := fun T i => ∑ j, toTensorRepMat Λ i j * T j,
map_add' := fun T S => by
funext i
trans ∑ j, (toTensorRepMat Λ i j * T j + toTensorRepMat Λ i j * S j)
· refine Finset.sum_congr rfl (fun j _ => ?_)
erw [mul_add]
· rw [Finset.sum_add_distrib]
rfl
map_smul' := fun a T => by
funext i
simp only [ RingHom.id_apply]
trans ∑ j, a * (toTensorRepMat Λ i j * T j)
· refine Finset.sum_congr rfl (fun j _ => ?_)
rw [← mul_assoc, mul_comm a _, mul_assoc]
rfl
· rw [← Finset.mul_sum]
rfl}
map_one' := by
ext T
simp only [map_one, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.one_apply]
funext i
rw [Finset.sum_eq_single_of_mem i]
simp only [Matrix.one_apply_eq, one_mul]
exact Finset.mem_univ i
exact fun j _ hij => mul_eq_zero.mpr (Or.inl (Matrix.one_apply_ne' hij))
map_mul' Λ Λ' := by
ext T
simp only
funext i
trans ∑ j, ∑ k : IndexValue d c, (∏ x, colorMatrix (c x) Λ (i x) (k x) *
colorMatrix (c x) Λ' (k x) (j x)) * T j
· refine Finset.sum_congr rfl (fun j _ => ?_)
rw [toTensorRepMat_mul', Finset.sum_mul]
· rw [Finset.sum_comm]
refine Finset.sum_congr rfl (fun j _ => ?_)
simp only [LinearMap.coe_mk, AddHom.coe_mk, Finset.mul_sum, toTensorRepMat, IndexValue]
refine Finset.sum_congr rfl (fun k _ => ?_)
rw [← mul_assoc, Finset.prod_mul_distrib]
rfl
scoped[RealLorentzTensor] infixl:78 " • " => lorentzAction
/-- The Lorentz action commutes with `mapIso`. -/
lemma lorentzAction_mapIso (e : X ≃ Y) {f1 : X → Color}
{f2 : Y → Color} (h : f1 = f2 ∘ e) (Λ : LorentzGroup d)
(T : RealLorentzTensor d f1) : mapIso e h (Λ • T) = Λ • (mapIso e h T) := by
funext i
rw [mapIso_apply, lorentzAction_apply_apply, lorentzAction_apply_apply]
rw [← Equiv.sum_comp (indexValueIso d e h)]
refine Finset.sum_congr rfl (fun j _ => Mathlib.Tactic.Ring.mul_congr ?_ ?_ rfl)
· rw [← Equiv.prod_comp e]
apply Finset.prod_congr rfl (fun x _ => ?_)
erw [colorMatrix_cast (congrFun h x)]
rw [Matrix.reindex_apply]
simp
apply colorMatrix_ext
rw [indexValueIso_eq_symm, indexValueIso_symm_apply']
erw [← Equiv.eq_symm_apply]
simp only [Function.comp_apply, Equiv.symm_symm_apply, colorIndexCast, Equiv.cast_symm,
Equiv.cast_apply, cast_cast, cast_eq]
rw [indexValueIso_eq_symm, indexValueIso_symm_apply']
simp [colorIndexCast]
symm
refine cast_eq_iff_heq.mpr ?_
rw [Equiv.symm_apply_apply]
· rw [mapIso_apply]
apply congrArg
rw [← Equiv.trans_apply]
simp
section
variable {d : } {X Y Y' X' : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
[Fintype Y'] [DecidableEq Y'] [Fintype X'] [DecidableEq X']
(cX : X → Color) (cY : Y → Color)
lemma lorentzAction_basis (Λ : LorentzGroup d) (i : IndexValue d cX) :
Λ • (basis cX i) = ∑ j, toTensorRepMat Λ j i • basis cX j := by
funext k
simp only [lorentzAction, MonoidHom.coe_mk, OneHom.coe_mk,
LinearMap.coe_mk, AddHom.coe_mk]
rw [Finset.sum_apply]
rw [Finset.sum_eq_single_of_mem i, Finset.sum_eq_single_of_mem k]
change _ = toTensorRepMat Λ k i * (Pi.basisFun (IndexValue d cX)) k k
rw [basis]
erw [Pi.basisFun_apply, Pi.basisFun_apply]
simp
exact Finset.mem_univ k
intro b _ hbk
change toTensorRepMat Λ b i • (basis cX) b k = 0
erw [basis, Pi.basisFun_apply]
simp [hbk]
exact Finset.mem_univ i
intro b hb hbk
erw [basis, Pi.basisFun_apply]
simp [hbk]
intro a
subst a
simp_all only [Finset.mem_univ, ne_eq, not_true_eq_false]
end
end RealLorentzTensor

View file

@ -1,195 +0,0 @@
/-
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.LorentzTensor.Real.Constructors
/-!
# Unit of multiplication of Real Lorentz Tensors
The definition of the unit is akin to the definition given in
[Raynor][raynor2021graphical]
for modular operads.
The main results of this file are:
- `mulUnit_right`: The multiplicative unit acts as a right unit for the multiplication of Lorentz
tensors.
- `mulUnit_left`: The multiplicative unit acts as a left unit for the multiplication of Lorentz
tensors.
- `mulUnit_lorentzAction`: The multiplicative unit is invariant under Lorentz transformations.
-/
namespace RealLorentzTensor
variable {d : } {X Y : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
(T : RealLorentzTensor d X) (c : X → Color) (Λ Λ' : LorentzGroup d) {μ : Color}
open Marked
/-!
## Unit of multiplication
-/
/-- The unit for the multiplication of Lorentz tensors. -/
def mulUnit (d : ) (μ : Color) : Marked d (Fin 1) 1 :=
match μ with
| .up => mapIso d ((Equiv.emptySum Empty (Fin (1 + 1))).trans finSumFinEquiv.symm)
(ofMatUpDown 1)
| .down => mapIso d ((Equiv.emptySum Empty (Fin (1 + 1))).trans finSumFinEquiv.symm)
(ofMatDownUp 1)
lemma mulUnit_up_coord : (mulUnit d Colors.up).coord = fun i =>
(1 : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (i (Sum.inl 0)) (i (Sum.inr 0)) := by
rfl
lemma mulUnit_down_coord : (mulUnit d Colors.down).coord = fun i =>
(1 : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (i (Sum.inl 0)) (i (Sum.inr 0)) := by
rfl
@[simp]
lemma mulUnit_markedColor (μ : Color) : (mulUnit d μ).markedColor 0 = τ μ := by
cases μ
case up => rfl
case down => rfl
lemma mulUnit_dual_markedColor (μ : Color) : τ ((mulUnit d μ).markedColor 0) = μ := by
cases μ
case up => rfl
case down => rfl
@[simp]
lemma mulUnit_unmarkedColor (μ : Color) : (mulUnit d μ).unmarkedColor 0 = μ := by
cases μ
case up => rfl
case down => rfl
lemma mulUnit_unmarkedColor_eq_dual_marked (μ : Color) :
(mulUnit d μ).unmarkedColor = τ ∘ (mulUnit d μ).markedColor := by
funext x
fin_cases x
simp only [Fin.zero_eta, Fin.isValue, mulUnit_unmarkedColor, Function.comp_apply,
mulUnit_markedColor]
exact color_eq_dual_symm rfl
lemma mulUnit_coord_diag (μ : Color) (i : (mulUnit d μ).UnmarkedIndexValue) :
(mulUnit d μ).coord (splitIndexValue.symm (i,
indexValueDualIso d (mulUnit_unmarkedColor_eq_dual_marked μ) i)) = 1 := by
cases μ
case' up => rw [mulUnit_up_coord]
case' down => rw [mulUnit_down_coord]
all_goals
simp only [mulUnit]
symm
simp_all only [Fin.isValue, Matrix.one_apply]
split
rfl
next h => exact False.elim (h rfl)
lemma mulUnit_coord_off_diag (μ : Color) (i: (mulUnit d μ).UnmarkedIndexValue)
(b : (mulUnit d μ).MarkedIndexValue)
(hb : b ≠ indexValueDualIso d (mulUnit_unmarkedColor_eq_dual_marked μ) i) :
(mulUnit d μ).coord (splitIndexValue.symm (i, b)) = 0 := by
match μ with
| Colors.up =>
rw [mulUnit_up_coord]
simp only [mulUnit, Matrix.one_apply, Fin.isValue, ite_eq_right_iff, one_ne_zero, imp_false,
ne_eq]
by_contra h
have h1 : (indexValueDualIso d (mulUnit_unmarkedColor_eq_dual_marked (Colors.up)) i) = b := by
funext a
fin_cases a
exact h
exact hb (id (Eq.symm h1))
| Colors.down =>
rw [mulUnit_down_coord]
simp only [mulUnit, Matrix.one_apply, Fin.isValue, ite_eq_right_iff, one_ne_zero, imp_false,
ne_eq]
by_contra h
have h1 : (indexValueDualIso d (mulUnit_unmarkedColor_eq_dual_marked (Colors.down)) i) = b := by
funext a
fin_cases a
exact h
exact hb (id (Eq.symm h1))
lemma mulUnit_right (μ : Color) (T : Marked d X 1) (h : T.markedColor 0 = μ) :
multMarked T (mulUnit d μ) (h.trans (mulUnit_dual_markedColor μ).symm) = T := by
refine ext ?_ ?_
· funext a
match a with
| .inl _ => rfl
| .inr 0 =>
simp only [Fin.isValue, mul_color, Sum.elim_inr, mulUnit_unmarkedColor]
exact h.symm
funext i
rw [mulMarked_indexValue_right]
change ∑ j,
T.coord (splitIndexValue.symm ((indexValueTensorator i).1, _)) *
(mulUnit d μ).coord (splitIndexValue.symm ((indexValueTensorator i).2, j)) = _
let y := indexValueDualIso d (mulUnit_unmarkedColor_eq_dual_marked μ) (indexValueTensorator i).2
erw [Finset.sum_eq_single_of_mem y]
rw [mulUnit_coord_diag]
simp only [Fin.isValue, mul_one]
apply congrArg
funext a
match a with
| .inl a =>
change (indexValueTensorator i).1 a = _
rfl
| .inr 0 =>
change oneMarkedIndexValue
((colorIndexDualCast (Eq.trans h (Eq.symm (mulUnit_dual_markedColor μ)))).symm
(oneMarkedIndexValue.symm y)) 0 = _
rw [indexValueIso_eq_symm, indexValueIso_symm_apply']
simp only [Fin.isValue, oneMarkedIndexValue, colorIndexDualCast, colorIndexCast,
Equiv.coe_fn_symm_mk, indexValueDualIso_apply, Equiv.trans_apply, Equiv.cast_apply,
Equiv.symm_trans_apply, Equiv.cast_symm, Equiv.symm_symm, Equiv.apply_symm_apply, cast_cast,
Equiv.coe_fn_mk, Equiv.refl_symm, Equiv.coe_refl, Function.comp_apply, id_eq, mul_color,
Sum.elim_inr, Equiv.refl_apply, cast_inj, y]
rfl
exact Finset.mem_univ y
intro b _ hab
rw [mul_eq_zero]
apply Or.inr
exact mulUnit_coord_off_diag μ (indexValueTensorator i).2 b hab
lemma mulUnit_left (μ : Color) (T : Marked d X 1) (h : T.markedColor 0 = μ) :
multMarked (mulUnit d μ) T ((mulUnit_markedColor μ).trans (congrArg τ h.symm)) =
mapIso d (Equiv.sumComm X (Fin 1)) T := by
rw [← mult_symmd_symm, mulUnit_right]
exact h
lemma mulUnit_lorentzAction (μ : Color) (Λ : LorentzGroup d) :
Λ • mulUnit d μ = mulUnit d μ := by
match μ with
| Colors.up =>
rw [mulUnit]
simp only [Nat.reduceAdd]
rw [← lorentzAction_mapIso]
rw [lorentzAction_ofMatUpDown]
repeat apply congrArg
rw [mul_one]
change (Λ * Λ⁻¹).1 = 1
rw [mul_inv_self Λ]
rfl
| Colors.down =>
rw [mulUnit]
simp only [Nat.reduceAdd]
rw [← lorentzAction_mapIso]
rw [lorentzAction_ofMatDownUp]
repeat apply congrArg
rw [mul_one]
change ((LorentzGroup.transpose Λ⁻¹) * LorentzGroup.transpose Λ).1 = 1
have inv_transpose : (LorentzGroup.transpose Λ⁻¹) = (LorentzGroup.transpose Λ)⁻¹ := by
simp [LorentzGroup.transpose]
rw [inv_transpose]
rw [inv_mul_self]
rfl
end RealLorentzTensor

View file

@ -1,143 +0,0 @@
/-
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.LorentzTensor.Real.LorentzAction
import Mathlib.LinearAlgebra.TensorProduct.Basic
import Mathlib.LinearAlgebra.TensorProduct.Basis
/-!
# Tensor products of real Lorentz Tensors
-/
noncomputable section
namespace RealLorentzTensor
open TensorProduct
open Set LinearMap Submodule
variable {d : } {X Y Y' X' : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
[Fintype Y'] [DecidableEq Y'] [Fintype X'] [DecidableEq X']
(cX : X → Color) (cY : Y → Color)
/-!
## The tensorator and properties thereof
-/
/-- An equivalence between `ColorFiber d (Sum.elim cX cY)` and
`ColorFiber d cX ⊗[] ColorFiber d cY`. This is related to the `tensorator` of
the monoidal functor defined by `ColorFiber`, hence the terminology. -/
def tensorator {cX : X → Color} {cY : Y → Color} :
RealLorentzTensor d (Sum.elim cX cY) ≃ₗ[] RealLorentzTensor d cX ⊗[] RealLorentzTensor d cY :=
(basis (Sum.elim cX cY)).equiv (basisProd cX cY) indexValueTensorator
lemma tensorator_symm_apply {cX : X → Color} {cY : Y → Color}
(X : RealLorentzTensor d cX ⊗[] RealLorentzTensor d cY) (i : IndexValue d (Sum.elim cX cY)) :
(tensorator.symm X) i = (basisProd cX cY).repr X (indexValueTensorator i) := by
rfl
/-- The naturality condition for the `tensorator`. -/
lemma tensorator_mapIso {cX : X → Color} {cY : Y → Color} {cX' : X' → Color}
{cY' : Y' → Color} (eX : X ≃ X') (eY : Y ≃ Y') (hX : cX = cX' ∘ eX) (hY : cY = cY' ∘ eY) :
(@mapIso _ _ d (Equiv.sumCongr eX eY) _ _ (tensorator_mapIso_cond hX hY)).trans tensorator =
tensorator.trans (TensorProduct.congr (mapIso eX hX) (mapIso eY hY)) := by
apply (basis (Sum.elim cX cY)).ext'
intro i
simp
nth_rewrite 2 [tensorator]
simp
rw [Basis.tensorProduct_apply, TensorProduct.congr_tmul, mapIsoFiber_basis]
rw [tensorator]
simp only [basisProd, Basis.equiv_apply]
rw [Basis.tensorProduct_apply, mapIsoFiber_basis, mapIsoFiber_basis]
congr
rw [← Equiv.trans_apply, indexValueTensorator_indexValueMapIso]
rfl
exact hY
rw [← Equiv.trans_apply, indexValueTensorator_indexValueMapIso]
rfl
exact hX
lemma tensorator_lorentzAction (Λ : LorentzGroup d) :
(tensorator).toLinearMap ∘ₗ lorentzAction Λ
= (TensorProduct.map (lorentzAction Λ) (lorentzAction Λ) ) ∘ₗ
((@tensorator d X Y _ _ _ _ cX cY).toLinearMap) := by
apply (basis (Sum.elim cX cY)).ext
intro i
nth_rewrite 2 [tensorator]
simp only [coe_comp, LinearEquiv.coe_coe, Function.comp_apply, Basis.equiv_apply]
rw [basisProd, Basis.tensorProduct_apply, TensorProduct.map_tmul, lorentzAction_basis,
lorentzAction_basis, lorentzAction_basis, map_sum, tmul_sum]
simp only [sum_tmul]
rw [← Equiv.sum_comp (indexValueTensorator).symm, Fintype.sum_prod_type, Finset.sum_comm]
refine Finset.sum_congr rfl (fun j _ => (Finset.sum_congr rfl (fun k _ => ?_)))
rw [tmul_smul, smul_tmul, tmul_smul, smul_smul, map_smul]
congr 1
· rw [toTensorRepMat_of_indexValueSumEquiv, Equiv.apply_symm_apply, mul_comm]
· simp [tensorator]
lemma tensorator_lorentzAction_apply (Λ : LorentzGroup d) (T : RealLorentzTensor d (Sum.elim cX cY)) :
tensorator (Λ • T) =
TensorProduct.map (lorentzAction Λ) (lorentzAction Λ) (tensorator T) := by
erw [← LinearMap.comp_apply, tensorator_lorentzAction]
rfl
/-!
## Decomposing tensors based on embeddings
-/
def decompEmbedSet (f : Y ↪ X) :
X ≃ {x // x ∈ (Finset.image f Finset.univ)ᶜ} ⊕ Y :=
(Equiv.Set.sumCompl (Set.range ⇑f)).symm.trans <|
(Equiv.sumComm _ _).trans <|
Equiv.sumCongr ((Equiv.subtypeEquivRight (by simp))) <|
(Function.Embedding.toEquivRange f).symm
def decompEmbedColorLeft (f : Y ↪ X) : {x // x ∈ (Finset.image f Finset.univ)ᶜ} → Color :=
(cX ∘ (decompEmbedSet f).symm) ∘ Sum.inl
def decompEmbedColorRight (f : Y ↪ X) : Y → Color :=
(cX ∘ (decompEmbedSet f).symm) ∘ Sum.inr
/-- Decomposes a tensor into a tensor product based on an embedding. -/
def decompEmbed (f : Y ↪ X) :
RealLorentzTensor d cX ≃ₗ[] RealLorentzTensor d (decompEmbedColorLeft cX f) ⊗[] RealLorentzTensor d (cX ∘ f) :=
(@mapIso _ _ d (decompEmbedSet f) cX (Sum.elim (decompEmbedColorLeft cX f)
(decompEmbedColorRight cX f)) (by
simpa [decompEmbedColorLeft, decompEmbedColorRight] using (Equiv.comp_symm_eq _ _ _).mp rfl)).trans
tensorator
lemma decompEmbed_lorentzAction_apply {f : Y ↪ X} (Λ : LorentzGroup d) (T : RealLorentzTensor d cX) :
decompEmbed cX f (Λ • T) =
TensorProduct.map (lorentzAction Λ) (lorentzAction Λ) (decompEmbed cX f T) := by
rw [decompEmbed]
simp
rw [lorentzAction_mapIso]
exact tensorator_lorentzAction_apply (decompEmbedColorLeft cX f) (decompEmbedColorRight cX f) Λ
((mapIso (decompEmbedSet f) (decompEmbed.proof_2 cX f)) T)
def decompEmbedProd (f : X' ↪ X) (g : Y' ↪ Y) :
RealLorentzTensor d cX ⊗[] RealLorentzTensor d cY ≃ₗ[]
RealLorentzTensor d (Sum.elim (decompEmbedColorLeft cX f) (decompEmbedColorLeft cY g))
⊗[] (RealLorentzTensor d (cX ∘ f) ⊗[] RealLorentzTensor d (cY ∘ g)) :=
(TensorProduct.congr (decompEmbed cX f) (decompEmbed cY g)).trans <|
(TensorProduct.assoc _ _ _).trans <|
(TensorProduct.congr (LinearEquiv.refl _)
((TensorProduct.assoc _ _ _).symm.trans <|
(TensorProduct.congr (TensorProduct.comm _ _ _) (LinearEquiv.refl _)).trans <|
(TensorProduct.assoc _ _ _))).trans <|
(TensorProduct.assoc _ _ _).symm.trans <|
(TensorProduct.congr tensorator.symm (LinearEquiv.refl _))
end RealLorentzTensor
end