Skip to content

Commit

Permalink
chore: future-proof proofs in Category/Ring/Limits (#11839)
Browse files Browse the repository at this point in the history
These proofs will become slow after leanprover/lean4#3807, but with these changes there is no regression. 🤷‍♀️

Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
  • Loading branch information
2 people authored and atarnoam committed Apr 16, 2024
1 parent e4710b2 commit a2d22f6
Showing 1 changed file with 15 additions and 22 deletions.
37 changes: 15 additions & 22 deletions Mathlib/Algebra/Category/Ring/Limits.lean
Original file line number Diff line number Diff line change
Expand Up @@ -388,16 +388,16 @@ All we need to do is notice that the limit point has a `Ring` instance available
and then reuse the existing limit.
-/
instance : CreatesLimit F (forget₂ RingCat.{u} SemiRingCat.{u}) :=
letI : ReflectsIsomorphisms (forget₂ RingCat SemiRingCat) :=
have : ReflectsIsomorphisms (forget₂ RingCat SemiRingCat) :=
CategoryTheory.reflectsIsomorphisms_forget₂ _ _
letI : Small.{u} (Functor.sections ((F ⋙ forget₂ _ SemiRingCat) ⋙ forget _)) :=
have : Small.{u} (Functor.sections ((F ⋙ forget₂ _ SemiRingCat) ⋙ forget _)) :=
inferInstanceAs <| Small.{u} (Functor.sections (F ⋙ forget _))
letI c : Cone F :=
let c : Cone F :=
{ pt := RingCat.of (Types.Small.limitCone (F ⋙ forget _)).pt
π :=
{ app := fun x => SemiRingCat.ofHom _
naturality := (SemiRingCat.HasLimits.limitCone
(F ⋙ forget₂ RingCat SemiRingCat.{u})).π.naturality } }
{ app := fun x => ofHom <| SemiRingCat.limitπRingHom.{v, u} (F ⋙ forget₂ _ SemiRingCat) x
naturality := fun _ _ f => RingHom.coe_inj
((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } }
createsLimitOfReflectsIso fun c' t =>
{ liftedCone := c
validLift := by apply IsLimit.uniqueUpToIso (SemiRingCat.HasLimits.limitConeIsLimit _) t
Expand Down Expand Up @@ -546,28 +546,21 @@ instance :
but it seems this would introduce additional identity morphisms in `limit.π`.
-/
-- Porting note: need to add these instances manually
letI : ReflectsIsomorphisms (forget₂ CommRingCat.{u} RingCat.{u}) :=
have : ReflectsIsomorphisms (forget₂ CommRingCat.{u} RingCat.{u}) :=
CategoryTheory.reflectsIsomorphisms_forget₂ _ _
letI : Small.{u}
(Functor.sections ((F ⋙ forget₂ CommRingCat RingCat) ⋙ forget RingCat)) :=
have : Small.{u} (Functor.sections ((F ⋙ forget₂ CommRingCat RingCat) ⋙ forget RingCat)) :=
inferInstanceAs <| Small.{u} (Functor.sections (F ⋙ forget _))
letI : Small.{u}
(Functor.sections ((F ⋙ forget₂ _ RingCat ⋙ forget₂ _ SemiRingCat) ⋙ forget _)) :=
let F' := F ⋙ forget₂ CommRingCat.{u} RingCat.{u} ⋙ forget₂ RingCat.{u} SemiRingCat.{u}
have : Small.{u} (Functor.sections (F' ⋙ forget _)) :=
inferInstanceAs <| Small.{u} (F ⋙ forget _).sections
letI c : Cone F :=
let c : Cone F :=
{ pt := CommRingCat.of (Types.Small.limitCone (F ⋙ forget _)).pt
π :=
{ app := fun x => ofHom <|
SemiRingCat.limitπRingHom.{v, u}
(F ⋙ forget₂ CommRingCat.{u} RingCat.{u} ⋙
forget₂ RingCat.{u} SemiRingCat.{u}) x
{ app := fun x => ofHom <| SemiRingCat.limitπRingHom.{v, u} F' x
naturality :=
(SemiRingCat.HasLimits.limitCone.{v, u}
(F ⋙
forget₂ _ RingCat.{u} ⋙
forget₂ _ SemiRingCat.{u})).π.naturality } }
createsLimitOfReflectsIso
fun _ t =>
fun _ _ f => RingHom.coe_inj
((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } }
createsLimitOfReflectsIso fun _ t =>
{ liftedCone := c
validLift := IsLimit.uniqueUpToIso (RingCat.limitConeIsLimit.{v, u} _) t
makesLimit :=
Expand Down

0 comments on commit a2d22f6

Please sign in to comment.