Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 01a8a9f

Browse files
committed
feat(convex/specific_functions): specific case of Jensen's inequality for powers (#16186)
Proves a specific case of Jensen's inequality: a powers of a sum divided by the cardinality of the `finset` is less than or equal to the sum of the powers.
1 parent 3c93c65 commit 01a8a9f

File tree

1 file changed

+25
-1
lines changed

1 file changed

+25
-1
lines changed

src/analysis/convex/specific_functions.lean

Lines changed: 25 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -31,7 +31,7 @@ For `p : ℝ`, prove that `λ x, x ^ p` is concave when `0 ≤ p ≤ 1` and stri
3131
-/
3232

3333
open real set
34-
open_locale big_operators
34+
open_locale big_operators nnreal
3535

3636
/-- `exp` is strictly convex on the whole real line. -/
3737
lemma strict_convex_on_exp : strict_convex_on ℝ univ exp :=
@@ -82,6 +82,30 @@ begin
8282
nat.sub_pos_of_lt hn) (nat.cast_pos.2 $ zero_lt_two.trans_le hn),
8383
end
8484

85+
/-- Specific case of Jensen's inequality for sums of powers -/
86+
lemma real.pow_sum_div_card_le_sum_pow {α : Type*} {s : finset α} {f : α → ℝ} (n : ℕ)
87+
(hf : ∀ a ∈ s, 0 ≤ f a) : (∑ x in s, f x) ^ (n + 1) / s.card ^ n ≤ ∑ x in s, (f x) ^ (n + 1) :=
88+
begin
89+
by_cases hs0 : s = ∅,
90+
{ simp_rw [hs0, finset.sum_empty, zero_pow' _ (nat.succ_ne_zero n), zero_div] },
91+
{ have hs : s.card ≠ 0 := hs0 ∘ finset.card_eq_zero.1,
92+
have hs' : (s.card : ℝ) ≠ 0 := (nat.cast_ne_zero.2 hs),
93+
have hs'' : 0 < (s.card : ℝ) := nat.cast_pos.2 (nat.pos_of_ne_zero hs),
94+
suffices : (∑ x in s, f x / s.card) ^ (n + 1) ≤ ∑ x in s, (f x ^ (n + 1) / s.card),
95+
by rwa [← finset.sum_div, ← finset.sum_div, div_pow, pow_succ' (s.card : ℝ),
96+
← div_div, div_le_iff hs'', div_mul, div_self hs', div_one] at this,
97+
have := @convex_on.map_sum_le ℝ ℝ ℝ α _ _ _ _ _ _ (set.Ici 0) (λ x, x ^ (n + 1)) s
98+
(λ _, 1 / s.card) (coe ∘ f) (convex_on_pow (n + 1)) _ _ (λ i hi, set.mem_Ici.2 (hf i hi)),
99+
{ simpa only [inv_mul_eq_div, one_div, algebra.id.smul_eq_mul] using this },
100+
{ simp only [one_div, inv_nonneg, nat.cast_nonneg, implies_true_iff] },
101+
{ simpa only [one_div, finset.sum_const, nsmul_eq_mul] using mul_inv_cancel hs' }}
102+
end
103+
104+
lemma nnreal.pow_sum_div_card_le_sum_pow {α : Type*} (s : finset α) (f : α → ℝ≥0) (n : ℕ) :
105+
(∑ x in s, f x) ^ (n + 1) / s.card ^ n ≤ ∑ x in s, (f x) ^ (n + 1) :=
106+
by simpa only [← nnreal.coe_le_coe, nnreal.coe_sum, nonneg.coe_div, nnreal.coe_pow] using
107+
@real.pow_sum_div_card_le_sum_pow α s (coe ∘ f) n (λ _ _, nnreal.coe_nonneg _)
108+
85109
lemma finset.prod_nonneg_of_card_nonpos_even
86110
{α β : Type*} [linear_ordered_comm_ring β]
87111
{f : α → β} [decidable_pred (λ x, f x ≤ 0)]

0 commit comments

Comments
 (0)