Skip to content

Commit e4e10f9

Browse files
feat(Topology/Algebra/InfiniteSum/Basic): add some lemmas on tsums (#10038)
This is the fifth PR in a sequence that adds auxiliary lemmas from the [EulerProducts](https://github.com/MichaelStollBayreuth/EulerProducts) project to Mathlib. It adds three lemmas on `tsum`s: ```lean lemma HasSum.tsum_fiberwise {α β γ : Type*} [AddCommGroup α] [UniformSpace α] [UniformAddGroup α] [T2Space α] [RegularSpace α] [CompleteSpace α] {f : β → α} {a : α} (hf : HasSum f a) (g : β → γ) : HasSum (fun c : γ ↦ ∑' b : g ⁻¹' {c}, f b) a lemma tsum_setProd_singleton_left {α β γ : Type*} [AddCommMonoid γ] [TopologicalSpace γ] [T2Space γ] (a : α) (t : Set β) (f : α × β → γ) : (∑' x : {a} ×ˢ t, f x) = ∑' b : t, f (a, b) lemma tsum_setProd_singleton_right {α β γ : Type*} [AddCommMonoid γ] [TopologicalSpace γ] [T2Space γ] (s : Set α) (b : β) (f : α × β → γ) : (∑' x : s ×ˢ {b}, f x) = ∑' a : s, f (a, b) ``` and the necessary equivalences ```lean def prod_singleton_left {α β : Type*} (a : α) (t : Set β) : ↑({a} ×ˢ t) ≃ ↑t def prod_singleton_right {α β : Type*} (s : Set α) (b : β) : ↑(s ×ˢ {b}) ≃ ↑s ```
1 parent 0bcd6be commit e4e10f9

File tree

2 files changed

+37
-0
lines changed

2 files changed

+37
-0
lines changed

Mathlib/Data/Set/Prod.lean

Lines changed: 18 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1002,3 +1002,21 @@ theorem sumPiEquivProdPi_symm_preimage_univ_pi (π : ι ⊕ ι' → Type*) (t :
10021002
· rintro ⟨h₁, h₂⟩ (i|i) <;> simp <;> apply_assumption
10031003

10041004
end Equiv
1005+
1006+
namespace Equiv.Set
1007+
1008+
/-- The canonical equivalence between `{a} ×ˢ t` and `t`, considered as types. -/
1009+
def prod_singleton_left {α β : Type*} (a : α) (t : Set β) : ↑({a} ×ˢ t) ≃ ↑t where
1010+
toFun := fun x ↦ ⟨x.val.snd, (Set.mem_prod.mp x.prop).2
1011+
invFun := fun b ↦ ⟨(a, b.val), Set.mem_prod.mpr ⟨Set.mem_singleton a, Subtype.mem b⟩⟩
1012+
left_inv := by simp [Function.LeftInverse]
1013+
right_inv := by simp [Function.RightInverse, Function.LeftInverse]
1014+
1015+
/-- The canonical equivalence between `s ×ˢ {b}` and `s`, considered as types. -/
1016+
def prod_singleton_right {α β : Type*} (s : Set α) (b : β) : ↑(s ×ˢ {b}) ≃ ↑s where
1017+
toFun := fun x ↦ ⟨x.val.fst, (Set.mem_prod.mp x.prop).1
1018+
invFun := fun a ↦ ⟨(a.val, b), Set.mem_prod.mpr ⟨Subtype.mem a, Set.mem_singleton b⟩⟩
1019+
left_inv := by simp [Function.LeftInverse]
1020+
right_inv := by simp [Function.RightInverse, Function.LeftInverse]
1021+
1022+
end Equiv.Set

Mathlib/Topology/Algebra/InfiniteSum/Basic.lean

Lines changed: 19 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1284,8 +1284,27 @@ theorem sum_add_tsum_subtype_compl [T2Space α] {f : β → α} (hf : Summable f
12841284
rfl
12851285
#align sum_add_tsum_subtype_compl sum_add_tsum_subtype_compl
12861286

1287+
lemma HasSum.tsum_fiberwise [T2Space α] {f : β → α} {a : α} (hf : HasSum f a) (g : β → γ) :
1288+
HasSum (fun c : γ ↦ ∑' b : g ⁻¹' {c}, f b) a :=
1289+
(((Equiv.sigmaFiberEquiv g).hasSum_iff).mpr hf).sigma <|
1290+
fun _ ↦ ((hf.summable.subtype _).hasSum_iff).mpr rfl
1291+
12871292
end UniformGroup
12881293

1294+
section prod_singleton
1295+
1296+
variable [AddCommMonoid γ] [TopologicalSpace γ] [T2Space γ]
1297+
1298+
lemma tsum_setProd_singleton_left (a : α) (t : Set β) (f : α × β → γ) :
1299+
(∑' x : {a} ×ˢ t, f x) = ∑' b : t, f (a, b) :=
1300+
(Equiv.Set.prod_singleton_left a t |>.symm.tsum_eq <| ({a} ×ˢ t).restrict f).symm
1301+
1302+
lemma tsum_setProd_singleton_right (s : Set α) (b : β) (f : α × β → γ) :
1303+
(∑' x : s ×ˢ {b}, f x) = ∑' a : s, f (a, b) :=
1304+
(Equiv.Set.prod_singleton_right s b |>.symm.tsum_eq <| (s ×ˢ {b}).restrict f).symm
1305+
1306+
end prod_singleton
1307+
12891308
section TopologicalGroup
12901309

12911310
variable {G : Type*} [TopologicalSpace G] [AddCommGroup G] [TopologicalAddGroup G] {f : α → G}

0 commit comments

Comments
 (0)