refactor(hott): move homotopy hits to new homotopy folder

This commit is contained in:
Ulrik Buchholtz 2015-09-13 14:58:11 -04:00
parent 97cf839665
commit 384a366e0f
15 changed files with 264 additions and 10 deletions

View file

@ -137,6 +137,8 @@ Chapter 7: Homotopy n-types
Chapter 8: Homotopy theory
---------
Unless otherwise noted, the files are in the folder [homotopy](homotopy/homotopy.md)
- 8.1 (π_1(S^1)): [hit.circle](hit/circle.hlean) (only one of the proofs)
- 8.2 (Connectedness of suspensions): not formalized
- 8.3 (πk≤n of an n-connected space and π_{k<n}(S^n)): not formalized

View file

@ -223,6 +223,16 @@ namespace function
variable {f}
-- Lemma 3.11.7
definition is_contr_retract {A B : Type} (r : A → B) [H : is_retraction r]
: is_contr A → is_contr B :=
begin
intro CA,
apply is_contr.mk (r (center A)),
intro b,
exact ap r (center_eq (is_retraction.sect r b)) ⬝ (is_retraction.right_inverse r b)
end
local attribute is_hprop_is_retraction_prod_is_section [instance]
definition is_retraction_prod_is_section_equiv_is_equiv
: (is_retraction f × is_section f) ≃ is_equiv f :=

View file

@ -8,6 +8,10 @@ primitive are n-truncation and quotients. These are defined in
[init.hit](../init/hit.hlean) and they have definitional computation
rules on the point constructors.
Here we find hits related to the basic structure theory of HoTT. The
hits related to homotopy theory are defined in
[homotopy](../homotopy/homotopy.md).
Files in this folder:
* [quotient](quotient.hlean) (quotients, primitive)
@ -15,9 +19,4 @@ Files in this folder:
* [colimit](colimit.hlean) (Colimits of arbitrary diagrams and sequential colimits, defined using quotients)
* [pushout](pushout.hlean) (Pushouts, defined using quotients)
* [coeq](coeq.hlean) (Co-equalizers, defined using quotients)
* [cylinder](cylinder.hlean) (Mapping cylinders, defined using quotients)
* [set_quotient](set_quotient.hlean) (Set-quotients, defined using quotients and set-truncation)
* [suspension](suspension.hlean) (Suspensions, defined using pushouts)
* [sphere](sphere.hlean) (Higher spheres, defined recursively using suspensions)
* [circle](circle.hlean) (defined as sphere 1)
* [interval](interval.hlean) (defined as the suspension of unit)

View file

