The category TopModuleCat R
of topological modules #
We define TopModuleCat R
, the category of topological modules, and show that
it has all limits and colimits.
We also provide various adjunctions:
TopModuleCat.withModuleTopologyAdj
: equipping the module topology is left adjoint to the forgetful functor intoModuleCat R
.TopModuleCat.indiscreteAdj
: equipping the indiscrete topology is right adjoint to the forgetful functor intoModuleCat R
.TopModuleCat.freeAdj
: the free-forgetful adjunction betweenTopModuleCat R
andTopCat
.
Future projects #
Show that the forgetful functor to TopCat
preserves filtered colimits.
The category of topological modules.
- isAddCommGroup : AddCommGroup ↑self.toModuleCat
- isModule : Module R ↑self.toModuleCat
- topologicalSpace : TopologicalSpace ↑self.toModuleCat
The underlying topological space.
- isTopologicalAddGroup : IsTopologicalAddGroup ↑self.toModuleCat
- continuousSMul : ContinuousSMul R ↑self.toModuleCat
Instances For
Equations
- TopModuleCat.instCoeSortType R = { coe := fun (M : TopModuleCat R) => ↑M.toModuleCat }
Make an object in TopModuleCat R
from an unbundled topological module.
Equations
- TopModuleCat.of R M = { toModuleCat := ModuleCat.of R M, topologicalSpace := inst✝², isTopologicalAddGroup := ⋯, continuousSMul := inst✝ }
Instances For
Homs in TopModuleCat
as one field structures over ContinuousLinearMap
.
- ofHom' :: (
The underlying continuous linear map. Use
hom
instead.- )
Instances For
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Cast a hom in TopModuleCat
into a continuous linear map.
Equations
Instances For
Construct a hom in TopModuleCat
from a continuous linear map.
Equations
Instances For
Use the ConcreteCategory.hom
projection for @[simps]
lemmas.
Equations
- TopModuleCat.Hom.Simps.hom R A B f = f.hom
Instances For
Construct an iso in TopModuleCat
from a continuous linear equiv.
Equations
- TopModuleCat.ofIso e = { hom := TopModuleCat.ofHom ↑e, inv := TopModuleCat.ofHom ↑e.symm, hom_inv_id := ⋯, inv_hom_id := ⋯ }
Instances For
Cast an iso in TopModuleCat
into a continuous linear equiv.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Equations
- TopModuleCat.instPreadditive R = { homGroup := inferInstance, add_comp := ⋯, comp_add := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- TopModuleCat.instLinear = { homModule := inferInstance, smul_comp := ⋯, comp_smul := ⋯ }
Equations
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
The coinduced topology on M
from a family of continuous linear map into M
, which is the
finest topology that makes it into a topological module and makes every map continuous.
Equations
Instances For
The maps into the coinduced topology as homs in TopModuleCat R
.
Equations
- TopModuleCat.toCoinduced f i = TopModuleCat.ofHom { toLinearMap := ModuleCat.Hom.hom (f i), cont := ⋯ }
Instances For
The cocone of topological modules associated to a cocone over the underlying modules, where the cocone point is given the coinduced topology. This is colimiting when the given cocone is.
Equations
- TopModuleCat.ofCocone c = { pt := TopModuleCat.coinduced c.ι.app, ι := { app := TopModuleCat.toCoinduced c.ι.app, naturality := ⋯ } }
Instances For
Given a colimit cocone over the underlying modules, equipping the cocone point with
the coinduced topology gives a colimit cocone in TopModuleCat R
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The induced topology on M
from a family of continuous linear map from M
, which is the
coarsest topology that makes every map continuous.
Equations
- TopModuleCat.induced f = TopModuleCat.of R ↑M
Instances For
The maps from the induced topology as homs in TopModuleCat R
.
Equations
- TopModuleCat.fromInduced f i = TopModuleCat.ofHom { toLinearMap := ModuleCat.Hom.hom (f i), cont := ⋯ }
Instances For
The cone of topological modules associated to a cone over the underlying modules, where the cone point is given the induced topology. This is limiting when the given cone is.
Equations
- TopModuleCat.ofCone c = { pt := TopModuleCat.induced c.π.app, π := { app := TopModuleCat.fromInduced c.π.app, naturality := ⋯ } }
Instances For
Given a limit cone over the underlying modules, equipping the cone point with
the induced topology gives a limit cone in TopModuleCat R
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The functor equipping a module over a topological ring with the finest possible topology making it into a topological module. This is left adjoint to the forgetful functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The adjunction between withModuleTopology
and the forgetful functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The functor equipping a module with the indiscrete topology. This is right adjoint to the forgetful functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The adjunction between the forgetful functor and the indiscrete topology functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The free topological module over a topological space.
Equations
- TopModuleCat.freeObj R X = TopModuleCat.of R (↑X →₀ R)
Instances For
The free topological module over a topological space is functorial.
Equations
- TopModuleCat.freeMap R f = CategoryTheory.ConcreteCategory.ofHom { toLinearMap := Finsupp.lmapDomain R R ⇑(TopCat.Hom.hom f), cont := ⋯ }
Instances For
The free topological module over a topological space as a functor. This is left adjoint to the forgetful functor.
Equations
- TopModuleCat.free R = { obj := TopModuleCat.freeObj R, map := fun {X Y : TopCat} (f : X ⟶ Y) => TopModuleCat.freeMap R f, map_id := ⋯, map_comp := ⋯ }
Instances For
The free-forgetful adjoint for TopModuleCat R
.
Equations
- One or more equations did not get rendered due to their size.