2015-02-26 18:19:54 +00:00
|
|
|
/-
|
2015-04-25 04:20:59 +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
|
|
|
Authors: Floris van Doorn, Jakob von Raumer
|
|
|
|
|
|
|
|
Category of hsets
|
2015-02-26 18:19:54 +00:00
|
|
|
-/
|
|
|
|
|
2015-09-25 19:02:14 +00:00
|
|
|
import ..category types.equiv ..functor types.lift ..limits
|
2015-02-26 18:19:54 +00:00
|
|
|
|
2015-04-25 04:20:59 +00:00
|
|
|
open eq category equiv iso is_equiv is_trunc function sigma
|
2015-02-26 18:19:54 +00:00
|
|
|
|
|
|
|
namespace category
|
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
definition precategory_hset.{u} [reducible] [constructor] : precategory hset.{u} :=
|
2015-04-25 04:20:59 +00:00
|
|
|
precategory.mk (λx y : hset, x → y)
|
|
|
|
(λx y z g f a, g (f a))
|
|
|
|
(λx a, a)
|
|
|
|
(λx y z w h g f, eq_of_homotopy (λa, idp))
|
|
|
|
(λx y f, eq_of_homotopy (λa, idp))
|
|
|
|
(λx y f, eq_of_homotopy (λa, idp))
|
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
definition Precategory_hset [reducible] [constructor] : Precategory :=
|
2015-04-25 04:20:59 +00:00
|
|
|
Precategory.mk hset precategory_hset
|
|
|
|
|
2015-02-28 06:16:20 +00:00
|
|
|
namespace set
|
2015-03-03 21:38:18 +00:00
|
|
|
local attribute is_equiv_subtype_eq [instance]
|
2015-09-03 04:46:11 +00:00
|
|
|
definition iso_of_equiv [constructor] {A B : Precategory_hset} (f : A ≃ B) : A ≅ B :=
|
2015-03-03 21:38:18 +00:00
|
|
|
iso.MK (to_fun f)
|
2015-09-02 23:41:19 +00:00
|
|
|
(to_inv f)
|
2015-04-27 19:39:36 +00:00
|
|
|
(eq_of_homotopy (left_inv (to_fun f)))
|
|
|
|
(eq_of_homotopy (right_inv (to_fun f)))
|
2015-03-03 21:38:18 +00:00
|
|
|
|
2015-09-03 04:46:11 +00:00
|
|
|
definition equiv_of_iso [constructor] {A B : Precategory_hset} (f : A ≅ B) : A ≃ B :=
|
2015-06-27 00:09:50 +00:00
|
|
|
begin
|
|
|
|
apply equiv.MK (to_hom f) (iso.to_inv f),
|
2015-07-29 12:17:16 +00:00
|
|
|
exact ap10 (to_right_inverse f),
|
|
|
|
exact ap10 (to_left_inverse f)
|
2015-06-27 00:09:50 +00:00
|
|
|
end
|
2015-03-03 21:38:18 +00:00
|
|
|
|
2015-09-03 04:46:11 +00:00
|
|
|
definition is_equiv_iso_of_equiv [constructor] (A B : Precategory_hset)
|
|
|
|
: is_equiv (@iso_of_equiv A B) :=
|
2015-03-03 21:38:18 +00:00
|
|
|
adjointify _ (λf, equiv_of_iso f)
|
2015-05-30 23:44:26 +00:00
|
|
|
(λf, proof iso_eq idp qed)
|
2015-04-27 21:29:56 +00:00
|
|
|
(λf, equiv_eq idp)
|
2015-05-30 23:44:26 +00:00
|
|
|
|
2015-03-03 21:38:18 +00:00
|
|
|
local attribute is_equiv_iso_of_equiv [instance]
|
|
|
|
|
|
|
|
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 :=
|
|
|
|
eq_of_homotopy (λp, eq.rec_on p idp)
|
|
|
|
|
2015-02-28 06:16:20 +00:00
|
|
|
definition equiv_equiv_iso (A B : Precategory_hset) : (A ≃ B) ≃ (A ≅ B) :=
|
2015-03-03 21:38:18 +00:00
|
|
|
equiv.MK (λf, iso_of_equiv f)
|
2015-06-27 00:09:50 +00:00
|
|
|
(λf, proof equiv.MK (to_hom f)
|
2015-02-28 06:16:20 +00:00
|
|
|
(iso.to_inv f)
|
2015-07-29 12:17:16 +00:00
|
|
|
(ap10 (to_right_inverse f))
|
|
|
|
(ap10 (to_left_inverse f)) qed)
|
2015-06-27 00:09:50 +00:00
|
|
|
(λf, proof iso_eq idp qed)
|
|
|
|
(λf, proof equiv_eq idp qed)
|
2015-02-28 06:16:20 +00:00
|
|
|
|
|
|
|
definition equiv_eq_iso (A B : Precategory_hset) : (A ≃ B) = (A ≅ B) :=
|
|
|
|
ua !equiv_equiv_iso
|
|
|
|
|
2015-05-14 02:01:48 +00:00
|
|
|
definition is_univalent_hset (A B : Precategory_hset) : is_equiv (iso_of_eq : A = B → A ≅ B) :=
|
2015-06-10 23:26:32 +00:00
|
|
|
assert H₁ : is_equiv (@iso_of_equiv A B ∘ @equiv_of_eq A B ∘ subtype_eq_inv _ _ ∘
|
|
|
|
@ap _ _ (to_fun (trunctype.sigma_char 0)) A B), from
|
|
|
|
@is_equiv_compose _ _ _ _ _
|
2015-03-03 21:38:18 +00:00
|
|
|
(@is_equiv_compose _ _ _ _ _
|
2015-06-10 23:26:32 +00:00
|
|
|
(@is_equiv_compose _ _ _ _ _
|
|
|
|
_
|
|
|
|
(@is_equiv_subtype_eq_inv _ _ _ _ _))
|
|
|
|
!univalence)
|
|
|
|
!is_equiv_iso_of_equiv,
|
2015-09-25 19:02:14 +00:00
|
|
|
let H₂ := (iso_of_eq_eq_compose A B)⁻¹ in
|
2015-06-10 23:26:32 +00:00
|
|
|
begin
|
|
|
|
rewrite H₂ at H₁,
|
|
|
|
assumption
|
|
|
|
end
|
2015-03-03 21:38:18 +00:00
|
|
|
end set
|
2015-02-26 18:19:54 +00:00
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
definition category_hset [instance] [constructor] : category hset :=
|
2015-04-25 04:20:59 +00:00
|
|
|
category.mk precategory_hset set.is_univalent_hset
|
2015-02-26 18:19:54 +00:00
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
definition Category_hset [reducible] [constructor] : Category :=
|
2015-02-26 18:19:54 +00:00
|
|
|
Category.mk hset category_hset
|
|
|
|
|
2015-09-02 23:41:19 +00:00
|
|
|
abbreviation set [constructor] := Category_hset
|
|
|
|
|
|
|
|
open functor lift
|
|
|
|
definition lift_functor.{u v} [constructor] : set.{u} ⇒ set.{max u v} :=
|
|
|
|
functor.mk tlift
|
|
|
|
(λ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))
|
2015-09-25 19:02:14 +00:00
|
|
|
|
|
|
|
open pi sigma.ops
|
|
|
|
local attribute Category.to.precategory [unfold 1]
|
|
|
|
local attribute category.to_precategory [unfold 2]
|
|
|
|
|
2015-09-25 20:25:46 +00:00
|
|
|
definition is_complete_set_cone.{u v w} [constructor]
|
2015-09-25 19:02:14 +00:00
|
|
|
(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
|
|
|
|
|
2015-09-25 23:44:04 +00:00
|
|
|
definition is_complete_set.{u v w} [instance] : is_complete.{(max u v w)+1 (max u v w) v w} set :=
|
2015-09-25 19:02:14 +00:00
|
|
|
begin
|
|
|
|
intro I F, fapply has_terminal_object.mk,
|
2015-09-25 20:25:46 +00:00
|
|
|
{ exact is_complete_set_cone.{u v w} I F},
|
2015-09-25 19:02:14 +00:00
|
|
|
{ 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
|
2015-09-25 23:44:04 +00:00
|
|
|
|
2015-02-26 18:19:54 +00:00
|
|
|
end category
|