2015-03-13 22:28:19 +00:00
|
|
|
|
/-
|
2015-03-17 00:08:45 +00:00
|
|
|
|
Copyright (c) 2015 Floris van Doorn. All rights reserved.
|
2015-03-13 22:28:19 +00:00
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Authors: Floris van Doorn
|
|
|
|
|
-/
|
|
|
|
|
|
2015-08-31 16:23:34 +00:00
|
|
|
|
import algebra.category.constructions function arity
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
2015-08-31 16:23:34 +00:00
|
|
|
|
open category functor nat_trans eq is_trunc iso equiv prod trunc function pi is_equiv
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
2015-04-29 00:48:39 +00:00
|
|
|
|
namespace category
|
2015-08-31 16:23:34 +00:00
|
|
|
|
variables {C D : Precategory} {F : C ⇒ D} {G : D ⇒ C}
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
2015-09-01 22:00:11 +00:00
|
|
|
|
-- TODO: define a structure "adjoint" and then define
|
2015-04-29 00:48:39 +00:00
|
|
|
|
-- structure is_left_adjoint (F : C ⇒ D) :=
|
2015-09-01 22:00:11 +00:00
|
|
|
|
-- (G : D ⇒ C) -- G
|
|
|
|
|
-- (is_adjoint : adjoint F G)
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
2015-04-29 00:48:39 +00:00
|
|
|
|
structure is_left_adjoint [class] (F : C ⇒ D) :=
|
|
|
|
|
(G : D ⇒ C)
|
2015-08-31 16:23:34 +00:00
|
|
|
|
(η : 1 ⟹ G ∘f F)
|
|
|
|
|
(ε : F ∘f G ⟹ 1)
|
2015-04-29 00:48:39 +00:00
|
|
|
|
(H : Π(c : C), (ε (F c)) ∘ (F (η c)) = ID (F c))
|
|
|
|
|
(K : Π(d : D), (G (ε d)) ∘ (η (G d)) = ID (G d))
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
2015-04-29 00:48:39 +00:00
|
|
|
|
abbreviation right_adjoint := @is_left_adjoint.G
|
|
|
|
|
abbreviation unit := @is_left_adjoint.η
|
|
|
|
|
abbreviation counit := @is_left_adjoint.ε
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
2015-04-29 00:48:39 +00:00
|
|
|
|
structure is_equivalence [class] (F : C ⇒ D) extends is_left_adjoint F :=
|
|
|
|
|
mk' ::
|
|
|
|
|
(is_iso_unit : is_iso η)
|
|
|
|
|
(is_iso_counit : is_iso ε)
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
2015-08-31 16:23:34 +00:00
|
|
|
|
abbreviation inverse := @is_equivalence.G
|
|
|
|
|
postfix `⁻¹` := inverse
|
|
|
|
|
--a second notation for the inverse, which is not overloaded
|
|
|
|
|
postfix [parsing-only] `⁻¹F`:std.prec.max_plus := inverse
|
|
|
|
|
|
2015-04-29 00:48:39 +00:00
|
|
|
|
--TODO: review and change
|
2015-08-31 16:23:34 +00:00
|
|
|
|
definition faithful [class] (F : C ⇒ D) := Π⦃c c' : C⦄ ⦃f f' : c ⟶ c'⦄, F f = F f' → f = f'
|
|
|
|
|
definition full [class] (F : C ⇒ D) := Π⦃c c' : C⦄, is_surjective (@(to_fun_hom F) c c')
|
|
|
|
|
definition fully_faithful [class] (F : C ⇒ D) := Π(c c' : C), is_equiv (@(to_fun_hom F) c c')
|
|
|
|
|
definition split_essentially_surjective [class] (F : C ⇒ D) := Π(d : D), Σ(c : C), F c ≅ d
|
|
|
|
|
definition essentially_surjective [class] (F : C ⇒ D) := Π(d : D), ∃(c : C), F c ≅ d
|
|
|
|
|
definition is_weak_equivalence [class] (F : C ⇒ D) := fully_faithful F × essentially_surjective F
|
|
|
|
|
definition is_isomorphism [class] (F : C ⇒ D) := fully_faithful F × is_equiv (to_fun_ob F)
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
2015-09-01 22:00:11 +00:00
|
|
|
|
structure equivalence (C D : Precategory) :=
|
|
|
|
|
(to_functor : C ⇒ D)
|
|
|
|
|
(struct : is_equivalence to_functor)
|
|
|
|
|
|
2015-04-29 00:48:39 +00:00
|
|
|
|
structure isomorphism (C D : Precategory) :=
|
|
|
|
|
(to_functor : C ⇒ D)
|
|
|
|
|
(struct : is_isomorphism to_functor)
|
2015-03-17 00:08:45 +00:00
|
|
|
|
-- infix `⊣`:55 := adjoint
|
|
|
|
|
|
2015-04-29 00:48:39 +00:00
|
|
|
|
infix `⋍`:25 := equivalence -- \backsimeq or \equiv
|
|
|
|
|
infix `≌`:25 := isomorphism -- \backcong or \iso
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
2015-08-31 16:23:34 +00:00
|
|
|
|
definition is_equiv_of_fully_faithful [instance] (F : C ⇒ D) [H : fully_faithful F] (c c' : C)
|
|
|
|
|
: is_equiv (@(to_fun_hom F) c c') :=
|
|
|
|
|
!H
|
|
|
|
|
|
2015-09-03 04:46:11 +00:00
|
|
|
|
definition hom_equiv_F_hom_F [constructor] (F : C ⇒ D)
|
|
|
|
|
[H : fully_faithful F] (c c' : C) : (c ⟶ c') ≃ (F c ⟶ F c') :=
|
|
|
|
|
equiv.mk _ !H
|
|
|
|
|
|
|
|
|
|
definition iso_of_F_iso_F (F : C ⇒ D)
|
|
|
|
|
[H : fully_faithful F] (c c' : C) (g : F c ≅ F c') : c ≅ c' :=
|
|
|
|
|
begin
|
|
|
|
|
induction g with g G, induction G with h p q, fapply iso.MK,
|
|
|
|
|
{ rexact (@(to_fun_hom F) c c')⁻¹ᶠ g},
|
|
|
|
|
{ rexact (@(to_fun_hom F) c' c)⁻¹ᶠ h},
|
|
|
|
|
{ exact abstract begin
|
|
|
|
|
apply eq_of_fn_eq_fn' (@(to_fun_hom F) c c),
|
|
|
|
|
rewrite [respect_comp, respect_id,
|
|
|
|
|
right_inv (@(to_fun_hom F) c c'), right_inv (@(to_fun_hom F) c' c), p],
|
|
|
|
|
end end},
|
|
|
|
|
{ exact abstract begin
|
|
|
|
|
apply eq_of_fn_eq_fn' (@(to_fun_hom F) c' c'),
|
|
|
|
|
rewrite [respect_comp, respect_id,
|
|
|
|
|
right_inv (@(to_fun_hom F) c c'), right_inv (@(to_fun_hom F) c' c), q],
|
|
|
|
|
end end}
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition iso_equiv_F_iso_F [constructor] (F : C ⇒ D)
|
|
|
|
|
[H : fully_faithful F] (c c' : C) : (c ≅ c') ≃ (F c ≅ F c') :=
|
|
|
|
|
begin
|
|
|
|
|
fapply equiv.MK,
|
|
|
|
|
{ exact preserve_iso F},
|
|
|
|
|
{ apply iso_of_F_iso_F},
|
|
|
|
|
{ exact abstract begin
|
|
|
|
|
intro f, induction f with f F', induction F' with g p q, apply iso_eq,
|
|
|
|
|
esimp [iso_of_F_iso_F], apply right_inv end end},
|
|
|
|
|
{ exact abstract begin
|
|
|
|
|
intro f, induction f with f F', induction F' with g p q, apply iso_eq,
|
|
|
|
|
esimp [iso_of_F_iso_F], apply right_inv end end},
|
|
|
|
|
end
|
|
|
|
|
|
2015-09-01 22:00:11 +00:00
|
|
|
|
definition is_iso_unit [instance] (F : C ⇒ D) [H : is_equivalence F] : is_iso (unit F) :=
|
2015-08-31 16:23:34 +00:00
|
|
|
|
!is_equivalence.is_iso_unit
|
|
|
|
|
|
2015-09-01 22:00:11 +00:00
|
|
|
|
definition is_iso_counit [instance] (F : C ⇒ D) [H : is_equivalence F] : is_iso (counit F) :=
|
2015-08-31 16:23:34 +00:00
|
|
|
|
!is_equivalence.is_iso_counit
|
|
|
|
|
|
2015-09-01 23:57:49 +00:00
|
|
|
|
theorem is_hprop_is_left_adjoint {C : Category} {D : Precategory} (F : C ⇒ D)
|
|
|
|
|
: is_hprop (is_left_adjoint F) :=
|
|
|
|
|
begin
|
|
|
|
|
apply is_hprop.mk,
|
|
|
|
|
intro G G', cases G with G η ε H K, cases G' with G' η' ε' H' K',
|
|
|
|
|
assert lem₁ : Π(p : G = G'), p ▸ η = η' → p ▸ ε = ε'
|
|
|
|
|
→ is_left_adjoint.mk G η ε H K = is_left_adjoint.mk G' η' ε' H' K',
|
|
|
|
|
{ intros p q r, induction p, induction q, induction r, esimp,
|
|
|
|
|
apply apd011 (is_left_adjoint.mk G η ε) !is_hprop.elim !is_hprop.elim},
|
|
|
|
|
assert lem₂ : Π (d : carrier D),
|
|
|
|
|
(to_fun_hom G (natural_map ε' d) ∘
|
|
|
|
|
natural_map η (to_fun_ob G' d)) ∘
|
|
|
|
|
to_fun_hom G' (natural_map ε d) ∘
|
|
|
|
|
natural_map η' (to_fun_ob G d) = id,
|
|
|
|
|
{ intro d, esimp,
|
2015-09-03 04:46:11 +00:00
|
|
|
|
rewrite [category.assoc],
|
2015-09-01 23:57:49 +00:00
|
|
|
|
rewrite [-assoc (G (ε' d))],
|
|
|
|
|
esimp, rewrite [nf_fn_eq_fn_nf_pt' G' ε η d],
|
2015-09-03 04:46:11 +00:00
|
|
|
|
esimp, rewrite [category.assoc],
|
|
|
|
|
esimp, rewrite [-category.assoc],
|
2015-09-01 23:57:49 +00:00
|
|
|
|
rewrite [↑functor.compose, -respect_comp G],
|
|
|
|
|
rewrite [nf_fn_eq_fn_nf_pt ε ε' d,nf_fn_eq_fn_nf_pt η' η (G d),▸*],
|
|
|
|
|
rewrite [respect_comp G],
|
2015-09-03 04:46:11 +00:00
|
|
|
|
rewrite [category.assoc,▸*,-category.assoc (G (ε d))],
|
2015-09-01 23:57:49 +00:00
|
|
|
|
rewrite [↑functor.compose, -respect_comp G],
|
|
|
|
|
rewrite [H' (G d)],
|
2015-09-03 04:46:11 +00:00
|
|
|
|
rewrite [respect_id,▸*,category.id_right],
|
2015-09-01 23:57:49 +00:00
|
|
|
|
apply K},
|
|
|
|
|
assert lem₃ : Π (d : carrier D),
|
|
|
|
|
(to_fun_hom G' (natural_map ε d) ∘
|
|
|
|
|
natural_map η' (to_fun_ob G d)) ∘
|
|
|
|
|
to_fun_hom G (natural_map ε' d) ∘
|
|
|
|
|
natural_map η (to_fun_ob G' d) = id,
|
|
|
|
|
{ intro d, esimp,
|
2015-09-03 04:46:11 +00:00
|
|
|
|
rewrite [category.assoc, -assoc (G' (ε d))],
|
2015-09-01 23:57:49 +00:00
|
|
|
|
esimp, rewrite [nf_fn_eq_fn_nf_pt' G ε' η' d],
|
2015-09-03 04:46:11 +00:00
|
|
|
|
esimp, rewrite [category.assoc], esimp, rewrite [-category.assoc],
|
2015-09-01 23:57:49 +00:00
|
|
|
|
rewrite [↑functor.compose, -respect_comp G'],
|
|
|
|
|
rewrite [nf_fn_eq_fn_nf_pt ε' ε d,nf_fn_eq_fn_nf_pt η η' (G' d)],
|
|
|
|
|
esimp,
|
|
|
|
|
rewrite [respect_comp G'],
|
2015-09-03 04:46:11 +00:00
|
|
|
|
rewrite [category.assoc,▸*,-category.assoc (G' (ε' d))],
|
2015-09-01 23:57:49 +00:00
|
|
|
|
rewrite [↑functor.compose, -respect_comp G'],
|
|
|
|
|
rewrite [H (G' d)],
|
2015-09-03 04:46:11 +00:00
|
|
|
|
rewrite [respect_id,▸*,category.id_right],
|
2015-09-01 23:57:49 +00:00
|
|
|
|
apply K'},
|
|
|
|
|
fapply lem₁,
|
|
|
|
|
{ fapply functor.eq_of_pointwise_iso,
|
|
|
|
|
{ fapply change_natural_map,
|
|
|
|
|
{ exact (G' ∘fn1 ε) ∘n !assoc_natural_rev ∘n (η' ∘1nf G)},
|
|
|
|
|
{ intro d, exact (G' (ε d) ∘ η' (G d))},
|
|
|
|
|
{ intro d, exact ap (λx, _ ∘ x) !id_left}},
|
|
|
|
|
{ intro d, fconstructor,
|
|
|
|
|
{ exact (G (ε' d) ∘ η (G' d))},
|
|
|
|
|
{ exact lem₂ d },
|
|
|
|
|
{ exact lem₃ d }}},
|
|
|
|
|
{ clear lem₁, refine transport_hom_of_eq_right _ η ⬝ _,
|
|
|
|
|
krewrite hom_of_eq_compose_right,
|
|
|
|
|
rewrite functor.hom_of_eq_eq_of_pointwise_iso,
|
|
|
|
|
apply nat_trans_eq, intro c, esimp,
|
|
|
|
|
refine !assoc⁻¹ ⬝ ap (λx, _ ∘ x) (nf_fn_eq_fn_nf_pt η η' c) ⬝ !assoc ⬝ _,
|
2015-09-03 04:46:11 +00:00
|
|
|
|
esimp, rewrite [-respect_comp G',H c,respect_id G',▸*,category.id_left]},
|
2015-09-01 23:57:49 +00:00
|
|
|
|
{ clear lem₁, refine transport_hom_of_eq_left _ ε ⬝ _,
|
|
|
|
|
krewrite inv_of_eq_compose_left,
|
|
|
|
|
rewrite functor.inv_of_eq_eq_of_pointwise_iso,
|
|
|
|
|
apply nat_trans_eq, intro d, esimp,
|
2015-09-03 04:46:11 +00:00
|
|
|
|
krewrite [respect_comp],
|
|
|
|
|
rewrite [assoc,nf_fn_eq_fn_nf_pt ε' ε d,-assoc,▸*,H (G' d),id_right]}
|
2015-09-01 23:57:49 +00:00
|
|
|
|
end
|
2015-08-31 16:23:34 +00:00
|
|
|
|
|
2015-03-13 22:28:19 +00:00
|
|
|
|
definition full_of_fully_faithful (H : fully_faithful F) : full F :=
|
2015-08-31 16:23:34 +00:00
|
|
|
|
λc c', is_surjective.mk (λg, tr (fiber.mk ((@(to_fun_hom F) c c')⁻¹ᶠ g) !right_inv))
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
|
|
|
|
definition faithful_of_fully_faithful (H : fully_faithful F) : faithful F :=
|
2015-04-29 00:48:39 +00:00
|
|
|
|
λc c' f f' p, is_injective_of_is_embedding p
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
|
|
|
|
definition fully_faithful_of_full_of_faithful (H : faithful F) (K : full F) : fully_faithful F :=
|
2015-08-31 16:23:34 +00:00
|
|
|
|
begin
|
|
|
|
|
intro c c',
|
|
|
|
|
apply is_equiv_of_is_surjective_of_is_embedding,
|
|
|
|
|
{ apply is_embedding_of_is_injective,
|
|
|
|
|
intros f f' p, exact H p},
|
|
|
|
|
{ apply K}
|
|
|
|
|
end
|
|
|
|
|
|
2015-09-01 22:00:11 +00:00
|
|
|
|
definition split_essentially_surjective_of_is_equivalence (F : C ⇒ D) [H : is_equivalence F]
|
|
|
|
|
: split_essentially_surjective F :=
|
|
|
|
|
begin
|
|
|
|
|
intro d, fconstructor,
|
|
|
|
|
{ exact F⁻¹ d},
|
|
|
|
|
{ exact componentwise_iso (@(iso.mk (counit F)) !is_iso_counit) d}
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
/-
|
|
|
|
|
|
2015-08-31 16:23:34 +00:00
|
|
|
|
definition fully_faithful_of_is_equivalence (F : C ⇒ D) [H : is_equivalence F]
|
|
|
|
|
: fully_faithful F :=
|
|
|
|
|
begin
|
|
|
|
|
intro c c',
|
|
|
|
|
fapply adjointify,
|
|
|
|
|
{ intro g, exact natural_map (@(iso.inverse (unit F)) !is_iso_unit) c' ∘ F⁻¹ g ∘ unit F c},
|
2015-09-01 22:00:11 +00:00
|
|
|
|
{ intro g, rewrite [+respect_comp,▸*],
|
|
|
|
|
krewrite [natural_map_inverse], xrewrite [respect_inv'],
|
|
|
|
|
apply inverse_comp_eq_of_eq_comp,
|
|
|
|
|
exact sorry /-this is basically the naturality of the counit-/ },
|
2015-08-31 16:23:34 +00:00
|
|
|
|
{ exact sorry},
|
|
|
|
|
end
|
|
|
|
|
|
2015-09-01 22:00:11 +00:00
|
|
|
|
section
|
|
|
|
|
variables (F G)
|
|
|
|
|
variables (η : G ∘f F ≅ 1) (ε : F ∘f G ≅ 1)
|
|
|
|
|
include η ε
|
|
|
|
|
--definition inverse_of_unit_counit
|
2015-08-31 16:23:34 +00:00
|
|
|
|
|
2015-09-01 22:00:11 +00:00
|
|
|
|
definition is_equivalence.mk : is_equivalence F :=
|
|
|
|
|
begin
|
|
|
|
|
exact sorry
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
end
|
2015-03-13 22:28:19 +00:00
|
|
|
|
|
|
|
|
|
definition fully_faithful_equiv (F : C ⇒ D) : fully_faithful F ≃ (faithful F × full F) :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_equivalence_equiv (F : C ⇒ D)
|
|
|
|
|
: is_equivalence F ≃ (fully_faithful F × split_essentially_surjective F) :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_hprop_is_weak_equivalence (F : C ⇒ D) : is_hprop (is_weak_equivalence F) :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_hprop_is_equivalence {C D : Category} (F : C ⇒ D) : is_hprop (is_equivalence F) :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_equivalence_equiv_is_weak_equivalence {C D : Category} (F : C ⇒ D)
|
|
|
|
|
: is_equivalence F ≃ is_weak_equivalence F :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_hprop_is_isomorphism (F : C ⇒ D) : is_hprop (is_isomorphism F) :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_isomorphism_equiv1 (F : C ⇒ D) : is_equivalence F
|
2015-08-31 16:23:34 +00:00
|
|
|
|
≃ Σ(G : D ⇒ C) (η : 1 = G ∘f F) (ε : F ∘f G = 1),
|
2015-05-01 03:23:12 +00:00
|
|
|
|
sorry ▸ ap (λ(H : C ⇒ C), F ∘f H) η = ap (λ(H : D ⇒ D), H ∘f F) ε⁻¹ :=
|
2015-03-13 22:28:19 +00:00
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_isomorphism_equiv2 (F : C ⇒ D) : is_equivalence F
|
2015-08-31 16:23:34 +00:00
|
|
|
|
≃ ∃(G : D ⇒ C), 1 = G ∘f F × F ∘f G = 1 :=
|
2015-03-13 22:28:19 +00:00
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_equivalence_of_isomorphism (H : is_isomorphism F) : is_equivalence F :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_isomorphism_of_is_equivalence {C D : Category} {F : C ⇒ D} (H : is_equivalence F)
|
|
|
|
|
: is_isomorphism F :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition isomorphism_of_eq {C D : Precategory} (p : C = D) : C ≌ D :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_equiv_isomorphism_of_eq (C D : Precategory) : is_equiv (@isomorphism_of_eq C D) :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition equivalence_of_eq {C D : Precategory} (p : C = D) : C ⋍ D :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_equiv_equivalence_of_eq (C D : Category) : is_equiv (@equivalence_of_eq C D) :=
|
|
|
|
|
sorry
|
2015-07-27 15:34:11 +00:00
|
|
|
|
-/
|
2015-03-13 22:28:19 +00:00
|
|
|
|
end category
|