@@ -11,13 +11,18 @@ import logic.equiv.basic
11
11
12
12
We define typeclasses `full` and `faithful`, decorating functors.
13
13
14
- Use `F.map_injective` to retrieve the fact that `F.map` is injective when `[faithful F]`,
15
- and `F.preimage` to obtain preimages of morphisms when `[full F]`.
16
-
17
- We prove some basic "cancellation" lemmas for full and/or faithful functors.
18
-
19
- See `category_theory.equivalence` for the fact that a functor is an equivalence if and only if
20
- it is fully faithful and essentially surjective.
14
+ ## Main definitions and results
15
+ * Use `F.map_injective` to retrieve the fact that `F.map` is injective when `[faithful F]`.
16
+ * Similarly, `F.map_surjective` states that `F.map` is surjective when `[full F]`.
17
+ * Use `F.preimage` to obtain preimages of morphisms when `[full F]`.
18
+ * We prove some basic "cancellation" lemmas for full and/or faithful functors, as well as a
19
+ construction for "dividing" a functor by a faithful functor, see `faithful.div`.
20
+ * `full F` carries data, so definitional properties of the preimage can be used when using
21
+ `F.preimage`. To obtain an instance of `full F` non-constructively, you can use `full_of_exists`
22
+ and `full_of_surjective`.
23
+
24
+ See `category_theory.equivalence.of_fully_faithful_ess_surj` for the fact that a functor is an
25
+ equivalence if and only if it is fully faithful and essentially surjective.
21
26
22
27
-/
23
28
@@ -55,8 +60,7 @@ restate_axiom faithful.map_injective'
55
60
namespace functor
56
61
variables {X Y : C}
57
62
58
- lemma map_injective (F : C ⥤ D) [faithful F] :
59
- function.injective $ @functor.map _ _ _ _ F X Y :=
63
+ lemma map_injective (F : C ⥤ D) [faithful F] : function.injective $ @functor.map _ _ _ _ F X Y :=
60
64
faithful.map_injective F
61
65
62
66
lemma map_iso_injective (F : C ⥤ D) [faithful F] :
@@ -70,6 +74,19 @@ full.preimage.{v₁ v₂} f
70
74
F.map (preimage F f) = f :=
71
75
by unfold preimage; obviously
72
76
77
+ lemma map_surjective (F : C ⥤ D) [full F] : function.surjective (@functor.map _ _ _ _ F X Y) :=
78
+ λ f, ⟨F.preimage f, F.image_preimage f⟩
79
+
80
+ /-- Deduce that `F` is full from the existence of preimages, using choice. -/
81
+ noncomputable def full_of_exists (F : C ⥤ D)
82
+ (h : ∀ (X Y : C) (f : F.obj X ⟶ F.obj Y), ∃ p, F.map p = f) : full F :=
83
+ by { choose p hp using h, exact ⟨p, hp⟩ }
84
+
85
+ /-- Deduce that `F` is full from surjectivity of `F.map`, using choice. -/
86
+ noncomputable def full_of_surjective (F : C ⥤ D)
87
+ (h : ∀ (X Y : C), function.surjective (@functor.map _ _ _ _ F X Y)) : full F :=
88
+ full_of_exists _ h
89
+
73
90
end functor
74
91
75
92
section
0 commit comments