refactor: Simp lemmas

This commit is contained in:
jstoobysmith 2024-10-12 07:57:35 +00:00
parent cd1b30c069
commit 1651b265e7
16 changed files with 116 additions and 86 deletions

View file

@ -152,7 +152,7 @@ lemma map_tprod {X Y : OverColor Color} (p : (i : X.left) → (colorToRep (X.hom
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']
simp only [map', mapToLinearEquiv', Functor.id_obj]
erw [LinearEquiv.trans_apply]
change (PiTensorProduct.congr fun i => colorToRepCongr _)
((PiTensorProduct.reindex (fun x => _) (OverColor.Hom.toEquiv f))
@ -276,7 +276,9 @@ def μ (X Y : OverColor Color) : colorFun.obj X ⊗ colorFun.obj Y ≅ colorFun.
inv := {
hom := (μModEquiv X Y).symm.toLinearMap
comm := fun M => by
simp [CategoryStruct.comp]
simp only [Action.instMonoidalCategory_tensorObj_V, CategoryStruct.comp,
colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Action.tensor_ρ']
erw [LinearEquiv.eq_comp_toLinearMap_symm,LinearMap.comp_assoc,
LinearEquiv.toLinearMap_symm_comp_eq]
refine HepLean.PiTensorProduct.induction_tmul (fun p q => ?_)
@ -428,7 +430,14 @@ lemma left_unitality (X : OverColor Color) : (leftUnitor (colorFun.obj X)).hom =
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]
simp only [Functor.id_obj, lid_tmul, colorFun_obj_V_carrier,
OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorUnit_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom, PiTensorProduct.isEmptyEquiv,
LinearEquiv.coe_symm_mk]
rw [TensorProduct.smul_tmul, TensorProduct.tmul_smul]
erw [LinearMap.map_smul, LinearMap.map_smul]
apply congrArg
@ -454,7 +463,14 @@ lemma right_unitality (X : OverColor Color) : (MonoidalCategory.rightUnitor (col
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]
simp only [Functor.id_obj, rid_tmul, colorFun_obj_V_carrier,
OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorUnit_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom, PiTensorProduct.isEmptyEquiv,
LinearEquiv.coe_symm_mk, tmul_smul]
erw [LinearMap.map_smul, LinearMap.map_smul]
apply congrArg
change _ = (colorFun.map (ρ_ X).hom).hom ((μ X (𝟙_ (OverColor Color))).hom.hom

View file

@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Tensors.OverColor.Basic
import HepLean.Tensors.OverColor.Functors
import HepLean.Tensors.ComplexLorentz.ColorFun
import HepLean.Mathematics.PiTensorProduct
/-!
@ -59,7 +58,7 @@ lemma mapToLinearEquiv'_tprod {f g : OverColor Color} (m : f ⟶ g)
PiTensorProduct.tprod fun i =>
(TensorProduct.congr (colorToRepCongr (OverColor.Hom.toEquiv_symm_apply m i))
(colorToRepCongr (mapToLinearEquiv'.proof_4 m i))) (x ((OverColor.Hom.toEquiv m).symm i)) := by
simp [mapToLinearEquiv']
simp only [mapToLinearEquiv', Functor.id_obj, LinearEquiv.trans_apply]
change (PiTensorProduct.congr fun i => TensorProduct.congr (colorToRepCongr _)
(colorToRepCongr _)) ((PiTensorProduct.reindex
(fun x => ↑(colorToRep (f.hom x)).V ⊗[] ↑(colorToRep (τ (f.hom x))).V)

View file

@ -5,7 +5,6 @@ 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
@ -33,6 +32,8 @@ def upDown : Fin 2 → complexLorentzTensor.C
variable (T S : complexLorentzTensor.F.obj (OverColor.mk upDown))
/-
import HepLean.Tensors.Tree.Elab
#check {T | i m ⊗ S | m l}ᵀ.dot
#check {T | i m ⊗ S | l τ(m)}ᵀ.dot
-/