Documentation

Mathlib.CategoryTheory.Monoidal.ExternalProduct.KanExtension

Preservation of pointwise left Kan extensions by external products #

We prove that if a functor H' : D' ⥤ V is a pointwise left Kan extension of H : D ⥤ V along L : D ⥤ D', and if K : E ⥤ V is any functor such that for any e : E, the functor tensorRight (K.obj e) commutes with colimits of shape CostructuredArrow L d, then the functor H' ⊠ K is a pointwise left kan extension of H ⊠ K along L.prod (𝟭 E).

We also prove a similar criterion to establish that K ⊠ H' is a pointwise left Kan extension of K ⊠ H along (𝟭 E).prod L.

@[reducible, inline]

Given an extension α : H ⟶ L ⋙ H', this is the canonical extension H ⊠ K ⟶ L.prod (𝟭 E) ⋙ H' ⊠ K it induces through bifunctoriality of the external product.

Equations
Instances For
    @[reducible, inline]

    Given an extension α : H ⟶ L ⋙ H', this is the canonical extension K ⊠ H ⟶ (𝟭 E).prod L ⋙ K ⊠ H' it induces through bifunctoriality of the external product.

    Equations
    Instances For

      If H' : D' ⥤ V is a pointwise left Kan extension along L : D ⥤ D' at (d : D') and if tensoring right with an object preserves colimits in V, then H' ⊠ K : D' × E ⥤ V is a pointwise left Kan extension along L × (𝟭 E) at (d, e) for every e : E.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For

        If H' : D' ⥤ V is a pointwise left Kan extension along L : D ⥤ D', and if tensoring right with an object preserves colimits in V then H' ⊠ K : D' × E ⥤ V is a pointwise left Kan extension along L × (𝟭 E).

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For

          If H' : D' ⥤ V is a pointwise left Kan extension along L : D ⥤ D' at d : D' and if tensoring left with an object preserves colimits in V, then K ⊠ H' : D' × E ⥤ V is a pointwise left Kan extension along (𝟭 E) × L at (e, d) for every e.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For

            If H' : D' ⥤ V is a pointwise left Kan extension along L : D ⥤ D' and if tensoring left with an object preserves colimits in V, then K ⊠ H' : D' × E ⥤ V is a pointwise left Kan extension along (𝟭 E) × L.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For