Limits of inverse systems indexed by well-ordered types #
Given a functor F : Jᵒᵖ ⥤ Type v
where J
is a well-ordered type,
we introduce a structure F.WellOrderInductionData
which allows
to show that the map F.sections → F.obj (op ⊥)
is surjective.
The data and properties in F.WellOrderInductionData
consist of a
section to the maps F.obj (op (Order.succ j)) → F.obj (op j)
when j
is not maximal,
and, when j
is limit, a section to the canonical map from F.obj (op j)
to the type of compatible families of elements in F.obj (op i)
for i < j
.
In other words, from val₀ : F.obj (op ⊥)
, a term d : F.WellOrderInductionData
allows the construction, by transfinite induction, of a section of F
which restricts to val₀
.
Given a functor F : Jᵒᵖ ⥤ Type v
where J
is a well-ordered type, this data
allows to construct a section of F
from an element in F.obj (op ⊥)
,
see WellOrderInductionData.sectionsMk
.
- map_succ (j : J) (hj : ¬IsMax j) (x : F.obj (Opposite.op j)) : F.map (CategoryTheory.homOfLE ⋯).op (self.succ j hj x) = x
- lift (j : J) (hj : Order.IsSuccLimit j) (x : ↑(⋯.functor.op.comp F).sections) : F.obj (Opposite.op j)
When
j
is a limit element, andx
is a compatible family of elements inF.obj (op i)
for alli < j
, this is a lifting toF.obj (op j)
. - map_lift (j : J) (hj : Order.IsSuccLimit j) (x : ↑(⋯.functor.op.comp F).sections) (i : J) (hi : i < j) : F.map (CategoryTheory.homOfLE ⋯).op (self.lift j hj x) = ↑x (Opposite.op ⟨i, hi⟩)
Instances For
Given d : F.WellOrderInductionData
, val₀ : F.obj (op ⊥)
and j : J
,
this is the data of an element val : F.obj (op j)
such that the induced
compatible family of elements in all F.obj (op i)
for i ≤ j
is determined by val₀
and the choice of "liftings" given by d
.
- val : F.obj (Opposite.op j)
An element in
F.obj (op j)
, which, by restriction, induces elements inF.obj (op i)
for alli ≤ j
. - map_zero : F.map (CategoryTheory.homOfLE ⋯).op self.val = val₀
- map_succ (i : J) (hi : i < j) : F.map (CategoryTheory.homOfLE ⋯).op self.val = d.succ i ⋯ (F.map (CategoryTheory.homOfLE ⋯).op self.val)
- map_limit (i : J) (hi : Order.IsSuccLimit i) (hij : i ≤ j) : F.map (CategoryTheory.homOfLE hij).op self.val = d.lift i hi ⟨fun (x : { x : J // x ∈ Set.Iio i }ᵒᵖ) => match x with | Opposite.op ⟨k, hk⟩ => F.map (CategoryTheory.homOfLE ⋯).op self.val, ⋯⟩
Instances For
An element in d.Extension val₀ j
induces an element in d.Extension val₀ i
when i ≤ j
.
Equations
- e.ofLE hij = { val := F.map (CategoryTheory.homOfLE hij).op e.val, map_zero := ⋯, map_succ := ⋯, map_limit := ⋯ }
Instances For
The obvious element in d.Extension val₀ ⊥
.
Equations
- CategoryTheory.Functor.WellOrderInductionData.Extension.zero d val₀ = { val := val₀, map_zero := ⋯, map_succ := ⋯, map_limit := ⋯ }
Instances For
The element in d.Extension val₀ (Order.succ j)
obtained by extending
an element in d.Extension val₀ j
when j
is not maximal.
Equations
- e.succ hj = { val := d.succ j hj e.val, map_zero := ⋯, map_succ := ⋯, map_limit := ⋯ }
Instances For
When j
is a limit element, this is the exntesion to d.Extension val₀ j
of a family of elements in d.Extension val₀ i
for all i < j
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
When J
is a well-ordered type, F : Jᵒᵖ ⥤ Type v
, and d : F.WellOrderInductionData
,
this is the section of F
that is determined by val₀ : F.obj (op ⊥)