feat: Three tensor nodes
This commit is contained in:
parent
6bca8c295c
commit
ed162d7e79
5 changed files with 68 additions and 4 deletions
|
@ -69,6 +69,14 @@ def pairIsoSep {c1 c2 : C} : F.obj (Discrete.mk c1) ⊗ F.obj (Discrete.mk c2)
|
|||
fin_cases x
|
||||
rfl
|
||||
|
||||
def tripleIsoSep {c1 c2 c3 : C} :
|
||||
F.obj (Discrete.mk c1) ⊗ F.obj (Discrete.mk c2) ⊗ F.obj (Discrete.mk c3) ≅
|
||||
(lift.obj F).obj (OverColor.mk ![c1,c2,c3]) :=
|
||||
(whiskerLeftIso (F.obj (Discrete.mk c1)) (pairIsoSep F (c1 := c2) (c2 := c3))).trans <|
|
||||
(whiskerRightIso (forgetLiftApp F c1).symm _).trans <|
|
||||
((lift.obj F).μIso _ _).trans <|
|
||||
(lift.obj F).mapIso fin3Iso'.symm
|
||||
|
||||
lemma pairIsoSep_tmul {c1 c2 : C} (x : F.obj (Discrete.mk c1)) (y : F.obj (Discrete.mk c2)) :
|
||||
(pairIsoSep F).hom.hom (x ⊗ₜ[k] y) =
|
||||
PiTensorProduct.tprod k (Fin.cases x (Fin.cases y (fun i => Fin.elim0 i))) := by
|
||||
|
|
|
@ -76,6 +76,33 @@ def fin2Iso {c : Fin 2 → C} : mk c ≅ mk ![c 0] ⊗ mk ![c 1] := by
|
|||
fin_cases x
|
||||
rfl
|
||||
|
||||
def fin3Iso {c : Fin 3 → C} : mk c ≅ mk ![c 0] ⊗ mk ![c 1, c 2] := by
|
||||
let e1 : Fin 3 ≃ Fin 1 ⊕ Fin 2 := (finSumFinEquiv (n := 2)).symm
|
||||
apply (equivToIso e1).trans
|
||||
apply (mkSum _).trans
|
||||
refine tensorIso (mkIso ?_) (mkIso ?_)
|
||||
· funext x
|
||||
fin_cases x
|
||||
rfl
|
||||
· funext x
|
||||
fin_cases x
|
||||
rfl
|
||||
rfl
|
||||
|
||||
|
||||
def fin3Iso' {c1 c2 c3 : C} : mk ![c1, c2, c3] ≅ mk (fun (_ : Fin 1) => c1) ⊗ mk ![c2, c3] := by
|
||||
let e1 : Fin 3 ≃ Fin 1 ⊕ Fin 2 := (finSumFinEquiv (n := 2)).symm
|
||||
apply (equivToIso e1).trans
|
||||
apply (mkSum _).trans
|
||||
refine tensorIso (mkIso ?_) (mkIso ?_)
|
||||
· funext x
|
||||
fin_cases x
|
||||
rfl
|
||||
· funext x
|
||||
fin_cases x
|
||||
rfl
|
||||
rfl
|
||||
|
||||
/-- Removes a given `i : Fin n.succ.succ` from a morphism in `OverColor C`. -/
|
||||
def extractOne {n : ℕ} (i : Fin n.succ.succ)
|
||||
{c1 c2 : Fin n.succ.succ → C} (σ : mk c1 ⟶ mk c2) :
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue