move more and update after changes

This commit is contained in:
Floris van Doorn 2018-09-11 17:06:46 +02:00
parent f7e0ce0e20
commit 68345f75ce
25 changed files with 316 additions and 1486 deletions

View file

@ -19,35 +19,37 @@ namespace group
definition ppi_inv [constructor] {A : Type*} {B : A → Type*} (f : Π*a, Ω (B a)) : Π*a, Ω (B a) :=
proof ppi.mk (λa, (f a)⁻¹ᵖ) (respect_pt f)⁻² qed
definition inf_group_ppi [constructor] [instance] {A : Type*} (B : A → Type*) :
inf_group (Π*a, Ω (B a)) :=
definition inf_pgroup_pppi [constructor] {A : Type*} (B : A → Type*) :
inf_pgroup (Π*a, Ω (B a)) :=
begin
fapply inf_group.mk,
fapply inf_pgroup.mk,
{ exact ppi_mul },
{ intro f g h, apply eq_of_phomotopy, fapply phomotopy.mk,
{ intro a, exact con.assoc (f a) (g a) (h a) },
{ symmetry, rexact eq_of_square (con2_assoc (respect_pt f) (respect_pt g) (respect_pt h)) }},
{ apply ppi_const },
{ exact ppi_inv },
{ intros f, apply eq_of_phomotopy, fapply phomotopy.mk,
{ intro a, exact one_mul (f a) },
{ symmetry, apply eq_of_square, refine _ ⬝vp !ap_id, apply natural_square_tr }},
{ intros f, apply eq_of_phomotopy, fapply phomotopy.mk,
{ intro a, exact mul_one (f a) },
{ reflexivity }},
{ exact ppi_inv },
{ intro f, apply eq_of_phomotopy, fapply phomotopy.mk,
{ intro a, exact con.left_inv (f a) },
{ exact !con_left_inv_idp }},
end
definition group_trunc_ppi [constructor] [instance] {A : Type*} (B : A → Type*) :
group (trunc 0 (Π*a, Ω (B a))) :=
!trunc_group
definition inf_group_ppi [constructor] {A : Type*} (B : A → Type*) :
inf_group (Π*a, Ω (B a)) :=
@inf_group_of_inf_pgroup _ (inf_pgroup_pppi B)
definition gppi_loop [constructor] {A : Type*} (B : A → Type*) : InfGroup :=
InfGroup.mk (Π*a, Ω (B a)) (inf_group_ppi B)
definition Group_trunc_ppi [reducible] [constructor] {A : Type*} (B : A → Type*) : Group :=
Group.mk (trunc 0 (Π*a, Ω (B a))) _
gtrunc (gppi_loop B)
definition ab_inf_group_ppi [constructor] [instance] {A : Type*} (B : A → Type*) :
definition ab_inf_group_ppi [constructor] {A : Type*} (B : A → Type*) :
ab_inf_group (Π*a, Ω (Ω (B a))) :=
⦃ab_inf_group, inf_group_ppi (λa, Ω (B a)), mul_comm :=
begin
@ -56,23 +58,20 @@ namespace group
{ symmetry, rexact eq_of_square (eckmann_hilton_con2 (respect_pt f) (respect_pt g)) }
end⦄
definition ab_group_trunc_ppi [constructor] [instance] {A : Type*} (B : A → Type*) :
ab_group (trunc 0 (Π*a, Ω (Ω (B a)))) :=
!trunc_ab_group
definition agppi_loop [constructor] {A : Type*} (B : A → Type*) : AbInfGroup :=
AbInfGroup.mk (Π*a, Ω (Ω (B a))) (ab_inf_group_ppi B)
definition AbGroup_trunc_ppi [reducible] [constructor] {A : Type*} (B : A → Type*) : AbGroup :=
AbGroup.mk (trunc 0 (Π*a, Ω (Ω (B a)))) _
agtrunc (agppi_loop B)
definition trunc_ppi_isomorphic_pmap (A B : Type*)
: Group.mk (trunc 0 (Π*(a : A), Ω B)) !trunc_group
≃g Group.mk (trunc 0 (A →* Ω B)) !trunc_group :=
begin
reflexivity,
-- apply trunc_isomorphism_of_equiv (pppi_equiv_pmap A (Ω B)),
-- intro h k, induction h with h h_pt, induction k with k k_pt, reflexivity
end
section
-- definition trunc_ppi_isomorphic_pmap (A B : Type*)
-- : Group.mk (trunc 0 (Π*(a : A), Ω B)) !group_trunc
-- ≃g Group.mk (trunc 0 (A →* Ω B)) !group_trunc :=
-- begin
-- reflexivity,
-- -- apply trunc_isomorphism_of_equiv (pppi_equiv_pmap A (Ω B)),
-- -- intro h k, induction h with h h_pt, induction k with k k_pt, reflexivity
-- end
universe variables u v
@ -101,23 +100,22 @@ namespace group
variable (k)
definition trunc_ppi_loop_isomorphism_lemma
: isomorphism.{(max u v) (max u v)}
(Group.mk (trunc 0 (k = k)) (@trunc_group (k = k) !inf_group_loop))
(Group.mk (trunc 0 (Π*(a : A), Ω (pType.mk (B a) (k a)))) !trunc_group) :=
definition gloop_ppi_isomorphism :
Ωg (pointed.Mk k) ≃∞g gppi_loop (λ a, pointed.Mk (ppi.to_fun k a)) :=
begin
apply @trunc_isomorphism_of_equiv _ _ !inf_group_loop !inf_group_ppi (ppi_loop_equiv k),
apply inf_isomorphism_of_equiv (ppi_loop_equiv k),
intro f g, induction k with k p, induction p,
apply trans (phomotopy_of_eq_homomorphism f g),
exact ppi_mul_loop (phomotopy_of_eq f) (phomotopy_of_eq g)
end
end
definition trunc_ppi_loop_isomorphism_lemma :
gtrunc (gloop (pointed.Mk k)) ≃g gtrunc (gppi_loop (λa, pointed.Mk (k a))) :=
gtrunc_isomorphism_gtrunc (gloop_ppi_isomorphism k)
definition trunc_ppi_loop_isomorphism {A : Type*} (B : A → Type*)
: Group.mk (trunc 0 (Ω (Π*(a : A), B a))) !trunc_group
≃g Group.mk (trunc 0 (Π*(a : A), Ω (B a))) !trunc_group :=
trunc_ppi_loop_isomorphism_lemma (ppi_const B)
definition trunc_ppi_loop_isomorphism {A : Type*} (B : A → Type*) :
gtrunc (gloop (Π*(a : A), B a)) ≃g gtrunc (gppi_loop B) :=
proof trunc_ppi_loop_isomorphism_lemma (ppi_const B) qed
/- We first define the group structure on A →* Ω B (except for truncatedness).
@ -171,21 +169,24 @@ namespace group
loop_susp_intro (pmap_mul f g) ~* pmap_mul (loop_susp_intro f) (loop_susp_intro g) :=
pwhisker_right _ !ap1_pmap_mul ⬝* !pmap_mul_pcompose
definition inf_group_pmap [constructor] [instance] (A B : Type*) : inf_group (A →* Ω B) :=
!inf_group_ppi
definition InfGroup_pmap [reducible] [constructor] (A B : Type*) : InfGroup :=
InfGroup.mk (A →* Ω B) !inf_group_ppi
definition group_trunc_pmap [constructor] [instance] (A B : Type*) : group (trunc 0 (A →* Ω B)) :=
!trunc_group
definition InfGroup_pmap' [reducible] [constructor] (A : Type*) {B C : Type*} (e : Ω C ≃* B) :
InfGroup :=
InfGroup.mk (A →* B)
(@inf_group_of_inf_pgroup _ (inf_pgroup_pequiv_closed (ppmap_pequiv_ppmap_right e)
!inf_pgroup_pppi))
definition Group_trunc_pmap [reducible] [constructor] (A B : Type*) : Group :=
Group.mk (trunc 0 (A →* Ω B)) _
Group.mk (trunc 0 (A →* Ω B)) (@group_trunc _ !inf_group_ppi)
definition Group_trunc_pmap_homomorphism [constructor] {A A' B : Type*} (f : A' →* A) :
Group_trunc_pmap A B →g Group_trunc_pmap A' B :=
begin
fapply homomorphism.mk,
{ apply trunc_functor, intro g, exact g ∘* f},
{ intro g h, induction g with g, induction h with h, apply ap tr,
{ intro g h, induction g with g, esimp, induction h with h, apply ap tr,
apply eq_of_phomotopy, fapply phomotopy.mk,
{ intro a, reflexivity },
{ symmetry, refine _ ⬝ !idp_con⁻¹,
@ -242,16 +243,11 @@ namespace group
definition ab_inf_group_pmap [constructor] [instance] (A B : Type*) :
ab_inf_group (A →* Ω (Ω B)) :=
⦃ab_inf_group, inf_group_pmap A (Ω B), mul_comm :=
begin
intro f g, apply eq_of_phomotopy, fapply phomotopy.mk,
{ intro a, exact eckmann_hilton (f a) (g a) },
{ symmetry, rexact eq_of_square (eckmann_hilton_con2 (respect_pt f) (respect_pt g)) }
end⦄
!ab_inf_group_ppi
definition ab_group_trunc_pmap [constructor] [instance] (A B : Type*) :
ab_group (trunc 0 (A →* Ω (Ω B))) :=
!trunc_ab_group
!ab_group_trunc
definition AbGroup_trunc_pmap [reducible] [constructor] (A B : Type*) : AbGroup :=
AbGroup.mk (trunc 0 (A →* Ω (Ω B))) _

View file

@ -165,13 +165,13 @@ namespace group
fapply is_equiv.adjointify,
exact dirsum_functor (λ i, (f i)⁻¹ᵍ),
{ intro ds,
refine (homomorphism_comp_compute (dirsum_functor (λ i, f i)) (dirsum_functor (λ i, (f i)⁻¹ᵍ)) _)⁻¹ ⬝ _,
refine (homomorphism_compose_eq (dirsum_functor (λ i, f i)) (dirsum_functor (λ i, (f i)⁻¹ᵍ)) _)⁻¹ ⬝ _,
refine dirsum_functor_compose (λ i, f i) (λ i, (f i)⁻¹ᵍ) ds ⬝ _,
refine dirsum_functor_homotopy _ (λ i, !gid) (λ i, to_right_inv (equiv_of_isomorphism (f i))) ds ⬝ _,
exact !dirsum_functor_gid
},
{ intro ds,
refine (homomorphism_comp_compute (dirsum_functor (λ i, (f i)⁻¹ᵍ)) (dirsum_functor (λ i, f i)) _)⁻¹ ⬝ _,
refine (homomorphism_compose_eq (dirsum_functor (λ i, (f i)⁻¹ᵍ)) (dirsum_functor (λ i, f i)) _)⁻¹ ⬝ _,
refine dirsum_functor_compose (λ i, (f i)⁻¹ᵍ) (λ i, f i) ds ⬝ _,
refine dirsum_functor_homotopy _ (λ i, !gid) (λ i x,
proof
@ -197,19 +197,19 @@ namespace group
fapply adjointify,
{ exact from_hom },
{ intro ds,
refine (homomorphism_comp_compute to_hom from_hom ds)⁻¹ ⬝ _,
refine (homomorphism_compose_eq to_hom from_hom ds)⁻¹ ⬝ _,
refine @dirsum_homotopy I _ Y (dirsum Y) (to_hom ∘g from_hom) !gid _ ds,
intro i y,
refine homomorphism_comp_compute to_hom from_hom _ ⬝ _,
refine homomorphism_compose_eq to_hom from_hom _ ⬝ _,
refine ap to_hom (dirsum_elim_compute (λi, dirsum_incl (Y ∘ down.{u v}) (up.{u v} i)) i y) ⬝ _,
refine dirsum_elim_compute _ (up.{u v} i) y ⬝ _,
reflexivity
},
{ intro ds,
refine (homomorphism_comp_compute from_hom to_hom ds)⁻¹ ⬝ _,
refine (homomorphism_compose_eq from_hom to_hom ds)⁻¹ ⬝ _,
refine @dirsum_homotopy _ _ (Y ∘ down.{u v}) (dirsum (Y ∘ down.{u v})) (from_hom ∘g to_hom) !gid _ ds,
intro i y, induction i with i,
refine homomorphism_comp_compute from_hom to_hom _ ⬝ _,
refine homomorphism_compose_eq from_hom to_hom _ ⬝ _,
refine ap from_hom (dirsum_elim_compute (λi, dirsum_incl Y (down.{u v} i)) (up.{u v} i) y) ⬝ _,
refine dirsum_elim_compute _ i y ⬝ _,
reflexivity

View file

@ -234,7 +234,7 @@ definition is_trunc_rlist {n : ℕ₋₂} {X : Type} (H : is_trunc (n.+2) X) :
is_trunc (n.+2) (rlist X) :=
begin
apply is_trunc_sigma, { apply is_trunc_list, apply is_trunc_sum },
intro l, apply is_trunc_succ_of_is_prop
intro l, exact is_trunc_succ_of_is_prop _ _ _
end
definition is_reduced_invert (v : X ⊎ X) : is_reduced (v::l) → is_reduced l :=

View file

@ -367,7 +367,7 @@ end
isomorphism.mk (to_lminv φ) !is_equiv_inv
definition isomorphism.trans [trans] [constructor] (φ : M₁ ≃lm M₂) (ψ : M₂ ≃lm M₃) : M₁ ≃lm M₃ :=
isomorphism.mk (ψ ∘lm φ) !is_equiv_compose
isomorphism.mk (ψ ∘lm φ) (is_equiv_compose ψ φ _ _)
definition isomorphism.eq_trans [trans] [constructor]
{M₁ M₂ : LeftModule R} {M₃ : LeftModule R} (φ : M₁ = M₂) (ψ : M₂ ≃lm M₃) : M₁ ≃lm M₃ :=

View file

@ -72,7 +72,7 @@ namespace group
φ (x.1 * y.1) * ψ (x.2 * y.2) = (φ x.1 * φ y.1) * (ψ x.2 * ψ y.2)
: by exact ap011 mul (to_respect_mul φ x.1 y.1) (to_respect_mul ψ x.2 y.2)
... = (φ x.1 * ψ x.2) * (φ y.1 * ψ y.2)
: by exact interchange I (φ x.1) (φ y.1) (ψ x.2) (ψ y.2)) end
: by exact mul.comm4 (φ x.1) (φ y.1) (ψ x.2) (ψ y.2)) end
definition product_functor [constructor] {G G' H H' : Group} (φ : G →g H) (ψ : G' →g H') :
G ×g G' →g H ×g H' :=

View file

@ -331,7 +331,7 @@ namespace group
have htpy : Π (a : A), K a ≃ L a,
begin
intro a,
apply @equiv_of_is_prop (a ∈ K) (a ∈ L) _ _ (K_in_L a) (L_in_K a),
exact @equiv_of_is_prop (a ∈ K) (a ∈ L) (K_in_L a) (L_in_K a) _ _,
end,
exact subgroup_rel_eq' htpy,
end
@ -561,11 +561,12 @@ definition ab_group_kernel_quotient_to_image_codomain_triangle {A B : AbGroup} (
fapply is_prop.elimo
end
-- set_option pp.all true
-- print algebra._trans_of_Group_of_AbGroup_2
definition is_surjective_kernel_quotient_to_image {A B : AbGroup} (f : A →g B)
: is_surjective (ab_group_kernel_quotient_to_image f) :=
begin
fapply is_surjective_factor (group_fun (ab_qg_map (kernel f))),
exact image_lift f,
refine is_surjective_factor (ab_qg_map (kernel f)) (image_lift f) _ _,
apply @quotient_group_compute _ _ _ (@is_normal_subgroup_ab _ (kernel f) _),
exact is_surjective_image_lift f
end
@ -573,7 +574,7 @@ definition is_surjective_kernel_quotient_to_image {A B : AbGroup} (f : A →g B)
definition is_embedding_kernel_quotient_to_image {A B : AbGroup} (f : A →g B)
: is_embedding (ab_group_kernel_quotient_to_image f) :=
begin
fapply is_embedding_factor (ab_group_kernel_quotient_to_image f) (image_incl f) (kernel_quotient_extension f),
fapply is_embedding_factor (image_incl f) (kernel_quotient_extension f),
exact ab_group_kernel_quotient_to_image_codomain_triangle f,
exact is_embedding_kernel_quotient_extension f
end

View file

@ -164,7 +164,7 @@ definition is_prop_submodule (S : property M) [is_submodule M S] [H : is_prop M]
begin apply @is_trunc_sigma, exact H end
local attribute is_prop_submodule [instance]
definition is_contr_submodule [instance] (S : property M) [is_submodule M S] [is_contr M] : is_contr (submodule S) :=
is_contr_of_inhabited_prop 0
is_contr_of_inhabited_prop 0 _
definition submodule_isomorphism [constructor] (S : property M) [is_submodule M S] (h : Πg, g ∈ S) :
submodule S ≃lm M :=
@ -244,7 +244,7 @@ local attribute is_prop_quotient_module [instance]
definition is_contr_quotient_module [instance] (S : property M) [is_submodule M S] [is_contr M] :
is_contr (quotient_module S) :=
is_contr_of_inhabited_prop 0
is_contr_of_inhabited_prop 0 _
definition rel_of_is_contr_quotient_module (S : property M) [is_submodule M S]
(H : is_contr (quotient_module S)) (m : M) : S m :=

View file

@ -18,7 +18,7 @@ equiv.mk _ (H A)
definition has_choice_of_succ (X : Type) (H : Πk, has_choice (k.+1) X) (n : ℕ₋₂) : has_choice n X :=
begin
cases n with n,
{ intro A, apply is_equiv_of_is_contr },
{ intro A, exact is_equiv_of_is_contr _ _ _ },
{ exact H n }
end

View file

@ -85,7 +85,6 @@ end
definition cohomology_isomorphism_shomotopy_group_sp_cotensor (X : Type*) (Y : spectrum) {n m : }
(p : -m = n) : H^n[X, Y] ≃g πₛ[m] (sp_cotensor X Y) :=
begin
refine !trunc_ppi_isomorphic_pmap⁻¹ᵍ ⬝g _,
refine parametrized_cohomology_isomorphism_shomotopy_group_spi (λx, Y) p ⬝g _,
apply shomotopy_group_isomorphism_of_pequiv, intro k,
apply pppi_pequiv_ppmap
@ -297,7 +296,7 @@ begin
{ exfalso, apply H, reflexivity },
{ apply is_conn_of_le, apply zero_le_of_nat n, exact is_conn_EMadd1 G n, }},
{ apply is_trunc_trunc_of_is_trunc, apply @is_contr_loop_of_is_trunc (n+1) (K G 0),
apply is_trunc_of_le _ (zero_le_of_nat n) }
apply is_trunc_of_le _ (zero_le_of_nat n) _ }
end
theorem EM_dimension (G : AbGroup) (n : ) (H : n ≠ 0) :

View file

@ -85,13 +85,15 @@ begin
exact pfiber_postnikov_map A n
end
set_option formatter.hide_full_terms false
definition pfiber_postnikov_map_pred' (A : spectrum) (n k l : ) (p : n + k = l) :
pfiber (postnikov_map_pred (A k) (maxm2 l)) ≃* EM_spectrum (πₛ[n] A) l :=
begin
cases l with l l,
{ refine pfiber_postnikov_map_pred (A k) l ⬝e* _,
exact EM_type_pequiv_EM A p },
{ apply pequiv_of_is_contr, apply is_contr_pfiber_pid,
{ refine pequiv_of_is_contr _ _ _ _, apply is_contr_pfiber_pid,
apply is_contr_EM_spectrum_neg }
end

View file

@ -322,8 +322,8 @@ variable (f)
definition is_contr_seq_colim {A : → Type} (f : seq_diagram A)
[Πk, is_contr (A k)] : is_contr (seq_colim f) :=
begin
apply @is_trunc_is_equiv_closed (A 0),
apply is_equiv_inclusion0, intro n, apply is_equiv_of_is_contr
refine is_contr_is_equiv_closed (ι' f 0) _ _,
apply is_equiv_inclusion0, intro n, exact is_equiv_of_is_contr _ _ _
end
definition seq_colim_equiv_of_is_equiv [constructor] {n : } (H : Πk, k ≥ n → is_equiv (@f k)) :
@ -882,7 +882,7 @@ theorem is_trunc_fun_inclusion (k : ℕ₋₂) (H : Πn, is_trunc_fun k (@f n))
is_trunc_fun k (ι' f 0) :=
begin
intro x,
apply @(is_trunc_equiv_closed_rev k (fiber_inclusion x)),
apply is_trunc_equiv_closed_rev k (fiber_inclusion x),
apply is_trunc_fun_seq_colim_functor,
intro n, apply is_trunc_fun_lrep, exact H
end
@ -913,11 +913,11 @@ definition lrep_seq_diagram_fin {n : } (x : fin n) :
begin
induction x with k H, esimp, induction H with n H p,
reflexivity,
exact ap (@lift_succ2 _) p
exact ap (@lift_succ _) p
end
definition lrep_seq_diagram_fin_lift_succ {n : } (x : fin n) :
lrep_seq_diagram_fin (lift_succ2 x) = ap (@lift_succ2 _) (lrep_seq_diagram_fin x) :=
lrep_seq_diagram_fin (lift_succ x) = ap (@lift_succ _) (lrep_seq_diagram_fin x) :=
begin
induction x with k H, reflexivity
end

View file

@ -6,7 +6,7 @@ Authors: Floris van Doorn, Egbert Rijke
import ..move_to_lib types.fin types.trunc
open nat eq equiv sigma sigma.ops is_equiv is_trunc trunc prod fiber function
open nat eq equiv sigma sigma.ops is_equiv is_trunc trunc prod fiber function is_conn
namespace seq_colim
@ -77,7 +77,7 @@ namespace seq_colim
begin
induction H with m H Hlrepf,
{ apply is_equiv_id },
{ exact is_equiv_compose (@f _) (lrep f H) },
{ exact is_equiv_compose (@f _) (lrep f H) _ _ },
end
local attribute is_equiv_lrep [instance]
@ -211,7 +211,7 @@ namespace seq_colim
open fin
definition seq_diagram_fin [unfold_full] : seq_diagram fin :=
lift_succ2
lift_succ
definition id0_seq [unfold_full] (a₁ a₂ : A 0) : → Type :=
λ k, rep0 f k a₁ = rep0 f k a₂

100
component.hlean Normal file
View file

@ -0,0 +1,100 @@
-- Author: Floris van Doorn
import homotopy.connectedness .move_to_lib
open eq equiv pointed is_conn is_trunc sigma prod trunc function group nat fiber
namespace is_conn
open sigma.ops pointed trunc_index
/- this is equivalent to pfiber (A → ∥A∥₀) ≡ connect 0 A -/
definition component [constructor] (A : Type*) : Type* :=
pType.mk (Σ(a : A), merely (pt = a)) ⟨pt, tr idp⟩
lemma is_conn_component [instance] (A : Type*) : is_conn 0 (component A) :=
is_conn_zero_pointed'
begin intro x, induction x with a p, induction p with p, induction p, exact tidp end
definition component_incl [constructor] (A : Type*) : component A →* A :=
pmap.mk pr1 idp
definition is_embedding_component_incl [instance] (A : Type*) : is_embedding (component_incl A) :=
is_embedding_pr1 _
definition component_intro [constructor] {A B : Type*} (f : A →* B) (H : merely_constant f) :
A →* component B :=
begin
fapply pmap.mk,
{ intro a, refine ⟨f a, _⟩, exact tinverse (merely_constant_pmap H a) },
exact subtype_eq !respect_pt
end
definition component_functor [constructor] {A B : Type*} (f : A →* B) :
component A →* component B :=
component_intro (f ∘* component_incl A) !merely_constant_of_is_conn
-- definition component_elim [constructor] {A B : Type*} (f : A →* B) (H : merely_constant f) :
-- A →* component B :=
-- begin
-- fapply pmap.mk,
-- { intro a, refine ⟨f a, _⟩, exact tinverse (merely_constant_pmap H a) },
-- exact subtype_eq !respect_pt
-- end
definition loop_component (A : Type*) : Ω (component A) ≃* Ω A :=
loop_pequiv_loop_of_is_embedding (component_incl A)
lemma loopn_component (n : ) (A : Type*) : Ω[n+1] (component A) ≃* Ω[n+1] A :=
!loopn_succ_in ⬝e* loopn_pequiv_loopn n (loop_component A) ⬝e* !loopn_succ_in⁻¹ᵉ*
-- lemma fundamental_group_component (A : Type*) : π₁ (component A) ≃g π₁ A :=
-- isomorphism_of_equiv (trunc_equiv_trunc 0 (loop_component A)) _
lemma homotopy_group_component (n : ) (A : Type*) : πg[n+1] (component A) ≃g πg[n+1] A :=
homotopy_group_isomorphism_of_is_embedding (n+1) (component_incl A)
definition is_trunc_component [instance] (n : ℕ₋₂) (A : Type*) [is_trunc n A] :
is_trunc n (component A) :=
begin
apply @is_trunc_sigma, intro a, cases n with n,
{ refine is_contr_of_inhabited_prop _ _, exact tr !is_prop.elim },
{ exact is_trunc_succ_of_is_prop _ _ _ },
end
definition ptrunc_component' (n : ℕ₋₂) (A : Type*) :
ptrunc (n.+2) (component A) ≃* component (ptrunc (n.+2) A) :=
begin
fapply pequiv.MK',
{ exact ptrunc.elim (n.+2) (component_functor !ptr) },
{ intro x, cases x with x p, induction x with a,
refine tr ⟨a, _⟩,
note q := trunc_functor -1 !tr_eq_tr_equiv p,
exact trunc_trunc_equiv_left _ !minus_one_le_succ q },
{ exact sorry },
{ exact sorry }
end
definition ptrunc_component (n : ℕ₋₂) (A : Type*) :
ptrunc n (component A) ≃* component (ptrunc n A) :=
begin
cases n with n, exact sorry,
cases n with n, exact sorry,
exact ptrunc_component' n A
end
definition break_into_components (A : Type) : A ≃ Σ(x : trunc 0 A), Σ(a : A), ∥ tr a = x ∥ :=
calc
A ≃ Σ(a : A) (x : trunc 0 A), tr a = x :
by exact (@sigma_equiv_of_is_contr_right _ _ (λa, !is_contr_sigma_eq))⁻¹ᵉ
... ≃ Σ(x : trunc 0 A) (a : A), tr a = x :
by apply sigma_comm_equiv
... ≃ Σ(x : trunc 0 A), Σ(a : A), ∥ tr a = x ∥ :
by exact sigma_equiv_sigma_right (λx, sigma_equiv_sigma_right (λa, !trunc_equiv⁻¹ᵉ))
definition pfiber_pequiv_component_of_is_contr [constructor] {A B : Type*} (f : A →* B)
[is_contr B]
/- extra condition, something like trunc_functor 0 f is an embedding -/ : pfiber f ≃* component A :=
sorry
end is_conn

View file

@ -6,7 +6,7 @@ Authors: Ulrik Buchholtz, Egbert Rijke, Floris van Doorn
Formalization of the higher groups paper
-/
import .homotopy.EM
import .homotopy.EM algebra.category.constructions.pullback
open eq is_conn pointed is_trunc trunc equiv is_equiv trunc_index susp nat algebra
prod.ops sigma sigma.ops category EM
namespace higher_group

View file

@ -102,8 +102,8 @@ namespace homology
fapply is_equiv_of_equiv_of_homotopy,
{ exact equiv_of_isomorphism iso },
{ refine dirsum_homotopy _, intro i y,
refine homomorphism_comp_compute iso3 (iso2 ∘g iso1) _ ⬝ _,
refine ap iso3 (homomorphism_comp_compute iso2 iso1 _) ⬝ _,
refine homomorphism_compose_eq iso3 (iso2 ∘g iso1) _ ⬝ _,
refine ap iso3 (homomorphism_compose_eq iso2 iso1 _) ⬝ _,
refine ap (iso3 ∘ iso2) _ ⬝ _,
{ exact dirsum_incl (λ i, HH theory n (X (down i))) (up i) y },
{ refine _ ⬝ dirsum_elim_compute (λi, dirsum_incl (λ i, HH theory n (X (down.{0 u} i))) (up i)) i y,

View file

@ -197,35 +197,35 @@ namespace EM
definition homotopy_group_functor_EMadd1_functor' {G H : AbGroup} (φ : G →g H) (n : ) :
hsquare (ghomotopy_group_EMadd1' G n)⁻¹ᵍ (ghomotopy_group_EMadd1' H n)⁻¹ᵍ
(π→g[n+1] (EMadd1_functor φ n)) φ :=
(homotopy_group_functor_EMadd1_functor φ n)⁻¹ʰᵗʸʰ
begin apply hhinverse, exact (homotopy_group_functor_EMadd1_functor φ n) end
definition EM1_pmap_natural {G H : Group} {X Y : Type*} (f : X →* Y) (eX : G → Ω X)
(eY : H → Ω Y) (rX : Πg h, eX (g * h) = eX g ⬝ eX h) (rY : Πg h, eY (g * h) = eY g ⬝ eY h)
[H2 : is_trunc 1 X] [is_trunc 1 Y] (φ : G →g H) (p : hsquare eX eY φ (Ω→ f)) :
psquare (EM1_pmap eX rX) (EM1_pmap eY rY) (EM1_functor φ) f :=
section infgroup
open infgroup
definition EM1_pmap_natural {G H : Group} {X Y : Type*} (f : X →* Y) (eX : G →∞g Ωg X)
(eY : H →∞g Ωg Y) [H2 : is_trunc 1 X] [is_trunc 1 Y] (φ : G →g H)
(p : hsquare eX eY φ (Ωg→ f)) : psquare (EM1_pmap eX) (EM1_pmap eY) (EM1_functor φ) f :=
begin
fapply phomotopy.mk,
{ intro x, induction x using EM.set_rec,
{ exact respect_pt f },
{ apply eq_pathover,
refine ap_compose f _ _ ⬝ph _ ⬝hp (ap_compose (EM1_pmap eY rY) _ _)⁻¹,
refine ap_compose f _ _ ⬝ph _ ⬝hp (ap_compose (EM1_pmap eY) _ _)⁻¹,
refine ap02 _ !elim_pth ⬝ph _ ⬝hp ap02 _ !elim_pth⁻¹,
refine _ ⬝hp !elim_pth⁻¹, apply transpose, exact square_of_eq_bot (p g) }},
{ exact !idp_con⁻¹ }
end
definition EM1_pequiv'_natural {G H : Group} {X Y : Type*} (f : X →* Y) (eX : G ≃* Ω X)
(eY : H ≃* Ω Y) (rX : Πg h, eX (g * h) = eX g ⬝ eX h) (rY : Πg h, eY (g * h) = eY g ⬝ eY h)
[H1 : is_conn 0 X] [H2 : is_trunc 1 X] [is_conn 0 Y] [is_trunc 1 Y]
definition EM1_pequiv'_natural {G H : Group} {X Y : Type*} (f : X →* Y) (eX : G ≃∞g Ωg X)
(eY : H ≃∞g Ωg Y) [H1 : is_conn 0 X] [H2 : is_trunc 1 X] [is_conn 0 Y] [is_trunc 1 Y]
(φ : G →g H) (p : hsquare eX eY φ (Ω→ f)) :
psquare (EM1_pequiv' eX rX) (EM1_pequiv' eY rY) (EM1_functor φ) f :=
EM1_pmap_natural f eX eY rX rY φ p
psquare (EM1_pequiv' eX) (EM1_pequiv' eY) (EM1_functor φ) f :=
EM1_pmap_natural f eX eY φ p
definition EM1_pequiv_natural {G H : Group} {X Y : Type*} (f : X →* Y) (eX : G ≃g π₁ X)
(eY : H ≃g π₁ Y) [H1 : is_conn 0 X] [H2 : is_trunc 1 X] [is_conn 0 Y] [is_trunc 1 Y]
(φ : G →g H) (p : hsquare eX eY φ (π→g[1] f)) :
psquare (EM1_pequiv eX) (EM1_pequiv eY) (EM1_functor φ) f :=
EM1_pequiv'_natural f _ _ _ _ φ
EM1_pequiv'_natural f _ _ φ
begin
assert p' : ptrunc_functor 0 (Ω→ f) ∘* pequiv_of_isomorphism eX ~*
pequiv_of_isomorphism eY ∘* pmap_of_homomorphism φ, exact phomotopy_of_homotopy p,
@ -238,38 +238,35 @@ namespace EM
begin refine EM1_pequiv_natural f _ _ _ _, exact homotopy.rfl end
definition EM_up_natural {G H : AbGroup} (φ : G →g H) {X Y : Type*} (f : X →* Y) {n : }
(eX : G →* Ω[succ (succ n)] X) (eY : H →* Ω[succ (succ n)] Y)
(p : hsquare eX eY φ (Ω→[succ (succ n)] f))
(eX : G →∞g Ωg[succ (succ n)] X) (eY : H →∞g Ωg[succ (succ n)] Y)
(p : hsquare eX eY φ (Ωg→[succ (succ n)] f))
: hsquare (EM_up eX) (EM_up eY) φ (Ω→[succ n] (Ω→ f)) :=
begin
refine p ⬝htyh hsquare_of_psquare (loopn_succ_in_natural (succ n) f),
end
p ⬝htyh hsquare_of_psquare (loopn_succ_in_natural (succ n) f)
definition EMadd1_pmap_natural {G H : AbGroup} {X Y : Type*} (f : X →* Y) (n : )
(eX : G →* Ω[succ n] X) (eY : H →* Ω[succ n] Y) (rX : Π(g h : G), eX (g * h) = eX g ⬝ eX h)
(rY : Π(g h : H), eY (g * h) = eY g ⬝ eY h) [HX : is_trunc (n.+1) X] [HY : is_trunc (n.+1) Y]
(φ : G →g H) (p : hsquare eX eY φ (Ω→[succ n] f)) :
psquare (EMadd1_pmap n eX rX) (EMadd1_pmap n eY rY) (EMadd1_functor φ n) f :=
(eX : G →∞g Ωg[succ n] X) (eY : H →∞g Ωg[succ n] Y)
[HX : is_trunc (n.+1) X] [HY : is_trunc (n.+1) Y]
(φ : G →g H) (p : hsquare eX eY φ (Ωg→[succ n] f)) :
psquare (EMadd1_pmap n eX) (EMadd1_pmap n eY) (EMadd1_functor φ n) f :=
begin
revert X Y f eX eY rX rY HX HY p, induction n with n IH: intros,
revert X Y f eX eY HX HY p, induction n with n IH: intros,
{ apply EM1_pmap_natural, exact p },
{ do 2 rewrite [EMadd1_pmap_succ], refine _ ⬝* pwhisker_left _ !EMadd1_functor_succ⁻¹*,
refine (ptrunc_elim_pcompose ((succ n).+1) _ _)⁻¹* ⬝* _ ⬝*
(ptrunc_elim_ptrunc_functor ((succ n).+1) _ _)⁻¹*,
apply ptrunc_elim_phomotopy,
refine _ ⬝* !susp_elim_susp_functor⁻¹*,
refine _ ⬝* susp_elim_phomotopy (IH _ _ _ _ _ (is_homomorphism_EM_up eX rX) _
refine _ ⬝* susp_elim_phomotopy (IH _ _ _ _ _
(@is_trunc_loop _ _ HX) _ (EM_up_natural φ f eX eY p)),
apply susp_elim_natural }
end
definition EMadd1_pequiv'_natural {G H : AbGroup} {X Y : Type*} (f : X →* Y) (n : )
(eX : G ≃* Ω[succ n] X) (eY : H ≃* Ω[succ n] Y) (rX : Π(g h : G), eX (g * h) = eX g ⬝ eX h)
(rY : Π(g h : H), eY (g * h) = eY g ⬝ eY h)
(eX : G ≃∞g Ωg[succ n] X) (eY : H ≃∞g Ωg[succ n] Y)
[H1 : is_conn n X] [H2 : is_trunc (n.+1) X] [is_conn n Y] [is_trunc (n.+1) Y]
(φ : G →g H) (p : hsquare eX eY φ (Ω→[succ n] f)) :
psquare (EMadd1_pequiv' n eX rX) (EMadd1_pequiv' n eY rY) (EMadd1_functor φ n) f :=
EMadd1_pmap_natural f n eX eY rX rY φ p
(φ : G →g H) (p : hsquare eX eY φ (Ωg→[succ n] f)) :
psquare (EMadd1_pequiv' n eX) (EMadd1_pequiv' n eY) (EMadd1_functor φ n) f :=
EMadd1_pmap_natural f n eX eY φ p
definition EMadd1_pequiv_natural_local_instance {X : Type*} (n : ) [H : is_trunc (n.+1) X] :
is_set (Ω[succ n] X) :=
@ -277,14 +274,17 @@ namespace EM
local attribute EMadd1_pequiv_natural_local_instance [instance]
definition EMadd1_pequiv_natural {G H : AbGroup} {X Y : Type*} (f : X →* Y) (n : ) (eX : G ≃g πg[n+1] X)
(eY : H ≃g πg[n+1] Y) [H1 : is_conn n X] [H2 : is_trunc (n.+1) X] [H3 : is_conn n Y]
definition EMadd1_pequiv_natural {G H : AbGroup} {X Y : Type*} (f : X →* Y) (n : )
(eX : G ≃g πg[n+1] X) (eY : H ≃g πg[n+1] Y)
[H1 : is_conn n X] [H2 : is_trunc (n.+1) X] [H3 : is_conn n Y]
[H4 : is_trunc (n.+1) Y] (φ : G →g H) (p : hsquare eX eY φ (π→g[n+1] f)) :
psquare (EMadd1_pequiv n eX) (EMadd1_pequiv n eY) (EMadd1_functor φ n) f :=
EMadd1_pequiv'_natural f n
(pequiv_of_isomorphism eX ⬝e* ptrunc_pequiv 0 (Ω[succ n] X))
(pequiv_of_isomorphism eY ⬝e* ptrunc_pequiv 0 (Ω[succ n] Y))
_ _ φ (p ⬝htyh hsquare_of_psquare (ptrunc_pequiv_natural 0 (Ω→[succ n] f)))
_ --(inf_isomorphism_of_isomorphism eX ⬝∞g gtrunc_isomorphism (Ω[succ n] X))
_ --(inf_isomorphism_of_isomorphism eY ⬝∞g gtrunc_isomorphism (Ω[succ n] Y))
φ (p ⬝htyh hsquare_of_psquare (ptrunc_pequiv_natural 0 (Ω→[succ n] f)))
end infgroup
definition EMadd1_pequiv_succ_natural {G H : AbGroup} {X Y : Type*} (f : X →* Y) (n : )
(eX : G ≃g πag[n+2] X) (eY : H ≃g πag[n+2] Y) [is_conn (n.+1) X] [is_trunc (n.+2) X]
@ -343,6 +343,24 @@ namespace EM
{ refine !loopn_succ_in ⬝e* Ω≃[n] (loop_EMadd1 G (m + n))⁻¹ᵉ* ⬝e* e }
end
/- properties about EM -/
definition gEM (G : AbGroup) (n : ) : InfGroup :=
InfGroup.mk (EM G n) (inf_group_equiv_closed (loop_EM G n) _)
definition gEM_functor {G H : AbGroup} (φ : G →g H) (n : ) : gEM G n →∞g gEM H n :=
inf_homomorphism.mk (EM_functor φ n) sorry
definition EM_pmap [unfold 8] {G : AbGroup} (X : InfGroup) (n : )
(e : AbInfGroup_of_AbGroup G →∞g Ωg'[n] X) [H : is_trunc n X] : EM G n →* X :=
begin
cases n with n,
{ exact pmap_of_inf_homomorphism e },
{ have is_trunc (n.+1) X, from H, exact EMadd1_pmap n e }
end
-- definition gEM_gfunctor {G H : AbGroup} (n : ) : (G →gg H) →∞g (gEM G n →∞g gEM H n) :=
-- inf_homomorphism.mk (EM_functor _ n) sorry
/- The Eilenberg-MacLane space K(G,n) with the same homotopy group as X on level n.
On paper this is written K(πₙ(X), n). The problem is that for
@ -508,14 +526,14 @@ namespace EM
definition homotopy_group_cfunctor : cType*[0] ⇒ Grp :=
functor.mk
(λX, πg[1] X)
(λX Y f, π→g[1] f)
(λX Y (f : X →* Y), π→g[1] f)
begin intro, apply homomorphism_eq, exact to_homotopy !homotopy_group_functor_pid end
begin intros, apply homomorphism_eq, exact to_homotopy !homotopy_group_functor_pcompose end
definition ab_homotopy_group_cfunctor (n : ) : cType*[n.+1] ⇒ AbGrp :=
functor.mk
(λX, πag[n+2] X)
(λX Y f, π→g[n+2] f)
(λX Y (f : X →* Y), by rexact π→g[n+2] f)
begin intro, apply homomorphism_eq, exact to_homotopy !homotopy_group_functor_pid end
begin intros, apply homomorphism_eq, exact to_homotopy !homotopy_group_functor_pcompose end
@ -593,8 +611,12 @@ namespace EM
EM1_functor_gid H end
definition is_contr_EM1 {G : Group} (H : is_contr G) : is_contr (EM1 G) :=
is_contr_of_is_conn_of_is_trunc
(is_trunc_succ_succ_of_is_trunc_loop _ _ (is_trunc_equiv_closed _ !loop_EM1 _) _) !is_conn_EM1
begin
refine is_contr_of_is_conn_of_is_trunc _ !is_conn_EM1,
refine is_trunc_succ_succ_of_is_trunc_loop _ _ _ _,
refine is_trunc_equiv_closed _ !loop_EM1 _,
apply is_trunc_succ, exact H
end
definition is_contr_EMadd1 (n : ) {G : AbGroup} (H : is_contr G) : is_contr (EMadd1 G n) :=
begin
@ -642,7 +664,7 @@ namespace EM
definition is_conn_fiber_EM1_functor {G H : Group} (φ : G →g H) :
is_conn -1 (pfiber (EM1_functor φ)) :=
begin
apply is_conn_fiber
apply is_conn_fiber_of_is_conn
end
definition is_trunc_fiber_EMadd1_functor {G H : AbGroup} (φ : G →g H) (n : ) :
@ -654,7 +676,7 @@ namespace EM
definition is_conn_fiber_EMadd1_functor {G H : AbGroup} (φ : G →g H) (n : ) :
is_conn (n.-1) (pfiber (EMadd1_functor φ n)) :=
begin
apply is_conn_fiber, apply is_conn_of_is_conn_succ, apply is_conn_EMadd1,
apply is_conn_fiber_of_is_conn, apply is_conn_of_is_conn_succ, apply is_conn_EMadd1,
apply is_conn_EMadd1
end
@ -667,7 +689,7 @@ namespace EM
definition is_conn_fiber_EM_functor {G H : AbGroup} (φ : G →g H) (n : ) :
is_conn (n.-2) (pfiber (EM_functor φ n)) :=
begin
apply is_conn_fiber, apply is_conn_of_is_conn_succ
apply is_conn_fiber_of_is_conn, apply is_conn_of_is_conn_succ
end
section

View file

@ -1,6 +1,6 @@
-- Authors: Floris van Doorn
import .EM .smash_adjoint ..algebra.ring
import .EM .smash_adjoint ..algebra.ring ..algebra.arrow_group
open algebra eq EM is_equiv equiv is_trunc is_conn pointed trunc susp smash group nat
namespace EM
@ -10,44 +10,58 @@ definition EM1product_adj {R : Ring} :
begin
have is_trunc 1 (ppmap (EM1 (AbGroup_of_Ring R)) (EMadd1 (AbGroup_of_Ring R) 1)),
from is_trunc_pmap_of_is_conn _ _ _ _ _ _ (le.refl 2) !is_trunc_EMadd1,
fapply EM1_pmap,
apply EM1_pmap, fapply inf_homomorphism.mk,
{ intro r, refine pfunext _ _, exact !loop_EM2⁻¹ᵉ* ∘* EM1_functor (ring_right_action r), },
{ intro r r', exact sorry }
end
definition EMproduct_map {G H K : AbGroup} (φ : G → H →g K) (n m : ) (g : G) :
EMadd1 H n →* EMadd1 K n :=
definition EMproduct_map {A B C : AbGroup} (φ : A → B →g C) (n m : ) (a : A) :
EMadd1 B n →* EMadd1 C n :=
begin
fapply EMadd1_functor (φ g) n
fapply EMadd1_functor (φ a) n
end
definition EM0EMadd1product {G H K : AbGroup} (φ : G →g H →gg K) (n : ) :
G →* EMadd1 H n →** EMadd1 K n :=
EMadd1_pfunctor H K n ∘* pmap_of_homomorphism φ
definition EM0EMadd1product {A B C : AbGroup} (φ : A →g B →gg C) (n : ) :
A →* EMadd1 B n →** EMadd1 C n :=
EMadd1_pfunctor B C n ∘* pmap_of_homomorphism φ
definition EMadd1product {G H K : AbGroup} (φ : G →g H →gg K) (n m : ) :
EMadd1 G n →* EMadd1 H m →** EMadd1 K (m + succ n) :=
definition EMadd1product {A B C : AbGroup} (φ : A →g B →gg C) (n m : ) :
EMadd1 A n →* EMadd1 B m →** EMadd1 C (m + succ n) :=
begin
assert H1 : is_trunc n.+1 (EMadd1 H m →** EMadd1 K (m + succ n)),
assert H1 : is_trunc n.+1 (EMadd1 B m →** EMadd1 C (m + succ n)),
{ refine is_trunc_pmap_of_is_conn _ (m.-1) !is_conn_EMadd1 _ _ _ _ !is_trunc_EMadd1,
exact le_of_eq (trunc_index.of_nat_add_plus_two_of_nat m n)⁻¹ᵖ },
fapply EMadd1_pmap,
{ refine (loopn_ppmap_pequiv _ _ _)⁻¹ᵉ* ∘* ppcompose_left !loopn_EMadd1_add⁻¹ᵉ* ∘*
EM0EMadd1product φ m },
{ exact sorry }
apply EMadd1_pmap,
exact sorry
/- the underlying pointed map is: -/
-- refine (loopn_ppmap_pequiv _ _ _)⁻¹ᵉ* ∘* ppcompose_left !loopn_EMadd1_add⁻¹ᵉ* ∘*
-- EM0EMadd1product φ m
end
definition EMproduct {G H K : AbGroup} (φ : G →g H →gg K) (n m : ) :
EM G n →* EM H m →** EM K (m + n) :=
definition EMproduct {A B C : AbGroup} (φ : A →g B →gg C) (n m : ) :
EM A n →* EM B m →** EM C (m + n) :=
begin
cases n with n,
{ cases m with m,
{ exact pmap_of_homomorphism2 φ },
{ exact EM0EMadd1product φ m }},
{ cases m with m,
{ exact ppcompose_left (ptransport (EMadd1 K) (zero_add n)⁻¹) ∘*
{ exact ppcompose_left (ptransport (EMadd1 C) (zero_add n)⁻¹) ∘*
pmap_swap_map (EM0EMadd1product (homomorphism_swap φ) n) },
{ exact ppcompose_left (ptransport (EMadd1 K) !succ_add⁻¹) ∘* EMadd1product φ n m }}
{ exact ppcompose_left (ptransport (EMadd1 C) !succ_add⁻¹) ∘* EMadd1product φ n m }}
end
definition EMproduct' {A B C : AbGroup} (φ : A →g B →gg C) (n m : ) :
EM A n →* EM B m →** EM C (m + n) :=
begin
assert H1 : is_trunc n (InfGroup_pmap' (EM B m) (loop_EM C (m + n))),
{ exact is_trunc_pmap_of_is_conn_nat _ m !is_conn_EM _ _ _ !le.refl !is_trunc_EM },
apply EM_pmap (InfGroup_pmap' (EM B m) (loop_EM C (m + n))) n,
exact sorry
-- exact _ /- (loopn_ppmap_pequiv _ _ _)⁻¹ᵉ* -/ ∘∞g _ /-ppcompose_left !loopn_EMadd1_add⁻¹ᵉ*-/ ∘∞g
-- _ ∘∞g inf_homomorphism_of_homomorphism φ
end
end EM

View file

@ -472,7 +472,7 @@ namespace smash
definition smash_pequiv [constructor] (f : A ≃* C) (g : B ≃* D) : A ∧ B ≃* C ∧ D :=
begin
fapply pequiv_of_pmap (f ∧→ g),
refine @homotopy_closed _ _ _ _ _ (smash_functor_homotopy_pushout_functor f g)⁻¹ʰᵗʸ,
refine homotopy_closed _ (smash_functor_homotopy_pushout_functor f g)⁻¹ʰᵗʸ _,
apply pushout.is_equiv_functor
end
@ -1054,7 +1054,7 @@ namespace smash
begin
unfold [smash, cofiber, smash'], symmetry,
fapply pushout_vcompose_equiv wedge_of_sum,
{ symmetry, apply equiv_unit_of_is_contr, apply is_contr_pushout_wedge_of_sum },
{ symmetry, refine equiv_unit_of_is_contr _ _, apply is_contr_pushout_wedge_of_sum },
{ intro x, reflexivity },
{ apply prod_of_wedge_of_sum }
end

View file

@ -71,7 +71,7 @@ namespace susp
begin
intro p,
refine susp_functor_psquare p ⬝v* _,
exact psquare_transpose (loop_susp_counit_natural f₁₂)
exact ptranspose (loop_susp_counit_natural f₁₂)
end
open pushout unit prod sigma sigma.ops

View file

@ -21,7 +21,7 @@ section
open trunc_index
definition propext {p q : Prop} (h : p ↔ q) : p = q :=
tua (equiv_of_iff_of_is_prop h)
tua (equiv_of_iff_of_is_prop h _ _)
end

File diff suppressed because it is too large Load diff

View file

@ -269,7 +269,7 @@ namespace pointed
fapply phomotopy.mk,
{ intro x, reflexivity },
{ refine !idp_con ⬝ _, esimp,
refine whisker_right _ !ap_sigma_functor_eq_dpair ⬝ _,
refine whisker_right _ !ap_sigma_functor_sigma_eq ⬝ _,
induction f₁ with f₁ f₁₀, induction f₂ with f₂ f₂₀, induction A₂ with A₂ a₂₀,
induction A₃ with A₃ a₃₀, esimp at * ⊢, induction p₁, induction p₂, reflexivity }
end
@ -977,10 +977,11 @@ open is_trunc is_conn
namespace is_conn
section
/- todo: reorder arguments and make some implicit -/
variables (A : Type*) (n : ℕ₋₂) (H : is_conn (n.+1) A)
include H
definition is_contr_ppi_match (P : A → Type*) (H : Πa, is_trunc (n.+1) (P a))
definition is_contr_ppi_match (P : A → Type*) (H2 : Πa, is_trunc (n.+1) (P a))
: is_contr (Π*(a : A), P a) :=
begin
apply is_contr.mk pt,
@ -997,7 +998,7 @@ namespace is_conn
(P : A → Type*) (H3 : Πa, is_trunc l (P a)) :
is_trunc k.+1 (Π*(a : A), P a) :=
begin
have H4 : Πa, is_trunc (n.+1+2+k) (P a), from λa, is_trunc_of_le (P a) H2,
have H4 : Πa, is_trunc (n.+1+2+k) (P a), from λa, is_trunc_of_le (P a) H2 _,
clear H2 H3, revert P H4,
induction k with k IH: intro P H4,
{ apply is_prop_of_imp_is_contr, intro f,
@ -1012,11 +1013,24 @@ namespace is_conn
definition is_trunc_pmap_of_is_conn (k l : ℕ₋₂) (B : Type*) (H2 : l ≤ n.+1+2+k)
(H3 : is_trunc l B) : is_trunc k.+1 (A →* B) :=
is_trunc_equiv_closed k.+1 (pppi_equiv_pmap A B)
(is_trunc_ppi_of_is_conn A n H k l H2 (λ a, B) _)
is_trunc_ppi_of_is_conn A n H k l H2 (λ a, B) _
end
open trunc_index algebra nat
definition is_trunc_ppi_of_is_conn_nat
(A : Type*) (n : ) (H : is_conn (n.-1) A) (k l : ) (H2 : l ≤ n + k)
(P : A → Type*) (H3 : Πa, is_trunc l (P a)) :
is_trunc k (Π*(a : A), P a) :=
begin
refine is_trunc_ppi_of_is_conn A (n.-2) H (k.-1) l _ P H3,
refine le.trans (of_nat_le_of_nat H2) (le_of_eq !sub_one_add_plus_two_sub_one⁻¹)
end
definition is_trunc_pmap_of_is_conn_nat (A : Type*) (n : ) (H : is_conn (n.-1) A) (k l : )
(B : Type*) (H2 : l ≤ n + k) (H3 : is_trunc l B) : is_trunc k (A →* B) :=
is_trunc_ppi_of_is_conn_nat A n H k l H2 (λ a, B) _
-- this is probably much easier to prove directly
definition is_trunc_ppi (A : Type*) (n k : ℕ₋₂) (H : n ≤ k) (P : A → Type*)
(H2 : Πa, is_trunc n (P a)) : is_trunc k (Π*(a : A), P a) :=

View file

@ -19,7 +19,7 @@ definition smash_prespectrum_fun {X X' : Type*} {Y Y' : prespectrum} (f : X →*
smap.mk (λn, smash_functor f (g n)) begin
intro n,
refine susp_to_loop_psquare _ _ _ _ _,
refine pvconcat (psquare_transpose (phinverse (smash_susp_natural f (g n)))) _,
refine pvconcat (ptranspose (phinverse (smash_susp_natural f (g n)))) _,
refine vconcat_phomotopy _ (smash_functor_split f (g (S n))),
refine phomotopy_vconcat (smash_functor_split f (susp_functor (g n))) _,
refine phconcat _ _,

View file

@ -107,8 +107,7 @@ begin induction p, exact id end
definition is_strunc_of_le {k l : } (E : spectrum) (H : k ≤ l)
: is_strunc k E → is_strunc l E :=
begin
intro T, intro n, exact is_trunc_of_le (E n)
(maxm2_monotone (algebra.add_le_add_right H n))
intro T, intro n, exact is_trunc_of_le (E n) (maxm2_monotone (algebra.add_le_add_right H n)) _
end
definition is_strunc_pequiv_closed {k : } {E F : spectrum} (H : Πn, E n ≃* F n)

View file

@ -28,7 +28,7 @@ section univ_subcat
definition is_univalent_domain_of_fully_faithful_embedding : is_univalent C :=
begin
intros,
apply homotopy_closed eq_equiv_iso_of_fully_faithful eq_equiv_iso_of_fully_faithful_homot
exact homotopy_closed eq_equiv_iso_of_fully_faithful eq_equiv_iso_of_fully_faithful_homot _
end
end univ_subcat
@ -216,7 +216,7 @@ end
(m, (i, o)) = (m', (i', o')) ≃ (m ~2 m') :=
begin
have is_set A, from pr1 H,
apply equiv_of_is_prop,
refine equiv_of_is_prop _ _ _ _,
{ intro p, exact apd100 (eq_pr1 p)},
{ intro p, apply prod_eq (eq_of_homotopy2 p),
apply prod_eq: esimp [Group_props] at *; esimp,