lean2/hott/homotopy/homotopy_group.hlean
Ulrik Buchholtz bd9e47c82c feat(hott): functoriality of pushout; connectedness in is_conn namespace
other changes:
- move result about connectedness of susp to homotopy.susp
- improved definition of circle multiplication
- improved the interface to join
2016-03-23 09:22:55 -07:00

51 lines
1.6 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
Copyright (c) 2016 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn, Clive Newstead
-/
import algebra.homotopy_group .sphere
open eq is_trunc trunc_index pointed algebra trunc nat is_conn fiber pointed
namespace is_trunc
-- Lemma 8.3.1
theorem trivial_homotopy_group_of_is_trunc (A : Type*) (n k : ) [is_trunc n A] (H : n ≤ k)
: is_contr (πg[k+1] A) :=
begin
apply is_trunc_trunc_of_is_trunc,
apply is_contr_loop_of_is_trunc,
apply @is_trunc_of_le A n _,
rewrite [succ_sub_two_succ k],
exact of_nat_le_of_nat H,
end
-- Lemma 8.3.2
theorem trivial_homotopy_group_of_is_conn (A : Type*) {k n : } (H : k ≤ n) [is_conn n A]
: is_contr (π[k] A) :=
begin
have H3 : is_contr (ptrunc k A), from is_conn_of_le A (of_nat_le_of_nat H),
have H4 : is_contr (Ω[k](ptrunc k A)), from !is_trunc_loop_of_is_trunc,
apply is_trunc_equiv_closed_rev,
{ apply equiv_of_pequiv (phomotopy_group_pequiv_loop_ptrunc k A)}
end
-- Corollary 8.3.3
section
open sphere sphere.ops sphere_index
theorem homotopy_group_sphere_le (n k : ) (H : k < n) : is_contr (π[k] (S. n)) :=
begin
cases n with n,
{ exfalso, apply not_lt_zero, exact H},
{ have H2 : k ≤ n, from le_of_lt_succ H,
apply @(trivial_homotopy_group_of_is_conn _ H2) }
end
end
theorem is_contr_HG_fiber_of_is_connected {A B : Type*} (k n : ) (f : A →* B)
[H : is_conn_fun n f] (H2 : k ≤ n) : is_contr (π[k] (pfiber f)) :=
@(trivial_homotopy_group_of_is_conn (pfiber f) H2) (H pt)
end is_trunc