feat(category.constructions.hset): prove that the category of sets is complete

This commit is contained in:
Floris van Doorn 2015-09-25 15:02:14 -04:00 committed by Leonardo de Moura
parent f82d1bd566
commit 51edd0ad1d
7 changed files with 97 additions and 29 deletions

View file

@ -37,7 +37,7 @@ namespace category
: is_hprop (f = g) :=
_
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_hom.f f = cone_hom.f f') : f = f' :=
begin
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

View file

@ -6,7 +6,7 @@ Authors: Floris van Doorn, Jakob von Raumer
Category of hsets
-/
import ..category types.equiv ..functor types.lift
import ..category types.equiv ..functor types.lift ..limits
open eq category equiv iso is_equiv is_trunc function sigma
@ -46,16 +46,6 @@ namespace category
local attribute is_equiv_iso_of_equiv [instance]
-- TODO: move
open sigma.ops
definition subtype_eq_inv {A : Type} {B : A → Type} [H : Πa, is_hprop (B a)] (u v : Σa, B a)
: u = v → u.1 = v.1 :=
(subtype_eq u v)⁻¹ᶠ
local attribute subtype_eq_inv [reducible]
definition is_equiv_subtype_eq_inv {A : Type} {B : A → Type} [H : Πa, is_hprop (B a)] (u v : Σa, B a)
: is_equiv (subtype_eq_inv u v) :=
_
definition iso_of_eq_eq_compose (A B : hset) : @iso_of_eq _ _ A B =
@iso_of_equiv A B ∘ @equiv_of_eq A B ∘ subtype_eq_inv _ _ ∘
@ap _ _ (to_fun (trunctype.sigma_char 0)) A B :=
@ -83,7 +73,7 @@ namespace category
(@is_equiv_subtype_eq_inv _ _ _ _ _))
!univalence)
!is_equiv_iso_of_equiv,
assert H₂ : _, from (iso_of_eq_eq_compose A B)⁻¹,
let H₂ := (iso_of_eq_eq_compose A B)⁻¹ in
begin
rewrite H₂ at H₁,
assumption
@ -104,4 +94,47 @@ namespace category
(λa b, lift_functor)
(λa, eq_of_homotopy (λx, by induction x; reflexivity))
(λa b c g f, eq_of_homotopy (λx, by induction x; reflexivity))
open pi sigma.ops
local attribute Category.to.precategory [unfold 1]
local attribute category.to_precategory [unfold 2]
definition is_complete_cone.{u v w} [constructor]
(I : Precategory.{v w}) (F : I ⇒ set.{max u v w}) : cone_obj F :=
begin
fapply cone_obj.mk,
{ fapply trunctype.mk,
{ exact Σ(s : Π(i : I), trunctype.carrier (F i)),
Π{i j : I} (f : i ⟶ j), F f (s i) = (s j)},
{ exact abstract begin apply is_trunc_sigma, intro s,
apply is_trunc_pi, intro i,
apply is_trunc_pi, intro j,
apply is_trunc_pi, intro f,
apply is_trunc_eq end end}},
{ fapply nat_trans.mk,
{ intro i x, esimp at x, exact x.1 i},
{ intro i j f, esimp, apply eq_of_homotopy, intro x, esimp at x, induction x with s p,
esimp, apply p}}
end
definition is_complete_set.{u v w} : is_complete.{(max u v w)+1 (max u v w) v w} set :=
begin
intro I F, fapply has_terminal_object.mk,
{ exact is_complete_cone.{u v w} I F},
{ intro c, esimp at *, induction c with X η, induction η with η p, esimp at *,
fapply is_contr.mk,
{ fapply cone_hom.mk,
{ intro x, esimp at *, fapply sigma.mk,
{ intro i, exact η i x},
{ intro i j f, exact ap10 (p f) x}},
{ intro i, reflexivity}},
{ esimp, intro h, induction h with f q, apply cone_hom_eq, esimp at *,
apply eq_of_homotopy, intro x, fapply sigma_eq: esimp,
{ apply eq_of_homotopy, intro i, exact (ap10 (q i) x)⁻¹},
{ apply is_hprop.elimo,
apply is_trunc_pi, intro i,
apply is_trunc_pi, intro j,
apply is_trunc_pi, intro f,
apply is_trunc_eq}}}
end
end category

View file

@ -48,7 +48,7 @@ namespace iso
split_epi.mk !right_inverse
definition is_iso_id [instance] [priority 500] (a : ob) : is_iso (ID a) :=
is_iso.mk !id_comp !id_comp
is_iso.mk !id_id !id_id
definition is_iso_inverse [instance] [priority 200] (f : a ⟶ b) [H : is_iso f] : is_iso f⁻¹ :=
is_iso.mk !right_inverse !left_inverse
@ -89,13 +89,13 @@ namespace iso
by cases p;apply inverse_unique
definition retraction_id (a : ob) : retraction_of (ID a) = id :=
retraction_eq !id_comp
retraction_eq !id_id
definition section_id (a : ob) : section_of (ID a) = id :=
section_eq !id_comp
section_eq !id_id
definition id_inverse (a : ob) [H : is_iso (ID a)] : (ID a)⁻¹ = id :=
inverse_eq_left !id_comp
inverse_eq_left !id_id
definition split_mono_comp [instance] [priority 150] (g : b ⟶ c) (f : a ⟶ b)
[Hf : split_mono f] [Hg : split_mono g] : split_mono (g ∘ f) :=
@ -202,7 +202,7 @@ namespace iso
definition iso_of_eq_con (p : a = b) (q : b = c)
: iso_of_eq (p ⬝ q) = iso.trans (iso_of_eq p) (iso_of_eq q) :=
eq.rec_on q (eq.rec_on p (iso_eq !id_comp⁻¹))
eq.rec_on q (eq.rec_on p (iso_eq !id_id⁻¹))
section
open funext

