Documentation

Mathlib.Topology.ContinuousFunction.CompactlySupported

Compactly supported continuous functions #

In this file, we define the type C_c(α, β) of compactly supported continuous functions and the class CompactlySupportedContinuousMapClass, and prove basic properties.

Main definitions and results #

This file contains various instances such as Add, Mul, SMul F C_c(α, β) when F is a class of continuous functions. When β has more structures, C_c(α, β) inherits such structures as AddCommGroup, NonUnitalRing and StarRing.

When the domain α is compact, ContinuousMap.liftCompactlySupported gives the identification C(α, β) ≃ C_c(α, β).

structure CompactlySupportedContinuousMap (α : Type u_5) (β : Type u_6) [TopologicalSpace α] [Zero β] [TopologicalSpace β] extends ContinuousMap :
Type (max u_5 u_6)

C_c(α, β) is the type of continuous functions α → β with compact support from a topological space to a topological space with a zero element.

When possible, instead of parametrizing results over f : C_c(α, β), you should parametrize over {F : Type*} [CompactlySupportedContinuousMapClass F α β] (f : F).

When you extend this structure, make sure to extend CompactlySupportedContinuousMapClass.

  • toFun : αβ
  • continuous_toFun : Continuous self.toFun
  • hasCompactSupport' : HasCompactSupport self.toFun

    The function has compact support .

