lean2/hott/homotopy/cellcomplex.hlean
Floris van Doorn ddef24223b make pointed suspensions, wedges and spheres the default (in contrast to the unpointed ones), remove sphere_index
All HITs which automatically have a point are pointed without a 'p' in front. HITs which do not automatically have a point do still have a p (e.g. pushout/ppushout).

There were a lot of annoyances with spheres being indexed by N_{-1} with almost no extra generality. We now index the spheres by nat, making sphere 0 = pbool.
2017-07-20 15:02:09 +01:00

49 lines
1.7 KiB
Text
Raw Permalink 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) 2015 Ulrik Buchholtz. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Ulrik Buchholtz
-/
import types.trunc homotopy.sphere hit.pushout
open eq is_trunc is_equiv nat equiv trunc prod pushout sigma unit pointed
-- where should this be?
definition family : Type := ΣX, X → Type
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 set
-- with realization the identity map
⟨Set , λ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 set of (succ n)-cells A,
and an attaching map f : A × sphere n → |X| -/
⟨Σ X : pr1 fdcc_family_n , Σ A : Set, 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 → Set :=
nat.cases (λA, A) (λn T, pr1 (pr2 T))
end cellcomplex