/- 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.MulActionTensor /-! # Rising and Lowering of indices We use the term `dualize` to describe the more general version of rising and lowering of indices. In particular, rising and lowering indices corresponds taking the color of that index to its dual. -/ noncomputable section open TensorProduct namespace TensorColor variable {๐“’ : TensorColor} [DecidableEq ๐“’.Color] [Fintype ๐“’.Color] variable {d : โ„•} {X Y Y' Z W C P : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] [Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W] [Fintype C] [DecidableEq C] [Fintype P] [DecidableEq P] /-! ## Dual maps -/ namespace ColorMap variable (cX : ๐“’.ColorMap X) /-- Given an equivalence `C โŠ• P โ‰ƒ X` the color map obtained by `cX` by dualising all indices in `C`. -/ def partDual (e : C โŠ• P โ‰ƒ X) : ๐“’.ColorMap X := (Sum.elim (๐“’.ฯ„ โˆ˜ cX โˆ˜ e โˆ˜ Sum.inl) (cX โˆ˜ e โˆ˜ Sum.inr) โˆ˜ e.symm) /-- Two color maps are said to be dual if their quotents are dual. -/ def DualMap (cโ‚ cโ‚‚ : ๐“’.ColorMap X) : Prop := ๐“’.colorQuot โˆ˜ cโ‚ = ๐“’.colorQuot โˆ˜ cโ‚‚ namespace DualMap variable {cโ‚ cโ‚‚ cโ‚ƒ : ๐“’.ColorMap X} variable {n : โ„•} /-- The bool which if `๐“’.colorQuot (cโ‚ i) = ๐“’.colorQuot (cโ‚‚ i)` is true for all `i`. -/ def boolFin (cโ‚ cโ‚‚ : ๐“’.ColorMap (Fin n)) : Bool := (Fin.list n).all fun i => if ๐“’.colorQuot (cโ‚ i) = ๐“’.colorQuot (cโ‚‚ i) then true else false lemma boolFin_DualMap {cโ‚ cโ‚‚ : ๐“’.ColorMap (Fin n)} (h : boolFin cโ‚ cโ‚‚ = true) : DualMap cโ‚ cโ‚‚ := by simp [boolFin] at h simp [DualMap] funext x have h2 {n : โ„•} (m : Fin n) : m โˆˆ Fin.list n := by have h1' : (Fin.list n)[m] = m := by erw [Fin.getElem_list] rfl rw [โ† h1'] apply List.getElem_mem exact h x (h2 _) /-- The bool which is ture if `๐“’.colorQuot (cโ‚ i) = ๐“’.colorQuot (cโ‚‚ i)` is true for all `i`. -/ def boolFin' (cโ‚ cโ‚‚ : ๐“’.ColorMap (Fin n)) : Bool := โˆ€ (i : Fin n), ๐“’.colorQuot (cโ‚ i) = ๐“’.colorQuot (cโ‚‚ i) lemma boolFin'_DualMap {cโ‚ cโ‚‚ : ๐“’.ColorMap (Fin n)} (h : boolFin' cโ‚ cโ‚‚ = true) : DualMap cโ‚ cโ‚‚ := by simp [boolFin'] at h simp [DualMap] funext x exact h x lemma refl : DualMap cโ‚ cโ‚ := rfl lemma symm (h : DualMap cโ‚ cโ‚‚) : DualMap cโ‚‚ cโ‚ := by rw [DualMap] at h โŠข exact h.symm lemma trans (h : DualMap cโ‚ cโ‚‚) (h' : DualMap cโ‚‚ cโ‚ƒ) : DualMap cโ‚ cโ‚ƒ := by rw [DualMap] at h h' โŠข exact h.trans h' /-- The splitting of `X` given two color maps based on the equality of the color. -/ def split (cโ‚ cโ‚‚ : ๐“’.ColorMap X) : { x // cโ‚ x โ‰  cโ‚‚ x} โŠ• { x // cโ‚ x = cโ‚‚ x} โ‰ƒ X := ((Equiv.Set.sumCompl {x | cโ‚ x = cโ‚‚ x}).symm.trans (Equiv.sumComm _ _)).symm lemma dual_eq_of_neq (h : DualMap cโ‚ cโ‚‚) {x : X} (h' : cโ‚ x โ‰  cโ‚‚ x) : ๐“’.ฯ„ (cโ‚ x) = cโ‚‚ x := by rw [DualMap] at h have h1 := congrFun h x simp [colorQuot, HasEquiv.Equiv, Setoid.r, colorRel] at h1 simp_all only [ne_eq, false_or] exact ๐“’.ฯ„_involutive (cโ‚‚ x) @[simp] lemma split_dual (h : DualMap cโ‚ cโ‚‚) : cโ‚.partDual (split cโ‚ cโ‚‚) = cโ‚‚ := by rw [partDual, Equiv.comp_symm_eq] funext x match x with | Sum.inl x => exact h.dual_eq_of_neq x.2 | Sum.inr x => exact x.2 @[simp] lemma split_dual' (h : DualMap cโ‚ cโ‚‚) : cโ‚‚.partDual (split cโ‚ cโ‚‚) = cโ‚ := by rw [partDual, Equiv.comp_symm_eq] funext x match x with | Sum.inl x => change ๐“’.ฯ„ (cโ‚‚ x) = cโ‚ x rw [โ† h.dual_eq_of_neq x.2] exact (๐“’.ฯ„_involutive (cโ‚ x)) | Sum.inr x => exact x.2.symm end DualMap end ColorMap end TensorColor variable {R : Type} [CommSemiring R] namespace TensorStructure variable (๐“ฃ : TensorStructure R) variable {d : โ„•} {X Y Y' Z W C P : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] [Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W] [Fintype C] [DecidableEq C] [Fintype P] [DecidableEq P] {cX cX2 : ๐“ฃ.ColorMap X} {cY : ๐“ฃ.ColorMap Y} {cZ : ๐“ฃ.ColorMap Z} {cW : ๐“ฃ.ColorMap W} {cY' : ๐“ฃ.ColorMap Y'} {ฮผ ฮฝ: ๐“ฃ.Color} variable {G H : Type} [Group G] [Group H] [MulActionTensor G ๐“ฃ] local infixl:101 " โ€ข " => ๐“ฃ.rep /-! ## Properties of the unit -/ /-! TODO: Move -/ lemma unit_lhs_eq (x : ๐“ฃ.ColorModule ฮผ) (y : ๐“ฃ.ColorModule (๐“ฃ.ฯ„ ฮผ) โŠ—[R] ๐“ฃ.ColorModule ฮผ) : contrLeftAux (๐“ฃ.contrDual ฮผ) (x โŠ—โ‚œ[R] y) = (contrRightAux (๐“ฃ.contrDual (๐“ฃ.ฯ„ ฮผ))) ((TensorProduct.comm R _ _) y โŠ—โ‚œ[R] (๐“ฃ.colorModuleCast (๐“ฃ.ฯ„_involutive ฮผ).symm) x) := by refine TensorProduct.induction_on y (by rfl) ?_ (fun z1 z2 h1 h2 => ?_) ยท intro x1 x2 simp only [contrRightAux, LinearEquiv.refl_toLinearMap, comm_tmul, colorModuleCast, Equiv.cast_symm, LinearEquiv.coe_mk, Equiv.cast_apply, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, assoc_tmul, map_tmul, LinearMap.id_coe, id_eq, contrDual_symm', cast_cast, cast_eq, rid_tmul] rfl ยท simp only [map_add, add_tmul] rw [โ† h1, โ† h2, tmul_add, LinearMap.map_add] @[simp] lemma unit_lid : (contrRightAux (๐“ฃ.contrDual (๐“ฃ.ฯ„ ฮผ))) ((TensorProduct.comm R _ _) (๐“ฃ.unit ฮผ) โŠ—โ‚œ[R] (๐“ฃ.colorModuleCast (๐“ฃ.ฯ„_involutive ฮผ).symm) x) = x := by have h1 := ๐“ฃ.unit_rid ฮผ x rw [โ† unit_lhs_eq] exact h1 /-! ## Properties of the metric -/ @[simp] lemma metric_cast (h : ฮผ = ฮฝ) : (TensorProduct.congr (๐“ฃ.colorModuleCast h) (๐“ฃ.colorModuleCast h)) (๐“ฃ.metric ฮผ) = ๐“ฃ.metric ฮฝ := by subst h erw [congr_refl_refl] rfl @[simp] lemma metric_contrRight_unit (ฮผ : ๐“ฃ.Color) (x : ๐“ฃ.ColorModule ฮผ) : (contrRightAux (๐“ฃ.contrDual ฮผ)) (๐“ฃ.metric ฮผ โŠ—โ‚œ[R] ((contrRightAux (๐“ฃ.contrDual (๐“ฃ.ฯ„ ฮผ))) (๐“ฃ.metric (๐“ฃ.ฯ„ ฮผ) โŠ—โ‚œ[R] (๐“ฃ.colorModuleCast (๐“ฃ.ฯ„_involutive ฮผ).symm x)))) = x := by change (contrRightAux (๐“ฃ.contrDual ฮผ) โˆ˜โ‚— TensorProduct.map (LinearMap.id) (contrRightAux (๐“ฃ.contrDual (๐“ฃ.ฯ„ ฮผ)))) (๐“ฃ.metric ฮผ โŠ—โ‚œ[R] ๐“ฃ.metric (๐“ฃ.ฯ„ ฮผ) โŠ—โ‚œ[R] (๐“ฃ.colorModuleCast (๐“ฃ.ฯ„_involutive ฮผ).symm x)) = _ rw [contrRightAux_comp] simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, assoc_symm_tmul, map_tmul, LinearMap.id_coe, id_eq] rw [๐“ฃ.metric_dual] exact unit_lid ๐“ฃ /-! ## Dualizing -/ /-- Takes a vector with index with dual color to a vector with index the underlying color. Obtained by contraction with the metric. -/ def dualizeSymm (ฮผ : ๐“ฃ.Color) : ๐“ฃ.ColorModule (๐“ฃ.ฯ„ ฮผ) โ†’โ‚—[R] ๐“ฃ.ColorModule ฮผ := contrRightAux (๐“ฃ.contrDual ฮผ) โˆ˜โ‚— TensorProduct.lTensorHomToHomLTensor R _ _ _ (๐“ฃ.metric ฮผ โŠ—โ‚œ[R] LinearMap.id) /-- Takes a vector to a vector with the dual color index. Obtained by contraction with the metric. -/ def dualizeFun (ฮผ : ๐“ฃ.Color) : ๐“ฃ.ColorModule ฮผ โ†’โ‚—[R] ๐“ฃ.ColorModule (๐“ฃ.ฯ„ ฮผ) := ๐“ฃ.dualizeSymm (๐“ฃ.ฯ„ ฮผ) โˆ˜โ‚— (๐“ฃ.colorModuleCast (๐“ฃ.ฯ„_involutive ฮผ).symm).toLinearMap /-- Equivalence between the module with a color `ฮผ` and the module with color `๐“ฃ.ฯ„ ฮผ` obtained by contraction with the metric. -/ def dualizeModule (ฮผ : ๐“ฃ.Color) : ๐“ฃ.ColorModule ฮผ โ‰ƒโ‚—[R] ๐“ฃ.ColorModule (๐“ฃ.ฯ„ ฮผ) := by refine LinearEquiv.ofLinear (๐“ฃ.dualizeFun ฮผ) (๐“ฃ.dualizeSymm ฮผ) ?_ ?_ ยท apply LinearMap.ext intro x simp [dualizeFun, dualizeSymm, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, lTensorHomToHomLTensor_apply, LinearMap.id_coe, id_eq, contrDual_symm_contrRightAux_apply_tmul, metric_cast] ยท apply LinearMap.ext intro x simp only [dualizeSymm, dualizeFun, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, lTensorHomToHomLTensor_apply, LinearMap.id_coe, id_eq, metric_contrRight_unit] @[simp] lemma dualizeModule_equivariant (g : G) : (๐“ฃ.dualizeModule ฮผ) โˆ˜โ‚— ((MulActionTensor.repColorModule ฮผ) g) = (MulActionTensor.repColorModule (๐“ฃ.ฯ„ ฮผ) g) โˆ˜โ‚— (๐“ฃ.dualizeModule ฮผ) := by apply LinearMap.ext intro x simp only [dualizeModule, dualizeFun, dualizeSymm, LinearEquiv.ofLinear_toLinearMap, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, colorModuleCast_equivariant_apply, lTensorHomToHomLTensor_apply, LinearMap.id_coe, id_eq] nth_rewrite 1 [โ† MulActionTensor.metric_inv (๐“ฃ.ฯ„ ฮผ) g] exact contrRightAux_contrDual_equivariant_tmul ๐“ฃ g (๐“ฃ.metric (๐“ฃ.ฯ„ ฮผ)) ((๐“ฃ.colorModuleCast (dualizeFun.proof_3 ๐“ฃ ฮผ)) x) @[simp] lemma dualizeModule_equivariant_apply (g : G) (x : ๐“ฃ.ColorModule ฮผ) : (๐“ฃ.dualizeModule ฮผ) ((MulActionTensor.repColorModule ฮผ) g x) = (MulActionTensor.repColorModule (๐“ฃ.ฯ„ ฮผ) g) (๐“ฃ.dualizeModule ฮผ x) := by trans ((๐“ฃ.dualizeModule ฮผ) โˆ˜โ‚— ((MulActionTensor.repColorModule ฮผ) g)) x ยท rfl ยท rw [dualizeModule_equivariant] rfl /-- Dualizes the color of all indicies of a tensor by contraction with the metric. -/ def dualizeAll : ๐“ฃ.Tensor cX โ‰ƒโ‚—[R] ๐“ฃ.Tensor (๐“ฃ.ฯ„ โˆ˜ cX) := by refine LinearEquiv.ofLinear (PiTensorProduct.map (fun x => (๐“ฃ.dualizeModule (cX x)).toLinearMap)) (PiTensorProduct.map (fun x => (๐“ฃ.dualizeModule (cX x)).symm.toLinearMap)) ?_ ?_ all_goals apply LinearMap.ext refine fun x โ†ฆ PiTensorProduct.induction_on' x ?_ (by intro a b hx a simp [map_add, add_tmul, hx] simp_all only [Function.comp_apply, LinearMap.coe_comp, LinearMap.id_coe, id_eq]) intro rx fx simp only [Function.comp_apply, PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMapClass.map_smul, LinearMap.coe_comp, LinearMap.id_coe, id_eq] apply congrArg change (PiTensorProduct.map _) ((PiTensorProduct.map _) ((PiTensorProduct.tprod R) fx)) = _ rw [PiTensorProduct.map_tprod, PiTensorProduct.map_tprod] apply congrArg simp @[simp] lemma dualizeAll_equivariant (g : G) : (๐“ฃ.dualizeAll.toLinearMap) โˆ˜โ‚— (@rep R _ G _ ๐“ฃ _ X cX g) = ๐“ฃ.rep g โˆ˜โ‚— (๐“ฃ.dualizeAll.toLinearMap) := by apply LinearMap.ext intro x simp [dualizeAll] refine PiTensorProduct.induction_on' x ?_ (by intro a b hx a simp [map_add, add_tmul, hx] simp_all only [Function.comp_apply, LinearMap.coe_comp, LinearMap.id_coe, id_eq]) intro rx fx simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMapClass.map_smul, rep_tprod] apply congrArg change (PiTensorProduct.map _) ((PiTensorProduct.tprod R) _) = (๐“ฃ.rep g) ((PiTensorProduct.map _) ((PiTensorProduct.tprod R) fx)) rw [PiTensorProduct.map_tprod, PiTensorProduct.map_tprod] simp lemma dualize_cond (e : C โŠ• P โ‰ƒ X) : cX = Sum.elim (cX โˆ˜ e โˆ˜ Sum.inl) (cX โˆ˜ e โˆ˜ Sum.inr) โˆ˜ e.symm := by rw [Equiv.eq_comp_symm] funext x match x with | Sum.inl x => rfl | Sum.inr x => rfl lemma dualize_cond' (e : C โŠ• P โ‰ƒ X) : Sum.elim (๐“ฃ.ฯ„ โˆ˜ cX โˆ˜ โ‡‘e โˆ˜ Sum.inl) (cX โˆ˜ โ‡‘e โˆ˜ Sum.inr) = (Sum.elim (๐“ฃ.ฯ„ โˆ˜ cX โˆ˜ โ‡‘e โˆ˜ Sum.inl) (cX โˆ˜ โ‡‘e โˆ˜ Sum.inr) โˆ˜ โ‡‘e.symm) โˆ˜ โ‡‘e := by funext x match x with | Sum.inl x => simp | Sum.inr x => simp /-- Given an equivalence `C โŠ• P โ‰ƒ X` dualizes those indices of a tensor which correspond to `C` whilst leaving the indices `P` invariant. -/ def dualize (e : C โŠ• P โ‰ƒ X) : ๐“ฃ.Tensor cX โ‰ƒโ‚—[R] ๐“ฃ.Tensor (cX.partDual e) := ๐“ฃ.mapIso e.symm (๐“ฃ.dualize_cond e) โ‰ชโ‰ซโ‚— (๐“ฃ.tensoratorEquiv _ _).symm โ‰ชโ‰ซโ‚— TensorProduct.congr ๐“ฃ.dualizeAll (LinearEquiv.refl _ _) โ‰ชโ‰ซโ‚— (๐“ฃ.tensoratorEquiv _ _) โ‰ชโ‰ซโ‚— ๐“ฃ.mapIso e (๐“ฃ.dualize_cond' e) /-- Dualizing indices is equivariant with respect to the group action. This is the applied version of this statement. -/ @[simp] lemma dualize_equivariant_apply (e : C โŠ• P โ‰ƒ X) (g : G) (x : ๐“ฃ.Tensor cX) : ๐“ฃ.dualize e (g โ€ข x) = g โ€ข (๐“ฃ.dualize e x) := by simp only [dualize, TensorProduct.congr, LinearEquiv.refl_toLinearMap, LinearEquiv.refl_symm, LinearEquiv.trans_apply, rep_mapIso_apply, rep_tensoratorEquiv_symm_apply, LinearEquiv.ofLinear_apply] rw [โ† LinearMap.comp_apply (TensorProduct.map _ _), โ† TensorProduct.map_comp] simp only [dualizeAll_equivariant, LinearMap.id_comp] have h1 {M N A B C : Type} [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid A] [AddCommMonoid B] [AddCommMonoid C] [Module R M] [Module R N] [Module R A] [Module R B] [Module R C] (f : M โ†’โ‚—[R] N) (g : A โ†’โ‚—[R] B) (h : B โ†’โ‚—[R] C) : TensorProduct.map (h โˆ˜โ‚— g) f = TensorProduct.map h f โˆ˜โ‚— TensorProduct.map g (LinearMap.id) := ext rfl rw [h1, LinearMap.coe_comp, Function.comp_apply] simp only [tensoratorEquiv_equivariant_apply, rep_mapIso_apply] end TensorStructure end