Documentation

Mathlib.Data.LocallyFinsupp

Type of functions with locally finite support #

This file defines functions with locally finite support, provides supporting API. For suitable targets, it establishes functions with locally finite support as an instance of a lattice ordered commutative group.

Throughout the present file, X denotes a topologically space and U a subset of X.

Definition, coercion to functions and basic extensionality lemmas #

A function with locally finite support within U is a function X → Y whose support is locally finite within U and entirely contained in U. For T1-spaces, the theorem supportDiscreteWithin_iff_locallyFiniteWithin shows that the first condition is equivalent to the condition that the support f is discrete within U.

structure Function.locallyFinsuppWithin {X : Type u_1} [TopologicalSpace X] (U : Set X) (Y : Type u_2) [Zero Y] :
Type (max u_1 u_2)

A function with locally finite support within U is a triple as specified below.

Instances For
    def Function.locallyFinsupp (X : Type u_1) [TopologicalSpace X] (Y : Type u_2) [Zero Y] :
    Type (max u_1 u_2)

    A function with locally finite support is a function with locally finite support within ⊤ : Set X.

    Equations
    Instances For
      theorem supportDiscreteWithin_iff_locallyFiniteWithin {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [T1Space X] [Zero Y] {f : XY} (h : Function.support f U) :
      f =ᶠ[Filter.codiscreteWithin U] 0 zU, tnhds z, (t Function.support f).Finite

      For T1 spaces, the condition supportLocallyFiniteWithinDomain' is equivalent to saying that the support is codiscrete within U.

      Functions with locally finite support within U are FunLike: the coercion to functions is injective.

      Equations
      @[reducible, inline]
      abbrev Function.locallyFinsuppWithin.support {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [Zero Y] (D : locallyFinsuppWithin U Y) :
      Set X

      This allows writing D.support instead of Function.support D

      Equations
      Instances For
        theorem Function.locallyFinsuppWithin.supportLocallyFiniteWithinDomain {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [Zero Y] (D : locallyFinsuppWithin U Y) (z : X) :
        z Utnhds z, (t D.support).Finite
        theorem Function.locallyFinsuppWithin.ext {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [Zero Y] {D₁ D₂ : locallyFinsuppWithin U Y} (h : ∀ (a : X), D₁ a = D₂ a) :
        D₁ = D₂
        theorem Function.locallyFinsuppWithin.coe_injective {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [Zero Y] :
        Injective fun (x : locallyFinsuppWithin U Y) => x

        Elementary properties of the support #

        @[simp]
        theorem Function.locallyFinsuppWithin.apply_eq_zero_of_not_mem {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [Zero Y] {z : X} (D : locallyFinsuppWithin U Y) (hz : zU) :
        D z = 0

        Simplifier lemma: Functions with locally finite support within U evaluate to zero outside of U.

        On a T1 space, the support of a functions with locally finite support within U is discrete.

        If X is T1 and if U is closed, then the support of support of a function with locally finite support within U is also closed.

        If X is T2 and if U is compact, then the support of a function with locally finite support within U is finite.

        Lattice ordered group structure #

        If X is a suitable instance, this section equips functions with locally finite support within U with the standard structure of a lattice ordered group, where addition, comparison, min and max are defined pointwise.

        Functions with locally finite support within U form an additive subgroup of functions X → Y.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For

          Assign a function with locally finite support within U to a function in the subgroup.

          Equations
          Instances For
            @[simp]
            theorem Function.locallyFinsuppWithin.mk_of_mem_toFun {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [AddCommGroup Y] (f : XY) (hf : f locallyFinsuppWithin.addSubgroup U) (a✝ : X) :
            (mk_of_mem f hf) a✝ = f a✝
            @[simp]
            theorem Function.locallyFinsuppWithin.coe_zero {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [AddCommGroup Y] :
            0 = 0
            @[simp]
            theorem Function.locallyFinsuppWithin.coe_add {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [AddCommGroup Y] (D₁ D₂ : locallyFinsuppWithin U Y) :
            ⇑(D₁ + D₂) = D₁ + D₂
            @[simp]
            theorem Function.locallyFinsuppWithin.coe_neg {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [AddCommGroup Y] (D : locallyFinsuppWithin U Y) :
            ⇑(-D) = -D
            @[simp]
            theorem Function.locallyFinsuppWithin.coe_sub {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [AddCommGroup Y] (D₁ D₂ : locallyFinsuppWithin U Y) :
            ⇑(D₁ - D₂) = D₁ - D₂
            @[simp]
            theorem Function.locallyFinsuppWithin.coe_nsmul {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [AddCommGroup Y] (D : locallyFinsuppWithin U Y) (n : ) :
            ⇑(n D) = n D
            @[simp]
            theorem Function.locallyFinsuppWithin.coe_zsmul {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [AddCommGroup Y] (D : locallyFinsuppWithin U Y) (n : ) :
            ⇑(n D) = n D
            instance Function.locallyFinsuppWithin.instLE {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [LE Y] [Zero Y] :
            Equations
            theorem Function.locallyFinsuppWithin.le_def {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [LE Y] [Zero Y] {D₁ D₂ : locallyFinsuppWithin U Y} :
            D₁ D₂ D₁ D₂
            Equations
            theorem Function.locallyFinsuppWithin.lt_def {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [Preorder Y] [Zero Y] {D₁ D₂ : locallyFinsuppWithin U Y} :
            D₁ < D₂ D₁ < D₂
            Equations
            • One or more equations did not get rendered due to their size.
            @[simp]
            theorem Function.locallyFinsuppWithin.max_apply {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [SemilatticeSup Y] [Zero Y] {D₁ D₂ : locallyFinsuppWithin U Y} {x : X} :
            (D₁ D₂) x = D₁ x D₂ x
            Equations
            • One or more equations did not get rendered due to their size.
            @[simp]
            theorem Function.locallyFinsuppWithin.min_apply {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [SemilatticeInf Y] [Zero Y] {D₁ D₂ : locallyFinsuppWithin U Y} {x : X} :
            (D₁ D₂) x = D₁ x D₂ x

            Functions with locally finite support within U form an ordered commutative group.

            Equations

            Restriction #

            noncomputable def Function.locallyFinsuppWithin.restrict {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [Zero Y] {V : Set X} (D : locallyFinsuppWithin U Y) (h : V U) :

            If V is a subset of U, then functions with locally finite support within U restrict to functions with locally finite support within V, by setting their values to zero outside of V.

            Equations
            • D.restrict h = { toFun := fun (z : X) => if hz : z V then D z else 0, supportWithinDomain' := , supportLocallyFiniteWithinDomain' := }
            Instances For
              theorem Function.locallyFinsuppWithin.restrict_apply {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [Zero Y] {V : Set X} (D : locallyFinsuppWithin U Y) (h : V U) (z : X) :
              (D.restrict h) z = if z V then D z else 0
              theorem Function.locallyFinsuppWithin.restrict_eqOn {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [Zero Y] {V : Set X} (D : locallyFinsuppWithin U Y) (h : V U) :
              Set.EqOn (⇑(D.restrict h)) (⇑D) V
              theorem Function.locallyFinsuppWithin.restrict_eqOn_compl {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [Zero Y] {V : Set X} (D : locallyFinsuppWithin U Y) (h : V U) :
              Set.EqOn (⇑(D.restrict h)) 0 V
              noncomputable def Function.locallyFinsuppWithin.restrictMonoidHom {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [AddCommGroup Y] {V : Set X} (h : V U) :

              Restriction as a group morphism

              Equations
              Instances For
                @[simp]
                noncomputable def Function.locallyFinsuppWithin.restrictLatticeHom {X : Type u_1} [TopologicalSpace X] {U : Set X} {Y : Type u_2} [AddCommGroup Y] [Lattice Y] {V : Set X} (h : V U) :

                Restriction as a lattice morphism

                Equations
                Instances For
                  @[simp]