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

Commit a5ff45a

Browse files
committed
chore(category_theory/abelian): backport removal of abelian.has_finite_biproducts instance (#18740)
This backports a proposed removal of the `abelian.has_finite_biproducts` global instance, instead enabling it locally in the files that need it. The reason for removing it is that it triggers the ~~dreaded~~ leanprover/lean4#2055 during the simpNF linter in leanprover-community/mathlib4#2769, the mathlib4 port of `category_theory.abelian.basic`. This backport verifies that we won't run into further problems downstream if we (hopefully temporarily) remove these instances in mathlib4. Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent b17e272 commit a5ff45a

File tree

2 files changed

+12
-2
lines changed

2 files changed

+12
-2
lines changed

src/category_theory/abelian/basic.lean

Lines changed: 7 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -257,10 +257,15 @@ namespace category_theory.abelian
257257
variables {C : Type u} [category.{v} C] [abelian C]
258258

259259
/-- An abelian category has finite biproducts. -/
260-
@[priority 100]
261-
instance has_finite_biproducts : has_finite_biproducts C :=
260+
-- Porting note: this should be an instance,
261+
-- but triggers https://github.com/leanprover/lean4/issues/2055
262+
-- We set it as a local instance instead.
263+
-- @[priority 100] instance
264+
theorem has_finite_biproducts : has_finite_biproducts C :=
262265
limits.has_finite_biproducts.of_has_finite_products
263266

267+
local attribute [instance] has_finite_biproducts
268+
264269
@[priority 100]
265270
instance has_binary_biproducts : has_binary_biproducts C :=
266271
limits.has_binary_biproducts_of_finite_biproducts _

src/category_theory/abelian/opposite.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -25,6 +25,11 @@ variables (C : Type*) [category C] [abelian C]
2525
local attribute [instance]
2626
has_finite_limits_of_has_equalizers_and_finite_products
2727
has_finite_colimits_of_has_coequalizers_and_finite_coproducts
28+
-- Porting note:
29+
-- This should have been a global instance,
30+
-- but triggers https://github.com/leanprover/lean4/issues/2055
31+
-- when ported to mathlib4.
32+
abelian.has_finite_biproducts
2833

2934
instance : abelian Cᵒᵖ :=
3035
{ normal_mono_of_mono := λ X Y f m, by exactI

0 commit comments

Comments
 (0)