diff --git a/src/analysis/analytic/basic.lean b/src/analysis/analytic/basic.lean index 67fcad15219a6..bb02eaffddf97 100644 --- a/src/analysis/analytic/basic.lean +++ b/src/analysis/analytic/basic.lean @@ -26,7 +26,7 @@ space is analytic, as well as the inverse on invertible operators. Let `p` be a formal multilinear series from `E` to `F`, i.e., `p n` is a multilinear map on `E^n` for `n : ℕ`. -* `p.radius`: the largest `r : ennreal` such that `∥p n∥ * r^n` grows subexponentially, defined as +* `p.radius`: the largest `r : ℝ≥0∞` such that `∥p n∥ * r^n` grows subexponentially, defined as a liminf. * `p.le_radius_of_bound`, `p.le_radius_of_bound_nnreal`, `p.le_radius_of_is_O`: if `∥p n∥ * r ^ n` is bounded above, then `r ≤ p.radius`; @@ -73,7 +73,7 @@ variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] {F : Type*} [normed_group F] [normed_space 𝕜 F] {G : Type*} [normed_group G] [normed_space 𝕜 G] -open_locale topological_space classical big_operators nnreal filter +open_locale topological_space classical big_operators nnreal filter ennreal open set filter asymptotics /-! ### The radius of a formal multilinear series -/ @@ -84,17 +84,17 @@ variables (p : formal_multilinear_series 𝕜 E F) {r : ℝ≥0} /-- The radius of a formal multilinear series is the largest `r` such that the sum `Σ pₙ yⁿ` converges for all `∥y∥ < r`. -/ -def radius (p : formal_multilinear_series 𝕜 E F) : ennreal := -⨆ (r : ℝ≥0) (C : ℝ) (hr : ∀ n, ∥p n∥ * r ^ n ≤ C), (r : ennreal) +def radius (p : formal_multilinear_series 𝕜 E F) : ℝ≥0∞ := +⨆ (r : ℝ≥0) (C : ℝ) (hr : ∀ n, ∥p n∥ * r ^ n ≤ C), (r : ℝ≥0∞) /-- If `∥pₙ∥ rⁿ` is bounded in `n`, then the radius of `p` is at least `r`. -/ lemma le_radius_of_bound (p : formal_multilinear_series 𝕜 E F) (C : ℝ) {r : ℝ≥0} - (h : ∀ (n : ℕ), ∥p n∥ * r^n ≤ C) : (r : ennreal) ≤ p.radius := -le_supr_of_le r $ le_supr_of_le C $ (le_supr (λ _, (r : ennreal)) h) + (h : ∀ (n : ℕ), ∥p n∥ * r^n ≤ C) : (r : ℝ≥0∞) ≤ p.radius := +le_supr_of_le r $ le_supr_of_le C $ (le_supr (λ _, (r : ℝ≥0∞)) h) /-- If `∥pₙ∥ rⁿ` is bounded in `n`, then the radius of `p` is at least `r`. -/ lemma le_radius_of_bound_nnreal (p : formal_multilinear_series 𝕜 E F) (C : ℝ≥0) {r : ℝ≥0} - (h : ∀ (n : ℕ), nnnorm (p n) * r^n ≤ C) : (r : ennreal) ≤ p.radius := + (h : ∀ (n : ℕ), nnnorm (p n) * r^n ≤ C) : (r : ℝ≥0∞) ≤ p.radius := p.le_radius_of_bound C $ λ n, by exact_mod_cast (h n) /-- If `∥pₙ∥ rⁿ = O(1)`, as `n → ∞`, then the radius of `p` is at least `r`. -/ @@ -166,19 +166,19 @@ end /-- For `r` strictly smaller than the radius of `p`, then `∥pₙ∥ rⁿ` is bounded. -/ lemma norm_mul_pow_le_of_lt_radius (p : formal_multilinear_series 𝕜 E F) {r : ℝ≥0} - (h : (r : ennreal) < p.radius) : ∃ C > 0, ∀ n, ∥p n∥ * r^n ≤ C := + (h : (r : ℝ≥0∞) < p.radius) : ∃ C > 0, ∀ n, ∥p n∥ * r^n ≤ C := let ⟨a, ha, C, hC, h⟩ := p.norm_mul_pow_le_mul_pow_of_lt_radius h in ⟨C, hC, λ n, (h n).trans $ mul_le_of_le_one_right hC.lt.le (pow_le_one _ ha.1.le ha.2.le)⟩ /-- For `r` strictly smaller than the radius of `p`, then `∥pₙ∥ rⁿ` is bounded. -/ lemma norm_le_div_pow_of_pos_of_lt_radius (p : formal_multilinear_series 𝕜 E F) {r : ℝ≥0} - (h0 : 0 < r) (h : (r : ennreal) < p.radius) : ∃ C > 0, ∀ n, ∥p n∥ ≤ C / r ^ n := + (h0 : 0 < r) (h : (r : ℝ≥0∞) < p.radius) : ∃ C > 0, ∀ n, ∥p n∥ ≤ C / r ^ n := let ⟨C, hC, hp⟩ := p.norm_mul_pow_le_of_lt_radius h in ⟨C, hC, λ n, iff.mpr (le_div_iff (pow_pos h0 _)) (hp n)⟩ /-- For `r` strictly smaller than the radius of `p`, then `∥pₙ∥ rⁿ` is bounded. -/ lemma nnnorm_mul_pow_le_of_lt_radius (p : formal_multilinear_series 𝕜 E F) {r : ℝ≥0} - (h : (r : ennreal) < p.radius) : ∃ C > 0, ∀ n, nnnorm (p n) * r^n ≤ C := + (h : (r : ℝ≥0∞) < p.radius) : ∃ C > 0, ∀ n, nnnorm (p n) * r^n ≤ C := let ⟨C, hC, hp⟩ := p.norm_mul_pow_le_of_lt_radius h in ⟨⟨C, hC.lt.le⟩, hC, by exact_mod_cast hp⟩ @@ -216,13 +216,13 @@ end formal_multilinear_series /-! ### Expanding a function as a power series -/ section -variables {f g : E → F} {p pf pg : formal_multilinear_series 𝕜 E F} {x : E} {r r' : ennreal} +variables {f g : E → F} {p pf pg : formal_multilinear_series 𝕜 E F} {x : E} {r r' : ℝ≥0∞} /-- Given a function `f : E → F` and a formal multilinear series `p`, we say that `f` has `p` as a power series on the ball of radius `r > 0` around `x` if `f (x + y) = ∑' pₙ yⁿ` for all `∥y∥ < r`. -/ structure has_fpower_series_on_ball - (f : E → F) (p : formal_multilinear_series 𝕜 E F) (x : E) (r : ennreal) : Prop := + (f : E → F) (p : formal_multilinear_series 𝕜 E F) (x : E) (r : ℝ≥0∞) : Prop := (r_le : r ≤ p.radius) (r_pos : 0 < r) (has_sum : ∀ {y}, y ∈ emetric.ball (0 : E) r → has_sum (λn:ℕ, p n (λ(i : fin n), y)) (f (x + y))) @@ -270,7 +270,7 @@ lemma has_fpower_series_on_ball.mono ⟨le_trans hr hf.1, r'_pos, λ y hy, hf.has_sum (emetric.ball_subset_ball hr hy)⟩ protected lemma has_fpower_series_at.eventually (hf : has_fpower_series_at f p x) : - ∀ᶠ r : ennreal in 𝓝[Ioi 0] 0, has_fpower_series_on_ball f p x r := + ∀ᶠ r : ℝ≥0∞ in 𝓝[Ioi 0] 0, has_fpower_series_on_ball f p x r := let ⟨r, hr⟩ := hf in mem_sets_of_superset (Ioo_mem_nhds_within_Ioi (left_mem_Ico.2 hr.r_pos)) $ λ r' hr', hr.mono hr'.1 hr'.2.le @@ -344,7 +344,7 @@ sums of this power series on strict subdisks of the disk of convergence. This version provides an upper estimate that decreases both in `∥y∥` and `n`. See also `has_fpower_series_on_ball.uniform_geometric_approx` for a weaker version. -/ lemma has_fpower_series_on_ball.uniform_geometric_approx' {r' : ℝ≥0} - (hf : has_fpower_series_on_ball f p x r) (h : (r' : ennreal) < r) : + (hf : has_fpower_series_on_ball f p x r) (h : (r' : ℝ≥0∞) < r) : ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), (∀ y ∈ metric.ball (0 : E) r', ∀ n, ∥f (x + y) - p.partial_sum n y∥ ≤ C * (a * (∥y∥ / r')) ^ n) := begin @@ -376,7 +376,7 @@ end /-- If a function admits a power series expansion, then it is exponentially close to the partial sums of this power series on strict subdisks of the disk of convergence. -/ lemma has_fpower_series_on_ball.uniform_geometric_approx {r' : ℝ≥0} - (hf : has_fpower_series_on_ball f p x r) (h : (r' : ennreal) < r) : + (hf : has_fpower_series_on_ball f p x r) (h : (r' : ℝ≥0∞) < r) : ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), (∀ y ∈ metric.ball (0 : E) r', ∀ n, ∥f (x + y) - p.partial_sum n y∥ ≤ C * a ^ n) := begin @@ -496,7 +496,7 @@ end partial sums of this power series on strict subdisks of the disk of convergence, i.e., `f (x + y)` is the uniform limit of `p.partial_sum n y` there. -/ lemma has_fpower_series_on_ball.tendsto_uniformly_on {r' : ℝ≥0} - (hf : has_fpower_series_on_ball f p x r) (h : (r' : ennreal) < r) : + (hf : has_fpower_series_on_ball f p x r) (h : (r' : ℝ≥0∞) < r) : tendsto_uniformly_on (λ n y, p.partial_sum n y) (λ y, f (x + y)) at_top (metric.ball (0 : E) r') := begin @@ -532,7 +532,7 @@ end partial sums of this power series on strict subdisks of the disk of convergence, i.e., `f y` is the uniform limit of `p.partial_sum n (y - x)` there. -/ lemma has_fpower_series_on_ball.tendsto_uniformly_on' {r' : ℝ≥0} - (hf : has_fpower_series_on_ball f p x r) (h : (r' : ennreal) < r) : + (hf : has_fpower_series_on_ball f p x r) (h : (r' : ℝ≥0∞) < r) : tendsto_uniformly_on (λ n y, p.partial_sum n (y - x)) f at_top (metric.ball (x : E) r') := begin convert (hf.tendsto_uniformly_on h).comp (λ y, y - x), @@ -575,7 +575,7 @@ lemma formal_multilinear_series.has_fpower_series_on_ball [complete_space F] r_pos := h, has_sum := λ y hy, begin rw zero_add, - replace hy : (nnnorm y : ennreal) < p.radius, + replace hy : (nnnorm y : ℝ≥0∞) < p.radius, by { convert hy, exact (edist_eq_coe_nnnorm _).symm }, obtain ⟨a, ha : a ∈ Ioo (0 : ℝ) 1, C, hC : 0 < C, hp⟩ := p.norm_mul_pow_le_mul_pow_of_lt_radius hy, @@ -647,7 +647,7 @@ def change_origin (x : E) : formal_multilinear_series 𝕜 E F := -- Note here and below it is necessary to use `@` and provide implicit arguments using `_`, -- so that it is possible to use pattern matching in the lambda. -- Overall this seems a good trade-off in readability. -lemma change_origin_summable_aux1 (h : (nnnorm x + r : ennreal) < p.radius) : +lemma change_origin_summable_aux1 (h : (nnnorm x + r : ℝ≥0∞) < p.radius) : @summable ℝ _ _ _ ((λ ⟨n, s⟩, ∥p n∥ * ∥x∥ ^ (n - s.card) * r ^ s.card) : (Σ (n : ℕ), finset (fin n)) → ℝ) := begin @@ -667,7 +667,7 @@ end /-- Auxiliary lemma controlling the summability of the sequence appearing in the definition of `p.change_origin`, second version. -/ -lemma change_origin_summable_aux2 (h : (nnnorm x + r : ennreal) < p.radius) : +lemma change_origin_summable_aux2 (h : (nnnorm x + r : ℝ≥0∞) < p.radius) : @summable ℝ _ _ _ ((λ ⟨k, n, s, hs⟩, ∥(p n).restr s hs x∥ * ↑r ^ k) : (Σ (k : ℕ) (n : ℕ), {s : finset (fin n) // finset.card s = k}) → ℝ) := begin @@ -713,11 +713,11 @@ end /-- Auxiliary lemma controlling the summability of the sequence appearing in the definition of `p.change_origin`, third version. -/ -lemma change_origin_summable_aux3 (k : ℕ) (h : (nnnorm x : ennreal) < p.radius) : +lemma change_origin_summable_aux3 (k : ℕ) (h : (nnnorm x : ℝ≥0∞) < p.radius) : @summable ℝ _ _ _ (λ ⟨n, s, hs⟩, ∥(p n).restr s hs x∥ : (Σ (n : ℕ), {s : finset (fin n) // finset.card s = k}) → ℝ) := begin - obtain ⟨r, rpos, hr⟩ : ∃ (r : ℝ≥0), 0 < r ∧ ((nnnorm x + r) : ennreal) < p.radius := + obtain ⟨r, rpos, hr⟩ : ∃ (r : ℝ≥0), 0 < r ∧ ((nnnorm x + r) : ℝ≥0∞) < p.radius := ennreal.lt_iff_exists_add_pos_lt.mp h, have S : @summable ℝ _ _ _ ((λ ⟨n, s, hs⟩, ∥(p n).restr s hs x∥ * (r : ℝ) ^ k) : (Σ (n : ℕ), {s : finset (fin n) // finset.card s = k}) → ℝ), @@ -782,7 +782,7 @@ end variable [complete_space F] /-- The `k`-th coefficient of `p.change_origin` is the sum of a summable series. -/ -lemma change_origin_has_sum (k : ℕ) (h : (nnnorm x : ennreal) < p.radius) : +lemma change_origin_has_sum (k : ℕ) (h : (nnnorm x : ℝ≥0∞) < p.radius) : @has_sum (E [×k]→L[𝕜] F) _ _ _ ((λ i, (p i.1).restr i.2.1 i.2.2 x) : (Σ (n : ℕ), {s : finset (fin n) // finset.card s = k}) → (E [×k]→L[𝕜] F)) (p.change_origin x k) := @@ -795,7 +795,7 @@ begin end /-- Summing the series `p.change_origin x` at a point `y` gives back `p (x + y)`-/ -theorem change_origin_eval (h : (nnnorm x + nnnorm y : ennreal) < p.radius) : +theorem change_origin_eval (h : (nnnorm x + nnnorm y : ℝ≥0∞) < p.radius) : has_sum ((λk:ℕ, p.change_origin x k (λ (i : fin k), y))) (p.sum (x + y)) := begin /- The series on the left is a series of series. If we order the terms differently, we get back @@ -856,7 +856,7 @@ begin -- defining `p.change_origin`, by definition have J : ∀k, has_sum (λ c, A ⟨k, c⟩) (p.change_origin x k (λ(i : fin k), y)), { assume k, - have : (nnnorm x : ennreal) < radius p := lt_of_le_of_lt (le_add_right (le_refl _)) h, + have : (nnnorm x : ℝ≥0∞) < radius p := lt_of_le_of_lt (le_add_right (le_refl _)) h, convert continuous_multilinear_map.has_sum_eval (p.change_origin_has_sum k this) (λ(i : fin k), y), ext ⟨_, _, _⟩, @@ -869,13 +869,13 @@ end formal_multilinear_series section variables [complete_space F] {f : E → F} {p : formal_multilinear_series 𝕜 E F} {x y : E} -{r : ennreal} +{r : ℝ≥0∞} /-- If a function admits a power series expansion `p` on a ball `B (x, r)`, then it also admits a power series on any subball of this ball (even with a different center), given by `p.change_origin`. -/ theorem has_fpower_series_on_ball.change_origin - (hf : has_fpower_series_on_ball f p x r) (h : (nnnorm y : ennreal) < r) : + (hf : has_fpower_series_on_ball f p x r) (h : (nnnorm y : ℝ≥0∞) < r) : has_fpower_series_on_ball f (p.change_origin y) (x + y) (r - nnnorm y) := { r_le := begin apply le_trans _ p.change_origin_radius, @@ -884,7 +884,7 @@ theorem has_fpower_series_on_ball.change_origin r_pos := by simp [h], has_sum := begin assume z hz, - have A : (nnnorm y : ennreal) + nnnorm z < r, + have A : (nnnorm y : ℝ≥0∞) + nnnorm z < r, { have : edist z 0 < r - ↑(nnnorm y) := hz, rwa [edist_eq_coe_nnnorm, ennreal.lt_sub_iff_add_lt, add_comm] at this }, convert p.change_origin_eval (lt_of_lt_of_le A hf.r_le), @@ -899,7 +899,7 @@ lemma has_fpower_series_on_ball.analytic_at_of_mem (hf : has_fpower_series_on_ball f p x r) (h : y ∈ emetric.ball x r) : analytic_at 𝕜 f y := begin - have : (nnnorm (y - x) : ennreal) < r, by simpa [edist_eq_coe_nnnorm_sub] using h, + have : (nnnorm (y - x) : ℝ≥0∞) < r, by simpa [edist_eq_coe_nnnorm_sub] using h, have := hf.change_origin this, rw [add_sub_cancel'_right] at this, exact this.analytic_at diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index 85bd9f2a439d7..2c76ac090845b 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -72,7 +72,7 @@ variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] {H : Type*} [normed_group H] [normed_space 𝕜 H] open filter list -open_locale topological_space big_operators classical nnreal +open_locale topological_space big_operators classical nnreal ennreal /-! ### Composing formal multilinear series -/ @@ -493,7 +493,7 @@ summability over all compositions. -/ theorem le_comp_radius_of_summable (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) (r : ℝ≥0) (hr : summable (λ i : (Σ n, composition n), nnnorm (q.comp_along_composition p i.2) * r ^ i.1)) : - (r : ennreal) ≤ (q.comp p).radius := + (r : ℝ≥0∞) ≤ (q.comp p).radius := begin refine le_radius_of_bound_nnreal _ (∑' i : (Σ n, composition n), nnnorm (comp_along_composition q p i.snd) * r ^ i.fst) (λ n, _), @@ -708,7 +708,7 @@ begin `f (x + y)` is close enough to `f x` to be in the disk where `g` is well behaved. Let `min (r, rf, δ)` be this new radius.-/ have : continuous_at f x := Hf.analytic_at.continuous_at, - obtain ⟨δ, δpos, hδ⟩ : ∃ (δ : ennreal) (H : 0 < δ), + obtain ⟨δ, δpos, hδ⟩ : ∃ (δ : ℝ≥0∞) (H : 0 < δ), ∀ {z : E}, z ∈ emetric.ball x δ → f z ∈ emetric.ball (f x) rg, { have : emetric.ball (f x) rg ∈ 𝓝 (f x) := emetric.ball_mem_nhds _ Hg.r_pos, rcases emetric.mem_nhds_iff.1 (Hf.analytic_at.continuous_at this) with ⟨δ, δpos, Hδ⟩, diff --git a/src/analysis/analytic/radius_liminf.lean b/src/analysis/analytic/radius_liminf.lean index 63ed1376f1436..24217e24ae3dc 100644 --- a/src/analysis/analytic/radius_liminf.lean +++ b/src/analysis/analytic/radius_liminf.lean @@ -17,7 +17,7 @@ variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {F : Type*} [normed_group F] [normed_space 𝕜 F] -open_locale topological_space classical big_operators nnreal +open_locale topological_space classical big_operators nnreal ennreal open filter asymptotics namespace formal_multilinear_series @@ -30,7 +30,7 @@ coercions. -/ lemma radius_eq_liminf : p.radius = liminf at_top (λ n, 1/((nnnorm (p n)) ^ (1 / (n : ℝ)) : ℝ≥0)) := begin have : ∀ (r : ℝ≥0) {n : ℕ}, 0 < n → - ((r : ennreal) ≤ 1 / ↑(nnnorm (p n) ^ (1 / (n : ℝ))) ↔ nnnorm (p n) * r ^ n ≤ 1), + ((r : ℝ≥0∞) ≤ 1 / ↑(nnnorm (p n) ^ (1 / (n : ℝ))) ↔ nnnorm (p n) * r ^ n ≤ 1), { intros r n hn, have : 0 < (n : ℝ) := nat.cast_pos.2 hn, conv_lhs {rw [one_div, ennreal.le_inv_iff_mul_le, ← ennreal.coe_mul, diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index b06c7d6499e30..57ed44fd845f3 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -82,7 +82,7 @@ See the explanations there. universes u v w noncomputable theory -open_locale classical topological_space big_operators filter +open_locale classical topological_space big_operators filter ennreal open filter asymptotics set open continuous_linear_map (smul_right smul_right_one_eq_iff) @@ -549,7 +549,7 @@ end linear_map section analytic -variables {p : formal_multilinear_series 𝕜 𝕜 F} {r : ennreal} +variables {p : formal_multilinear_series 𝕜 𝕜 F} {r : ℝ≥0∞} protected lemma has_fpower_series_at.has_strict_deriv_at (h : has_fpower_series_at f p x) : has_strict_deriv_at f (p 1 (λ _, 1)) x := diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv.lean index 8e5648b279971..d7c5a7a84cda1 100644 --- a/src/analysis/calculus/fderiv.lean +++ b/src/analysis/calculus/fderiv.lean @@ -114,7 +114,7 @@ derivative, differentiable, Fréchet, calculus -/ open filter asymptotics continuous_linear_map set metric -open_locale topological_space classical nnreal asymptotics filter +open_locale topological_space classical nnreal asymptotics filter ennreal noncomputable theory @@ -913,7 +913,7 @@ end continuous_linear_map section analytic -variables {p : formal_multilinear_series 𝕜 E F} {r : ennreal} +variables {p : formal_multilinear_series 𝕜 E F} {r : ℝ≥0∞} lemma has_fpower_series_at.has_strict_fderiv_at (h : has_fpower_series_at f p x) : has_strict_fderiv_at f (continuous_multilinear_curry_fin1 𝕜 E F (p 1)) x := diff --git a/src/analysis/mean_inequalities.lean b/src/analysis/mean_inequalities.lean index 5f0ca9bcfbb75..28a7614392056 100644 --- a/src/analysis/mean_inequalities.lean +++ b/src/analysis/mean_inequalities.lean @@ -74,17 +74,17 @@ $$ \sum_{i\in s} a_ib_i ≤ \sqrt[p]{\sum_{i\in s} a_i^p}\sqrt[q]{\sum_{i\in s} b_i^q}. $$ -We give versions of this result in `real`, `nnreal` and `ennreal`. +We give versions of this result in `ℝ`, `ℝ≥0` and `ℝ≥0∞`. There are at least two short proofs of this inequality. In one proof we prenormalize both vectors, then apply Young's inequality to each $a_ib_i$. We use a different proof deducing this inequality from the generalized mean inequality for well-chosen vectors and weights. -Hölder's inequality for the Lebesgue integral of ennreal and nnreal functions: we prove +Hölder's inequality for the Lebesgue integral of `ℝ≥0∞` and `ℝ≥0` functions: we prove `∫ (f * g) ∂μ ≤ (∫ f^p ∂μ) ^ (1/p) * (∫ g^q ∂μ) ^ (1/q)` for `p`, `q` conjugate real exponents and `α→(e)nnreal` functions in two cases, -* `ennreal.lintegral_mul_le_Lp_mul_Lq` : ennreal functions, -* `nnreal.lintegral_mul_le_Lp_mul_Lq` : nnreal functions. +* `ennreal.lintegral_mul_le_Lp_mul_Lq` : ℝ≥0∞ functions, +* `nnreal.lintegral_mul_le_Lp_mul_Lq` : ℝ≥0 functions. ### Minkowski's inequality @@ -94,14 +94,14 @@ $$ $$ satisfies the triangle inequality $\|a+b\|_p\le \|a\|_p+\|b\|_p$. -We give versions of this result in `real`, `nnreal` and `ennreal`. +We give versions of this result in `real`, `ℝ≥0` and `ℝ≥0∞`. We deduce this inequality from Hölder's inequality. Namely, Hölder inequality implies that $\|a\|_p$ is the maximum of the inner product $\sum_{i\in s}a_ib_i$ over `b` such that $\|b\|_q\le 1$. Now Minkowski's inequality follows from the fact that the maximum value of the sum of two functions is less than or equal to the sum of the maximum values of the summands. -Minkowski's inequality for the Lebesgue integral of measurable functions with `ennreal` values: +Minkowski's inequality for the Lebesgue integral of measurable functions with `ℝ≥0∞` values: we prove `(∫ (f + g)^p ∂μ) ^ (1/p) ≤ (∫ f^p ∂μ) ^ (1/p) + (∫ g^p ∂μ) ^ (1/p)` for `1 ≤ p`. ## TODO @@ -117,7 +117,7 @@ we prove `(∫ (f + g)^p ∂μ) ^ (1/p) ≤ (∫ f^p ∂μ) ^ (1/p) + (∫ g^p universes u v open finset -open_locale classical nnreal big_operators +open_locale classical big_operators nnreal ennreal noncomputable theory variables {ι : Type u} (s : finset ι) @@ -257,9 +257,9 @@ end nnreal namespace ennreal -/-- Weighted generalized mean inequality, version for sums over finite sets, with `ennreal`-valued +/-- Weighted generalized mean inequality, version for sums over finite sets, with `ℝ≥0∞`-valued functions and real exponents. -/ -theorem rpow_arith_mean_le_arith_mean_rpow (w z : ι → ennreal) (hw' : ∑ i in s, w i = 1) {p : ℝ} +theorem rpow_arith_mean_le_arith_mean_rpow (w z : ι → ℝ≥0∞) (hw' : ∑ i in s, w i = 1) {p : ℝ} (hp : 1 ≤ p) : (∑ i in s, w i * z i) ^ p ≤ ∑ i in s, (w i * z i ^ p) := begin @@ -310,9 +310,9 @@ begin rwa [←coe_eq_coe, ←h_sum_nnreal], }, end -/-- Weighted generalized mean inequality, version for two elements of `ennreal` and real +/-- Weighted generalized mean inequality, version for two elements of `ℝ≥0∞` and real exponents. -/ -theorem rpow_arith_mean_le_arith_mean2_rpow (w₁ w₂ z₁ z₂ : ennreal) (hw' : w₁ + w₂ = 1) {p : ℝ} +theorem rpow_arith_mean_le_arith_mean2_rpow (w₁ w₂ z₁ z₂ : ℝ≥0∞) (hw' : w₁ + w₂ = 1) {p : ℝ} (hp : 1 ≤ p) : (w₁ * z₁ + w₂ * z₂) ^ p ≤ w₁ * z₁ ^ p + w₂ * z₂ ^ p := begin @@ -532,8 +532,8 @@ end real namespace ennreal -/-- Young's inequality, `ennreal` version with real conjugate exponents. -/ -theorem young_inequality (a b : ennreal) {p q : ℝ} (hpq : p.is_conjugate_exponent q) : +/-- Young's inequality, `ℝ≥0∞` version with real conjugate exponents. -/ +theorem young_inequality (a b : ℝ≥0∞) {p q : ℝ} (hpq : p.is_conjugate_exponent q) : a * b ≤ a ^ p / ennreal.of_real p + b ^ q / ennreal.of_real q := begin by_cases h : a = ⊤ ∨ b = ⊤, @@ -548,11 +548,11 @@ begin exact nnreal.young_inequality_real a.to_nnreal b.to_nnreal hpq, end -variables (f g : ι → ennreal) {p q : ℝ} +variables (f g : ι → ℝ≥0∞) {p q : ℝ} /-- Hölder inequality: the scalar product of two functions is bounded by the product of their `L^p` and `L^q` norms when `p` and `q` are conjugate exponents. Version for sums over finite sets, -with `ennreal`-valued functions. -/ +with `ℝ≥0∞`-valued functions. -/ theorem inner_le_Lp_mul_Lq (hpq : p.is_conjugate_exponent q) : (∑ i in s, f i * g i) ≤ (∑ i in s, (f i)^p) ^ (1/p) * (∑ i in s, (g i)^q) ^ (1/q) := begin @@ -581,7 +581,7 @@ begin end /-- Minkowski inequality: the `L_p` seminorm of the sum of two vectors is less than or equal -to the sum of the `L_p`-seminorms of the summands. A version for `ennreal` valued nonnegative +to the sum of the `L_p`-seminorms of the summands. A version for `ℝ≥0∞` valued nonnegative functions. -/ theorem Lp_add_le (hp : 1 ≤ p) : (∑ i in s, (f i + g i) ^ p)^(1/p) ≤ (∑ i in s, (f i)^p) ^ (1/p) + (∑ i in s, (g i)^p) ^ (1/p) := @@ -600,17 +600,17 @@ begin { apply finset.sum_congr rfl (λ i hi, _), simp [H'.1 i hi, H'.2 i hi] } end -private lemma add_rpow_le_one_of_add_le_one {p : ℝ} (a b : ennreal) (hab : a + b ≤ 1) +private lemma add_rpow_le_one_of_add_le_one {p : ℝ} (a b : ℝ≥0∞) (hab : a + b ≤ 1) (hp1 : 1 ≤ p) : a ^ p + b ^ p ≤ 1 := begin - have h_le_one : ∀ x : ennreal, x ≤ 1 → x ^ p ≤ x, from λ x hx, rpow_le_self_of_le_one hx hp1, + have h_le_one : ∀ x : ℝ≥0∞, x ≤ 1 → x ^ p ≤ x, from λ x hx, rpow_le_self_of_le_one hx hp1, have ha : a ≤ 1, from (self_le_add_right a b).trans hab, have hb : b ≤ 1, from (self_le_add_left b a).trans hab, exact (add_le_add (h_le_one a ha) (h_le_one b hb)).trans hab, end -lemma add_rpow_le_rpow_add {p : ℝ} (a b : ennreal) (hp1 : 1 ≤ p) : +lemma add_rpow_le_rpow_add {p : ℝ} (a b : ℝ≥0∞) (hp1 : 1 ≤ p) : a ^ p + b ^ p ≤ (a + b) ^ p := begin have hp_pos : 0 < p := lt_of_lt_of_le zero_lt_one hp1, @@ -634,7 +634,7 @@ begin ←mul_assoc, mul_inv_cancel hab_0 hab_top, one_mul, one_mul] at h_mul, end -lemma rpow_add_rpow_le_add {p : ℝ} (a b : ennreal) (hp1 : 1 ≤ p) : +lemma rpow_add_rpow_le_add {p : ℝ} (a b : ℝ≥0∞) (hp1 : 1 ≤ p) : (a ^ p + b ^ p) ^ (1/p) ≤ a + b := begin rw ←@ennreal.le_rpow_one_div_iff _ _ (1/p) (by simp [lt_of_lt_of_le zero_lt_one hp1]), @@ -642,10 +642,10 @@ begin exact add_rpow_le_rpow_add _ _ hp1, end -theorem rpow_add_rpow_le {p q : ℝ} (a b : ennreal) (hp_pos : 0 < p) (hpq : p ≤ q) : +theorem rpow_add_rpow_le {p q : ℝ} (a b : ℝ≥0∞) (hp_pos : 0 < p) (hpq : p ≤ q) : (a ^ q + b ^ q) ^ (1/q) ≤ (a ^ p + b ^ p) ^ (1/p) := begin - have h_rpow : ∀ a : ennreal, a^q = (a^p)^(q/p), + have h_rpow : ∀ a : ℝ≥0∞, a^q = (a^p)^(q/p), from λ a, by rw [←ennreal.rpow_mul, div_eq_inv_mul, ←mul_assoc, _root_.mul_inv_cancel hp_pos.ne.symm, one_mul], have h_rpow_add_rpow_le_add : ((a^p)^(q/p) + (b^p)^(q/p)) ^ (1/(q/p)) ≤ a^p + b^p, @@ -656,7 +656,7 @@ begin rwa one_div_div at h_rpow_add_rpow_le_add, end -lemma rpow_add_le_add_rpow {p : ℝ} (a b : ennreal) (hp_pos : 0 < p) (hp1 : p ≤ 1) : +lemma rpow_add_le_add_rpow {p : ℝ} (a b : ℝ≥0∞) (hp_pos : 0 < p) (hp1 : p ≤ 1) : (a + b) ^ p ≤ a ^ p + b ^ p := begin have h := rpow_add_rpow_le a b hp_pos hp1, @@ -669,16 +669,16 @@ end ennreal section lintegral /-! -### Hölder's inequality for the Lebesgue integral of ennreal and nnreal functions +### Hölder's inequality for the Lebesgue integral of ℝ≥0∞ and nnreal functions We prove `∫ (f * g) ∂μ ≤ (∫ f^p ∂μ) ^ (1/p) * (∫ g^q ∂μ) ^ (1/q)` for `p`, `q` conjugate real exponents and `α→(e)nnreal` functions in several cases, the first two being useful only to prove the more general results: -* `ennreal.lintegral_mul_le_one_of_lintegral_rpow_eq_one` : ennreal functions for which the +* `ennreal.lintegral_mul_le_one_of_lintegral_rpow_eq_one` : ℝ≥0∞ functions for which the integrals on the right are equal to 1, -* `ennreal.lintegral_mul_le_Lp_mul_Lq_of_ne_zero_of_ne_top` : ennreal functions for which the +* `ennreal.lintegral_mul_le_Lp_mul_Lq_of_ne_zero_of_ne_top` : ℝ≥0∞ functions for which the integrals on the right are neither ⊤ nor 0, -* `ennreal.lintegral_mul_le_Lp_mul_Lq` : ennreal functions, +* `ennreal.lintegral_mul_le_Lp_mul_Lq` : ℝ≥0∞ functions, * `nnreal.lintegral_mul_le_Lp_mul_Lq` : nnreal functions. -/ @@ -688,7 +688,7 @@ variables {α : Type*} [measurable_space α] {μ : measure α} namespace ennreal lemma lintegral_mul_le_one_of_lintegral_rpow_eq_one {p q : ℝ} (hpq : p.is_conjugate_exponent q) - {f g : α → ennreal} (hf : ae_measurable f μ) (hg : ae_measurable g μ) + {f g : α → ℝ≥0∞} (hf : ae_measurable f μ) (hg : ae_measurable g μ) (hf_norm : ∫⁻ a, (f a)^p ∂μ = 1) (hg_norm : ∫⁻ a, (g a)^q ∂μ = 1) : ∫⁻ a, (f * g) a ∂μ ≤ 1 := begin @@ -708,15 +708,15 @@ begin end /-- Function multiplied by the inverse of its p-seminorm `(∫⁻ f^p ∂μ) ^ 1/p`-/ -def fun_mul_inv_snorm (f : α → ennreal) (p : ℝ) (μ : measure α) : α → ennreal := +def fun_mul_inv_snorm (f : α → ℝ≥0∞) (p : ℝ) (μ : measure α) : α → ℝ≥0∞ := λ a, (f a) * ((∫⁻ c, (f c) ^ p ∂μ) ^ (1 / p))⁻¹ -lemma fun_eq_fun_mul_inv_snorm_mul_snorm {p : ℝ} (f : α → ennreal) +lemma fun_eq_fun_mul_inv_snorm_mul_snorm {p : ℝ} (f : α → ℝ≥0∞) (hf_nonzero : ∫⁻ a, (f a) ^ p ∂μ ≠ 0) (hf_top : ∫⁻ a, (f a) ^ p ∂μ ≠ ⊤) {a : α} : f a = (fun_mul_inv_snorm f p μ a) * (∫⁻ c, (f c)^p ∂μ)^(1/p) := by simp [fun_mul_inv_snorm, mul_assoc, inv_mul_cancel, hf_nonzero, hf_top] -lemma fun_mul_inv_snorm_rpow {p : ℝ} (hp0 : 0 < p) {f : α → ennreal} {a : α} : +lemma fun_mul_inv_snorm_rpow {p : ℝ} (hp0 : 0 < p) {f : α → ℝ≥0∞} {a : α} : (fun_mul_inv_snorm f p μ a) ^ p = (f a)^p * (∫⁻ c, (f c) ^ p ∂μ)⁻¹ := begin rw [fun_mul_inv_snorm, mul_rpow_of_nonneg _ _ (le_of_lt hp0)], @@ -726,7 +726,7 @@ begin _root_.inv_mul_cancel (ne_of_lt hp0).symm, rpow_one], end -lemma lintegral_rpow_fun_mul_inv_snorm_eq_one {p : ℝ} (hp0_lt : 0 < p) {f : α → ennreal} +lemma lintegral_rpow_fun_mul_inv_snorm_eq_one {p : ℝ} (hp0_lt : 0 < p) {f : α → ℝ≥0∞} (hf : ae_measurable f μ) (hf_nonzero : ∫⁻ a, (f a)^p ∂μ ≠ 0) (hf_top : ∫⁻ a, (f a)^p ∂μ ≠ ⊤) : ∫⁻ c, (fun_mul_inv_snorm f p μ c)^p ∂μ = 1 := begin @@ -736,7 +736,7 @@ end /-- Hölder's inequality in case of finite non-zero integrals -/ lemma lintegral_mul_le_Lp_mul_Lq_of_ne_zero_of_ne_top {p q : ℝ} (hpq : p.is_conjugate_exponent q) - {f g : α → ennreal} (hf : ae_measurable f μ) (hg : ae_measurable g μ) + {f g : α → ℝ≥0∞} (hf : ae_measurable f μ) (hg : ae_measurable g μ) (hf_nontop : ∫⁻ a, (f a)^p ∂μ ≠ ⊤) (hg_nontop : ∫⁻ a, (g a)^q ∂μ ≠ ⊤) (hf_nonzero : ∫⁻ a, (f a)^p ∂μ ≠ 0) (hg_nonzero : ∫⁻ a, (g a)^q ∂μ ≠ 0) : ∫⁻ a, (f * g) a ∂μ ≤ (∫⁻ a, (f a)^p ∂μ)^(1/p) * (∫⁻ a, (g a)^q ∂μ)^(1/q) := @@ -764,7 +764,7 @@ begin end end -lemma ae_eq_zero_of_lintegral_rpow_eq_zero {p : ℝ} (hp0_lt : 0 < p) {f : α → ennreal} +lemma ae_eq_zero_of_lintegral_rpow_eq_zero {p : ℝ} (hp0_lt : 0 < p) {f : α → ℝ≥0∞} (hf : ae_measurable f μ) (hf_zero : ∫⁻ a, (f a)^p ∂μ = 0) : f =ᵐ[μ] 0 := begin @@ -780,7 +780,7 @@ begin end lemma lintegral_mul_eq_zero_of_lintegral_rpow_eq_zero {p : ℝ} (hp0_lt : 0 < p) - {f g : α → ennreal} (hf : ae_measurable f μ) (hf_zero : ∫⁻ a, (f a)^p ∂μ = 0) : + {f g : α → ℝ≥0∞} (hf : ae_measurable f μ) (hf_zero : ∫⁻ a, (f a)^p ∂μ = 0) : ∫⁻ a, (f * g) a ∂μ = 0 := begin rw ←@lintegral_zero_fun α _ μ, @@ -791,7 +791,7 @@ begin end lemma lintegral_mul_le_Lp_mul_Lq_of_ne_zero_of_eq_top {p q : ℝ} (hp0_lt : 0 < p) (hq0 : 0 ≤ q) - {f g : α → ennreal} (hf_top : ∫⁻ a, (f a)^p ∂μ = ⊤) (hg_nonzero : ∫⁻ a, (g a)^q ∂μ ≠ 0) : + {f g : α → ℝ≥0∞} (hf_top : ∫⁻ a, (f a)^p ∂μ = ⊤) (hg_nonzero : ∫⁻ a, (g a)^q ∂μ ≠ 0) : ∫⁻ a, (f * g) a ∂μ ≤ (∫⁻ a, (f a)^p ∂μ) ^ (1/p) * (∫⁻ a, (g a)^q ∂μ) ^ (1/q) := begin refine le_trans le_top (le_of_eq _), @@ -800,11 +800,11 @@ begin simp [hq0, hg_nonzero], end -/-- Hölder's inequality for functions `α → ennreal`. The integral of the product of two functions +/-- Hölder's inequality for functions `α → ℝ≥0∞`. The integral of the product of two functions is bounded by the product of their `ℒp` and `ℒq` seminorms when `p` and `q` are conjugate exponents. -/ theorem lintegral_mul_le_Lp_mul_Lq (μ : measure α) {p q : ℝ} (hpq : p.is_conjugate_exponent q) - {f g : α → ennreal} (hf : ae_measurable f μ) (hg : ae_measurable g μ) : + {f g : α → ℝ≥0∞} (hf : ae_measurable f μ) (hg : ae_measurable g μ) : ∫⁻ a, (f * g) a ∂μ ≤ (∫⁻ a, (f a)^p ∂μ) ^ (1/p) * (∫⁻ a, (g a)^q ∂μ) ^ (1/q) := begin by_cases hf_zero : ∫⁻ a, (f a) ^ p ∂μ = 0, @@ -825,28 +825,28 @@ begin end lemma lintegral_rpow_add_lt_top_of_lintegral_rpow_lt_top {p : ℝ} - {f g : α → ennreal} (hf : ae_measurable f μ) (hf_top : ∫⁻ a, (f a) ^ p ∂μ < ⊤) + {f g : α → ℝ≥0∞} (hf : ae_measurable f μ) (hf_top : ∫⁻ a, (f a) ^ p ∂μ < ⊤) (hg : ae_measurable g μ) (hg_top : ∫⁻ a, (g a) ^ p ∂μ < ⊤) (hp1 : 1 ≤ p) : ∫⁻ a, ((f + g) a) ^ p ∂μ < ⊤ := begin have hp0_lt : 0 < p, from lt_of_lt_of_le zero_lt_one hp1, have hp0 : 0 ≤ p, from le_of_lt hp0_lt, calc ∫⁻ (a : α), (f a + g a) ^ p ∂μ - ≤ ∫⁻ a, ((2:ennreal)^(p-1) * (f a) ^ p + (2:ennreal)^(p-1) * (g a) ^ p) ∂ μ : + ≤ ∫⁻ a, ((2:ℝ≥0∞)^(p-1) * (f a) ^ p + (2:ℝ≥0∞)^(p-1) * (g a) ^ p) ∂ μ : begin refine lintegral_mono (λ a, _), dsimp only, - have h_zero_lt_half_rpow : (0 : ennreal) < (1 / 2) ^ p, + have h_zero_lt_half_rpow : (0 : ℝ≥0∞) < (1 / 2) ^ p, { rw [←ennreal.zero_rpow_of_pos hp0_lt], exact ennreal.rpow_lt_rpow (by simp [zero_lt_one]) hp0_lt, }, - have h_rw : (1 / 2) ^ p * (2:ennreal) ^ (p - 1) = 1 / 2, + have h_rw : (1 / 2) ^ p * (2:ℝ≥0∞) ^ (p - 1) = 1 / 2, { rw [sub_eq_add_neg, ennreal.rpow_add _ _ ennreal.two_ne_zero ennreal.coe_ne_top, ←mul_assoc, ←ennreal.mul_rpow_of_nonneg _ _ hp0, one_div, ennreal.inv_mul_cancel ennreal.two_ne_zero ennreal.coe_ne_top, ennreal.one_rpow, one_mul, ennreal.rpow_neg_one], }, rw ←ennreal.mul_le_mul_left (ne_of_lt h_zero_lt_half_rpow).symm _, { rw [mul_add, ← mul_assoc, ← mul_assoc, h_rw, ←ennreal.mul_rpow_of_nonneg _ _ hp0, mul_add], - refine ennreal.rpow_arith_mean_le_arith_mean2_rpow (1/2 : ennreal) (1/2 : ennreal) + refine ennreal.rpow_arith_mean_le_arith_mean2_rpow (1/2 : ℝ≥0∞) (1/2 : ℝ≥0∞) (f a) (g a) _ hp1, rw [ennreal.div_add_div_same, one_add_one_eq_two, ennreal.div_self ennreal.two_ne_zero ennreal.coe_ne_top], }, @@ -859,7 +859,7 @@ begin begin rw [lintegral_add', lintegral_const_mul'' _ hf.ennreal_rpow_const, lintegral_const_mul'' _ hg.ennreal_rpow_const, ennreal.add_lt_top], - { have h_two : (2 : ennreal) ^ (p - 1) < ⊤, + { have h_two : (2 : ℝ≥0∞) ^ (p - 1) < ⊤, from ennreal.rpow_lt_top_of_nonneg (by simp [hp1]) ennreal.coe_ne_top, repeat {rw ennreal.mul_lt_top_iff}, simp [hf_top, hg_top, h_two], }, @@ -871,7 +871,7 @@ begin end lemma lintegral_Lp_mul_le_Lq_mul_Lr {α} [measurable_space α] {p q r : ℝ} (hp0_lt : 0 < p) - (hpq : p < q) (hpqr : 1/p = 1/q + 1/r) (μ : measure α) {f g : α → ennreal} + (hpq : p < q) (hpqr : 1/p = 1/q + 1/r) (μ : measure α) {f g : α → ℝ≥0∞} (hf : ae_measurable f μ) (hg : ae_measurable g μ) : (∫⁻ a, ((f * g) a)^p ∂μ) ^ (1/p) ≤ (∫⁻ a, (f a)^q ∂μ) ^ (1/q) * (∫⁻ a, (g a)^r ∂μ) ^ (1/r) := begin @@ -912,7 +912,7 @@ begin end lemma lintegral_mul_rpow_le_lintegral_rpow_mul_lintegral_rpow {p q : ℝ} - (hpq : p.is_conjugate_exponent q) {f g : α → ennreal} + (hpq : p.is_conjugate_exponent q) {f g : α → ℝ≥0∞} (hf : ae_measurable f μ) (hg : ae_measurable g μ) (hf_top : ∫⁻ a, (f a) ^ p ∂μ ≠ ⊤) : ∫⁻ a, (f a) * (g a) ^ (p - 1) ∂μ ≤ (∫⁻ a, (f a)^p ∂μ) ^ (1/p) * (∫⁻ a, (g a)^p ∂μ) ^ (1/q) := begin @@ -933,7 +933,7 @@ begin end lemma lintegral_rpow_add_le_add_snorm_mul_lintegral_rpow_add {p q : ℝ} - (hpq : p.is_conjugate_exponent q) {f g : α → ennreal} (hf : ae_measurable f μ) + (hpq : p.is_conjugate_exponent q) {f g : α → ℝ≥0∞} (hf : ae_measurable f μ) (hf_top : ∫⁻ a, (f a) ^ p ∂μ ≠ ⊤) (hg : ae_measurable g μ) (hg_top : ∫⁻ a, (g a) ^ p ∂μ ≠ ⊤) : ∫⁻ a, ((f + g) a)^p ∂ μ ≤ ((∫⁻ a, (f a)^p ∂μ) ^ (1/p) + (∫⁻ a, (g a)^p ∂μ) ^ (1/p)) @@ -974,7 +974,7 @@ begin end private lemma lintegral_Lp_add_le_aux {p q : ℝ} - (hpq : p.is_conjugate_exponent q) {f g : α → ennreal} (hf : ae_measurable f μ) + (hpq : p.is_conjugate_exponent q) {f g : α → ℝ≥0∞} (hf : ae_measurable f μ) (hf_top : ∫⁻ a, (f a) ^ p ∂μ ≠ ⊤) (hg : ae_measurable g μ) (hg_top : ∫⁻ a, (g a) ^ p ∂μ ≠ ⊤) (h_add_zero : ∫⁻ a, ((f+g) a) ^ p ∂ μ ≠ 0) (h_add_top : ∫⁻ a, ((f+g) a) ^ p ∂ μ ≠ ⊤) : (∫⁻ a, ((f + g) a)^p ∂ μ) ^ (1/p) ≤ (∫⁻ a, (f a)^p ∂μ) ^ (1/p) + (∫⁻ a, (g a)^p ∂μ) ^ (1/p) := @@ -1002,9 +1002,9 @@ begin rwa [←mul_assoc, ennreal.mul_le_mul_right h_add_zero h_add_top, mul_comm] at h, end -/-- Minkowski's inequality for functions `α → ennreal`: the `ℒp` seminorm of the sum of two +/-- Minkowski's inequality for functions `α → ℝ≥0∞`: the `ℒp` seminorm of the sum of two functions is bounded by the sum of their `ℒp` seminorms. -/ -theorem lintegral_Lp_add_le {p : ℝ} {f g : α → ennreal} +theorem lintegral_Lp_add_le {p : ℝ} {f g : α → ℝ≥0∞} (hf : ae_measurable f μ) (hg : ae_measurable g μ) (hp1 : 1 ≤ p) : (∫⁻ a, ((f + g) a)^p ∂ μ) ^ (1/p) ≤ (∫⁻ a, (f a)^p ∂μ) ^ (1/p) + (∫⁻ a, (g a)^p ∂μ) ^ (1/p) := begin diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index ce55bd92cbb41..ab566806a4514 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -15,7 +15,7 @@ variables {α : Type*} {β : Type*} {γ : Type*} {ι : Type*} noncomputable theory open filter metric -open_locale topological_space big_operators nnreal +open_locale topological_space big_operators nnreal ennreal /-- Auxiliary class, endowing a type `α` with a function `norm : α → ℝ`. This class is designed to be extended in more interesting classes specifying the properties of the norm. -/ @@ -313,16 +313,16 @@ nnreal.eq $ norm_neg g lemma nndist_nnnorm_nnnorm_le (g h : α) : nndist (nnnorm g) (nnnorm h) ≤ nnnorm (g - h) := nnreal.coe_le_coe.2 $ dist_norm_norm_le g h -lemma of_real_norm_eq_coe_nnnorm (x : β) : ennreal.of_real ∥x∥ = (nnnorm x : ennreal) := +lemma of_real_norm_eq_coe_nnnorm (x : β) : ennreal.of_real ∥x∥ = (nnnorm x : ℝ≥0∞) := ennreal.of_real_eq_coe_nnreal _ -lemma edist_eq_coe_nnnorm_sub (x y : β) : edist x y = (nnnorm (x - y) : ennreal) := +lemma edist_eq_coe_nnnorm_sub (x y : β) : edist x y = (nnnorm (x - y) : ℝ≥0∞) := by rw [edist_dist, dist_eq_norm, of_real_norm_eq_coe_nnnorm] -lemma edist_eq_coe_nnnorm (x : β) : edist x 0 = (nnnorm x : ennreal) := +lemma edist_eq_coe_nnnorm (x : β) : edist x 0 = (nnnorm x : ℝ≥0∞) := by rw [edist_eq_coe_nnnorm_sub, _root_.sub_zero] -lemma mem_emetric_ball_0_iff {x : β} {r : ennreal} : x ∈ emetric.ball (0 : β) r ↔ ↑(nnnorm x) < r := +lemma mem_emetric_ball_0_iff {x : β} {r : ℝ≥0∞} : x ∈ emetric.ball (0 : β) r ↔ ↑(nnnorm x) < r := by rw [emetric.mem_ball, edist_eq_coe_nnnorm] lemma nndist_add_add_le (g₁ g₂ h₁ h₂ : α) : @@ -871,7 +871,7 @@ by { ext, exact norm_of_nonneg (zero_le x) } lemma nnnorm_of_nonneg {x : ℝ} (hx : 0 ≤ x) : nnnorm x = ⟨x, hx⟩ := @nnnorm_coe_eq_self ⟨x, hx⟩ -lemma ennnorm_eq_of_real {x : ℝ} (hx : 0 ≤ x) : (nnnorm x : ennreal) = ennreal.of_real x := +lemma ennnorm_eq_of_real {x : ℝ} (hx : 0 ≤ x) : (nnnorm x : ℝ≥0∞) = ennreal.of_real x := by { rw [← of_real_norm_eq_coe_nnnorm, norm_of_nonneg hx] } end real diff --git a/src/analysis/normed_space/enorm.lean b/src/analysis/normed_space/enorm.lean index 76ffebdb3f8a1..32db241a8cd3f 100644 --- a/src/analysis/normed_space/enorm.lean +++ b/src/analysis/normed_space/enorm.lean @@ -31,11 +31,12 @@ normed space, extended norm -/ local attribute [instance, priority 1001] classical.prop_decidable +open_locale ennreal /-- Extended norm on a vector space. As in the case of normed spaces, we require only `∥c • x∥ ≤ ∥c∥ * ∥x∥` in the definition, then prove an equality in `map_smul`. -/ structure enorm (𝕜 : Type*) (V : Type*) [normed_field 𝕜] [add_comm_group V] [vector_space 𝕜 V] := -(to_fun : V → ennreal) +(to_fun : V → ℝ≥0∞) (eq_zero' : ∀ x, to_fun x = 0 → x = 0) (map_add_le' : ∀ x y : V, to_fun (x + y) ≤ to_fun x + to_fun y) (map_smul_le' : ∀ (c : 𝕜) (x : V), to_fun (c • x) ≤ nnnorm c * to_fun x) @@ -63,7 +64,7 @@ injective_coe_fn.eq_iff le_antisymm (e.map_smul_le' c x) $ begin by_cases hc : c = 0, { simp [hc] }, - calc (nnnorm c : ennreal) * e x = nnnorm c * e (c⁻¹ • c • x) : by rw [inv_smul_smul' hc] + calc (nnnorm c : ℝ≥0∞) * e x = nnnorm c * e (c⁻¹ • c • x) : by rw [inv_smul_smul' hc] ... ≤ nnnorm c * (nnnorm (c⁻¹) * e (c • x)) : _ ... = e (c • x) : _, { exact ennreal.mul_le_mul (le_refl _) (e.map_smul_le' _ _) }, diff --git a/src/analysis/p_series.lean b/src/analysis/p_series.lean index 136329c868303..eb5ff99c071cd 100644 --- a/src/analysis/p_series.lean +++ b/src/analysis/p_series.lean @@ -89,7 +89,7 @@ end finset namespace ennreal -variable {f : ℕ → ennreal} +variable {f : ℕ → ℝ≥0∞} lemma le_tsum_condensed (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f m) : ∑' k, f k ≤ f 0 + ∑' k : ℕ, (2 ^ k) * f (2 ^ k) := @@ -120,10 +120,10 @@ begin simp only [← ennreal.tsum_coe_ne_top_iff_summable, ne.def, not_iff_not, ennreal.coe_mul, ennreal.coe_pow, ennreal.coe_two], split; intro h, - { replace hf : ∀ m n, 1 < m → m ≤ n → (f n : ennreal) ≤ f m := + { replace hf : ∀ m n, 1 < m → m ≤ n → (f n : ℝ≥0∞) ≤ f m := λ m n hm hmn, ennreal.coe_le_coe.2 (hf (zero_lt_one.trans hm) hmn), simpa [h, ennreal.add_eq_top] using (ennreal.tsum_condensed_le hf) }, - { replace hf : ∀ m n, 0 < m → m ≤ n → (f n : ennreal) ≤ f m := + { replace hf : ∀ m n, 0 < m → m ≤ n → (f n : ℝ≥0∞) ≤ f m := λ m n hm hmn, ennreal.coe_le_coe.2 (hf hm hmn), simpa [h, ennreal.add_eq_top] using (ennreal.le_tsum_condensed hf) } end diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 20b0cca34c9fa..ee58e1834c9cd 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -8,20 +8,20 @@ import analysis.special_functions.trigonometric import analysis.calculus.extend_deriv /-! -# Power function on `ℂ`, `ℝ`, `ℝ≥0`, and `ennreal` +# Power function on `ℂ`, `ℝ`, `ℝ≥0`, and `ℝ≥0∞` We construct the power functions `x ^ y` where * `x` and `y` are complex numbers, * or `x` and `y` are real numbers, * or `x` is a nonnegative real number and `y` is a real number; -* or `x` is a number from `[0, +∞]` (a.k.a. `ennreal`) and `y` is a real number. +* or `x` is a number from `[0, +∞]` (a.k.a. `ℝ≥0∞`) and `y` is a real number. We also prove basic properties of these functions. -/ noncomputable theory -open_locale classical real topological_space nnreal +open_locale classical real topological_space nnreal ennreal namespace complex @@ -1067,45 +1067,45 @@ end nnreal namespace ennreal -/-- The real power function `x^y` on extended nonnegative reals, defined for `x : ennreal` and +/-- The real power function `x^y` on extended nonnegative reals, defined for `x : ℝ≥0∞` and `y : ℝ` as the restriction of the real power function if `0 < x < ⊤`, and with the natural values for `0` and `⊤` (i.e., `0 ^ x = 0` for `x > 0`, `1` for `x = 0` and `⊤` for `x < 0`, and `⊤ ^ x = 1 / 0 ^ x`). -/ -noncomputable def rpow : ennreal → ℝ → ennreal +noncomputable def rpow : ℝ≥0∞ → ℝ → ℝ≥0∞ | (some x) y := if x = 0 ∧ y < 0 then ⊤ else (x ^ y : ℝ≥0) | none y := if 0 < y then ⊤ else if y = 0 then 1 else 0 -noncomputable instance : has_pow ennreal ℝ := ⟨rpow⟩ +noncomputable instance : has_pow ℝ≥0∞ ℝ := ⟨rpow⟩ -@[simp] lemma rpow_eq_pow (x : ennreal) (y : ℝ) : rpow x y = x ^ y := rfl +@[simp] lemma rpow_eq_pow (x : ℝ≥0∞) (y : ℝ) : rpow x y = x ^ y := rfl -@[simp] lemma rpow_zero {x : ennreal} : x ^ (0 : ℝ) = 1 := +@[simp] lemma rpow_zero {x : ℝ≥0∞} : x ^ (0 : ℝ) = 1 := by cases x; { dsimp only [(^), rpow], simp [lt_irrefl] } -lemma top_rpow_def (y : ℝ) : (⊤ : ennreal) ^ y = if 0 < y then ⊤ else if y = 0 then 1 else 0 := +lemma top_rpow_def (y : ℝ) : (⊤ : ℝ≥0∞) ^ y = if 0 < y then ⊤ else if y = 0 then 1 else 0 := rfl -@[simp] lemma top_rpow_of_pos {y : ℝ} (h : 0 < y) : (⊤ : ennreal) ^ y = ⊤ := +@[simp] lemma top_rpow_of_pos {y : ℝ} (h : 0 < y) : (⊤ : ℝ≥0∞) ^ y = ⊤ := by simp [top_rpow_def, h] -@[simp] lemma top_rpow_of_neg {y : ℝ} (h : y < 0) : (⊤ : ennreal) ^ y = 0 := +@[simp] lemma top_rpow_of_neg {y : ℝ} (h : y < 0) : (⊤ : ℝ≥0∞) ^ y = 0 := by simp [top_rpow_def, asymm h, ne_of_lt h] -@[simp] lemma zero_rpow_of_pos {y : ℝ} (h : 0 < y) : (0 : ennreal) ^ y = 0 := +@[simp] lemma zero_rpow_of_pos {y : ℝ} (h : 0 < y) : (0 : ℝ≥0∞) ^ y = 0 := begin rw [← ennreal.coe_zero, ← ennreal.some_eq_coe], dsimp only [(^), rpow], simp [h, asymm h, ne_of_gt h], end -@[simp] lemma zero_rpow_of_neg {y : ℝ} (h : y < 0) : (0 : ennreal) ^ y = ⊤ := +@[simp] lemma zero_rpow_of_neg {y : ℝ} (h : y < 0) : (0 : ℝ≥0∞) ^ y = ⊤ := begin rw [← ennreal.coe_zero, ← ennreal.some_eq_coe], dsimp only [(^), rpow], simp [h, ne_of_gt h], end -lemma zero_rpow_def (y : ℝ) : (0 : ennreal) ^ y = if 0 < y then 0 else if y = 0 then 1 else ⊤ := +lemma zero_rpow_def (y : ℝ) : (0 : ℝ≥0∞) ^ y = if 0 < y then 0 else if y = 0 then 1 else ⊤ := begin rcases lt_trichotomy 0 y with H|rfl|H, { simp [H, ne_of_gt, zero_rpow_of_pos, lt_irrefl] }, @@ -1114,7 +1114,7 @@ begin end @[norm_cast] lemma coe_rpow_of_ne_zero {x : ℝ≥0} (h : x ≠ 0) (y : ℝ) : - (x : ennreal) ^ y = (x ^ y : ℝ≥0) := + (x : ℝ≥0∞) ^ y = (x ^ y : ℝ≥0) := begin rw [← ennreal.some_eq_coe], dsimp only [(^), rpow], @@ -1122,7 +1122,7 @@ begin end @[norm_cast] lemma coe_rpow_of_nonneg (x : ℝ≥0) {y : ℝ} (h : 0 ≤ y) : - (x : ennreal) ^ y = (x ^ y : ℝ≥0) := + (x : ℝ≥0∞) ^ y = (x ^ y : ℝ≥0) := begin by_cases hx : x = 0, { rcases le_iff_eq_or_lt.1 h with H|H, @@ -1132,15 +1132,15 @@ begin end lemma coe_rpow_def (x : ℝ≥0) (y : ℝ) : - (x : ennreal) ^ y = if x = 0 ∧ y < 0 then ⊤ else (x ^ y : ℝ≥0) := rfl + (x : ℝ≥0∞) ^ y = if x = 0 ∧ y < 0 then ⊤ else (x ^ y : ℝ≥0) := rfl -@[simp] lemma rpow_one (x : ennreal) : x ^ (1 : ℝ) = x := +@[simp] lemma rpow_one (x : ℝ≥0∞) : x ^ (1 : ℝ) = x := by cases x; dsimp only [(^), rpow]; simp [zero_lt_one, not_lt_of_le zero_le_one] -@[simp] lemma one_rpow (x : ℝ) : (1 : ennreal) ^ x = 1 := +@[simp] lemma one_rpow (x : ℝ) : (1 : ℝ≥0∞) ^ x = 1 := by { rw [← coe_one, coe_rpow_of_ne_zero one_ne_zero], simp } -@[simp] lemma rpow_eq_zero_iff {x : ennreal} {y : ℝ} : +@[simp] lemma rpow_eq_zero_iff {x : ℝ≥0∞} {y : ℝ} : x ^ y = 0 ↔ (x = 0 ∧ 0 < y) ∨ (x = ⊤ ∧ y < 0) := begin cases x, @@ -1152,7 +1152,7 @@ begin { simp [coe_rpow_of_ne_zero h, h] } } end -@[simp] lemma rpow_eq_top_iff {x : ennreal} {y : ℝ} : +@[simp] lemma rpow_eq_top_iff {x : ℝ≥0∞} {y : ℝ} : x ^ y = ⊤ ↔ (x = 0 ∧ y < 0) ∨ (x = ⊤ ∧ 0 < y) := begin cases x, @@ -1164,10 +1164,10 @@ begin { simp [coe_rpow_of_ne_zero h, h] } } end -lemma rpow_eq_top_iff_of_pos {x : ennreal} {y : ℝ} (hy : 0 < y) : x ^ y = ⊤ ↔ x = ⊤ := +lemma rpow_eq_top_iff_of_pos {x : ℝ≥0∞} {y : ℝ} (hy : 0 < y) : x ^ y = ⊤ ↔ x = ⊤ := by simp [rpow_eq_top_iff, hy, asymm hy] -lemma rpow_eq_top_of_nonneg (x : ennreal) {y : ℝ} (hy0 : 0 ≤ y) : x ^ y = ⊤ → x = ⊤ := +lemma rpow_eq_top_of_nonneg (x : ℝ≥0∞) {y : ℝ} (hy0 : 0 ≤ y) : x ^ y = ⊤ → x = ⊤ := begin rw ennreal.rpow_eq_top_iff, intro h, @@ -1176,20 +1176,20 @@ begin { exact h.left, }, end -lemma rpow_ne_top_of_nonneg {x : ennreal} {y : ℝ} (hy0 : 0 ≤ y) (h : x ≠ ⊤) : x ^ y ≠ ⊤ := +lemma rpow_ne_top_of_nonneg {x : ℝ≥0∞} {y : ℝ} (hy0 : 0 ≤ y) (h : x ≠ ⊤) : x ^ y ≠ ⊤ := mt (ennreal.rpow_eq_top_of_nonneg x hy0) h -lemma rpow_lt_top_of_nonneg {x : ennreal} {y : ℝ} (hy0 : 0 ≤ y) (h : x ≠ ⊤) : x ^ y < ⊤ := +lemma rpow_lt_top_of_nonneg {x : ℝ≥0∞} {y : ℝ} (hy0 : 0 ≤ y) (h : x ≠ ⊤) : x ^ y < ⊤ := ennreal.lt_top_iff_ne_top.mpr (ennreal.rpow_ne_top_of_nonneg hy0 h) -lemma rpow_add {x : ennreal} (y z : ℝ) (hx : x ≠ 0) (h'x : x ≠ ⊤) : x ^ (y + z) = x ^ y * x ^ z := +lemma rpow_add {x : ℝ≥0∞} (y z : ℝ) (hx : x ≠ 0) (h'x : x ≠ ⊤) : x ^ (y + z) = x ^ y * x ^ z := begin cases x, { exact (h'x rfl).elim }, have : x ≠ 0 := λ h, by simpa [h] using hx, simp [coe_rpow_of_ne_zero this, nnreal.rpow_add this] end -lemma rpow_neg (x : ennreal) (y : ℝ) : x ^ -y = (x ^ y)⁻¹ := +lemma rpow_neg (x : ℝ≥0∞) (y : ℝ) : x ^ -y = (x ^ y)⁻¹ := begin cases x, { rcases lt_trichotomy y 0 with H|H|H; @@ -1201,10 +1201,10 @@ begin simp [coe_rpow_of_ne_zero h, ← coe_inv A, nnreal.rpow_neg] } } end -lemma rpow_neg_one (x : ennreal) : x ^ (-1 : ℝ) = x ⁻¹ := +lemma rpow_neg_one (x : ℝ≥0∞) : x ^ (-1 : ℝ) = x ⁻¹ := by simp [rpow_neg] -lemma rpow_mul (x : ennreal) (y z : ℝ) : x ^ (y * z) = (x ^ y) ^ z := +lemma rpow_mul (x : ℝ≥0∞) (y z : ℝ) : x ^ (y * z) = (x ^ y) ^ z := begin cases x, { rcases lt_trichotomy y 0 with Hy|Hy|Hy; @@ -1220,7 +1220,7 @@ begin simp [coe_rpow_of_ne_zero h, coe_rpow_of_ne_zero this, nnreal.rpow_mul] } } end -@[simp, norm_cast] lemma rpow_nat_cast (x : ennreal) (n : ℕ) : x ^ (n : ℝ) = x ^ n := +@[simp, norm_cast] lemma rpow_nat_cast (x : ℝ≥0∞) (n : ℕ) : x ^ (n : ℝ) = x ^ n := begin cases x, { cases n; @@ -1229,14 +1229,14 @@ begin end @[norm_cast] lemma coe_mul_rpow (x y : ℝ≥0) (z : ℝ) : - ((x : ennreal) * y) ^ z = x^z * y^z := + ((x : ℝ≥0∞) * y) ^ z = x^z * y^z := begin rcases lt_trichotomy z 0 with H|H|H, { by_cases hx : x = 0; by_cases hy : y = 0, { simp [hx, hy, zero_rpow_of_neg, H] }, - { have : (y : ennreal) ^ z ≠ 0, by simp [rpow_eq_zero_iff, hy], + { have : (y : ℝ≥0∞) ^ z ≠ 0, by simp [rpow_eq_zero_iff, hy], simp [hx, hy, zero_rpow_of_neg, H, with_top.top_mul this] }, - { have : (x : ennreal) ^ z ≠ 0, by simp [rpow_eq_zero_iff, hx], + { have : (x : ℝ≥0∞) ^ z ≠ 0, by simp [rpow_eq_zero_iff, hx], simp [hx, hy, zero_rpow_of_neg H, with_top.mul_top this] }, { rw [← coe_mul, coe_rpow_of_ne_zero, nnreal.mul_rpow, coe_mul, coe_rpow_of_ne_zero hx, coe_rpow_of_ne_zero hy], @@ -1244,16 +1244,16 @@ begin { simp [H] }, { by_cases hx : x = 0; by_cases hy : y = 0, { simp [hx, hy, zero_rpow_of_pos, H] }, - { have : (y : ennreal) ^ z ≠ 0, by simp [rpow_eq_zero_iff, hy], + { have : (y : ℝ≥0∞) ^ z ≠ 0, by simp [rpow_eq_zero_iff, hy], simp [hx, hy, zero_rpow_of_pos H, with_top.top_mul this] }, - { have : (x : ennreal) ^ z ≠ 0, by simp [rpow_eq_zero_iff, hx], + { have : (x : ℝ≥0∞) ^ z ≠ 0, by simp [rpow_eq_zero_iff, hx], simp [hx, hy, zero_rpow_of_pos H, with_top.mul_top this] }, { rw [← coe_mul, coe_rpow_of_ne_zero, nnreal.mul_rpow, coe_mul, coe_rpow_of_ne_zero hx, coe_rpow_of_ne_zero hy], simp [hx, hy] } }, end -lemma mul_rpow_of_ne_top {x y : ennreal} (hx : x ≠ ⊤) (hy : y ≠ ⊤) (z : ℝ) : +lemma mul_rpow_of_ne_top {x y : ℝ≥0∞} (hx : x ≠ ⊤) (hy : y ≠ ⊤) (z : ℝ) : (x * y) ^ z = x^z * y^z := begin lift x to ℝ≥0 using hx, @@ -1261,7 +1261,7 @@ begin exact coe_mul_rpow x y z end -lemma mul_rpow_of_ne_zero {x y : ennreal} (hx : x ≠ 0) (hy : y ≠ 0) (z : ℝ) : +lemma mul_rpow_of_ne_zero {x y : ℝ≥0∞} (hx : x ≠ 0) (hy : y ≠ 0) (z : ℝ) : (x * y) ^ z = x ^ z * y ^ z := begin rcases lt_trichotomy z 0 with H|H|H, @@ -1292,7 +1292,7 @@ begin simp [hx', hy'] } } end -lemma mul_rpow_of_nonneg (x y : ennreal) {z : ℝ} (hz : 0 ≤ z) : +lemma mul_rpow_of_nonneg (x y : ℝ≥0∞) {z : ℝ} (hz : 0 ≤ z) : (x * y) ^ z = x ^ z * y ^ z := begin rcases le_iff_eq_or_lt.1 hz with H|H, { simp [← H] }, @@ -1302,7 +1302,7 @@ begin exact mul_rpow_of_ne_zero h.1 h.2 z end -lemma inv_rpow_of_pos {x : ennreal} {y : ℝ} (hy : 0 < y) : (x⁻¹) ^ y = (x ^ y)⁻¹ := +lemma inv_rpow_of_pos {x : ℝ≥0∞} {y : ℝ} (hy : 0 < y) : (x⁻¹) ^ y = (x ^ y)⁻¹ := begin by_cases h0 : x = 0, { rw [h0, zero_rpow_of_pos hy, inv_zero, top_rpow_of_pos hy], }, @@ -1318,7 +1318,7 @@ begin { simp [h], }, end -lemma div_rpow_of_nonneg (x y : ennreal) {z : ℝ} (hz : 0 ≤ z) : +lemma div_rpow_of_nonneg (x y : ℝ≥0∞) {z : ℝ} (hz : 0 ≤ z) : (x / y) ^ z = x ^ z / y ^ z := begin by_cases h0 : z = 0, @@ -1328,7 +1328,7 @@ begin rw [div_eq_mul_inv, mul_rpow_of_nonneg x y⁻¹ hz, inv_rpow_of_pos hz_pos, ←div_eq_mul_inv], end -lemma rpow_le_rpow {x y : ennreal} {z : ℝ} (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x^z ≤ y^z := +lemma rpow_le_rpow {x y : ℝ≥0∞} {z : ℝ} (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x^z ≤ y^z := begin rcases le_iff_eq_or_lt.1 h₂ with H|H, { simp [← H, le_refl] }, cases y, { simp [top_rpow_of_pos H] }, @@ -1337,7 +1337,7 @@ begin simp [coe_rpow_of_nonneg _ h₂, nnreal.rpow_le_rpow h₁ h₂] end -lemma rpow_lt_rpow {x y : ennreal} {z : ℝ} (h₁ : x < y) (h₂ : 0 < z) : x^z < y^z := +lemma rpow_lt_rpow {x y : ℝ≥0∞} {z : ℝ} (h₁ : x < y) (h₂ : 0 < z) : x^z < y^z := begin cases x, { exact (not_top_lt h₁).elim }, cases y, { simp [top_rpow_of_pos h₂, coe_rpow_of_nonneg _ (le_of_lt h₂)] }, @@ -1345,7 +1345,7 @@ begin simp [coe_rpow_of_nonneg _ (le_of_lt h₂), nnreal.rpow_lt_rpow h₁ h₂] end -lemma rpow_le_rpow_iff {x y : ennreal} {z : ℝ} (hz : 0 < z) : x ^ z ≤ y ^ z ↔ x ≤ y := +lemma rpow_le_rpow_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z ≤ y ^ z ↔ x ≤ y := begin refine ⟨λ h, _, λ h, rpow_le_rpow h (le_of_lt hz)⟩, rw [←rpow_one x, ←rpow_one y, ←@_root_.mul_inv_cancel _ _ z (ne_of_lt hz).symm, rpow_mul, @@ -1353,7 +1353,7 @@ begin exact rpow_le_rpow h (by simp [le_of_lt hz]), end -lemma rpow_lt_rpow_iff {x y : ennreal} {z : ℝ} (hz : 0 < z) : x ^ z < y ^ z ↔ x < y := +lemma rpow_lt_rpow_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z < y ^ z ↔ x < y := begin refine ⟨λ h_lt, _, λ h, rpow_lt_rpow h hz⟩, rw [←rpow_one x, ←rpow_one y, ←@_root_.mul_inv_cancel _ _ z (ne_of_lt hz).symm, rpow_mul, @@ -1361,21 +1361,21 @@ begin exact rpow_lt_rpow h_lt (by simp [hz]), end -lemma le_rpow_one_div_iff {x y : ennreal} {z : ℝ} (hz : 0 < z) : x ≤ y ^ (1 / z) ↔ x ^ z ≤ y := +lemma le_rpow_one_div_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ≤ y ^ (1 / z) ↔ x ^ z ≤ y := begin nth_rewrite 0 ←rpow_one x, nth_rewrite 0 ←@_root_.mul_inv_cancel _ _ z (ne_of_lt hz).symm, rw [rpow_mul, ←one_div, @rpow_le_rpow_iff _ _ (1/z) (by simp [hz])], end -lemma lt_rpow_one_div_iff {x y : ennreal} {z : ℝ} (hz : 0 < z) : x < y ^ (1 / z) ↔ x ^ z < y := +lemma lt_rpow_one_div_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x < y ^ (1 / z) ↔ x ^ z < y := begin nth_rewrite 0 ←rpow_one x, nth_rewrite 0 ←@_root_.mul_inv_cancel _ _ z (ne_of_lt hz).symm, rw [rpow_mul, ←one_div, @rpow_lt_rpow_iff _ _ (1/z) (by simp [hz])], end -lemma rpow_lt_rpow_of_exponent_lt {x : ennreal} {y z : ℝ} (hx : 1 < x) (hx' : x ≠ ⊤) (hyz : y < z) : +lemma rpow_lt_rpow_of_exponent_lt {x : ℝ≥0∞} {y z : ℝ} (hx : 1 < x) (hx' : x ≠ ⊤) (hyz : y < z) : x^y < x^z := begin lift x to ℝ≥0 using hx', @@ -1384,7 +1384,7 @@ begin nnreal.rpow_lt_rpow_of_exponent_lt hx hyz] end -lemma rpow_le_rpow_of_exponent_le {x : ennreal} {y z : ℝ} (hx : 1 ≤ x) (hyz : y ≤ z) : x^y ≤ x^z := +lemma rpow_le_rpow_of_exponent_le {x : ℝ≥0∞} {y z : ℝ} (hx : 1 ≤ x) (hyz : y ≤ z) : x^y ≤ x^z := begin cases x, { rcases lt_trichotomy y 0 with Hy|Hy|Hy; @@ -1396,7 +1396,7 @@ begin nnreal.rpow_le_rpow_of_exponent_le hx hyz] } end -lemma rpow_lt_rpow_of_exponent_gt {x : ennreal} {y z : ℝ} (hx0 : 0 < x) (hx1 : x < 1) (hyz : z < y) : +lemma rpow_lt_rpow_of_exponent_gt {x : ℝ≥0∞} {y z : ℝ} (hx0 : 0 < x) (hx1 : x < 1) (hyz : z < y) : x^y < x^z := begin lift x to ℝ≥0 using ne_of_lt (lt_of_lt_of_le hx1 le_top), @@ -1404,7 +1404,7 @@ begin simp [coe_rpow_of_ne_zero (ne_of_gt hx0), nnreal.rpow_lt_rpow_of_exponent_gt hx0 hx1 hyz] end -lemma rpow_le_rpow_of_exponent_ge {x : ennreal} {y z : ℝ} (hx1 : x ≤ 1) (hyz : z ≤ y) : +lemma rpow_le_rpow_of_exponent_ge {x : ℝ≥0∞} {y z : ℝ} (hx1 : x ≤ 1) (hyz : z ≤ y) : x^y ≤ x^z := begin lift x to ℝ≥0 using ne_of_lt (lt_of_le_of_lt hx1 coe_lt_top), @@ -1418,19 +1418,19 @@ begin nnreal.rpow_le_rpow_of_exponent_ge (bot_lt_iff_ne_bot.mpr h) hx1 hyz] } end -lemma rpow_le_self_of_le_one {x : ennreal} {z : ℝ} (hx : x ≤ 1) (h_one_le : 1 ≤ z) : x ^ z ≤ x := +lemma rpow_le_self_of_le_one {x : ℝ≥0∞} {z : ℝ} (hx : x ≤ 1) (h_one_le : 1 ≤ z) : x ^ z ≤ x := begin nth_rewrite 1 ←ennreal.rpow_one x, exact ennreal.rpow_le_rpow_of_exponent_ge hx h_one_le, end -lemma le_rpow_self_of_one_le {x : ennreal} {z : ℝ} (hx : 1 ≤ x) (h_one_le : 1 ≤ z) : x ≤ x ^ z := +lemma le_rpow_self_of_one_le {x : ℝ≥0∞} {z : ℝ} (hx : 1 ≤ x) (h_one_le : 1 ≤ z) : x ≤ x ^ z := begin nth_rewrite 0 ←ennreal.rpow_one x, exact ennreal.rpow_le_rpow_of_exponent_le hx h_one_le, end -lemma rpow_pos_of_nonneg {p : ℝ} {x : ennreal} (hx_pos : 0 < x) (hp_nonneg : 0 ≤ p) : 0 < x^p := +lemma rpow_pos_of_nonneg {p : ℝ} {x : ℝ≥0∞} (hx_pos : 0 < x) (hp_nonneg : 0 ≤ p) : 0 < x^p := begin by_cases hp_zero : p = 0, { simp [hp_zero, ennreal.zero_lt_one], }, @@ -1439,7 +1439,7 @@ begin rw ←zero_rpow_of_pos hp_pos, exact rpow_lt_rpow hx_pos hp_pos, }, end -lemma rpow_pos {p : ℝ} {x : ennreal} (hx_pos : 0 < x) (hx_ne_top : x ≠ ⊤) : 0 < x^p := +lemma rpow_pos {p : ℝ} {x : ℝ≥0∞} (hx_pos : 0 < x) (hx_ne_top : x ≠ ⊤) : 0 < x^p := begin cases lt_or_le 0 p with hp_pos hp_nonpos, { exact rpow_pos_of_nonneg hx_pos (le_of_lt hp_pos), }, @@ -1447,21 +1447,21 @@ begin exact rpow_ne_top_of_nonneg (by simp [hp_nonpos]) hx_ne_top, }, end -lemma rpow_lt_one {x : ennreal} {z : ℝ} (hx : x < 1) (hz : 0 < z) : x^z < 1 := +lemma rpow_lt_one {x : ℝ≥0∞} {z : ℝ} (hx : x < 1) (hz : 0 < z) : x^z < 1 := begin lift x to ℝ≥0 using ne_of_lt (lt_of_lt_of_le hx le_top), simp only [coe_lt_one_iff] at hx, simp [coe_rpow_of_nonneg _ (le_of_lt hz), nnreal.rpow_lt_one (zero_le x) hx hz], end -lemma rpow_le_one {x : ennreal} {z : ℝ} (hx : x ≤ 1) (hz : 0 ≤ z) : x^z ≤ 1 := +lemma rpow_le_one {x : ℝ≥0∞} {z : ℝ} (hx : x ≤ 1) (hz : 0 ≤ z) : x^z ≤ 1 := begin lift x to ℝ≥0 using ne_of_lt (lt_of_le_of_lt hx coe_lt_top), simp only [coe_le_one_iff] at hx, simp [coe_rpow_of_nonneg _ hz, nnreal.rpow_le_one hx hz], end -lemma rpow_lt_one_of_one_lt_of_neg {x : ennreal} {z : ℝ} (hx : 1 < x) (hz : z < 0) : x^z < 1 := +lemma rpow_lt_one_of_one_lt_of_neg {x : ℝ≥0∞} {z : ℝ} (hx : 1 < x) (hz : z < 0) : x^z < 1 := begin cases x, { simp [top_rpow_of_neg hz, ennreal.zero_lt_one] }, @@ -1470,7 +1470,7 @@ begin nnreal.rpow_lt_one_of_one_lt_of_neg hx hz] }, end -lemma rpow_le_one_of_one_le_of_neg {x : ennreal} {z : ℝ} (hx : 1 ≤ x) (hz : z < 0) : x^z ≤ 1 := +lemma rpow_le_one_of_one_le_of_neg {x : ℝ≥0∞} {z : ℝ} (hx : 1 ≤ x) (hz : z < 0) : x^z ≤ 1 := begin cases x, { simp [top_rpow_of_neg hz, ennreal.zero_lt_one] }, @@ -1479,7 +1479,7 @@ begin nnreal.rpow_le_one_of_one_le_of_nonpos hx (le_of_lt hz)] }, end -lemma one_lt_rpow {x : ennreal} {z : ℝ} (hx : 1 < x) (hz : 0 < z) : 1 < x^z := +lemma one_lt_rpow {x : ℝ≥0∞} {z : ℝ} (hx : 1 < x) (hz : 0 < z) : 1 < x^z := begin cases x, { simp [top_rpow_of_pos hz] }, @@ -1487,7 +1487,7 @@ begin simp [coe_rpow_of_nonneg _ (le_of_lt hz), nnreal.one_lt_rpow hx hz] } end -lemma one_le_rpow {x : ennreal} {z : ℝ} (hx : 1 ≤ x) (hz : 0 < z) : 1 ≤ x^z := +lemma one_le_rpow {x : ℝ≥0∞} {z : ℝ} (hx : 1 ≤ x) (hz : 0 < z) : 1 ≤ x^z := begin cases x, { simp [top_rpow_of_pos hz] }, @@ -1495,7 +1495,7 @@ begin simp [coe_rpow_of_nonneg _ (le_of_lt hz), nnreal.one_le_rpow hx (le_of_lt hz)] }, end -lemma one_lt_rpow_of_pos_of_lt_one_of_neg {x : ennreal} {z : ℝ} (hx1 : 0 < x) (hx2 : x < 1) +lemma one_lt_rpow_of_pos_of_lt_one_of_neg {x : ℝ≥0∞} {z : ℝ} (hx1 : 0 < x) (hx2 : x < 1) (hz : z < 0) : 1 < x^z := begin lift x to ℝ≥0 using ne_of_lt (lt_of_lt_of_le hx2 le_top), @@ -1503,7 +1503,7 @@ begin simp [coe_rpow_of_ne_zero (ne_of_gt hx1), nnreal.one_lt_rpow_of_pos_of_lt_one_of_neg hx1 hx2 hz], end -lemma one_le_rpow_of_pos_of_le_one_of_neg {x : ennreal} {z : ℝ} (hx1 : 0 < x) (hx2 : x ≤ 1) +lemma one_le_rpow_of_pos_of_le_one_of_neg {x : ℝ≥0∞} {z : ℝ} (hx1 : 0 < x) (hx2 : x ≤ 1) (hz : z < 0) : 1 ≤ x^z := begin lift x to ℝ≥0 using ne_of_lt (lt_of_le_of_lt hx2 coe_lt_top), @@ -1512,7 +1512,7 @@ begin nnreal.one_le_rpow_of_pos_of_le_one_of_nonpos hx1 hx2 (le_of_lt hz)], end -lemma to_nnreal_rpow (x : ennreal) (z : ℝ) : (x.to_nnreal) ^ z = (x ^ z).to_nnreal := +lemma to_nnreal_rpow (x : ℝ≥0∞) (z : ℝ) : (x.to_nnreal) ^ z = (x ^ z).to_nnreal := begin rcases lt_trichotomy z 0 with H|H|H, { cases x, { simp [H, ne_of_lt] }, @@ -1524,11 +1524,11 @@ begin simp [coe_rpow_of_nonneg _ (le_of_lt H)] } end -lemma to_real_rpow (x : ennreal) (z : ℝ) : (x.to_real) ^ z = (x ^ z).to_real := +lemma to_real_rpow (x : ℝ≥0∞) (z : ℝ) : (x.to_real) ^ z = (x ^ z).to_real := by rw [ennreal.to_real, ennreal.to_real, ←nnreal.coe_rpow, ennreal.to_nnreal_rpow] lemma rpow_left_injective {x : ℝ} (hx : x ≠ 0) : - function.injective (λ y : ennreal, y^x) := + function.injective (λ y : ℝ≥0∞, y^x) := begin intros y z hyz, dsimp only at hyz, @@ -1536,24 +1536,24 @@ begin end lemma rpow_left_surjective {x : ℝ} (hx : x ≠ 0) : - function.surjective (λ y : ennreal, y^x) := + function.surjective (λ y : ℝ≥0∞, y^x) := λ y, ⟨y ^ x⁻¹, by simp_rw [←rpow_mul, _root_.inv_mul_cancel hx, rpow_one]⟩ lemma rpow_left_bijective {x : ℝ} (hx : x ≠ 0) : - function.bijective (λ y : ennreal, y^x) := + function.bijective (λ y : ℝ≥0∞, y^x) := ⟨rpow_left_injective hx, rpow_left_surjective hx⟩ -lemma rpow_left_monotone_of_nonneg {x : ℝ} (hx : 0 ≤ x) : monotone (λ y : ennreal, y^x) := +lemma rpow_left_monotone_of_nonneg {x : ℝ} (hx : 0 ≤ x) : monotone (λ y : ℝ≥0∞, y^x) := λ y z hyz, rpow_le_rpow hyz hx -lemma rpow_left_strict_mono_of_pos {x : ℝ} (hx : 0 < x) : strict_mono (λ y : ennreal, y^x) := +lemma rpow_left_strict_mono_of_pos {x : ℝ} (hx : 0 < x) : strict_mono (λ y : ℝ≥0∞, y^x) := λ y z hyz, rpow_lt_rpow hyz hx end ennreal section measurability_ennreal -lemma ennreal.measurable_rpow : measurable (λ p : ennreal × ℝ, p.1 ^ p.2) := +lemma ennreal.measurable_rpow : measurable (λ p : ℝ≥0∞ × ℝ, p.1 ^ p.2) := begin refine ennreal.measurable_of_measurable_nnreal_prod _ _, { simp_rw ennreal.coe_rpow_def, @@ -1565,28 +1565,28 @@ begin exact measurable.ite (measurable_set_singleton 0) measurable_const measurable_const, }, end -lemma measurable.ennreal_rpow {α} [measurable_space α] {f : α → ennreal} (hf : measurable f) +lemma measurable.ennreal_rpow {α} [measurable_space α] {f : α → ℝ≥0∞} (hf : measurable f) {g : α → ℝ} (hg : measurable g) : measurable (λ a : α, (f a) ^ (g a)) := begin - change measurable ((λ p : ennreal × ℝ, p.1 ^ p.2) ∘ (λ a, (f a, g a))), + change measurable ((λ p : ℝ≥0∞ × ℝ, p.1 ^ p.2) ∘ (λ a, (f a, g a))), exact ennreal.measurable_rpow.comp (measurable.prod hf hg), end -lemma ennreal.measurable_rpow_const {y : ℝ} : measurable (λ a : ennreal, a ^ y) := +lemma ennreal.measurable_rpow_const {y : ℝ} : measurable (λ a : ℝ≥0∞, a ^ y) := begin - change measurable ((λ p : ennreal × ℝ, p.1 ^ p.2) ∘ (λ a, (a, y))), + change measurable ((λ p : ℝ≥0∞ × ℝ, p.1 ^ p.2) ∘ (λ a, (a, y))), refine ennreal.measurable_rpow.comp (measurable.prod measurable_id _), dsimp only, exact measurable_const, end -lemma measurable.ennreal_rpow_const {α} [measurable_space α] {f : α → ennreal} (hf : measurable f) +lemma measurable.ennreal_rpow_const {α} [measurable_space α] {f : α → ℝ≥0∞} (hf : measurable f) {y : ℝ} : measurable (λ a : α, (f a) ^ y) := hf.ennreal_rpow measurable_const -lemma ae_measurable.ennreal_rpow_const {α} [measurable_space α] {f : α → ennreal} +lemma ae_measurable.ennreal_rpow_const {α} [measurable_space α] {f : α → ℝ≥0∞} {μ : measure_theory.measure α} (hf : ae_measurable f μ) {y : ℝ} : ae_measurable (λ a : α, (f a) ^ y) μ := ennreal.measurable_rpow_const.comp_ae_measurable hf diff --git a/src/analysis/specific_limits.lean b/src/analysis/specific_limits.lean index 8c791081f70a8..86e1dfafaf7e9 100644 --- a/src/analysis/specific_limits.lean +++ b/src/analysis/specific_limits.lean @@ -17,7 +17,7 @@ import analysis.asymptotics noncomputable theory open classical set function filter finset metric asymptotics -open_locale classical topological_space nat big_operators uniformity nnreal +open_locale classical topological_space nat big_operators uniformity nnreal ennreal variables {α : Type*} {β : Type*} {ι : Type*} @@ -265,7 +265,7 @@ lemma nnreal.tendsto_pow_at_top_nhds_0_of_lt_1 {r : ℝ≥0} (hr : r < 1) : nnreal.tendsto_coe.1 $ by simp only [nnreal.coe_pow, nnreal.coe_zero, tendsto_pow_at_top_nhds_0_of_lt_1 r.coe_nonneg hr] -lemma ennreal.tendsto_pow_at_top_nhds_0_of_lt_1 {r : ennreal} (hr : r < 1) : +lemma ennreal.tendsto_pow_at_top_nhds_0_of_lt_1 {r : ℝ≥0∞} (hr : r < 1) : tendsto (λ n:ℕ, r^n) at_top (𝓝 0) := begin rcases ennreal.lt_iff_exists_coe.1 hr with ⟨r, rfl, hr'⟩, @@ -352,7 +352,7 @@ lemma tsum_geometric_nnreal {r : ℝ≥0} (hr : r < 1) : ∑'n:ℕ, r ^ n = (1 - /-- The series `pow r` converges to `(1-r)⁻¹`. For `r < 1` the RHS is a finite number, and for `1 ≤ r` the RHS equals `∞`. -/ -lemma ennreal.tsum_geometric (r : ennreal) : ∑'n:ℕ, r ^ n = (1 - r)⁻¹ := +lemma ennreal.tsum_geometric (r : ℝ≥0∞) : ∑'n:ℕ, r ^ n = (1 - r)⁻¹ := begin cases lt_or_le r 1 with hr hr, { rcases ennreal.lt_iff_exists_coe.1 hr with ⟨r, rfl, hr'⟩, @@ -363,7 +363,7 @@ begin refine λ a ha, (ennreal.exists_nat_gt (lt_top_iff_ne_top.1 ha)).imp (λ n hn, lt_of_lt_of_le hn _), have : ∀ k:ℕ, 1 ≤ r^k, by simpa using canonically_ordered_semiring.pow_le_pow_of_le_left hr, - calc (n:ennreal) = (∑ i in range n, 1) : by rw [sum_const, nsmul_one, card_range] + calc (n:ℝ≥0∞) = (∑ i in range n, 1) : by rw [sum_const, nsmul_one, card_range] ... ≤ ∑ i in range n, r ^ i : sum_le_sum (λ k _, this k) } end @@ -462,7 +462,7 @@ decaying terms. -/ section edist_le_geometric -variables [emetric_space α] (r C : ennreal) (hr : r < 1) (hC : C ≠ ⊤) {f : ℕ → α} +variables [emetric_space α] (r C : ℝ≥0∞) (hr : r < 1) (hC : C ≠ ⊤) {f : ℕ → α} (hu : ∀n, edist (f n) (f (n+1)) ≤ C * r^n) include hr hC hu @@ -498,7 +498,7 @@ end edist_le_geometric section edist_le_geometric_two -variables [emetric_space α] (C : ennreal) (hC : C ≠ ⊤) {f : ℕ → α} +variables [emetric_space α] (C : ℝ≥0∞) (hC : C ≠ ⊤) {f : ℕ → α} (hu : ∀n, edist (f n) (f (n+1)) ≤ C / 2^n) {a : α} (ha : tendsto f at_top (𝓝 a)) include hC hu @@ -724,8 +724,8 @@ end nnreal namespace ennreal -theorem exists_pos_sum_of_encodable {ε : ennreal} (hε : 0 < ε) (ι) [encodable ι] : - ∃ ε' : ι → ℝ≥0, (∀ i, 0 < ε' i) ∧ ∑' i, (ε' i : ennreal) < ε := +theorem exists_pos_sum_of_encodable {ε : ℝ≥0∞} (hε : 0 < ε) (ι) [encodable ι] : + ∃ ε' : ι → ℝ≥0, (∀ i, 0 < ε' i) ∧ ∑' i, (ε' i : ℝ≥0∞) < ε := begin rcases exists_between hε with ⟨r, h0r, hrε⟩, rcases lt_iff_exists_coe.1 hrε with ⟨x, rfl, hx⟩, diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 47d52df697c3e..b4e1ff1784b76 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -9,15 +9,15 @@ import data.set.intervals /-! # Extended non-negative reals -We define `ennreal := with_no ℝ≥0` to be the type of extended nonnegative real numbers, i.e., the -interval `[0, +∞]`. This type is used as the codomain of a `measure_theory.measure`, and of the -extended distance `edist` in a `emetric_space`. In this file we define some algebraic operations and -a linear order on `ennreal` and prove basic properties of these operations, order, and conversions -to/from `ℝ`, `ℝ≥0`, and `ℕ`. +We define `ennreal = ℝ≥0∞ := with_no ℝ≥0` to be the type of extended nonnegative real numbers, +i.e., the interval `[0, +∞]`. This type is used as the codomain of a `measure_theory.measure`, +and of the extended distance `edist` in a `emetric_space`. +In this file we define some algebraic operations and a linear order on `ℝ≥0∞` +and prove basic properties of these operations, order, and conversions to/from `ℝ`, `ℝ≥0`, and `ℕ`. ## Main definitions -* `ennreal`: the extended nonnegative real numbers `[0, ∞]`; defined as `with_top ℝ≥0`; it is +* `ℝ≥0∞`: the extended nonnegative real numbers `[0, ∞]`; defined as `with_top ℝ≥0`; it is equipped with the following structures: - coercion from `ℝ≥0` defined in the natural way; @@ -39,12 +39,12 @@ to/from `ℝ`, `ℝ≥0`, and `ℕ`. - `a / b` is defined as `a * b⁻¹`. The addition and multiplication defined this way together with `0 = ↑0` and `1 = ↑1` turn - `ennreal` into a canonically ordered commutative semiring of characteristic zero. + `ℝ≥0∞` into a canonically ordered commutative semiring of characteristic zero. * Coercions to/from other types: - - coercion `ℝ≥0 → ennreal` is defined as `has_coe`, so one can use `(p : ℝ≥0)` in a context that - expects `a : ennreal`, and Lean will apply `coe` automatically; + - coercion `ℝ≥0 → ℝ≥0∞` is defined as `has_coe`, so one can use `(p : ℝ≥0)` in a context that + expects `a : ℝ≥0∞`, and Lean will apply `coe` automatically; - `ennreal.to_nnreal` sends `↑p` to `p` and `∞` to `0`; @@ -52,19 +52,20 @@ to/from `ℝ`, `ℝ≥0`, and `ℕ`. - `ennreal.of_real := coe ∘ nnreal.of_real` sends `x : ℝ` to `↑⟨max x 0, _⟩` - - `ennreal.ne_top_equiv_nnreal` is an equivalence between `{a : ennreal // a ≠ 0}` and `ℝ≥0`. + - `ennreal.ne_top_equiv_nnreal` is an equivalence between `{a : ℝ≥0∞ // a ≠ 0}` and `ℝ≥0`. ## Implementation notes -We define a `can_lift ennreal ℝ≥0` instance, so one of the ways to prove theorems about an `ennreal` +We define a `can_lift ℝ≥0∞ ℝ≥0` instance, so one of the ways to prove theorems about an `ℝ≥0∞` number `a` is to consider the cases `a = ∞` and `a ≠ ∞`, and use the tactic `lift a to ℝ≥0 using ha` in the second case. This instance is even more useful if one already has `ha : a ≠ ∞` in the -context, or if we have `(f : α → ennreal) (hf : ∀ x, f x ≠ ∞)`. +context, or if we have `(f : α → ℝ≥0∞) (hf : ∀ x, f x ≠ ∞)`. ## Notations -* `ℝ≥0`: type of nonnegative real numbers `[0, ∞)`; defined in `data.real.nnreal`; -* `∞`: a localized notation in `ennreal` for `⊤ : ennreal`. +* `ℝ≥0∞`: the type of the extended nonnegative real numbers; +* `ℝ≥0`: the type of nonnegative real numbers `[0, ∞)`; defined in `data.real.nnreal`; +* `∞`: a localized notation in `ℝ≥0∞` for `⊤ : ℝ≥0∞`. -/ noncomputable theory @@ -79,41 +80,42 @@ variables {α : Type*} {β : Type*} derive nontrivial] def ennreal := with_top ℝ≥0 +localized "notation `ℝ≥0∞` := ennreal" in ennreal localized "notation `∞` := (⊤ : ennreal)" in ennreal namespace ennreal -variables {a b c d : ennreal} {r p q : ℝ≥0} +variables {a b c d : ℝ≥0∞} {r p q : ℝ≥0} -instance : inhabited ennreal := ⟨0⟩ +instance : inhabited ℝ≥0∞ := ⟨0⟩ -instance : has_coe ℝ≥0 ennreal := ⟨ option.some ⟩ +instance : has_coe ℝ≥0 ℝ≥0∞ := ⟨ option.some ⟩ -instance : can_lift ennreal ℝ≥0 := +instance : can_lift ℝ≥0∞ ℝ≥0 := { coe := coe, cond := λ r, r ≠ ∞, prf := λ x hx, ⟨option.get $ option.ne_none_iff_is_some.1 hx, option.some_get _⟩ } -@[simp] lemma none_eq_top : (none : ennreal) = ∞ := rfl -@[simp] lemma some_eq_coe (a : ℝ≥0) : (some a : ennreal) = (↑a : ennreal) := rfl +@[simp] lemma none_eq_top : (none : ℝ≥0∞) = ∞ := rfl +@[simp] lemma some_eq_coe (a : ℝ≥0) : (some a : ℝ≥0∞) = (↑a : ℝ≥0∞) := rfl /-- `to_nnreal x` returns `x` if it is real, otherwise 0. -/ -protected def to_nnreal : ennreal → ℝ≥0 +protected def to_nnreal : ℝ≥0∞ → ℝ≥0 | (some r) := r | none := 0 /-- `to_real x` returns `x` if it is real, `0` otherwise. -/ -protected def to_real (a : ennreal) : real := coe (a.to_nnreal) +protected def to_real (a : ℝ≥0∞) : real := coe (a.to_nnreal) /-- `of_real x` returns `x` if it is nonnegative, `0` otherwise. -/ -protected def of_real (r : real) : ennreal := coe (nnreal.of_real r) +protected def of_real (r : real) : ℝ≥0∞ := coe (nnreal.of_real r) -@[simp, norm_cast] lemma to_nnreal_coe : (r : ennreal).to_nnreal = r := rfl +@[simp, norm_cast] lemma to_nnreal_coe : (r : ℝ≥0∞).to_nnreal = r := rfl -@[simp] lemma coe_to_nnreal : ∀{a:ennreal}, a ≠ ∞ → ↑(a.to_nnreal) = a +@[simp] lemma coe_to_nnreal : ∀{a:ℝ≥0∞}, a ≠ ∞ → ↑(a.to_nnreal) = a | (some r) h := rfl | none h := (h rfl).elim -@[simp] lemma of_real_to_real {a : ennreal} (h : a ≠ ∞) : ennreal.of_real (a.to_real) = a := +@[simp] lemma of_real_to_real {a : ℝ≥0∞} (h : a ≠ ∞) : ennreal.of_real (a.to_real) = a := by simp [ennreal.to_real, ennreal.of_real, h] @[simp] lemma to_real_of_real {r : ℝ} (h : 0 ≤ r) : ennreal.to_real (ennreal.of_real r) = r := @@ -121,63 +123,63 @@ by simp [ennreal.to_real, ennreal.of_real, nnreal.coe_of_real _ h] lemma to_real_of_real' {r : ℝ} : ennreal.to_real (ennreal.of_real r) = max r 0 := rfl -lemma coe_to_nnreal_le_self : ∀{a:ennreal}, ↑(a.to_nnreal) ≤ a +lemma coe_to_nnreal_le_self : ∀{a:ℝ≥0∞}, ↑(a.to_nnreal) ≤ a | (some r) := by rw [some_eq_coe, to_nnreal_coe]; exact le_refl _ | none := le_top -lemma coe_nnreal_eq (r : ℝ≥0) : (r : ennreal) = ennreal.of_real r := +lemma coe_nnreal_eq (r : ℝ≥0) : (r : ℝ≥0∞) = ennreal.of_real r := by { rw [ennreal.of_real, nnreal.of_real], cases r with r h, congr, dsimp, rw max_eq_left h } lemma of_real_eq_coe_nnreal {x : ℝ} (h : 0 ≤ x) : - ennreal.of_real x = @coe ℝ≥0 ennreal _ (⟨x, h⟩ : ℝ≥0) := + ennreal.of_real x = @coe ℝ≥0 ℝ≥0∞ _ (⟨x, h⟩ : ℝ≥0) := by { rw [coe_nnreal_eq], refl } @[simp] lemma of_real_coe_nnreal : ennreal.of_real p = p := (coe_nnreal_eq p).symm -@[simp, norm_cast] lemma coe_zero : ↑(0 : ℝ≥0) = (0 : ennreal) := rfl -@[simp, norm_cast] lemma coe_one : ↑(1 : ℝ≥0) = (1 : ennreal) := rfl +@[simp, norm_cast] lemma coe_zero : ↑(0 : ℝ≥0) = (0 : ℝ≥0∞) := rfl +@[simp, norm_cast] lemma coe_one : ↑(1 : ℝ≥0) = (1 : ℝ≥0∞) := rfl -@[simp] lemma to_real_nonneg {a : ennreal} : 0 ≤ a.to_real := by simp [ennreal.to_real] +@[simp] lemma to_real_nonneg {a : ℝ≥0∞} : 0 ≤ a.to_real := by simp [ennreal.to_real] @[simp] lemma top_to_nnreal : ∞.to_nnreal = 0 := rfl @[simp] lemma top_to_real : ∞.to_real = 0 := rfl -@[simp] lemma one_to_real : (1 : ennreal).to_real = 1 := rfl -@[simp] lemma one_to_nnreal : (1 : ennreal).to_nnreal = 1 := rfl -@[simp] lemma coe_to_real (r : ℝ≥0) : (r : ennreal).to_real = r := rfl -@[simp] lemma zero_to_nnreal : (0 : ennreal).to_nnreal = 0 := rfl -@[simp] lemma zero_to_real : (0 : ennreal).to_real = 0 := rfl +@[simp] lemma one_to_real : (1 : ℝ≥0∞).to_real = 1 := rfl +@[simp] lemma one_to_nnreal : (1 : ℝ≥0∞).to_nnreal = 1 := rfl +@[simp] lemma coe_to_real (r : ℝ≥0) : (r : ℝ≥0∞).to_real = r := rfl +@[simp] lemma zero_to_nnreal : (0 : ℝ≥0∞).to_nnreal = 0 := rfl +@[simp] lemma zero_to_real : (0 : ℝ≥0∞).to_real = 0 := rfl @[simp] lemma of_real_zero : ennreal.of_real (0 : ℝ) = 0 := by simp [ennreal.of_real]; refl -@[simp] lemma of_real_one : ennreal.of_real (1 : ℝ) = (1 : ennreal) := +@[simp] lemma of_real_one : ennreal.of_real (1 : ℝ) = (1 : ℝ≥0∞) := by simp [ennreal.of_real] -lemma of_real_to_real_le {a : ennreal} : ennreal.of_real (a.to_real) ≤ a := +lemma of_real_to_real_le {a : ℝ≥0∞} : ennreal.of_real (a.to_real) ≤ a := if ha : a = ∞ then ha.symm ▸ le_top else le_of_eq (of_real_to_real ha) -lemma forall_ennreal {p : ennreal → Prop} : (∀a, p a) ↔ (∀r:ℝ≥0, p r) ∧ p ∞ := +lemma forall_ennreal {p : ℝ≥0∞ → Prop} : (∀a, p a) ↔ (∀r:ℝ≥0, p r) ∧ p ∞ := ⟨assume h, ⟨assume r, h _, h _⟩, assume ⟨h₁, h₂⟩ a, match a with some r := h₁ _ | none := h₂ end⟩ -lemma forall_ne_top {p : ennreal → Prop} : (∀ a ≠ ∞, p a) ↔ ∀ r : ℝ≥0, p r := +lemma forall_ne_top {p : ℝ≥0∞ → Prop} : (∀ a ≠ ∞, p a) ↔ ∀ r : ℝ≥0, p r := option.ball_ne_none -lemma exists_ne_top {p : ennreal → Prop} : (∃ a ≠ ∞, p a) ↔ ∃ r : ℝ≥0, p r := +lemma exists_ne_top {p : ℝ≥0∞ → Prop} : (∃ a ≠ ∞, p a) ↔ ∃ r : ℝ≥0, p r := option.bex_ne_none -lemma to_nnreal_eq_zero_iff (x : ennreal) : x.to_nnreal = 0 ↔ x = 0 ∨ x = ∞ := +lemma to_nnreal_eq_zero_iff (x : ℝ≥0∞) : x.to_nnreal = 0 ↔ x = 0 ∨ x = ∞ := ⟨begin cases x, { simp [none_eq_top] }, - { have A : some (0:ℝ≥0) = (0:ennreal) := rfl, + { have A : some (0:ℝ≥0) = (0:ℝ≥0∞) := rfl, simp [ennreal.to_nnreal, A] {contextual := tt} } end, by intro h; cases h; simp [h]⟩ -lemma to_real_eq_zero_iff (x : ennreal) : x.to_real = 0 ↔ x = 0 ∨ x = ∞ := +lemma to_real_eq_zero_iff (x : ℝ≥0∞) : x.to_real = 0 ↔ x = 0 ∨ x = ∞ := by simp [ennreal.to_real, to_nnreal_eq_zero_iff] -@[simp] lemma coe_ne_top : (r : ennreal) ≠ ∞ := with_top.coe_ne_top -@[simp] lemma top_ne_coe : ∞ ≠ (r : ennreal) := with_top.top_ne_coe +@[simp] lemma coe_ne_top : (r : ℝ≥0∞) ≠ ∞ := with_top.coe_ne_top +@[simp] lemma top_ne_coe : ∞ ≠ (r : ℝ≥0∞) := with_top.top_ne_coe @[simp] lemma of_real_ne_top {r : ℝ} : ennreal.of_real r ≠ ∞ := by simp [ennreal.of_real] @[simp] lemma of_real_lt_top {r : ℝ} : ennreal.of_real r < ∞ := lt_top_iff_ne_top.2 of_real_ne_top @[simp] lemma top_ne_of_real {r : ℝ} : ∞ ≠ ennreal.of_real r := by simp [ennreal.of_real] @@ -188,87 +190,87 @@ by simp [ennreal.to_real, to_nnreal_eq_zero_iff] @[simp] lemma one_ne_top : 1 ≠ ∞ := coe_ne_top @[simp] lemma top_ne_one : ∞ ≠ 1 := top_ne_coe -@[simp, norm_cast] lemma coe_eq_coe : (↑r : ennreal) = ↑q ↔ r = q := with_top.coe_eq_coe -@[simp, norm_cast] lemma coe_le_coe : (↑r : ennreal) ≤ ↑q ↔ r ≤ q := with_top.coe_le_coe -@[simp, norm_cast] lemma coe_lt_coe : (↑r : ennreal) < ↑q ↔ r < q := with_top.coe_lt_coe -lemma coe_mono : monotone (coe : ℝ≥0 → ennreal) := λ _ _, coe_le_coe.2 +@[simp, norm_cast] lemma coe_eq_coe : (↑r : ℝ≥0∞) = ↑q ↔ r = q := with_top.coe_eq_coe +@[simp, norm_cast] lemma coe_le_coe : (↑r : ℝ≥0∞) ≤ ↑q ↔ r ≤ q := with_top.coe_le_coe +@[simp, norm_cast] lemma coe_lt_coe : (↑r : ℝ≥0∞) < ↑q ↔ r < q := with_top.coe_lt_coe +lemma coe_mono : monotone (coe : ℝ≥0 → ℝ≥0∞) := λ _ _, coe_le_coe.2 -@[simp, norm_cast] lemma coe_eq_zero : (↑r : ennreal) = 0 ↔ r = 0 := coe_eq_coe -@[simp, norm_cast] lemma zero_eq_coe : 0 = (↑r : ennreal) ↔ 0 = r := coe_eq_coe -@[simp, norm_cast] lemma coe_eq_one : (↑r : ennreal) = 1 ↔ r = 1 := coe_eq_coe -@[simp, norm_cast] lemma one_eq_coe : 1 = (↑r : ennreal) ↔ 1 = r := coe_eq_coe -@[simp, norm_cast] lemma coe_nonneg : 0 ≤ (↑r : ennreal) ↔ 0 ≤ r := coe_le_coe -@[simp, norm_cast] lemma coe_pos : 0 < (↑r : ennreal) ↔ 0 < r := coe_lt_coe +@[simp, norm_cast] lemma coe_eq_zero : (↑r : ℝ≥0∞) = 0 ↔ r = 0 := coe_eq_coe +@[simp, norm_cast] lemma zero_eq_coe : 0 = (↑r : ℝ≥0∞) ↔ 0 = r := coe_eq_coe +@[simp, norm_cast] lemma coe_eq_one : (↑r : ℝ≥0∞) = 1 ↔ r = 1 := coe_eq_coe +@[simp, norm_cast] lemma one_eq_coe : 1 = (↑r : ℝ≥0∞) ↔ 1 = r := coe_eq_coe +@[simp, norm_cast] lemma coe_nonneg : 0 ≤ (↑r : ℝ≥0∞) ↔ 0 ≤ r := coe_le_coe +@[simp, norm_cast] lemma coe_pos : 0 < (↑r : ℝ≥0∞) ↔ 0 < r := coe_lt_coe -@[simp, norm_cast] lemma coe_add : ↑(r + p) = (r + p : ennreal) := with_top.coe_add -@[simp, norm_cast] lemma coe_mul : ↑(r * p) = (r * p : ennreal) := with_top.coe_mul +@[simp, norm_cast] lemma coe_add : ↑(r + p) = (r + p : ℝ≥0∞) := with_top.coe_add +@[simp, norm_cast] lemma coe_mul : ↑(r * p) = (r * p : ℝ≥0∞) := with_top.coe_mul -@[simp, norm_cast] lemma coe_bit0 : (↑(bit0 r) : ennreal) = bit0 r := coe_add -@[simp, norm_cast] lemma coe_bit1 : (↑(bit1 r) : ennreal) = bit1 r := by simp [bit1] -lemma coe_two : ((2:ℝ≥0) : ennreal) = 2 := by norm_cast +@[simp, norm_cast] lemma coe_bit0 : (↑(bit0 r) : ℝ≥0∞) = bit0 r := coe_add +@[simp, norm_cast] lemma coe_bit1 : (↑(bit1 r) : ℝ≥0∞) = bit1 r := by simp [bit1] +lemma coe_two : ((2:ℝ≥0) : ℝ≥0∞) = 2 := by norm_cast -protected lemma zero_lt_one : 0 < (1 : ennreal) := +protected lemma zero_lt_one : 0 < (1 : ℝ≥0∞) := canonically_ordered_semiring.zero_lt_one -@[simp] lemma one_lt_two : (1 : ennreal) < 2 := +@[simp] lemma one_lt_two : (1 : ℝ≥0∞) < 2 := coe_one ▸ coe_two ▸ by exact_mod_cast (@one_lt_two ℕ _ _) -@[simp] lemma zero_lt_two : (0:ennreal) < 2 := lt_trans ennreal.zero_lt_one one_lt_two -lemma two_ne_zero : (2:ennreal) ≠ 0 := (ne_of_lt zero_lt_two).symm -lemma two_ne_top : (2:ennreal) ≠ ∞ := coe_two ▸ coe_ne_top +@[simp] lemma zero_lt_two : (0:ℝ≥0∞) < 2 := lt_trans ennreal.zero_lt_one one_lt_two +lemma two_ne_zero : (2:ℝ≥0∞) ≠ 0 := (ne_of_lt zero_lt_two).symm +lemma two_ne_top : (2:ℝ≥0∞) ≠ ∞ := coe_two ▸ coe_ne_top -/-- The set of `ennreal` numbers that are not equal to `∞` is equivalent to `ℝ≥0`. -/ +/-- The set of numbers in `ℝ≥0∞` that are not equal to `∞` is equivalent to `ℝ≥0`. -/ def ne_top_equiv_nnreal : {a | a ≠ ∞} ≃ ℝ≥0 := { to_fun := λ x, ennreal.to_nnreal x, inv_fun := λ x, ⟨x, coe_ne_top⟩, left_inv := λ ⟨x, hx⟩, subtype.eq $ coe_to_nnreal hx, right_inv := λ x, to_nnreal_coe } -lemma cinfi_ne_top [has_Inf α] (f : ennreal → α) : (⨅ x : {x // x ≠ ∞}, f x) = ⨅ x : ℝ≥0, f x := +lemma cinfi_ne_top [has_Inf α] (f : ℝ≥0∞ → α) : (⨅ x : {x // x ≠ ∞}, f x) = ⨅ x : ℝ≥0, f x := eq.symm $ infi_congr _ ne_top_equiv_nnreal.symm.surjective $ λ x, rfl -lemma infi_ne_top [complete_lattice α] (f : ennreal → α) : (⨅ x ≠ ∞, f x) = ⨅ x : ℝ≥0, f x := +lemma infi_ne_top [complete_lattice α] (f : ℝ≥0∞ → α) : (⨅ x ≠ ∞, f x) = ⨅ x : ℝ≥0, f x := by rw [infi_subtype', cinfi_ne_top] -lemma csupr_ne_top [has_Sup α] (f : ennreal → α) : (⨆ x : {x // x ≠ ∞}, f x) = ⨆ x : ℝ≥0, f x := +lemma csupr_ne_top [has_Sup α] (f : ℝ≥0∞ → α) : (⨆ x : {x // x ≠ ∞}, f x) = ⨆ x : ℝ≥0, f x := @cinfi_ne_top (order_dual α) _ _ -lemma supr_ne_top [complete_lattice α] (f : ennreal → α) : (⨆ x ≠ ∞, f x) = ⨆ x : ℝ≥0, f x := +lemma supr_ne_top [complete_lattice α] (f : ℝ≥0∞ → α) : (⨆ x ≠ ∞, f x) = ⨆ x : ℝ≥0, f x := @infi_ne_top (order_dual α) _ _ -lemma infi_ennreal {α : Type*} [complete_lattice α] {f : ennreal → α} : +lemma infi_ennreal {α : Type*} [complete_lattice α] {f : ℝ≥0∞ → α} : (⨅ n, f n) = (⨅ n : ℝ≥0, f n) ⊓ f ∞ := le_antisymm (le_inf (le_infi $ assume i, infi_le _ _) (infi_le _ _)) (le_infi $ forall_ennreal.2 ⟨assume r, inf_le_left_of_le $ infi_le _ _, inf_le_right⟩) -lemma supr_ennreal {α : Type*} [complete_lattice α] {f : ennreal → α} : +lemma supr_ennreal {α : Type*} [complete_lattice α] {f : ℝ≥0∞ → α} : (⨆ n, f n) = (⨆ n : ℝ≥0, f n) ⊔ f ∞ := @infi_ennreal (order_dual α) _ _ @[simp] lemma add_top : a + ∞ = ∞ := with_top.add_top @[simp] lemma top_add : ∞ + a = ∞ := with_top.top_add -/-- Coercion `ℝ≥0 → ennreal` as a `ring_hom`. -/ -def of_nnreal_hom : ℝ≥0 →+* ennreal := +/-- Coercion `ℝ≥0 → ℝ≥0∞` as a `ring_hom`. -/ +def of_nnreal_hom : ℝ≥0 →+* ℝ≥0∞ := ⟨coe, coe_one, λ _ _, coe_mul, coe_zero, λ _ _, coe_add⟩ @[simp] lemma coe_of_nnreal_hom : ⇑of_nnreal_hom = coe := rfl @[simp, norm_cast] lemma coe_indicator {α} (s : set α) (f : α → ℝ≥0) (a : α) : - ((s.indicator f a : ℝ≥0) : ennreal) = s.indicator (λ x, f x) a := -(of_nnreal_hom : ℝ≥0 →+ ennreal).map_indicator _ _ _ + ((s.indicator f a : ℝ≥0) : ℝ≥0∞) = s.indicator (λ x, f x) a := +(of_nnreal_hom : ℝ≥0 →+ ℝ≥0∞).map_indicator _ _ _ -@[simp, norm_cast] lemma coe_pow (n : ℕ) : (↑(r^n) : ennreal) = r^n := +@[simp, norm_cast] lemma coe_pow (n : ℕ) : (↑(r^n) : ℝ≥0∞) = r^n := of_nnreal_hom.map_pow r n @[simp] lemma add_eq_top : a + b = ∞ ↔ a = ∞ ∨ b = ∞ := with_top.add_eq_top @[simp] lemma add_lt_top : a + b < ∞ ↔ a < ∞ ∧ b < ∞ := with_top.add_lt_top -lemma to_nnreal_add {r₁ r₂ : ennreal} (h₁ : r₁ < ∞) (h₂ : r₂ < ∞) : +lemma to_nnreal_add {r₁ r₂ : ℝ≥0∞} (h₁ : r₁ < ∞) (h₂ : r₂ < ∞) : (r₁ + r₂).to_nnreal = r₁.to_nnreal + r₂.to_nnreal := begin rw [← coe_eq_coe, coe_add, coe_to_nnreal, coe_to_nnreal, coe_to_nnreal]; - apply @ne_top_of_lt ennreal _ _ ∞, + apply @ne_top_of_lt ℝ≥0∞ _ _ ∞, exact h₂, exact h₁, exact add_lt_top.2 ⟨h₁, h₂⟩ @@ -279,7 +281,7 @@ end protected lemma lt_top_iff_ne_top : a < ∞ ↔ a ≠ ∞ := lt_top_iff_ne_top protected lemma bot_lt_iff_ne_bot : 0 < a ↔ a ≠ 0 := bot_lt_iff_ne_bot -lemma not_lt_top {x : ennreal} : ¬ x < ∞ ↔ x = ∞ := by rw [lt_top_iff_ne_top, not_not] +lemma not_lt_top {x : ℝ≥0∞} : ¬ x < ∞ ↔ x = ∞ := by rw [lt_top_iff_ne_top, not_not] lemma add_ne_top : a + b ≠ ∞ ↔ a ≠ ∞ ∧ b ≠ ∞ := by simpa only [lt_top_iff_ne_top] using add_lt_top @@ -317,7 +319,7 @@ by { rw [ennreal.lt_top_iff_ne_top] at h ⊢, exact ne_top_of_mul_ne_top_left h lemma lt_top_of_mul_lt_top_right (h : a * b < ∞) (ha : a ≠ 0) : b < ∞ := lt_top_of_mul_lt_top_left (by rwa [mul_comm]) ha -lemma mul_lt_top_iff {a b : ennreal} : a * b < ∞ ↔ (a < ∞ ∧ b < ∞) ∨ a = 0 ∨ b = 0 := +lemma mul_lt_top_iff {a b : ℝ≥0∞} : a * b < ∞ ↔ (a < ∞ ∧ b < ∞) ∨ a = 0 ∨ b = 0 := begin split, { intro h, rw [← or_assoc, or_iff_not_imp_right, or_iff_not_imp_right], intros hb ha, @@ -325,6 +327,9 @@ begin { rintro (⟨ha, hb⟩|rfl|rfl); [exact mul_lt_top ha hb, simp, simp] } end +lemma mul_self_lt_top_iff {a : ℝ≥0∞} : a * a < ⊤ ↔ a < ⊤ := +by { rw [ennreal.mul_lt_top_iff, and_self, or_self, or_iff_left_iff_imp], rintro rfl, norm_num } + @[simp] lemma mul_pos : 0 < a * b ↔ 0 < a ∧ 0 < b := by simp only [pos_iff_ne_zero, ne.def, mul_eq_zero, not_or_distrib] @@ -339,27 +344,27 @@ lemma pow_lt_top : a < ∞ → ∀ n:ℕ, a^n < ∞ := by simpa only [lt_top_iff_ne_top] using pow_ne_top @[simp, norm_cast] lemma coe_finset_sum {s : finset α} {f : α → ℝ≥0} : - ↑(∑ a in s, f a) = (∑ a in s, f a : ennreal) := + ↑(∑ a in s, f a) = (∑ a in s, f a : ℝ≥0∞) := of_nnreal_hom.map_sum f s @[simp, norm_cast] lemma coe_finset_prod {s : finset α} {f : α → ℝ≥0} : - ↑(∏ a in s, f a) = ((∏ a in s, f a) : ennreal) := + ↑(∏ a in s, f a) = ((∏ a in s, f a) : ℝ≥0∞) := of_nnreal_hom.map_prod f s section order -@[simp] lemma bot_eq_zero : (⊥ : ennreal) = 0 := rfl +@[simp] lemma bot_eq_zero : (⊥ : ℝ≥0∞) = 0 := rfl @[simp] lemma coe_lt_top : coe r < ∞ := with_top.coe_lt_top r @[simp] lemma not_top_le_coe : ¬ ∞ ≤ ↑r := with_top.not_top_le_coe r -lemma zero_lt_coe_iff : 0 < (↑p : ennreal) ↔ 0 < p := coe_lt_coe -@[simp, norm_cast] lemma one_le_coe_iff : (1:ennreal) ≤ ↑r ↔ 1 ≤ r := coe_le_coe -@[simp, norm_cast] lemma coe_le_one_iff : ↑r ≤ (1:ennreal) ↔ r ≤ 1 := coe_le_coe -@[simp, norm_cast] lemma coe_lt_one_iff : (↑p : ennreal) < 1 ↔ p < 1 := coe_lt_coe -@[simp, norm_cast] lemma one_lt_coe_iff : 1 < (↑p : ennreal) ↔ 1 < p := coe_lt_coe -@[simp, norm_cast] lemma coe_nat (n : ℕ) : ((n : ℝ≥0) : ennreal) = n := with_top.coe_nat n +lemma zero_lt_coe_iff : 0 < (↑p : ℝ≥0∞) ↔ 0 < p := coe_lt_coe +@[simp, norm_cast] lemma one_le_coe_iff : (1:ℝ≥0∞) ≤ ↑r ↔ 1 ≤ r := coe_le_coe +@[simp, norm_cast] lemma coe_le_one_iff : ↑r ≤ (1:ℝ≥0∞) ↔ r ≤ 1 := coe_le_coe +@[simp, norm_cast] lemma coe_lt_one_iff : (↑p : ℝ≥0∞) < 1 ↔ p < 1 := coe_lt_coe +@[simp, norm_cast] lemma one_lt_coe_iff : 1 < (↑p : ℝ≥0∞) ↔ 1 < p := coe_lt_coe +@[simp, norm_cast] lemma coe_nat (n : ℕ) : ((n : ℝ≥0) : ℝ≥0∞) = n := with_top.coe_nat n @[simp] lemma of_real_coe_nat (n : ℕ) : ennreal.of_real n = n := by simp [ennreal.of_real] -@[simp] lemma nat_ne_top (n : ℕ) : (n : ennreal) ≠ ∞ := with_top.nat_ne_top n +@[simp] lemma nat_ne_top (n : ℕ) : (n : ℝ≥0∞) ≠ ∞ := with_top.nat_ne_top n @[simp] lemma top_ne_nat (n : ℕ) : ∞ ≠ n := with_top.top_ne_nat n @[simp] lemma one_lt_top : 1 < ∞ := coe_lt_top @@ -369,7 +374,7 @@ lemma coe_le_iff : ↑r ≤ a ↔ (∀p:ℝ≥0, a = p → r ≤ p) := with_top. lemma lt_iff_exists_coe : a < b ↔ (∃p:ℝ≥0, a = p ∧ ↑p < b) := with_top.lt_iff_exists_coe @[simp, norm_cast] lemma coe_finset_sup {s : finset α} {f : α → ℝ≥0} : - ↑(s.sup f) = s.sup (λ x, (f x : ennreal)) := + ↑(s.sup f) = s.sup (λ x, (f x : ℝ≥0∞)) := finset.comp_sup_eq_sup_comp_of_is_total _ coe_mono rfl lemma pow_le_pow {n m : ℕ} (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := @@ -411,7 +416,7 @@ with_top.add_lt_add_iff_right lemma lt_add_right (ha : a < ∞) (hb : 0 < b) : a < a + b := by rwa [← add_lt_add_iff_left ha, add_zero] at hb -lemma le_of_forall_pos_le_add : ∀{a b : ennreal}, (∀ε:ℝ≥0, 0 < ε → b < ∞ → a ≤ b + ε) → a ≤ b +lemma le_of_forall_pos_le_add : ∀{a b : ℝ≥0∞}, (∀ε:ℝ≥0, 0 < ε → b < ∞ → a ≤ b + ε) → a ≤ b | a none h := le_top | none (some a) h := have ∞ ≤ ↑a + ↑(1:ℝ≥0), from h 1 zero_lt_one coe_lt_top, @@ -421,7 +426,7 @@ lemma le_of_forall_pos_le_add : ∀{a b : ennreal}, (∀ε:ℝ≥0, 0 < ε → b at *; exact nnreal.le_of_forall_pos_le_add h lemma lt_iff_exists_rat_btwn : - a < b ↔ (∃q:ℚ, 0 ≤ q ∧ a < nnreal.of_real q ∧ (nnreal.of_real q:ennreal) < b) := + a < b ↔ (∃q:ℚ, 0 ≤ q ∧ a < nnreal.of_real q ∧ (nnreal.of_real q:ℝ≥0∞) < b) := ⟨λ h, begin rcases lt_iff_exists_coe.1 h with ⟨p, rfl, _⟩, @@ -433,13 +438,13 @@ lemma lt_iff_exists_rat_btwn : λ ⟨q, q0, qa, qb⟩, lt_trans qa qb⟩ lemma lt_iff_exists_real_btwn : - a < b ↔ (∃r:ℝ, 0 ≤ r ∧ a < ennreal.of_real r ∧ (ennreal.of_real r:ennreal) < b) := + a < b ↔ (∃r:ℝ, 0 ≤ r ∧ a < ennreal.of_real r ∧ (ennreal.of_real r:ℝ≥0∞) < b) := ⟨λ h, let ⟨q, q0, aq, qb⟩ := ennreal.lt_iff_exists_rat_btwn.1 h in ⟨q, rat.cast_nonneg.2 q0, aq, qb⟩, λ ⟨q, q0, qa, qb⟩, lt_trans qa qb⟩ lemma lt_iff_exists_nnreal_btwn : - a < b ↔ (∃r:ℝ≥0, a < r ∧ (r : ennreal) < b) := + a < b ↔ (∃r:ℝ≥0, a < r ∧ (r : ℝ≥0∞) < b) := with_top.lt_iff_exists_coe_btwn lemma lt_iff_exists_add_pos_lt : a < b ↔ (∃ r : ℝ≥0, 0 < r ∧ a + r < b) := @@ -460,18 +465,18 @@ begin exact nnreal.sub_add_cancel_of_le (le_of_lt ad) end -lemma coe_nat_lt_coe {n : ℕ} : (n : ennreal) < r ↔ ↑n < r := ennreal.coe_nat n ▸ coe_lt_coe -lemma coe_lt_coe_nat {n : ℕ} : (r : ennreal) < n ↔ r < n := ennreal.coe_nat n ▸ coe_lt_coe -@[norm_cast] lemma coe_nat_lt_coe_nat {m n : ℕ} : (m : ennreal) < n ↔ m < n := +lemma coe_nat_lt_coe {n : ℕ} : (n : ℝ≥0∞) < r ↔ ↑n < r := ennreal.coe_nat n ▸ coe_lt_coe +lemma coe_lt_coe_nat {n : ℕ} : (r : ℝ≥0∞) < n ↔ r < n := ennreal.coe_nat n ▸ coe_lt_coe +@[norm_cast] lemma coe_nat_lt_coe_nat {m n : ℕ} : (m : ℝ≥0∞) < n ↔ m < n := ennreal.coe_nat n ▸ coe_nat_lt_coe.trans nat.cast_lt -lemma coe_nat_ne_top {n : ℕ} : (n : ennreal) ≠ ∞ := ennreal.coe_nat n ▸ coe_ne_top -lemma coe_nat_mono : strict_mono (coe : ℕ → ennreal) := λ _ _, coe_nat_lt_coe_nat.2 -@[norm_cast] lemma coe_nat_le_coe_nat {m n : ℕ} : (m : ennreal) ≤ n ↔ m ≤ n := +lemma coe_nat_ne_top {n : ℕ} : (n : ℝ≥0∞) ≠ ∞ := ennreal.coe_nat n ▸ coe_ne_top +lemma coe_nat_mono : strict_mono (coe : ℕ → ℝ≥0∞) := λ _ _, coe_nat_lt_coe_nat.2 +@[norm_cast] lemma coe_nat_le_coe_nat {m n : ℕ} : (m : ℝ≥0∞) ≤ n ↔ m ≤ n := coe_nat_mono.le_iff_le -instance : char_zero ennreal := ⟨coe_nat_mono.injective⟩ +instance : char_zero ℝ≥0∞ := ⟨coe_nat_mono.injective⟩ -protected lemma exists_nat_gt {r : ennreal} (h : r ≠ ∞) : ∃n:ℕ, r < n := +protected lemma exists_nat_gt {r : ℝ≥0∞} (h : r ≠ ∞) : ∃n:ℕ, r < n := begin lift r to ℝ≥0 using h, rcases exists_nat_gt r with ⟨n, hn⟩, @@ -487,13 +492,13 @@ begin exact add_lt_add ac bd end -@[norm_cast] lemma coe_min : ((min r p:ℝ≥0):ennreal) = min r p := +@[norm_cast] lemma coe_min : ((min r p:ℝ≥0):ℝ≥0∞) = min r p := coe_mono.map_min -@[norm_cast] lemma coe_max : ((max r p:ℝ≥0):ennreal) = max r p := +@[norm_cast] lemma coe_max : ((max r p:ℝ≥0):ℝ≥0∞) = max r p := coe_mono.map_max -lemma le_of_top_imp_top_of_to_nnreal_le {a b : ennreal} (h : a = ⊤ → b = ⊤) +lemma le_of_top_imp_top_of_to_nnreal_le {a b : ℝ≥0∞} (h : a = ⊤ → b = ⊤) (h_nnreal : a ≠ ⊤ → b ≠ ⊤ → a.to_nnreal ≤ b.to_nnreal) : a ≤ b := begin @@ -511,14 +516,14 @@ end order section complete_lattice -lemma coe_Sup {s : set ℝ≥0} : bdd_above s → (↑(Sup s) : ennreal) = (⨆a∈s, ↑a) := with_top.coe_Sup -lemma coe_Inf {s : set ℝ≥0} : s.nonempty → (↑(Inf s) : ennreal) = (⨅a∈s, ↑a) := with_top.coe_Inf +lemma coe_Sup {s : set ℝ≥0} : bdd_above s → (↑(Sup s) : ℝ≥0∞) = (⨆a∈s, ↑a) := with_top.coe_Sup +lemma coe_Inf {s : set ℝ≥0} : s.nonempty → (↑(Inf s) : ℝ≥0∞) = (⨅a∈s, ↑a) := with_top.coe_Inf -@[simp] lemma top_mem_upper_bounds {s : set ennreal} : ∞ ∈ upper_bounds s := +@[simp] lemma top_mem_upper_bounds {s : set ℝ≥0∞} : ∞ ∈ upper_bounds s := assume x hx, le_top lemma coe_mem_upper_bounds {s : set ℝ≥0} : - ↑r ∈ upper_bounds ((coe : ℝ≥0 → ennreal) '' s) ↔ r ∈ upper_bounds s := + ↑r ∈ upper_bounds ((coe : ℝ≥0 → ℝ≥0∞) '' s) ↔ r ∈ upper_bounds s := by simp [upper_bounds, ball_image_iff, -mem_image, *] {contextual := tt} end complete_lattice @@ -580,13 +585,13 @@ mul_comm c a ▸ mul_comm c b ▸ mul_lt_mul_left end mul section sub -instance : has_sub ennreal := ⟨λa b, Inf {d | a ≤ d + b}⟩ +instance : has_sub ℝ≥0∞ := ⟨λa b, Inf {d | a ≤ d + b}⟩ -@[norm_cast] lemma coe_sub : ↑(p - r) = (↑p:ennreal) - r := +@[norm_cast] lemma coe_sub : ↑(p - r) = (↑p:ℝ≥0∞) - r := le_antisymm (le_Inf $ assume b (hb : ↑p ≤ b + r), coe_le_iff.2 $ by rintros d rfl; rwa [← coe_add, coe_le_coe, ← nnreal.sub_le_iff_le_add] at hb) - (Inf_le $ show (↑p : ennreal) ≤ ↑(p - r) + ↑r, + (Inf_le $ show (↑p : ℝ≥0∞) ≤ ↑(p - r) + ↑r, by rw [← coe_add, coe_le_coe, ← nnreal.sub_le_iff_le_add]) @[simp] lemma top_sub_coe : ∞ - ↑r = ∞ := @@ -609,7 +614,7 @@ Inf_le_Inf $ assume e (h : b ≤ e + d), ... ≤ e + d : h ... ≤ e + c : add_le_add (le_refl _) h₂ -@[simp] lemma add_sub_self : ∀{a b : ennreal}, b < ∞ → (a + b) - b = a +@[simp] lemma add_sub_self : ∀{a b : ℝ≥0∞}, b < ∞ → (a + b) - b = a | a none := by simp [none_eq_top] | none (some b) := by simp [none_eq_top, some_eq_coe] | (some a) (some b) := @@ -624,7 +629,7 @@ lemma add_right_inj (h : a < ∞) : a + b = a + c ↔ b = c := lemma add_left_inj (h : a < ∞) : b + a = c + a ↔ b = c := by rw [add_comm, add_comm c, add_right_inj h] -@[simp] lemma sub_add_cancel_of_le : ∀{a b : ennreal}, b ≤ a → (a - b) + b = a := +@[simp] lemma sub_add_cancel_of_le : ∀{a b : ℝ≥0∞}, b ≤ a → (a - b) + b = a := begin simp [forall_ennreal, le_coe_iff, -add_comm] {contextual := tt}, rintros r p x rfl h, @@ -686,7 +691,7 @@ begin rw [← coe_add, ← coe_sub, coe_lt_coe, coe_lt_coe, nnreal.lt_sub_iff_add_lt], end -lemma sub_le_self (a b : ennreal) : a - b ≤ a := +lemma sub_le_self (a b : ℝ≥0∞) : a - b ≤ a := ennreal.sub_le_iff_le_add.2 $ le_add_right (le_refl a) @[simp] lemma sub_zero : a - 0 = a := @@ -703,7 +708,7 @@ lemma sub_sub_cancel (h : a < ∞) (h2 : b ≤ a) : a - (a - b) = b := by rw [← add_left_inj (lt_of_le_of_lt (sub_le_self _ _) h), sub_add_cancel_of_le (sub_le_self _ _), add_sub_cancel_of_le h2] -lemma sub_right_inj {a b c : ennreal} (ha : a < ∞) (hb : b ≤ a) (hc : c ≤ a) : +lemma sub_right_inj {a b c : ℝ≥0∞} (ha : a < ∞) (hb : b ≤ a) (hc : c ≤ a) : a - b = a - c ↔ b = c := iff.intro begin @@ -748,27 +753,27 @@ section sum open finset /-- A product of finite numbers is still finite -/ -lemma prod_lt_top {s : finset α} {f : α → ennreal} (h : ∀a∈s, f a < ∞) : (∏ a in s, f a) < ∞ := +lemma prod_lt_top {s : finset α} {f : α → ℝ≥0∞} (h : ∀a∈s, f a < ∞) : (∏ a in s, f a) < ∞ := with_top.prod_lt_top h /-- A sum of finite numbers is still finite -/ -lemma sum_lt_top {s : finset α} {f : α → ennreal} : +lemma sum_lt_top {s : finset α} {f : α → ℝ≥0∞} : (∀a∈s, f a < ∞) → ∑ a in s, f a < ∞ := with_top.sum_lt_top /-- A sum of finite numbers is still finite -/ -lemma sum_lt_top_iff {s : finset α} {f : α → ennreal} : +lemma sum_lt_top_iff {s : finset α} {f : α → ℝ≥0∞} : ∑ a in s, f a < ∞ ↔ (∀a∈s, f a < ∞) := with_top.sum_lt_top_iff /-- A sum of numbers is infinite iff one of them is infinite -/ -lemma sum_eq_top_iff {s : finset α} {f : α → ennreal} : +lemma sum_eq_top_iff {s : finset α} {f : α → ℝ≥0∞} : (∑ x in s, f x) = ∞ ↔ (∃a∈s, f a = ∞) := with_top.sum_eq_top_iff -/-- seeing `ennreal` as `ℝ≥0` does not change their sum, unless one of the `ennreal` is +/-- seeing `ℝ≥0∞` as `ℝ≥0` does not change their sum, unless one of the `ℝ≥0∞` is infinity -/ -lemma to_nnreal_sum {s : finset α} {f : α → ennreal} (hf : ∀a∈s, f a < ∞) : +lemma to_nnreal_sum {s : finset α} {f : α → ℝ≥0∞} (hf : ∀a∈s, f a < ∞) : ennreal.to_nnreal (∑ a in s, f a) = ∑ a in s, ennreal.to_nnreal (f a) := begin rw [← coe_eq_coe, coe_to_nnreal, coe_finset_sum, sum_congr], @@ -777,8 +782,8 @@ begin { exact (sum_lt_top hf).ne } end -/-- seeing `ennreal` as `real` does not change their sum, unless one of the `ennreal` is infinity -/ -lemma to_real_sum {s : finset α} {f : α → ennreal} (hf : ∀a∈s, f a < ∞) : +/-- seeing `ℝ≥0∞` as `real` does not change their sum, unless one of the `ℝ≥0∞` is infinity -/ +lemma to_real_sum {s : finset α} {f : α → ℝ≥0∞} (hf : ∀a∈s, f a < ∞) : ennreal.to_real (∑ a in s, f a) = ∑ a in s, ennreal.to_real (f a) := by { rw [ennreal.to_real, to_nnreal_sum hf, nnreal.coe_sum], refl } @@ -793,7 +798,7 @@ end sum section interval -variables {x y z : ennreal} {ε ε₁ ε₂ : ennreal} {s : set ennreal} +variables {x y z : ℝ≥0∞} {ε ε₁ ε₂ : ℝ≥0∞} {s : set ℝ≥0∞} protected lemma Ico_eq_Iio : (Ico 0 y) = (Iio y) := ext $ assume a, iff.intro @@ -849,38 +854,38 @@ by unfold bit1; rw add_eq_top; simp end bit section inv -instance : has_inv ennreal := ⟨λa, Inf {b | 1 ≤ a * b}⟩ +instance : has_inv ℝ≥0∞ := ⟨λa, Inf {b | 1 ≤ a * b}⟩ -instance : div_inv_monoid ennreal := +instance : div_inv_monoid ℝ≥0∞ := { inv := has_inv.inv, - .. (infer_instance : monoid ennreal) } + .. (infer_instance : monoid ℝ≥0∞) } -@[simp] lemma inv_zero : (0 : ennreal)⁻¹ = ∞ := -show Inf {b : ennreal | 1 ≤ 0 * b} = ∞, by simp; refl +@[simp] lemma inv_zero : (0 : ℝ≥0∞)⁻¹ = ∞ := +show Inf {b : ℝ≥0∞ | 1 ≤ 0 * b} = ∞, by simp; refl @[simp] lemma inv_top : ∞⁻¹ = 0 := bot_unique $ le_of_forall_le_of_dense $ λ a (h : a > 0), Inf_le $ by simp [*, ne_of_gt h, top_mul] -@[simp, norm_cast] lemma coe_inv (hr : r ≠ 0) : (↑r⁻¹ : ennreal) = (↑r)⁻¹ := +@[simp, norm_cast] lemma coe_inv (hr : r ≠ 0) : (↑r⁻¹ : ℝ≥0∞) = (↑r)⁻¹ := le_antisymm (le_Inf $ assume b (hb : 1 ≤ ↑r * b), coe_le_iff.2 $ by rintros b rfl; rwa [← coe_mul, ← coe_one, coe_le_coe, ← nnreal.inv_le hr] at hb) (Inf_le $ by simp; rw [← coe_mul, mul_inv_cancel hr]; exact le_refl 1) -lemma coe_inv_le : (↑r⁻¹ : ennreal) ≤ (↑r)⁻¹ := +lemma coe_inv_le : (↑r⁻¹ : ℝ≥0∞) ≤ (↑r)⁻¹ := if hr : r = 0 then by simp only [hr, inv_zero, coe_zero, le_top] else by simp only [coe_inv hr, le_refl] -@[norm_cast] lemma coe_inv_two : ((2⁻¹:ℝ≥0):ennreal) = 2⁻¹ := +@[norm_cast] lemma coe_inv_two : ((2⁻¹:ℝ≥0):ℝ≥0∞) = 2⁻¹ := by rw [coe_inv (ne_of_gt _root_.zero_lt_two), coe_two] -@[simp, norm_cast] lemma coe_div (hr : r ≠ 0) : (↑(p / r) : ennreal) = p / r := +@[simp, norm_cast] lemma coe_div (hr : r ≠ 0) : (↑(p / r) : ℝ≥0∞) = p / r := by rw [div_eq_mul_inv, div_eq_mul_inv, coe_mul, coe_inv hr] -@[simp] lemma inv_one : (1:ennreal)⁻¹ = 1 := +@[simp] lemma inv_one : (1:ℝ≥0∞)⁻¹ = 1 := by simpa only [coe_inv one_ne_zero, coe_one] using coe_eq_coe.2 inv_one -@[simp] lemma div_one {a : ennreal} : a / 1 = a := +@[simp] lemma div_one {a : ℝ≥0∞} : a / 1 = a := by rw [div_eq_mul_inv, inv_one, mul_one] protected lemma inv_pow {n : ℕ} : (a^n)⁻¹ = (a⁻¹)^n := @@ -894,10 +899,10 @@ end by by_cases a = 0; cases a; simp [*, none_eq_top, some_eq_coe, -coe_inv, (coe_inv _).symm] at * -lemma inv_involutive : function.involutive (λ a:ennreal, a⁻¹) := +lemma inv_involutive : function.involutive (λ a:ℝ≥0∞, a⁻¹) := λ a, ennreal.inv_inv -lemma inv_bijective : function.bijective (λ a:ennreal, a⁻¹) := +lemma inv_bijective : function.bijective (λ a:ℝ≥0∞, a⁻¹) := ennreal.inv_involutive.bijective @[simp] lemma inv_eq_inv : a⁻¹ = b⁻¹ ↔ a = b := inv_bijective.1.eq_iff @@ -907,10 +912,10 @@ inv_zero ▸ inv_eq_inv lemma inv_ne_top : a⁻¹ ≠ ∞ ↔ a ≠ 0 := by simp -@[simp] lemma inv_lt_top {x : ennreal} : x⁻¹ < ∞ ↔ 0 < x := +@[simp] lemma inv_lt_top {x : ℝ≥0∞} : x⁻¹ < ∞ ↔ 0 < x := by { simp only [lt_top_iff_ne_top, inv_ne_top, pos_iff_ne_zero] } -lemma div_lt_top {x y : ennreal} (h1 : x < ∞) (h2 : 0 < y) : x / y < ∞ := +lemma div_lt_top {x y : ℝ≥0∞} (h1 : x < ∞) (h2 : 0 < y) : x / y < ∞ := mul_lt_top h1 (inv_lt_top.mpr h2) @[simp] lemma inv_eq_zero : a⁻¹ = 0 ↔ a = ∞ := @@ -987,7 +992,7 @@ begin have : a * ∞ = ∞, by simp [ennreal.mul_eq_top, ha], simp [this, ht] } }, by_cases hb : b ≠ 0, - { have : (b : ennreal) ≠ 0, by simp [hb], + { have : (b : ℝ≥0∞) ≠ 0, by simp [hb], rw [← ennreal.mul_le_mul_left this coe_ne_top], suffices : ↑b * a ≤ (↑b * ↑b⁻¹) * c ↔ a * ↑b ≤ c, { simpa [some_eq_coe, div_eq_mul_inv, hb, mul_left_comm, mul_comm, mul_assoc] }, @@ -1043,20 +1048,20 @@ end lemma inv_mul_cancel (h0 : a ≠ 0) (ht : a ≠ ∞) : a⁻¹ * a = 1 := mul_comm a a⁻¹ ▸ mul_inv_cancel h0 ht -lemma mul_le_iff_le_inv {a b r : ennreal} (hr₀ : r ≠ 0) (hr₁ : r ≠ ∞) : (r * a ≤ b ↔ a ≤ r⁻¹ * b) := +lemma mul_le_iff_le_inv {a b r : ℝ≥0∞} (hr₀ : r ≠ 0) (hr₁ : r ≠ ∞) : (r * a ≤ b ↔ a ≤ r⁻¹ * b) := by rw [← @ennreal.mul_le_mul_left _ a _ hr₀ hr₁, ← mul_assoc, mul_inv_cancel hr₀ hr₁, one_mul] -lemma le_of_forall_nnreal_lt {x y : ennreal} (h : ∀ r : ℝ≥0, ↑r < x → ↑r ≤ y) : x ≤ y := +lemma le_of_forall_nnreal_lt {x y : ℝ≥0∞} (h : ∀ r : ℝ≥0, ↑r < x → ↑r ≤ y) : x ≤ y := begin refine le_of_forall_ge_of_dense (λ r hr, _), lift r to ℝ≥0 using ne_top_of_lt hr, exact h r hr end -lemma eq_top_of_forall_nnreal_le {x : ennreal} (h : ∀ r : ℝ≥0, ↑r ≤ x) : x = ∞ := +lemma eq_top_of_forall_nnreal_le {x : ℝ≥0∞} (h : ∀ r : ℝ≥0, ↑r ≤ x) : x = ∞ := top_unique $ le_of_forall_nnreal_lt $ λ r hr, h r -lemma div_add_div_same {a b c : ennreal} : a / c + b / c = (a + b) / c := +lemma div_add_div_same {a b c : ℝ≥0∞} : a / c + b / c = (a + b) / c := eq.symm $ right_distrib a b (c⁻¹) lemma div_self (h0 : a ≠ 0) (hI : a ≠ ∞) : a / a = 1 := @@ -1075,10 +1080,10 @@ begin rw mul_div_cancel' h0 hI, exact le_refl b end -lemma inv_two_add_inv_two : (2:ennreal)⁻¹ + 2⁻¹ = 1 := +lemma inv_two_add_inv_two : (2:ℝ≥0∞)⁻¹ + 2⁻¹ = 1 := by rw [← two_mul, ← div_eq_mul_inv, div_self two_ne_zero two_ne_top] -lemma add_halves (a : ennreal) : a / 2 + a / 2 = a := +lemma add_halves (a : ℝ≥0∞) : a / 2 + a / 2 = a := by rw [div_eq_mul_inv, ← mul_add, inv_two_add_inv_two, mul_one] @[simp] lemma div_zero_iff : a / b = 0 ↔ a = 0 ∨ b = ∞ := @@ -1087,15 +1092,15 @@ by simp [div_eq_mul_inv] @[simp] lemma div_pos_iff : 0 < a / b ↔ a ≠ 0 ∧ b ≠ ∞ := by simp [pos_iff_ne_zero, not_or_distrib] -lemma half_pos {a : ennreal} (h : 0 < a) : 0 < a / 2 := +lemma half_pos {a : ℝ≥0∞} (h : 0 < a) : 0 < a / 2 := by simp [ne_of_gt h] -lemma one_half_lt_one : (2⁻¹:ennreal) < 1 := inv_lt_one.2 $ one_lt_two +lemma one_half_lt_one : (2⁻¹:ℝ≥0∞) < 1 := inv_lt_one.2 $ one_lt_two -lemma half_lt_self {a : ennreal} (hz : a ≠ 0) (ht : a ≠ ∞) : a / 2 < a := +lemma half_lt_self {a : ℝ≥0∞} (hz : a ≠ 0) (ht : a ≠ ∞) : a / 2 < a := begin lift a to ℝ≥0 using ht, - have h : (2 : ennreal) = ((2 : ℝ≥0) : ennreal), from rfl, + have h : (2 : ℝ≥0∞) = ((2 : ℝ≥0) : ℝ≥0∞), from rfl, have h' : (2 : ℝ≥0) ≠ 0, from _root_.two_ne_zero', rw [h, ← coe_div h', coe_lt_coe], -- `norm_cast` fails to apply `coe_div` norm_cast at hz, @@ -1108,11 +1113,11 @@ begin exact sub_eq_of_add_eq (mul_ne_top coe_ne_top $ by simp) (add_halves a) end -lemma one_sub_inv_two : (1:ennreal) - 2⁻¹ = 2⁻¹ := +lemma one_sub_inv_two : (1:ℝ≥0∞) - 2⁻¹ = 2⁻¹ := by simpa only [div_eq_mul_inv, one_mul] using sub_half one_ne_top -lemma exists_inv_nat_lt {a : ennreal} (h : a ≠ 0) : - ∃n:ℕ, (n:ennreal)⁻¹ < a := +lemma exists_inv_nat_lt {a : ℝ≥0∞} (h : a ≠ 0) : + ∃n:ℕ, (n:ℝ≥0∞)⁻¹ < a := @inv_inv a ▸ by simp only [inv_lt_inv, ennreal.exists_nat_gt (inv_ne_top.2 h)] lemma exists_nat_pos_mul_gt (ha : a ≠ 0) (hb : b ≠ ∞) : @@ -1120,7 +1125,7 @@ lemma exists_nat_pos_mul_gt (ha : a ≠ 0) (hb : b ≠ ∞) : begin have : b / a ≠ ∞, from mul_ne_top hb (inv_ne_top.2 ha), refine (ennreal.exists_nat_gt this).imp (λ n hn, _), - have : 0 < (n : ennreal), from (zero_le _).trans_lt hn, + have : 0 < (n : ℝ≥0∞), from (zero_le _).trans_lt hn, refine ⟨coe_nat_lt_coe_nat.1 this, _⟩, rwa [← ennreal.div_lt_iff (or.inl ha) (or.inr hb)] end @@ -1130,10 +1135,10 @@ lemma exists_nat_mul_gt (ha : a ≠ 0) (hb : b ≠ ∞) : (exists_nat_pos_mul_gt ha hb).imp $ λ n, Exists.snd lemma exists_nat_pos_inv_mul_lt (ha : a ≠ ∞) (hb : b ≠ 0) : - ∃ n > 0, ((n : ℕ) : ennreal)⁻¹ * a < b := + ∃ n > 0, ((n : ℕ) : ℝ≥0∞)⁻¹ * a < b := begin rcases exists_nat_pos_mul_gt hb ha with ⟨n, npos, hn⟩, - have : (n : ennreal) ≠ 0 := nat.cast_ne_zero.2 npos.lt.ne', + have : (n : ℝ≥0∞) ≠ 0 := nat.cast_ne_zero.2 npos.lt.ne', use [n, npos], rwa [← one_mul b, ← inv_mul_cancel this coe_nat_ne_top, mul_assoc, mul_lt_mul_left (inv_ne_zero.2 coe_nat_ne_top) (inv_ne_top.2 this)] @@ -1204,7 +1209,7 @@ lemma to_real_pos_iff : 0 < a.to_real ↔ (0 < a ∧ a ≠ ∞):= lemma of_real_le_of_real {p q : ℝ} (h : p ≤ q) : ennreal.of_real p ≤ ennreal.of_real q := by simp [ennreal.of_real, nnreal.of_real_le_of_real h] -lemma of_real_le_of_le_to_real {a : ℝ} {b : ennreal} (h : a ≤ ennreal.to_real b) : +lemma of_real_le_of_le_to_real {a : ℝ} {b : ℝ≥0∞} (h : a ≤ ennreal.to_real b) : ennreal.of_real a ≤ b := (of_real_le_of_real h).trans of_real_to_real_le @@ -1226,33 +1231,33 @@ by simp [ennreal.of_real] @[simp] lemma of_real_eq_zero {p : ℝ} : ennreal.of_real p = 0 ↔ p ≤ 0 := by simp [ennreal.of_real] -lemma of_real_le_iff_le_to_real {a : ℝ} {b : ennreal} (hb : b ≠ ∞) : +lemma of_real_le_iff_le_to_real {a : ℝ} {b : ℝ≥0∞} (hb : b ≠ ∞) : ennreal.of_real a ≤ b ↔ a ≤ ennreal.to_real b := begin lift b to ℝ≥0 using hb, simpa [ennreal.of_real, ennreal.to_real] using nnreal.of_real_le_iff_le_coe end -lemma of_real_lt_iff_lt_to_real {a : ℝ} {b : ennreal} (ha : 0 ≤ a) (hb : b ≠ ∞) : +lemma of_real_lt_iff_lt_to_real {a : ℝ} {b : ℝ≥0∞} (ha : 0 ≤ a) (hb : b ≠ ∞) : ennreal.of_real a < b ↔ a < ennreal.to_real b := begin lift b to ℝ≥0 using hb, simpa [ennreal.of_real, ennreal.to_real] using nnreal.of_real_lt_iff_lt_coe ha end -lemma le_of_real_iff_to_real_le {a : ennreal} {b : ℝ} (ha : a ≠ ∞) (hb : 0 ≤ b) : +lemma le_of_real_iff_to_real_le {a : ℝ≥0∞} {b : ℝ} (ha : a ≠ ∞) (hb : 0 ≤ b) : a ≤ ennreal.of_real b ↔ ennreal.to_real a ≤ b := begin lift a to ℝ≥0 using ha, simpa [ennreal.of_real, ennreal.to_real] using nnreal.le_of_real_iff_coe_le hb end -lemma to_real_le_of_le_of_real {a : ennreal} {b : ℝ} (hb : 0 ≤ b) (h : a ≤ ennreal.of_real b) : +lemma to_real_le_of_le_of_real {a : ℝ≥0∞} {b : ℝ} (hb : 0 ≤ b) (h : a ≤ ennreal.of_real b) : ennreal.to_real a ≤ b := have ha : a ≠ ∞, from ne_top_of_le_ne_top of_real_ne_top h, (le_of_real_iff_to_real_le ha hb).1 h -lemma lt_of_real_iff_to_real_lt {a : ennreal} {b : ℝ} (ha : a ≠ ∞) : +lemma lt_of_real_iff_to_real_lt {a : ℝ≥0∞} {b : ℝ} (ha : a ≠ ∞) : a < ennreal.of_real b ↔ ennreal.to_real a < b := begin lift a to ℝ≥0 using ha, @@ -1273,7 +1278,7 @@ lemma of_real_div_of_pos {x y : ℝ} (hy : 0 < y) : by rw [div_eq_inv_mul, div_eq_mul_inv, of_real_mul (inv_nonneg.2 hy.le), of_real_inv_of_pos hy, mul_comm] -lemma to_real_of_real_mul (c : ℝ) (a : ennreal) (h : 0 ≤ c) : +lemma to_real_of_real_mul (c : ℝ) (a : ℝ≥0∞) (h : 0 ≤ c) : ennreal.to_real ((ennreal.of_real c) * a) = c * ennreal.to_real a := begin cases a, @@ -1286,20 +1291,20 @@ begin congr, apply nnreal.coe_of_real, exact h } end -@[simp] lemma to_nnreal_mul_top (a : ennreal) : ennreal.to_nnreal (a * ∞) = 0 := +@[simp] lemma to_nnreal_mul_top (a : ℝ≥0∞) : ennreal.to_nnreal (a * ∞) = 0 := begin by_cases h : a = 0, { rw [h, zero_mul, zero_to_nnreal] }, { rw [mul_top, if_neg h, top_to_nnreal] } end -@[simp] lemma to_nnreal_top_mul (a : ennreal) : ennreal.to_nnreal (∞ * a) = 0 := +@[simp] lemma to_nnreal_top_mul (a : ℝ≥0∞) : ennreal.to_nnreal (∞ * a) = 0 := by rw [mul_comm, to_nnreal_mul_top] -@[simp] lemma to_real_mul_top (a : ennreal) : ennreal.to_real (a * ∞) = 0 := +@[simp] lemma to_real_mul_top (a : ℝ≥0∞) : ennreal.to_real (a * ∞) = 0 := by rw [ennreal.to_real, to_nnreal_mul_top, nnreal.coe_zero] -@[simp] lemma to_real_top_mul (a : ennreal) : ennreal.to_real (∞ * a) = 0 := +@[simp] lemma to_real_top_mul (a : ℝ≥0∞) : ennreal.to_real (∞ * a) = 0 := by { rw mul_comm, exact to_real_mul_top _ } lemma to_real_eq_to_real (ha : a < ∞) (hb : b < ∞) : @@ -1311,33 +1316,33 @@ begin end /-- `ennreal.to_nnreal` as a `monoid_hom`. -/ -def to_nnreal_hom : ennreal →* ℝ≥0 := +def to_nnreal_hom : ℝ≥0∞ →* ℝ≥0 := { to_fun := ennreal.to_nnreal, map_one' := to_nnreal_coe, map_mul' := by rintro (_|x) (_|y); simp only [← coe_mul, none_eq_top, some_eq_coe, to_nnreal_top_mul, to_nnreal_mul_top, top_to_nnreal, mul_zero, zero_mul, to_nnreal_coe] } -lemma to_nnreal_mul {a b : ennreal}: (a * b).to_nnreal = a.to_nnreal * b.to_nnreal := +lemma to_nnreal_mul {a b : ℝ≥0∞}: (a * b).to_nnreal = a.to_nnreal * b.to_nnreal := to_nnreal_hom.map_mul a b -lemma to_nnreal_pow (a : ennreal) (n : ℕ) : (a ^ n).to_nnreal = a.to_nnreal ^ n := +lemma to_nnreal_pow (a : ℝ≥0∞) (n : ℕ) : (a ^ n).to_nnreal = a.to_nnreal ^ n := to_nnreal_hom.map_pow a n -lemma to_nnreal_prod {ι : Type*} {s : finset ι} {f : ι → ennreal} : +lemma to_nnreal_prod {ι : Type*} {s : finset ι} {f : ι → ℝ≥0∞} : (∏ i in s, f i).to_nnreal = ∏ i in s, (f i).to_nnreal := to_nnreal_hom.map_prod _ _ /-- `ennreal.to_real` as a `monoid_hom`. -/ -def to_real_hom : ennreal →* ℝ := +def to_real_hom : ℝ≥0∞ →* ℝ := (nnreal.to_real_hom : ℝ≥0 →* ℝ).comp to_nnreal_hom lemma to_real_mul : (a * b).to_real = a.to_real * b.to_real := to_real_hom.map_mul a b -lemma to_real_pow (a : ennreal) (n : ℕ) : (a ^ n).to_real = a.to_real ^ n := +lemma to_real_pow (a : ℝ≥0∞) (n : ℕ) : (a ^ n).to_real = a.to_real ^ n := to_real_hom.map_pow a n -lemma to_real_prod {ι : Type*} {s : finset ι} {f : ι → ennreal} : +lemma to_real_prod {ι : Type*} {s : finset ι} {f : ι → ℝ≥0∞} : (∏ i in s, f i).to_real = ∏ i in s, (f i).to_real := to_real_hom.map_prod _ _ @@ -1351,7 +1356,7 @@ end end real section infi -variables {ι : Sort*} {f g : ι → ennreal} +variables {ι : Sort*} {f g : ι → ℝ≥0∞} lemma infi_add : infi f + a = ⨅i, f i + a := le_antisymm @@ -1370,10 +1375,10 @@ begin simp [ennreal.sub_le_iff_le_add, sub_eq_add_neg, add_comm], end -lemma Inf_add {s : set ennreal} : Inf s + a = ⨅b∈s, b + a := +lemma Inf_add {s : set ℝ≥0∞} : Inf s + a = ⨅b∈s, b + a := by simp [Inf_eq_infi, infi_add] -lemma add_infi {a : ennreal} : a + infi f = ⨅b, a + f b := +lemma add_infi {a : ℝ≥0∞} : a + infi f = ⨅b, a + f b := by rw [add_comm, infi_add]; simp [add_comm] lemma infi_add_infi (h : ∀i j, ∃k, f k + g k ≤ f i + g j) : infi f + infi g = (⨅a, f a + g a) := @@ -1385,7 +1390,7 @@ calc (⨅a, f a + g a) ≤ (⨅ a a', f a + g a') : ... ≤ infi f + infi g : by simp [add_infi, infi_add, -add_comm, -le_infi_iff]; exact le_refl _ -lemma infi_sum {f : ι → α → ennreal} {s : finset α} [nonempty ι] +lemma infi_sum {f : ι → α → ℝ≥0∞} {s : finset α} [nonempty ι] (h : ∀(t : finset α) (i j : ι), ∃k, ∀a∈t, f k a ≤ f i a ∧ f k a ≤ f j a) : (⨅i, ∑ a in s, f i a) = ∑ a in s, ⨅i, f i a := finset.induction_on s (by simp) $ assume a s ha ih, @@ -1397,7 +1402,7 @@ finset.induction_on s (by simp) $ assume a s ha ih, by simp [ha, ih.symm, infi_add_infi this] -lemma infi_mul {ι} [nonempty ι] {f : ι → ennreal} {x : ennreal} (h : x ≠ ∞) : +lemma infi_mul {ι} [nonempty ι] {f : ι → ℝ≥0∞} {x : ℝ≥0∞} (h : x ≠ ∞) : infi f * x = ⨅i, f i * x := begin by_cases h2 : x = 0, simp only [h2, mul_zero, infi_const], @@ -1407,7 +1412,7 @@ begin λ i, (div_le_iff_le_mul (or.inl h2) $ or.inl h).mpr $ infi_le _ _) end -lemma mul_infi {ι} [nonempty ι] {f : ι → ennreal} {x : ennreal} (h : x ≠ ∞) : +lemma mul_infi {ι} [nonempty ι] {f : ι → ℝ≥0∞} {x : ℝ≥0∞} (h : x ≠ ∞) : x * infi f = ⨅i, x * f i := by { rw [mul_comm, infi_mul h], simp only [mul_comm], assumption } @@ -1417,14 +1422,14 @@ end infi section supr -lemma supr_coe_nat : (⨆n:ℕ, (n : ennreal)) = ∞ := +lemma supr_coe_nat : (⨆n:ℕ, (n : ℝ≥0∞)) = ∞ := (supr_eq_top _).2 $ assume b hb, ennreal.exists_nat_gt (lt_top_iff_ne_top.1 hb) end supr /-- `le_of_add_le_add_left` is normally applicable to `ordered_cancel_add_comm_monoid`, -but it holds in `ennreal` with the additional assumption that `a < ∞`. -/ -lemma le_of_add_le_add_left {a b c : ennreal} : a < ∞ → +but it holds in `ℝ≥0∞` with the additional assumption that `a < ∞`. -/ +lemma le_of_add_le_add_left {a b c : ℝ≥0∞} : a < ∞ → a + b ≤ a + c → b ≤ c := by cases a; cases b; cases c; simp [← ennreal.coe_add, ennreal.coe_le_coe] diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index 44402e60b7b08..c5590e5b6981c 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -25,7 +25,7 @@ real, ereal, complete lattice ## TODO -abs : ereal → ennreal +abs : ereal → ℝ≥0∞ In Isabelle they define + - * and / (making junk choices for things like -∞ + ∞) and then prove whatever bits of the ordered ring/field axioms still hold. They diff --git a/src/measure_theory/ae_eq_fun.lean b/src/measure_theory/ae_eq_fun.lean index c547df92456bf..92afac9e16350 100644 --- a/src/measure_theory/ae_eq_fun.lean +++ b/src/measure_theory/ae_eq_fun.lean @@ -46,7 +46,7 @@ See `l1_space.lean` for `L¹` space. If `β` is an `emetric_space`, then `L⁰` can be made into an `emetric_space`, where `edist [f] [g]` is defined to be `∫⁻ a, edist (f a) (g a)`. - The integral used here is `lintegral : (α → ennreal) → ennreal`, which is defined in the file + The integral used here is `lintegral : (α → ℝ≥0∞) → ℝ≥0∞`, which is defined in the file `integration.lean`. See `edist_mk_mk` and `edist_to_fun`. @@ -70,7 +70,7 @@ function space, almost everywhere equal, `L⁰`, ae_eq_fun -/ noncomputable theory -open_locale classical +open_locale classical ennreal open set filter topological_space ennreal emetric measure_theory function variables {α β γ δ : Type*} [measurable_space α] {μ ν : measure α} @@ -405,32 +405,32 @@ end semimodule open ennreal -/-- For `f : α → ennreal`, define `∫ [f]` to be `∫ f` -/ -def lintegral (f : α →ₘ[μ] ennreal) : ennreal := -quotient.lift_on' f (λf, ∫⁻ a, (f : α → ennreal) a ∂μ) (assume f g, lintegral_congr_ae) +/-- For `f : α → ℝ≥0∞`, define `∫ [f]` to be `∫ f` -/ +def lintegral (f : α →ₘ[μ] ℝ≥0∞) : ℝ≥0∞ := +quotient.lift_on' f (λf, ∫⁻ a, (f : α → ℝ≥0∞) a ∂μ) (assume f g, lintegral_congr_ae) -@[simp] lemma lintegral_mk (f : α → ennreal) (hf) : - (mk f hf : α →ₘ[μ] ennreal).lintegral = ∫⁻ a, f a ∂μ := rfl +@[simp] lemma lintegral_mk (f : α → ℝ≥0∞) (hf) : + (mk f hf : α →ₘ[μ] ℝ≥0∞).lintegral = ∫⁻ a, f a ∂μ := rfl -lemma lintegral_coe_fn (f : α →ₘ[μ] ennreal) : ∫⁻ a, f a ∂μ = f.lintegral := +lemma lintegral_coe_fn (f : α →ₘ[μ] ℝ≥0∞) : ∫⁻ a, f a ∂μ = f.lintegral := by rw [← lintegral_mk, mk_coe_fn] -@[simp] lemma lintegral_zero : lintegral (0 : α →ₘ[μ] ennreal) = 0 := lintegral_zero +@[simp] lemma lintegral_zero : lintegral (0 : α →ₘ[μ] ℝ≥0∞) = 0 := lintegral_zero -@[simp] lemma lintegral_eq_zero_iff {f : α →ₘ[μ] ennreal} : lintegral f = 0 ↔ f = 0 := +@[simp] lemma lintegral_eq_zero_iff {f : α →ₘ[μ] ℝ≥0∞} : lintegral f = 0 ↔ f = 0 := induction_on f $ λ f hf, (lintegral_eq_zero_iff' hf).trans mk_eq_mk.symm -lemma lintegral_add (f g : α →ₘ[μ] ennreal) : lintegral (f + g) = lintegral f + lintegral g := +lemma lintegral_add (f g : α →ₘ[μ] ℝ≥0∞) : lintegral (f + g) = lintegral f + lintegral g := induction_on₂ f g $ λ f hf g hg, by simp [lintegral_add' hf hg] -lemma lintegral_mono {f g : α →ₘ[μ] ennreal} : f ≤ g → lintegral f ≤ lintegral g := +lemma lintegral_mono {f g : α →ₘ[μ] ℝ≥0∞} : f ≤ g → lintegral f ≤ lintegral g := induction_on₂ f g $ λ f hf g hg hfg, lintegral_mono_ae hfg section variables [emetric_space γ] [second_countable_topology γ] [opens_measurable_space γ] /-- `comp_edist [f] [g] a` will return `edist (f a) (g a)` -/ -protected def edist (f g : α →ₘ[μ] γ) : α →ₘ[μ] ennreal := comp₂ edist measurable_edist f g +protected def edist (f g : α →ₘ[μ] γ) : α →ₘ[μ] ℝ≥0∞ := comp₂ edist measurable_edist f g protected lemma edist_comm (f g : α →ₘ[μ] γ) : f.edist g = g.edist f := induction_on₂ f g $ λ f hf g hg, mk_eq_mk.2 $ eventually_of_forall $ λ x, edist_comm (f x) (g x) diff --git a/src/measure_theory/bochner_integration.lean b/src/measure_theory/bochner_integration.lean index f9fab5870524f..efaba24ee2750 100644 --- a/src/measure_theory/bochner_integration.lean +++ b/src/measure_theory/bochner_integration.lean @@ -22,7 +22,7 @@ The Bochner integral is defined following these steps: where `E` is a real normed space. (See `simple_func.bintegral` and section `bintegral` for details. Also see `simple_func.integral` - for the integral on simple functions of the type `simple_func α ennreal`.) + for the integral on simple functions of the type `simple_func α ℝ≥0∞`.) 2. Use `α →ₛ E` to cut out the simple functions from L1 functions, and define integral on these. The type of simple functions in L1 space is written as `α →₁ₛ[μ] E`. @@ -58,7 +58,7 @@ The Bochner integral is defined following these steps: * `integral_nonpos` : `f ≤ 0 → ∫ x, f x ∂μ ≤ 0` * `integral_mono` : `f ≤ᵐ[μ] g → ∫ x, f x ∂μ ≤ ∫ x, g x ∂μ` -3. Propositions connecting the Bochner integral with the integral on `ennreal`-valued functions, +3. Propositions connecting the Bochner integral with the integral on `ℝ≥0∞`-valued functions, which is called `lintegral` and has the notation `∫⁻`. * `integral_eq_lintegral_max_sub_lintegral_min` : `∫ x, f x ∂μ = ∫⁻ x, f⁺ x ∂μ - ∫⁻ x, f⁻ x ∂μ`, @@ -78,7 +78,7 @@ you to prove something for an arbitrary measurable + integrable function. Another method is using the following steps. See `integral_eq_lintegral_max_sub_lintegral_min` for a complicated example, which proves that `∫ f = ∫⁻ f⁺ - ∫⁻ f⁻`, with the first integral sign being the Bochner integral of a real-valued -function `f : α → ℝ`, and second and third integral sign being the integral on ennreal-valued +function `f : α → ℝ`, and second and third integral sign being the integral on `ℝ≥0∞`-valued functions (called `lintegral`). The proof of `integral_eq_lintegral_max_sub_lintegral_min` is scattered in sections with the name `pos_part`. @@ -123,7 +123,7 @@ Bochner integral, simple function, function space, Lebesgue dominated convergenc -/ noncomputable theory -open_locale classical topological_space big_operators nnreal +open_locale classical topological_space big_operators nnreal ennreal namespace measure_theory @@ -191,10 +191,10 @@ variables {μ : measure α} finite volume support. -/ lemma integrable_iff_fin_meas_supp {f : α →ₛ E} {μ : measure α} : integrable f μ ↔ f.fin_meas_supp μ := -calc integrable f μ ↔ ∫⁻ x, f.map (coe ∘ nnnorm : E → ennreal) x ∂μ < ⊤ : +calc integrable f μ ↔ ∫⁻ x, f.map (coe ∘ nnnorm : E → ℝ≥0∞) x ∂μ < ⊤ : and_iff_right f.ae_measurable -... ↔ (f.map (coe ∘ nnnorm : E → ennreal)).lintegral μ < ⊤ : by rw lintegral_eq_lintegral -... ↔ (f.map (coe ∘ nnnorm : E → ennreal)).fin_meas_supp μ : iff.symm $ +... ↔ (f.map (coe ∘ nnnorm : E → ℝ≥0∞)).lintegral μ < ⊤ : by rw lintegral_eq_lintegral +... ↔ (f.map (coe ∘ nnnorm : E → ℝ≥0∞)).fin_meas_supp μ : iff.symm $ fin_meas_supp.iff_lintegral_lt_top $ eventually_of_forall $ λ x, coe_lt_top ... ↔ _ : fin_meas_supp.map_iff $ λ b, coe_eq_zero.trans nnnorm_eq_zero @@ -253,9 +253,9 @@ begin end /-- `simple_func.integral` and `simple_func.lintegral` agree when the integrand has type - `α →ₛ ennreal`. But since `ennreal` is not a `normed_space`, we need some form of coercion. + `α →ₛ ℝ≥0∞`. But since `ℝ≥0∞` is not a `normed_space`, we need some form of coercion. See `integral_eq_lintegral` for a simpler version. -/ -lemma integral_eq_lintegral' {f : α →ₛ E} {g : E → ennreal} (hf : integrable f μ) (hg0 : g 0 = 0) +lemma integral_eq_lintegral' {f : α →ₛ E} {g : E → ℝ≥0∞} (hf : integrable f μ) (hg0 : g 0 = 0) (hgt : ∀b, g b < ⊤): (f.map (ennreal.to_real ∘ g)).integral μ = ennreal.to_real (∫⁻ a, g (f a) ∂μ) := begin @@ -292,7 +292,7 @@ begin end /-- `simple_func.bintegral` and `simple_func.integral` agree when the integrand has type - `α →ₛ ennreal`. But since `ennreal` is not a `normed_space`, we need some form of coercion. -/ + `α →ₛ ℝ≥0∞`. But since `ℝ≥0∞` is not a `normed_space`, we need some form of coercion. -/ lemma integral_eq_lintegral {f : α →ₛ ℝ} (hf : integrable f μ) (h_pos : 0 ≤ᵐ[μ] f) : f.integral μ = ennreal.to_real (∫⁻ a, ennreal.of_real (f a) ∂μ) := begin @@ -1071,7 +1071,7 @@ begin end lemma ennnorm_integral_le_lintegral_ennnorm (f : α → E) : - (nnnorm (∫ a, f a ∂μ) : ennreal) ≤ ∫⁻ a, (nnnorm (f a)) ∂μ := + (nnnorm (∫ a, f a ∂μ) : ℝ≥0∞) ≤ ∫⁻ a, (nnnorm (f a)) ∂μ := by { simp_rw [← of_real_norm_eq_coe_nnnorm], apply ennreal.of_real_le_of_le_to_real, exact norm_integral_le_lintegral_norm f } @@ -1255,7 +1255,7 @@ begin rw [← lt_top_iff_ne_top], convert hfi.has_finite_integral, ext1 x, rw [real.nnnorm_coe_eq_self] end -lemma integral_to_real {f : α → ennreal} (hfm : ae_measurable f μ) (hf : ∀ᵐ x ∂μ, f x < ⊤) : +lemma integral_to_real {f : α → ℝ≥0∞} (hfm : ae_measurable f μ) (hf : ∀ᵐ x ∂μ, f x < ⊤) : ∫ a, (f a).to_real ∂μ = (∫⁻ a, f a ∂μ).to_real := begin rw [integral_eq_lintegral_of_nonneg_ae _ hfm.to_real], @@ -1448,7 +1448,7 @@ end @[simp] lemma integral_zero_measure (f : α → E) : ∫ x, f x ∂0 = 0 := norm_le_zero_iff.1 $ le_trans (norm_integral_le_lintegral_norm f) $ by simp -private lemma integral_smul_measure_aux {f : α → E} {c : ennreal} +private lemma integral_smul_measure_aux {f : α → E} {c : ℝ≥0∞} (h0 : 0 < c) (hc : c < ⊤) (fmeas : measurable f) (hfi : integrable f μ) : ∫ x, f x ∂(c • μ) = c.to_real • ∫ x, f x ∂μ := begin @@ -1459,7 +1459,7 @@ begin ennreal.to_real_mul] end -@[simp] lemma integral_smul_measure (f : α → E) (c : ennreal) : +@[simp] lemma integral_smul_measure (f : α → E) (c : ℝ≥0∞) : ∫ x, f x ∂(c • μ) = c.to_real • ∫ x, f x ∂μ := begin -- First we consider “degenerate” cases: diff --git a/src/measure_theory/borel_space.lean b/src/measure_theory/borel_space.lean index 559cd033096ea..fba733a162b87 100644 --- a/src/measure_theory/borel_space.lean +++ b/src/measure_theory/borel_space.lean @@ -20,7 +20,7 @@ import topology.G_delta * `class opens_measurable_space` : a space with `topological_space` and `measurable_space` structures such that all open sets are measurable; equivalently, `borel α ≤ ‹measurable_space α›`. * `borel_space` instances on `empty`, `unit`, `bool`, `nat`, `int`, `rat`; -* `measurable` and `borel_space` instances on `ℝ`, `ℝ≥0`, `ennreal`. +* `measurable` and `borel_space` instances on `ℝ`, `ℝ≥0`, `ℝ≥0∞`. * A measure is `regular` if it is finite on compact sets, inner regular and outer regular. ## Main statements @@ -32,13 +32,13 @@ import topology.G_delta * `measurable.add` etc : dot notation for arithmetic operations on `measurable` predicates, and similarly for `dist` and `edist`; * `ae_measurable.add` : similar dot notation for almost everywhere measurable functions; -* `measurable.ennreal*` : special cases for arithmetic operations on `ennreal`s. +* `measurable.ennreal*` : special cases for arithmetic operations on `ℝ≥0∞`. -/ noncomputable theory open classical set filter measure_theory -open_locale classical big_operators topological_space nnreal +open_locale classical big_operators topological_space nnreal ennreal universes u v w x y variables {α β γ γ₂ δ : Type*} {ι : Sort y} {s t u : set α} @@ -917,8 +917,8 @@ instance real.borel_space : borel_space ℝ := ⟨rfl⟩ instance nnreal.measurable_space : measurable_space ℝ≥0 := borel ℝ≥0 instance nnreal.borel_space : borel_space ℝ≥0 := ⟨rfl⟩ -instance ennreal.measurable_space : measurable_space ennreal := borel ennreal -instance ennreal.borel_space : borel_space ennreal := ⟨rfl⟩ +instance ennreal.measurable_space : measurable_space ℝ≥0∞ := borel ℝ≥0∞ +instance ennreal.borel_space : borel_space ℝ≥0∞ := ⟨rfl⟩ instance complex.measurable_space : measurable_space ℂ := borel ℂ instance complex.borel_space : borel_space ℂ := ⟨rfl⟩ @@ -970,7 +970,7 @@ end metric_space section emetric_space variables [emetric_space α] [measurable_space α] [opens_measurable_space α] -variables [measurable_space β] {x : α} {ε : ennreal} +variables [measurable_space β] {x : α} {ε : ℝ≥0∞} open emetric @@ -1077,50 +1077,50 @@ lemma measurable.nnreal_coe {f : α → ℝ≥0} (hf : measurable f) : nnreal.measurable_coe.comp hf lemma measurable.ennreal_coe {f : α → ℝ≥0} (hf : measurable f) : - measurable (λ x, (f x : ennreal)) := + measurable (λ x, (f x : ℝ≥0∞)) := ennreal.continuous_coe.measurable.comp hf lemma ae_measurable.ennreal_coe {f : α → ℝ≥0} {μ : measure α} (hf : ae_measurable f μ) : - ae_measurable (λ x, (f x : ennreal)) μ := + ae_measurable (λ x, (f x : ℝ≥0∞)) μ := ennreal.continuous_coe.measurable.comp_ae_measurable hf lemma measurable.ennreal_of_real {f : α → ℝ} (hf : measurable f) : measurable (λ x, ennreal.of_real (f x)) := ennreal.continuous_of_real.measurable.comp hf -/-- The set of finite `ennreal` numbers is `measurable_equiv` to `ℝ≥0`. -/ -def measurable_equiv.ennreal_equiv_nnreal : {r : ennreal | r ≠ ⊤} ≃ᵐ ℝ≥0 := +/-- The set of finite `ℝ≥0∞` numbers is `measurable_equiv` to `ℝ≥0`. -/ +def measurable_equiv.ennreal_equiv_nnreal : {r : ℝ≥0∞ | r ≠ ⊤} ≃ᵐ ℝ≥0 := ennreal.ne_top_homeomorph_nnreal.to_measurable_equiv namespace ennreal -lemma measurable_coe : measurable (coe : ℝ≥0 → ennreal) := +lemma measurable_coe : measurable (coe : ℝ≥0 → ℝ≥0∞) := measurable_id.ennreal_coe -lemma measurable_of_measurable_nnreal {f : ennreal → α} +lemma measurable_of_measurable_nnreal {f : ℝ≥0∞ → α} (h : measurable (λ p : ℝ≥0, f p)) : measurable f := measurable_of_measurable_on_compl_singleton ⊤ (measurable_equiv.ennreal_equiv_nnreal.symm.measurable_coe_iff.1 h) -/-- `ennreal` is `measurable_equiv` to `ℝ≥0 ⊕ unit`. -/ -def ennreal_equiv_sum : ennreal ≃ᵐ ℝ≥0 ⊕ unit := +/-- `ℝ≥0∞` is `measurable_equiv` to `ℝ≥0 ⊕ unit`. -/ +def ennreal_equiv_sum : ℝ≥0∞ ≃ᵐ ℝ≥0 ⊕ unit := { measurable_to_fun := measurable_of_measurable_nnreal measurable_inl, - measurable_inv_fun := measurable_sum measurable_coe (@measurable_const ennreal unit _ _ ⊤), + measurable_inv_fun := measurable_sum measurable_coe (@measurable_const ℝ≥0∞ unit _ _ ⊤), .. equiv.option_equiv_sum_punit ℝ≥0 } open function (uncurry) lemma measurable_of_measurable_nnreal_prod [measurable_space β] [measurable_space γ] - {f : ennreal × β → γ} (H₁ : measurable (λ p : ℝ≥0 × β, f (p.1, p.2))) + {f : ℝ≥0∞ × β → γ} (H₁ : measurable (λ p : ℝ≥0 × β, f (p.1, p.2))) (H₂ : measurable (λ x, f (⊤, x))) : measurable f := -let e : ennreal × β ≃ᵐ ℝ≥0 × β ⊕ unit × β := +let e : ℝ≥0∞ × β ≃ᵐ ℝ≥0 × β ⊕ unit × β := (ennreal_equiv_sum.prod_congr (measurable_equiv.refl β)).trans (measurable_equiv.sum_prod_distrib _ _ _) in e.symm.measurable_coe_iff.1 $ measurable_sum H₁ (H₂.comp measurable_id.snd) lemma measurable_of_measurable_nnreal_nnreal [measurable_space β] - {f : ennreal × ennreal → β} (h₁ : measurable (λ p : ℝ≥0 × ℝ≥0, f (p.1, p.2))) + {f : ℝ≥0∞ × ℝ≥0∞ → β} (h₁ : measurable (λ p : ℝ≥0 × ℝ≥0, f (p.1, p.2))) (h₂ : measurable (λ r : ℝ≥0, f (⊤, r))) (h₃ : measurable (λ r : ℝ≥0, f (r, ⊤))) : measurable f := measurable_of_measurable_nnreal_prod @@ -1136,7 +1136,7 @@ ennreal.measurable_of_measurable_nnreal nnreal.measurable_coe lemma measurable_to_nnreal : measurable ennreal.to_nnreal := ennreal.measurable_of_measurable_nnreal measurable_id -lemma measurable_mul : measurable (λ p : ennreal × ennreal, p.1 * p.2) := +lemma measurable_mul : measurable (λ p : ℝ≥0∞ × ℝ≥0∞, p.1 * p.2) := begin apply measurable_of_measurable_nnreal_nnreal, { simp only [← ennreal.coe_mul, measurable_mul.ennreal_coe] }, @@ -1146,55 +1146,55 @@ begin exact measurable_const.piecewise (measurable_set_singleton _) measurable_const } end -lemma measurable_sub : measurable (λ p : ennreal × ennreal, p.1 - p.2) := +lemma measurable_sub : measurable (λ p : ℝ≥0∞ × ℝ≥0∞, p.1 - p.2) := by apply measurable_of_measurable_nnreal_nnreal; simp [← ennreal.coe_sub, continuous_sub.measurable.ennreal_coe] -lemma measurable_inv : measurable (has_inv.inv : ennreal → ennreal) := +lemma measurable_inv : measurable (has_inv.inv : ℝ≥0∞ → ℝ≥0∞) := ennreal.continuous_inv.measurable -lemma measurable_div : measurable (λ p : ennreal × ennreal, p.1 / p.2) := +lemma measurable_div : measurable (λ p : ℝ≥0∞ × ℝ≥0∞, p.1 / p.2) := ennreal.measurable_mul.comp $ measurable_fst.prod_mk $ ennreal.measurable_inv.comp measurable_snd end ennreal -lemma measurable.to_nnreal {f : α → ennreal} (hf : measurable f) : +lemma measurable.to_nnreal {f : α → ℝ≥0∞} (hf : measurable f) : measurable (λ x, (f x).to_nnreal) := ennreal.measurable_to_nnreal.comp hf lemma measurable_ennreal_coe_iff {f : α → ℝ≥0} : - measurable (λ x, (f x : ennreal)) ↔ measurable f := + measurable (λ x, (f x : ℝ≥0∞)) ↔ measurable f := ⟨λ h, h.to_nnreal, λ h, h.ennreal_coe⟩ -lemma measurable.to_real {f : α → ennreal} (hf : measurable f) : +lemma measurable.to_real {f : α → ℝ≥0∞} (hf : measurable f) : measurable (λ x, ennreal.to_real (f x)) := ennreal.measurable_to_real.comp hf -lemma ae_measurable.to_real {f : α → ennreal} {μ : measure α} (hf : ae_measurable f μ) : +lemma ae_measurable.to_real {f : α → ℝ≥0∞} {μ : measure α} (hf : ae_measurable f μ) : ae_measurable (λ x, ennreal.to_real (f x)) μ := ennreal.measurable_to_real.comp_ae_measurable hf -lemma measurable.ennreal_mul {f g : α → ennreal} (hf : measurable f) (hg : measurable g) : +lemma measurable.ennreal_mul {f g : α → ℝ≥0∞} (hf : measurable f) (hg : measurable g) : measurable (λ a, f a * g a) := ennreal.measurable_mul.comp (hf.prod_mk hg) -lemma ae_measurable.ennreal_mul {f g : α → ennreal} {μ : measure α} +lemma ae_measurable.ennreal_mul {f g : α → ℝ≥0∞} {μ : measure α} (hf : ae_measurable f μ) (hg : ae_measurable g μ) : ae_measurable (λ a, f a * g a) μ := ennreal.measurable_mul.comp_ae_measurable (hf.prod_mk hg) -lemma measurable.ennreal_sub {f g : α → ennreal} (hf : measurable f) (hg : measurable g) : +lemma measurable.ennreal_sub {f g : α → ℝ≥0∞} (hf : measurable f) (hg : measurable g) : measurable (λ a, f a - g a) := ennreal.measurable_sub.comp (hf.prod_mk hg) -/-- note: `ennreal` can probably be generalized in a future version of this lemma. -/ -lemma measurable.ennreal_tsum {ι} [encodable ι] {f : ι → α → ennreal} (h : ∀ i, measurable (f i)) : +/-- note: `ℝ≥0∞` can probably be generalized in a future version of this lemma. -/ +lemma measurable.ennreal_tsum {ι} [encodable ι] {f : ι → α → ℝ≥0∞} (h : ∀ i, measurable (f i)) : measurable (λ x, ∑' i, f i x) := by { simp_rw [ennreal.tsum_eq_supr_sum], apply measurable_supr, exact λ s, s.measurable_sum h } -lemma measurable.ennreal_inv {f : α → ennreal} (hf : measurable f) : measurable (λ a, (f a)⁻¹) := +lemma measurable.ennreal_inv {f : α → ℝ≥0∞} (hf : measurable f) : measurable (λ a, (f a)⁻¹) := ennreal.measurable_inv.comp hf -lemma measurable.ennreal_div {f g : α → ennreal} (hf : measurable f) (hg : measurable g) : +lemma measurable.ennreal_div {f g : α → ℝ≥0∞} (hf : measurable f) (hg : measurable g) : measurable (λ a, f a / g a) := ennreal.measurable_div.comp $ hf.prod_mk hg @@ -1222,15 +1222,15 @@ lemma ae_measurable.nnnorm {f : β → α} {μ : measure β} (hf : ae_measurable ae_measurable (λ a, nnnorm (f a)) μ := measurable_nnnorm.comp_ae_measurable hf -lemma measurable_ennnorm : measurable (λ x : α, (nnnorm x : ennreal)) := +lemma measurable_ennnorm : measurable (λ x : α, (nnnorm x : ℝ≥0∞)) := measurable_nnnorm.ennreal_coe lemma measurable.ennnorm {f : β → α} (hf : measurable f) : - measurable (λ a, (nnnorm (f a) : ennreal)) := + measurable (λ a, (nnnorm (f a) : ℝ≥0∞)) := hf.nnnorm.ennreal_coe lemma ae_measurable.ennnorm {f : β → α} {μ : measure β} (hf : ae_measurable f μ) : - ae_measurable (λ a, (nnnorm (f a) : ennreal)) μ := + ae_measurable (λ a, (nnnorm (f a) : ℝ≥0∞)) μ := measurable_ennnorm.comp_ae_measurable hf end normed_group @@ -1249,10 +1249,10 @@ lemma measurable_of_tendsto_nnreal' {ι ι'} {f : ι → α → ℝ≥0} {g : α {s : ι' → set ι} (hu : u.has_countable_basis p s) (hs : ∀ i, (s i).countable) : measurable g := begin rw [tendsto_pi] at lim, rw [← measurable_ennreal_coe_iff], - have : ∀ x, liminf u (λ n, (f n x : ennreal)) = (g x : ennreal) := + have : ∀ x, liminf u (λ n, (f n x : ℝ≥0∞)) = (g x : ℝ≥0∞) := λ x, ((ennreal.continuous_coe.tendsto (g x)).comp (lim x)).liminf_eq, simp_rw [← this], - show measurable (λ x, liminf u (λ n, (f n x : ennreal))), + show measurable (λ x, liminf u (λ n, (f n x : ℝ≥0∞))), exact measurable_liminf' (λ i, (hf i).ennreal_coe) hu hs, end @@ -1429,7 +1429,7 @@ begin rw [map_apply hf hK.measurable_set] } end -protected lemma smul (hμ : μ.regular) {x : ennreal} (hx : x < ⊤) : +protected lemma smul (hμ : μ.regular) {x : ℝ≥0∞} (hx : x < ⊤) : (x • μ).regular := begin split, diff --git a/src/measure_theory/category/Meas.lean b/src/measure_theory/category/Meas.lean index 7b410c8c9ff47..e9c09585ebeaa 100644 --- a/src/measure_theory/category/Meas.lean +++ b/src/measure_theory/category/Meas.lean @@ -27,6 +27,7 @@ measurable space, giry monad, borel noncomputable theory open category_theory measure_theory +open_locale ennreal universes u v /-- The category of measurable spaces and measurable functions. -/ @@ -81,11 +82,11 @@ instance : category_theory.monad Measure.{u} := /-- An example for an algebra on `Measure`: the nonnegative Lebesgue integral is a hom, behaving nicely under the monad operations. -/ def Integral : monad.algebra Measure := -{ A := Meas.of ennreal , - a := ⟨λm:measure ennreal, ∫⁻ x, x ∂m, measure.measurable_lintegral measurable_id ⟩, - unit' := subtype.eq $ funext $ assume r:ennreal, lintegral_dirac' _ measurable_id, - assoc' := subtype.eq $ funext $ assume μ : measure (measure ennreal), - show ∫⁻ x, x ∂ μ.join = ∫⁻ x, x ∂ (measure.map (λm:measure ennreal, ∫⁻ x, x ∂m) μ), +{ A := Meas.of ℝ≥0∞ , + a := ⟨λm:measure ℝ≥0∞, ∫⁻ x, x ∂m, measure.measurable_lintegral measurable_id ⟩, + unit' := subtype.eq $ funext $ assume r:ℝ≥0∞, lintegral_dirac' _ measurable_id, + assoc' := subtype.eq $ funext $ assume μ : measure (measure ℝ≥0∞), + show ∫⁻ x, x ∂ μ.join = ∫⁻ x, x ∂ (measure.map (λm:measure ℝ≥0∞, ∫⁻ x, x ∂m) μ), by rw [measure.lintegral_join, lintegral_map]; apply_rules [measurable_id, measure.measurable_lintegral] } diff --git a/src/measure_theory/content.lean b/src/measure_theory/content.lean index 15348e7871cdc..d4e9a25c2f862 100644 --- a/src/measure_theory/content.lean +++ b/src/measure_theory/content.lean @@ -12,7 +12,7 @@ import topology.compacts # Contents In this file we work with *contents*. A content `λ` is a function from a certain class of subsets -(such as the the compact subsets) to `ennreal` (or `ℝ≥0`) that is +(such as the the compact subsets) to `ℝ≥0∞` (or `ℝ≥0`) that is * additive: If `K₁` and `K₂` are disjoint sets in the domain of `λ`, then `λ(K₁ ∪ K₂) = λ(K₁) + λ(K₂)`; * subadditive: If `K₁` and `K₂` are in the domain of `λ`, then `λ(K₁ ∪ K₂) ≤ λ(K₁) + λ(K₂)`; @@ -42,7 +42,7 @@ universe variables u v w noncomputable theory open set topological_space -open_locale nnreal +open_locale nnreal ennreal namespace measure_theory @@ -51,26 +51,26 @@ variables {G : Type w} [topological_space G] /-- Constructing the inner content of a content. From a content defined on the compact sets, we obtain a function defined on all open sets, by taking the supremum of the content of all compact subsets. -/ -def inner_content (μ : compacts G → ennreal) (U : opens G) : ennreal := +def inner_content (μ : compacts G → ℝ≥0∞) (U : opens G) : ℝ≥0∞ := ⨆ (K : compacts G) (h : K.1 ⊆ U), μ K -lemma le_inner_content {μ : compacts G → ennreal} (K : compacts G) (U : opens G) +lemma le_inner_content {μ : compacts G → ℝ≥0∞} (K : compacts G) (U : opens G) (h2 : K.1 ⊆ U) : μ K ≤ inner_content μ U := le_supr_of_le K $ le_supr _ h2 -lemma inner_content_le {μ : compacts G → ennreal} +lemma inner_content_le {μ : compacts G → ℝ≥0∞} (h : ∀ (K₁ K₂ : compacts G), K₁.1 ⊆ K₂.1 → μ K₁ ≤ μ K₂) (U : opens G) (K : compacts G) (h2 : (U : set G) ⊆ K.1) : inner_content μ U ≤ μ K := bsupr_le $ λ K' hK', h _ _ (subset.trans hK' h2) -lemma inner_content_of_is_compact {μ : compacts G → ennreal} +lemma inner_content_of_is_compact {μ : compacts G → ℝ≥0∞} (h : ∀ (K₁ K₂ : compacts G), K₁.1 ⊆ K₂.1 → μ K₁ ≤ μ K₂) {K : set G} (h1K : is_compact K) (h2K : is_open K) : inner_content μ ⟨K, h2K⟩ = μ ⟨K, h1K⟩ := le_antisymm (bsupr_le $ λ K' hK', h _ ⟨K, h1K⟩ hK') (le_inner_content _ _ subset.rfl) -lemma inner_content_empty {μ : compacts G → ennreal} (h : μ ⊥ = 0) : +lemma inner_content_empty {μ : compacts G → ℝ≥0∞} (h : μ ⊥ = 0) : inner_content μ ∅ = 0 := begin refine le_antisymm _ (zero_le _), rw ←h, @@ -79,11 +79,11 @@ begin end /-- This is "unbundled", because that it required for the API of `induced_outer_measure`. -/ -lemma inner_content_mono {μ : compacts G → ennreal} ⦃U V : set G⦄ (hU : is_open U) (hV : is_open V) +lemma inner_content_mono {μ : compacts G → ℝ≥0∞} ⦃U V : set G⦄ (hU : is_open U) (hV : is_open V) (h2 : U ⊆ V) : inner_content μ ⟨U, hU⟩ ≤ inner_content μ ⟨V, hV⟩ := supr_le_supr $ λ K, supr_le_supr_const $ λ hK, subset.trans hK h2 -lemma inner_content_exists_compact {μ : compacts G → ennreal} {U : opens G} +lemma inner_content_exists_compact {μ : compacts G → ℝ≥0∞} {U : opens G} (hU : inner_content μ U < ⊤) {ε : ℝ≥0} (hε : 0 < ε) : ∃ K : compacts G, K.1 ⊆ U ∧ inner_content μ U ≤ μ K + ε := begin @@ -98,7 +98,7 @@ end /-- The inner content of a supremum of opens is at most the sum of the individual inner contents. -/ -lemma inner_content_Sup_nat [t2_space G] {μ : compacts G → ennreal} +lemma inner_content_Sup_nat [t2_space G] {μ : compacts G → ℝ≥0∞} (h1 : μ ⊥ = 0) (h2 : ∀ (K₁ K₂ : compacts G), μ (K₁ ⊔ K₂) ≤ μ K₁ + μ K₂) (U : ℕ → opens G) : inner_content μ (⨆ (i : ℕ), U i) ≤ ∑' (i : ℕ), inner_content μ (U i) := @@ -124,14 +124,14 @@ end /-- The inner content of a union of sets is at most the sum of the individual inner contents. This is the "unbundled" version of `inner_content_Sup_nat`. It required for the API of `induced_outer_measure`. -/ -lemma inner_content_Union_nat [t2_space G] {μ : compacts G → ennreal} +lemma inner_content_Union_nat [t2_space G] {μ : compacts G → ℝ≥0∞} (h1 : μ ⊥ = 0) (h2 : ∀ (K₁ K₂ : compacts G), μ (K₁ ⊔ K₂) ≤ μ K₁ + μ K₂) ⦃U : ℕ → set G⦄ (hU : ∀ (i : ℕ), is_open (U i)) : inner_content μ ⟨⋃ (i : ℕ), U i, is_open_Union hU⟩ ≤ ∑' (i : ℕ), inner_content μ ⟨U i, hU i⟩ := by { have := inner_content_Sup_nat h1 h2 (λ i, ⟨U i, hU i⟩), rwa [opens.supr_def] at this } -lemma inner_content_comap {μ : compacts G → ennreal} (f : G ≃ₜ G) +lemma inner_content_comap {μ : compacts G → ℝ≥0∞} (f : G ≃ₜ G) (h : ∀ ⦃K : compacts G⦄, μ (K.map f f.continuous) = μ K) (U : opens G) : inner_content μ (U.comap f.continuous) = inner_content μ U := begin @@ -142,14 +142,14 @@ begin end @[to_additive] -lemma is_mul_left_invariant_inner_content [group G] [topological_group G] {μ : compacts G → ennreal} +lemma is_mul_left_invariant_inner_content [group G] [topological_group G] {μ : compacts G → ℝ≥0∞} (h : ∀ (g : G) {K : compacts G}, μ (K.map _ $ continuous_mul_left g) = μ K) (g : G) (U : opens G) : inner_content μ (U.comap $ continuous_mul_left g) = inner_content μ U := by convert inner_content_comap (homeomorph.mul_left g) (λ K, h g) U -- @[to_additive] (fails for now) lemma inner_content_pos_of_is_mul_left_invariant [t2_space G] [group G] [topological_group G] - {μ : compacts G → ennreal} + {μ : compacts G → ℝ≥0∞} (h1 : μ ⊥ = 0) (h2 : ∀ (K₁ K₂ : compacts G), μ (K₁ ⊔ K₂) ≤ μ K₁ + μ K₂) (h3 : ∀ (g : G) {K : compacts G}, μ (K.map _ $ continuous_mul_left g) = μ K) @@ -168,7 +168,7 @@ begin simp only [is_mul_left_invariant_inner_content h3, finset.sum_const, nsmul_eq_mul, le_refl] end -lemma inner_content_mono' {μ : compacts G → ennreal} ⦃U V : set G⦄ +lemma inner_content_mono' {μ : compacts G → ℝ≥0∞} ⦃U V : set G⦄ (hU : is_open U) (hV : is_open V) (h2 : U ⊆ V) : inner_content μ ⟨U, hU⟩ ≤ inner_content μ ⟨V, hV⟩ := supr_le_supr $ λ K, supr_le_supr_const $ λ hK, subset.trans hK h2 @@ -176,10 +176,10 @@ supr_le_supr $ λ K, supr_le_supr_const $ λ hK, subset.trans hK h2 namespace outer_measure /-- Extending a content on compact sets to an outer measure on all sets. -/ -def of_content [t2_space G] (μ : compacts G → ennreal) (h1 : μ ⊥ = 0) : outer_measure G := +def of_content [t2_space G] (μ : compacts G → ℝ≥0∞) (h1 : μ ⊥ = 0) : outer_measure G := induced_outer_measure (λ U hU, inner_content μ ⟨U, hU⟩) is_open_empty (inner_content_empty h1) -variables [t2_space G] {μ : compacts G → ennreal} {h1 : μ ⊥ = 0} +variables [t2_space G] {μ : compacts G → ℝ≥0∞} {h1 : μ ⊥ = 0} (h2 : ∀ (K₁ K₂ : compacts G), μ (K₁ ⊔ K₂) ≤ μ K₁ + μ K₂) include h2 diff --git a/src/measure_theory/decomposition.lean b/src/measure_theory/decomposition.lean index 29e7e530690ff..8e15c7eb3e4c2 100644 --- a/src/measure_theory/decomposition.lean +++ b/src/measure_theory/decomposition.lean @@ -12,7 +12,7 @@ TODO: import measure_theory.measure_space open set filter -open_locale classical topological_space +open_locale classical topological_space ennreal namespace measure_theory @@ -34,9 +34,9 @@ begin have hμ : ∀s, μ s < ⊤ := assume s, lt_of_le_of_lt (measure_mono $ subset_univ _) hμ, have hν : ∀s, ν s < ⊤ := assume s, lt_of_le_of_lt (measure_mono $ subset_univ _) hν, - have to_nnreal_μ : ∀s, ((μ s).to_nnreal : ennreal) = μ s := + have to_nnreal_μ : ∀s, ((μ s).to_nnreal : ℝ≥0∞) = μ s := (assume s, ennreal.coe_to_nnreal $ ne_top_of_lt $ hμ _), - have to_nnreal_ν : ∀s, ((ν s).to_nnreal : ennreal) = ν s := + have to_nnreal_ν : ∀s, ((ν s).to_nnreal : ℝ≥0∞) = ν s := (assume s, ennreal.coe_to_nnreal $ ne_top_of_lt $ hν _), have d_empty : d ∅ = 0, { simp [d], rw [measure_empty, measure_empty], simp }, diff --git a/src/measure_theory/ess_sup.lean b/src/measure_theory/ess_sup.lean index 935fa072aaae5..53050c29605b6 100644 --- a/src/measure_theory/ess_sup.lean +++ b/src/measure_theory/ess_sup.lean @@ -12,7 +12,7 @@ We define the essential supremum and infimum of a function `f : α → β` with `μ` on `α`. The essential supremum is the infimum of the constants `c : β` such that `f x ≤ c` almost everywhere. -TODO: The essential supremum of functions `α → ennreal` is used in particular to define the norm in +TODO: The essential supremum of functions `α → ℝ≥0∞` is used in particular to define the norm in the `L∞` space (see measure_theory/lp_space.lean). There is a different quantity which is sometimes also called essential supremum: the least @@ -27,6 +27,7 @@ sense). We do not define that quantity here, which is simply the supremum of a m -/ open measure_theory +open_locale ennreal variables {α β : Type*} [measurable_space α] {μ : measure α} @@ -107,17 +108,17 @@ end complete_linear_order namespace ennreal -lemma ae_le_ess_sup (f : α → ennreal) : ∀ᵐ y ∂μ, f y ≤ ess_sup f μ := +lemma ae_le_ess_sup (f : α → ℝ≥0∞) : ∀ᵐ y ∂μ, f y ≤ ess_sup f μ := eventually_le_limsup f -lemma ess_sup_eq_zero_iff {f : α → ennreal} : ess_sup f μ = 0 ↔ f =ᵐ[μ] 0 := +lemma ess_sup_eq_zero_iff {f : α → ℝ≥0∞} : ess_sup f μ = 0 ↔ f =ᵐ[μ] 0 := limsup_eq_zero_iff -lemma ess_sup_const_mul {f : α → ennreal} {a : ennreal} : +lemma ess_sup_const_mul {f : α → ℝ≥0∞} {a : ℝ≥0∞} : ess_sup (λ (x : α), a * (f x)) μ = a * ess_sup f μ := limsup_const_mul -lemma ess_sup_add_le (f g : α → ennreal) : ess_sup (f + g) μ ≤ ess_sup f μ + ess_sup g μ := +lemma ess_sup_add_le (f g : α → ℝ≥0∞) : ess_sup (f + g) μ ≤ ess_sup f μ + ess_sup g μ := limsup_add_le f g end ennreal diff --git a/src/measure_theory/giry_monad.lean b/src/measure_theory/giry_monad.lean index 31f20bd99e36d..30ec040b36595 100644 --- a/src/measure_theory/giry_monad.lean +++ b/src/measure_theory/giry_monad.lean @@ -28,7 +28,7 @@ giry monad -/ noncomputable theory -open_locale classical big_operators +open_locale classical big_operators ennreal open classical set filter @@ -42,7 +42,7 @@ variables [measurable_space α] [measurable_space β] /-- Measurability structure on `measure`: Measures are measurable w.r.t. all projections -/ instance : measurable_space (measure α) := -⨆ (s : set α) (hs : measurable_set s), (borel ennreal).comap (λμ, μ s) +⨆ (s : set α) (hs : measurable_set s), (borel ℝ≥0∞).comap (λμ, μ s) lemma measurable_coe {s : set α} (hs : measurable_set s) : measurable (λμ : measure α, μ s) := measurable.of_comap_le $ le_supr_of_le s $ le_supr_of_le hs $ le_refl _ @@ -72,7 +72,7 @@ measurable_of_measurable_coe _ $ assume s hs, exact measurable_one.indicator hs end -lemma measurable_lintegral {f : α → ennreal} (hf : measurable f) : +lemma measurable_lintegral {f : α → ℝ≥0∞} (hf : measurable f) : measurable (λμ : measure α, ∫⁻ x, f x ∂μ) := begin simp only [lintegral_eq_supr_eapprox_lintegral, hf, simple_func.lintegral], @@ -102,7 +102,7 @@ lemma measurable_join : measurable (join : measure (measure α) → measure α) measurable_of_measurable_coe _ $ assume s hs, by simp only [join_apply hs]; exact measurable_lintegral (measurable_coe hs) -lemma lintegral_join {m : measure (measure α)} {f : α → ennreal} (hf : measurable f) : +lemma lintegral_join {m : measure (measure α)} {f : α → ℝ≥0∞} (hf : measurable f) : ∫⁻ x, f x ∂(join m) = ∫⁻ μ, ∫⁻ x, f x ∂μ ∂m := begin rw [lintegral_eq_supr_eapprox_lintegral hf], @@ -112,7 +112,7 @@ begin assume n x, join_apply (simple_func.measurable_set_preimage _ _), simp only [simple_func.lintegral, this], transitivity, - have : ∀(s : ℕ → finset ennreal) (f : ℕ → ennreal → measure α → ennreal) + have : ∀(s : ℕ → finset ℝ≥0∞) (f : ℕ → ℝ≥0∞ → measure α → ℝ≥0∞) (hf : ∀n r, measurable (f n r)) (hm : monotone (λn μ, ∑ r in s n, r * f n r μ)), (⨆n:ℕ, ∑ r in s n, r * ∫⁻ μ, f n r μ ∂m) = ∫⁻ μ, ⨆n:ℕ, ∑ r in s n, r * f n r μ ∂m, @@ -160,7 +160,7 @@ by rw [bind, join_apply hs, lintegral_map (measurable_coe hs) hf] lemma measurable_bind' {g : α → measure β} (hg : measurable g) : measurable (λm, bind m g) := measurable_join.comp (measurable_map _ hg) -lemma lintegral_bind {m : measure α} {μ : α → measure β} {f : β → ennreal} +lemma lintegral_bind {m : measure α} {μ : α → measure β} {f : β → ℝ≥0∞} (hμ : measurable μ) (hf : measurable f) : ∫⁻ x, f x ∂ (bind m μ) = ∫⁻ a, ∫⁻ x, f x ∂(μ a) ∂m:= (lintegral_join hf).trans (lintegral_map (measurable_lintegral hf) hμ) diff --git a/src/measure_theory/group.lean b/src/measure_theory/group.lean index e2d92b00ddaea..5018e461c019d 100644 --- a/src/measure_theory/group.lean +++ b/src/measure_theory/group.lean @@ -17,6 +17,7 @@ We develop some properties of measures on (topological) groups noncomputable theory +open_locale ennreal open has_inv set function measure_theory.measure namespace measure_theory @@ -35,7 +36,7 @@ variables [measurable_space G] [has_mul G] if the measure of left translations of a set are equal to the measure of the set itself. To left translate sets we use preimage under left addition, since preimages are nicer to work with than images."] -def is_mul_left_invariant (μ : set G → ennreal) : Prop := +def is_mul_left_invariant (μ : set G → ℝ≥0∞) : Prop := ∀ (g : G) {A : set G} (h : measurable_set A), μ ((λ h, g * h) ⁻¹' A) = μ A /-- A measure `μ` on a topological group is right invariant @@ -46,7 +47,7 @@ def is_mul_left_invariant (μ : set G → ennreal) : Prop := if the measure of right translations of a set are equal to the measure of the set itself. To right translate sets we use preimage under right addition, since preimages are nicer to work with than images."] -def is_mul_right_invariant (μ : set G → ennreal) : Prop := +def is_mul_right_invariant (μ : set G → ℝ≥0∞) : Prop := ∀ (g : G) {A : set G} (h : measurable_set A), μ ((λ h, h * g) ⁻¹' A) = μ A end @@ -187,13 +188,13 @@ by simp_rw [← ne_empty_iff_nonempty, ne.def, h2μ.null_iff_empty hμ h3μ hs] `f` is 0 iff `f` is 0. -/ -- @[to_additive] (fails for now) lemma lintegral_eq_zero_of_is_mul_left_invariant (hμ : regular μ) - (h2μ : is_mul_left_invariant μ) (h3μ : μ ≠ 0) {f : G → ennreal} (hf : continuous f) : + (h2μ : is_mul_left_invariant μ) (h3μ : μ ≠ 0) {f : G → ℝ≥0∞} (hf : continuous f) : ∫⁻ x, f x ∂μ = 0 ↔ f = 0 := begin split, swap, { rintro rfl, simp_rw [pi.zero_apply, lintegral_zero] }, intro h, contrapose h, simp_rw [funext_iff, not_forall, pi.zero_apply] at h, cases h with x hx, - obtain ⟨r, h1r, h2r⟩ : ∃ r : ennreal, 0 < r ∧ r < f x := + obtain ⟨r, h1r, h2r⟩ : ∃ r : ℝ≥0∞, 0 < r ∧ r < f x := exists_between (pos_iff_ne_zero.mpr hx), have h3r := hf.is_open_preimage (Ioi r) is_open_Ioi, let s := Ioi r, diff --git a/src/measure_theory/haar_measure.lean b/src/measure_theory/haar_measure.lean index 098f6f3f210a7..7170642554d93 100644 --- a/src/measure_theory/haar_measure.lean +++ b/src/measure_theory/haar_measure.lean @@ -54,7 +54,7 @@ where `ᵒ` denotes the interior. noncomputable theory open set has_inv function topological_space measurable_space -open_locale nnreal classical +open_locale nnreal classical ennreal variables {G : Type*} [group G] @@ -412,8 +412,8 @@ begin { apply continuous_iff_is_closed.mp this, exact is_closed_singleton }, end -/-- The function `chaar` interpreted in `ennreal` -/ -@[reducible] def echaar (K₀ : positive_compacts G) (K : compacts G) : ennreal := +/-- The function `chaar` interpreted in `ℝ≥0∞` -/ +@[reducible] def echaar (K₀ : positive_compacts G) (K : compacts G) : ℝ≥0∞ := show nnreal, from ⟨chaar K₀ K, chaar_nonneg _ _⟩ /-! We only prove the properties for `echaar` that we use at least twice below. -/ diff --git a/src/measure_theory/integration.lean b/src/measure_theory/integration.lean index a267482573734..f545d35236412 100644 --- a/src/measure_theory/integration.lean +++ b/src/measure_theory/integration.lean @@ -9,33 +9,33 @@ import data.indicator_function import data.support /-! -# Lebesgue integral for `ennreal`-valued functions +# Lebesgue integral for `ℝ≥0∞`-valued functions -We define simple functions and show that each Borel measurable function on `ennreal` can be +We define simple functions and show that each Borel measurable function on `ℝ≥0∞` can be approximated by a sequence of simple functions. -To prove something for an arbitrary measurable function into `ennreal`, the theorem +To prove something for an arbitrary measurable function into `ℝ≥0∞`, the theorem `measurable.ennreal_induction` shows that is it sufficient to show that the property holds for (multiples of) characteristic functions and is closed under addition and supremum of increasing sequences of functions. ## Notation -We introduce the following notation for the lower Lebesgue integral of a function `f : α → ennreal`. +We introduce the following notation for the lower Lebesgue integral of a function `f : α → ℝ≥0∞`. -* `∫⁻ x, f x ∂μ`: integral of a function `f : α → ennreal` with respect to a measure `μ`; -* `∫⁻ x, f x`: integral of a function `f : α → ennreal` with respect to the canonical measure +* `∫⁻ x, f x ∂μ`: integral of a function `f : α → ℝ≥0∞` with respect to a measure `μ`; +* `∫⁻ x, f x`: integral of a function `f : α → ℝ≥0∞` with respect to the canonical measure `volume` on `α`; -* `∫⁻ x in s, f x ∂μ`: integral of a function `f : α → ennreal` over a set `s` with respect +* `∫⁻ x in s, f x ∂μ`: integral of a function `f : α → ℝ≥0∞` over a set `s` with respect to a measure `μ`, defined as `∫⁻ x, f x ∂(μ.restrict s)`; -* `∫⁻ x in s, f x`: integral of a function `f : α → ennreal` over a set `s` with respect +* `∫⁻ x in s, f x`: integral of a function `f : α → ℝ≥0∞` over a set `s` with respect to the canonical measure `volume`, defined as `∫⁻ x, f x ∂(volume.restrict s)`. -/ noncomputable theory open set (hiding restrict restrict_apply) filter ennreal -open_locale classical topological_space big_operators nnreal +open_locale classical topological_space big_operators nnreal ennreal namespace measure_theory @@ -399,7 +399,7 @@ if hs : measurable_set s then by simp [hs, set.indicator_comp_of_zero hg] else by simp [restrict_of_not_measurable hs, hg] theorem map_coe_ennreal_restrict (f : α →ₛ ℝ≥0) (s : set α) : - (f.restrict s).map (coe : ℝ≥0 → ennreal) = (f.map coe).restrict s := + (f.restrict s).map (coe : ℝ≥0 → ℝ≥0∞) = (f.map coe).restrict s := map_restrict_of_zero ennreal.coe_zero _ _ theorem map_coe_nnreal_restrict (f : α →ₛ ℝ≥0) (s : set α) : @@ -442,7 +442,7 @@ section variables [semilattice_sup_bot β] [has_zero β] /-- Fix a sequence `i : ℕ → β`. Given a function `α → β`, its `n`-th approximation -by simple functions is defined so that in case `β = ennreal` it sends each `a` to the supremum +by simple functions is defined so that in case `β = ℝ≥0∞` it sends each `a` to the supremum of the set `{i k | k ≤ n ∧ i k ≤ f a}`, see `approx_apply` and `supr_approx_apply` for details. -/ def approx (i : ℕ → β) (f : α → β) (n : ℕ) : α →ₛ β := (finset.range n).sup (λk, restrict (const α (i k)) {a:α | i k ≤ f a}) @@ -494,29 +494,29 @@ end approx section eapprox -/-- A sequence of `ennreal`s such that its range is the set of non-negative rational numbers. -/ -def ennreal_rat_embed (n : ℕ) : ennreal := +/-- A sequence of `ℝ≥0∞`s such that its range is the set of non-negative rational numbers. -/ +def ennreal_rat_embed (n : ℕ) : ℝ≥0∞ := ennreal.of_real ((encodable.decode ℚ n).get_or_else (0 : ℚ)) lemma ennreal_rat_embed_encode (q : ℚ) : ennreal_rat_embed (encodable.encode q) = nnreal.of_real q := by rw [ennreal_rat_embed, encodable.encodek]; refl -/-- Approximate a function `α → ennreal` by a sequence of simple functions. -/ -def eapprox : (α → ennreal) → ℕ → α →ₛ ennreal := +/-- Approximate a function `α → ℝ≥0∞` by a sequence of simple functions. -/ +def eapprox : (α → ℝ≥0∞) → ℕ → α →ₛ ℝ≥0∞ := approx ennreal_rat_embed -@[mono] lemma monotone_eapprox (f : α → ennreal) : monotone (eapprox f) := +@[mono] lemma monotone_eapprox (f : α → ℝ≥0∞) : monotone (eapprox f) := monotone_approx _ f -lemma supr_eapprox_apply (f : α → ennreal) (hf : measurable f) (a : α) : - (⨆n, (eapprox f n : α →ₛ ennreal) a) = f a := +lemma supr_eapprox_apply (f : α → ℝ≥0∞) (hf : measurable f) (a : α) : + (⨆n, (eapprox f n : α →ₛ ℝ≥0∞) a) = f a := begin rw [eapprox, supr_approx_apply ennreal_rat_embed f a hf rfl], refine le_antisymm (supr_le $ assume i, supr_le $ assume hi, hi) (le_of_not_gt _), assume h, rcases ennreal.lt_iff_exists_rat_btwn.1 h with ⟨q, hq, lt_q, q_lt⟩, - have : (nnreal.of_real q : ennreal) ≤ + have : (nnreal.of_real q : ℝ≥0∞) ≤ (⨆ (k : ℕ) (h : ennreal_rat_embed k ≤ f a), ennreal_rat_embed k), { refine le_supr_of_le (encodable.encode q) _, rw [ennreal_rat_embed_encode q], @@ -525,9 +525,9 @@ begin exact lt_irrefl _ (lt_of_le_of_lt this lt_q) end -lemma eapprox_comp [measurable_space γ] {f : γ → ennreal} {g : α → γ} {n : ℕ} +lemma eapprox_comp [measurable_space γ] {f : γ → ℝ≥0∞} {g : α → γ} {n : ℕ} (hf : measurable f) (hg : measurable g) : - (eapprox (f ∘ g) n : α → ennreal) = (eapprox f n : γ →ₛ ennreal) ∘ g := + (eapprox (f ∘ g) n : α → ℝ≥0∞) = (eapprox f n : γ →ₛ ℝ≥0∞) ∘ g := funext $ assume a, approx_comp a hf hg end eapprox @@ -537,11 +537,11 @@ end measurable section measure variables [measurable_space α] {μ : measure α} -/-- Integral of a simple function whose codomain is `ennreal`. -/ -def lintegral (f : α →ₛ ennreal) (μ : measure α) : ennreal := +/-- Integral of a simple function whose codomain is `ℝ≥0∞`. -/ +def lintegral (f : α →ₛ ℝ≥0∞) (μ : measure α) : ℝ≥0∞ := ∑ x in f.range, x * μ (f ⁻¹' {x}) -lemma lintegral_eq_of_subset (f : α →ₛ ennreal) {s : finset ennreal} +lemma lintegral_eq_of_subset (f : α →ₛ ℝ≥0∞) {s : finset ℝ≥0∞} (hs : ∀ x, f x ≠ 0 → μ (f ⁻¹' {f x}) ≠ 0 → f x ∈ s) : f.lintegral μ = ∑ x in s, x * μ (f ⁻¹' {x}) := begin @@ -555,8 +555,8 @@ begin { intros, refl } end -/-- Calculate the integral of `(g ∘ f)`, where `g : β → ennreal` and `f : α →ₛ β`. -/ -lemma map_lintegral (g : β → ennreal) (f : α →ₛ β) : +/-- Calculate the integral of `(g ∘ f)`, where `g : β → ℝ≥0∞` and `f : α →ₛ β`. -/ +lemma map_lintegral (g : β → ℝ≥0∞) (f : α →ₛ β) : (f.map g).lintegral μ = ∑ x in f.range, g x * μ (f ⁻¹' {x}) := begin simp only [lintegral, range_map], @@ -568,7 +568,7 @@ begin { assume x, simp only [finset.mem_filter], rintro ⟨_, h⟩, rw h }, end -lemma add_lintegral (f g : α →ₛ ennreal) : (f + g).lintegral μ = f.lintegral μ + g.lintegral μ := +lemma add_lintegral (f g : α →ₛ ℝ≥0∞) : (f + g).lintegral μ = f.lintegral μ + g.lintegral μ := calc (f + g).lintegral μ = ∑ x in (pair f g).range, (x.1 * μ (pair f g ⁻¹' {x}) + x.2 * μ (pair f g ⁻¹' {x})) : by rw [add_eq_map₂, map_lintegral]; exact finset.sum_congr rfl (assume a ha, add_mul _ _ _) @@ -578,7 +578,7 @@ calc (f + g).lintegral μ = by rw [map_lintegral, map_lintegral] ... = lintegral f μ + lintegral g μ : rfl -lemma const_mul_lintegral (f : α →ₛ ennreal) (x : ennreal) : +lemma const_mul_lintegral (f : α →ₛ ℝ≥0∞) (x : ℝ≥0∞) : (const α x * f).lintegral μ = x * f.lintegral μ := calc (f.map (λa, x * a)).lintegral μ = ∑ r in f.range, x * r * μ (f ⁻¹' {r}) : map_lintegral _ _ @@ -587,8 +587,8 @@ calc (f.map (λa, x * a)).lintegral μ = ∑ r in f.range, x * r * μ (f ⁻¹' ... = x * f.lintegral μ : finset.mul_sum.symm -/-- Integral of a simple function `α →ₛ ennreal` as a bilinear map. -/ -def lintegralₗ : (α →ₛ ennreal) →ₗ[ennreal] measure α →ₗ[ennreal] ennreal := +/-- Integral of a simple function `α →ₛ ℝ≥0∞` as a bilinear map. -/ +def lintegralₗ : (α →ₛ ℝ≥0∞) →ₗ[ℝ≥0∞] measure α →ₗ[ℝ≥0∞] ℝ≥0∞ := { to_fun := λ f, { to_fun := lintegral f, map_add' := by simp [lintegral, mul_add, finset.sum_add_distrib], @@ -596,21 +596,21 @@ def lintegralₗ : (α →ₛ ennreal) →ₗ[ennreal] measure α →ₗ[ennreal map_add' := λ f g, linear_map.ext (λ μ, add_lintegral f g), map_smul' := λ c f, linear_map.ext (λ μ, const_mul_lintegral f c) } -@[simp] lemma zero_lintegral : (0 : α →ₛ ennreal).lintegral μ = 0 := +@[simp] lemma zero_lintegral : (0 : α →ₛ ℝ≥0∞).lintegral μ = 0 := linear_map.ext_iff.1 lintegralₗ.map_zero μ -lemma lintegral_add {ν} (f : α →ₛ ennreal) : f.lintegral (μ + ν) = f.lintegral μ + f.lintegral ν := +lemma lintegral_add {ν} (f : α →ₛ ℝ≥0∞) : f.lintegral (μ + ν) = f.lintegral μ + f.lintegral ν := (lintegralₗ f).map_add μ ν -lemma lintegral_smul (f : α →ₛ ennreal) (c : ennreal) : +lemma lintegral_smul (f : α →ₛ ℝ≥0∞) (c : ℝ≥0∞) : f.lintegral (c • μ) = c • f.lintegral μ := (lintegralₗ f).map_smul c μ -@[simp] lemma lintegral_zero (f : α →ₛ ennreal) : +@[simp] lemma lintegral_zero (f : α →ₛ ℝ≥0∞) : f.lintegral 0 = 0 := (lintegralₗ f).map_zero -lemma lintegral_sum {ι} (f : α →ₛ ennreal) (μ : ι → measure α) : +lemma lintegral_sum {ι} (f : α →ₛ ℝ≥0∞) (μ : ι → measure α) : f.lintegral (measure.sum μ) = ∑' i, f.lintegral (μ i) := begin simp only [lintegral, measure.sum_apply, f.measurable_set_preimage, ← finset.tsum_subtype, @@ -618,7 +618,7 @@ begin apply ennreal.tsum_comm end -lemma restrict_lintegral (f : α →ₛ ennreal) {s : set α} (hs : measurable_set s) : +lemma restrict_lintegral (f : α →ₛ ℝ≥0∞) {s : set α} (hs : measurable_set s) : (restrict f s).lintegral μ = ∑ r in f.range, r * μ (f ⁻¹' {r} ∩ s) := calc (restrict f s).lintegral μ = ∑ r in f.range, r * μ (restrict f s ⁻¹' {r}) : lintegral_eq_of_subset _ $ λ x hx, if hxs : x ∈ s @@ -628,16 +628,16 @@ calc (restrict f s).lintegral μ = ∑ r in f.range, r * μ (restrict f s ⁻¹' finset.sum_congr rfl $ forall_range_iff.2 $ λ b, if hb : f b = 0 then by simp only [hb, zero_mul] else by rw [restrict_preimage_singleton _ hs hb, inter_comm] -lemma lintegral_restrict (f : α →ₛ ennreal) (s : set α) (μ : measure α) : +lemma lintegral_restrict (f : α →ₛ ℝ≥0∞) (s : set α) (μ : measure α) : f.lintegral (μ.restrict s) = ∑ y in f.range, y * μ (f ⁻¹' {y} ∩ s) := by simp only [lintegral, measure.restrict_apply, f.measurable_set_preimage] -lemma restrict_lintegral_eq_lintegral_restrict (f : α →ₛ ennreal) {s : set α} +lemma restrict_lintegral_eq_lintegral_restrict (f : α →ₛ ℝ≥0∞) {s : set α} (hs : measurable_set s) : (restrict f s).lintegral μ = f.lintegral (μ.restrict s) := by rw [f.restrict_lintegral hs, lintegral_restrict] -lemma const_lintegral (c : ennreal) : (const α c).lintegral μ = c * μ univ := +lemma const_lintegral (c : ℝ≥0∞) : (const α c).lintegral μ = c * μ univ := begin rw [lintegral], by_cases ha : nonempty α, @@ -645,15 +645,15 @@ begin { simp [μ.eq_zero_of_not_nonempty ha] } end -lemma const_lintegral_restrict (c : ennreal) (s : set α) : +lemma const_lintegral_restrict (c : ℝ≥0∞) (s : set α) : (const α c).lintegral (μ.restrict s) = c * μ s := by rw [const_lintegral, measure.restrict_apply measurable_set.univ, univ_inter] -lemma restrict_const_lintegral (c : ennreal) {s : set α} (hs : measurable_set s) : +lemma restrict_const_lintegral (c : ℝ≥0∞) {s : set α} (hs : measurable_set s) : ((const α c).restrict s).lintegral μ = c * μ s := by rw [restrict_lintegral_eq_lintegral_restrict _ hs, const_lintegral_restrict] -lemma le_sup_lintegral (f g : α →ₛ ennreal) : f.lintegral μ ⊔ g.lintegral μ ≤ (f ⊔ g).lintegral μ := +lemma le_sup_lintegral (f g : α →ₛ ℝ≥0∞) : f.lintegral μ ⊔ g.lintegral μ ≤ (f ⊔ g).lintegral μ := calc f.lintegral μ ⊔ g.lintegral μ = ((pair f g).map prod.fst).lintegral μ ⊔ ((pair f g).map prod.snd).lintegral μ : rfl ... ≤ ∑ x in (pair f g).range, (x.1 ⊔ x.2) * μ (pair f g ⁻¹' {x}) : @@ -667,7 +667,7 @@ calc f.lintegral μ ⊔ g.lintegral μ = ... = (f ⊔ g).lintegral μ : by rw [sup_eq_map₂, map_lintegral] /-- `simple_func.lintegral` is monotone both in function and in measure. -/ -@[mono] lemma lintegral_mono {f g : α →ₛ ennreal} (hfg : f ≤ g) {μ ν : measure α} (hμν : μ ≤ ν) : +@[mono] lemma lintegral_mono {f g : α →ₛ ℝ≥0∞} (hfg : f ≤ g) {μ ν : measure α} (hμν : μ ≤ ν) : f.lintegral μ ≤ g.lintegral ν := calc f.lintegral μ ≤ f.lintegral μ ⊔ g.lintegral μ : le_sup_left ... ≤ (f ⊔ g).lintegral μ : le_sup_lintegral _ _ @@ -676,7 +676,7 @@ calc f.lintegral μ ≤ f.lintegral μ ⊔ g.lintegral μ : le_sup_left hμν _ (g.measurable_set_preimage _) /-- `simple_func.lintegral` depends only on the measures of `f ⁻¹' {y}`. -/ -lemma lintegral_eq_of_measure_preimage [measurable_space β] {f : α →ₛ ennreal} {g : β →ₛ ennreal} +lemma lintegral_eq_of_measure_preimage [measurable_space β] {f : α →ₛ ℝ≥0∞} {g : β →ₛ ℝ≥0∞} {ν : measure β} (H : ∀ y, μ (f ⁻¹' {y}) = ν (g ⁻¹' {y})) : f.lintegral μ = g.lintegral ν := begin @@ -688,12 +688,12 @@ begin end /-- If two simple functions are equal a.e., then their `lintegral`s are equal. -/ -lemma lintegral_congr {f g : α →ₛ ennreal} (h : f =ᵐ[μ] g) : +lemma lintegral_congr {f g : α →ₛ ℝ≥0∞} (h : f =ᵐ[μ] g) : f.lintegral μ = g.lintegral μ := lintegral_eq_of_measure_preimage $ λ y, measure_congr $ eventually.set_eq $ h.mono $ λ x hx, by simp [hx] -lemma lintegral_map {β} [measurable_space β] {μ' : measure β} (f : α →ₛ ennreal) (g : β →ₛ ennreal) +lemma lintegral_map {β} [measurable_space β] {μ' : measure β} (f : α →ₛ ℝ≥0∞) (g : β →ₛ ℝ≥0∞) (m : α → β) (eq : ∀a:α, f a = g (m a)) (h : ∀s:set β, measurable_set s → μ' s = μ (m ⁻¹' s)) : f.lintegral μ = g.lintegral μ' := lintegral_eq_of_measure_preimage $ λ y, @@ -771,7 +771,7 @@ protected lemma mul {β} [monoid_with_zero β] {f g : α →ₛ β} (hf : f.fin_ (f * g).fin_meas_supp μ := by { rw [mul_eq_map₂], exact hf.map₂ hg (zero_mul 0) } -lemma lintegral_lt_top {f : α →ₛ ennreal} (hm : f.fin_meas_supp μ) (hf : ∀ᵐ a ∂μ, f a < ⊤) : +lemma lintegral_lt_top {f : α →ₛ ℝ≥0∞} (hm : f.fin_meas_supp μ) (hf : ∀ᵐ a ∂μ, f a < ⊤) : f.lintegral μ < ⊤ := begin refine sum_lt_top (λ a ha, _), @@ -783,7 +783,7 @@ begin { exact mul_lt_top ha (fin_meas_supp_iff.1 hm _ ha0) } } end -lemma of_lintegral_lt_top {f : α →ₛ ennreal} (h : f.lintegral μ < ⊤) : f.fin_meas_supp μ := +lemma of_lintegral_lt_top {f : α →ₛ ℝ≥0∞} (h : f.lintegral μ < ⊤) : f.fin_meas_supp μ := begin refine fin_meas_supp_iff.2 (λ b hb, _), rw [lintegral, sum_lt_top_iff] at h, @@ -798,7 +798,7 @@ begin exact with_top.zero_lt_top } end -lemma iff_lintegral_lt_top {f : α →ₛ ennreal} (hf : ∀ᵐ a ∂μ, f a < ⊤) : +lemma iff_lintegral_lt_top {f : α →ₛ ℝ≥0∞} (hf : ∀ᵐ a ∂μ, f a < ⊤) : f.fin_meas_supp μ ↔ f.lintegral μ < ⊤ := ⟨λ h, h.lintegral_lt_top hf, λ h, of_lintegral_lt_top h⟩ @@ -847,8 +847,8 @@ open simple_func variables [measurable_space α] {μ : measure α} /-- The lower Lebesgue integral of a function `f` with respect to a measure `μ`. -/ -def lintegral (μ : measure α) (f : α → ennreal) : ennreal := -⨆ (g : α →ₛ ennreal) (hf : ⇑g ≤ f), g.lintegral μ +def lintegral (μ : measure α) (f : α → ℝ≥0∞) : ℝ≥0∞ := +⨆ (g : α →ₛ ℝ≥0∞) (hf : ⇑g ≤ f), g.lintegral μ /-! In the notation for integrals, an expression like `∫⁻ x, g ∥x∥ ∂μ` will not be parsed correctly, and needs parentheses. We do not set the binding power of `r` to `0`, because then @@ -859,17 +859,17 @@ notation `∫⁻` binders ` in ` s `, ` r:(scoped:60 f, f) ` ∂` μ:70 := lintegral (measure.restrict μ s) r notation `∫⁻` binders ` in ` s `, ` r:(scoped:60 f, lintegral (measure.restrict volume s) f) := r -theorem simple_func.lintegral_eq_lintegral (f : α →ₛ ennreal) (μ : measure α) : +theorem simple_func.lintegral_eq_lintegral (f : α →ₛ ℝ≥0∞) (μ : measure α) : ∫⁻ a, f a ∂ μ = f.lintegral μ := le_antisymm (bsupr_le $ λ g hg, lintegral_mono hg $ le_refl _) (le_supr_of_le f $ le_supr_of_le (le_refl _) (le_refl _)) -@[mono] lemma lintegral_mono' ⦃μ ν : measure α⦄ (hμν : μ ≤ ν) ⦃f g : α → ennreal⦄ (hfg : f ≤ g) : +@[mono] lemma lintegral_mono' ⦃μ ν : measure α⦄ (hμν : μ ≤ ν) ⦃f g : α → ℝ≥0∞⦄ (hfg : f ≤ g) : ∫⁻ a, f a ∂μ ≤ ∫⁻ a, g a ∂ν := supr_le_supr $ λ φ, supr_le_supr2 $ λ hφ, ⟨le_trans hφ hfg, lintegral_mono (le_refl φ) hμν⟩ -lemma lintegral_mono ⦃f g : α → ennreal⦄ (hfg : f ≤ g) : +lemma lintegral_mono ⦃f g : α → ℝ≥0∞⦄ (hfg : f ≤ g) : ∫⁻ a, f a ∂μ ≤ ∫⁻ a, g a ∂μ := lintegral_mono' (le_refl μ) hfg @@ -885,31 +885,31 @@ end lemma monotone_lintegral (μ : measure α) : monotone (lintegral μ) := lintegral_mono -@[simp] lemma lintegral_const (c : ennreal) : ∫⁻ a, c ∂μ = c * μ univ := +@[simp] lemma lintegral_const (c : ℝ≥0∞) : ∫⁻ a, c ∂μ = c * μ univ := by rw [← simple_func.const_lintegral, ← simple_func.lintegral_eq_lintegral, simple_func.coe_const] -@[simp] lemma lintegral_one : ∫⁻ a, (1 : ennreal) ∂μ = μ univ := +@[simp] lemma lintegral_one : ∫⁻ a, (1 : ℝ≥0∞) ∂μ = μ univ := by rw [lintegral_const, one_mul] -lemma set_lintegral_const (s : set α) (c : ennreal) : ∫⁻ a in s, c ∂μ = c * μ s := +lemma set_lintegral_const (s : set α) (c : ℝ≥0∞) : ∫⁻ a in s, c ∂μ = c * μ s := by rw [lintegral_const, measure.restrict_apply_univ] lemma set_lintegral_one (s) : ∫⁻ a in s, 1 ∂μ = μ s := by rw [set_lintegral_const, one_mul] /-- `∫⁻ a in s, f a ∂μ` is defined as the supremum of integrals of simple functions -`φ : α →ₛ ennreal` such that `φ ≤ f`. This lemma says that it suffices to take +`φ : α →ₛ ℝ≥0∞` such that `φ ≤ f`. This lemma says that it suffices to take functions `φ : α →ₛ ℝ≥0`. -/ -lemma lintegral_eq_nnreal (f : α → ennreal) (μ : measure α) : +lemma lintegral_eq_nnreal (f : α → ℝ≥0∞) (μ : measure α) : (∫⁻ a, f a ∂μ) = (⨆ (φ : α →ₛ ℝ≥0) (hf : ∀ x, ↑(φ x) ≤ f x), - (φ.map (coe : ℝ≥0 → ennreal)).lintegral μ) := + (φ.map (coe : ℝ≥0 → ℝ≥0∞)).lintegral μ) := begin refine le_antisymm (bsupr_le $ assume φ hφ, _) - (supr_le_supr2 $ λ φ, ⟨φ.map (coe : ℝ≥0 → ennreal), le_refl _⟩), + (supr_le_supr2 $ λ φ, ⟨φ.map (coe : ℝ≥0 → ℝ≥0∞), le_refl _⟩), by_cases h : ∀ᵐ a ∂μ, φ a ≠ ⊤, { let ψ := φ.map ennreal.to_nnreal, - replace h : ψ.map (coe : ℝ≥0 → ennreal) =ᵐ[μ] φ := + replace h : ψ.map (coe : ℝ≥0 → ℝ≥0∞) =ᵐ[μ] φ := h.mono (λ a, ennreal.coe_to_nnreal), have : ∀ x, ↑(ψ x) ≤ f x := λ x, le_trans ennreal.coe_to_nnreal_le_self (hφ x), exact le_supr_of_le (φ.map ennreal.to_nnreal) @@ -926,8 +926,8 @@ begin simp only [hx, le_top] } end -lemma exists_simple_func_forall_lintegral_sub_lt_of_pos {f : α → ennreal} (h : ∫⁻ x, f x ∂μ < ⊤) - {ε : ennreal} (hε : 0 < ε) : +lemma exists_simple_func_forall_lintegral_sub_lt_of_pos {f : α → ℝ≥0∞} (h : ∫⁻ x, f x ∂μ < ⊤) + {ε : ℝ≥0∞} (hε : 0 < ε) : ∃ φ : α →ₛ ℝ≥0, (∀ x, ↑(φ x) ≤ f x) ∧ ∀ ψ : α →ₛ ℝ≥0, (∀ x, ↑(ψ x) ≤ f x) → (map coe (ψ - φ)).lintegral μ < ε := begin @@ -944,26 +944,26 @@ begin simp only [add_apply, sub_apply, nnreal.add_sub_eq_max] end -theorem supr_lintegral_le {ι : Sort*} (f : ι → α → ennreal) : +theorem supr_lintegral_le {ι : Sort*} (f : ι → α → ℝ≥0∞) : (⨆i, ∫⁻ a, f i a ∂μ) ≤ (∫⁻ a, ⨆i, f i a ∂μ) := begin simp only [← supr_apply], exact (monotone_lintegral μ).le_map_supr end -theorem supr2_lintegral_le {ι : Sort*} {ι' : ι → Sort*} (f : Π i, ι' i → α → ennreal) : +theorem supr2_lintegral_le {ι : Sort*} {ι' : ι → Sort*} (f : Π i, ι' i → α → ℝ≥0∞) : (⨆i (h : ι' i), ∫⁻ a, f i h a ∂μ) ≤ (∫⁻ a, ⨆i (h : ι' i), f i h a ∂μ) := by { convert (monotone_lintegral μ).le_map_supr2 f, ext1 a, simp only [supr_apply] } -theorem le_infi_lintegral {ι : Sort*} (f : ι → α → ennreal) : +theorem le_infi_lintegral {ι : Sort*} (f : ι → α → ℝ≥0∞) : (∫⁻ a, ⨅i, f i a ∂μ) ≤ (⨅i, ∫⁻ a, f i a ∂μ) := by { simp only [← infi_apply], exact (monotone_lintegral μ).map_infi_le } -theorem le_infi2_lintegral {ι : Sort*} {ι' : ι → Sort*} (f : Π i, ι' i → α → ennreal) : +theorem le_infi2_lintegral {ι : Sort*} {ι' : ι → Sort*} (f : Π i, ι' i → α → ℝ≥0∞) : (∫⁻ a, ⨅ i (h : ι' i), f i h a ∂μ) ≤ (⨅ i (h : ι' i), ∫⁻ a, f i h a ∂μ) := by { convert (monotone_lintegral μ).map_infi2_le f, ext1 a, simp only [infi_apply] } -lemma lintegral_mono_ae {f g : α → ennreal} (h : ∀ᵐ a ∂μ, f a ≤ g a) : +lemma lintegral_mono_ae {f g : α → ℝ≥0∞} (h : ∀ᵐ a ∂μ, f a ≤ g a) : (∫⁻ a, f a ∂μ) ≤ (∫⁻ a, g a ∂μ) := begin rcases exists_measurable_superset_of_null h with ⟨t, hts, ht, ht0⟩, @@ -979,15 +979,15 @@ begin exact (hnt hat).elim } end -lemma lintegral_congr_ae {f g : α → ennreal} (h : f =ᵐ[μ] g) : +lemma lintegral_congr_ae {f g : α → ℝ≥0∞} (h : f =ᵐ[μ] g) : (∫⁻ a, f a ∂μ) = (∫⁻ a, g a ∂μ) := le_antisymm (lintegral_mono_ae $ h.le) (lintegral_mono_ae $ h.symm.le) -lemma lintegral_congr {f g : α → ennreal} (h : ∀ a, f a = g a) : +lemma lintegral_congr {f g : α → ℝ≥0∞} (h : ∀ a, f a = g a) : (∫⁻ a, f a ∂μ) = (∫⁻ a, g a ∂μ) := by simp only [h] -lemma set_lintegral_congr {f : α → ennreal} {s t : set α} (h : s =ᵐ[μ] t) : +lemma set_lintegral_congr {f : α → ℝ≥0∞} {s t : set α} (h : s =ᵐ[μ] t) : ∫⁻ x in s, f x ∂μ = ∫⁻ x in t, f x ∂μ := by rw [restrict_congr_set h] @@ -995,10 +995,10 @@ by rw [restrict_congr_set h] See `lintegral_supr_directed` for a more general form. -/ theorem lintegral_supr - {f : ℕ → α → ennreal} (hf : ∀n, measurable (f n)) (h_mono : monotone f) : + {f : ℕ → α → ℝ≥0∞} (hf : ∀n, measurable (f n)) (h_mono : monotone f) : (∫⁻ a, ⨆n, f n a ∂μ) = (⨆n, ∫⁻ a, f n a ∂μ) := begin - set c : ℝ≥0 → ennreal := coe, + set c : ℝ≥0 → ℝ≥0∞ := coe, set F := λ a:α, ⨆n, f n a, have hF : measurable F := measurable_supr hf, refine le_antisymm _ (supr_lintegral_le _), @@ -1008,7 +1008,7 @@ begin rcases ennreal.lt_iff_exists_coe.1 ha with ⟨r, rfl, ha⟩, have ha : r < 1 := ennreal.coe_lt_coe.1 ha, let rs := s.map (λa, r * a), - have eq_rs : (const α r : α →ₛ ennreal) * map c s = rs.map c, + have eq_rs : (const α r : α →ₛ ℝ≥0∞) * map c s = rs.map c, { ext1 a, exact ennreal.coe_mul.symm }, have eq : ∀p, (rs.map c) ⁻¹' {p} = (⋃n, (rs.map c) ⁻¹' {p} ∩ {a | p ≤ f n a}), { assume p, @@ -1024,13 +1024,13 @@ begin exact mul_lt_mul_of_pos_right ha (pos_iff_ne_zero.2 this) }, rcases lt_supr_iff.1 this with ⟨i, hi⟩, exact mem_Union.2 ⟨i, le_of_lt hi⟩ }, - have mono : ∀r:ennreal, monotone (λn, (rs.map c) ⁻¹' {r} ∩ {a | r ≤ f n a}), + have mono : ∀r:ℝ≥0∞, monotone (λn, (rs.map c) ⁻¹' {r} ∩ {a | r ≤ f n a}), { assume r i j h, refine inter_subset_inter (subset.refl _) _, assume x hx, exact le_trans hx (h_mono h x) }, have h_meas : ∀n, measurable_set {a : α | ⇑(map c rs) a ≤ f n a} := assume n, measurable_set_le (simple_func.measurable _) (hf n), - calc (r:ennreal) * (s.map c).lintegral μ = ∑ r in (rs.map c).range, r * μ ((rs.map c) ⁻¹' {r}) : + calc (r:ℝ≥0∞) * (s.map c).lintegral μ = ∑ r in (rs.map c).range, r * μ ((rs.map c) ⁻¹' {r}) : by rw [← const_mul_lintegral, eq_rs, simple_func.lintegral] ... ≤ ∑ r in (rs.map c).range, r * μ (⋃n, (rs.map c) ⁻¹' {r} ∩ {a | r ≤ f n a}) : le_of_eq (finset.sum_congr rfl $ assume x hx, by rw ← eq) @@ -1072,12 +1072,12 @@ end /-- Monotone convergence theorem -- sometimes called Beppo-Levi convergence. Version with ae_measurable functions. -/ -theorem lintegral_supr' {f : ℕ → α → ennreal} (hf : ∀n, ae_measurable (f n) μ) +theorem lintegral_supr' {f : ℕ → α → ℝ≥0∞} (hf : ∀n, ae_measurable (f n) μ) (h_mono : ∀ᵐ x ∂μ, monotone (λ n, f n x)) : (∫⁻ a, ⨆n, f n a ∂μ) = (⨆n, ∫⁻ a, f n a ∂μ) := begin simp_rw ←supr_apply, - let p : α → (ℕ → ennreal) → Prop := λ x f', monotone f', + let p : α → (ℕ → ℝ≥0∞) → Prop := λ x f', monotone f', have hp : ∀ᵐ x ∂μ, p x (λ i, f i x), from h_mono, have h_ae_seq_mono : monotone (ae_seq hf p), { intros n m hnm x, @@ -1092,11 +1092,11 @@ begin exact funext (λ n, lintegral_congr_ae (ae_seq.ae_seq_n_eq_fun_n_ae hf hp n)), end -lemma lintegral_eq_supr_eapprox_lintegral {f : α → ennreal} (hf : measurable f) : +lemma lintegral_eq_supr_eapprox_lintegral {f : α → ℝ≥0∞} (hf : measurable f) : (∫⁻ a, f a ∂μ) = (⨆n, (eapprox f n).lintegral μ) := -calc (∫⁻ a, f a ∂μ) = (∫⁻ a, ⨆n, (eapprox f n : α → ennreal) a ∂μ) : +calc (∫⁻ a, f a ∂μ) = (∫⁻ a, ⨆n, (eapprox f n : α → ℝ≥0∞) a ∂μ) : by congr; ext a; rw [supr_eapprox_apply f hf] -... = (⨆n, ∫⁻ a, (eapprox f n : α → ennreal) a ∂μ) : +... = (⨆n, ∫⁻ a, (eapprox f n : α → ℝ≥0∞) a ∂μ) : begin rw [lintegral_supr], { assume n, exact (eapprox f n).measurable }, @@ -1106,8 +1106,8 @@ end /-- If `f` has finite integral, then `∫⁻ x in s, f x ∂μ` is absolutely continuous in `s`: it tends to zero as `μ s` tends to zero. This lemma states states this fact in terms of `ε` and `δ`. -/ -lemma exists_pos_set_lintegral_lt_of_measure_lt {f : α → ennreal} (h : ∫⁻ x, f x ∂μ < ⊤) - {ε : ennreal} (hε : 0 < ε) : +lemma exists_pos_set_lintegral_lt_of_measure_lt {f : α → ℝ≥0∞} (h : ∫⁻ x, f x ∂μ < ⊤) + {ε : ℝ≥0∞} (hε : 0 < ε) : ∃ δ > 0, ∀ s, μ s < δ → ∫⁻ x in s, f x ∂μ < ε := begin rcases exists_between hε with ⟨ε₂, hε₂0, hε₂ε⟩, rcases exists_between hε₂0 with ⟨ε₁, hε₁0, hε₁₂⟩, @@ -1129,7 +1129,7 @@ begin refine add_le_add le_rfl (le_trans _ (hφ _ hψ).le), exact simple_func.lintegral_mono le_rfl measure.restrict_le_self end - ... ≤ (simple_func.const α (C : ennreal)).lintegral (μ.restrict s) + ε₁ : + ... ≤ (simple_func.const α (C : ℝ≥0∞)).lintegral (μ.restrict s) + ε₁ : by { mono*, exacts [λ x, coe_le_coe.2 (hC x), le_rfl, le_rfl] } ... = C * μ s + ε₁ : by simp [← simple_func.lintegral_eq_lintegral] ... ≤ C * ((ε₂ - ε₁) / C) + ε₁ : by { mono*, exacts [le_rfl, hs.le, le_rfl] } @@ -1139,7 +1139,7 @@ end /-- If `f` has finite integral, then `∫⁻ x in s, f x ∂μ` is absolutely continuous in `s`: it tends to zero as `μ s` tends to zero. -/ -lemma tendsto_set_lintegral_zero {ι} {f : α → ennreal} (h : ∫⁻ x, f x ∂μ < ⊤) +lemma tendsto_set_lintegral_zero {ι} {f : α → ℝ≥0∞} (h : ∫⁻ x, f x ∂μ < ⊤) {l : filter ι} {s : ι → set α} (hl : tendsto (μ ∘ s) l (𝓝 0)) : tendsto (λ i, ∫⁻ x in s i, f x ∂μ) l (𝓝 0) := begin @@ -1150,12 +1150,12 @@ begin exact (hl δ δ0).mono (λ i, hδ _) end -@[simp] lemma lintegral_add {f g : α → ennreal} (hf : measurable f) (hg : measurable g) : +@[simp] lemma lintegral_add {f g : α → ℝ≥0∞} (hf : measurable f) (hg : measurable g) : (∫⁻ a, f a + g a ∂μ) = (∫⁻ a, f a ∂μ) + (∫⁻ a, g a ∂μ) := calc (∫⁻ a, f a + g a ∂μ) = - (∫⁻ a, (⨆n, (eapprox f n : α → ennreal) a) + (⨆n, (eapprox g n : α → ennreal) a) ∂μ) : + (∫⁻ a, (⨆n, (eapprox f n : α → ℝ≥0∞) a) + (⨆n, (eapprox g n : α → ℝ≥0∞) a) ∂μ) : by simp only [supr_eapprox_apply, hf, hg] - ... = (∫⁻ a, (⨆n, (eapprox f n + eapprox g n : α → ennreal) a) ∂μ) : + ... = (∫⁻ a, (⨆n, (eapprox f n + eapprox g n : α → ℝ≥0∞) a) ∂μ) : begin congr, funext a, rw [ennreal.supr_add_supr_of_monotone], { refl }, @@ -1177,7 +1177,7 @@ calc (∫⁻ a, f a + g a ∂μ) = ... = (∫⁻ a, f a ∂μ) + (∫⁻ a, g a ∂μ) : by rw [lintegral_eq_supr_eapprox_lintegral hf, lintegral_eq_supr_eapprox_lintegral hg] -lemma lintegral_add' {f g : α → ennreal} (hf : ae_measurable f μ) (hg : ae_measurable g μ) : +lemma lintegral_add' {f g : α → ℝ≥0∞} (hf : ae_measurable f μ) (hg : ae_measurable g μ) : (∫⁻ a, f a + g a ∂μ) = (∫⁻ a, f a ∂μ) + (∫⁻ a, g a ∂μ) := calc (∫⁻ a, f a + g a ∂μ) = (∫⁻ a, hf.mk f a + hg.mk g a ∂μ) : lintegral_congr_ae (eventually_eq.add hf.ae_eq_mk hg.ae_eq_mk) @@ -1190,13 +1190,13 @@ end lemma lintegral_zero : (∫⁻ a:α, 0 ∂μ) = 0 := by simp -lemma lintegral_zero_fun : (∫⁻ a:α, (0 : α → ennreal) a ∂μ) = 0 := by simp +lemma lintegral_zero_fun : (∫⁻ a:α, (0 : α → ℝ≥0∞) a ∂μ) = 0 := by simp -@[simp] lemma lintegral_smul_measure (c : ennreal) (f : α → ennreal) : +@[simp] lemma lintegral_smul_measure (c : ℝ≥0∞) (f : α → ℝ≥0∞) : ∫⁻ a, f a ∂ (c • μ) = c * ∫⁻ a, f a ∂μ := by simp only [lintegral, supr_subtype', simple_func.lintegral_smul, ennreal.mul_supr, smul_eq_mul] -@[simp] lemma lintegral_sum_measure {ι} (f : α → ennreal) (μ : ι → measure α) : +@[simp] lemma lintegral_sum_measure {ι} (f : α → ℝ≥0∞) (μ : ι → measure α) : ∫⁻ a, f a ∂(measure.sum μ) = ∑' i, ∫⁻ a, f a ∂(μ i) := begin simp only [lintegral, supr_subtype', simple_func.lintegral_sum, ennreal.tsum_eq_supr_sum], @@ -1211,14 +1211,14 @@ begin (finset.sum_le_sum $ λ j hj, simple_func.lintegral_mono le_sup_right (le_refl _))⟩ end -@[simp] lemma lintegral_add_measure (f : α → ennreal) (μ ν : measure α) : +@[simp] lemma lintegral_add_measure (f : α → ℝ≥0∞) (μ ν : measure α) : ∫⁻ a, f a ∂ (μ + ν) = ∫⁻ a, f a ∂μ + ∫⁻ a, f a ∂ν := by simpa [tsum_fintype] using lintegral_sum_measure f (λ b, cond b μ ν) -@[simp] lemma lintegral_zero_measure (f : α → ennreal) : ∫⁻ a, f a ∂0 = 0 := +@[simp] lemma lintegral_zero_measure (f : α → ℝ≥0∞) : ∫⁻ a, f a ∂0 = 0 := bot_unique $ by simp [lintegral] -lemma lintegral_finset_sum (s : finset β) {f : β → α → ennreal} (hf : ∀b, measurable (f b)) : +lemma lintegral_finset_sum (s : finset β) {f : β → α → ℝ≥0∞} (hf : ∀b, measurable (f b)) : (∫⁻ a, ∑ b in s, f b a ∂μ) = ∑ b in s, ∫⁻ a, f b a ∂μ := begin refine finset.induction_on s _ _, @@ -1228,7 +1228,7 @@ begin rw [lintegral_add (hf _) (s.measurable_sum hf), ih] } end -@[simp] lemma lintegral_const_mul (r : ennreal) {f : α → ennreal} (hf : measurable f) : +@[simp] lemma lintegral_const_mul (r : ℝ≥0∞) {f : α → ℝ≥0∞} (hf : measurable f) : (∫⁻ a, r * f a ∂μ) = r * (∫⁻ a, f a ∂μ) := calc (∫⁻ a, r * f a ∂μ) = (∫⁻ a, (⨆n, (const α r * eapprox f n) a) ∂μ) : by { congr, funext a, rw [← supr_eapprox_apply f hf, ennreal.mul_supr], refl } @@ -1243,7 +1243,7 @@ calc (∫⁻ a, r * f a ∂μ) = (∫⁻ a, (⨆n, (const α r * eapprox f n) a) end ... = r * (∫⁻ a, f a ∂μ) : by rw [← ennreal.mul_supr, lintegral_eq_supr_eapprox_lintegral hf] -lemma lintegral_const_mul'' (r : ennreal) {f : α → ennreal} (hf : ae_measurable f μ) : +lemma lintegral_const_mul'' (r : ℝ≥0∞) {f : α → ℝ≥0∞} (hf : ae_measurable f μ) : (∫⁻ a, r * f a ∂μ) = r * (∫⁻ a, f a ∂μ) := begin have A : ∫⁻ a, f a ∂μ = ∫⁻ a, hf.mk f a ∂μ := lintegral_congr_ae hf.ae_eq_mk, @@ -1252,7 +1252,7 @@ begin rw [A, B, lintegral_const_mul _ hf.measurable_mk], end -lemma lintegral_const_mul_le (r : ennreal) (f : α → ennreal) : +lemma lintegral_const_mul_le (r : ℝ≥0∞) (f : α → ℝ≥0∞) : r * (∫⁻ a, f a ∂μ) ≤ (∫⁻ a, r * f a ∂μ) := begin rw [lintegral, ennreal.mul_supr], @@ -1265,7 +1265,7 @@ begin exact canonically_ordered_semiring.mul_le_mul (le_refl _) (hs x) end -lemma lintegral_const_mul' (r : ennreal) (f : α → ennreal) (hr : r ≠ ⊤) : +lemma lintegral_const_mul' (r : ℝ≥0∞) (f : α → ℝ≥0∞) (hr : r ≠ ⊤) : (∫⁻ a, r * f a ∂μ) = r * (∫⁻ a, f a ∂μ) := begin by_cases h : r = 0, @@ -1279,41 +1279,41 @@ begin using canonically_ordered_semiring.mul_le_mul (le_refl r) this end -lemma lintegral_mul_const (r : ennreal) {f : α → ennreal} (hf : measurable f) : +lemma lintegral_mul_const (r : ℝ≥0∞) {f : α → ℝ≥0∞} (hf : measurable f) : ∫⁻ a, f a * r ∂μ = ∫⁻ a, f a ∂μ * r := by simp_rw [mul_comm, lintegral_const_mul r hf] -lemma lintegral_mul_const'' (r : ennreal) {f : α → ennreal} (hf : ae_measurable f μ) : +lemma lintegral_mul_const'' (r : ℝ≥0∞) {f : α → ℝ≥0∞} (hf : ae_measurable f μ) : ∫⁻ a, f a * r ∂μ = ∫⁻ a, f a ∂μ * r := by simp_rw [mul_comm, lintegral_const_mul'' r hf] -lemma lintegral_mul_const_le (r : ennreal) (f : α → ennreal) : +lemma lintegral_mul_const_le (r : ℝ≥0∞) (f : α → ℝ≥0∞) : ∫⁻ a, f a ∂μ * r ≤ ∫⁻ a, f a * r ∂μ := by simp_rw [mul_comm, lintegral_const_mul_le r f] -lemma lintegral_mul_const' (r : ennreal) (f : α → ennreal) (hr : r ≠ ⊤): +lemma lintegral_mul_const' (r : ℝ≥0∞) (f : α → ℝ≥0∞) (hr : r ≠ ⊤): ∫⁻ a, f a * r ∂μ = ∫⁻ a, f a ∂μ * r := by simp_rw [mul_comm, lintegral_const_mul' r f hr] /- A double integral of a product where each factor contains only one variable is a product of integrals -/ lemma lintegral_lintegral_mul {β} [measurable_space β] {ν : measure β} - {f : α → ennreal} {g : β → ennreal} (hf : ae_measurable f μ) (hg : ae_measurable g ν) : + {f : α → ℝ≥0∞} {g : β → ℝ≥0∞} (hf : ae_measurable f μ) (hg : ae_measurable g ν) : ∫⁻ x, ∫⁻ y, f x * g y ∂ν ∂μ = ∫⁻ x, f x ∂μ * ∫⁻ y, g y ∂ν := by simp [lintegral_const_mul'' _ hg, lintegral_mul_const'' _ hf] -- TODO: Need a better way of rewriting inside of a integral -lemma lintegral_rw₁ {f f' : α → β} (h : f =ᵐ[μ] f') (g : β → ennreal) : +lemma lintegral_rw₁ {f f' : α → β} (h : f =ᵐ[μ] f') (g : β → ℝ≥0∞) : (∫⁻ a, g (f a) ∂μ) = (∫⁻ a, g (f' a) ∂μ) := lintegral_congr_ae $ h.mono $ λ a h, by rw h -- TODO: Need a better way of rewriting inside of a integral lemma lintegral_rw₂ {f₁ f₁' : α → β} {f₂ f₂' : α → γ} (h₁ : f₁ =ᵐ[μ] f₁') - (h₂ : f₂ =ᵐ[μ] f₂') (g : β → γ → ennreal) : + (h₂ : f₂ =ᵐ[μ] f₂') (g : β → γ → ℝ≥0∞) : (∫⁻ a, g (f₁ a) (f₂ a) ∂μ) = (∫⁻ a, g (f₁' a) (f₂' a) ∂μ) := lintegral_congr_ae $ h₁.mp $ h₂.mono $ λ _ h₂ h₁, by rw [h₁, h₂] -@[simp] lemma lintegral_indicator (f : α → ennreal) {s : set α} (hs : measurable_set s) : +@[simp] lemma lintegral_indicator (f : α → ℝ≥0∞) {s : set α} (hs : measurable_set s) : ∫⁻ a, s.indicator f a ∂μ = ∫⁻ a in s, f a ∂μ := begin simp only [lintegral, ← restrict_lintegral_eq_lintegral_restrict _ hs, supr_subtype'], @@ -1329,7 +1329,7 @@ begin end /-- Chebyshev's inequality -/ -lemma mul_meas_ge_le_lintegral {f : α → ennreal} (hf : measurable f) (ε : ennreal) : +lemma mul_meas_ge_le_lintegral {f : α → ℝ≥0∞} (hf : measurable f) (ε : ℝ≥0∞) : ε * μ {x | ε ≤ f x} ≤ ∫⁻ a, f a ∂μ := begin have : measurable_set {a : α | ε ≤ f a }, from hf measurable_set_Ici, @@ -1339,13 +1339,13 @@ begin split_ifs; [assumption, exact zero_le _] end -lemma meas_ge_le_lintegral_div {f : α → ennreal} (hf : measurable f) {ε : ennreal} +lemma meas_ge_le_lintegral_div {f : α → ℝ≥0∞} (hf : measurable f) {ε : ℝ≥0∞} (hε : ε ≠ 0) (hε' : ε ≠ ⊤) : μ {x | ε ≤ f x} ≤ (∫⁻ a, f a ∂μ) / ε := (ennreal.le_div_iff_mul_le (or.inl hε) (or.inl hε')).2 $ by { rw [mul_comm], exact mul_meas_ge_le_lintegral hf ε } -@[simp] lemma lintegral_eq_zero_iff {f : α → ennreal} (hf : measurable f) : +@[simp] lemma lintegral_eq_zero_iff {f : α → ℝ≥0∞} (hf : measurable f) : ∫⁻ a, f a ∂μ = 0 ↔ (f =ᵐ[μ] 0) := begin refine iff.intro (assume h, _) (assume h, _), @@ -1366,7 +1366,7 @@ begin ... = 0 : lintegral_zero } end -@[simp] lemma lintegral_eq_zero_iff' {f : α → ennreal} (hf : ae_measurable f μ) : +@[simp] lemma lintegral_eq_zero_iff' {f : α → ℝ≥0∞} (hf : ae_measurable f μ) : ∫⁻ a, f a ∂μ = 0 ↔ (f =ᵐ[μ] 0) := begin have : ∫⁻ a, f a ∂μ = ∫⁻ a, hf.mk f a ∂μ := lintegral_congr_ae hf.ae_eq_mk, @@ -1374,12 +1374,12 @@ begin exact ⟨λ H, hf.ae_eq_mk.trans H, λ H, hf.ae_eq_mk.symm.trans H⟩ end -lemma lintegral_pos_iff_support {f : α → ennreal} (hf : measurable f) : +lemma lintegral_pos_iff_support {f : α → ℝ≥0∞} (hf : measurable f) : 0 < ∫⁻ a, f a ∂μ ↔ 0 < μ (function.support f) := by simp [pos_iff_ne_zero, hf, filter.eventually_eq, ae_iff, function.support] /-- Weaker version of the monotone convergence theorem-/ -lemma lintegral_supr_ae {f : ℕ → α → ennreal} (hf : ∀n, measurable (f n)) +lemma lintegral_supr_ae {f : ℕ → α → ℝ≥0∞} (hf : ∀n, measurable (f n)) (h_mono : ∀n, ∀ᵐ a ∂μ, f n a ≤ f n.succ a) : (∫⁻ a, ⨆n, f n a ∂μ) = (⨆n, ∫⁻ a, f n a ∂μ) := let ⟨s, hs⟩ := exists_measurable_superset_of_null @@ -1403,7 +1403,7 @@ calc ... = ⨆n, (∫⁻ a, f n a ∂μ) : by simp only [lintegral_congr_ae (g_eq_f.mono $ λ a ha, ha _)] -lemma lintegral_sub {f g : α → ennreal} (hf : measurable f) (hg : measurable g) +lemma lintegral_sub {f g : α → ℝ≥0∞} (hf : measurable f) (hg : measurable g) (hg_fin : ∫⁻ a, g a ∂μ < ⊤) (h_le : g ≤ᵐ[μ] f) : ∫⁻ a, f a - g a ∂μ = ∫⁻ a, f a ∂μ - ∫⁻ a, g a ∂μ := begin @@ -1416,7 +1416,7 @@ end /-- Monotone convergence theorem for nonincreasing sequences of functions -/ lemma lintegral_infi_ae - {f : ℕ → α → ennreal} (h_meas : ∀n, measurable (f n)) + {f : ℕ → α → ℝ≥0∞} (h_meas : ∀n, measurable (f n)) (h_mono : ∀n:ℕ, f n.succ ≤ᵐ[μ] f n) (h_fin : ∫⁻ a, f 0 a ∂μ < ⊤) : ∫⁻ a, ⨅n, f n a ∂μ = ⨅n, ∫⁻ a, f n a ∂μ := have fn_le_f0 : ∫⁻ a, ⨅n, f n a ∂μ ≤ ∫⁻ a, f 0 a ∂μ, from @@ -1452,13 +1452,13 @@ calc /-- Monotone convergence theorem for nonincreasing sequences of functions -/ lemma lintegral_infi - {f : ℕ → α → ennreal} (h_meas : ∀n, measurable (f n)) + {f : ℕ → α → ℝ≥0∞} (h_meas : ∀n, measurable (f n)) (h_mono : ∀ ⦃m n⦄, m ≤ n → f n ≤ f m) (h_fin : ∫⁻ a, f 0 a ∂μ < ⊤) : ∫⁻ a, ⨅n, f n a ∂μ = ⨅n, ∫⁻ a, f n a ∂μ := lintegral_infi_ae h_meas (λ n, ae_of_all _ $ h_mono $ le_of_lt n.lt_succ_self) h_fin /-- Known as Fatou's lemma, version with `ae_measurable` functions -/ -lemma lintegral_liminf_le' {f : ℕ → α → ennreal} (h_meas : ∀n, ae_measurable (f n) μ) : +lemma lintegral_liminf_le' {f : ℕ → α → ℝ≥0∞} (h_meas : ∀n, ae_measurable (f n) μ) : ∫⁻ a, liminf at_top (λ n, f n a) ∂μ ≤ liminf at_top (λ n, ∫⁻ a, f n a ∂μ) := calc ∫⁻ a, liminf at_top (λ n, f n a) ∂μ = ∫⁻ a, ⨆n:ℕ, ⨅i≥n, f i a ∂μ : @@ -1472,11 +1472,11 @@ calc ... = at_top.liminf (λ n, ∫⁻ a, f n a ∂μ) : filter.liminf_eq_supr_infi_of_nat.symm /-- Known as Fatou's lemma -/ -lemma lintegral_liminf_le {f : ℕ → α → ennreal} (h_meas : ∀n, measurable (f n)) : +lemma lintegral_liminf_le {f : ℕ → α → ℝ≥0∞} (h_meas : ∀n, measurable (f n)) : ∫⁻ a, liminf at_top (λ n, f n a) ∂μ ≤ liminf at_top (λ n, ∫⁻ a, f n a ∂μ) := lintegral_liminf_le' (λ n, (h_meas n).ae_measurable) -lemma limsup_lintegral_le {f : ℕ → α → ennreal} {g : α → ennreal} +lemma limsup_lintegral_le {f : ℕ → α → ℝ≥0∞} {g : α → ℝ≥0∞} (hf_meas : ∀ n, measurable (f n)) (h_bound : ∀n, f n ≤ᵐ[μ] g) (h_fin : ∫⁻ a, g a ∂μ < ⊤) : limsup at_top (λn, ∫⁻ a, f n a ∂μ) ≤ ∫⁻ a, limsup at_top (λn, f n a) ∂μ := calc @@ -1498,7 +1498,7 @@ calc /-- Dominated convergence theorem for nonnegative functions -/ lemma tendsto_lintegral_of_dominated_convergence - {F : ℕ → α → ennreal} {f : α → ennreal} (bound : α → ennreal) + {F : ℕ → α → ℝ≥0∞} {f : α → ℝ≥0∞} (bound : α → ℝ≥0∞) (hF_meas : ∀n, measurable (F n)) (h_bound : ∀n, F n ≤ᵐ[μ] bound) (h_fin : ∫⁻ a, bound a ∂μ < ⊤) (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, F n a) at_top (𝓝 (f a))) : @@ -1514,7 +1514,7 @@ tendsto_of_le_liminf_of_limsup_le /-- Dominated convergence theorem for nonnegative functions which are just almost everywhere measurable. -/ lemma tendsto_lintegral_of_dominated_convergence' - {F : ℕ → α → ennreal} {f : α → ennreal} (bound : α → ennreal) + {F : ℕ → α → ℝ≥0∞} {f : α → ℝ≥0∞} (bound : α → ℝ≥0∞) (hF_meas : ∀n, ae_measurable (F n) μ) (h_bound : ∀n, F n ≤ᵐ[μ] bound) (h_fin : ∫⁻ a, bound a ∂μ < ⊤) (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, F n a) at_top (𝓝 (f a))) : @@ -1539,7 +1539,7 @@ end /-- Dominated convergence theorem for filters with a countable basis -/ lemma tendsto_lintegral_filter_of_dominated_convergence {ι} {l : filter ι} - {F : ι → α → ennreal} {f : α → ennreal} (bound : α → ennreal) + {F : ι → α → ℝ≥0∞} {f : α → ℝ≥0∞} (bound : α → ℝ≥0∞) (hl_cb : l.is_countably_generated) (hF_meas : ∀ᶠ n in l, measurable (F n)) (h_bound : ∀ᶠ n in l, ∀ᵐ a ∂μ, F n a ≤ bound a) @@ -1570,7 +1570,7 @@ section open encodable /-- Monotone convergence for a suprema over a directed family and indexed by an encodable type -/ -theorem lintegral_supr_directed [encodable β] {f : β → α → ennreal} +theorem lintegral_supr_directed [encodable β] {f : β → α → ℝ≥0∞} (hf : ∀b, measurable (f b)) (h_directed : directed (≤) f) : ∫⁻ a, ⨆b, f b a ∂μ = ⨆b, ∫⁻ a, f b a ∂μ := begin @@ -1596,7 +1596,7 @@ end end -lemma lintegral_tsum [encodable β] {f : β → α → ennreal} (hf : ∀i, measurable (f i)) : +lemma lintegral_tsum [encodable β] {f : β → α → ℝ≥0∞} (hf : ∀i, measurable (f i)) : ∫⁻ a, ∑' i, f i a ∂μ = ∑' i, ∫⁻ a, f i a ∂μ := begin simp only [ennreal.tsum_eq_supr_sum], @@ -1613,18 +1613,18 @@ end open measure lemma lintegral_Union [encodable β] {s : β → set α} (hm : ∀ i, measurable_set (s i)) - (hd : pairwise (disjoint on s)) (f : α → ennreal) : + (hd : pairwise (disjoint on s)) (f : α → ℝ≥0∞) : ∫⁻ a in ⋃ i, s i, f a ∂μ = ∑' i, ∫⁻ a in s i, f a ∂μ := by simp only [measure.restrict_Union hd hm, lintegral_sum_measure] -lemma lintegral_Union_le [encodable β] (s : β → set α) (f : α → ennreal) : +lemma lintegral_Union_le [encodable β] (s : β → set α) (f : α → ℝ≥0∞) : ∫⁻ a in ⋃ i, s i, f a ∂μ ≤ ∑' i, ∫⁻ a in s i, f a ∂μ := begin rw [← lintegral_sum_measure], exact lintegral_mono' restrict_Union_le (le_refl _) end -lemma lintegral_map [measurable_space β] {f : β → ennreal} {g : α → β} +lemma lintegral_map [measurable_space β] {f : β → ℝ≥0∞} {g : α → β} (hf : measurable f) (hg : measurable g) : ∫⁻ a, f a ∂(map g μ) = ∫⁻ a, f (g a) ∂μ := begin simp only [lintegral_eq_supr_eapprox_lintegral, hf, hf.comp hg], @@ -1634,7 +1634,7 @@ begin { assume s hs, exact map_apply hg hs } }, end -lemma lintegral_map' [measurable_space β] {f : β → ennreal} {g : α → β} +lemma lintegral_map' [measurable_space β] {f : β → ℝ≥0∞} {g : α → β} (hf : ae_measurable f (measure.map g μ)) (hg : measurable g) : ∫⁻ a, f a ∂(measure.map g μ) = ∫⁻ a, f (g a) ∂μ := calc ∫⁻ a, f a ∂(measure.map g μ) = ∫⁻ a, hf.mk f a ∂(measure.map g μ) : @@ -1642,24 +1642,24 @@ calc ∫⁻ a, f a ∂(measure.map g μ) = ∫⁻ a, hf.mk f a ∂(measure.map g ... = ∫⁻ a, hf.mk f (g a) ∂μ : lintegral_map hf.measurable_mk hg ... = ∫⁻ a, f (g a) ∂μ : lintegral_congr_ae (ae_eq_comp hg hf.ae_eq_mk.symm) -lemma lintegral_comp [measurable_space β] {f : β → ennreal} {g : α → β} +lemma lintegral_comp [measurable_space β] {f : β → ℝ≥0∞} {g : α → β} (hf : measurable f) (hg : measurable g) : lintegral μ (f ∘ g) = ∫⁻ a, f a ∂(map g μ) := (lintegral_map hf hg).symm -lemma set_lintegral_map [measurable_space β] {f : β → ennreal} {g : α → β} +lemma set_lintegral_map [measurable_space β] {f : β → ℝ≥0∞} {g : α → β} {s : set β} (hs : measurable_set s) (hf : measurable f) (hg : measurable g) : ∫⁻ y in s, f y ∂(map g μ) = ∫⁻ x in g ⁻¹' s, f (g x) ∂μ := by rw [restrict_map hg hs, lintegral_map hf hg] -lemma lintegral_dirac' (a : α) {f : α → ennreal} (hf : measurable f) : +lemma lintegral_dirac' (a : α) {f : α → ℝ≥0∞} (hf : measurable f) : ∫⁻ a, f a ∂(dirac a) = f a := by simp [lintegral_congr_ae (ae_eq_dirac' hf)] -lemma lintegral_dirac [measurable_singleton_class α] (a : α) (f : α → ennreal) : +lemma lintegral_dirac [measurable_singleton_class α] (a : α) (f : α → ℝ≥0∞) : ∫⁻ a, f a ∂(dirac a) = f a := by simp [lintegral_congr_ae (ae_eq_dirac f)] -lemma lintegral_count' {f : α → ennreal} (hf : measurable f) : +lemma lintegral_count' {f : α → ℝ≥0∞} (hf : measurable f) : ∫⁻ a, f a ∂count = ∑' a, f a := begin rw [count, lintegral_sum_measure], @@ -1667,7 +1667,7 @@ begin exact funext (λ a, lintegral_dirac' a hf), end -lemma lintegral_count [measurable_singleton_class α] (f : α → ennreal) : +lemma lintegral_count [measurable_singleton_class α] (f : α → ℝ≥0∞) : ∫⁻ a, f a ∂count = ∑' a, f a := begin rw [count, lintegral_sum_measure], @@ -1675,7 +1675,7 @@ begin exact funext (λ a, lintegral_dirac a f), end -lemma ae_lt_top {f : α → ennreal} (hf : measurable f) (h2f : ∫⁻ x, f x ∂μ < ⊤) : +lemma ae_lt_top {f : α → ℝ≥0∞} (hf : measurable f) (h2f : ∫⁻ x, f x ∂μ < ⊤) : ∀ᵐ x ∂μ, f x < ⊤ := begin simp_rw [ae_iff, ennreal.not_lt_top], by_contra h, rw [← not_le] at h2f, apply h2f, @@ -1685,12 +1685,12 @@ begin rw [lintegral_indicator _ (hf (measurable_set_singleton ⊤))], simp [ennreal.top_mul, preimage, h] end -/-- Given a measure `μ : measure α` and a function `f : α → ennreal`, `μ.with_density f` is the +/-- Given a measure `μ : measure α` and a function `f : α → ℝ≥0∞`, `μ.with_density f` is the measure such that for a measurable set `s` we have `μ.with_density f s = ∫⁻ a in s, f a ∂μ`. -/ -def measure.with_density (μ : measure α) (f : α → ennreal) : measure α := +def measure.with_density (μ : measure α) (f : α → ℝ≥0∞) : measure α := measure.of_measurable (λs hs, ∫⁻ a in s, f a ∂μ) (by simp) (λ s hs hd, lintegral_Union hs hd _) -@[simp] lemma with_density_apply (f : α → ennreal) {s : set α} (hs : measurable_set s) : +@[simp] lemma with_density_apply (f : α → ℝ≥0∞) {s : set α} (hs : measurable_set s) : μ.with_density f s = ∫⁻ a in s, f a ∂μ := measure.of_measurable_apply s hs @@ -1699,7 +1699,7 @@ end lintegral end measure_theory open measure_theory measure_theory.simple_func -/-- To prove something for an arbitrary measurable function into `ennreal`, it suffices to show +/-- To prove something for an arbitrary measurable function into `ℝ≥0∞`, it suffices to show that the property holds for (multiples of) characteristic functions and is closed under addition and supremum of increasing sequences of functions. @@ -1708,13 +1708,13 @@ can be added once we need them (for example in `h_sum` it is only necessary to c a simple function with a multiple of a characteristic function and that the intersection of their images is a subset of `{0}`. -/ @[elab_as_eliminator] -theorem measurable.ennreal_induction {α} [measurable_space α] {P : (α → ennreal) → Prop} - (h_ind : ∀ (c : ennreal) ⦃s⦄, measurable_set s → P (indicator s (λ _, c))) - (h_sum : ∀ ⦃f g : α → ennreal⦄, set.univ ⊆ f ⁻¹' {0} ∪ g ⁻¹' {0} → measurable f → measurable g → +theorem measurable.ennreal_induction {α} [measurable_space α] {P : (α → ℝ≥0∞) → Prop} + (h_ind : ∀ (c : ℝ≥0∞) ⦃s⦄, measurable_set s → P (indicator s (λ _, c))) + (h_sum : ∀ ⦃f g : α → ℝ≥0∞⦄, set.univ ⊆ f ⁻¹' {0} ∪ g ⁻¹' {0} → measurable f → measurable g → P f → P g → P (f + g)) - (h_supr : ∀ ⦃f : ℕ → α → ennreal⦄ (hf : ∀n, measurable (f n)) (h_mono : monotone f) + (h_supr : ∀ ⦃f : ℕ → α → ℝ≥0∞⦄ (hf : ∀n, measurable (f n)) (h_mono : monotone f) (hP : ∀ n, P (f n)), P (λ x, ⨆ n, f n x)) - ⦃f : α → ennreal⦄ (hf : measurable f) : P f := + ⦃f : α → ℝ≥0∞⦄ (hf : measurable f) : P f := begin convert h_supr (λ n, (eapprox f n).measurable) (monotone_eapprox f) _, { ext1 x, rw [supr_eapprox_apply f hf] }, @@ -1734,7 +1734,7 @@ of [wasserman2004]). Thus, this method shows how to one can calculate expectatio and other moments as a function of the probability density function. -/ lemma lintegral_with_density_eq_lintegral_mul {α} [measurable_space α] (μ : measure α) - {f : α → ennreal} (h_mf : measurable f) : ∀ {g : α → ennreal}, measurable g → + {f : α → ℝ≥0∞} (h_mf : measurable f) : ∀ {g : α → ℝ≥0∞}, measurable g → ∫⁻ a, g a ∂(μ.with_density f) = ∫⁻ a, (f * g) a ∂μ := begin apply measurable.ennreal_induction, diff --git a/src/measure_theory/interval_integral.lean b/src/measure_theory/interval_integral.lean index 0668d33c581e4..d4aa9754ca55c 100644 --- a/src/measure_theory/interval_integral.lean +++ b/src/measure_theory/interval_integral.lean @@ -128,7 +128,7 @@ noncomputable theory open topological_space (second_countable_topology) open measure_theory set classical filter -open_locale classical topological_space filter +open_locale classical topological_space filter ennreal variables {α β 𝕜 E F : Type*} [linear_order α] [measurable_space α] [measurable_space E] [normed_group E] @@ -361,7 +361,7 @@ lemma integral_const {a b : ℝ} (c : E) : (∫ (x : ℝ) in a..b, c) = (b - a) by simp only [integral_const', real.volume_Ioc, ennreal.to_real_of_real', ← neg_sub b, max_zero_sub_eq_self] -lemma integral_smul_measure (c : ennreal) : +lemma integral_smul_measure (c : ℝ≥0∞) : ∫ x in a..b, f x ∂(c • μ) = c.to_real • ∫ x in a..b, f x ∂μ := by simp only [interval_integral, measure.restrict_smul, integral_smul_measure, smul_sub] diff --git a/src/measure_theory/l1_space.lean b/src/measure_theory/l1_space.lean index 655c3b8eea045..ad19aa9d577f9 100644 --- a/src/measure_theory/l1_space.lean +++ b/src/measure_theory/l1_space.lean @@ -56,7 +56,7 @@ integrable, function space, l1 -/ noncomputable theory -open_locale classical topological_space big_operators +open_locale classical topological_space big_operators ennreal open set filter topological_space ennreal emetric measure_theory @@ -195,7 +195,7 @@ h.mono_measure $ measure.le_add_left $ le_refl _ has_finite_integral f (μ + ν) ↔ has_finite_integral f μ ∧ has_finite_integral f ν := ⟨λ h, ⟨h.left_of_add_measure, h.right_of_add_measure⟩, λ h, h.1.add_measure h.2⟩ -lemma has_finite_integral.smul_measure {f : α → β} (h : has_finite_integral f μ) {c : ennreal} +lemma has_finite_integral.smul_measure {f : α → β} (h : has_finite_integral f μ) {c : ℝ≥0∞} (hc : c < ⊤) : has_finite_integral f (c • μ) := begin simp only [has_finite_integral, lintegral_smul_measure] at *, @@ -220,7 +220,7 @@ by simpa [has_finite_integral] using hfi lemma has_finite_integral.norm {f : α → β} (hfi : has_finite_integral f μ) : has_finite_integral (λa, ∥f a∥) μ := -have eq : (λa, (nnnorm ∥f a∥ : ennreal)) = λa, (nnnorm (f a) : ennreal), +have eq : (λa, (nnnorm ∥f a∥ : ℝ≥0∞)) = λa, (nnnorm (f a) : ℝ≥0∞), by { funext, rw nnnorm_norm }, by { rwa [has_finite_integral, eq] } @@ -438,7 +438,7 @@ h.mono_measure $ measure.le_add_left $ le_refl _ integrable f (μ + ν) ↔ integrable f μ ∧ integrable f ν := ⟨λ h, ⟨h.left_of_add_measure, h.right_of_add_measure⟩, λ h, h.1.add_measure h.2⟩ -lemma integrable.smul_measure {f : α → β} (h : integrable f μ) {c : ennreal} (hc : c < ⊤) : +lemma integrable.smul_measure {f : α → β} (h : integrable f μ) {c : ℝ≥0∞} (hc : c < ⊤) : integrable f (c • μ) := ⟨h.ae_measurable.smul_measure c, h.has_finite_integral.smul_measure hc⟩ @@ -557,7 +557,7 @@ begin simp_rw [integrable, ae_measurable_smul_const hc, and.congr_right_iff, has_finite_integral, nnnorm_smul, ennreal.coe_mul], intro hf, rw [lintegral_mul_const' _ _ ennreal.coe_ne_top, ennreal.mul_lt_top_iff], - have : ∀ x : ennreal, x = 0 → x < ⊤ := by simp, + have : ∀ x : ℝ≥0∞, x = 0 → x < ⊤ := by simp, simp [hc, or_iff_left_of_imp (this _)] end end normed_space_over_complete_field @@ -819,14 +819,14 @@ variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β] lemma smul_to_fun (c : 𝕜) (f : α →₁[μ] β) : ⇑(c • f) =ᵐ[μ] c • f := ae_eq_fun.coe_fn_smul _ _ -lemma norm_eq_lintegral (f : α →₁[μ] β) : ∥f∥ = (∫⁻ x, (nnnorm (f x) : ennreal) ∂μ).to_real := +lemma norm_eq_lintegral (f : α →₁[μ] β) : ∥f∥ = (∫⁻ x, (nnnorm (f x) : ℝ≥0∞) ∂μ).to_real := by simp [l1.norm_eq, ae_eq_fun.edist_zero_eq_coe, ← edist_eq_coe_nnnorm] /-- Computing the norm of a difference between two L¹-functions. Note that this is not a special case of `norm_eq_lintegral` since `(f - g) x` and `f x - g x` are not equal (but only a.e.-equal). -/ lemma norm_sub_eq_lintegral (f g : α →₁[μ] β) : - ∥f - g∥ = (∫⁻ x, (nnnorm (f x - g x) : ennreal) ∂μ).to_real := + ∥f - g∥ = (∫⁻ x, (nnnorm (f x - g x) : ℝ≥0∞) ∂μ).to_real := begin simp_rw [l1.norm_eq, ae_eq_fun.edist_zero_eq_coe, ← edist_eq_coe_nnnorm], rw lintegral_congr_ae, @@ -835,7 +835,7 @@ begin end lemma of_real_norm_eq_lintegral (f : α →₁[μ] β) : - ennreal.of_real ∥f∥ = ∫⁻ x, (nnnorm (f x) : ennreal) ∂μ := + ennreal.of_real ∥f∥ = ∫⁻ x, (nnnorm (f x) : ℝ≥0∞) ∂μ := by { rw [norm_eq_lintegral, ennreal.of_real_to_real], rw [← ennreal.lt_top_iff_ne_top], exact f.has_finite_integral } @@ -843,7 +843,7 @@ by { rw [norm_eq_lintegral, ennreal.of_real_to_real], rw [← ennreal.lt_top_iff special case of `of_real_norm_eq_lintegral` since `(f - g) x` and `f x - g x` are not equal (but only a.e.-equal). -/ lemma of_real_norm_sub_eq_lintegral (f g : α →₁[μ] β) : - ennreal.of_real ∥f - g∥ = ∫⁻ x, (nnnorm (f x - g x) : ennreal) ∂μ := + ennreal.of_real ∥f - g∥ = ∫⁻ x, (nnnorm (f x - g x) : ℝ≥0∞) ∂μ := begin simp_rw [of_real_norm_eq_lintegral, ← edist_eq_coe_nnnorm], apply lintegral_congr_ae, diff --git a/src/measure_theory/lebesgue_measure.lean b/src/measure_theory/lebesgue_measure.lean index c43c3ad308b9d..cf59be0d14022 100644 --- a/src/measure_theory/lebesgue_measure.lean +++ b/src/measure_theory/lebesgue_measure.lean @@ -22,7 +22,7 @@ namespace measure_theory /-- Length of an interval. This is the largest monotonic function which correctly measures all intervals. -/ -def lebesgue_length (s : set ℝ) : ennreal := ⨅a b (h : s ⊆ Ico a b), of_real (b - a) +def lebesgue_length (s : set ℝ) : ℝ≥0∞ := ⨅a b (h : s ⊆ Ico a b), of_real (b - a) @[simp] lemma lebesgue_length_empty : lebesgue_length ∅ = 0 := nonpos_iff_eq_zero.1 $ infi_le_of_le 0 $ infi_le_of_le 0 $ by simp @@ -103,11 +103,11 @@ outer_measure.of_function_le _ lemma lebesgue_length_subadditive {a b : ℝ} {c d : ℕ → ℝ} (ss : Icc a b ⊆ ⋃i, Ioo (c i) (d i)) : - (of_real (b - a) : ennreal) ≤ ∑' i, of_real (d i - c i) := + (of_real (b - a) : ℝ≥0∞) ≤ ∑' i, of_real (d i - c i) := begin suffices : ∀ (s:finset ℕ) b (cv : Icc a b ⊆ ⋃ i ∈ (↑s:set ℕ), Ioo (c i) (d i)), - (of_real (b - a) : ennreal) ≤ ∑ i in s, of_real (d i - c i), + (of_real (b - a) : ℝ≥0∞) ≤ ∑ i in s, of_real (d i - c i), { rcases compact_Icc.elim_finite_subcover_image (λ (i : ℕ) (_ : i ∈ univ), @is_open_Ioo _ _ _ _ (c i) (d i)) (by simpa using ss) with ⟨s, su, hf, hs⟩, have e : (⋃ i ∈ (↑hf.to_finset:set ℕ), @@ -143,7 +143,7 @@ begin refine le_trans _ (add_le_add_left (le_of_lt hε) _), rw ← ennreal.tsum_add, choose g hg using show - ∀ i, ∃ p:ℝ×ℝ, f i ⊆ Ioo p.1 p.2 ∧ (of_real (p.2 - p.1) : ennreal) < + ∀ i, ∃ p:ℝ×ℝ, f i ⊆ Ioo p.1 p.2 ∧ (of_real (p.2 - p.1) : ℝ≥0∞) < lebesgue_length (f i) + ε' i, { intro i, have := (ennreal.lt_add_right (lt_of_le_of_lt (ennreal.le_tsum i) h) @@ -266,7 +266,7 @@ by rw [interval, volume_Icc, max_sub_min_eq_abs] @[simp] lemma volume_Ioi {a : ℝ} : volume (Ioi a) = ∞ := top_unique $ le_of_tendsto' ennreal.tendsto_nat_nhds_top $ λ n, -calc (n : ennreal) = volume (Ioo a (a + n)) : by simp +calc (n : ℝ≥0∞) = volume (Ioo a (a + n)) : by simp ... ≤ volume (Ioi a) : measure_mono Ioo_subset_Ioi_self @[simp] lemma volume_Ici {a : ℝ} : volume (Ici a) = ∞ := @@ -274,7 +274,7 @@ by simp [← measure_congr Ioi_ae_eq_Ici] @[simp] lemma volume_Iio {a : ℝ} : volume (Iio a) = ∞ := top_unique $ le_of_tendsto' ennreal.tendsto_nat_nhds_top $ λ n, -calc (n : ennreal) = volume (Ioo (a - n) a) : by simp +calc (n : ℝ≥0∞) = volume (Ioo (a - n) a) : by simp ... ≤ volume (Iio a) : measure_mono Ioo_subset_Iio_self @[simp] lemma volume_Iic {a : ℝ} : volume (Iic a) = ∞ := @@ -373,7 +373,7 @@ end real open_locale topological_space lemma filter.eventually.volume_pos_of_nhds_real {p : ℝ → Prop} {a : ℝ} (h : ∀ᶠ x in 𝓝 a, p x) : - (0 : ennreal) < volume {x | p x} := + (0 : ℝ≥0∞) < volume {x | p x} := begin rcases h.exists_Ioo_subset with ⟨l, u, hx, hs⟩, refine lt_of_lt_of_le _ (measure_mono hs), diff --git a/src/measure_theory/lp_space.lean b/src/measure_theory/lp_space.lean index fb453c38d20fa..f7322f0b524af 100644 --- a/src/measure_theory/lp_space.lean +++ b/src/measure_theory/lp_space.lean @@ -11,7 +11,7 @@ import analysis.mean_inequalities # ℒp space and Lp space This file describes properties of almost everywhere measurable functions with finite seminorm, -denoted by `snorm f p μ` and defined for `p:ennreal` as `0` if `p=0`, `(∫ ∥f a∥^p ∂μ) ^ (1/p)` for +denoted by `snorm f p μ` and defined for `p:ℝ≥0∞` as `0` if `p=0`, `(∫ ∥f a∥^p ∂μ) ^ (1/p)` for `0 < p < ∞` and `ess_sup ∥f∥ μ` for `p=∞`. The Prop-valued `mem_ℒp f p μ` states that a function `f : α → E` has finite seminorm. @@ -25,7 +25,7 @@ TODO: prove that Lp is complete. * `snorm' f p μ` : `(∫ ∥f a∥^p ∂μ) ^ (1/p)` for `f : α → F` and `p : ℝ`, where `α` is a measurable space and `F` is a normed group. * `snorm_ess_sup f μ` : seminorm in `ℒ∞`, equal to the essential supremum `ess_sup ∥f∥ μ`. -* `snorm f p μ` : for `p : ennreal`, seminorm in `ℒp`, equal to `0` for `p=0`, to `snorm' f p μ` +* `snorm f p μ` : for `p : ℝ≥0∞`, seminorm in `ℒp`, equal to `0` for `p=0`, to `snorm' f p μ` for `0 < p < ∞` and to `snorm_ess_sup f μ` for `p = ∞`. * `mem_ℒp f p μ` : property that the function `f` is almost everywhere measurable and has finite @@ -37,6 +37,8 @@ TODO: prove that Lp is complete. noncomputable theory +open_locale ennreal + namespace measure_theory section ℒp @@ -44,20 +46,20 @@ section ℒp variables {α E F : Type*} [measurable_space α] {μ : measure α} [measurable_space E] [normed_group E] [normed_group F] - {p : ℝ} {q : ennreal} + {p : ℝ} {q : ℝ≥0∞} section ℒp_space_definition /-- `(∫ ∥f a∥^p ∂μ) ^ (1/p)`, which is a seminorm on the space of measurable functions for which this quantity is finite -/ -def snorm' (f : α → F) (p : ℝ) (μ : measure α) : ennreal := (∫⁻ a, (nnnorm (f a))^p ∂μ) ^ (1/p) +def snorm' (f : α → F) (p : ℝ) (μ : measure α) : ℝ≥0∞ := (∫⁻ a, (nnnorm (f a))^p ∂μ) ^ (1/p) /-- seminorm for `ℒ∞`, equal to the essential supremum of `∥f∥`. -/ -def snorm_ess_sup (f : α → F) (μ : measure α) := ess_sup (λ x, (nnnorm (f x) : ennreal)) μ +def snorm_ess_sup (f : α → F) (μ : measure α) := ess_sup (λ x, (nnnorm (f x) : ℝ≥0∞)) μ /-- `ℒp` seminorm, equal to `0` for `p=0`, to `(∫ ∥f a∥^p ∂μ) ^ (1/p)` for `0 < p < ∞` and to `ess_sup ∥f∥ μ` for `p = ∞`. -/ -def snorm (f : α → F) (q : ennreal) (μ : measure α) : ennreal := +def snorm (f : α → F) (q : ℝ≥0∞) (μ : measure α) : ℝ≥0∞ := if q = 0 then 0 else (if q = ⊤ then snorm_ess_sup f μ else snorm' f (ennreal.to_real q) μ) lemma snorm_eq_snorm' (hq_ne_zero : q ≠ 0) (hq_ne_top : q ≠ ⊤) {f : α → F} : @@ -67,7 +69,7 @@ by simp [snorm, hq_ne_zero, hq_ne_top] @[simp] lemma snorm_exponent_top {f : α → F} : snorm f ⊤ μ = snorm_ess_sup f μ := by simp [snorm] /-- The property that `f:α→E` is ae_measurable and `(∫ ∥f a∥^p ∂μ)^(1/p)` is finite -/ -def mem_ℒp (f : α → E) (p : ennreal) (μ : measure α) : Prop := +def mem_ℒp (f : α → E) (p : ℝ≥0∞) (μ : measure α) : Prop := ae_measurable f μ ∧ snorm f p μ < ⊤ lemma lintegral_rpow_nnnorm_eq_rpow_snorm' {f : α → F} (hp0_lt : 0 < p) : @@ -168,7 +170,7 @@ end zero section const lemma snorm'_const (c : F) (hp_pos : 0 < p) : - snorm' (λ x : α , c) p μ = (nnnorm c : ennreal) * (μ set.univ) ^ (1/p) := + snorm' (λ x : α , c) p μ = (nnnorm c : ℝ≥0∞) * (μ set.univ) ^ (1/p) := begin rw [snorm', lintegral_const, @ennreal.mul_rpow_of_nonneg _ _ (1/p) (by simp [le_of_lt hp_pos])], congr, @@ -178,7 +180,7 @@ begin end lemma snorm'_const' [finite_measure μ] (c : F) (hc_ne_zero : c ≠ 0) (hp_ne_zero : p ≠ 0) : - snorm' (λ x : α , c) p μ = (nnnorm c : ennreal) * (μ set.univ) ^ (1/p) := + snorm' (λ x : α , c) p μ = (nnnorm c : ℝ≥0∞) * (μ set.univ) ^ (1/p) := begin rw [snorm', lintegral_const, ennreal.mul_rpow_of_ne_top _ (measure_ne_top μ set.univ)], { congr, @@ -193,15 +195,15 @@ begin end lemma snorm_ess_sup_const (c : F) (hμ : μ ≠ 0) : - snorm_ess_sup (λ x : α, c) μ = (nnnorm c : ennreal) := + snorm_ess_sup (λ x : α, c) μ = (nnnorm c : ℝ≥0∞) := by rw [snorm_ess_sup, ess_sup_const _ hμ] lemma snorm'_const_of_probability_measure (c : F) (hp_pos : 0 < p) [probability_measure μ] : - snorm' (λ x : α , c) p μ = (nnnorm c : ennreal) := + snorm' (λ x : α , c) p μ = (nnnorm c : ℝ≥0∞) := by simp [snorm'_const c hp_pos, measure_univ] lemma snorm_const (c : F) (h0 : q ≠ 0) (hμ : μ ≠ 0) : - snorm (λ x : α , c) q μ = (nnnorm c : ennreal) * (μ set.univ) ^ (1/(ennreal.to_real q)) := + snorm (λ x : α , c) q μ = (nnnorm c : ℝ≥0∞) * (μ set.univ) ^ (1/(ennreal.to_real q)) := begin by_cases h_top : q = ⊤, { simp [h_top, snorm_ess_sup_const c hμ], }, @@ -210,7 +212,7 @@ begin end lemma snorm_const' (c : F) (h0 : q ≠ 0) (h_top: q ≠ ⊤) : - snorm (λ x : α , c) q μ = (nnnorm c : ennreal) * (μ set.univ) ^ (1/(ennreal.to_real q)) := + snorm (λ x : α , c) q μ = (nnnorm c : ℝ≥0∞) * (μ set.univ) ^ (1/(ennreal.to_real q)) := begin simp [snorm_eq_snorm' h0 h_top, snorm'_const, ennreal.to_real_pos_iff.mpr ⟨lt_of_le_of_ne (zero_le _) h0.symm, h_top⟩], @@ -299,8 +301,8 @@ lemma snorm'_eq_zero_iff (hp0_lt : 0 < p) {f : α → E} (hf : ae_measurable f ⟨ae_eq_zero_of_snorm'_eq_zero (le_of_lt hp0_lt) hf, snorm'_eq_zero_of_ae_zero hp0_lt⟩ lemma coe_nnnorm_ae_le_snorm_ess_sup (f : α → F) (μ : measure α) : - ∀ᵐ x ∂μ, (nnnorm (f x) : ennreal) ≤ snorm_ess_sup f μ := -ennreal.ae_le_ess_sup (λ x, (nnnorm (f x) : ennreal)) + ∀ᵐ x ∂μ, (nnnorm (f x) : ℝ≥0∞) ≤ snorm_ess_sup f μ := +ennreal.ae_le_ess_sup (λ x, (nnnorm (f x) : ℝ≥0∞)) lemma snorm_ess_sup_eq_zero_iff {f : α → F} : snorm_ess_sup f μ = 0 ↔ f =ᵐ[μ] 0 := begin @@ -349,7 +351,7 @@ begin { rw [hpq_eq, sub_self, ennreal.rpow_zero, mul_one], exact le_refl _, }, have hpq : p < q, from lt_of_le_of_ne hpq hpq_eq, - let g := λ a : α, (1 : ennreal), + let g := λ a : α, (1 : ℝ≥0∞), have h_rw : ∫⁻ a, ↑(nnnorm (f a))^p ∂ μ = ∫⁻ a, (nnnorm (f a) * (g a))^p ∂ μ, from lintegral_congr (λ a, by simp), repeat {rw snorm'}, @@ -394,7 +396,7 @@ lemma snorm'_le_snorm_ess_sup (hp_pos : 0 < p) {f : α → F} [probability_measu snorm' f p μ ≤ snorm_ess_sup f μ := le_trans (snorm'_le_snorm_ess_sup_mul_rpow_measure_univ hp_pos) (le_of_eq (by simp [measure_univ])) -lemma snorm_le_snorm_of_exponent_le {p q : ennreal} (hpq : p ≤ q) [probability_measure μ] +lemma snorm_le_snorm_of_exponent_le {p q : ℝ≥0∞} (hpq : p ≤ q) [probability_measure μ] {f : α → E} (hf : ae_measurable f μ) : snorm f p μ ≤ snorm f q μ := begin @@ -447,7 +449,7 @@ begin end end -lemma mem_ℒp.mem_ℒp_of_exponent_le {p q : ennreal} [finite_measure μ] {f : α → E} +lemma mem_ℒp.mem_ℒp_of_exponent_le {p q : ℝ≥0∞} [finite_measure μ] {f : α → E} (hfq : mem_ℒp f q μ) (hpq : p ≤ q) : mem_ℒp f p μ := begin @@ -488,8 +490,8 @@ mem_ℒp_one_iff_integrable.mp (hfq.mem_ℒp_of_exponent_le hq1) lemma snorm'_add_le {f g : α → E} (hf : ae_measurable f μ) (hg : ae_measurable g μ) (hp1 : 1 ≤ p) : snorm' (f + g) p μ ≤ snorm' f p μ + snorm' g p μ := calc (∫⁻ a, ↑(nnnorm ((f + g) a)) ^ p ∂μ) ^ (1 / p) - ≤ (∫⁻ a, (((λ a, (nnnorm (f a) : ennreal)) - + (λ a, (nnnorm (g a) : ennreal))) a) ^ p ∂μ) ^ (1 / p) : + ≤ (∫⁻ a, (((λ a, (nnnorm (f a) : ℝ≥0∞)) + + (λ a, (nnnorm (g a) : ℝ≥0∞))) a) ^ p ∂μ) ^ (1 / p) : begin refine @ennreal.rpow_le_rpow _ _ (1/p) _ (by simp [le_trans zero_le_one hp1]), refine lintegral_mono (λ a, ennreal.rpow_le_rpow _ (le_trans zero_le_one hp1)), @@ -529,14 +531,14 @@ lemma snorm'_add_lt_top_of_le_one {f g : α → E} (hf : ae_measurable f μ) (hg (hf_snorm : snorm' f p μ < ⊤) (hg_snorm : snorm' g p μ < ⊤) (hp_pos : 0 < p) (hp1 : p ≤ 1) : snorm' (f + g) p μ < ⊤ := calc (∫⁻ a, ↑(nnnorm ((f + g) a)) ^ p ∂μ) ^ (1 / p) - ≤ (∫⁻ a, (((λ a, (nnnorm (f a) : ennreal)) - + (λ a, (nnnorm (g a) : ennreal))) a) ^ p ∂μ) ^ (1 / p) : + ≤ (∫⁻ a, (((λ a, (nnnorm (f a) : ℝ≥0∞)) + + (λ a, (nnnorm (g a) : ℝ≥0∞))) a) ^ p ∂μ) ^ (1 / p) : begin refine @ennreal.rpow_le_rpow _ _ (1/p) _ (by simp [hp_pos.le]), refine lintegral_mono (λ a, ennreal.rpow_le_rpow _ hp_pos.le), simp [←ennreal.coe_add, nnnorm_add_le], end -... ≤ (∫⁻ a, (nnnorm (f a) : ennreal) ^ p + (nnnorm (g a) : ennreal) ^ p ∂μ) ^ (1 / p) : +... ≤ (∫⁻ a, (nnnorm (f a) : ℝ≥0∞) ^ p + (nnnorm (g a) : ℝ≥0∞) ^ p ∂μ) ^ (1 / p) : begin refine @ennreal.rpow_le_rpow _ _ (1/p) (lintegral_mono (λ a, _)) (by simp [hp_pos.le]), exact ennreal.rpow_add_le_add_rpow _ _ hp_pos hp1, @@ -588,13 +590,13 @@ section normed_space variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 E] [normed_space 𝕜 F] lemma snorm'_const_smul {f : α → F} (c : 𝕜) (hp0_lt : 0 < p) : - snorm' (c • f) p μ = (nnnorm c : ennreal) * snorm' f p μ := + snorm' (c • f) p μ = (nnnorm c : ℝ≥0∞) * snorm' f p μ := begin rw snorm', simp_rw [pi.smul_apply, nnnorm_smul, ennreal.coe_mul, ennreal.mul_rpow_of_nonneg _ _ (le_of_lt hp0_lt)], suffices h_integral : ∫⁻ a, ↑(nnnorm c) ^ p * ↑(nnnorm (f a)) ^ p ∂μ - = (nnnorm c : ennreal)^p * ∫⁻ a, (nnnorm (f a)) ^ p ∂μ, + = (nnnorm c : ℝ≥0∞)^p * ∫⁻ a, (nnnorm (f a)) ^ p ∂μ, { apply_fun (λ x, x ^ (1/p)) at h_integral, rw [h_integral, @ennreal.mul_rpow_of_nonneg _ _ (1/p) (by simp [le_of_lt hp0_lt])], congr, @@ -606,11 +608,11 @@ begin end lemma snorm_ess_sup_const_smul {f : α → F} (c : 𝕜) : - snorm_ess_sup (c • f) μ = (nnnorm c : ennreal) * snorm_ess_sup f μ := + snorm_ess_sup (c • f) μ = (nnnorm c : ℝ≥0∞) * snorm_ess_sup f μ := by simp_rw [snorm_ess_sup, pi.smul_apply, nnnorm_smul, ennreal.coe_mul, ennreal.ess_sup_const_mul] lemma snorm_const_smul {f : α → F} (c : 𝕜) : - snorm (c • f) q μ = (nnnorm c : ennreal) * snorm f q μ := + snorm (c • f) q μ = (nnnorm c : ℝ≥0∞) * snorm f q μ := begin by_cases h0 : q = 0, { simp [h0], }, @@ -648,19 +650,19 @@ The space of equivalence classes of measurable functions for which `snorm f p μ -/ @[simp] lemma snorm_ae_eq_fun {α E : Type*} [measurable_space α] {μ : measure α} - [measurable_space E] [normed_group E] {p : ennreal} {f : α → E} (hf : ae_measurable f μ) : + [measurable_space E] [normed_group E] {p : ℝ≥0∞} {f : α → E} (hf : ae_measurable f μ) : snorm (ae_eq_fun.mk f hf) p μ = snorm f p μ := snorm_congr_ae (ae_eq_fun.coe_fn_mk _ _) lemma mem_ℒp.snorm_mk_lt_top {α E : Type*} [measurable_space α] {μ : measure α} - [measurable_space E] [normed_group E] {p : ennreal} {f : α → E} (hfp : mem_ℒp f p μ) : + [measurable_space E] [normed_group E] {p : ℝ≥0∞} {f : α → E} (hfp : mem_ℒp f p μ) : snorm (ae_eq_fun.mk f hfp.1) p μ < ⊤ := by simp [hfp.2] /-- Lp space -/ def Lp {α} (E : Type*) [measurable_space α] [measurable_space E] [normed_group E] [borel_space E] [topological_space.second_countable_topology E] - (p : ennreal) (μ : measure α) : add_subgroup (α →ₘ[μ] E) := + (p : ℝ≥0∞) (μ : measure α) : add_subgroup (α →ₘ[μ] E) := { carrier := {f | snorm f p μ < ⊤}, zero_mem' := by simp [snorm_congr_ae ae_eq_fun.coe_fn_zero, snorm_zero], add_mem' := λ f g hf hg, by simp [snorm_congr_ae (ae_eq_fun.coe_fn_add _ _), @@ -671,22 +673,22 @@ def Lp {α} (E : Type*) [measurable_space α] [measurable_space E] [normed_group /-- make an element of Lp from a function verifying `mem_ℒp` -/ def mem_ℒp.to_Lp {α E} [measurable_space α] [measurable_space E] [normed_group E] [borel_space E] [topological_space.second_countable_topology E] - (f : α → E) {p : ennreal} {μ : measure α} (h_mem_ℒp : mem_ℒp f p μ) : Lp E p μ := + (f : α → E) {p : ℝ≥0∞} {μ : measure α} (h_mem_ℒp : mem_ℒp f p μ) : Lp E p μ := ⟨ae_eq_fun.mk f h_mem_ℒp.1, h_mem_ℒp.snorm_mk_lt_top⟩ lemma mem_ℒp.coe_fn_to_Lp {α E} [measurable_space α] [measurable_space E] [normed_group E] - [borel_space E] [topological_space.second_countable_topology E] {μ : measure α} {p : ennreal} + [borel_space E] [topological_space.second_countable_topology E] {μ : measure α} {p : ℝ≥0∞} {f : α → E} (hf : mem_ℒp f p μ) : hf.to_Lp f =ᵐ[μ] f := ae_eq_fun.coe_fn_mk _ _ namespace Lp variables {α E F : Type*} [measurable_space α] {μ : measure α} [measurable_space E] [normed_group E] - [borel_space E] [topological_space.second_countable_topology E] {p : ennreal} + [borel_space E] [topological_space.second_countable_topology E] {p : ℝ≥0∞} lemma mem_Lp_iff_snorm_lt_top {f : α →ₘ[μ] E} : f ∈ Lp E p μ ↔ snorm f p μ < ⊤ := iff.refl _ -lemma antimono [finite_measure μ] {p q : ennreal} (hpq : p ≤ q) : Lp E q μ ≤ Lp E p μ := +lemma antimono [finite_measure μ] {p q : ℝ≥0∞} (hpq : p ≤ q) : Lp E q μ ≤ Lp E p μ := λ f hf, (mem_ℒp.mem_ℒp_of_exponent_le ⟨f.ae_measurable, hf⟩ hpq).2 lemma coe_fn_mk {f : α →ₘ[μ] E} (hf : snorm f p μ < ⊤) : ⇑(⟨f, hf⟩ : Lp E p μ) =ᵐ[μ] f := diff --git a/src/measure_theory/measure_space.lean b/src/measure_theory/measure_space.lean index cf71539d58267..d8776985e91e7 100644 --- a/src/measure_theory/measure_space.lean +++ b/src/measure_theory/measure_space.lean @@ -23,7 +23,7 @@ In this file a measure is defined to be an outer measure that is countably addit measurable sets, with the additional assumption that the outer measure is the canonical extension of the restricted measure. -Measures on `α` form a complete lattice, and are closed under scalar multiplication with `ennreal`. +Measures on `α` form a complete lattice, and are closed under scalar multiplication with `ℝ≥0∞`. We introduce the following typeclasses for measures: @@ -85,7 +85,7 @@ measure, almost everywhere, measure space, completion, null set, null measurable noncomputable theory open classical set filter (hiding map) function measurable_space -open_locale classical topological_space big_operators filter +open_locale classical topological_space big_operators filter ennreal variables {α β γ δ ι : Type*} @@ -107,7 +107,7 @@ But we can extend this to _all_ sets, but using the outer measure. This gives us subadditivity for all sets. -/ instance measure.has_coe_to_fun [measurable_space α] : has_coe_to_fun (measure α) := -⟨λ _, set α → ennreal, λ m, m.to_outer_measure⟩ +⟨λ _, set α → ℝ≥0∞, λ m, m.to_outer_measure⟩ section @@ -118,7 +118,7 @@ namespace measure /-! ### General facts about measures -/ /-- Obtain a measure by giving a countably additive function that sends `∅` to `0`. -/ -def of_measurable (m : Π (s : set α), measurable_set s → ennreal) +def of_measurable (m : Π (s : set α), measurable_set s → ℝ≥0∞) (m0 : m ∅ measurable_set.empty = 0) (mU : ∀ {{f : ℕ → set α}} (h : ∀ i, measurable_set (f i)), pairwise (disjoint on f) → m (⋃ i, f i) (measurable_set.Union h) = ∑' i, m (f i) (h i)) : measure α := @@ -136,7 +136,7 @@ def of_measurable (m : Π (s : set α), measurable_set s → ennreal) end, ..induced_outer_measure m _ m0 } -lemma of_measurable_apply {m : Π (s : set α), measurable_set s → ennreal} +lemma of_measurable_apply {m : Π (s : set α), measurable_set s → ℝ≥0∞} {m0 : m ∅ measurable_set.empty = 0} {mU : ∀ {{f : ℕ → set α}} (h : ∀ i, measurable_set (f i)), pairwise (disjoint on f) → m (⋃ i, f i) (measurable_set.Union h) = ∑' i, m (f i) (h i)} @@ -534,7 +534,7 @@ protected lemma caratheodory (μ : measure α) (hs : measurable_set s) : μ (t ∩ s) + μ (t \ s) = μ t := (le_to_outer_measure_caratheodory μ s hs t).symm -/-! ### The `ennreal`-module of measures -/ +/-! ### The `ℝ≥0∞`-module of measures -/ instance : has_zero (measure α) := ⟨{ to_outer_measure := 0, @@ -569,24 +569,24 @@ instance add_comm_monoid : add_comm_monoid (measure α) := to_outer_measure_injective.add_comm_monoid to_outer_measure zero_to_outer_measure add_to_outer_measure -instance : has_scalar ennreal (measure α) := +instance : has_scalar ℝ≥0∞ (measure α) := ⟨λ c μ, { to_outer_measure := c • μ.to_outer_measure, m_Union := λ s hs hd, by simp [measure_Union, *, ennreal.tsum_mul_left], trimmed := by rw [outer_measure.trim_smul, μ.trimmed] }⟩ -@[simp] theorem smul_to_outer_measure (c : ennreal) (μ : measure α) : +@[simp] theorem smul_to_outer_measure (c : ℝ≥0∞) (μ : measure α) : (c • μ).to_outer_measure = c • μ.to_outer_measure := rfl -@[simp, norm_cast] theorem coe_smul (c : ennreal) (μ : measure α) : ⇑(c • μ) = c • μ := +@[simp, norm_cast] theorem coe_smul (c : ℝ≥0∞) (μ : measure α) : ⇑(c • μ) = c • μ := rfl -theorem smul_apply (c : ennreal) (μ : measure α) (s : set α) : (c • μ) s = c * μ s := +theorem smul_apply (c : ℝ≥0∞) (μ : measure α) (s : set α) : (c • μ) s = c * μ s := rfl -instance : semimodule ennreal (measure α) := -injective.semimodule ennreal ⟨to_outer_measure, zero_to_outer_measure, add_to_outer_measure⟩ +instance : semimodule ℝ≥0∞ (measure α) := +injective.semimodule ℝ≥0∞ ⟨to_outer_measure, zero_to_outer_measure, add_to_outer_measure⟩ to_outer_measure_injective smul_to_outer_measure /-! ### The complete lattice of measures -/ @@ -692,23 +692,23 @@ lemma nonpos_iff_eq_zero' : μ ≤ 0 ↔ μ = 0 := /-- Lift a linear map between `outer_measure` spaces such that for each measure `μ` every measurable set is caratheodory-measurable w.r.t. `f μ` to a linear map between `measure` spaces. -/ -def lift_linear (f : outer_measure α →ₗ[ennreal] outer_measure β) +def lift_linear (f : outer_measure α →ₗ[ℝ≥0∞] outer_measure β) (hf : ∀ μ : measure α, ‹_› ≤ (f μ.to_outer_measure).caratheodory) : - measure α →ₗ[ennreal] measure β := + measure α →ₗ[ℝ≥0∞] measure β := { to_fun := λ μ, (f μ.to_outer_measure).to_measure (hf μ), map_add' := λ μ₁ μ₂, ext $ λ s hs, by simp [hs], map_smul' := λ c μ, ext $ λ s hs, by simp [hs] } -@[simp] lemma lift_linear_apply {f : outer_measure α →ₗ[ennreal] outer_measure β} (hf) +@[simp] lemma lift_linear_apply {f : outer_measure α →ₗ[ℝ≥0∞] outer_measure β} (hf) {s : set β} (hs : measurable_set s) : lift_linear f hf μ s = f μ.to_outer_measure s := to_measure_apply _ _ hs -lemma le_lift_linear_apply {f : outer_measure α →ₗ[ennreal] outer_measure β} (hf) (s : set β) : +lemma le_lift_linear_apply {f : outer_measure α →ₗ[ℝ≥0∞] outer_measure β} (hf) (s : set β) : f μ.to_outer_measure s ≤ lift_linear f hf μ s := le_to_measure_apply _ _ s /-- The pushforward of a measure. It is defined to be `0` if `f` is not a measurable function. -/ -def map (f : α → β) : measure α →ₗ[ennreal] measure β := +def map (f : α → β) : measure α →ₗ[ℝ≥0∞] measure β := if hf : measurable f then lift_linear (outer_measure.map f) $ λ μ s hs t, le_to_outer_measure_caratheodory μ _ (hf hs) (f ⁻¹' t) @@ -747,7 +747,7 @@ nonpos_iff_eq_zero.mp $ (le_map_apply hf s).trans_eq hs /-- Pullback of a `measure`. If `f` sends each `measurable` set to a `measurable` set, then for each measurable set `s` we have `comap f μ s = μ (f '' s)`. -/ -def comap (f : α → β) : measure β →ₗ[ennreal] measure α := +def comap (f : α → β) : measure β →ₗ[ℝ≥0∞] measure α := if hf : injective f ∧ ∀ s, measurable_set s → measurable_set (f '' s) then lift_linear (outer_measure.comap f) $ λ μ s hs t, begin @@ -768,8 +768,8 @@ end /-! ### Restricting a measure -/ -/-- Restrict a measure `μ` to a set `s` as an `ennreal`-linear map. -/ -def restrictₗ (s : set α) : measure α →ₗ[ennreal] measure α := +/-- Restrict a measure `μ` to a set `s` as an `ℝ≥0∞`-linear map. -/ +def restrictₗ (s : set α) : measure α →ₗ[ℝ≥0∞] measure α := lift_linear (outer_measure.restrict s) $ λ μ s' hs' t, begin suffices : μ (s ∩ t) = μ (s ∩ t ∩ s') + μ (s ∩ t \ s'), @@ -801,7 +801,7 @@ by { rw [restrict, restrictₗ], convert le_lift_linear_apply _ t, simp } @[simp] lemma restrict_zero (s : set α) : (0 : measure α).restrict s = 0 := (restrictₗ s).map_zero -@[simp] lemma restrict_smul (c : ennreal) (μ : measure α) (s : set α) : +@[simp] lemma restrict_smul (c : ℝ≥0∞) (μ : measure α) (s : set α) : (c • μ).restrict s = c • μ.restrict s := (restrictₗ s).map_smul c μ @@ -1094,7 +1094,7 @@ to_measure_apply _ _ hs @[simp] lemma dirac_apply_of_mem {a : α} (h : a ∈ s) : dirac a s = 1 := begin - have : ∀ t : set α, a ∈ t → t.indicator (1 : α → ennreal) a = 1, + have : ∀ t : set α, a ∈ t → t.indicator (1 : α → ℝ≥0∞) a = 1, from λ t ht, indicator_of_mem ht 1, refine le_antisymm (this univ trivial ▸ _) (this s h ▸ le_dirac_apply), rw [← dirac_apply' a measurable_set.univ], @@ -1158,8 +1158,8 @@ ext $ λ t ht, by simp only [sum_apply, restrict_apply, ht, ht.inter hs] /-- Counting measure on any measurable space. -/ def count : measure α := sum dirac -lemma le_count_apply : (∑' i : s, 1 : ennreal) ≤ count s := -calc (∑' i : s, 1 : ennreal) = ∑' i, indicator s 1 i : tsum_subtype s 1 +lemma le_count_apply : (∑' i : s, 1 : ℝ≥0∞) ≤ count s := +calc (∑' i : s, 1 : ℝ≥0∞) = ∑' i, indicator s 1 i : tsum_subtype s 1 ... ≤ ∑' i, dirac i s : ennreal.tsum_le_tsum $ λ x, le_dirac_apply ... ≤ count s : le_sum_apply _ _ @@ -1181,7 +1181,7 @@ lemma count_apply_infinite (hs : s.infinite) : count s = ⊤ := begin refine top_unique (le_of_tendsto' ennreal.tendsto_nat_nhds_top $ λ n, _), rcases hs.exists_subset_card_eq n with ⟨t, ht, rfl⟩, - calc (t.card : ennreal) = ∑ i in t, 1 : by simp + calc (t.card : ℝ≥0∞) = ∑ i in t, 1 : by simp ... = ∑' i : (t : set α), 1 : (t.tsum_subtype 1).symm ... ≤ count (t : set α) : le_count_apply ... ≤ count s : measure_mono ht @@ -1358,10 +1358,10 @@ begin congr' with x, simp [and_comm] end -lemma ae_smul_measure {p : α → Prop} (h : ∀ᵐ x ∂μ, p x) (c : ennreal) : ∀ᵐ x ∂(c • μ), p x := +lemma ae_smul_measure {p : α → Prop} (h : ∀ᵐ x ∂μ, p x) (c : ℝ≥0∞) : ∀ᵐ x ∂(c • μ), p x := ae_iff.2 $ by rw [smul_apply, ae_iff.1 h, mul_zero] -lemma ae_smul_measure_iff {p : α → Prop} {c : ennreal} (hc : c ≠ 0) : +lemma ae_smul_measure_iff {p : α → Prop} {c : ℝ≥0∞} (hc : c ≠ 0) : (∀ᵐ x ∂(c • μ), p x) ↔ ∀ᵐ x ∂μ, p x := by simp [ae_iff, hc] @@ -1732,7 +1732,7 @@ lemma measure.finite_at_nhds [topological_space α] (μ : measure α) locally_finite_measure.finite_at_nhds x lemma measure.smul_finite {α : Type*} [measurable_space α] (μ : measure α) [finite_measure μ] - {c : ennreal} (hc : c < ⊤) : + {c : ℝ≥0∞} (hc : c < ⊤) : finite_measure (c • μ) := begin refine ⟨_⟩, @@ -2027,9 +2027,9 @@ theorem measurable_set.diff_null (hs : measurable_set s) (hz : μ z = 0) : begin rw measure_eq_infi at hz, choose f hf using show ∀ q : {q : ℚ // q > 0}, ∃ t : set α, - z ⊆ t ∧ measurable_set t ∧ μ t < (nnreal.of_real q.1 : ennreal), + z ⊆ t ∧ measurable_set t ∧ μ t < (nnreal.of_real q.1 : ℝ≥0∞), { rintro ⟨ε, ε0⟩, - have : 0 < (nnreal.of_real ε : ennreal), { simpa using ε0 }, + have : 0 < (nnreal.of_real ε : ℝ≥0∞), { simpa using ε0 }, rw ← hz at this, simpa [infi_lt_iff] }, refine null_measurable_set_iff.2 ⟨s \ Inter f, diff_subset_diff_right (subset_Inter (λ i, (hf i).1)), @@ -2311,7 +2311,7 @@ begin ... = 0 : hν.ae_eq_mk } end -lemma smul_measure (h : ae_measurable f μ) (c : ennreal) : +lemma smul_measure (h : ae_measurable f μ) (c : ℝ≥0∞) : ae_measurable f (c • μ) := ⟨h.mk f, h.measurable_mk, ae_smul_measure h.ae_eq_mk c⟩ @@ -2354,7 +2354,7 @@ lemma ae_measurable_congr (h : f =ᵐ[μ] g) : @[simp] lemma ae_measurable_const {b : β} : ae_measurable (λ a : α, b) μ := measurable_const.ae_measurable -@[simp] lemma ae_measurable_smul_measure_iff {c : ennreal} (hc : c ≠ 0) : +@[simp] lemma ae_measurable_smul_measure_iff {c : ℝ≥0∞} (hc : c ≠ 0) : ae_measurable f (c • μ) ↔ ae_measurable f μ := ⟨λ h, ⟨h.mk f, h.measurable_mk, (ae_smul_measure_iff hc).1 h.ae_eq_mk⟩, λ h, ⟨h.mk f, h.measurable_mk, (ae_smul_measure_iff hc).2 h.ae_eq_mk⟩⟩ diff --git a/src/measure_theory/outer_measure.lean b/src/measure_theory/outer_measure.lean index 7fa01a1d10392..bba97b6f14bbd 100644 --- a/src/measure_theory/outer_measure.lean +++ b/src/measure_theory/outer_measure.lean @@ -10,7 +10,7 @@ import topology.algebra.infinite_sum /-! # Outer Measures -An outer measure is a function `μ : set α → ennreal`, from the powerset of a type to the extended +An outer measure is a function `μ : set α → ℝ≥0∞`, from the powerset of a type to the extended nonnegative real numbers that satisfies the following conditions: 1. `μ ∅ = 0`; 2. `μ` is monotone; @@ -21,7 +21,7 @@ Note that we do not need `α` to be measurable to define an outer measure. The outer measures on a type `α` form a complete lattice. -Given an arbitrary function `m : set α → ennreal` that sends `∅` to `0` we can define an outer +Given an arbitrary function `m : set α → ℝ≥0∞` that sends `∅` to `0` we can define an outer measure on `α` that on `s` is defined to be the infimum of `∑ᵢ, m (sᵢ)` for all collections of sets `sᵢ` that cover `s`. This is the unique maximal outer measure that is at most the given function. We also define this for functions `m` defined on a subset of `set α`, by treating the function as @@ -52,13 +52,13 @@ outer measure, Carathéodory-measurable, Carathéodory's criterion noncomputable theory open set finset function filter encodable -open_locale classical big_operators nnreal topological_space +open_locale classical big_operators nnreal topological_space ennreal namespace measure_theory /-- An outer measure is a countably subadditive monotone function that sends `∅` to `0`. -/ structure outer_measure (α : Type*) := -(measure_of : set α → ennreal) +(measure_of : set α → ℝ≥0∞) (empty : measure_of ∅ = 0) (mono : ∀{s₁ s₂}, s₁ ⊆ s₂ → measure_of s₁ ≤ measure_of s₂) (Union_nat : ∀(s:ℕ → set α), measure_of (⋃i, s i) ≤ ∑'i, measure_of (s i)) @@ -146,23 +146,23 @@ theorem add_apply (m₁ m₂ : outer_measure α) (s : set α) : (m₁ + m₂) s instance add_comm_monoid : add_comm_monoid (outer_measure α) := { zero := 0, add := (+), - .. injective.add_comm_monoid (show outer_measure α → set α → ennreal, from coe_fn) + .. injective.add_comm_monoid (show outer_measure α → set α → ℝ≥0∞, from coe_fn) injective_coe_fn rfl (λ _ _, rfl) } -instance : has_scalar ennreal (outer_measure α) := +instance : has_scalar ℝ≥0∞ (outer_measure α) := ⟨λ c m, { measure_of := λ s, c * m s, empty := by simp, mono := λ s t h, ennreal.mul_left_mono $ m.mono h, Union_nat := λ s, by { rw [ennreal.tsum_mul_left], exact ennreal.mul_left_mono (m.Union _) } }⟩ -@[simp] lemma coe_smul (c : ennreal) (m : outer_measure α) : ⇑(c • m) = c • m := rfl +@[simp] lemma coe_smul (c : ℝ≥0∞) (m : outer_measure α) : ⇑(c • m) = c • m := rfl -lemma smul_apply (c : ennreal) (m : outer_measure α) (s : set α) : (c • m) s = c * m s := rfl +lemma smul_apply (c : ℝ≥0∞) (m : outer_measure α) (s : set α) : (c • m) s = c * m s := rfl -instance : semimodule ennreal (outer_measure α) := +instance : semimodule ℝ≥0∞ (outer_measure α) := { smul := (•), - .. injective.semimodule ennreal ⟨show outer_measure α → set α → ennreal, from coe_fn, coe_zero, + .. injective.semimodule ℝ≥0∞ ⟨show outer_measure α → set α → ℝ≥0∞, from coe_fn, coe_zero, coe_add⟩ injective_coe_fn coe_smul } instance : has_bot (outer_measure α) := ⟨0⟩ @@ -210,7 +210,7 @@ by have := supr_apply (λ b, cond b m₁ m₂) s; end supremum /-- The pushforward of `m` along `f`. The outer measure on `s` is defined to be `m (f ⁻¹' s)`. -/ -def map {β} (f : α → β) : outer_measure α →ₗ[ennreal] outer_measure β := +def map {β} (f : α → β) : outer_measure α →ₗ[ℝ≥0∞] outer_measure β := { to_fun := λ m, { measure_of := λs, m (f ⁻¹' s), empty := m.empty, @@ -243,7 +243,7 @@ def dirac (a : α) : outer_measure α := mono := λ s t h, indicator_le_indicator_of_subset h (λ _, zero_le _) a, Union_nat := λ s, if hs : a ∈ ⋃ n, s n then let ⟨i, hi⟩ := mem_Union.1 hs in - calc indicator (⋃ n, s n) (λ _, (1 : ennreal)) a = 1 : indicator_of_mem hs _ + calc indicator (⋃ n, s n) (λ _, (1 : ℝ≥0∞)) a = 1 : indicator_of_mem hs _ ... = indicator (s i) (λ _, 1) a : (indicator_of_mem hi _).symm ... ≤ ∑' n, indicator (s n) (λ _, 1) a : ennreal.le_tsum _ else by simp only [indicator_of_not_mem hs, zero_le]} @@ -262,12 +262,12 @@ def sum {ι} (f : ι → outer_measure α) : outer_measure α := @[simp] theorem sum_apply {ι} (f : ι → outer_measure α) (s : set α) : sum f s = ∑' i, f i s := rfl -theorem smul_dirac_apply (a : ennreal) (b : α) (s : set α) : +theorem smul_dirac_apply (a : ℝ≥0∞) (b : α) (s : set α) : (a • dirac b) s = indicator s (λ _, a) b := by simp /-- Pullback of an `outer_measure`: `comap f μ s = μ (f '' s)`. -/ -def comap {β} (f : α → β) : outer_measure β →ₗ[ennreal] outer_measure α := +def comap {β} (f : α → β) : outer_measure β →ₗ[ℝ≥0∞] outer_measure α := { to_fun := λ m, { measure_of := λ s, m (f '' s), empty := by simp, @@ -281,7 +281,7 @@ def comap {β} (f : α → β) : outer_measure β →ₗ[ennreal] outer_measure rfl /-- Restrict an `outer_measure` to a set. -/ -def restrict (s : set α) : outer_measure α →ₗ[ennreal] outer_measure α := +def restrict (s : set α) : outer_measure α →ₗ[ℝ≥0∞] outer_measure α := (map coe).comp (comap (coe : s → α)) @[simp] lemma restrict_apply (s t : set α) (m : outer_measure α) : @@ -290,14 +290,14 @@ by simp [restrict] theorem top_apply {s : set α} (h : s.nonempty) : (⊤ : outer_measure α) s = ⊤ := let ⟨a, as⟩ := h in -top_unique $ le_trans (by simp [smul_dirac_apply, as]) (le_bsupr ((⊤ : ennreal) • dirac a) trivial) +top_unique $ le_trans (by simp [smul_dirac_apply, as]) (le_bsupr ((⊤ : ℝ≥0∞) • dirac a) trivial) end basic section of_function set_option eqn_compiler.zeta true -variables {α : Type*} (m : set α → ennreal) (m_empty : m ∅ = 0) +variables {α : Type*} (m : set α → ℝ≥0∞) (m_empty : m ∅ = 0) include m_empty /-- Given any function `m` assigning measures to sets satisying `m ∅ = 0`, there is @@ -360,7 +360,7 @@ end of_function section bounded_by -variables {α : Type*} (m : set α → ennreal) +variables {α : Type*} (m : set α → ℝ≥0∞) /-- Given any function `m` assigning measures to sets, there is a unique maximal outer measure `μ` satisfying `μ s ≤ m s` for all `s : set α`. This is the same as `outer_measure.of_function`, @@ -516,7 +516,7 @@ end caratheodory_measurable variables {α : Type*} -lemma of_function_caratheodory {m : set α → ennreal} {s : set α} +lemma of_function_caratheodory {m : set α → ℝ≥0∞} {s : set α} {h₀ : m ∅ = 0} (hs : ∀t, m (t ∩ s) + m (t \ s) ≤ m t) : (outer_measure.of_function m h₀).caratheodory.measurable_set' s := begin @@ -530,7 +530,7 @@ begin { rw ← ennreal.tsum_add, exact ennreal.tsum_le_tsum (λ i, hs _) } end -lemma bounded_by_caratheodory {m : set α → ennreal} {s : set α} +lemma bounded_by_caratheodory {m : set α → ℝ≥0∞} {s : set α} (hs : ∀t, m (t ∩ s) + m (t \ s) ≤ m t) : (bounded_by m).caratheodory.measurable_set' s := begin apply of_function_caratheodory, intro t, @@ -556,7 +556,7 @@ theorem le_sum_caratheodory {ι} (m : ι → outer_measure α) : λ s h t, by simp [λ i, measurable_space.measurable_set_infi.1 h i t, ennreal.tsum_add] -theorem le_smul_caratheodory (a : ennreal) (m : outer_measure α) : +theorem le_smul_caratheodory (a : ℝ≥0∞) (m : outer_measure α) : m.caratheodory ≤ (a • m).caratheodory := λ s h t, by simp [h t, mul_add] @@ -572,7 +572,7 @@ section Inf_gen infimum of `μ(s)` for the outer measures `μ` in the collection. We ensure that this function is defined to be `0` on `∅`, even if the collection of outer measures is empty. The outer measure generated by this function is the infimum of the given outer measures. -/ -def Inf_gen (m : set (outer_measure α)) (s : set α) : ennreal := +def Inf_gen (m : set (outer_measure α)) (s : set α) : ℝ≥0∞ := ⨅ (μ : outer_measure α) (h : μ ∈ m), μ s lemma Inf_gen_def (m : set (outer_measure α)) (t : set α) : @@ -655,11 +655,11 @@ open outer_measure at the end. -/ section extend variables {α : Type*} {P : α → Prop} -variables (m : Π (s : α), P s → ennreal) +variables (m : Π (s : α), P s → ℝ≥0∞) -/-- We can trivially extend a function defined on a subclass of objects (with codomain `ennreal`) +/-- We can trivially extend a function defined on a subclass of objects (with codomain `ℝ≥0∞`) to all objects by defining it to be `∞` on the objects not in the class. -/ -def extend (s : α) : ennreal := ⨅ h : P s, m s h +def extend (s : α) : ℝ≥0∞ := ⨅ h : P s, m s h lemma extend_eq {s : α} (h : P s) : extend m s = m s h := by simp [extend, h] @@ -672,7 +672,7 @@ end extend section extend_set variables {α : Type*} {P : set α → Prop} -variables {m : Π (s : set α), P s → ennreal} +variables {m : Π (s : set α), P s → ℝ≥0∞} variables (P0 : P ∅) (m0 : m ∅ P0 = 0) variables (PU : ∀{{f : ℕ → set α}} (hm : ∀i, P (f i)), P (⋃i, f i)) variables (mU : ∀ {{f : ℕ → set α}} (hm : ∀i, P (f i)), pairwise (disjoint on f) → @@ -813,7 +813,7 @@ end extend_set section measurable_space variables {α : Type*} [measurable_space α] -variables {m : Π (s : set α), measurable_set s → ennreal} +variables {m : Π (s : set α), measurable_set s → ℝ≥0∞} variables (m0 : m ∅ measurable_set.empty = 0) variable (mU : ∀ {{f : ℕ → set α}} (hm : ∀i, measurable_set (f i)), pairwise (disjoint on f) → m (⋃i, f i) (measurable_set.Union hm) = ∑'i, m (f i) (hm i)) @@ -943,7 +943,7 @@ begin exact ⟨t, hst, ht, h ▸ hm⟩ end -theorem trim_smul (c : ennreal) (m : outer_measure α) : +theorem trim_smul (c : ℝ≥0∞) (m : outer_measure α) : (c • m).trim = c • m.trim := begin ext1 s, diff --git a/src/measure_theory/pi.lean b/src/measure_theory/pi.lean index 4c0092ccc7ab1..90ba72e5ebf67 100644 --- a/src/measure_theory/pi.lean +++ b/src/measure_theory/pi.lean @@ -45,7 +45,7 @@ finitary product measure noncomputable theory open function set measure_theory.outer_measure filter -open_locale classical big_operators topological_space +open_locale classical big_operators topological_space ennreal namespace measure_theory @@ -55,7 +55,7 @@ variables {ι : Type*} [fintype ι] {α : ι → Type*} {m : Π i, outer_measure It is defined to by taking the image of the set under all projections, and taking the product of the measures of these images. For measurable boxes it is equal to the correct measure. -/ -@[simp] def pi_premeasure (m : Π i, outer_measure (α i)) (s : set (Π i, α i)) : ennreal := +@[simp] def pi_premeasure (m : Π i, outer_measure (α i)) (s : set (Π i, α i)) : ℝ≥0∞ := ∏ i, m i (eval i '' s) lemma pi_premeasure_pi {s : Π i, set (α i)} (hs : (pi univ s).nonempty) : @@ -69,7 +69,7 @@ begin { rcases univ_pi_eq_empty_iff.mp h with ⟨i, hi⟩, have : ∃ i, m i (s i) = 0 := ⟨i, by simp [hi]⟩, simpa [h, finset.card_univ, zero_pow (fintype.card_pos_iff.mpr _inst_2), - @eq_comm _ (0 : ennreal), finset.prod_eq_zero_iff] }, + @eq_comm _ (0 : ℝ≥0∞), finset.prod_eq_zero_iff] }, { simp [h] } end diff --git a/src/measure_theory/probability_mass_function.lean b/src/measure_theory/probability_mass_function.lean index 207332bb7cf7e..f540b5fb6e58a 100644 --- a/src/measure_theory/probability_mass_function.lean +++ b/src/measure_theory/probability_mass_function.lean @@ -25,7 +25,7 @@ probability mass function, discrete probability measure, bernoulli distribution -/ noncomputable theory variables {α : Type*} {β : Type*} {γ : Type*} -open_locale classical big_operators nnreal +open_locale classical big_operators nnreal ennreal /-- A probability mass function, or discrete probability measures is a function `α → ℝ≥0` such that the values have (infinite) sum `1`. -/ @@ -81,7 +81,7 @@ def bind (p : pmf α) (f : α → pmf β) : pmf β := rfl lemma coe_bind_apply (p : pmf α) (f : α → pmf β) (b : β) : - (p.bind f b : ennreal) = ∑'a, p a * f a b := + (p.bind f b : ℝ≥0∞) = ∑'a, p a * f a b := eq.trans (ennreal.coe_tsum $ bind.summable p f b) $ by simp @[simp] lemma pure_bind (a : α) (f : α → pmf β) : (pure a).bind f = f a := diff --git a/src/measure_theory/prod.lean b/src/measure_theory/prod.lean index ee235396636e6..25f119c341819 100644 --- a/src/measure_theory/prod.lean +++ b/src/measure_theory/prod.lean @@ -28,8 +28,8 @@ We also prove Tonelli's theorem and Fubini's theorem. * `measure_theory.measure.prod_prod` states `μ.prod ν (s.prod t) = μ s * ν t` for measurable sets `s` and `t`. * `measure_theory.lintegral_prod`: Tonelli's theorem. It states that for a measurable function - `α × β → ennreal` we have `∫⁻ z, f z ∂(μ.prod ν) = ∫⁻ x, ∫⁻ y, f (x, y) ∂ν ∂μ`. The version - for functions `α → β → ennreal` is reversed, and called `lintegral_lintegral`. Both versions have + `α × β → ℝ≥0∞` we have `∫⁻ z, f z ∂(μ.prod ν) = ∫⁻ x, ∫⁻ y, f (x, y) ∂ν ∂μ`. The version + for functions `α → β → ℝ≥0∞` is reversed, and called `lintegral_lintegral`. Both versions have a variant with `_symm` appended, where the order of integration is reversed. The lemma `measurable.lintegral_prod_right'` states that the inner integral of the right-hand side is measurable. @@ -58,7 +58,7 @@ product measure, Fubini's theorem, Tonelli's theorem, Fubini-Tonelli theorem -/ noncomputable theory -open_locale classical topological_space +open_locale classical topological_space ennreal open set function real ennreal open measure_theory measurable_space measure_theory.measure open topological_space (hiding generate_from) @@ -198,7 +198,7 @@ end /-- The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) Tonelli's theorem is measurable. -/ lemma measurable.lintegral_prod_right' [sigma_finite ν] : - ∀ {f : α × β → ennreal} (hf : measurable f), measurable (λ x, ∫⁻ y, f (x, y) ∂ν) := + ∀ {f : α × β → ℝ≥0∞} (hf : measurable f), measurable (λ x, ∫⁻ y, f (x, y) ∂ν) := begin have m := @measurable_prod_mk_left, refine measurable.ennreal_induction _ _ _, @@ -217,20 +217,20 @@ end /-- The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) Tonelli's theorem is measurable. This version has the argument `f` in curried form. -/ -lemma measurable.lintegral_prod_right [sigma_finite ν] {f : α → β → ennreal} +lemma measurable.lintegral_prod_right [sigma_finite ν] {f : α → β → ℝ≥0∞} (hf : measurable (uncurry f)) : measurable (λ x, ∫⁻ y, f x y ∂ν) := hf.lintegral_prod_right' /-- The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) the symmetric version of Tonelli's theorem is measurable. -/ -lemma measurable.lintegral_prod_left' [sigma_finite μ] {f : α × β → ennreal} +lemma measurable.lintegral_prod_left' [sigma_finite μ] {f : α × β → ℝ≥0∞} (hf : measurable f) : measurable (λ y, ∫⁻ x, f (x, y) ∂μ) := (measurable_swap_iff.mpr hf).lintegral_prod_right' /-- The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) the symmetric version of Tonelli's theorem is measurable. This version has the argument `f` in curried form. -/ -lemma measurable.lintegral_prod_left [sigma_finite μ] {f : α → β → ennreal} +lemma measurable.lintegral_prod_left [sigma_finite μ] {f : α → β → ℝ≥0∞} (hf : measurable (uncurry f)) : measurable (λ y, ∫⁻ x, f x y ∂μ) := hf.lintegral_prod_left' @@ -578,14 +578,14 @@ namespace measure_theory variables [sigma_finite ν] -lemma lintegral_prod_swap [sigma_finite μ] (f : α × β → ennreal) +lemma lintegral_prod_swap [sigma_finite μ] (f : α × β → ℝ≥0∞) (hf : ae_measurable f (μ.prod ν)) : ∫⁻ z, f z.swap ∂(ν.prod μ) = ∫⁻ z, f z ∂(μ.prod ν) := by { rw ← prod_swap at hf, rw [← lintegral_map' hf measurable_swap, prod_swap] } -/-- Tonelli's Theorem: For `ennreal`-valued measurable functions on `α × β`, +/-- Tonelli's Theorem: For `ℝ≥0∞`-valued measurable functions on `α × β`, the integral of `f` is equal to the iterated integral. -/ lemma lintegral_prod_of_measurable : - ∀ (f : α × β → ennreal) (hf : measurable f), ∫⁻ z, f z ∂(μ.prod ν) = ∫⁻ x, ∫⁻ y, f (x, y) ∂ν ∂μ := + ∀ (f : α × β → ℝ≥0∞) (hf : measurable f), ∫⁻ z, f z ∂(μ.prod ν) = ∫⁻ x, ∫⁻ y, f (x, y) ∂ν ∂μ := begin have m := @measurable_prod_mk_left, refine measurable.ennreal_induction _ _ _, @@ -603,9 +603,9 @@ begin simp only [lintegral_supr hf h2f, lintegral_supr (kf _), k2f, lintegral_supr lf l2f, h3f] }, end -/-- Tonelli's Theorem: For `ennreal`-valued almost everywhere measurable functions on `α × β`, +/-- Tonelli's Theorem: For `ℝ≥0∞`-valued almost everywhere measurable functions on `α × β`, the integral of `f` is equal to the iterated integral. -/ -lemma lintegral_prod (f : α × β → ennreal) (hf : ae_measurable f (μ.prod ν)) : +lemma lintegral_prod (f : α × β → ℝ≥0∞) (hf : ae_measurable f (μ.prod ν)) : ∫⁻ z, f z ∂(μ.prod ν) = ∫⁻ x, ∫⁻ y, f (x, y) ∂ν ∂μ := begin have A : ∫⁻ z, f z ∂(μ.prod ν) = ∫⁻ z, hf.mk f z ∂(μ.prod ν) := @@ -619,39 +619,39 @@ begin apply_instance end -/-- The symmetric verion of Tonelli's Theorem: For `ennreal`-valued almost everywhere measurable +/-- The symmetric verion of Tonelli's Theorem: For `ℝ≥0∞`-valued almost everywhere measurable functions on `α × β`, the integral of `f` is equal to the iterated integral, in reverse order. -/ -lemma lintegral_prod_symm' [sigma_finite μ] (f : α × β → ennreal) +lemma lintegral_prod_symm' [sigma_finite μ] (f : α × β → ℝ≥0∞) (hf : ae_measurable f (μ.prod ν)) : ∫⁻ z, f z ∂(μ.prod ν) = ∫⁻ y, ∫⁻ x, f (x, y) ∂μ ∂ν := by { simp_rw [← lintegral_prod_swap f hf], exact lintegral_prod _ hf.prod_swap } -/-- The symmetric verion of Tonelli's Theorem: For `ennreal`-valued measurable +/-- The symmetric verion of Tonelli's Theorem: For `ℝ≥0∞`-valued measurable functions on `α × β`, the integral of `f` is equal to the iterated integral, in reverse order. -/ -lemma lintegral_prod_symm [sigma_finite μ] (f : α × β → ennreal) +lemma lintegral_prod_symm [sigma_finite μ] (f : α × β → ℝ≥0∞) (hf : ae_measurable f (μ.prod ν)) : ∫⁻ z, f z ∂(μ.prod ν) = ∫⁻ y, ∫⁻ x, f (x, y) ∂μ ∂ν := lintegral_prod_symm' f hf /-- The reversed version of Tonelli's Theorem. In this version `f` is in curried form, which makes it easier for the elaborator to figure out `f` automatically. -/ -lemma lintegral_lintegral ⦃f : α → β → ennreal⦄ +lemma lintegral_lintegral ⦃f : α → β → ℝ≥0∞⦄ (hf : ae_measurable (uncurry f) (μ.prod ν)) : ∫⁻ x, ∫⁻ y, f x y ∂ν ∂μ = ∫⁻ z, f z.1 z.2 ∂(μ.prod ν) := (lintegral_prod _ hf).symm /-- The reversed version of Tonelli's Theorem (symmetric version). In this version `f` is in curried form, which makes it easier for the elaborator to figure out `f` automatically. -/ -lemma lintegral_lintegral_symm [sigma_finite μ] ⦃f : α → β → ennreal⦄ +lemma lintegral_lintegral_symm [sigma_finite μ] ⦃f : α → β → ℝ≥0∞⦄ (hf : ae_measurable (uncurry f) (μ.prod ν)) : ∫⁻ x, ∫⁻ y, f x y ∂ν ∂μ = ∫⁻ z, f z.2 z.1 ∂(ν.prod μ) := (lintegral_prod_symm _ hf.prod_swap).symm /-- Change the order of Lebesgue integration. -/ -lemma lintegral_lintegral_swap [sigma_finite μ] ⦃f : α → β → ennreal⦄ +lemma lintegral_lintegral_swap [sigma_finite μ] ⦃f : α → β → ℝ≥0∞⦄ (hf : ae_measurable (uncurry f) (μ.prod ν)) : ∫⁻ x, ∫⁻ y, f x y ∂ν ∂μ = ∫⁻ y, ∫⁻ x, f x y ∂μ ∂ν := (lintegral_lintegral hf).trans (lintegral_prod_symm _ hf) -lemma lintegral_prod_mul {f : α → ennreal} {g : β → ennreal} +lemma lintegral_prod_mul {f : α → ℝ≥0∞} {g : β → ℝ≥0∞} (hf : ae_measurable f μ) (hg : ae_measurable g ν) : ∫⁻ z, f z.1 * g z.2 ∂(μ.prod ν) = ∫⁻ x, f x ∂μ * ∫⁻ y, g y ∂ν := by simp [lintegral_prod _ (hf.fst.ennreal_mul hg.snd), lintegral_lintegral_mul hf hg] @@ -794,7 +794,7 @@ end /-- Integrals commute with subtraction inside a lower Lebesgue integral. `F` can be any function. -/ lemma lintegral_fn_integral_sub ⦃f g : α × β → E⦄ - (F : E → ennreal) (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : + (F : E → ℝ≥0∞) (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : ∫⁻ x, F (∫ y, f (x, y) - g (x, y) ∂ν) ∂μ = ∫⁻ x, F (∫ y, f (x, y) ∂ν - ∫ y, g (x, y) ∂ν) ∂μ := begin refine lintegral_congr_ae _, @@ -838,13 +838,13 @@ begin refine tendsto_integral_of_l1 _ g.integrable.integral_prod_left (eventually_of_forall $ λ h, h.integrable.integral_prod_left) _, simp_rw [edist_eq_coe_nnnorm_sub, - ← lintegral_fn_integral_sub (λ x, (nnnorm x : ennreal)) (l1.integrable _) g.integrable], + ← lintegral_fn_integral_sub (λ x, (nnnorm x : ℝ≥0∞)) (l1.integrable _) g.integrable], refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds _ (λ i, zero_le _) _, { exact λ i, ∫⁻ x, ∫⁻ y, nnnorm (i (x, y) - g (x, y)) ∂ν ∂μ }, swap, { exact λ i, lintegral_mono (λ x, ennnorm_integral_le_lintegral_ennnorm _) }, show tendsto (λ (i : α × β →₁[μ.prod ν] E), ∫⁻ x, ∫⁻ (y : β), nnnorm (i (x, y) - g (x, y)) ∂ν ∂μ) (𝓝 g) (𝓝 0), - have : ∀ (i : α × β →₁[μ.prod ν] E), measurable (λ z, (nnnorm (i z - g z) : ennreal)) := + have : ∀ (i : α × β →₁[μ.prod ν] E), measurable (λ z, (nnnorm (i z - g z) : ℝ≥0∞)) := λ i, (i.measurable.sub g.measurable).ennnorm, simp_rw [← lintegral_prod_of_measurable _ (this _), ← l1.of_real_norm_sub_eq_lintegral, ← of_real_zero], diff --git a/src/measure_theory/prod_group.lean b/src/measure_theory/prod_group.lean index 6fcc23a49c0f3..788393c000ad3 100644 --- a/src/measure_theory/prod_group.lean +++ b/src/measure_theory/prod_group.lean @@ -35,7 +35,7 @@ https://math.stackexchange.com/questions/3974485/does-right-translation-preserve noncomputable theory open topological_space set (hiding prod_eq) function -open_locale classical +open_locale classical ennreal namespace measure_theory @@ -133,7 +133,7 @@ begin end lemma lintegral_lintegral_mul_inv (hμ : is_mul_left_invariant μ) (hν : is_mul_left_invariant ν) - (f : G → G → ennreal) (hf : ae_measurable (uncurry f) (μ.prod ν)) : + (f : G → G → ℝ≥0∞) (hf : ae_measurable (uncurry f) (μ.prod ν)) : ∫⁻ x, ∫⁻ y, f (y * x) x⁻¹ ∂ν ∂μ = ∫⁻ x, ∫⁻ y, f x y ∂ν ∂μ := begin have h : measurable (λ z : G × G, (z.2 * z.1, z.1⁻¹)) := @@ -169,7 +169,7 @@ lemma measure_mul_right_ne_zero (hμ : is_mul_left_invariant μ) {E : set G} (hE sets). -/ lemma measure_lintegral_div_measure [t2_space G] (hμ : is_mul_left_invariant μ) (hν : is_mul_left_invariant ν) (h2ν : regular ν) {E : set G} (hE : is_compact E) (h2E : ν E ≠ 0) - (f : G → ennreal) (hf : measurable f) : + (f : G → ℝ≥0∞) (hf : measurable f) : μ E * ∫⁻ y, f y⁻¹ / ν ((λ h, h * y⁻¹) ⁻¹' E) ∂ν = ∫⁻ x, f x ∂μ := begin have Em := hE.measurable_set, @@ -182,9 +182,9 @@ begin ← lintegral_lintegral_mul_inv hμ hν], swap, { exact (((measurable_const.indicator Em).comp measurable_fst).ennreal_mul (hg.comp measurable_snd)).ae_measurable }, - have mE : ∀ x : G, measurable (λ y, ((λ z, z * x) ⁻¹' E).indicator (λ z, (1 : ennreal)) y) := + have mE : ∀ x : G, measurable (λ y, ((λ z, z * x) ⁻¹' E).indicator (λ z, (1 : ℝ≥0∞)) y) := λ x, measurable_const.indicator (measurable_mul_right _ Em), - have : ∀ x y, E.indicator (λ (z : G), (1 : ennreal)) (y * x) = + have : ∀ x y, E.indicator (λ (z : G), (1 : ℝ≥0∞)) (y * x) = ((λ z, z * x) ⁻¹' E).indicator (λ (b : G), 1) y, { intros x y, symmetry, convert indicator_comp_right (λ y, y * x), ext1 z, simp }, have h3E : ∀ y, ν ((λ x, x * y) ⁻¹' E) ≠ ⊤ := diff --git a/src/measure_theory/set_integral.lean b/src/measure_theory/set_integral.lean index b017fcfe4792d..d78b561096a40 100644 --- a/src/measure_theory/set_integral.lean +++ b/src/measure_theory/set_integral.lean @@ -50,7 +50,7 @@ migrated to the new definition. noncomputable theory open set filter topological_space measure_theory function -open_locale classical topological_space interval big_operators filter +open_locale classical topological_space interval big_operators filter ennreal variables {α β E F : Type*} [measurable_space α] @@ -350,8 +350,8 @@ begin by_cases hc : c = 0, { subst hc, convert h_ind 0 measurable_set.empty (by simp) using 1, simp [const] }, apply h_ind c hs, - have : (nnnorm c : ennreal) * μ s < ⊤, - { have := @comp_indicator _ _ _ _ (λ x : E, (nnnorm x : ennreal)) (const α c) s, + have : (nnnorm c : ℝ≥0∞) * μ s < ⊤, + { have := @comp_indicator _ _ _ _ (λ x : E, (nnnorm x : ℝ≥0∞)) (const α c) s, dsimp only at this, have h' := h.has_finite_integral, simpa [has_finite_integral, this, lintegral_indicator, hs] using h' }, @@ -492,7 +492,7 @@ variables {ι : Type*} [measurable_space E] [normed_group E] /-- Fundamental theorem of calculus for set integrals: if `μ` is a measure that is finite at a filter `l` and `f` is a measurable function that has a finite limit `b` at `l ⊓ μ.ae`, then `∫ x in s i, f x ∂μ = μ (s i) • b + o(μ (s i))` at a filter `li` provided that `s i` tends to `l.lift' -powerset` along `li`. Since `μ (s i)` is an `ennreal` number, we use `(μ (s i)).to_real` in the +powerset` along `li`. Since `μ (s i)` is an `ℝ≥0∞` number, we use `(μ (s i)).to_real` in the actual statement. Often there is a good formula for `(μ (s i)).to_real`, so the formalization can take an optional @@ -526,7 +526,7 @@ end /-- Fundamental theorem of calculus for set integrals, `nhds_within` version: if `μ` is a locally finite measure and `f` is an almost everywhere measurable function that is continuous at a point `a` within a measurable set `t`, then `∫ x in s i, f x ∂μ = μ (s i) • f a + o(μ (s i))` at a filter `li` -provided that `s i` tends to `(𝓝[t] a).lift' powerset` along `li`. Since `μ (s i)` is an `ennreal` +provided that `s i` tends to `(𝓝[t] a).lift' powerset` along `li`. Since `μ (s i)` is an `ℝ≥0∞` number, we use `(μ (s i)).to_real` in the actual statement. Often there is a good formula for `(μ (s i)).to_real`, so the formalization can take an optional @@ -549,7 +549,7 @@ exact (ha.mono_left inf_le_left).integral_sub_linear_is_o_ae /-- Fundamental theorem of calculus for set integrals, `nhds` version: if `μ` is a locally finite measure and `f` is an almost everywhere measurable function that is continuous at a point `a`, then `∫ x in s i, f x ∂μ = μ (s i) • f a + o(μ (s i))` at `li` provided that `s` tends to `(𝓝 a).lift' -powerset` along `li. Since `μ (s i)` is an `ennreal` number, we use `(μ (s i)).to_real` in the +powerset` along `li. Since `μ (s i)` is an `ℝ≥0∞` number, we use `(μ (s i)).to_real` in the actual statement. Often there is a good formula for `(μ (s i)).to_real`, so the formalization can take an optional @@ -601,7 +601,7 @@ exact (hft a ha).integrable_at_filter ⟨_, self_mem_nhds_within, hft.ae_measura /-- Fundamental theorem of calculus for set integrals, `nhds_within` version: if `μ` is a locally finite measure, `f` is continuous on a measurable set `t`, and `a ∈ t`, then `∫ x in (s i), f x ∂μ = μ (s i) • f a + o(μ (s i))` at `li` provided that `s i` tends to `(𝓝[t] a).lift' powerset` along -`li`. Since `μ (s i)` is an `ennreal` number, we use `(μ (s i)).to_real` in the actual statement. +`li`. Since `μ (s i)` is an `ℝ≥0∞` number, we use `(μ (s i)).to_real` in the actual statement. Often there is a good formula for `(μ (s i)).to_real`, so the formalization can take an optional argument `m` with this formula and a proof `of `(λ i, (μ (s i)).to_real) =ᶠ[li] m`. Without these diff --git a/src/order/filter/ennreal.lean b/src/order/filter/ennreal.lean index 744fc1fe6a407..da040c9193c1f 100644 --- a/src/order/filter/ennreal.lean +++ b/src/order/filter/ennreal.lean @@ -11,28 +11,28 @@ import order.liminf_limsup /-! # Order properties of extended non-negative reals -This file compiles filter-related results about `ennreal` (see data/real/ennreal.lean). +This file compiles filter-related results about `ℝ≥0∞` (see data/real/ennreal.lean). -/ open filter -open_locale filter +open_locale filter ennreal namespace ennreal variables {α : Type*} {f : filter α} -lemma eventually_le_limsup [countable_Inter_filter f] (u : α → ennreal) : +lemma eventually_le_limsup [countable_Inter_filter f] (u : α → ℝ≥0∞) : ∀ᶠ y in f, u y ≤ f.limsup u := begin by_cases hx_top : f.limsup u = ⊤, { simp_rw hx_top, exact eventually_of_forall (λ a, le_top), }, - have h_forall_le : ∀ᶠ y in f, ∀ n : ℕ, u y < f.limsup u + (1:ennreal)/n, + have h_forall_le : ∀ᶠ y in f, ∀ n : ℕ, u y < f.limsup u + (1:ℝ≥0∞)/n, { rw eventually_countable_forall, refine λ n, eventually_lt_of_limsup_lt _, nth_rewrite 0 ←add_zero (f.limsup u), exact (add_lt_add_iff_left (lt_top_iff_ne_top.mpr hx_top)).mpr (by simp), }, refine h_forall_le.mono (λ y hy, le_of_forall_pos_le_add (λ r hr_pos hx_top, _)), - have hr_ne_zero : (r : ennreal) ≠ 0, + have hr_ne_zero : (r : ℝ≥0∞) ≠ 0, { rw [ne.def, coe_eq_zero], exact (ne_of_lt hr_pos).symm, }, cases (exists_inv_nat_lt hr_ne_zero) with i hi, @@ -40,7 +40,7 @@ begin exact le_trans (le_of_lt (hy i)) (add_le_add_left hi.le (f.limsup u)), end -lemma limsup_eq_zero_iff [countable_Inter_filter f] {u : α → ennreal} : +lemma limsup_eq_zero_iff [countable_Inter_filter f] {u : α → ℝ≥0∞} : f.limsup u = 0 ↔ u =ᶠ[f] 0 := begin split; intro h, @@ -51,13 +51,13 @@ begin simp_rw [pi.zero_apply, ←ennreal.bot_eq_zero, limsup_const_bot] }, end -lemma limsup_const_mul_of_ne_top {u : α → ennreal} {a : ennreal} (ha_top : a ≠ ⊤) : +lemma limsup_const_mul_of_ne_top {u : α → ℝ≥0∞} {a : ℝ≥0∞} (ha_top : a ≠ ⊤) : f.limsup (λ (x : α), a * (u x)) = a * f.limsup u := begin by_cases ha_zero : a = 0, { simp_rw [ha_zero, zero_mul, ←ennreal.bot_eq_zero], exact limsup_const_bot, }, - let g := λ x : ennreal, a * x, + let g := λ x : ℝ≥0∞, a * x, have hg_bij : function.bijective g, from function.bijective_iff_has_inverse.mpr ⟨(λ x, a⁻¹ * x), ⟨λ x, by simp [←mul_assoc, inv_mul_cancel ha_zero ha_top], @@ -70,7 +70,7 @@ begin all_goals { by is_bounded_default }, end -lemma limsup_const_mul [countable_Inter_filter f] {u : α → ennreal} {a : ennreal} : +lemma limsup_const_mul [countable_Inter_filter f] {u : α → ℝ≥0∞} {a : ℝ≥0∞} : f.limsup (λ (x : α), a * (u x)) = a * f.limsup u := begin by_cases ha_top : a ≠ ⊤, @@ -84,18 +84,18 @@ begin simp only [limsup_congr hu, limsup_congr hau, pi.zero_apply, ← bot_eq_zero, limsup_const_bot], simp, }, { simp_rw [ha_top, top_mul], - have hu_mul : ∃ᶠ (x : α) in f, ⊤ ≤ ite (u x = 0) (0 : ennreal) ⊤, + have hu_mul : ∃ᶠ (x : α) in f, ⊤ ≤ ite (u x = 0) (0 : ℝ≥0∞) ⊤, { rw [eventually_eq, not_eventually] at hu, refine hu.mono (λ x hx, _), rw pi.zero_apply at hx, simp [hx], }, - have h_top_le : f.limsup (λ (x : α), ite (u x = 0) (0 : ennreal) ⊤) = ⊤, + have h_top_le : f.limsup (λ (x : α), ite (u x = 0) (0 : ℝ≥0∞) ⊤) = ⊤, from eq_top_iff.mpr (le_limsup_of_frequently_le hu_mul), have hfu : f.limsup u ≠ 0, from mt limsup_eq_zero_iff.1 hu, simp only [h_top_le, hfu, if_false], }, end -lemma limsup_add_le [countable_Inter_filter f] (u v : α → ennreal) : +lemma limsup_add_le [countable_Inter_filter f] (u v : α → ℝ≥0∞) : f.limsup (u + v) ≤ f.limsup u + f.limsup v := Inf_le ((eventually_le_limsup u).mp ((eventually_le_limsup v).mono (λ _ hxg hxf, add_le_add hxf hxg))) diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 136165551aa9c..2160c5421acc8 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -16,24 +16,24 @@ open_locale classical topological_space ennreal nnreal big_operators filter variables {α : Type*} {β : Type*} {γ : Type*} namespace ennreal -variables {a b c d : ennreal} {r p q : ℝ≥0} -variables {x y z : ennreal} {ε ε₁ ε₂ : ennreal} {s : set ennreal} +variables {a b c d : ℝ≥0∞} {r p q : ℝ≥0} +variables {x y z : ℝ≥0∞} {ε ε₁ ε₂ : ℝ≥0∞} {s : set ℝ≥0∞} section topological_space open topological_space -/-- Topology on `ennreal`. +/-- Topology on `ℝ≥0∞`. Note: this is different from the `emetric_space` topology. The `emetric_space` topology has `is_open {⊤}`, while this topology doesn't have singleton elements. -/ -instance : topological_space ennreal := preorder.topology ennreal +instance : topological_space ℝ≥0∞ := preorder.topology ℝ≥0∞ -instance : order_topology ennreal := ⟨rfl⟩ +instance : order_topology ℝ≥0∞ := ⟨rfl⟩ -instance : t2_space ennreal := by apply_instance -- short-circuit type class inference +instance : t2_space ℝ≥0∞ := by apply_instance -- short-circuit type class inference -instance : second_countable_topology ennreal := -⟨⟨⋃q ≥ (0:ℚ), {{a : ennreal | a < nnreal.of_real q}, {a : ennreal | ↑(nnreal.of_real q) < a}}, +instance : second_countable_topology ℝ≥0∞ := +⟨⟨⋃q ≥ (0:ℚ), {{a : ℝ≥0∞ | a < nnreal.of_real q}, {a : ℝ≥0∞ | ↑(nnreal.of_real q) < a}}, (countable_encodable _).bUnion $ assume a ha, (countable_singleton _).insert _, le_antisymm (le_generate_from $ by simp [or_imp_distrib, is_open_lt', is_open_gt'] {contextual := tt}) @@ -49,10 +49,10 @@ instance : second_countable_topology ennreal := exact generate_open.basic _ (mem_bUnion hq.1 $ by simp) } end)⟩⟩ -lemma embedding_coe : embedding (coe : ℝ≥0 → ennreal) := +lemma embedding_coe : embedding (coe : ℝ≥0 → ℝ≥0∞) := ⟨⟨begin refine le_antisymm _ _, - { rw [@order_topology.topology_eq_generate_intervals ennreal _, + { rw [@order_topology.topology_eq_generate_intervals ℝ≥0∞ _, ← coinduced_le_iff_le_induced], refine le_generate_from (assume s ha, _), rcases ha with ⟨a, rfl | rfl⟩, @@ -68,30 +68,30 @@ lemma embedding_coe : embedding (coe : ℝ≥0 → ennreal) := end⟩, assume a b, coe_eq_coe.1⟩ -lemma is_open_ne_top : is_open {a : ennreal | a ≠ ⊤} := is_open_ne +lemma is_open_ne_top : is_open {a : ℝ≥0∞ | a ≠ ⊤} := is_open_ne lemma is_open_Ico_zero : is_open (Ico 0 b) := by { rw ennreal.Ico_eq_Iio, exact is_open_Iio} -lemma coe_range_mem_nhds : range (coe : ℝ≥0 → ennreal) ∈ 𝓝 (r : ennreal) := -have {a : ennreal | a ≠ ⊤} = range (coe : ℝ≥0 → ennreal), +lemma coe_range_mem_nhds : range (coe : ℝ≥0 → ℝ≥0∞) ∈ 𝓝 (r : ℝ≥0∞) := +have {a : ℝ≥0∞ | a ≠ ⊤} = range (coe : ℝ≥0 → ℝ≥0∞), from set.ext $ assume a, by cases a; simp [none_eq_top, some_eq_coe], this ▸ mem_nhds_sets is_open_ne_top coe_ne_top @[norm_cast] lemma tendsto_coe {f : filter α} {m : α → ℝ≥0} {a : ℝ≥0} : - tendsto (λa, (m a : ennreal)) f (𝓝 ↑a) ↔ tendsto m f (𝓝 a) := + tendsto (λa, (m a : ℝ≥0∞)) f (𝓝 ↑a) ↔ tendsto m f (𝓝 a) := embedding_coe.tendsto_nhds_iff.symm -lemma continuous_coe : continuous (coe : ℝ≥0 → ennreal) := +lemma continuous_coe : continuous (coe : ℝ≥0 → ℝ≥0∞) := embedding_coe.continuous lemma continuous_coe_iff {α} [topological_space α] {f : α → ℝ≥0} : - continuous (λa, (f a : ennreal)) ↔ continuous f := + continuous (λa, (f a : ℝ≥0∞)) ↔ continuous f := embedding_coe.continuous_iff.symm -lemma nhds_coe {r : ℝ≥0} : 𝓝 (r : ennreal) = (𝓝 r).map coe := +lemma nhds_coe {r : ℝ≥0} : 𝓝 (r : ℝ≥0∞) = (𝓝 r).map coe := by rw [embedding_coe.induced, map_nhds_induced_eq coe_range_mem_nhds] -lemma nhds_coe_coe {r p : ℝ≥0} : 𝓝 ((r : ennreal), (p : ennreal)) = +lemma nhds_coe_coe {r p : ℝ≥0} : 𝓝 ((r : ℝ≥0∞), (p : ℝ≥0∞)) = (𝓝 (r, p)).map (λp:ℝ≥0×ℝ≥0, (p.1, p.2)) := begin rw [(embedding_coe.prod_mk embedding_coe).map_nhds_eq], @@ -106,7 +106,7 @@ lemma tendsto_of_real {f : filter α} {m : α → ℝ} {a : ℝ} (h : tendsto m tendsto (λa, ennreal.of_real (m a)) f (𝓝 (ennreal.of_real a)) := tendsto.comp (continuous.tendsto continuous_of_real _) h -lemma tendsto_to_nnreal {a : ennreal} : a ≠ ⊤ → +lemma tendsto_to_nnreal {a : ℝ≥0∞} : a ≠ ⊤ → tendsto (ennreal.to_nnreal) (𝓝 a) (𝓝 a.to_nnreal) := begin cases a; simp [some_eq_coe, none_eq_top, nhds_coe, tendsto_map'_iff, (∘)], @@ -117,17 +117,17 @@ lemma continuous_on_to_nnreal : continuous_on ennreal.to_nnreal {a | a ≠ ∞} continuous_on_iff_continuous_restrict.2 $ continuous_iff_continuous_at.2 $ λ x, (tendsto_to_nnreal x.2).comp continuous_at_subtype_coe -lemma tendsto_to_real {a : ennreal} : a ≠ ⊤ → tendsto (ennreal.to_real) (𝓝 a) (𝓝 a.to_real) := +lemma tendsto_to_real {a : ℝ≥0∞} : a ≠ ⊤ → tendsto (ennreal.to_real) (𝓝 a) (𝓝 a.to_real) := λ ha, tendsto.comp ((@nnreal.tendsto_coe _ (𝓝 a.to_nnreal) id (a.to_nnreal)).2 tendsto_id) (tendsto_to_nnreal ha) -/-- The set of finite `ennreal` numbers is homeomorphic to `ℝ≥0`. -/ +/-- The set of finite `ℝ≥0∞` numbers is homeomorphic to `ℝ≥0`. -/ def ne_top_homeomorph_nnreal : {a | a ≠ ∞} ≃ₜ ℝ≥0 := { continuous_to_fun := continuous_on_iff_continuous_restrict.1 continuous_on_to_nnreal, continuous_inv_fun := continuous_subtype_mk _ continuous_coe, .. ne_top_equiv_nnreal } -/-- The set of finite `ennreal` numbers is homeomorphic to `ℝ≥0`. -/ +/-- The set of finite `ℝ≥0∞` numbers is homeomorphic to `ℝ≥0`. -/ def lt_top_homeomorph_nnreal : {a | a < ∞} ≃ₜ ℝ≥0 := by refine (homeomorph.set_congr $ set.ext $ λ x, _).trans ne_top_homeomorph_nnreal; simp only [mem_set_of_eq, lt_top_iff_ne_top] @@ -138,17 +138,17 @@ nhds_top_order.trans $ by simp [lt_top_iff_ne_top, Ioi] lemma nhds_top' : 𝓝 ∞ = ⨅ r : ℝ≥0, 𝓟 (Ioi r) := nhds_top.trans $ infi_ne_top _ -lemma tendsto_nhds_top_iff_nnreal {m : α → ennreal} {f : filter α} : +lemma tendsto_nhds_top_iff_nnreal {m : α → ℝ≥0∞} {f : filter α} : tendsto m f (𝓝 ⊤) ↔ ∀ x : ℝ≥0, ∀ᶠ a in f, ↑x < m a := by simp only [nhds_top', tendsto_infi, tendsto_principal, mem_Ioi] -lemma tendsto_nhds_top_iff_nat {m : α → ennreal} {f : filter α} : +lemma tendsto_nhds_top_iff_nat {m : α → ℝ≥0∞} {f : filter α} : tendsto m f (𝓝 ⊤) ↔ ∀ n : ℕ, ∀ᶠ a in f, ↑n < m a := tendsto_nhds_top_iff_nnreal.trans ⟨λ h n, by simpa only [ennreal.coe_nat] using h n, λ h x, let ⟨n, hn⟩ := exists_nat_gt x in (h n).mono (λ y, lt_trans $ by rwa [← ennreal.coe_nat, coe_lt_coe])⟩ -lemma tendsto_nhds_top {m : α → ennreal} {f : filter α} +lemma tendsto_nhds_top {m : α → ℝ≥0∞} {f : filter α} (h : ∀ n : ℕ, ∀ᶠ a in f, ↑n < m a) : tendsto m f (𝓝 ⊤) := tendsto_nhds_top_iff_nat.2 h @@ -157,17 +157,17 @@ tendsto_nhds_top $ λ n, mem_at_top_sets.2 ⟨n+1, λ m hm, ennreal.coe_nat_lt_coe_nat.2 $ nat.lt_of_succ_le hm⟩ @[simp, norm_cast] lemma tendsto_coe_nhds_top {f : α → ℝ≥0} {l : filter α} : - tendsto (λ x, (f x : ennreal)) l (𝓝 ∞) ↔ tendsto f l at_top := + tendsto (λ x, (f x : ℝ≥0∞)) l (𝓝 ∞) ↔ tendsto f l at_top := by rw [tendsto_nhds_top_iff_nnreal, at_top_basis_Ioi.tendsto_right_iff]; [simp, apply_instance, apply_instance] -lemma nhds_zero : 𝓝 (0 : ennreal) = ⨅a ≠ 0, 𝓟 (Iio a) := +lemma nhds_zero : 𝓝 (0 : ℝ≥0∞) = ⨅a ≠ 0, 𝓟 (Iio a) := nhds_bot_order.trans $ by simp [bot_lt_iff_ne_bot, Iio] -@[instance] lemma nhds_within_Ioi_coe_ne_bot {r : ℝ≥0} : (𝓝[Ioi r] (r : ennreal)).ne_bot := +@[instance] lemma nhds_within_Ioi_coe_ne_bot {r : ℝ≥0} : (𝓝[Ioi r] (r : ℝ≥0∞)).ne_bot := nhds_within_Ioi_self_ne_bot' ennreal.coe_lt_top -@[instance] lemma nhds_within_Ioi_zero_ne_bot : (𝓝[Ioi 0] (0 : ennreal)).ne_bot := +@[instance] lemma nhds_within_Ioi_zero_ne_bot : (𝓝[Ioi 0] (0 : ℝ≥0∞)).ne_bot := nhds_within_Ioi_coe_ne_bot -- using Icc because @@ -209,17 +209,17 @@ begin assume y, rintros ⟨h₁, h₂⟩, rw xbx at h₂, calc y ≤ b : h₂ ... < a : ba }, end -/-- Characterization of neighborhoods for `ennreal` numbers. See also `tendsto_order` +/-- Characterization of neighborhoods for `ℝ≥0∞` numbers. See also `tendsto_order` for a version with strict inequalities. -/ -protected theorem tendsto_nhds {f : filter α} {u : α → ennreal} {a : ennreal} (ha : a ≠ ⊤) : +protected theorem tendsto_nhds {f : filter α} {u : α → ℝ≥0∞} {a : ℝ≥0∞} (ha : a ≠ ⊤) : tendsto u f (𝓝 a) ↔ ∀ ε > 0, ∀ᶠ x in f, (u x) ∈ Icc (a - ε) (a + ε) := by simp only [nhds_of_ne_top ha, tendsto_infi, tendsto_principal, mem_Icc] -protected lemma tendsto_at_top [nonempty β] [semilattice_sup β] {f : β → ennreal} {a : ennreal} +protected lemma tendsto_at_top [nonempty β] [semilattice_sup β] {f : β → ℝ≥0∞} {a : ℝ≥0∞} (ha : a ≠ ⊤) : tendsto f at_top (𝓝 a) ↔ ∀ε>0, ∃N, ∀n≥N, (f n) ∈ Icc (a - ε) (a + ε) := by simp only [ennreal.tendsto_nhds ha, mem_at_top_sets, mem_set_of_eq, filter.eventually] -instance : has_continuous_add ennreal := +instance : has_continuous_add ℝ≥0∞ := begin refine ⟨continuous_iff_continuous_at.2 _⟩, rintro ⟨(_|a), b⟩, @@ -231,8 +231,8 @@ begin end protected lemma tendsto_mul (ha : a ≠ 0 ∨ b ≠ ⊤) (hb : b ≠ 0 ∨ a ≠ ⊤) : - tendsto (λp:ennreal×ennreal, p.1 * p.2) (𝓝 (a, b)) (𝓝 (a * b)) := -have ht : ∀b:ennreal, b ≠ 0 → tendsto (λp:ennreal×ennreal, p.1 * p.2) (𝓝 ((⊤:ennreal), b)) (𝓝 ⊤), + tendsto (λp:ℝ≥0∞×ℝ≥0∞, p.1 * p.2) (𝓝 (a, b)) (𝓝 (a * b)) := +have ht : ∀b:ℝ≥0∞, b ≠ 0 → tendsto (λp:ℝ≥0∞×ℝ≥0∞, p.1 * p.2) (𝓝 ((⊤:ℝ≥0∞), b)) (𝓝 ⊤), begin refine assume b hb, tendsto_nhds_top_iff_nnreal.2 $ assume n, _, rcases lt_iff_exists_nnreal_btwn.1 (pos_iff_ne_zero.2 hb) with ⟨ε, hε, hεb⟩, @@ -247,52 +247,52 @@ begin cases a, {simp [none_eq_top] at hb, simp [none_eq_top, ht b hb, top_mul, hb] }, cases b, { simp [none_eq_top] at ha, - simp [*, nhds_swap (a : ennreal) ⊤, none_eq_top, some_eq_coe, top_mul, tendsto_map'_iff, (∘), + simp [*, nhds_swap (a : ℝ≥0∞) ⊤, none_eq_top, some_eq_coe, top_mul, tendsto_map'_iff, (∘), mul_comm] }, simp [some_eq_coe, nhds_coe_coe, tendsto_map'_iff, (∘)], simp only [coe_mul.symm, tendsto_coe, tendsto_mul] end -protected lemma tendsto.mul {f : filter α} {ma : α → ennreal} {mb : α → ennreal} {a b : ennreal} +protected lemma tendsto.mul {f : filter α} {ma : α → ℝ≥0∞} {mb : α → ℝ≥0∞} {a b : ℝ≥0∞} (hma : tendsto ma f (𝓝 a)) (ha : a ≠ 0 ∨ b ≠ ⊤) (hmb : tendsto mb f (𝓝 b)) (hb : b ≠ 0 ∨ a ≠ ⊤) : tendsto (λa, ma a * mb a) f (𝓝 (a * b)) := -show tendsto ((λp:ennreal×ennreal, p.1 * p.2) ∘ (λa, (ma a, mb a))) f (𝓝 (a * b)), from +show tendsto ((λp:ℝ≥0∞×ℝ≥0∞, p.1 * p.2) ∘ (λa, (ma a, mb a))) f (𝓝 (a * b)), from tendsto.comp (ennreal.tendsto_mul ha hb) (hma.prod_mk_nhds hmb) -protected lemma tendsto.const_mul {f : filter α} {m : α → ennreal} {a b : ennreal} +protected lemma tendsto.const_mul {f : filter α} {m : α → ℝ≥0∞} {a b : ℝ≥0∞} (hm : tendsto m f (𝓝 b)) (hb : b ≠ 0 ∨ a ≠ ⊤) : tendsto (λb, a * m b) f (𝓝 (a * b)) := by_cases (assume : a = 0, by simp [this, tendsto_const_nhds]) (assume ha : a ≠ 0, ennreal.tendsto.mul tendsto_const_nhds (or.inl ha) hm hb) -protected lemma tendsto.mul_const {f : filter α} {m : α → ennreal} {a b : ennreal} +protected lemma tendsto.mul_const {f : filter α} {m : α → ℝ≥0∞} {a b : ℝ≥0∞} (hm : tendsto m f (𝓝 a)) (ha : a ≠ 0 ∨ b ≠ ⊤) : tendsto (λx, m x * b) f (𝓝 (a * b)) := by simpa only [mul_comm] using ennreal.tendsto.const_mul hm ha -protected lemma continuous_at_const_mul {a b : ennreal} (h : a ≠ ⊤ ∨ b ≠ 0) : +protected lemma continuous_at_const_mul {a b : ℝ≥0∞} (h : a ≠ ⊤ ∨ b ≠ 0) : continuous_at ((*) a) b := tendsto.const_mul tendsto_id h.symm -protected lemma continuous_at_mul_const {a b : ennreal} (h : a ≠ ⊤ ∨ b ≠ 0) : +protected lemma continuous_at_mul_const {a b : ℝ≥0∞} (h : a ≠ ⊤ ∨ b ≠ 0) : continuous_at (λ x, x * a) b := tendsto.mul_const tendsto_id h.symm -protected lemma continuous_const_mul {a : ennreal} (ha : a ≠ ⊤) : continuous ((*) a) := +protected lemma continuous_const_mul {a : ℝ≥0∞} (ha : a ≠ ⊤) : continuous ((*) a) := continuous_iff_continuous_at.2 $ λ x, ennreal.continuous_at_const_mul (or.inl ha) -protected lemma continuous_mul_const {a : ennreal} (ha : a ≠ ⊤) : continuous (λ x, x * a) := +protected lemma continuous_mul_const {a : ℝ≥0∞} (ha : a ≠ ⊤) : continuous (λ x, x * a) := continuous_iff_continuous_at.2 $ λ x, ennreal.continuous_at_mul_const (or.inl ha) -lemma le_of_forall_lt_one_mul_le {x y : ennreal} (h : ∀ a < 1, a * x ≤ y) : x ≤ y := +lemma le_of_forall_lt_one_mul_le {x y : ℝ≥0∞} (h : ∀ a < 1, a * x ≤ y) : x ≤ y := begin have : tendsto (* x) (𝓝[Iio 1] 1) (𝓝 (1 * x)) := (ennreal.continuous_at_mul_const (or.inr one_ne_zero)).mono_left inf_le_left, rw one_mul at this, - haveI : (𝓝[Iio 1] (1 : ennreal)).ne_bot := nhds_within_Iio_self_ne_bot' ennreal.zero_lt_one, + haveI : (𝓝[Iio 1] (1 : ℝ≥0∞)).ne_bot := nhds_within_Iio_self_ne_bot' ennreal.zero_lt_one, exact le_of_tendsto this (eventually_nhds_within_iff.2 $ eventually_of_forall h) end -lemma infi_mul_left {ι} [nonempty ι] {f : ι → ennreal} {a : ennreal} +lemma infi_mul_left {ι} [nonempty ι] {f : ι → ℝ≥0∞} {a : ℝ≥0∞} (h : a = ⊤ → (⨅ i, f i) = 0 → ∃ i, f i = 0) : (⨅ i, a * f i) = a * ⨅ i, f i := begin @@ -305,12 +305,12 @@ begin ennreal.mul_left_mono).symm } end -lemma infi_mul_right {ι} [nonempty ι] {f : ι → ennreal} {a : ennreal} +lemma infi_mul_right {ι} [nonempty ι] {f : ι → ℝ≥0∞} {a : ℝ≥0∞} (h : a = ⊤ → (⨅ i, f i) = 0 → ∃ i, f i = 0) : (⨅ i, f i * a) = (⨅ i, f i) * a := by simpa only [mul_comm a] using infi_mul_left h -protected lemma continuous_inv : continuous (has_inv.inv : ennreal → ennreal) := +protected lemma continuous_inv : continuous (has_inv.inv : ℝ≥0∞ → ℝ≥0∞) := continuous_iff_continuous_at.2 $ λ a, tendsto_order.2 ⟨begin assume b hb, @@ -323,29 +323,29 @@ begin exact lt_mem_nhds (ennreal.inv_lt_iff_inv_lt.1 hb) end⟩ -@[simp] protected lemma tendsto_inv_iff {f : filter α} {m : α → ennreal} {a : ennreal} : +@[simp] protected lemma tendsto_inv_iff {f : filter α} {m : α → ℝ≥0∞} {a : ℝ≥0∞} : tendsto (λ x, (m x)⁻¹) f (𝓝 a⁻¹) ↔ tendsto m f (𝓝 a) := ⟨λ h, by simpa only [function.comp, ennreal.inv_inv] using (ennreal.continuous_inv.tendsto a⁻¹).comp h, (ennreal.continuous_inv.tendsto a).comp⟩ -protected lemma tendsto.div {f : filter α} {ma : α → ennreal} {mb : α → ennreal} {a b : ennreal} +protected lemma tendsto.div {f : filter α} {ma : α → ℝ≥0∞} {mb : α → ℝ≥0∞} {a b : ℝ≥0∞} (hma : tendsto ma f (𝓝 a)) (ha : a ≠ 0 ∨ b ≠ 0) (hmb : tendsto mb f (𝓝 b)) (hb : b ≠ ⊤ ∨ a ≠ ⊤) : tendsto (λa, ma a / mb a) f (𝓝 (a / b)) := by { apply tendsto.mul hma _ (ennreal.tendsto_inv_iff.2 hmb) _; simp [ha, hb] } -protected lemma tendsto.const_div {f : filter α} {m : α → ennreal} {a b : ennreal} +protected lemma tendsto.const_div {f : filter α} {m : α → ℝ≥0∞} {a b : ℝ≥0∞} (hm : tendsto m f (𝓝 b)) (hb : b ≠ ⊤ ∨ a ≠ ⊤) : tendsto (λb, a / m b) f (𝓝 (a / b)) := by { apply tendsto.const_mul (ennreal.tendsto_inv_iff.2 hm), simp [hb] } -protected lemma tendsto.div_const {f : filter α} {m : α → ennreal} {a b : ennreal} +protected lemma tendsto.div_const {f : filter α} {m : α → ℝ≥0∞} {a b : ℝ≥0∞} (hm : tendsto m f (𝓝 a)) (ha : a ≠ 0 ∨ b ≠ 0) : tendsto (λx, m x / b) f (𝓝 (a / b)) := by { apply tendsto.mul_const hm, simp [ha] } -protected lemma tendsto_inv_nat_nhds_zero : tendsto (λ n : ℕ, (n : ennreal)⁻¹) at_top (𝓝 0) := +protected lemma tendsto_inv_nat_nhds_zero : tendsto (λ n : ℕ, (n : ℝ≥0∞)⁻¹) at_top (𝓝 0) := ennreal.inv_top ▸ ennreal.tendsto_inv_iff.2 tendsto_nat_nhds_top -lemma bsupr_add {ι} {s : set ι} (hs : s.nonempty) {f : ι → ennreal} : +lemma bsupr_add {ι} {s : set ι} (hs : s.nonempty) {f : ι → ℝ≥0∞} : (⨆ i ∈ s, f i) + a = ⨆ i ∈ s, f i + a := begin simp only [← Sup_image], symmetry, @@ -355,19 +355,19 @@ begin tendsto.add (tendsto_id' inf_le_left) tendsto_const_nhds] end -lemma Sup_add {s : set ennreal} (hs : s.nonempty) : Sup s + a = ⨆b∈s, b + a := +lemma Sup_add {s : set ℝ≥0∞} (hs : s.nonempty) : Sup s + a = ⨆b∈s, b + a := by rw [Sup_eq_supr, bsupr_add hs] -lemma supr_add {ι : Sort*} {s : ι → ennreal} [h : nonempty ι] : supr s + a = ⨆b, s b + a := +lemma supr_add {ι : Sort*} {s : ι → ℝ≥0∞} [h : nonempty ι] : supr s + a = ⨆b, s b + a := let ⟨x⟩ := h in calc supr s + a = Sup (range s) + a : by rw Sup_range ... = (⨆b∈range s, b + a) : Sup_add ⟨s x, x, rfl⟩ ... = _ : supr_range -lemma add_supr {ι : Sort*} {s : ι → ennreal} [h : nonempty ι] : a + supr s = ⨆b, a + s b := +lemma add_supr {ι : Sort*} {s : ι → ℝ≥0∞} [h : nonempty ι] : a + supr s = ⨆b, a + s b := by rw [add_comm, supr_add]; simp [add_comm] -lemma supr_add_supr {ι : Sort*} {f g : ι → ennreal} (h : ∀i j, ∃k, f i + g j ≤ f k + g k) : +lemma supr_add_supr {ι : Sort*} {f g : ι → ℝ≥0∞} (h : ∀i j, ∃k, f i + g j ≤ f k + g k) : supr f + supr g = (⨆ a, f a + g a) := begin by_cases hι : nonempty ι, @@ -376,16 +376,16 @@ begin simpa [add_supr, supr_add] using λ i j:ι, show f i + g j ≤ ⨆ a, f a + g a, from let ⟨k, hk⟩ := h i j in le_supr_of_le k hk }, - { have : ∀f:ι → ennreal, (⨆i, f i) = 0 := λ f, bot_unique (supr_le $ assume i, (hι ⟨i⟩).elim), + { have : ∀f:ι → ℝ≥0∞, (⨆i, f i) = 0 := λ f, bot_unique (supr_le $ assume i, (hι ⟨i⟩).elim), rw [this, this, this, zero_add] } end lemma supr_add_supr_of_monotone {ι : Sort*} [semilattice_sup ι] - {f g : ι → ennreal} (hf : monotone f) (hg : monotone g) : + {f g : ι → ℝ≥0∞} (hf : monotone f) (hg : monotone g) : supr f + supr g = (⨆ a, f a + g a) := supr_add_supr $ assume i j, ⟨i ⊔ j, add_le_add (hf $ le_sup_left) (hg $ le_sup_right)⟩ -lemma finset_sum_supr_nat {α} {ι} [semilattice_sup ι] {s : finset α} {f : α → ι → ennreal} +lemma finset_sum_supr_nat {α} {ι} [semilattice_sup ι] {s : finset α} {f : α → ι → ℝ≥0∞} (hf : ∀a, monotone (f a)) : ∑ a in s, supr (f a) = (⨆ n, ∑ a in s, f a n) := begin @@ -399,9 +399,9 @@ begin exact (finset.sum_le_sum $ assume a ha, hf a h) } end -lemma mul_Sup {s : set ennreal} {a : ennreal} : a * Sup s = ⨆i∈s, a * i := +lemma mul_Sup {s : set ℝ≥0∞} {a : ℝ≥0∞} : a * Sup s = ⨆i∈s, a * i := begin - by_cases hs : ∀x∈s, x = (0:ennreal), + by_cases hs : ∀x∈s, x = (0:ℝ≥0∞), { have h₁ : Sup s = 0 := (bot_unique $ Sup_le $ assume a ha, (hs a ha).symm ▸ le_refl 0), have h₂ : (⨆i ∈ s, a * i) = 0 := (bot_unique $ supr_le $ assume a, supr_le $ assume ha, by simp [hs a ha]), @@ -419,13 +419,13 @@ begin rw [this.symm, Sup_image] } end -lemma mul_supr {ι : Sort*} {f : ι → ennreal} {a : ennreal} : a * supr f = ⨆i, a * f i := +lemma mul_supr {ι : Sort*} {f : ι → ℝ≥0∞} {a : ℝ≥0∞} : a * supr f = ⨆i, a * f i := by rw [← Sup_range, mul_Sup, supr_range] -lemma supr_mul {ι : Sort*} {f : ι → ennreal} {a : ennreal} : supr f * a = ⨆i, f i * a := +lemma supr_mul {ι : Sort*} {f : ι → ℝ≥0∞} {a : ℝ≥0∞} : supr f * a = ⨆i, f i * a := by rw [mul_comm, mul_supr]; congr; funext; rw [mul_comm] -protected lemma tendsto_coe_sub : ∀{b:ennreal}, tendsto (λb:ennreal, ↑r - b) (𝓝 b) (𝓝 (↑r - b)) := +protected lemma tendsto_coe_sub : ∀{b:ℝ≥0∞}, tendsto (λb:ℝ≥0∞, ↑r - b) (𝓝 b) (𝓝 (↑r - b)) := begin refine (forall_ennreal.2 $ and.intro (assume a, _) _), { simp [@nhds_coe a, tendsto_map'_iff, (∘), tendsto_coe, coe_sub.symm], @@ -435,7 +435,7 @@ begin by simp [le_of_lt] {contextual := tt})) tendsto_const_nhds end -lemma sub_supr {ι : Sort*} [hι : nonempty ι] {b : ι → ennreal} (hr : a < ⊤) : +lemma sub_supr {ι : Sort*} [hι : nonempty ι] {b : ι → ℝ≥0∞} (hr : a < ⊤) : a - (⨆i, b i) = (⨅i, a - b i) := let ⟨i⟩ := hι in let ⟨r, eq, _⟩ := lt_iff_exists_coe.mp hr in @@ -447,25 +447,25 @@ have Inf ((λb, ↑r - b) '' range b) = ↑r - (⨆i, b i), (tendsto.comp ennreal.tendsto_coe_sub (tendsto_id' inf_le_left)), by rw [eq, ←this]; simp [Inf_image, infi_range, -mem_range]; exact le_refl _ -lemma supr_eq_zero {ι : Sort*} {f : ι → ennreal} : (⨆ i, f i) = 0 ↔ ∀ i, f i = 0 := +lemma supr_eq_zero {ι : Sort*} {f : ι → ℝ≥0∞} : (⨆ i, f i) = 0 ↔ ∀ i, f i = 0 := by simp_rw [← nonpos_iff_eq_zero, supr_le_iff] end topological_space section tsum -variables {f g : α → ennreal} +variables {f g : α → ℝ≥0∞} @[norm_cast] protected lemma has_sum_coe {f : α → ℝ≥0} {r : ℝ≥0} : - has_sum (λa, (f a : ennreal)) ↑r ↔ has_sum f r := -have (λs:finset α, ∑ a in s, ↑(f a)) = (coe : ℝ≥0 → ennreal) ∘ (λs:finset α, ∑ a in s, f a), + has_sum (λa, (f a : ℝ≥0∞)) ↑r ↔ has_sum f r := +have (λs:finset α, ∑ a in s, ↑(f a)) = (coe : ℝ≥0 → ℝ≥0∞) ∘ (λs:finset α, ∑ a in s, f a), from funext $ assume s, ennreal.coe_finset_sum.symm, by unfold has_sum; rw [this, tendsto_coe] -protected lemma tsum_coe_eq {f : α → ℝ≥0} (h : has_sum f r) : ∑'a, (f a : ennreal) = r := +protected lemma tsum_coe_eq {f : α → ℝ≥0} (h : has_sum f r) : ∑'a, (f a : ℝ≥0∞) = r := (ennreal.has_sum_coe.2 h).tsum_eq -protected lemma coe_tsum {f : α → ℝ≥0} : summable f → ↑(tsum f) = ∑'a, (f a : ennreal) +protected lemma coe_tsum {f : α → ℝ≥0} : summable f → ↑(tsum f) = ∑'a, (f a : ℝ≥0∞) | ⟨r, hr⟩ := by rw [hr.tsum_eq, ennreal.tsum_coe_eq hr] protected lemma has_sum : has_sum f (⨆s:finset α, ∑ a in s, f a) := @@ -474,10 +474,10 @@ tendsto_at_top_supr $ λ s t, finset.sum_le_sum_of_subset @[simp] protected lemma summable : summable f := ⟨_, ennreal.has_sum⟩ lemma tsum_coe_ne_top_iff_summable {f : β → ℝ≥0} : - ∑' b, (f b:ennreal) ≠ ∞ ↔ summable f := + ∑' b, (f b:ℝ≥0∞) ≠ ∞ ↔ summable f := begin refine ⟨λ h, _, λ h, ennreal.coe_tsum h ▸ ennreal.coe_ne_top⟩, - lift (∑' b, (f b:ennreal)) to ℝ≥0 using h with a ha, + lift (∑' b, (f b:ℝ≥0∞)) to ℝ≥0 using h with a ha, refine ⟨a, ennreal.has_sum_coe.1 _⟩, rw ha, exact ennreal.summable.has_sum @@ -495,18 +495,18 @@ begin exact (finset.sum_mono_set f).supr_comp_eq hs end -protected lemma tsum_sigma {β : α → Type*} (f : Πa, β a → ennreal) : +protected lemma tsum_sigma {β : α → Type*} (f : Πa, β a → ℝ≥0∞) : ∑'p:Σa, β a, f p.1 p.2 = ∑'a b, f a b := tsum_sigma' (assume b, ennreal.summable) ennreal.summable -protected lemma tsum_sigma' {β : α → Type*} (f : (Σ a, β a) → ennreal) : +protected lemma tsum_sigma' {β : α → Type*} (f : (Σ a, β a) → ℝ≥0∞) : ∑'p:(Σa, β a), f p = ∑'a b, f ⟨a, b⟩ := tsum_sigma' (assume b, ennreal.summable) ennreal.summable -protected lemma tsum_prod {f : α → β → ennreal} : ∑'p:α×β, f p.1 p.2 = ∑'a, ∑'b, f a b := +protected lemma tsum_prod {f : α → β → ℝ≥0∞} : ∑'p:α×β, f p.1 p.2 = ∑'a, ∑'b, f a b := tsum_prod' ennreal.summable $ λ _, ennreal.summable -protected lemma tsum_comm {f : α → β → ennreal} : ∑'a, ∑'b, f a b = ∑'b, ∑'a, f a b := +protected lemma tsum_comm {f : α → β → ℝ≥0∞} : ∑'a, ∑'b, f a b = ∑'b, ∑'a, f a b := tsum_comm' ennreal.summable (λ _, ennreal.summable) (λ _, ennreal.summable) protected lemma tsum_add : ∑'a, (f a + g a) = (∑'a, f a) + (∑'a, g a) := @@ -515,17 +515,17 @@ tsum_add ennreal.summable ennreal.summable protected lemma tsum_le_tsum (h : ∀a, f a ≤ g a) : ∑'a, f a ≤ ∑'a, g a := tsum_le_tsum h ennreal.summable ennreal.summable -protected lemma sum_le_tsum {f : α → ennreal} (s : finset α) : ∑ x in s, f x ≤ ∑' x, f x := +protected lemma sum_le_tsum {f : α → ℝ≥0∞} (s : finset α) : ∑ x in s, f x ≤ ∑' x, f x := sum_le_tsum s (λ x hx, zero_le _) ennreal.summable -protected lemma tsum_eq_supr_nat' {f : ℕ → ennreal} {N : ℕ → ℕ} (hN : tendsto N at_top at_top) : +protected lemma tsum_eq_supr_nat' {f : ℕ → ℝ≥0∞} {N : ℕ → ℕ} (hN : tendsto N at_top at_top) : ∑'i:ℕ, f i = (⨆i:ℕ, ∑ a in finset.range (N i), f a) := ennreal.tsum_eq_supr_sum' _ $ λ t, let ⟨n, hn⟩ := t.exists_nat_subset_range, ⟨k, _, hk⟩ := exists_le_of_tendsto_at_top hN 0 n in ⟨k, finset.subset.trans hn (finset.range_mono hk)⟩ -protected lemma tsum_eq_supr_nat {f : ℕ → ennreal} : +protected lemma tsum_eq_supr_nat {f : ℕ → ℝ≥0∞} : ∑'i:ℕ, f i = (⨆i:ℕ, ∑ a in finset.range i, f a) := ennreal.tsum_eq_supr_sum' _ finset.exists_nat_subset_range @@ -553,7 +553,7 @@ has_sum.tsum_eq this protected lemma tsum_mul_right : (∑'i, f i * a) = (∑'i, f i) * a := by simp [mul_comm, ennreal.tsum_mul_left] -@[simp] lemma tsum_supr_eq {α : Type*} (a : α) {f : α → ennreal} : +@[simp] lemma tsum_supr_eq {α : Type*} (a : α) {f : α → ℝ≥0∞} : ∑'b:α, (⨆ (h : a = b), f b) = f a := le_antisymm (by rw [ennreal.tsum_eq_supr_sum]; exact supr_le (assume s, @@ -566,7 +566,7 @@ le_antisymm (calc f a ≤ (⨆ (h : a = a), f a) : le_supr (λh:a=a, f a) rfl ... ≤ (∑'b:α, ⨆ (h : a = b), f b) : ennreal.le_tsum _) -lemma has_sum_iff_tendsto_nat {f : ℕ → ennreal} (r : ennreal) : +lemma has_sum_iff_tendsto_nat {f : ℕ → ℝ≥0∞} (r : ℝ≥0∞) : has_sum f r ↔ tendsto (λn:ℕ, ∑ i in finset.range n, f i) at_top (𝓝 r) := begin refine ⟨has_sum.tendsto_sum_nat, assume h, _⟩, @@ -575,19 +575,19 @@ begin { exact assume s t hst, finset.sum_le_sum_of_subset (finset.range_subset.2 hst) } end -lemma to_nnreal_apply_of_tsum_ne_top {α : Type*} {f : α → ennreal} (hf : ∑' i, f i ≠ ∞) (x : α) : - (((ennreal.to_nnreal ∘ f) x : ℝ≥0) : ennreal) = f x := +lemma to_nnreal_apply_of_tsum_ne_top {α : Type*} {f : α → ℝ≥0∞} (hf : ∑' i, f i ≠ ∞) (x : α) : + (((ennreal.to_nnreal ∘ f) x : ℝ≥0) : ℝ≥0∞) = f x := coe_to_nnreal $ ennreal.ne_top_of_tsum_ne_top hf _ -lemma summable_to_nnreal_of_tsum_ne_top {α : Type*} {f : α → ennreal} (hf : ∑' i, f i ≠ ∞) : +lemma summable_to_nnreal_of_tsum_ne_top {α : Type*} {f : α → ℝ≥0∞} (hf : ∑' i, f i ≠ ∞) : summable (ennreal.to_nnreal ∘ f) := by simpa only [←tsum_coe_ne_top_iff_summable, to_nnreal_apply_of_tsum_ne_top hf] using hf -protected lemma tsum_apply {ι α : Type*} {f : ι → α → ennreal} {x : α} : +protected lemma tsum_apply {ι α : Type*} {f : ι → α → ℝ≥0∞} {x : α} : (∑' i, f i) x = ∑' i, f i x := tsum_apply $ pi.summable.mpr $ λ _, ennreal.summable -lemma tsum_sub {f : ℕ → ennreal} {g : ℕ → ennreal} (h₁ : ∑' i, g i < ∞) (h₂ : g ≤ f) : +lemma tsum_sub {f : ℕ → ℝ≥0∞} {g : ℕ → ℝ≥0∞} (h₁ : ∑' i, g i < ∞) (h₂ : g ≤ f) : ∑' i, (f i - g i) = (∑' i, f i) - (∑' i, g i) := begin have h₃: ∑' i, (f i - g i) = ∑' i, (f i - g i + g i) - ∑' i, g i, @@ -608,7 +608,7 @@ open_locale nnreal /-- Comparison test of convergence of `ℝ≥0`-valued series. -/ lemma exists_le_has_sum_of_le {f g : β → ℝ≥0} {r : ℝ≥0} (hgf : ∀b, g b ≤ f b) (hfr : has_sum f r) : ∃p≤r, has_sum g p := -have ∑'b, (g b : ennreal) ≤ r, +have ∑'b, (g b : ℝ≥0∞) ≤ r, begin refine has_sum_le (assume b, _) ennreal.summable.has_sum (ennreal.has_sum_coe.2 hfr), exact ennreal.coe_le_coe.2 (hgf _) @@ -687,10 +687,10 @@ end nnreal namespace ennreal -lemma tendsto_sum_nat_add (f : ℕ → ennreal) (hf : ∑' i, f i ≠ ∞) : +lemma tendsto_sum_nat_add (f : ℕ → ℝ≥0∞) (hf : ∑' i, f i ≠ ∞) : tendsto (λ i, ∑' k, f (k + i)) at_top (𝓝 0) := begin - have : ∀ i, ∑' k, (((ennreal.to_nnreal ∘ f) (k + i) : ℝ≥0) : ennreal) = + have : ∀ i, ∑' k, (((ennreal.to_nnreal ∘ f) (k + i) : ℝ≥0) : ℝ≥0∞) = (∑' k, (ennreal.to_nnreal ∘ f) (k + i) : ℝ≥0) := λ i, (ennreal.coe_tsum (nnreal.summable_nat_add _ (summable_to_nnreal_of_tsum_ne_top hf) _)).symm, @@ -775,7 +775,7 @@ variables [emetric_space β] open ennreal filter emetric /-- In an emetric ball, the distance between points is everywhere finite -/ -lemma edist_ne_top_of_mem_ball {a : β} {r : ennreal} (x y : ball a r) : edist x.1 y.1 ≠ ⊤ := +lemma edist_ne_top_of_mem_ball {a : β} {r : ℝ≥0∞} (x y : ball a r) : edist x.1 y.1 ≠ ⊤ := lt_top_iff_ne_top.1 $ calc edist x y ≤ edist a x + edist a y : edist_triangle_left x.1 y.1 a ... < r + r : by rw [edist_comm a x, edist_comm a y]; exact add_lt_add x.2 y.2 @@ -783,12 +783,12 @@ calc edist x y ≤ edist a x + edist a y : edist_triangle_left x.1 y.1 a /-- Each ball in an extended metric space gives us a metric space, as the edist is everywhere finite. -/ -def metric_space_emetric_ball (a : β) (r : ennreal) : metric_space (ball a r) := +def metric_space_emetric_ball (a : β) (r : ℝ≥0∞) : metric_space (ball a r) := emetric_space.to_metric_space edist_ne_top_of_mem_ball local attribute [instance] metric_space_emetric_ball -lemma nhds_eq_nhds_emetric_ball (a x : β) (r : ennreal) (h : x ∈ ball a r) : +lemma nhds_eq_nhds_emetric_ball (a x : β) (r : ℝ≥0∞) (h : x ∈ ball a r) : 𝓝 x = map (coe : ball a r → β) (𝓝 ⟨x, h⟩) := (map_nhds_subtype_coe_eq _ $ mem_nhds_sets emetric.is_open_ball h).symm end @@ -800,12 +800,12 @@ open emetric lemma tendsto_iff_edist_tendsto_0 {l : filter β} {f : β → α} {y : α} : tendsto f l (𝓝 y) ↔ tendsto (λ x, edist (f x) y) l (𝓝 0) := by simp only [emetric.nhds_basis_eball.tendsto_right_iff, emetric.mem_ball, - @tendsto_order ennreal β _ _, forall_prop_of_false ennreal.not_lt_zero, forall_const, true_and] + @tendsto_order ℝ≥0∞ β _ _, forall_prop_of_false ennreal.not_lt_zero, forall_const, true_and] /-- Yet another metric characterization of Cauchy sequences on integers. This one is often the most efficient. -/ lemma emetric.cauchy_seq_iff_le_tendsto_0 [nonempty β] [semilattice_sup β] {s : β → α} : - cauchy_seq s ↔ (∃ (b: β → ennreal), (∀ n m N : β, N ≤ n → N ≤ m → edist (s n) (s m) ≤ b N) + cauchy_seq s ↔ (∃ (b: β → ℝ≥0∞), (∀ n m N : β, N ≤ n → N ≤ m → edist (s n) (s m) ≤ b N) ∧ (tendsto b at_top (𝓝 0))) := ⟨begin assume hs, @@ -847,7 +847,7 @@ begin ... < ε : (hN _ (le_refl N)) ⟩ end⟩ -lemma continuous_of_le_add_edist {f : α → ennreal} (C : ennreal) +lemma continuous_of_le_add_edist {f : α → ℝ≥0∞} (C : ℝ≥0∞) (hC : C ≠ ⊤) (h : ∀x y, f x ≤ f y + C * edist x y) : continuous f := begin refine continuous_iff_continuous_at.2 (λx, tendsto_order.2 ⟨_, _⟩), @@ -924,7 +924,7 @@ theorem filter.tendsto.edist {f g : β → α} {x : filter β} {a b : α} tendsto (λx, edist (f x) (g x)) x (𝓝 (edist a b)) := (continuous_edist.tendsto (a, b)).comp (hf.prod_mk_nhds hg) -lemma cauchy_seq_of_edist_le_of_tsum_ne_top {f : ℕ → α} (d : ℕ → ennreal) +lemma cauchy_seq_of_edist_le_of_tsum_ne_top {f : ℕ → α} (d : ℕ → ℝ≥0∞) (hf : ∀ n, edist (f n) (f n.succ) ≤ d n) (hd : tsum d ≠ ∞) : cauchy_seq f := begin @@ -933,12 +933,12 @@ begin exact cauchy_seq_of_edist_le_of_summable d hf hd end -lemma emetric.is_closed_ball {a : α} {r : ennreal} : is_closed (closed_ball a r) := +lemma emetric.is_closed_ball {a : α} {r : ℝ≥0∞} : is_closed (closed_ball a r) := is_closed_le (continuous_id.edist continuous_const) continuous_const -/-- If `edist (f n) (f (n+1))` is bounded above by a function `d : ℕ → ennreal`, +/-- If `edist (f n) (f (n+1))` is bounded above by a function `d : ℕ → ℝ≥0∞`, then the distance from `f n` to the limit is bounded by `∑'_{k=n}^∞ d k`. -/ -lemma edist_le_tsum_of_edist_le_of_tendsto {f : ℕ → α} (d : ℕ → ennreal) +lemma edist_le_tsum_of_edist_le_of_tendsto {f : ℕ → α} (d : ℕ → ℝ≥0∞) (hf : ∀ n, edist (f n) (f n.succ) ≤ d n) {a : α} (ha : tendsto f at_top (𝓝 a)) (n : ℕ) : edist (f n) a ≤ ∑' m, d (n + m) := @@ -950,9 +950,9 @@ begin exact sum_le_tsum _ (λ _ _, zero_le _) ennreal.summable end -/-- If `edist (f n) (f (n+1))` is bounded above by a function `d : ℕ → ennreal`, +/-- If `edist (f n) (f (n+1))` is bounded above by a function `d : ℕ → ℝ≥0∞`, then the distance from `f 0` to the limit is bounded by `∑'_{k=0}^∞ d k`. -/ -lemma edist_le_tsum_of_edist_le_of_tendsto₀ {f : ℕ → α} (d : ℕ → ennreal) +lemma edist_le_tsum_of_edist_le_of_tendsto₀ {f : ℕ → α} (d : ℕ → ℝ≥0∞) (hf : ∀ n, edist (f n) (f n.succ) ≤ d n) {a : α} (ha : tendsto f at_top (𝓝 a)) : edist (f 0) a ≤ ∑' m, d m := diff --git a/src/topology/metric_space/antilipschitz.lean b/src/topology/metric_space/antilipschitz.lean index 596c61b844483..651461104cb4f 100644 --- a/src/topology/metric_space/antilipschitz.lean +++ b/src/topology/metric_space/antilipschitz.lean @@ -15,7 +15,7 @@ For a metric space, the latter inequality is equivalent to `dist x y ≤ K * dis ## Implementation notes The parameter `K` has type `ℝ≥0`. This way we avoid conjuction in the definition and have -coercions both to `ℝ` and `ennreal`. We do not require `0 < K` in the definition, mostly because +coercions both to `ℝ` and `ℝ≥0∞`. We do not require `0 < K` in the definition, mostly because we do not have a `posreal` type. -/ diff --git a/src/topology/metric_space/baire.lean b/src/topology/metric_space/baire.lean index 7a400cf2d405d..96c5b98d51dc9 100644 --- a/src/topology/metric_space/baire.lean +++ b/src/topology/metric_space/baire.lean @@ -26,7 +26,7 @@ has the countable intersection property. -/ noncomputable theory -open_locale classical topological_space filter +open_locale classical topological_space filter ennreal open filter encodable set @@ -42,7 +42,7 @@ encodable source space). -/ theorem dense_Inter_of_open_nat {f : ℕ → set α} (ho : ∀n, is_open (f n)) (hd : ∀n, dense (f n)) : dense (⋂n, f n) := begin - let B : ℕ → ennreal := λn, 1/2^n, + let B : ℕ → ℝ≥0∞ := λn, 1/2^n, have Bpos : ∀n, 0 < B n, { intro n, simp only [B, one_div, one_mul, ennreal.inv_pos], @@ -77,10 +77,10 @@ begin `closed_ball (c n) (r n)` is included in the previous ball and in `f n`, and such that `r n` is small enough to ensure that `c n` is a Cauchy sequence. Then `c n` converges to a limit which belongs to all the `f n`. -/ - let F : ℕ → (α × ennreal) := λn, nat.rec_on n (prod.mk x (min ε (B 0))) + let F : ℕ → (α × ℝ≥0∞) := λn, nat.rec_on n (prod.mk x (min ε (B 0))) (λn p, prod.mk (center n p.1 p.2) (radius n p.1 p.2)), let c : ℕ → α := λn, (F n).1, - let r : ℕ → ennreal := λn, (F n).2, + let r : ℕ → ℝ≥0∞ := λn, (F n).2, have rpos : ∀n, r n > 0, { assume n, induction n with n hn, diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 8d4bb37261991..857fb74ae726f 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -15,7 +15,7 @@ import topology.algebra.ordered open set filter classical topological_space noncomputable theory -open_locale uniformity topological_space big_operators filter nnreal +open_locale uniformity topological_space big_operators filter nnreal ennreal universes u v w variables {α : Type u} {β : Type v} {γ : Type w} @@ -63,7 +63,7 @@ class metric_space (α : Type u) extends has_dist α : Type u := (eq_of_dist_eq_zero : ∀ {x y : α}, dist x y = 0 → x = y) (dist_comm : ∀ x y : α, dist x y = dist y x) (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) -(edist : α → α → ennreal := λx y, ennreal.of_real (dist x y)) +(edist : α → α → ℝ≥0∞ := λx y, ennreal.of_real (dist x y)) (edist_dist : ∀ x y : α, edist x y = ennreal.of_real (dist x y) . control_laws_tac) (to_uniform_space : uniform_space α := uniform_space_of_dist dist dist_self dist_comm dist_triangle) (uniformity_dist : 𝓤 α = ⨅ ε>0, 𝓟 {p:α×α | dist p.1 p.2 < ε} . control_laws_tac) @@ -701,7 +701,7 @@ distance coincide. -/ /-- Expressing the uniformity in terms of `edist` -/ protected lemma metric.uniformity_basis_edist : - (𝓤 α).has_basis (λ ε:ennreal, 0 < ε) (λ ε, {p | edist p.1 p.2 < ε}) := + (𝓤 α).has_basis (λ ε:ℝ≥0∞, 0 < ε) (λ ε, {p | edist p.1 p.2 < ε}) := ⟨begin intro t, refine mem_uniformity_dist.trans ⟨_, _⟩; rintro ⟨ε, ε0, Hε⟩, @@ -1194,7 +1194,7 @@ begin show ∀ (x y : Π (b : β), π b), edist x y ≠ ⊤, { assume x y, rw ← lt_top_iff_ne_top, - have : (⊥ : ennreal) < ⊤ := ennreal.coe_lt_top, + have : (⊥ : ℝ≥0∞) < ⊤ := ennreal.coe_lt_top, simp [edist_pi_def, finset.sup_lt_iff this, edist_lt_top] }, show ∀ (x y : Π (b : β), π b), ↑(sup univ (λ (b : β), nndist (x b) (y b))) = ennreal.to_real (sup univ (λ (b : β), edist (x b) (y b))), diff --git a/src/topology/metric_space/closeds.lean b/src/topology/metric_space/closeds.lean index 8490a24394936..88f41f0682aac 100644 --- a/src/topology/metric_space/closeds.lean +++ b/src/topology/metric_space/closeds.lean @@ -22,8 +22,7 @@ always finite in this context. -/ noncomputable theory -open_locale classical -open_locale topological_space +open_locale classical topological_space ennreal universe u open classical set function topological_space filter @@ -89,8 +88,8 @@ begin `edist (s n) (s (n+1)) < 2^{-n}`, then it converges. This is enough to guarantee completeness, by a standard completeness criterion. We use the shorthand `B n = 2^{-n}` in ennreal. -/ - let B : ℕ → ennreal := λ n, (2⁻¹)^n, - have B_pos : ∀ n, (0:ennreal) < B n, + let B : ℕ → ℝ≥0∞ := λ n, (2⁻¹)^n, + have B_pos : ∀ n, (0:ℝ≥0∞) < B n, by simp [B, ennreal.pow_pos], have B_ne_top : ∀ n, B n ≠ ⊤, by simp [B, ennreal.pow_ne_top], diff --git a/src/topology/metric_space/contracting.lean b/src/topology/metric_space/contracting.lean index d7c1b89460515..29fc63fb6e250 100644 --- a/src/topology/metric_space/contracting.lean +++ b/src/topology/metric_space/contracting.lean @@ -27,7 +27,7 @@ of convergence, and some properties of the map sending a contracting map to its contracting map, fixed point, Banach fixed point theorem -/ -open_locale nnreal topological_space classical +open_locale nnreal topological_space classical ennreal open filter function variables {α : Type*} @@ -44,12 +44,12 @@ open emetric set lemma to_lipschitz_with (hf : contracting_with K f) : lipschitz_with K f := hf.2 -lemma one_sub_K_pos' (hf : contracting_with K f) : (0:ennreal) < 1 - K := by simp [hf.1] +lemma one_sub_K_pos' (hf : contracting_with K f) : (0:ℝ≥0∞) < 1 - K := by simp [hf.1] -lemma one_sub_K_ne_zero (hf : contracting_with K f) : (1:ennreal) - K ≠ 0 := +lemma one_sub_K_ne_zero (hf : contracting_with K f) : (1:ℝ≥0∞) - K ≠ 0 := ne_of_gt hf.one_sub_K_pos' -lemma one_sub_K_ne_top : (1:ennreal) - K ≠ ⊤ := +lemma one_sub_K_ne_top : (1:ℝ≥0∞) - K ≠ ⊤ := by { norm_cast, exact ennreal.coe_ne_top } lemma edist_inequality (hf : contracting_with K f) {x y} (h : edist x y < ⊤) : diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index ddbe13eca6a3a..04b4560f3c3b2 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -14,7 +14,7 @@ import topology.uniform_space.uniform_convergence This file is devoted to the definition and study of `emetric_spaces`, i.e., metric spaces in which the distance is allowed to take the value ∞. This extended distance is -called `edist`, and takes values in `ennreal`. +called `edist`, and takes values in `ℝ≥0∞`. Many definitions and theorems expected on emetric spaces are already introduced on uniform spaces and topological spaces. For example: open and closed sets, compactness, completeness, continuity and @@ -26,7 +26,7 @@ The class `emetric_space` therefore extends `uniform_space` (and `topological_sp open set filter classical noncomputable theory -open_locale uniformity topological_space big_operators filter nnreal +open_locale uniformity topological_space big_operators filter nnreal ennreal universes u v w variables {α : Type u} {β : Type v} {γ : Type w} @@ -41,12 +41,12 @@ le_antisymm (λ r ur, let ⟨ε, ε0, h⟩ := (H _).1 ur in mem_infi_sets ε $ mem_infi_sets ε0 $ mem_principal_sets.2 $ λ ⟨a, b⟩, h) -class has_edist (α : Type*) := (edist : α → α → ennreal) +class has_edist (α : Type*) := (edist : α → α → ℝ≥0∞) export has_edist (edist) /-- Creating a uniform space from an extended distance. -/ def uniform_space_of_edist - (edist : α → α → ennreal) + (edist : α → α → ℝ≥0∞) (edist_self : ∀ x : α, edist x x = 0) (edist_comm : ∀ x y : α, edist x y = edist y x) (edist_triangle : ∀ x y z : α, edist x z ≤ edist x y + edist y z) : uniform_space α := @@ -56,7 +56,7 @@ uniform_space.of_core { by simp [set.subset_def, id_rel, edist_self, (>)] {contextual := tt}, comp := le_infi $ assume ε, le_infi $ assume h, - have (2 : ennreal) = (2 : ℕ) := by simp, + have (2 : ℝ≥0∞) = (2 : ℕ) := by simp, have A : 0 < ε / 2 := ennreal.div_pos_iff.2 ⟨ne_of_gt h, by { convert ennreal.nat_ne_top 2 }⟩, lift'_le @@ -138,7 +138,7 @@ begin refine nat.le_induction _ _, { simp only [finset.sum_empty, finset.Ico.self_eq_empty, edist_self], -- TODO: Why doesn't Lean close this goal automatically? `apply le_refl` fails too. - exact le_refl (0:ennreal) }, + exact le_refl (0:ℝ≥0∞) }, { assume n hn hrec, calc edist (f m) (f (n+1)) ≤ edist (f m) (f n) + edist (f n) (f (n+1)) : edist_triangle _ _ _ ... ≤ ∑ i in finset.Ico m n, _ + _ : add_le_add hrec (le_refl _) @@ -154,7 +154,7 @@ finset.Ico.zero_bot n ▸ edist_le_Ico_sum_edist f (nat.zero_le n) /-- A version of `edist_le_Ico_sum_edist` with each intermediate distance replaced with an upper estimate. -/ lemma edist_le_Ico_sum_of_edist_le {f : ℕ → α} {m n} (hmn : m ≤ n) - {d : ℕ → ennreal} (hd : ∀ {k}, m ≤ k → k < n → edist (f k) (f (k + 1)) ≤ d k) : + {d : ℕ → ℝ≥0∞} (hd : ∀ {k}, m ≤ k → k < n → edist (f k) (f (k + 1)) ≤ d k) : edist (f m) (f n) ≤ ∑ i in finset.Ico m n, d i := le_trans (edist_le_Ico_sum_edist f hmn) $ finset.sum_le_sum $ λ k hk, hd (finset.Ico.mem.1 hk).1 (finset.Ico.mem.1 hk).2 @@ -162,7 +162,7 @@ finset.sum_le_sum $ λ k hk, hd (finset.Ico.mem.1 hk).1 (finset.Ico.mem.1 hk).2 /-- A version of `edist_le_range_sum_edist` with each intermediate distance replaced with an upper estimate. -/ lemma edist_le_range_sum_of_edist_le {f : ℕ → α} (n : ℕ) - {d : ℕ → ennreal} (hd : ∀ {k}, k < n → edist (f k) (f (k + 1)) ≤ d k) : + {d : ℕ → ℝ≥0∞} (hd : ∀ {k}, k < n → edist (f k) (f (k + 1)) ≤ d k) : edist (f 0) (f n) ≤ ∑ i in finset.range n, d i := finset.Ico.zero_bot n ▸ edist_le_Ico_sum_of_edist_le (zero_le n) (λ _ _, hd) @@ -176,7 +176,7 @@ theorem uniformity_edist : emetric_space.uniformity_edist theorem uniformity_basis_edist : - (𝓤 α).has_basis (λ ε : ennreal, 0 < ε) (λ ε, {p:α×α | edist p.1 p.2 < ε}) := + (𝓤 α).has_basis (λ ε : ℝ≥0∞, 0 < ε) (λ ε, {p:α×α | edist p.1 p.2 < ε}) := (@uniformity_edist α _).symm ▸ has_basis_binfi_principal (λ r hr p hp, ⟨min r p, lt_min hr hp, λ x hx, lt_of_lt_of_le hx (min_le_left _ _), @@ -188,12 +188,12 @@ theorem mem_uniformity_edist {s : set (α×α)} : s ∈ 𝓤 α ↔ (∃ε>0, ∀{a b:α}, edist a b < ε → (a, b) ∈ s) := uniformity_basis_edist.mem_uniformity_iff -/-- Given `f : β → ennreal`, if `f` sends `{i | p i}` to a set of positive numbers +/-- Given `f : β → ℝ≥0∞`, if `f` sends `{i | p i}` to a set of positive numbers accumulating to zero, then `f i`-neighborhoods of the diagonal form a basis of `𝓤 α`. For specific bases see `uniformity_basis_edist`, `uniformity_basis_edist'`, `uniformity_basis_edist_nnreal`, and `uniformity_basis_edist_inv_nat`. -/ -protected theorem emetric.mk_uniformity_basis {β : Type*} {p : β → Prop} {f : β → ennreal} +protected theorem emetric.mk_uniformity_basis {β : Type*} {p : β → Prop} {f : β → ℝ≥0∞} (hf₀ : ∀ x, p x → 0 < f x) (hf : ∀ ε, 0 < ε → ∃ x (hx : p x), f x ≤ ε) : (𝓤 α).has_basis p (λ x, {p:α×α | edist p.1 p.2 < f x}) := begin @@ -205,11 +205,11 @@ begin { exact λ ⟨i, hi, H⟩, ⟨f i, hf₀ i hi, H⟩ } end -/-- Given `f : β → ennreal`, if `f` sends `{i | p i}` to a set of positive numbers +/-- Given `f : β → ℝ≥0∞`, if `f` sends `{i | p i}` to a set of positive numbers accumulating to zero, then closed `f i`-neighborhoods of the diagonal form a basis of `𝓤 α`. For specific bases see `uniformity_basis_edist_le` and `uniformity_basis_edist_le'`. -/ -protected theorem emetric.mk_uniformity_basis_le {β : Type*} {p : β → Prop} {f : β → ennreal} +protected theorem emetric.mk_uniformity_basis_le {β : Type*} {p : β → Prop} {f : β → ℝ≥0∞} (hf₀ : ∀ x, p x → 0 < f x) (hf : ∀ ε, 0 < ε → ∃ x (hx : p x), f x ≤ ε) : (𝓤 α).has_basis p (λ x, {p:α×α | edist p.1 p.2 ≤ f x}) := begin @@ -223,17 +223,17 @@ begin end theorem uniformity_basis_edist_le : - (𝓤 α).has_basis (λ ε : ennreal, 0 < ε) (λ ε, {p:α×α | edist p.1 p.2 ≤ ε}) := + (𝓤 α).has_basis (λ ε : ℝ≥0∞, 0 < ε) (λ ε, {p:α×α | edist p.1 p.2 ≤ ε}) := emetric.mk_uniformity_basis_le (λ _, id) (λ ε ε₀, ⟨ε, ε₀, le_refl ε⟩) -theorem uniformity_basis_edist' (ε' : ennreal) (hε' : 0 < ε') : - (𝓤 α).has_basis (λ ε : ennreal, ε ∈ Ioo 0 ε') (λ ε, {p:α×α | edist p.1 p.2 < ε}) := +theorem uniformity_basis_edist' (ε' : ℝ≥0∞) (hε' : 0 < ε') : + (𝓤 α).has_basis (λ ε : ℝ≥0∞, ε ∈ Ioo 0 ε') (λ ε, {p:α×α | edist p.1 p.2 < ε}) := emetric.mk_uniformity_basis (λ _, and.left) (λ ε ε₀, let ⟨δ, hδ⟩ := exists_between hε' in ⟨min ε δ, ⟨lt_min ε₀ hδ.1, lt_of_le_of_lt (min_le_right _ _) hδ.2⟩, min_le_left _ _⟩) -theorem uniformity_basis_edist_le' (ε' : ennreal) (hε' : 0 < ε') : - (𝓤 α).has_basis (λ ε : ennreal, ε ∈ Ioo 0 ε') (λ ε, {p:α×α | edist p.1 p.2 ≤ ε}) := +theorem uniformity_basis_edist_le' (ε' : ℝ≥0∞) (hε' : 0 < ε') : + (𝓤 α).has_basis (λ ε : ℝ≥0∞, ε ∈ Ioo 0 ε') (λ ε, {p:α×α | edist p.1 p.2 ≤ ε}) := emetric.mk_uniformity_basis_le (λ _, and.left) (λ ε ε₀, let ⟨δ, hδ⟩ := exists_between hε' in ⟨min ε δ, ⟨lt_min ε₀ hδ.1, lt_of_le_of_lt (min_le_right _ _) hδ.2⟩, min_le_left _ _⟩) @@ -251,7 +251,7 @@ emetric.mk_uniformity_basis (λ ε ε₀, let ⟨n, hn⟩ := ennreal.exists_inv_nat_lt (ne_of_gt ε₀) in ⟨n, trivial, le_of_lt hn⟩) /-- Fixed size neighborhoods of the diagonal belong to the uniform structure -/ -theorem edist_mem_uniformity {ε:ennreal} (ε0 : 0 < ε) : +theorem edist_mem_uniformity {ε:ℝ≥0∞} (ε0 : 0 < ε) : {p:α×α | edist p.1 p.2 < ε} ∈ 𝓤 α := mem_uniformity_edist.2 ⟨ε, ε0, λ a b, id⟩ @@ -315,7 +315,7 @@ which satisfy a bound of the form `edist (u n) (u m) < B N` for all `n m ≥ N` converging. This is often applied for `B N = 2^{-N}`, i.e., with a very fast convergence to `0`, which makes it possible to use arguments of converging series, while this is impossible to do in general for arbitrary Cauchy sequences. -/ -theorem complete_of_convergent_controlled_sequences (B : ℕ → ennreal) (hB : ∀n, 0 < B n) +theorem complete_of_convergent_controlled_sequences (B : ℕ → ℝ≥0∞) (hB : ∀n, 0 < B n) (H : ∀u : ℕ → α, (∀N n m : ℕ, N ≤ n → N ≤ m → edist (u n) (u m) < B N) → ∃x, tendsto u at_top (𝓝 x)) : complete_space α := @@ -492,17 +492,17 @@ finset.sup_const univ_nonempty (edist a b) end pi namespace emetric -variables {x y z : α} {ε ε₁ ε₂ : ennreal} {s : set α} +variables {x y z : α} {ε ε₁ ε₂ : ℝ≥0∞} {s : set α} /-- `emetric.ball x ε` is the set of all points `y` with `edist y x < ε` -/ -def ball (x : α) (ε : ennreal) : set α := {y | edist y x < ε} +def ball (x : α) (ε : ℝ≥0∞) : set α := {y | edist y x < ε} @[simp] theorem mem_ball : y ∈ ball x ε ↔ edist y x < ε := iff.rfl theorem mem_ball' : y ∈ ball x ε ↔ edist x y < ε := by rw edist_comm; refl /-- `emetric.closed_ball x ε` is the set of all points `y` with `edist y x ≤ ε` -/ -def closed_ball (x : α) (ε : ennreal) := {y | edist y x ≤ ε} +def closed_ball (x : α) (ε : ℝ≥0∞) := {y | edist y x ≤ ε} @[simp] theorem mem_closed_ball : y ∈ closed_ball x ε ↔ edist y x ≤ ε := iff.rfl @@ -563,10 +563,10 @@ def edist_lt_top_setoid : setoid α := @[simp] lemma ball_zero : ball x 0 = ∅ := by rw [emetric.ball_eq_empty_iff] -theorem nhds_basis_eball : (𝓝 x).has_basis (λ ε:ennreal, 0 < ε) (ball x) := +theorem nhds_basis_eball : (𝓝 x).has_basis (λ ε:ℝ≥0∞, 0 < ε) (ball x) := nhds_basis_uniformity uniformity_basis_edist -theorem nhds_basis_closed_eball : (𝓝 x).has_basis (λ ε:ennreal, 0 < ε) (closed_ball x) := +theorem nhds_basis_closed_eball : (𝓝 x).has_basis (λ ε:ℝ≥0∞, 0 < ε) (closed_ball x) := nhds_basis_uniformity uniformity_basis_edist_le theorem nhds_eq : 𝓝 x = (⨅ε>0, 𝓟 (ball x ε)) := @@ -584,14 +584,14 @@ theorem is_closed_ball_top : is_closed (ball x ⊤) := is_open_iff.2 $ λ y hy, ⟨⊤, ennreal.coe_lt_top, subset_compl_iff_disjoint.2 $ ball_disjoint $ by { rw ennreal.top_add, exact le_of_not_lt hy }⟩ -theorem ball_mem_nhds (x : α) {ε : ennreal} (ε0 : 0 < ε) : ball x ε ∈ 𝓝 x := +theorem ball_mem_nhds (x : α) {ε : ℝ≥0∞} (ε0 : 0 < ε) : ball x ε ∈ 𝓝 x := mem_nhds_sets is_open_ball (mem_ball_self ε0) -theorem ball_prod_same [emetric_space β] (x : α) (y : β) (r : ennreal) : +theorem ball_prod_same [emetric_space β] (x : α) (y : β) (r : ℝ≥0∞) : (ball x r).prod (ball y r) = ball (x, y) r := ext $ λ z, max_lt_iff.symm -theorem closed_ball_prod_same [emetric_space β] (x : α) (y : β) (r : ennreal) : +theorem closed_ball_prod_same [emetric_space β] (x : α) (y : β) (r : ℝ≥0∞) : (closed_ball x r).prod (closed_ball y r) = closed_ball (x, y) r := ext $ λ z, max_le_iff.symm @@ -619,7 +619,7 @@ uniformity_basis_edist.cauchy_seq_iff /-- A variation around the emetric characterization of Cauchy sequences -/ theorem cauchy_seq_iff' [nonempty β] [semilattice_sup β] {u : β → α} : - cauchy_seq u ↔ ∀ε>(0 : ennreal), ∃N, ∀n≥N, edist (u n) (u N) < ε := + cauchy_seq u ↔ ∀ε>(0 : ℝ≥0∞), ∃N, ∀n≥N, edist (u n) (u N) < ε := uniformity_basis_edist.cauchy_seq_iff' /-- A variation of the emetric characterization of Cauchy sequences that deals with @@ -648,10 +648,10 @@ section compact lemma countable_closure_of_compact {α : Type u} [emetric_space α] {s : set α} (hs : is_compact s) : ∃ t ⊆ s, (countable t ∧ s = closure t) := begin - have A : ∀ (e:ennreal), e > 0 → ∃ t ⊆ s, (finite t ∧ s ⊆ (⋃x∈t, ball x e)) := + have A : ∀ (e:ℝ≥0∞), e > 0 → ∃ t ⊆ s, (finite t ∧ s ⊆ (⋃x∈t, ball x e)) := totally_bounded_iff'.1 (compact_iff_totally_bounded_complete.1 hs).1, -- assume e, finite_cover_balls_of_compact hs, - have B : ∀ (e:ennreal), ∃ t ⊆ s, finite t ∧ (e > 0 → s ⊆ (⋃x∈t, ball x e)), + have B : ∀ (e:ℝ≥0∞), ∃ t ⊆ s, finite t ∧ (e > 0 → s ⊆ (⋃x∈t, ball x e)), { intro e, cases le_or_gt e 0 with h, { exact ⟨∅, by finish⟩ }, @@ -667,7 +667,7 @@ begin apply mem_closure_iff.2, intros ε εpos, rcases ennreal.exists_inv_nat_lt (bot_lt_iff_ne_bot.1 εpos) with ⟨n, hn⟩, - have inv_n_pos : (0 : ennreal) < (n : ℕ)⁻¹ := by simp [ennreal.bot_lt_iff_ne_bot], + have inv_n_pos : (0 : ℝ≥0∞) < (n : ℕ)⁻¹ := by simp [ennreal.bot_lt_iff_ne_bot], have C : x ∈ (⋃y∈ T (n : ℕ)⁻¹, ball y (n : ℕ)⁻¹) := mem_of_mem_of_subset x_in_s ((finite_T (n : ℕ)⁻¹).2 inv_n_pos), rcases mem_Union.1 C with ⟨y, _, ⟨y_in_T, rfl⟩, Dxy⟩, @@ -710,7 +710,7 @@ section diam /-- The diameter of a set in an emetric space, named `emetric.diam` -/ def diam (s : set α) := ⨆ (x ∈ s) (y ∈ s), edist x y -lemma diam_le_iff_forall_edist_le {d : ennreal} : +lemma diam_le_iff_forall_edist_le {d : ℝ≥0∞} : diam s ≤ d ↔ ∀ (x ∈ s) (y ∈ s), edist x y ≤ d := by simp only [diam, supr_le_iff] @@ -720,7 +720,7 @@ diam_le_iff_forall_edist_le.1 (le_refl _) x hx y hy /-- If the distance between any two points in a set is bounded by some constant, this constant bounds the diameter. -/ -lemma diam_le_of_forall_edist_le {d : ennreal} (h : ∀ (x ∈ s) (y ∈ s), edist x y ≤ d) : +lemma diam_le_of_forall_edist_le {d : ℝ≥0∞} (h : ∀ (x ∈ s) (y ∈ s), edist x y ≤ d) : diam s ≤ d := diam_le_iff_forall_edist_le.2 h @@ -789,13 +789,13 @@ end lemma diam_union' {t : set α} (h : (s ∩ t).nonempty) : diam (s ∪ t) ≤ diam s + diam t := let ⟨x, ⟨xs, xt⟩⟩ := h in by simpa using diam_union xs xt -lemma diam_closed_ball {r : ennreal} : diam (closed_ball x r) ≤ 2 * r := +lemma diam_closed_ball {r : ℝ≥0∞} : diam (closed_ball x r) ≤ 2 * r := diam_le_of_forall_edist_le $ λa ha b hb, calc edist a b ≤ edist a x + edist b x : edist_triangle_right _ _ _ ... ≤ r + r : add_le_add ha hb ... = 2 * r : by simp [mul_two, mul_comm] -lemma diam_ball {r : ennreal} : diam (ball x r) ≤ 2 * r := +lemma diam_ball {r : ℝ≥0∞} : diam (ball x r) ≤ 2 * r := le_trans (diam_mono ball_subset_closed_ball) diam_closed_ball end diam diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index f64986d379062..973881da06f34 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -25,7 +25,7 @@ This files introduces: `Hausdorff_dist`. -/ noncomputable theory -open_locale classical nnreal +open_locale classical nnreal ennreal universes u v w open classical set function topological_space filter @@ -33,14 +33,14 @@ open classical set function topological_space filter namespace emetric section inf_edist -open_locale ennreal + variables {α : Type u} {β : Type v} [emetric_space α] [emetric_space β] {x y : α} {s t : set α} {Φ : α → β} -/-! ### Distance of a point to a set as a function into `ennreal`. -/ +/-! ### Distance of a point to a set as a function into `ℝ≥0∞`. -/ /-- The minimal edistance of a point to a set -/ -def inf_edist (x : α) (s : set α) : ennreal := Inf ((edist x) '' s) +def inf_edist (x : α) (s : set α) : ℝ≥0∞ := Inf ((edist x) '' s) @[simp] lemma inf_edist_empty : inf_edist x ∅ = ∞ := by unfold inf_edist; simp @@ -66,7 +66,7 @@ lemma inf_edist_le_inf_edist_of_subset (h : s ⊆ t) : inf_edist x t ≤ inf_edi Inf_le_Inf (image_subset _ h) /-- If the edist to a set is `< r`, there exists a point in the set at edistance `< r` -/ -lemma exists_edist_lt_of_inf_edist_lt {r : ennreal} (h : inf_edist x s < r) : +lemma exists_edist_lt_of_inf_edist_lt {r : ℝ≥0∞} (h : inf_edist x s < r) : ∃y∈s, edist x y < r := let ⟨t, ⟨ht, tr⟩⟩ := Inf_lt_iff.1 h in let ⟨y, ⟨ys, hy⟩⟩ := (mem_image _ _ _).1 ht in @@ -100,7 +100,7 @@ lemma inf_edist_closure : inf_edist x (closure s) = inf_edist x s := begin refine le_antisymm (inf_edist_le_inf_edist_of_subset subset_closure) _, refine ennreal.le_of_forall_pos_le_add (λε εpos h, _), - have εpos' : (0 : ennreal) < ε := by simpa, + have εpos' : (0 : ℝ≥0∞) < ε := by simpa, have : inf_edist x (closure s) < inf_edist x (closure s) + ε/2 := ennreal.lt_add_right h (ennreal.half_pos εpos'), rcases exists_edist_lt_of_inf_edist_lt this with ⟨y, ycs, hy⟩, @@ -145,11 +145,11 @@ end end inf_edist --section -/-! ### The Hausdorff distance as a function into `ennreal`. -/ +/-! ### The Hausdorff distance as a function into `ℝ≥0∞`. -/ /-- The Hausdorff edistance between two sets is the smallest `r` such that each set is contained in the `r`-neighborhood of the other one -/ -def Hausdorff_edist {α : Type u} [emetric_space α] (s t : set α) : ennreal := +def Hausdorff_edist {α : Type u} [emetric_space α] (s t : set α) : ℝ≥0∞ := Sup ((λx, inf_edist x t) '' s) ⊔ Sup ((λx, inf_edist x s) '' t) lemma Hausdorff_edist_def {α : Type u} [emetric_space α] (s t : set α) : @@ -158,7 +158,7 @@ lemma Hausdorff_edist_def {α : Type u} [emetric_space α] (s t : set α) : attribute [irreducible] Hausdorff_edist section Hausdorff_edist -open_locale ennreal + variables {α : Type u} {β : Type v} [emetric_space α] [emetric_space β] {x y : α} {s t u : set α} {Φ : α → β} @@ -176,7 +176,7 @@ by unfold Hausdorff_edist; apply sup_comm /-- Bounding the Hausdorff edistance by bounding the edistance of any point in each set to the other set -/ -lemma Hausdorff_edist_le_of_inf_edist {r : ennreal} +lemma Hausdorff_edist_le_of_inf_edist {r : ℝ≥0∞} (H1 : ∀x ∈ s, inf_edist x t ≤ r) (H2 : ∀x ∈ t, inf_edist x s ≤ r) : Hausdorff_edist s t ≤ r := begin @@ -186,7 +186,7 @@ end /-- Bounding the Hausdorff edistance by exhibiting, for any point in each set, another point in the other set at controlled distance -/ -lemma Hausdorff_edist_le_of_mem_edist {r : ennreal} +lemma Hausdorff_edist_le_of_mem_edist {r : ℝ≥0∞} (H1 : ∀x ∈ s, ∃y ∈ t, edist x y ≤ r) (H2 : ∀x ∈ t, ∃y ∈ s, edist x y ≤ r) : Hausdorff_edist s t ≤ r := begin @@ -209,7 +209,7 @@ end /-- If the Hausdorff distance is `