|
| 1 | +/- |
| 2 | +Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Antoine Chambert-Loir |
| 5 | +-/ |
| 6 | + |
| 7 | +import Mathlib.LinearAlgebra.DirectSum.Finsupp |
| 8 | +import Mathlib.Algebra.MvPolynomial.Basic |
| 9 | +import Mathlib.RingTheory.TensorProduct.Basic |
| 10 | +import Mathlib.Algebra.MvPolynomial.Equiv |
| 11 | +/-! |
| 12 | +
|
| 13 | +# Tensor Product of (multivariate) polynomial rings |
| 14 | +
|
| 15 | +Let `Semiring R`, `Algebra R S` and `Module R N`. |
| 16 | +
|
| 17 | +* `MvPolynomial.rTensor` gives the linear equivalence |
| 18 | + `MvPolynomial σ S ⊗[R] N ≃ₗ[R] (σ →₀ ℕ) →₀ (S ⊗[R] N)` characterized, |
| 19 | + for `p : MvPolynomial σ S`, `n : N` and `d : σ →₀ ℕ`, by |
| 20 | + `rTensor (p ⊗ₜ[R] n) d = (coeff d p) ⊗ₜ[R] n` |
| 21 | +* `MvPolynomial.scalarRTensor` gives the linear equivalence |
| 22 | + `MvPolynomial σ R ⊗[R] N ≃ₗ[R] (σ →₀ ℕ) →₀ N` |
| 23 | + such that `MvPolynomial.scalarRTensor (p ⊗ₜ[R] n) d = coeff d p • n` |
| 24 | + for `p : MvPolynomial σ R`, `n : N` and `d : σ →₀ ℕ`, by |
| 25 | +
|
| 26 | +* `MvPolynomial.rTensorAlgHom`, the algebra morphism from the tensor product |
| 27 | + of a polynomial algebra by an algebra to a polynomial algebra |
| 28 | +* `MvPolynomial.rTensorAlgEquiv`, `MvPolynomial.scalarRTensorAlgEquiv`, |
| 29 | + the tensor product of a polynomial algebra by an algebra |
| 30 | + is algebraically equivalent to a polynomial algebra |
| 31 | +
|
| 32 | +## TODO : |
| 33 | +* `MvPolynomial.rTensor` could be phrased in terms of `AddMonoidAlgebra`, and |
| 34 | + `MvPolynomial.rTensor` then has `smul` by the polynomial algebra. |
| 35 | +* `MvPolynomial.rTensorAlgHom` and `MvPolynomial.scalarRTensorAlgEquiv` |
| 36 | + are morphisms for the algebra structure by `MvPolynomial σ R`. |
| 37 | +-/ |
| 38 | + |
| 39 | + |
| 40 | +universe u v w |
| 41 | + |
| 42 | +noncomputable section |
| 43 | + |
| 44 | +namespace MvPolynomial |
| 45 | + |
| 46 | +open DirectSum TensorProduct |
| 47 | + |
| 48 | +open Set LinearMap Submodule |
| 49 | + |
| 50 | +variable {R : Type u} {M : Type v} {N : Type w} |
| 51 | + [CommSemiring R] [AddCommMonoid M] [Module R M] |
| 52 | + |
| 53 | +variable {σ : Type*} [DecidableEq σ] |
| 54 | + |
| 55 | +variable {S : Type*} [CommSemiring S] [Algebra R S] |
| 56 | + |
| 57 | +section Module |
| 58 | + |
| 59 | +variable [AddCommMonoid N] [Module R N] |
| 60 | + |
| 61 | +/-- The tensor product of a polynomial ring by a module is |
| 62 | + linearly equivalent to a Finsupp of a tensor product -/ |
| 63 | +noncomputable def rTensor : |
| 64 | + MvPolynomial σ S ⊗[R] N ≃ₗ[S] (σ →₀ ℕ) →₀ (S ⊗[R] N) := |
| 65 | + TensorProduct.finsuppLeft' _ _ _ _ _ |
| 66 | + |
| 67 | +lemma rTensor_apply_tmul (p : MvPolynomial σ S) (n : N) : |
| 68 | + rTensor (p ⊗ₜ[R] n) = p.sum (fun i m ↦ Finsupp.single i (m ⊗ₜ[R] n)) := |
| 69 | + TensorProduct.finsuppLeft_apply_tmul p n |
| 70 | + |
| 71 | +lemma rTensor_apply_tmul_apply (p : MvPolynomial σ S) (n : N) (d : σ →₀ ℕ) : |
| 72 | + rTensor (p ⊗ₜ[R] n) d = (coeff d p) ⊗ₜ[R] n := |
| 73 | + TensorProduct.finsuppLeft_apply_tmul_apply p n d |
| 74 | + |
| 75 | +lemma rTensor_apply_monomial_tmul (e : σ →₀ ℕ) (s : S) (n : N) (d : σ →₀ ℕ) : |
| 76 | + rTensor (monomial e s ⊗ₜ[R] n) d = if e = d then s ⊗ₜ[R] n else 0 := by |
| 77 | + simp only [rTensor_apply_tmul_apply, coeff_monomial, ite_tmul] |
| 78 | + |
| 79 | +lemma rTensor_apply_X_tmul (s : σ) (n : N) (d : σ →₀ ℕ) : |
| 80 | + rTensor (X s ⊗ₜ[R] n) d = if Finsupp.single s 1 = d then (1 : S) ⊗ₜ[R] n else 0 := by |
| 81 | + rw [rTensor_apply_tmul_apply, coeff_X', ite_tmul] |
| 82 | + |
| 83 | +lemma rTensor_apply (t : MvPolynomial σ S ⊗[R] N) (d : σ →₀ ℕ) : |
| 84 | + rTensor t d = ((lcoeff S d).restrictScalars R).rTensor N t := |
| 85 | + TensorProduct.finsuppLeft_apply t d |
| 86 | + |
| 87 | +@[simp] |
| 88 | +lemma rTensor_symm_apply_single (d : σ →₀ ℕ) (s : S) (n : N) : |
| 89 | + rTensor.symm (Finsupp.single d (s ⊗ₜ n)) = |
| 90 | + (monomial d s) ⊗ₜ[R] n := |
| 91 | + TensorProduct.finsuppLeft_symm_apply_single (R := R) d s n |
| 92 | + |
| 93 | +/-- The tensor product of the polynomial algebra by a module |
| 94 | + is linearly equivalent to a Finsupp of that module -/ |
| 95 | +noncomputable def scalarRTensor : |
| 96 | + MvPolynomial σ R ⊗[R] N ≃ₗ[R] (σ →₀ ℕ) →₀ N := |
| 97 | + TensorProduct.finsuppScalarLeft _ _ _ |
| 98 | + |
| 99 | +lemma scalarRTensor_apply_tmul (p : MvPolynomial σ R) (n : N) : |
| 100 | + scalarRTensor (p ⊗ₜ[R] n) = p.sum (fun i m ↦ Finsupp.single i (m • n)) := |
| 101 | + TensorProduct.finsuppScalarLeft_apply_tmul p n |
| 102 | + |
| 103 | +lemma scalarRTensor_apply_tmul_apply (p : MvPolynomial σ R) (n : N) (d : σ →₀ ℕ): |
| 104 | + scalarRTensor (p ⊗ₜ[R] n) d = coeff d p • n := |
| 105 | + TensorProduct.finsuppScalarLeft_apply_tmul_apply p n d |
| 106 | + |
| 107 | +lemma scalarRTensor_apply_monomial_tmul (e : σ →₀ ℕ) (r : R) (n : N) (d : σ →₀ ℕ): |
| 108 | + scalarRTensor (monomial e r ⊗ₜ[R] n) d = if e = d then r • n else 0 := by |
| 109 | + rw [scalarRTensor_apply_tmul_apply, coeff_monomial, ite_smul, zero_smul] |
| 110 | + |
| 111 | + lemma scalarRTensor_apply_X_tmul_apply (s : σ) (n : N) (d : σ →₀ ℕ): |
| 112 | + scalarRTensor (X s ⊗ₜ[R] n) d = if Finsupp.single s 1 = d then n else 0 := by |
| 113 | + rw [scalarRTensor_apply_tmul_apply, coeff_X', ite_smul, one_smul, zero_smul] |
| 114 | + |
| 115 | +lemma scalarRTensor_symm_apply_single (d : σ →₀ ℕ) (n : N) : |
| 116 | + scalarRTensor.symm (Finsupp.single d n) = (monomial d 1) ⊗ₜ[R] n := |
| 117 | + TensorProduct.finsuppScalarLeft_symm_apply_single d n |
| 118 | + |
| 119 | +end Module |
| 120 | + |
| 121 | +section Algebra |
| 122 | + |
| 123 | +variable [CommSemiring N] [Algebra R N] |
| 124 | + |
| 125 | +/-- The algebra morphism from a tensor product of a polynomial algebra |
| 126 | + by an algebra to a polynomial algebra -/ |
| 127 | +noncomputable def rTensorAlgHom : |
| 128 | + (MvPolynomial σ S) ⊗[R] N →ₐ[S] MvPolynomial σ (S ⊗[R] N) := |
| 129 | + Algebra.TensorProduct.lift |
| 130 | + (mapAlgHom Algebra.TensorProduct.includeLeft) |
| 131 | + ((IsScalarTower.toAlgHom R (S ⊗[R] N) _).comp Algebra.TensorProduct.includeRight) |
| 132 | + (fun p n => by simp [commute_iff_eq, algebraMap_eq, mul_comm]) |
| 133 | + |
| 134 | +@[simp] |
| 135 | +lemma coeff_rTensorAlgHom_tmul |
| 136 | + (p : MvPolynomial σ S) (n : N) (d : σ →₀ ℕ) : |
| 137 | + coeff d (rTensorAlgHom (p ⊗ₜ[R] n)) = (coeff d p) ⊗ₜ[R] n := by |
| 138 | + rw [rTensorAlgHom, Algebra.TensorProduct.lift_tmul] |
| 139 | + rw [AlgHom.coe_comp, IsScalarTower.coe_toAlgHom', Function.comp_apply, |
| 140 | + Algebra.TensorProduct.includeRight_apply] |
| 141 | + rw [algebraMap_eq, mul_comm, coeff_C_mul] |
| 142 | + simp [mapAlgHom, coeff_map] |
| 143 | + |
| 144 | +lemma coeff_rTensorAlgHom_monomial_tmul |
| 145 | + (e : σ →₀ ℕ) (s : S) (n : N) (d : σ →₀ ℕ) : |
| 146 | + coeff d (rTensorAlgHom (monomial e s ⊗ₜ[R] n)) = |
| 147 | + if e = d then s ⊗ₜ[R] n else 0 := by |
| 148 | + simp [ite_tmul] |
| 149 | + |
| 150 | +lemma rTensorAlgHom_toLinearMap : |
| 151 | + (rTensorAlgHom : |
| 152 | + MvPolynomial σ S ⊗[R] N →ₐ[S] MvPolynomial σ (S ⊗[R] N)).toLinearMap = |
| 153 | + rTensor.toLinearMap := by |
| 154 | + ext d n e |
| 155 | + dsimp only [AlgebraTensorModule.curry_apply, TensorProduct.curry_apply, |
| 156 | + LinearMap.coe_restrictScalars, AlgHom.toLinearMap_apply] |
| 157 | + simp only [coe_comp, Function.comp_apply, AlgebraTensorModule.curry_apply, curry_apply, |
| 158 | + LinearMap.coe_restrictScalars, AlgHom.toLinearMap_apply] |
| 159 | + rw [coeff_rTensorAlgHom_tmul] |
| 160 | + simp only [coeff] |
| 161 | + erw [finsuppLeft_apply_tmul_apply] |
| 162 | + |
| 163 | +lemma rTensorAlgHom_apply_eq (p : MvPolynomial σ S ⊗[R] N) : |
| 164 | + rTensorAlgHom (S := S) p = rTensor p := by |
| 165 | + rw [← AlgHom.toLinearMap_apply, rTensorAlgHom_toLinearMap] |
| 166 | + rfl |
| 167 | + |
| 168 | +/-- The tensor product of a polynomial algebra by an algebra |
| 169 | + is algebraically equivalent to a polynomial algebra -/ |
| 170 | +noncomputable def rTensorAlgEquiv : |
| 171 | + (MvPolynomial σ S) ⊗[R] N ≃ₐ[S] MvPolynomial σ (S ⊗[R] N) := by |
| 172 | + apply AlgEquiv.ofLinearEquiv rTensor |
| 173 | + · simp only [Algebra.TensorProduct.one_def] |
| 174 | + apply symm |
| 175 | + rw [← LinearEquiv.symm_apply_eq] |
| 176 | + exact finsuppLeft_symm_apply_single (R := R) (0 : σ →₀ ℕ) (1 : S) (1 : N) |
| 177 | + · intro x y |
| 178 | + erw [← rTensorAlgHom_apply_eq (S := S)] |
| 179 | + simp only [_root_.map_mul, rTensorAlgHom_apply_eq] |
| 180 | + rfl |
| 181 | + |
| 182 | +/-- The tensor product of the polynomial algebra by an algebra |
| 183 | + is algebraically equivalent to a polynomial algebra with |
| 184 | + coefficients in that algegra -/ |
| 185 | +noncomputable def scalarRTensorAlgEquiv : |
| 186 | + MvPolynomial σ R ⊗[R] N ≃ₐ[R] MvPolynomial σ N := |
| 187 | + rTensorAlgEquiv.trans (mapAlgEquiv σ (Algebra.TensorProduct.lid R N)) |
| 188 | + |
| 189 | +end Algebra |
| 190 | + |
| 191 | +end MvPolynomial |
| 192 | + |
| 193 | +end |
0 commit comments