2015-04-25 04:20:59 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2015 Floris van Doorn. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Authors: Floris van Doorn, Jakob von Raumer
|
|
|
|
|
|
|
|
|
|
Functor precategory and category
|
|
|
|
|
-/
|
|
|
|
|
|
2016-07-15 16:20:42 +00:00
|
|
|
|
import .opposite ..functor.attributes
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
2015-10-02 23:54:27 +00:00
|
|
|
|
open eq category is_trunc nat_trans iso is_equiv category.hom
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
2015-10-02 23:54:27 +00:00
|
|
|
|
namespace functor
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
2016-02-25 20:26:20 +00:00
|
|
|
|
definition precategory_functor [instance] [constructor] (D C : Precategory)
|
2015-04-25 04:20:59 +00:00
|
|
|
|
: precategory (functor C D) :=
|
|
|
|
|
precategory.mk (λa b, nat_trans a b)
|
|
|
|
|
(λ a b c g f, nat_trans.compose g f)
|
|
|
|
|
(λ a, nat_trans.id)
|
|
|
|
|
(λ a b c d h g f, !nat_trans.assoc)
|
|
|
|
|
(λ a b f, !nat_trans.id_left)
|
|
|
|
|
(λ a b f, !nat_trans.id_right)
|
|
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
|
definition Precategory_functor [reducible] [constructor] (D C : Precategory) : Precategory :=
|
2015-04-25 04:20:59 +00:00
|
|
|
|
precategory.Mk (precategory_functor D C)
|
|
|
|
|
|
2015-10-22 22:41:55 +00:00
|
|
|
|
infixr ` ^c `:80 := Precategory_functor
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
|
|
|
|
section
|
|
|
|
|
/- we prove that if a natural transformation is pointwise an iso, then it is an iso -/
|
|
|
|
|
variables {C D : Precategory} {F G : C ⇒ D} (η : F ⟹ G) [iso : Π(a : C), is_iso (η a)]
|
|
|
|
|
include iso
|
|
|
|
|
|
2015-08-31 16:23:34 +00:00
|
|
|
|
definition nat_trans_inverse [constructor] : G ⟹ F :=
|
2015-04-25 04:20:59 +00:00
|
|
|
|
nat_trans.mk
|
|
|
|
|
(λc, (η c)⁻¹)
|
|
|
|
|
(λc d f,
|
2015-09-28 04:38:35 +00:00
|
|
|
|
abstract begin
|
2015-04-25 04:20:59 +00:00
|
|
|
|
apply comp_inverse_eq_of_eq_comp,
|
2015-05-03 05:22:31 +00:00
|
|
|
|
transitivity (natural_map η d)⁻¹ ∘ to_fun_hom G f ∘ natural_map η c,
|
|
|
|
|
{apply eq_inverse_comp_of_comp_eq, symmetry, apply naturality},
|
|
|
|
|
{apply assoc}
|
2015-09-28 04:38:35 +00:00
|
|
|
|
end end)
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
2015-09-28 04:38:35 +00:00
|
|
|
|
definition nat_trans_left_inverse : nat_trans_inverse η ∘n η = 1 :=
|
2015-04-25 04:20:59 +00:00
|
|
|
|
begin
|
2016-06-23 20:49:54 +00:00
|
|
|
|
fapply (apdt011 nat_trans.mk),
|
2015-04-25 04:20:59 +00:00
|
|
|
|
apply eq_of_homotopy, intro c, apply left_inverse,
|
2016-06-23 20:49:54 +00:00
|
|
|
|
apply eq_of_homotopy3, intros, apply is_set.elim
|
2015-04-25 04:20:59 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-09-28 04:38:35 +00:00
|
|
|
|
definition nat_trans_right_inverse : η ∘n nat_trans_inverse η = 1 :=
|
2015-04-25 04:20:59 +00:00
|
|
|
|
begin
|
2016-06-23 20:49:54 +00:00
|
|
|
|
fapply (apdt011 nat_trans.mk),
|
2015-04-25 04:20:59 +00:00
|
|
|
|
apply eq_of_homotopy, intro c, apply right_inverse,
|
2016-06-23 20:49:54 +00:00
|
|
|
|
apply eq_of_homotopy3, intros, apply is_set.elim
|
2015-04-25 04:20:59 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-10-23 05:12:34 +00:00
|
|
|
|
definition is_natural_iso [constructor] : is_iso η :=
|
2015-10-22 22:41:55 +00:00
|
|
|
|
is_iso.mk _ (nat_trans_left_inverse η) (nat_trans_right_inverse η)
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
2015-08-31 16:23:34 +00:00
|
|
|
|
variable (iso)
|
2015-10-23 05:12:34 +00:00
|
|
|
|
definition natural_iso.mk [constructor] : F ≅ G :=
|
|
|
|
|
iso.mk _ (is_natural_iso η)
|
2015-08-31 16:23:34 +00:00
|
|
|
|
|
2015-10-16 19:15:44 +00:00
|
|
|
|
omit iso
|
|
|
|
|
|
2015-10-23 05:12:34 +00:00
|
|
|
|
variables (F G)
|
2015-10-16 19:15:44 +00:00
|
|
|
|
definition is_natural_inverse (η : Πc, F c ≅ G c)
|
|
|
|
|
(nat : Π⦃a b : C⦄ (f : hom a b), G f ∘ to_hom (η a) = to_hom (η b) ∘ F f)
|
|
|
|
|
{a b : C} (f : hom a b) : F f ∘ to_inv (η a) = to_inv (η b) ∘ G f :=
|
|
|
|
|
let η' : F ⟹ G := nat_trans.mk (λc, to_hom (η c)) @nat in
|
|
|
|
|
naturality (nat_trans_inverse η') f
|
|
|
|
|
|
|
|
|
|
definition is_natural_inverse' (η₁ : Πc, F c ≅ G c) (η₂ : F ⟹ G) (p : η₁ ~ η₂)
|
|
|
|
|
{a b : C} (f : hom a b) : F f ∘ to_inv (η₁ a) = to_inv (η₁ b) ∘ G f :=
|
|
|
|
|
is_natural_inverse F G η₁ abstract λa b g, (p a)⁻¹ ▸ (p b)⁻¹ ▸ naturality η₂ g end f
|
|
|
|
|
|
2015-10-23 05:12:34 +00:00
|
|
|
|
variables {F G}
|
|
|
|
|
definition natural_iso.MK [constructor]
|
|
|
|
|
(η : Πc, F c ⟶ G c) (p : Π(c c' : C) (f : c ⟶ c'), G f ∘ η c = η c' ∘ F f)
|
|
|
|
|
(θ : Πc, G c ⟶ F c) (r : Πc, θ c ∘ η c = id) (q : Πc, η c ∘ θ c = id) : F ≅ G :=
|
|
|
|
|
iso.mk (nat_trans.mk η p) (@(is_natural_iso _) (λc, is_iso.mk (θ c) (r c) (q c)))
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
2015-10-23 05:12:34 +00:00
|
|
|
|
end
|
2015-10-16 19:15:44 +00:00
|
|
|
|
|
2015-04-25 04:20:59 +00:00
|
|
|
|
section
|
|
|
|
|
/- and conversely, if a natural transformation is an iso, it is componentwise an iso -/
|
2015-09-28 04:38:35 +00:00
|
|
|
|
variables {A B C D : Precategory} {F G : C ⇒ D} (η : hom F G) [isoη : is_iso η] (c : C)
|
2015-04-25 04:20:59 +00:00
|
|
|
|
include isoη
|
2015-10-16 19:15:44 +00:00
|
|
|
|
definition componentwise_is_iso [constructor] : is_iso (η c) :=
|
2015-04-25 04:20:59 +00:00
|
|
|
|
@is_iso.mk _ _ _ _ _ (natural_map η⁻¹ c) (ap010 natural_map ( left_inverse η) c)
|
|
|
|
|
(ap010 natural_map (right_inverse η) c)
|
|
|
|
|
|
|
|
|
|
local attribute componentwise_is_iso [instance]
|
|
|
|
|
|
2015-09-01 22:00:11 +00:00
|
|
|
|
variable {isoη}
|
2015-04-25 04:20:59 +00:00
|
|
|
|
definition natural_map_inverse : natural_map η⁻¹ c = (η c)⁻¹ := idp
|
2015-09-01 22:00:11 +00:00
|
|
|
|
variable [isoη]
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
|
|
|
|
definition naturality_iso {c c' : C} (f : c ⟶ c') : G f = η c' ∘ F f ∘ (η c)⁻¹ :=
|
|
|
|
|
calc
|
2015-05-03 05:22:31 +00:00
|
|
|
|
G f = (G f ∘ η c) ∘ (η c)⁻¹ : by rewrite comp_inverse_cancel_right
|
2015-04-25 04:20:59 +00:00
|
|
|
|
... = (η c' ∘ F f) ∘ (η c)⁻¹ : by rewrite naturality
|
2015-05-03 05:22:31 +00:00
|
|
|
|
... = η c' ∘ F f ∘ (η c)⁻¹ : by rewrite assoc
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
|
|
|
|
definition naturality_iso' {c c' : C} (f : c ⟶ c') : (η c')⁻¹ ∘ G f ∘ η c = F f :=
|
|
|
|
|
calc
|
|
|
|
|
(η c')⁻¹ ∘ G f ∘ η c = (η c')⁻¹ ∘ η c' ∘ F f : by rewrite naturality
|
2015-05-03 05:22:31 +00:00
|
|
|
|
... = F f : by rewrite inverse_comp_cancel_left
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
|
|
|
|
omit isoη
|
|
|
|
|
|
|
|
|
|
definition componentwise_iso (η : F ≅ G) (c : C) : F c ≅ G c :=
|
2015-10-23 05:12:34 +00:00
|
|
|
|
iso.mk (natural_map (to_hom η) c)
|
|
|
|
|
(@componentwise_is_iso _ _ _ _ (to_hom η) (struct η) c)
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
|
|
|
|
definition componentwise_iso_id (c : C) : componentwise_iso (iso.refl F) c = iso.refl (F c) :=
|
2015-04-27 21:29:56 +00:00
|
|
|
|
iso_eq (idpath (ID (F c)))
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
|
|
|
|
definition componentwise_iso_iso_of_eq (p : F = G) (c : C)
|
|
|
|
|
: componentwise_iso (iso_of_eq p) c = iso_of_eq (ap010 to_fun_ob p c) :=
|
|
|
|
|
eq.rec_on p !componentwise_iso_id
|
|
|
|
|
|
2015-10-16 19:15:44 +00:00
|
|
|
|
theorem naturality_iso_id {F : C ⇒ C} (η : F ≅ 1) (c : C)
|
|
|
|
|
: componentwise_iso η (F c) = F (componentwise_iso η c) :=
|
|
|
|
|
comp.cancel_left (to_hom (componentwise_iso η c))
|
|
|
|
|
((naturality (to_hom η)) (to_hom (componentwise_iso η c)))
|
|
|
|
|
|
2015-04-25 04:20:59 +00:00
|
|
|
|
definition natural_map_hom_of_eq (p : F = G) (c : C)
|
|
|
|
|
: natural_map (hom_of_eq p) c = hom_of_eq (ap010 to_fun_ob p c) :=
|
|
|
|
|
eq.rec_on p idp
|
|
|
|
|
|
|
|
|
|
definition natural_map_inv_of_eq (p : F = G) (c : C)
|
|
|
|
|
: natural_map (inv_of_eq p) c = hom_of_eq (ap010 to_fun_ob p c)⁻¹ :=
|
|
|
|
|
eq.rec_on p idp
|
|
|
|
|
|
2015-09-28 04:38:35 +00:00
|
|
|
|
definition hom_of_eq_compose_right {H : B ⇒ C} (p : F = G)
|
2015-08-31 16:23:34 +00:00
|
|
|
|
: hom_of_eq (ap (λx, x ∘f H) p) = hom_of_eq p ∘nf H :=
|
|
|
|
|
eq.rec_on p idp
|
|
|
|
|
|
2015-09-28 04:38:35 +00:00
|
|
|
|
definition inv_of_eq_compose_right {H : B ⇒ C} (p : F = G)
|
2015-08-31 16:23:34 +00:00
|
|
|
|
: inv_of_eq (ap (λx, x ∘f H) p) = inv_of_eq p ∘nf H :=
|
|
|
|
|
eq.rec_on p idp
|
|
|
|
|
|
2015-09-28 04:38:35 +00:00
|
|
|
|
definition hom_of_eq_compose_left {H : D ⇒ C} (p : F = G)
|
2015-08-31 16:23:34 +00:00
|
|
|
|
: hom_of_eq (ap (λx, H ∘f x) p) = H ∘fn hom_of_eq p :=
|
|
|
|
|
by induction p; exact !fn_id⁻¹
|
|
|
|
|
|
2015-09-28 04:38:35 +00:00
|
|
|
|
definition inv_of_eq_compose_left {H : D ⇒ C} (p : F = G)
|
2015-08-31 16:23:34 +00:00
|
|
|
|
: inv_of_eq (ap (λx, H ∘f x) p) = H ∘fn inv_of_eq p :=
|
|
|
|
|
by induction p; exact !fn_id⁻¹
|
|
|
|
|
|
|
|
|
|
definition assoc_natural [constructor] (H : C ⇒ D) (G : B ⇒ C) (F : A ⇒ B)
|
|
|
|
|
: H ∘f (G ∘f F) ⟹ (H ∘f G) ∘f F :=
|
|
|
|
|
change_natural_map (hom_of_eq !functor.assoc)
|
|
|
|
|
(λa, id)
|
|
|
|
|
(λa, !natural_map_hom_of_eq ⬝ ap hom_of_eq !ap010_assoc)
|
|
|
|
|
|
|
|
|
|
definition assoc_natural_rev [constructor] (H : C ⇒ D) (G : B ⇒ C) (F : A ⇒ B)
|
|
|
|
|
: (H ∘f G) ∘f F ⟹ H ∘f (G ∘f F) :=
|
|
|
|
|
change_natural_map (inv_of_eq !functor.assoc)
|
|
|
|
|
(λa, id)
|
|
|
|
|
(λa, !natural_map_inv_of_eq ⬝ ap (λx, hom_of_eq x⁻¹) !ap010_assoc)
|
|
|
|
|
|
|
|
|
|
definition id_left_natural [constructor] (F : C ⇒ D) : functor.id ∘f F ⟹ F :=
|
|
|
|
|
change_natural_map
|
|
|
|
|
(hom_of_eq !functor.id_left)
|
|
|
|
|
(λc, id)
|
|
|
|
|
(λc, by induction F; exact !natural_map_hom_of_eq ⬝ ap hom_of_eq !ap010_functor_mk_eq_constant)
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
definition id_left_natural_rev [constructor] (F : C ⇒ D) : F ⟹ functor.id ∘f F :=
|
|
|
|
|
change_natural_map
|
|
|
|
|
(inv_of_eq !functor.id_left)
|
|
|
|
|
(λc, id)
|
|
|
|
|
(λc, by induction F; exact !natural_map_inv_of_eq ⬝
|
|
|
|
|
ap (λx, hom_of_eq x⁻¹) !ap010_functor_mk_eq_constant)
|
|
|
|
|
|
|
|
|
|
definition id_right_natural [constructor] (F : C ⇒ D) : F ∘f functor.id ⟹ F :=
|
|
|
|
|
change_natural_map
|
|
|
|
|
(hom_of_eq !functor.id_right)
|
|
|
|
|
(λc, id)
|
|
|
|
|
(λc, by induction F; exact !natural_map_hom_of_eq ⬝ ap hom_of_eq !ap010_functor_mk_eq_constant)
|
|
|
|
|
|
|
|
|
|
definition id_right_natural_rev [constructor] (F : C ⇒ D) : F ⟹ F ∘f functor.id :=
|
|
|
|
|
change_natural_map
|
|
|
|
|
(inv_of_eq !functor.id_right)
|
|
|
|
|
(λc, id)
|
|
|
|
|
(λc, by induction F; exact !natural_map_inv_of_eq ⬝
|
|
|
|
|
ap (λx, hom_of_eq x⁻¹) !ap010_functor_mk_eq_constant)
|
|
|
|
|
|
2015-04-25 04:20:59 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-10-22 22:41:55 +00:00
|
|
|
|
section
|
|
|
|
|
variables {C D E : Precategory} {G G' : D ⇒ E} {F F' : C ⇒ D} {J : D ⇒ D}
|
|
|
|
|
|
|
|
|
|
definition is_iso_nf_compose [constructor] (G : D ⇒ E) (η : F ⟹ F') [H : is_iso η]
|
|
|
|
|
: is_iso (G ∘fn η) :=
|
|
|
|
|
is_iso.mk
|
|
|
|
|
(G ∘fn @inverse (C ⇒ D) _ _ _ η _)
|
|
|
|
|
abstract !fn_n_distrib⁻¹ ⬝ ap (λx, G ∘fn x) (@left_inverse (C ⇒ D) _ _ _ η _) ⬝ !fn_id end
|
|
|
|
|
abstract !fn_n_distrib⁻¹ ⬝ ap (λx, G ∘fn x) (@right_inverse (C ⇒ D) _ _ _ η _) ⬝ !fn_id end
|
|
|
|
|
|
|
|
|
|
definition is_iso_fn_compose [constructor] (η : G ⟹ G') (F : C ⇒ D) [H : is_iso η]
|
|
|
|
|
: is_iso (η ∘nf F) :=
|
|
|
|
|
is_iso.mk
|
|
|
|
|
(@inverse (D ⇒ E) _ _ _ η _ ∘nf F)
|
|
|
|
|
abstract !n_nf_distrib⁻¹ ⬝ ap (λx, x ∘nf F) (@left_inverse (D ⇒ E) _ _ _ η _) ⬝ !id_nf end
|
|
|
|
|
abstract !n_nf_distrib⁻¹ ⬝ ap (λx, x ∘nf F) (@right_inverse (D ⇒ E) _ _ _ η _) ⬝ !id_nf end
|
|
|
|
|
|
|
|
|
|
definition functor_iso_compose [constructor] (G : D ⇒ E) (η : F ≅ F') : G ∘f F ≅ G ∘f F' :=
|
2015-10-23 05:12:34 +00:00
|
|
|
|
iso.mk _ (is_iso_nf_compose G (to_hom η))
|
2015-10-22 22:41:55 +00:00
|
|
|
|
|
|
|
|
|
definition iso_functor_compose [constructor] (η : G ≅ G') (F : C ⇒ D) : G ∘f F ≅ G' ∘f F :=
|
2015-10-23 05:12:34 +00:00
|
|
|
|
iso.mk _ (is_iso_fn_compose (to_hom η) F)
|
2015-10-22 22:41:55 +00:00
|
|
|
|
|
|
|
|
|
infixr ` ∘fi ` :62 := functor_iso_compose
|
|
|
|
|
infixr ` ∘if ` :62 := iso_functor_compose
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
/- TODO: also needs n_nf_distrib and id_nf for these compositions
|
|
|
|
|
definition nidf_compose [constructor] (η : J ⟹ 1) (F : C ⇒ D) [H : is_iso η]
|
|
|
|
|
: is_iso (η ∘n1f F) :=
|
|
|
|
|
is_iso.mk
|
|
|
|
|
(@inverse (D ⇒ D) _ _ _ η _ ∘1nf F)
|
|
|
|
|
abstract _ end
|
|
|
|
|
_
|
|
|
|
|
|
|
|
|
|
definition idnf_compose [constructor] (η : 1 ⟹ J) (F : C ⇒ D) [H : is_iso η]
|
|
|
|
|
: is_iso (η ∘1nf F) :=
|
|
|
|
|
is_iso.mk _
|
|
|
|
|
_
|
|
|
|
|
_
|
|
|
|
|
|
|
|
|
|
definition fnid_compose [constructor] (F : D ⇒ E) (η : J ⟹ 1) [H : is_iso η]
|
|
|
|
|
: is_iso (F ∘fn1 η) :=
|
|
|
|
|
is_iso.mk _
|
|
|
|
|
_
|
|
|
|
|
_
|
|
|
|
|
|
|
|
|
|
definition fidn_compose [constructor] (F : D ⇒ E) (η : 1 ⟹ J) [H : is_iso η]
|
|
|
|
|
: is_iso (F ∘f1n η) :=
|
|
|
|
|
is_iso.mk _
|
|
|
|
|
_
|
|
|
|
|
_
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
end
|
|
|
|
|
|
2015-04-25 04:20:59 +00:00
|
|
|
|
namespace functor
|
|
|
|
|
|
|
|
|
|
variables {C : Precategory} {D : Category} {F G : D ^c C}
|
|
|
|
|
definition eq_of_iso_ob (η : F ≅ G) (c : C) : F c = G c :=
|
|
|
|
|
by apply eq_of_iso; apply componentwise_iso; exact η
|
|
|
|
|
|
2016-02-26 01:42:44 +00:00
|
|
|
|
local attribute functor.to_fun_hom [reducible]
|
2015-04-25 04:20:59 +00:00
|
|
|
|
definition eq_of_iso (η : F ≅ G) : F = G :=
|
|
|
|
|
begin
|
|
|
|
|
fapply functor_eq,
|
|
|
|
|
{exact (eq_of_iso_ob η)},
|
2015-05-07 22:29:02 +00:00
|
|
|
|
{intro c c' f,
|
|
|
|
|
esimp [eq_of_iso_ob, inv_of_eq, hom_of_eq, eq_of_iso],
|
|
|
|
|
rewrite [*right_inv iso_of_eq],
|
2015-06-27 00:09:50 +00:00
|
|
|
|
symmetry, apply @naturality_iso _ _ _ _ _ (iso.struct _)
|
2015-05-07 22:29:02 +00:00
|
|
|
|
}
|
2015-04-25 04:20:59 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition iso_of_eq_eq_of_iso (η : F ≅ G) : iso_of_eq (eq_of_iso η) = η :=
|
|
|
|
|
begin
|
2015-04-27 21:29:56 +00:00
|
|
|
|
apply iso_eq,
|
|
|
|
|
apply nat_trans_eq,
|
2015-04-25 04:20:59 +00:00
|
|
|
|
intro c,
|
|
|
|
|
rewrite natural_map_hom_of_eq, esimp [eq_of_iso],
|
|
|
|
|
rewrite ap010_functor_eq, esimp [hom_of_eq,eq_of_iso_ob],
|
2015-04-27 19:39:36 +00:00
|
|
|
|
rewrite (right_inv iso_of_eq),
|
2015-04-25 04:20:59 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition eq_of_iso_iso_of_eq (p : F = G) : eq_of_iso (iso_of_eq p) = p :=
|
|
|
|
|
begin
|
|
|
|
|
apply functor_eq2,
|
|
|
|
|
intro c,
|
|
|
|
|
esimp [eq_of_iso],
|
|
|
|
|
rewrite ap010_functor_eq,
|
|
|
|
|
esimp [eq_of_iso_ob],
|
|
|
|
|
rewrite componentwise_iso_iso_of_eq,
|
2015-04-27 19:39:36 +00:00
|
|
|
|
rewrite (left_inv iso_of_eq)
|
2015-04-25 04:20:59 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition is_univalent (D : Category) (C : Precategory) : is_univalent (D ^c C) :=
|
|
|
|
|
λF G, adjointify _ eq_of_iso
|
|
|
|
|
iso_of_eq_eq_of_iso
|
|
|
|
|
eq_of_iso_iso_of_eq
|
|
|
|
|
|
|
|
|
|
end functor
|
|
|
|
|
|
2015-09-03 04:46:11 +00:00
|
|
|
|
definition category_functor [instance] [constructor] (D : Category) (C : Precategory)
|
2015-04-29 00:48:39 +00:00
|
|
|
|
: category (D ^c C) :=
|
|
|
|
|
category.mk (D ^c C) (functor.is_univalent D C)
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
2015-09-03 04:46:11 +00:00
|
|
|
|
definition Category_functor [constructor] (D : Category) (C : Precategory) : Category :=
|
2015-04-29 00:48:39 +00:00
|
|
|
|
category.Mk (D ^c C) !category_functor
|
|
|
|
|
|
|
|
|
|
--this definition is only useful if the exponent is a category,
|
|
|
|
|
-- and the elaborator has trouble with inserting the coercion
|
2015-09-03 04:46:11 +00:00
|
|
|
|
definition Category_functor' [constructor] (D C : Category) : Category :=
|
2015-04-29 00:48:39 +00:00
|
|
|
|
Category_functor D C
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
|
|
|
|
namespace ops
|
2015-10-01 19:52:28 +00:00
|
|
|
|
infixr ` ^c2 `:35 := Category_functor
|
2015-04-25 04:20:59 +00:00
|
|
|
|
end ops
|
|
|
|
|
|
2015-08-31 16:23:34 +00:00
|
|
|
|
namespace functor
|
|
|
|
|
variables {C : Precategory} {D : Category} {F G : D ^c C}
|
|
|
|
|
|
|
|
|
|
definition eq_of_pointwise_iso (η : F ⟹ G) (iso : Π(a : C), is_iso (η a)) : F = G :=
|
2015-10-23 05:12:34 +00:00
|
|
|
|
eq_of_iso (natural_iso.mk η iso)
|
2015-08-31 16:23:34 +00:00
|
|
|
|
|
|
|
|
|
definition iso_of_eq_eq_of_pointwise_iso (η : F ⟹ G) (iso : Π(c : C), is_iso (η c))
|
2015-10-23 05:12:34 +00:00
|
|
|
|
: iso_of_eq (eq_of_pointwise_iso η iso) = natural_iso.mk η iso :=
|
2015-08-31 16:23:34 +00:00
|
|
|
|
!iso_of_eq_eq_of_iso
|
|
|
|
|
|
|
|
|
|
definition hom_of_eq_eq_of_pointwise_iso (η : F ⟹ G) (iso : Π(c : C), is_iso (η c))
|
|
|
|
|
: hom_of_eq (eq_of_pointwise_iso η iso) = η :=
|
|
|
|
|
!hom_of_eq_eq_of_iso
|
|
|
|
|
|
|
|
|
|
definition inv_of_eq_eq_of_pointwise_iso (η : F ⟹ G) (iso : Π(c : C), is_iso (η c))
|
|
|
|
|
: inv_of_eq (eq_of_pointwise_iso η iso) = nat_trans_inverse η :=
|
|
|
|
|
!inv_of_eq_eq_of_iso
|
|
|
|
|
|
|
|
|
|
end functor
|
2015-04-25 04:20:59 +00:00
|
|
|
|
|
2015-10-22 22:41:55 +00:00
|
|
|
|
/-
|
|
|
|
|
functors involving only the functor category
|
|
|
|
|
(see ..functor.curry for some other functors involving also products)
|
|
|
|
|
-/
|
|
|
|
|
|
2015-10-02 23:54:27 +00:00
|
|
|
|
variables {C D I : Precategory}
|
|
|
|
|
definition constant2_functor [constructor] (F : I ⇒ D ^c C) (c : C) : I ⇒ D :=
|
|
|
|
|
functor.mk (λi, to_fun_ob (F i) c)
|
|
|
|
|
(λi j f, natural_map (F f) c)
|
|
|
|
|
abstract (λi, ap010 natural_map !respect_id c ⬝ proof idp qed) end
|
|
|
|
|
abstract (λi j k g f, ap010 natural_map !respect_comp c) end
|
|
|
|
|
|
|
|
|
|
definition constant2_functor_natural [constructor] (F : I ⇒ D ^c C) {c d : C} (f : c ⟶ d)
|
|
|
|
|
: constant2_functor F c ⟹ constant2_functor F d :=
|
|
|
|
|
nat_trans.mk (λi, to_fun_hom (F i) f)
|
|
|
|
|
(λi j k, (naturality (F k) f)⁻¹)
|
|
|
|
|
|
|
|
|
|
definition functor_flip [constructor] (F : I ⇒ D ^c C) : C ⇒ D ^c I :=
|
|
|
|
|
functor.mk (constant2_functor F)
|
|
|
|
|
@(constant2_functor_natural F)
|
|
|
|
|
abstract begin intros, apply nat_trans_eq, intro i, esimp, apply respect_id end end
|
|
|
|
|
abstract begin intros, apply nat_trans_eq, intro i, esimp, apply respect_comp end end
|
|
|
|
|
|
2015-10-22 22:41:55 +00:00
|
|
|
|
definition eval_functor [constructor] (C D : Precategory) (d : D) : C ^c D ⇒ C :=
|
|
|
|
|
begin
|
|
|
|
|
fapply functor.mk: esimp,
|
|
|
|
|
{ intro F, exact F d},
|
|
|
|
|
{ intro G F η, exact η d},
|
|
|
|
|
{ intro F, reflexivity},
|
|
|
|
|
{ intro H G F η θ, reflexivity},
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition precomposition_functor [constructor] {C D} (E) (F : C ⇒ D)
|
|
|
|
|
: E ^c D ⇒ E ^c C :=
|
|
|
|
|
begin
|
|
|
|
|
fapply functor.mk: esimp,
|
|
|
|
|
{ intro G, exact G ∘f F},
|
|
|
|
|
{ intro G H η, exact η ∘nf F},
|
|
|
|
|
{ intro G, reflexivity},
|
|
|
|
|
{ intro G H I η θ, reflexivity},
|
|
|
|
|
end
|
|
|
|
|
|
2016-07-18 15:10:32 +00:00
|
|
|
|
definition faithful_precomposition_functor [instance]
|
|
|
|
|
{C D E} {H : C ⇒ D} [Hs : essentially_surjective H] : faithful (precomposition_functor E H) :=
|
2016-07-15 16:20:42 +00:00
|
|
|
|
begin
|
|
|
|
|
intro F G γ δ Hγδ, apply nat_trans_eq, intro b,
|
2016-07-18 15:10:32 +00:00
|
|
|
|
induction Hs b with Hb, induction Hb with a f,
|
2016-07-15 16:20:42 +00:00
|
|
|
|
refine naturality_iso_right γ f ⬝ _ ⬝ (naturality_iso_right δ f)⁻¹,
|
|
|
|
|
apply ap (λ x, _ ∘ natural_map x a ∘ _) Hγδ,
|
|
|
|
|
end
|
|
|
|
|
|
2016-07-18 15:10:32 +00:00
|
|
|
|
open sigma sigma.ops
|
|
|
|
|
section fully_faithful_precomposition
|
|
|
|
|
variables {E : Precategory} {H : C ⇒ D} [Hs : essentially_surjective H] [Hf : full H]
|
|
|
|
|
{F G : D ⇒ E} (γ : F ∘f H ⟹ G ∘f H)
|
|
|
|
|
include Hs Hf
|
|
|
|
|
|
|
|
|
|
private definition fully_faithful_precomposition_functor_prop [instance] (b) :
|
|
|
|
|
is_prop (Σ g, Π a (f : H a ≅ b), γ a = G f⁻¹ⁱ ∘ g ∘ F f) :=
|
|
|
|
|
begin
|
|
|
|
|
fapply is_prop.mk, intros g h, cases g with g Hg, cases h with h Hh,
|
|
|
|
|
fapply sigma.dpair_eq_dpair,
|
|
|
|
|
{ induction Hs b with Hb, induction Hb with a0 f,
|
|
|
|
|
apply comp.cancel_right (F f), apply comp.cancel_left (G f⁻¹ⁱ),
|
|
|
|
|
apply (Hg a0 f)⁻¹ ⬝ (Hh a0 f) },
|
|
|
|
|
apply is_prop.elimo
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
private definition fully_faithful_precomposition_functor_pair [reducible] (b) :
|
|
|
|
|
Σ g, Π a (f : H a ≅ b), γ a = G f⁻¹ⁱ ∘ g ∘ F f :=
|
|
|
|
|
begin
|
|
|
|
|
induction Hs b with Hb, induction Hb with a0 h, fconstructor,
|
|
|
|
|
exact G h ∘ γ a0 ∘ F h⁻¹ⁱ, intro a f,
|
|
|
|
|
induction Hf (to_hom (f ⬝i h⁻¹ⁱ)) with k Ek,
|
|
|
|
|
have is_iso (H k), by rewrite Ek; apply _,
|
|
|
|
|
refine _ ⬝ !assoc⁻¹, refine _ ⬝ ap (λ x, x ∘ F f) !assoc⁻¹, refine _ ⬝ !assoc,
|
|
|
|
|
refine _ ⬝ ap (λ x, (G f⁻¹ⁱ ∘ G h) ∘ x) !assoc,
|
|
|
|
|
do 2 krewrite [-respect_comp], esimp,
|
|
|
|
|
apply eq_comp_of_inverse_comp_eq,
|
|
|
|
|
exact ap (λ x, G x ∘ γ a) Ek⁻¹ ⬝ naturality γ k ⬝ ap (λ x, γ a0 ∘ F x) Ek
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
private definition fully_faithful_precomposition_naturality {b b' : carrier D}
|
|
|
|
|
(f : hom b b') : to_fun_hom G f ∘ (fully_faithful_precomposition_functor_pair γ b).1
|
|
|
|
|
= (fully_faithful_precomposition_functor_pair γ b').1 ∘ to_fun_hom F f :=
|
|
|
|
|
begin
|
|
|
|
|
esimp[fully_faithful_precomposition_functor_pair],
|
|
|
|
|
induction Hs b with Hb, induction Hb with a h,
|
|
|
|
|
induction Hs b' with Hb', induction Hb' with a' h',
|
|
|
|
|
induction Hf (to_hom h'⁻¹ⁱ ∘ f ∘ to_hom h) with k Ek,
|
|
|
|
|
apply concat, apply assoc,
|
|
|
|
|
apply concat, apply ap (λ x, x ∘ _),
|
|
|
|
|
apply concat, apply !respect_comp⁻¹,
|
|
|
|
|
apply concat, apply ap (λ x, to_fun_hom G x), apply inverse,
|
|
|
|
|
apply comp_eq_of_eq_inverse_comp, apply Ek, apply respect_comp,
|
|
|
|
|
apply concat, apply !assoc⁻¹,
|
|
|
|
|
apply concat, apply ap (λ x, _ ∘ x), apply concat, apply assoc,
|
|
|
|
|
apply concat, apply ap (λ x, x ∘ _), apply naturality γ, apply !assoc⁻¹,
|
|
|
|
|
apply concat, apply ap (λ x, _ ∘ _ ∘ x), apply concat, esimp, apply !respect_comp⁻¹,
|
|
|
|
|
apply concat, apply ap (λ x, to_fun_hom F x),
|
|
|
|
|
apply comp_inverse_eq_of_eq_comp, apply Ek ⬝ !assoc, apply respect_comp,
|
|
|
|
|
apply concat, apply assoc, apply concat, apply assoc,
|
|
|
|
|
apply ap (λ x, x ∘ _) !assoc⁻¹
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition fully_faithful_precomposition_functor [instance] :
|
|
|
|
|
fully_faithful (precomposition_functor E H) :=
|
|
|
|
|
begin
|
|
|
|
|
apply fully_faithful_of_full_of_faithful,
|
|
|
|
|
{ apply faithful_precomposition_functor },
|
|
|
|
|
{ intro F G γ, esimp at *, fapply image.mk,
|
|
|
|
|
fconstructor,
|
|
|
|
|
{ intro b, apply (fully_faithful_precomposition_functor_pair γ b).1 },
|
|
|
|
|
{ intro b b' f, apply fully_faithful_precomposition_naturality },
|
|
|
|
|
{ fapply nat_trans_eq, intro a, esimp,
|
|
|
|
|
apply inverse,
|
|
|
|
|
induction (fully_faithful_precomposition_functor_pair γ (to_fun_ob H a)) with g Hg,
|
|
|
|
|
esimp, apply concat, apply Hg a (iso.refl (H a)), esimp,
|
|
|
|
|
apply concat, apply ap (λ x, x ∘ _), apply respect_id, apply concat, apply id_left,
|
|
|
|
|
apply concat, apply ap (λ x, _ ∘ x), apply respect_id, apply id_right } }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
end fully_faithful_precomposition
|
|
|
|
|
|
2016-07-20 14:03:01 +00:00
|
|
|
|
section essentially_surjective_precomposition
|
|
|
|
|
variables {E : Category}
|
|
|
|
|
{H : C ⇒ D} [He : is_weak_equivalence H]
|
|
|
|
|
(F : C ⇒ E) (b : carrier D)
|
|
|
|
|
include E H He F b
|
|
|
|
|
|
|
|
|
|
structure essentially_surj_precomp_X : Type :=
|
|
|
|
|
(c : carrier E)
|
|
|
|
|
(k : Π (a : carrier C) (h : H a ≅ b), F a ≅ c)
|
|
|
|
|
(k_coh : Π {a a'} h h' (f : hom a a'), to_hom h' ∘ (to_fun_hom H f) = to_hom h
|
|
|
|
|
→ to_hom (k a' h') ∘ to_fun_hom F f = to_hom (k a h))
|
|
|
|
|
|
|
|
|
|
section
|
|
|
|
|
variables {c c' : carrier E} (p : c = c')
|
|
|
|
|
{k : Π (a : carrier C) (h : H a ≅ b), F a ≅ c}
|
|
|
|
|
{k' : Π (a : carrier C) (h : H a ≅ b), F a ≅ c'}
|
|
|
|
|
(q : Π (a : carrier C) (h : H a ≅ b), to_hom (k a h ⬝i iso_of_eq p) = to_hom (k' a h))
|
|
|
|
|
{k_coh : Π {a a'} h h' (f : hom a a'), to_hom h' ∘ (to_fun_hom H f) = to_hom h
|
|
|
|
|
→ to_hom (k a' h') ∘ to_fun_hom F f = to_hom (k a h)}
|
|
|
|
|
{k'_coh : Π {a a'} h h' (f : hom a a'), to_hom h' ∘ (to_fun_hom H f) = to_hom h
|
|
|
|
|
→ to_hom (k' a' h') ∘ to_fun_hom F f = to_hom (k' a h)}
|
|
|
|
|
include c c' p k k' q
|
|
|
|
|
|
|
|
|
|
private definition essentially_surj_precomp_X_eq :
|
|
|
|
|
essentially_surj_precomp_X.mk c k @k_coh =
|
|
|
|
|
essentially_surj_precomp_X.mk c' k' @k'_coh :=
|
|
|
|
|
begin
|
|
|
|
|
cases p,
|
|
|
|
|
assert q' : k = k',
|
|
|
|
|
{ apply eq_of_homotopy, intro a, apply eq_of_homotopy, intro h,
|
|
|
|
|
apply iso_eq, apply !id_left⁻¹ ⬝ q a h },
|
|
|
|
|
cases q',
|
|
|
|
|
apply ap (essentially_surj_precomp_X.mk c' k'),
|
|
|
|
|
apply is_prop.elim
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
open prod.ops
|
|
|
|
|
private definition essentially_surj_precomp_X_prop [instance] :
|
|
|
|
|
is_prop (@essentially_surj_precomp_X C D E H He F b) :=
|
|
|
|
|
begin
|
|
|
|
|
induction He.2 b with Hb, cases Hb with a0 Ha0,
|
|
|
|
|
fapply is_prop.mk, intros f g, cases f with cf kf kf_coh, cases g with cg kg kg_coh,
|
|
|
|
|
fapply essentially_surj_precomp_X_eq,
|
|
|
|
|
{ apply eq_of_iso, apply iso.trans, apply iso.symm, apply kf a0 Ha0,
|
|
|
|
|
apply kg a0 Ha0 },
|
|
|
|
|
{ intro a h,
|
|
|
|
|
assert fHf : Σ f : hom a a0, to_hom Ha0 ∘ (to_fun_hom H f) = to_hom h,
|
|
|
|
|
{ fconstructor, apply hom_inv, apply He.1, exact to_hom Ha0⁻¹ⁱ ∘ to_hom h,
|
|
|
|
|
apply concat, apply ap (λ x, _ ∘ x), apply right_inv (to_fun_hom H),
|
|
|
|
|
apply comp_inverse_cancel_left },
|
|
|
|
|
apply concat, apply ap (λ x, to_hom x ∘ _), apply iso_of_eq_eq_of_iso,
|
|
|
|
|
apply concat, apply ap (λ x, _ ∘ x), apply (kf_coh h Ha0 fHf.1 fHf.2)⁻¹,
|
|
|
|
|
apply concat, rotate 1, apply kg_coh h Ha0 fHf.1 fHf.2,
|
|
|
|
|
apply concat, apply assoc, apply ap (λ x, x ∘ _),
|
|
|
|
|
apply concat, apply !assoc⁻¹,
|
|
|
|
|
apply concat, apply ap (λ x, _ ∘ x), apply comp.left_inverse,
|
|
|
|
|
apply id_right },
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
private definition essentially_surj_precomp_X_inh :
|
|
|
|
|
@essentially_surj_precomp_X C D E H He F b :=
|
|
|
|
|
begin
|
|
|
|
|
induction He.2 b with Hb, cases Hb with a0 Ha0,
|
|
|
|
|
fconstructor, exact F a0,
|
|
|
|
|
{ intro a h, apply to_fun_iso, apply iso_of_F_iso_F H,
|
|
|
|
|
exact h ⬝i Ha0⁻¹ⁱ },
|
|
|
|
|
{ intros a a' h h' f HH, esimp[iso_of_F_iso_F],
|
|
|
|
|
apply concat, apply !respect_comp⁻¹, apply ap (to_fun_hom F),
|
|
|
|
|
rewrite [-HH],
|
|
|
|
|
apply concat, apply ap (λ x, _ ∘ x), apply inverse, apply left_inv (to_fun_hom H),
|
|
|
|
|
apply concat, apply !hom_inv_respect_comp⁻¹, apply ap (hom_inv H),
|
|
|
|
|
apply !assoc⁻¹ }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition essentially_surjective_precomposition_functor :
|
|
|
|
|
essentially_surjective (precomposition_functor E H) :=
|
|
|
|
|
begin
|
|
|
|
|
intro F, esimp,
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
end essentially_surjective_precomposition
|
|
|
|
|
|
2015-10-22 22:41:55 +00:00
|
|
|
|
definition postcomposition_functor [constructor] {C D} (E) (F : C ⇒ D)
|
|
|
|
|
: C ^c E ⇒ D ^c E :=
|
|
|
|
|
begin
|
|
|
|
|
fapply functor.mk: esimp,
|
|
|
|
|
{ intro G, exact F ∘f G},
|
|
|
|
|
{ intro G H η, exact F ∘fn η},
|
|
|
|
|
{ intro G, apply fn_id},
|
|
|
|
|
{ intro G H I η θ, apply fn_n_distrib},
|
|
|
|
|
end
|
|
|
|
|
|
2015-10-23 05:12:34 +00:00
|
|
|
|
definition constant_diagram [constructor] (C D) : C ⇒ C ^c D :=
|
|
|
|
|
begin
|
|
|
|
|
fapply functor.mk: esimp,
|
|
|
|
|
{ intro c, exact constant_functor D c},
|
|
|
|
|
{ intro c d f, exact constant_nat_trans D f},
|
|
|
|
|
{ intro c, fapply nat_trans_eq, reflexivity},
|
|
|
|
|
{ intro c d e g f, fapply nat_trans_eq, reflexivity},
|
|
|
|
|
end
|
|
|
|
|
|
2015-11-16 21:21:21 +00:00
|
|
|
|
definition opposite_functor_opposite_left [constructor] (C D : Precategory)
|
|
|
|
|
: (C ^c D)ᵒᵖ ⇒ Cᵒᵖ ^c Dᵒᵖ :=
|
|
|
|
|
begin
|
|
|
|
|
fapply functor.mk: esimp,
|
|
|
|
|
{ exact opposite_functor},
|
|
|
|
|
{ intro F G, exact opposite_nat_trans},
|
|
|
|
|
{ intro F, apply nat_trans_eq, reflexivity},
|
|
|
|
|
{ intro u v w g f, apply nat_trans_eq, reflexivity}
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition opposite_functor_opposite_right [constructor] (C D : Precategory)
|
|
|
|
|
: Cᵒᵖ ^c Dᵒᵖ ⇒ (C ^c D)ᵒᵖ :=
|
|
|
|
|
begin
|
|
|
|
|
fapply functor.mk: esimp,
|
|
|
|
|
{ exact opposite_functor_rev},
|
|
|
|
|
{ apply @opposite_rev_nat_trans},
|
|
|
|
|
{ intro F, apply nat_trans_eq, intro d, reflexivity},
|
|
|
|
|
{ intro F G H η θ, apply nat_trans_eq, intro d, reflexivity}
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition constant_diagram_opposite [constructor] (C D)
|
|
|
|
|
: (constant_diagram C D)ᵒᵖᶠ = opposite_functor_opposite_right C D ∘f constant_diagram Cᵒᵖ Dᵒᵖ :=
|
|
|
|
|
begin
|
|
|
|
|
fapply functor_eq,
|
|
|
|
|
{ reflexivity},
|
|
|
|
|
{ intro c c' f, esimp at *, refine !nat_trans.id_right ⬝ !nat_trans.id_left ⬝ _,
|
|
|
|
|
apply nat_trans_eq, intro d, reflexivity}
|
|
|
|
|
end
|
2015-10-22 22:41:55 +00:00
|
|
|
|
|
2015-10-02 23:54:27 +00:00
|
|
|
|
|
|
|
|
|
end functor
|