2015-02-26 18:19:54 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2014 Jakob von Raumer. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
|
|
|
|
|
Module: algebra.precategory.iso
|
|
|
|
|
Authors: Floris van Doorn, Jakob von Raumer
|
|
|
|
|
-/
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2014-12-12 19:19:06 +00:00
|
|
|
|
import .basic .morphism types.sigma
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-02-26 18:19:54 +00:00
|
|
|
|
open eq category sigma sigma.ops equiv is_equiv function is_trunc
|
2014-12-12 04:14:53 +00:00
|
|
|
|
open prod
|
|
|
|
|
|
|
|
|
|
namespace morphism
|
|
|
|
|
variables {ob : Type} [C : precategory ob] include C
|
|
|
|
|
variables {a b c : ob} {g : b ⟶ c} {f : a ⟶ b} {h : b ⟶ a}
|
|
|
|
|
|
|
|
|
|
-- "is_iso f" is equivalent to a certain sigma type
|
2014-12-12 19:19:06 +00:00
|
|
|
|
protected definition sigma_char (f : hom a b) :
|
|
|
|
|
(Σ (g : hom b a), (g ∘ f = id) × (f ∘ g = id)) ≃ is_iso f :=
|
2014-12-12 04:14:53 +00:00
|
|
|
|
begin
|
|
|
|
|
fapply (equiv.mk),
|
2015-02-24 21:27:57 +00:00
|
|
|
|
{intro S, apply is_iso.mk,
|
|
|
|
|
exact (pr₁ S.2),
|
|
|
|
|
exact (pr₂ S.2)},
|
|
|
|
|
{fapply adjointify,
|
|
|
|
|
{intro H, cases H with (g, η, ε),
|
|
|
|
|
exact (sigma.mk g (pair η ε))},
|
|
|
|
|
{intro H, cases H, apply idp},
|
|
|
|
|
{intro S, cases S with (g, ηε), cases ηε, apply idp}},
|
2014-12-12 04:14:53 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
-- The structure for isomorphism can be characterized up to equivalence
|
|
|
|
|
-- by a sigma type.
|
|
|
|
|
definition sigma_is_iso_equiv ⦃a b : ob⦄ : (Σ (f : hom a b), is_iso f) ≃ (a ≅ b) :=
|
|
|
|
|
begin
|
|
|
|
|
fapply (equiv.mk),
|
2015-02-24 21:27:57 +00:00
|
|
|
|
{intro S, apply isomorphic.mk, apply (S.2)},
|
|
|
|
|
{fapply adjointify,
|
|
|
|
|
{intro p, cases p with (f, H), exact (sigma.mk f H)},
|
|
|
|
|
{intro p, cases p, apply idp},
|
|
|
|
|
{intro S, cases S, apply idp}},
|
2014-12-12 04:14:53 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-02-07 01:27:56 +00:00
|
|
|
|
set_option apply.class_instance false -- disable class instance resolution in the apply tactic
|
|
|
|
|
|
2014-12-12 04:14:53 +00:00
|
|
|
|
-- The statement "f is an isomorphism" is a mere proposition
|
|
|
|
|
definition is_hprop_of_is_iso : is_hset (is_iso f) :=
|
|
|
|
|
begin
|
2015-02-21 00:30:32 +00:00
|
|
|
|
apply is_trunc_is_equiv_closed,
|
2014-12-12 04:14:53 +00:00
|
|
|
|
apply (equiv.to_is_equiv (!sigma_char)),
|
2015-02-21 00:30:32 +00:00
|
|
|
|
apply is_trunc_sigma,
|
2014-12-12 04:14:53 +00:00
|
|
|
|
apply (!homH),
|
2015-02-24 21:27:57 +00:00
|
|
|
|
{intro g, apply is_trunc_prod,
|
|
|
|
|
repeat (apply is_trunc_eq; apply is_trunc_succ; apply (!homH))},
|
2014-12-12 04:14:53 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
-- The type of isomorphisms between two objects is a set
|
|
|
|
|
definition is_hset_iso : is_hset (a ≅ b) :=
|
|
|
|
|
begin
|
2015-02-21 00:30:32 +00:00
|
|
|
|
apply is_trunc_is_equiv_closed,
|
2014-12-12 04:14:53 +00:00
|
|
|
|
apply (equiv.to_is_equiv (!sigma_is_iso_equiv)),
|
2015-02-24 21:27:57 +00:00
|
|
|
|
apply is_trunc_sigma,
|
|
|
|
|
apply homH,
|
|
|
|
|
{intro f, apply is_hprop_of_is_iso},
|
2014-12-12 04:14:53 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
-- In a precategory, equal objects are isomorphic
|
2015-02-26 18:19:54 +00:00
|
|
|
|
definition iso_of_path (p : a = b) : a ≅ b :=
|
2014-12-12 19:19:06 +00:00
|
|
|
|
eq.rec_on p (isomorphic.mk id)
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
|
|
|
|
end morphism
|