Skip to content

Commit

Permalink
feat(topology/algebra/infinite_sum): generalize `tsum_le_of_sum_range…
Browse files Browse the repository at this point in the history
…_le` (#16951)
  • Loading branch information
JasonKYi committed Oct 14, 2022
1 parent e623a7e commit f6b1276
Show file tree
Hide file tree
Showing 2 changed files with 18 additions and 4 deletions.
11 changes: 11 additions & 0 deletions src/topology/algebra/infinite_sum.lean
Expand Up @@ -1395,6 +1395,17 @@ end

end topological_group

section preorder

variables {E : Type*} [preorder E] [add_comm_monoid E]
[topological_space E] [order_closed_topology E] [t2_space E]

lemma tsum_le_of_sum_range_le {f : ℕ → E} {c : E} (hsum : summable f)
(h : ∀ n, ∑ i in finset.range n, f i ≤ c) : ∑' n, f n ≤ c :=
let ⟨l, hl⟩ := hsum in hl.tsum_eq.symm ▸ le_of_tendsto' hl.tendsto_sum_nat h

end preorder

section linear_order

/-! For infinite sums taking values in a linearly ordered monoid, the existence of a least upper
Expand Down
11 changes: 7 additions & 4 deletions src/topology/instances/ennreal.lean
Expand Up @@ -1055,7 +1055,7 @@ end

lemma tsum_le_of_sum_range_le {f : ℕ → ℝ≥0} {c : ℝ≥0}
(h : ∀ n, ∑ i in finset.range n, f i ≤ c) : ∑' n, f n ≤ c :=
le_of_tendsto' (has_sum_iff_tendsto_nat.1 (summable_of_sum_range_le h).has_sum) h
tsum_le_of_sum_range_le (summable_of_sum_range_le h) h

lemma tsum_comp_le_tsum_of_inj {β : Type*} {f : α → ℝ≥0} (hf : summable f)
{i : β → α} (hi : function.injective i) : ∑' x, f (i x) ≤ ∑' x, f x :=
Expand Down Expand Up @@ -1145,6 +1145,10 @@ begin
exact_mod_cast nnreal.tendsto_sum_nat_add f
end

lemma tsum_le_of_sum_range_le {f : ℕ → ℝ≥0∞} {c : ℝ≥0∞}
(h : ∀ n, ∑ i in finset.range n, f i ≤ c) : ∑' n, f n ≤ c :=
tsum_le_of_sum_range_le ennreal.summable h

end ennreal

lemma tsum_comp_le_tsum_of_inj {β : Type*} {f : α → ℝ} (hf : summable f) (hn : ∀ a, 0 ≤ f a)
Expand Down Expand Up @@ -1209,10 +1213,9 @@ begin
exact lt_irrefl _ (hn.trans_le (h n)),
end

lemma tsum_le_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n)
lemma real.tsum_le_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n)
(h : ∀ n, ∑ i in finset.range n, f i ≤ c) : ∑' n, f n ≤ c :=
le_of_tendsto' ((has_sum_iff_tendsto_nat_of_nonneg hf _).1
(summable_of_sum_range_le hf h).has_sum) h
tsum_le_of_sum_range_le (summable_of_sum_range_le hf h) h

/-- If a sequence `f` with non-negative terms is dominated by a sequence `g` with summable
series and at least one term of `f` is strictly smaller than the corresponding term in `g`,
Expand Down

0 comments on commit f6b1276

Please sign in to comment.