2015-02-21 00:30:32 +00:00
|
|
|
|
/-
|
2015-03-17 00:08:45 +00:00
|
|
|
|
Copyright (c) 2015 Floris van Doorn. All rights reserved.
|
2015-02-21 00:30:32 +00:00
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
2015-02-28 00:02:18 +00:00
|
|
|
|
Authors: Floris van Doorn
|
2015-02-21 00:30:32 +00:00
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
--note: modify definition in category.set
|
2015-04-25 04:20:59 +00:00
|
|
|
|
import .constructions.functor .constructions.hset .constructions.product .constructions.opposite
|
2015-09-03 04:46:11 +00:00
|
|
|
|
.adjoint
|
2015-02-21 00:30:32 +00:00
|
|
|
|
|
2015-03-13 14:32:48 +00:00
|
|
|
|
open category eq category.ops functor prod.ops is_trunc iso
|
2015-02-21 00:30:32 +00:00
|
|
|
|
|
|
|
|
|
namespace yoneda
|
2015-09-02 23:41:19 +00:00
|
|
|
|
-- set_option class.conservative false
|
2015-02-24 22:09:20 +00:00
|
|
|
|
|
2015-02-26 18:19:54 +00:00
|
|
|
|
definition representable_functor_assoc [C : Precategory] {a1 a2 a3 a4 a5 a6 : C}
|
|
|
|
|
(f1 : hom a5 a6) (f2 : hom a4 a5) (f3 : hom a3 a4) (f4 : hom a2 a3) (f5 : hom a1 a2)
|
|
|
|
|
: (f1 ∘ f2) ∘ f3 ∘ (f4 ∘ f5) = f1 ∘ (f2 ∘ f3 ∘ f4) ∘ f5 :=
|
2015-02-21 00:30:32 +00:00
|
|
|
|
calc
|
2015-03-05 19:10:34 +00:00
|
|
|
|
_ = f1 ∘ f2 ∘ f3 ∘ f4 ∘ f5 : by rewrite -assoc
|
|
|
|
|
... = f1 ∘ (f2 ∘ f3) ∘ f4 ∘ f5 : by rewrite -assoc
|
|
|
|
|
... = f1 ∘ ((f2 ∘ f3) ∘ f4) ∘ f5 : by rewrite -(assoc (f2 ∘ f3) _ _)
|
|
|
|
|
... = _ : by rewrite (assoc f2 f3 f4)
|
2015-02-21 00:30:32 +00:00
|
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
|
definition hom_functor.{u v} [constructor] (C : Precategory.{u v}) : Cᵒᵖ ×c C ⇒ set.{v} :=
|
2015-06-27 00:09:50 +00:00
|
|
|
|
functor.mk
|
|
|
|
|
(λ (x : Cᵒᵖ ×c C), @homset (Cᵒᵖ) C x.1 x.2)
|
|
|
|
|
(λ (x y : Cᵒᵖ ×c C) (f : @category.precategory.hom (Cᵒᵖ ×c C) (Cᵒᵖ ×c C) x y) (h : @homset (Cᵒᵖ) C x.1 x.2),
|
2015-09-02 23:41:19 +00:00
|
|
|
|
f.2 ∘[C] (h ∘[C] f.1))
|
2015-06-27 00:09:50 +00:00
|
|
|
|
(λ x, @eq_of_homotopy _ _ _ (ID (@homset Cᵒᵖ C x.1 x.2)) (λ h, concat (by apply @id_left) (by apply @id_right)))
|
|
|
|
|
(λ x y z g f,
|
|
|
|
|
eq_of_homotopy (by intros; apply @representable_functor_assoc))
|
2015-02-21 00:30:32 +00:00
|
|
|
|
end yoneda
|
|
|
|
|
|
2015-02-26 18:19:54 +00:00
|
|
|
|
open is_equiv equiv
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
|
|
|
|
namespace functor
|
|
|
|
|
open prod nat_trans
|
2015-02-27 05:45:21 +00:00
|
|
|
|
variables {C D E : Precategory} (F : C ×c D ⇒ E) (G : C ⇒ E ^c D)
|
2015-09-02 23:41:19 +00:00
|
|
|
|
definition functor_curry_ob [reducible] [constructor] (c : C) : E ^c D :=
|
2015-02-24 00:54:16 +00:00
|
|
|
|
functor.mk (λd, F (c,d))
|
|
|
|
|
(λd d' g, F (id, g))
|
|
|
|
|
(λd, !respect_id)
|
2015-03-05 19:10:34 +00:00
|
|
|
|
(λd₁ d₂ d₃ g' g, calc
|
2015-03-13 00:07:27 +00:00
|
|
|
|
F (id, g' ∘ g) = F (id ∘ id, g' ∘ g) : by rewrite id_comp
|
|
|
|
|
... = F ((id,g') ∘ (id, g)) : by esimp
|
2015-03-13 03:06:25 +00:00
|
|
|
|
... = F (id,g') ∘ F (id, g) : by rewrite respect_comp)
|
2015-03-05 19:10:34 +00:00
|
|
|
|
|
2015-02-24 00:54:16 +00:00
|
|
|
|
local abbreviation Fob := @functor_curry_ob
|
2015-02-27 05:45:21 +00:00
|
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
|
definition functor_curry_hom [constructor] ⦃c c' : C⦄ (f : c ⟶ c') : Fob F c ⟹ Fob F c' :=
|
2015-06-27 00:09:50 +00:00
|
|
|
|
begin
|
|
|
|
|
fapply @nat_trans.mk,
|
|
|
|
|
{intro d, exact F (f, id)},
|
|
|
|
|
{intro d d' g, calc
|
|
|
|
|
F (id, g) ∘ F (f, id) = F (id ∘ f, g ∘ id) : respect_comp F
|
2015-03-13 00:07:27 +00:00
|
|
|
|
... = F (f, g ∘ id) : by rewrite id_left
|
|
|
|
|
... = F (f, g) : by rewrite id_right
|
|
|
|
|
... = F (f ∘ id, g) : by rewrite id_right
|
|
|
|
|
... = F (f ∘ id, id ∘ g) : by rewrite id_left
|
2015-06-27 00:09:50 +00:00
|
|
|
|
... = F (f, id) ∘ F (id, g) : (respect_comp F (f, id) (id, g))⁻¹ᵖ
|
|
|
|
|
}
|
|
|
|
|
end
|
2015-02-27 05:45:21 +00:00
|
|
|
|
local abbreviation Fhom := @functor_curry_hom
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
2015-03-13 14:32:48 +00:00
|
|
|
|
theorem functor_curry_hom_def ⦃c c' : C⦄ (f : c ⟶ c') (d : D) :
|
2015-02-28 00:50:06 +00:00
|
|
|
|
(Fhom F f) d = to_fun_hom F (f, id) := idp
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
2015-02-28 06:16:20 +00:00
|
|
|
|
theorem functor_curry_id (c : C) : Fhom F (ID c) = nat_trans.id :=
|
2015-04-27 21:29:56 +00:00
|
|
|
|
nat_trans_eq (λd, respect_id F _)
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
2015-02-28 06:16:20 +00:00
|
|
|
|
theorem functor_curry_comp ⦃c c' c'' : C⦄ (f' : c' ⟶ c'') (f : c ⟶ c')
|
2015-02-27 05:45:21 +00:00
|
|
|
|
: Fhom F (f' ∘ f) = Fhom F f' ∘n Fhom F f :=
|
2015-06-27 00:09:50 +00:00
|
|
|
|
begin
|
|
|
|
|
apply @nat_trans_eq,
|
|
|
|
|
intro d, calc
|
2015-06-18 22:52:08 +00:00
|
|
|
|
natural_map (Fhom F (f' ∘ f)) d = F (f' ∘ f, id) : by rewrite functor_curry_hom_def
|
2015-03-13 00:07:27 +00:00
|
|
|
|
... = F (f' ∘ f, id ∘ id) : by rewrite id_comp
|
|
|
|
|
... = F ((f',id) ∘ (f, id)) : by esimp
|
2015-06-18 22:52:08 +00:00
|
|
|
|
... = F (f',id) ∘ F (f, id) : by rewrite [respect_comp F]
|
2015-06-27 00:09:50 +00:00
|
|
|
|
... = natural_map ((Fhom F f') ∘ (Fhom F f)) d : by esimp
|
|
|
|
|
end
|
2015-02-27 05:45:21 +00:00
|
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
|
definition functor_curry [reducible] [constructor] : C ⇒ E ^c D :=
|
2015-02-24 00:54:16 +00:00
|
|
|
|
functor.mk (functor_curry_ob F)
|
2015-02-27 05:45:21 +00:00
|
|
|
|
(functor_curry_hom F)
|
2015-02-24 00:54:16 +00:00
|
|
|
|
(functor_curry_id F)
|
|
|
|
|
(functor_curry_comp F)
|
|
|
|
|
|
2015-02-27 05:45:21 +00:00
|
|
|
|
definition functor_uncurry_ob [reducible] (p : C ×c D) : E :=
|
2015-02-28 00:50:06 +00:00
|
|
|
|
to_fun_ob (G p.1) p.2
|
2015-02-27 05:45:21 +00:00
|
|
|
|
local abbreviation Gob := @functor_uncurry_ob
|
|
|
|
|
|
|
|
|
|
definition functor_uncurry_hom ⦃p p' : C ×c D⦄ (f : hom p p') : Gob G p ⟶ Gob G p' :=
|
2015-02-28 00:50:06 +00:00
|
|
|
|
to_fun_hom (to_fun_ob G p'.1) f.2 ∘ natural_map (to_fun_hom G f.1) p.2
|
2015-02-27 05:45:21 +00:00
|
|
|
|
local abbreviation Ghom := @functor_uncurry_hom
|
|
|
|
|
|
2015-02-28 06:16:20 +00:00
|
|
|
|
theorem functor_uncurry_id (p : C ×c D) : Ghom G (ID p) = id :=
|
2015-02-27 05:45:21 +00:00
|
|
|
|
calc
|
2015-03-13 00:07:27 +00:00
|
|
|
|
Ghom G (ID p) = to_fun_hom (to_fun_ob G p.1) id ∘ natural_map (to_fun_hom G id) p.2 : by esimp
|
|
|
|
|
... = id ∘ natural_map (to_fun_hom G id) p.2 : by rewrite respect_id
|
|
|
|
|
... = id ∘ natural_map nat_trans.id p.2 : by rewrite respect_id
|
|
|
|
|
... = id : id_comp
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
2015-02-28 06:16:20 +00:00
|
|
|
|
theorem functor_uncurry_comp ⦃p p' p'' : C ×c D⦄ (f' : p' ⟶ p'') (f : p ⟶ p')
|
2015-02-27 05:45:21 +00:00
|
|
|
|
: Ghom G (f' ∘ f) = Ghom G f' ∘ Ghom G f :=
|
|
|
|
|
calc
|
|
|
|
|
Ghom G (f' ∘ f)
|
2015-03-13 00:07:27 +00:00
|
|
|
|
= to_fun_hom (to_fun_ob G p''.1) (f'.2 ∘ f.2) ∘ natural_map (to_fun_hom G (f'.1 ∘ f.1)) p.2 : by esimp
|
2015-02-28 00:50:06 +00:00
|
|
|
|
... = (to_fun_hom (to_fun_ob G p''.1) f'.2 ∘ to_fun_hom (to_fun_ob G p''.1) f.2)
|
2015-09-02 23:41:19 +00:00
|
|
|
|
∘ natural_map (to_fun_hom G (f'.1 ∘ f.1)) p.2 : by rewrite respect_comp
|
2015-02-28 00:50:06 +00:00
|
|
|
|
... = (to_fun_hom (to_fun_ob G p''.1) f'.2 ∘ to_fun_hom (to_fun_ob G p''.1) f.2)
|
2015-09-02 23:41:19 +00:00
|
|
|
|
∘ natural_map (to_fun_hom G f'.1 ∘ to_fun_hom G f.1) p.2 : by rewrite respect_comp
|
2015-02-28 00:50:06 +00:00
|
|
|
|
... = (to_fun_hom (to_fun_ob G p''.1) f'.2 ∘ to_fun_hom (to_fun_ob G p''.1) f.2)
|
2015-09-02 23:41:19 +00:00
|
|
|
|
∘ (natural_map (to_fun_hom G f'.1) p.2 ∘ natural_map (to_fun_hom G f.1) p.2) : by esimp
|
2015-02-28 00:50:06 +00:00
|
|
|
|
... = (to_fun_hom (to_fun_ob G p''.1) f'.2 ∘ natural_map (to_fun_hom G f'.1) p'.2)
|
|
|
|
|
∘ (to_fun_hom (to_fun_ob G p'.1) f.2 ∘ natural_map (to_fun_hom G f.1) p.2) :
|
2015-06-18 22:52:08 +00:00
|
|
|
|
by rewrite [square_prepostcompose (!naturality⁻¹ᵖ) _ _]
|
2015-03-13 00:07:27 +00:00
|
|
|
|
... = Ghom G f' ∘ Ghom G f : by esimp
|
2015-02-27 05:45:21 +00:00
|
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
|
definition functor_uncurry [reducible] [constructor] : C ×c D ⇒ E :=
|
2015-02-27 05:45:21 +00:00
|
|
|
|
functor.mk (functor_uncurry_ob G)
|
|
|
|
|
(functor_uncurry_hom G)
|
|
|
|
|
(functor_uncurry_id G)
|
2015-03-13 16:13:50 +00:00
|
|
|
|
(functor_uncurry_comp G)
|
2015-03-13 14:32:48 +00:00
|
|
|
|
|
|
|
|
|
theorem functor_uncurry_functor_curry : functor_uncurry (functor_curry F) = F :=
|
2015-03-13 22:27:29 +00:00
|
|
|
|
functor_eq (λp, ap (to_fun_ob F) !prod.eta)
|
2015-02-27 05:45:21 +00:00
|
|
|
|
begin
|
2015-04-30 18:00:39 +00:00
|
|
|
|
intro cd cd' fg,
|
|
|
|
|
cases cd with c d, cases cd' with c' d', cases fg with f g,
|
2015-05-03 05:22:31 +00:00
|
|
|
|
transitivity to_fun_hom (functor_uncurry (functor_curry F)) (f, g),
|
|
|
|
|
apply id_leftright,
|
2015-03-13 14:32:48 +00:00
|
|
|
|
show (functor_uncurry (functor_curry F)) (f, g) = F (f,g),
|
2015-02-27 05:45:21 +00:00
|
|
|
|
from calc
|
2015-03-13 00:07:27 +00:00
|
|
|
|
(functor_uncurry (functor_curry F)) (f, g) = to_fun_hom F (id, g) ∘ to_fun_hom F (f, id) : by esimp
|
2015-09-03 15:11:46 +00:00
|
|
|
|
... = F (id ∘ f, g ∘ id) : by krewrite [-respect_comp F (id,g) (f,id)]
|
2015-05-03 05:22:31 +00:00
|
|
|
|
... = F (f, g ∘ id) : by rewrite id_left
|
|
|
|
|
... = F (f,g) : by rewrite id_right,
|
2015-02-27 05:45:21 +00:00
|
|
|
|
end
|
2015-03-13 14:32:48 +00:00
|
|
|
|
|
|
|
|
|
definition functor_curry_functor_uncurry_ob (c : C)
|
|
|
|
|
: functor_curry (functor_uncurry G) c = G c :=
|
2015-02-27 05:45:21 +00:00
|
|
|
|
begin
|
2015-03-13 22:27:29 +00:00
|
|
|
|
fapply functor_eq,
|
2015-05-03 05:22:31 +00:00
|
|
|
|
{intro d, reflexivity},
|
2015-04-30 18:00:39 +00:00
|
|
|
|
{intro d d' g,
|
2015-03-13 14:32:48 +00:00
|
|
|
|
apply concat, apply id_leftright,
|
|
|
|
|
show to_fun_hom (functor_curry (functor_uncurry G) c) g = to_fun_hom (G c) g,
|
|
|
|
|
from calc
|
|
|
|
|
to_fun_hom (functor_curry (functor_uncurry G) c) g
|
|
|
|
|
= to_fun_hom (G c) g ∘ natural_map (to_fun_hom G (ID c)) d : by esimp
|
2015-05-03 05:22:31 +00:00
|
|
|
|
... = to_fun_hom (G c) g ∘ natural_map (ID (G c)) d : by rewrite respect_id
|
|
|
|
|
... = to_fun_hom (G c) g ∘ id : by reflexivity
|
2015-06-18 22:52:08 +00:00
|
|
|
|
... = to_fun_hom (G c) g : by rewrite id_right}
|
2015-02-27 05:45:21 +00:00
|
|
|
|
end
|
2015-02-24 00:54:16 +00:00
|
|
|
|
|
2015-03-13 14:32:48 +00:00
|
|
|
|
theorem functor_curry_functor_uncurry : functor_curry (functor_uncurry G) = G :=
|
|
|
|
|
begin
|
2015-03-13 22:27:29 +00:00
|
|
|
|
fapply functor_eq, exact (functor_curry_functor_uncurry_ob G),
|
2015-04-30 18:00:39 +00:00
|
|
|
|
intro c c' f,
|
2015-04-27 21:29:56 +00:00
|
|
|
|
fapply nat_trans_eq,
|
2015-03-13 14:32:48 +00:00
|
|
|
|
intro d,
|
|
|
|
|
apply concat,
|
|
|
|
|
{apply (ap (λx, x ∘ _)),
|
2015-03-13 22:27:29 +00:00
|
|
|
|
apply concat, apply natural_map_hom_of_eq, apply (ap hom_of_eq), apply ap010_functor_eq},
|
2015-05-03 05:22:31 +00:00
|
|
|
|
apply concat,
|
2015-03-13 14:32:48 +00:00
|
|
|
|
{apply (ap (λx, _ ∘ x)), apply (ap (λx, _ ∘ x)),
|
|
|
|
|
apply concat, apply natural_map_inv_of_eq,
|
2015-03-13 22:27:29 +00:00
|
|
|
|
apply (ap (λx, hom_of_eq x⁻¹)), apply ap010_functor_eq},
|
2015-03-13 14:32:48 +00:00
|
|
|
|
apply concat, apply id_leftright,
|
|
|
|
|
apply concat, apply (ap (λx, x ∘ _)), apply respect_id,
|
|
|
|
|
apply id_left
|
|
|
|
|
end
|
|
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
|
definition prod_functor_equiv_functor_functor [constructor] (C D E : Precategory)
|
2015-03-13 14:32:48 +00:00
|
|
|
|
: (C ×c D ⇒ E) ≃ (C ⇒ E ^c D) :=
|
2015-02-27 05:45:21 +00:00
|
|
|
|
equiv.MK functor_curry
|
|
|
|
|
functor_uncurry
|
|
|
|
|
functor_curry_functor_uncurry
|
|
|
|
|
functor_uncurry_functor_curry
|
|
|
|
|
|
|
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
|
definition functor_prod_flip [constructor] (C D : Precategory) : C ×c D ⇒ D ×c C :=
|
2015-03-13 14:32:48 +00:00
|
|
|
|
functor.mk (λp, (p.2, p.1))
|
|
|
|
|
(λp p' h, (h.2, h.1))
|
|
|
|
|
(λp, idp)
|
|
|
|
|
(λp p' p'' h' h, idp)
|
2015-02-27 05:45:21 +00:00
|
|
|
|
|
2015-03-13 14:32:48 +00:00
|
|
|
|
definition functor_prod_flip_functor_prod_flip (C D : Precategory)
|
|
|
|
|
: functor_prod_flip D C ∘f (functor_prod_flip C D) = functor.id :=
|
|
|
|
|
begin
|
2015-03-13 22:27:29 +00:00
|
|
|
|
fapply functor_eq, {intro p, apply prod.eta},
|
2015-04-30 18:00:39 +00:00
|
|
|
|
intro p p' h, cases p with c d, cases p' with c' d',
|
2015-03-13 14:32:48 +00:00
|
|
|
|
apply id_leftright,
|
|
|
|
|
end
|
|
|
|
|
end functor
|
|
|
|
|
open functor
|
|
|
|
|
|
|
|
|
|
namespace yoneda
|
2015-09-02 23:41:19 +00:00
|
|
|
|
open category.set nat_trans lift
|
2015-09-10 22:32:52 +00:00
|
|
|
|
|
|
|
|
|
/-
|
|
|
|
|
These attributes make sure that the fields of the category "set" reduce to the right things
|
|
|
|
|
However, we don't want to have them globally, because that will unfold the composition g ∘ f
|
|
|
|
|
in a Category to category.category.comp g f
|
|
|
|
|
-/
|
|
|
|
|
local attribute Category.to.precategory category.to_precategory [constructor]
|
|
|
|
|
|
2015-09-03 04:46:11 +00:00
|
|
|
|
-- should this be defined as "yoneda_embedding Cᵒᵖ"?
|
2015-03-13 14:32:48 +00:00
|
|
|
|
definition contravariant_yoneda_embedding (C : Precategory) : Cᵒᵖ ⇒ set ^c C :=
|
|
|
|
|
functor_curry !hom_functor
|
2015-02-27 05:45:21 +00:00
|
|
|
|
|
2015-03-13 14:32:48 +00:00
|
|
|
|
definition yoneda_embedding (C : Precategory) : C ⇒ set ^c Cᵒᵖ :=
|
|
|
|
|
functor_curry (!hom_functor ∘f !functor_prod_flip)
|
2015-02-27 05:45:21 +00:00
|
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
|
notation `ɏ` := yoneda_embedding _
|
|
|
|
|
|
|
|
|
|
definition yoneda_lemma_hom [constructor] {C : Precategory} (c : C) (F : Cᵒᵖ ⇒ set)
|
|
|
|
|
(x : trunctype.carrier (F c)) : ɏ c ⟹ F :=
|
|
|
|
|
begin
|
|
|
|
|
fapply nat_trans.mk,
|
|
|
|
|
{ intro c', esimp [yoneda_embedding], intro f, exact F f x},
|
|
|
|
|
{ intro c' c'' f, esimp [yoneda_embedding], apply eq_of_homotopy, intro f',
|
|
|
|
|
refine _ ⬝ ap (λy, to_fun_hom F y x) !(@id_left _ C)⁻¹,
|
|
|
|
|
exact ap10 !(@respect_comp Cᵒᵖ set)⁻¹ x}
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition yoneda_lemma {C : Precategory} (c : C) (F : Cᵒᵖ ⇒ set) :
|
|
|
|
|
homset (ɏ c) F ≅ lift_functor (F c) :=
|
|
|
|
|
begin
|
|
|
|
|
apply iso_of_equiv, esimp,
|
|
|
|
|
fapply equiv.MK,
|
|
|
|
|
{ intro η, exact up (η c id)},
|
|
|
|
|
{ intro x, induction x with x, exact yoneda_lemma_hom c F x},
|
2015-09-03 04:46:11 +00:00
|
|
|
|
{ exact abstract begin intro x, induction x with x, esimp, apply ap up,
|
|
|
|
|
exact ap10 !respect_id x end end},
|
|
|
|
|
{ exact abstract begin intro η, esimp, apply nat_trans_eq,
|
2015-09-02 23:41:19 +00:00
|
|
|
|
intro c', esimp, apply eq_of_homotopy,
|
|
|
|
|
intro f, esimp [yoneda_embedding] at f,
|
|
|
|
|
transitivity (F f ∘ η c) id, reflexivity,
|
2015-09-03 04:46:11 +00:00
|
|
|
|
rewrite naturality, esimp [yoneda_embedding], rewrite [id_left], apply ap _ !id_left end end},
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
theorem yoneda_lemma_natural_ob {C : Precategory} (F : Cᵒᵖ ⇒ set) {c c' : C} (f : c' ⟶ c)
|
|
|
|
|
(η : ɏ c ⟹ F) :
|
|
|
|
|
to_fun_hom (lift_functor ∘f F) f (to_hom (yoneda_lemma c F) η) =
|
|
|
|
|
proof to_hom (yoneda_lemma c' F) (η ∘n to_fun_hom ɏ f) qed :=
|
|
|
|
|
begin
|
|
|
|
|
esimp [yoneda_lemma,yoneda_embedding], apply ap up,
|
|
|
|
|
transitivity (F f ∘ η c) id, reflexivity,
|
|
|
|
|
rewrite naturality,
|
|
|
|
|
esimp [yoneda_embedding],
|
|
|
|
|
apply ap (η c'),
|
|
|
|
|
esimp [yoneda_embedding, Opposite],
|
|
|
|
|
rewrite [+id_left,+id_right],
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
theorem yoneda_lemma_natural_functor.{u v} {C : Precategory.{u v}} (c : C) (F F' : Cᵒᵖ ⇒ set)
|
|
|
|
|
(θ : F ⟹ F') (η : to_fun_ob ɏ c ⟹ F) :
|
|
|
|
|
proof (lift_functor.{v u} ∘fn θ) c (to_hom (yoneda_lemma c F) η) qed =
|
|
|
|
|
(to_hom (yoneda_lemma c F') proof (θ ∘n η : (to_fun_ob ɏ c : Cᵒᵖ ⇒ set) ⟹ F') qed) :=
|
|
|
|
|
by reflexivity
|
|
|
|
|
|
|
|
|
|
definition fully_faithful_yoneda_embedding [instance] (C : Precategory) :
|
|
|
|
|
fully_faithful (ɏ : C ⇒ set ^c Cᵒᵖ) :=
|
|
|
|
|
begin
|
|
|
|
|
intro c c',
|
|
|
|
|
fapply is_equiv_of_equiv_of_homotopy,
|
|
|
|
|
{ symmetry, transitivity _, apply @equiv_of_iso (homset _ _),
|
|
|
|
|
rexact yoneda_lemma c (ɏ c'), esimp [yoneda_embedding], exact !equiv_lift⁻¹ᵉ},
|
|
|
|
|
{ intro f, apply nat_trans_eq, intro c, apply eq_of_homotopy, intro f',
|
|
|
|
|
esimp [equiv.symm,equiv.trans],
|
|
|
|
|
esimp [yoneda_lemma,yoneda_embedding,Opposite],
|
|
|
|
|
rewrite [id_left,id_right]}
|
|
|
|
|
end
|
|
|
|
|
|
2015-09-03 15:11:46 +00:00
|
|
|
|
definition embedding_on_objects_yoneda_embedding (C : Category) :
|
2015-09-03 04:46:11 +00:00
|
|
|
|
is_embedding (ɏ : C → Cᵒᵖ ⇒ set) :=
|
|
|
|
|
begin
|
2015-09-10 22:32:52 +00:00
|
|
|
|
intro c c', fapply is_equiv_of_equiv_of_homotopy,
|
2015-09-03 04:46:11 +00:00
|
|
|
|
{ exact !eq_equiv_iso ⬝e !iso_equiv_F_iso_F ⬝e !eq_equiv_iso⁻¹ᵉ},
|
|
|
|
|
{ intro p, induction p, esimp [equiv.trans, equiv.symm],
|
2015-09-10 22:32:52 +00:00
|
|
|
|
esimp [to_fun_iso],
|
2015-09-03 04:46:11 +00:00
|
|
|
|
rewrite -eq_of_iso_refl,
|
|
|
|
|
apply ap eq_of_iso, apply iso_eq, esimp,
|
|
|
|
|
apply nat_trans_eq, intro c',
|
|
|
|
|
apply eq_of_homotopy, esimp [yoneda_embedding], intro f,
|
|
|
|
|
rewrite [category.category.id_left], apply id_right}
|
2015-09-02 23:41:19 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-09-03 15:11:46 +00:00
|
|
|
|
definition is_representable {C : Precategory} (F : Cᵒᵖ ⇒ set) := Σ(c : C), ɏ c ≅ F
|
|
|
|
|
|
|
|
|
|
definition is_hprop_representable {C : Category} (F : Cᵒᵖ ⇒ set)
|
|
|
|
|
: is_hprop (is_representable F) :=
|
|
|
|
|
begin
|
|
|
|
|
fapply is_trunc_equiv_closed,
|
|
|
|
|
{ transitivity _, rotate 1,
|
|
|
|
|
{ apply sigma.sigma_equiv_sigma_id, intro c, exact !eq_equiv_iso},
|
|
|
|
|
{ apply fiber.sigma_char}},
|
|
|
|
|
{ apply function.is_hprop_fiber_of_is_embedding,
|
|
|
|
|
apply embedding_on_objects_yoneda_embedding}
|
|
|
|
|
end
|
|
|
|
|
|
2015-03-13 14:32:48 +00:00
|
|
|
|
end yoneda
|