@@ -49,7 +49,7 @@ instance semiringObj (F : J ⥤ SemiRingCatMax.{v, u}) (j) :
49
49
set_option linter.uppercaseLean3 false in
50
50
#align SemiRing.semiring_obj SemiRingCat.semiringObj
51
51
52
- /-- The flat sections of a functor into `SemiRing ` form a subsemiring of all sections.
52
+ /-- The flat sections of a functor into `SemiRingCat ` form a subsemiring of all sections.
53
53
-/
54
54
def sectionsSubsemiring (F : J ⥤ SemiRingCatMax.{v, u}) : Subsemiring.{max v u} (∀ j, F.obj j) :=
55
55
-- Porting note : if `f` and `g` were inlined, it does not compile
@@ -68,7 +68,7 @@ instance limitSemiring (F : J ⥤ SemiRingCatMax.{v, u}) :
68
68
set_option linter.uppercaseLean3 false in
69
69
#align SemiRing.limit_semiring SemiRingCat.limitSemiring
70
70
71
- /-- `limit.π (F ⋙ forget SemiRing ) j` as a `ring_hom `. -/
71
+ /-- `limit.π (F ⋙ forget SemiRingCat ) j` as a `RingHom `. -/
72
72
def limitπRingHom (F : J ⥤ SemiRingCatMax.{v, u}) (j) :
73
73
(Types.limitCone.{v, u} (F ⋙ forget SemiRingCat)).pt →+* (F ⋙ forget SemiRingCat).obj j :=
74
74
-- Porting note : if `f` and `g` were inlined, it does not compile
@@ -82,10 +82,10 @@ set_option linter.uppercaseLean3 false in
82
82
83
83
namespace HasLimits
84
84
85
- -- The next two definitions are used in the construction of `has_limits SemiRing `.
85
+ -- The next two definitions are used in the construction of `HasLimits SemiRingCat `.
86
86
-- After that, the limits should be constructed using the generic limits API,
87
- -- e.g. `limit F`, `limit.cone F`, and `limit.is_limit F`.
88
- /-- Construction of a limit cone in `SemiRing `.
87
+ -- e.g. `limit F`, `limit.cone F`, and `limit.isLimit F`.
88
+ /-- Construction of a limit cone in `SemiRingCat `.
89
89
(Internal use only; use the limits API.)
90
90
-/
91
91
def limitCone (F : J ⥤ SemiRingCatMax.{v, u}) : Cone F where
@@ -97,7 +97,7 @@ def limitCone (F : J ⥤ SemiRingCatMax.{v, u}) : Cone F where
97
97
set_option linter.uppercaseLean3 false in
98
98
#align SemiRing.has_limits.limit_cone SemiRingCat.HasLimits.limitCone
99
99
100
- /-- Witness that the limit cone in `SemiRing ` is a limit cone.
100
+ /-- Witness that the limit cone in `SemiRingCat ` is a limit cone.
101
101
(Internal use only; use the limits API.)
102
102
-/
103
103
def limitConeIsLimit (F : J ⥤ SemiRingCatMax.{v, u}) : IsLimit (limitCone F) := by
@@ -139,8 +139,8 @@ set_option linter.uppercaseLean3 false in
139
139
/-- The forgetful functor from semirings to additive commutative monoids preserves all limits.
140
140
-/
141
141
instance forget₂AddCommMonPreservesLimitsOfSize :
142
- PreservesLimitsOfSize.{v, v} (forget₂ SemiRingCat AddCommMonCat.{max v u})
143
- where preservesLimitsOfShape {_ _} :=
142
+ PreservesLimitsOfSize.{v, v} (forget₂ SemiRingCat AddCommMonCat.{max v u}) where
143
+ preservesLimitsOfShape {_ _} :=
144
144
{ preservesLimit := fun {F} =>
145
145
preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F)
146
146
(forget₂AddCommMonPreservesLimitsAux F) }
@@ -164,8 +164,8 @@ set_option linter.uppercaseLean3 false in
164
164
/-- The forgetful functor from semirings to monoids preserves all limits.
165
165
-/
166
166
instance forget₂MonPreservesLimitsOfSize :
167
- PreservesLimitsOfSize.{v, v} (forget₂ SemiRingCat MonCat.{max v u})
168
- where preservesLimitsOfShape {_ _} :=
167
+ PreservesLimitsOfSize.{v, v} (forget₂ SemiRingCat MonCat.{max v u}) where
168
+ preservesLimitsOfShape {_ _} :=
169
169
{ preservesLimit := fun {F} =>
170
170
preservesLimitOfPreservesLimitCone (limitConeIsLimit F)
171
171
(forget₂MonPreservesLimitsAux.{v, u} F) }
@@ -179,8 +179,9 @@ set_option linter.uppercaseLean3 false in
179
179
180
180
/-- The forgetful functor from semirings to types preserves all limits.
181
181
-/
182
- instance forgetPreservesLimitsOfSize : PreservesLimitsOfSize.{v, v} (forget SemiRingCat.{max v u})
183
- where preservesLimitsOfShape {_ _} :=
182
+ instance forgetPreservesLimitsOfSize :
183
+ PreservesLimitsOfSize.{v, v} (forget SemiRingCat.{max v u}) where
184
+ preservesLimitsOfShape {_ _} :=
184
185
{ preservesLimit := fun {F} =>
185
186
preservesLimitOfPreservesLimitCone (limitConeIsLimit F)
186
187
(Types.limitConeIsLimit.{v, u} (F ⋙ forget _)) }
@@ -217,15 +218,15 @@ instance limitCommSemiring (F : J ⥤ CommSemiRingCatMax.{v, u}) :
217
218
set_option linter.uppercaseLean3 false in
218
219
#align CommSemiRing.limit_comm_semiring CommSemiRingCat.limitCommSemiring
219
220
220
- /-- We show that the forgetful functor `CommSemiRing ⥤ SemiRing ` creates limits.
221
+ /-- We show that the forgetful functor `CommSemiRingCat ⥤ SemiRingCat ` creates limits.
221
222
222
- All we need to do is notice that the limit point has a `comm_semiring ` instance available,
223
+ All we need to do is notice that the limit point has a `CommSemiring ` instance available,
223
224
and then reuse the existing limit.
224
225
-/
225
226
instance (F : J ⥤ CommSemiRingCatMax.{v, u}) :
226
227
CreatesLimit F (forget₂ CommSemiRingCatMax.{v, u} SemiRingCatMax.{v, u}) :=
227
- -- Porting note : `CommSemiRing ⥤ Type` reflecting isomorphism is needed to make Lean see that
228
- -- `CommSemiRingCat ⥤ SemiRingCat` reflects isomorphism. `CommSemiRing ⥤ Type` reflecting
228
+ -- Porting note : `CommSemiRingCat ⥤ Type` reflecting isomorphism is needed to make Lean see that
229
+ -- `CommSemiRingCat ⥤ SemiRingCat` reflects isomorphism. `CommSemiRingCat ⥤ Type` reflecting
229
230
-- isomorphism is added manually since Lean can't see it, but even with this addition Lean can not
230
231
-- see `CommSemiRingCat ⥤ SemiRingCat` reflects isomorphism, so this instance is also added.
231
232
letI : ReflectsIsomorphisms (forget CommSemiRingCatMax.{v, u}) :=
@@ -253,7 +254,7 @@ instance (F : J ⥤ CommSemiRingCatMax.{v, u}) :
253
254
fun x y => Subtype.ext <| funext fun j => by exact (s.π.app j).map_add x y⟩)
254
255
fun s => rfl }
255
256
256
- /-- A choice of limit cone for a functor into `CommSemiRing `.
257
+ /-- A choice of limit cone for a functor into `CommSemiRingCat `.
257
258
(Generally, you'll just want to use `limit F`.)
258
259
-/
259
260
def limitCone (F : J ⥤ CommSemiRingCatMax.{v, u}) : Cone F :=
@@ -272,8 +273,7 @@ set_option linter.uppercaseLean3 false in
272
273
/- ./././Mathport/Syntax/Translate/Command.lean:322:38: unsupported irreducible non-definition -/
273
274
/-- The category of rings has all limits. -/
274
275
instance hasLimitsOfSize : HasLimitsOfSize.{v, v} CommSemiRingCatMax.{v, u} :=
275
- {
276
- has_limits_of_shape := fun _ _ =>
276
+ { has_limits_of_shape := fun _ _ =>
277
277
{ has_limit := fun F =>
278
278
hasLimit_of_created F (forget₂ CommSemiRingCat SemiRingCatMax.{v, u}) } }
279
279
set_option linter.uppercaseLean3 false in
@@ -287,8 +287,8 @@ set_option linter.uppercaseLean3 false in
287
287
/-- The forgetful functor from rings to semirings preserves all limits.
288
288
-/
289
289
instance forget₂SemiRingPreservesLimitsOfSize :
290
- PreservesLimitsOfSize.{v, v} (forget₂ CommSemiRingCat SemiRingCat.{max v u})
291
- where preservesLimitsOfShape {_ _} :=
290
+ PreservesLimitsOfSize.{v, v} (forget₂ CommSemiRingCat SemiRingCat.{max v u}) where
291
+ preservesLimitsOfShape {_ _} :=
292
292
{ preservesLimit := fun {F} =>
293
293
preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F)
294
294
(SemiRingCat.HasLimits.limitConeIsLimit _) }
@@ -305,8 +305,8 @@ set_option linter.uppercaseLean3 false in
305
305
types could have been computed instead as limits in the category of types.)
306
306
-/
307
307
instance forgetPreservesLimitsOfSize :
308
- PreservesLimitsOfSize.{v, v} (forget CommSemiRingCatMax.{v, u})
309
- where preservesLimitsOfShape {_ _} :=
308
+ PreservesLimitsOfSize.{v, v} (forget CommSemiRingCatMax.{v, u}) where
309
+ preservesLimitsOfShape {_ _} :=
310
310
{ preservesLimit := fun {F} =>
311
311
preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F)
312
312
(Types.limitConeIsLimit.{v, u} _) }
@@ -321,7 +321,7 @@ set_option linter.uppercaseLean3 false in
321
321
end CommSemiRingCat
322
322
323
323
-- Porting note: typemax hack to fix universe complaints
324
- /-- An alias for `Ring .{max u v}`, to deal around unification issues. -/
324
+ /-- An alias for `RingCat .{max u v}`, to deal around unification issues. -/
325
325
@[nolint checkUnivs]
326
326
abbrev RingCatMax. {u1, u2} := RingCat.{max u1 u2}
327
327
@@ -335,7 +335,7 @@ instance ringObj (F : J ⥤ RingCatMax.{v, u}) (j) : Ring ((F ⋙ forget RingCat
335
335
set_option linter.uppercaseLean3 false in
336
336
#align Ring.ring_obj RingCat.ringObj
337
337
338
- /-- The flat sections of a functor into `Ring ` form a subring of all sections.
338
+ /-- The flat sections of a functor into `RingCat ` form a subring of all sections.
339
339
-/
340
340
def sectionsSubring (F : J ⥤ RingCatMax.{v, u}) : Subring.{max v u} (∀ j, F.obj j) :=
341
341
letI f : J ⥤ AddGroupCat.{max v u} :=
@@ -354,13 +354,13 @@ instance limitRing (F : J ⥤ RingCatMax.{v, u}) :
354
354
set_option linter.uppercaseLean3 false in
355
355
#align Ring.limit_ring RingCat.limitRing
356
356
357
- /-- We show that the forgetful functor `CommRing ⥤ Ring ` creates limits.
357
+ /-- We show that the forgetful functor `CommRingCat ⥤ RingCat ` creates limits.
358
358
359
- All we need to do is notice that the limit point has a `ring ` instance available,
359
+ All we need to do is notice that the limit point has a `Ring ` instance available,
360
360
and then reuse the existing limit.
361
361
-/
362
362
instance (F : J ⥤ RingCatMax.{v, u}) :
363
- CreatesLimit F (forget₂ RingCatMax.{v, u} SemiRingCatMax.{v, u}) :=
363
+ CreatesLimit F (forget₂ RingCatMax.{v, u} SemiRingCatMax.{v, u}) :=
364
364
letI : ReflectsIsomorphisms (forget₂ RingCatMax SemiRingCatMax) :=
365
365
CategoryTheory.reflectsIsomorphisms_forget₂ _ _
366
366
letI c : Cone F :=
@@ -376,7 +376,7 @@ instance (F : J ⥤ RingCatMax.{v, u}) :
376
376
IsLimit.ofFaithful (forget₂ RingCat SemiRingCat.{max v u})
377
377
(by apply SemiRingCat.HasLimits.limitConeIsLimit _) (fun s => _) fun s => rfl }
378
378
379
- /-- A choice of limit cone for a functor into `Ring `.
379
+ /-- A choice of limit cone for a functor into `RingCat `.
380
380
(Generally, you'll just want to use `limit F`.)
381
381
-/
382
382
def limitCone (F : J ⥤ RingCatMax.{v, u}) : Cone F :=
@@ -409,8 +409,8 @@ set_option linter.uppercaseLean3 false in
409
409
/-- The forgetful functor from rings to semirings preserves all limits.
410
410
-/
411
411
instance forget₂SemiRingPreservesLimitsOfSize :
412
- PreservesLimitsOfSize.{v, v} (forget₂ RingCat SemiRingCat.{max v u})
413
- where preservesLimitsOfShape {_ _} :=
412
+ PreservesLimitsOfSize.{v, v} (forget₂ RingCat SemiRingCat.{max v u}) where
413
+ preservesLimitsOfShape {_ _} :=
414
414
{ preservesLimit := fun {F} =>
415
415
preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F)
416
416
(SemiRingCat.HasLimits.limitConeIsLimit.{v, u} _) }
@@ -435,8 +435,8 @@ set_option linter.uppercaseLean3 false in
435
435
/-- The forgetful functor from rings to additive commutative groups preserves all limits.
436
436
-/
437
437
instance forget₂AddCommGroupPreservesLimitsOfSize :
438
- PreservesLimitsOfSize.{v, v} (forget₂ RingCatMax.{v, u} AddCommGroupCat.{max v u})
439
- where preservesLimitsOfShape {_ _} :=
438
+ PreservesLimitsOfSize.{v, v} (forget₂ RingCatMax.{v, u} AddCommGroupCat.{max v u}) where
439
+ preservesLimitsOfShape {_ _} :=
440
440
{ preservesLimit := fun {F} =>
441
441
preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F)
442
442
(forget₂AddCommGroupPreservesLimitsAux F) }
@@ -452,8 +452,8 @@ set_option linter.uppercaseLean3 false in
452
452
/-- The forgetful functor from rings to types preserves all limits. (That is, the underlying
453
453
types could have been computed instead as limits in the category of types.)
454
454
-/
455
- instance forgetPreservesLimitsOfSize : PreservesLimitsOfSize.{v, v} (forget RingCatMax.{v, u})
456
- where preservesLimitsOfShape {_ _} :=
455
+ instance forgetPreservesLimitsOfSize : PreservesLimitsOfSize.{v, v} (forget RingCatMax.{v, u}) where
456
+ preservesLimitsOfShape {_ _} :=
457
457
{ preservesLimit := fun {F} =>
458
458
preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F)
459
459
(Types.limitConeIsLimit.{v, u} _) }
@@ -468,7 +468,7 @@ set_option linter.uppercaseLean3 false in
468
468
end RingCat
469
469
470
470
-- Porting note: typemax hack to fix universe complaints
471
- /-- An alias for `CommRing .{max u v}`, to deal around unification issues. -/
471
+ /-- An alias for `CommRingCat .{max u v}`, to deal around unification issues. -/
472
472
@[nolint checkUnivs]
473
473
abbrev CommRingCatMax. {u1, u2} := CommRingCat.{max u1 u2}
474
474
@@ -490,17 +490,17 @@ instance limitCommRing (F : J ⥤ CommRingCatMax.{v, u}) :
490
490
set_option linter.uppercaseLean3 false in
491
491
#align CommRing.limit_comm_ring CommRingCat.limitCommRing
492
492
493
- /-- We show that the forgetful functor `CommRing ⥤ Ring ` creates limits.
493
+ /-- We show that the forgetful functor `CommRingCat ⥤ RingCat ` creates limits.
494
494
495
- All we need to do is notice that the limit point has a `comm_ring ` instance available,
495
+ All we need to do is notice that the limit point has a `CommRing ` instance available,
496
496
and then reuse the existing limit.
497
497
-/
498
498
instance (F : J ⥤ CommRingCatMax.{v, u}) :
499
499
CreatesLimit F (forget₂ CommRingCatMax.{v, u} RingCatMax.{v, u}) :=
500
500
/-
501
501
A terse solution here would be
502
502
```
503
- creates_limit_of_fully_faithful_of_iso (CommRing .of (limit (F ⋙ forget _))) (iso .refl _)
503
+ createsLimitOfFullyFaithfulOfIso (CommRingCat .of (limit (F ⋙ forget _))) (Iso .refl _)
504
504
```
505
505
but it seems this would introduce additional identity morphisms in `limit.π`.
506
506
-/
@@ -531,7 +531,7 @@ instance (F : J ⥤ CommRingCatMax.{v, u}) :
531
531
(F ⋙ forget₂ CommRingCatMax.{v, u} RingCatMax.{v, u})).lift
532
532
((forget₂ _ RingCatMax.{v, u}).mapCone s)) fun _ => rfl }
533
533
534
- /-- A choice of limit cone for a functor into `CommRing `.
534
+ /-- A choice of limit cone for a functor into `CommRingCat `.
535
535
(Generally, you'll just want to use `limit F`.)
536
536
-/
537
537
def limitCone (F : J ⥤ CommRingCatMax.{v, u}) : Cone F :=
@@ -569,8 +569,8 @@ set_option linter.uppercaseLean3 false in
569
569
(That is, the underlying rings could have been computed instead as limits in the category of rings.)
570
570
-/
571
571
instance forget₂RingPreservesLimitsOfSize :
572
- PreservesLimitsOfSize.{v, v} (forget₂ CommRingCat RingCat.{max v u})
573
- where preservesLimitsOfShape {_ _} :=
572
+ PreservesLimitsOfSize.{v, v} (forget₂ CommRingCat RingCat.{max v u}) where
573
+ preservesLimitsOfShape {_ _} :=
574
574
{ preservesLimit := fun {F} =>
575
575
preservesLimitOfPreservesLimitCone.{v, v} (limitConeIsLimit.{v, u} F)
576
576
(RingCat.limitConeIsLimit.{v, u} _) }
@@ -595,8 +595,8 @@ set_option linter.uppercaseLean3 false in
595
595
in the category of commutative semirings.)
596
596
-/
597
597
instance forget₂CommSemiRingPreservesLimitsOfSize :
598
- PreservesLimitsOfSize.{v, v} (forget₂ CommRingCat CommSemiRingCat.{max v u})
599
- where preservesLimitsOfShape {_ _} :=
598
+ PreservesLimitsOfSize.{v, v} (forget₂ CommRingCat CommSemiRingCat.{max v u}) where
599
+ preservesLimitsOfShape {_ _} :=
600
600
{ preservesLimit := fun {F} =>
601
601
preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F)
602
602
(forget₂CommSemiRingPreservesLimitsAux.{v, u} F) }
@@ -612,8 +612,9 @@ set_option linter.uppercaseLean3 false in
612
612
/-- The forgetful functor from commutative rings to types preserves all limits.
613
613
(That is, the underlying types could have been computed instead as limits in the category of types.)
614
614
-/
615
- instance forgetPreservesLimitsOfSize : PreservesLimitsOfSize.{v, v} (forget CommRingCat.{max v u})
616
- where preservesLimitsOfShape {_ _} :=
615
+ instance forgetPreservesLimitsOfSize :
616
+ PreservesLimitsOfSize.{v, v} (forget CommRingCat.{max v u}) where
617
+ preservesLimitsOfShape {_ _} :=
617
618
{ preservesLimit := fun {F} =>
618
619
preservesLimitOfPreservesLimitCone.{v, v} (limitConeIsLimit.{v, u} F)
619
620
(Types.limitConeIsLimit.{v, u} _) }
0 commit comments