2016-03-06 23:09:32 +00:00
|
|
|
/-
|
|
|
|
Copyright (c) 2016 Ulrik Buchholtz and Egbert Rijke. All rights reserved.
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
2016-04-22 19:12:25 +00:00
|
|
|
Authors: Ulrik Buchholtz, Egbert Rijke, Floris van Doorn
|
2016-03-06 23:09:32 +00:00
|
|
|
|
|
|
|
The H-space structure on S¹ and the complex Hopf fibration
|
|
|
|
(the standard one).
|
|
|
|
-/
|
|
|
|
|
2016-04-22 19:12:25 +00:00
|
|
|
import .hopf .circle types.fin
|
2016-03-06 23:09:32 +00:00
|
|
|
|
2016-04-22 19:12:25 +00:00
|
|
|
open eq equiv is_equiv circle is_conn trunc is_trunc sphere susp pointed fiber sphere.ops function
|
2017-07-20 14:01:40 +00:00
|
|
|
join
|
2016-03-06 23:09:32 +00:00
|
|
|
|
|
|
|
namespace hopf
|
|
|
|
|
|
|
|
definition circle_h_space [instance] : h_space S¹ :=
|
|
|
|
⦃ h_space, one := base, mul := circle_mul,
|
|
|
|
one_mul := circle_base_mul, mul_one := circle_mul_base ⦄
|
|
|
|
|
|
|
|
definition circle_assoc (x y z : S¹) : (x * y) * z = x * (y * z) :=
|
|
|
|
begin
|
|
|
|
induction x,
|
|
|
|
{ reflexivity },
|
|
|
|
{ apply eq_pathover, induction y,
|
2016-11-23 22:59:13 +00:00
|
|
|
{ exact natural_square
|
2016-03-06 23:09:32 +00:00
|
|
|
(λa : S¹, ap (λb : S¹, b * z) (circle_mul_base a))
|
|
|
|
loop },
|
2017-07-20 14:01:40 +00:00
|
|
|
{ apply is_prop.elimo, apply is_trunc_square }}
|
2016-03-06 23:09:32 +00:00
|
|
|
end
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
definition complex_hopf' : S 3 → S 2 :=
|
2016-03-06 23:09:32 +00:00
|
|
|
begin
|
|
|
|
intro x, apply @sigma.pr1 (susp S¹) (hopf S¹),
|
2017-07-20 14:01:40 +00:00
|
|
|
apply inv (hopf.total S¹), exact (join_sphere 1 1)⁻¹ᵉ x
|
2016-03-06 23:09:32 +00:00
|
|
|
end
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
definition complex_hopf [constructor] : S 3 →* S 2 :=
|
|
|
|
proof pmap.mk complex_hopf' idp qed
|
2016-04-22 19:12:25 +00:00
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
definition pfiber_complex_hopf : pfiber complex_hopf ≃* S 1 :=
|
2016-04-22 19:12:25 +00:00
|
|
|
begin
|
|
|
|
fapply pequiv_of_equiv,
|
2017-07-20 14:01:40 +00:00
|
|
|
{ esimp, unfold [complex_hopf'],
|
2016-04-22 19:12:25 +00:00
|
|
|
refine fiber.equiv_precompose (sigma.pr1 ∘ (hopf.total S¹)⁻¹ᵉ)
|
2017-07-20 14:01:40 +00:00
|
|
|
(join_sphere 1 1)⁻¹ᵉ _ ⬝e _,
|
2016-04-22 19:12:25 +00:00
|
|
|
refine fiber.equiv_precompose _ (hopf.total S¹)⁻¹ᵉ _ ⬝e _,
|
2017-07-20 14:01:40 +00:00
|
|
|
apply fiber_pr1 },
|
|
|
|
{ reflexivity }
|
2016-04-22 19:12:25 +00:00
|
|
|
end
|
|
|
|
|
2016-03-06 23:09:32 +00:00
|
|
|
end hopf
|