Skip to content

Commit e84fd56

Browse files
committed
feat: Port Topology.Algebra.InfiniteSum.Ring (#2640)
1 parent 41b3bbf commit e84fd56

File tree

2 files changed

+250
-0
lines changed

2 files changed

+250
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1241,6 +1241,7 @@ import Mathlib.Topology.Algebra.Group.Basic
12411241
import Mathlib.Topology.Algebra.Group.Compact
12421242
import Mathlib.Topology.Algebra.GroupWithZero
12431243
import Mathlib.Topology.Algebra.InfiniteSum.Basic
1244+
import Mathlib.Topology.Algebra.InfiniteSum.Ring
12441245
import Mathlib.Topology.Algebra.Monoid
12451246
import Mathlib.Topology.Algebra.MulAction
12461247
import Mathlib.Topology.Algebra.Order.Archimedean
Lines changed: 249 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,249 @@
1+
/-
2+
Copyright (c) 2017 Johannes Hölzl. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Johannes Hölzl
5+
6+
! This file was ported from Lean 3 source module topology.algebra.infinite_sum.ring
7+
! leanprover-community/mathlib commit 9a59dcb7a2d06bf55da57b9030169219980660cd
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.Algebra.BigOperators.NatAntidiagonal
12+
import Mathlib.Topology.Algebra.InfiniteSum.Basic
13+
import Mathlib.Topology.Algebra.Ring.Basic
14+
15+
/-!
16+
# Infinite sum in a ring
17+
18+
This file provides lemmas about the interaction between infinite sums and multiplication.
19+
20+
## Main results
21+
22+
* `tsum_mul_tsum_eq_tsum_sum_antidiagonal`: Cauchy product formula
23+
-/
24+
25+
26+
open Filter Finset Function
27+
28+
open BigOperators Classical
29+
30+
variable {ι κ R α : Type _}
31+
32+
section NonUnitalNonAssocSemiring
33+
34+
variable [NonUnitalNonAssocSemiring α] [TopologicalSpace α] [TopologicalSemiring α] {f g : ι → α}
35+
{a a₁ a₂ : α}
36+
37+
theorem HasSum.mul_left (a₂) (h : HasSum f a₁) : HasSum (fun i => a₂ * f i) (a₂ * a₁) := by
38+
simpa only using h.map (AddMonoidHom.mulLeft a₂) (continuous_const.mul continuous_id)
39+
#align has_sum.mul_left HasSum.mul_left
40+
41+
theorem HasSum.mul_right (a₂) (hf : HasSum f a₁) : HasSum (fun i => f i * a₂) (a₁ * a₂) := by
42+
simpa only using hf.map (AddMonoidHom.mulRight a₂) (continuous_id.mul continuous_const)
43+
#align has_sum.mul_right HasSum.mul_right
44+
45+
theorem Summable.mul_left (a) (hf : Summable f) : Summable fun i => a * f i :=
46+
(hf.hasSum.mul_left _).summable
47+
#align summable.mul_left Summable.mul_left
48+
49+
theorem Summable.mul_right (a) (hf : Summable f) : Summable fun i => f i * a :=
50+
(hf.hasSum.mul_right _).summable
51+
#align summable.mul_right Summable.mul_right
52+
53+
section tsum
54+
55+
variable [T2Space α]
56+
57+
theorem Summable.tsum_mul_left (a) (hf : Summable f) : (∑' i, a * f i) = a * ∑' i, f i :=
58+
(hf.hasSum.mul_left _).tsum_eq
59+
#align summable.tsum_mul_left Summable.tsum_mul_left
60+
61+
theorem Summable.tsum_mul_right (a) (hf : Summable f) : (∑' i, f i * a) = (∑' i, f i) * a :=
62+
(hf.hasSum.mul_right _).tsum_eq
63+
#align summable.tsum_mul_right Summable.tsum_mul_right
64+
65+
theorem Commute.tsum_right (a) (h : ∀ i, _root_.Commute a (f i)) : _root_.Commute a (∑' i, f i) :=
66+
if hf : Summable f then
67+
(hf.tsum_mul_left a).symm.trans ((congr_arg _ <| funext h).trans (hf.tsum_mul_right a))
68+
else (tsum_eq_zero_of_not_summable hf).symm ▸ Commute.zero_right _
69+
#align commute.tsum_right Commute.tsum_right
70+
71+
theorem Commute.tsum_left (a) (h : ∀ i, _root_.Commute (f i) a) : _root_.Commute (∑' i, f i) a :=
72+
(Commute.tsum_right _ fun i => (h i).symm).symm
73+
#align commute.tsum_left Commute.tsum_left
74+
75+
end tsum
76+
77+
end NonUnitalNonAssocSemiring
78+
79+
section DivisionSemiring
80+
81+
variable [DivisionSemiring α] [TopologicalSpace α] [TopologicalSemiring α] {f g : ι → α}
82+
{a a₁ a₂ : α}
83+
84+
theorem HasSum.div_const (h : HasSum f a) (b : α) : HasSum (fun i => f i / b) (a / b) := by
85+
simp only [div_eq_mul_inv, h.mul_right b⁻¹]
86+
#align has_sum.div_const HasSum.div_const
87+
88+
theorem Summable.div_const (h : Summable f) (b : α) : Summable fun i => f i / b :=
89+
(h.hasSum.div_const _).summable
90+
#align summable.div_const Summable.div_const
91+
92+
theorem hasSum_mul_left_iff (h : a₂ ≠ 0) : HasSum (fun i => a₂ * f i) (a₂ * a₁) ↔ HasSum f a₁ :=
93+
fun H => by simpa only [inv_mul_cancel_left₀ h] using H.mul_left a₂⁻¹, HasSum.mul_left _⟩
94+
#align has_sum_mul_left_iff hasSum_mul_left_iff
95+
96+
theorem hasSum_mul_right_iff (h : a₂ ≠ 0) : HasSum (fun i => f i * a₂) (a₁ * a₂) ↔ HasSum f a₁ :=
97+
fun H => by simpa only [mul_inv_cancel_right₀ h] using H.mul_right a₂⁻¹, HasSum.mul_right _⟩
98+
#align has_sum_mul_right_iff hasSum_mul_right_iff
99+
100+
theorem hasSum_div_const_iff (h : a₂ ≠ 0) : HasSum (fun i => f i / a₂) (a₁ / a₂) ↔ HasSum f a₁ := by
101+
simpa only [div_eq_mul_inv] using hasSum_mul_right_iff (inv_ne_zero h)
102+
#align has_sum_div_const_iff hasSum_div_const_iff
103+
104+
theorem summable_mul_left_iff (h : a ≠ 0) : (Summable fun i => a * f i) ↔ Summable f :=
105+
fun H => by simpa only [inv_mul_cancel_left₀ h] using H.mul_left a⁻¹, fun H => H.mul_left _⟩
106+
#align summable_mul_left_iff summable_mul_left_iff
107+
108+
theorem summable_mul_right_iff (h : a ≠ 0) : (Summable fun i => f i * a) ↔ Summable f :=
109+
fun H => by simpa only [mul_inv_cancel_right₀ h] using H.mul_right a⁻¹, fun H => H.mul_right _⟩
110+
#align summable_mul_right_iff summable_mul_right_iff
111+
112+
theorem summable_div_const_iff (h : a ≠ 0) : (Summable fun i => f i / a) ↔ Summable f := by
113+
simpa only [div_eq_mul_inv] using summable_mul_right_iff (inv_ne_zero h)
114+
#align summable_div_const_iff summable_div_const_iff
115+
116+
theorem tsum_mul_left [T2Space α] : (∑' x, a * f x) = a * ∑' x, f x :=
117+
if hf : Summable f then hf.tsum_mul_left a
118+
else if ha : a = 0 then by simp [ha]
119+
else by rw [tsum_eq_zero_of_not_summable hf,
120+
tsum_eq_zero_of_not_summable (mt (summable_mul_left_iff ha).mp hf), mul_zero]
121+
#align tsum_mul_left tsum_mul_left
122+
123+
theorem tsum_mul_right [T2Space α] : (∑' x, f x * a) = (∑' x, f x) * a :=
124+
if hf : Summable f then hf.tsum_mul_right a
125+
else if ha : a = 0 then by simp [ha]
126+
else by rw [tsum_eq_zero_of_not_summable hf,
127+
tsum_eq_zero_of_not_summable (mt (summable_mul_right_iff ha).mp hf), zero_mul]
128+
#align tsum_mul_right tsum_mul_right
129+
130+
theorem tsum_div_const [T2Space α] : (∑' x, f x / a) = (∑' x, f x) / a := by
131+
simpa only [div_eq_mul_inv] using tsum_mul_right
132+
#align tsum_div_const tsum_div_const
133+
134+
end DivisionSemiring
135+
136+
/-!
137+
### Multipliying two infinite sums
138+
139+
In this section, we prove various results about `(∑' x : ι, f x) * (∑' y : κ, g y)`. Note that we
140+
always assume that the family `λ x : ι × κ, f x.1 * g x.2` is summable, since there is no way to
141+
deduce this from the summmabilities of `f` and `g` in general, but if you are working in a normed
142+
space, you may want to use the analogous lemmas in `Analysis/NormedSpace/Basic`
143+
(e.g `tsum_mul_tsum_of_summable_norm`).
144+
145+
We first establish results about arbitrary index types, `ι` and `κ`, and then we specialize to
146+
`ι = κ = ℕ` to prove the Cauchy product formula (see `tsum_mul_tsum_eq_tsum_sum_antidiagonal`).
147+
148+
#### Arbitrary index types
149+
-/
150+
151+
152+
section tsum_mul_tsum
153+
154+
variable [TopologicalSpace α] [T3Space α] [NonUnitalNonAssocSemiring α] [TopologicalSemiring α]
155+
{f : ι → α} {g : κ → α} {s t u : α}
156+
157+
theorem HasSum.mul_eq (hf : HasSum f s) (hg : HasSum g t)
158+
(hfg : HasSum (fun x : ι × κ => f x.1 * g x.2) u) : s * t = u :=
159+
have key₁ : HasSum (fun i => f i * t) (s * t) := hf.mul_right t
160+
have this : ∀ i : ι, HasSum (fun c : κ => f i * g c) (f i * t) := fun i => hg.mul_left (f i)
161+
have key₂ : HasSum (fun i => f i * t) u := HasSum.prod_fiberwise hfg this
162+
key₁.unique key₂
163+
#align has_sum.mul_eq HasSum.mul_eq
164+
165+
theorem HasSum.mul (hf : HasSum f s) (hg : HasSum g t)
166+
(hfg : Summable fun x : ι × κ => f x.1 * g x.2) :
167+
HasSum (fun x : ι × κ => f x.1 * g x.2) (s * t) :=
168+
let ⟨_u, hu⟩ := hfg
169+
(hf.mul_eq hg hu).symm ▸ hu
170+
#align has_sum.mul HasSum.mul
171+
172+
/-- Product of two infinites sums indexed by arbitrary types.
173+
See also `tsum_mul_tsum_of_summable_norm` if `f` and `g` are abolutely summable. -/
174+
theorem tsum_mul_tsum (hf : Summable f) (hg : Summable g)
175+
(hfg : Summable fun x : ι × κ => f x.1 * g x.2) :
176+
((∑' x, f x) * ∑' y, g y) = ∑' z : ι × κ, f z.1 * g z.2 :=
177+
hf.hasSum.mul_eq hg.hasSum hfg.hasSum
178+
#align tsum_mul_tsum tsum_mul_tsum
179+
180+
end tsum_mul_tsum
181+
182+
/-!
183+
#### `ℕ`-indexed families (Cauchy product)
184+
185+
We prove two versions of the Cauchy product formula. The first one is
186+
`tsum_mul_tsum_eq_tsum_sum_range`, where the `n`-th term is a sum over `Finset.range (n+1)`
187+
involving `Nat` subtraction.
188+
In order to avoid `Nat` subtraction, we also provide `tsum_mul_tsum_eq_tsum_sum_antidiagonal`,
189+
where the `n`-th term is a sum over all pairs `(k, l)` such that `k+l=n`, which corresponds to the
190+
`Finset` `Finset.Nat.antidiagonal n`
191+
-/
192+
193+
194+
section CauchyProduct
195+
196+
variable [TopologicalSpace α] [NonUnitalNonAssocSemiring α] {f g : ℕ → α}
197+
198+
/- The family `(k, l) : ℕ × ℕ ↦ f k * g l` is summable if and only if the family
199+
`(n, k, l) : Σ (n : ℕ), Nat.antidiagonal n ↦ f k * g l` is summable. -/
200+
theorem summable_mul_prod_iff_summable_mul_sigma_antidiagonal :
201+
(Summable fun x : ℕ × ℕ => f x.1 * g x.2) ↔
202+
Summable fun x : Σn : ℕ, Nat.antidiagonal n => f (x.2 : ℕ × ℕ).1 * g (x.2 : ℕ × ℕ).2 :=
203+
Nat.sigmaAntidiagonalEquivProd.summable_iff.symm
204+
#align summable_mul_prod_iff_summable_mul_sigma_antidiagonal summable_mul_prod_iff_summable_mul_sigma_antidiagonal
205+
206+
variable [T3Space α] [TopologicalSemiring α]
207+
208+
theorem summable_sum_mul_antidiagonal_of_summable_mul
209+
(h : Summable fun x : ℕ × ℕ => f x.1 * g x.2) :
210+
Summable fun n => ∑ kl in Nat.antidiagonal n, f kl.1 * g kl.2 := by
211+
rw [summable_mul_prod_iff_summable_mul_sigma_antidiagonal] at h
212+
conv => congr; ext; rw [← Finset.sum_finset_coe, ← tsum_fintype]
213+
exact h.sigma' fun n => (hasSum_fintype _).summable
214+
#align summable_sum_mul_antidiagonal_of_summable_mul summable_sum_mul_antidiagonal_of_summable_mul
215+
216+
/-- The **Cauchy product formula** for the product of two infinites sums indexed by `ℕ`, expressed
217+
by summing on `Finset.Nat.antidiagonal`.
218+
219+
See also `tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm` if `f` and `g` are absolutely
220+
summable. -/
221+
theorem tsum_mul_tsum_eq_tsum_sum_antidiagonal (hf : Summable f) (hg : Summable g)
222+
(hfg : Summable fun x : ℕ × ℕ => f x.1 * g x.2) :
223+
((∑' n, f n) * ∑' n, g n) = ∑' n, ∑ kl in Nat.antidiagonal n, f kl.1 * g kl.2 := by
224+
conv_rhs => congr; ext; rw [← Finset.sum_finset_coe, ← tsum_fintype]
225+
rw [tsum_mul_tsum hf hg hfg, ← Nat.sigmaAntidiagonalEquivProd.tsum_eq (_ : ℕ × ℕ → α)]
226+
exact
227+
tsum_sigma' (fun n => (hasSum_fintype _).summable)
228+
(summable_mul_prod_iff_summable_mul_sigma_antidiagonal.mp hfg)
229+
#align tsum_mul_tsum_eq_tsum_sum_antidiagonal tsum_mul_tsum_eq_tsum_sum_antidiagonal
230+
231+
theorem summable_sum_mul_range_of_summable_mul (h : Summable fun x : ℕ × ℕ => f x.1 * g x.2) :
232+
Summable fun n => ∑ k in range (n + 1), f k * g (n - k) := by
233+
simp_rw [← Nat.sum_antidiagonal_eq_sum_range_succ fun k l => f k * g l]
234+
exact summable_sum_mul_antidiagonal_of_summable_mul h
235+
#align summable_sum_mul_range_of_summable_mul summable_sum_mul_range_of_summable_mul
236+
237+
/-- The **Cauchy product formula** for the product of two infinites sums indexed by `ℕ`, expressed
238+
by summing on `Finset.range`.
239+
240+
See also `tsum_mul_tsum_eq_tsum_sum_range_of_summable_norm` if `f` and `g` are absolutely summable.
241+
-/
242+
theorem tsum_mul_tsum_eq_tsum_sum_range (hf : Summable f) (hg : Summable g)
243+
(hfg : Summable fun x : ℕ × ℕ => f x.1 * g x.2) :
244+
((∑' n, f n) * ∑' n, g n) = ∑' n, ∑ k in range (n + 1), f k * g (n - k) := by
245+
simp_rw [← Nat.sum_antidiagonal_eq_sum_range_succ fun k l => f k * g l]
246+
exact tsum_mul_tsum_eq_tsum_sum_antidiagonal hf hg hfg
247+
#align tsum_mul_tsum_eq_tsum_sum_range tsum_mul_tsum_eq_tsum_sum_range
248+
249+
end CauchyProduct

0 commit comments

Comments
 (0)