Documentation

Mathlib.RingTheory.MvPowerSeries.Evaluation

Evaluation of multivariate power series #

Let σ, R, S be types, with CommRing R, CommRing S. One assumes that IsTopologicalRing R and UniformAddGroup R, and that S is a complete and separated topological R-algebra, with IsLinearTopology S S, which means there is a basis of neighborhoods of 0 consisting of ideals.

Given φ : R →+* S, a : σ → S, and f : MvPowerSeries σ R, MvPowerSeries.eval₂ f φ a is the evaluation of the multivariate power series f at a. It f is (the coercion of) a polynomial, it coincides with the evaluation of that polynomial. Otherwise, it is defined by density from polynomials; its values are irrelevant unless φ is continuous and a satisfies two conditions bundled in MvPowerSeries.HasEval a :

Under Continuous φ and HasEval a, the following lemmas furnish the properties of evaluation:

structure MvPowerSeries.HasEval {σ : Type u_1} {S : Type u_3} [CommRing S] [TopologicalSpace S] (a : σS) :

Families at which power series can be consistently evaluated

Instances For
    theorem MvPowerSeries.HasEval.add {σ : Type u_1} {S : Type u_3} [CommRing S] [TopologicalSpace S] [ContinuousAdd S] [IsLinearTopology S S] {a b : σS} (ha : HasEval a) (hb : HasEval b) :
    HasEval (a + b)
    theorem MvPowerSeries.HasEval.mul_left {σ : Type u_1} {S : Type u_3} [CommRing S] [TopologicalSpace S] [IsLinearTopology S S] (c : σS) {x : σS} (hx : HasEval x) :
    HasEval (c * x)
    theorem MvPowerSeries.HasEval.mul_right {σ : Type u_1} {S : Type u_3} [CommRing S] [TopologicalSpace S] [IsLinearTopology S S] (c : σS) {x : σS} (hx : HasEval x) :
    HasEval (x * c)
    theorem MvPowerSeries.HasEval.map {σ : Type u_1} {R : Type u_2} [CommRing R] [TopologicalSpace R] {S : Type u_3} [CommRing S] [TopologicalSpace S] {φ : R →+* S} ( : Continuous φ) {a : σR} (ha : HasEval a) :
    HasEval fun (s : σ) => φ (a s)

    Bourbaki, Algebra, chap. 4, §4, n°3, Prop. 4 (i) (a & b).

    theorem MvPowerSeries.HasEval.X {σ : Type u_1} {R : Type u_2} [CommRing R] [TopologicalSpace R] :
    HasEval fun (s : σ) => X s

    The domain of evaluation of MvPowerSeries, as an ideal

    Equations
    Instances For

      The inclusion of polynomials into power series has dense image

      theorem MvPolynomial.toMvPowerSeries_uniformContinuous {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {φ : R →+* S} {a : σS} [UniformAddGroup R] [UniformAddGroup S] [IsLinearTopology S S] ( : Continuous φ) (ha : MvPowerSeries.HasEval a) :
      noncomputable def MvPowerSeries.eval₂ {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] (φ : R →+* S) (a : σS) (f : MvPowerSeries σ R) :
      S

      Evaluation of a multivariate power series at f at a point a : σ → S.

      It coincides with the evaluation of f as a polynomial if f is the coercion of a polynomial. Otherwise, it is only relevant if φ is continuous and HasEval a.

      Equations
      Instances For
        @[simp]
        theorem MvPowerSeries.eval₂_coe {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] (φ : R →+* S) (a : σS) (f : MvPolynomial σ R) :
        eval₂ φ a f = MvPolynomial.eval₂ φ a f
        @[simp]
        theorem MvPowerSeries.eval₂_C {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] (φ : R →+* S) (a : σS) (r : R) :
        eval₂ φ a ((C σ R) r) = φ r
        @[simp]
        theorem MvPowerSeries.eval₂_X {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] (φ : R →+* S) (a : σS) (s : σ) :
        eval₂ φ a (X s) = a s
        noncomputable def MvPowerSeries.eval₂Hom {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {φ : R →+* S} {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] ( : Continuous φ) (ha : HasEval a) :

        Evaluation of power series at adequate elements, as a RingHom

        Equations
        Instances For
          theorem MvPowerSeries.eval₂Hom_eq_extend {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {φ : R →+* S} {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] ( : Continuous φ) (ha : HasEval a) (f : MvPowerSeries σ R) :
          (eval₂Hom ha) f = .extend (MvPolynomial.eval₂ φ a) f
          theorem MvPowerSeries.coe_eval₂Hom {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {φ : R →+* S} {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] ( : Continuous φ) (ha : HasEval a) :
          (eval₂Hom ha) = eval₂ φ a
          theorem MvPowerSeries.continuous_eval₂ {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {φ : R →+* S} {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] ( : Continuous φ) (ha : HasEval a) :
          theorem MvPowerSeries.hasSum_eval₂ {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {φ : R →+* S} {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] ( : Continuous φ) (ha : HasEval a) (f : MvPowerSeries σ R) :
          HasSum (fun (d : σ →₀ ) => φ ((coeff R d) f) * d.prod fun (s : σ) (e : ) => a s ^ e) (eval₂ φ a f)
          theorem MvPowerSeries.eval₂_eq_tsum {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {φ : R →+* S} {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] ( : Continuous φ) (ha : HasEval a) (f : MvPowerSeries σ R) :
          eval₂ φ a f = ∑' (d : σ →₀ ), φ ((coeff R d) f) * d.prod fun (s : σ) (e : ) => a s ^ e
          theorem MvPowerSeries.eval₂_unique {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {φ : R →+* S} {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] ( : Continuous φ) (ha : HasEval a) {ε : MvPowerSeries σ RS} ( : Continuous ε) (h : ∀ (p : MvPolynomial σ R), ε p = MvPolynomial.eval₂ φ a p) :
          ε = eval₂ φ a
          theorem MvPowerSeries.comp_eval₂ {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {φ : R →+* S} {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] ( : Continuous φ) (ha : HasEval a) {T : Type u_4} [UniformSpace T] [CompleteSpace T] [T2Space T] [CommRing T] [IsTopologicalRing T] [IsLinearTopology T T] [UniformAddGroup T] {ε : S →+* T} ( : Continuous ε) :
          ε eval₂ φ a = eval₂ (ε.comp φ) (ε a)
          noncomputable def MvPowerSeries.aeval {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] [Algebra R S] [ContinuousSMul R S] (ha : HasEval a) :

          Evaluation of power series at adequate elements, as an AlgHom

          Equations
          Instances For
            theorem MvPowerSeries.coe_aeval {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] [Algebra R S] [ContinuousSMul R S] (ha : HasEval a) :
            (aeval ha) = eval₂ (algebraMap R S) a
            @[simp]
            theorem MvPowerSeries.aeval_coe {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] [Algebra R S] [ContinuousSMul R S] (ha : HasEval a) (p : MvPolynomial σ R) :
            (aeval ha) p = (MvPolynomial.aeval a) p
            theorem MvPowerSeries.hasSum_aeval {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] [Algebra R S] [ContinuousSMul R S] (ha : HasEval a) (f : MvPowerSeries σ R) :
            HasSum (fun (d : σ →₀ ) => (coeff R d) f d.prod fun (s : σ) (e : ) => a s ^ e) ((aeval ha) f)
            theorem MvPowerSeries.aeval_eq_sum {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] [Algebra R S] [ContinuousSMul R S] (ha : HasEval a) (f : MvPowerSeries σ R) :
            (aeval ha) f = ∑' (d : σ →₀ ), (coeff R d) f d.prod fun (s : σ) (e : ) => a s ^ e
            theorem MvPowerSeries.comp_aeval {σ : Type u_1} {R : Type u_2} [CommRing R] [UniformSpace R] {S : Type u_3} [CommRing S] [UniformSpace S] {a : σS} [IsTopologicalSemiring R] [UniformAddGroup R] [UniformAddGroup S] [CompleteSpace S] [T2Space S] [IsTopologicalRing S] [IsLinearTopology S S] [Algebra R S] [ContinuousSMul R S] (ha : HasEval a) {T : Type u_4} [CommRing T] [UniformSpace T] [UniformAddGroup T] [IsTopologicalRing T] [IsLinearTopology T T] [T2Space T] [Algebra R T] [ContinuousSMul R T] [CompleteSpace T] {ε : S →ₐ[R] T} ( : Continuous ε) :
            ε.comp (aeval ha) = aeval