Documentation

Mathlib.Topology.Algebra.Module.ClosedSubmodule

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) :
    x = y
    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} :
    x = y (↑x).carrier = (↑y).carrier
    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) :
    (↑s).carrier = s
    @[simp]
    theorem ClosedSubmodule.mem_mk {R : Type u_2} {M : Type u_3} [Semiring R] [AddCommMonoid M] [TopologicalSpace M] [Module R M] {x : M} {s : Submodule R M} {hs : IsClosed s.carrier} :
    x { toSubmodule := s, isClosed' := hs } x s
    @[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) :
    s = s
    @[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) :
    s = s
    theorem ClosedSubmodule.isClosed {R : Type u_2} {M : Type u_3} [Semiring R] [AddCommMonoid M] [TopologicalSpace M] [Module R M] (s : ClosedSubmodule R M) :
    @[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} :
    s t s t
    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) :

    The preimage of a closed submodule under a continuous linear map as a closed submodule.

    Equations
    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) :
      (comap f s) = f ⁻¹' s
      @[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} :
      x comap f s f x s
      @[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) :
      (comap f s) = Submodule.comap f s
      @[simp]
      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) :
      comap f (comap g s) = comap (g.comp f) s
      instance ClosedSubmodule.instInf {R : Type u_2} {M : Type u_3} [Semiring R] [AddCommMonoid M] [TopologicalSpace M] [Module R M] :
      Equations
      Equations
      @[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)) :
      (sInf S) = sS, s
      @[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) :
      (⨅ (i : ι), f i) = ⨅ (i : ι), (f i)
      @[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)) :
      (sInf S) = sS, s
      @[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) :
      (⨅ (i : ι), f i) = ⨅ (i : ι), (f i)
      @[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)} :
      x sInf S sS, x s
      @[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} :
      x ⨅ (i : ι), f i ∀ (i : ι), x f i
      @[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) :
      (st) = st
      @[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) :
      (st) = st
      @[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} :
      x st x s x t
      instance ClosedSubmodule.instTop {R : Type u_2} {M : Type u_3} [Semiring R] [AddCommMonoid M] [TopologicalSpace M] [Module R M] :
      Equations
      @[simp]
      @[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} :
      Equations
      @[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] :
      = {0}
      @[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] :
      x x = 0

      The closure of a submodule as a closed submodule.

      Equations
      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) :
        s.closure = closure s
        @[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} :
        s.closure t s t

        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
        Instances For
          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} :
          map f s t s comap f t