applications: prove almost completely that S^3 and S^2 have the same high enough homotopy groups

There is one missing fact, which is that the equivalence between S^1 and the fiber of the hopf fibration respects the basepoint
This commit is contained in:
Floris van Doorn 2016-04-11 23:17:10 -04:00
parent 9a476fecfe
commit a716ef2108
2 changed files with 239 additions and 15 deletions

View file

@ -5,8 +5,9 @@ Authors: Floris van Doorn
-/
import .LES_of_homotopy_groups homotopy.connectedness homotopy.homotopy_group homotopy.join
open eq is_trunc pointed is_conn is_equiv fiber equiv trunc nat chain_complex prod fin algebra
group trunc_index function join pushout
homotopy.complex_hopf
open eq is_trunc pointed is_conn is_equiv fiber equiv trunc nat chain_complex fin algebra
group trunc_index function join pushout prod sigma sigma.ops
namespace nat
open sigma sum
@ -69,17 +70,132 @@ namespace is_conn
: square (ap f p) (ap g q) (h a b) (h a' b') :=
by induction p; induction q; exact hrfl
-- TODO: move
section
open sphere sphere_index
end is_conn
definition add_plus_one_minus_one (n : ℕ₋₁) : n +1+ -1 = n := idp
definition add_plus_one_succ (n m : ℕ₋₁) : n +1+ (m.+1) = (n +1+ m).+1 := idp
definition minus_one_add_plus_one (n : ℕ₋₁) : -1 +1+ n = n :=
begin induction n with n IH, reflexivity, exact ap succ IH end
definition succ_add_plus_one (n m : ℕ₋₁) : (n.+1) +1+ m = (n +1+ m).+1 :=
begin induction m with m IH, reflexivity, exact ap succ IH end
namespace sigma
definition ppr1 {A : Type*} {B : A → Type*} : (Σ*(x : A), B x) →* A :=
pmap.mk pr1 idp
definition ppr2 {A : Type*} (B : A → Type*) (v : (Σ*(x : A), B x)) : B (ppr1 v) :=
pr2 v
end sigma
namespace hopf
open sphere.ops susp circle sphere_index
-- definition complex_phopf : S. 3 →* S. 2 :=
-- proof pmap.mk complex_hopf idp qed
-- variable (A : Type)
-- variables [h_space A] [is_conn 0 A]
attribute hopf [unfold 4]
-- definition phopf (x : psusp A) : Type* :=
-- pointed.MK (hopf A x)
-- begin
-- induction x with a: esimp,
-- do 2 exact 1,
-- apply pathover_of_tr_eq, rewrite [↑hopf, elim_type_merid, ▸*, mul_one],
-- end
-- maybe define this as a composition of pointed maps?
definition complex_phopf [constructor] : S. 3 →* S. 2 :=
proof pmap.mk complex_hopf idp qed
definition fiber_pr1_fun {A : Type} {B : A → Type} {a : A} (b : B a)
: fiber_pr1 B a (fiber.mk ⟨a, b⟩ idp) = b :=
idp
--TODO: in fiber.equiv_precompose, make f explicit
open sphere sphere.ops
definition add_plus_one_of_nat (n m : ) : (n +1+ m) = sphere_index.of_nat (n + m + 1) :=
begin
induction m with m IH,
{ reflexivity },
{ exact ap succ IH}
end
end is_conn
-- definition pjoin_pspheres (n m : ) : join (S. n) (S. m) ≃ (S. (n + m + 1)) :=
-- join.spheres n m ⬝e begin esimp, apply equiv_of_eq, apply ap S, apply add_plus_one_of_nat end
-- set_option pp.all true
-- definition pjoin_spheres_inv_base (n m : )
-- : (join.spheres 1 1)⁻¹ (cast proof idp qed (@sphere.base 3)) = inl base :=
-- begin
-- exact sorry
-- end
definition pfiber_complex_phopf : pfiber complex_phopf ≃* S. 1 :=
begin
fapply pequiv_of_equiv,
{ esimp, unfold [complex_hopf],
refine @fiber.equiv_precompose _ _ (sigma.pr1 ∘ (hopf.total S¹)⁻¹ᵉ) _ _
(join.spheres 1 1)⁻¹ᵉ _ ⬝e _,
refine fiber.equiv_precompose (hopf.total S¹)⁻¹ᵉ ⬝e _,
apply fiber_pr1},
{ exact sorry}
end
open int
definition one_le_succ (n : ) : 1 ≤ succ n :=
nat.succ_le_succ !zero_le
definition π2S2 : πg[1+1] (S. 2) = g :=
begin
refine _ ⬝ fundamental_group_of_circle,
refine _ ⬝ ap (λx, π₁ x) (eq_of_pequiv pfiber_complex_phopf),
fapply Group_eq,
{ fapply equiv.mk,
{ exact cc_to_fn (LES_of_homotopy_groups complex_phopf) (1, 2)},
{ refine @is_equiv_of_trivial _
_ _
(is_exact_LES_of_homotopy_groups _ (1, 1))
(is_exact_LES_of_homotopy_groups _ (1, 2))
_
_
(@pgroup_of_group _ (group_LES_of_homotopy_groups complex_phopf _ _) idp)
(@pgroup_of_group _ (group_LES_of_homotopy_groups complex_phopf _ _) idp)
_,
{ rewrite [LES_of_homotopy_groups_1, ▸*],
have H : 1 ≤[] 2, from !one_le_succ,
apply trivial_homotopy_group_of_is_conn, exact H, rexact is_conn_psphere 3},
{ refine tr_rev (λx, is_contr (ptrunctype._trans_of_to_pType x))
(LES_of_homotopy_groups_1 complex_phopf 2) _,
apply trivial_homotopy_group_of_is_conn, apply le.refl, rexact is_conn_psphere 3},
{ exact homomorphism.struct (homomorphism_LES_of_homotopy_groups_fun _ (0, 2))}}},
{ exact homomorphism.struct (homomorphism_LES_of_homotopy_groups_fun _ (0, 2))}
end
open circle
definition πnS3_eq_πnS2 (n : ) : πg[n+2 +1] (S. 3) = πg[n+2 +1] (S. 2) :=
begin
fapply Group_eq,
{ fapply equiv.mk,
{ exact cc_to_fn (LES_of_homotopy_groups complex_phopf) (n+3, 0)},
{ have H : is_trunc 1 (pfiber complex_phopf),
from @(is_trunc_equiv_closed_rev _ pfiber_complex_phopf) is_trunc_circle,
refine @is_equiv_of_trivial _
_ _
(is_exact_LES_of_homotopy_groups _ (n+2, 2))
(is_exact_LES_of_homotopy_groups _ (n+3, 0))
_
_
(@pgroup_of_group _ (group_LES_of_homotopy_groups complex_phopf _ _) idp)
(@pgroup_of_group _ (group_LES_of_homotopy_groups complex_phopf _ _) idp)
_,
{ rewrite [▸*, LES_of_homotopy_groups_2 _ (n +[] 2)],
have H : 1 ≤[] n + 1, from !one_le_succ,
apply trivial_homotopy_group_of_is_trunc _ _ _ H},
{ refine tr_rev (λx, is_contr (ptrunctype._trans_of_to_pType x))
(LES_of_homotopy_groups_2 complex_phopf _) _,
have H : 1 ≤[] n + 2, from !one_le_succ,
apply trivial_homotopy_group_of_is_trunc _ _ _ H},
{ exact homomorphism.struct (homomorphism_LES_of_homotopy_groups_fun _ (n+2, 0))}}},
{ exact homomorphism.struct (homomorphism_LES_of_homotopy_groups_fun _ (n+2, 0))}
end
end hopf

