@@ -43,9 +43,6 @@ class PreservesFiniteLimits (F : C ⥤ D) : Prop where
43
43
44
44
attribute [instance] PreservesFiniteLimits.preservesFiniteLimits
45
45
46
- instance (F : C ⥤ D) : Subsingleton (PreservesFiniteLimits F) :=
47
- ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩
48
-
49
46
/-- Preserving finite limits also implies preserving limits over finite shapes in higher universes,
50
47
though through a noncomputable instance. -/
51
48
instance (priority := 100 ) preservesLimitsOfShapeOfPreservesFiniteLimits (F : C ⥤ D)
@@ -93,29 +90,25 @@ lemma preservesFiniteLimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesFi
93
90
94
91
/- Porting note: adding this class because quantified classes don't behave well
95
92
https://github.com/leanprover-community/mathlib4/pull/2764 -/
96
- /-- A functor `F` preserves finite products if it preserves all from `Discrete J`
97
- for `Fintype J` -/
93
+ /-- A functor `F` preserves finite products if it preserves all from `Discrete J` for `Finite J`.
94
+ We require this for `J = Fin n` in the definition,
95
+ then generalize to `J : Type u` in the instance. -/
98
96
class PreservesFiniteProducts (F : C ⥤ D) : Prop where
99
- preserves : ∀ (J : Type ) [Fintype J], PreservesLimitsOfShape (Discrete J) F
100
-
101
- attribute [instance] PreservesFiniteProducts.preserves
102
-
103
- instance (F : C ⥤ D) : Subsingleton (PreservesFiniteProducts F) :=
104
- ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩
97
+ preserves : ∀ n, PreservesLimitsOfShape (Discrete (Fin n)) F
105
98
106
99
instance (priority := 100 ) (F : C ⥤ D) (J : Type u) [Finite J]
107
100
[PreservesFiniteProducts F] : PreservesLimitsOfShape (Discrete J) F := by
108
- apply Nonempty.some
109
101
obtain ⟨n, ⟨e⟩⟩ := Finite.exists_equiv_fin J
110
- exact ⟨preservesLimitsOfShape_of_equiv (Discrete.equivalence e.symm) F⟩
102
+ have := PreservesFiniteProducts.preserves (F := F) n
103
+ exact preservesLimitsOfShape_of_equiv (Discrete.equivalence e.symm) F
111
104
112
105
instance comp_preservesFiniteProducts (F : C ⥤ D) (G : D ⥤ E)
113
106
[PreservesFiniteProducts F] [PreservesFiniteProducts G] :
114
107
PreservesFiniteProducts (F ⋙ G) where
115
- preserves _ _ := inferInstance
108
+ preserves _ := inferInstance
116
109
117
110
instance (F : C ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteProducts F where
118
- preserves _ _ := inferInstance
111
+ preserves _ := inferInstance
119
112
120
113
/--
121
114
A functor is said to reflect finite limits, if it reflects all limits of shape `J`,
@@ -127,20 +120,20 @@ class ReflectsFiniteLimits (F : C ⥤ D) : Prop where
127
120
128
121
attribute [instance] ReflectsFiniteLimits.reflects
129
122
130
- instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteLimits F) :=
131
- ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩
132
-
133
123
/- Similarly to preserving finite products, quantified classes don't behave well. -/
134
124
/--
135
- A functor `F` preserves finite products if it reflects limits of shape `Discrete J` for finite `J`
125
+ A functor `F` preserves finite products if it reflects limits of shape `Discrete J` for finite `J`.
126
+ We require this for `J = Fin n` in the definition,
127
+ then generalize to `J : Type u` in the instance.
136
128
-/
137
129
class ReflectsFiniteProducts (F : C ⥤ D) : Prop where
138
- reflects : ∀ (J : Type ) [Fintype J] , ReflectsLimitsOfShape (Discrete J ) F
130
+ reflects : ∀ n , ReflectsLimitsOfShape (Discrete (Fin n) ) F
139
131
140
- attribute [instance] ReflectsFiniteProducts.reflects
141
-
142
- instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteProducts F) :=
143
- ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩
132
+ instance (priority := 100 ) (F : C ⥤ D) [ReflectsFiniteProducts F] (J : Type u) [Finite J] :
133
+ ReflectsLimitsOfShape (Discrete J) F :=
134
+ let ⟨n, ⟨e⟩⟩ := Finite.exists_equiv_fin J
135
+ have := ReflectsFiniteProducts.reflects (F := F) n
136
+ reflectsLimitsOfShape_of_equiv (Discrete.equivalence e.symm) _
144
137
145
138
-- This is a dangerous instance as it has unbound universe variables.
146
139
/-- If we reflect limits of some arbitrary size, then we reflect all finite limits. -/
@@ -174,7 +167,7 @@ finite products.
174
167
-/
175
168
lemma preservesFiniteProducts_of_reflects_of_preserves (F : C ⥤ D) (G : D ⥤ E)
176
169
[PreservesFiniteProducts (F ⋙ G)] [ReflectsFiniteProducts G] : PreservesFiniteProducts F where
177
- preserves _ _ := preservesLimitsOfShape_of_reflects_of_preserves F G
170
+ preserves _ := preservesLimitsOfShape_of_reflects_of_preserves F G
178
171
179
172
instance reflectsFiniteLimits_of_reflectsIsomorphisms (F : C ⥤ D)
180
173
[F.ReflectsIsomorphisms] [HasFiniteLimits C] [PreservesFiniteLimits F] :
@@ -184,15 +177,15 @@ instance reflectsFiniteLimits_of_reflectsIsomorphisms (F : C ⥤ D)
184
177
instance reflectsFiniteProducts_of_reflectsIsomorphisms (F : C ⥤ D)
185
178
[F.ReflectsIsomorphisms] [HasFiniteProducts C] [PreservesFiniteProducts F] :
186
179
ReflectsFiniteProducts F where
187
- reflects _ _ := reflectsLimitsOfShape_of_reflectsIsomorphisms
180
+ reflects _ := reflectsLimitsOfShape_of_reflectsIsomorphisms
188
181
189
182
instance comp_reflectsFiniteProducts (F : C ⥤ D) (G : D ⥤ E)
190
183
[ReflectsFiniteProducts F] [ReflectsFiniteProducts G] :
191
184
ReflectsFiniteProducts (F ⋙ G) where
192
- reflects _ _ := inferInstance
185
+ reflects _ := inferInstance
193
186
194
187
instance (F : C ⥤ D) [ReflectsFiniteLimits F] : ReflectsFiniteProducts F where
195
- reflects _ _ := inferInstance
188
+ reflects _ := inferInstance
196
189
197
190
/-- A functor is said to preserve finite colimits, if it preserves all colimits of
198
191
shape `J`, where `J : Type` is a finite category.
@@ -204,9 +197,6 @@ class PreservesFiniteColimits (F : C ⥤ D) : Prop where
204
197
205
198
attribute [instance] PreservesFiniteColimits.preservesFiniteColimits
206
199
207
- instance (F : C ⥤ D) : Subsingleton (PreservesFiniteColimits F) :=
208
- ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩
209
-
210
200
/--
211
201
Preserving finite colimits also implies preserving colimits over finite shapes in higher
212
202
universes.
@@ -257,44 +247,36 @@ lemma preservesFiniteColimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [Preserves
257
247
258
248
/- Porting note: adding this class because quantified classes don't behave well
259
249
https://github.com/leanprover-community/mathlib4/pull/2764 -/
260
- /-- A functor `F` preserves finite products if it preserves all from `Discrete J`
261
- for `Fintype J` -/
250
+ /-- A functor `F` preserves finite products if it preserves all from `Discrete J` for `Fintype J`.
251
+ We require this for `J = Fin n` in the definition,
252
+ then generalize to `J : Type u` in the instance. -/
262
253
class PreservesFiniteCoproducts (F : C ⥤ D) : Prop where
263
- /-- preservation of colimits indexed by `Discrete J` when `[Fintype J]` -/
264
- preserves : ∀ (J : Type ) [Fintype J], PreservesColimitsOfShape (Discrete J) F
265
-
266
- attribute [instance] PreservesFiniteCoproducts.preserves
267
-
268
- instance (F : C ⥤ D) : Subsingleton (PreservesFiniteCoproducts F) :=
269
- ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩
254
+ /-- preservation of colimits indexed by `Discrete (Fin n)`. -/
255
+ preserves : ∀ n, PreservesColimitsOfShape (Discrete (Fin n)) F
270
256
271
257
instance (priority := 100 ) (F : C ⥤ D) (J : Type u) [Finite J]
272
- [PreservesFiniteCoproducts F] : PreservesColimitsOfShape (Discrete J) F := by
273
- apply Nonempty.some
274
- obtain ⟨n, ⟨e⟩⟩ := Finite.exists_equiv_fin J
275
- exact ⟨ preservesColimitsOfShape_of_equiv (Discrete.equivalence e.symm) F⟩
258
+ [PreservesFiniteCoproducts F] : PreservesColimitsOfShape (Discrete J) F :=
259
+ let ⟨n, ⟨e⟩⟩ := Finite.exists_equiv_fin J
260
+ have := PreservesFiniteCoproducts.preserves (F := F) n
261
+ preservesColimitsOfShape_of_equiv (Discrete.equivalence e.symm) F
276
262
277
263
instance comp_preservesFiniteCoproducts (F : C ⥤ D) (G : D ⥤ E)
278
264
[PreservesFiniteCoproducts F] [PreservesFiniteCoproducts G] :
279
265
PreservesFiniteCoproducts (F ⋙ G) where
280
- preserves _ _ := inferInstance
266
+ preserves _ := inferInstance
281
267
282
268
instance (F : C ⥤ D) [PreservesFiniteColimits F] : PreservesFiniteCoproducts F where
283
- preserves _ _ := inferInstance
269
+ preserves _ := inferInstance
284
270
285
271
/--
286
272
A functor is said to reflect finite colimits, if it reflects all colimits of shape `J`,
287
273
where `J : Type` is a finite category.
288
274
-/
289
275
class ReflectsFiniteColimits (F : C ⥤ D) : Prop where
290
- reflects : ∀ (J : Type ) [SmallCategory J] [FinCategory J], ReflectsColimitsOfShape J F := by
291
- infer_instance
276
+ [reflects : ∀ (J : Type ) [SmallCategory J] [FinCategory J], ReflectsColimitsOfShape J F]
292
277
293
278
attribute [instance] ReflectsFiniteColimits.reflects
294
279
295
- instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteColimits F) :=
296
- ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩
297
-
298
280
-- This is a dangerous instance as it has unbound universe variables.
299
281
/-- If we reflect colimits of some arbitrary size, then we reflect all finite colimits. -/
300
282
lemma ReflectsColimitsOfSize.reflectsFiniteColimits
@@ -315,16 +297,20 @@ instance (priority := 120) (F : C ⥤ D)
315
297
316
298
/- Similarly to preserving finite coproducts, quantified classes don't behave well. -/
317
299
/--
318
- A functor `F` preserves finite coproducts if it reflects colimits of shape `Discrete J` for
319
- finite `J`
300
+ A functor `F` preserves finite coproducts if it reflects colimits of shape `Discrete J`
301
+ for finite `J`.
302
+
303
+ We require this for `J = Fin n` in the definition,
304
+ then generalize to `J : Type u` in the instance.
320
305
-/
321
306
class ReflectsFiniteCoproducts (F : C ⥤ D) : Prop where
322
- reflects : ∀ (J : Type ) [Fintype J], ReflectsColimitsOfShape (Discrete J) F
323
-
324
- attribute [instance] ReflectsFiniteCoproducts.reflects
307
+ reflects : ∀ n, ReflectsColimitsOfShape (Discrete (Fin n)) F
325
308
326
- instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteCoproducts F) :=
327
- ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩
309
+ instance (priority := 100 ) (F : C ⥤ D) [ReflectsFiniteCoproducts F] (J : Type u) [Finite J] :
310
+ ReflectsColimitsOfShape (Discrete J) F :=
311
+ let ⟨n, ⟨e⟩⟩ := Finite.exists_equiv_fin J
312
+ have := ReflectsFiniteCoproducts.reflects (F := F) n
313
+ reflectsColimitsOfShape_of_equiv (Discrete.equivalence e.symm) _
328
314
329
315
/--
330
316
If `F ⋙ G` preserves finite colimits and `G` reflects finite colimits, then `F` preserves finite
@@ -341,7 +327,7 @@ finite coproducts.
341
327
lemma preservesFiniteCoproducts_of_reflects_of_preserves (F : C ⥤ D) (G : D ⥤ E)
342
328
[PreservesFiniteCoproducts (F ⋙ G)] [ReflectsFiniteCoproducts G] :
343
329
PreservesFiniteCoproducts F where
344
- preserves _ _ := preservesColimitsOfShape_of_reflects_of_preserves F G
330
+ preserves _ := preservesColimitsOfShape_of_reflects_of_preserves F G
345
331
346
332
instance reflectsFiniteColimitsOfReflectsIsomorphisms (F : C ⥤ D)
347
333
[F.ReflectsIsomorphisms] [HasFiniteColimits C] [PreservesFiniteColimits F] :
@@ -351,14 +337,14 @@ instance reflectsFiniteColimitsOfReflectsIsomorphisms (F : C ⥤ D)
351
337
instance reflectsFiniteCoproductsOfReflectsIsomorphisms (F : C ⥤ D)
352
338
[F.ReflectsIsomorphisms] [HasFiniteCoproducts C] [PreservesFiniteCoproducts F] :
353
339
ReflectsFiniteCoproducts F where
354
- reflects _ _ := reflectsColimitsOfShape_of_reflectsIsomorphisms
340
+ reflects _ := reflectsColimitsOfShape_of_reflectsIsomorphisms
355
341
356
342
instance comp_reflectsFiniteCoproducts (F : C ⥤ D) (G : D ⥤ E)
357
343
[ReflectsFiniteCoproducts F] [ReflectsFiniteCoproducts G] :
358
344
ReflectsFiniteCoproducts (F ⋙ G) where
359
- reflects _ _ := inferInstance
345
+ reflects _ := inferInstance
360
346
361
347
instance (F : C ⥤ D) [ReflectsFiniteColimits F] : ReflectsFiniteCoproducts F where
362
- reflects _ _ := inferInstance
348
+ reflects _ := inferInstance
363
349
364
350
end CategoryTheory.Limits
0 commit comments