Documentation

Mathlib.Algebra.MvPolynomial.Equiv

Equivalences between polynomial rings #

This file establishes a number of equivalences between polynomial rings, based on equivalences between the underlying types.

Notation #

As in other polynomial files, we typically use the notation:

Tags #

equivalence, isomorphism, morphism, ring hom, hom

The ring isomorphism between multivariable polynomials in a single variable and polynomials over the ground ring.

Equations
    Instances For
      def MvPolynomial.mapEquiv {S₁ : Type v} {S₂ : Type w} (σ : Type u_1) [CommSemiring S₁] [CommSemiring S₂] (e : S₁ ≃+* S₂) :

      If e : A ≃+* B is an isomorphism of rings, then so is map e.

      Equations
        Instances For
          @[simp]
          theorem MvPolynomial.mapEquiv_apply {S₁ : Type v} {S₂ : Type w} (σ : Type u_1) [CommSemiring S₁] [CommSemiring S₂] (e : S₁ ≃+* S₂) (a : MvPolynomial σ S₁) :
          (mapEquiv σ e) a = (map e) a
          @[simp]
          theorem MvPolynomial.mapEquiv_symm {S₁ : Type v} {S₂ : Type w} (σ : Type u_1) [CommSemiring S₁] [CommSemiring S₂] (e : S₁ ≃+* S₂) :
          @[simp]
          theorem MvPolynomial.mapEquiv_trans {S₁ : Type v} {S₂ : Type w} {S₃ : Type x} (σ : Type u_1) [CommSemiring S₁] [CommSemiring S₂] [CommSemiring S₃] (e : S₁ ≃+* S₂) (f : S₂ ≃+* S₃) :
          (mapEquiv σ e).trans (mapEquiv σ f) = mapEquiv σ (e.trans f)
          def MvPolynomial.mapAlgEquiv {R : Type u} (σ : Type u_1) [CommSemiring R] {A₁ : Type u_2} {A₂ : Type u_3} [CommSemiring A₁] [CommSemiring A₂] [Algebra R A₁] [Algebra R A₂] (e : A₁ ≃ₐ[R] A₂) :

          If e : A ≃ₐ[R] B is an isomorphism of R-algebras, then so is map e.

          Equations
            Instances For
              @[simp]
              theorem MvPolynomial.mapAlgEquiv_apply {R : Type u} (σ : Type u_1) [CommSemiring R] {A₁ : Type u_2} {A₂ : Type u_3} [CommSemiring A₁] [CommSemiring A₂] [Algebra R A₁] [Algebra R A₂] (e : A₁ ≃ₐ[R] A₂) (a : MvPolynomial σ A₁) :
              (mapAlgEquiv σ e) a = (map e) a
              @[simp]
              theorem MvPolynomial.mapAlgEquiv_refl {R : Type u} (σ : Type u_1) [CommSemiring R] {A₁ : Type u_2} [CommSemiring A₁] [Algebra R A₁] :
              @[simp]
              theorem MvPolynomial.mapAlgEquiv_symm {R : Type u} (σ : Type u_1) [CommSemiring R] {A₁ : Type u_2} {A₂ : Type u_3} [CommSemiring A₁] [CommSemiring A₂] [Algebra R A₁] [Algebra R A₂] (e : A₁ ≃ₐ[R] A₂) :
              @[simp]
              theorem MvPolynomial.mapAlgEquiv_trans {R : Type u} (σ : Type u_1) [CommSemiring R] {A₁ : Type u_2} {A₂ : Type u_3} {A₃ : Type u_4} [CommSemiring A₁] [CommSemiring A₂] [CommSemiring A₃] [Algebra R A₁] [Algebra R A₂] [Algebra R A₃] (e : A₁ ≃ₐ[R] A₂) (f : A₂ ≃ₐ[R] A₃) :
              theorem MvPolynomial.eval₂_pUnitAlgEquiv_symm {R : Type u_2} {S : Type u_3} [CommSemiring R] [CommSemiring S] {f : Polynomial R} {φ : R →+* S} {a : UnitS} :
              theorem MvPolynomial.eval₂_const_pUnitAlgEquiv_symm {R : Type u_2} {S : Type u_3} [CommSemiring R] [CommSemiring S] {f : Polynomial R} {φ : R →+* S} {a : S} :
              eval₂ φ (fun (x : PUnit.{1}) => a) ((pUnitAlgEquiv R).symm f) = Polynomial.eval₂ φ a f
              theorem MvPolynomial.eval₂_const_pUnitAlgEquiv {R : Type u_2} {S : Type u_3} [CommSemiring R] [CommSemiring S] {f : MvPolynomial PUnit.{u_4 + 1} R} {φ : R →+* S} {a : S} :
              Polynomial.eval₂ φ a ((pUnitAlgEquiv R) f) = eval₂ φ (fun (x : PUnit.{u_4 + 1}) => a) f
              def MvPolynomial.sumToIter (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] :
              MvPolynomial (S₁ S₂) R →+* MvPolynomial S₁ (MvPolynomial S₂ R)

              The function from multivariable polynomials in a sum of two types, to multivariable polynomials in one of the types, with coefficients in multivariable polynomials in the other type.

              See sumRingEquiv for the ring isomorphism.

              Equations
                Instances For
                  @[simp]
                  theorem MvPolynomial.sumToIter_C (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] (a : R) :
                  (sumToIter R S₁ S₂) (C a) = C (C a)
                  @[simp]
                  theorem MvPolynomial.sumToIter_Xl (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] (b : S₁) :
                  (sumToIter R S₁ S₂) (X (Sum.inl b)) = X b
                  @[simp]
                  theorem MvPolynomial.sumToIter_Xr (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] (c : S₂) :
                  (sumToIter R S₁ S₂) (X (Sum.inr c)) = C (X c)
                  def MvPolynomial.iterToSum (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] :
                  MvPolynomial S₁ (MvPolynomial S₂ R) →+* MvPolynomial (S₁ S₂) R

                  The function from multivariable polynomials in one type, with coefficients in multivariable polynomials in another type, to multivariable polynomials in the sum of the two types.

                  See sumRingEquiv for the ring isomorphism.

                  Equations
                    Instances For
                      @[simp]
                      theorem MvPolynomial.iterToSum_C_C (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] (a : R) :
                      (iterToSum R S₁ S₂) (C (C a)) = C a
                      @[simp]
                      theorem MvPolynomial.iterToSum_X (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] (b : S₁) :
                      (iterToSum R S₁ S₂) (X b) = X (Sum.inl b)
                      @[simp]
                      theorem MvPolynomial.iterToSum_C_X (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] (c : S₂) :
                      (iterToSum R S₁ S₂) (C (X c)) = X (Sum.inr c)

                      The algebra isomorphism between multivariable polynomials in no variables and the ground ring.

                      Equations
                        Instances For
                          @[simp]
                          theorem MvPolynomial.isEmptyAlgEquiv_apply (R : Type u) (σ : Type u_1) [CommSemiring R] [IsEmpty σ] (a : MvPolynomial σ R) :
                          (isEmptyAlgEquiv R σ) a = (aeval fun (a : σ) => isEmptyElim a) a
                          @[simp]
                          theorem MvPolynomial.aeval_injective_iff_of_isEmpty {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [IsEmpty σ] [CommSemiring S₁] [Algebra R S₁] {f : σS₁} :

                          The ring isomorphism between multivariable polynomials in no variables and the ground ring.

                          Equations
                            Instances For
                              @[simp]
                              theorem MvPolynomial.isEmptyRingEquiv_apply (R : Type u) (σ : Type u_1) [CommSemiring R] [IsEmpty σ] (a : MvPolynomial σ R) :
                              (isEmptyRingEquiv R σ) a = (aeval fun (a : σ) => isEmptyElim a) a
                              @[simp]
                              theorem MvPolynomial.isEmptyRingEquiv_symm_apply (R : Type u) {σ : Type u_1} [CommSemiring R] [IsEmpty σ] (r : R) :
                              def MvPolynomial.mvPolynomialEquivMvPolynomial (R : Type u) (S₁ : Type v) (S₂ : Type w) (S₃ : Type x) [CommSemiring R] [CommSemiring S₃] (f : MvPolynomial S₁ R →+* MvPolynomial S₂ S₃) (g : MvPolynomial S₂ S₃ →+* MvPolynomial S₁ R) (hfgC : (f.comp g).comp C = C) (hfgX : ∀ (n : S₂), f (g (X n)) = X n) (hgfC : (g.comp f).comp C = C) (hgfX : ∀ (n : S₁), g (f (X n)) = X n) :

                              A helper function for sumRingEquiv.

                              Equations
                                Instances For
                                  @[simp]
                                  theorem MvPolynomial.mvPolynomialEquivMvPolynomial_symm_apply (R : Type u) (S₁ : Type v) (S₂ : Type w) (S₃ : Type x) [CommSemiring R] [CommSemiring S₃] (f : MvPolynomial S₁ R →+* MvPolynomial S₂ S₃) (g : MvPolynomial S₂ S₃ →+* MvPolynomial S₁ R) (hfgC : (f.comp g).comp C = C) (hfgX : ∀ (n : S₂), f (g (X n)) = X n) (hgfC : (g.comp f).comp C = C) (hgfX : ∀ (n : S₁), g (f (X n)) = X n) (a : MvPolynomial S₂ S₃) :
                                  (mvPolynomialEquivMvPolynomial R S₁ S₂ S₃ f g hfgC hfgX hgfC hgfX).symm a = g a
                                  @[simp]
                                  theorem MvPolynomial.mvPolynomialEquivMvPolynomial_apply (R : Type u) (S₁ : Type v) (S₂ : Type w) (S₃ : Type x) [CommSemiring R] [CommSemiring S₃] (f : MvPolynomial S₁ R →+* MvPolynomial S₂ S₃) (g : MvPolynomial S₂ S₃ →+* MvPolynomial S₁ R) (hfgC : (f.comp g).comp C = C) (hfgX : ∀ (n : S₂), f (g (X n)) = X n) (hgfC : (g.comp f).comp C = C) (hgfX : ∀ (n : S₁), g (f (X n)) = X n) (a : MvPolynomial S₁ R) :
                                  (mvPolynomialEquivMvPolynomial R S₁ S₂ S₃ f g hfgC hfgX hgfC hgfX) a = f a
                                  def MvPolynomial.sumRingEquiv (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] :
                                  MvPolynomial (S₁ S₂) R ≃+* MvPolynomial S₁ (MvPolynomial S₂ R)

                                  The ring isomorphism between multivariable polynomials in a sum of two types, and multivariable polynomials in one of the types, with coefficients in multivariable polynomials in the other type.

                                  Equations
                                    Instances For
                                      def MvPolynomial.sumAlgEquiv (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] :
                                      MvPolynomial (S₁ S₂) R ≃ₐ[R] MvPolynomial S₁ (MvPolynomial S₂ R)

                                      The algebra isomorphism between multivariable polynomials in a sum of two types, and multivariable polynomials in one of the types, with coefficients in multivariable polynomials in the other type.

                                      Equations
                                        Instances For
                                          @[simp]
                                          theorem MvPolynomial.sumAlgEquiv_symm_apply (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] (a : MvPolynomial S₁ (MvPolynomial S₂ R)) :
                                          (sumAlgEquiv R S₁ S₂).symm a = (iterToSum R S₁ S₂) a
                                          @[simp]
                                          theorem MvPolynomial.sumAlgEquiv_apply (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] (a : MvPolynomial (S₁ S₂) R) :
                                          (sumAlgEquiv R S₁ S₂) a = (sumToIter R S₁ S₂) a
                                          theorem MvPolynomial.sumAlgEquiv_comp_rename_inl (R : Type u) (S₁ : Type v) (S₂ : Type w) [CommSemiring R] :
                                          noncomputable def MvPolynomial.commAlgEquiv (R : Type u_2) (S₁ : Type u_3) (S₂ : Type u_4) [CommSemiring R] :

                                          The algebra isomorphism between multivariable polynomials in variables S₁ of multivariable polynomials in variables S₂ and multivariable polynomials in variables S₂ of multivariable polynomials in variables S₁.

                                          Equations
                                            Instances For
                                              @[simp]
                                              theorem MvPolynomial.commAlgEquiv_C {R : Type u_2} {S₁ : Type u_3} {S₂ : Type u_4} [CommSemiring R] (p : MvPolynomial S₂ R) :
                                              (commAlgEquiv R S₁ S₂) (C p) = (map C) p
                                              theorem MvPolynomial.commAlgEquiv_C_X {R : Type u_2} {S₁ : Type u_3} {S₂ : Type u_4} [CommSemiring R] (i : S₂) :
                                              (commAlgEquiv R S₁ S₂) (C (X i)) = X i
                                              @[simp]
                                              theorem MvPolynomial.commAlgEquiv_X {R : Type u_2} {S₁ : Type u_3} {S₂ : Type u_4} [CommSemiring R] (i : S₁) :
                                              (commAlgEquiv R S₁ S₂) (X i) = C (X i)

                                              The algebra isomorphism between multivariable polynomials in Option S₁ and polynomials with coefficients in MvPolynomial S₁ R.

                                              Equations
                                                Instances For
                                                  theorem MvPolynomial.optionEquivLeft_apply (R : Type u) (S₁ : Type v) [CommSemiring R] (a : MvPolynomial (Option S₁) R) :
                                                  (optionEquivLeft R S₁) a = (aeval fun (o : Option S₁) => o.elim Polynomial.X fun (s : S₁) => Polynomial.C (X s)) a
                                                  theorem MvPolynomial.optionEquivLeft_X_some (R : Type u) (S₁ : Type v) [CommSemiring R] (x : S₁) :
                                                  (optionEquivLeft R S₁) (X (some x)) = Polynomial.C (X x)
                                                  theorem MvPolynomial.optionEquivLeft_C (R : Type u) (S₁ : Type v) [CommSemiring R] (r : R) :
                                                  (optionEquivLeft R S₁) (C r) = Polynomial.C (C r)
                                                  theorem MvPolynomial.optionEquivLeft_monomial (R : Type u) (S₁ : Type v) [CommSemiring R] (m : Option S₁ →₀ ) (r : R) :
                                                  theorem MvPolynomial.optionEquivLeft_coeff_coeff (R : Type u) (S₁ : Type v) [CommSemiring R] (n : Option S₁ →₀ ) (f : MvPolynomial (Option S₁) R) :
                                                  coeff n.some (((optionEquivLeft R S₁) f).coeff (n none)) = coeff n f

                                                  The coefficient of n.some in the n none-th coefficient of optionEquivLeft R S₁ f equals the coefficient of n in f

                                                  theorem MvPolynomial.optionEquivLeft_elim_eval (R : Type u) (S₁ : Type v) [CommSemiring R] (s : S₁R) (y : R) (f : MvPolynomial (Option S₁) R) :
                                                  (eval fun (x : Option S₁) => x.elim y s) f = Polynomial.eval y (Polynomial.map (eval s) ((optionEquivLeft R S₁) f))
                                                  @[simp]

                                                  The algebra isomorphism between multivariable polynomials in Option S₁ and multivariable polynomials with coefficients in polynomials.

                                                  Equations
                                                    Instances For
                                                      @[simp]
                                                      theorem MvPolynomial.optionEquivRight_symm_apply (R : Type u) (S₁ : Type v) [CommSemiring R] (a : MvPolynomial S₁ (Polynomial R)) :
                                                      (optionEquivRight R S₁).symm a = (aevalTower (Polynomial.aeval (X none)) fun (i : S₁) => X (some i)) a
                                                      @[simp]
                                                      theorem MvPolynomial.optionEquivRight_apply (R : Type u) (S₁ : Type v) [CommSemiring R] (a : MvPolynomial (Option S₁) R) :
                                                      (optionEquivRight R S₁) a = (aeval fun (o : Option S₁) => o.elim (C Polynomial.X) X) a
                                                      theorem MvPolynomial.optionEquivRight_X_some (R : Type u) (S₁ : Type v) [CommSemiring R] (x : S₁) :
                                                      (optionEquivRight R S₁) (X (some x)) = X x
                                                      theorem MvPolynomial.optionEquivRight_C (R : Type u) (S₁ : Type v) [CommSemiring R] (r : R) :
                                                      (optionEquivRight R S₁) (C r) = C (Polynomial.C r)

                                                      The algebra isomorphism between multivariable polynomials in Fin (n + 1) and polynomials over multivariable polynomials in Fin n.

                                                      Equations
                                                        Instances For
                                                          theorem MvPolynomial.finSuccEquiv_eq (R : Type u) [CommSemiring R] (n : ) :
                                                          (finSuccEquiv R n) = eval₂Hom (Polynomial.C.comp C) fun (i : Fin (n + 1)) => Fin.cases Polynomial.X (fun (k : Fin n) => Polynomial.C (X k)) i
                                                          theorem MvPolynomial.finSuccEquiv_apply (R : Type u) [CommSemiring R] (n : ) (p : MvPolynomial (Fin (n + 1)) R) :
                                                          (finSuccEquiv R n) p = (eval₂Hom (Polynomial.C.comp C) fun (i : Fin (n + 1)) => Fin.cases Polynomial.X (fun (k : Fin n) => Polynomial.C (X k)) i) p
                                                          theorem MvPolynomial.finSuccEquiv_coeff_coeff {R : Type u} [CommSemiring R] {n : } (m : Fin n →₀ ) (f : MvPolynomial (Fin (n + 1)) R) (i : ) :
                                                          coeff m (((finSuccEquiv R n) f).coeff i) = coeff (Finsupp.cons i m) f

                                                          The coefficient of m in the i-th coefficient of finSuccEquiv R n f equals the coefficient of Finsupp.cons i m in f.

                                                          theorem MvPolynomial.eval_eq_eval_mv_eval' {R : Type u} [CommSemiring R] {n : } (s : Fin nR) (y : R) (f : MvPolynomial (Fin (n + 1)) R) :
                                                          theorem MvPolynomial.coeff_eval_eq_eval_coeff {R : Type u} (S₁ : Type v) [CommSemiring R] (s' : S₁R) (f : Polynomial (MvPolynomial S₁ R)) (i : ) :
                                                          (Polynomial.map (eval s') f).coeff i = (eval s') (f.coeff i)
                                                          theorem MvPolynomial.totalDegree_coeff_finSuccEquiv_add_le {R : Type u} [CommSemiring R] {n : } (f : MvPolynomial (Fin (n + 1)) R) (i : ) (hi : ((finSuccEquiv R n) f).coeff i 0) :

                                                          The totalDegree of a multivariable polynomial p is at least i more than the totalDegree of the ith coefficient of finSuccEquiv applied to p, if this is nonzero.

                                                          theorem MvPolynomial.support_finSuccEquiv {R : Type u} [CommSemiring R] {n : } (f : MvPolynomial (Fin (n + 1)) R) :
                                                          ((finSuccEquiv R n) f).support = Finset.image (fun (m : Fin (n + 1) →₀ ) => m 0) f.support
                                                          theorem MvPolynomial.mem_support_finSuccEquiv {R : Type u} [CommSemiring R] {n : } {f : MvPolynomial (Fin (n + 1)) R} {x : } :
                                                          x ((finSuccEquiv R n) f).support x (fun (m : Fin (n + 1) →₀ ) => m 0) '' f.support
                                                          theorem MvPolynomial.image_support_finSuccEquiv {R : Type u} [CommSemiring R] {n : } {f : MvPolynomial (Fin (n + 1)) R} {i : } :
                                                          Finset.image (Finsupp.cons i) (((finSuccEquiv R n) f).coeff i).support = {mf.support | m 0 = i}
                                                          theorem MvPolynomial.mem_image_support_coeff_finSuccEquiv {R : Type u} [CommSemiring R] {n : } {f : MvPolynomial (Fin (n + 1)) R} {i : } {x : Fin (n + 1) →₀ } :
                                                          x Finsupp.cons i '' (((finSuccEquiv R n) f).coeff i).support x f.support x 0 = i
                                                          theorem MvPolynomial.degree_finSuccEquiv {R : Type u} [CommSemiring R] {n : } {f : MvPolynomial (Fin (n + 1)) R} (h : f 0) :
                                                          ((finSuccEquiv R n) f).degree = (degreeOf 0 f)
                                                          theorem MvPolynomial.degreeOf_coeff_finSuccEquiv {R : Type u} [CommSemiring R] {n : } (p : MvPolynomial (Fin (n + 1)) R) (j : Fin n) (i : ) :

                                                          Consider a multivariate polynomial φ whose variables are indexed by Option σ, and suppose that σ ≃ Fin n. Then one may view φ as a polynomial over MvPolynomial (Fin n) R, by

                                                          1. renaming the variables via Option σ ≃ Fin (n+1), and then singling out the 0-th variable via MvPolynomial.finSuccEquiv;
                                                          2. first viewing it as polynomial over MvPolynomial σ R via MvPolynomial.optionEquivLeft, and then renaming the variables.

                                                          This lemma shows that both constructions are the same.

                                                          @[simp]
                                                          theorem MvPolynomial.rename_polynomial_aeval_X (R : Type u) [CommSemiring R] {σ : Type u_2} {τ : Type u_3} (f : στ) (i : σ) (p : Polynomial R) :
                                                          (rename f) ((Polynomial.aeval (X i)) p) = (Polynomial.aeval (X (f i))) p
                                                          noncomputable def Polynomial.toMvPolynomial {R : Type u_1} {σ : Type u_3} [CommSemiring R] (i : σ) :

                                                          The embedding of R[X] into R[Xᵢ] as an R-algebra homomorphism.

                                                          Equations
                                                            Instances For
                                                              @[simp]
                                                              theorem Polynomial.toMvPolynomial_C {R : Type u_1} {σ : Type u_3} [CommSemiring R] (i : σ) (r : R) :
                                                              @[simp]
                                                              theorem Polynomial.toMvPolynomial_X {R : Type u_1} {σ : Type u_3} [CommSemiring R] (i : σ) :
                                                              @[simp]
                                                              theorem MvPolynomial.eval_comp_toMvPolynomial {R : Type u_1} {σ : Type u_3} [CommSemiring R] (f : σR) (i : σ) :
                                                              @[simp]
                                                              theorem MvPolynomial.eval_toMvPolynomial {R : Type u_1} {σ : Type u_3} [CommSemiring R] (f : σR) (i : σ) (p : Polynomial R) :
                                                              @[simp]
                                                              theorem MvPolynomial.aeval_comp_toMvPolynomial {R : Type u_1} {S : Type u_2} {σ : Type u_3} [CommSemiring R] [CommSemiring S] [Algebra R S] (f : σS) (i : σ) :
                                                              @[simp]
                                                              theorem MvPolynomial.aeval_toMvPolynomial {R : Type u_1} {S : Type u_2} {σ : Type u_3} [CommSemiring R] [CommSemiring S] [Algebra R S] (f : σS) (i : σ) (p : Polynomial R) :
                                                              @[simp]
                                                              theorem MvPolynomial.rename_comp_toMvPolynomial {R : Type u_1} {σ : Type u_3} {τ : Type u_4} [CommSemiring R] (f : στ) (a : σ) :
                                                              @[simp]
                                                              theorem MvPolynomial.rename_toMvPolynomial {R : Type u_1} {σ : Type u_3} {τ : Type u_4} [CommSemiring R] (f : στ) (a : σ) (p : Polynomial R) :