Skip to content

Commit 74c7a5c

Browse files
committed
feat(Topology/Algebra/InfiniteSum/NatInt): add more pnat tsum lemmas (#27841)
1 parent 8d736b5 commit 74c7a5c

File tree

4 files changed

+36
-8
lines changed

4 files changed

+36
-8
lines changed

Mathlib/Analysis/SpecialFunctions/Trigonometric/Cotangent.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -226,7 +226,7 @@ lemma cot_series_rep' (hz : x ∈ ℂ_ℤ) : π * cot (π * x) - 1 / x =
226226
/-- The cotangent infinite sum representation. -/
227227
theorem cot_series_rep (hz : x ∈ ℂ_ℤ) :
228228
π * cot (π * x) = 1 / x + ∑' n : ℕ+, (1 / (x - n) + 1 / (x + n)) := by
229-
have h0 := tsum_pnat_eq_tsum_succ fun n ↦ 1 / (x - n) + 1 / (x + n)
229+
have h0 := tsum_pnat_eq_tsum_succ (f := fun n ↦ 1 / (x - n) + 1 / (x + n))
230230
have h1 := cot_series_rep' hz
231231
simp only [one_div, Nat.cast_add, Nat.cast_one] at *
232232
rw [h0, ← h1]

Mathlib/NumberTheory/TsumDivsorsAntidiagonal.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -113,7 +113,7 @@ lemma tsum_pow_div_one_sub_eq_tsum_sigma {r : 𝕜} (hr : ‖r‖ < 1) (k : ℕ)
113113
have (m : ℕ) [NeZero m] := tsum_geometric_of_norm_lt_one (ξ := r ^ m)
114114
(by simpa using pow_lt_one₀ (by simp) hr (NeZero.ne _))
115115
simp only [div_eq_mul_inv, ← this, ← tsum_mul_left, mul_assoc, ← _root_.pow_succ',
116-
fun (n : ℕ) ↦ tsum_pnat_eq_tsum_succ (fun m ↦ n ^ k * (r ^ n) ^ m)]
116+
fun (n : ℕ) ↦ tsum_pnat_eq_tsum_succ (f := fun m ↦ n ^ k * (r ^ n) ^ m)]
117117
have h00 := tsum_prod_pow_eq_tsum_sigma k hr
118118
rw [Summable.tsum_comm (by apply (summable_prod_mul_pow k hr).prod_symm)] at h00
119119
rw [← h00]

Mathlib/Topology/Algebra/InfiniteSum/Basic.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -461,6 +461,11 @@ theorem Function.Injective.tprod_eq {g : γ → β} (hg : Injective g) {f : β
461461
theorem Equiv.tprod_eq (e : γ ≃ β) (f : β → α) : ∏' c, f (e c) = ∏' b, f b :=
462462
e.injective.tprod_eq <| by simp
463463

464+
@[to_additive (attr := simp)]
465+
theorem tprod_comp_neg {β : Type*} [InvolutiveNeg β] (f : β → α) :
466+
∏' d, f (-d) = ∏' d, f d :=
467+
(Equiv.neg β).tprod_eq f
468+
464469
/-! ### `tprod` on subsets - part 1 -/
465470

466471
@[to_additive]

Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean

Lines changed: 29 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -545,15 +545,38 @@ end IsUniformGroup
545545

546546
end Int
547547

548-
section pnat
548+
section PNat
549549

550550
@[to_additive]
551-
theorem pnat_multipliable_iff_multipliable_succ : Type*} [TopologicalSpace α] [CommMonoid α]
552-
{f : ℕ → α} : Multipliable (fun x : ℕ+ => f x) ↔ Multipliable fun x : ℕ => f (x + 1) :=
551+
theorem multipliable_pnat_iff_multipliable_succ {f : ℕ → M} :
552+
Multipliable (fun x : ℕ+ f x) ↔ Multipliable fun x f (x + 1) :=
553553
Equiv.pnatEquivNat.symm.multipliable_iff.symm
554554

555+
@[deprecated (since := "2025-09-31")]
556+
alias pnat_multipliable_iff_multipliable_succ := multipliable_pnat_iff_multipliable_succ
557+
555558
@[to_additive]
556-
theorem tprod_pnat_eq_tprod_succ {α : Type*} [TopologicalSpace α] [CommMonoid α] (f : ℕ → α) :
557-
∏' n : ℕ+, f n = ∏' n, f (n + 1) := (Equiv.pnatEquivNat.symm.tprod_eq _).symm
559+
theorem tprod_pnat_eq_tprod_succ {f : ℕ → M} : ∏' n : ℕ+, f n = ∏' n, f (n + 1) :=
560+
(Equiv.pnatEquivNat.symm.tprod_eq _).symm
558561

559-
end pnat
562+
@[to_additive]
563+
lemma tprod_zero_pnat_eq_tprod_nat [TopologicalSpace G] [IsTopologicalGroup G] [T2Space G]
564+
{f : ℕ → G} (hf : Multipliable f) :
565+
f 0 * ∏' n : ℕ+, f ↑n = ∏' n, f n := by
566+
simpa [hf.tprod_eq_zero_mul] using tprod_pnat_eq_tprod_succ
567+
568+
@[to_additive tsum_int_eq_zero_add_two_mul_tsum_pnat]
569+
theorem tprod_int_eq_zero_mul_tprod_pnat_sq [UniformSpace G] [IsUniformGroup G] [CompleteSpace G]
570+
[T2Space G] {f : ℤ → G} (hf : ∀ n : ℤ, f (-n) = f n) (hf2 : Multipliable f) :
571+
∏' n, f n = f 0 * (∏' n : ℕ+, f n) ^ 2 := by
572+
have hf3 : Multipliable fun n : ℕ ↦ f n :=
573+
(multipliable_int_iff_multipliable_nat_and_neg.mp hf2).1
574+
have hf4 : Multipliable fun n : ℕ+ ↦ f n := by
575+
rwa [multipliable_pnat_iff_multipliable_succ (f := (f ·)),
576+
multipliable_nat_add_iff 1 (f := (f ·))]
577+
have := tprod_nat_mul_neg hf2
578+
rw [← tprod_zero_pnat_eq_tprod_nat (by simpa [hf] using hf3.mul hf3), mul_comm _ (f 0)] at this
579+
simp only [hf, Nat.cast_zero, mul_assoc, mul_right_inj] at this
580+
rw [← this, mul_right_inj, hf4.tprod_mul hf4, sq]
581+
582+
end PNat

0 commit comments

Comments
 (0)