Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit f48cbb1

Browse files
committed
feat(category_theory/limits): reindexing (co/bi)products (#14193)
Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent ca5930d commit f48cbb1

File tree

2 files changed

+67
-0
lines changed

2 files changed

+67
-0
lines changed

src/category_theory/limits/shapes/biproducts.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1395,6 +1395,20 @@ by { ext, simp, }
13951395

13961396
end
13971397

1398+
/-- Reindex a categorical biproduct via an equivalence of the index types. -/
1399+
@[simps]
1400+
def biproduct.reindex {β γ : Type v} [fintype β] [decidable_eq β] [decidable_eq γ]
1401+
(ε : β ≃ γ) (f : γ → C) [has_biproduct f] [has_biproduct (f ∘ ε)] : (⨁ (f ∘ ε)) ≅ (⨁ f) :=
1402+
{ hom := biproduct.desc (λ b, biproduct.ι f (ε b)),
1403+
inv := biproduct.lift (λ b, biproduct.π f (ε b)),
1404+
hom_inv_id' := by { ext b b', by_cases h : b = b', { subst h, simp, }, { simp [h], }, },
1405+
inv_hom_id' := begin
1406+
ext g g',
1407+
by_cases h : g = g';
1408+
simp [preadditive.sum_comp, preadditive.comp_sum, biproduct.ι_π, biproduct.ι_π_assoc, comp_dite,
1409+
equiv.apply_eq_iff_eq_symm_apply, finset.sum_dite_eq' finset.univ (ε.symm g') _, h],
1410+
end, }
1411+
13981412
/--
13991413
In a preadditive category, we can construct a binary biproduct for `X Y : C` from
14001414
any binary bicone `b` satisfying `total : b.fst ≫ b.inl + b.snd ≫ b.inr = 𝟙 b.X`.

src/category_theory/limits/shapes/products.lean

Lines changed: 53 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -191,4 +191,57 @@ abbreviation has_products := Π (J : Type v), has_limits_of_shape (discrete J) C
191191
/-- An abbreviation for `Π J, has_colimits_of_shape (discrete J) C` -/
192192
abbreviation has_coproducts := Π (J : Type v), has_colimits_of_shape (discrete J) C
193193

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+
194247
end category_theory.limits

0 commit comments

Comments
 (0)