Instances For

    The function has compact support .

    C_c(α, β) is the type of continuous functions α → β with compact support from a topological space to a topological space with a zero element.

    When possible, instead of parametrizing results over f : C_c(α, β), you should parametrize over {F : Type*} [CompactlySupportedContinuousMapClass F α β] (f : F).

    When you extend this structure, make sure to extend CompactlySupportedContinuousMapClass.

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

      C_c(α, β) is the type of continuous functions α → β with compact support from a topological space to a topological space with a zero element.

      When possible, instead of parametrizing results over f : C_c(α, β), you should parametrize over {F : Type*} [CompactlySupportedContinuousMapClass F α β] (f : F).

      When you extend this structure, make sure to extend CompactlySupportedContinuousMapClass.

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

        CompactlySupportedContinuousMapClass F α β states that F is a type of continuous maps with compact support.

        You should also extend this typeclass when you extend CompactlySupportedContinuousMap.

        • map_continuous : ∀ (f : F), Continuous f
        • hasCompactSupport : ∀ (f : F), HasCompactSupport f

          Each member of the class has compact support.

        Instances

          Each member of the class has compact support.

          Equations
          @[simp]
          theorem CompactlySupportedContinuousMap.coe_toContinuousMap {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [Zero β] (f : CompactlySupportedContinuousMap α β) :
          f.toContinuousMap = f
          theorem CompactlySupportedContinuousMap.ext {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [Zero β] {f : CompactlySupportedContinuousMap α β} {g : CompactlySupportedContinuousMap α β} (h : ∀ (x : α), f x = g x) :
          f = g
          @[simp]
          theorem CompactlySupportedContinuousMap.coe_mk {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [Zero β] (f : C(α, β)) (h : HasCompactSupport f) :
          { toContinuousMap := f, hasCompactSupport' := h } = f
          def CompactlySupportedContinuousMap.copy {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [Zero β] (f : CompactlySupportedContinuousMap α β) (f' : αβ) (h : f' = f) :

          Copy of a CompactlySupportedContinuousMap with a new toFun equal to the old one. Useful to fix definitional equalities.

          Equations
          • f.copy f' h = { toFun := f', continuous_toFun := , hasCompactSupport' := }
          Instances For
            @[simp]
            theorem CompactlySupportedContinuousMap.coe_copy {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [Zero β] (f : CompactlySupportedContinuousMap α β) (f' : αβ) (h : f' = f) :
            (f.copy f' h) = f'
            theorem CompactlySupportedContinuousMap.copy_eq {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [Zero β] (f : CompactlySupportedContinuousMap α β) (f' : αβ) (h : f' = f) :
            f.copy f' h = f
            @[simp]
            theorem CompactlySupportedContinuousMap.ContinuousMap.liftCompactlySupported_symm_apply {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [Zero β] [CompactSpace α] (f : CompactlySupportedContinuousMap α β) :
            CompactlySupportedContinuousMap.ContinuousMap.liftCompactlySupported.symm f = f
            @[simp]
            theorem CompactlySupportedContinuousMap.ContinuousMap.liftCompactlySupported_apply_toFun {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [Zero β] [CompactSpace α] (f : C(α, β)) (a : α) :
            (CompactlySupportedContinuousMap.ContinuousMap.liftCompactlySupported f) a = f a

            A continuous function on a compact space automatically has compact support.

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

              Algebraic structure #

              Whenever β has the structure of continuous additive monoid and a compatible topological structure, then C_c(α, β) inherits a corresponding algebraic structure. The primary exception to this is that C_c(α, β) will not have a multiplicative identity.

              Equations
              • CompactlySupportedContinuousMap.instZero = { zero := { toFun := 0, continuous_toFun := , hasCompactSupport' := } }
              Equations
              • CompactlySupportedContinuousMap.instInhabited = { default := 0 }
              @[simp]
              theorem CompactlySupportedContinuousMap.coe_zero {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [Zero β] :
              0 = 0
              theorem CompactlySupportedContinuousMap.zero_apply {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] (x : α) [Zero β] :
              0 x = 0
              Equations
              • CompactlySupportedContinuousMap.instMulOfContinuousMul = { mul := fun (f g : CompactlySupportedContinuousMap α β) => { toContinuousMap := f * g, hasCompactSupport' := } }

              the product of f : F assuming ContinuousMapClass F α γ and ContinuousSMul γ β and g : C_c(α, β) is in C_c(α, β)

              Equations
              • One or more equations did not get rendered due to their size.
              @[simp]
              theorem CompactlySupportedContinuousMap.coe_smulc {α : Type u_2} {β : Type u_3} {γ : Type u_4} [TopologicalSpace α] [TopologicalSpace β] [Zero β] [TopologicalSpace γ] [SMulZeroClass γ β] [ContinuousSMul γ β] {F : Type u_5} [FunLike F α γ] [ContinuousMapClass F α γ] (f : F) (g : CompactlySupportedContinuousMap α β) :
              (f g) = fun (x : α) => f x g x
              theorem CompactlySupportedContinuousMap.smulc_apply {α : Type u_2} {β : Type u_3} {γ : Type u_4} [TopologicalSpace α] [TopologicalSpace β] [Zero β] [TopologicalSpace γ] [SMulZeroClass γ β] [ContinuousSMul γ β] {F : Type u_5} [FunLike F α γ] [ContinuousMapClass F α γ] (f : F) (g : CompactlySupportedContinuousMap α β) (x : α) :
              (f g) x = f x g x
              Equations
              Equations
              Equations
              • CompactlySupportedContinuousMap.instAddOfContinuousAdd = { add := fun (f g : CompactlySupportedContinuousMap α β) => { toContinuousMap := f + g, hasCompactSupport' := } }
              Equations

              Coercion to a function as a AddMonoidHom. Similar to AddMonoidHom.coeFn.

              Equations
              Instances For
                Equations
                • One or more equations did not get rendered due to their size.
                @[simp]
                theorem CompactlySupportedContinuousMap.coe_smul {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [Zero β] {R : Type u_5} [SMulZeroClass R β] [ContinuousConstSMul R β] (r : R) (f : CompactlySupportedContinuousMap α β) :
                (r f) = r f
                theorem CompactlySupportedContinuousMap.smul_apply {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [Zero β] {R : Type u_5} [SMulZeroClass R β] [ContinuousConstSMul R β] (r : R) (f : CompactlySupportedContinuousMap α β) (x : α) :
                (r f) x = r f x
                Equations
                Equations
                @[simp]
                theorem CompactlySupportedContinuousMap.coe_sum {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [AddCommMonoid β] [ContinuousAdd β] {ι : Type u_5} (s : Finset ι) (f : ιCompactlySupportedContinuousMap α β) :
                (is, f i) = is, (f i)
                theorem CompactlySupportedContinuousMap.sum_apply {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [AddCommMonoid β] [ContinuousAdd β] {ι : Type u_5} (s : Finset ι) (f : ιCompactlySupportedContinuousMap α β) (a : α) :
                (is, f i) a = is, (f i) a
                Equations
                • CompactlySupportedContinuousMap.instNeg = { neg := fun (f : CompactlySupportedContinuousMap α β) => { toFun := -f.toContinuousMap, continuous_toFun := , hasCompactSupport' := } }
                Equations
                • One or more equations did not get rendered due to their size.
                Equations
                Equations
                Equations
                Equations
                Equations
                • CompactlySupportedContinuousMap.instModuleOfContinuousConstSMul = Function.Injective.module R { toFun := DFunLike.coe, map_zero' := , map_add' := }
                Equations
                Equations
                Equations
                Equations
                Equations
                Equations

                Star structure #

                It is possible to equip C_c(α, β) with a pointwise star operation whenever there is a continuous star : β → β for which star (0 : β) = 0. We don't have quite this weak a typeclass, but StarAddMonoid is close enough.

                The StarAddMonoid class on C_c(α, β) is inherited from their counterparts on α →ᵇ β.

                Equations
                • One or more equations did not get rendered due to their size.
                Equations
                Equations
                • =
                Equations
                • CompactlySupportedContinuousMap.instStarRing = let __src := CompactlySupportedContinuousMap.instStarAddMonoid; StarRing.mk

                C_c as a functor #

                For each β with sufficient structure, there is a contravariant functor C_c(-, β) from the category of topological spaces with morphisms given by CocompactMaps.

                Composition of a continuous function with compact support with a cocompact map yields another continuous function with compact support.

                Equations
                • f.comp g = { toContinuousMap := (f).comp g, hasCompactSupport' := }
                Instances For
                  @[simp]
                  theorem CompactlySupportedContinuousMap.coe_comp_to_continuous_fun {β : Type u_3} {γ : Type u_4} {δ : Type u_5} [TopologicalSpace β] [TopologicalSpace γ] [TopologicalSpace δ] [Zero δ] (f : CompactlySupportedContinuousMap γ δ) (g : CocompactMap β γ) :
                  (f.comp g) = f g
                  @[simp]
                  theorem CompactlySupportedContinuousMap.comp_assoc {α : Type u_2} {β : Type u_3} {γ : Type u_4} [TopologicalSpace α] {δ : Type u_5} [TopologicalSpace β] [TopologicalSpace γ] [TopologicalSpace δ] [Zero δ] (f : CompactlySupportedContinuousMap γ δ) (g : CocompactMap β γ) (h : CocompactMap α β) :
                  (f.comp g).comp h = f.comp (g.comp h)

                  Composition as an additive monoid homomorphism.

                  Equations
                  Instances For

                    Composition as a semigroup homomorphism.

                    Equations
                    Instances For

                      Composition as a linear map.

                      Equations
                      Instances For

                        Composition as a non-unital algebra homomorphism.

                        Equations
                        Instances For
                          Equations
                          • CompactlySupportedContinuousMapClass.instCoeTCCompactlySupportedContinuousMap = { coe := fun (f : F) => { toFun := f, continuous_toFun := , hasCompactSupport' := } }

                          A continuous function on a compact space has automatically compact support. This is not an instance to avoid type class loops.