Preliminaries for Nöbeling's theorem #
This file constructs basic objects and results concerning them that are needed in the proof of
Nöbeling's theorem, which is in Mathlib.Topology.Category.Profinite.Nobeling.Induction
.
See the section docstrings for more information.
Proof idea #
We follow the proof of theorem 5.4 in [Sch19], in which the idea is to embed S
in
a product of I
copies of Bool
for some sufficiently large I
, and then to choose a
well-ordering on I
and use ordinal induction over that well-order. Here we can let I
be
the set of clopen subsets of S
since S
is totally separated.
The above means it suffices to prove the following statement: For a closed subset C
of I → Bool
,
the ℤ
-module LocallyConstant C ℤ
is free.
For i : I
, let e C i : LocallyConstant C ℤ
denote the map fun f ↦ (if f.val i then 1 else 0)
.
The basis will consist of products e C iᵣ * ⋯ * e C i₁
with iᵣ > ⋯ > i₁
which cannot be written
as linear combinations of lexicographically smaller products. We call this set GoodProducts C
.
What is proved by ordinal induction (in Mathlib.Topology.Category.Profinite.Nobeling.ZeroLimit
and
Mathlib.Topology.Category.Profinite.Nobeling.Successor
) is that this set is linearly independent.
The fact that it spans is proved directly in Mathlib.Topology.Category.Profinite.Nobeling.Span
.
References #
- [Sch19], Theorem 5.4.
Projection maps #
The purpose of this section is twofold.
Firstly, in the proof that the set GoodProducts C
spans the whole module LocallyConstant C ℤ
,
we need to project C
down to finite discrete subsets and write C
as a cofiltered limit of those.
Secondly, in the inductive argument, we need to project C
down to "smaller" sets satisfying the
inductive hypothesis.
In this section we define the relevant projection maps and prove some compatibility results.
Main definitions #
Let
J : I → Prop
. ThenProj J : (I → Bool) → (I → Bool)
is the projection mapping everything that satisfiesJ i
to itself, and everything else tofalse
.The image of
C
underProj J
is denotedπ C J
and the corresponding mapC → π C J
is calledProjRestrict
. IfJ
impliesK
we have a mapProjRestricts : π C K → π C J
.spanCone_isLimit
establishes that whenC
is compact, it can be written as a limit of its images under the mapsProj (· ∈ s)
wheres : Finset I
.
A variant of ProjRestrict
with domain of the form π C K
Equations
Instances For
The objectwise map in the isomorphism spanFunctor ≅ Profinite.indexFunctor
.
Equations
- Profinite.NobelingProof.iso_map C J = { toFun := fun (x : ↑(Profinite.NobelingProof.π C J)) => ⟨fun (i : { i : I // J i }) => ↑x ↑i, ⋯⟩, continuous_toFun := ⋯ }
Instances For
For a given compact subset C
of I → Bool
, spanFunctor
is the functor from the poset of finsets
of I
to Profinite
, sending a finite subset set J
to the image of C
under the projection
Proj J
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The limit cone on spanFunctor
with point C
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Defining the basis #
Our proposed basis consists of products e C iᵣ * ⋯ * e C i₁
with iᵣ > ⋯ > i₁
which cannot be
written as linear combinations of lexicographically smaller products. See below for the definition
of e
.
Main definitions #
For
i : I
, we lete C i : LocallyConstant C ℤ
denote the mapfun f ↦ (if f.val i then 1 else 0)
.Products I
is the type of lists of decreasing elements ofI
, so a typical element is[i₁, i₂,..., iᵣ]
withi₁ > i₂ > ... > iᵣ
.Products.eval C
is theC
-evaluation of a list. It takes a term[i₁, i₂,..., iᵣ] : Products I
and returns the actual producte C i₁ ··· e C iᵣ : LocallyConstant C ℤ
.GoodProducts C
is the set ofProducts I
such that theirC
-evaluation cannot be written as a linear combination of evaluations of lexicographically smaller lists.
Main results #
Products.evalFacProp
andProducts.evalFacProps
establish the fact thatProducts.eval
interacts nicely with the projection maps from the previous section.GoodProducts.span_iff_products
: the good products spanLocallyConstant C ℤ
iff all the products spanLocallyConstant C ℤ
.
Products I
is the type of lists of decreasing elements of I
, so a typical element is
[i₁, i₂, ...]
with i₁ > i₂ > ...
. We order Products I
lexicographically, so [] < [i₁, ...]
,
and [i₁, i₂, ...] < [j₁, j₂, ...]
if either i₁ < j₁
, or i₁ = j₁
and [i₂, ...] < [j₂, ...]
.
Terms m = [i₁, i₂, ..., iᵣ]
of this type will be used to represent products of the form
e C i₁ ··· e C iᵣ : LocallyConstant C ℤ
. The function associated to m
is m.eval
.
Equations
- Profinite.NobelingProof.Products I = { l : List I // List.Chain' (fun (x1 x2 : I) => x1 > x2) l }
Instances For
Equations
- Profinite.NobelingProof.Products.instLinearOrder = inferInstanceAs (LinearOrder { l : List I // List.Chain' (fun (x1 x2 : I) => x1 > x2) l })
The evaluation e C i₁ ··· e C iᵣ : C → ℤ
of a formal product [i₁, i₂, ..., iᵣ]
.
Equations
Instances For
The predicate on products which we prove picks out a basis of LocallyConstant C ℤ
. We call such a
product "good".
Equations
- One or more equations did not get rendered due to their size.
Instances For
The set of good products.
Equations
Instances For
Evaluation of good products.
Equations
Instances For
The image of the good products in the module LocallyConstant C ℤ
.
Equations
Instances For
The type of good products is equivalent to its image.
Equations
Instances For
The good products span LocallyConstant C ℤ
if and only all the products do.
Relating elements of the well-order I
with ordinals #
We choose a well-ordering on I
. This amounts to regarding I
as an ordinal, and as such it
can be regarded as the set of all strictly smaller ordinals, allowing to apply ordinal induction.
Main definitions #
ord I i
is the termi
ofI
regarded as an ordinal.term I ho
is a sufficiently small ordinal regarded as a term ofI
.contained C o
is a predicate saying thatC
is "small" enough in relation to the ordinalo
to satisfy the inductive hypothesis.P I
is the predicate on ordinals about linear independence of good products, which the rest of this file is spent on proving by induction.
A term of I
regarded as an ordinal.
Equations
- Profinite.NobelingProof.ord I i = (Ordinal.typein fun (x1 x2 : I) => x1 < x2).toRelEmbedding i
Instances For
An ordinal regarded as a term of I
.
Equations
- Profinite.NobelingProof.term I ho = (Ordinal.enum fun (x1 x2 : I) => x1 < x2) ⟨o, ho⟩
Instances For
A predicate saying that C
is "small" enough to satisfy the inductive hypothesis.
Equations
- Profinite.NobelingProof.contained C o = ∀ f ∈ C, ∀ (i : I), f i = true → Profinite.NobelingProof.ord I i < o
Instances For
The predicate on ordinals which we prove by induction, see GoodProducts.P0
,
GoodProducts.Plimit
and GoodProducts.linearIndependentAux
in the section Induction
below
Equations
- One or more equations did not get rendered due to their size.
Instances For
ℤ
-linear maps induced by projections #
We define injective ℤ
-linear maps between modules of the form LocallyConstant C ℤ
induced by
precomposition with the projections defined in the section Projections
.
Main definitions #
πs
andπs'
are theℤ
-linear maps corresponding toProjRestrict
andProjRestricts
respectively.
Main result #
- We prove that
πs
andπs'
interact well withProducts.eval
and the main application is the theoremisGood_mono
which says that the propertyisGood
is "monotone" on ordinals.
The ℤ
-linear map induced by precomposition of the projection C → π C (ord I · < o)
.
Equations
- Profinite.NobelingProof.πs C o = LocallyConstant.comapₗ ℤ { toFun := Profinite.NobelingProof.ProjRestrict C fun (x : I) => Profinite.NobelingProof.ord I x < o, continuous_toFun := ⋯ }
Instances For
The ℤ
-linear map induced by precomposition of the projection
π C (ord I · < o₂) → π C (ord I · < o₁)
for o₁ ≤ o₂
.
Equations
- Profinite.NobelingProof.πs' C h = LocallyConstant.comapₗ ℤ { toFun := Profinite.NobelingProof.ProjRestricts C ⋯, continuous_toFun := ⋯ }
Instances For
If l
is good w.r.t. π C (ord I · < o₁)
and o₁ ≤ o₂
, then it is good w.r.t.
π C (ord I · < o₂)