Documentation

Mathlib.Tactic.CategoryTheory.Coherence

A coherence tactic for monoidal categories, and ⊗≫ (composition up to associators) #

We provide a coherence tactic, which proves equations where the two sides differ by replacing strings of monoidal structural morphisms with other such strings. (The replacements are always equalities by the monoidal coherence theorem.)

A simpler version of this tactic is pure_coherence, which proves that any two morphisms (with the same source and target) in a monoidal category which are built out of associators and unitors are equal.

We also provide f ⊗≫ g, the monoidal_comp operation, which automatically inserts associators and unitors as needed to make the target of f match the source of g.

class Mathlib.Tactic.Coherence.LiftObj {C : Type u} (X : C) :

A typeclass carrying a choice of lift of an object from C to FreeMonoidalCategory C. It must be the case that projectObj id (LiftObj.lift x) = x by defeq.

Instances

    A typeclass carrying a choice of lift of a morphism from C to FreeMonoidalCategory C. It must be the case that projectMap id _ _ (LiftHom.lift f) = f by defeq.

    Instances

      A typeclass carrying a choice of monoidal structural isomorphism between two objects. Used by the ⊗≫ monoidal composition operator, and the coherence tactic.

      Instances

        Construct an isomorphism between two objects in a monoidal category out of unitors and associators.

        Instances For

          Compose two morphisms in a monoidal category, inserting unitors and associators between as necessary.

          Instances For

            Compose two morphisms in a monoidal category, inserting unitors and associators between as necessary.

            Instances For

              Compose two isomorphisms in a monoidal category, inserting unitors and associators between as necessary.

              Instances For

                Compose two isomorphisms in a monoidal category, inserting unitors and associators between as necessary.

                Instances For

                  Helper function for throwing exceptions.

                  Instances For

                    Helper function for throwing exceptions with respect to the main goal.

                    Instances For

                      Coherence tactic for monoidal categories.

                      Instances For

                        Coherence tactic for monoidal categories. Use pure_coherence instead, which is a frontend to this one.

                        Instances For

                          pure_coherence uses the coherence theorem for monoidal categories to prove the goal. It can prove any equality made up only of associators, unitors, and identities.

                          example {C : Type} [Category C] [MonoidalCategory C] :
                            (λ_ (𝟙_ C)).hom = (ρ_ (𝟙_ C)).hom :=
                          by pure_coherence
                          

                          Users will typically just use the coherence tactic, which can also cope with identities of the form a ≫ f ≫ b ≫ g ≫ c = a' ≫ f ≫ b' ≫ g ≫ c' where a = a', b = b', and c = c' can be proved using pure_coherence

                          Instances For

                            Auxiliary simp lemma for the coherence tactic: this moves brackets to the left in order to expose a maximal prefix built out of unitors and associators.

                            Internal tactic used in coherence.

                            Rewrites an equation f = g as f₀ ≫ f₁ = g₀ ≫ g₁, where f₀ and g₀ are maximal prefixes of f and g (possibly after reassociating) which are "liftable" (i.e. expressible as compositions of unitors and associators).

                            Instances For

                              If either the lhs or rhs is not a composition, compose it on the right with an identity.

                              Instances For

                                The main part of coherence tactic.

                                Equations
                                Instances For

                                  Use the coherence theorem for monoidal categories to solve equations in a monoidal equation, where the two sides only differ by replacing strings of monoidal structural morphisms (that is, associators, unitors, and identities) with different strings of structural morphisms with the same source and target.

                                  That is, coherence can handle goals of the form a ≫ f ≫ b ≫ g ≫ c = a' ≫ f ≫ b' ≫ g ≫ c' where a = a', b = b', and c = c' can be proved using pure_coherence.

                                  (If you have very large equations on which coherence is unexpectedly failing, you may need to increase the typeclass search depth, using e.g. set_option synthInstance.maxSize 500.)

                                  Instances For