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

Commit fcd0f11

Browse files
committed
feat(category_theory/flat_functor): Generalize results into algebraic categories. (#10735)
Also proves that the identity is flat, and compositions of flat functors are flat.
1 parent 0690542 commit fcd0f11

File tree

1 file changed

+79
-6
lines changed

1 file changed

+79
-6
lines changed

src/category_theory/flat_functors.lean

Lines changed: 79 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -83,6 +83,7 @@ end structured_arrow_cone
8383

8484
section representably_flat
8585
variables {C : Type u₁} [category.{v₁} C] {D : Type u₂} [category.{v₂} D]
86+
variables {E : Type u₃} [category.{v₃} E]
8687

8788
/--
8889
A functor `F : C ⥤ D` is representably-flat functor if the comma category `(X/F)`
@@ -93,6 +94,74 @@ class representably_flat (F : C ⥤ D) : Prop :=
9394

9495
attribute [instance] representably_flat.cofiltered
9596

97+
local attribute [instance] is_cofiltered.nonempty
98+
99+
instance representably_flat.id : representably_flat (𝟭 C) :=
100+
begin
101+
constructor,
102+
intro X,
103+
haveI : nonempty (structured_arrow X (𝟭 C)) := ⟨structured_arrow.mk (𝟙 _)⟩,
104+
suffices : is_cofiltered_or_empty (structured_arrow X (𝟭 C)),
105+
{ resetI, constructor },
106+
constructor,
107+
{ intros Y Z,
108+
use structured_arrow.mk (𝟙 _),
109+
use structured_arrow.hom_mk Y.hom (by erw [functor.id_map, category.id_comp]),
110+
use structured_arrow.hom_mk Z.hom (by erw [functor.id_map, category.id_comp]) },
111+
{ intros Y Z f g,
112+
use structured_arrow.mk (𝟙 _),
113+
use structured_arrow.hom_mk Y.hom (by erw [functor.id_map, category.id_comp]),
114+
ext,
115+
transitivity Z.hom; simp }
116+
end
117+
118+
instance representably_flat.comp (F : C ⥤ D) (G : D ⥤ E)
119+
[representably_flat F] [representably_flat G] : representably_flat (F ⋙ G) :=
120+
begin
121+
constructor,
122+
intro X,
123+
haveI : nonempty (structured_arrow X (F ⋙ G)),
124+
{ have f₁ : structured_arrow X G := nonempty.some infer_instance,
125+
have f₂ : structured_arrow f₁.right F := nonempty.some infer_instance,
126+
exact ⟨structured_arrow.mk (f₁.hom ≫ G.map f₂.hom)⟩ },
127+
suffices : is_cofiltered_or_empty (structured_arrow X (F ⋙ G)),
128+
{ resetI, constructor },
129+
constructor,
130+
{ intros Y Z,
131+
let W := @is_cofiltered.min (structured_arrow X G) _ _
132+
(structured_arrow.mk Y.hom) (structured_arrow.mk Z.hom),
133+
let Y' : W ⟶ _ := is_cofiltered.min_to_left _ _,
134+
let Z' : W ⟶ _ := is_cofiltered.min_to_right _ _,
135+
136+
let W' := @is_cofiltered.min (structured_arrow W.right F) _ _
137+
(structured_arrow.mk Y'.right) (structured_arrow.mk Z'.right),
138+
let Y'' : W' ⟶ _ := is_cofiltered.min_to_left _ _,
139+
let Z'' : W' ⟶ _ := is_cofiltered.min_to_right _ _,
140+
141+
use structured_arrow.mk (W.hom ≫ G.map W'.hom),
142+
use structured_arrow.hom_mk Y''.right (by simp [← G.map_comp]),
143+
use structured_arrow.hom_mk Z''.right (by simp [← G.map_comp]) },
144+
{ intros Y Z f g,
145+
let W := @is_cofiltered.eq (structured_arrow X G) _ _
146+
(structured_arrow.mk Y.hom) (structured_arrow.mk Z.hom)
147+
(structured_arrow.hom_mk (F.map f.right) (structured_arrow.w f))
148+
(structured_arrow.hom_mk (F.map g.right) (structured_arrow.w g)),
149+
let h : W ⟶ _ := is_cofiltered.eq_hom _ _,
150+
let h_cond : h ≫ _ = h ≫ _ := is_cofiltered.eq_condition _ _,
151+
152+
let W' := @is_cofiltered.eq (structured_arrow W.right F) _ _
153+
(structured_arrow.mk h.right) (structured_arrow.mk (h.right ≫ F.map f.right))
154+
(structured_arrow.hom_mk f.right rfl)
155+
(structured_arrow.hom_mk g.right (congr_arg comma_morphism.right h_cond).symm),
156+
let h' : W' ⟶ _ := is_cofiltered.eq_hom _ _,
157+
let h'_cond : h' ≫ _ = h' ≫ _ := is_cofiltered.eq_condition _ _,
158+
159+
use structured_arrow.mk (W.hom ≫ G.map W'.hom),
160+
use structured_arrow.hom_mk h'.right (by simp [← G.map_comp]),
161+
ext,
162+
exact (congr_arg comma_morphism.right h'_cond : _) }
163+
end
164+
96165
end representably_flat
97166

98167
section has_limit
@@ -207,14 +276,14 @@ end has_limit
207276

208277

209278
section small_category
210-
variables {C D : Type u₁} [small_category C] [small_category D]
279+
variables {C D : Type u₁} [small_category C] [small_category D] (E : Type u₂) [category.{u₁} E]
211280

212281
/--
213282
(Implementation)
214283
The evaluation of `Lan F` at `X` is the colimit over the costructured arrows over `X`.
215284
-/
216285
noncomputable
217-
def Lan_evaluation_iso_colim (E : Type u₂) [category.{u₁} E] (F : C ⥤ D) (X : D)
286+
def Lan_evaluation_iso_colim (F : C ⥤ D) (X : D)
218287
[∀ (X : D), has_colimits_of_shape (costructured_arrow F X) E] :
219288
Lan F ⋙ (evaluation D E).obj X ≅
220289
((whiskering_left _ _ E).obj (costructured_arrow.proj F X)) ⋙ colim :=
@@ -232,30 +301,34 @@ begin
232301
rw [costructured_arrow.map_mk, category.id_comp, costructured_arrow.mk]
233302
end
234303

304+
variables [concrete_category.{u₁} E] [has_limits E] [has_colimits E]
305+
variables [reflects_limits (forget E)] [preserves_filtered_colimits (forget E)]
306+
variables [preserves_limits (forget E)]
307+
235308
/--
236309
If `F : C ⥤ D` is a representably flat functor between small categories, then the functor
237310
`Lan F.op` that takes presheaves over `C` to presheaves over `D` preserves finite limits.
238311
-/
239312
noncomputable
240313
instance Lan_preserves_finite_limits_of_flat (F : C ⥤ D) [representably_flat F] :
241-
preserves_finite_limits (Lan F.op : _ ⥤ (Dᵒᵖ ⥤ Type u₁)) :=
314+
preserves_finite_limits (Lan F.op : _ ⥤ (Dᵒᵖ ⥤ E)) :=
242315
⟨λ J _ _, begin
243316
resetI,
244-
apply preserves_limits_of_shape_of_evaluation (Lan F.op : (Cᵒᵖ ⥤ Type u₁) ⥤ (Dᵒᵖ ⥤ Type u₁)) J,
317+
apply preserves_limits_of_shape_of_evaluation (Lan F.op : (Cᵒᵖ ⥤ E) ⥤ (Dᵒᵖ ⥤ E)) J,
245318
intro K,
246319
haveI : is_filtered (costructured_arrow F.op K) :=
247320
is_filtered.of_equivalence (structured_arrow_op_equivalence F (unop K)),
248321
exact preserves_limits_of_shape_of_nat_iso (Lan_evaluation_iso_colim _ _ _).symm
249322
end
250323

251324
instance Lan_flat_of_flat (F : C ⥤ D) [representably_flat F] :
252-
representably_flat (Lan F.op : _ ⥤ (Dᵒᵖ ⥤ Type u₁)) := flat_of_preserves_finite_limits _
325+
representably_flat (Lan F.op : _ ⥤ (Dᵒᵖ ⥤ E)) := flat_of_preserves_finite_limits _
253326

254327
variable [has_finite_limits C]
255328

256329
noncomputable
257330
instance Lan_preserves_finite_limits_of_preserves_finite_limits (F : C ⥤ D)
258-
[preserves_finite_limits F] : preserves_finite_limits (Lan F.op : _ ⥤ (Dᵒᵖ ⥤ Type u₁)) :=
331+
[preserves_finite_limits F] : preserves_finite_limits (Lan F.op : _ ⥤ (Dᵒᵖ ⥤ E)) :=
259332
begin
260333
haveI := flat_of_preserves_finite_limits F,
261334
apply_instance

0 commit comments

Comments
 (0)