Documentation

Mathlib.Topology.Algebra.Module.LinearPMap

Partially defined linear operators over topological vector spaces #

We define basic notions of partially defined linear operators, which we call unbounded operators for short. In this file we prove all elementary properties of unbounded operators that do not assume that the underlying spaces are normed.

Main definitions #

Main statements #

References #

Tags #

Unbounded operators, closed operators

Closed and closable operators #

def LinearPMap.IsClosed {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] (f : E →ₗ.[R] F) :

An unbounded operator is closed iff its graph is closed.

Equations
Instances For

    An unbounded operator is closable iff the closure of its graph is a graph.

    Equations
    • f.IsClosable = ∃ (f' : E →ₗ.[R] F), f.graph.topologicalClosure = f'.graph
    Instances For
      theorem LinearPMap.IsClosed.isClosable {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} (hf : f.IsClosed) :
      f.IsClosable

      A closed operator is trivially closable.

      theorem LinearPMap.IsClosable.leIsClosable {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} {g : E →ₗ.[R] F} (hf : f.IsClosable) (hfg : g f) :
      g.IsClosable

      If g has a closable extension f, then g itself is closable.

      theorem LinearPMap.IsClosable.existsUnique {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} (hf : f.IsClosable) :
      ∃! f' : E →ₗ.[R] F, f.graph.topologicalClosure = f'.graph

      The closure is unique.

      noncomputable def LinearPMap.closure {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] (f : E →ₗ.[R] F) :

      If f is closable, then f.closure is the closure. Otherwise it is defined as f.closure = f.

      Equations
      Instances For
        theorem LinearPMap.closure_def {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} (hf : f.IsClosable) :
        f.closure = Exists.choose hf
        theorem LinearPMap.closure_def' {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} (hf : ¬f.IsClosable) :
        f.closure = f
        theorem LinearPMap.IsClosable.graph_closure_eq_closure_graph {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} (hf : f.IsClosable) :
        f.graph.topologicalClosure = f.closure.graph

        The closure (as a submodule) of the graph is equal to the graph of the closure (as a LinearPMap).

        theorem LinearPMap.le_closure {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] (f : E →ₗ.[R] F) :
        f f.closure

        A LinearPMap is contained in its closure.

        theorem LinearPMap.IsClosable.closure_mono {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} {g : E →ₗ.[R] F} (hg : g.IsClosable) (h : f g) :
        f.closure g.closure
        theorem LinearPMap.IsClosable.closure_isClosed {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} (hf : f.IsClosable) :
        f.closure.IsClosed

        If f is closable, then the closure is closed.

        theorem LinearPMap.IsClosable.closureIsClosable {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} (hf : f.IsClosable) :
        f.closure.IsClosable

        If f is closable, then the closure is closable.

        theorem LinearPMap.isClosable_iff_exists_closed_extension {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} :
        f.IsClosable ∃ (g : E →ₗ.[R] F), g.IsClosed f g

        The core of a linear operator #

        structure LinearPMap.HasCore {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] (f : E →ₗ.[R] F) (S : Submodule R E) :

        A submodule S is a core of f if the closure of the restriction of f to S is f.

        • le_domain : S f.domain
        • closure_eq : (f.domRestrict S).closure = f
        Instances For
          theorem LinearPMap.HasCore.le_domain {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} {S : Submodule R E} (self : f.HasCore S) :
          S f.domain
          theorem LinearPMap.HasCore.closure_eq {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} {S : Submodule R E} (self : f.HasCore S) :
          (f.domRestrict S).closure = f
          theorem LinearPMap.hasCore_def {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} {S : Submodule R E} (h : f.HasCore S) :
          (f.domRestrict S).closure = f
          theorem LinearPMap.closureHasCore {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] (f : E →ₗ.[R] F) :
          f.closure.HasCore f.domain

          For every unbounded operator f the submodule f.domain is a core of its closure.

          Note that we don't require that f is closable, due to the definition of the closure.

          Topological properties of the inverse #

          theorem LinearPMap.closure_inverse_graph {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} (hf : LinearMap.ker f.toFun = ) (hf' : f.IsClosable) (hcf : LinearMap.ker f.closure.toFun = ) :
          f.closure.inverse.graph = f.inverse.graph.topologicalClosure

          If f is invertible and closable as well as its closure being invertible, then the graph of the inverse of the closure is given by the closure of the graph of the inverse.

          theorem LinearPMap.inverse_isClosable_iff {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} (hf : LinearMap.ker f.toFun = ) (hf' : f.IsClosable) :
          f.inverse.IsClosable LinearMap.ker f.closure.toFun =

          Assuming that f is invertible and closable, then the closure is invertible if and only if the inverse of f is closable.

          theorem LinearPMap.inverse_closure {R : Type u_1} {E : Type u_2} {F : Type u_3} [CommRing R] [AddCommGroup E] [AddCommGroup F] [Module R E] [Module R F] [TopologicalSpace E] [TopologicalSpace F] [ContinuousAdd E] [ContinuousAdd F] [TopologicalSpace R] [ContinuousSMul R E] [ContinuousSMul R F] {f : E →ₗ.[R] F} (hf : LinearMap.ker f.toFun = ) (hf' : f.IsClosable) (hcf : LinearMap.ker f.closure.toFun = ) :
          f.inverse.closure = f.closure.inverse

          If f is invertible and closable, then taking the closure and the inverse commute.