feat(category): start on proof of yoneda preserves limits and limit functor is left adjoint

This commit is contained in:
Floris van Doorn 2015-10-27 18:02:00 -04:00 committed by Leonardo de Moura
parent a99a99f047
commit e14754a337
7 changed files with 243 additions and 63 deletions

View file

@ -3,7 +3,7 @@ Copyright (c) 2015 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn Authors: Floris van Doorn
Cones Cones of a diagram in a category
-/ -/
import ..nat_trans ..category import ..nat_trans ..category
@ -16,38 +16,48 @@ namespace category
(c : C) (c : C)
(η : constant_functor I c ⟹ F) (η : constant_functor I c ⟹ F)
local attribute cone_obj.c [coercion] variables {I C D : Precategory} {F : I ⇒ C} {x y z : cone_obj F} {i : I}
definition cone_to_obj [unfold 4] := @cone_obj.c
definition cone_to_nat [unfold 4] (c : cone_obj F) : constant_functor I (cone_to_obj c) ⟹ F :=
cone_obj.η c
local attribute cone_to_obj [coercion]
variables {I C : Precategory} {F : I ⇒ C} {x y z : cone_obj F}
structure cone_hom (x y : cone_obj F) := structure cone_hom (x y : cone_obj F) :=
(f : x ⟶ y) (f : x ⟶ y)
(p : Πi, cone_obj.η y i ∘ f = cone_obj.η x i) (p : Πi, cone_to_nat y i ∘ f = cone_to_nat x i)
local attribute cone_hom.f [coercion] definition cone_to_hom [unfold 6] := @cone_hom.f
definition cone_to_eq [unfold 6] (f : cone_hom x y) (i : I)
: cone_to_nat y i ∘ (cone_to_hom f) = cone_to_nat x i :=
cone_hom.p f i
local attribute cone_to_hom [coercion]
definition cone_id [constructor] (x : cone_obj F) : cone_hom x x := definition cone_id [constructor] (x : cone_obj F) : cone_hom x x :=
cone_hom.mk id cone_hom.mk id
(λi, !id_right) (λi, !id_right)
definition cone_comp [constructor] (g : cone_hom y z) (f : cone_hom x y) : cone_hom x z := definition cone_comp [constructor] (g : cone_hom y z) (f : cone_hom x y) : cone_hom x z :=
cone_hom.mk (cone_hom.f g ∘ cone_hom.f f) cone_hom.mk (cone_to_hom g ∘ cone_to_hom f)
abstract λi, by rewrite [assoc, +cone_hom.p] end abstract λi, by rewrite [assoc, +cone_to_eq] end
definition cone_obj_eq (p : cone_obj.c x = cone_obj.c y) definition cone_obj_eq (p : cone_to_obj x = cone_to_obj y)
(q : Πi, cone_obj.η x i = cone_obj.η y i ∘ hom_of_eq p) : x = y := (q : Πi, cone_to_nat x i = cone_to_nat y i ∘ hom_of_eq p) : x = y :=
begin begin
induction x, induction y, esimp at *, induction p, apply ap (cone_obj.mk c), induction x, induction y, esimp at *, induction p, apply ap (cone_obj.mk c),
apply nat_trans_eq, intro i, exact q i ⬝ !id_right apply nat_trans_eq, intro i, exact q i ⬝ !id_right
end end
theorem c_cone_obj_eq (p : cone_obj.c x = cone_obj.c y) theorem c_cone_obj_eq (p : cone_to_obj x = cone_to_obj y)
(q : Πi, cone_obj.η x i = cone_obj.η y i ∘ hom_of_eq p) : ap cone_obj.c (cone_obj_eq p q) = p := (q : Πi, cone_to_nat x i = cone_to_nat y i ∘ hom_of_eq p) : ap cone_to_obj (cone_obj_eq p q) = p :=
begin begin
induction x, induction y, esimp at *, induction p, induction x, induction y, esimp at *, induction p,
esimp [cone_obj_eq], rewrite [-ap_compose,↑function.compose,ap_constant] esimp [cone_obj_eq], rewrite [-ap_compose,↑function.compose,ap_constant]
end end
theorem cone_hom_eq {f f' : cone_hom x y} (q : cone_hom.f f = cone_hom.f f') : f = f' := theorem cone_hom_eq {f f' : cone_hom x y} (q : cone_to_hom f = cone_to_hom f') : f = f' :=
begin begin
induction f, induction f', esimp at *, induction q, apply ap (cone_hom.mk f), induction f, induction f', esimp at *, induction q, apply ap (cone_hom.mk f),
apply @is_hprop.elim, apply pi.is_trunc_pi, intro x, apply is_trunc_eq, -- type class fails apply @is_hprop.elim, apply pi.is_trunc_pi, intro x, apply is_trunc_eq, -- type class fails
@ -59,7 +69,7 @@ namespace category
@precategory.mk _ cone_hom @precategory.mk _ cone_hom
abstract begin abstract begin
intro x y, intro x y,
assert H : cone_hom x y ≃ Σ(f : x ⟶ y), Πi, cone_obj.η y i ∘ f = cone_obj.η x i, assert H : cone_hom x y ≃ Σ(f : x ⟶ y), Πi, cone_to_nat y i ∘ f = cone_to_nat x i,
{ fapply equiv.MK, { fapply equiv.MK,
{ intro f, induction f, constructor, assumption}, { intro f, induction f, constructor, assumption},
{ intro v, induction v, constructor, assumption}, { intro v, induction v, constructor, assumption},
@ -81,16 +91,16 @@ namespace category
precategory.Mk (precategory_cone F) precategory.Mk (precategory_cone F)
variable {F} variable {F}
definition cone_iso_pr1 (h : x ≅ y) : cone_obj.c x ≅ cone_obj.c y := definition cone_iso_pr1 [constructor] (h : x ≅ y) : cone_to_obj x ≅ cone_to_obj y :=
iso.MK iso.MK
(cone_hom.f (to_hom h)) (cone_to_hom (to_hom h))
(cone_hom.f (to_inv h)) (cone_to_hom (to_inv h))
(ap cone_hom.f (to_left_inverse h)) (ap cone_to_hom (to_left_inverse h))
(ap cone_hom.f (to_right_inverse h)) (ap cone_to_hom (to_right_inverse h))
definition cone_iso.mk (f : cone_obj.c x ≅ cone_obj.c y) definition cone_iso.mk [constructor] (f : cone_to_obj x ≅ cone_to_obj y)
(p : Πi, cone_obj.η y i ∘ to_hom f = cone_obj.η x i) : x ≅ y := (p : Πi, cone_to_nat y i ∘ to_hom f = cone_to_nat x i) : x ≅ y :=
begin begin
fapply iso.MK, fapply iso.MK,
{ exact !cone_hom.mk p}, { exact !cone_hom.mk p},
@ -102,11 +112,11 @@ namespace category
end end
variables (x y) variables (x y)
definition cone_iso_equiv [constructor] : (x ≅ y) ≃ Σ(f : cone_obj.c x ≅ cone_obj.c y), definition cone_iso_equiv [constructor] : (x ≅ y) ≃ Σ(f : cone_to_obj x ≅ cone_to_obj y),
Πi, cone_obj.η y i ∘ to_hom f = cone_obj.η x i := Πi, cone_to_nat y i ∘ to_hom f = cone_to_nat x i :=
begin begin
fapply equiv.MK, fapply equiv.MK,
{ intro h, exact ⟨cone_iso_pr1 h, cone_hom.p (to_hom h)⟩}, { intro h, exact ⟨cone_iso_pr1 h, cone_to_eq (to_hom h)⟩},
{ intro v, exact cone_iso.mk v.1 v.2}, { intro v, exact cone_iso.mk v.1 v.2},
{ intro v, induction v with f p, fapply sigma_eq: esimp, { intro v, induction v with f p, fapply sigma_eq: esimp,
{ apply iso_eq, reflexivity}, { apply iso_eq, reflexivity},
@ -114,12 +124,11 @@ namespace category
{ intro h, esimp, apply iso_eq, apply cone_hom_eq, reflexivity}, { intro h, esimp, apply iso_eq, apply cone_hom_eq, reflexivity},
end end
--set_option pp.implicit true definition cone_eq_equiv : (x = y) ≃ Σ(f : cone_to_obj x = cone_to_obj y),
definition cone_eq_equiv : (x = y) ≃ Σ(f : cone_obj.c x = cone_obj.c y), Πi, cone_to_nat y i ∘ hom_of_eq f = cone_to_nat x i :=
Πi, cone_obj.η y i ∘ hom_of_eq f = cone_obj.η x i :=
begin begin
fapply equiv.MK, fapply equiv.MK,
{ intro r, fapply sigma.mk, exact ap cone_obj.c r, induction r, intro i, apply id_right}, { intro r, fapply sigma.mk, exact ap cone_to_obj r, induction r, intro i, apply id_right},
{ intro v, induction v with p q, induction x with c η, induction y with c' η', esimp at *, { intro v, induction v with p q, induction x with c η, induction y with c' η', esimp at *,
apply cone_obj_eq p, esimp, intro i, exact (q i)⁻¹}, apply cone_obj_eq p, esimp, intro i, exact (q i)⁻¹},
{ intro v, induction v with p q, induction x with c η, induction y with c' η', esimp at *, { intro v, induction v with p q, induction x with c η, induction y with c' η', esimp at *,
@ -137,9 +146,9 @@ namespace category
intro x y, intro x y,
fapply is_equiv_of_equiv_of_homotopy, fapply is_equiv_of_equiv_of_homotopy,
{ exact calc { exact calc
(x = y) ≃ (Σ(f : cone_obj.c x = cone_obj.c y), Πi, cone_obj.η y i ∘ hom_of_eq f = cone_obj.η x i) (x = y) ≃ (Σ(f : cone_to_obj x = cone_to_obj y), Πi, cone_to_nat y i ∘ hom_of_eq f = cone_to_nat x i)
: cone_eq_equiv : cone_eq_equiv
... ≃ (Σ(f : cone_obj.c x ≅ cone_obj.c y), Πi, cone_obj.η y i ∘ to_hom f = cone_obj.η x i) ... ≃ (Σ(f : cone_to_obj x ≅ cone_to_obj y), Πi, cone_to_nat y i ∘ to_hom f = cone_to_nat x i)
: sigma_equiv_sigma !eq_equiv_iso (λa, !equiv.refl) : sigma_equiv_sigma !eq_equiv_iso (λa, !equiv.refl)
... ≃ (x ≅ y) : cone_iso_equiv }, ... ≃ (x ≅ y) : cone_iso_equiv },
{ intro p, induction p, esimp [equiv.trans,equiv.symm], esimp [sigma_functor], { intro p, induction p, esimp [equiv.trans,equiv.symm], esimp [sigma_functor],
@ -156,5 +165,15 @@ namespace category
end is_univalent end is_univalent
definition cone_obj_compose [constructor] (G : C ⇒ D) (x : cone_obj F) : cone_obj (G ∘f F) :=
begin
fapply cone_obj.mk,
{ exact G x},
{ fapply change_natural_map,
{ refine ((G ∘fn cone_to_nat x) ∘n _), apply nat_trans_of_eq, fapply functor_eq: esimp,
intro i j k, esimp, rewrite [id_leftright,respect_id]},
{ intro i, esimp, exact G (cone_to_nat x i)},
{ intro i, esimp, rewrite [ap010_functor_eq, ▸*, id_right]}}
end
end category end category

View file

@ -123,7 +123,7 @@ namespace category
{ exact εn}, { exact εn},
{ exact adjointify_adjH}, { exact adjointify_adjH},
{ exact adjointify_adjK}, { exact adjointify_adjK},
{ exact @(is_iso_nat_trans _) (λc, !is_iso_inverse)}, { exact @(is_natural_iso _) (λc, !is_iso_inverse)},
{ unfold εn, apply iso.struct, }, { unfold εn, apply iso.struct, },
end end

View file

@ -17,8 +17,7 @@ namespace category
include C include C
definition is_terminal [class] (c : ob) := Πd, is_contr (d ⟶ c) definition is_terminal [class] (c : ob) := Πd, is_contr (d ⟶ c)
definition is_contr_of_is_terminal (c d : ob) [H : is_terminal d] definition is_contr_of_is_terminal (c d : ob) [H : is_terminal d] : is_contr (c ⟶ d) :=
: is_contr (c ⟶ d) :=
H c H c
local attribute is_contr_of_is_terminal [instance] local attribute is_contr_of_is_terminal [instance]
@ -32,7 +31,8 @@ namespace category
definition eq_terminal_morphism [H : is_terminal c'] (f : c ⟶ c') : f = terminal_morphism c c' := definition eq_terminal_morphism [H : is_terminal c'] (f : c ⟶ c') : f = terminal_morphism c c' :=
!is_hprop.elim !is_hprop.elim
definition terminal_iso_terminal {c c' : ob} (H : is_terminal c) (K : is_terminal c') : c ≅ c' := definition terminal_iso_terminal (c c' : ob) [H : is_terminal c] [K : is_terminal c']
: c ≅ c' :=
iso.MK !terminal_morphism !terminal_morphism !hom_terminal_eq !hom_terminal_eq iso.MK !terminal_morphism !terminal_morphism !hom_terminal_eq !hom_terminal_eq
local attribute is_terminal [reducible] local attribute is_terminal [reducible]
@ -51,15 +51,14 @@ namespace category
variable {D} variable {D}
definition terminal_object_iso_terminal_object (H₁ H₂ : has_terminal_object D) definition terminal_object_iso_terminal_object (H₁ H₂ : has_terminal_object D)
: @terminal_object D H₁ ≅ @terminal_object D H₂ := : @terminal_object D H₁ ≅ @terminal_object D H₂ :=
terminal_iso_terminal (@has_terminal_object.is_terminal D H₁) !terminal_iso_terminal
(@has_terminal_object.is_terminal D H₂)
theorem is_hprop_has_terminal_object [instance] (D : Category) theorem is_hprop_has_terminal_object [instance] (D : Category)
: is_hprop (has_terminal_object D) := : is_hprop (has_terminal_object D) :=
begin begin
apply is_hprop.mk, intro t₁ t₂, induction t₁ with d₁ H₁, induction t₂ with d₂ H₂, apply is_hprop.mk, intro t₁ t₂, induction t₁ with d₁ H₁, induction t₂ with d₂ H₂,
assert p : d₁ = d₂, assert p : d₁ = d₂,
{ apply eq_of_iso, apply terminal_iso_terminal H₁ H₂}, { apply eq_of_iso, apply terminal_iso_terminal},
induction p, exact ap _ !is_hprop.elim induction p, exact ap _ !is_hprop.elim
end end
@ -101,11 +100,73 @@ namespace category
definition is_terminal_limit_cone [instance] : is_terminal (limit_cone F) := definition is_terminal_limit_cone [instance] : is_terminal (limit_cone F) :=
has_terminal_object.is_terminal _ has_terminal_object.is_terminal _
section specific_limit
omit H
variable {F}
variables (x : cone_obj F) [K : is_terminal x]
include K
definition to_limit_object : D :=
cone_to_obj x
definition to_limit_nat_trans : constant_functor I (to_limit_object x) ⟹ F :=
cone_to_nat x
definition to_limit_morphism (i : I) : to_limit_object x ⟶ F i :=
to_limit_nat_trans x i
theorem to_limit_commute {i j : I} (f : i ⟶ j)
: to_fun_hom F f ∘ to_limit_morphism x i = to_limit_morphism x j :=
naturality (to_limit_nat_trans x) f ⬝ !id_right
definition to_limit_cone_obj [constructor] {d : D} {η : Πi, d ⟶ F i}
(p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) : cone_obj F :=
cone_obj.mk d (nat_trans.mk η (λa b f, p f ⬝ !id_right⁻¹))
definition to_hom_limit {d : D} (η : Πi, d ⟶ F i)
(p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) : d ⟶ to_limit_object x :=
cone_to_hom (terminal_morphism (to_limit_cone_obj x p) x)
theorem to_hom_limit_commute {d : D} (η : Πi, d ⟶ F i)
(p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) (i : I)
: to_limit_morphism x i ∘ to_hom_limit x η p = η i :=
cone_to_eq (terminal_morphism (to_limit_cone_obj x p) x) i
definition to_limit_cone_hom [constructor] {d : D} {η : Πi, d ⟶ F i}
(p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) {h : d ⟶ to_limit_object x}
(q : Πi, to_limit_morphism x i ∘ h = η i)
: cone_hom (to_limit_cone_obj x p) x :=
cone_hom.mk h q
variable {x}
theorem to_eq_hom_limit {d : D} {η : Πi, d ⟶ F i}
(p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) {h : d ⟶ to_limit_object x}
(q : Πi, to_limit_morphism x i ∘ h = η i) : h = to_hom_limit x η p :=
ap cone_to_hom (eq_terminal_morphism (to_limit_cone_hom x p q))
theorem to_limit_cone_unique {d : D} {η : Πi, d ⟶ F i}
(p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j)
{h₁ : d ⟶ to_limit_object x} (q₁ : Πi, to_limit_morphism x i ∘ h₁ = η i)
{h₂ : d ⟶ to_limit_object x} (q₂ : Πi, to_limit_morphism x i ∘ h₂ = η i): h₁ = h₂ :=
to_eq_hom_limit p q₁ ⬝ (to_eq_hom_limit p q₂)⁻¹
omit K
definition to_limit_object_iso_to_limit_object [constructor] (x y : cone_obj F)
[K : is_terminal x] [L : is_terminal y] : to_limit_object x ≅ to_limit_object y :=
cone_iso_pr1 !terminal_iso_terminal
end specific_limit
/-
TODO: relate below definitions to above definitions.
However, type class resolution seems to fail...
-/
definition limit_object : D := definition limit_object : D :=
cone_obj.c (limit_cone F) cone_to_obj (limit_cone F)
definition limit_nat_trans : constant_functor I (limit_object F) ⟹ F := definition limit_nat_trans : constant_functor I (limit_object F) ⟹ F :=
cone_obj.η (limit_cone F) cone_to_nat (limit_cone F)
definition limit_morphism (i : I) : limit_object F ⟶ F i := definition limit_morphism (i : I) : limit_object F ⟶ F i :=
limit_nat_trans F i limit_nat_trans F i
@ -123,12 +184,12 @@ namespace category
variable {H} variable {H}
definition hom_limit {d : D} (η : Πi, d ⟶ F i) definition hom_limit {d : D} (η : Πi, d ⟶ F i)
(p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) : d ⟶ limit_object F := (p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) : d ⟶ limit_object F :=
cone_hom.f (@(terminal_morphism (limit_cone_obj F p) _) (is_terminal_limit_cone _)) cone_to_hom (@(terminal_morphism (limit_cone_obj F p) _) (is_terminal_limit_cone _))
theorem hom_limit_commute {d : D} (η : Πi, d ⟶ F i) theorem hom_limit_commute {d : D} (η : Πi, d ⟶ F i)
(p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) (i : I) (p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) (i : I)
: limit_morphism F i ∘ hom_limit F η p = η i := : limit_morphism F i ∘ hom_limit F η p = η i :=
cone_hom.p (@(terminal_morphism (limit_cone_obj F p) _) (is_terminal_limit_cone _)) i cone_to_eq (@(terminal_morphism (limit_cone_obj F p) _) (is_terminal_limit_cone _)) i
definition limit_cone_hom [constructor] {d : D} {η : Πi, d ⟶ F i} definition limit_cone_hom [constructor] {d : D} {η : Πi, d ⟶ F i}
(p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) {h : d ⟶ limit_object F} (p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) {h : d ⟶ limit_object F}
@ -139,7 +200,7 @@ namespace category
theorem eq_hom_limit {d : D} {η : Πi, d ⟶ F i} theorem eq_hom_limit {d : D} {η : Πi, d ⟶ F i}
(p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) {h : d ⟶ limit_object F} (p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) {h : d ⟶ limit_object F}
(q : Πi, limit_morphism F i ∘ h = η i) : h = hom_limit F η p := (q : Πi, limit_morphism F i ∘ h = η i) : h = hom_limit F η p :=
ap cone_hom.f (@eq_terminal_morphism _ _ _ _ (is_terminal_limit_cone _) (limit_cone_hom F p q)) ap cone_to_hom (@eq_terminal_morphism _ _ _ _ (is_terminal_limit_cone _) (limit_cone_hom F p q))
theorem limit_cone_unique {d : D} {η : Πi, d ⟶ F i} theorem limit_cone_unique {d : D} {η : Πi, d ⟶ F i}
(p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j) (p : Π⦃i j : I⦄ (f : i ⟶ j), to_fun_hom F f ∘ η i = η j)
@ -153,25 +214,10 @@ namespace category
omit H omit H
-- notation `noinstances` t:max := by+ with_options [elaborator.ignore_instances true] (exact t)
-- definition noinstance (t : tactic) : tactic := with_options [elaborator.ignore_instances true] t
variable (F) variable (F)
definition limit_object_iso_limit_object [constructor] (H₁ H₂ : has_limits_of_shape D I) : definition limit_object_iso_limit_object [constructor] (H₁ H₂ : has_limits_of_shape D I) :
@(limit_object F) H₁ ≅ @(limit_object F) H₂ := @(limit_object F) H₁ ≅ @(limit_object F) H₂ :=
begin cone_iso_pr1 !terminal_object_iso_terminal_object
fapply iso.MK,
{ apply hom_limit, apply @(limit_commute F) H₁},
{ apply @(hom_limit F) H₁, apply limit_commute},
{ exact abstract begin fapply limit_cone_unique,
{ apply limit_commute},
{ intro i, rewrite [assoc, hom_limit_commute], apply hom_limit_commute},
{ intro i, apply id_right} end end},
{ exact abstract begin fapply limit_cone_unique,
{ apply limit_commute},
{ intro i, rewrite [assoc, hom_limit_commute], apply hom_limit_commute},
{ intro i, apply id_right} end end}
end
definition limit_functor [constructor] (D I : Precategory) [H : has_limits_of_shape D I] definition limit_functor [constructor] (D I : Precategory) [H : has_limits_of_shape D I]
: D ^c I ⇒ D := : D ^c I ⇒ D :=

View file

@ -0,0 +1,32 @@
/-
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
colimit_functor ⊣ Δ ⊣ limit_functor
-/
import ..colimits ..functor.adjoint2
open functor category is_trunc prod yoneda --remove
namespace category
definition limit_functor_adjoint [constructor] (D I : Precategory) [H : has_limits_of_shape D I] :
limit_functor D I ⊣ constant_diagram D I :=
adjoint.mk'
begin
fapply natural_iso.MK: esimp,
{ intro Fd f, induction Fd with F d, esimp at *,
exact sorry},
{ exact sorry},
{ exact sorry},
{ exact sorry},
{ exact sorry}
end
-- set_option pp.universes true
-- print Precategory_functor
-- print limit_functor_adjoint
-- print adjoint.mk'
end category

View file

@ -0,0 +1,82 @@
/-
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
Functors preserving limits
-/
import ..colimits ..yoneda
open functor yoneda is_trunc nat_trans
namespace category
variables {I C D : Precategory} {F : I ⇒ C} {G : C ⇒ D}
/- notions of preservation of limits -/
definition preserves_limits_of_shape [class] (G : C ⇒ D) (I : Precategory)
[H : has_limits_of_shape C I] :=
Π(F : I ⇒ C), is_terminal (cone_obj_compose G (limit_cone F))
definition preserves_existing_limits_of_shape [class] (G : C ⇒ D) (I : Precategory) :=
Π(F : I ⇒ C) [H : has_terminal_object (cone F)],
is_terminal (cone_obj_compose G (terminal_object (cone F)))
definition preserves_existing_limits [class] (G : C ⇒ D) :=
Π(I : Precategory) (F : I ⇒ C) [H : has_terminal_object (cone F)],
is_terminal (cone_obj_compose G (terminal_object (cone F)))
definition preserves_limits [class] (G : C ⇒ D) [H : is_complete C] :=
Π(I : Precategory) [H : has_limits_of_shape C I] (F : I ⇒ C),
is_terminal (cone_obj_compose G (limit_cone F))
definition preserves_chosen_limits_of_shape [class] (G : C ⇒ D) (I : Precategory)
[H : has_limits_of_shape C I] [H : has_limits_of_shape D I] :=
Π(F : I ⇒ C), cone_obj_compose G (limit_cone F) = limit_cone (G ∘f F)
definition preserves_chosen_limits [class] (G : C ⇒ D)
[H : is_complete C] [H : is_complete D] :=
Π(I : Precategory) (F : I ⇒ C), cone_obj_compose G (limit_cone F) = limit_cone (G ∘f F)
/- basic instances -/
definition preserves_limits_of_shape_of_preserves_limits [instance] (G : C ⇒ D)
(I : Precategory) [H : is_complete C] [H : preserves_limits G]
: preserves_limits_of_shape G I := H I
definition preserves_chosen_limits_of_shape_of_preserves_chosen_limits [instance] (G : C ⇒ D)
(I : Precategory) [H : is_complete C] [H : is_complete D] [K : preserves_chosen_limits G]
: preserves_chosen_limits_of_shape G I := K I
/- yoneda preserves existing limits -/
definition preserves_existing_limits_yoneda_embedding_lemma [constructor] (x : cone_obj F)
[H : is_terminal x] {G : Cᵒᵖ ⇒ cset} (η : constant_functor I G ⟹ ɏ ∘f F) :
G ⟹ to_fun_ob ɏ (cone_to_obj x) :=
begin
fapply nat_trans.mk: esimp,
{ intro c x, fapply to_hom_limit,
{ intro i, exact η i c x},
{ intro i j k, exact sorry}},
{ intro c c' f, apply eq_of_homotopy, intro x, exact sorry}
end
theorem preserves_existing_limits_yoneda_embedding (C : Precategory)
: preserves_existing_limits (yoneda_embedding C) :=
begin
intro I F H Gη, induction H with x H, induction Gη with G η, esimp at *,
fapply is_contr.mk,
{ fapply cone_hom.mk: esimp,
{ exact sorry
-- fapply nat_trans.mk: esimp,
-- { intro c x, esimp [yoneda_embedding], fapply to_hom_limit,
-- { apply has_terminal_object.is_terminal}, --this should be solved by type class res.
-- { intro i, induction dη with d η, esimp at *, },
-- { intro i j k, }},
-- { }
},
{ exact sorry}},
{ exact sorry}
end
end category

View file

@ -178,7 +178,7 @@ namespace nat_trans
definition nf_id (η : G ⟹ H) (c : C) : (η ∘nf 1) c = η c := definition nf_id (η : G ⟹ H) (c : C) : (η ∘nf 1) c = η c :=
idp idp
definition nat_trans_of_eq [reducible] (p : F = G) : F ⟹ G := definition nat_trans_of_eq [reducible] [constructor] (p : F = G) : F ⟹ G :=
nat_trans.mk (λc, hom_of_eq (ap010 to_fun_ob p c)) nat_trans.mk (λc, hom_of_eq (ap010 to_fun_ob p c))
(λa b f, eq.rec_on p (!id_right ⬝ !id_left⁻¹)) (λa b f, eq.rec_on p (!id_right ⬝ !id_left⁻¹))

View file

@ -38,10 +38,11 @@ namespace yoneda
local attribute Category.to.precategory category.to_precategory [constructor] local attribute Category.to.precategory category.to_precategory [constructor]
-- should this be defined as "yoneda_embedding Cᵒᵖ"? -- should this be defined as "yoneda_embedding Cᵒᵖ"?
definition contravariant_yoneda_embedding [reducible] (C : Precategory) : Cᵒᵖ ⇒ cset ^c C := definition contravariant_yoneda_embedding [constructor] [reducible]
(C : Precategory) : Cᵒᵖ ⇒ cset ^c C :=
functor_curry !hom_functor functor_curry !hom_functor
definition yoneda_embedding (C : Precategory) : C ⇒ cset ^c Cᵒᵖ := definition yoneda_embedding [constructor] (C : Precategory) : C ⇒ cset ^c Cᵒᵖ :=
functor_curry (!hom_functor ∘f !prod_flip_functor) functor_curry (!hom_functor ∘f !prod_flip_functor)
notation `ɏ` := yoneda_embedding _ notation `ɏ` := yoneda_embedding _
@ -66,7 +67,7 @@ namespace yoneda
exact ap10 !respect_id x end end}, exact ap10 !respect_id x end end},
{ exact abstract begin intro η, esimp, apply nat_trans_eq, { exact abstract begin intro η, esimp, apply nat_trans_eq,
intro c', esimp, apply eq_of_homotopy, intro c', esimp, apply eq_of_homotopy,
intro f, esimp [yoneda_embedding] at f, intro f,
transitivity (F f ∘ η c) id, reflexivity, transitivity (F f ∘ η c) id, reflexivity,
rewrite naturality, esimp [yoneda_embedding], rewrite [id_left], apply ap _ !id_left end end}, rewrite naturality, esimp [yoneda_embedding], rewrite [id_left], apply ap _ !id_left end end},
end end
@ -137,8 +138,8 @@ namespace yoneda
rewrite -eq_of_iso_refl, rewrite -eq_of_iso_refl,
apply ap eq_of_iso, apply iso_eq, esimp, apply ap eq_of_iso, apply iso_eq, esimp,
apply nat_trans_eq, intro c', apply nat_trans_eq, intro c',
apply eq_of_homotopy, esimp [yoneda_embedding], intro f, apply eq_of_homotopy, intro f,
rewrite [category.category.id_left], apply id_right} rewrite [▸*, category.category.id_left], apply id_right}
end end
definition is_representable {C : Precategory} (F : Cᵒᵖ ⇒ cset) := Σ(c : C), ɏ c ≅ F definition is_representable {C : Precategory} (F : Cᵒᵖ ⇒ cset) := Σ(c : C), ɏ c ≅ F