Documentation

Mathlib.Analysis.Normed.Unbundled.InvariantExtension

algNormOfAlgEquiv and invariantExtension #

Let K be a nonarchimedean normed field and L/K be a finite algebraic extension. In the comments, ‖ ⬝ ‖ denotes any power-multiplicative K-algebra norm on L extending the norm on K.

Main Definitions #

Main Results #

References #

Tags #

algNormOfAlgEquiv, invariantExtension, norm, nonarchimedean

def IsUltrametricDist.algNormOfAlgEquiv {K : Type u_1} [NormedField K] {L : Type u_2} [Field L] [Algebra K L] [h_fin : FiniteDimensional K L] [hu : IsUltrametricDist K] (σ : L ≃ₐ[K] L) :

Given a normed field K, a finite algebraic extension L/K and σ : L ≃ₐ[K] L, the function L → ℝ sending x : L to ‖ σ x ‖, where ‖ ⬝ ‖ is any power-multiplicative algebra norm on L extending the norm on K, is an algebra norm on K.

Equations
Instances For
    theorem IsUltrametricDist.algNormOfAlgEquiv_apply {K : Type u_1} [NormedField K] {L : Type u_2} [Field L] [Algebra K L] [h_fin : FiniteDimensional K L] [hu : IsUltrametricDist K] (σ : L ≃ₐ[K] L) (x : L) :
    theorem IsUltrametricDist.isPowMul_algNormOfAlgEquiv {K : Type u_1} [NormedField K] {L : Type u_2} [Field L] [Algebra K L] [h_fin : FiniteDimensional K L] [hu : IsUltrametricDist K] (σ : L ≃ₐ[K] L) :

    The algebra norm algNormOfAlgEquiv is power-multiplicative.

    The algebra norm algNormOfAlgEquiv is nonarchimedean.

    theorem IsUltrametricDist.algNormOfAlgEquiv_extends {K : Type u_1} [NormedField K] {L : Type u_2} [Field L] [Algebra K L] [h_fin : FiniteDimensional K L] [hu : IsUltrametricDist K] (σ : L ≃ₐ[K] L) (x : K) :

    The algebra norm algNormOfAlgEquiv extends the norm on K.

    The function L → ℝ sending x : L to the maximum of algNormOfAlgEquiv hna σ over all σ : L ≃ₐ[K] L is an algebra norm on L.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem IsUltrametricDist.invariantExtension_apply (K : Type u_1) [NormedField K] (L : Type u_2) [Field L] [Algebra K L] [h_fin : FiniteDimensional K L] [hu : IsUltrametricDist K] (x : L) :
      (invariantExtension K L) x = ⨆ (σ : L ≃ₐ[K] L), (algNormOfAlgEquiv σ) x

      The algebra norm invariantExtension is power-multiplicative.

      The algebra norm invariantExtension is nonarchimedean.

      theorem IsUltrametricDist.invariantExtension_extends (K : Type u_1) [NormedField K] (L : Type u_2) [Field L] [Algebra K L] [h_fin : FiniteDimensional K L] [hu : IsUltrametricDist K] (x : K) :

      The algebra norm invariantExtension extends the norm on K.