The Orthogonal-reflection construction #
Given W : MorphismProperty C (which should be small) and assuming the existence
of certain colimits in C, we construct a morphism toSucc W Z : Z ⟶ succ W Z for
any Z : C. This morphism belongs to W.isLocal.isLocal and
is an isomorphism iff Z belongs to W.isLocal (see the lemma isIso_toSucc_iff).
The morphism toSucc W Z : Z ⟶ succ W Z is defined as a composition
of two morphisms that are roughly described as follows:
toStep W Z : Z ⟶ step W Z: for any morphismf : X ⟶ YsatisfyingWand any morphismX ⟶ Z, we "attach" a morphismY ⟶ step W Z(using coproducts and a pushout in essentially the same way as it is done in the fileMathlib/CategoryTheory/SmallObject/Construction.leanfor the small object argument);fromStep W Z : step W Z ⟶ succ W Z: this morphism coequalizes all pairs of morphismsg₁ g₂ : Y ⟶ step W Zsuch that there is af : X ⟶ YsatisfyingWsuch thatf ≫ g₁ = f ≫ g₂.
The morphism toSucc W Z : Z ⟶ succ W Z is a variant of the (wrong) definition
p. 32 in the book by Adámek and Rosický. In this book, a slightly different object
as succ W Z is defined directly as a colimit of an intricate diagram, but
contrary to what is stated on p. 33, it does not satisfy isIso_toSucc_iff.
The author of this file was unable to not understand the attempt of the authors
to fix this mistake in the errata to this book. This led to the definition
in two steps outlined above.
Main results #
The morphisms described above toSucc W Z : Z ⟶ succ W Z for all Z : C allow to
define succStruct W Z₀ : SuccStruct C for any Z₀ : C. By applying
a transfinite iteration to this SuccStruct, we obtain the following results
under the assumption that W : MorphismProperty C is a w-small property
of morphisms in a locally κ-presentable category C (with κ : Cardinal.{w}
a regular cardinal) such that the domains and codomains of the morphisms
satisfying W are κ-presentable :
MorphismProperty.isRightAdjoint_ι_isLocal: existence of the left adjoint of the inclusionW.isLocal ⥤ C;MorphismProperty.isLocallyPresentable_isLocal: the full subcategoryW.isLocalis locally presentable.
This is essentially the implication (i) → (ii) in Theorem 1.39 (and the corollary 1.40)
in the book by Adámek and Rosický (note that according to the
errata to this book, the implication (ii) → (i) is wrong when κ = ℵ₀).
References #
Given W : MorphismProperty C and Z : C, this is the index type
parametrising the data of a morphism f : X ⟶ Y satisfying W
and a morphism X ⟶ Z.
Equations
- CategoryTheory.OrthogonalReflection.D₁ W Z = ((f : ↑W.toSet) × ((↑f).left ⟶ Z))
Instances For
If d : D₁ W Z corresponds to the data of f : X ⟶ Y satisfying W and
of a morphism X ⟶ Z, this is the object X.
Instances For
If d : D₁ W Z corresponds to the data of f : X ⟶ Y satisfying W and
of a morphism X ⟶ Z, this is the object Y.
Instances For
Considering all diagrams consisting of a morphism f : X ⟶ Y satisfying W
and of a morphism d : X ⟶ Z, this is the morphism from the coproduct of
all these X objects to Z given by these morphisms d.
Equations
Instances For
The inclusion of a summand in ∐ obj₁.
Equations
Instances For
The coproduct of all the morphisms f indexed by all diagrams
consisting of a morphism f : X ⟶ Y satisfying W and of a morphism d : X ⟶ Z.
Equations
- CategoryTheory.OrthogonalReflection.D₁.t W Z = CategoryTheory.Limits.Sigma.map fun (d : CategoryTheory.OrthogonalReflection.D₁ W Z) => (↑d.fst).hom
Instances For
The inclusion of a summand in ∐ obj₂.
Equations
Instances For
The intermediate object in the definition of the morphism toSucc W Z : Z ⟶ succ W Z.
It is the pushout of the following square:
∐ D₁.obj₁ ⟶ ∐ D₁.obj₂
| |
v v
Z ⟶ step W Z
where the coproduct is taken over all the diagram consisting of a morphism f : X ⟶ Y
satisfying W and a morphism X ⟶ Z. The top map is the coproduct of all of these f.
Equations
Instances For
The canonical map from Z to the pushout of D₁.t W Z and D₁.l W Z.
Equations
Instances For
The index type parametrising the data of two morphisms g₁ g₂ : Y ⟶ step W Z, and
a map f : X ⟶ Y satisfying W such that f ≫ g₁ = f ≫ g₂.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The shape of the multicoequalizer of all pairs of morphisms g₁ g₂ : Y ⟶ step W Z with
a f : X ⟶ Y satisfying W such that f ≫ g₁ = f ≫ g₂.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The diagram of the multicoequalizer of all pair of morphisms g₁ g₂ : Y ⟶ step W Z with
a f : X ⟶ Y satisfying W such that f ≫ g₁ = f ≫ g₂.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The object succ W Z is the multicoequalizer of all pairs of morphisms
g₁ g₂ : Y ⟶ step W Z with a f : X ⟶ Y satisfying W such that f ≫ g₁ = f ≫ g₂.
Equations
Instances For
The projection from Z to the multicoequalizer of all morphisms g₁ g₂ : Y ⟶ step W Z with
a f : X ⟶ Y satisfying W such that f ≫ g₁ = f ≫ g₂.
Equations
Instances For
The morphism Z ⟶ succ W Z.
Equations
Instances For
Alias of CategoryTheory.OrthogonalReflection.isLocal_isLocal_toSucc.
The successor structure of the orthogonal-reflection construction.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The transfinite iteration of succStruct W Z to the power κ.ord.toType.
Equations
Instances For
The map which shall exhibit reflectionObj W Z κ as the image of Z by
the left adjoint of the inclusion of W.isLocal, see corepresentableBy.
Equations
Instances For
The morphism reflection W Z κ : Z ⟶ reflectionObj W Z κ is a transfinite
compositions of morphisms in LeftBousfield.W W.isLocal.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The functor κ.ord.toType ⥤ C that is the diagram of the
transfinite composition transfiniteCompositionOfShapeReflection.
Equations
Instances For
(iteration W Z κ).obj (Order.succ j) identifies to the image of
(iteration W Z κ).obj j by succ.
Equations
Instances For
The morphism reflection W Z κ : Z ⟶ reflectionObj W Z κ exhibits reflectionObj W Z κ
as the image of Z by the left adjoint of the inclusion W.isLocal.ι.
Equations
- One or more equations did not get rendered due to their size.