Closed submodules of a topological module #
This files builds the frame of closed R
-submodules of a topological module M
.
One can turn s : Submodule R E
+ hs : IsClosed s
into s : ClosedSubmodule R E
in a tactic
block by doing lift s to ClosedSubmodule R E using hs
.
TODO #
Actually provide the Order.Frame (ClosedSubmodule R M)
instance.
structure
ClosedSubmodule
(R : Type u_2)
(M : Type u_3)
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
extends Submodule R M, TopologicalSpace.Closeds M :
Type u_3
The type of closed submodules of a topological module.
Instances For
theorem
ClosedSubmodule.ext
{R : Type u_2}
{M : Type u_3}
{inst✝ : Semiring R}
{inst✝¹ : AddCommMonoid M}
{inst✝² : TopologicalSpace M}
{inst✝³ : Module R M}
{x y : ClosedSubmodule R M}
(carrier : (↑x).carrier = (↑y).carrier)
:
theorem
ClosedSubmodule.ext_iff
{R : Type u_2}
{M : Type u_3}
{inst✝ : Semiring R}
{inst✝¹ : AddCommMonoid M}
{inst✝² : TopologicalSpace M}
{inst✝³ : Module R M}
{x y : ClosedSubmodule R M}
:
theorem
ClosedSubmodule.toSubmodule_injective
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
instance
ClosedSubmodule.instSetLike
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
SetLike (ClosedSubmodule R M) M
Equations
- ClosedSubmodule.instSetLike = { coe := fun (s : ClosedSubmodule R M) => ↑↑s, coe_injective' := ⋯ }
theorem
ClosedSubmodule.toCloseds_injective
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
instance
ClosedSubmodule.instAddSubmonoidClass
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
AddSubmonoidClass (ClosedSubmodule R M) M
instance
ClosedSubmodule.instSMulMemClass
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
SMulMemClass (ClosedSubmodule R M) R M
instance
ClosedSubmodule.instCoeSubmodule
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
Coe (ClosedSubmodule R M) (Submodule R M)
Equations
@[simp]
theorem
ClosedSubmodule.carrier_eq_coe
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
(s : ClosedSubmodule R M)
:
@[simp]
theorem
ClosedSubmodule.coe_toSubmodule
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
(s : ClosedSubmodule R M)
:
@[simp]
theorem
ClosedSubmodule.coe_toCloseds
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
(s : ClosedSubmodule R M)
:
theorem
ClosedSubmodule.isClosed
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
(s : ClosedSubmodule R M)
:
IsClosed ↑s
instance
ClosedSubmodule.instCanLiftSubmoduleToSubmoduleIsClosedCoe
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
CanLift (Submodule R M) (ClosedSubmodule R M) toSubmodule fun (x : Submodule R M) => IsClosed ↑x
@[simp]
theorem
ClosedSubmodule.toSubmodule_le_toSubmodule
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
{s t : ClosedSubmodule R M}
:
def
ClosedSubmodule.comap
{R : Type u_2}
{M : Type u_3}
{N : Type u_4}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[AddCommMonoid N]
[TopologicalSpace N]
[Module R N]
(f : M →L[R] N)
(s : ClosedSubmodule R N)
:
ClosedSubmodule R M
The preimage of a closed submodule under a continuous linear map as a closed submodule.
Equations
- ClosedSubmodule.comap f s = { toSubmodule := Submodule.comap f ↑s, isClosed' := ⋯ }
Instances For
@[simp]
theorem
ClosedSubmodule.coe_comap
{R : Type u_2}
{M : Type u_3}
{N : Type u_4}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[AddCommMonoid N]
[TopologicalSpace N]
[Module R N]
(f : M →L[R] N)
(s : ClosedSubmodule R N)
:
@[simp]
theorem
ClosedSubmodule.mem_comap
{R : Type u_2}
{M : Type u_3}
{N : Type u_4}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[AddCommMonoid N]
[TopologicalSpace N]
[Module R N]
{f : M →L[R] N}
{s : ClosedSubmodule R N}
{x : M}
:
@[simp]
theorem
ClosedSubmodule.toSubmodule_comap
{R : Type u_2}
{M : Type u_3}
{N : Type u_4}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[AddCommMonoid N]
[TopologicalSpace N]
[Module R N]
(f : M →L[R] N)
(s : ClosedSubmodule R N)
:
@[simp]
theorem
ClosedSubmodule.comap_id
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
(s : ClosedSubmodule R M)
:
theorem
ClosedSubmodule.comap_comap
{R : Type u_2}
{M : Type u_3}
{N : Type u_4}
{O : Type u_5}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[AddCommMonoid N]
[TopologicalSpace N]
[Module R N]
[AddCommMonoid O]
[TopologicalSpace O]
[Module R O]
(g : N →L[R] O)
(f : M →L[R] N)
(s : ClosedSubmodule R O)
:
instance
ClosedSubmodule.instInf
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
Min (ClosedSubmodule R M)
Equations
- ClosedSubmodule.instInf = { min := fun (s t : ClosedSubmodule R M) => { toSubmodule := ↑s ⊓ ↑t, isClosed' := ⋯ } }
instance
ClosedSubmodule.instInfSet
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
InfSet (ClosedSubmodule R M)
Equations
- ClosedSubmodule.instInfSet = { sInf := fun (S : Set (ClosedSubmodule R M)) => { toSubmodule := ⨅ s ∈ S, ↑s, isClosed' := ⋯ } }
@[simp]
theorem
ClosedSubmodule.toSubmodule_sInf
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
(S : Set (ClosedSubmodule R M))
:
@[simp]
theorem
ClosedSubmodule.toSubmodule_iInf
{ι : Sort u_1}
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
(f : ι → ClosedSubmodule R M)
:
@[simp]
theorem
ClosedSubmodule.coe_sInf
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
(S : Set (ClosedSubmodule R M))
:
@[simp]
theorem
ClosedSubmodule.coe_iInf
{ι : Sort u_1}
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
(f : ι → ClosedSubmodule R M)
:
@[simp]
theorem
ClosedSubmodule.mem_sInf
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
{x : M}
{S : Set (ClosedSubmodule R M)}
:
@[simp]
theorem
ClosedSubmodule.mem_iInf
{ι : Sort u_1}
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
{x : M}
{f : ι → ClosedSubmodule R M}
:
instance
ClosedSubmodule.instSemilatticeInf
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
@[simp]
theorem
ClosedSubmodule.toSubmodule_inf
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
(s t : ClosedSubmodule R M)
:
@[simp]
theorem
ClosedSubmodule.coe_inf
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
(s t : ClosedSubmodule R M)
:
@[simp]
theorem
ClosedSubmodule.mem_inf
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
{s t : ClosedSubmodule R M}
{x : M}
:
instance
ClosedSubmodule.instTop
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
Top (ClosedSubmodule R M)
@[simp]
theorem
ClosedSubmodule.toSubmodule_top
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
@[simp]
theorem
ClosedSubmodule.coe_top
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
:
@[simp]
theorem
ClosedSubmodule.mem_top
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
{x : M}
:
instance
ClosedSubmodule.instOrderBot
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[T1Space M]
:
OrderBot (ClosedSubmodule R M)
@[simp]
theorem
ClosedSubmodule.toSubmodule_bot
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[T1Space M]
:
@[simp]
theorem
ClosedSubmodule.coe_bot
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[T1Space M]
:
@[simp]
theorem
ClosedSubmodule.mem_bot
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
{x : M}
[T1Space M]
:
def
Submodule.closure
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[ContinuousAdd M]
[ContinuousConstSMul R M]
(s : Submodule R M)
:
ClosedSubmodule R M
The closure of a submodule as a closed submodule.
Equations
- s.closure = { toSubmodule := s.topologicalClosure, isClosed' := ⋯ }
Instances For
@[simp]
theorem
Submodule.coe_closure
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[ContinuousAdd M]
[ContinuousConstSMul R M]
(s : Submodule R M)
:
@[simp]
theorem
Submodule.closure_le
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[ContinuousAdd M]
[ContinuousConstSMul R M]
{s : Submodule R M}
{t : ClosedSubmodule R M}
:
def
ClosedSubmodule.map
{R : Type u_2}
{M : Type u_3}
{N : Type u_4}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[AddCommMonoid N]
[TopologicalSpace N]
[Module R N]
[ContinuousAdd N]
[ContinuousConstSMul R N]
(f : M →L[R] N)
(s : ClosedSubmodule R M)
:
ClosedSubmodule R N
The closure of the image of a closed submodule under a continuous linear map is a closed submodule.
ClosedSubmodule.map f
is left-adjoint to ClosedSubmodule.comap f
.
See ClosedSubmodule.gc_map_comap
.
Equations
- ClosedSubmodule.map f s = (Submodule.map f ↑s).closure
Instances For
@[simp]
theorem
ClosedSubmodule.map_id
{R : Type u_2}
{M : Type u_3}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[ContinuousAdd M]
[ContinuousConstSMul R M]
(s : ClosedSubmodule R M)
:
theorem
ClosedSubmodule.map_le_iff_le_comap
{R : Type u_2}
{M : Type u_3}
{N : Type u_4}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[AddCommMonoid N]
[TopologicalSpace N]
[Module R N]
[ContinuousAdd N]
[ContinuousConstSMul R N]
{f : M →L[R] N}
{s : ClosedSubmodule R M}
{t : ClosedSubmodule R N}
:
theorem
ClosedSubmodule.gc_map_comap
{R : Type u_2}
{M : Type u_3}
{N : Type u_4}
[Semiring R]
[AddCommMonoid M]
[TopologicalSpace M]
[Module R M]
[AddCommMonoid N]
[TopologicalSpace N]
[Module R N]
[ContinuousAdd N]
[ContinuousConstSMul R N]
{f : M →L[R] N}
:
GaloisConnection (map f) (comap f)