mathlib3 documentation

algebraic_topology.dold_kan.homotopies

Construction of homotopies for the Dold-Kan correspondence #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

TODO (@joelriou) continue adding the various files referenced below

(The general strategy of proof of the Dold-Kan correspondence is explained in equivalence.lean.)

The purpose of the files homotopies.lean, faces.lean, projections.lean and p_infty.lean is to construct an idempotent endomorphism P_infty : K[X] ⟶ K[X] of the alternating face map complex for each X : simplicial_object C when C is a preadditive category. In the case C is abelian, this P_infty shall be the projection on the normalized Moore subcomplex of K[X] associated to the decomposition of the complex K[X] as a direct sum of this normalized subcomplex and of the degenerate subcomplex.

In p_infty.lean, this endomorphism P_infty shall be obtained by passing to the limit idempotent endomorphisms P q for all (q : ℕ). These endomorphisms P q are defined by induction. The idea is to start from the identity endomorphism P 0 of K[X] and to ensure by induction that the q higher face maps (except $d_0$) vanish on the image of P q. Then, in a certain degree n, the image of P q for a big enough q will be contained in the normalized subcomplex. This construction is done in projections.lean.

It would be easy to define the P q degreewise (similarly as it is done in Simplicial Homotopy Theory by Goerrs-Jardine p. 149), but then we would have to prove that they are compatible with the differential (i.e. they are chain complex maps), and also that they are homotopic to the identity. These two verifications are quite technical. In order to reduce the number of such technical lemmas, the strategy that is followed here is to define a series of null homotopic maps Hσ q (attached to families of maps ) and use these in order to construct P q : the endomorphisms P q shall basically be obtained by altering the identity endomorphism by adding null homotopic maps, so that we get for free that they are morphisms of chain complexes and that they are homotopic to the identity. The most technical verifications that are needed about the null homotopic maps are obtained in faces.lean.

In this file homotopies.lean, we define the null homotopic maps Hσ q : K[X] ⟶ K[X], show that they are natural (see nat_trans_Hσ) and compatible the application of additive functors (see map_Hσ).

References #

@[reducible]

As we are using chain complexes indexed by , we shall need the relation c such c m n if and only if n+1=m.

Helper when we need some c.rel i j (i.e. complex_shape.down), e.g. c_mk n (n+1) rfl

This lemma is meant to be used with null_homotopic_map'_f_of_not_rel_left

The sequence of maps which gives the null homotopic maps that shall be in the inductive construction of the projections P q : K[X] ⟶ K[X]

Equations

The maps hσ' q n m hnm are compatible with the application of additive functors.

The null homotopic maps are compatible with the application of additive functors.