/- 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.Basic import Mathlib.RepresentationTheory.Basic /-! # Group actions on tensor structures -/ noncomputable section open TensorProduct variable {R : Type} [CommSemiring R] /-! TODO: Add preservation of the unit, and the metric. -/ /-- A multiplicative action on a tensor structure (e.g. the action of the Lorentz group on real Lorentz tensors). -/ class MulActionTensor (G : Type) [Monoid G] (๐“ฃ : TensorStructure R) where /-- For each color `ฮผ` a representation of `G` on `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 ฮผ /-- The invariance of the metric under the group action. -/ metric_inv : โˆ€ ฮผ g, (TensorProduct.map (repColorModule ฮผ g) (repColorModule ฮผ g)) (๐“ฃ.metric ฮผ) = ๐“ฃ.metric ฮผ namespace MulActionTensor variable {G H : Type} [Group G] [Group H] 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` -/ /-- The `MulActionTensor` defined by restriction along a group homomorphism. -/ 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] metric_inv ฮผ h := by simp only [MonoidHom.coe_comp, Function.comp_apply] rw [metric_inv] /-- The trivial `MulActionTensor` defined via trivial representations. -/ def trivial : MulActionTensor G ๐“ฃ where repColorModule ฮผ := Representation.trivial R contrDual_inv ฮผ g := by simp only [Representation.trivial, MonoidHom.one_apply, TensorProduct.map_one] rfl metric_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} /-! # Equivariance properties involving modules -/ @[simp] lemma contrDual_equivariant_tmul (g : G) (x : ๐“ฃ.ColorModule ฮผ) (y : ๐“ฃ.ColorModule (๐“ฃ.ฯ„ ฮผ)) : (๐“ฃ.contrDual ฮผ ((repColorModule ฮผ g x) โŠ—โ‚œ[R] (repColorModule (๐“ฃ.ฯ„ ฮผ) g y))) = ๐“ฃ.contrDual ฮผ (x โŠ—โ‚œ[R] y) := by trans (๐“ฃ.contrDual ฮผ โˆ˜โ‚— TensorProduct.map (repColorModule ฮผ g) (repColorModule (๐“ฃ.ฯ„ ฮผ) g)) (x โŠ—โ‚œ[R] y) rfl rw [contrDual_inv] @[simp] lemma colorModuleCast_equivariant_apply (h : ฮผ = ฮฝ) (g : G) (x : ๐“ฃ.ColorModule ฮผ) : (๐“ฃ.colorModuleCast h) (repColorModule ฮผ g x) = (repColorModule ฮฝ g) (๐“ฃ.colorModuleCast h x) := by subst h simp [colorModuleCast] @[simp] lemma contrRightAux_contrDual_equivariant_tmul (g : G) (m : ๐“ฃ.ColorModule ฮฝ โŠ—[R] ๐“ฃ.ColorModule ฮผ) (x : ๐“ฃ.ColorModule (๐“ฃ.ฯ„ ฮผ)) : (contrRightAux (๐“ฃ.contrDual ฮผ)) ((TensorProduct.map (repColorModule ฮฝ g) (repColorModule ฮผ g) m) โŠ—โ‚œ[R] (repColorModule (๐“ฃ.ฯ„ ฮผ) g x)) = repColorModule ฮฝ g ((contrRightAux (๐“ฃ.contrDual ฮผ)) (m โŠ—โ‚œ[R] x)) := by refine TensorProduct.induction_on m (by simp) ?_ ?_ ยท intro y z simp [contrRightAux] ยท intro x z h1 h2 simp [add_tmul, LinearMap.map_add, h1, h2] /-! ## 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) map_one' := by simp_all only [_root_.map_one, PiTensorProduct.map_one] map_mul' g g' := by simp_all only [_root_.map_mul] exact PiTensorProduct.map_mul _ _ local infixl:101 " โ€ข " => ๐“ฃ.rep @[simp] lemma repColorModule_colorModuleCast (h : ฮผ = ฮฝ) (g : G) : (repColorModule ฮฝ g) โˆ˜โ‚— (๐“ฃ.colorModuleCast h).toLinearMap = (๐“ฃ.colorModuleCast h).toLinearMap โˆ˜โ‚— (repColorModule ฮผ g) := by apply LinearMap.ext intro x simp [colorModuleCast_equivariant_apply] @[simp] lemma rep_mapIso (e : X โ‰ƒ Y) (h : cX = cY โˆ˜ e) (g : G) : (๐“ฃ.rep g) โˆ˜โ‚— (๐“ฃ.mapIso e h).toLinearMap = (๐“ฃ.mapIso e h).toLinearMap โˆ˜โ‚— ๐“ฃ.rep g := by apply PiTensorProduct.ext apply MultilinearMap.ext intro x simp only [LinearMap.compMultilinearMap_apply, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply] erw [mapIso_tprod] simp [rep, colorModuleCast_equivariant_apply] 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] apply congrArg funext i subst h simp [colorModuleCast_equivariant_apply] @[simp] lemma rep_mapIso_apply (e : X โ‰ƒ Y) (h : cX = cY โˆ˜ e) (g : G) (x : ๐“ฃ.Tensor cX) : (๐“ฃ.mapIso e h) (g โ€ข x) = g โ€ข (๐“ฃ.mapIso e h x) := by trans ((๐“ฃ.rep g) โˆ˜โ‚— (๐“ฃ.mapIso e h).toLinearMap) x simp rfl @[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 simp [rep] change (PiTensorProduct.map _) ((PiTensorProduct.tprod R) f) = _ rw [PiTensorProduct.map_tprod] /-! ## Group acting on tensor products -/ lemma tensoratorEquiv_equivariant (g : G) : (๐“ฃ.tensoratorEquiv cX cY) โˆ˜โ‚— (TensorProduct.map (๐“ฃ.rep g) (๐“ฃ.rep g)) = ๐“ฃ.rep g โˆ˜โ‚— (๐“ฃ.tensoratorEquiv cX cY).toLinearMap := by apply tensorProd_piTensorProd_ext intro p q simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, map_tmul, rep_tprod, tensoratorEquiv_tmul_tprod] apply congrArg funext x match x with | Sum.inl x => rfl | Sum.inr x => rfl @[simp] lemma tensoratorEquiv_equivariant_apply (g : G) (x : ๐“ฃ.Tensor cX โŠ—[R] ๐“ฃ.Tensor cY) : (๐“ฃ.tensoratorEquiv cX cY) ((TensorProduct.map (๐“ฃ.rep g) (๐“ฃ.rep g)) x) = (๐“ฃ.rep g) ((๐“ฃ.tensoratorEquiv cX cY) x) := by trans ((๐“ฃ.tensoratorEquiv cX cY) โˆ˜โ‚— (TensorProduct.map (๐“ฃ.rep g) (๐“ฃ.rep g))) x rfl rw [tensoratorEquiv_equivariant] rfl lemma rep_tensoratorEquiv_tmul (g : G) (x : ๐“ฃ.Tensor cX) (y : ๐“ฃ.Tensor cY) : (๐“ฃ.tensoratorEquiv cX cY) ((g โ€ข x) โŠ—โ‚œ[R] (g โ€ข y)) = g โ€ข ((๐“ฃ.tensoratorEquiv cX cY) (x โŠ—โ‚œ[R] y)) := by nth_rewrite 1 [โ† tensoratorEquiv_equivariant_apply] rfl lemma rep_tensoratorEquiv_symm (g : G) : (๐“ฃ.tensoratorEquiv cX cY).symm โˆ˜โ‚— ๐“ฃ.rep g = (TensorProduct.map (๐“ฃ.rep g) (๐“ฃ.rep g)) โˆ˜โ‚— (๐“ฃ.tensoratorEquiv cX cY).symm.toLinearMap := by rw [LinearEquiv.eq_comp_toLinearMap_symm, LinearMap.comp_assoc, LinearEquiv.toLinearMap_symm_comp_eq] exact Eq.symm (tensoratorEquiv_equivariant ๐“ฃ g) @[simp] lemma rep_tensoratorEquiv_symm_apply (g : G) (x : ๐“ฃ.Tensor (Sum.elim cX cY)) : (๐“ฃ.tensoratorEquiv cX cY).symm ((๐“ฃ.rep g) x) = (TensorProduct.map (๐“ฃ.rep g) (๐“ฃ.rep g)) ((๐“ฃ.tensoratorEquiv cX cY).symm x) := by trans ((๐“ฃ.tensoratorEquiv cX cY).symm โˆ˜โ‚— ๐“ฃ.rep g) x rfl rw [rep_tensoratorEquiv_symm] rfl @[simp] lemma rep_lid (g : G) : TensorProduct.lid R (๐“ฃ.Tensor cX) โˆ˜โ‚— (TensorProduct.map (LinearMap.id) (๐“ฃ.rep g)) = (๐“ฃ.rep g) โˆ˜โ‚— (TensorProduct.lid R (๐“ฃ.Tensor cX)).toLinearMap := by apply TensorProduct.ext' intro r y simp @[simp] lemma rep_lid_apply (g : G) (x : R โŠ—[R] ๐“ฃ.Tensor cX) : (TensorProduct.lid R (๐“ฃ.Tensor cX)) ((TensorProduct.map (LinearMap.id) (๐“ฃ.rep g)) x) = (๐“ฃ.rep g) ((TensorProduct.lid R (๐“ฃ.Tensor cX)).toLinearMap x) := by trans ((TensorProduct.lid R (๐“ฃ.Tensor cX)) โˆ˜โ‚— (TensorProduct.map (LinearMap.id) (๐“ฃ.rep g))) x rfl rw [rep_lid] rfl end TensorStructure end