Skip to content

Commit

Permalink
feat(category_theory/limits): preserving pullbacks (#5668)
Browse files Browse the repository at this point in the history
This touches multiple files but it's essentially the same thing as all my other PRs for preserving limits of special shapes - I can split it up if you'd like but hopefully this is alright?
  • Loading branch information
b-mehta committed Jan 9, 2021
1 parent ce34ae6 commit 1294500
Show file tree
Hide file tree
Showing 4 changed files with 162 additions and 8 deletions.
4 changes: 2 additions & 2 deletions src/category_theory/abelian/non_preadditive.lean
Expand Up @@ -131,7 +131,7 @@ let ⟨b', hb'⟩ := kernel_fork.is_limit.lift' i' (kernel.ι (prod.lift f g)) $
... = (0 : kernel (prod.lift f g) ⟶ P ⨯ Q) ≫ limits.prod.snd : by rw kernel.condition_assoc
... = 0 : zero_comp in
has_limit.mk { cone := pullback_cone.mk a' b' $ by { simp at ha' hb', rw [ha', hb'] },
is_limit := pullback_cone.is_limit.mk _ _ _
is_limit := pullback_cone.is_limit.mk _
(λ s, kernel.lift (prod.lift f g) (pullback_cone.snd s ≫ b) $ prod.hom_ext
(calc ((pullback_cone.snd s ≫ b) ≫ prod.lift f g) ≫ limits.prod.fst
= pullback_cone.snd s ≫ b ≫ f : by simp only [prod.lift_fst, category.assoc]
Expand Down Expand Up @@ -173,7 +173,7 @@ let ⟨b', hb'⟩ := cokernel_cofork.is_colimit.desc' i' (cokernel.π (coprod.de
... = 0 : has_zero_morphisms.comp_zero _ _ in
has_colimit.mk
{ cocone := pushout_cocone.mk a' b' $ by { simp only [cofork.π_of_π] at ha' hb', rw [ha', hb'] },
is_colimit := pushout_cocone.is_colimit.mk _ _ _
is_colimit := pushout_cocone.is_colimit.mk _
(λ s, cokernel.desc (coprod.desc f g) (b ≫ pushout_cocone.inr s) $ coprod.hom_ext
(calc coprod.inl ≫ coprod.desc f g ≫ b ≫ pushout_cocone.inr s
= f ≫ b ≫ pushout_cocone.inr s : by rw coprod.inl_desc_assoc
Expand Down
109 changes: 109 additions & 0 deletions src/category_theory/limits/preserves/shapes/pullbacks.lean
@@ -0,0 +1,109 @@
/-
Copyright (c) 2020 Bhavik Mehta. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Bhavik Mehta
-/
import category_theory.limits.shapes.pullbacks
import category_theory.limits.preserves.basic

/-!
# Preserving pullbacks
Constructions to relate the notions of preserving pullbacks and reflecting pullbacks to concrete
pullback cones.
In particular, we show that `pullback_comparison G f g` is an isomorphism iff `G` preserves
the pullback of `f` and `g`.
## TODO
* Dualise to pushouts
* Generalise to wide pullbacks
-/

noncomputable theory

universes v u₁ u₂

open category_theory category_theory.category category_theory.limits

variables {C : Type u₁} [category.{v} C]
variables {D : Type u₂} [category.{v} D]
variables (G : C ⥤ D)

namespace category_theory.limits

variables {W X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} {h : W ⟶ X} {k : W ⟶ Y} (comm : h ≫ f = k ≫ g)

/--
The map of a pullback cone is a limit iff the fork consisting of the mapped morphisms is a limit.
This essentially lets us commute `pullback_cone.mk` with `functor.map_cone`.
-/
def is_limit_map_cone_pullback_cone_equiv :
is_limit (G.map_cone (pullback_cone.mk h k comm)) ≃
is_limit (pullback_cone.mk (G.map h) (G.map k) (by simp only [← G.map_comp, comm])
: pullback_cone (G.map f) (G.map g)) :=
(is_limit.postcompose_hom_equiv (diagram_iso_cospan _) _).symm.trans
(is_limit.equiv_iso_limit (cones.ext (iso.refl _) (by { rintro (_ | _ | _), tidy })))

/-- The property of preserving pullbacks expressed in terms of binary fans. -/
def is_limit_pullback_cone_map_of_is_limit [preserves_limit (cospan f g) G]
(l : is_limit (pullback_cone.mk h k comm)) :
is_limit (pullback_cone.mk (G.map h) (G.map k) _) :=
is_limit_map_cone_pullback_cone_equiv G comm (preserves_limit.preserves l)

/-- The property of reflecting pullbacks expressed in terms of binary fans. -/
def is_limit_of_is_limit_pullback_cone_map [reflects_limit (cospan f g) G]
(l : is_limit (pullback_cone.mk (G.map h) (G.map k) _)) :
is_limit (pullback_cone.mk h k comm) :=
reflects_limit.reflects ((is_limit_map_cone_pullback_cone_equiv G comm).symm l)

variables (f g) [has_pullback f g]

/--
If `G` preserves pullbacks and `C` has them, then the pullback cone constructed of the mapped
morphisms of the pullback cone is a limit.
-/
def is_limit_of_has_pullback_of_preserves_limit
[preserves_limit (cospan f g) G] :
is_limit (pullback_cone.mk (G.map pullback.fst) (G.map pullback.snd) _) :=
is_limit_pullback_cone_map_of_is_limit G _ (pullback_is_pullback f g)

variables [has_pullback (G.map f) (G.map g)]

/--
If the pullback comparison map for `G` at `(f,g)` is an isomorphism, then `G` preserves the
pullback of `(f,g)`.
-/
def preserves_pullback.of_iso_comparison [i : is_iso (pullback_comparison G f g)] :
preserves_limit (cospan f g) G :=
begin
apply preserves_limit_of_preserves_limit_cone (pullback_is_pullback f g),
apply (is_limit_map_cone_pullback_cone_equiv _ _).symm _,
apply is_limit.of_point_iso (limit.is_limit (cospan (G.map f) (G.map g))),
apply i,
end

variables [preserves_limit (cospan f g) G]
/--
If `G` preserves the pullback of `(f,g)`, then the pullback comparison map for `G` at `(f,g)` is
an isomorphism.
-/
def preserves_pullback.iso :
G.obj (pullback f g) ≅ pullback (G.map f) (G.map g) :=
is_limit.cone_point_unique_up_to_iso
(is_limit_of_has_pullback_of_preserves_limit G f g)
(limit.is_limit _)

@[simp]
lemma preserves_pullback.iso_hom :
(preserves_pullback.iso G f g).hom = pullback_comparison G f g := rfl

instance : is_iso (pullback_comparison G f g) :=
begin
rw ← preserves_pullback.iso_hom,
apply_instance
end

end category_theory.limits
Expand Up @@ -28,7 +28,7 @@ lemma has_limit_cospan_of_has_limit_pair_of_has_limit_parallel_pair
let π₁ : X ⨯ Y ⟶ X := prod.fst, π₂ : X ⨯ Y ⟶ Y := prod.snd, e := equalizer.ι (π₁ ≫ f) (π₂ ≫ g) in
has_limit.mk
{ cone := pullback_cone.mk (e ≫ π₁) (e ≫ π₂) $ by simp only [category.assoc, equalizer.condition],
is_limit := pullback_cone.is_limit.mk _ _ _
is_limit := pullback_cone.is_limit.mk _
(λ s, equalizer.lift (prod.lift (s.π.app walking_cospan.left)
(s.π.app walking_cospan.right)) $ by
rw [←category.assoc, limit.lift_π, ←category.assoc, limit.lift_π];
Expand Down Expand Up @@ -61,7 +61,7 @@ let ι₁ : Y ⟶ Y ⨿ Z := coprod.inl, ι₂ : Z ⟶ Y ⨿ Z := coprod.inr,
has_colimit.mk
{ cocone := pushout_cocone.mk (ι₁ ≫ c) (ι₂ ≫ c) $
by rw [←category.assoc, ←category.assoc, coequalizer.condition],
is_colimit := pushout_cocone.is_colimit.mk _ _ _
is_colimit := pushout_cocone.is_colimit.mk _
(λ s, coequalizer.desc (coprod.desc (s.ι.app walking_span.left)
(s.ι.app walking_span.right)) $ by
rw [category.assoc, colimit.ι_desc, category.assoc, colimit.ι_desc];
Expand Down
53 changes: 49 additions & 4 deletions src/category_theory/limits/shapes/pullbacks.lean
Expand Up @@ -26,7 +26,7 @@ open category_theory

namespace category_theory.limits

universes v u
universes v u u₂

local attribute [tidy] tactic.case_bash

Expand Down Expand Up @@ -131,11 +131,13 @@ lemma span_map_id {X Y Z : C} (f : X ⟶ Y) (g : X ⟶ Z) (w : walking_span) :
(span f g).map (walking_span.hom.id w) = 𝟙 _ := rfl

/-- Every diagram indexing an pullback is naturally isomorphic (actually, equal) to a `cospan` -/
@[simps {rhs_md := semireducible}]
def diagram_iso_cospan (F : walking_cospan ⥤ C) :
F ≅ cospan (F.map inl) (F.map inr) :=
nat_iso.of_components (λ j, eq_to_iso (by tidy)) (by tidy)

/-- Every diagram indexing a pushout is naturally isomorphic (actually, equal) to a `span` -/
@[simps {rhs_md := semireducible}]
def diagram_iso_span (F : walking_span ⥤ C) :
F ≅ span (F.map fst) (F.map snd) :=
nat_iso.of_components (λ j, eq_to_iso (by tidy)) (by tidy)
Expand Down Expand Up @@ -228,7 +230,7 @@ def is_limit.lift' {t : pullback_cone f g} (ht : is_limit t) {W : C} (h : W ⟶
This is a more convenient formulation to show that a `pullback_cone` constructed using
`pullback_cone.mk` is a limit cone.
-/
def is_limit.mk {W : C} (fst : W ⟶ X) (snd : W ⟶ Y) (eq : fst ≫ f = snd ≫ g)
def is_limit.mk {W : C} {fst : W ⟶ X} {snd : W ⟶ Y} (eq : fst ≫ f = snd ≫ g)
(lift : Π (s : pullback_cone f g), s.X ⟶ W)
(fac_left : ∀ (s : pullback_cone f g), lift s ≫ fst = s.fst)
(fac_right : ∀ (s : pullback_cone f g), lift s ≫ snd = s.snd)
Expand Down Expand Up @@ -257,7 +259,7 @@ shown in `mono_of_pullback_is_id`.
-/
def is_limit_mk_id_id (f : X ⟶ Y) [mono f] :
is_limit (mk (𝟙 X) (𝟙 X) rfl : pullback_cone f f) :=
is_limit.mk _ _ _
is_limit.mk _
(λ s, s.fst)
(λ s, category.comp_id _)
(λ s, by rw [←cancel_mono f, category.comp_id, s.condition])
Expand Down Expand Up @@ -359,7 +361,7 @@ def is_colimit.desc' {t : pushout_cocone f g} (ht : is_colimit t) {W : C} (h : Y
This is a more convenient formulation to show that a `pushout_cocone` constructed using
`pushout_cocone.mk` is a colimit cocone.
-/
def is_colimit.mk {W : C} (inl : Y ⟶ W) (inr : Z ⟶ W) (eq : f ≫ inl = g ≫ inr)
def is_colimit.mk {W : C} {inl : Y ⟶ W} {inr : Z ⟶ W} (eq : f ≫ inl = g ≫ inr)
(desc : Π (s : pushout_cocone f g), W ⟶ s.X)
(fac_left : ∀ (s : pushout_cocone f g), inl ≫ desc s = s.inl)
(fac_right : ∀ (s : pushout_cocone f g), inr ≫ desc s = s.inr)
Expand Down Expand Up @@ -527,6 +529,12 @@ pushout_cocone.condition _
(h₁ : k ≫ pullback.snd = l ≫ pullback.snd) : k = l :=
limit.hom_ext $ pullback_cone.equalizer_ext _ h₀ h₁

/-- The pullback cone built from the pullback projections is a pullback. -/
def pullback_is_pullback {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) [has_pullback f g] :
is_limit (pullback_cone.mk (pullback.fst : pullback f g ⟶ _) pullback.snd pullback.condition) :=
pullback_cone.is_limit.mk _ (λ s, pullback.lift s.fst s.snd s.condition)
(by simp) (by simp) (by tidy)

/-- The pullback of a monomorphism is a monomorphism -/
instance pullback.fst_of_mono {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} [has_pullback f g]
[mono g] : mono (pullback.fst : pullback f g ⟶ X) :=
Expand Down Expand Up @@ -554,6 +562,43 @@ instance pushout.inr_of_epi {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} [has_pushout
epi (pushout.inr : Z ⟶ pushout f g) :=
⟨λ W u v h, pushout.hom_ext ((cancel_epi f).1 $ by simp [pushout.condition_assoc, h]) h⟩

section

variables {D : Type u₂} [category.{v} D] (G : C ⥤ D)

/--
The comparison morphism for the pullback of `f,g`.
This is an isomorphism iff `G` preserves the pullback of `f,g`; see
`category_theory/limits/preserves/shapes/pullbacks.lean`
-/
def pullback_comparison (f : X ⟶ Z) (g : Y ⟶ Z)
[has_pullback f g] [has_pullback (G.map f) (G.map g)] :
G.obj (pullback f g) ⟶ pullback (G.map f) (G.map g) :=
pullback.lift (G.map pullback.fst) (G.map pullback.snd)
(by simp only [←G.map_comp, pullback.condition])

@[simp, reassoc]
lemma pullback_comparison_comp_fst (f : X ⟶ Z) (g : Y ⟶ Z)
[has_pullback f g] [has_pullback (G.map f) (G.map g)] :
pullback_comparison G f g ≫ pullback.fst = G.map pullback.fst :=
pullback.lift_fst _ _ _

@[simp, reassoc]
lemma pullback_comparison_comp_snd (f : X ⟶ Z) (g : Y ⟶ Z)
[has_pullback f g] [has_pullback (G.map f) (G.map g)] :
pullback_comparison G f g ≫ pullback.snd = G.map pullback.snd :=
pullback.lift_snd _ _ _

@[simp, reassoc]
lemma map_lift_pullback_comparison (f : X ⟶ Z) (g : Y ⟶ Z)
[has_pullback f g] [has_pullback (G.map f) (G.map g)]
{W : C} {h : W ⟶ X} {k : W ⟶ Y} (w : h ≫ f = k ≫ g) :
G.map (pullback.lift _ _ w) ≫ pullback_comparison G f g =
pullback.lift (G.map h) (G.map k) (by simp only [←G.map_comp, w]) :=
by { ext; simp [← G.map_comp] }

end

variables (C)

/--
Expand Down

0 comments on commit 1294500

Please sign in to comment.