Documentation

Mathlib.Analysis.NormedSpace.ContinuousAffineMap

Continuous affine maps between normed spaces. #

This file develops the theory of continuous affine maps between affine spaces modelled on normed spaces.

In the particular case that the affine spaces are just normed vector spaces V, W, we define a norm on the space of continuous affine maps by defining the norm of f : V β†’A[π•œ] W to be β€–fβ€– = max β€–f 0β€– β€–f.cont_linearβ€–. This is chosen so that we have a linear isometry: (V β†’A[π•œ] W) ≃ₗᡒ[π•œ] W Γ— (V β†’L[π•œ] W).

The abstract picture is that for an affine space P modelled on a vector space V, together with a vector space W, there is an exact sequence of π•œ-modules: 0 β†’ C β†’ A β†’ L β†’ 0 where C, A are the spaces of constant and affine maps P β†’ W and L is the space of linear maps V β†’ W.

Any choice of a base point in P corresponds to a splitting of this sequence so in particular if we take P = V, using 0 : V as the base point provides a splitting, and we prove this is an isometric decomposition.

On the other hand, choosing a base point breaks the affine invariance so the norm fails to be submultiplicative: for a composition of maps, we have only β€–f.comp gβ€– ≀ β€–fβ€– * β€–gβ€– + β€–f 0β€–.

Main definitions: #

def ContinuousAffineMap.contLinear {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} {Q : Type u_7} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [MetricSpace Q] [NormedAddTorsor W Q] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : P β†’A[R] Q) :

The linear map underlying a continuous affine map is continuous.

Equations
  • f.contLinear = let __src := f.linear; { toFun := ⇑f.linear, map_add' := β‹―, map_smul' := β‹―, cont := β‹― }
