PhysLean/HepLean/Tensors/Tree/Basic.lean

526 lines
26 KiB
Text
Raw Normal View History

/-
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.Iso
2024-10-15 06:08:56 +00:00
import HepLean.Tensors.OverColor.Discrete
import HepLean.Tensors.OverColor.Lift
import Mathlib.CategoryTheory.Monoidal.NaturalTransformation
/-!
## Tensor trees
-/
open IndexNotation
open CategoryTheory
2024-10-09 16:57:41 +00:00
open MonoidalCategory
2024-10-08 07:52:55 +00:00
/-- The sturcture of a type of tensors e.g. Lorentz tensors, Einstien tensors,
complex Lorentz tensors.
Note: This structure is not fully defined yet. -/
structure TensorStruct where
2024-10-08 07:52:55 +00:00
/-- The colors of indices e.g. up or down. -/
C : Type
2024-10-08 07:52:55 +00:00
/-- The symmetry group acting on these tensor e.g. the Lorentz group or SL(2,). -/
G : Type
2024-10-08 07:52:55 +00:00
/-- An instance of `G` as a group. -/
G_group : Group G
2024-10-08 07:52:55 +00:00
/-- The field over which we want to consider the tensors to live in, usually `` or ``. -/
k : Type
2024-10-08 07:52:55 +00:00
/-- An instance of `k` as a commutative ring. -/
k_commRing : CommRing k
2024-10-08 07:52:55 +00:00
/-- A `MonoidalFunctor` from `OverColor C` giving the rep corresponding to a map of colors
`X → C`. -/
2024-10-15 06:08:56 +00:00
FDiscrete : Discrete C ⥤ Rep k G
2024-10-08 07:52:55 +00:00
/-- A map from `C` to `C`. An involution. -/
τ : C → C
2024-10-15 06:08:56 +00:00
τ_involution : Function.Involutive τ
/-- The natural transformation describing contraction. -/
2024-10-16 16:38:36 +00:00
contr : OverColor.Discrete.pairτ FDiscrete τ ⟶ 𝟙_ (Discrete C ⥤ Rep k G)
2024-10-15 06:08:56 +00:00
/-- The natural transformation describing the metric. -/
2024-10-16 16:38:36 +00:00
metric : 𝟙_ (Discrete C ⥤ Rep k G) ⟶ OverColor.Discrete.pair FDiscrete
2024-10-15 06:08:56 +00:00
/-- The natural transformation describing the unit. -/
unit : 𝟙_ (Discrete C ⥤ Rep k G) ⟶ OverColor.Discrete.τPair FDiscrete τ
2024-10-08 07:52:55 +00:00
/-- A specification of the dimension of each color in C. This will be used for explicit
evaluation of tensors. -/
evalNo : C →
2024-10-15 06:08:56 +00:00
noncomputable section
namespace TensorStruct
2024-10-18 16:08:17 +00:00
open OverColor
variable (S : TensorStruct)
instance : CommRing S.k := S.k_commRing
instance : Group S.G := S.G_group
2024-10-15 06:08:56 +00:00
/-- The lift of the functor `S.F` to a monoidal functor. -/
def F : MonoidalFunctor (OverColor S.C) (Rep S.k S.G) := (OverColor.lift).obj S.FDiscrete
2024-10-18 16:08:17 +00:00
lemma F_def : F S = (OverColor.lift).obj S.FDiscrete := rfl
lemma perm_contr_cond {n : } {c : Fin n.succ.succ → S.C} {c1 : Fin n.succ.succ → S.C}
{i : Fin n.succ.succ} {j : Fin n.succ}
(h : c1 (i.succAbove j) = S.τ (c1 i)) (σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) :
c (Fin.succAbove ((Hom.toEquiv σ).symm i) ((Hom.toEquiv (extractOne i σ)).symm j)) =
S.τ (c ((Hom.toEquiv σ).symm i)) := by
have h1 := Hom.toEquiv_comp_apply σ
2024-10-19 09:19:29 +00:00
simp only [Nat.succ_eq_add_one, Functor.const_obj_obj, mk_hom] at h1
2024-10-18 16:08:17 +00:00
rw [h1, h1]
2024-10-19 09:19:29 +00:00
simp only [Nat.succ_eq_add_one, extractOne_homToEquiv, Equiv.apply_symm_apply]
2024-10-18 16:08:17 +00:00
rw [← h]
congr
2024-10-19 09:19:29 +00:00
simp only [Nat.succ_eq_add_one, HepLean.Fin.finExtractOnePerm, HepLean.Fin.finExtractOnPermHom,
HepLean.Fin.finExtractOne_symm_inr_apply, Equiv.symm_apply_apply, Equiv.coe_fn_symm_mk]
2024-10-18 16:08:17 +00:00
erw [Equiv.apply_symm_apply]
rw [HepLean.Fin.succsAbove_predAboveI]
erw [Equiv.apply_symm_apply]
2024-10-19 09:19:29 +00:00
simp only [Nat.succ_eq_add_one, ne_eq]
2024-10-18 16:08:17 +00:00
erw [Equiv.apply_eq_iff_eq]
exact (Fin.succAbove_ne i j).symm
2024-10-15 06:08:56 +00:00
/-- The isomorphism between the image of a map `Fin 1 ⊕ Fin 1 → S.C` contructed by `finExtractTwo`
2024-10-19 09:19:29 +00:00
under `S.F.obj`, and an object in the image of `OverColor.Discrete.pairτ S.FDiscrete`. -/
def contrFin1Fin1 {n : } (c : Fin n.succ.succ → S.C)
2024-10-15 06:08:56 +00:00
(i : Fin n.succ.succ) (j : Fin n.succ) (h : c (i.succAbove j) = S.τ (c i)) :
S.F.obj (OverColor.mk ((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inl)) ≅
(OverColor.Discrete.pairτ S.FDiscrete S.τ).obj { as := c i } := by
2024-10-19 09:47:23 +00:00
apply (S.F.mapIso
(OverColor.mkSum (((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inl)))).trans
2024-10-15 06:08:56 +00:00
apply (S.F.μIso _ _).symm.trans
apply tensorIso ?_ ?_
· symm
apply (OverColor.forgetLiftApp S.FDiscrete (c i)).symm.trans
apply S.F.mapIso
apply OverColor.mkIso
funext x
fin_cases x
rfl
· symm
apply (OverColor.forgetLiftApp S.FDiscrete (S.τ (c i))).symm.trans
apply S.F.mapIso
apply OverColor.mkIso
funext x
fin_cases x
simp [h]
lemma contrFin1Fin1_inv_tmul {n : } (c : Fin n.succ.succ → S.C)
(i : Fin n.succ.succ) (j : Fin n.succ) (h : c (i.succAbove j) = S.τ (c i))
(x : S.FDiscrete.obj { as := c i })
(y : S.FDiscrete.obj { as := S.τ (c i) }) :
(S.contrFin1Fin1 c i j h).inv.hom (x ⊗ₜ[S.k] y) =
PiTensorProduct.tprod S.k (fun k =>
match k with | Sum.inl 0 => x | Sum.inr 0 => (S.FDiscrete.map
(eqToHom (by simp [h]))).hom y) := by
2024-10-19 09:19:29 +00:00
simp only [Nat.succ_eq_add_one, contrFin1Fin1, Functor.comp_obj, Discrete.functor_obj_eq_as,
Function.comp_apply, Iso.trans_symm, Iso.symm_symm_eq, Iso.trans_inv, tensorIso_inv,
Iso.symm_inv, Functor.mapIso_hom, tensor_comp, MonoidalFunctor.μIso_hom, Category.assoc,
LaxMonoidalFunctor.μ_natural, Functor.mapIso_inv, Action.comp_hom,
Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorHom_hom,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, ModuleCat.coe_comp, Functor.id_obj, mk_hom,
Fin.isValue]
change (S.F.map (OverColor.mkSum ((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inl)).inv).hom
((S.F.map ((OverColor.mkIso _).hom ⊗ (OverColor.mkIso _).hom)).hom
2024-10-19 09:47:23 +00:00
((S.F.μ (OverColor.mk fun _ => c i) (OverColor.mk fun _ => S.τ (c i))).hom
((((OverColor.forgetLiftApp S.FDiscrete (c i)).inv.hom x) ⊗ₜ[S.k]
((OverColor.forgetLiftApp S.FDiscrete (S.τ (c i))).inv.hom y))))) = _
2024-10-19 09:19:29 +00:00
simp only [Nat.succ_eq_add_one, Action.instMonoidalCategory_tensorObj_V, Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
forgetLiftApp, Action.mkIso_inv_hom, LinearEquiv.toModuleIso_inv, Fin.isValue]
2024-10-19 09:47:23 +00:00
erw [OverColor.forgetLiftAppV_symm_apply,
OverColor.forgetLiftAppV_symm_apply S.FDiscrete (S.τ (c i))]
change ((OverColor.lift.obj S.FDiscrete).map (OverColor.mkSum
((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inl)).inv).hom
(((OverColor.lift.obj S.FDiscrete).map ((OverColor.mkIso _).hom ⊗ (OverColor.mkIso _).hom)).hom
2024-10-19 09:47:23 +00:00
(((OverColor.lift.obj S.FDiscrete).μ (OverColor.mk fun _ => c i)
(OverColor.mk fun _ => S.τ (c i))).hom
(((PiTensorProduct.tprod S.k) fun _ => x) ⊗ₜ[S.k] (PiTensorProduct.tprod S.k) fun _ => y))) = _
rw [OverColor.lift.obj_μ_tprod_tmul S.FDiscrete]
2024-10-19 09:47:23 +00:00
change ((OverColor.lift.obj S.FDiscrete).map
(OverColor.mkSum ((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inl)).inv).hom
(((OverColor.lift.obj S.FDiscrete).map ((OverColor.mkIso _).hom ⊗ (OverColor.mkIso _).hom)).hom
2024-10-19 09:47:23 +00:00
((PiTensorProduct.tprod S.k) _)) = _
rw [OverColor.lift.map_tprod S.FDiscrete]
2024-10-19 09:47:23 +00:00
change ((OverColor.lift.obj S.FDiscrete).map
(OverColor.mkSum ((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inl)).inv).hom
((PiTensorProduct.tprod S.k _)) = _
rw [OverColor.lift.map_tprod S.FDiscrete]
apply congrArg
funext r
match r with
| Sum.inl 0 =>
2024-10-19 09:19:29 +00:00
simp only [Nat.succ_eq_add_one, mk_hom, Fin.isValue, Function.comp_apply,
instMonoidalCategoryStruct_tensorObj_left, mkSum_inv_homToEquiv, Equiv.refl_symm,
instMonoidalCategoryStruct_tensorObj_hom, Functor.id_obj, lift.discreteSumEquiv, Sum.elim_inl,
Sum.elim_inr, HepLean.PiTensorProduct.elimPureTensor]
simp only [Fin.isValue, lift.discreteFunctorMapEqIso, eqToIso_refl, Functor.mapIso_refl,
Iso.refl_hom, Action.id_hom, Iso.refl_inv, LinearEquiv.ofLinear_apply]
rfl
| Sum.inr 0 =>
2024-10-19 09:19:29 +00:00
simp only [Nat.succ_eq_add_one, mk_hom, Fin.isValue, Function.comp_apply,
instMonoidalCategoryStruct_tensorObj_left, mkSum_inv_homToEquiv, Equiv.refl_symm,
instMonoidalCategoryStruct_tensorObj_hom, lift.discreteFunctorMapEqIso, eqToIso_refl,
Functor.mapIso_refl, Iso.refl_hom, Action.id_hom, Iso.refl_inv, Functor.mapIso_hom,
eqToIso.hom, Functor.mapIso_inv, eqToIso.inv, Functor.id_obj, lift.discreteSumEquiv,
Sum.elim_inl, Sum.elim_inr, HepLean.PiTensorProduct.elimPureTensor,
LinearEquiv.ofLinear_apply]
rfl
2024-10-18 16:08:17 +00:00
lemma contrFin1Fin1_hom_hom_tprod {n : } (c : Fin n.succ.succ → S.C)
(i : Fin n.succ.succ) (j : Fin n.succ) (h : c (i.succAbove j) = S.τ (c i))
2024-10-19 09:47:23 +00:00
(x : (k : Fin 1 ⊕ Fin 1) → (S.FDiscrete.obj
{ as := (OverColor.mk ((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inl)).hom k })) :
2024-10-18 16:08:17 +00:00
(S.contrFin1Fin1 c i j h).hom.hom (PiTensorProduct.tprod S.k x) =
x (Sum.inl 0) ⊗ₜ[S.k] ((S.FDiscrete.map (eqToHom (by simp [h]))).hom (x (Sum.inr 0))) := by
change ((Action.forget _ _).mapIso (S.contrFin1Fin1 c i j h)).hom _ = _
2024-10-19 09:47:23 +00:00
trans ((Action.forget _ _).mapIso (S.contrFin1Fin1 c i j h)).toLinearEquiv
(PiTensorProduct.tprod S.k x)
2024-10-18 16:08:17 +00:00
· congr
erw [← LinearEquiv.eq_symm_apply]
erw [contrFin1Fin1_inv_tmul]
congr
funext i
match i with
| Sum.inl 0 =>
simp
| Sum.inr 0 =>
2024-10-19 09:19:29 +00:00
simp only [Nat.succ_eq_add_one, Fin.isValue, mk_hom, Function.comp_apply,
Discrete.functor_obj_eq_as]
2024-10-18 16:08:17 +00:00
change _ = ((S.FDiscrete.map (eqToHom _)) ≫ (S.FDiscrete.map (eqToHom _))).hom (x (Sum.inr 0))
rw [← Functor.map_comp]
simp
exact h
/-- The isomorphism of objects in `Rep S.k S.G` given an `i` in `Fin n.succ.succ` and
a `j` in `Fin n.succ` allowing us to undertake contraction. -/
def contrIso {n : } (c : Fin n.succ.succ → S.C)
(i : Fin n.succ.succ) (j : Fin n.succ) (h : c (i.succAbove j) = S.τ (c i)) :
S.F.obj (OverColor.mk c) ≅ ((OverColor.Discrete.pairτ S.FDiscrete S.τ).obj
(Discrete.mk (c i))) ⊗
(OverColor.lift.obj S.FDiscrete).obj (OverColor.mk (c ∘ i.succAbove ∘ j.succAbove)) :=
(S.F.mapIso (OverColor.equivToIso (HepLean.Fin.finExtractTwo i j))).trans <|
(S.F.mapIso (OverColor.mkSum (c ∘ (HepLean.Fin.finExtractTwo i j).symm))).trans <|
(S.F.μIso _ _).symm.trans <| by
refine tensorIso (S.contrFin1Fin1 c i j h) (S.F.mapIso (OverColor.mkIso (by ext x; simp)))
2024-10-18 16:08:17 +00:00
lemma contrIso_hom_hom {n : } {c1 : Fin n.succ.succ → S.C}
2024-10-19 09:19:29 +00:00
{i : Fin n.succ.succ} {j : Fin n.succ} {h : c1 (i.succAbove j) = S.τ (c1 i)} :
(S.contrIso c1 i j h).hom.hom =
(S.F.map (equivToIso (HepLean.Fin.finExtractTwo i j)).hom).hom ≫
(S.F.map (mkSum (c1 ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)).hom).hom ≫
(S.F.μIso (OverColor.mk ((c1 ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inl))
(OverColor.mk ((c1 ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inr))).inv.hom ≫
((S.contrFin1Fin1 c1 i j h).hom.hom ⊗
(S.F.map (mkIso (contrIso.proof_1 S c1 i j)).hom).hom) := by
rw [contrIso]
2024-10-19 09:19:29 +00:00
simp [Nat.succ_eq_add_one, Action.instMonoidalCategory_tensorObj_V, Action.comp_hom,
extractOne_homToEquiv, Action.instMonoidalCategory_tensorHom_hom]
2024-10-18 16:08:17 +00:00
/-- `contrMap` is a function that takes a natural number `n`, a function `c` from
2024-10-16 16:38:36 +00:00
`Fin n.succ.succ` to `S.C`, an index `i` of type `Fin n.succ.succ`, an index `j` of type
`Fin n.succ`, and a proof `h` that `c (i.succAbove j) = S.τ (c i)`. It returns a morphism
corresponding to the contraction of the `i`th index with the `i.succAbove j` index.
--/
def contrMap {n : } (c : Fin n.succ.succ → S.C)
2024-10-15 06:08:56 +00:00
(i : Fin n.succ.succ) (j : Fin n.succ) (h : c (i.succAbove j) = S.τ (c i)) :
S.F.obj (OverColor.mk c) ⟶
2024-10-19 08:33:49 +00:00
S.F.obj (OverColor.mk (c ∘ i.succAbove ∘ j.succAbove)) :=
2024-10-15 06:08:56 +00:00
(S.contrIso c i j h).hom ≫
(tensorHom (S.contr.app (Discrete.mk (c i))) (𝟙 _)) ≫
(MonoidalCategory.leftUnitor _).hom
2024-10-19 10:07:03 +00:00
/-- Casts an element of the monoidal unit of `Rep S.k S.G` to the field `S.k`. -/
2024-10-19 09:19:29 +00:00
def castToField (v : (↑((𝟙_ (Discrete S.C ⥤ Rep S.k S.G)).obj { as := c }).V)) : S.k := v
2024-10-18 16:08:17 +00:00
lemma contrMap_tprod {n : } (c : Fin n.succ.succ → S.C)
(i : Fin n.succ.succ) (j : Fin n.succ) (h : c (i.succAbove j) = S.τ (c i))
(x : (i : Fin n.succ.succ) → S.FDiscrete.obj (Discrete.mk (c i))) :
2024-10-19 09:19:29 +00:00
(S.contrMap c i j h).hom (PiTensorProduct.tprod S.k x) =
2024-10-18 16:08:17 +00:00
(S.castToField ((S.contr.app (Discrete.mk (c i))).hom ((x i) ⊗ₜ[S.k]
2024-10-19 09:19:29 +00:00
(S.FDiscrete.map (Discrete.eqToHom h)).hom (x (i.succAbove j)))) : S.k)
2024-10-19 09:47:23 +00:00
• (PiTensorProduct.tprod S.k (fun k => x (i.succAbove (j.succAbove k))) :
S.F.obj (OverColor.mk (c ∘ i.succAbove ∘ j.succAbove))) := by
2024-10-18 16:08:17 +00:00
rw [contrMap, contrIso]
simp only [Nat.succ_eq_add_one, S.F_def, Iso.trans_hom, Functor.mapIso_hom, Iso.symm_hom,
tensorIso_hom, Monoidal.tensorUnit_obj, tensorHom_id,
Category.assoc, Action.comp_hom, Action.instMonoidalCategory_tensorObj_V,
Action.instMonoidalCategory_tensorHom_hom, Action.instMonoidalCategory_tensorUnit_V,
Action.instMonoidalCategory_whiskerRight_hom, Functor.id_obj, mk_hom, ModuleCat.coe_comp,
Function.comp_apply, Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, Functor.comp_obj, Discrete.functor_obj_eq_as]
change (λ_ ((lift.obj S.FDiscrete).obj _)).hom.hom
2024-10-19 09:47:23 +00:00
(((S.contr.app { as := c i }).hom ▷ ((lift.obj S.FDiscrete).obj
(OverColor.mk (c ∘ i.succAbove ∘ j.succAbove))).V)
(((S.contrFin1Fin1 c i j h).hom.hom ⊗ ((lift.obj S.FDiscrete).map (mkIso _).hom).hom)
(((lift.obj S.FDiscrete).μIso (OverColor.mk ((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)
∘ Sum.inl))
(OverColor.mk ((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inr))).inv.hom
(((lift.obj S.FDiscrete).map (mkSum (c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)).hom).hom
(((lift.obj S.FDiscrete).map (equivToIso (HepLean.Fin.finExtractTwo i j)).hom).hom
((PiTensorProduct.tprod S.k) x)))))) = _
2024-10-18 16:08:17 +00:00
rw [lift.map_tprod]
change (λ_ ((lift.obj S.FDiscrete).obj (OverColor.mk (c ∘ i.succAbove ∘ j.succAbove)))).hom.hom
2024-10-19 09:47:23 +00:00
(((S.contr.app { as := c i }).hom ▷
((lift.obj S.FDiscrete).obj (OverColor.mk (c ∘ i.succAbove ∘ j.succAbove))).V)
(((S.contrFin1Fin1 c i j h).hom.hom ⊗ ((lift.obj S.FDiscrete).map (mkIso _).hom).hom)
(((lift.obj S.FDiscrete).μIso (OverColor.mk
((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inl))
(OverColor.mk ((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inr))).inv.hom
(((lift.obj S.FDiscrete).map (mkSum (c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)).hom).hom
((PiTensorProduct.tprod S.k) fun i_1 =>
(lift.discreteFunctorMapEqIso S.FDiscrete _)
(x ((Hom.toEquiv (equivToIso (HepLean.Fin.finExtractTwo i j)).hom).symm i_1))))))) = _
2024-10-18 16:08:17 +00:00
rw [lift.map_tprod]
change (λ_ ((lift.obj S.FDiscrete).obj (OverColor.mk (c ∘ i.succAbove ∘ j.succAbove)))).hom.hom
2024-10-19 09:47:23 +00:00
(((S.contr.app { as := c i }).hom ▷ ((lift.obj S.FDiscrete).obj
(OverColor.mk (c ∘ i.succAbove ∘ j.succAbove))).V)
(((S.contrFin1Fin1 c i j h).hom.hom ⊗ ((lift.obj S.FDiscrete).map (mkIso _).hom).hom)
(((lift.obj S.FDiscrete).μIso
(OverColor.mk ((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inl))
(OverColor.mk ((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inr))).inv.hom
((PiTensorProduct.tprod S.k) fun i_1 =>
(lift.discreteFunctorMapEqIso S.FDiscrete _)
((lift.discreteFunctorMapEqIso S.FDiscrete _)
(x ((Hom.toEquiv (equivToIso (HepLean.Fin.finExtractTwo i j)).hom).symm
((Hom.toEquiv (mkSum (c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)).hom).symm i_1)))))))) = _
2024-10-18 16:08:17 +00:00
rw [lift.μIso_inv_tprod]
change (λ_ ((lift.obj S.FDiscrete).obj (OverColor.mk (c ∘ i.succAbove ∘ j.succAbove)))).hom.hom
2024-10-19 09:47:23 +00:00
(((S.contr.app { as := c i }).hom ▷ ((lift.obj S.FDiscrete).obj
(OverColor.mk (c ∘ i.succAbove ∘ j.succAbove))).V)
((TensorProduct.map (S.contrFin1Fin1 c i j h).hom.hom
((lift.obj S.FDiscrete).map (mkIso _).hom).hom)
(((PiTensorProduct.tprod S.k) fun i_1 =>
(lift.discreteFunctorMapEqIso S.FDiscrete _)
((lift.discreteFunctorMapEqIso S.FDiscrete _) (x
((Hom.toEquiv (equivToIso (HepLean.Fin.finExtractTwo i j)).hom).symm
((Hom.toEquiv (mkSum (c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)).hom).symm
(Sum.inl i_1)))))) ⊗ₜ[S.k] (PiTensorProduct.tprod S.k) fun i_1 =>
(lift.discreteFunctorMapEqIso S.FDiscrete _) ((lift.discreteFunctorMapEqIso S.FDiscrete _)
2024-10-19 09:19:29 +00:00
(x ((Hom.toEquiv (equivToIso (HepLean.Fin.finExtractTwo i j)).hom).symm
((Hom.toEquiv
(mkSum (c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)).hom).symm (Sum.inr i_1)))))))) = _
2024-10-18 16:08:17 +00:00
rw [TensorProduct.map_tmul]
rw [contrFin1Fin1_hom_hom_tprod]
simp only [Nat.succ_eq_add_one, Action.instMonoidalCategory_tensorObj_V,
Action.instMonoidalCategory_tensorUnit_V, Fin.isValue, mk_hom, Function.comp_apply,
Discrete.functor_obj_eq_as, instMonoidalCategoryStruct_tensorObj_left, mkSum_homToEquiv,
Equiv.refl_symm, Functor.id_obj, ModuleCat.MonoidalCategory.whiskerRight_apply]
rw [Action.instMonoidalCategory_leftUnitor_hom_hom]
2024-10-19 09:19:29 +00:00
simp only [Monoidal.tensorUnit_obj, Action.instMonoidalCategory_tensorUnit_V, Fin.isValue,
ModuleCat.MonoidalCategory.leftUnitor_hom_apply]
2024-10-18 16:08:17 +00:00
congr 1
/- The contraction. -/
2024-10-19 09:19:29 +00:00
· simp only [Fin.isValue, castToField]
2024-10-18 16:08:17 +00:00
congr 2
2024-10-19 09:19:29 +00:00
· simp only [Fin.isValue, lift.discreteFunctorMapEqIso, eqToIso_refl, Functor.mapIso_refl,
Iso.refl_hom, Action.id_hom, Iso.refl_inv, LinearEquiv.ofLinear_apply]
2024-10-18 16:08:17 +00:00
rfl
2024-10-19 09:47:23 +00:00
· simp only [Fin.isValue, lift.discreteFunctorMapEqIso, eqToIso_refl, Functor.mapIso_refl,
Iso.refl_hom, Action.id_hom, Iso.refl_inv, LinearEquiv.ofLinear_apply]
2024-10-18 16:08:17 +00:00
change (S.FDiscrete.map (eqToHom _)).hom
(x (((HepLean.Fin.finExtractTwo i j)).symm ((Sum.inl (Sum.inr 0))))) = _
2024-10-19 09:19:29 +00:00
simp only [Nat.succ_eq_add_one, Fin.isValue]
2024-10-18 16:08:17 +00:00
have h1' {a b d: Fin n.succ.succ} (hbd : b =d) (h : c d = S.τ (c a)) (h' : c b = S.τ (c a)) :
2024-10-19 09:19:29 +00:00
(S.FDiscrete.map (Discrete.eqToHom (h))).hom (x d) =
(S.FDiscrete.map (Discrete.eqToHom h')).hom (x b) := by
2024-10-18 16:08:17 +00:00
subst hbd
rfl
refine h1' ?_ ?_ ?_
2024-10-19 09:19:29 +00:00
simp only [Nat.succ_eq_add_one, Fin.isValue, HepLean.Fin.finExtractTwo_symm_inl_inr_apply]
2024-10-18 16:08:17 +00:00
simp [h]
/- The tensor. -/
· erw [lift.map_tprod]
apply congrArg
funext d
2024-10-19 09:19:29 +00:00
simp only [mk_hom, Function.comp_apply, lift.discreteFunctorMapEqIso, Functor.mapIso_hom,
eqToIso.hom, Functor.mapIso_inv, eqToIso.inv, eqToIso_refl, Functor.mapIso_refl, Iso.refl_hom,
Action.id_hom, Iso.refl_inv, LinearEquiv.ofLinear_apply]
2024-10-18 16:08:17 +00:00
change (S.FDiscrete.map (eqToHom _)).hom
2024-10-19 09:19:29 +00:00
((x ((HepLean.Fin.finExtractTwo i j).symm (Sum.inr (d))))) = _
simp only [Nat.succ_eq_add_one]
2024-10-19 09:47:23 +00:00
have h1 : ((HepLean.Fin.finExtractTwo i j).symm (Sum.inr d))
= (i.succAbove (j.succAbove d)) := HepLean.Fin.finExtractTwo_symm_inr_apply i j d
2024-10-18 16:08:17 +00:00
have h1' {a b : Fin n.succ.succ} (h : a = b) :
(S.FDiscrete.map (eqToHom (by rw [h]))).hom (x a) = x b := by
subst h
simp
exact h1' h1
end TensorStruct
2024-10-08 07:52:55 +00:00
/-- A syntax tree for tensor expressions. -/
inductive TensorTree (S : TensorStruct) : ∀ {n : }, (Fin n → S.C) → Type where
2024-10-16 16:38:36 +00:00
/-- A general tensor node. -/
2024-10-08 07:52:55 +00:00
| tensorNode {n : } {c : Fin n → S.C} (T : S.F.obj (OverColor.mk c)) : TensorTree S c
2024-10-16 16:38:36 +00:00
/-- A node consisting of a single vector. -/
| vecNode {c : S.C} (v : S.FDiscrete.obj (Discrete.mk c)) : TensorTree S ![c]
/-- A node consisting of a two tensor. -/
| twoNode {c1 c2 : S.C}
(v : (S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)).V) :
2024-10-16 16:38:36 +00:00
TensorTree S ![c1, c2]
/-- A node consisting of a three tensor. -/
| threeNode {c1 c2 c3 : S.C}
(v : S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2) ⊗
S.FDiscrete.obj (Discrete.mk c3)) : TensorTree S ![c1, c2, c3]
/-- A general constant node. -/
| constNode {n : } {c : Fin n → S.C} (T : 𝟙_ (Rep S.k S.G) ⟶ S.F.obj (OverColor.mk c)) :
TensorTree S c
/-- A constant vector. -/
| constVecNode {c : S.C} (v : 𝟙_ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c)) :
TensorTree S ![c]
/-- A constant two tensor (e.g. metric and unit). -/
| constTwoNode {c1 c2 : S.C}
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)) :
TensorTree S ![c1, c2]
/-- A constant three tensor (e.g. Pauli-matrices). -/
| constThreeNode {c1 c2 c3 : S.C}
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2) ⊗
S.FDiscrete.obj (Discrete.mk c3)) : TensorTree S ![c1, c2, c3]
/-- A node corresponding to the addition of two tensors. -/
| add {n : } {c : Fin n → S.C} : TensorTree S c → TensorTree S c → TensorTree S c
2024-10-16 16:38:36 +00:00
/-- A node corresponding to the permutation of indices of a tensor. -/
| perm {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C}
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) (t : TensorTree S c) : TensorTree S c1
| prod {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C}
(t : TensorTree S c) (t1 : TensorTree S c1) : TensorTree S (Sum.elim c c1 ∘ finSumFinEquiv.symm)
| smul {n : } {c : Fin n → S.C} : S.k → TensorTree S c → TensorTree S c
/-- The negative of a node. -/
| neg {n : } {c : Fin n → S.C} : TensorTree S c → TensorTree S c
2024-10-16 16:38:36 +00:00
| contr {n : } {c : Fin n.succ.succ → S.C} : (i : Fin n.succ.succ) →
(j : Fin n.succ) → (h : c (i.succAbove j) = S.τ (c i)) → TensorTree S c →
TensorTree S (c ∘ Fin.succAbove i ∘ Fin.succAbove j)
2024-10-08 07:26:23 +00:00
| eval {n : } {c : Fin n.succ → S.C} :
(i : Fin n.succ) → (x : Fin (S.evalNo (c i))) → TensorTree S c →
2024-10-08 07:52:55 +00:00
TensorTree S (c ∘ Fin.succAbove i)
namespace TensorTree
variable {S : TensorStruct} {n : } {c : Fin n → S.C} (T : TensorTree S c)
open MonoidalCategory
open TensorProduct
2024-10-16 16:38:36 +00:00
/-- The node `twoNode` of a tensor tree, with all arguments explicit. -/
abbrev twoNodeE (S : TensorStruct) (c1 c2 : S.C)
(v : (S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)).V) :
2024-10-16 16:38:36 +00:00
TensorTree S ![c1, c2] := twoNode v
/-- The node `constTwoNodeE` of a tensor tree, with all arguments explicit. -/
abbrev constTwoNodeE (S : TensorStruct) (c1 c2 : S.C)
2024-10-16 16:42:20 +00:00
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)) :
2024-10-16 16:38:36 +00:00
TensorTree S ![c1, c2] := constTwoNode v
/-- The node `constThreeNodeE` of a tensor tree, with all arguments explicit. -/
2024-10-16 16:42:20 +00:00
abbrev constThreeNodeE (S : TensorStruct) (c1 c2 c3 : S.C)
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2) ⊗
2024-10-16 16:38:36 +00:00
S.FDiscrete.obj (Discrete.mk c3)) : TensorTree S ![c1, c2, c3] :=
constThreeNode v
2024-10-08 07:52:55 +00:00
/-- The number of nodes in a tensor tree. -/
def size : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → := fun
| tensorNode _ => 1
2024-10-16 16:38:36 +00:00
| vecNode _ => 1
| twoNode _ => 1
| threeNode _ => 1
| constNode _ => 1
| constVecNode _ => 1
| constTwoNode _ => 1
| constThreeNode _ => 1
| add t1 t2 => t1.size + t2.size + 1
| perm _ t => t.size + 1
| neg t => t.size + 1
| smul _ t => t.size + 1
| prod t1 t2 => t1.size + t2.size + 1
2024-10-16 16:38:36 +00:00
| contr _ _ _ t => t.size + 1
2024-10-08 07:26:23 +00:00
| eval _ _ t => t.size + 1
noncomputable section
2024-10-08 07:52:55 +00:00
/-- The underlying tensor a tensor tree corresponds to.
Note: This function is not fully defined yet. -/
def tensor : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → S.F.obj (OverColor.mk c) := fun
| tensorNode t => t
| constTwoNode t => (OverColor.Discrete.pairIsoSep S.FDiscrete).hom.hom (t.hom (1 : S.k))
| add t1 t2 => t1.tensor + t2.tensor
| perm σ t => (S.F.map σ).hom t.tensor
| neg t => - t.tensor
| smul a t => a • t.tensor
| prod t1 t2 => (S.F.map (OverColor.equivToIso finSumFinEquiv).hom).hom
((S.F.μ _ _).hom (t1.tensor ⊗ₜ t2.tensor))
2024-10-16 16:42:20 +00:00
| contr i j h t => (S.contrMap _ i j h).hom t.tensor
| _ => 0
/-!
## Tensor on different nodes.
-/
@[simp]
lemma tensoreNode_tensor {c : Fin n → S.C} (T : S.F.obj (OverColor.mk c)) :
(tensorNode T).tensor = T := rfl
@[simp]
lemma constTwoNode_tensor {c1 c2 : S.C}
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)) :
2024-10-19 09:47:23 +00:00
(constTwoNode v).tensor =
(OverColor.Discrete.pairIsoSep S.FDiscrete).hom.hom (v.hom (1 : S.k)) :=
rfl
2024-10-19 09:47:23 +00:00
lemma prod_tensor {c1 : Fin n → S.C} {c2 : Fin m → S.C} (t1 : TensorTree S c1)
(t2 : TensorTree S c2) :
(prod t1 t2).tensor = (S.F.map (OverColor.equivToIso finSumFinEquiv).hom).hom
2024-10-19 09:19:29 +00:00
((S.F.μ _ _).hom (t1.tensor ⊗ₜ t2.tensor)) := rfl
lemma add_tensor (t1 t2 : TensorTree S c) : (add t1 t2).tensor = t1.tensor + t2.tensor := rfl
lemma perm_tensor (σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) (t : TensorTree S c) :
(perm σ t).tensor = (S.F.map σ).hom t.tensor := rfl
2024-10-19 09:47:23 +00:00
lemma contr_tensor {n : } {c : Fin n.succ.succ → S.C} {i : Fin n.succ.succ} {j : Fin n.succ}
{h : c (i.succAbove j) = S.τ (c i)} (t : TensorTree S c) :
(contr i j h t).tensor = (S.contrMap c i j h).hom t.tensor := rfl
lemma neg_tensor (t : TensorTree S c) : (neg t).tensor = - t.tensor := rfl
/-!
## Equality of tensors and rewrites.
-/
lemma contr_tensor_eq {n : } {c : Fin n.succ.succ → S.C} {T1 T2 : TensorTree S c}
(h : T1.tensor = T2.tensor) {i : Fin n.succ.succ} {j : Fin n.succ}
{h' : c (i.succAbove j) = S.τ (c i)} :
(contr i j h' T1).tensor = (contr i j h' T2).tensor := by
simp only [Nat.succ_eq_add_one, contr_tensor]
rw [h]
lemma prod_tensor_eq_fst {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C}
2024-10-19 09:19:29 +00:00
{T1 T1' : TensorTree S c} { T2 : TensorTree S c1}
(h : T1.tensor = T1'.tensor) :
(prod T1 T2).tensor = (prod T1' T2).tensor := by
2024-10-19 09:19:29 +00:00
simp only [prod_tensor, Functor.id_obj, OverColor.mk_hom, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj]
rw [h]
lemma prod_tensor_eq_snd {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C}
{T1 : TensorTree S c} {T2 T2' : TensorTree S c1}
(h : T2.tensor = T2'.tensor) :
(prod T1 T2).tensor = (prod T1 T2').tensor := by
2024-10-19 09:19:29 +00:00
simp only [prod_tensor, Functor.id_obj, OverColor.mk_hom, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj]
rw [h]
end
end TensorTree
2024-10-15 06:08:56 +00:00
end