Built with Alectryon, running Coq+SerAPI v8.19.0+0.19.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
[Loading ML file number_string_notation_plugin.cmxs (using legacy method) ... done]
Require Import WildCat.Core WildCat.Bifunctor WildCat.Prod WildCat.Equiv.Require Import WildCat.NatTrans WildCat.Square WildCat.Opposite.(** * Building Symmetric Monoidal Categories *)(** The following construction is what we call the "twist construction". It is a way to build a symmetric monoidal category from simpler pieces than the axioms ask for.The core observation is that the associator can be broken up into a [braid] and what we call a [twist] map. The twist map takes a right associated triple [(A, (B, C))] and swaps the first two factors [(B, (A, C)]. Together with functoriality of the tensor and the braiding, here termed [braid] we can simplify the axioms we ask for.For instance, the hexagon identity is about associators, but if we unfold the definition and simplify the diagram, we get a diagram about only twists and braids.This means in practice, you can show a category has a symmetric monoidal structure by proving some simpler axioms. This idea has been used in TriJoin.v to show the associativity of join for example. *)SectionTwistConstruction.(** The aim of this section is to build a symmetric monoidal category. We do this piecewise so that the separate steps are useful in and of themselves. Our basic starting assumption is that we have a category with equivalences, a bifunctor called the tensor product, and a unit object.*)Context (A : Type) `{HasEquivs A}
(cat_tensor : A -> A -> A) (cat_tensor_unit : A)
`{!Is0Bifunctor cat_tensor, !Is1Bifunctor cat_tensor}
(** Next we postulate the existence of a [braid] map. This takes a tensor pair and swaps the factors. We also postulate that [braid] is natural in both factors and self-inverse. *)
(braid : SymmetricBraiding cat_tensor)
(** We postulate the existence of a [twist] map. This takes a right associated triple [(A, (B, C))] and swaps the first two factors [(B, (A, C)]. We also postulate that [twist] is natural in all three factors and self-inverse. *)
(twist : forallabc, cat_tensor a (cat_tensor b c) $-> cat_tensor b (cat_tensor a c))
(twist_twist : forallabc, twist a b c $o twist b a c $== Id _)
(twist_nat : forallaa'bb'cc' (f : a $-> a') (g : b $-> b') (h : c $-> c'),
twist a' b' c' $o fmap11 cat_tensor f (fmap11 cat_tensor g h)
$== fmap11 cat_tensor g (fmap11 cat_tensor f h) $o twist a b c)
(** We assume that there is a natural isomorphism [right_unitor] witnessing the right unit law. The left unit law will be derived from this one. We also assume a coherence called [twist_unitor] which determines how the right_unitor interacts with [braid] and [twist]. This is the basis of the triangle axiom. *)
(right_unitor : RightUnitor cat_tensor cat_tensor_unit)
(twist_unitor : forallab, fmap01 cat_tensor a (right_unitor b)
$== braid b a $o fmap01 cat_tensor b (right_unitor a) $o twist a b cat_tensor_unit)
(** The hexagon identity is about the interaction of associators and braids. We will derive this axiom from an analogous one for twists and braids. *)
(twist_hexagon : forallabc,
fmap01 cat_tensor c (braid b a) $o twist b c a $o fmap01 cat_tensor b (braid a c)
$== twist a c b $o fmap01 cat_tensor a (braid b c) $o twist b a c)
(** The 9-gon identity. TODO: explain this *)
(twist_9_gon : forallabcd,
fmap01 cat_tensor c (braid (cat_tensor a b) d)
$o twist (cat_tensor a b) c d
$o braid (cat_tensor c d) (cat_tensor a b)
$o twist a (cat_tensor c d) b
$o fmap01 cat_tensor a (braid b (cat_tensor c d))
$== fmap01 cat_tensor c (twist a d b)
$o fmap01 cat_tensor c (fmap01 cat_tensor a (braid b d))
$o twist a c (cat_tensor b d)
$o fmap01 cat_tensor a (twist b c d))
.(** *** Setup *)(** Before starting the proofs, we need to setup some useful definitions and helpful lemmas for working with diagrams. *)(** We give notations and abbreviations to the morphisms that will appear in diagrams. This helps us read the goal and understand what is happening, otherwise it is too verbose. *)Declare Scope monoidal_scope.LocalInfix"⊗" := cat_tensor (at level40) : monoidal_scope.LocalInfix"⊗R" := (fmap01 cat_tensor) (at level40) : monoidal_scope.LocalInfix"⊗L" := (fmap10 cat_tensor) (at level40) : monoidal_scope.LocalNotation"f ∘ g" := (f $o g) (at level61, left associativity, format"f '/' '∘' g") : monoidal_scope.(** TODO: in 8.19 this notation causes issues, when updating to 8.20 we should uncomment this. *)(* Local Notation "f $== g :> A" := (GpdHom (A := A) f g) (at level 70, format "'[v' '[v' f ']' '/' $== '/' '[v' g ']' '/' :> '[' A ']' ']'") : long_path_scope. *)LocalOpen Scope monoidal_scope.(** [twist] is an equivalence which we will call [twiste]. *)Local Definitiontwisteabc
: cat_tensor a (cat_tensor b c) $<~> cat_tensor b (cat_tensor a c)
:= cate_adjointify (twist a b c) (twist b a c)
(twist_twist a b c) (twist_twist b a c).(** *** Finer naturality *)(** The naturality postulates we have for [twist] are natural in all their arguments simultaneously. We show the finer naturality of [twist] in each argument separately as this becomes more useful in practice. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, a': A f: a $-> a' b, c: A
twist a' b c ∘ f ⊗L (b ⊗ c) $== b ⊗R (f ⊗L c)
∘ twist a b c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, a': A f: a $-> a' b, c: A
twist a' b c ∘ f ⊗L (b ⊗ c) $== b ⊗R (f ⊗L c)
∘ twist a b c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, a': A f: a $-> a' b, c: A
fmap11 cat_tensor f (fmap11 cat_tensor (Id b) (Id c)) $->
f ⊗L (b ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, a': A f: a $-> a' b, c: A
fmap11 cat_tensor (Id b) (fmap11 cat_tensor f (Id c)) $==
b ⊗R (f ⊗L c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, a': A f: a $-> a' b, c: A
fmap11 cat_tensor f (fmap11 cat_tensor (Id b) (Id c)) $->
f ⊗L (b ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, a': A f: a $-> a' b, c: A
fmap11 cat_tensor (Id b) (Id c) $== Id (b ⊗ c)
rapply fmap11_id.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, a': A f: a $-> a' b, c: A
fmap11 cat_tensor (Id b) (fmap11 cat_tensor f (Id c)) $==
b ⊗R (f ⊗L c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, a': A f: a $-> a' b, c: A
fmap11 cat_tensor f (Id c) $== f ⊗L c
rapply fmap10_is_fmap11.Defined.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, b': A g: b $-> b' c: A
twist a b' c ∘ a ⊗R (g ⊗L c) $== g ⊗L (a ⊗ c)
∘ twist a b c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, b': A g: b $-> b' c: A
twist a b' c ∘ a ⊗R (g ⊗L c) $== g ⊗L (a ⊗ c)
∘ twist a b c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, b': A g: b $-> b' c: A
fmap11 cat_tensor (Id a) (fmap11 cat_tensor g (Id c)) $->
a ⊗R (g ⊗L c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, b': A g: b $-> b' c: A
fmap11 cat_tensor g (fmap11 cat_tensor (Id a) (Id c)) $==
g ⊗L (a ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, b': A g: b $-> b' c: A
fmap11 cat_tensor (Id a) (fmap11 cat_tensor g (Id c)) $->
a ⊗R (g ⊗L c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, b': A g: b $-> b' c: A
fmap11 cat_tensor g (Id c) $== g ⊗L c
rapply fmap10_is_fmap11.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, b': A g: b $-> b' c: A
fmap11 cat_tensor g (fmap11 cat_tensor (Id a) (Id c)) $==
g ⊗L (a ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, b': A g: b $-> b' c: A
fmap11 cat_tensor (Id a) (Id c) $== Id (a ⊗ c)
rapply fmap11_id.Defined.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, c': A h: c $-> c'
twist a b c' ∘ a ⊗R (b ⊗R h) $== b ⊗R (a ⊗R h)
∘ twist a b c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, c': A h: c $-> c'
twist a b c' ∘ a ⊗R (b ⊗R h) $== b ⊗R (a ⊗R h)
∘ twist a b c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, c': A h: c $-> c'
fmap11 cat_tensor (Id a) (fmap11 cat_tensor (Id b) h) $->
a ⊗R (b ⊗R h)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, c': A h: c $-> c'
fmap11 cat_tensor (Id b) (fmap11 cat_tensor (Id a) h) $==
b ⊗R (a ⊗R h)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, c': A h: c $-> c'
fmap11 cat_tensor (Id a) (fmap11 cat_tensor (Id b) h) $->
a ⊗R (b ⊗R h)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, c': A h: c $-> c'
fmap11 cat_tensor (Id b) h $== b ⊗R h
rapply fmap01_is_fmap11.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, c': A h: c $-> c'
fmap11 cat_tensor (Id b) (fmap11 cat_tensor (Id a) h) $==
b ⊗R (a ⊗R h)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, c': A h: c $-> c'
fmap11 cat_tensor (Id a) h $== a ⊗R h
rapply fmap01_is_fmap11.Defined.(** *** Movement lemmas *)(** Here we collect lemmas about moving morphisms around in a diagram. We could have created [cate_moveL_eM]-style lemmas for [CatIsEquiv] but this leads to a lot of unnecessary unfolding and duplication. It is typically easier to use a hand crafted lemma for each situation. *)(** TODO: A lot of these proofs are copy and pasted between lemmas. We need to work out an efficient way of proving them. *)(** **** Moving [twist] *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: d $-> a ⊗ (b ⊗ c) g: d $-> b ⊗ (a ⊗ c)
twist a b c ∘ f $== g -> f $== twist b a c ∘ g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: d $-> a ⊗ (b ⊗ c) g: d $-> b ⊗ (a ⊗ c)
twist a b c ∘ f $== g -> f $== twist b a c ∘ g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: d $-> a ⊗ (b ⊗ c) g: d $-> b ⊗ (a ⊗ c) p: twist a b c ∘ f $== g
f $== twist b a c ∘ g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: d $-> a ⊗ (b ⊗ c) g: d $-> b ⊗ (a ⊗ c) p: twist a b c ∘ f $== g
twiste a b c ∘ f $== twiste a b c ∘ (twist b a c ∘ g)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: d $-> a ⊗ (b ⊗ c) g: d $-> b ⊗ (a ⊗ c) p: twist a b c ∘ f $== g
g $== twiste a b c ∘ twist b a c ∘ g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: d $-> a ⊗ (b ⊗ c) g: d $-> b ⊗ (a ⊗ c) p: twist a b c ∘ f $== g
twiste a b c ∘ twist b a c $-> Id (b ⊗ (a ⊗ c))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: d $-> a ⊗ (b ⊗ c) g: d $-> b ⊗ (a ⊗ c) p: twist a b c ∘ f $== g
twist a b c ∘ twist b a c $== Id (b ⊗ (a ⊗ c))
apply twist_twist.Defined.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d
f ∘ twist a b c $== g -> f $== g ∘ twist b a c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d
f ∘ twist a b c $== g -> f $== g ∘ twist b a c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d p: f ∘ twist a b c $== g
f $== g ∘ twist b a c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d p: f ∘ twist a b c $== g
f ∘ twiste a b c $== g ∘ twist b a c ∘ twiste a b c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d p: f ∘ twist a b c $== g
f ∘ twiste a b c $== g ∘ (twist b a c ∘ twiste a b c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d p: f ∘ twist a b c $== g
f ∘ twiste a b c $== g ∘ ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d p: f ∘ twist a b c $== g
twist b a c ∘ twist a b c $== ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d p: f ∘ twist a b c $== g
f ∘ twiste a b c $== g ∘ Id (a ⊗ (b ⊗ c))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d p: f ∘ twist a b c $== g
twiste a b c $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d p: f ∘ twist a b c $== g
f ∘ ?Goal $== g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d p: f ∘ twist a b c $== g
f ∘ twist a b c $== g
exact p.Defined.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: d $-> a ⊗ (b ⊗ c) g: d $-> b ⊗ (a ⊗ c)
f $== twist b a c ∘ g -> twist a b c ∘ f $== g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: d $-> a ⊗ (b ⊗ c) g: d $-> b ⊗ (a ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d
f $== g ∘ twist b a c -> f ∘ twist a b c $== g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A f: b ⊗ (a ⊗ c) $-> d g: a ⊗ (b ⊗ c) $-> d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d))
a ⊗R twist b c d ∘ f $== g ->
f $== a ⊗R twist c b d ∘ g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d))
a ⊗R twist b c d ∘ f $== g ->
f $== a ⊗R twist c b d ∘ g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
f $== a ⊗R twist c b d ∘ g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
emap01 cat_tensor a (twiste b c d) ∘ f $==
emap01 cat_tensor a (twiste b c d)
∘ (a ⊗R twist c b d ∘ g)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
emap01 cat_tensor a (twiste b c d) ∘ f $==
emap01 cat_tensor a (twiste b c d) ∘ a ⊗R twist c b d
∘ g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
emap01 cat_tensor a (twiste b c d) ∘ f $== ?Goal0 ∘ g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
?Goal0 $== emap01 cat_tensor a (twiste b c d)
∘ a ⊗R twist c b d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
?Goal0 $== emap01 cat_tensor a (twiste b c d)
∘ a ⊗R twist c b d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
?Goal0 $== ?Goal2 ∘ a ⊗R twist c b d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
emap01 cat_tensor a (twiste b c d) $-> ?Goal2
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
?Goal0 $== a ⊗R twiste b c d ∘ a ⊗R twist c b d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
Id (c ⊗ (b ⊗ d)) $== twiste b c d ∘ twist c b d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
?Goal1 ∘ twist c b d $-> Id (c ⊗ (b ⊗ d))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
twiste b c d $-> ?Goal1
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
twist b c d ∘ twist c b d $-> Id (c ⊗ (b ⊗ d))
apply twist_twist.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
emap01 cat_tensor a (twiste b c d) ∘ f $==
Id (a ⊗ (c ⊗ (b ⊗ d))) ∘ g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
emap01 cat_tensor a (twiste b c d) $==
a ⊗R twist b c d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d)) p: a ⊗R twist b c d ∘ f $== g
twiste b c d $== twist b c d
apply cate_buildequiv_fun.Defined.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e
f ∘ a ⊗R twist b c d $== g ->
f $== g ∘ a ⊗R twist c b d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e
f ∘ a ⊗R twist b c d $== g ->
f $== g ∘ a ⊗R twist c b d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
f $== g ∘ a ⊗R twist c b d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
f ∘ emap01 cat_tensor a (twiste b c d) $== g
∘ a ⊗R twist c b d
∘ emap01 cat_tensor a (twiste b c d)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
f ∘ emap01 cat_tensor a (twiste b c d) $== g
∘ (a ⊗R twist c b d
∘ emap01 cat_tensor a (twiste b c d))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
f ∘ emap01 cat_tensor a (twiste b c d) $== g ∘ ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
?Goal0 $== a ⊗R twist c b d
∘ emap01 cat_tensor a (twiste b c d)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
?Goal0 $== a ⊗R twist c b d
∘ emap01 cat_tensor a (twiste b c d)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
a ⊗R twist c b d ∘ ?Goal2 $-> ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
emap01 cat_tensor a (twiste b c d) $-> ?Goal2
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
a ⊗R twist c b d ∘ a ⊗R twiste b c d $-> ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
twist c b d ∘ twiste b c d $== Id (b ⊗ (c ⊗ d))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
twiste b c d $== ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
twist c b d ∘ ?Goal0 $== Id (b ⊗ (c ⊗ d))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
twist c b d ∘ twist b c d $== Id (b ⊗ (c ⊗ d))
apply twist_twist.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
f ∘ emap01 cat_tensor a (twiste b c d) $== g
∘ Id (a ⊗ (b ⊗ (c ⊗ d)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
emap01 cat_tensor a (twiste b c d) $==
a ⊗R twist b c d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e p: f ∘ a ⊗R twist b c d $== g
twiste b c d $== twist b c d
apply cate_buildequiv_fun.Defined.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d))
f $== a ⊗R twist c b d ∘ g ->
a ⊗R twist b c d ∘ f $== g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: e $-> a ⊗ (b ⊗ (c ⊗ d)) g: e $-> a ⊗ (c ⊗ (b ⊗ d))
f $== a ⊗R twist c b d ∘ g ->
a ⊗R twist b c d ∘ f $== g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e
f $== g ∘ a ⊗R twist c b d ->
f ∘ a ⊗R twist b c d $== g
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d, e: A f: a ⊗ (c ⊗ (b ⊗ d)) $-> e g: a ⊗ (b ⊗ (c ⊗ d)) $-> e
f $== g ∘ a ⊗R twist c b d ->
f ∘ a ⊗R twist b c d $== g
intros p; symmetry; apply moveL_fmap01_twistR; symmetry; exact p.Defined.(** *** The associator *)(** Using [braide] and [twiste] we can build an associator. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
a ⊗ (b ⊗ c) $<~> (a ⊗ b) ⊗ c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
a ⊗ (b ⊗ c) $<~> (a ⊗ b) ⊗ c
(** We can build the associator out of [braide] and [twiste]. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
a ⊗ (b ⊗ c) $<~> c ⊗ (a ⊗ b)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
a ⊗ (b ⊗ c) $<~> a ⊗ (c ⊗ b)
exact (emap01 cat_tensor a (braide _ _)).Defined.(** We would like to be able to unfold [associator_twist'] to the underlying morphisms. We use this lemma to make that process easier. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
associator_twist' a b c $== braid c (a ⊗ b)
∘ (twist a c b ∘ a ⊗R braid b c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
associator_twist' a b c $== braid c (a ⊗ b)
∘ (twist a c b ∘ a ⊗R braid b c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
twiste a c b $oE emap01 cat_tensor a (braide b c) $==
twist a c b ∘ a ⊗R braid b c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
emap01 cat_tensor a (braide b c) $== a ⊗R braid b c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braide b c $== braid b c
apply cate_buildequiv_fun.Defined.(** Now we can use [associator_twist'] and show that it is a natural equivalence in each variable. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
Associator cat_tensor
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
Associator cat_tensor
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
forallabc : A, a ⊗ (b ⊗ c) $<~> (a ⊗ b) ⊗ c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
Is1Natural
(funpat : A * A * A =>
letx := pat inletfst := fst x inletc := snd x inletfst0 := fst inleta := Overture.fst fst0 inletb := snd fst0 in a ⊗ (b ⊗ c))
(funpat : A * A * A =>
letx := pat inletfst := fst x inletc := snd x inletfst0 := fst inleta := Overture.fst fst0 inletb := snd fst0 in (a ⊗ b) ⊗ c)
(funpat : A * A * A =>
letx := pat inletfst := fst x inletc := snd x inletfst0 := fst inleta := Overture.fst fst0 inletb := snd fst0 in?associator a b c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
forallabc : A, a ⊗ (b ⊗ c) $<~> (a ⊗ b) ⊗ c
exact associator_twist'.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
Is1Natural
(funpat : A * A * A =>
letx := pat inletfst := fst x inletc := snd x inletfst0 := fst inleta := Overture.fst fst0 inletb := snd fst0 in a ⊗ (b ⊗ c))
(funpat : A * A * A =>
letx := pat inletfst := fst x inletc := snd x inletfst0 := fst inleta := Overture.fst fst0 inletb := snd fst0 in (a ⊗ b) ⊗ c)
(funpat : A * A * A =>
letx := pat inletfst := fst x inletc := snd x inletfst0 := fst inleta := Overture.fst fst0 inletb := snd fst0 in associator_twist' a b c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
forall (aa' : A * A * A) (f : a $-> a'),
(funpat : A * A * A =>
letx := pat inletfst := fst x inletc := snd x inletfst0 := fst inleta0 := Overture.fst fst0 inletb := snd fst0 in
cate_fun (associator_twist' a0 b c)) a'
∘ fmap
(uncurry (uncurry (funxyz : A => x ⊗ (y ⊗ z))))
f $==
fmap
(uncurry (uncurry (funxyz : A => (x ⊗ y) ⊗ z))) f
∘ (funpat : A * A * A =>
letx := pat inletfst := fst x inletc := snd x inletfst0 := fst inleta0 := Overture.fst fst0 inletb := snd fst0 in
cate_fun (associator_twist' a0 b c)) a
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
associator_twist' a' b' c'
∘ fmap11 cat_tensor f (fmap11 cat_tensor g h) $==
fmap (uncurry cat_tensor)
(fmap (uncurry cat_tensor) (f, g), h)
∘ associator_twist' a b c
(** To prove naturality it will be easier to reason about squares. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square (associator_twist' a b c)
(associator_twist' a' b' c')
(fmap11 cat_tensor f (fmap11 cat_tensor g h))
(fmap (uncurry cat_tensor)
(fmap (uncurry cat_tensor) (f, g), h))
(** First we remove all the equivalences from the equation. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
associator_twist' a b c $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square ?Goal (associator_twist' a' b' c')
(fmap11 cat_tensor f (fmap11 cat_tensor g h))
(fmap (uncurry cat_tensor)
(fmap (uncurry cat_tensor) (f, g), h))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square
(braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c))
(associator_twist' a' b' c')
(fmap11 cat_tensor f (fmap11 cat_tensor g h))
(fmap (uncurry cat_tensor)
(fmap (uncurry cat_tensor) (f, g), h))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square
(braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c))
?Goal (fmap11 cat_tensor f (fmap11 cat_tensor g h))
(fmap (uncurry cat_tensor)
(fmap (uncurry cat_tensor) (f, g), h))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
associator_twist' a' b' c' $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square
(braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c))
(braid c' (a' ⊗ b')
∘ (twist a' c' b' ∘ a' ⊗R braid b' c'))
(fmap11 cat_tensor f (fmap11 cat_tensor g h))
(fmap (uncurry cat_tensor)
(fmap (uncurry cat_tensor) (f, g), h))
(** The first square involving [braid] on its own is a naturality square. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square (twist a c b ∘ a ⊗R braid b c)
(twist a' c' b' ∘ a' ⊗R braid b' c')
(fmap11 cat_tensor f (fmap11 cat_tensor g h)) ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square (braid c (a ⊗ b))
(braid c' (a' ⊗ b'))
?Goal
(fmap (uncurry cat_tensor)
(fmap (uncurry cat_tensor) (f, g), h))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square (twist a c b ∘ a ⊗R braid b c)
(twist a' c' b' ∘ a' ⊗R braid b' c')
(fmap11 cat_tensor f (fmap11 cat_tensor g h))
(fmap11 cat_tensor h
(fmap (uncurry cat_tensor) (f, g)))
(** The second square is just the naturality of twist. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square (a ⊗R braid b c) (a' ⊗R braid b' c')
(fmap11 cat_tensor f (fmap11 cat_tensor g h)) ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square (twist a c b) (twist a' c' b')
?Goal
(fmap11 cat_tensor h
(fmap (uncurry cat_tensor) (f, g)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square (a ⊗R braid b c) (a' ⊗R braid b' c')
(fmap11 cat_tensor f (fmap11 cat_tensor g h))
(fmap11 cat_tensor f (fmap11 cat_tensor h g))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
a ⊗R braid b c $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square ?Goal (a' ⊗R braid b' c')
(fmap11 cat_tensor f (fmap11 cat_tensor g h))
(fmap11 cat_tensor f (fmap11 cat_tensor h g))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
a ⊗R braid b c $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square ?Goal?Goal1
(fmap11 cat_tensor f (fmap11 cat_tensor g h))
(fmap11 cat_tensor f (fmap11 cat_tensor h g))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
a' ⊗R braid b' c' $== ?Goal1
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square (fmap11 cat_tensor (Id a) (braid b c))
(fmap11 cat_tensor (Id a') (braid b' c'))
(fmap11 cat_tensor f (fmap11 cat_tensor g h))
(fmap11 cat_tensor f (fmap11 cat_tensor h g))
(** Leaving us with a square with a functor application. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square (Id a) (Id a') f f
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square (braid b c) (braid b' c')
(fmap11 cat_tensor g h)
(fmap11 cat_tensor h g)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, a', b', c': A f: a $-> a' g: b $-> b' h: c $-> c'
Square (braid b c) (braid b' c')
(fmap11 cat_tensor g h) (fmap11 cat_tensor h g)
(** We are finally left with the naturality of braid. *)apply braid_nat.Defined.(** We abbreviate the associator to [α] for the remainder of the section. *)Local Notationα := associator_twist.(** *** Unitors *)(** Since we assume the [right_unitor] exists, we can derive the [left_unitor] from it together with [braid]. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
LeftUnitor cat_tensor cat_tensor_unit
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
LeftUnitor cat_tensor cat_tensor_unit
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
NatTrans (cat_tensor cat_tensor_unit) idmap
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
foralla : A, CatIsEquiv (?alpha a)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
NatTrans (cat_tensor cat_tensor_unit) idmap
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
cat_tensor cat_tensor_unit $=> idmap
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
cat_tensor cat_tensor_unit $=> idmap
exact (funa => right_unitor a $o braid cat_tensor_unit a).
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
Is1Natural (cat_tensor cat_tensor_unit) idmap
(funa : A => right_unitor a
∘ braid cat_tensor_unit a)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
forall (aa' : A) (f : a $-> a'),
(funa0 : A => right_unitor a0
∘ braid cat_tensor_unit a0) a'
∘ fmap (cat_tensor cat_tensor_unit) f $== fmap idmap f
∘ (funa0 : A => right_unitor a0
∘ braid cat_tensor_unit a0) a
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A f: a $-> b
(funa : A => right_unitor a ∘ braid cat_tensor_unit a)
b ∘ fmap (cat_tensor cat_tensor_unit) f $==
fmap idmap f
∘ (funa : A => right_unitor a
∘ braid cat_tensor_unit a) a
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A f: a $-> b
Square
((funa : A => right_unitor a
∘ braid cat_tensor_unit a) a)
((funa : A => right_unitor a
∘ braid cat_tensor_unit a) b)
(fmap (cat_tensor cat_tensor_unit) f) (fmap idmap f)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A f: a $-> b
Square (braid cat_tensor_unit a)
(braid cat_tensor_unit b)
(fmap (cat_tensor cat_tensor_unit) f) ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A f: a $-> b
Square (right_unitor a)
(right_unitor b) ?Goal
(fmap idmap f)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A f: a $-> b
Square (braid cat_tensor_unit a)
(braid cat_tensor_unit b)
(fmap (cat_tensor cat_tensor_unit) f)
(fmap (flip cat_tensor cat_tensor_unit) f)
napply braid_nat_r.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
foralla : A,
CatIsEquiv
({|
trans_nattrans :=
funa0 : A => right_unitor a0
∘ braid cat_tensor_unit a0;
is1natural_nattrans :=
Build_Is1Natural
(funa0 : A => right_unitor a0
∘ braid cat_tensor_unit a0)
(fun (a0b : A) (f : a0 $-> b) =>
braid_nat_r f $@v
is1natural_natequiv right_unitor a0 b f)
|} a)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a: A
CatIsEquiv
({|
trans_nattrans :=
funa : A => right_unitor a
∘ braid cat_tensor_unit a;
is1natural_nattrans :=
Build_Is1Natural
(funa : A => right_unitor a
∘ braid cat_tensor_unit a)
(fun (ab : A) (f : a $-> b) =>
braid_nat_r f $@v
is1natural_natequiv right_unitor a b f)
|} a)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a: A
CatIsEquiv (braid cat_tensor_unit a)
exact (catie_braid _ _).Defined.(** *** Triangle *)(** The triangle identity can easily be proven by rearranging the diagram, cancelling and using naturality of [braid]. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
TriangleIdentity cat_tensor cat_tensor_unit
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
TriangleIdentity cat_tensor cat_tensor_unit
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A
a ⊗R left_unitor_twist b $== right_unitor a ⊗L b
∘ α a cat_tensor_unit b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A
a ⊗R left_unitor_twist b $== right_unitor a ⊗L b
∘ ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A
α a cat_tensor_unit b $== ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A
a ⊗R left_unitor_twist b $== right_unitor a ⊗L b
∘ (braid b (a ⊗ cat_tensor_unit)
∘ (twist a b cat_tensor_unit
∘ a ⊗R braid cat_tensor_unit b))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A
a ⊗R (right_unitor b ∘ braid cat_tensor_unit b) $==
right_unitor a ⊗L b
∘ (braid b (a ⊗ cat_tensor_unit)
∘ (twist a b cat_tensor_unit
∘ a ⊗R braid cat_tensor_unit b))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A
a ⊗R right_unitor b ∘ a ⊗R braid cat_tensor_unit b $==
right_unitor a ⊗L b
∘ (braid b (a ⊗ cat_tensor_unit)
∘ (twist a b cat_tensor_unit
∘ a ⊗R braid cat_tensor_unit b))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A
a ⊗R right_unitor b ∘ a ⊗R braid cat_tensor_unit b $==
right_unitor a ⊗L b ∘ braid b (a ⊗ cat_tensor_unit)
∘ twist a b cat_tensor_unit
∘ a ⊗R braid cat_tensor_unit b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b: A
braid b a ∘ b ⊗R right_unitor a $==
right_unitor a ⊗L b ∘ braid b (a ⊗ cat_tensor_unit)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
PentagonIdentity cat_tensor
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
PentagonIdentity cat_tensor
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
PentagonIdentity cat_tensor
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
α (a ⊗ b) c d ∘ α a b (c ⊗ d) $== α a b c ⊗L d
∘ α a (b ⊗ c) d ∘ a ⊗R α b c d
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
α a b (c ⊗ d) $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
α (a ⊗ b) c d $== ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
?Goal0 ∘ ?Goal $== ?Goal9 ∘ ?Goal8 ∘ ?Goal4
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
α b c d $== ?Goal5
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
a ⊗R ?Goal5 $== ?Goal4
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
α a (b ⊗ c) d $== ?Goal8
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
α a b c $== ?Goal11
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
?Goal11 ⊗L d $== ?Goal9
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c)
∘ (twist (a ⊗ b) d c ∘ (a ⊗ b) ⊗R braid c d)
∘ (braid (c ⊗ d) (a ⊗ b)
∘ (twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d))) $==
?Goal2
∘ (braid d (a ⊗ (b ⊗ c))
∘ (twist a d (b ⊗ c) ∘ a ⊗R braid (b ⊗ c) d))
∘ ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
a
⊗R (braid d (b ⊗ c) ∘ (twist b d c ∘ b ⊗R braid c d)) $==
?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
(braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c))
⊗L d $== ?Goal2
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c)
∘ (twist (a ⊗ b) d c ∘ (a ⊗ b) ⊗R braid c d)
∘ (braid (c ⊗ d) (a ⊗ b)
∘ (twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d))) $==
?Goal0
∘ (braid d (a ⊗ (b ⊗ c))
∘ (twist a d (b ⊗ c) ∘ a ⊗R braid (b ⊗ c) d))
∘ (a ⊗R braid d (b ⊗ c)
∘ (a ⊗R twist b d c ∘ a ⊗R (b ⊗R braid c d)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
(braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c))
⊗L d $== ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c)
∘ (twist (a ⊗ b) d c ∘ (a ⊗ b) ⊗R braid c d)
∘ (braid (c ⊗ d) (a ⊗ b)
∘ (twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d))) $==
braid c (a ⊗ b) ⊗L d
∘ (twist a c b ⊗L d ∘ (a ⊗R braid b c) ⊗L d)
∘ (braid d (a ⊗ (b ⊗ c))
∘ (twist a d (b ⊗ c) ∘ a ⊗R braid (b ⊗ c) d))
∘ (a ⊗R braid d (b ⊗ c)
∘ (a ⊗R twist b d c ∘ a ⊗R (b ⊗R braid c d)))
(** We use a notation defined above that shows the base type of the groupoid hom and formats the equation in a way that is easier to read. *)(** Normalize brackets on LHS *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c)
∘ (twist (a ⊗ b) d c ∘ (a ⊗ b) ⊗R braid c d
∘ (braid (c ⊗ d) (a ⊗ b)
∘ (twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)))) $==
braid c (a ⊗ b) ⊗L d
∘ (twist a c b ⊗L d ∘ (a ⊗R braid b c) ⊗L d)
∘ (braid d (a ⊗ (b ⊗ c))
∘ (twist a d (b ⊗ c) ∘ a ⊗R braid (b ⊗ c) d))
∘ (a ⊗R braid d (b ⊗ c)
∘ (a ⊗R twist b d c ∘ a ⊗R (b ⊗R braid c d)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c)
∘ (twist (a ⊗ b) d c
∘ ((a ⊗ b) ⊗R braid c d
∘ (braid (c ⊗ d) (a ⊗ b)
∘ (twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d))))) $==
braid c (a ⊗ b) ⊗L d
∘ (twist a c b ⊗L d ∘ (a ⊗R braid b c) ⊗L d)
∘ (braid d (a ⊗ (b ⊗ c))
∘ (twist a d (b ⊗ c) ∘ a ⊗R braid (b ⊗ c) d))
∘ (a ⊗R braid d (b ⊗ c)
∘ (a ⊗R twist b d c ∘ a ⊗R (b ⊗R braid c d)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d) $==
braid c (a ⊗ b) ⊗L d
∘ (twist a c b ⊗L d ∘ (a ⊗R braid b c) ⊗L d)
∘ (braid d (a ⊗ (b ⊗ c))
∘ (twist a d (b ⊗ c) ∘ a ⊗R braid (b ⊗ c) d))
∘ (a ⊗R braid d (b ⊗ c)
∘ (a ⊗R twist b d c ∘ a ⊗R (b ⊗R braid c d)))
(** Normalize brackets on RHS *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d) $==
braid c (a ⊗ b) ⊗L d ∘ twist a c b ⊗L d
∘ (a ⊗R braid b c) ⊗L d
∘ (braid d (a ⊗ (b ⊗ c))
∘ (twist a d (b ⊗ c) ∘ a ⊗R braid (b ⊗ c) d))
∘ (a ⊗R braid d (b ⊗ c)
∘ (a ⊗R twist b d c ∘ a ⊗R (b ⊗R braid c d)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d) $==
braid c (a ⊗ b) ⊗L d ∘ twist a c b ⊗L d
∘ (a ⊗R braid b c) ⊗L d ∘ braid d (a ⊗ (b ⊗ c))
∘ twist a d (b ⊗ c) ∘ a ⊗R braid (b ⊗ c) d
∘ (a ⊗R braid d (b ⊗ c)
∘ (a ⊗R twist b d c ∘ a ⊗R (b ⊗R braid c d)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d) $==
braid c (a ⊗ b) ⊗L d ∘ twist a c b ⊗L d
∘ (a ⊗R braid b c) ⊗L d ∘ braid d (a ⊗ (b ⊗ c))
∘ twist a d (b ⊗ c) ∘ a ⊗R braid (b ⊗ c) d
∘ a ⊗R braid d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R (b ⊗R braid c d)
(** Cancel two braids next to each other. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) $== braid c (a ⊗ b) ⊗L d
∘ twist a c b ⊗L d ∘ (a ⊗R braid b c) ⊗L d
∘ braid d (a ⊗ (b ⊗ c)) ∘ twist a d (b ⊗ c)
∘ a ⊗R braid (b ⊗ c) d ∘ a ⊗R braid d (b ⊗ c)
∘ a ⊗R twist b d c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c $==
braid c (a ⊗ b) ⊗L d ∘ twist a c b ⊗L d
∘ (a ⊗R braid b c) ⊗L d ∘ braid d (a ⊗ (b ⊗ c))
∘ twist a d (b ⊗ c) ∘ a ⊗R braid (b ⊗ c) d
∘ a ⊗R braid d (b ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c $==
braid c (a ⊗ b) ⊗L d ∘ twist a c b ⊗L d
∘ (a ⊗R braid b c) ⊗L d ∘ braid d (a ⊗ (b ⊗ c))
∘ twist a d (b ⊗ c)
∘ (a ⊗R braid (b ⊗ c) d ∘ a ⊗R braid d (b ⊗ c))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c $==
braid c (a ⊗ b) ⊗L d ∘ twist a c b ⊗L d
∘ (a ⊗R braid b c) ⊗L d ∘ braid d (a ⊗ (b ⊗ c))
∘ twist a d (b ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
a ⊗R braid (b ⊗ c) d ∘
a ⊗R braid d (b ⊗ c) $==
Id (a ⊗ (d ⊗ (b ⊗ c)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c $==
braid c (a ⊗ b) ⊗L d ∘ twist a c b ⊗L d
∘ (a ⊗R braid b c) ⊗L d ∘ braid d (a ⊗ (b ⊗ c))
∘ twist a d (b ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid (b ⊗ c) d ∘ braid d (b ⊗ c) $== Id (d ⊗ (b ⊗ c))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c $==
braid c (a ⊗ b) ⊗L d ∘ twist a c b ⊗L d
∘ (a ⊗R braid b c) ⊗L d ∘ braid d (a ⊗ (b ⊗ c))
∘ twist a d (b ⊗ c)
(** *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) $== braid c (a ⊗ b) ⊗L d
∘ twist a c b ⊗L d ∘ (a ⊗R braid b c) ⊗L d
∘ braid d (a ⊗ (b ⊗ c))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) $== braid c (a ⊗ b) ⊗L d
∘ twist a c b ⊗L d
∘ ((a ⊗R braid b c) ⊗L d ∘ braid d (a ⊗ (b ⊗ c)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) $== braid c (a ⊗ b) ⊗L d
∘ twist a c b ⊗L d ∘ ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
?Goal0 $== (a ⊗R braid b c) ⊗L d
∘ braid d (a ⊗ (b ⊗ c))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) $== braid c (a ⊗ b) ⊗L d
∘ twist a c b ⊗L d
∘ (braid d (a ⊗ (c ⊗ b)) ∘ d ⊗R (a ⊗R braid b c))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) $== braid c (a ⊗ b) ⊗L d
∘ twist a c b ⊗L d ∘ braid d (a ⊗ (c ⊗ b))
∘ d ⊗R (a ⊗R braid b c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) ∘ d ⊗R (a ⊗R braid c b) $==
braid c (a ⊗ b) ⊗L d ∘ twist a c b ⊗L d
∘ braid d (a ⊗ (c ⊗ b))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) ∘ d ⊗R (a ⊗R braid c b) $==
braid c (a ⊗ b) ⊗L d
∘ (twist a c b ⊗L d ∘ braid d (a ⊗ (c ⊗ b)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) ∘ d ⊗R (a ⊗R braid c b) $==
braid c (a ⊗ b) ⊗L d ∘ ?Goal0
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
?Goal0 $== twist a c b ⊗L d ∘ braid d (a ⊗ (c ⊗ b))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) ∘ d ⊗R (a ⊗R braid c b) $==
braid c (a ⊗ b) ⊗L d
∘ (braid d (c ⊗ (a ⊗ b)) ∘ d ⊗R twist a c b)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) ∘ d ⊗R (a ⊗R braid c b) $==
braid c (a ⊗ b) ⊗L d ∘ braid d (c ⊗ (a ⊗ b))
∘ d ⊗R twist a c b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) ∘ d ⊗R (a ⊗R braid c b)
∘ d ⊗R twist c a b $== braid c (a ⊗ b) ⊗L d
∘ braid d (c ⊗ (a ⊗ b))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) ∘ d ⊗R (a ⊗R braid c b)
∘ d ⊗R twist c a b $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
?Goal $== braid c (a ⊗ b) ⊗L d ∘ braid d (c ⊗ (a ⊗ b))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) ∘ d ⊗R (a ⊗R braid c b)
∘ d ⊗R twist c a b $== braid d ((a ⊗ b) ⊗ c)
∘ d ⊗R braid c (a ⊗ b)
(** Putting things back. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) ∘ d ⊗R (a ⊗R braid c b) $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) $== braid d ((a ⊗ b) ⊗ c)
∘ d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) ∘ a ⊗R twist d b c $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R braid b (c ⊗ d)
∘ a ⊗R (b ⊗R braid d c) $== braid d ((a ⊗ b) ⊗ c)
∘ d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c) ∘ twist a d (b ⊗ c)
∘ a ⊗R twist b d c
(** There are two braids on the RHS of the LHS that can be swapped. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ (a ⊗R braid b (c ⊗ d) ∘ a ⊗R (b ⊗R braid d c)) $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
a ⊗R braid b (c ⊗ d) ∘ a ⊗R (b ⊗R braid d c) $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘
twist (a ⊗ b) d c ∘ (a ⊗ b) ⊗R braid c d
∘ braid (c ⊗ d) (a ⊗ b) ∘
twist a (c ⊗ d) b ∘ ?Goal $==
braid d ((a ⊗ b) ⊗ c) ∘
d ⊗R braid c (a ⊗ b) ∘
d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid b (c ⊗ d) ∘ b ⊗R braid d c $== ?Goal2 ∘ ?Goal3
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘
twist (a ⊗ b) d c ∘ (a ⊗ b) ⊗R braid c d
∘ braid (c ⊗ d) (a ⊗ b) ∘
twist a (c ⊗ d) b ∘ (a ⊗R ?Goal2 ∘ a ⊗R ?Goal3) $==
braid d ((a ⊗ b) ⊗ c) ∘
d ⊗R braid c (a ⊗ b) ∘
d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ (a ⊗R (braid d c ⊗L b) ∘ a ⊗R braid b (d ⊗ c)) $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R (braid d c ⊗L b)
∘ a ⊗R braid b (d ⊗ c) $== braid d ((a ⊗ b) ⊗ c)
∘ d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c) ∘ twist a d (b ⊗ c)
∘ a ⊗R twist b d c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b ∘ a ⊗R (braid d c ⊗L b) $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b
(** Naturality of twist on the RHS of the LHS. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ (twist a (c ⊗ d) b ∘ a ⊗R (braid d c ⊗L b)) $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
twist a (c ⊗ d) b ∘ a ⊗R (braid d c ⊗L b) $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘
twist (a ⊗ b) d c ∘ (a ⊗ b) ⊗R braid c d
∘ braid (c ⊗ d) (a ⊗ b) ∘
?Goal $== braid d ((a ⊗ b) ⊗ c) ∘
d ⊗R braid c (a ⊗ b) ∘
d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ (braid d c ⊗L (a ⊗ b) ∘ twist a (d ⊗ c) b) $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ braid d c ⊗L (a ⊗ b) ∘ twist a (d ⊗ c) b $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b
(** Moving some things to the RHS so that we can braid and cancel on the LHS. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ braid (c ⊗ d) (a ⊗ b)
∘ braid d c ⊗L (a ⊗ b) $== braid d ((a ⊗ b) ⊗ c)
∘ d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c) ∘ twist a d (b ⊗ c)
∘ a ⊗R twist b d c ∘ a ⊗R braid (d ⊗ c) b
∘ twist (d ⊗ c) a b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d
∘ (braid (c ⊗ d) (a ⊗ b) ∘ braid d c ⊗L (a ⊗ b)) $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b ∘ twist (d ⊗ c) a b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid (c ⊗ d) (a ⊗ b) ∘ braid d c ⊗L (a ⊗ b) $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘
twist (a ⊗ b) d c ∘ (a ⊗ b) ⊗R braid c d ∘
?Goal $== braid d ((a ⊗ b) ⊗ c) ∘
d ⊗R braid c (a ⊗ b) ∘
d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b ∘
twist (d ⊗ c) a b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d
∘ ((a ⊗ b) ⊗R braid d c ∘ braid (d ⊗ c) (a ⊗ b)) $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b ∘ twist (d ⊗ c) a b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ (a ⊗ b) ⊗R braid d c
∘ braid (d ⊗ c) (a ⊗ b) $== braid d ((a ⊗ b) ⊗ c)
∘ d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c) ∘ twist a d (b ⊗ c)
∘ a ⊗R twist b d c ∘ a ⊗R braid (d ⊗ c) b
∘ twist (d ⊗ c) a b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ (a ⊗ b) ⊗R braid c d ∘ (a ⊗ b) ⊗R braid d c $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b ∘ twist (d ⊗ c) a b
∘ braid (a ⊗ b) (d ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ ((a ⊗ b) ⊗R braid c d ∘ (a ⊗ b) ⊗R braid d c) $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b ∘ twist (d ⊗ c) a b
∘ braid (a ⊗ b) (d ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
(a ⊗ b) ⊗R braid c d ∘ (a ⊗ b) ⊗R braid d c $==
Id ((a ⊗ b) ⊗ (d ⊗ c))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘
twist (a ⊗ b) d c $== braid d ((a ⊗ b) ⊗ c)
∘ d ⊗R braid c (a ⊗ b) ∘
d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b ∘
twist (d ⊗ c) a b ∘ braid (a ⊗ b) (d ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid c d ∘ braid d c $== Id (d ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘
twist (a ⊗ b) d c $== braid d ((a ⊗ b) ⊗ c)
∘ d ⊗R braid c (a ⊗ b) ∘
d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b ∘
twist (d ⊗ c) a b ∘ braid (a ⊗ b) (d ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b ∘ twist (d ⊗ c) a b
∘ braid (a ⊗ b) (d ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ braid (d ⊗ c) (a ⊗ b) $== braid d ((a ⊗ b) ⊗ c)
∘ d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c) ∘ twist a d (b ⊗ c)
∘ a ⊗R twist b d c ∘ a ⊗R braid (d ⊗ c) b
∘ twist (d ⊗ c) a b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ braid (d ⊗ c) (a ⊗ b) ∘ twist a (d ⊗ c) b $==
braid d ((a ⊗ b) ⊗ c) ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
∘ a ⊗R braid (d ⊗ c) b
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ braid (d ⊗ c) (a ⊗ b) ∘ twist a (d ⊗ c) b
∘ a ⊗R braid b (d ⊗ c) $== braid d ((a ⊗ b) ⊗ c)
∘ d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c) ∘ twist a d (b ⊗ c)
∘ a ⊗R twist b d c
(** We are almost at the desired 9-gon. Now we cancel the inner braid on the LHS. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c) ∘ twist (a ⊗ b) d c
∘ braid (d ⊗ c) (a ⊗ b) ∘ twist a (d ⊗ c) b
∘ a ⊗R braid b (d ⊗ c) $== braid d ((a ⊗ b) ⊗ c)
∘ (d ⊗R braid c (a ⊗ b)
∘ (d ⊗R twist a c b
∘ (d ⊗R (a ⊗R braid b c)
∘ (twist a d (b ⊗ c) ∘ a ⊗R twist b d c))))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid d ((a ⊗ b) ⊗ c)
∘ (twist (a ⊗ b) d c
∘ (braid (d ⊗ c) (a ⊗ b)
∘ (twist a (d ⊗ c) b ∘ a ⊗R braid b (d ⊗ c)))) $==
braid d ((a ⊗ b) ⊗ c)
∘ (d ⊗R braid c (a ⊗ b)
∘ (d ⊗R twist a c b
∘ (d ⊗R (a ⊗R braid b c)
∘ (twist a d (b ⊗ c) ∘ a ⊗R twist b d c))))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
twist (a ⊗ b) d c
∘ (braid (d ⊗ c) (a ⊗ b)
∘ (twist a (d ⊗ c) b ∘ a ⊗R braid b (d ⊗ c))) $==
d ⊗R braid c (a ⊗ b)
∘ (d ⊗R twist a c b
∘ (d ⊗R (a ⊗R braid b c)
∘ (twist a d (b ⊗ c) ∘ a ⊗R twist b d c)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid (d ⊗ c) (a ⊗ b)
∘ (twist a (d ⊗ c) b ∘ a ⊗R braid b (d ⊗ c)) $==
twist d (a ⊗ b) c
∘ (d ⊗R braid c (a ⊗ b)
∘ (d ⊗R twist a c b
∘ (d ⊗R (a ⊗R braid b c)
∘ (twist a d (b ⊗ c) ∘ a ⊗R twist b d c))))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid (d ⊗ c) (a ⊗ b)
∘ (twist a (d ⊗ c) b ∘ a ⊗R braid b (d ⊗ c)) $==
twist d (a ⊗ b) c ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid (d ⊗ c) (a ⊗ b) ∘ twist a (d ⊗ c) b
∘ a ⊗R braid b (d ⊗ c) $== twist d (a ⊗ b) c
∘ d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c) ∘ twist a d (b ⊗ c)
∘ a ⊗R twist b d c
(** Now we move terms around in order to get a homotopy in [a ⊗ (b ⊗ (d ⊗ c)) $-> d ⊗ (c ⊗ (a ⊗ b))]. *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid (d ⊗ c) (a ⊗ b) ∘ twist a (d ⊗ c) b
∘ a ⊗R braid b (d ⊗ c) ∘ a ⊗R twist d b c $==
twist d (a ⊗ b) c ∘ d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid (d ⊗ c) (a ⊗ b) ∘ twist a (d ⊗ c) b
∘ a ⊗R braid b (d ⊗ c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) $== twist d (a ⊗ b) c
∘ d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid (d ⊗ c) (a ⊗ b) ∘ twist a (d ⊗ c) b
∘ a ⊗R braid b (d ⊗ c) ∘ a ⊗R twist d b c
∘ twist d a (b ⊗ c) $== twist d (a ⊗ b) c
∘ (d ⊗R braid c (a ⊗ b)
∘ (d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
braid (d ⊗ c) (a ⊗ b)
∘ (twist a (d ⊗ c) b
∘ (a ⊗R braid b (d ⊗ c)
∘ (a ⊗R twist d b c ∘ twist d a (b ⊗ c)))) $==
twist d (a ⊗ b) c
∘ (d ⊗R braid c (a ⊗ b)
∘ (d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
twist (a ⊗ b) d c
∘ (braid (d ⊗ c) (a ⊗ b)
∘ (twist a (d ⊗ c) b
∘ (a ⊗R braid b (d ⊗ c)
∘ (a ⊗R twist d b c ∘ twist d a (b ⊗ c))))) $==
d ⊗R braid c (a ⊗ b)
∘ (d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
twist (a ⊗ b) d c
∘ (braid (d ⊗ c) (a ⊗ b)
∘ (twist a (d ⊗ c) b
∘ (a ⊗R braid b (d ⊗ c)
∘ (a ⊗R twist d b c ∘ twist d a (b ⊗ c))))) $==
d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
twist (a ⊗ b) d c ∘ braid (d ⊗ c) (a ⊗ b)
∘ twist a (d ⊗ c) b ∘ a ⊗R braid b (d ⊗ c)
∘ a ⊗R twist d b c ∘ twist d a (b ⊗ c) $==
d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
twist (a ⊗ b) d c ∘ braid (d ⊗ c) (a ⊗ b)
∘ twist a (d ⊗ c) b ∘ a ⊗R braid b (d ⊗ c)
∘ a ⊗R twist d b c $== d ⊗R braid c (a ⊗ b)
∘ d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
twist (a ⊗ b) d c ∘ braid (d ⊗ c) (a ⊗ b)
∘ twist a (d ⊗ c) b ∘ a ⊗R braid b (d ⊗ c) $==
d ⊗R braid c (a ⊗ b) ∘ d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c) ∘ twist a d (b ⊗ c)
∘ a ⊗R twist b d c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
twist (a ⊗ b) d c ∘ braid (d ⊗ c) (a ⊗ b)
∘ twist a (d ⊗ c) b ∘ a ⊗R braid b (d ⊗ c) $==
d ⊗R braid c (a ⊗ b)
∘ (d ⊗R twist a c b
∘ (d ⊗R (a ⊗R braid b c)
∘ (twist a d (b ⊗ c) ∘ a ⊗R twist b d c)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
twist (a ⊗ b) d c
∘ (braid (d ⊗ c) (a ⊗ b)
∘ (twist a (d ⊗ c) b ∘ a ⊗R braid b (d ⊗ c))) $==
d ⊗R braid c (a ⊗ b)
∘ (d ⊗R twist a c b
∘ (d ⊗R (a ⊗R braid b c)
∘ (twist a d (b ⊗ c) ∘ a ⊗R twist b d c)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
d ⊗R braid (a ⊗ b) c
∘ (twist (a ⊗ b) d c
∘ (braid (d ⊗ c) (a ⊗ b)
∘ (twist a (d ⊗ c) b ∘ a ⊗R braid b (d ⊗ c)))) $==
d ⊗R twist a c b
∘ (d ⊗R (a ⊗R braid b c)
∘ (twist a d (b ⊗ c) ∘ a ⊗R twist b d c))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
d ⊗R braid (a ⊗ b) c
∘ (twist (a ⊗ b) d c
∘ (braid (d ⊗ c) (a ⊗ b)
∘ (twist a (d ⊗ c) b ∘ a ⊗R braid b (d ⊗ c)))) $==
d ⊗R twist a c b ∘ d ⊗R (a ⊗R braid b c)
∘ twist a d (b ⊗ c) ∘ a ⊗R twist b d c
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
d ⊗R braid (a ⊗ b) c ∘ twist (a ⊗ b) d c
∘ braid (d ⊗ c) (a ⊗ b) ∘ twist a (d ⊗ c) b
∘ a ⊗R braid b (d ⊗ c) $== d ⊗R twist a c b
∘ d ⊗R (a ⊗R braid b c) ∘ twist a d (b ⊗ c)
∘ a ⊗R twist b d c
(** And finally, this is the 9-gon we asked for. *)apply twist_9_gon.Defined.LocalClose Scope long_path_scope.(** *** Hexagon *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
HexagonIdentity cat_tensor
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d
HexagonIdentity cat_tensor
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid b a ⊗L c ∘ α b a c ∘ b ⊗R braid c a $== α a b c
∘ braid (b ⊗ c) a ∘ α b c a
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
α b a c $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid b a ⊗L c ∘ ?Goal ∘
b ⊗R braid c a $== ?Goal4 ∘
braid (b ⊗ c) a ∘ ?Goal2
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
α b c a $== ?Goal2
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
α a b c $== ?Goal4
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid b a ⊗L c
∘ (braid c (b ⊗ a) ∘ (twist b c a ∘ b ⊗R braid a c))
∘ b ⊗R braid c a $== braid c (a ⊗ b)
∘ (twist a c b ∘ a ⊗R braid b c) ∘ braid (b ⊗ c) a
∘ (braid a (b ⊗ c) ∘ (twist b a c ∘ b ⊗R braid c a))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid b a ⊗L c ∘ braid c (b ⊗ a) ∘ twist b c a
∘ b ⊗R braid a c ∘ b ⊗R braid c a $== braid c (a ⊗ b)
∘ (twist a c b ∘ a ⊗R braid b c) ∘ braid (b ⊗ c) a
∘ (braid a (b ⊗ c) ∘ (twist b a c ∘ b ⊗R braid c a))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
b ⊗R braid a c ∘ b ⊗R braid c a $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid b a ⊗L c ∘ braid c (b ⊗ a) ∘
twist b c a ∘ ?Goal $==
braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c)
∘ braid (b ⊗ c) a
∘ (braid a (b ⊗ c) ∘ (twist b a c ∘ b ⊗R braid c a))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
b ⊗R braid a c ∘ b ⊗R braid c a $== ?Goal
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid a c ∘ braid c a $== Id (c ⊗ a)
apply braid_braid.
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid b a ⊗L c ∘ braid c (b ⊗ a) ∘ twist b c a
∘ Id (b ⊗ (c ⊗ a)) $== braid c (a ⊗ b)
∘ (twist a c b ∘ a ⊗R braid b c) ∘ braid (b ⊗ c) a
∘ (braid a (b ⊗ c) ∘ (twist b a c ∘ b ⊗R braid c a))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid b a ⊗L c ∘ braid c (b ⊗ a) ∘ twist b c a $==
braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c)
∘ braid (b ⊗ c) a
∘ (braid a (b ⊗ c) ∘ (twist b a c ∘ b ⊗R braid c a))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid b a ⊗L c ∘ braid c (b ⊗ a) ∘ twist b c a $==
braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c)
∘ braid (b ⊗ c) a ∘ braid a (b ⊗ c)
∘ (twist b a c ∘ b ⊗R braid c a)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid b a ⊗L c ∘ braid c (b ⊗ a) ∘ twist b c a $==
braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c)
∘ (braid (b ⊗ c) a ∘ braid a (b ⊗ c))
∘ (twist b a c ∘ b ⊗R braid c a)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid b a ⊗L c ∘ braid c (b ⊗ a) ∘ twist b c a $==
braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c)
∘ (twist b a c ∘ b ⊗R braid c a)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid (b ⊗ c) a ∘ braid a (b ⊗ c) $-> Id (a ⊗ (b ⊗ c))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid b a ⊗L c ∘ braid c (b ⊗ a) ∘ twist b c a $==
braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c)
∘ (twist b a c ∘ b ⊗R braid c a)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
braid c (a ⊗ b) ∘ c ⊗R braid b a ∘ twist b c a $==
braid c (a ⊗ b) ∘ (twist a c b ∘ a ⊗R braid b c)
∘ (twist b a c ∘ b ⊗R braid c a)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
c ⊗R braid b a ∘ twist b c a $== twist a c b
∘ a ⊗R braid b c ∘ (twist b a c ∘ b ⊗R braid c a)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
c ⊗R braid b a ∘ twist b c a $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c ∘ b ⊗R braid c a
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c: A
c ⊗R braid b a ∘ twist b c a ∘ b ⊗R braid a c $==
twist a c b ∘ a ⊗R braid b c ∘ twist b a c
apply twist_hexagon.Defined.(** *** Conclusion *)(** In conclusion, we have proven the following: *)(** There is a monoidal structure on [A]. *)Instanceismonoidal_twist
: IsMonoidal A cat_tensor cat_tensor_unit
:= {}.(** There is a symmetric monoidal category on [A]. *)Instanceissymmetricmonoidal_twist
: IsSymmetricMonoidal A cat_tensor cat_tensor_unit
:= {}.(** TODO: WIP *)(** Here is a hexagon involving only twist *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d $== twist b c (a ⊗ d)
∘ b ⊗R twist a c d ∘ twist a b (c ⊗ d)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d $== twist b c (a ⊗ d)
∘ b ⊗R twist a c d ∘ twist a b (c ⊗ d)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A p: d ⊗R braid a c ∘ twist a d c ∘
a ⊗R braid c d $== twist c d a
∘ (c ⊗R braid a d ∘ twist a c d)
c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d $== twist b c (a ⊗ d)
∘ b ⊗R twist a c d ∘ twist a b (c ⊗ d)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A p: twist d c a
∘ (d ⊗R braid a c ∘ twist a d c ∘ a ⊗R braid c d) $==
c ⊗R braid a d ∘ twist a c d
c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d $== twist b c (a ⊗ d)
∘ b ⊗R twist a c d ∘ twist a b (c ⊗ d)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A p: c ⊗R braid d a
∘ (twist d c a
∘ (d ⊗R braid a c ∘
twist a d c ∘ a ⊗R braid c d)) $==
twist a c d
c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d $== twist b c (a ⊗ d)
∘ b ⊗R twist a c d ∘ twist a b (c ⊗ d)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A p: b
⊗R (c ⊗R braid d a
∘ (twist d c a
∘ (d ⊗R braid a c ∘
twist a d c ∘
a ⊗R braid c d))) $==
b ⊗R twist a c d
c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d $== twist b c (a ⊗ d)
∘ b ⊗R twist a c d ∘ twist a b (c ⊗ d)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d $== twist b c (a ⊗ d)
∘ b
⊗R (c ⊗R braid d a
∘ (twist d c a
∘ (d ⊗R braid a c ∘ twist a d c
∘ a ⊗R braid c d))) ∘ twist a b (c ⊗ d)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d ∘ twist b a (c ⊗ d) $==
twist b c (a ⊗ d)
∘ b
⊗R (c ⊗R braid d a
∘ (twist d c a
∘ (d ⊗R braid a c ∘ twist a d c
∘ a ⊗R braid c d)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
twist c b (a ⊗ d)
∘ (c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d ∘ twist b a (c ⊗ d)) $==
b
⊗R (c ⊗R braid d a
∘ (twist d c a
∘ (d ⊗R braid a c ∘ twist a d c
∘ a ⊗R braid c d)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
twist c b (a ⊗ d)
∘ (c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d ∘ twist b a (c ⊗ d)) $==
b ⊗R (c ⊗R braid d a)
∘ b
⊗R (twist d c a
∘ (d ⊗R braid a c ∘ twist a d c
∘ a ⊗R braid c d))
(** TODO simplify *)
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d ∘ twist b a (c ⊗ d) $==
twist b c (a ⊗ d)
∘ (b ⊗R (c ⊗R braid d a)
∘ b
⊗R (twist d c a
∘ (d ⊗R braid a c ∘ twist a d c
∘ a ⊗R braid c d)))
A: Type IsGraph0: IsGraph A Is2Graph0: Is2Graph A Is01Cat0: Is01Cat A H: Is1Cat A H0: HasEquivs A cat_tensor: A -> A -> A cat_tensor_unit: A Is0Bifunctor0: Is0Bifunctor cat_tensor Is1Bifunctor0: Is1Bifunctor cat_tensor braid: SymmetricBraiding cat_tensor twist: forallabc : A, a ⊗ (b ⊗ c) $-> b ⊗ (a ⊗ c) twist_twist: forallabc : A,
twist a b c ∘ twist b a c $==
Id (b ⊗ (a ⊗ c)) twist_nat: forall (aa'bb'cc' : A)
(f : a $-> a') (g : b $-> b')
(h : c $-> c'),
twist a' b' c'
∘ fmap11 cat_tensor f
(fmap11 cat_tensor g h) $==
fmap11 cat_tensor g
(fmap11 cat_tensor f h) ∘
twist a b c right_unitor: RightUnitor cat_tensor cat_tensor_unit twist_unitor: forallab : A,
a ⊗R right_unitor b $==
braid b a ∘ b ⊗R right_unitor a
∘ twist a b cat_tensor_unit twist_hexagon: forallabc : A,
c ⊗R braid b a ∘ twist b c a
∘ b ⊗R braid a c $== twist a c b
∘ a ⊗R braid b c ∘ twist b a c twist_9_gon: forallabcd : A,
c ⊗R braid (a ⊗ b) d ∘
twist (a ⊗ b) c d
∘ braid (c ⊗ d) (a ⊗ b)
∘ twist a (c ⊗ d) b
∘ a ⊗R braid b (c ⊗ d) $==
c ⊗R twist a d b ∘ c ⊗R (a ⊗R braid b d)
∘ twist a c (b ⊗ d) ∘ a ⊗R twist b c d a, b, c, d: A
c ⊗R twist a b d ∘ twist a c (b ⊗ d)
∘ a ⊗R twist b c d ∘ twist b a (c ⊗ d) $==
twist b c (a ⊗ d) ∘ b ⊗R (c ⊗R braid d a)
∘ b
⊗R (twist d c a
∘ (d ⊗R braid a c ∘ twist a d c
∘ a ⊗R braid c d))