@ -0,0 +1,57 @@
/-
Copyright (c) 2015 Ulrik Buchholtz. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Ulrik Buchholtz
-/
import types.trunc hit.sphere hit.pushout
open eq is_trunc is_equiv nat equiv trunc prod pushout sigma sphere_index unit
-- where should this be?
definition family : Type := ΣX, X → Type
-- this should be in init!
namespace nat
definition cases {M : → Type} (mz : M zero) (ms : Πn, M (succ n)) : Πn, M n :=
nat.rec mz (λn dummy, ms n)
end nat
namespace cellcomplex
/-
define by recursion on
both the type of fdccs of dimension n
and the realization map fdcc n → Type
in other words, we define a function
fdcc : → family
an alternative to the approach here (perhaps necessary) is to
define relative cell complexes relative to a type A, and then use
spherical indexing, so a -1-dimensional relative cell complex is
just star : unit with realization A
-/
definition fdcc_family [reducible] : → family :=
nat.rec
-- a zero-dimensional cell complex is just an hset
-- with realization the identity map
⟨hset , λA, trunctype.carrier A⟩
(λn fdcc_family_n, -- sigma.rec (λ fdcc_n realize_n,
/- a (succ n)-dimensional cell complex is a triple of
an n-dimensional cell complex X, an hset of (succ n)-cells A,
and an attaching map f : A × sphere n → |X| -/
⟨Σ X : pr1 fdcc_family_n , Σ A : hset, A × sphere n → pr2 fdcc_family_n X ,
/- the realization of such is the pushout of f with
canonical map A × sphere n → unit -/
sigma.rec (λX , sigma.rec (λA f, pushout (λx , star) f))
⟩)
definition fdcc (n : ) : Type := pr1 (fdcc_family n)
definition cell : Πn, fdcc n → hset :=
nat.cases (λA, A) (λn T, pr1 (pr2 T))
end cellcomplex

View file

@ -0,0 +1,76 @@
/-
Copyright (c) 2015 Ulrik Buchholtz. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Ulrik Buchholtz
-/
import types.trunc
open eq is_trunc is_equiv nat equiv trunc function
-- TODO(Ulrik) move this to somewhere else (cannot be sigma b/c dep. on fiber)
namespace sigma
variables {A : Type} {P Q : A → Type}
definition total [reducible] (f : Πa, P a → Q a) : (Σa, P a) → (Σa, Q a) :=
sigma.rec (λa p, ⟨a , f a p⟩)
-- Theorem 4.7.6
--definition fiber_total_equiv (f : Πa, P a → Q a) {a : A} (q : Q a)
-- : fiber (total f) ⟨a , q⟩ ≃ fiber (f a) q :=
--sorry
end sigma
-- TODO(Ulrik) move this to somewhere else (cannot be unit b/c dep. on fiber)
namespace unit
definition fiber_star_equiv (A : Type) : fiber (λx : A, star) star ≃ A :=
begin
fapply equiv.MK,
{ intro f, cases f with a H, exact a },
{ intro a, apply fiber.mk a, reflexivity },
{ intro a, reflexivity },
{ intro f, cases f with a H, change fiber.mk a (refl star) = fiber.mk a H,
rewrite [is_hset.elim H (refl star)] }
end
end unit
namespace homotopy
definition is_conn (n : trunc_index) (A : Type) : Type :=
is_contr (trunc n A)
definition is_conn_map (n : trunc_index) {A B : Type} (f : A → B) : Type :=
Πb : B, is_conn n (fiber f b)
definition is_conn_of_map_to_unit (n : trunc_index) (A : Type)
: is_conn_map n (λx : A, unit.star) → is_conn n A :=
begin
intro H, unfold is_conn_map at H,
rewrite [-(ua (unit.fiber_star_equiv A))],
exact (H unit.star)
end
-- Lemma 7.5.2
definition minus_one_conn_of_surjective {A B : Type} (f : A → B)
: is_surjective f → is_conn_map -1 f :=
begin
intro H, intro b,
exact @is_contr_of_inhabited_hprop (∥fiber f b∥) (is_trunc_trunc -1 (fiber f b)) (H b),
end
definition is_surjection_of_minus_one_conn {A B : Type} (f : A → B)
: is_conn_map -1 f → is_surjective f :=
begin
intro H, intro b,
exact @center (∥fiber f b∥) (H b),
end
definition merely_of_minus_one_conn {A : Type} : is_conn -1 A → ∥ A ∥ :=
λH, @center (∥A∥) H
definition minus_one_conn_of_merely {A : Type} : ∥A∥ → is_conn -1 A :=
@is_contr_of_inhabited_hprop (∥A∥) (is_trunc_trunc -1 A)
end homotopy

View file

@ -6,7 +6,7 @@ Authors: Floris van Doorn
Declaration of mapping cylinders
-/
import .quotient
import hit.quotient
open quotient eq sum equiv equiv.ops

17
hott/homotopy/homotopy.md Normal file
View file

@ -0,0 +1,17 @@
homotopy
========
Development of Homotopy Theory, including basic hits (higher inductive
types; see also [hit](../hit/hit.md)). The following files are in this
folder (sorted such that files only import previous files).
* [cylinder](cylinder.hlean) (Mapping cylinders, defined using quotients)
* [susp](susp.hlean) (Suspensions, defined using pushouts)
* [red_susp](red_susp.hlean) (Reduced suspensions)
* [sphere](sphere.hlean) (Higher spheres, defined recursively using suspensions)
* [circle](circle.hlean) (defined as sphere 1)
* [torus](torus.hlean) (defined as a two-quotient)
* [interval](interval.hlean) (defined as the suspension of unit)
* [cellcomplex](cellcomplex.hlean) (general cell complexes)
* [connectedness](connectedness.hlean)

View file

@ -6,7 +6,7 @@ Authors: Floris van Doorn
Declaration of the reduced suspension
-/
import .two_quotient types.pointed algebra.e_closure
import hit.two_quotient types.pointed algebra.e_closure
open simple_two_quotient eq unit pointed e_closure

View file

@ -13,7 +13,7 @@ open eq nat susp bool is_trunc unit pointed
/-
We can define spheres with the following possible indices:
- trunc_index (defining S^-2 = S^-1 = empty)
- nat (forgetting that S^1 = empty)
- nat (forgetting that S^-1 = empty)
- nat, but counting wrong (S^0 = empty, S^1 = bool, ...)
- some new type "integers >= -1"
We choose the last option here.

View file

@ -6,7 +6,7 @@ Authors: Floris van Doorn
Declaration of suspension
-/
import .pushout types.pointed cubical.square
import hit.pushout types.pointed cubical.square
open pushout unit eq equiv equiv.ops

View file

@ -7,7 +7,7 @@ Authors: Floris van Doorn
Declaration of the torus
-/
import .two_quotient
import hit.two_quotient
open two_quotient eq bool unit relation

92
hott/types/arrow_2.hlean Normal file
View file

@ -0,0 +1,92 @@
/-
Copyright (c) 2015 Ulrik Buchholtz. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Ulrik Buchholtz
-/
import ..function
open eq is_equiv function
namespace arrow
structure arrow :=
(dom : Type)
(cod : Type)
(arrow : dom → cod)
abbreviation dom [unfold 2] := @arrow.dom
abbreviation cod [unfold 2] := @arrow.cod
structure morphism (A B : Type) :=
(mor : A → B)
definition morphism_of_arrow [coercion] (f : arrow) : morphism (dom f) (cod f) :=
morphism.mk (arrow.arrow f)
attribute morphism.mor [coercion]
structure arrow_hom (f g : arrow) :=
(on_dom : dom f → dom g)
(on_cod : cod f → cod g)
(commute : Π(x : dom f), g (on_dom x) = on_cod (f x))
abbreviation on_dom [unfold 2] := @arrow_hom.on_dom
abbreviation on_cod [unfold 2] := @arrow_hom.on_cod
abbreviation commute [unfold 2] := @arrow_hom.commute
variables {f g : arrow}
definition on_fiber [reducible] (r : arrow_hom f g) (y : cod f)
: fiber f y → fiber g (on_cod r y) :=
fiber.rec (λx p, fiber.mk (on_dom r x) (commute r x ⬝ ap (on_cod r) p))
structure is_retraction [class] (r : arrow_hom f g) : Type :=
(sect : arrow_hom g f)
(right_inverse_dom : Π(a : dom g), on_dom r (on_dom sect a) = a)
(right_inverse_cod : Π(b : cod g), on_cod r (on_cod sect b) = b)
(cohere : Π(a : dom g), commute r (on_dom sect a) ⬝ ap (on_cod r) (commute sect a)
= ap g (right_inverse_dom a) ⬝ (right_inverse_cod (g a))⁻¹)
definition retraction_on_fiber [reducible] (r : arrow_hom f g) [H : is_retraction r]
(b : cod g) : fiber f (on_cod (is_retraction.sect r) b) → fiber g b :=
fiber.rec (λx q, fiber.mk (on_dom r x) (commute r x ⬝ ap (on_cod r) q ⬝ is_retraction.right_inverse_cod r b))
definition retraction_on_fiber_right_inverse' (r : arrow_hom f g) [H : is_retraction r]
(a : dom g) (b : cod g) (p : g a = b)
: retraction_on_fiber r b (on_fiber (is_retraction.sect r) b (fiber.mk a p)) = fiber.mk a p :=
begin
induction p, unfold on_fiber, unfold retraction_on_fiber,
apply @fiber.fiber_eq _ _ g (g a)
(fiber.mk
(on_dom r (on_dom (is_retraction.sect r) a))
(commute r (on_dom (is_retraction.sect r) a)
⬝ ap (on_cod r) (commute (is_retraction.sect r) a)
⬝ is_retraction.right_inverse_cod r (g a)))
(fiber.mk a (refl (g a)))
(is_retraction.right_inverse_dom r a), -- everything but this field should be inferred
unfold fiber.point_eq,
rewrite [is_retraction.cohere r a],
apply inv_con_cancel_right
end
definition retraction_on_fiber_right_inverse (r : arrow_hom f g) [H : is_retraction r]
: Π(b : cod g), Π(z : fiber g b), retraction_on_fiber r b (on_fiber (is_retraction.sect r) b z) = z :=
λb, fiber.rec (λa p, retraction_on_fiber_right_inverse' r a b p)
-- Lemma 4.7.3
definition retraction_on_fiber_is_retraction [instance] (r : arrow_hom f g) [H : is_retraction r]
(b : cod g) : _root_.is_retraction (retraction_on_fiber r b) :=
_root_.is_retraction.mk (on_fiber (is_retraction.sect r) b) (retraction_on_fiber_right_inverse r b)
-- Theorem 4.7.4
definition retract_of_equivalence_is_equivalence (r : arrow_hom f g) [H : is_retraction r]
[K : is_equiv f] : is_equiv g :=
begin
apply @is_equiv_of_is_contr_fun _ _ g,
intro b,
apply is_contr_retract (retraction_on_fiber r b),
exact is_contr_fun_of_is_equiv f (on_cod (is_retraction.sect r) b)
end
end arrow

View file

@ -13,6 +13,7 @@ Types (not necessarily HoTT-related):
* [pi](pi.hlean)
* [arrow](arrow.hlean)
* [W](W.hlean): W-types (not loaded by default)
* [arrow_2](arrow_2.hlean): alternative development of properties of arrows
HoTT types