2015-02-26 18:19:54 +00:00
|
|
|
|
/-
|
2015-03-17 00:08:45 +00:00
|
|
|
|
Copyright (c) 2015 Floris van Doorn. All rights reserved.
|
2015-02-26 18:19:54 +00:00
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
|
2015-04-25 04:20:59 +00:00
|
|
|
|
Module: algebra.category.functor
|
2015-02-26 18:19:54 +00:00
|
|
|
|
Authors: Floris van Doorn, Jakob von Raumer
|
|
|
|
|
-/
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-04-25 04:20:59 +00:00
|
|
|
|
import .iso types.pi
|
2014-12-12 19:19:06 +00:00
|
|
|
|
|
2015-03-13 14:32:48 +00:00
|
|
|
|
open function category eq prod equiv is_equiv sigma sigma.ops is_trunc funext iso
|
2015-01-01 00:30:17 +00:00
|
|
|
|
open pi
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2014-12-18 02:49:37 +00:00
|
|
|
|
structure functor (C D : Precategory) : Type :=
|
2015-02-28 00:50:06 +00:00
|
|
|
|
(to_fun_ob : C → D)
|
|
|
|
|
(to_fun_hom : Π ⦃a b : C⦄, hom a b → hom (to_fun_ob a) (to_fun_ob b))
|
|
|
|
|
(respect_id : Π (a : C), to_fun_hom (ID a) = ID (to_fun_ob a))
|
2014-12-18 02:49:37 +00:00
|
|
|
|
(respect_comp : Π {a b c : C} (g : hom b c) (f : hom a b),
|
2015-02-28 00:50:06 +00:00
|
|
|
|
to_fun_hom (g ∘ f) = to_fun_hom g ∘ to_fun_hom f)
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
|
|
|
|
namespace functor
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
|
|
|
|
infixl `⇒`:25 := functor
|
2015-03-13 22:27:29 +00:00
|
|
|
|
variables {A B C D E : Precategory}
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-02-28 00:50:06 +00:00
|
|
|
|
attribute to_fun_ob [coercion]
|
|
|
|
|
attribute to_fun_hom [coercion]
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-02-24 00:54:16 +00:00
|
|
|
|
-- The following lemmas will later be used to prove that the type of
|
|
|
|
|
-- precategories forms a precategory itself
|
2015-02-27 05:45:21 +00:00
|
|
|
|
protected definition compose [reducible] (G : functor D E) (F : functor C D) : functor C E :=
|
2015-02-24 00:54:16 +00:00
|
|
|
|
functor.mk
|
|
|
|
|
(λ x, G (F x))
|
|
|
|
|
(λ a b f, G (F f))
|
|
|
|
|
(λ a, calc
|
2015-03-13 03:52:00 +00:00
|
|
|
|
G (F (ID a)) = G (ID (F a)) : by rewrite respect_id
|
|
|
|
|
... = ID (G (F a)) : by rewrite respect_id)
|
2015-02-24 00:54:16 +00:00
|
|
|
|
(λ a b c g f, calc
|
2015-03-13 03:13:40 +00:00
|
|
|
|
G (F (g ∘ f)) = G (F g ∘ F f) : by rewrite respect_comp
|
|
|
|
|
... = G (F g) ∘ G (F f) : by rewrite respect_comp)
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
|
|
|
|
infixr `∘f`:60 := compose
|
|
|
|
|
|
2015-02-27 05:45:21 +00:00
|
|
|
|
protected definition id [reducible] {C : Precategory} : functor C C :=
|
|
|
|
|
mk (λa, a) (λ a b f, f) (λ a, idp) (λ a b c f g, idp)
|
|
|
|
|
|
|
|
|
|
protected definition ID [reducible] (C : Precategory) : functor C C := id
|
|
|
|
|
|
2015-03-13 22:27:29 +00:00
|
|
|
|
definition functor_mk_eq' {F₁ F₂ : C → D} {H₁ : Π(a b : C), hom a b → hom (F₁ a) (F₁ b)}
|
2015-02-24 00:54:16 +00:00
|
|
|
|
{H₂ : Π(a b : C), hom a b → hom (F₂ a) (F₂ b)} (id₁ id₂ comp₁ comp₂)
|
|
|
|
|
(pF : F₁ = F₂) (pH : pF ▹ H₁ = H₂)
|
|
|
|
|
: functor.mk F₁ H₁ id₁ comp₁ = functor.mk F₂ H₂ id₂ comp₂ :=
|
2015-04-24 21:00:32 +00:00
|
|
|
|
apd01111 functor.mk pF pH !is_hprop.elim !is_hprop.elim
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
2015-03-13 22:27:29 +00:00
|
|
|
|
definition functor_eq' {F₁ F₂ : C ⇒ D}
|
|
|
|
|
: Π(p : to_fun_ob F₁ = to_fun_ob F₂),
|
|
|
|
|
(transport (λx, Πa b f, hom (x a) (x b)) p (to_fun_hom F₁) = to_fun_hom F₂) → F₁ = F₂ :=
|
|
|
|
|
functor.rec_on F₁ (λO₁ H₁ id₁ comp₁, functor.rec_on F₂ (λO₂ H₂ id₂ comp₂ p, !functor_mk_eq'))
|
|
|
|
|
|
|
|
|
|
definition functor_mk_eq {F₁ F₂ : C → D} {H₁ : Π(a b : C), hom a b → hom (F₁ a) (F₁ b)}
|
2015-03-13 14:32:48 +00:00
|
|
|
|
{H₂ : Π(a b : C), hom a b → hom (F₂ a) (F₂ b)} (id₁ id₂ comp₁ comp₂) (pF : F₁ ∼ F₂)
|
|
|
|
|
(pH : Π(a b : C) (f : hom a b), hom_of_eq (pF b) ∘ H₁ a b f ∘ inv_of_eq (pF a) = H₂ a b f)
|
2015-02-24 00:54:16 +00:00
|
|
|
|
: functor.mk F₁ H₁ id₁ comp₁ = functor.mk F₂ H₂ id₂ comp₂ :=
|
2015-03-13 22:27:29 +00:00
|
|
|
|
functor_mk_eq' id₁ id₂ comp₁ comp₂ (eq_of_homotopy pF)
|
2015-02-24 00:54:16 +00:00
|
|
|
|
(eq_of_homotopy (λc, eq_of_homotopy (λc', eq_of_homotopy (λf,
|
|
|
|
|
begin
|
2015-03-13 14:32:48 +00:00
|
|
|
|
apply concat, rotate_left 1, exact (pH c c' f),
|
|
|
|
|
apply concat, rotate_left 1, apply transport_hom,
|
|
|
|
|
apply concat, rotate_left 1,
|
|
|
|
|
exact (pi_transport_constant (eq_of_homotopy pF) (H₁ c c') f),
|
2015-04-24 21:00:32 +00:00
|
|
|
|
apply (apd10' f),
|
2015-04-27 18:20:15 +00:00
|
|
|
|
apply concat, rotate_left 1, esimp,
|
2015-03-13 14:32:48 +00:00
|
|
|
|
exact (pi_transport_constant (eq_of_homotopy pF) (H₁ c) c'),
|
2015-04-24 21:00:32 +00:00
|
|
|
|
apply (apd10' c'),
|
2015-04-27 18:20:15 +00:00
|
|
|
|
apply concat, rotate_left 1, esimp,
|
2015-03-13 14:32:48 +00:00
|
|
|
|
exact (pi_transport_constant (eq_of_homotopy pF) H₁ c),
|
2015-05-03 05:22:31 +00:00
|
|
|
|
reflexivity
|
2015-02-24 00:54:16 +00:00
|
|
|
|
end))))
|
|
|
|
|
|
2015-03-13 22:27:29 +00:00
|
|
|
|
definition functor_eq {F₁ F₂ : C ⇒ D} : Π(p : to_fun_ob F₁ ∼ to_fun_ob F₂),
|
|
|
|
|
(Π(a b : C) (f : hom a b), hom_of_eq (p b) ∘ F₁ f ∘ inv_of_eq (p a) = F₂ f) → F₁ = F₂ :=
|
|
|
|
|
functor.rec_on F₁ (λO₁ H₁ id₁ comp₁, functor.rec_on F₂ (λO₂ H₂ id₂ comp₂ p, !functor_mk_eq))
|
|
|
|
|
|
|
|
|
|
definition functor_mk_eq_constant {F : C → D} {H₁ : Π(a b : C), hom a b → hom (F a) (F b)}
|
2015-02-24 00:54:16 +00:00
|
|
|
|
{H₂ : Π(a b : C), hom a b → hom (F a) (F b)} (id₁ id₂ comp₁ comp₂)
|
|
|
|
|
(pH : Π(a b : C) (f : hom a b), H₁ a b f = H₂ a b f)
|
|
|
|
|
: functor.mk F H₁ id₁ comp₁ = functor.mk F H₂ id₂ comp₂ :=
|
2015-03-13 22:27:29 +00:00
|
|
|
|
functor_eq (λc, idp) (λa b f, !id_leftright ⬝ !pH)
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
2015-03-19 20:35:38 +00:00
|
|
|
|
protected definition preserve_iso (F : C ⇒ D) {a b : C} (f : hom a b) [H : is_iso f] :
|
|
|
|
|
is_iso (F f) :=
|
|
|
|
|
begin
|
|
|
|
|
fapply @is_iso.mk, apply (F (f⁻¹)),
|
2015-05-03 05:22:31 +00:00
|
|
|
|
repeat (apply concat ; symmetry ; apply (respect_comp F) ;
|
2015-03-19 20:35:38 +00:00
|
|
|
|
apply concat ; apply (ap (λ x, to_fun_hom F x)) ;
|
2015-04-06 16:24:09 +00:00
|
|
|
|
(apply left_inverse | apply right_inverse);
|
2015-03-19 20:35:38 +00:00
|
|
|
|
apply (respect_id F) ),
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
attribute preserve_iso [instance]
|
|
|
|
|
|
2015-03-19 21:25:20 +00:00
|
|
|
|
protected definition respect_inv (F : C ⇒ D) {a b : C} (f : hom a b)
|
|
|
|
|
[H : is_iso f] [H' : is_iso (F f)] :
|
2015-03-19 20:35:38 +00:00
|
|
|
|
F (f⁻¹) = (F f)⁻¹ :=
|
|
|
|
|
begin
|
|
|
|
|
fapply @left_inverse_eq_right_inverse, apply (F f),
|
2015-05-03 05:22:31 +00:00
|
|
|
|
transitivity to_fun_hom F (f⁻¹ ∘ f),
|
|
|
|
|
{symmetry, apply (respect_comp F)},
|
|
|
|
|
{transitivity to_fun_hom F category.id,
|
|
|
|
|
{congruence, apply left_inverse},
|
|
|
|
|
{apply respect_id}},
|
|
|
|
|
apply right_inverse
|
2015-03-19 20:35:38 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-03-13 22:27:29 +00:00
|
|
|
|
protected definition assoc (H : C ⇒ D) (G : B ⇒ C) (F : A ⇒ B) :
|
2015-02-24 00:54:16 +00:00
|
|
|
|
H ∘f (G ∘f F) = (H ∘f G) ∘f F :=
|
2015-03-13 22:27:29 +00:00
|
|
|
|
!functor_mk_eq_constant (λa b f, idp)
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
2015-03-13 22:27:29 +00:00
|
|
|
|
protected definition id_left (F : C ⇒ D) : id ∘f F = F :=
|
|
|
|
|
functor.rec_on F (λF1 F2 F3 F4, !functor_mk_eq_constant (λa b f, idp))
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
2015-03-13 22:27:29 +00:00
|
|
|
|
protected definition id_right (F : C ⇒ D) : F ∘f id = F :=
|
|
|
|
|
functor.rec_on F (λF1 F2 F3 F4, !functor_mk_eq_constant (λa b f, idp))
|
|
|
|
|
|
|
|
|
|
protected definition comp_id_eq_id_comp (F : C ⇒ D) : F ∘f functor.id = functor.id ∘f F :=
|
|
|
|
|
!functor.id_right ⬝ !functor.id_left⁻¹
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
2014-12-18 02:49:37 +00:00
|
|
|
|
-- "functor C D" is equivalent to a certain sigma type
|
|
|
|
|
protected definition sigma_char :
|
2015-02-28 00:50:06 +00:00
|
|
|
|
(Σ (to_fun_ob : C → D)
|
|
|
|
|
(to_fun_hom : Π ⦃a b : C⦄, hom a b → hom (to_fun_ob a) (to_fun_ob b)),
|
|
|
|
|
(Π (a : C), to_fun_hom (ID a) = ID (to_fun_ob a)) ×
|
2014-12-18 02:49:37 +00:00
|
|
|
|
(Π {a b c : C} (g : hom b c) (f : hom a b),
|
2015-02-28 00:50:06 +00:00
|
|
|
|
to_fun_hom (g ∘ f) = to_fun_hom g ∘ to_fun_hom f)) ≃ (functor C D) :=
|
2014-12-18 02:49:37 +00:00
|
|
|
|
begin
|
2015-02-24 00:54:16 +00:00
|
|
|
|
fapply equiv.MK,
|
2015-02-24 21:27:57 +00:00
|
|
|
|
{intro S, fapply functor.mk,
|
2014-12-18 02:49:37 +00:00
|
|
|
|
exact (S.1), exact (S.2.1),
|
2015-04-06 20:23:38 +00:00
|
|
|
|
-- TODO(Leo): investigate why we need to use relaxed-exact (rexact) tactic here
|
|
|
|
|
exact (pr₁ S.2.2), rexact (pr₂ S.2.2)},
|
2015-02-24 00:54:16 +00:00
|
|
|
|
{intro F,
|
2015-04-30 18:00:39 +00:00
|
|
|
|
cases F with d1 d2 d3 d4,
|
2015-03-17 00:08:45 +00:00
|
|
|
|
exact ⟨d1, d2, (d3, @d4)⟩},
|
2015-02-24 00:54:16 +00:00
|
|
|
|
{intro F,
|
|
|
|
|
cases F,
|
2015-05-03 05:22:31 +00:00
|
|
|
|
reflexivity},
|
2015-02-24 00:54:16 +00:00
|
|
|
|
{intro S,
|
2015-04-30 18:00:39 +00:00
|
|
|
|
cases S with d1 S2,
|
|
|
|
|
cases S2 with d2 P1,
|
2015-02-24 00:54:16 +00:00
|
|
|
|
cases P1,
|
2015-05-03 05:22:31 +00:00
|
|
|
|
reflexivity},
|
2014-12-18 02:49:37 +00:00
|
|
|
|
end
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-04-25 03:04:24 +00:00
|
|
|
|
section
|
|
|
|
|
local attribute precategory.is_hset_hom [priority 1001]
|
|
|
|
|
protected theorem is_hset_functor [instance]
|
|
|
|
|
[HD : is_hset D] : is_hset (functor C D) :=
|
2015-04-25 04:20:59 +00:00
|
|
|
|
by apply is_trunc_equiv_closed; apply functor.sigma_char
|
2015-01-03 06:29:11 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-03-13 22:27:29 +00:00
|
|
|
|
definition functor_mk_eq'_idp (F : C → D) (H : Π(a b : C), hom a b → hom (F a) (F b))
|
|
|
|
|
(id comp) : functor_mk_eq' id id comp comp (idpath F) (idpath H) = idp :=
|
2015-03-13 14:32:48 +00:00
|
|
|
|
begin
|
2015-04-24 21:00:32 +00:00
|
|
|
|
fapply (apd011 (apd01111 functor.mk idp idp)),
|
2015-03-13 22:27:29 +00:00
|
|
|
|
apply is_hset.elim,
|
|
|
|
|
apply is_hset.elim
|
2015-03-13 16:13:50 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-03-13 22:27:29 +00:00
|
|
|
|
definition functor_eq'_idp (F : C ⇒ D) : functor_eq' idp idp = (idpath F) :=
|
|
|
|
|
by (cases F; apply functor_mk_eq'_idp)
|
|
|
|
|
|
|
|
|
|
definition functor_eq_eta' {F₁ F₂ : C ⇒ D} (p : F₁ = F₂)
|
2015-04-24 21:00:32 +00:00
|
|
|
|
: functor_eq' (ap to_fun_ob p) (!transport_compose⁻¹ ⬝ apd to_fun_hom p) = p :=
|
2015-03-13 16:13:50 +00:00
|
|
|
|
begin
|
2015-03-13 22:27:29 +00:00
|
|
|
|
cases p, cases F₁,
|
|
|
|
|
apply concat, rotate_left 1, apply functor_eq'_idp,
|
2015-05-03 05:22:31 +00:00
|
|
|
|
esimp
|
2015-03-13 14:32:48 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-03-13 22:27:29 +00:00
|
|
|
|
definition functor_eq2' {F₁ F₂ : C ⇒ D} {p₁ p₂ : to_fun_ob F₁ = to_fun_ob F₂} (q₁ q₂)
|
|
|
|
|
(r : p₁ = p₂) : functor_eq' p₁ q₁ = functor_eq' p₂ q₂ :=
|
|
|
|
|
by cases r; apply (ap (functor_eq' p₂)); apply is_hprop.elim
|
|
|
|
|
|
2015-03-13 14:32:48 +00:00
|
|
|
|
definition functor_eq2 {F₁ F₂ : C ⇒ D} (p q : F₁ = F₂) (r : ap010 to_fun_ob p ∼ ap010 to_fun_ob q)
|
|
|
|
|
: p = q :=
|
|
|
|
|
begin
|
2015-04-30 18:00:39 +00:00
|
|
|
|
cases F₁ with ob₁ hom₁ id₁ comp₁,
|
|
|
|
|
cases F₂ with ob₂ hom₂ id₂ comp₂,
|
2015-03-13 22:27:29 +00:00
|
|
|
|
rewrite [-functor_eq_eta' p, -functor_eq_eta' q],
|
|
|
|
|
apply functor_eq2',
|
|
|
|
|
apply ap_eq_ap_of_homotopy,
|
|
|
|
|
exact r,
|
2015-03-13 14:32:48 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-04-27 21:29:56 +00:00
|
|
|
|
-- definition ap010_functor_eq' {F₁ F₂ : C ⇒ D} (p : to_fun_ob F₁ = to_fun_ob F₂)
|
2015-03-13 14:32:48 +00:00
|
|
|
|
-- (q : p ▹ F₁ = F₂) (c : C) :
|
2015-04-27 21:29:56 +00:00
|
|
|
|
-- ap to_fun_ob (functor_eq (apd10 p) (λa b f, _)) = p := sorry
|
2015-03-13 14:32:48 +00:00
|
|
|
|
-- begin
|
2015-04-30 18:00:39 +00:00
|
|
|
|
-- cases F₂, revert q, apply (homotopy.rec_on p), clear p, esimp, intros p q,
|
|
|
|
|
-- cases p, clear e_1 e_2,
|
2015-03-13 14:32:48 +00:00
|
|
|
|
-- end
|
|
|
|
|
|
|
|
|
|
-- TODO: remove sorry
|
2015-03-13 22:27:29 +00:00
|
|
|
|
definition ap010_functor_eq {F₁ F₂ : C ⇒ D} (p : to_fun_ob F₁ ∼ to_fun_ob F₂)
|
2015-03-13 14:32:48 +00:00
|
|
|
|
(q : (λ(a b : C) (f : hom a b), hom_of_eq (p b) ∘ F₁ f ∘ inv_of_eq (p a)) ∼3 to_fun_hom F₂) (c : C) :
|
2015-03-13 22:27:29 +00:00
|
|
|
|
ap010 to_fun_ob (functor_eq p q) c = p c :=
|
2015-03-13 14:32:48 +00:00
|
|
|
|
begin
|
2015-05-01 22:07:28 +00:00
|
|
|
|
cases F₂, revert q, eapply (homotopy.rec_on p), clear p, esimp, intro p q,
|
2015-03-13 14:32:48 +00:00
|
|
|
|
apply sorry,
|
2015-03-13 22:27:29 +00:00
|
|
|
|
--apply (homotopy3.rec_on q), clear q, intro q,
|
|
|
|
|
--cases p, --TODO: report: this fails
|
2015-03-13 14:32:48 +00:00
|
|
|
|
end
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-03-13 22:27:29 +00:00
|
|
|
|
definition ap010_functor_mk_eq_constant {F : C → D} {H₁ : Π(a b : C), hom a b → hom (F a) (F b)}
|
|
|
|
|
{H₂ : Π(a b : C), hom a b → hom (F a) (F b)} {id₁ id₂ comp₁ comp₂}
|
|
|
|
|
(pH : Π(a b : C) (f : hom a b), H₁ a b f = H₂ a b f) (c : C) :
|
|
|
|
|
ap010 to_fun_ob (functor_mk_eq_constant id₁ id₂ comp₁ comp₂ pH) c = idp :=
|
|
|
|
|
!ap010_functor_eq
|
|
|
|
|
|
|
|
|
|
--do we need this theorem?
|
|
|
|
|
definition compose_pentagon (K : D ⇒ E) (H : C ⇒ D) (G : B ⇒ C) (F : A ⇒ B) :
|
|
|
|
|
(calc K ∘f H ∘f G ∘f F = (K ∘f H) ∘f G ∘f F : functor.assoc
|
|
|
|
|
... = ((K ∘f H) ∘f G) ∘f F : functor.assoc)
|
|
|
|
|
=
|
|
|
|
|
(calc K ∘f H ∘f G ∘f F = K ∘f (H ∘f G) ∘f F : ap (λx, K ∘f x) !functor.assoc
|
|
|
|
|
... = (K ∘f H ∘f G) ∘f F : functor.assoc
|
|
|
|
|
... = ((K ∘f H) ∘f G) ∘f F : ap (λx, x ∘f F) !functor.assoc) :=
|
|
|
|
|
sorry
|
2015-03-13 14:32:48 +00:00
|
|
|
|
-- begin
|
2015-03-13 22:27:29 +00:00
|
|
|
|
-- apply functor_eq2,
|
|
|
|
|
-- intro a,
|
|
|
|
|
-- rewrite +ap010_con,
|
|
|
|
|
-- -- rewrite +ap010_ap,
|
|
|
|
|
-- -- apply sorry
|
|
|
|
|
-- /-to prove this we need a stronger ap010-lemma, something like
|
|
|
|
|
-- ap010 (λy, to_fun_ob (f y)) (functor_mk_eq_constant ...) c = idp
|
|
|
|
|
-- or something another way of getting ap out of ap010
|
|
|
|
|
-- -/
|
|
|
|
|
-- --rewrite +ap010_ap,
|
|
|
|
|
-- --unfold functor.assoc,
|
|
|
|
|
-- --rewrite ap010_functor_mk_eq_constant,
|
2015-03-13 14:32:48 +00:00
|
|
|
|
-- end
|
|
|
|
|
|
|
|
|
|
end functor
|