Lifting of bifunctors #
In this file, in the context of the localization of categories, we extend the notion
of lifting of functors to the case of bifunctors. As the localization of categories
behaves well with respect to finite products of categories (when the classes of
morphisms contain identities), all the definitions for bifunctors C₁ ⥤ C₂ ⥤ E
are obtained by reducing to the case of functors (C₁ × C₂) ⥤ E
by using
currying and uncurrying.
Given morphism properties W₁ : MorphismProperty C₁
and W₂ : MorphismProperty C₂
,
and a functor F : C₁ ⥤ C₂ ⥤ E
, we define MorphismProperty.IsInvertedBy₂ W₁ W₂ F
as the condition that the functor uncurry.obj F : C₁ × C₂ ⥤ E
inverts W₁.prod W₂
.
If L₁ : C₁ ⥤ D₁
and L₂ : C₂ ⥤ D₂
are localization functors for W₁
and W₂
respectively, and F : C₁ ⥤ C₂ ⥤ E
satisfies MorphismProperty.IsInvertedBy₂ W₁ W₂ F
,
we introduce Localization.lift₂ F hF L₁ L₂ : D₁ ⥤ D₂ ⥤ E
which is a bifunctor
which lifts F
.
Classes of morphisms W₁ : MorphismProperty C₁
and W₂ : MorphismProperty C₂
are said
to be inverted by F : C₁ ⥤ C₂ ⥤ E
if W₁.prod W₂
is inverted by
the functor uncurry.obj F : C₁ × C₂ ⥤ E
.
Equations
- W₁.IsInvertedBy₂ W₂ F = (W₁.prod W₂).IsInvertedBy (CategoryTheory.uncurry.obj F)
Instances For
Given functors L₁ : C₁ ⥤ D₁
, L₂ : C₂ ⥤ D₂
, morphisms properties W₁
on C₁
and W₂
on C₂
, and functors F : C₁ ⥤ C₂ ⥤ E
and F' : D₁ ⥤ D₂ ⥤ E
, we say
Lifting₂ L₁ L₂ W₁ W₂ F F'
holds if F
is induced by F'
, up to an isomorphism.
- iso' : (((CategoryTheory.whiskeringLeft₂ E).obj L₁).obj L₂).obj F' ≅ F
the isomorphism
(((whiskeringLeft₂ E).obj L₁).obj L₂).obj F' ≅ F
expressing thatF
is induced byF'
up to an isomorphism
Instances
The isomorphism (((whiskeringLeft₂ E).obj L₁).obj L₂).obj F' ≅ F
when
Lifting₂ L₁ L₂ W₁ W₂ F F'
holds.
Equations
- CategoryTheory.Localization.Lifting₂.iso L₁ L₂ W₁ W₂ F F' = CategoryTheory.Localization.Lifting₂.iso' W₁ W₂
Instances For
If Lifting₂ L₁ L₂ W₁ W₂ F F'
holds, then Lifting L₂ W₂ (F.obj X₁) (F'.obj (L₁.obj X₁))
holds for any X₁ : C₁
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- CategoryTheory.Localization.Lifting₂.flip L₁ L₂ W₁ W₂ F F' = { iso' := (CategoryTheory.flipFunctor C₁ C₂ E).mapIso (CategoryTheory.Localization.Lifting₂.iso L₁ L₂ W₁ W₂ F F') }
If Lifting₂ L₁ L₂ W₁ W₂ F F'
holds, then
Lifting L₁ W₁ (F.flip.obj X₂) (F'.flip.obj (L₂.obj X₂))
holds for any X₂ : C₂
.
Equations
- CategoryTheory.Localization.Lifting₂.snd L₁ L₂ W₁ W₂ F F' X₂ = CategoryTheory.Localization.Lifting₂.fst L₂ L₁ W₂ W₁ F.flip F'.flip X₂
Instances For
Equations
- CategoryTheory.Localization.Lifting₂.uncurry L₁ L₂ W₁ W₂ F F' = { iso' := CategoryTheory.uncurry.mapIso (CategoryTheory.Localization.Lifting₂.iso L₁ L₂ W₁ W₂ F F') }
Given localization functor L₁ : C₁ ⥤ D₁
and L₂ : C₂ ⥤ D₂
with respect
to W₁ : MorphismProperty C₁
and W₂ : MorphismProperty C₂
respectively,
and a bifunctor F : C₁ ⥤ C₂ ⥤ E
which inverts W₁
and W₂
, this is
the induced localized bifunctor D₁ ⥤ D₂ ⥤ E
.
Equations
- CategoryTheory.Localization.lift₂ F hF L₁ L₂ = CategoryTheory.curry.obj (CategoryTheory.Localization.lift (CategoryTheory.uncurry.obj F) hF (L₁.prod L₂))
Instances For
Equations
- One or more equations did not get rendered due to their size.
Equations
- CategoryTheory.Localization.Lifting₂.liftingLift₂ F hF L₁ L₂ X₁ = CategoryTheory.Localization.Lifting₂.fst L₁ L₂ W₁ W₂ F (CategoryTheory.Localization.lift₂ F hF L₁ L₂) X₁
Equations
- CategoryTheory.Localization.Lifting₂.liftingLift₂Flip F hF L₁ L₂ X₂ = CategoryTheory.Localization.Lifting₂.snd L₁ L₂ W₁ W₂ F (CategoryTheory.Localization.lift₂ F hF L₁ L₂) X₂
The natural transformation F₁' ⟶ F₂'
of bifunctors induced by a
natural transformation τ : F₁ ⟶ F₂
when Lifting₂ L₁ L₂ W₁ W₂ F₁ F₁'
and Lifting₂ L₁ L₂ W₁ W₂ F₂ F₂'
hold.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The natural isomorphism F₁' ≅ F₂'
of bifunctors induced by a
natural isomorphism e : F₁ ≅ F₂
when Lifting₂ L₁ L₂ W₁ W₂ F₁ F₁'
and Lifting₂ L₁ L₂ W₁ W₂ F₂ F₂'
hold.
Equations
- One or more equations did not get rendered due to their size.