higher groups: prove equivalence of categories for 0-Grp
This commit is contained in:
parent
85b04639cb
commit
9914352e10
3 changed files with 85 additions and 36 deletions
|
@ -336,40 +336,28 @@ end
|
|||
|
||||
local attribute [instance] is_set_Grp_hom
|
||||
|
||||
definition Grp_precategory [constructor] (k : ℕ) : precategory [0;k]Grp :=
|
||||
precategory.mk (λG H, Grp_hom G H) (λX Y Z g f, g ∘* f) (λX, pid (B X))
|
||||
begin intros, apply eq_of_phomotopy, exact !passoc⁻¹* end
|
||||
begin intros, apply eq_of_phomotopy, apply pid_pcompose end
|
||||
begin intros, apply eq_of_phomotopy, apply pcompose_pid end
|
||||
|
||||
definition cGrp [constructor] (k : ℕ) : Precategory :=
|
||||
Precategory.mk _ (Grp_precategory k)
|
||||
pb_Precategory (cptruncconntype' (k.-1))
|
||||
(Grp_equiv 0 k ⬝e equiv_ap (λx, x-Type*[k.-1]) (ap of_nat (zero_add k)) ⬝e
|
||||
(ptruncconntype'_equiv_ptruncconntype (k.-1))⁻¹ᵉ)
|
||||
|
||||
example (k : ℕ) : Precategory.carrier (cGrp k) = [0;k]Grp := by reflexivity
|
||||
|
||||
definition cGrp_equivalence_cType [constructor] (k : ℕ) : cGrp k ≃c cType*[k.-1] :=
|
||||
sorry
|
||||
!pb_Precategory_equivalence_of_equiv
|
||||
|
||||
-- set_option pp.all true
|
||||
definition cGrp_equivalence_Grp [constructor] : cGrp 1 ≃c category.Grp :=
|
||||
sorry
|
||||
|
||||
-- equivalence.trans
|
||||
-- _
|
||||
-- (equivalence.symm Grp_equivalence_cptruncconntype')
|
||||
-- begin
|
||||
-- transitivity cptruncconntype'.{u} 0,
|
||||
-- exact sorry,
|
||||
-- symmetry, exact Grp_equivalence_cptruncconntype'
|
||||
-- end
|
||||
-- category.equivalence.{u+1 u u+1 u} (category.Category.to_Precategory.{u+1 u} category.Grp.{u})
|
||||
-- (EM.cptruncconntype'.{u} (@zero.{0} trunc_index has_zero_trunc_index))
|
||||
-- equivalence.trans
|
||||
-- _
|
||||
-- (equivalence.symm Grp_equivalence_cptruncconntype')
|
||||
definition cGrp_equivalence_Grp [constructor] : cGrp.{u} 1 ≃c category.Grp.{u} :=
|
||||
equivalence.trans !pb_Precategory_equivalence_of_equiv
|
||||
(equivalence.trans (equivalence.symm Grp_equivalence_cptruncconntype')
|
||||
proof equivalence.refl _ qed)
|
||||
|
||||
definition cGrp_equivalence_AbGrp [constructor] (k : ℕ) : cGrp.{u} (k+2) ≃c category.AbGrp.{u} :=
|
||||
equivalence.trans !pb_Precategory_equivalence_of_equiv
|
||||
(equivalence.trans (equivalence.symm (AbGrp_equivalence_cptruncconntype' k))
|
||||
proof equivalence.refl _ qed)
|
||||
|
||||
--has sorry
|
||||
print axioms ωstabilization
|
||||
print axioms cGrp_equivalence_Grp
|
||||
|
||||
-- no sorry's
|
||||
print axioms Decat_adjoint_Disc
|
||||
|
@ -380,5 +368,7 @@ print axioms Stabilize_adjoint_Forget
|
|||
print axioms Stabilize_adjoint_Forget_natural
|
||||
print axioms stabilization
|
||||
print axioms is_trunc_Grp
|
||||
print axioms cGrp_equivalence_Grp
|
||||
print axioms cGrp_equivalence_AbGrp
|
||||
|
||||
end higher_group
|
||||
|
|
|
@ -4,7 +4,7 @@ import homotopy.EM algebra.category.functor.equivalence types.pointed2 ..pointed
|
|||
..move_to_lib .susp ..algebra.exactness ..univalent_subcategory
|
||||
|
||||
open eq equiv is_equiv algebra group nat pointed EM.ops is_trunc trunc susp function is_conn nat
|
||||
|
||||
universe variable u
|
||||
/- TODO: try to fix the compilation time of this file -/
|
||||
|
||||
namespace EM
|
||||
|
@ -315,7 +315,7 @@ namespace EM
|
|||
| 1 := EM1 (π₁ X)
|
||||
| (n+2) := EMadd1 (πag[n+2] X) (n+1)
|
||||
|
||||
definition EM_type_pequiv.{u} {X Y : pType.{u}} (n : ℕ) [Hn : is_succ n] (e : πg[n] Y ≃g πg[n] X)
|
||||
definition EM_type_pequiv {X Y : pType.{u}} (n : ℕ) [Hn : is_succ n] (e : πg[n] Y ≃g πg[n] X)
|
||||
[H1 : is_conn (n.-1) X] [H2 : is_trunc n X] : EM_type Y n ≃* X :=
|
||||
begin
|
||||
induction Hn with n, cases n with n,
|
||||
|
@ -399,6 +399,16 @@ namespace EM
|
|||
attribute ptruncconntype'.A [coercion]
|
||||
attribute ptruncconntype'.H1 ptruncconntype'.H2 [instance]
|
||||
|
||||
definition ptruncconntype'_equiv_ptruncconntype (n : ℕ₋₂) :
|
||||
(ptruncconntype' n : Type.{u+1}) ≃ ((n+1)-Type*[n] : Type.{u+1}) :=
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
{ intro X, exact ptruncconntype.mk (ptruncconntype'.A X) _ pt _ },
|
||||
{ intro X, exact ptruncconntype'.mk X _ _ },
|
||||
{ intro X, induction X with X H1 x₀ H2, reflexivity },
|
||||
{ intro X, induction X with X H1 H2, induction X with X x₀, reflexivity }
|
||||
end
|
||||
|
||||
definition EM1_pequiv_ptruncconntype' (X : ptruncconntype' 0) : EM1 (πg[1] X) ≃* X :=
|
||||
@(EM1_pequiv_type X) _ (ptruncconntype'.H2 X)
|
||||
|
||||
|
@ -416,7 +426,7 @@ namespace EM
|
|||
|
||||
open category functor nat_trans
|
||||
|
||||
definition precategory_ptruncconntype'.{u} [constructor] (n : ℕ₋₂) :
|
||||
definition precategory_ptruncconntype' [constructor] (n : ℕ₋₂) :
|
||||
precategory.{u+1 u} (ptruncconntype' n) :=
|
||||
begin
|
||||
fapply precategory.mk,
|
||||
|
@ -468,7 +478,7 @@ namespace EM
|
|||
begin intro, apply homomorphism_eq, exact to_homotopy !homotopy_group_functor_pid end
|
||||
begin intros, apply homomorphism_eq, exact to_homotopy !homotopy_group_functor_compose end
|
||||
|
||||
definition is_equivalence_EM1_cfunctor.{u} : is_equivalence EM1_cfunctor.{u} :=
|
||||
definition is_equivalence_EM1_cfunctor : is_equivalence EM1_cfunctor.{u} :=
|
||||
begin
|
||||
fapply is_equivalence.mk,
|
||||
{ exact homotopy_group_cfunctor.{u} },
|
||||
|
@ -492,7 +502,7 @@ namespace EM
|
|||
{ apply eq_of_phomotopy, apply pright_inv }}}
|
||||
end
|
||||
|
||||
definition is_equivalence_EM_cfunctor.{u} (n : ℕ) : is_equivalence (EM_cfunctor.{u} (n+2)) :=
|
||||
definition is_equivalence_EM_cfunctor (n : ℕ) : is_equivalence (EM_cfunctor.{u} (n+2)) :=
|
||||
begin
|
||||
fapply is_equivalence.mk,
|
||||
{ exact ab_homotopy_group_cfunctor.{u} n },
|
||||
|
@ -516,10 +526,10 @@ namespace EM
|
|||
{ apply eq_of_phomotopy, apply pright_inv }}}
|
||||
end
|
||||
|
||||
definition Grp_equivalence_cptruncconntype'.{u} [constructor] : Grp.{u} ≃c cType*[0] :=
|
||||
definition Grp_equivalence_cptruncconntype' [constructor] : Grp.{u} ≃c cType*[0] :=
|
||||
equivalence.mk EM1_cfunctor.{u} is_equivalence_EM1_cfunctor.{u}
|
||||
|
||||
definition AbGrp_equivalence_cptruncconntype'.{u} [constructor] (n : ℕ) : AbGrp.{u} ≃c cType*[n.+1] :=
|
||||
definition AbGrp_equivalence_cptruncconntype' [constructor] (n : ℕ) : AbGrp.{u} ≃c cType*[n.+1] :=
|
||||
equivalence.mk (EM_cfunctor.{u} (n+2)) (is_equivalence_EM_cfunctor.{u} n)
|
||||
end category
|
||||
|
||||
|
@ -623,7 +633,7 @@ namespace EM
|
|||
open chain_complex prod fin
|
||||
|
||||
/- TODO: other cases -/
|
||||
definition LES_isomorphism_kernel_of_trivial.{u}
|
||||
definition LES_isomorphism_kernel_of_trivial
|
||||
{X Y : pType.{u}} (f : X →* Y) (n : ℕ) [H : is_succ n]
|
||||
(H1 : is_contr (πg[n+1] Y)) : πg[n] (pfiber f) ≃g Kernel (π→g[n] f) :=
|
||||
begin
|
||||
|
@ -639,7 +649,7 @@ namespace EM
|
|||
end
|
||||
|
||||
open group algebra is_trunc
|
||||
definition homotopy_group_fiber_EM1_functor.{u} {G H : Group.{u}} (φ : G →g H) :
|
||||
definition homotopy_group_fiber_EM1_functor {G H : Group.{u}} (φ : G →g H) :
|
||||
π₁ (pfiber (EM1_functor φ)) ≃g Kernel φ :=
|
||||
have H1 : is_trunc 1 (EM1 H), from sorry,
|
||||
have H2 : 1 <[ℕ] 1 + 1, from sorry,
|
||||
|
|
|
@ -1,6 +1,6 @@
|
|||
-- definitions, theorems and attributes which should be moved to files in the HoTT library
|
||||
|
||||
import homotopy.sphere2 homotopy.cofiber homotopy.wedge hit.prop_trunc hit.set_quotient eq2 types.pointed2 algebra.graph
|
||||
import homotopy.sphere2 homotopy.cofiber homotopy.wedge hit.prop_trunc hit.set_quotient eq2 types.pointed2 algebra.graph algebra.category.functor.equivalence
|
||||
|
||||
open eq nat int susp pointed sigma is_equiv equiv fiber algebra trunc pi group
|
||||
is_trunc function unit prod bool
|
||||
|
@ -1890,3 +1890,52 @@ begin
|
|||
{ apply decidable.inr, intro p, cases p, apply np, reflexivity }},
|
||||
end
|
||||
end decidable
|
||||
|
||||
namespace category
|
||||
open functor
|
||||
/- shortening pullback to pb to keep names relatively short -/
|
||||
definition pb_precategory [constructor] {A B : Type} (f : A → B) (C : precategory B) :
|
||||
precategory A :=
|
||||
precategory.mk (λa a', hom (f a) (f a')) (λa a' a'' h g, h ∘ g) (λa, ID (f a))
|
||||
(λa a' a'' a''' k h g, assoc k h g) (λa a' g, id_left g) (λa a' g, id_right g)
|
||||
|
||||
definition pb_Precategory [constructor] {A : Type} (C : Precategory) (f : A → C) :
|
||||
Precategory :=
|
||||
Precategory.mk A (pb_precategory f C)
|
||||
|
||||
definition pb_Precategory_functor [constructor] {A : Type} (C : Precategory) (f : A → C) :
|
||||
pb_Precategory C f ⇒ C :=
|
||||
functor.mk f (λa a' g, g) proof (λa, idp) qed proof (λa a' a'' h g, idp) qed
|
||||
|
||||
definition fully_faithful_pb_Precategory_functor {A : Type} (C : Precategory)
|
||||
(f : A → C) : fully_faithful (pb_Precategory_functor C f) :=
|
||||
begin intro a a', apply is_equiv_id end
|
||||
|
||||
definition split_essentially_surjective_pb_Precategory_functor {A : Type} (C : Precategory)
|
||||
(f : A → C) (H : is_split_surjective f) :
|
||||
split_essentially_surjective (pb_Precategory_functor C f) :=
|
||||
begin intro c, cases H c with a p, exact ⟨a, iso.iso_of_eq p⟩ end
|
||||
|
||||
definition is_equivalence_pb_Precategory_functor {A : Type} (C : Precategory)
|
||||
(f : A → C) (H : is_split_surjective f) : is_equivalence (pb_Precategory_functor C f) :=
|
||||
@(is_equivalence_of_fully_faithful_of_split_essentially_surjective _)
|
||||
(fully_faithful_pb_Precategory_functor C f)
|
||||
(split_essentially_surjective_pb_Precategory_functor C f H)
|
||||
|
||||
definition pb_Precategory_equivalence [constructor] {A : Type} (C : Precategory) (f : A → C)
|
||||
(H : is_split_surjective f) : pb_Precategory C f ≃c C :=
|
||||
equivalence.mk _ (is_equivalence_pb_Precategory_functor C f H)
|
||||
|
||||
definition pb_Precategory_equivalence_of_equiv [constructor] {A : Type} (C : Precategory)
|
||||
(f : A ≃ C) : pb_Precategory C f ≃c C :=
|
||||
pb_Precategory_equivalence C f (is_split_surjective_of_is_retraction f)
|
||||
|
||||
definition is_isomorphism_pb_Precategory_functor [constructor] {A : Type} (C : Precategory)
|
||||
(f : A ≃ C) : is_isomorphism (pb_Precategory_functor C f) :=
|
||||
(fully_faithful_pb_Precategory_functor C f, to_is_equiv f)
|
||||
|
||||
definition pb_Precategory_isomorphism [constructor] {A : Type} (C : Precategory) (f : A ≃ C) :
|
||||
pb_Precategory C f ≅c C :=
|
||||
isomorphism.mk _ (is_isomorphism_pb_Precategory_functor C f)
|
||||
|
||||
end category
|
||||
|
|
Loading…
Reference in a new issue