View file

@ -6,8 +6,8 @@ Authors: Floris van Doorn
Limits in a category
-/
import .constructions.cone .groupoid .constructions.discrete .constructions.product
.constructions.finite_cats
import .constructions.cone .constructions.discrete .constructions.product
.constructions.finite_cats .category
open is_trunc functor nat_trans eq
@ -33,6 +33,10 @@ namespace category
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
local attribute is_terminal [reducible]
theorem is_hprop_is_terminal [instance] : is_hprop (is_terminal c) :=
_
omit C
structure has_terminal_object [class] (D : Precategory) :=
@ -42,16 +46,39 @@ namespace category
abbreviation terminal_object [constructor] := @has_terminal_object.d
attribute has_terminal_object.is_term [instance]
variable {D}
definition terminal_object_iso_terminal_object (H₁ H₂ : has_terminal_object D)
: @terminal_object D H₁ ≅ @terminal_object D H₂ :=
terminal_iso_terminal (@has_terminal_object.is_term D H₁) (@has_terminal_object.is_term D H₂)
theorem is_hprop_has_terminal_object [instance] (D : Category)
: is_hprop (has_terminal_object D) :=
begin
apply is_hprop.mk, intro t₁ t₂, induction t₁ with d₁ H₁, induction t₂ with d₂ H₂,
assert p : d₁ = d₂,
{ apply eq_of_iso, apply terminal_iso_terminal H₁ H₂},
induction p, exact ap _ !is_hprop.elim
end
variable (D)
definition has_limits_of_shape [class] := Π(F : I ⇒ D), has_terminal_object (cone F)
/-
The next definitions states that a category is complete with respect to diagrams
in a certain universe. "is_complete.{o₁ h₁ o₂ h₂}" means that D is complete
with respect to diagrams of type Precategory.{o₂ h₂}
-/
definition is_complete.{o₁ h₁ o₂ h₂} [class] (D : Precategory.{o₁ h₁}) :=
Π(I : Precategory.{o₂ h₂}), has_limits_of_shape D I
definition has_limits_of_shape_of_is_complete [instance] [H : is_complete D] (I : Precategory)
: has_limits_of_shape D I := H I
variables {I D}
definition has_terminal_object_of_has_limits_of_shape [instance] [H : has_limits_of_shape D I]
(F : I ⇒ D) : has_terminal_object (cone F) :=
H F
definition has_terminal_object_cone [H : has_limits_of_shape D I]
(F : I ⇒ D) : has_terminal_object (cone F) := H F
local attribute has_terminal_object_cone [instance]
variables (F : I ⇒ D) [H : has_limits_of_shape D I] {i j : I}
include H

View file

@ -65,9 +65,7 @@ namespace category
variables {a b c d : ob} {h : c ⟶ d} {g : hom b c} {f f' : hom a b} {i : a ⟶ a}
include C
definition id [reducible] := ID a
definition id_comp (a : ob) : ID a ∘ ID a = ID a := !id_left
definition id [reducible] [unfold 2] := ID a
definition id_leftright (f : hom a b) : id ∘ f ∘ id = f := !id_left ⬝ !id_right
definition comp_id_eq_id_comp (f : hom a b) : f ∘ id = id ∘ f := !id_right ⬝ !id_left⁻¹

View file

@ -42,7 +42,7 @@ namespace functor
(λd d' g, F (id, g))
(λd, !respect_id)
(λd₁ d₂ d₃ g' g, calc
F (id, g' ∘ g) = F (id ∘ id, g' ∘ g) : by rewrite id_comp
F (id, g' ∘ g) = F (id ∘ id, g' ∘ g) : by rewrite id_id
... = F ((id,g') ∘ (id, g)) : by esimp
... = F (id,g') ∘ F (id, g) : by rewrite respect_comp)
@ -75,7 +75,7 @@ namespace functor
apply @nat_trans_eq,
intro d, calc
natural_map (Fhom F (f' ∘ f)) d = F (f' ∘ f, id) : by rewrite functor_curry_hom_def
... = F (f' ∘ f, id ∘ id) : by rewrite id_comp
... = F (f' ∘ f, id ∘ id) : by rewrite id_id
... = F ((f',id) ∘ (f, id)) : by esimp
... = F (f',id) ∘ F (f, id) : by rewrite [respect_comp F]
... = natural_map ((Fhom F f') ∘ (Fhom F f)) d : by esimp
@ -100,7 +100,7 @@ namespace functor
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
... = id : id_id
theorem functor_uncurry_comp ⦃p p' p'' : C ×c D⦄ (f' : p' ⟶ p'') (f : p ⟶ p')
: Ghom G (f' ∘ f) = Ghom G f' ∘ Ghom G f :=

View file

@ -412,6 +412,16 @@ namespace sigma
definition equiv_subtype [H : Πa, is_hprop (B a)] (u v : {a | B a}) : (u.1 = v.1) ≃ (u = v) :=
equiv.mk !subtype_eq _
definition subtype_eq_inv {A : Type} {B : A → Type} [H : Πa, is_hprop (B a)] (u v : Σa, B a)
: u = v → u.1 = v.1 :=
(subtype_eq u v)⁻¹ᶠ
local attribute subtype_eq_inv [reducible]
definition is_equiv_subtype_eq_inv {A : Type} {B : A → Type} [H : Πa, is_hprop (B a)]
(u v : Σa, B a) : is_equiv (subtype_eq_inv u v) :=
_
/- truncatedness -/
theorem is_trunc_sigma (B : A → Type) (n : trunc_index)
[HA : is_trunc n A] [HB : Πa, is_trunc n (B a)] : is_trunc n (Σa, B a) :=