@@ -12,6 +12,8 @@ open import Categories.Adjoint.RAPL using (rapl)
12
12
open import Categories.Category using (Category; _[_,_]; _[_≈_]; _[_∘_])
13
13
open import Categories.Category.Product using (_⁂_; _⁂ⁿⁱ_)
14
14
open import Categories.Category.Construction.Comma using (CommaObj; Comma⇒; _↙_)
15
+ open import Categories.Diagram.Cone using (Cone; Cone⇒)
16
+ open import Categories.Diagram.Cone.Properties using (F-map-Coneʳ)
15
17
open import Categories.Diagram.Limit using (Limit)
16
18
open import Categories.Diagram.Colimit using (Colimit)
17
19
open import Categories.Functor using (Functor; _∘F_) renaming (id to idF)
@@ -34,12 +36,12 @@ import Categories.Yoneda.Properties as YP using (yoneda-NI)
34
36
import Categories.Diagram.Duality as Duality using (coLimit⇒Colimit; Colimit⇒coLimit)
35
37
36
38
import Categories.Morphism.Reasoning as MR using (pushʳ; pullˡ; pushˡ; elimʳ; center; center⁻¹;
37
- elimˡ; cancelˡ; pullʳ; cancelʳ; id-comm-sym)
39
+ elimˡ; cancelˡ; pullʳ; cancelʳ; id-comm-sym; extendʳ )
38
40
39
41
private
40
42
variable
41
43
o ℓ e : Level
42
- C D E J : Category o ℓ e
44
+ C D E J K : Category o ℓ e
43
45
44
46
-- if the left adjoint functor is a partial application of bifunctor, then it uniquely
45
47
-- determines a bifunctor compatible with the right adjoint functor.
@@ -169,6 +171,73 @@ module _ {L : Functor C D} {R : Functor D C} (L⊣R : L ⊣ R) (F : Functor J C)
169
171
lapc : Colimit F → Colimit (L ∘F F)
170
172
lapc col = Duality.coLimit⇒Colimit D (rapl (Adjoint.op L⊣R) F.op (Duality.Colimit⇒coLimit C col))
171
173
174
+ -- left adjoint preserves diagrams in limits
175
+ module _ {L : Functor J K} {R : Functor K J} (L⊣R : L ⊣ R) {F : Functor K C} (Lm : Limit F) where
176
+
177
+ private
178
+ module L = Functor L
179
+ module R = Functor R
180
+ module F = Functor F
181
+ open Adjoint L⊣R
182
+ open Category C
183
+ open HomReasoning
184
+ open MR C
185
+ open Limit Lm
186
+
187
+ private module _ {A : Cone (F ∘F L)} where
188
+
189
+ private module A = Cone A
190
+
191
+ A′ : Cone F
192
+ A′ = record
193
+ { N = A.N
194
+ ; apex = record
195
+ { ψ = λ X → F.₁ (counit.η X) ∘ A.ψ (R.₀ X)
196
+ ; commute = λ {X} {Y} f → begin
197
+ F.₁ f ∘ F.₁ (counit.η X) ∘ A.ψ (R.₀ X) ≈⟨ extendʳ ([ F ]-resp-square (counit.sym-commute f)) ⟩
198
+ F.₁ (counit.η Y) ∘ F.₁ (L.₁ (R.₁ f)) ∘ A.ψ (R.₀ X) ≈⟨ refl⟩∘⟨ A.commute (R.₁ f) ⟩
199
+ F.₁ (counit.η Y) ∘ A.ψ (R.₀ Y) ∎
200
+ }
201
+ }
202
+
203
+ !cone : Cone⇒ (F ∘F L) A (F-map-Coneʳ L limit)
204
+ !cone = record
205
+ { arr = rep A′
206
+ ; commute = λ {X} → begin
207
+ proj (L.₀ X) ∘ rep A′ ≈⟨ commute ⟩
208
+ F.₁ (counit.η (L.₀ X)) ∘ A.ψ (R.₀ (L.₀ X)) ≈⟨ refl⟩∘⟨ A.commute (unit.η X) ⟨
209
+ F.₁ (counit.η (L.₀ X)) ∘ F.₁ (L.₁ (unit.η X)) ∘ A.ψ X ≈⟨ cancelˡ ([ F ]-resp-∘ zig ○ F.identity) ⟩
210
+ A.ψ X ∎
211
+ }
212
+
213
+ private
214
+ !cone-unique : ∀ {A} (f : Cone⇒ (F ∘F L) A (F-map-Coneʳ L limit)) → rep (A′ {A}) ≈ Cone⇒.arr f
215
+ !cone-unique {A} f = terminal.!-unique {A′ {A}} record
216
+ { arr = f.arr
217
+ ; commute = λ {X} → begin
218
+ proj X ∘ f.arr ≈⟨ pushˡ (⟺ (limit-commute (counit.η X))) ⟩
219
+ F.₁ (counit.η X) ∘ proj (L.₀ (R.₀ X)) ∘ f.arr ≈⟨ (refl⟩∘⟨ f.commute) ⟩
220
+ F.₁ (counit.η X) ∘ A.ψ (R.₀ X) ∎
221
+ }
222
+ where
223
+ module A = Cone A
224
+ module f = Cone⇒ f
225
+
226
+ la-preserves-diagram : Limit (F ∘F L)
227
+ la-preserves-diagram = record
228
+ { terminal = record
229
+ { ⊤ = F-map-Coneʳ L limit
230
+ ; ⊤-is-terminal = record
231
+ { ! = !cone
232
+ ; !-unique = !cone-unique
233
+ }
234
+ }
235
+ }
236
+
237
+ -- right adjoint preserves diagrams in colimits
238
+ ra-preserves-diagram : {L : Functor J K} {R : Functor K J} (L⊣R : L ⊣ R) {F : Functor J C} → Colimit F → Colimit (F ∘F R)
239
+ ra-preserves-diagram L⊣R colim = Duality.coLimit⇒Colimit _ (la-preserves-diagram (Adjoint.op L⊣R) (Duality.Colimit⇒coLimit _ colim))
240
+
172
241
-- adjoint functors induce monads and comonads
173
242
module _ {L : Functor C D} {R : Functor D C} (L⊣R : L ⊣ R) where
174
243
private
0 commit comments