Expressions for monoidal categories #
This file converts lean expressions representing morphisms in monoidal categories into Mor₂Iso
or Mor
terms. The converted expressions are used in the coherence tactics and the string diagram
widgets.
The domain of a morphism.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The codomain of a morphism.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The domain of an isomorphism.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The codomain of an isomorphism.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The context for evaluating expressions.
- level₂ : Lean.Level
The level for morphisms.
- level₁ : Lean.Level
The level for objects.
- C : Q(Type udummy._uniq.15)
The expression for the underlying category.
- instCat : Q(CategoryTheory.Category.{udummy._uniq.16, udummy._uniq.17} unknown_1)
The category instance.
- instMonoidal? : Option Q(CategoryTheory.MonoidalCategory unknown_1)
The monoidal category instance.
Instances For
Populate a context
object for evaluating e
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
The monad for the normalization of 2-morphisms.
Equations
Instances For
Throw an error if the monoidal category instance is not found.
Equations
- Mathlib.Tactic.Monoidal.synthMonoidalError = Lean.throwError (Lean.toMessageData "failed to find monoidal category instance")
Instances For
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Check that e
is definitionally equal to 𝟙_ C
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Return (f, g)
if e
is definitionally equal to f ⊗ g
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Construct a Mor₁
expression from a Lean expression.
Construct a Mor₂Iso
term from a Lean expression.
Construct a Mor₂
term from a Lean expression.