125
125
126
126
/-- Every functor indexing a (co)equalizer is naturally isomorphic (actually, equal) to a
127
127
`parallel_pair` -/
128
+ @[simps {rhs_md := semireducible}]
128
129
def diagram_iso_parallel_pair (F : walking_parallel_pair ⥤ C) :
129
130
F ≅ parallel_pair (F.map left) (F.map right) :=
130
131
nat_iso.of_components (λ j, eq_to_iso $ by cases j; tidy) $ by tidy
@@ -138,19 +139,19 @@ abbreviation cofork (f g : X ⟶ Y) := cocone (parallel_pair f g)
138
139
variables {f g : X ⟶ Y}
139
140
140
141
@[simp, reassoc] lemma fork.app_zero_left (s : fork f g) :
141
- ( s.π) .app zero ≫ f = ( s.π) .app one :=
142
+ s.π.app zero ≫ f = s.π.app one :=
142
143
by rw [←s.w left, parallel_pair_map_left]
143
144
144
145
@[simp, reassoc] lemma fork.app_zero_right (s : fork f g) :
145
- ( s.π) .app zero ≫ g = ( s.π) .app one :=
146
+ s.π.app zero ≫ g = s.π.app one :=
146
147
by rw [←s.w right, parallel_pair_map_right]
147
148
148
149
@[simp, reassoc] lemma cofork.left_app_one (s : cofork f g) :
149
- f ≫ ( s.ι) .app one = ( s.ι) .app zero :=
150
+ f ≫ s.ι.app one = s.ι.app zero :=
150
151
by rw [←s.w left, parallel_pair_map_left]
151
152
152
153
@[simp, reassoc] lemma cofork.right_app_one (s : cofork f g) :
153
- g ≫ ( s.ι) .app one = ( s.ι) .app zero :=
154
+ g ≫ s.ι.app one = s.ι.app zero :=
154
155
by rw [←s.w right, parallel_pair_map_right]
155
156
156
157
/-- A fork on `f g : X ⟶ Y` is determined by the morphism `ι : P ⟶ X` satisfying `ι ≫ f = ι ≫ g`.
@@ -208,42 +209,13 @@ abbreviation cofork.π (t : cofork f g) := t.ι.app one
208
209
lemma fork.ι_eq_app_zero (t : fork f g) : fork.ι t = t.π.app zero := rfl
209
210
lemma cofork.π_eq_app_one (t : cofork f g) : cofork.π t = t.ι.app one := rfl
210
211
211
- lemma fork.condition (t : fork f g) : (fork.ι t) ≫ f = (fork.ι t) ≫ g :=
212
- begin
213
- erw [t.w left, ← t.w right], refl
214
- end
215
- lemma cofork.condition (t : cofork f g) : f ≫ (cofork.π t) = g ≫ (cofork.π t) :=
216
- begin
217
- erw [t.w left, ← t.w right], refl
218
- end
212
+ @[reassoc]
213
+ lemma fork.condition (t : fork f g) : fork.ι t ≫ f = fork.ι t ≫ g :=
214
+ by rw [t.app_zero_left, t.app_zero_right]
215
+ @[reassoc]
216
+ lemma cofork.condition (t : cofork f g) : f ≫ cofork.π t = g ≫ cofork.π t :=
217
+ by rw [t.left_app_one, t.right_app_one]
219
218
220
- /--
221
- To construct an isomorphism between forks,
222
- it suffices to give an isomorphism between the cone points
223
- and check that it commutes with the `ι` morphisms.
224
- -/
225
- def fork.ext {s t : fork f g} (i : s.X ≅ t.X) (w : s.ι = i.hom ≫ t.ι) : s ≅ t :=
226
- cones.ext i
227
- begin
228
- rintro ⟨⟩,
229
- { exact w, },
230
- { rw [←fork.app_zero_left, ←fork.ι_eq_app_zero, w],
231
- simp, },
232
- end
233
-
234
- /--
235
- To construct an isomorphism between coforks,
236
- it suffices to give an isomorphism between the cocone points
237
- and check that it commutes with the `π` morphisms.
238
- -/
239
- def cofork.ext {s t : cofork f g} (i : s.X ≅ t.X) (w : s.π = t.π ≫ i.inv) : s ≅ t :=
240
- cocones.ext i
241
- begin
242
- rintro ⟨⟩,
243
- { rw [←cofork.left_app_one, ←cofork.π_eq_app_one, w],
244
- simp, },
245
- { rw [←cofork.π_eq_app_one, w], simp },
246
- end
247
219
248
220
/-- To check whether two maps are equalized by both maps of a fork, it suffices to check it for the
249
221
first map -/
@@ -390,6 +362,7 @@ def cofork.of_cocone
390
362
/--
391
363
Helper function for constructing morphisms between equalizer forks.
392
364
-/
365
+ @[simps]
393
366
def fork.mk_hom {s t : fork f g} (k : s.X ⟶ t.X) (w : k ≫ t.ι = s.ι) : s ⟶ t :=
394
367
{ hom := k,
395
368
w' :=
@@ -399,6 +372,38 @@ def fork.mk_hom {s t : fork f g} (k : s.X ⟶ t.X) (w : k ≫ t.ι = s.ι) : s
399
372
simpa using w =≫ f,
400
373
end }
401
374
375
+ /--
376
+ To construct an isomorphism between forks,
377
+ it suffices to give an isomorphism between the cone points
378
+ and check that it commutes with the `ι` morphisms.
379
+ -/
380
+ @[simps]
381
+ def fork.ext {s t : fork f g} (i : s.X ≅ t.X) (w : i.hom ≫ t.ι = s.ι) : s ≅ t :=
382
+ { hom := fork.mk_hom i.hom w,
383
+ inv := fork.mk_hom i.inv (by rw [← w, iso.inv_hom_id_assoc]) }
384
+
385
+ /--
386
+ Helper function for constructing morphisms between coequalizer coforks.
387
+ -/
388
+ @[simps]
389
+ def cofork.mk_hom {s t : cofork f g} (k : s.X ⟶ t.X) (w : s.π ≫ k = t.π) : s ⟶ t :=
390
+ { hom := k,
391
+ w' :=
392
+ begin
393
+ rintro ⟨_|_⟩,
394
+ simpa using f ≫= w,
395
+ exact w,
396
+ end }
397
+
398
+ /--
399
+ To construct an isomorphism between coforks,
400
+ it suffices to give an isomorphism between the cocone points
401
+ and check that it commutes with the `π` morphisms.
402
+ -/
403
+ def cofork.ext {s t : cofork f g} (i : s.X ≅ t.X) (w : s.π ≫ i.hom = t.π) : s ≅ t :=
404
+ { hom := cofork.mk_hom i.hom w,
405
+ inv := cofork.mk_hom i.inv (by rw [iso.comp_inv_eq, w]) }
406
+
402
407
variables (f g)
403
408
404
409
section
@@ -433,6 +438,10 @@ abbreviation equalizer.fork : fork f g := limit.cone (parallel_pair f g)
433
438
@[reassoc] lemma equalizer.condition : equalizer.ι f g ≫ f = equalizer.ι f g ≫ g :=
434
439
fork.condition $ limit.cone $ parallel_pair f g
435
440
441
+ /-- The equalizer built from `equalizer.ι f g` is limiting. -/
442
+ def equalizer_is_equalizer : is_limit (fork.of_ι (equalizer.ι f g) (equalizer.condition f g)) :=
443
+ is_limit.of_iso_limit (limit.is_limit _) (fork.ext (iso.refl _) (by tidy))
444
+
436
445
variables {f g}
437
446
438
447
/-- A morphism `k : W ⟶ X` satisfying `k ≫ f = k ≫ g` factors through the equalizer of `f` and `g`
528
537
(equalizer.iso_source_of_self f).inv = equalizer.lift (𝟙 X) (by simp) :=
529
538
rfl
530
539
531
- /--
532
- Helper function for constructing morphisms between coequalizer coforks.
533
- -/
534
- def cofork.mk_hom {s t : cofork f g} (k : s.X ⟶ t.X) (w : s.π ≫ k = t.π) : s ⟶ t :=
535
- { hom := k,
536
- w' :=
537
- begin
538
- rintro ⟨_|_⟩,
539
- simpa using f ≫= w,
540
- exact w,
541
- end }
542
-
543
540
section
544
541
/--
545
542
`has_coequalizer f g` represents a particular choice of colimiting cocone
@@ -572,6 +569,11 @@ abbreviation coequalizer.cofork : cofork f g := colimit.cocone (parallel_pair f
572
569
@[reassoc] lemma coequalizer.condition : f ≫ coequalizer.π f g = g ≫ coequalizer.π f g :=
573
570
cofork.condition $ colimit.cocone $ parallel_pair f g
574
571
572
+ /-- The cofork built from `coequalizer.π f g` is colimiting. -/
573
+ def coequalizer_is_coequalizer :
574
+ is_colimit (cofork.of_π (coequalizer.π f g) (coequalizer.condition f g)) :=
575
+ is_colimit.of_iso_colimit (colimit.is_colimit _) (cofork.ext (iso.refl _) (by tidy))
576
+
575
577
variables {f g}
576
578
577
579
/-- Any morphism `k : Y ⟶ W` satisfying `f ≫ k = g ≫ k` factors through the coequalizer of `f`
0 commit comments