feat(hott) formalize book lemma 9.9.1: essentially surjective functors induce faithful functors in the functor category

This commit is contained in:
Jakob von Raumer 2016-07-15 18:20:42 +02:00 committed by Leonardo de Moura
parent 3de39200a4
commit 0ff8a96be1
3 changed files with 18 additions and 2 deletions

View file

@ -6,7 +6,7 @@ Authors: Floris van Doorn, Jakob von Raumer
Functor precategory and category
-/
import ..nat_trans ..category .opposite
import .opposite ..functor.attributes
open eq category is_trunc nat_trans iso is_equiv category.hom
@ -372,6 +372,15 @@ namespace functor
{ intro G H I η θ, reflexivity},
end
definition faithful_precomposition_of_essentially_surjective [instance]
{C D E} {H : C ⇒ D} [HH : essentially_surjective H] : faithful (precomposition_functor E H) :=
begin
intro F G γ δ Hγδ, apply nat_trans_eq, intro b,
induction HH b with Hb, induction Hb with a f,
refine naturality_iso_right γ f ⬝ _ ⬝ (naturality_iso_right δ f)⁻¹,
apply ap (λ x, _ ∘ natural_map x a ∘ _) Hγδ,
end
definition postcomposition_functor [constructor] {C D} (E) (F : C ⇒ D)
: C ^c E ⇒ D ^c E :=
begin

View file

@ -8,7 +8,7 @@ Attributes of functors (full, faithful, split essentially surjective, ...)
Adjoint functors, isomorphisms and equivalences have their own file
-/
import ..constructions.functor function arity
import .basic function arity
open eq functor trunc prod is_equiv iso equiv function is_trunc

View file

@ -47,6 +47,13 @@ namespace nat_trans
(g : d ⟶ d') : constant_functor C d ⟹ constant_functor C d' :=
mk (λc, g) (λc c' f, !id_comp_eq_comp_id)
open iso
definition naturality_iso_left (η : F ⟹ G) {a b : C} (f : a ≅ b) : η a = (G f)⁻¹ ∘ η b ∘ F f :=
by apply eq_inverse_comp_of_comp_eq; apply naturality
definition naturality_iso_right (η : F ⟹ G) {a b : C} (f : a ≅ b) : η b = G f ∘ η a ∘ (F f)⁻¹ :=
by refine _⁻¹ ⬝ !assoc⁻¹; apply comp_inverse_eq_of_eq_comp; apply naturality
definition nat_trans_mk_eq {η₁ η₂ : Π (a : C), hom (F a) (G a)}
(nat₁ : Π (a b : C) (f : hom a b), G f ∘ η₁ a = η₁ b ∘ F f)
(nat₂ : Π (a b : C) (f : hom a b), G f ∘ η₂ a = η₂ b ∘ F f)