|
| 1 | +/- |
| 2 | +Copyright (c) 2020 Bhavik Mehta. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Bhavik Mehta |
| 5 | +-/ |
| 6 | +import category_theory.limits.preserves.limits |
| 7 | +import category_theory.limits.shapes |
| 8 | + |
| 9 | +/-! |
| 10 | +# Preserving terminal object |
| 11 | +
|
| 12 | +Constructions to relate the notions of preserving terminal objects and reflecting terminal objects |
| 13 | +to concrete objects. |
| 14 | +
|
| 15 | +In particular, we show that `terminal_comparison G` is an isomorphism iff `G` preserves terminal |
| 16 | +objects. |
| 17 | +-/ |
| 18 | + |
| 19 | +universes v u₁ u₂ |
| 20 | + |
| 21 | +noncomputable theory |
| 22 | + |
| 23 | +open category_theory category_theory.category category_theory.limits |
| 24 | + |
| 25 | +variables {C : Type u₁} [category.{v} C] |
| 26 | +variables {D : Type u₂} [category.{v} D] |
| 27 | +variables (G : C ⥤ D) |
| 28 | + |
| 29 | +namespace category_theory.limits |
| 30 | + |
| 31 | +variables (X : C) |
| 32 | + |
| 33 | +/-- |
| 34 | +The map of an empty cone is a limit iff the mapped object is terminal. |
| 35 | +-/ |
| 36 | +def is_limit_map_cone_empty_cone_equiv : |
| 37 | + is_limit (G.map_cone (as_empty_cone X)) ≃ is_terminal (G.obj X) := |
| 38 | +(is_limit.postcompose_hom_equiv (functor.empty_ext _ _) _).symm.trans |
| 39 | + (is_limit.equiv_iso_limit (cones.ext (iso.refl _) (by tidy))) |
| 40 | + |
| 41 | +/-- The property of preserving terminal objects expressed in terms of `is_terminal`. -/ |
| 42 | +def is_terminal_obj_of_is_terminal [preserves_limit (functor.empty C) G] |
| 43 | + (l : is_terminal X) : is_terminal (G.obj X) := |
| 44 | +is_limit_map_cone_empty_cone_equiv G X (preserves_limit.preserves l) |
| 45 | + |
| 46 | +/-- The property of reflecting terminal objects expressed in terms of `is_terminal`. -/ |
| 47 | +def is_terminal_of_is_terminal_obj [reflects_limit (functor.empty C) G] |
| 48 | + (l : is_terminal (G.obj X)) : is_terminal X := |
| 49 | +reflects_limit.reflects ((is_limit_map_cone_empty_cone_equiv G X).symm l) |
| 50 | + |
| 51 | +variables [has_terminal C] |
| 52 | +/-- |
| 53 | +If `G` preserves the terminal object and `C` has a terminal object, then the image of the terminal |
| 54 | +object is terminal. |
| 55 | +-/ |
| 56 | +def is_limit_of_has_terminal_of_preserves_limit [preserves_limit (functor.empty C) G] : |
| 57 | + is_terminal (G.obj (⊤_ C)) := |
| 58 | +is_terminal_obj_of_is_terminal G (⊤_ C) terminal_is_terminal |
| 59 | + |
| 60 | +/-- |
| 61 | +If `C` has a terminal object and `G` preserves terminal objects, then `D` has a terminal object |
| 62 | +also. |
| 63 | +Note this property is somewhat unique to (co)limits of the empty diagram: for general `J`, if `C` |
| 64 | +has limits of shape `J` and `G` preserves them, then `D` does not necessarily have limits of shape |
| 65 | +`J`. |
| 66 | +-/ |
| 67 | +lemma has_terminal_of_has_terminal_of_preserves_limit [preserves_limit (functor.empty C) G] : |
| 68 | + has_terminal D := |
| 69 | +⟨λ F, |
| 70 | +begin |
| 71 | + haveI := has_limit.mk ⟨_, is_limit_of_has_terminal_of_preserves_limit G⟩, |
| 72 | + apply has_limit_of_iso F.unique_from_empty.symm, |
| 73 | +end⟩ |
| 74 | + |
| 75 | +variable [has_terminal D] |
| 76 | +/-- |
| 77 | +If the terminal comparison map for `G` is an isomorphism, then `G` preserves terminal objects. |
| 78 | +-/ |
| 79 | +def preserves_terminal.of_iso_comparison |
| 80 | + [i : is_iso (terminal_comparison G)] : preserves_limit (functor.empty C) G := |
| 81 | +begin |
| 82 | + apply preserves_limit_of_preserves_limit_cone terminal_is_terminal, |
| 83 | + apply (is_limit_map_cone_empty_cone_equiv _ _).symm _, |
| 84 | + apply is_limit.of_point_iso (limit.is_limit (functor.empty D)), |
| 85 | + apply i, |
| 86 | +end |
| 87 | + |
| 88 | +/-- If there is any isomorphism `G.obj ⊤ ⟶ ⊤`, then `G` preserves terminal objects. -/ |
| 89 | +def preserves_terminal_of_is_iso |
| 90 | + (f : G.obj (⊤_ C) ⟶ ⊤_ D) [i : is_iso f] : preserves_limit (functor.empty C) G := |
| 91 | +begin |
| 92 | + rw subsingleton.elim f (terminal_comparison G) at i, |
| 93 | + exactI preserves_terminal.of_iso_comparison G, |
| 94 | +end |
| 95 | + |
| 96 | +/-- If there is any isomorphism `G.obj ⊤ ≅ ⊤`, then `G` preserves terminal objects. -/ |
| 97 | +def preserves_terminal_of_iso |
| 98 | + (f : G.obj (⊤_ C) ≅ ⊤_ D) : preserves_limit (functor.empty C) G := |
| 99 | +preserves_terminal_of_is_iso G f.hom |
| 100 | + |
| 101 | +variables [preserves_limit (functor.empty C) G] |
| 102 | + |
| 103 | +/-- If `G` preserves terminal objects, then the terminal comparison map for `G` an isomorphism. -/ |
| 104 | +def preserves_terminal.iso : G.obj (⊤_ C) ≅ ⊤_ D := |
| 105 | +(is_limit_of_has_terminal_of_preserves_limit G).cone_point_unique_up_to_iso (limit.is_limit _) |
| 106 | + |
| 107 | +@[simp] |
| 108 | +lemma preserves_terminal.iso_hom : (preserves_terminal.iso G).hom = terminal_comparison G := |
| 109 | +rfl |
| 110 | + |
| 111 | +instance : is_iso (terminal_comparison G) := |
| 112 | +begin |
| 113 | + rw ← preserves_terminal.iso_hom, |
| 114 | + apply_instance, |
| 115 | +end |
| 116 | + |
| 117 | +end category_theory.limits |
0 commit comments