Documentation

Mathlib.Topology.Algebra.Module.Multilinear.Basic

Continuous multilinear maps #

We define continuous multilinear maps as maps from (i : ι) → M₁ i to M₂ which are multilinear and continuous, by extending the space of multilinear maps with a continuity assumption. Here, M₁ i and M₂ are modules over a ring R, and ι is an arbitrary type, and all these spaces are also topological spaces.

Main definitions #

Implementation notes #

We mostly follow the API of multilinear maps.

Notation #

We introduce the notation M [×n]→L[R] M' for the space of continuous n-multilinear maps from M^n to M'. This is a particular case of the general notion (where we allow varying dependent types as the arguments of our continuous multilinear maps), but arguably the most important one, especially when defining iterated derivatives.

structure ContinuousMultilinearMap (R : Type u) {ι : Type v} (M₁ : ιType w₁) (M₂ : Type w₂) [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] extends MultilinearMap :
Type (max (max v w₁) w₂)

Continuous multilinear maps over the ring R, from ∀ i, M₁ i to M₂ where M₁ i and M₂ are modules over R with a topological structure. In applications, there will be compatibility conditions between the algebraic and the topological structures, but this is not needed for the definition.

  • toFun : ((i : ι) → M₁ i)M₂
  • map_add' : ∀ [inst : DecidableEq ι] (m : (i : ι) → M₁ i) (i : ι) (x y : M₁ i), self.toFun (Function.update m i (x + y)) = self.toFun (Function.update m i x) + self.toFun (Function.update m i y)
  • map_smul' : ∀ [inst : DecidableEq ι] (m : (i : ι) → M₁ i) (i : ι) (c : R) (x : M₁ i), self.toFun (Function.update m i (c x)) = c self.toFun (Function.update m i x)
  • cont : Continuous self.toFun

    Continuous multilinear maps over the ring R, from ∀ i, M₁ i to M₂ where M₁ i and M₂ are modules over R with a topological structure. In applications, there will be compatibility conditions between the algebraic and the topological structures, but this is not needed for the definition.

