|
| 1 | +/- |
| 2 | +Copyright (c) 2022 Yuma Mizuno. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Yuma Mizuno |
| 5 | +-/ |
| 6 | +import category_theory.bicategory.basic |
| 7 | + |
| 8 | +/-! |
| 9 | +# Oplax functors |
| 10 | +
|
| 11 | +An oplax functor `F` between bicategories `B` and `C` consists of |
| 12 | +* a function between objects `F.obj : B ⟶ C`, |
| 13 | +* a family of functions between 1-morphisms `F.map : (a ⟶ b) → (obj a ⟶ obj b)`, |
| 14 | +* a family of functions between 2-morphisms `F.map₂ : (f ⟶ g) → (map f ⟶ map g)`, |
| 15 | +* a family of 2-morphisms `F.map_id a : F.map (𝟙 a) ⟶ 𝟙 (F.obj a)`, |
| 16 | +* a family of 2-morphisms `F.map_comp f g : F.map (f ≫ g) ⟶ F.map f ≫ F.map g`, and |
| 17 | +* certain consistency conditions on them. |
| 18 | +
|
| 19 | +## Main definitions |
| 20 | +
|
| 21 | +* `oplax_functor B C` : an oplax functor between bicategories `B` and `C` |
| 22 | +* `oplax_functor.comp F G` : the composition of oplax functors |
| 23 | +
|
| 24 | +## Future work |
| 25 | +
|
| 26 | +There are two types of functors between bicategories, called lax and oplax functors, depending on |
| 27 | +the directions of `map_id` and `map_comp`. We may need both in mathlib in the future, but for |
| 28 | +now we only define oplax functors. |
| 29 | +
|
| 30 | +In addition, if we require `map_id` and `map_comp` to be isomorphisms, we obtain the definition |
| 31 | +of pseudofunctors. There are several possible design choices to implement pseudofunctors, |
| 32 | +but the choice is left to future development. |
| 33 | +-/ |
| 34 | + |
| 35 | +set_option old_structure_cmd true |
| 36 | + |
| 37 | +namespace category_theory |
| 38 | + |
| 39 | +open category bicategory |
| 40 | +open_locale bicategory |
| 41 | + |
| 42 | +universes w₁ w₂ w₃ v₁ v₂ v₃ u₁ u₂ u₃ |
| 43 | + |
| 44 | +section |
| 45 | +variables (B : Type u₁) [quiver.{v₁+1} B] [∀ a b : B, quiver.{w₁+1} (a ⟶ b)] |
| 46 | +variables (C : Type u₂) [quiver.{v₂+1} C] [∀ a b : C, quiver.{w₂+1} (a ⟶ b)] |
| 47 | + |
| 48 | +/-- |
| 49 | +A prelax functor between bicategories consists of functions between objects, |
| 50 | +1-morphisms, and 2-morphisms. This structure will be extended to define `oplax_functor`. |
| 51 | +-/ |
| 52 | +structure prelax_functor extends prefunctor B C : Type (max w₁ w₂ v₁ v₂ u₁ u₂) := |
| 53 | +(map₂ {a b : B} {f g : a ⟶ b} : (f ⟶ g) → (map f ⟶ map g)) |
| 54 | + |
| 55 | +/-- The prefunctor between the underlying quivers. -/ |
| 56 | +add_decl_doc prelax_functor.to_prefunctor |
| 57 | + |
| 58 | +namespace prelax_functor |
| 59 | + |
| 60 | +variables {B C} {D : Type u₃} [quiver.{v₃+1} D] [∀ a b : D, quiver.{w₃+1} (a ⟶ b)] |
| 61 | +variables (F : prelax_functor B C) (G : prelax_functor C D) |
| 62 | + |
| 63 | +@[simp] lemma to_prefunctor_obj : F.to_prefunctor.obj = F.obj := rfl |
| 64 | +@[simp] lemma to_prefunctor_map : F.to_prefunctor.map = F.map := rfl |
| 65 | + |
| 66 | +variables (B) |
| 67 | + |
| 68 | +/-- The identity prelax functor. -/ |
| 69 | +@[simps] |
| 70 | +def id : prelax_functor B B := |
| 71 | +{ map₂ := λ a b f g η, η, .. prefunctor.id B } |
| 72 | + |
| 73 | +instance : inhabited (prelax_functor B B) := ⟨prelax_functor.id B⟩ |
| 74 | + |
| 75 | +variables {B} |
| 76 | + |
| 77 | +/-- Composition of prelax functors. -/ |
| 78 | +@[simps] |
| 79 | +def comp : prelax_functor B D := |
| 80 | +{ map₂ := λ a b f g η, G.map₂ (F.map₂ η), .. F.to_prefunctor.comp G.to_prefunctor } |
| 81 | + |
| 82 | +end prelax_functor |
| 83 | + |
| 84 | +end |
| 85 | + |
| 86 | +section |
| 87 | +variables {B : Type u₁} [bicategory.{w₁ v₁} B] {C : Type u₂} [bicategory.{w₂ v₂} C] |
| 88 | + |
| 89 | +/-- |
| 90 | +This auxiliary definition states that oplax functors preserve the associators |
| 91 | +modulo some adjustments of domains and codomains of 2-morphisms. |
| 92 | +-/ |
| 93 | +/- |
| 94 | +We use this auxiliary definition instead of writing it directly in the definition |
| 95 | +of oplax functors because doing so will cause a timeout. |
| 96 | +-/ |
| 97 | +@[simp] |
| 98 | +def oplax_functor.map₂_associator_aux |
| 99 | + (obj : B → C) (map : Π {X Y : B}, (X ⟶ Y) → (obj X ⟶ obj Y)) |
| 100 | + (map₂ : Π {a b : B} {f g : a ⟶ b}, (f ⟶ g) → (map f ⟶ map g)) |
| 101 | + (map_comp : Π {a b c : B} (f : a ⟶ b) (g : b ⟶ c), map (f ≫ g) ⟶ map f ≫ map g) |
| 102 | + {a b c d : B} (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) : Prop := |
| 103 | +map₂ (α_ f g h).hom ≫ map_comp f (g ≫ h) ≫ (map f ◁ map_comp g h) = |
| 104 | + map_comp (f ≫ g) h ≫ (map_comp f g ▷ map h) ≫ (α_ (map f) (map g) (map h)).hom |
| 105 | + |
| 106 | +variables (B C) |
| 107 | + |
| 108 | +/-- |
| 109 | +An oplax functor `F` between bicategories `B` and `C` consists of functions between objects, |
| 110 | +1-morphisms, and 2-morphisms. |
| 111 | +
|
| 112 | +Unlike functors between categories, functions between 1-morphisms do not need to strictly commute |
| 113 | +with compositions, and do not need to strictly preserve the identity. Instead, there are |
| 114 | +specified 2-morphisms `F.map (𝟙 a) ⟶ 𝟙 (F.obj a)` and `F.map (f ≫ g) ⟶ F.map f ≫ F.map g`. |
| 115 | +
|
| 116 | +Functions between 2-morphisms strictly commute with compositions and preserve the identity. |
| 117 | +They also preserve the associator, the left unitor, and the right unitor modulo some adjustments |
| 118 | +of domains and codomains of 2-morphisms. |
| 119 | +-/ |
| 120 | +structure oplax_functor extends prelax_functor B C : Type (max w₁ w₂ v₁ v₂ u₁ u₂) := |
| 121 | +(map_id (a : B) : map (𝟙 a) ⟶ 𝟙 (obj a)) |
| 122 | +(map_comp {a b c : B} (f : a ⟶ b) (g : b ⟶ c) : map (f ≫ g) ⟶ map f ≫ map g) |
| 123 | +(map_comp_naturality_left' : ∀ {a b c : B} {f f' : a ⟶ b} (η : f ⟶ f') (g : b ⟶ c), |
| 124 | + map₂ (η ▷ g) ≫ map_comp f' g = map_comp f g ≫ (map₂ η ▷ map g) . obviously) |
| 125 | +(map_comp_naturality_right' : ∀ {a b c : B} (f : a ⟶ b) {g g' : b ⟶ c} (η : g ⟶ g'), |
| 126 | + map₂ (f ◁ η) ≫ map_comp f g' = map_comp f g ≫ (map f ◁ map₂ η) . obviously) |
| 127 | +(map₂_id' : ∀ {a b : B} (f : a ⟶ b), map₂ (𝟙 f) = 𝟙 (map f) . obviously) |
| 128 | +(map₂_comp' : ∀ {a b : B} {f g h : a ⟶ b} (η : f ⟶ g) (θ : g ⟶ h), |
| 129 | + map₂ (η ≫ θ) = map₂ η ≫ map₂ θ . obviously) |
| 130 | +(map₂_associator' : ∀ {a b c d : B} (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d), |
| 131 | + oplax_functor.map₂_associator_aux obj (λ a b, map) (λ a b f g, map₂) (λ a b c, map_comp) f g h |
| 132 | + . obviously) |
| 133 | +(map₂_left_unitor' : ∀ {a b : B} (f : a ⟶ b), |
| 134 | + map₂ (λ_ f).hom = map_comp (𝟙 a) f ≫ (map_id a ▷ map f) ≫ (λ_ (map f)).hom . obviously) |
| 135 | +(map₂_right_unitor' : ∀ {a b : B} (f : a ⟶ b), |
| 136 | + map₂ (ρ_ f).hom = map_comp f (𝟙 b) ≫ (map f ◁ map_id b) ≫ (ρ_ (map f)).hom . obviously) |
| 137 | + |
| 138 | +restate_axiom oplax_functor.map_comp_naturality_left' |
| 139 | +restate_axiom oplax_functor.map_comp_naturality_right' |
| 140 | +restate_axiom oplax_functor.map₂_id' |
| 141 | +restate_axiom oplax_functor.map₂_comp' |
| 142 | +restate_axiom oplax_functor.map₂_associator' |
| 143 | +restate_axiom oplax_functor.map₂_left_unitor' |
| 144 | +restate_axiom oplax_functor.map₂_right_unitor' |
| 145 | +attribute [simp] |
| 146 | + oplax_functor.map_comp_naturality_left oplax_functor.map_comp_naturality_right |
| 147 | + oplax_functor.map₂_id oplax_functor.map₂_associator |
| 148 | +attribute [reassoc] |
| 149 | + oplax_functor.map_comp_naturality_left oplax_functor.map_comp_naturality_right |
| 150 | + oplax_functor.map₂_comp oplax_functor.map₂_associator |
| 151 | + oplax_functor.map₂_left_unitor oplax_functor.map₂_right_unitor |
| 152 | +attribute [simp] |
| 153 | + oplax_functor.map₂_comp oplax_functor.map₂_left_unitor oplax_functor.map₂_right_unitor |
| 154 | + |
| 155 | +namespace oplax_functor |
| 156 | + |
| 157 | +variables {B} {C} {D : Type u₃} [bicategory.{w₃ v₃} D] |
| 158 | +variables (F : oplax_functor B C) (G : oplax_functor C D) |
| 159 | + |
| 160 | +/-- Function between 1-morphisms as a functor. -/ |
| 161 | +@[simps] |
| 162 | +def map_functor (a b : B) : (a ⟶ b) ⥤ (F.obj a ⟶ F.obj b) := |
| 163 | +{ obj := λ f, F.map f, |
| 164 | + map := λ f g η, F.map₂ η } |
| 165 | + |
| 166 | +/-- The prelax functor between the underlying quivers. -/ |
| 167 | +add_decl_doc oplax_functor.to_prelax_functor |
| 168 | + |
| 169 | +@[simp] lemma to_prelax_functor_obj : F.to_prelax_functor.obj = F.obj := rfl |
| 170 | +@[simp] lemma to_prelax_functor_map : F.to_prelax_functor.map = F.map := rfl |
| 171 | +@[simp] lemma to_prelax_functor_map₂ : F.to_prelax_functor.map₂ = F.map₂ := rfl |
| 172 | + |
| 173 | +variables (B) |
| 174 | + |
| 175 | +/-- The identity oplax functor. -/ |
| 176 | +@[simps] |
| 177 | +def id : oplax_functor B B := |
| 178 | +{ map_id := λ a, 𝟙 (𝟙 a), |
| 179 | + map_comp := λ a b c f g, 𝟙 (f ≫ g), |
| 180 | + .. prelax_functor.id B } |
| 181 | + |
| 182 | +instance : inhabited (oplax_functor B B) := ⟨id B⟩ |
| 183 | + |
| 184 | +variables {B} |
| 185 | + |
| 186 | +/-- Composition of oplax functors. -/ |
| 187 | +@[simps] |
| 188 | +def comp : oplax_functor B D := |
| 189 | +{ map_id := λ a, |
| 190 | + (G.map_functor _ _).map (F.map_id a) ≫ G.map_id (F.obj a), |
| 191 | + map_comp := λ a b c f g, |
| 192 | + (G.map_functor _ _).map (F.map_comp f g) ≫ G.map_comp (F.map f) (F.map g), |
| 193 | + map_comp_naturality_left' := λ a b c f f' η g, by |
| 194 | + { dsimp, |
| 195 | + rw [←map₂_comp_assoc, map_comp_naturality_left, map₂_comp_assoc, map_comp_naturality_left, |
| 196 | + assoc] }, |
| 197 | + map_comp_naturality_right' := λ a b c f g g' η, by |
| 198 | + { dsimp, |
| 199 | + rw [←map₂_comp_assoc, map_comp_naturality_right, map₂_comp_assoc, map_comp_naturality_right, |
| 200 | + assoc] }, |
| 201 | + map₂_associator' := λ a b c d f g h, by |
| 202 | + { dsimp, |
| 203 | + simp only [map₂_associator, ←map₂_comp_assoc, ←map_comp_naturality_right_assoc, |
| 204 | + whisker_left_comp, assoc], |
| 205 | + simp only [map₂_associator, map₂_comp, map_comp_naturality_left_assoc, |
| 206 | + whisker_right_comp, assoc] }, |
| 207 | + map₂_left_unitor' := λ a b f, by |
| 208 | + { dsimp, |
| 209 | + simp only [map₂_left_unitor, map₂_comp, map_comp_naturality_left_assoc, |
| 210 | + whisker_right_comp, assoc] }, |
| 211 | + map₂_right_unitor' := λ a b f, by |
| 212 | + { dsimp, |
| 213 | + simp only [map₂_right_unitor, map₂_comp, map_comp_naturality_right_assoc, |
| 214 | + whisker_left_comp, assoc] }, |
| 215 | + .. F.to_prelax_functor.comp G.to_prelax_functor } |
| 216 | + |
| 217 | +end oplax_functor |
| 218 | + |
| 219 | +end |
| 220 | + |
| 221 | +end category_theory |
0 commit comments