Datatypes for bicategory like structures #
This file defines the basic datatypes for bicategory like structures. We will use these datatypes to write tactics that can be applied to both monoidal categories and bicategories:
Obj
: objects typeAtom₁
: atomic 1-morphisms typeMor₁
: 1-morphisms typeAtom
: atomic non-structural 2-morphisms typeMor₂
: 2-morphisms typeAtomIso
: atomic non-structural 2-isomorphisms typeMor₂Iso
: 2-isomorphisms typeNormalizedHom
: normalized 1-morphisms type
A term of these datatypes wraps the corresponding Expr
term, which can be extracted by
e.g. η.e
for η : Mor₂
.
The operations of these datatypes are defined in a monad m
with the corresponding typeclasses:
For example, a monad m
with [MonadMor₂ m]
provides the operation
MonadMor₂.comp₂M : Mor₂Iso → Mor₂Iso → m Mor₂Iso
, which constructs the expression for the
composition η ≫ θ
of 2-morphisms η
and θ
in the monad m
.
Equations
- Mathlib.Tactic.BicategoryLike.instInhabitedObj = { default := { e? := default } }
Equations
- Mathlib.Tactic.BicategoryLike.instInhabitedAtom₁ = { default := { e := default, src := default, tgt := default } }
The underlying lean expression of a 1-morphism.
Equations
- (Mathlib.Tactic.BicategoryLike.Mor₁.id e a).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₁.comp e a a_1).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₁.of a).e = a.e
Instances For
The domain of a 1-morphism.
Equations
- (Mathlib.Tactic.BicategoryLike.Mor₁.id e a).src = a
- (Mathlib.Tactic.BicategoryLike.Mor₁.comp e a a_1).src = a.src
- (Mathlib.Tactic.BicategoryLike.Mor₁.of a).src = a.src
Instances For
The codomain of a 1-morphism.
Equations
- (Mathlib.Tactic.BicategoryLike.Mor₁.id e a).tgt = a
- (Mathlib.Tactic.BicategoryLike.Mor₁.comp e a a_1).tgt = a_1.tgt
- (Mathlib.Tactic.BicategoryLike.Mor₁.of a).tgt = a.tgt
Instances For
Converts a 1-morphism into a list of its components.
Equations
Instances For
Expressions for coherence isomorphisms (i.e., structural 2-morphisms
giveb by BicategorycalCoherence.iso
).
- e : Lean.Expr
The underlying lean expression of a coherence isomorphism.
- src : Mor₁
The domain of a coherence isomorphism.
- tgt : Mor₁
The codomain of a coherence isomorphism.
- inst : Lean.Expr
The
BicategoricalCoherence
instance. - unfold : Lean.Expr
Extract the structural 2-isomorphism.
Instances For
Equations
- Mathlib.Tactic.BicategoryLike.instInhabitedAtomIso = { default := { e := default, src := default, tgt := default } }
Expressions for atomic structural 2-morphisms.
- associator
(e : Lean.Expr)
(f g h : Mor₁)
: StructuralAtom
The expression for the associator
α_ f g h
. - leftUnitor
(e : Lean.Expr)
(f : Mor₁)
: StructuralAtom
The expression for the left unitor
λ_ f
. - rightUnitor
(e : Lean.Expr)
(f : Mor₁)
: StructuralAtom
The expression for the right unitor
ρ_ f
. - id (e : Lean.Expr) (f : Mor₁) : StructuralAtom
- coherenceHom (α : CoherenceHom) : StructuralAtom
Instances For
Expressions for 2-isomorphisms.
- structuralAtom (α : StructuralAtom) : Mor₂Iso
- comp (e : Lean.Expr) (f g h : Mor₁) (η θ : Mor₂Iso) : Mor₂Iso
- whiskerLeft (e : Lean.Expr) (f g h : Mor₁) (η : Mor₂Iso) : Mor₂Iso
- whiskerRight (e : Lean.Expr) (f g : Mor₁) (η : Mor₂Iso) (h : Mor₁) : Mor₂Iso
- horizontalComp (e : Lean.Expr) (f₁ g₁ f₂ g₂ : Mor₁) (η θ : Mor₂Iso) : Mor₂Iso
- inv (e : Lean.Expr) (f g : Mor₁) (η : Mor₂Iso) : Mor₂Iso
- coherenceComp (e : Lean.Expr) (f g h i : Mor₁) (α : CoherenceHom) (η θ : Mor₂Iso) : Mor₂Iso
- of (η : AtomIso) : Mor₂Iso
Instances For
A monad equipped with the ability to unfold BicategoricalCoherence.iso
.
- unfoldM (α : CoherenceHom) : m Mor₂Iso
Unfold a coherence isomorphism.
Instances
The underlying lean expression of a 2-isomorphism.
Equations
- (Mathlib.Tactic.BicategoryLike.StructuralAtom.associator e f g h).e = e
- (Mathlib.Tactic.BicategoryLike.StructuralAtom.leftUnitor e f).e = e
- (Mathlib.Tactic.BicategoryLike.StructuralAtom.rightUnitor e f).e = e
- (Mathlib.Tactic.BicategoryLike.StructuralAtom.id e f).e = e
- (Mathlib.Tactic.BicategoryLike.StructuralAtom.coherenceHom α).e = α.e
Instances For
The domain of a 2-isomorphism.
Equations
- One or more equations did not get rendered due to their size.
- (Mathlib.Tactic.BicategoryLike.StructuralAtom.id e f).srcM = pure f
- (Mathlib.Tactic.BicategoryLike.StructuralAtom.coherenceHom α).srcM = pure α.src
Instances For
The codomain of a 2-isomorphism.
Equations
- One or more equations did not get rendered due to their size.
- (Mathlib.Tactic.BicategoryLike.StructuralAtom.leftUnitor e f).tgtM = pure f
- (Mathlib.Tactic.BicategoryLike.StructuralAtom.rightUnitor e f).tgtM = pure f
- (Mathlib.Tactic.BicategoryLike.StructuralAtom.id e f).tgtM = pure f
- (Mathlib.Tactic.BicategoryLike.StructuralAtom.coherenceHom α).tgtM = pure α.tgt
Instances For
The underlying lean expression of a 2-isomorphism.
Equations
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.structuralAtom α).e = α.e
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.comp e f g h η θ).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.whiskerLeft e f g h η).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.whiskerRight e f g η h).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.horizontalComp e f₁ g₁ f₂ g₂ η θ).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.inv e f g η).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.coherenceComp e f g h i α η θ).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.of η).e = η.e
Instances For
The domain of a 2-isomorphism.
Equations
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.structuralAtom α).srcM = α.srcM
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.comp e f g h η θ).srcM = pure f
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.whiskerLeft e f g h η).srcM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M f g
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.whiskerRight e f g η h).srcM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M f h
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.horizontalComp e f₁ g₁ f₂ g₂ η θ).srcM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M f₁ f₂
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.inv e f g η).srcM = pure g
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.coherenceComp e f g h i α η θ).srcM = pure f
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.of η).srcM = pure η.src
Instances For
The codomain of a 2-isomorphism.
Equations
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.structuralAtom α).tgtM = α.tgtM
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.comp e f g h η θ).tgtM = pure h
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.whiskerLeft e f g h η).tgtM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M f h
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.whiskerRight e f g η h).tgtM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M g h
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.horizontalComp e f₁ g₁ f₂ g₂ η θ).tgtM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M g₁ g₂
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.inv e f g η).tgtM = pure f
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.coherenceComp e f g h i α η θ).tgtM = pure i
- (Mathlib.Tactic.BicategoryLike.Mor₂Iso.of η).tgtM = pure η.tgt
Instances For
A monad equipped with the ability to construct Mor₂Iso
terms.
- associatorM (f g h : Mor₁) : m StructuralAtom
The expression for the associator
α_ f g h
. - leftUnitorM (f : Mor₁) : m StructuralAtom
The expression for the left unitor
λ_ f
. - rightUnitorM (f : Mor₁) : m StructuralAtom
The expression for the right unitor
ρ_ f
. - id₂M (f : Mor₁) : m StructuralAtom
The expression for the identity
Iso.refl f
. The expression for the coherence isomorphism
⊗𝟙 : f ⟶ g
.The expression for the composition
η ≪≫ θ
.The expression for the left whiskering
whiskerLeftIso f η
.The expression for the right whiskering
whiskerRightIso η h
.The expression for the horizontal composition
η ◫ θ
.The expression for the inverse
Iso.symm η
.The expression for the coherence composition
η ≪⊗≫ θ := η ≪≫ α ≪≫ θ
.
Instances
The expression for the associator α_ f g h
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The expression for the left unitor λ_ f
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The expression for the right unitor ρ_ f
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The expression for the identity Iso.refl f
.
Equations
- Mathlib.Tactic.BicategoryLike.MonadMor₂Iso.id₂M' f = do let __do_lift ← Mathlib.Tactic.BicategoryLike.MonadMor₂Iso.id₂M f pure (Mathlib.Tactic.BicategoryLike.Mor₂Iso.structuralAtom __do_lift)
Instances For
The expression for the coherence isomorphism ⊗𝟙 : f ⟶ g
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Mathlib.Tactic.BicategoryLike.instInhabitedAtom = { default := { e := default, src := default, tgt := default } }
Mor₂
expressions defined below will have the isoLift? : Option IsoLift
field.
For η : Mor₂
such that η.isoLift? = .some isoLift
, we have the following data:
isoLift.e
: an expression for a 2-isomorphismη'
, given as aMor₂Iso
term,isoLift.eq
: a lean expression for the proof thatη'.hom = η
.
- e : Mor₂Iso
The expression for the 2-isomorphism.
- eq : Lean.Expr
The expression for the proof that the forward direction of the 2-isomorphism is equal to the original 2-morphism.
Instances For
Expressions for 2-morphisms.
- isoHom
(e : Lean.Expr)
(isoLift : IsoLift)
(iso : Mor₂Iso)
: Mor₂
The expression for
Iso.hom
. - isoInv
(e : Lean.Expr)
(isoLift : IsoLift)
(iso : Mor₂Iso)
: Mor₂
The expression for
Iso.inv
. - id
(e : Lean.Expr)
(isoLift : IsoLift)
(f : Mor₁)
: Mor₂
The expression for the identity
𝟙 f
. - comp
(e : Lean.Expr)
(isoLift? : Option IsoLift)
(f g h : Mor₁)
(η θ : Mor₂)
: Mor₂
The expression for the composition
η ≫ θ
. - whiskerLeft
(e : Lean.Expr)
(isoLift? : Option IsoLift)
(f g h : Mor₁)
(η : Mor₂)
: Mor₂
The expression for the left whiskering
f ◁ η
withη : g ⟶ h
. - whiskerRight
(e : Lean.Expr)
(isoLift? : Option IsoLift)
(f g : Mor₁)
(η : Mor₂)
(h : Mor₁)
: Mor₂
The expression for the right whiskering
η ▷ h
withη : f ⟶ g
. - horizontalComp
(e : Lean.Expr)
(isoLift? : Option IsoLift)
(f₁ g₁ f₂ g₂ : Mor₁)
(η θ : Mor₂)
: Mor₂
The expression for the horizontal composition
η ◫ θ
withη : f₁ ⟶ g₁
andθ : f₂ ⟶ g₂
. - coherenceComp
(e : Lean.Expr)
(isoLift? : Option IsoLift)
(f g h i : Mor₁)
(α : CoherenceHom)
(η θ : Mor₂)
: Mor₂
The expression for the coherence composition
η ⊗≫ θ := η ≫ α ≫ θ
withη : f ⟶ g
andθ : h ⟶ i
. - of
(η : Atom)
: Mor₂
The expression for an atomic non-structural 2-morphism.
Instances For
The underlying lean expression of a 2-morphism.
Equations
- (Mathlib.Tactic.BicategoryLike.Mor₂.isoHom e isoLift iso).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂.isoInv e isoLift iso).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂.id e isoLift f).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂.comp e isoLift? f g h η θ).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂.whiskerLeft e isoLift? f g h η).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂.whiskerRight e isoLift? f g η h).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂.horizontalComp e isoLift? f₁ g₁ f₂ g₂ η θ).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂.coherenceComp e isoLift? f g h i α η θ).e = e
- (Mathlib.Tactic.BicategoryLike.Mor₂.of η).e = η.e
Instances For
η.isoLift?
is a pair of a 2-isomorphism η'
and a proof that η'.hom = η
. If no such η'
is found, returns none
. This function does not seek IsIso
instance.
Equations
- (Mathlib.Tactic.BicategoryLike.Mor₂.isoHom e isoLift iso).isoLift? = some isoLift
- (Mathlib.Tactic.BicategoryLike.Mor₂.isoInv e isoLift iso).isoLift? = some isoLift
- (Mathlib.Tactic.BicategoryLike.Mor₂.id e isoLift f).isoLift? = some isoLift
- (Mathlib.Tactic.BicategoryLike.Mor₂.comp e isoLift? f g h η θ).isoLift? = isoLift?
- (Mathlib.Tactic.BicategoryLike.Mor₂.whiskerLeft e isoLift? f g h η).isoLift? = isoLift?
- (Mathlib.Tactic.BicategoryLike.Mor₂.whiskerRight e isoLift? f g η h).isoLift? = isoLift?
- (Mathlib.Tactic.BicategoryLike.Mor₂.horizontalComp e isoLift? f₁ g₁ f₂ g₂ η θ).isoLift? = isoLift?
- (Mathlib.Tactic.BicategoryLike.Mor₂.coherenceComp e isoLift? f g h i α η θ).isoLift? = isoLift?
- (Mathlib.Tactic.BicategoryLike.Mor₂.of η).isoLift? = none
Instances For
The domain of a 2-morphism.
Equations
- (Mathlib.Tactic.BicategoryLike.Mor₂.isoHom e isoLift iso).srcM = iso.srcM
- (Mathlib.Tactic.BicategoryLike.Mor₂.isoInv e isoLift iso).srcM = iso.tgtM
- (Mathlib.Tactic.BicategoryLike.Mor₂.id e isoLift f).srcM = pure f
- (Mathlib.Tactic.BicategoryLike.Mor₂.comp e isoLift? f g h η θ).srcM = pure f
- (Mathlib.Tactic.BicategoryLike.Mor₂.whiskerLeft e isoLift? f g h η).srcM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M f g
- (Mathlib.Tactic.BicategoryLike.Mor₂.whiskerRight e isoLift? f g η h).srcM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M f h
- (Mathlib.Tactic.BicategoryLike.Mor₂.horizontalComp e isoLift? f₁ g₁ f₂ g₂ η θ).srcM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M f₁ f₂
- (Mathlib.Tactic.BicategoryLike.Mor₂.coherenceComp e isoLift? f g h i α η θ).srcM = pure f
- (Mathlib.Tactic.BicategoryLike.Mor₂.of η).srcM = pure η.src
Instances For
The codomain of a 2-morphism.
Equations
- (Mathlib.Tactic.BicategoryLike.Mor₂.isoHom e isoLift iso).tgtM = iso.tgtM
- (Mathlib.Tactic.BicategoryLike.Mor₂.isoInv e isoLift iso).tgtM = iso.srcM
- (Mathlib.Tactic.BicategoryLike.Mor₂.id e isoLift f).tgtM = pure f
- (Mathlib.Tactic.BicategoryLike.Mor₂.comp e isoLift? f g h η θ).tgtM = pure h
- (Mathlib.Tactic.BicategoryLike.Mor₂.whiskerLeft e isoLift? f g h η).tgtM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M f h
- (Mathlib.Tactic.BicategoryLike.Mor₂.whiskerRight e isoLift? f g η h).tgtM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M g h
- (Mathlib.Tactic.BicategoryLike.Mor₂.horizontalComp e isoLift? f₁ g₁ f₂ g₂ η θ).tgtM = Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M g₁ g₂
- (Mathlib.Tactic.BicategoryLike.Mor₂.coherenceComp e isoLift? f g h i α η θ).tgtM = pure i
- (Mathlib.Tactic.BicategoryLike.Mor₂.of η).tgtM = pure η.tgt
Instances For
A monad equipped with the ability to manipulate 2-morphisms.
The expression for
Iso.hom η
.The expression for
Iso.hom η
.The expression for
Iso.inv η
.The expression for
Iso.inv η
.The expression for the identity
𝟙 f
.The expression for the composition
η ≫ θ
.The expression for the left whiskering
f ◁ η
.The expression for the right whiskering
η ▷ h
.The expression for the horizontal composition
η ◫ θ
.The expression for the coherence composition
η ⊗≫ θ := η ≫ α ≫ θ
.
Instances
Type of normalized 1-morphisms ((... ≫ h) ≫ g) ≫ f
.
- nil
(e : Mor₁)
(a : Obj)
: NormalizedHom
The identity 1-morphism
𝟙 a
. - cons
(e : Mor₁)
: NormalizedHom → Atom₁ → NormalizedHom
The
cons
composes an atomic 1-morphism at the end of a normalized 1-morphism.
Instances For
The underlying expression of a normalized 1-morphism.
Equations
Instances For
The domain of a normalized 1-morphism.
Equations
Instances For
The codomain of a normalized 1-morphism.
Equations
- (Mathlib.Tactic.BicategoryLike.NormalizedHom.nil e a).tgt = a
- (Mathlib.Tactic.BicategoryLike.NormalizedHom.cons e a a_1).tgt = a_1.tgt
Instances For
Construct the NormalizedHom.nil
term in m
.
Equations
- Mathlib.Tactic.BicategoryLike.normalizedHom.nilM a = do let __do_lift ← Mathlib.Tactic.BicategoryLike.MonadMor₁.id₁M a pure (Mathlib.Tactic.BicategoryLike.NormalizedHom.nil __do_lift a)
Instances For
Construct a NormalizedHom.cons
term in m
.
Equations
- p.consM f = do let __do_lift ← Mathlib.Tactic.BicategoryLike.MonadMor₁.comp₁M p.e (Mathlib.Tactic.BicategoryLike.Mor₁.of f) pure (Mathlib.Tactic.BicategoryLike.NormalizedHom.cons __do_lift p f)
Instances For
Context ρ
provides the context for manipulating 2-morphisms in a monoidal category or
bicategory. In particular, we will store MonoidalCategory
or Bicategory
instance in a context,
and use this through a reader monad when we construct the lean expressions for 2-morphisms.
- mkContext? : Lean.Expr → Lean.MetaM (Option ρ)
Construct a context from a lean expression for a 2-morphism.
Instances
Construct a context from a lean expression for a 2-morphism.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The state for the CoherenceM ρ
monad.
- cache : Lean.PersistentExprMap Mor₁
The cache for evaluating lean expressions of 1-morphisms into
Mor₁
terms.
Instances For
The monad for manipulating 2-morphisms in a monoidal category or bicategory.
Equations
Instances For
Run the CoherenceM ρ
monad.
Equations
- x.run ctx s = Prod.fst <$> ReaderT.run x ctx s