Instances For
    theorem ContinuousMultilinearMap.cont {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (self : ContinuousMultilinearMap R M₁ M₂) :
    Continuous self.toFun

    Continuous multilinear maps over the ring R, from ∀ i, M₁ i to M₂ where M₁ i and M₂ are modules over R with a topological structure. In applications, there will be compatibility conditions between the algebraic and the topological structures, but this is not needed for the definition.

    Continuous multilinear maps over the ring R, from ∀ i, M₁ i to M₂ where M₁ i and M₂ are modules over R with a topological structure. In applications, there will be compatibility conditions between the algebraic and the topological structures, but this is not needed for the definition.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      theorem ContinuousMultilinearMap.toMultilinearMap_injective {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] :
      Function.Injective ContinuousMultilinearMap.toMultilinearMap
      instance ContinuousMultilinearMap.funLike {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] :
      FunLike (ContinuousMultilinearMap R M₁ M₂) ((i : ι) → M₁ i) M₂
      Equations
      • ContinuousMultilinearMap.funLike = { coe := fun (f : ContinuousMultilinearMap R M₁ M₂) => f.toFun, coe_injective' := }
      instance ContinuousMultilinearMap.continuousMapClass {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] :
      ContinuousMapClass (ContinuousMultilinearMap R M₁ M₂) ((i : ι) → M₁ i) M₂
      Equations
      • =
      instance ContinuousMultilinearMap.instCoeFunForallForall {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] :
      CoeFun (ContinuousMultilinearMap R M₁ M₂) fun (x : ContinuousMultilinearMap R M₁ M₂) => ((i : ι) → M₁ i)M₂
      Equations
      def ContinuousMultilinearMap.Simps.apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (L₁ : ContinuousMultilinearMap R M₁ M₂) (v : (i : ι) → M₁ i) :
      M₂

      See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections.

      Equations
      Instances For
        theorem ContinuousMultilinearMap.coe_continuous {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) :
        @[simp]
        theorem ContinuousMultilinearMap.coe_coe {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) :
        f.toMultilinearMap = f
        theorem ContinuousMultilinearMap.ext {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] {f : ContinuousMultilinearMap R M₁ M₂} {f' : ContinuousMultilinearMap R M₁ M₂} (H : ∀ (x : (i : ι) → M₁ i), f x = f' x) :
        f = f'
        theorem ContinuousMultilinearMap.ext_iff {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] {f : ContinuousMultilinearMap R M₁ M₂} {f' : ContinuousMultilinearMap R M₁ M₂} :
        f = f' ∀ (x : (i : ι) → M₁ i), f x = f' x
        @[simp]
        theorem ContinuousMultilinearMap.map_add {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [DecidableEq ι] (m : (i : ι) → M₁ i) (i : ι) (x : M₁ i) (y : M₁ i) :
        f (Function.update m i (x + y)) = f (Function.update m i x) + f (Function.update m i y)
        @[simp]
        theorem ContinuousMultilinearMap.map_smul {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [DecidableEq ι] (m : (i : ι) → M₁ i) (i : ι) (c : R) (x : M₁ i) :
        f (Function.update m i (c x)) = c f (Function.update m i x)
        theorem ContinuousMultilinearMap.map_coord_zero {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) {m : (i : ι) → M₁ i} (i : ι) (h : m i = 0) :
        f m = 0
        @[simp]
        theorem ContinuousMultilinearMap.map_zero {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [Nonempty ι] :
        f 0 = 0
        instance ContinuousMultilinearMap.instZero {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] :
        Equations
        • ContinuousMultilinearMap.instZero = { zero := let __src := 0; { toMultilinearMap := __src, cont := } }
        instance ContinuousMultilinearMap.instInhabited {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] :
        Equations
        • ContinuousMultilinearMap.instInhabited = { default := 0 }
        @[simp]
        theorem ContinuousMultilinearMap.zero_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (m : (i : ι) → M₁ i) :
        0 m = 0
        @[simp]
        theorem ContinuousMultilinearMap.toMultilinearMap_zero {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] :
        instance ContinuousMultilinearMap.instSMul {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] {R' : Type u_1} {A : Type u_3} [Monoid R'] [Semiring A] [(i : ι) → Module A (M₁ i)] [Module A M₂] [DistribMulAction R' M₂] [ContinuousConstSMul R' M₂] [SMulCommClass A R' M₂] :
        Equations
        • One or more equations did not get rendered due to their size.
        @[simp]
        theorem ContinuousMultilinearMap.smul_apply {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] {R' : Type u_1} {A : Type u_3} [Monoid R'] [Semiring A] [(i : ι) → Module A (M₁ i)] [Module A M₂] [DistribMulAction R' M₂] [ContinuousConstSMul R' M₂] [SMulCommClass A R' M₂] (f : ContinuousMultilinearMap A M₁ M₂) (c : R') (m : (i : ι) → M₁ i) :
        (c f) m = c f m
        @[simp]
        theorem ContinuousMultilinearMap.toMultilinearMap_smul {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] {R' : Type u_1} {A : Type u_3} [Monoid R'] [Semiring A] [(i : ι) → Module A (M₁ i)] [Module A M₂] [DistribMulAction R' M₂] [ContinuousConstSMul R' M₂] [SMulCommClass A R' M₂] (c : R') (f : ContinuousMultilinearMap A M₁ M₂) :
        (c f).toMultilinearMap = c f.toMultilinearMap
        instance ContinuousMultilinearMap.instSMulCommClass {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] {R' : Type u_1} {R'' : Type u_2} {A : Type u_3} [Monoid R'] [Monoid R''] [Semiring A] [(i : ι) → Module A (M₁ i)] [Module A M₂] [DistribMulAction R' M₂] [ContinuousConstSMul R' M₂] [SMulCommClass A R' M₂] [DistribMulAction R'' M₂] [ContinuousConstSMul R'' M₂] [SMulCommClass A R'' M₂] [SMulCommClass R' R'' M₂] :
        Equations
        • =
        instance ContinuousMultilinearMap.instIsScalarTower {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] {R' : Type u_1} {R'' : Type u_2} {A : Type u_3} [Monoid R'] [Monoid R''] [Semiring A] [(i : ι) → Module A (M₁ i)] [Module A M₂] [DistribMulAction R' M₂] [ContinuousConstSMul R' M₂] [SMulCommClass A R' M₂] [DistribMulAction R'' M₂] [ContinuousConstSMul R'' M₂] [SMulCommClass A R'' M₂] [SMul R' R''] [IsScalarTower R' R'' M₂] :
        Equations
        • =
        instance ContinuousMultilinearMap.instIsCentralScalar {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] {R' : Type u_1} {A : Type u_3} [Monoid R'] [Semiring A] [(i : ι) → Module A (M₁ i)] [Module A M₂] [DistribMulAction R' M₂] [ContinuousConstSMul R' M₂] [SMulCommClass A R' M₂] [DistribMulAction R'ᵐᵒᵖ M₂] [IsCentralScalar R' M₂] :
        Equations
        • =
        instance ContinuousMultilinearMap.instMulAction {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] {R' : Type u_1} {A : Type u_3} [Monoid R'] [Semiring A] [(i : ι) → Module A (M₁ i)] [Module A M₂] [DistribMulAction R' M₂] [ContinuousConstSMul R' M₂] [SMulCommClass A R' M₂] :
        Equations
        instance ContinuousMultilinearMap.instAdd {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [ContinuousAdd M₂] :
        Equations
        • ContinuousMultilinearMap.instAdd = { add := fun (f f' : ContinuousMultilinearMap R M₁ M₂) => { toMultilinearMap := f.toMultilinearMap + f'.toMultilinearMap, cont := } }
        @[simp]
        theorem ContinuousMultilinearMap.add_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) (f' : ContinuousMultilinearMap R M₁ M₂) [ContinuousAdd M₂] (m : (i : ι) → M₁ i) :
        (f + f') m = f m + f' m
        @[simp]
        theorem ContinuousMultilinearMap.toMultilinearMap_add {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [ContinuousAdd M₂] (f : ContinuousMultilinearMap R M₁ M₂) (g : ContinuousMultilinearMap R M₁ M₂) :
        (f + g).toMultilinearMap = f.toMultilinearMap + g.toMultilinearMap
        instance ContinuousMultilinearMap.addCommMonoid {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [ContinuousAdd M₂] :
        Equations
        def ContinuousMultilinearMap.applyAddHom {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [ContinuousAdd M₂] (m : (i : ι) → M₁ i) :

        Evaluation of a ContinuousMultilinearMap at a vector as an AddMonoidHom.

        Equations
        Instances For
          @[simp]
          theorem ContinuousMultilinearMap.sum_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [ContinuousAdd M₂] {α : Type u_1} (f : αContinuousMultilinearMap R M₁ M₂) (m : (i : ι) → M₁ i) {s : Finset α} :
          (s.sum fun (a : α) => f a) m = s.sum fun (a : α) => (f a) m
          def ContinuousMultilinearMap.toContinuousLinearMap {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [DecidableEq ι] (m : (i : ι) → M₁ i) (i : ι) :
          M₁ i →L[R] M₂

          If f is a continuous multilinear map, then f.toContinuousLinearMap m i is the continuous linear map obtained by fixing all coordinates but i equal to those of m, and varying the i-th coordinate.

          Equations
          • f.toContinuousLinearMap m i = let __src := f.toLinearMap m i; { toLinearMap := __src, cont := }
          Instances For
            def ContinuousMultilinearMap.prod {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} {M₃ : Type w₃} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [AddCommMonoid M₃] [(i : ι) → Module R (M₁ i)] [Module R M₂] [Module R M₃] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [TopologicalSpace M₃] (f : ContinuousMultilinearMap R M₁ M₂) (g : ContinuousMultilinearMap R M₁ M₃) :
            ContinuousMultilinearMap R M₁ (M₂ × M₃)

            The cartesian product of two continuous multilinear maps, as a continuous multilinear map.

            Equations
            • f.prod g = let __src := f.prod g.toMultilinearMap; { toMultilinearMap := __src, cont := }
            Instances For
              @[simp]
              theorem ContinuousMultilinearMap.prod_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} {M₃ : Type w₃} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [AddCommMonoid M₃] [(i : ι) → Module R (M₁ i)] [Module R M₂] [Module R M₃] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [TopologicalSpace M₃] (f : ContinuousMultilinearMap R M₁ M₂) (g : ContinuousMultilinearMap R M₁ M₃) (m : (i : ι) → M₁ i) :
              (f.prod g) m = (f m, g m)
              def ContinuousMultilinearMap.pi {R : Type u} {ι : Type v} {M₁ : ιType w₁} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [(i : ι) → Module R (M₁ i)] [(i : ι) → TopologicalSpace (M₁ i)] {ι' : Type u_1} {M' : ι'Type u_2} [(i : ι') → AddCommMonoid (M' i)] [(i : ι') → TopologicalSpace (M' i)] [(i : ι') → Module R (M' i)] (f : (i : ι') → ContinuousMultilinearMap R M₁ (M' i)) :
              ContinuousMultilinearMap R M₁ ((i : ι') → M' i)

              Combine a family of continuous multilinear maps with the same domain and codomains M' i into a continuous multilinear map taking values in the space of functions ∀ i, M' i.

              Equations
              Instances For
                @[simp]
                theorem ContinuousMultilinearMap.coe_pi {R : Type u} {ι : Type v} {M₁ : ιType w₁} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [(i : ι) → Module R (M₁ i)] [(i : ι) → TopologicalSpace (M₁ i)] {ι' : Type u_1} {M' : ι'Type u_2} [(i : ι') → AddCommMonoid (M' i)] [(i : ι') → TopologicalSpace (M' i)] [(i : ι') → Module R (M' i)] (f : (i : ι') → ContinuousMultilinearMap R M₁ (M' i)) :
                (ContinuousMultilinearMap.pi f) = fun (m : (i : ι) → M₁ i) (j : ι') => (f j) m
                theorem ContinuousMultilinearMap.pi_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [(i : ι) → Module R (M₁ i)] [(i : ι) → TopologicalSpace (M₁ i)] {ι' : Type u_1} {M' : ι'Type u_2} [(i : ι') → AddCommMonoid (M' i)] [(i : ι') → TopologicalSpace (M' i)] [(i : ι') → Module R (M' i)] (f : (i : ι') → ContinuousMultilinearMap R M₁ (M' i)) (m : (i : ι) → M₁ i) (j : ι') :
                @[simp]
                theorem ContinuousMultilinearMap.codRestrict_toMultilinearMap {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) (p : Submodule R M₂) (h : ∀ (v : (i : ι) → M₁ i), f v p) :
                (f.codRestrict p h).toMultilinearMap = f.codRestrict p h
                @[simp]
                theorem ContinuousMultilinearMap.codRestrict_apply_coe {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) (p : Submodule R M₂) (h : ∀ (v : (i : ι) → M₁ i), f v p) (v : (i : ι) → M₁ i) :
                ((f.codRestrict p h) v) = f v
                def ContinuousMultilinearMap.codRestrict {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) (p : Submodule R M₂) (h : ∀ (v : (i : ι) → M₁ i), f v p) :

                Restrict the codomain of a continuous multilinear map to a submodule.

                Equations
                • f.codRestrict p h = { toMultilinearMap := f.codRestrict p h, cont := }
                Instances For
                  @[simp]
                  theorem ContinuousMultilinearMap.ofSubsingleton_symm_apply_apply (R : Type u) {ι : Type v} (M₂ : Type w₂) (M₃ : Type w₃) [Semiring R] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M₂] [Module R M₃] [TopologicalSpace M₂] [TopologicalSpace M₃] [Subsingleton ι] (i : ι) (f : ContinuousMultilinearMap R (fun (x : ι) => M₂) M₃) (x : M₂) :
                  ((ContinuousMultilinearMap.ofSubsingleton R M₂ M₃ i).symm f) x = f fun (x_1 : ι) => x
                  @[simp]
                  theorem ContinuousMultilinearMap.ofSubsingleton_apply_toMultilinearMap (R : Type u) {ι : Type v} (M₂ : Type w₂) (M₃ : Type w₃) [Semiring R] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M₂] [Module R M₃] [TopologicalSpace M₂] [TopologicalSpace M₃] [Subsingleton ι] (i : ι) (f : M₂ →L[R] M₃) :
                  ((ContinuousMultilinearMap.ofSubsingleton R M₂ M₃ i) f).toMultilinearMap = (MultilinearMap.ofSubsingleton R M₂ M₃ i) f
                  @[simp]
                  theorem ContinuousMultilinearMap.ofSubsingleton_apply_apply (R : Type u) {ι : Type v} (M₂ : Type w₂) (M₃ : Type w₃) [Semiring R] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M₂] [Module R M₃] [TopologicalSpace M₂] [TopologicalSpace M₃] [Subsingleton ι] (i : ι) (f : M₂ →L[R] M₃) (x : ιM₂) :
                  ((ContinuousMultilinearMap.ofSubsingleton R M₂ M₃ i) f) x = f (x i)
                  def ContinuousMultilinearMap.ofSubsingleton (R : Type u) {ι : Type v} (M₂ : Type w₂) (M₃ : Type w₃) [Semiring R] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M₂] [Module R M₃] [TopologicalSpace M₂] [TopologicalSpace M₃] [Subsingleton ι] (i : ι) :
                  (M₂ →L[R] M₃) ContinuousMultilinearMap R (fun (x : ι) => M₂) M₃

                  The natural equivalence between continuous linear maps from M₂ to M₃ and continuous 1-multilinear maps from M₂ to M₃.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem ContinuousMultilinearMap.constOfIsEmpty_apply (R : Type u) {ι : Type v} (M₁ : ιType w₁) {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [IsEmpty ι] (m : M₂) :
                    ∀ (a : (i : ι) → M₁ i), (ContinuousMultilinearMap.constOfIsEmpty R M₁ m) a = m
                    @[simp]
                    theorem ContinuousMultilinearMap.constOfIsEmpty_toMultilinearMap (R : Type u) {ι : Type v} (M₁ : ιType w₁) {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [IsEmpty ι] (m : M₂) :
                    def ContinuousMultilinearMap.constOfIsEmpty (R : Type u) {ι : Type v} (M₁ : ιType w₁) {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [IsEmpty ι] (m : M₂) :

                    The constant map is multilinear when ι is empty.

                    Equations
                    Instances For
                      def ContinuousMultilinearMap.compContinuousLinearMap {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₁' : ιType w₁'} {M₄ : Type w₄} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [(i : ι) → AddCommMonoid (M₁' i)] [AddCommMonoid M₄] [(i : ι) → Module R (M₁ i)] [(i : ι) → Module R (M₁' i)] [Module R M₄] [(i : ι) → TopologicalSpace (M₁ i)] [(i : ι) → TopologicalSpace (M₁' i)] [TopologicalSpace M₄] (g : ContinuousMultilinearMap R M₁' M₄) (f : (i : ι) → M₁ i →L[R] M₁' i) :

                      If g is continuous multilinear and f is a collection of continuous linear maps, then g (f₁ m₁, ..., fₙ mₙ) is again a continuous multilinear map, that we call g.compContinuousLinearMap f.

                      Equations
                      • g.compContinuousLinearMap f = let __src := g.compLinearMap fun (i : ι) => (f i); { toMultilinearMap := __src, cont := }
                      Instances For
                        @[simp]
                        theorem ContinuousMultilinearMap.compContinuousLinearMap_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₁' : ιType w₁'} {M₄ : Type w₄} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [(i : ι) → AddCommMonoid (M₁' i)] [AddCommMonoid M₄] [(i : ι) → Module R (M₁ i)] [(i : ι) → Module R (M₁' i)] [Module R M₄] [(i : ι) → TopologicalSpace (M₁ i)] [(i : ι) → TopologicalSpace (M₁' i)] [TopologicalSpace M₄] (g : ContinuousMultilinearMap R M₁' M₄) (f : (i : ι) → M₁ i →L[R] M₁' i) (m : (i : ι) → M₁ i) :
                        (g.compContinuousLinearMap f) m = g fun (i : ι) => (f i) (m i)
                        def ContinuousLinearMap.compContinuousMultilinearMap {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} {M₃ : Type w₃} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [AddCommMonoid M₃] [(i : ι) → Module R (M₁ i)] [Module R M₂] [Module R M₃] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [TopologicalSpace M₃] (g : M₂ →L[R] M₃) (f : ContinuousMultilinearMap R M₁ M₂) :

                        Composing a continuous multilinear map with a continuous linear map gives again a continuous multilinear map.

                        Equations
                        • g.compContinuousMultilinearMap f = let __src := (g).compMultilinearMap f.toMultilinearMap; { toMultilinearMap := __src, cont := }
                        Instances For
                          @[simp]
                          theorem ContinuousLinearMap.compContinuousMultilinearMap_coe {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} {M₃ : Type w₃} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [AddCommMonoid M₃] [(i : ι) → Module R (M₁ i)] [Module R M₂] [Module R M₃] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [TopologicalSpace M₃] (g : M₂ →L[R] M₃) (f : ContinuousMultilinearMap R M₁ M₂) :
                          (g.compContinuousMultilinearMap f) = g f
                          @[simp]
                          theorem ContinuousMultilinearMap.piEquiv_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [(i : ι) → Module R (M₁ i)] [(i : ι) → TopologicalSpace (M₁ i)] {ι' : Type u_1} {M' : ι'Type u_2} [(i : ι') → AddCommMonoid (M' i)] [(i : ι') → TopologicalSpace (M' i)] [(i : ι') → Module R (M' i)] (f : (i : ι') → ContinuousMultilinearMap R M₁ (M' i)) :
                          ContinuousMultilinearMap.piEquiv f = ContinuousMultilinearMap.pi f
                          @[simp]
                          theorem ContinuousMultilinearMap.piEquiv_symm_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [(i : ι) → Module R (M₁ i)] [(i : ι) → TopologicalSpace (M₁ i)] {ι' : Type u_1} {M' : ι'Type u_2} [(i : ι') → AddCommMonoid (M' i)] [(i : ι') → TopologicalSpace (M' i)] [(i : ι') → Module R (M' i)] (f : ContinuousMultilinearMap R M₁ ((i : ι') → M' i)) (i : ι') :
                          ContinuousMultilinearMap.piEquiv.symm f i = (ContinuousLinearMap.proj i).compContinuousMultilinearMap f
                          def ContinuousMultilinearMap.piEquiv {R : Type u} {ι : Type v} {M₁ : ιType w₁} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [(i : ι) → Module R (M₁ i)] [(i : ι) → TopologicalSpace (M₁ i)] {ι' : Type u_1} {M' : ι'Type u_2} [(i : ι') → AddCommMonoid (M' i)] [(i : ι') → TopologicalSpace (M' i)] [(i : ι') → Module R (M' i)] :
                          ((i : ι') → ContinuousMultilinearMap R M₁ (M' i)) ContinuousMultilinearMap R M₁ ((i : ι') → M' i)

                          ContinuousMultilinearMap.pi as an Equiv.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          Instances For
                            @[simp]
                            theorem ContinuousMultilinearMap.domDomCongr_apply {R : Type u} {ι : Type v} {M₂ : Type w₂} {M₃ : Type w₃} [Semiring R] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M₂] [Module R M₃] [TopologicalSpace M₂] [TopologicalSpace M₃] {ι' : Type u_1} (e : ι ι') (f : ContinuousMultilinearMap R (fun (x : ι) => M₂) M₃) (v : ι'M₂) :
                            (ContinuousMultilinearMap.domDomCongr e f) v = f fun (i : ι) => v (e i)
                            @[simp]
                            theorem ContinuousMultilinearMap.domDomCongr_toMultilinearMap {R : Type u} {ι : Type v} {M₂ : Type w₂} {M₃ : Type w₃} [Semiring R] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M₂] [Module R M₃] [TopologicalSpace M₂] [TopologicalSpace M₃] {ι' : Type u_1} (e : ι ι') (f : ContinuousMultilinearMap R (fun (x : ι) => M₂) M₃) :
                            (ContinuousMultilinearMap.domDomCongr e f).toMultilinearMap = MultilinearMap.domDomCongr e f.toMultilinearMap
                            def ContinuousMultilinearMap.domDomCongr {R : Type u} {ι : Type v} {M₂ : Type w₂} {M₃ : Type w₃} [Semiring R] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M₂] [Module R M₃] [TopologicalSpace M₂] [TopologicalSpace M₃] {ι' : Type u_1} (e : ι ι') (f : ContinuousMultilinearMap R (fun (x : ι) => M₂) M₃) :
                            ContinuousMultilinearMap R (fun (x : ι') => M₂) M₃

                            An equivalence of the index set defines an equivalence between the spaces of continuous multilinear maps. This is the forward map of this equivalence.

                            Equations
                            Instances For
                              @[simp]
                              theorem ContinuousMultilinearMap.domDomCongrEquiv_symm_apply {R : Type u} {ι : Type v} {M₂ : Type w₂} {M₃ : Type w₃} [Semiring R] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M₂] [Module R M₃] [TopologicalSpace M₂] [TopologicalSpace M₃] {ι' : Type u_1} (e : ι ι') (f : ContinuousMultilinearMap R (fun (x : ι') => M₂) M₃) :
                              @[simp]
                              theorem ContinuousMultilinearMap.domDomCongrEquiv_apply {R : Type u} {ι : Type v} {M₂ : Type w₂} {M₃ : Type w₃} [Semiring R] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M₂] [Module R M₃] [TopologicalSpace M₂] [TopologicalSpace M₃] {ι' : Type u_1} (e : ι ι') (f : ContinuousMultilinearMap R (fun (x : ι) => M₂) M₃) :
                              def ContinuousMultilinearMap.domDomCongrEquiv {R : Type u} {ι : Type v} {M₂ : Type w₂} {M₃ : Type w₃} [Semiring R] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M₂] [Module R M₃] [TopologicalSpace M₂] [TopologicalSpace M₃] {ι' : Type u_1} (e : ι ι') :
                              ContinuousMultilinearMap R (fun (x : ι) => M₂) M₃ ContinuousMultilinearMap R (fun (x : ι') => M₂) M₃

                              An equivalence of the index set defines an equivalence between the spaces of continuous multilinear maps. In case of normed spaces, this is a linear isometric equivalence, see ContinuousMultilinearMap.domDomCongrₗᵢ.

                              Equations
                              Instances For
                                def ContinuousMultilinearMap.linearDeriv {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [ContinuousAdd M₂] [DecidableEq ι] [Fintype ι] (x : (i : ι) → M₁ i) :
                                ((i : ι) → M₁ i) →L[R] M₂

                                The derivative of a continuous multilinear map, as a continuous linear map from ∀ i, M₁ i to M₂; see ContinuousMultilinearMap.hasFDerivAt.

                                Equations
                                Instances For
                                  @[simp]
                                  theorem ContinuousMultilinearMap.linearDeriv_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [ContinuousAdd M₂] [DecidableEq ι] [Fintype ι] (x : (i : ι) → M₁ i) (y : (i : ι) → M₁ i) :
                                  (f.linearDeriv x) y = Finset.univ.sum fun (i : ι) => f (Function.update x i (y i))
                                  theorem ContinuousMultilinearMap.cons_add {R : Type u} {n : } {M : Fin n.succType w} {M₂ : Type w₂} [Semiring R] [(i : Fin n.succ) → AddCommMonoid (M i)] [AddCommMonoid M₂] [(i : Fin n.succ) → Module R (M i)] [Module R M₂] [(i : Fin n.succ) → TopologicalSpace (M i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M M₂) (m : (i : Fin n) → M i.succ) (x : M 0) (y : M 0) :
                                  f (Fin.cons (x + y) m) = f (Fin.cons x m) + f (Fin.cons y m)

                                  In the specific case of continuous multilinear maps on spaces indexed by Fin (n+1), where one can build an element of (i : Fin (n+1)) → M i using cons, one can express directly the additivity of a multilinear map along the first variable.

                                  theorem ContinuousMultilinearMap.cons_smul {R : Type u} {n : } {M : Fin n.succType w} {M₂ : Type w₂} [Semiring R] [(i : Fin n.succ) → AddCommMonoid (M i)] [AddCommMonoid M₂] [(i : Fin n.succ) → Module R (M i)] [Module R M₂] [(i : Fin n.succ) → TopologicalSpace (M i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M M₂) (m : (i : Fin n) → M i.succ) (c : R) (x : M 0) :
                                  f (Fin.cons (c x) m) = c f (Fin.cons x m)

                                  In the specific case of continuous multilinear maps on spaces indexed by Fin (n+1), where one can build an element of (i : Fin (n+1)) → M i using cons, one can express directly the multiplicativity of a multilinear map along the first variable.

                                  theorem ContinuousMultilinearMap.map_piecewise_add {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [DecidableEq ι] (m : (i : ι) → M₁ i) (m' : (i : ι) → M₁ i) (t : Finset ι) :
                                  f (t.piecewise (m + m') m') = t.powerset.sum fun (s : Finset ι) => f (s.piecewise m m')
                                  theorem ContinuousMultilinearMap.map_add_univ {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [DecidableEq ι] [Fintype ι] (m : (i : ι) → M₁ i) (m' : (i : ι) → M₁ i) :
                                  f (m + m') = Finset.univ.sum fun (s : Finset ι) => f (s.piecewise m m')

                                  Additivity of a continuous multilinear map along all coordinates at the same time, writing f (m + m') as the sum of f (s.piecewise m m') over all sets s.

                                  theorem ContinuousMultilinearMap.map_sum_finset {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) {α : ιType u_1} [Fintype ι] (g : (i : ι) → α iM₁ i) (A : (i : ι) → Finset (α i)) [DecidableEq ι] :
                                  (f fun (i : ι) => (A i).sum fun (j : α i) => g i j) = (Fintype.piFinset A).sum fun (r : (a : ι) → α a) => f fun (i : ι) => g i (r i)

                                  If f is continuous multilinear, then f (Σ_{j₁ ∈ A₁} g₁ j₁, ..., Σ_{jₙ ∈ Aₙ} gₙ jₙ) is the sum of f (g₁ (r 1), ..., gₙ (r n)) where r ranges over all functions with r 1 ∈ A₁, ..., r n ∈ Aₙ. This follows from multilinearity by expanding successively with respect to each coordinate.

                                  theorem ContinuousMultilinearMap.map_sum {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) {α : ιType u_1} [Fintype ι] (g : (i : ι) → α iM₁ i) [DecidableEq ι] [(i : ι) → Fintype (α i)] :
                                  (f fun (i : ι) => Finset.univ.sum fun (j : α i) => g i j) = Finset.univ.sum fun (r : (i : ι) → α i) => f fun (i : ι) => g i (r i)

                                  If f is continuous multilinear, then f (Σ_{j₁} g₁ j₁, ..., Σ_{jₙ} gₙ jₙ) is the sum of f (g₁ (r 1), ..., gₙ (r n)) where r ranges over all functions r. This follows from multilinearity by expanding successively with respect to each coordinate.

                                  def ContinuousMultilinearMap.restrictScalars (R : Type u) {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] {A : Type u_1} [Semiring A] [SMul R A] [(i : ι) → Module A (M₁ i)] [Module A M₂] [∀ (i : ι), IsScalarTower R A (M₁ i)] [IsScalarTower R A M₂] (f : ContinuousMultilinearMap A M₁ M₂) :

                                  Reinterpret an A-multilinear map as an R-multilinear map, if A is an algebra over R and their actions on all involved modules agree with the action of R on A.

                                  Equations
                                  Instances For
                                    @[simp]
                                    theorem ContinuousMultilinearMap.coe_restrictScalars (R : Type u) {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Semiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] {A : Type u_1} [Semiring A] [SMul R A] [(i : ι) → Module A (M₁ i)] [Module A M₂] [∀ (i : ι), IsScalarTower R A (M₁ i)] [IsScalarTower R A M₂] (f : ContinuousMultilinearMap A M₁ M₂) :
                                    @[simp]
                                    theorem ContinuousMultilinearMap.map_sub {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Ring R] [(i : ι) → AddCommGroup (M₁ i)] [AddCommGroup M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [DecidableEq ι] (m : (i : ι) → M₁ i) (i : ι) (x : M₁ i) (y : M₁ i) :
                                    f (Function.update m i (x - y)) = f (Function.update m i x) - f (Function.update m i y)
                                    instance ContinuousMultilinearMap.instNeg {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Ring R] [(i : ι) → AddCommGroup (M₁ i)] [AddCommGroup M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [TopologicalAddGroup M₂] :
                                    Equations
                                    • ContinuousMultilinearMap.instNeg = { neg := fun (f : ContinuousMultilinearMap R M₁ M₂) => let __src := -f.toMultilinearMap; { toMultilinearMap := __src, cont := } }
                                    @[simp]
                                    theorem ContinuousMultilinearMap.neg_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Ring R] [(i : ι) → AddCommGroup (M₁ i)] [AddCommGroup M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [TopologicalAddGroup M₂] (m : (i : ι) → M₁ i) :
                                    (-f) m = -f m
                                    instance ContinuousMultilinearMap.instSub {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Ring R] [(i : ι) → AddCommGroup (M₁ i)] [AddCommGroup M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [TopologicalAddGroup M₂] :
                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    @[simp]
                                    theorem ContinuousMultilinearMap.sub_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Ring R] [(i : ι) → AddCommGroup (M₁ i)] [AddCommGroup M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) (f' : ContinuousMultilinearMap R M₁ M₂) [TopologicalAddGroup M₂] (m : (i : ι) → M₁ i) :
                                    (f - f') m = f m - f' m
                                    instance ContinuousMultilinearMap.instAddCommGroup {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [Ring R] [(i : ι) → AddCommGroup (M₁ i)] [AddCommGroup M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [TopologicalAddGroup M₂] :
                                    Equations
                                    theorem ContinuousMultilinearMap.map_piecewise_smul {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [CommSemiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [DecidableEq ι] (c : ιR) (m : (i : ι) → M₁ i) (s : Finset ι) :
                                    f (s.piecewise (fun (i : ι) => c i m i) m) = (s.prod fun (i : ι) => c i) f m
                                    theorem ContinuousMultilinearMap.map_smul_univ {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [CommSemiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] (f : ContinuousMultilinearMap R M₁ M₂) [Fintype ι] (c : ιR) (m : (i : ι) → M₁ i) :
                                    (f fun (i : ι) => c i m i) = (Finset.univ.prod fun (i : ι) => c i) f m

                                    Multiplicativity of a continuous multilinear map along all coordinates at the same time, writing f (fun i ↦ c i • m i) as (∏ i, c i) • f m.

                                    instance ContinuousMultilinearMap.instDistribMulAction {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} {R' : Type u_1} {A : Type u_3} [Monoid R'] [Semiring A] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [(i : ι) → Module A (M₁ i)] [Module A M₂] [DistribMulAction R' M₂] [ContinuousConstSMul R' M₂] [SMulCommClass A R' M₂] [ContinuousAdd M₂] :
                                    Equations
                                    • ContinuousMultilinearMap.instDistribMulAction = Function.Injective.distribMulAction { toFun := ContinuousMultilinearMap.toMultilinearMap, map_zero' := , map_add' := }
                                    instance ContinuousMultilinearMap.instModule {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} {R' : Type u_1} {A : Type u_2} [Semiring R'] [Semiring A] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [ContinuousAdd M₂] [(i : ι) → Module A (M₁ i)] [Module A M₂] [Module R' M₂] [ContinuousConstSMul R' M₂] [SMulCommClass A R' M₂] :

                                    The space of continuous multilinear maps over an algebra over R is a module over R, for the pointwise addition and scalar multiplication.

                                    Equations
                                    • ContinuousMultilinearMap.instModule = Function.Injective.module R' { toFun := ContinuousMultilinearMap.toMultilinearMap, map_zero' := , map_add' := }
                                    @[simp]
                                    theorem ContinuousMultilinearMap.toMultilinearMapLinear_apply {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} {R' : Type u_1} {A : Type u_2} [Semiring R'] [Semiring A] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [ContinuousAdd M₂] [(i : ι) → Module A (M₁ i)] [Module A M₂] [Module R' M₂] [ContinuousConstSMul R' M₂] [SMulCommClass A R' M₂] (self : ContinuousMultilinearMap A M₁ M₂) :
                                    ContinuousMultilinearMap.toMultilinearMapLinear self = self.toMultilinearMap
                                    def ContinuousMultilinearMap.toMultilinearMapLinear {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} {R' : Type u_1} {A : Type u_2} [Semiring R'] [Semiring A] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [ContinuousAdd M₂] [(i : ι) → Module A (M₁ i)] [Module A M₂] [Module R' M₂] [ContinuousConstSMul R' M₂] [SMulCommClass A R' M₂] :

                                    Linear map version of the map toMultilinearMap associating to a continuous multilinear map the corresponding multilinear map.

                                    Equations
                                    • ContinuousMultilinearMap.toMultilinearMapLinear = { toFun := ContinuousMultilinearMap.toMultilinearMap, map_add' := , map_smul' := }
                                    Instances For
                                      @[simp]
                                      theorem ContinuousMultilinearMap.piLinearEquiv_apply {ι : Type v} {M₁ : ιType w₁} {R' : Type u_1} {A : Type u_2} [Semiring R'] [Semiring A] [(i : ι) → AddCommMonoid (M₁ i)] [(i : ι) → TopologicalSpace (M₁ i)] [(i : ι) → Module A (M₁ i)] {ι' : Type u_3} {M' : ι'Type u_4} [(i : ι') → AddCommMonoid (M' i)] [(i : ι') → TopologicalSpace (M' i)] [∀ (i : ι'), ContinuousAdd (M' i)] [(i : ι') → Module R' (M' i)] [(i : ι') → Module A (M' i)] [∀ (i : ι'), SMulCommClass A R' (M' i)] [∀ (i : ι'), ContinuousConstSMul R' (M' i)] :
                                      ∀ (a : (i : ι') → ContinuousMultilinearMap A M₁ (M' i)), ContinuousMultilinearMap.piLinearEquiv a = ContinuousMultilinearMap.pi a
                                      @[simp]
                                      theorem ContinuousMultilinearMap.piLinearEquiv_symm_apply {ι : Type v} {M₁ : ιType w₁} {R' : Type u_1} {A : Type u_2} [Semiring R'] [Semiring A] [(i : ι) → AddCommMonoid (M₁ i)] [(i : ι) → TopologicalSpace (M₁ i)] [(i : ι) → Module A (M₁ i)] {ι' : Type u_3} {M' : ι'Type u_4} [(i : ι') → AddCommMonoid (M' i)] [(i : ι') → TopologicalSpace (M' i)] [∀ (i : ι'), ContinuousAdd (M' i)] [(i : ι') → Module R' (M' i)] [(i : ι') → Module A (M' i)] [∀ (i : ι'), SMulCommClass A R' (M' i)] [∀ (i : ι'), ContinuousConstSMul R' (M' i)] :
                                      ∀ (a : ContinuousMultilinearMap A M₁ ((i : ι') → M' i)) (i : ι'), ContinuousMultilinearMap.piLinearEquiv.symm a i = (ContinuousLinearMap.proj i).compContinuousMultilinearMap a
                                      def ContinuousMultilinearMap.piLinearEquiv {ι : Type v} {M₁ : ιType w₁} {R' : Type u_1} {A : Type u_2} [Semiring R'] [Semiring A] [(i : ι) → AddCommMonoid (M₁ i)] [(i : ι) → TopologicalSpace (M₁ i)] [(i : ι) → Module A (M₁ i)] {ι' : Type u_3} {M' : ι'Type u_4} [(i : ι') → AddCommMonoid (M' i)] [(i : ι') → TopologicalSpace (M' i)] [∀ (i : ι'), ContinuousAdd (M' i)] [(i : ι') → Module R' (M' i)] [(i : ι') → Module A (M' i)] [∀ (i : ι'), SMulCommClass A R' (M' i)] [∀ (i : ι'), ContinuousConstSMul R' (M' i)] :
                                      ((i : ι') → ContinuousMultilinearMap A M₁ (M' i)) ≃ₗ[R'] ContinuousMultilinearMap A M₁ ((i : ι') → M' i)

                                      ContinuousMultilinearMap.pi as a LinearEquiv.

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For
                                        def ContinuousMultilinearMap.mkPiAlgebra (R : Type u) (ι : Type v) (A : Type u_1) [Fintype ι] [CommSemiring R] [CommSemiring A] [Algebra R A] [TopologicalSpace A] [ContinuousMul A] :
                                        ContinuousMultilinearMap R (fun (x : ι) => A) A

                                        The continuous multilinear map on A^ι, where A is a normed commutative algebra over 𝕜, associating to m the product of all the m i.

                                        See also ContinuousMultilinearMap.mkPiAlgebraFin.

                                        Equations
                                        Instances For
                                          @[simp]
                                          theorem ContinuousMultilinearMap.mkPiAlgebra_apply (R : Type u) (ι : Type v) (A : Type u_1) [Fintype ι] [CommSemiring R] [CommSemiring A] [Algebra R A] [TopologicalSpace A] [ContinuousMul A] (m : ιA) :
                                          (ContinuousMultilinearMap.mkPiAlgebra R ι A) m = Finset.univ.prod fun (i : ι) => m i

                                          The continuous multilinear map on A^n, where A is a normed algebra over 𝕜, associating to m the product of all the m i.

                                          See also: ContinuousMultilinearMap.mkPiAlgebra.

                                          Equations
                                          Instances For
                                            @[simp]
                                            theorem ContinuousMultilinearMap.smulRight_apply {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [CommSemiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [TopologicalSpace R] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [ContinuousSMul R M₂] (f : ContinuousMultilinearMap R M₁ R) (z : M₂) :
                                            ∀ (a : (i : ι) → M₁ i), (f.smulRight z) a = f a z
                                            @[simp]
                                            theorem ContinuousMultilinearMap.smulRight_toMultilinearMap {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [CommSemiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [TopologicalSpace R] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [ContinuousSMul R M₂] (f : ContinuousMultilinearMap R M₁ R) (z : M₂) :
                                            (f.smulRight z).toMultilinearMap = f.smulRight z
                                            def ContinuousMultilinearMap.smulRight {R : Type u} {ι : Type v} {M₁ : ιType w₁} {M₂ : Type w₂} [CommSemiring R] [(i : ι) → AddCommMonoid (M₁ i)] [AddCommMonoid M₂] [(i : ι) → Module R (M₁ i)] [Module R M₂] [TopologicalSpace R] [(i : ι) → TopologicalSpace (M₁ i)] [TopologicalSpace M₂] [ContinuousSMul R M₂] (f : ContinuousMultilinearMap R M₁ R) (z : M₂) :

                                            Given a continuous R-multilinear map f taking values in R, f.smulRight z is the continuous multilinear map sending m to f m • z.

                                            Equations
                                            • f.smulRight z = { toMultilinearMap := f.smulRight z, cont := }
                                            Instances For
                                              def ContinuousMultilinearMap.mkPiRing (R : Type u) (ι : Type v) {M : Type u_1} [Fintype ι] [CommRing R] [AddCommMonoid M] [Module R M] [TopologicalSpace R] [TopologicalSpace M] [ContinuousMul R] [ContinuousSMul R M] (z : M) :
                                              ContinuousMultilinearMap R (fun (x : ι) => R) M

                                              The canonical continuous multilinear map on R^ι, associating to m the product of all the m i (multiplied by a fixed reference element z in the target module)

                                              Equations
                                              Instances For
                                                @[simp]
                                                theorem ContinuousMultilinearMap.mkPiRing_apply {R : Type u} {ι : Type v} {M : Type u_1} [Fintype ι] [CommRing R] [AddCommMonoid M] [Module R M] [TopologicalSpace R] [TopologicalSpace M] [ContinuousMul R] [ContinuousSMul R M] (z : M) (m : ιR) :
                                                (ContinuousMultilinearMap.mkPiRing R ι z) m = (Finset.univ.prod fun (i : ι) => m i) z