Instances For
    @[simp]
    theorem ContinuousAffineMap.coe_contLinear {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} {Q : Type u_7} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [MetricSpace Q] [NormedAddTorsor W Q] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : P β†’A[R] Q) :
    ⇑f.contLinear = ⇑f.linear
    @[simp]
    theorem ContinuousAffineMap.coe_contLinear_eq_linear {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} {Q : Type u_7} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [MetricSpace Q] [NormedAddTorsor W Q] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : P β†’A[R] Q) :
    ↑f.contLinear = f.linear
    @[simp]
    theorem ContinuousAffineMap.coe_mk_const_linear_eq_linear {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} {Q : Type u_7} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [MetricSpace Q] [NormedAddTorsor W Q] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : P →ᡃ[R] Q) (h : Continuous f.toFun) :
    ⇑{ toAffineMap := f, cont := h }.contLinear = ⇑f.linear
    theorem ContinuousAffineMap.coe_linear_eq_coe_contLinear {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} {Q : Type u_7} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [MetricSpace Q] [NormedAddTorsor W Q] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : P β†’A[R] Q) :
    ⇑f.linear = ⇑f.contLinear
    @[simp]
    theorem ContinuousAffineMap.comp_contLinear {R : Type u_2} {V : Type u_3} {W : Type u_4} {Wβ‚‚ : Type u_5} {P : Type u_6} {Q : Type u_7} {Qβ‚‚ : Type u_8} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [MetricSpace Q] [NormedAddTorsor W Q] [NormedAddCommGroup Wβ‚‚] [MetricSpace Qβ‚‚] [NormedAddTorsor Wβ‚‚ Qβ‚‚] [NormedField R] [NormedSpace R V] [NormedSpace R W] [NormedSpace R Wβ‚‚] (f : P β†’A[R] Q) (g : Q β†’A[R] Qβ‚‚) :
    (g.comp f).contLinear = g.contLinear.comp f.contLinear
    @[simp]
    theorem ContinuousAffineMap.map_vadd {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} {Q : Type u_7} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [MetricSpace Q] [NormedAddTorsor W Q] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : P β†’A[R] Q) (p : P) (v : V) :
    f (v +α΅₯ p) = f.contLinear v +α΅₯ f p
    @[simp]
    theorem ContinuousAffineMap.contLinear_map_vsub {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} {Q : Type u_7} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [MetricSpace Q] [NormedAddTorsor W Q] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : P β†’A[R] Q) (p₁ : P) (pβ‚‚ : P) :
    f.contLinear (p₁ -α΅₯ pβ‚‚) = f p₁ -α΅₯ f pβ‚‚
    @[simp]
    theorem ContinuousAffineMap.const_contLinear {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} {Q : Type u_7} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [MetricSpace Q] [NormedAddTorsor W Q] [NormedField R] [NormedSpace R V] [NormedSpace R W] (q : Q) :
    (ContinuousAffineMap.const R P q).contLinear = 0
    theorem ContinuousAffineMap.contLinear_eq_zero_iff_exists_const {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} {Q : Type u_7} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [MetricSpace Q] [NormedAddTorsor W Q] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : P β†’A[R] Q) :
    f.contLinear = 0 ↔ βˆƒ (q : Q), f = ContinuousAffineMap.const R P q
    @[simp]
    theorem ContinuousAffineMap.to_affine_map_contLinear {R : Type u_2} {V : Type u_3} {W : Type u_4} [NormedAddCommGroup V] [NormedAddCommGroup W] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : V β†’L[R] W) :
    f.toContinuousAffineMap.contLinear = f
    @[simp]
    theorem ContinuousAffineMap.add_contLinear {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : P β†’A[R] W) (g : P β†’A[R] W) :
    (f + g).contLinear = f.contLinear + g.contLinear
    @[simp]
    theorem ContinuousAffineMap.sub_contLinear {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : P β†’A[R] W) (g : P β†’A[R] W) :
    (f - g).contLinear = f.contLinear - g.contLinear
    @[simp]
    theorem ContinuousAffineMap.neg_contLinear {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : P β†’A[R] W) :
    (-f).contLinear = -f.contLinear
    @[simp]
    theorem ContinuousAffineMap.smul_contLinear {R : Type u_2} {V : Type u_3} {W : Type u_4} {P : Type u_6} [NormedAddCommGroup V] [MetricSpace P] [NormedAddTorsor V P] [NormedAddCommGroup W] [NormedField R] [NormedSpace R V] [NormedSpace R W] (t : R) (f : P β†’A[R] W) :
    (t β€’ f).contLinear = t β€’ f.contLinear
    theorem ContinuousAffineMap.decomp {R : Type u_2} {V : Type u_3} {W : Type u_4} [NormedAddCommGroup V] [NormedAddCommGroup W] [NormedField R] [NormedSpace R V] [NormedSpace R W] (f : V β†’A[R] W) :
    ⇑f = ⇑f.contLinear + Function.const V (f 0)
    noncomputable instance ContinuousAffineMap.hasNorm {π•œ : Type u_1} {V : Type u_3} {W : Type u_4} [NormedAddCommGroup V] [NormedAddCommGroup W] [NontriviallyNormedField π•œ] [NormedSpace π•œ V] [NormedSpace π•œ W] :
    Norm (V β†’A[π•œ] W)

    Note that unlike the operator norm for linear maps, this norm is not submultiplicative: we do not necessarily have β€–f.comp gβ€– ≀ β€–fβ€– * β€–gβ€–. See norm_comp_le for what we can say.

    Equations
    theorem ContinuousAffineMap.norm_def {π•œ : Type u_1} {V : Type u_3} {W : Type u_4} [NormedAddCommGroup V] [NormedAddCommGroup W] [NontriviallyNormedField π•œ] [NormedSpace π•œ V] [NormedSpace π•œ W] (f : V β†’A[π•œ] W) :
    theorem ContinuousAffineMap.norm_contLinear_le {π•œ : Type u_1} {V : Type u_3} {W : Type u_4} [NormedAddCommGroup V] [NormedAddCommGroup W] [NontriviallyNormedField π•œ] [NormedSpace π•œ V] [NormedSpace π•œ W] (f : V β†’A[π•œ] W) :
    theorem ContinuousAffineMap.norm_image_zero_le {π•œ : Type u_1} {V : Type u_3} {W : Type u_4} [NormedAddCommGroup V] [NormedAddCommGroup W] [NontriviallyNormedField π•œ] [NormedSpace π•œ V] [NormedSpace π•œ W] (f : V β†’A[π•œ] W) :
    @[simp]
    theorem ContinuousAffineMap.norm_eq {π•œ : Type u_1} {V : Type u_3} {W : Type u_4} [NormedAddCommGroup V] [NormedAddCommGroup W] [NontriviallyNormedField π•œ] [NormedSpace π•œ V] [NormedSpace π•œ W] (f : V β†’A[π•œ] W) (h : f 0 = 0) :
    noncomputable instance ContinuousAffineMap.instNormedAddCommGroup {π•œ : Type u_1} {V : Type u_3} {W : Type u_4} [NormedAddCommGroup V] [NormedAddCommGroup W] [NontriviallyNormedField π•œ] [NormedSpace π•œ V] [NormedSpace π•œ W] :
    Equations
    • One or more equations did not get rendered due to their size.
    instance ContinuousAffineMap.instNormedSpace {π•œ : Type u_1} {V : Type u_3} {W : Type u_4} [NormedAddCommGroup V] [NormedAddCommGroup W] [NontriviallyNormedField π•œ] [NormedSpace π•œ V] [NormedSpace π•œ W] :
    NormedSpace π•œ (V β†’A[π•œ] W)
    Equations
    theorem ContinuousAffineMap.norm_comp_le {π•œ : Type u_1} {V : Type u_3} {W : Type u_4} {Wβ‚‚ : Type u_5} [NormedAddCommGroup V] [NormedAddCommGroup W] [NormedAddCommGroup Wβ‚‚] [NontriviallyNormedField π•œ] [NormedSpace π•œ V] [NormedSpace π•œ W] [NormedSpace π•œ Wβ‚‚] (f : V β†’A[π•œ] W) (g : Wβ‚‚ β†’A[π•œ] V) :
    def ContinuousAffineMap.toConstProdContinuousLinearMap (π•œ : Type u_1) (V : Type u_3) (W : Type u_4) [NormedAddCommGroup V] [NormedAddCommGroup W] [NontriviallyNormedField π•œ] [NormedSpace π•œ V] [NormedSpace π•œ W] :
    (V β†’A[π•œ] W) ≃ₗᡒ[π•œ] W Γ— (V β†’L[π•œ] W)

    The space of affine maps between two normed spaces is linearly isometric to the product of the codomain with the space of linear maps, by taking the value of the affine map at (0 : V) and the linear part.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem ContinuousAffineMap.toConstProdContinuousLinearMap_fst (π•œ : Type u_1) (V : Type u_3) (W : Type u_4) [NormedAddCommGroup V] [NormedAddCommGroup W] [NontriviallyNormedField π•œ] [NormedSpace π•œ V] [NormedSpace π•œ W] (f : V β†’A[π•œ] W) :
      @[simp]
      theorem ContinuousAffineMap.toConstProdContinuousLinearMap_snd (π•œ : Type u_1) (V : Type u_3) (W : Type u_4) [NormedAddCommGroup V] [NormedAddCommGroup W] [NontriviallyNormedField π•œ] [NormedSpace π•œ V] [NormedSpace π•œ W] (f : V β†’A[π•œ] W) :
      ((ContinuousAffineMap.toConstProdContinuousLinearMap π•œ V W) f).2 = f.contLinear