View file

@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn
-/
import homotopy.wedge algebra.homotopy_group homotopy.sphere types.nat
import homotopy.wedge homotopy.homotopy_group homotopy.circle .LES_applications
open eq is_conn is_trunc pointed susp nat pi equiv is_equiv trunc fiber trunc_index
@ -295,9 +295,9 @@ namespace sphere
pcast (phomotopy_group_succ_in2 A (succ n)) g * pcast (phomotopy_group_succ_in2 A (succ n)) h :=
begin
induction g with p, induction h with q, esimp,
rewrite [-+tr_eq_cast_ap, ↑phomotopy_group_succ_in2, -+tr_compose],
rewrite [-+ tr_eq_cast_ap, ↑phomotopy_group_succ_in2, -+ tr_compose],
refine ap (transport _ _) !tr_mul_tr ⬝ _,
rewrite [+trunc_transport],
rewrite [+ trunc_transport],
apply ap tr, apply loop_space_succ_eq_in_concat,
end
@ -314,4 +314,112 @@ namespace sphere
}
end
theorem mul_two {A : Type} [semiring A] (a : A) : a * 2 = a + a :=
by rewrite [-one_add_one_eq_two, left_distrib, +mul_one]
theorem two_mul {A : Type} [semiring A] (a : A) : 2 * a = a + a :=
by rewrite [-one_add_one_eq_two, right_distrib, +one_mul]
definition two_le_succ_succ (n : ) : 2 ≤ succ (succ n) :=
nat.succ_le_succ (nat.succ_le_succ !zero_le)
open int circle hopf
definition πnSn (n : ) : πg[n+1] (S. (succ n)) = g :=
begin
cases n with n IH,
{ exact fundamental_group_of_circle},
{ induction n with n IH,
{ exact π2S2},
{ refine _ ⬝ IH, apply stability_eq,
calc
succ n + 3 = succ (succ n) + 2 : !succ_add⁻¹
... ≤ succ (succ n) + (succ (succ n)) : add_le_add_left !two_le_succ_succ
... = 2 * (succ (succ n)) : !two_mul⁻¹ }}
end
attribute group_integers [constructor]
theorem not_is_trunc_sphere (n : ) : ¬is_trunc n (S. (succ n)) :=
begin
intro H,
note H2 := trivial_homotopy_group_of_is_trunc (S. (succ n)) n n !le.refl,
rewrite [πnSn at H2, ▸* at H2],
have H3 : (0 : ) ≠ (1 : ), from dec_star,
apply H3,
apply is_prop.elim,
end
definition transport11 {A B : Type} (P : A → B → Type) {a a' : A} {b b' : B}
(p : a = a') (q : b = b') (z : P a b) : P a' b' :=
transport (P a') q (p ▸ z)
section
open sphere_index
definition add_plus_one_minus_one (n : ℕ₋₁) : n +1+ -1 = n := idp
definition add_plus_one_succ (n m : ℕ₋₁) : n +1+ (m.+1) = (n +1+ m).+1 := idp
definition minus_one_add_plus_one (n : ℕ₋₁) : -1 +1+ n = n :=
begin induction n with n IH, reflexivity, exact ap succ IH end
definition succ_add_plus_one (n m : ℕ₋₁) : (n.+1) +1+ m = (n +1+ m).+1 :=
begin induction m with m IH, reflexivity, exact ap succ IH end
definition nat_of_sphere_index : ℕ₋₁ → :=
sphere_index.rec 0 (λx, succ)
definition trunc_index_of_nat_of_sphere_index (n : ℕ₋₁)
: trunc_index.of_nat (nat_of_sphere_index n) = (of_sphere_index n).+1 :=
begin
induction n with n IH,
{ reflexivity},
{ exact ap succ IH}
end
definition sphere_index_of_nat_of_sphere_index (n : ℕ₋₁)
: sphere_index.of_nat (nat_of_sphere_index n) = n.+1 :=
begin
induction n with n IH,
{ reflexivity},
{ exact ap succ IH}
end
definition of_sphere_index_succ (n : ℕ₋₁)
: of_sphere_index (n.+1) = (of_sphere_index n).+1 :=
begin
induction n with n IH,
{ reflexivity},
{ exact ap succ IH}
end
definition sphere_index.of_nat_succ (n : )
: sphere_index.of_nat (succ n) = (sphere_index.of_nat n).+1 :=
begin
induction n with n IH,
{ reflexivity},
{ exact ap succ IH}
end
definition nat_of_sphere_index_succ (n : ℕ₋₁)
: nat_of_sphere_index (n.+1) = succ (nat_of_sphere_index n) :=
begin
induction n with n IH,
{ reflexivity},
{ exact ap succ IH}
end
definition not_is_trunc_sphere' (n : ℕ₋₁) : ¬is_trunc n (S (n.+1)) :=
begin
cases n with n,
{ esimp [sphere.ops.S, sphere], intro H,
have H2 : is_prop bool, from @(is_trunc_equiv_closed -1 sphere_equiv_bool) H,
have H3 : bool.tt ≠ bool.ff, from dec_star, apply H3, apply is_prop.elim},
{ intro H, apply not_is_trunc_sphere (nat_of_sphere_index n),
rewrite [▸*, trunc_index_of_nat_of_sphere_index, -nat_of_sphere_index_succ,
sphere_index_of_nat_of_sphere_index],
exact H}
end
end
definition π3S2 : πg[2+1] (S. 2) = g :=
(πnS3_eq_πnS2 0)⁻¹ ⬝ πnSn 2
end sphere