Documentation

Mathlib.RingTheory.TensorProduct.MvPolynomial

Tensor Product of (multivariate) polynomial rings #

Let Semiring R, Algebra R S and Module R N.

TODO : #

noncomputable def MvPolynomial.rTensor {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] {S : Type u_2} [CommSemiring S] [Algebra R S] [AddCommMonoid N] [Module R N] :

The tensor product of a polynomial ring by a module is linearly equivalent to a Finsupp of a tensor product

Equations
Instances For
    theorem MvPolynomial.rTensor_apply_tmul {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] {S : Type u_2} [CommSemiring S] [Algebra R S] [AddCommMonoid N] [Module R N] (p : MvPolynomial σ S) (n : N) :
    MvPolynomial.rTensor (p ⊗ₜ[R] n) = Finsupp.sum p fun (i : σ →₀ ) (m : S) => Finsupp.single i (m ⊗ₜ[R] n)
    theorem MvPolynomial.rTensor_apply_tmul_apply {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] {S : Type u_2} [CommSemiring S] [Algebra R S] [AddCommMonoid N] [Module R N] (p : MvPolynomial σ S) (n : N) (d : σ →₀ ) :
    (MvPolynomial.rTensor (p ⊗ₜ[R] n)) d = MvPolynomial.coeff d p ⊗ₜ[R] n
    theorem MvPolynomial.rTensor_apply_monomial_tmul {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] {S : Type u_2} [CommSemiring S] [Algebra R S] [AddCommMonoid N] [Module R N] (e : σ →₀ ) (s : S) (n : N) (d : σ →₀ ) :
    (MvPolynomial.rTensor ((MvPolynomial.monomial e) s ⊗ₜ[R] n)) d = if e = d then s ⊗ₜ[R] n else 0
    theorem MvPolynomial.rTensor_apply_X_tmul {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] {S : Type u_2} [CommSemiring S] [Algebra R S] [AddCommMonoid N] [Module R N] (s : σ) (n : N) (d : σ →₀ ) :
    (MvPolynomial.rTensor (MvPolynomial.X s ⊗ₜ[R] n)) d = if Finsupp.single s 1 = d then 1 ⊗ₜ[R] n else 0
    theorem MvPolynomial.rTensor_apply {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] {S : Type u_2} [CommSemiring S] [Algebra R S] [AddCommMonoid N] [Module R N] (t : TensorProduct R (MvPolynomial σ S) N) (d : σ →₀ ) :
    (MvPolynomial.rTensor t) d = (LinearMap.rTensor N (R (MvPolynomial.lcoeff S d))) t
    @[simp]
    theorem MvPolynomial.rTensor_symm_apply_single {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] {S : Type u_2} [CommSemiring S] [Algebra R S] [AddCommMonoid N] [Module R N] (d : σ →₀ ) (s : S) (n : N) :
    MvPolynomial.rTensor.symm (Finsupp.single d (s ⊗ₜ[R] n)) = (MvPolynomial.monomial d) s ⊗ₜ[R] n
    noncomputable def MvPolynomial.scalarRTensor {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] [AddCommMonoid N] [Module R N] :

    The tensor product of the polynomial algebra by a module is linearly equivalent to a Finsupp of that module

    Equations
    Instances For
      theorem MvPolynomial.scalarRTensor_apply_tmul {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] [AddCommMonoid N] [Module R N] (p : MvPolynomial σ R) (n : N) :
      MvPolynomial.scalarRTensor (p ⊗ₜ[R] n) = Finsupp.sum p fun (i : σ →₀ ) (m : R) => Finsupp.single i (m n)
      theorem MvPolynomial.scalarRTensor_apply_tmul_apply {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] [AddCommMonoid N] [Module R N] (p : MvPolynomial σ R) (n : N) (d : σ →₀ ) :
      (MvPolynomial.scalarRTensor (p ⊗ₜ[R] n)) d = MvPolynomial.coeff d p n
      theorem MvPolynomial.scalarRTensor_apply_monomial_tmul {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] [AddCommMonoid N] [Module R N] (e : σ →₀ ) (r : R) (n : N) (d : σ →₀ ) :
      (MvPolynomial.scalarRTensor ((MvPolynomial.monomial e) r ⊗ₜ[R] n)) d = if e = d then r n else 0
      theorem MvPolynomial.scalarRTensor_apply_X_tmul_apply {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] [AddCommMonoid N] [Module R N] (s : σ) (n : N) (d : σ →₀ ) :
      (MvPolynomial.scalarRTensor (MvPolynomial.X s ⊗ₜ[R] n)) d = if Finsupp.single s 1 = d then n else 0
      theorem MvPolynomial.scalarRTensor_symm_apply_single {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] [AddCommMonoid N] [Module R N] (d : σ →₀ ) (n : N) :
      MvPolynomial.scalarRTensor.symm (Finsupp.single d n) = (MvPolynomial.monomial d) 1 ⊗ₜ[R] n
      noncomputable def MvPolynomial.rTensorAlgHom {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} {S : Type u_2} [CommSemiring S] [Algebra R S] [CommSemiring N] [Algebra R N] :

      The algebra morphism from a tensor product of a polynomial algebra by an algebra to a polynomial algebra

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[simp]
        theorem MvPolynomial.coeff_rTensorAlgHom_tmul {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} {S : Type u_2} [CommSemiring S] [Algebra R S] [CommSemiring N] [Algebra R N] (p : MvPolynomial σ S) (n : N) (d : σ →₀ ) :
        MvPolynomial.coeff d (MvPolynomial.rTensorAlgHom (p ⊗ₜ[R] n)) = MvPolynomial.coeff d p ⊗ₜ[R] n
        theorem MvPolynomial.coeff_rTensorAlgHom_monomial_tmul {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] {S : Type u_2} [CommSemiring S] [Algebra R S] [CommSemiring N] [Algebra R N] (e : σ →₀ ) (s : S) (n : N) (d : σ →₀ ) :
        MvPolynomial.coeff d (MvPolynomial.rTensorAlgHom ((MvPolynomial.monomial e) s ⊗ₜ[R] n)) = if e = d then s ⊗ₜ[R] n else 0
        theorem MvPolynomial.rTensorAlgHom_toLinearMap {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] {S : Type u_2} [CommSemiring S] [Algebra R S] [CommSemiring N] [Algebra R N] :
        MvPolynomial.rTensorAlgHom.toLinearMap = MvPolynomial.rTensor
        theorem MvPolynomial.rTensorAlgHom_apply_eq {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] {S : Type u_2} [CommSemiring S] [Algebra R S] [CommSemiring N] [Algebra R N] (p : TensorProduct R (MvPolynomial σ S) N) :
        MvPolynomial.rTensorAlgHom p = MvPolynomial.rTensor p
        noncomputable def MvPolynomial.rTensorAlgEquiv {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] {S : Type u_2} [CommSemiring S] [Algebra R S] [CommSemiring N] [Algebra R N] :

        The tensor product of a polynomial algebra by an algebra is algebraically equivalent to a polynomial algebra

        Equations
        Instances For
          noncomputable def MvPolynomial.scalarRTensorAlgEquiv {R : Type u} {N : Type w} [CommSemiring R] {σ : Type u_1} [DecidableEq σ] [CommSemiring N] [Algebra R N] :

          The tensor product of the polynomial algebra by an algebra is algebraically equivalent to a polynomial algebra with coefficients in that algegra

          Equations
          Instances For
            noncomputable def MvPolynomial.algebraTensorAlgEquiv (R : Type u) [CommSemiring R] {σ : Type u_1} (A : Type u_3) [CommSemiring A] [Algebra R A] :

            Tensoring MvPolynomial σ R on the left by an R-algebra A is algebraically equivalent to M̀vPolynomial σ A.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]