feat: Update contraction for index notation

This commit is contained in:
jstoobysmith 2024-10-10 08:57:22 +00:00
parent e90e41751e
commit 3f5eb58db4
10 changed files with 401 additions and 153 deletions

View file

@ -0,0 +1,508 @@
/-
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.Tensors.OverColor.Basic
import HepLean.Mathematics.PiTensorProduct
/-!
## Monodial functor from color cat.
-/
namespace Fermion
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open IndexNotation
open CategoryTheory
/-- The colors associated with complex representations of SL(2, ) of intrest to physics. -/
inductive Color
| upL : Color
| downL : Color
| upR : Color
| downR : Color
| up : Color
| down : Color
def τ : Color → Color
| Color.upL => Color.downL
| Color.downL => Color.upL
| Color.upR => Color.downR
| Color.downR => Color.upR
| Color.up => Color.down
| Color.down => Color.up
def evalNo : Color →
| Color.upL => 2
| Color.downL => 2
| Color.upR => 2
| Color.downR => 2
| Color.up => 4
| Color.down => 4
/-- The corresponding representations associated with a color. -/
def colorToRep (c : Color) : Rep SL(2, ) :=
match c with
| Color.upL => altLeftHanded
| Color.downL => leftHanded
| Color.upR => altRightHanded
| Color.downR => rightHanded
| Color.up => Lorentz.complexContr
| Color.down => Lorentz.complexCo
/-- The linear equivalence between `colorToRep c1` and `colorToRep c2` when `c1 = c2`. -/
def colorToRepCongr {c1 c2 : Color} (h : c1 = c2) : colorToRep c1 ≃ₗ[] colorToRep c2 where
toFun := Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)
invFun := (Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)).symm
map_add' x y := by
subst h
rfl
map_smul' x y := by
subst h
rfl
left_inv x := Equiv.symm_apply_apply (Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)) x
right_inv x := Equiv.apply_symm_apply (Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)) x
lemma colorToRepCongr_comm_ρ {c1 c2 : Color} (h : c1 = c2) (M : SL(2, )) (x : (colorToRep c1)) :
(colorToRepCongr h) ((colorToRep c1).ρ M x) = (colorToRep c2).ρ M ((colorToRepCongr h) x) := by
subst h
rfl
namespace colorFun
/-- Given a object in `OverColor Color` the correpsonding tensor product of representations. -/
def obj' (f : OverColor Color) : Rep SL(2, ) := Rep.of {
toFun := fun M => PiTensorProduct.map (fun x => (colorToRep (f.hom x)).ρ M),
map_one' := by
simp
map_mul' := fun M N => by
simp only [CategoryTheory.Functor.id_obj, _root_.map_mul]
ext x : 2
simp only [LinearMap.compMultilinearMap_apply, PiTensorProduct.map_tprod, LinearMap.mul_apply]}
lemma obj'_ρ (f : OverColor Color) (M : SL(2, )) : (obj' f).ρ M =
PiTensorProduct.map (fun x => (colorToRep (f.hom x)).ρ M) := rfl
lemma obj'_ρ_tprod (f : OverColor Color) (M : SL(2, ))
(x : (i : f.left) → CoeSort.coe (colorToRep (f.hom i))) :
(obj' f).ρ M ((PiTensorProduct.tprod ) x) =
PiTensorProduct.tprod (fun i => (colorToRep (f.hom i)).ρ M (x i)) := by
rw [obj'_ρ]
change (PiTensorProduct.map fun x => (colorToRep (f.hom x)).ρ M) ((PiTensorProduct.tprod ) x) =
(PiTensorProduct.tprod ) fun i => ((colorToRep (f.hom i)).ρ M) (x i)
rw [PiTensorProduct.map_tprod]
/-- Given a morphism in `OverColor Color` the corresopnding linear equivalence between `obj' _`
induced by reindexing. -/
def mapToLinearEquiv' {f g : OverColor Color} (m : f ⟶ g) : (obj' f).V ≃ₗ[] (obj' g).V :=
(PiTensorProduct.reindex (fun x => colorToRep (f.hom x)) (OverColor.Hom.toEquiv m)).trans
(PiTensorProduct.congr (fun i => colorToRepCongr (OverColor.Hom.toEquiv_symm_apply m i)))
lemma mapToLinearEquiv'_tprod {f g : OverColor Color} (m : f ⟶ g)
(x : (i : f.left) → CoeSort.coe (colorToRep (f.hom i))) :
mapToLinearEquiv' m (PiTensorProduct.tprod x) =
PiTensorProduct.tprod (fun i => (colorToRepCongr (OverColor.Hom.toEquiv_symm_apply m i))
(x ((OverColor.Hom.toEquiv m).symm i))) := by
rw [mapToLinearEquiv']
simp only [CategoryTheory.Functor.id_obj, LinearEquiv.trans_apply]
change (PiTensorProduct.congr fun i => colorToRepCongr _)
((PiTensorProduct.reindex (fun x => CoeSort.coe (colorToRep (f.hom x)))
(OverColor.Hom.toEquiv m)) ((PiTensorProduct.tprod ) x)) = _
rw [PiTensorProduct.reindex_tprod, PiTensorProduct.congr_tprod]
rfl
/-- Given a morphism in `OverColor Color` the corresopnding map of representations induced by
reindexing. -/
def map' {f g : OverColor Color} (m : f ⟶ g) : obj' f ⟶ obj' g where
hom := (mapToLinearEquiv' m).toLinearMap
comm M := by
ext x : 2
refine PiTensorProduct.induction_on' x ?_ (by
intro x y hx hy
simp only [CategoryTheory.Functor.id_obj, map_add, hx, ModuleCat.coe_comp,
Function.comp_apply, hy])
intro r x
simp only [CategoryTheory.Functor.id_obj, PiTensorProduct.tprodCoeff_eq_smul_tprod,
_root_.map_smul, ModuleCat.coe_comp, Function.comp_apply]
apply congrArg
change (mapToLinearEquiv' m) (((obj' f).ρ M) ((PiTensorProduct.tprod ) x)) =
((obj' g).ρ M) ((mapToLinearEquiv' m) ((PiTensorProduct.tprod ) x))
rw [mapToLinearEquiv'_tprod, obj'_ρ_tprod]
erw [mapToLinearEquiv'_tprod, obj'_ρ_tprod]
apply congrArg
funext i
rw [colorToRepCongr_comm_ρ]
end colorFun
/-- The functor between `OverColor Color` and `Rep SL(2, )` taking a map of colors
to the corresponding tensor product representation. -/
@[simps! obj_V_carrier]
def colorFun : OverColor Color ⥤ Rep SL(2, ) where
obj := colorFun.obj'
map := colorFun.map'
map_id f := by
ext x
refine PiTensorProduct.induction_on' x (fun r x => ?_) (fun x y hx hy => by
simp only [CategoryTheory.Functor.id_obj, map_add, hx, ModuleCat.coe_comp,
Function.comp_apply, hy])
simp only [CategoryTheory.Functor.id_obj, PiTensorProduct.tprodCoeff_eq_smul_tprod,
_root_.map_smul, Action.id_hom, ModuleCat.id_apply]
apply congrArg
erw [colorFun.mapToLinearEquiv'_tprod]
exact congrArg _ (funext (fun i => rfl))
map_comp {X Y Z} f g := by
ext x
refine PiTensorProduct.induction_on' x (fun r x => ?_) (fun x y hx hy => by
simp only [CategoryTheory.Functor.id_obj, map_add, hx, ModuleCat.coe_comp,
Function.comp_apply, hy])
simp only [Functor.id_obj, PiTensorProduct.tprodCoeff_eq_smul_tprod, _root_.map_smul,
Action.comp_hom, ModuleCat.coe_comp, Function.comp_apply]
apply congrArg
rw [colorFun.map', colorFun.map', colorFun.map']
change (colorFun.mapToLinearEquiv' (CategoryTheory.CategoryStruct.comp f g))
((PiTensorProduct.tprod ) x) =
(colorFun.mapToLinearEquiv' g) ((colorFun.mapToLinearEquiv' f) ((PiTensorProduct.tprod ) x))
rw [colorFun.mapToLinearEquiv'_tprod, colorFun.mapToLinearEquiv'_tprod]
erw [colorFun.mapToLinearEquiv'_tprod]
refine congrArg _ (funext (fun i => ?_))
simp only [colorToRepCongr, Function.comp_apply, Equiv.cast_symm, LinearEquiv.coe_mk,
Equiv.cast_apply, cast_cast, cast_inj]
rfl
namespace colorFun
open CategoryTheory
open MonoidalCategory
lemma map_tprod {X Y : OverColor Color} (p : (i : X.left) → (colorToRep (X.hom i)))
(f : X ⟶ Y) : (colorFun.map f).hom (PiTensorProduct.tprod p) =
PiTensorProduct.tprod fun (i : Y.left) => colorToRepCongr
(OverColor.Hom.toEquiv_comp_inv_apply f i) (p ((OverColor.Hom.toEquiv f).symm i)) := by
change (colorFun.map' f).hom ((PiTensorProduct.tprod ) p) = _
simp [colorFun.map', mapToLinearEquiv']
erw [LinearEquiv.trans_apply]
change (PiTensorProduct.congr fun i => colorToRepCongr _)
((PiTensorProduct.reindex (fun x => _) (OverColor.Hom.toEquiv f))
((PiTensorProduct.tprod ) p)) = _
rw [PiTensorProduct.reindex_tprod, PiTensorProduct.congr_tprod]
@[simp]
lemma obj_ρ_empty (g : SL(2, )) : (colorFun.obj (𝟙_ (OverColor Color))).ρ g = LinearMap.id := by
erw [colorFun.obj'_ρ]
ext x
refine PiTensorProduct.induction_on' x (fun r x => ?_) <| fun x y hx hy => by
simp only [CategoryTheory.Functor.id_obj, map_add, hx, ModuleCat.coe_comp,
Function.comp_apply, hy]
erw [hx, hy]
rfl
simp only [OverColor.instMonoidalCategoryStruct_tensorUnit_left, Functor.id_obj,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom, PiTensorProduct.tprodCoeff_eq_smul_tprod,
_root_.map_smul, PiTensorProduct.map_tprod, LinearMap.id_coe, id_eq]
apply congrArg
apply congrArg
funext i
exact Empty.elim i
/-- The unit natural isomorphism. -/
def ε : 𝟙_ (Rep SL(2, )) ≅ colorFun.obj (𝟙_ (OverColor Color)) where
hom := {
hom := (PiTensorProduct.isEmptyEquiv Empty).symm.toLinearMap
comm := fun M => by
refine LinearMap.ext (fun x => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorUnit_left,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom,
Action.instMonoidalCategory_tensorUnit_V, Action.tensorUnit_ρ', Functor.id_obj,
Category.id_comp, LinearEquiv.coe_coe]
erw [obj_ρ_empty M]
rfl}
inv := {
hom := (PiTensorProduct.isEmptyEquiv Empty).toLinearMap
comm := fun M => by
refine LinearMap.ext (fun x => ?_)
simp only [Action.instMonoidalCategory_tensorUnit_V, colorFun_obj_V_carrier,
OverColor.instMonoidalCategoryStruct_tensorUnit_left,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom, Functor.id_obj, Action.tensorUnit_ρ']
erw [obj_ρ_empty M]
rfl}
hom_inv_id := by
ext1
simp only [Action.instMonoidalCategory_tensorUnit_V, CategoryStruct.comp,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom,
OverColor.instMonoidalCategoryStruct_tensorUnit_left, Functor.id_obj, Action.Hom.comp_hom,
colorFun_obj_V_carrier, LinearEquiv.comp_coe, LinearEquiv.symm_trans_self,
LinearEquiv.refl_toLinearMap, Action.id_hom]
rfl
inv_hom_id := by
ext1
simp only [CategoryStruct.comp, OverColor.instMonoidalCategoryStruct_tensorUnit_hom,
OverColor.instMonoidalCategoryStruct_tensorUnit_left, Functor.id_obj, Action.Hom.comp_hom,
colorFun_obj_V_carrier, Action.instMonoidalCategory_tensorUnit_V, LinearEquiv.comp_coe,
LinearEquiv.self_trans_symm, LinearEquiv.refl_toLinearMap, Action.id_hom]
rfl
/-- An auxillary equivalence, and trivial, of modules needed to define `μModEquiv`. -/
def colorToRepSumEquiv {X Y : OverColor Color} (i : X.left ⊕ Y.left) :
Sum.elim (fun i => colorToRep (X.hom i)) (fun i => colorToRep (Y.hom i)) i ≃ₗ[]
colorToRep (Sum.elim X.hom Y.hom i) :=
match i with
| Sum.inl _ => LinearEquiv.refl _ _
| Sum.inr _ => LinearEquiv.refl _ _
/-- The equivalence of modules corresonding to the tensorate. -/
def μModEquiv (X Y : OverColor Color) :
(colorFun.obj X ⊗ colorFun.obj Y).V ≃ₗ[] colorFun.obj (X ⊗ Y) :=
HepLean.PiTensorProduct.tmulEquiv ≪≫ₗ PiTensorProduct.congr colorToRepSumEquiv
lemma μModEquiv_tmul_tprod {X Y : OverColor Color}(p : (i : X.left) → (colorToRep (X.hom i)))
(q : (i : Y.left) → (colorToRep (Y.hom i))) :
(μModEquiv X Y) ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q) =
(PiTensorProduct.tprod ) fun i =>
(colorToRepSumEquiv i) (HepLean.PiTensorProduct.elimPureTensor p q i) := by
rw [μModEquiv]
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Action.instMonoidalCategory_tensorObj_V,
Functor.id_obj, Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj]
rw [LinearEquiv.trans_apply]
erw [HepLean.PiTensorProduct.tmulEquiv_tmul_tprod]
change (PiTensorProduct.congr colorToRepSumEquiv) ((PiTensorProduct.tprod )
(HepLean.PiTensorProduct.elimPureTensor p q)) = _
rw [PiTensorProduct.congr_tprod]
rfl
/-- The natural isomorphism corresponding to the tensorate. -/
def μ (X Y : OverColor Color) : colorFun.obj X ⊗ colorFun.obj Y ≅ colorFun.obj (X ⊗ Y) where
hom := {
hom := (μModEquiv X Y).toLinearMap
comm := fun M => by
refine HepLean.PiTensorProduct.induction_tmul (fun p q => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Functor.id_obj, CategoryStruct.comp,
Action.instMonoidalCategory_tensorObj_V, Action.tensor_ρ', LinearMap.coe_comp,
Function.comp_apply]
change (μModEquiv X Y) (((((colorFun.obj X).ρ M) (PiTensorProduct.tprod p)) ⊗ₜ[]
(((colorFun.obj Y).ρ M) (PiTensorProduct.tprod q)))) = ((colorFun.obj (X ⊗ Y)).ρ M)
((μModEquiv X Y) ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
rw [μModEquiv_tmul_tprod]
erw [obj'_ρ_tprod, obj'_ρ_tprod, obj'_ρ_tprod]
rw [μModEquiv_tmul_tprod]
apply congrArg
funext i
match i with
| Sum.inl i =>
rfl
| Sum.inr i =>
rfl
}
inv := {
hom := (μModEquiv X Y).symm.toLinearMap
comm := fun M => by
simp [CategoryStruct.comp]
erw [LinearEquiv.eq_comp_toLinearMap_symm,LinearMap.comp_assoc,
LinearEquiv.toLinearMap_symm_comp_eq]
refine HepLean.PiTensorProduct.induction_tmul (fun p q => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Functor.id_obj, CategoryStruct.comp,
Action.instMonoidalCategory_tensorObj_V, Action.tensor_ρ', LinearMap.coe_comp,
Function.comp_apply]
symm
change (μModEquiv X Y) (((((colorFun.obj X).ρ M) (PiTensorProduct.tprod p)) ⊗ₜ[]
(((colorFun.obj Y).ρ M) (PiTensorProduct.tprod q)))) = ((colorFun.obj (X ⊗ Y)).ρ M)
((μModEquiv X Y) ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
rw [μModEquiv_tmul_tprod]
erw [obj'_ρ_tprod, obj'_ρ_tprod, obj'_ρ_tprod]
rw [μModEquiv_tmul_tprod]
apply congrArg
funext i
match i with
| Sum.inl i =>
rfl
| Sum.inr i =>
rfl}
hom_inv_id := by
ext1
simp only [Action.instMonoidalCategory_tensorObj_V, CategoryStruct.comp, Action.Hom.comp_hom,
colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, LinearEquiv.comp_coe,
LinearEquiv.self_trans_symm, LinearEquiv.refl_toLinearMap, Action.id_hom]
rfl
inv_hom_id := by
ext1
simp only [CategoryStruct.comp, Action.instMonoidalCategory_tensorObj_V, Action.Hom.comp_hom,
colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, LinearEquiv.comp_coe,
LinearEquiv.symm_trans_self, LinearEquiv.refl_toLinearMap, Action.id_hom]
rfl
lemma μ_tmul_tprod {X Y : OverColor Color} (p : (i : X.left) → (colorToRep (X.hom i)))
(q : (i : Y.left) → (colorToRep (Y.hom i))) :
(μ X Y).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q) =
(PiTensorProduct.tprod ) fun i =>
(colorToRepSumEquiv i) (HepLean.PiTensorProduct.elimPureTensor p q i) := by
exact μModEquiv_tmul_tprod p q
lemma μ_natural_left {X Y : OverColor Color} (f : X ⟶ Y) (Z : OverColor Color) :
MonoidalCategory.whiskerRight (colorFun.map f) (colorFun.obj Z) ≫ (μ Y Z).hom =
(μ X Z).hom ≫ colorFun.map (MonoidalCategory.whiskerRight f Z) := by
ext1
refine HepLean.PiTensorProduct.induction_tmul (fun p q => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Functor.id_obj, CategoryStruct.comp,
Action.Hom.comp_hom, Action.instMonoidalCategory_tensorObj_V,
Action.instMonoidalCategory_whiskerRight_hom, LinearMap.coe_comp, Function.comp_apply]
change _ = (colorFun.map (MonoidalCategory.whiskerRight f Z)).hom
((μ X Z).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
rw [μ_tmul_tprod]
change _ = (colorFun.map (f ▷ Z)).hom
((PiTensorProduct.tprod ) fun i => (colorToRepSumEquiv i)
(HepLean.PiTensorProduct.elimPureTensor p q i))
rw [colorFun.map_tprod]
have h1 : (((colorFun.map f).hom ▷ (colorFun.obj Z).V) ((PiTensorProduct.tprod ) p ⊗ₜ[]
(PiTensorProduct.tprod ) q)) = ((colorFun.map f).hom
((PiTensorProduct.tprod ) p) ⊗ₜ[] ((PiTensorProduct.tprod ) q)) := by rfl
erw [h1]
rw [colorFun.map_tprod]
change (μ Y Z).hom.hom (((PiTensorProduct.tprod ) fun i => (colorToRepCongr _)
(p ((OverColor.Hom.toEquiv f).symm i))) ⊗ₜ[] (PiTensorProduct.tprod ) q) = _
rw [μ_tmul_tprod]
apply congrArg
funext i
match i with
| Sum.inl i => rfl
| Sum.inr i => rfl
lemma μ_natural_right {X Y : OverColor Color} (X' : OverColor Color) (f : X ⟶ Y) :
MonoidalCategory.whiskerLeft (colorFun.obj X') (colorFun.map f) ≫ (μ X' Y).hom =
(μ X' X).hom ≫ colorFun.map (MonoidalCategory.whiskerLeft X' f) := by
ext1
refine HepLean.PiTensorProduct.induction_tmul (fun p q => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Functor.id_obj, CategoryStruct.comp,
Action.Hom.comp_hom, Action.instMonoidalCategory_tensorObj_V,
Action.instMonoidalCategory_whiskerLeft_hom, LinearMap.coe_comp, Function.comp_apply]
change _ = (colorFun.map (X' ◁ f)).hom ((μ X' X).hom.hom
((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
rw [μ_tmul_tprod]
change _ = (colorFun.map (X' ◁ f)).hom ((PiTensorProduct.tprod ) fun i =>
(colorToRepSumEquiv i) (HepLean.PiTensorProduct.elimPureTensor p q i))
rw [map_tprod]
have h1 : (((colorFun.obj X').V ◁ (colorFun.map f).hom)
((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
= ((PiTensorProduct.tprod ) p ⊗ₜ[] (colorFun.map f).hom ((PiTensorProduct.tprod ) q)) := by
rfl
erw [h1]
rw [map_tprod]
change (μ X' Y).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) fun i =>
(colorToRepCongr _) (q ((OverColor.Hom.toEquiv f).symm i))) = _
rw [μ_tmul_tprod]
apply congrArg
funext i
match i with
| Sum.inl i => rfl
| Sum.inr i => rfl
lemma associativity (X Y Z : OverColor Color) :
whiskerRight (μ X Y).hom (colorFun.obj Z) ≫
(μ (X ⊗ Y) Z).hom ≫ colorFun.map (associator X Y Z).hom =
(associator (colorFun.obj X) (colorFun.obj Y) (colorFun.obj Z)).hom ≫
whiskerLeft (colorFun.obj X) (μ Y Z).hom ≫ (μ X (Y ⊗ Z)).hom := by
ext1
refine HepLean.PiTensorProduct.induction_assoc' (fun p q m => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Functor.id_obj, CategoryStruct.comp,
Action.Hom.comp_hom, Action.instMonoidalCategory_tensorObj_V,
Action.instMonoidalCategory_whiskerRight_hom, LinearMap.coe_comp, Function.comp_apply,
Action.instMonoidalCategory_whiskerLeft_hom, Action.instMonoidalCategory_associator_hom_hom]
change (colorFun.map (α_ X Y Z).hom).hom ((μ (X ⊗ Y) Z).hom.hom
((((μ X Y).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[]
(PiTensorProduct.tprod ) q)) ⊗ₜ[] (PiTensorProduct.tprod ) m))) =
(μ X (Y ⊗ Z)).hom.hom ((((PiTensorProduct.tprod ) p ⊗ₜ[] ((μ Y Z).hom.hom
((PiTensorProduct.tprod ) q ⊗ₜ[] (PiTensorProduct.tprod ) m)))))
rw [μ_tmul_tprod, μ_tmul_tprod]
change (colorFun.map (α_ X Y Z).hom).hom ((μ (X ⊗ Y) Z).hom.hom
(((PiTensorProduct.tprod ) fun i => (colorToRepSumEquiv i)
(HepLean.PiTensorProduct.elimPureTensor p q i)) ⊗ₜ[] (PiTensorProduct.tprod ) m)) =
(μ X (Y ⊗ Z)).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) fun i =>
(colorToRepSumEquiv i) (HepLean.PiTensorProduct.elimPureTensor q m i))
rw [μ_tmul_tprod, μ_tmul_tprod]
erw [map_tprod]
apply congrArg
funext i
match i with
| Sum.inl i => rfl
| Sum.inr (Sum.inl i) => rfl
| Sum.inr (Sum.inr i) => rfl
lemma left_unitality (X : OverColor Color) : (leftUnitor (colorFun.obj X)).hom =
whiskerRight colorFun.ε.hom (colorFun.obj X) ≫
(μ (𝟙_ (OverColor Color)) X).hom ≫ colorFun.map (leftUnitor X).hom := by
ext1
apply HepLean.PiTensorProduct.induction_mod_tmul (fun x q => ?_)
simp only [colorFun_obj_V_carrier, Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
Action.instMonoidalCategory_tensorUnit_V, Functor.id_obj,
Action.instMonoidalCategory_leftUnitor_hom_hom, CategoryStruct.comp, Action.Hom.comp_hom,
Action.instMonoidalCategory_tensorObj_V, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorUnit_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom,
Action.instMonoidalCategory_whiskerRight_hom, LinearMap.coe_comp, Function.comp_apply]
change TensorProduct.lid (colorFun.obj X) (x ⊗ₜ[] (PiTensorProduct.tprod ) q) =
(colorFun.map (λ_ X).hom).hom ((μ (𝟙_ (OverColor Color)) X).hom.hom
((((PiTensorProduct.isEmptyEquiv Empty).symm x) ⊗ₜ[] (PiTensorProduct.tprod ) q)))
simp [PiTensorProduct.isEmptyEquiv]
rw [TensorProduct.smul_tmul, TensorProduct.tmul_smul]
erw [LinearMap.map_smul, LinearMap.map_smul]
apply congrArg
change _ = (colorFun.map (λ_ X).hom).hom ((μ (𝟙_ (OverColor Color)) X).hom.hom
((PiTensorProduct.tprod ) _ ⊗ₜ[] (PiTensorProduct.tprod ) q))
rw [μ_tmul_tprod]
erw [map_tprod]
rfl
lemma right_unitality (X : OverColor Color) : (MonoidalCategory.rightUnitor (colorFun.obj X)).hom =
whiskerLeft (colorFun.obj X) ε.hom ≫
(μ X (𝟙_ (OverColor Color))).hom ≫ colorFun.map (MonoidalCategory.rightUnitor X).hom := by
ext1
apply HepLean.PiTensorProduct.induction_tmul_mod (fun p x => ?_)
simp only [colorFun_obj_V_carrier, Functor.id_obj, Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_rightUnitor_hom_hom,
CategoryStruct.comp, Action.Hom.comp_hom, Action.instMonoidalCategory_tensorObj_V,
OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorUnit_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Action.instMonoidalCategory_whiskerLeft_hom,
LinearMap.coe_comp, Function.comp_apply]
change TensorProduct.rid (colorFun.obj X) ((PiTensorProduct.tprod ) p ⊗ₜ[] x) =
(colorFun.map (ρ_ X).hom).hom ((μ X (𝟙_ (OverColor Color))).hom.hom
((((PiTensorProduct.tprod ) p ⊗ₜ[] ((PiTensorProduct.isEmptyEquiv Empty).symm x)))))
simp [PiTensorProduct.isEmptyEquiv]
erw [LinearMap.map_smul, LinearMap.map_smul]
apply congrArg
change _ = (colorFun.map (ρ_ X).hom).hom ((μ X (𝟙_ (OverColor Color))).hom.hom
((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) _))
rw [μ_tmul_tprod]
erw [map_tprod]
rfl
end colorFun
/-- The monoidal functor between `OverColor Color` and `Rep SL(2, )` taking a map of colors
to the corresponding tensor product representation. -/
def colorFunMon : MonoidalFunctor (OverColor Color) (Rep SL(2, )) where
toFunctor := colorFun
ε := colorFun.ε.hom
μ X Y := (colorFun.μ X Y).hom
μ_natural_left := colorFun.μ_natural_left
μ_natural_right := colorFun.μ_natural_right
associativity := colorFun.associativity
left_unitality := colorFun.left_unitality
right_unitality := colorFun.right_unitality
end
end Fermion

View file

@ -0,0 +1,38 @@
/-
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.Tensors.Tree.Basic
import HepLean.Tensors.ComplexLorentz.TensorStruct
import HepLean.Tensors.Tree.Elab
/-!
## The tensor structure for complex Lorentz tensors
-/
open IndexNotation
open CategoryTheory
open MonoidalCategory
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open IndexNotation
open CategoryTheory
noncomputable section
namespace complexLorentzTensor
def upDown : Fin 2 → complexLorentzTensor.C
| 0 => Fermion.Color.up
| 1 => Fermion.Color.down
variable (T S : complexLorentzTensor.F.obj (OverColor.mk upDown))
#check {T | i i}ᵀ.dot
end complexLorentzTensor
end

View file

@ -0,0 +1,36 @@
/-
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.Tensors.Tree.Basic
import HepLean.Tensors.ComplexLorentz.ColorFun
/-!
## The tensor structure for complex Lorentz tensors
-/
open IndexNotation
open CategoryTheory
open MonoidalCategory
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open IndexNotation
open CategoryTheory
noncomputable section
def complexLorentzTensor : TensorStruct where
C := Fermion.Color
G := SL(2, )
G_group := inferInstance
k :=
k_commRing := inferInstance
F := Fermion.colorFunMon
τ := Fermion.τ
evalNo := Fermion.evalNo
end

View file

@ -13,7 +13,7 @@ import HepLean.SpaceTime.WeylFermion.Basic
import HepLean.SpaceTime.LorentzVector.Complex
/-!
## Over category.
## Over color category.
-/
@ -65,6 +65,16 @@ lemma toEquiv_comp_inv_apply (m : f ⟶ g) (i : g.left) :
f.hom ((OverColor.Hom.toEquiv m).symm i) = g.hom i := by
simpa [toEquiv, types_comp] using congrFun m.inv.w i
def toIso (m : f ⟶ g) : f ≅ g := {
hom := m,
inv := m.symm,
hom_inv_id := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
simp only [CategoryStruct.comp, Iso.self_symm_id, Iso.refl_hom, Over.id_left, types_id_apply]
rfl,
inv_hom_id := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
simp only [CategoryStruct.comp, Iso.symm_self_id, Iso.refl_hom, Over.id_left, types_id_apply]
rfl}
end Hom
section monoidal
@ -234,145 +244,6 @@ def mk (f : X → C) : OverColor C := Over.mk f
open MonoidalCategory
/-- The monoidal functor from `OverColor C` to `OverColor D` constructed from a map
`f : C → D`. -/
def map {C D : Type} (f : C → D) : MonoidalFunctor (OverColor C) (OverColor D) where
toFunctor := Core.functorToCore (Core.inclusion (Over C) ⋙ (Over.map f))
ε := Over.isoMk (Iso.refl _) (by
ext x
exact Empty.elim x)
μ X Y := Over.isoMk (Iso.refl _) (by
ext x
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl)
μ_natural_left X Y := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
μ_natural_right X Y := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
associativity X Y Z := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr x => rfl
left_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
right_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
/-- The tensor product on `OverColor C` as a monoidal functor. -/
def tensor (C : Type) : MonoidalFunctor (OverColor C × OverColor C) (OverColor C) where
toFunctor := MonoidalCategory.tensor (OverColor C)
ε := Over.isoMk (Equiv.sumEmpty Empty Empty).symm.toIso (by
ext x
exact Empty.elim x)
μ X Y := Over.isoMk (Equiv.sumSumSumComm X.1.left X.2.left Y.1.left Y.2.left).toIso (by
ext x
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr (Sum.inl x) => rfl
| Sum.inr (Sum.inr x) => rfl)
μ_natural_left X Y := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr (Sum.inl x) => rfl
| Sum.inr (Sum.inr x) => rfl
μ_natural_right X Y := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr (Sum.inl x) => rfl
| Sum.inr (Sum.inr x) => rfl
associativity X Y Z := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl (Sum.inl x)) => rfl
| Sum.inl (Sum.inl (Sum.inr x)) => rfl
| Sum.inl (Sum.inr (Sum.inl x)) => rfl
| Sum.inl (Sum.inr (Sum.inr x)) => rfl
| Sum.inr (Sum.inl x) => rfl
| Sum.inr (Sum.inr x) => rfl
left_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => exact Empty.elim x
| Sum.inr (Sum.inl x)=> rfl
| Sum.inr (Sum.inr x)=> rfl
right_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr x => exact Empty.elim x
def diag (C : Type) : MonoidalFunctor (OverColor C) (OverColor C × OverColor C) :=
MonoidalFunctor.diag (OverColor C)
def const (C : Type) : MonoidalFunctor (OverColor C) (OverColor C) where
toFunctor := (Functor.const (OverColor C)).obj (𝟙_ (OverColor C))
ε := 𝟙 (𝟙_ (OverColor C))
μ _ _:= (λ_ (𝟙_ (OverColor C))).hom
μ_natural_left _ _ := by
simp only [Functor.const_obj_obj, Functor.const_obj_map, MonoidalCategory.whiskerRight_id,
Category.id_comp, Iso.hom_inv_id, Category.comp_id]
μ_natural_right _ _ := by
simp only [Functor.const_obj_obj, Functor.const_obj_map, MonoidalCategory.whiskerLeft_id,
Category.id_comp, Category.comp_id]
associativity X Y Z := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun i => by
match i with
| Sum.inl (Sum.inl i) => rfl
| Sum.inl (Sum.inr i) => rfl
| Sum.inr i => rfl
left_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun i => by
match i with
| Sum.inl i => exact Empty.elim i
| Sum.inr i => exact Empty.elim i
right_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun i => by
match i with
| Sum.inl i => exact Empty.elim i
| Sum.inr i => exact Empty.elim i
/-!
## Useful equivalences.
-/
/-- The isomorphism between `c : X → C` and `c ∘ e.symm` as objects in `OverColor C` for an
equivalence `e`. -/
def equivToIso {c : X → C} (e : X ≃ Y) : mk c ≅ mk (c ∘ e.symm) := {
hom := Over.isoMk e.toIso ((Iso.eq_inv_comp e.toIso).mp rfl),
inv := (Over.isoMk e.toIso ((Iso.eq_inv_comp e.toIso).mp rfl)).symm,
hom_inv_id := by
apply CategoryTheory.Iso.ext
erw [CategoryTheory.Iso.trans_hom]
simp only [Functor.id_obj, Over.mk_left, Over.mk_hom, Iso.symm_hom, Iso.hom_inv_id]
rfl,
inv_hom_id := by
apply CategoryTheory.Iso.ext
erw [CategoryTheory.Iso.trans_hom]
simp only [Iso.symm_hom, Iso.inv_hom_id]
rfl}
/-- The equivalence between `Fin n.succ` and `Fin 1 ⊕ Fin n` extracting the
`i`th component. -/
def finExtractOne {n : } (i : Fin n.succ) : Fin n.succ ≃ Fin 1 ⊕ Fin n :=
(finCongr (by omega : n.succ = i + 1 + (n - i))).trans <|
finSumFinEquiv.symm.trans <|
(Equiv.sumCongr (finSumFinEquiv.symm.trans (Equiv.sumComm (Fin i) (Fin 1)))
(Equiv.refl (Fin (n-i)))).trans <|
(Equiv.sumAssoc (Fin 1) (Fin i) (Fin (n - i))).trans <|
Equiv.sumCongr (Equiv.refl (Fin 1)) (finSumFinEquiv.trans (finCongr (by omega)))
end OverColor
end IndexNotation

View file

@ -0,0 +1,128 @@
/-
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.Tensors.OverColor.Basic
/-!
## Functors related to the OverColor category
-/
namespace IndexNotation
namespace OverColor
open CategoryTheory
open MonoidalCategory
/-- The monoidal functor from `OverColor C` to `OverColor D` constructed from a map
`f : C → D`. -/
def map {C D : Type} (f : C → D) : MonoidalFunctor (OverColor C) (OverColor D) where
toFunctor := Core.functorToCore (Core.inclusion (Over C) ⋙ (Over.map f))
ε := Over.isoMk (Iso.refl _) (by
ext x
exact Empty.elim x)
μ X Y := Over.isoMk (Iso.refl _) (by
ext x
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl)
μ_natural_left X Y := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
μ_natural_right X Y := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
associativity X Y Z := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr x => rfl
left_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
right_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
/-- The tensor product on `OverColor C` as a monoidal functor. -/
def tensor (C : Type) : MonoidalFunctor (OverColor C × OverColor C) (OverColor C) where
toFunctor := MonoidalCategory.tensor (OverColor C)
ε := Over.isoMk (Equiv.sumEmpty Empty Empty).symm.toIso (by
ext x
exact Empty.elim x)
μ X Y := Over.isoMk (Equiv.sumSumSumComm X.1.left X.2.left Y.1.left Y.2.left).toIso (by
ext x
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr (Sum.inl x) => rfl
| Sum.inr (Sum.inr x) => rfl)
μ_natural_left X Y := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr (Sum.inl x) => rfl
| Sum.inr (Sum.inr x) => rfl
μ_natural_right X Y := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr (Sum.inl x) => rfl
| Sum.inr (Sum.inr x) => rfl
associativity X Y Z := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl (Sum.inl x)) => rfl
| Sum.inl (Sum.inl (Sum.inr x)) => rfl
| Sum.inl (Sum.inr (Sum.inl x)) => rfl
| Sum.inl (Sum.inr (Sum.inr x)) => rfl
| Sum.inr (Sum.inl x) => rfl
| Sum.inr (Sum.inr x) => rfl
left_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => exact Empty.elim x
| Sum.inr (Sum.inl x)=> rfl
| Sum.inr (Sum.inr x)=> rfl
right_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr x => exact Empty.elim x
def diag (C : Type) : MonoidalFunctor (OverColor C) (OverColor C × OverColor C) :=
MonoidalFunctor.diag (OverColor C)
def const (C : Type) : MonoidalFunctor (OverColor C) (OverColor C) where
toFunctor := (Functor.const (OverColor C)).obj (𝟙_ (OverColor C))
ε := 𝟙 (𝟙_ (OverColor C))
μ _ _:= (λ_ (𝟙_ (OverColor C))).hom
μ_natural_left _ _ := by
simp only [Functor.const_obj_obj, Functor.const_obj_map, MonoidalCategory.whiskerRight_id,
Category.id_comp, Iso.hom_inv_id, Category.comp_id]
μ_natural_right _ _ := by
simp only [Functor.const_obj_obj, Functor.const_obj_map, MonoidalCategory.whiskerLeft_id,
Category.id_comp, Category.comp_id]
associativity X Y Z := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun i => by
match i with
| Sum.inl (Sum.inl i) => rfl
| Sum.inl (Sum.inr i) => rfl
| Sum.inr i => rfl
left_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun i => by
match i with
| Sum.inl i => exact Empty.elim i
| Sum.inr i => exact Empty.elim i
right_unitality X := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun i => by
match i with
| Sum.inl i => exact Empty.elim i
| Sum.inr i => exact Empty.elim i
def contrPair (C : Type) (τ : C → C) : MonoidalFunctor (OverColor C) (OverColor C) :=
OverColor.diag C
⊗⋙ (MonoidalFunctor.prod (MonoidalFunctor.id (OverColor C)) (OverColor.map τ))
⊗⋙ OverColor.tensor C
end OverColor
end IndexNotation

View file

@ -0,0 +1,163 @@
/-
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.Tensors.OverColor.Functors
/-!
## Isomorphisms in the OverColor category
-/
namespace IndexNotation
namespace OverColor
open CategoryTheory
open MonoidalCategory
/-!
## Useful equivalences.
-/
/-- The isomorphism between `c : X → C` and `c ∘ e.symm` as objects in `OverColor C` for an
equivalence `e`. -/
def equivToIso {c : X → C} (e : X ≃ Y) : mk c ≅ mk (c ∘ e.symm) :=
Hom.toIso (Over.isoMk e.toIso ((Iso.eq_inv_comp e.toIso).mp rfl))
def mkSum (c : X ⊕ Y → C) : mk c ≅ mk (c ∘ Sum.inl) ⊗ mk (c ∘ Sum.inr) :=
Hom.toIso (Over.isoMk (Equiv.refl _).toIso (by
ext x
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl))
def mkIso {c1 c2 : X → C} (h : c1 = c2) : mk c1 ≅ mk c2 :=
Hom.toIso (Over.isoMk (Equiv.refl _).toIso (by
subst h
rfl))
/-- The equivalence between `Fin n.succ` and `Fin 1 ⊕ Fin n` extracting the
`i`th component. -/
def finExtractOne {n : } (i : Fin n.succ) : Fin n.succ ≃ Fin 1 ⊕ Fin n :=
(finCongr (by omega : n.succ = i + 1 + (n - i))).trans <|
finSumFinEquiv.symm.trans <|
(Equiv.sumCongr (finSumFinEquiv.symm.trans (Equiv.sumComm (Fin i) (Fin 1)))
(Equiv.refl (Fin (n-i)))).trans <|
(Equiv.sumAssoc (Fin 1) (Fin i) (Fin (n - i))).trans <|
Equiv.sumCongr (Equiv.refl (Fin 1)) (finSumFinEquiv.trans (finCongr (by omega)))
lemma finExtractOne_symm_inr {n : } (i : Fin n.succ) :
(finExtractOne i).symm ∘ Sum.inr = i.succAbove := by
ext x
simp only [Nat.succ_eq_add_one, finExtractOne, Function.comp_apply, Equiv.symm_trans_apply,
finCongr_symm, Equiv.symm_symm, Equiv.sumCongr_symm, Equiv.refl_symm, Equiv.sumCongr_apply,
Equiv.coe_refl, Sum.map_inr, finCongr_apply, Fin.coe_cast]
change (finSumFinEquiv (Sum.map (⇑(finSumFinEquiv.symm.trans (Equiv.sumComm (Fin ↑i) (Fin 1))).symm) id
((Equiv.sumAssoc (Fin 1) (Fin ↑i) (Fin (n - i))).symm (Sum.inr (finSumFinEquiv.symm (Fin.cast (finExtractOne.proof_2 i).symm x)))))).val = _
by_cases hi : x.1 < i.1
· have h1 : (finSumFinEquiv.symm (Fin.cast (finExtractOne.proof_2 i).symm x)) =
Sum.inl ⟨x, hi⟩ := by
rw [← finSumFinEquiv_symm_apply_castAdd]
apply congrArg
ext
simp
rw [h1]
simp only [Nat.succ_eq_add_one, Equiv.sumAssoc_symm_apply_inr_inl, Sum.map_inl,
Equiv.symm_trans_apply, Equiv.symm_symm, Equiv.sumComm_symm, Equiv.sumComm_apply,
Sum.swap_inr, finSumFinEquiv_apply_left, Fin.castAdd_mk]
rw [Fin.succAbove]
split
· rfl
rename_i hn
simp_all only [Nat.succ_eq_add_one, not_lt, Fin.le_def, Fin.coe_castSucc, Fin.val_succ,
self_eq_add_right, one_ne_zero]
omega
· have h1 : (finSumFinEquiv.symm (Fin.cast (finExtractOne.proof_2 i).symm x)) =
Sum.inr ⟨x - i, by omega⟩ := by
rw [← finSumFinEquiv_symm_apply_natAdd]
apply congrArg
ext
simp
omega
rw [h1, Fin.succAbove]
split
· rename_i hn
simp_all [Fin.lt_def]
simp only [Nat.succ_eq_add_one, Equiv.sumAssoc_symm_apply_inr_inr, Sum.map_inr, id_eq,
finSumFinEquiv_apply_right, Fin.natAdd_mk, Fin.val_succ]
omega
@[simp]
lemma finExtractOne_symm_inr_apply {n : } (i : Fin n.succ) (x : Fin n) :
(finExtractOne i).symm (Sum.inr x) = i.succAbove x := calc
_ = ((finExtractOne i).symm ∘ Sum.inr) x := rfl
_ = i.succAbove x := by rw [finExtractOne_symm_inr]
@[simp]
lemma finExtractOne_symm_inl_apply {n : } (i : Fin n.succ) :
(finExtractOne i).symm (Sum.inl 0) = i := by
simp only [Nat.succ_eq_add_one, finExtractOne, Fin.isValue, Equiv.symm_trans_apply, finCongr_symm,
Equiv.symm_symm, Equiv.sumCongr_symm, Equiv.refl_symm, Equiv.sumCongr_apply, Equiv.coe_refl,
Sum.map_inl, id_eq, Equiv.sumAssoc_symm_apply_inl, Equiv.sumComm_symm, Equiv.sumComm_apply,
Sum.swap_inl, finSumFinEquiv_apply_right, finSumFinEquiv_apply_left, finCongr_apply]
ext
rfl
def finExtractTwo {n : } (i : Fin n.succ.succ) (j : Fin n.succ) :
Fin n.succ.succ ≃ (Fin 1 ⊕ Fin 1) ⊕ Fin n :=
(finExtractOne i).trans <|
(Equiv.sumCongr (Equiv.refl (Fin 1)) (finExtractOne j)).trans <|
(Equiv.sumAssoc (Fin 1) (Fin 1) (Fin n)).symm
lemma finExtractTwo_symm_inr {n : } (i : Fin n.succ.succ) (j : Fin n.succ) :
(finExtractTwo i j).symm ∘ Sum.inr = i.succAbove ∘ j.succAbove := by
rw [finExtractTwo]
ext1 x
simp
@[simp]
lemma finExtractTwo_symm_inr_apply {n : } (i : Fin n.succ.succ) (j : Fin n.succ) (x : Fin n) :
(finExtractTwo i j).symm (Sum.inr x) = i.succAbove (j.succAbove x) := by
rw [finExtractTwo]
simp
@[simp]
lemma finExtractTwo_symm_inl_inr_apply {n : } (i : Fin n.succ.succ) (j : Fin n.succ) :
(finExtractTwo i j).symm (Sum.inl (Sum.inr 0)) = i.succAbove j := by
rw [finExtractTwo]
simp
@[simp]
lemma finExtractTwo_symm_inl_inl_apply {n : } (i : Fin n.succ.succ) (j : Fin n.succ) :
(finExtractTwo i j).symm (Sum.inl (Sum.inl 0)) = i := by
rw [finExtractTwo]
simp
def contrPairFin1Fin1 (τ : C → C) (c : Fin 1 ⊕ Fin 1 → C)
(h : c (Sum.inr 0) = τ (c (Sum.inl 0))) :
OverColor.mk c ≅ (contrPair C τ).obj (OverColor.mk (fun (_ : Fin 1) => c (Sum.inl 0))) :=
Hom.toIso (Over.isoMk (Equiv.refl _).toIso (by
ext x
match x with
| Sum.inl x =>
fin_cases x
rfl
| Sum.inr x =>
fin_cases x
change _ = c (Sum.inr 0)
rw [h]
rfl))
def contrPairEquiv {n : } (τ : C → C) (c : Fin n.succ.succ → C) (i : Fin n.succ.succ)
(j : Fin n.succ) (h : c (i.succAbove j) = τ (c i)) :
OverColor.mk c ≅ ((contrPair C τ).obj (Over.mk (fun (_ : Fin 1) => c i))) ⊗
(OverColor.mk (c ∘ i.succAbove ∘ j.succAbove)) :=
(equivToIso (finExtractTwo i j)).trans <|
(OverColor.mkSum (c ∘ ⇑(finExtractTwo i j).symm)).trans <|
tensorIso
(contrPairFin1Fin1 τ ((c ∘ ⇑(finExtractTwo i j).symm) ∘ Sum.inl) (by simpa using h)) <|
mkIso (by ext x; simp)
end OverColor
end IndexNotation

View file

@ -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.Tensors.ColorCat.Basic
import HepLean.Tensors.OverColor.Iso
import Mathlib.CategoryTheory.Monoidal.NaturalTransformation
/-!
@ -34,12 +34,9 @@ structure TensorStruct where
F : MonoidalFunctor (OverColor C) (Rep k G)
/-- A map from `C` to `C`. An involution. -/
τ : C → C
/-- Contraction natural transformation. -/
ηContr : (OverColor.diag C) ⊗⋙ (MonoidalFunctor.prod (MonoidalFunctor.id (OverColor C)) (OverColor.map τ))
⊗⋙ OverColor.tensor C ⊗⋙ F ⟶ @constMonoFunc (OverColor C) G k
/-- A monoidal natural isomorphism from OverColor.map τ ⊗⋙ F to F.
This will allow us to, for example, rise and lower indices. -/
dual : (OverColor.map τ ⊗⋙ F) ≅ F
/-
μContr : OverColor.contrPair C τ ⊗⋙ F ⟶ OverColor.const C ⊗⋙ F
dual : (OverColor.map τ ⊗⋙ F) ≅ F-/
/-- A specification of the dimension of each color in C. This will be used for explicit
evaluation of tensors. -/
evalNo : C →
@ -67,7 +64,7 @@ inductive TensorTree (S : TensorStruct) : ∀ {n : }, (Fin n → S.C) → Typ
(i : Fin n.succ) → (j : Fin m.succ) → TensorTree S c → TensorTree S c1 →
TensorTree S (Sum.elim (c ∘ Fin.succAbove i) (c1 ∘ Fin.succAbove j) ∘ finSumFinEquiv.symm)
| contr {n : } {c : Fin n.succ.succ → S.C} : (i : Fin n.succ.succ) →
(j : Fin n.succ) → TensorTree S c → TensorTree S (c ∘ Fin.succAbove i ∘ Fin.succAbove j)
(j : Fin n.succ) → (h : c (i.succAbove j) = S.τ (c i)) → TensorTree S c → TensorTree S (c ∘ Fin.succAbove i ∘ Fin.succAbove j)
| jiggle {n : } {c : Fin n → S.C} : (i : Fin n) → TensorTree S c →
TensorTree S (Function.update c i (S.τ (c i)))
| eval {n : } {c : Fin n.succ → S.C} :
@ -89,7 +86,7 @@ def size : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → := fun
| smul _ t => t.size + 1
| prod t1 t2 => t1.size + t2.size + 1
| mult _ _ t1 t2 => t1.size + t2.size + 1
| contr _ _ t => t.size + 1
| contr _ _ _ t => t.size + 1
| jiggle _ t => t.size + 1
| eval _ _ t => t.size + 1

View file

@ -53,7 +53,7 @@ def dotString (m : ) (nt : ) : ∀ {n : } {c : Fin n → S.C}, TensorTr
let jiggleNode := " node" ++ toString m ++ " [label=\"τ\", shape=box];\n"
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
jiggleNode ++ dotString (m + 1) nt t1 ++ edge1
| contr i j t1 =>
| contr i j _ t1 =>
let contrNode := " node" ++ toString m ++ " [label=\"contr " ++ toString i ++ " "
++ toString j ++ "\", shape=box];\n"
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"

View file

@ -192,7 +192,7 @@ def withoutContr (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
/-- For each element of `l : List ( × )` applies `TensorTree.contr` to the given term. -/
def contrSyntax (l : List ( × )) (T : Term) : Term :=
l.foldl (fun T' (x0, x1) => Syntax.mkApp (mkIdent ``TensorTree.contr)
#[Syntax.mkNumLit (toString x1), Syntax.mkNumLit (toString x0), T']) T
#[Syntax.mkNumLit (toString x1), Syntax.mkNumLit (toString x0), mkIdent ``rfl, T']) T
/-- Creates the syntax associated with a tensor node. -/
def syntaxFull (stx : Syntax) : TermElabM Term := do
@ -256,7 +256,7 @@ def withoutContr (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
/-- For each element of `l : List ( × )` applies `TensorTree.contr` to the given term. -/
def contrSyntax (l : List ( × )) (T : Term) : Term :=
l.foldl (fun T' (x0, x1) => Syntax.mkApp (mkIdent ``TensorTree.contr)
#[Syntax.mkNumLit (toString x1), Syntax.mkNumLit (toString x0), T']) T
#[Syntax.mkNumLit (toString x1), Syntax.mkNumLit (toString x0), mkIdent ``rfl, T']) T
/-- The syntax associated with a product of tensors. -/
def prodSyntax (T1 T2 : Term) : Term :=