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-10-02 23:54:27 +00:00
|
|
|
|
|
|
|
Yoneda embedding and Yoneda lemma
|
2015-02-21 00:30:32 +00:00
|
|
|
-/
|
|
|
|
|
2015-10-27 23:02:42 +00:00
|
|
|
import .examples .attributes
|
2015-02-21 00:30:32 +00:00
|
|
|
|
2015-10-02 23:54:27 +00:00
|
|
|
open category eq functor prod.ops is_trunc iso is_equiv equiv category.set nat_trans lift
|
2015-02-21 00:30:32 +00:00
|
|
|
|
|
|
|
namespace yoneda
|
2015-02-24 22:09:20 +00:00
|
|
|
|
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-10-27 22:02:00 +00:00
|
|
|
definition contravariant_yoneda_embedding [constructor] [reducible]
|
|
|
|
(C : Precategory) : Cᵒᵖ ⇒ cset ^c C :=
|
2015-03-13 14:32:48 +00:00
|
|
|
functor_curry !hom_functor
|
2015-02-27 05:45:21 +00:00
|
|
|
|
2015-10-28 02:03:59 +00:00
|
|
|
/-
|
|
|
|
we use (change_fun) to make sure that (to_fun_ob (yoneda_embedding C) c) will reduce to
|
|
|
|
(hom_functor_left c) instead of (functor_curry_rev_ob (hom_functor C) c)
|
|
|
|
-/
|
2015-10-27 22:02:00 +00:00
|
|
|
definition yoneda_embedding [constructor] (C : Precategory) : C ⇒ cset ^c Cᵒᵖ :=
|
2015-10-28 02:03:59 +00:00
|
|
|
--(functor_curry_rev !hom_functor)
|
|
|
|
change_fun
|
|
|
|
(functor_curry_rev !hom_functor)
|
|
|
|
hom_functor_left
|
|
|
|
nat_trans_hom_functor_left
|
|
|
|
idp
|
|
|
|
idpo
|
2015-10-27 23:02:42 +00:00
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
notation `ɏ` := yoneda_embedding _
|
|
|
|
|
2015-10-02 23:54:27 +00:00
|
|
|
definition yoneda_lemma_hom [constructor] {C : Precategory} (c : C) (F : Cᵒᵖ ⇒ cset)
|
2015-09-02 23:41:19 +00:00
|
|
|
(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)⁻¹,
|
2015-10-02 23:54:27 +00:00
|
|
|
exact ap10 !(@respect_comp Cᵒᵖ cset)⁻¹ x}
|
2015-09-02 23:41:19 +00:00
|
|
|
end
|
|
|
|
|
2015-09-22 17:11:33 +00:00
|
|
|
definition yoneda_lemma_equiv [constructor] {C : Precategory} (c : C)
|
2015-10-02 23:54:27 +00:00
|
|
|
(F : Cᵒᵖ ⇒ cset) : hom (ɏ c) F ≃ lift (F c) :=
|
2015-09-02 23:41:19 +00:00
|
|
|
begin
|
|
|
|
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,
|
2015-10-27 22:02:00 +00:00
|
|
|
intro f,
|
2015-09-02 23:41:19 +00:00
|
|
|
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
|
|
|
|
|
2015-10-02 23:54:27 +00:00
|
|
|
definition yoneda_lemma {C : Precategory} (c : C) (F : Cᵒᵖ ⇒ cset) :
|
2015-09-22 17:11:33 +00:00
|
|
|
homset (ɏ c) F ≅ lift_functor (F c) :=
|
|
|
|
begin
|
|
|
|
apply iso_of_equiv, esimp, apply yoneda_lemma_equiv,
|
|
|
|
end
|
|
|
|
|
2015-10-02 23:54:27 +00:00
|
|
|
theorem yoneda_lemma_natural_ob {C : Precategory} (F : Cᵒᵖ ⇒ cset) {c c' : C} (f : c' ⟶ c)
|
2015-09-03 04:46:11 +00:00
|
|
|
(η : ɏ c ⟹ F) :
|
|
|
|
to_fun_hom (lift_functor ∘f F) f (to_hom (yoneda_lemma c F) η) =
|
2015-09-22 17:11:33 +00:00
|
|
|
to_hom (yoneda_lemma c' F) (η ∘n to_fun_hom ɏ f) :=
|
2015-09-03 04:46:11 +00:00
|
|
|
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
|
|
|
|
|
2015-09-22 17:11:33 +00:00
|
|
|
-- TODO: Investigate what is the bottleneck to type check the next theorem
|
|
|
|
|
|
|
|
-- attribute yoneda_lemma lift_functor Precategory_hset precategory_hset homset
|
|
|
|
-- yoneda_embedding nat_trans.compose functor_nat_trans_compose [reducible]
|
|
|
|
-- attribute tlift functor.compose [reducible]
|
2015-10-02 23:54:27 +00:00
|
|
|
theorem yoneda_lemma_natural_functor.{u v} {C : Precategory.{u v}} (c : C) (F F' : Cᵒᵖ ⇒ cset)
|
2015-09-03 04:46:11 +00:00
|
|
|
(θ : F ⟹ F') (η : to_fun_ob ɏ c ⟹ F) :
|
2015-09-22 17:11:33 +00:00
|
|
|
(lift_functor.{v u} ∘fn θ) c (to_hom (yoneda_lemma c F) η) =
|
|
|
|
proof to_hom (yoneda_lemma c F') (θ ∘n η) qed :=
|
2015-09-03 04:46:11 +00:00
|
|
|
by reflexivity
|
|
|
|
|
2015-09-22 17:11:33 +00:00
|
|
|
-- theorem xx.{u v} {C : Precategory.{u v}} (c : C) (F F' : Cᵒᵖ ⇒ set)
|
|
|
|
-- (θ : F ⟹ F') (η : to_fun_ob ɏ c ⟹ F) :
|
|
|
|
-- proof _ qed =
|
|
|
|
-- to_hom (yoneda_lemma c F') (θ ∘n η) :=
|
|
|
|
-- by reflexivity
|
|
|
|
|
|
|
|
-- theorem yy.{u v} {C : Precategory.{u v}} (c : C) (F F' : Cᵒᵖ ⇒ set)
|
|
|
|
-- (θ : F ⟹ F') (η : to_fun_ob ɏ c ⟹ F) :
|
|
|
|
-- (lift_functor.{v u} ∘fn θ) c (to_hom (yoneda_lemma c F) η) =
|
|
|
|
-- proof _ qed :=
|
|
|
|
-- by reflexivity
|
|
|
|
|
2015-09-03 04:46:11 +00:00
|
|
|
definition fully_faithful_yoneda_embedding [instance] (C : Precategory) :
|
2015-10-02 23:54:27 +00:00
|
|
|
fully_faithful (ɏ : C ⇒ cset ^c Cᵒᵖ) :=
|
2015-09-03 04:46:11 +00:00
|
|
|
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-22 17:11:33 +00:00
|
|
|
definition is_embedding_yoneda_embedding (C : Category) :
|
2015-10-02 23:54:27 +00:00
|
|
|
is_embedding (ɏ : C → Cᵒᵖ ⇒ cset) :=
|
2015-09-03 04:46:11 +00:00
|
|
|
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⁻¹ᵉ},
|
2015-10-23 05:12:34 +00:00
|
|
|
{ intro p, induction p, esimp [equiv.trans, equiv.symm, to_fun_iso], -- to_fun_iso not unfolded
|
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',
|
2015-10-27 22:02:00 +00:00
|
|
|
apply eq_of_homotopy, intro f,
|
|
|
|
rewrite [▸*, category.category.id_left], apply id_right}
|
2015-09-02 23:41:19 +00:00
|
|
|
end
|
|
|
|
|
2015-10-02 23:54:27 +00:00
|
|
|
definition is_representable {C : Precategory} (F : Cᵒᵖ ⇒ cset) := Σ(c : C), ɏ c ≅ F
|
2015-09-03 15:11:46 +00:00
|
|
|
|
2015-10-23 05:12:34 +00:00
|
|
|
section
|
|
|
|
set_option apply.class_instance false
|
|
|
|
definition is_hprop_representable {C : Category} (F : Cᵒᵖ ⇒ cset)
|
|
|
|
: is_hprop (is_representable F) :=
|
|
|
|
begin
|
|
|
|
fapply is_trunc_equiv_closed,
|
|
|
|
{ exact proof fiber.sigma_char ɏ F qed ⬝e sigma.sigma_equiv_sigma_id (λc, !eq_equiv_iso)},
|
|
|
|
{ apply function.is_hprop_fiber_of_is_embedding, apply is_embedding_yoneda_embedding}
|
|
|
|
end
|
2015-09-03 15:11:46 +00:00
|
|
|
end
|
|
|
|
|
2015-10-23 05:12:34 +00:00
|
|
|
|
|
|
|
|
2015-03-13 14:32:48 +00:00
|
|
|
end yoneda
|