Results on finitely supported functions. #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
The tensor product of ι →₀ M and κ →₀ N is linearly equivalent to (ι × κ) →₀ (M ⊗ N).
noncomputable
def
finsupp_tensor_finsupp
(R : Type u_1)
(M : Type u_2)
(N : Type u_3)
(ι : Type u_4)
(κ : Type u_5)
[comm_ring R]
[add_comm_group M]
[module R M]
[add_comm_group N]
[module R N] :
tensor_product R (ι →₀ M) (κ →₀ N) ≃ₗ[R] ι × κ →₀ tensor_product R M N
The tensor product of ι →₀ M and κ →₀ N is linearly equivalent to (ι × κ) →₀ (M ⊗ N).
Equations
- finsupp_tensor_finsupp R M N ι κ = (tensor_product.congr (finsupp_lequiv_direct_sum R M ι) (finsupp_lequiv_direct_sum R N κ)).trans ((tensor_product.direct_sum R (λ (_x : ι), M) (λ (_x : κ), N)).trans (finsupp_lequiv_direct_sum R (tensor_product R M N) (ι × κ)).symm)
@[simp]
theorem
finsupp_tensor_finsupp_single
(R : Type u_1)
(M : Type u_2)
(N : Type u_3)
(ι : Type u_4)
(κ : Type u_5)
[comm_ring R]
[add_comm_group M]
[module R M]
[add_comm_group N]
[module R N]
(i : ι)
(m : M)
(k : κ)
(n : N) :
⇑(finsupp_tensor_finsupp R M N ι κ) (finsupp.single i m ⊗ₜ[R] finsupp.single k n) = finsupp.single (i, k) (m ⊗ₜ[R] n)
@[simp]
theorem
finsupp_tensor_finsupp_apply
(R : Type u_1)
(M : Type u_2)
(N : Type u_3)
(ι : Type u_4)
(κ : Type u_5)
[comm_ring R]
[add_comm_group M]
[module R M]
[add_comm_group N]
[module R N]
(f : ι →₀ M)
(g : κ →₀ N)
(i : ι)
(k : κ) :
@[simp]
theorem
finsupp_tensor_finsupp_symm_single
(R : Type u_1)
(M : Type u_2)
(N : Type u_3)
(ι : Type u_4)
(κ : Type u_5)
[comm_ring R]
[add_comm_group M]
[module R M]
[add_comm_group N]
[module R N]
(i : ι × κ)
(m : M)
(n : N) :
⇑((finsupp_tensor_finsupp R M N ι κ).symm) (finsupp.single i (m ⊗ₜ[R] n)) = finsupp.single i.fst m ⊗ₜ[R] finsupp.single i.snd n
noncomputable
def
finsupp_tensor_finsupp'
(S : Type u_1)
[comm_ring S]
(α : Type u_2)
(β : Type u_3) :
A variant of finsupp_tensor_finsupp
where both modules are the ground ring.
Equations
- finsupp_tensor_finsupp' S α β = (finsupp_tensor_finsupp S S S α β).trans (finsupp.lcongr (equiv.refl (α × β)) (tensor_product.lid S S))
@[simp]
theorem
finsupp_tensor_finsupp'_single_tmul_single
(S : Type u_1)
[comm_ring S]
(α : Type u_2)
(β : Type u_3)
(a : α)
(b : β)
(r₁ r₂ : S) :
⇑(finsupp_tensor_finsupp' S α β) (finsupp.single a r₁ ⊗ₜ[S] finsupp.single b r₂) = finsupp.single (a, b) (r₁ * r₂)