@@ -191,4 +191,57 @@ abbreviation has_products := Π (J : Type v), has_limits_of_shape (discrete J) C
191
191
/-- An abbreviation for `Π J, has_colimits_of_shape (discrete J) C` -/
192
192
abbreviation has_coproducts := Π (J : Type v), has_colimits_of_shape (discrete J) C
193
193
194
+ section reindex
195
+ variables {C} {γ : Type v} (ε : β ≃ γ) (f : γ → C)
196
+
197
+ section
198
+ variables [has_product f] [has_product (f ∘ ε)]
199
+
200
+ /-- Reindex a categorical product via an equivalence of the index types. -/
201
+ def pi.reindex : pi_obj (f ∘ ε) ≅ pi_obj f :=
202
+ has_limit.iso_of_equivalence (discrete.equivalence ε) (discrete.nat_iso (λ i, iso.refl _))
203
+
204
+ @[simp, reassoc]
205
+ lemma pi.reindex_hom_π (b : β) : (pi.reindex ε f).hom ≫ pi.π f (ε b) = pi.π (f ∘ ε) b :=
206
+ begin
207
+ dsimp [pi.reindex],
208
+ simp only [has_limit.iso_of_equivalence_hom_π, discrete.nat_iso_inv_app,
209
+ equivalence.equivalence_mk'_counit, discrete.equivalence_counit_iso, discrete.nat_iso_hom_app,
210
+ eq_to_iso.hom, eq_to_hom_map],
211
+ dsimp,
212
+ simpa using limit.w (discrete.functor (f ∘ ε)) (eq_to_hom (ε.symm_apply_apply b)),
213
+ end
214
+
215
+ @[simp, reassoc]
216
+ lemma pi.reindex_inv_π (b : β) : (pi.reindex ε f).inv ≫ pi.π (f ∘ ε) b = pi.π f (ε b) :=
217
+ by simp [iso.inv_comp_eq]
218
+
219
+ end
220
+
221
+ section
222
+ variables [has_coproduct f] [has_coproduct (f ∘ ε)]
223
+
224
+ /-- Reindex a categorical coproduct via an equivalence of the index types. -/
225
+ def sigma.reindex : sigma_obj (f ∘ ε) ≅ sigma_obj f :=
226
+ has_colimit.iso_of_equivalence (discrete.equivalence ε) (discrete.nat_iso (λ i, iso.refl _))
227
+
228
+ @[simp, reassoc]
229
+ lemma sigma.ι_reindex_hom (b : β) : sigma.ι (f ∘ ε) b ≫ (sigma.reindex ε f).hom = sigma.ι f (ε b) :=
230
+ begin
231
+ dsimp [sigma.reindex],
232
+ simp only [has_colimit.iso_of_equivalence_hom_π, equivalence.equivalence_mk'_unit,
233
+ discrete.equivalence_unit_iso, discrete.nat_iso_hom_app, eq_to_iso.hom, eq_to_hom_map,
234
+ discrete.nat_iso_inv_app],
235
+ dsimp,
236
+ simp [←colimit.w (discrete.functor f) (eq_to_hom (ε.apply_symm_apply (ε b)))],
237
+ end
238
+
239
+ @[simp, reassoc]
240
+ lemma sigma.ι_reindex_inv (b : β) : sigma.ι f (ε b) ≫ (sigma.reindex ε f).inv = sigma.ι (f ∘ ε) b :=
241
+ by simp [iso.comp_inv_eq]
242
+
243
+ end
244
+
245
+ end reindex
246
+
194
247
end category_theory.limits
0 commit comments