From 0f0f5ea15c51247db5fe45f033a2c24d2ad694d5 Mon Sep 17 00:00:00 2001 From: madvorak Date: Fri, 13 Sep 2024 10:19:11 +0200 Subject: [PATCH 1/4] by exact --- Carleson/DoublingMeasure.lean | 2 +- Carleson/GridStructure.lean | 2 +- Carleson/RealInterpolation.lean | 24 ++++++++++++------------ Carleson/TileExistence.lean | 10 +++++----- Carleson/ToMathlib/Misc.lean | 6 +++--- 5 files changed, 22 insertions(+), 22 deletions(-) diff --git a/Carleson/DoublingMeasure.lean b/Carleson/DoublingMeasure.lean index 67713a51..eb509b2c 100644 --- a/Carleson/DoublingMeasure.lean +++ b/Carleson/DoublingMeasure.lean @@ -111,7 +111,7 @@ lemma measure_ball_le_pow_two {x : X} {r : ℝ} {n : ℕ} : _ ≤ A * μ.real (ball x (2 ^ m * r)) := by rw[mul_assoc]; norm_cast; exact measure_real_ball_two_le_same _ _ _ ≤ A * (↑(A ^ m) * μ.real (ball x r)) := by - exact mul_le_mul_of_nonneg_left hm (by exact zero_le_coe) + exact mul_le_mul_of_nonneg_left hm zero_le_coe _ = A^(m.succ) * μ.real (ball x r) := by rw [NNReal.coe_pow,← mul_assoc, pow_succ'] lemma measure_ball_le_pow_two' {x : X} {r : ℝ} {n : ℕ} : diff --git a/Carleson/GridStructure.lean b/Carleson/GridStructure.lean index 557525c4..4c2c2400 100644 --- a/Carleson/GridStructure.lean +++ b/Carleson/GridStructure.lean @@ -114,7 +114,7 @@ postfix:max "ᵒ" => Grid.int /-- An auxiliary measure used in the well-foundedness of `Ω` in Lemma `tile_structure`. -/ def opSize (i : Grid X) : ℕ := (S - s i).toNat -lemma int_subset : i.int ⊆ i := by exact ball_subset_Grid +lemma int_subset : i.int ⊆ i := ball_subset_Grid end Grid end Generic diff --git a/Carleson/RealInterpolation.lean b/Carleson/RealInterpolation.lean index b9339dfc..de0a116e 100644 --- a/Carleson/RealInterpolation.lean +++ b/Carleson/RealInterpolation.lean @@ -1693,7 +1693,7 @@ lemma rpow_le_rpow_of_exponent_le_base_le {a b t γ : ℝ} (ht : t > 0) (htγ : lemma rpow_le_rpow_of_exponent_le_base_ge {a b t γ : ℝ} (hγ : γ > 0) (htγ : γ ≤ t) (hab : a ≤ b) : ENNReal.ofReal (t ^ a) ≤ ENNReal.ofReal (t ^ b) * ENNReal.ofReal (γ ^ (a - b)) := by rw [mul_comm] - have t_pos : 0 < t := by exact gt_of_ge_of_gt htγ hγ + have t_pos : 0 < t := gt_of_ge_of_gt htγ hγ rw [Real.rpow_sub hγ] refine (ENNReal.mul_le_mul_left (a := ENNReal.ofReal (γ ^ (-a) )) ?_ coe_ne_top).mp ?_ · exact (ofReal_pos.mpr (Real.rpow_pos_of_pos hγ (-a))).ne.symm @@ -1799,7 +1799,7 @@ lemma estimate_eLpNorm_trunc {p q : ℝ≥0∞} have p_ne_top : p ≠ ⊤ := LT.lt.ne_top hpq.2 have : p ≠ 0 := hpq.1.ne.symm have : q ≠ 0 := (lt_trans hpq.1 hpq.2).ne.symm - have q_toReal_pos : q.toReal > 0 := by exact toReal_pos this hq + have q_toReal_pos : q.toReal > 0 := toReal_pos this hq split_ifs · contradiction · calc @@ -1939,7 +1939,7 @@ lemma res'comp (j : Bool) (β : ℝ) (hβ : β > 0) : exact ⟨lt_trans hβ h, fun _ ↦ h⟩ · ext x simp only [Ioi_diff_Ici, mem_Ioo] - · have : j = false := by exact eq_false_of_ne_true h₀ + · have : j = false := eq_false_of_ne_true h₀ rw [this] at h₂ simp at h₂ @@ -2649,7 +2649,7 @@ lemma estimate_trnc {p₀ q₀ q : ℝ} {spf : ScaledPowerFunction} {j : Bool} unfold eLpNorm eLpNorm' let tc := spf_to_tc spf split_ifs with is_p₀pos is_p₀top - · have : p₀ ≤ 0 := by exact ofReal_eq_zero.mp is_p₀pos + · have : p₀ ≤ 0 := ofReal_eq_zero.mp is_p₀pos contrapose! this; exact hp₀ · contrapose! is_p₀top; exact coe_ne_top · have coe_p₀ : (ENNReal.ofReal p₀).toReal = p₀ := by @@ -3515,7 +3515,7 @@ lemma estimate_norm_rpow_range_operator' apply setLIntegral_mono' measurableSet_Ioi intro s (s_pos : s > 0) simp only [is_q₀top, mem_Ioi, false_or] at hq₀' - have : q₀ = ⊤ := by exact not_lt_top.mp is_q₀top + have : q₀ = ⊤ := not_lt_top.mp is_q₀top rw [hq₀' this s s_pos, zero_mul, add_zero] gcongr apply weaktype_estimate_trunc p_pos <;> try assumption @@ -3523,14 +3523,14 @@ lemma estimate_norm_rpow_range_operator' apply setLIntegral_mono' measurableSet_Ioi intro s (s_pos : s > 0) simp only [is_q₁top, mem_Ioi, false_or] at hq₁' - have : q₁ = ⊤ := by exact not_lt_top.mp is_q₁top + have : q₁ = ⊤ := not_lt_top.mp is_q₁top rw [hq₁' this s s_pos, zero_mul, zero_add] gcongr apply weaktype_estimate_trunc_compl (p₀ := p₀) <;> try assumption · exact LT.lt.ne_top hp₁p · exact tc.ran_ton s s_pos · simp only [zero_mul, add_zero, nonpos_iff_eq_zero] - have : ∫⁻ (_ : ℝ) in Ioi 0, 0 = 0 := by exact lintegral_zero + have : ∫⁻ (_ : ℝ) in Ioi 0, 0 = 0 := lintegral_zero rw [← this] apply lintegral_congr_ae filter_upwards [self_mem_ae_restrict measurableSet_Ioi] with s (s_pos) @@ -3562,7 +3562,7 @@ lemma simplify_factor₀ {D : ℝ} have q₁pos : q₁ > 0 := lt_of_lt_of_le hp₁.1 hp₁.2 have p₀ne_top : p₀ ≠ ⊤ := ne_top_of_le_ne_top hq₀' hp₀.2 have hp' : p ∈ Ioo 0 ⊤ := interp_exp_in_Ioo_zero_top ht p₀pos p₁pos (by left; exact p₀ne_top) hp - have p_toReal_pos : p.toReal > 0 := by exact toReal_pos_of_Ioo hp' + have p_toReal_pos : p.toReal > 0 := toReal_pos_of_Ioo hp' rw [hD] unfold d rw [← ofReal_rpow_of_pos] @@ -3811,7 +3811,7 @@ lemma combine_estimates₀ {A : ℝ} (hA : A > 0) have q₀pos : q₀ > 0 := lt_of_lt_of_le hp₀.1 hp₀.2 have p₁pos : p₁ > 0 := hp₁.1 have q₁pos : q₁ > 0 := lt_of_lt_of_le hp₁.1 hp₁.2 - have p_pos : p > 0 := by exact interpolated_pos' one_le_p₀ one_le_p1 hp + have p_pos : p > 0 := interpolated_pos' one_le_p₀ one_le_p1 hp have : SigmaFinite (μ.restrict (Function.support f)) := by apply support_sigma_finite_from_Memℒp (p := p) hf · exact interp_exp_ne_top (ne_of_lt hp₀p₁) ht hp @@ -3997,7 +3997,7 @@ lemma simplify_factor₃ [NormedAddCommGroup E₁] (hp₀ : p₀ > 0) (hp₀' : (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) (hp₀p₁ : p₀ = p₁) : C₀ ^ q₀.toReal * (eLpNorm f p μ ^ p.toReal) ^ (q₀.toReal / p₀.toReal) = (↑C₀ * eLpNorm f p μ) ^ q₀.toReal := by - have hp' : p = p₀ := by exact Eq.symm (interp_exp_eq hp₀p₁ ht hp) -- TODO abstract + have hp' : p = p₀ := (interp_exp_eq hp₀p₁ ht hp).symm -- TODO abstract rw [hp'] rw [ENNReal.mul_rpow_of_nonneg, ← ENNReal.rpow_mul, ← mul_div_assoc, mul_div_cancel_left₀] <;> try positivity @@ -4597,7 +4597,7 @@ lemma Subadditive_trunc_from_SubadditiveOn_Lp₀p₁ {p₀ p₁ p : ℝ≥0∞} · exact interpolated_pos' hp₀ hp₁ hp · exact (interp_exp_between hp₀ hp₁ p₀lt_p₁ ht hp).2 · left - have : p₀ = p := by exact interp_exp_eq p₀eq_p₁ ht hp + have : p₀ = p := interp_exp_eq p₀eq_p₁ ht hp rw [this] exact trunc_preserves_Lp hf · left @@ -4613,7 +4613,7 @@ lemma Subadditive_trunc_from_SubadditiveOn_Lp₀p₁ {p₀ p₁ p : ℝ≥0∞} · exact hp₀ · exact (interp_exp_between hp₀ hp₁ p₀lt_p₁ ht hp).1 · left - have : p₀ = p := by exact interp_exp_eq p₀eq_p₁ ht hp + have : p₀ = p := interp_exp_eq p₀eq_p₁ ht hp rw [this] exact trunc_compl_preserves_Lp hf · right diff --git a/Carleson/TileExistence.lean b/Carleson/TileExistence.lean index db536b7e..e1fae1a9 100644 --- a/Carleson/TileExistence.lean +++ b/Carleson/TileExistence.lean @@ -399,7 +399,7 @@ mutual rw [I2,dif_pos hk_s] exact measurableSet_ball else - let hk'' : -S < k := by exact lt_of_le_of_ne hk fun a_1 ↦ hk_s (id (Eq.symm a_1)) + let hk'' : -S < k := lt_of_le_of_ne hk fun a_1 ↦ hk_s (id (Eq.symm a_1)) have : (S + (k-1)).toNat < (S + k).toNat := by rw [Int.lt_toNat,Int.toNat_of_nonneg (by linarith)] linarith @@ -534,7 +534,7 @@ mutual else simp_rw [dif_neg hk_s] intro x hx - have : -S < k := by exact lt_of_le_of_ne hk fun a_1 ↦ hk_s (id (Eq.symm a_1)) + have : -S < k := lt_of_le_of_ne hk fun a_1 ↦ hk_s (id (Eq.symm a_1)) have : ((2 * (S + (k - 1))).toNat : ℤ) + 1 < 2 * (S + k) := by rw [Int.toNat_of_nonneg (by linarith)] linarith @@ -1583,8 +1583,8 @@ lemma boundary_measure {k:ℤ} (hk:-S ≤ k) (y:Yk X k) {t:ℝ≥0} (ht:t∈ Set simp only at this nth_rw 1 [NNReal.val_eq_coe] at this simp_rw [← Real.rpow_intCast] at this - · rw [← ENNReal.ofReal_le_ofReal_iff (Real.rpow_nonneg (realD_nonneg) _), ENNReal.ofReal_mul (by exact ht.left.le), - ENNReal.ofReal_coe_nnreal, + · rw [← ENNReal.ofReal_le_ofReal_iff (Real.rpow_nonneg (realD_nonneg) _), + ENNReal.ofReal_mul (by exact ht.left.le), ENNReal.ofReal_coe_nnreal, ← ENNReal.ofReal_rpow_of_pos (defaultD_pos a),← ENNReal.ofReal_rpow_of_pos (defaultD_pos a), ENNReal.ofReal_natCast, ENNReal.rpow_intCast, ENNReal.rpow_intCast] at this exact this @@ -2137,7 +2137,7 @@ lemma Ω_RFD {p q : 𝔓 X} (h𝓘 : 𝓘 p ≤ 𝓘 q) : Disjoint (Ω p) (Ω q) rw [mem_iUnion] at zi; obtain ⟨a, ma⟩ := zi -- Paper's `q'` is `⟨J, a⟩` have nmaxJ : ¬IsMax J := by by_contra maxJ; rw [Grid.isMax_iff] at maxJ - rw [maxJ, show s topCube = S by exact s_topCube (X := X)] at sJ + rw [maxJ, show s topCube = S from s_topCube (X := X)] at sJ have : 𝔰 q ≤ S := (range_s_subset ⟨q.1, rfl⟩).2 omega have succJ : J.succ = q.1 := (Grid.succ_def nmaxJ).mpr ⟨ubJ, by change 𝔰 q = _; omega⟩ diff --git a/Carleson/ToMathlib/Misc.lean b/Carleson/ToMathlib/Misc.lean index 3d14826c..9492f2b2 100644 --- a/Carleson/ToMathlib/Misc.lean +++ b/Carleson/ToMathlib/Misc.lean @@ -66,13 +66,13 @@ section ENNReal lemma tsum_one_eq' {α : Type*} (s : Set α) : ∑' (_:s), (1 : ℝ≥0∞) = s.encard := by if hfin : s.Finite then - have hfin' : Finite s := by exact hfin + have hfin' : Finite s := hfin rw [tsum_def] simp only [ENNReal.summable, ↓reduceDIte] have hsup: support (fun (_ : s) ↦ (1 : ℝ≥0∞)) = Set.univ := by ext i simp only [mem_support, ne_eq, one_ne_zero, not_false_eq_true, mem_univ] - have hsupfin: (Set.univ : Set s).Finite := by exact finite_univ + have hsupfin: (Set.univ : Set s).Finite := finite_univ rw [← hsup] at hsupfin rw [if_pos hsupfin] rw [hfin.encard_eq_coe_toFinset_card] @@ -94,7 +94,7 @@ lemma tsum_one_eq' {α : Type*} (s : Set α) : ∑' (_:s), (1 : ℝ≥0∞) = s. simp only [Finite.mem_toFinset, mem_support, ne_eq, one_ne_zero, not_false_eq_true, exists_const] else - have : Infinite s := by exact infinite_coe_iff.mpr hfin + have : Infinite s := infinite_coe_iff.mpr hfin rw [ENNReal.tsum_const_eq_top_of_ne_zero (by norm_num)] rw [Set.encard_eq_top_iff.mpr hfin] simp only [ENat.toENNReal_top] From 2bdc6c03a2f8d19609c6d68a4c0fd4e9b6bb5d12 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Fri, 13 Sep 2024 14:41:03 +0400 Subject: [PATCH 2/4] State lemmas in chapter 7 (#125) Also: * Use the simpler statement of classical Carleson as the main theorem * Move `stackSize` --- CONTRIBUTING.md | 16 +- Carleson/Classical/ClassicalCarleson.lean | 8 +- Carleson/Discrete/Defs.lean | 1 - Carleson/Discrete/Forests.lean | 29 +- Carleson/FinitaryCarleson.lean | 5 +- Carleson/Forest.lean | 142 ++--- Carleson/ForestOperator.lean | 610 +++++++++++++++++++--- Carleson/HardyLittlewood.lean | 2 +- Carleson/TileStructure.lean | 44 ++ blueprint/src/chapter/main.tex | 142 +++-- 10 files changed, 770 insertions(+), 229 deletions(-) diff --git a/CONTRIBUTING.md b/CONTRIBUTING.md index 5ead4932..e4911181 100644 --- a/CONTRIBUTING.md +++ b/CONTRIBUTING.md @@ -37,6 +37,8 @@ Below, I will try to give a translation of some notation/conventions. We use mat | `T_Q f(x)` | `linearizedCarlesonOperator Q K f x` | The linearized generalized Carleson operator | | `T_𝓝^θ f(x)` | `nontangentialMaximalFunction θ f x` | | | `Tₚ f(x)` | `carlesonOn p f x` | | +| `T_ℭ f(x)` | `carlesonSum ℭ f x` | The sum of Tₚ f(x) for p ∈ ℭ. In the blueprint only used in chapter 7, but in the formalization we will use it more. | +| `Tₚ* f(x)` | `adjointCarleson p f x` | | | `e(x)` | `Complex.exp (Complex.I * x)` | | | `𝔓(I)` | `𝓘 ⁻¹' {I}` | | | `I ⊆ J` | `I ≤ J` | We noticed recently that we cannot (easily) assume that the coercion `Grid X → Set X` is injective. Therefore, Lean introduces two orders on `Grid X`: `I ⊆ J` means that the underlying sets satisfy this relation, and `I ≤ J` means *additionally* that `s I ≤ s J`. The order is what you should use in (almost?) all cases. | @@ -48,8 +50,12 @@ Below, I will try to give a translation of some notation/conventions. We use mat | `M` | `globalMaximalFunction volume 1` | | | `I_i(x)` | `cubeOf i x` | | | `R_Q(θ, x)` | `upperRadius Q θ x` | | -| `S_{1,𝔲} f(x)` | `boundaryOperator1 t u f x` | | -| `S_{2,𝔲} g(x)` | `boundaryOperator2 t u g x` | | -| `` | `` | | -| `` | `` | | -| `` | `` | | +| `S_{1,𝔲} f(x)` | `boundaryOperator t u f x` | | +| `S_{2,𝔲} g(x)` | `adjointBoundaryOperator t u g x` | | +| `𝔘` | `t` | `t` is the (implicit) forest, sometimes `(t : Set (𝔓 X))` is required. It is equivalent to `t.𝔘` | +| `u ∈ 𝔘` | `u ∈ t` | `t` is the (implicit) forest, and this notation is equivalent to `u ∈ t.𝔘` | +| `𝔗(u)` | `t u` | `t` is the (implicit) forest, and this notation is equivalent to `t.𝔗 u` | +| `𝔘ⱼ` | `rowDecomp t j` | sometimes `(rowDecomp t j : Set (𝔓 X))` | +| `𝔗ⱼ(u)` | `rowDecomp t j u` | | +| `E` | `E` | | +| `Eⱼ` | `rowSupport t j` | | diff --git a/Carleson/Classical/ClassicalCarleson.lean b/Carleson/Classical/ClassicalCarleson.lean index 13dec313..678e174f 100644 --- a/Carleson/Classical/ClassicalCarleson.lean +++ b/Carleson/Classical/ClassicalCarleson.lean @@ -12,7 +12,7 @@ noncomputable section local notation "S_" => partialFourierSum /- Theorem 1.1 (Classical Carleson) -/ -theorem classical_carleson {f : ℝ → ℂ} +theorem exceptional_set_carleson {f : ℝ → ℂ} (cont_f : Continuous f) (periodic_f : f.Periodic (2 * π)) {ε : ℝ} (εpos : 0 < ε) : ∃ E ⊆ Set.Icc 0 (2 * π), MeasurableSet E ∧ volume.real E ≤ ε ∧ @@ -105,8 +105,8 @@ theorem carleson_interval {f : ℝ → ℂ} (cont_f : Continuous f) (periodic_f rw [one_mul] norm_num - -- Main step: Apply classical_carleson to get a family of exceptional sets parameterized by ε. - choose Eε hEε_subset _ hEε_measure hEε using (@classical_carleson f cont_f periodic_f) + -- Main step: Apply exceptional_set_carleson to get a family of exceptional sets parameterized by ε. + choose Eε hEε_subset _ hEε_measure hEε using (@exceptional_set_carleson f cont_f periodic_f) have Eεmeasure {ε : ℝ} (hε : 0 < ε) : volume (Eε hε) ≤ ENNReal.ofReal ε := by rw [ENNReal.le_ofReal_iff_toReal_le _ hε.le] @@ -204,7 +204,7 @@ lemma Function.Periodic.ae_of_ae_restrict {T : ℝ} (hT : 0 < T) {a : ℝ} {P : end section /- Carleson's theorem asserting a.e. convergence of the partial Fourier sums for continous functions. -/ -theorem carleson {f : ℝ → ℂ} (cont_f : Continuous f) (periodic_f : f.Periodic (2 * π)) : +theorem classical_carleson {f : ℝ → ℂ} (cont_f : Continuous f) (periodic_f : f.Periodic (2 * π)) : ∀ᵐ x, Filter.Tendsto (S_ · f x) Filter.atTop (nhds (f x)) := by -- Reduce to a.e. convergence on [0,2π] apply @Function.Periodic.ae_of_ae_restrict _ Real.two_pi_pos 0 diff --git a/Carleson/Discrete/Defs.lean b/Carleson/Discrete/Defs.lean index 99691fb2..53d35179 100644 --- a/Carleson/Discrete/Defs.lean +++ b/Carleson/Discrete/Defs.lean @@ -1,4 +1,3 @@ -import Carleson.Forest import Carleson.MinLayerTiles open MeasureTheory Measure NNReal Metric Set diff --git a/Carleson/Discrete/Forests.lean b/Carleson/Discrete/Forests.lean index 4d2048d7..0e257d72 100644 --- a/Carleson/Discrete/Forests.lean +++ b/Carleson/Discrete/Forests.lean @@ -1,4 +1,5 @@ import Carleson.Discrete.ExceptionalSet +import Carleson.Forest open MeasureTheory Measure NNReal Metric Complex Set open scoped ENNReal @@ -128,13 +129,13 @@ lemma dens1_le_dens' {P : Set (𝔓 X)} (hP : P ⊆ TilesAt k) : dens₁ P ≤ d apply absurd _ this.2; use J, sl.1.trans lJ /-- Lemma 5.3.12 -/ -lemma dens1_le {A : Set (𝔓 X)} (hA : A ⊆ ℭ k n) : dens₁ A ≤ 2 ^ (4 * a - n + 1 : ℤ) := +lemma dens1_le {A : Set (𝔓 X)} (hA : A ⊆ ℭ k n) : dens₁ A ≤ 2 ^ (4 * (a : ℝ) - n + 1) := calc _ ≤ dens' k A := dens1_le_dens' (hA.trans ℭ_subset_TilesAt) _ ≤ dens' k (ℭ (X := X) k n) := iSup_le_iSup_of_subset hA _ ≤ _ := by rw [dens'_iSup, iSup₂_le_iff]; intro p mp - rw [ℭ, mem_setOf] at mp; exact mp.2.2 + rw [ℭ, mem_setOf] at mp; exact_mod_cast mp.2.2 /-! ## Section 5.4 and Lemma 5.1.2 -/ @@ -584,14 +585,14 @@ variable (k n j l) in def forest : Forest X n where 𝔘 := 𝔘₄ k n j l 𝔗 := 𝔗₂ k n j - nonempty {u} hu := sorry - ordConnected {u} hu := forest_convex - 𝓘_ne_𝓘 hu hp := sorry - smul_four_le {u} hu := forest_geometry <| 𝔘₄_subset_𝔘₃ hu - stackSize_le {x} := stackSize_𝔘₄_le x - dens₁_𝔗_le {u} hu := dens1_le <| 𝔗₂_subset_ℭ₆.trans ℭ₆_subset_ℭ - lt_dist hu hu' huu' p hp := forest_separation (𝔘₄_subset_𝔘₃ hu) (𝔘₄_subset_𝔘₃ hu') huu' hp - ball_subset hu p hp := forest_inner (𝔘₄_subset_𝔘₃ hu) hp + nonempty' {u} hu := sorry + ordConnected' {u} hu := forest_convex + 𝓘_ne_𝓘' hu hp := sorry + smul_four_le' {u} hu := forest_geometry <| 𝔘₄_subset_𝔘₃ hu + stackSize_le' {x} := stackSize_𝔘₄_le x + dens₁_𝔗_le' {u} hu := dens1_le <| 𝔗₂_subset_ℭ₆.trans ℭ₆_subset_ℭ + lt_dist' hu hu' huu' p hp := forest_separation (𝔘₄_subset_𝔘₃ hu) (𝔘₄_subset_𝔘₃ hu') huu' hp + ball_subset' hu p hp := forest_inner (𝔘₄_subset_𝔘₃ hu) hp /-- The constant used in Lemma 5.1.2, with value `2 ^ (235 * a ^ 3) / (q - 1) ^ 4` -/ -- todo: redefine in terms of other constants @@ -600,7 +601,7 @@ def C5_1_2 (a : ℝ) (q : ℝ≥0) : ℝ≥0 := 2 ^ (235 * a ^ 3) / (q - 1) ^ 4 lemma C5_1_2_pos : C5_1_2 a nnq > 0 := sorry lemma forest_union {f : X → ℂ} (hf : ∀ x, ‖f x‖ ≤ F.indicator 1 x) : - ∫⁻ x in G \ G', ‖∑ p ∈ { p | p ∈ 𝔓₁ }, carlesonOn p f x‖₊ ≤ + ∫⁻ x in G \ G', ‖carlesonSum 𝔓₁ f x‖₊ ≤ C5_1_2 a nnq * volume G ^ (1 - q⁻¹) * volume F ^ q⁻¹ := by sorry @@ -944,8 +945,8 @@ def C5_1_3 (a : ℝ) (q : ℝ≥0) : ℝ≥0 := 2 ^ (210 * a ^ 3) / (q - 1) ^ 5 lemma C5_1_3_pos : C5_1_3 a nnq > 0 := sorry lemma forest_complement {f : X → ℂ} (hf : ∀ x, ‖f x‖ ≤ F.indicator 1 x) : - ∫⁻ x in G \ G', ‖∑ p ∈ { p | p ∉ 𝔓₁ }, carlesonOn p f x‖₊ ≤ - C5_1_2 a nnq * volume G ^ (1 - q⁻¹) * volume F ^ q⁻¹ := by + ∫⁻ x in G \ G', ‖carlesonSum 𝔓₁ᶜ f x‖₊ ≤ + C5_1_2 a nnq * volume G ^ (1 - q⁻¹) * volume F ^ q⁻¹ := by sorry /-! ## Proposition 2.0.2 -/ @@ -960,5 +961,5 @@ variable (X) in theorem discrete_carleson : ∃ G', MeasurableSet G' ∧ 2 * volume G' ≤ volume G ∧ ∀ f : X → ℂ, Measurable f → (∀ x, ‖f x‖ ≤ F.indicator 1 x) → - ∫⁻ x in G \ G', ‖∑ p, carlesonOn p f x‖₊ ≤ + ∫⁻ x in G \ G', ‖carlesonSum univ f x‖₊ ≤ C2_0_2 a nnq * volume G ^ (1 - q⁻¹) * volume F ^ q⁻¹ := by sorry diff --git a/Carleson/FinitaryCarleson.lean b/Carleson/FinitaryCarleson.lean index 92745395..8e3a29f3 100644 --- a/Carleson/FinitaryCarleson.lean +++ b/Carleson/FinitaryCarleson.lean @@ -40,7 +40,7 @@ private lemma 𝔓_biUnion : @Finset.univ (𝔓 X) _ = (Icc (-S : ℤ) S).toFins private lemma sum_eq_zero_of_nmem_Icc {f : X → ℂ} {x : X} (s : ℤ) (hs : s ∈ (Icc (-S : ℤ) S).toFinset.filter (fun t ↦ t ∉ Icc (σ₁ x) (σ₂ x))) : - ∑ i ∈ Finset.filter (fun p ↦ 𝔰 p = s) Finset.univ, carlesonOn i f x = 0 := by + ∑ i ∈ Finset.univ.filter (fun p ↦ 𝔰 p = s), carlesonOn i f x = 0 := by refine Finset.sum_eq_zero (fun p hp ↦ ?_) simp only [Finset.mem_filter, Finset.mem_univ, true_and] at hp simp only [mem_Icc, not_and, not_le, toFinset_Icc, Finset.mem_filter, Finset.mem_Icc] at hs @@ -111,7 +111,8 @@ theorem finitary_carleson : ∃ G', MeasurableSet G' ∧ 2 * volume G' ≤ volum rcases discrete_carleson X with ⟨G', hG', h2G', hfG'⟩ refine ⟨G', hG', h2G', fun f meas_f h2f ↦ le_of_eq_of_le ?_ (hfG' f meas_f h2f)⟩ refine setLIntegral_congr_fun (measurableSet_G.diff hG') (ae_of_all volume fun x hx ↦ ?_) - simp_rw [tile_sum_operator hx, mul_sub, exp_sub, mul_div, div_eq_mul_inv, + simp_rw [carlesonSum, mem_univ, Finset.filter_True, tile_sum_operator hx, mul_sub, exp_sub, + mul_div, div_eq_mul_inv, ← smul_eq_mul (a' := _⁻¹), integral_smul_const, ← Finset.sum_smul, nnnorm_smul] suffices ‖(cexp (I * ((Q x) x : ℂ)))⁻¹‖₊ = 1 by rw [this, mul_one] simp [← coe_eq_one, mul_comm I] diff --git a/Carleson/Forest.lean b/Carleson/Forest.lean index 53e87b45..0f44176c 100644 --- a/Carleson/Forest.lean +++ b/Carleson/Forest.lean @@ -7,52 +7,9 @@ noncomputable section open scoped ShortVariables variable {X : Type*} [PseudoMetricSpace X] {a : ℕ} {q : ℝ} {K : X → X → ℂ} {σ₁ σ₂ : X → ℤ} {F G : Set X} [ProofData a q K σ₁ σ₂ F G] -variable [TileStructure Q D κ S o] {p p' : 𝔓 X} {f g : Θ X} +variable [TileStructure Q D κ S o] {u u' p p' : 𝔓 X} {f g : Θ X} {C C' : Set (𝔓 X)} {x x' : X} -/-- The number of tiles `p` in `s` whose underlying cube `𝓘 p` contains `x`. -/ -def stackSize (C : Set (𝔓 X)) (x : X) : ℕ := - ∑ p ∈ { p | p ∈ C }, (𝓘 p : Set X).indicator 1 x - -lemma stackSize_setOf_add_stackSize_setOf_not {P : 𝔓 X → Prop} : - stackSize {p ∈ C | P p} x + stackSize {p ∈ C | ¬ P p} x = stackSize C x := by - classical - simp_rw [stackSize] - conv_rhs => rw [← Finset.sum_filter_add_sum_filter_not _ P] - simp_rw [Finset.filter_filter] - congr - -lemma stackSize_congr (h : ∀ p ∈ C, x ∈ (𝓘 p : Set X) ↔ x' ∈ (𝓘 p : Set X)) : - stackSize C x = stackSize C x' := by - refine Finset.sum_congr rfl fun p hp ↦ ?_ - simp_rw [Finset.mem_filter, Finset.mem_univ, true_and] at hp - simp_rw [indicator, h p hp, Pi.one_apply] - -lemma stackSize_mono (h : C ⊆ C') : stackSize C x ≤ stackSize C' x := by - apply Finset.sum_le_sum_of_subset (fun x ↦ ?_) - simp [iff_true_intro (@h x)] - --- Simplify the cast of `stackSize C x` from `ℕ` to `ℝ` -lemma stackSize_real (C : Set (𝔓 X)) (x : X) : (stackSize C x : ℝ) = - ∑ p ∈ { p | p ∈ C }, (𝓘 p : Set X).indicator (1 : X → ℝ) x := by - rw [stackSize, Nat.cast_sum] - refine Finset.sum_congr rfl (fun u _ ↦ ?_) - by_cases hx : x ∈ (𝓘 u : Set X) <;> simp [hx] - -lemma stackSize_measurable : Measurable fun x ↦ (stackSize C x : ℝ≥0∞) := by - simp_rw [stackSize, Nat.cast_sum, indicator, Nat.cast_ite] - refine Finset.measurable_sum _ fun _ _ ↦ Measurable.ite coeGrid_measurable ?_ ?_ <;> simp - -/-! We might want to develop some API about partitioning a set. -But maybe `Set.PairwiseDisjoint` and `Set.Union` are enough. -Related, but not quite useful: `Setoid.IsPartition`. -/ - --- /-- `u` is partitioned into subsets in `C`. -/ --- class Set.IsPartition {α ι : Type*} (u : Set α) (s : Set ι) (C : ι → Set α) : Prop := --- pairwiseDisjoint : s.PairwiseDisjoint C --- iUnion_eq : ⋃ (i ∈ s), C i = u - - namespace TileStructure -- variable (X) in -- structure Tree where @@ -71,73 +28,58 @@ variable (X) in /-- An `n`-forest -/ structure Forest (n : ℕ) where 𝔘 : Set (𝔓 X) - 𝔗 : 𝔓 X → Set (𝔓 X) -- Is it a problem that we totalized this function? - nonempty {u} (hu : u ∈ 𝔘) : (𝔗 u).Nonempty - ordConnected {u} (hu : u ∈ 𝔘) : OrdConnected (𝔗 u) -- (2.0.33) - 𝓘_ne_𝓘 {u} (hu : u ∈ 𝔘) {p} (hp : p ∈ 𝔗 u) : 𝓘 p ≠ 𝓘 u - smul_four_le {u} (hu : u ∈ 𝔘) {p} (hp : p ∈ 𝔗 u) : smul 4 p ≤ smul 1 u -- (2.0.32) - stackSize_le {x} : stackSize 𝔘 x ≤ 2 ^ n -- (2.0.34), we formulate this a bit differently. - dens₁_𝔗_le {u} (hu : u ∈ 𝔘) : dens₁ (𝔗 u : Set (𝔓 X)) ≤ 2 ^ (4 * a - n + 1 : ℤ) -- (2.0.35) - lt_dist {u u'} (hu : u ∈ 𝔘) (hu' : u' ∈ 𝔘) (huu' : u ≠ u') {p} (hp : p ∈ 𝔗 u') + 𝔗 : 𝔓 X → Set (𝔓 X) + nonempty' {u} (hu : u ∈ 𝔘) : (𝔗 u).Nonempty + ordConnected' {u} (hu : u ∈ 𝔘) : OrdConnected (𝔗 u) -- (2.0.33) + 𝓘_ne_𝓘' {u} (hu : u ∈ 𝔘) {p} (hp : p ∈ 𝔗 u) : 𝓘 p ≠ 𝓘 u + smul_four_le' {u} (hu : u ∈ 𝔘) {p} (hp : p ∈ 𝔗 u) : smul 4 p ≤ smul 1 u -- (2.0.32) + stackSize_le' {x} : stackSize 𝔘 x ≤ 2 ^ n -- (2.0.34), we formulate this a bit differently. + dens₁_𝔗_le' {u} (hu : u ∈ 𝔘) : dens₁ (𝔗 u) ≤ 2 ^ (4 * (a : ℝ) - n + 1) -- (2.0.35) + lt_dist' {u u'} (hu : u ∈ 𝔘) (hu' : u' ∈ 𝔘) (huu' : u ≠ u') {p} (hp : p ∈ 𝔗 u') (h : 𝓘 p ≤ 𝓘 u) : 2 ^ (Z * (n + 1)) < dist_(p) (𝒬 p) (𝒬 u) -- (2.0.36) - ball_subset {u} (hu : u ∈ 𝔘) {p} (hp : p ∈ 𝔗 u) : ball (𝔠 p) (8 * D ^ 𝔰 p) ⊆ 𝓘 u -- (2.0.37) - -- old conditions - -- disjoint_I : ∀ {𝔗 𝔗'}, 𝔗 ∈ I → 𝔗' ∈ I → Disjoint 𝔗.carrier 𝔗'.carrier - -- top_finite (x : X) : {𝔗 ∈ I | x ∈ Grid (𝓘 𝔗.top)}.Finite - -- card_top_le (x : X) : Nat.card {𝔗 ∈ I | x ∈ Grid (𝓘 𝔗.top) } ≤ 2 ^ n * Real.log (n + 1) - -- density_le {𝔗} (h𝔗 : 𝔗 ∈ I) : density G Q 𝔗 ≤ (2 : ℝ) ^ (-n : ℤ) - -- delta_gt {j j'} (hj : j ∈ I) (hj' : j' ∈ I) (hjj' : j ≠ j') {p : 𝔓 X} (hp : p ∈ j) - -- (h2p : Grid (𝓘 p) ⊆ Grid (𝓘 j'.top)) : Δ p (Q j.top) > (2 : ℝ) ^ (3 * n / δ) - -end TileStructure - ---below is old - --- class Tree.IsThin (𝔗 : Tree X) : Prop where --- thin {p : 𝔓 X} (hp : p ∈ 𝔗) : ball (𝔠 p) (8 * a/-fix-/ * D ^ 𝔰 p) ⊆ Grid (𝓘 𝔗.top) + ball_subset' {u} (hu : u ∈ 𝔘) {p} (hp : p ∈ 𝔗 u) : ball (𝔠 p) (8 * D ^ 𝔰 p) ⊆ 𝓘 u -- (2.0.37) --- alias Tree.thin := Tree.IsThin.thin +namespace Forest --- def Δ (p : 𝔓 X) (Q' : C(X, ℝ)) : ℝ := localOscillation (Grid (𝓘 p)) (𝒬 p) Q' + 1 +variable {n : ℕ} (t : Forest X n) --- namespace Forest +instance : CoeHead (Forest X n) (Set (𝔓 X)) := ⟨Forest.𝔘⟩ +instance : Membership (𝔓 X) (Forest X n) := ⟨fun t x ↦ x ∈ (t : Set (𝔓 X))⟩ +instance : CoeFun (Forest X n) (fun _ ↦ 𝔓 X → Set (𝔓 X)) := ⟨fun t x ↦ t.𝔗 x⟩ -/- Do we want to treat a forest as a set of trees, or a set of elements from `𝔓 X`? -/ +@[simp] lemma mem_𝔘 : u ∈ t.𝔘 ↔ u ∈ t := .rfl +@[simp] lemma mem_𝔗 : p ∈ t.𝔗 u ↔ p ∈ t u := .rfl --- instance : SetLike (Forest G Q δ n) (Tree X) where --- coe s := s.I --- coe_injective' p q h := by cases p; cases q; congr +lemma nonempty (hu : u ∈ t) : (t u).Nonempty := t.nonempty' hu +lemma ordConnected (hu : u ∈ t) : OrdConnected (t u) := t.ordConnected' hu +lemma 𝓘_ne_𝓘 (hu : u ∈ t) (hp : p ∈ t u) : 𝓘 p ≠ 𝓘 u := t.𝓘_ne_𝓘' hu hp +lemma smul_four_le (hu : u ∈ t) (hp : p ∈ t u) : smul 4 p ≤ smul 1 u := t.smul_four_le' hu hp +lemma stackSize_le : stackSize t x ≤ 2 ^ n := t.stackSize_le' +lemma dens₁_𝔗_le (hu : u ∈ t) : dens₁ (t u) ≤ 2 ^ (4 * (a : ℝ) - n + 1) := t.dens₁_𝔗_le' hu +lemma lt_dist (hu : u ∈ t) (hu' : u' ∈ t) (huu' : u ≠ u') {p} (hp : p ∈ t u') (h : 𝓘 p ≤ 𝓘 u) : + 2 ^ (Z * (n + 1)) < dist_(p) (𝒬 p) (𝒬 u) := t.lt_dist' hu hu' huu' hp h +lemma ball_subset (hu : u ∈ t) (hp : p ∈ t u) : ball (𝔠 p) (8 * D ^ 𝔰 p) ⊆ 𝓘 u := + t.ball_subset' hu hp --- instance : PartialOrder (Forest G Q δ n) := PartialOrder.lift (↑) SetLike.coe_injective +end Forest --- class IsThin (𝔉 : Forest G Q δ n) : Prop where --- thin {𝔗} (h𝔗 : 𝔗 ∈ 𝔉.I) : 𝔗.IsThin - --- alias thin := Forest.IsThin.thin - --- /-- The union of all the trees in the forest. -/ --- def carrier (𝔉 : Forest G Q δ n) : Set (𝔓 X) := ⋃ 𝔗 ∈ 𝔉.I, 𝔗 +variable (X) in +/-- An `n`-row -/ +structure Row (n : ℕ) extends Forest X n where + pairwiseDisjoint' : 𝔘.PairwiseDisjoint 𝔗 ---end Forest +namespace Row --- set_option linter.unusedVariables false in --- variable (X) in --- class SmallBoundaryProperty (η : ℝ) : Prop where --- volume_diff_le : ∃ (C : ℝ) (hC : C > 0), ∀ (x : X) r (δ : ℝ), 0 < r → 0 < δ → δ < 1 → --- volume.real (ball x ((1 + δ) * r) \ ball x ((1 - δ) * r)) ≤ C * δ ^ η * volume.real (ball x r) +variable {n : ℕ} (t : Row X n) ---def boundedTiles (F : Set X) (t : ℝ) : Set (𝔓 X) := --- { p : 𝔓 X | ∃ x ∈ 𝓘 p, maximalFunction volume (Set.indicator F (1 : X → ℂ)) x ≤ t } +instance : CoeHead (Row X n) (Set (𝔓 X)) := ⟨fun t ↦ t.𝔘⟩ +instance : Membership (𝔓 X) (Row X n) := ⟨fun t x ↦ x ∈ (t : Set (𝔓 X))⟩ +instance : CoeFun (Row X n) (fun _ ↦ 𝔓 X → Set (𝔓 X)) := ⟨fun t x ↦ t.𝔗 x⟩ --- set_option linter.unusedVariables false in --- variable (X) in --- class SmallBoundaryProperty (η : ℝ) : Prop where --- volume_diff_le : ∃ (C : ℝ) (hC : C > 0), ∀ (x : X) r (δ : ℝ), 0 < r → 0 < δ → δ < 1 → --- volume.real (ball x ((1 + δ) * r) \ ball x ((1 - δ) * r)) ≤ C * δ ^ η * volume.real (ball x r) +@[simp] lemma mem_𝔘 : u ∈ t.𝔘 ↔ u ∈ t := .rfl +@[simp] lemma mem_𝔗 : p ∈ t.𝔗 u ↔ p ∈ t u := .rfl -/- This is defined to live in `ℝ≥0∞`. Use `ENNReal.toReal` to get a real number. -/ -/- def MB_p {ι : Type*} [Fintype ι] (p : ℝ) (ℬ : ι → X × ℝ) (u : X → ℂ) (x : X) : ℝ≥0∞ := - ⨆ (i : ι) , indicator (ball (ℬ i).1 (ℬ i).2) (1 : X → ℝ≥0∞) x / volume (ball (ℬ i).1 (ℬ i).2) * - (∫⁻ y in (ball (ℬ i).1 (ℬ i).2), ‖u y‖₊^p)^(1/p) +lemma pairwiseDisjoint : Set.PairwiseDisjoint t t := t.pairwiseDisjoint' -abbrev MB {ι : Type*} [Fintype ι] (ℬ : ι → X × ℝ) (u : X → ℂ) (x : X) := MB_p 1 ℬ u x -/ +end Row +end TileStructure diff --git a/Carleson/ForestOperator.lean b/Carleson/ForestOperator.lean index 96c04b06..ca360aab 100644 --- a/Carleson/ForestOperator.lean +++ b/Carleson/ForestOperator.lean @@ -1,11 +1,12 @@ import Carleson.Forest import Carleson.HardyLittlewood +import Carleson.Discrete.Forests open ShortVariables TileStructure variable {X : Type*} {a : ℕ} {q : ℝ} {K : X → X → ℂ} {σ₁ σ₂ : X → ℤ} {F G : Set X} [MetricSpace X] [ProofData a q K σ₁ σ₂ F G] [TileStructure Q D κ S o] - {n : ℕ} {t : Forest X n} {u : 𝔓 X} {x x' : X} {G : Set (𝔓 X)} {f g : X → ℂ} - {I J L : Grid X} + {n j j' : ℕ} {t : Forest X n} {u u₁ u₂ p : 𝔓 X} {x x' : X} {𝔖 : Set (𝔓 X)} + {f f₁ f₂ g g₁ g₂ : X → ℂ} {I J J' L : Grid X} variable {E' : Type*} [NormedAddCommGroup E'] [NormedSpace ℝ E'] noncomputable section @@ -19,50 +20,29 @@ namespace TileStructure.Forest variable (t) in /-- The definition `σ(u, x)` given in Section 7.1. -We may assume `u ∈ t.𝔘` whenever proving things about this definition. -/ -def σ (u : 𝔓 X) (x : X) : Finset ℤ := .image 𝔰 { p | p ∈ t.𝔗 u ∧ x ∈ E p } +We may assume `u ∈ t` whenever proving things about this definition. -/ +def σ (u : 𝔓 X) (x : X) : Finset ℤ := .image 𝔰 { p | p ∈ t u ∧ x ∈ E p } /- Maybe we should try to avoid using \overline{σ} and \underline{σ} in Lean: I don't think the set is always non-empty(?) -/ -- def σMax (u : 𝔓 X) (x : X) : ℤ := --- Finset.image 𝔰 { p | p ∈ t.𝔗 u ∧ x ∈ E p } |>.max' sorry +-- Finset.image 𝔰 { p | p ∈ t u ∧ x ∈ E p } |>.max' sorry -/-- Lemma 7.1.1, freely translated. -/ -lemma convex_scales (hu : u ∈ t.𝔘) : OrdConnected (t.σ u x : Set ℤ) := sorry - -/-- The definition of `𝓙₀(G), defined above Lemma 7.1.2 -/ -def 𝓙₀ (G : Set (𝔓 X)) : Set (Grid X) := - {J : Grid X | s J = - S ∨ ∀ p ∈ G, ¬ (𝓘 p : Set X) ⊆ ball (c J) (100 * D ^ (s J + 1)) } - -/-- The definition of `𝓙(G), defined above Lemma 7.1.2 -/ -def 𝓙 (G : Set (𝔓 X)) : Set (Grid X) := - {x | Maximal (· ∈ 𝓙₀ G) x} +/-- The definition of `𝓙₀(𝔖), defined above Lemma 7.1.2 -/ +def 𝓙₀ (𝔖 : Set (𝔓 X)) : Set (Grid X) := + { J : Grid X | s J = - S ∨ ∀ p ∈ 𝔖, ¬ (𝓘 p : Set X) ⊆ ball (c J) (100 * D ^ (s J + 1)) } -/-- The definition of `𝓛₀(G), defined above Lemma 7.1.2 -/ -def 𝓛₀ (G : Set (𝔓 X)) : Set (Grid X) := - { L : Grid X | s L = - S ∨ (∃ p ∈ G, L ≤ 𝓘 p) ∧ ∀ p ∈ G, ¬ 𝓘 p ≤ L } - -/-- The definition of `𝓛(G), defined above Lemma 7.1.2 -/ -def 𝓛 (G : Set (𝔓 X)) : Set (Grid X) := - {x | Maximal (· ∈ 𝓛₀ G) x} - -/-- Part of Lemma 7.1.2 -/ -@[simp] -lemma biUnion_𝓙 : ⋃ J ∈ 𝓙 G, J = ⋃ I : Grid X, (I : Set X) := by - sorry - -/-- Part of Lemma 7.1.2 -/ -lemma pairwiseDisjoint_𝓙 : (𝓙 G).PairwiseDisjoint (fun I ↦ (I : Set X)) := by - sorry +/-- The definition of `𝓙(𝔖), defined above Lemma 7.1.2 -/ +def 𝓙 (𝔖 : Set (𝔓 X)) : Set (Grid X) := + { x | Maximal (· ∈ 𝓙₀ 𝔖) x} -/-- Part of Lemma 7.1.2 -/ -@[simp] -lemma biUnion_𝓛 : ⋃ J ∈ 𝓛 G, J = ⋃ I : Grid X, (I : Set X) := by - sorry +/-- The definition of `𝓛₀(𝔖), defined above Lemma 7.1.2 -/ +def 𝓛₀ (𝔖 : Set (𝔓 X)) : Set (Grid X) := + { L : Grid X | s L = - S ∨ (∃ p ∈ 𝔖, L ≤ 𝓘 p) ∧ ∀ p ∈ 𝔖, ¬ 𝓘 p ≤ L } -/-- Part of Lemma 7.1.2 -/ -lemma pairwiseDisjoint_𝓛 : (𝓛 G).PairwiseDisjoint (fun I ↦ (I : Set X)) := by - sorry +/-- The definition of `𝓛(𝔖), defined above Lemma 7.1.2 -/ +def 𝓛 (𝔖 : Set (𝔓 X)) : Set (Grid X) := + { x | Maximal (· ∈ 𝓛₀ 𝔖) x} /-- The projection operator `P_𝓒 f(x)`, given above Lemma 7.1.3. In lemmas the `c` will be pairwise disjoint on `C`. -/ @@ -84,18 +64,40 @@ def nontangentialMaximalFunction (θ : Θ X) (f : X → ℂ) (x : X) : ℝ≥0 variable (t) in /-- The operator `S_{1,𝔲} f(x)`, given in (7.1.4). -/ -def boundaryOperator1 (u : 𝔓 X) (f : X → ℂ) (x : X) : ℝ≥0∞ := +def boundaryOperator (u : 𝔓 X) (f : X → ℂ) (x : X) : ℝ≥0∞ := ∑ I : Grid X, (I : Set X).indicator (x := x) fun _ ↦ - ∑ J ∈ { J | J ∈ 𝓙 (t.𝔗 u) ∧ (J : Set X) ⊆ ball (c I) (16 * D ^ (s I)) ∧ s J ≤ s I }, + ∑ J ∈ { J | J ∈ 𝓙 (t u) ∧ (J : Set X) ⊆ ball (c I) (16 * D ^ (s I)) ∧ s J ≤ s I }, D ^ ((s J - s I) / a) / volume (ball (c I) (16 * D ^ (s I))) * ∫⁻ y in J, ‖f y‖₊ /-- The indexing set for the collection of balls 𝓑, defined above Lemma 7.1.3. -/ def 𝓑 : Set (ℕ × Grid X) := Icc 0 (S + 5) ×ˢ univ -/-- The radius function for the collection of balls 𝓑, defined above Lemma 7.1.3. -/ +/-- The center function for the collection of balls 𝓑. -/ +def c𝓑 (z : ℕ × Grid X) : X := c z.2 + +/-- The radius function for the collection of balls 𝓑. -/ def r𝓑 (z : ℕ × Grid X) : ℝ := 2 ^ z.1 * D ^ s z.2 --- def 𝓑 : Set (Set X) := (fun (i, I) ↦ ball (c I) (2 ^ i * D ^ s I)) '' Icc 0 (S + 5) ×ˢ univ +/-- Lemma 7.1.1, freely translated. -/ +lemma convex_scales (hu : u ∈ t) : OrdConnected (t.σ u x : Set ℤ) := sorry + +/-- Part of Lemma 7.1.2 -/ +@[simp] +lemma biUnion_𝓙 : ⋃ J ∈ 𝓙 𝔖, J = ⋃ I : Grid X, (I : Set X) := by + sorry + +/-- Part of Lemma 7.1.2 -/ +lemma pairwiseDisjoint_𝓙 : (𝓙 𝔖).PairwiseDisjoint (fun I ↦ (I : Set X)) := by + sorry + +/-- Part of Lemma 7.1.2 -/ +@[simp] +lemma biUnion_𝓛 : ⋃ J ∈ 𝓛 𝔖, J = ⋃ I : Grid X, (I : Set X) := by + sorry + +/-- Part of Lemma 7.1.2 -/ +lemma pairwiseDisjoint_𝓛 : (𝓛 𝔖).PairwiseDisjoint (fun I ↦ (I : Set X)) := by + sorry /-- The constant used in `first_tree_pointwise`. Has value `10 * 2 ^ (105 * a ^ 3)` in the blueprint. -/ @@ -103,17 +105,17 @@ Has value `10 * 2 ^ (105 * a ^ 3)` in the blueprint. -/ def C7_1_4 (a : ℕ) : ℝ≥0 := 10 * 2 ^ (105 * (a : ℝ) ^ 3) /-- Lemma 7.1.4 -/ -lemma first_tree_pointwise (hu : u ∈ t.𝔘) (hL : L ∈ 𝓛 (t.𝔗 u)) (hx : x ∈ L) (hx' : x' ∈ L) +lemma first_tree_pointwise (hu : u ∈ t) (hL : L ∈ 𝓛 (t u)) (hx : x ∈ L) (hx' : x' ∈ L) (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : ‖∑ i in t.σ u x, ∫ y, (exp (.I * (- 𝒬 u y + Q x y + 𝒬 u x - Q x x)) - 1) * Ks i x y * f y ‖₊ ≤ - C7_1_4 a * MB volume 𝓑 (c ·.2) r𝓑 (approxOnCube (𝓙 (t.𝔗 u)) (‖f ·‖)) x' := by + C7_1_4 a * MB volume 𝓑 c𝓑 r𝓑 (approxOnCube (𝓙 (t u)) (‖f ·‖)) x' := by sorry /-- Lemma 7.1.5 -/ -lemma second_tree_pointwise (hu : u ∈ t.𝔘) (hL : L ∈ 𝓛 (t.𝔗 u)) (hx : x ∈ L) (hx' : x' ∈ L) +lemma second_tree_pointwise (hu : u ∈ t) (hL : L ∈ 𝓛 (t u)) (hx : x ∈ L) (hx' : x' ∈ L) (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : - ‖∑ i in t.σ u x, ∫ y, Ks i x y * approxOnCube (𝓙 (t.𝔗 u)) f y‖₊ ≤ - nontangentialMaximalFunction (𝒬 u) (approxOnCube (𝓙 (t.𝔗 u)) f) x' := by + ‖∑ i in t.σ u x, ∫ y, Ks i x y * approxOnCube (𝓙 (t u)) f y‖₊ ≤ + nontangentialMaximalFunction (𝒬 u) (approxOnCube (𝓙 (t u)) f) x' := by sorry /-- The constant used in `third_tree_pointwise`. @@ -122,10 +124,10 @@ Has value `2 ^ (151 * a ^ 3)` in the blueprint. -/ def C7_1_6 (a : ℕ) : ℝ≥0 := 2 ^ (151 * (a : ℝ) ^ 3) /-- Lemma 7.1.6 -/ -lemma third_tree_pointwise (hu : u ∈ t.𝔘) (hL : L ∈ 𝓛 (t.𝔗 u)) (hx : x ∈ L) (hx' : x' ∈ L) +lemma third_tree_pointwise (hu : u ∈ t) (hL : L ∈ 𝓛 (t u)) (hx : x ∈ L) (hx' : x' ∈ L) (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : - ‖∑ i in t.σ u x, ∫ y, Ks i x y * (f y - approxOnCube (𝓙 (t.𝔗 u)) f y)‖₊ ≤ - C7_1_6 a * t.boundaryOperator1 u (approxOnCube (𝓙 (t.𝔗 u)) (‖f ·‖)) x' := by + ‖∑ i in t.σ u x, ∫ y, Ks i x y * (f y - approxOnCube (𝓙 (t u)) f y)‖₊ ≤ + C7_1_6 a * t.boundaryOperator u (approxOnCube (𝓙 (t u)) (‖f ·‖)) x' := by sorry /-- The constant used in `pointwise_tree_estimate`. @@ -134,13 +136,13 @@ Has value `2 ^ (151 * a ^ 3)` in the blueprint. -/ def C7_1_3 (a : ℕ) : ℝ≥0 := 2 ^ (151 * (a : ℝ) ^ 3) /-- Lemma 7.1.3. -/ -lemma pointwise_tree_estimate (hu : u ∈ t.𝔘) (hL : L ∈ 𝓛 (t.𝔗 u)) (hx : x ∈ L) (hx' : x' ∈ L) +lemma pointwise_tree_estimate (hu : u ∈ t) (hL : L ∈ 𝓛 (t u)) (hx : x ∈ L) (hx' : x' ∈ L) (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : - ‖∑ p ∈ { p | p ∈ t.𝔗 u }, carlesonOn p (fun y ↦ exp (.I * - 𝒬 u y) * f y) x‖₊ ≤ - C7_1_3 a * (MB volume 𝓑 (c ·.2) r𝓑 (approxOnCube (𝓙 (t.𝔗 u)) (‖f ·‖)) x' + - t.boundaryOperator1 u (approxOnCube (𝓙 (t.𝔗 u)) (‖f ·‖)) x' + - nontangentialMaximalFunction (𝒬 u) (approxOnCube (𝓙 (t.𝔗 u)) f) x'):= by - set g := approxOnCube (𝓙 (t.𝔗 u)) (‖f ·‖) + ‖carlesonSum (t u) (fun y ↦ exp (.I * - 𝒬 u y) * f y) x‖₊ ≤ + C7_1_3 a * (MB volume 𝓑 c𝓑 r𝓑 (approxOnCube (𝓙 (t u)) (‖f ·‖)) x' + + t.boundaryOperator u (approxOnCube (𝓙 (t u)) (‖f ·‖)) x' + + nontangentialMaximalFunction (𝒬 u) (approxOnCube (𝓙 (t u)) f) x'):= by + set g := approxOnCube (𝓙 (t u)) (‖f ·‖) sorry @@ -153,7 +155,7 @@ def C7_2_2 (a : ℕ) : ℝ≥0 := 2 ^ (103 * (a : ℝ) ^ 3) /-- Lemma 7.2.2. -/ lemma nontangential_operator_bound - (hf : IsBounded (range f)) (h2f : HasCompactSupport f) (θ : Θ X) : + (hf : IsBounded (range f)) (h2f : HasCompactSupport f) (θ : Θ X) : eLpNorm (nontangentialMaximalFunction θ f · |>.toReal) 2 volume ≤ eLpNorm f 2 volume := by sorry @@ -165,8 +167,8 @@ lemma boundary_overlap (I : Grid X) : /-- Lemma 7.2.3. -/ lemma boundary_operator_bound - (hf : IsBounded (range f)) (h2f : HasCompactSupport f) {u : 𝔓 X} (hu : u ∈ t.𝔘) : - eLpNorm (boundaryOperator1 t u f · |>.toReal) 2 volume ≤ eLpNorm f 2 volume := by + (hf : IsBounded (range f)) (h2f : HasCompactSupport f) (hu : u ∈ t) : + eLpNorm (boundaryOperator t u f · |>.toReal) 2 volume ≤ eLpNorm f 2 volume := by sorry /-- The constant used in `nontangential_operator_bound`. @@ -176,39 +178,510 @@ def C7_2_1 (a : ℕ) : ℝ≥0 := 2 ^ (104 * (a : ℝ) ^ 3) /-- Lemma 7.2.1. -/ lemma tree_projection_estimate - (hf : IsBounded (range f)) (h2f : HasCompactSupport f) - (hg : IsBounded (range g)) (h2g : HasCompactSupport g) {u : 𝔓 X} (hu : u ∈ t.𝔘) : - ‖∫ x, ∑ p ∈ t.𝔗 u, conj (g x) * carlesonOn p f x‖₊ ≤ - C7_2_1 a * eLpNorm (approxOnCube (𝓙 (t.𝔗 u)) (‖f ·‖)) 2 volume * - eLpNorm (approxOnCube (𝓛 (t.𝔗 u)) (‖g ·‖)) 2 volume := by + (hf : IsBounded (range f)) (h2f : HasCompactSupport f) + (hg : IsBounded (range g)) (h2g : HasCompactSupport g) (hu : u ∈ t) : + ‖∫ x, conj (g x) * carlesonSum (t u) f x‖₊ ≤ + C7_2_1 a * eLpNorm (approxOnCube (𝓙 (t u)) (‖f ·‖)) 2 volume * + eLpNorm (approxOnCube (𝓛 (t u)) (‖g ·‖)) 2 volume := by sorry /-! ## Section 7.3 and Lemma 7.3.1 -/ +/-- The constant used in `local_dens1_tree_bound`. +Has value `2 ^ (101 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_3_2 (a : ℕ) : ℝ≥0 := 2 ^ (101 * (a : ℝ) ^ 3) + +/-- Lemma 7.3.2. -/ +lemma local_dens1_tree_bound (hu : u ∈ t) (hL : L ∈ 𝓛 (t u)) : + volume (L ∩ ⋃ (p ∈ t u), E p) ≤ C7_3_2 a * dens₁ (t u) * volume (L : Set X) := by + sorry + +/-- The constant used in `local_dens2_tree_bound`. +Has value `2 ^ (200 * a ^ 3 + 19)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +-- feel free to modify the constant to something simpler. +def C7_3_3 (a : ℕ) : ℝ≥0 := 2 ^ (201 * (a : ℝ) ^ 3) + +/-- Lemma 7.3.3. -/ +lemma local_dens2_tree_bound (hJ : J ∈ 𝓙 (t u)) {q : 𝔓 X} (hq : q ∈ t u) + (hJq : ¬ Disjoint (J : Set X) (𝓘 q)) : + volume (F ∩ J) ≤ C7_3_3 a * dens₂ (t u) * volume (J : Set X) := by + sorry + +/-- The constant used in `density_tree_bound1`. +Has value `2 ^ (155 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_3_1_1 (a : ℕ) : ℝ≥0 := 2 ^ (155 * (a : ℝ) ^ 3) + +/-- First part of Lemma 7.3.1. -/ +lemma density_tree_bound1 + (hf : IsBounded (range f)) (h2f : HasCompactSupport f) + (hg : IsBounded (range g)) (h2g : HasCompactSupport g) (hu : u ∈ t) : + ‖∫ x, conj (g x) * carlesonSum (t u) f x‖₊ ≤ + C7_3_1_1 a * dens₁ (t u) ^ (2 : ℝ)⁻¹ * eLpNorm f 2 volume * eLpNorm g 2 volume := by + sorry + +/-- The constant used in `density_tree_bound2`. +Has value `2 ^ (256 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_3_1_2 (a : ℕ) : ℝ≥0 := 2 ^ (256 * (a : ℝ) ^ 3) + +/-- Second part of Lemma 7.3.1. -/ +lemma density_tree_bound2 + (hf : IsBounded (range f)) (h2f : HasCompactSupport f) (h3f : ∀ x, ‖f x‖ ≤ F.indicator 1 x) + (hg : IsBounded (range g)) (h2g : HasCompactSupport g) (hu : u ∈ t) : + ‖∫ x, conj (g x) * carlesonSum (t u) f x‖₊ ≤ + C7_3_1_2 a * dens₁ (t u) ^ (2 : ℝ)⁻¹ * dens₂ (t u) ^ (2 : ℝ)⁻¹ * + eLpNorm f 2 volume * eLpNorm g 2 volume := by + sorry + +/-! ## Section 7.4 except Lemmas 4-6 -/ + +/-- The definition of `Tₚ*g(x), defined above Lemma 7.4.1 -/ +def adjointCarleson (p : 𝔓 X) (f : X → ℂ) (x : X) : ℂ := + ∫ y in E p, conj (Ks (𝔰 p) y x) * exp (.I * (Q y y - Q y x)) * f y + +/-- The definition of `T_ℭ*g(x), defined at the bottom of Section 7.4 -/ +def adjointCarlesonSum (ℭ : Set (𝔓 X)) (f : X → ℂ) (x : X) : ℂ := + ∑ p ∈ {p | p ∈ ℭ}, adjointCarleson p f x + +variable (t) in +/-- The operator `S_{2,𝔲} f(x)`, given above Lemma 7.4.3. -/ +def adjointBoundaryOperator (u : 𝔓 X) (f : X → ℂ) (x : X) : ℝ≥0∞ := + ‖adjointCarlesonSum (t u) f x‖₊ + MB volume 𝓑 c𝓑 r𝓑 f x + ‖f x‖₊ + +variable (t u₁ u₂) in +/-- The set `𝔖` defined in the proof of Lemma 7.4.4. +We append a subscript 0 to distinguish it from the section variable. -/ +def 𝔖₀ : Set (𝔓 X) := { p ∈ t u₁ ∪ t u₂ | 2 ^ ((Z : ℝ) * n / 2) ≤ dist_(p) (𝒬 u₁) (𝒬 u₂) } + +/-- Part 1 of Lemma 7.4.1. +Todo: update blueprint with precise properties needed on the function. -/ +lemma adjoint_tile_support1 (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : + adjointCarleson p f = + (ball (𝔠 p) (5 * D ^ 𝔰 p)).indicator (adjointCarleson p ((𝓘 p : Set X).indicator f)) := by + sorry + +/-- Part 2 of Lemma 7.4.1. +Todo: update blueprint with precise properties needed on the function. -/ +lemma adjoint_tile_support2 (hu : u ∈ t) (hp : p ∈ t u) + (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : + adjointCarleson p f = + (𝓘 p : Set X).indicator (adjointCarleson p ((𝓘 p : Set X).indicator f)) := by + sorry + +/-- The constant used in `adjoint_tree_estimate`. +Has value `2 ^ (155 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_4_2 (a : ℕ) : ℝ≥0 := 2 ^ (155 * (a : ℝ) ^ 3) +/-- Lemma 7.4.2. -/ +lemma adjoint_tree_estimate (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : + eLpNorm (adjointCarlesonSum (t u) f) 2 volume ≤ + C7_4_2 a * dens₁ (t u) ^ (2 : ℝ)⁻¹ * eLpNorm f 2 volume := by + sorry + +/-- The constant used in `adjoint_tree_control`. +Has value `2 ^ (156 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_4_3 (a : ℕ) : ℝ≥0 := 2 ^ (155 * (a : ℝ) ^ 3) -/-! ## Section 7.4 and Lemma 7.4.4 -/ +/-- Lemma 7.4.3. -/ +lemma adjoint_tree_control (hu : u ∈ t) (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : + eLpNorm (adjointBoundaryOperator t u f · |>.toReal) 2 volume ≤ + C7_4_3 a * eLpNorm f 2 volume := by + sorry + +/-- Part 2 of Lemma 7.4.7. -/ +lemma 𝔗_subset_𝔖₀ (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) : t u₁ ⊆ 𝔖₀ t u₁ u₂ := by + sorry + +/-- Part 1 of Lemma 7.4.7. -/ +lemma overlap_implies_distance (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hp : p ∈ t u₁ ∪ t u₂) + (hpu₁ : ¬ Disjoint (𝓘 p : Set X) (𝓘 u₁)) : p ∈ 𝔖₀ t u₁ u₂ := by + sorry +/-! ## Section 7.5 -/ + +variable (t u₁ u₂) in +/-- The definition `𝓙'` at the start of Section 7.5.1. +We use a different notation to distinguish it from the 𝓙' used in Section 7.6 -/ +def 𝓙₅ : Set (Grid X) := 𝓙 (𝔖₀ t u₁ u₂) ∩ Iic (𝓘 u₁) + +/-- The definition of χ-tilde, defined in the proof of Lemma 7.5.2 -/ +def χtilde (J : Grid X) (x : X) : ℝ≥0 := + 8 - D ^ (- s J) * dist x (c J) |>.toNNReal + +variable (t u₁ u₂) in +/-- The definition of χ, defined in the proof of Lemma 7.5.2 -/ +def χ (J : Grid X) (x : X) : ℝ≥0 := + χtilde J x / ∑ J' ∈ { I | I ∈ 𝓙₅ t u₁ u₂ }, χtilde J' x + +-- /-- The definition of `B`, defined in (7.5.1) -/ +-- protected def _root_.Grid.ball (I : Grid X) : Set X := ball (c I) (8 * D ^ s I) + +variable (t u₁ u₂) in +/-- The definition of h_J, defined in the proof of Section 7.5.2 -/ +def holderFunction (f₁ f₂ : X → ℂ) (J : Grid X) (x : X) : ℂ := + χ t u₁ u₂ J x * (exp (.I * 𝒬 u₁ x) * adjointCarlesonSum (t u₁) f₁ x) * + conj (exp (.I * 𝒬 u₂ x) * adjointCarlesonSum (t u₂ ∩ 𝔖₀ t u₁ u₂) f₂ x) -/-! ### Section 7.5 -/ /-! ### Subsection 7.5.1 and Lemma 7.5.2 -/ +/-- Part of Lemma 7.5.1. -/ +lemma union_𝓙₅ (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) : + ⋃ J ∈ 𝓙₅ t u₁ u₂, (J : Set X) = 𝓘 u₁ := by + sorry + +/-- Part of Lemma 7.5.1. -/ +lemma pairwiseDisjoint_𝓙₅ (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) : + (𝓙₅ t u₁ u₂).PairwiseDisjoint (fun I ↦ (I : Set X)) := by + sorry + +/-- Lemma 7.5.3 (stated somewhat differently). -/ +lemma moderate_scale_change (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hJ : J ∈ 𝓙₅ t u₁ u₂) (hJ' : J' ∈ 𝓙₅ t u₁ u₂) + (h : ¬ Disjoint (J : Set X) J') : s J + 1 ≤ s J' := by + sorry + +/-- The constant used in `dist_χ_χ_le`. +Has value `2 ^ (226 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_5_2 (a : ℕ) : ℝ≥0 := 2 ^ (226 * (a : ℝ) ^ 3) + +/-- Part of Lemma 7.5.2. -/ +lemma sum_χ (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (x : X) : + ∑ J ∈ { I | I ∈ 𝓙₅ t u₁ u₂ }, χ t u₁ u₂ J x = (𝓘 u₁ : Set X).indicator 1 x := by + sorry + +/-- Part of Lemma 7.5.2. -/ +lemma χ_mem_Icc (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hJ : J ∈ 𝓙₅ t u₁ u₂) (hx : x ∈ 𝓘 u₁) : + χ t u₁ u₂ J x ∈ Icc 0 ((ball (c I) (8 * D ^ s I)).indicator 1 x) := by + sorry + +/-- Part of Lemma 7.5.2. -/ +lemma dist_χ_χ_le (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hJ : J ∈ 𝓙₅ t u₁ u₂) (hx : x ∈ 𝓘 u₁) (hx' : x' ∈ 𝓘 u₁) : + dist (χ t u₁ u₂ J x) (χ t u₁ u₂ J x) ≤ C7_5_2 a * dist x x' / D ^ s J := by + sorry /-! ### Subsection 7.5.2 and Lemma 7.5.4 -/ +/-- The constant used in `holder_correlation_tile`. +Has value `2 ^ (151 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_5_5 (a : ℕ) : ℝ≥0 := 2 ^ (151 * (a : ℝ) ^ 3) + +/-- Lemma 7.5.5. -/ +lemma holder_correlation_tile (hu : u ∈ t) (hp : p ∈ t u) + (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : + nndist (exp (.I * 𝒬 u x) * adjointCarleson p f x) + (exp (.I * 𝒬 u x') * adjointCarleson p f x') ≤ + C7_5_5 a / volume (ball (𝔠 p) (4 * D ^ 𝔰 p)) * + (nndist x x' / D ^ (𝔰 p : ℝ)) ^ (a : ℝ)⁻¹ * ∫⁻ x in E p, ‖f x‖₊ := by + sorry + +/-- Lemma 7.5.6. -/ +lemma limited_scale_impact (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hp : p ∈ t u₂ \ 𝔖₀ t u₁ u₂) (hJ : J ∈ 𝓙₅ t u₁ u₂) + (h : ¬ Disjoint (ball (𝔠 p) (8 * D ^ 𝔰 p)) (ball (c J) (8⁻¹ * D ^ s J))) : + 𝔰 p ∈ Icc (s J) (s J + 3) := by + sorry + +/-- The constant used in `local_tree_control`. +Has value `2 ^ (104 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_5_7 (a : ℕ) : ℝ≥0 := 2 ^ (104 * (a : ℝ) ^ 3) + +/-- Lemma 7.5.7. -/ +lemma local_tree_control (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hJ : J ∈ 𝓙₅ t u₁ u₂) + (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : + ⨆ x ∈ ball (c J) (8⁻¹ * D ^ s J), ‖adjointCarlesonSum (t u₂ \ 𝔖₀ t u₁ u₂) f x‖₊ ≤ + C7_5_7 a * ⨅ x ∈ J, MB volume 𝓑 c𝓑 r𝓑 (‖f ·‖) x := by + sorry + +/-- Lemma 7.5.8. -/ +lemma scales_impacting_interval (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hJ : J ∈ 𝓙₅ t u₁ u₂) (hp : p ∈ t u₁ ∪ (t u₂ ∩ 𝔖₀ t u₁ u₂)) + (h : ¬ Disjoint (ball (𝔠 p) (8 * D ^ 𝔰 p)) (ball (c J) (8 * D ^ s J))) : s J ≤ 𝔰 p := by + sorry + +/-- The constant used in `global_tree_control1_1`. +Has value `2 ^ (154 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_5_9_1 (a : ℕ) : ℝ≥0 := 2 ^ (154 * (a : ℝ) ^ 3) + +/-- The constant used in `global_tree_control1_2`. +Has value `2 ^ (153 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_5_9_2 (a : ℕ) : ℝ≥0 := 2 ^ (153 * (a : ℝ) ^ 3) + +/-- Part 1 of Lemma 7.5.9 -/ +lemma global_tree_control1_1 (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (ℭ : Set (𝔓 X)) (hℭ : ℭ = t u₁ ∨ ℭ = t u₂ ∩ 𝔖₀ t u₁ u₂) + (hJ : J ∈ 𝓙₅ t u₁ u₂) (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : + ⨆ x ∈ ball (c J) (8 * D ^ s J), ‖adjointCarlesonSum ℭ f x‖₊ ≤ + (⨅ x ∈ ball (c J) (8⁻¹ * D ^ s J), ‖adjointCarlesonSum ℭ f x‖₊) + + C7_5_9_1 a * ⨅ x ∈ J, MB volume 𝓑 c𝓑 r𝓑 (‖f ·‖) x := by + sorry + +/-- Part 2 of Lemma 7.5.9 -/ +lemma global_tree_control1_2 (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (ℭ : Set (𝔓 X)) (hℭ : ℭ = t u₁ ∨ ℭ = t u₂ ∩ 𝔖₀ t u₁ u₂) + (hJ : J ∈ 𝓙₅ t u₁ u₂) (hf : IsBounded (range f)) (h2f : HasCompactSupport f) + (hx : x ∈ ball (c J) (8 * D ^ s J)) (hx' : x' ∈ ball (c J) (8 * D ^ s J)) : + nndist (exp (.I * 𝒬 u x) * adjointCarlesonSum ℭ f x) + (exp (.I * 𝒬 u x') * adjointCarlesonSum ℭ f x') ≤ + C7_5_9_1 a * (nndist x x' / D ^ (𝔰 p : ℝ)) ^ (a : ℝ)⁻¹ * + ⨅ x ∈ J, MB volume 𝓑 c𝓑 r𝓑 (‖f ·‖) x := by + sorry + +/-- The constant used in `global_tree_control2`. +Has value `2 ^ (155 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_5_10 (a : ℕ) : ℝ≥0 := 2 ^ (155 * (a : ℝ) ^ 3) + +/-- Lemma 7.5.10 -/ +lemma global_tree_control2 (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hJ : J ∈ 𝓙₅ t u₁ u₂) + (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : + ⨆ x ∈ ball (c J) (8 * D ^ s J), ‖adjointCarlesonSum (t u₂ ∩ 𝔖₀ t u₁ u₂) f x‖₊ ≤ + ⨅ x ∈ ball (c J) (8⁻¹ * D ^ s J), ‖adjointCarlesonSum (t u₂) f x‖₊ + + C7_5_10 a * ⨅ x ∈ J, MB volume 𝓑 c𝓑 r𝓑 (‖f ·‖) x := by + sorry + +/-- The constant used in `holder_correlation_tree`. +Has value `2 ^ (535 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_5_4 (a : ℕ) : ℝ≥0 := 2 ^ (535 * (a : ℝ) ^ 3) + +/-- Lemma 7.5.4. -/ +lemma holder_correlation_tree (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hJ : J ∈ 𝓙₅ t u₁ u₂) + (hf₁ : IsBounded (range f₁)) (h2f₁ : HasCompactSupport f₁) + (hf₂ : IsBounded (range f₂)) (h2f₂ : HasCompactSupport f₂) : + HolderOnWith (C7_5_4 a * + ((⨅ x ∈ ball (c J) (8⁻¹ * D ^ s J), ‖adjointCarlesonSum (t u₁) f₁ x‖₊) + + (⨅ x ∈ J, MB volume 𝓑 c𝓑 r𝓑 (‖f₁ ·‖) x).toNNReal) * + ((⨅ x ∈ ball (c J) (8⁻¹ * D ^ s J), ‖adjointCarlesonSum (t u₂) f₂ x‖₊) + + (⨅ x ∈ J, MB volume 𝓑 c𝓑 r𝓑 (‖f₂ ·‖) x).toNNReal)) + nnτ (holderFunction t u₁ u₂ f₁ f₂ J) (ball (c J) (8 * D ^ s J)) := by + sorry /-! ### Subsection 7.5.3 and Lemma 7.4.5 -/ +/-- The constant used in `lower_oscillation_bound`. +Has value `2 ^ (Z * n / 2 - 201 * a ^ 3)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_5_11 (a n : ℕ) : ℝ≥0 := 2 ^ (Z * n / 2 - 201 * (a : ℝ) ^ 3) +/-- Lemma 7.5.11 -/ +lemma lower_oscillation_bound (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hJ : J ∈ 𝓙₅ t u₁ u₂) : + C7_5_11 a n ≤ dist_{c J, 8 * D ^ s J} (𝒬 u₁) (𝒬 u₂) := by + sorry + +/-- The constant used in `correlation_distant_tree_parts`. +Has value `2 ^ (541 * a ^ 3 - Z * n / (4 * a ^ 2 + 2 * a ^ 3))` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_4_5 (a n : ℕ) : ℝ≥0 := 2 ^ (541 * (a : ℝ) ^ 3 - Z * n / (4 * a ^ 2 + 2 * a ^ 3)) + +/-- Lemma 7.4.5 -/ +lemma correlation_distant_tree_parts (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) + (hf₁ : IsBounded (range f₁)) (h2f₁ : HasCompactSupport f₁) + (hf₂ : IsBounded (range f₂)) (h2f₂ : HasCompactSupport f₂) : + ‖∫ x, adjointCarlesonSum (t u₁) g₁ x * conj (adjointCarlesonSum (t u₂ ∩ 𝔖₀ t u₁ u₂) g₂ x)‖₊ ≤ + C7_4_5 a n * + eLpNorm ((𝓘 u₁ : Set X).indicator (adjointBoundaryOperator t u₁ g₁) · |>.toReal) 2 volume * + eLpNorm ((𝓘 u₁ : Set X).indicator (adjointBoundaryOperator t u₂ g₂) · |>.toReal) 2 volume := by + sorry /-! ## Section 7.6 and Lemma 7.4.6 -/ +variable (t u₁) in +/-- The definition `𝓙'` at the start of Section 7.6. +We use a different notation to distinguish it from the 𝓙' used in Section 7.5 -/ +def 𝓙₆ : Set (Grid X) := 𝓙 (t u₁) ∩ Iic (𝓘 u₁) + +/-- Part of Lemma 7.6.1. -/ +lemma union_𝓙₆ (hu₁ : u₁ ∈ t) : + ⋃ J ∈ 𝓙₆ t u₁, (J : Set X) = 𝓘 u₁ := by + sorry + +/-- Part of Lemma 7.6.1. -/ +lemma pairwiseDisjoint_𝓙₆ (hu₁ : u₁ ∈ t) : + (𝓙₆ t u₁).PairwiseDisjoint (fun I ↦ (I : Set X)) := by + sorry + +/-- The constant used in `thin_scale_impact`. This is denoted `s₁` in the proof of Lemma 7.6.3. +Has value `Z * n / (202 * a ^ 3) - 2` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_6_3 (a n : ℕ) : ℝ := Z * n / (202 * a ^ 3) - 2 + +-- if needed +lemma C7_6_3_pos [ProofData a q K σ₁ σ₂ F G] (h : 1 ≤ n) : 0 < C7_6_3 a n := by + sorry + +/-- Lemma 7.6.3. -/ +lemma thin_scale_impact (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hp : p ∈ t u₂ \ 𝔖₀ t u₁ u₂) (hJ : J ∈ 𝓙₆ t u₁) + (h : ¬ Disjoint (ball (𝔠 p) (8 * D ^ 𝔰 p)) (ball (c J) (8 * D ^ s J))) : + 𝔰 p ≤ s J - C7_6_3 a n := by + sorry + +/-- The constant used in `square_function_count`. +Has value `Z * n / (202 * a ^ 3) - 2` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_6_4 (a : ℕ) (s : ℤ) : ℝ≥0 := 2 ^ (104 * (a : ℝ) ^ 2) * (8 * D ^ (- s)) ^ κ + +/-- Lemma 7.6.4. -/ +lemma square_function_count (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hp : p ∈ t u₂ \ 𝔖₀ t u₁ u₂) (hJ : J ∈ 𝓙₆ t u₁) (s' : ℤ /- ? -/) : + ⨍⁻ x : X, (∑ I ∈ {I : Grid X | s I = s J - s' ∧ Disjoint (I : Set X) (𝓘 u₁) ∧ + ¬ Disjoint (J : Set X) (ball (c I) (8 * D ^ s I)) }, + (ball (c I) (8 * D ^ s I)).indicator 1 x) ^ 2 ≤ C7_6_4 a s' := by + sorry + + + + +/-- The constant used in `bound_for_tree_projection`. +Has value `2 ^ (118 * a ^ 3 - 100 / (202 * a) * Z * n * κ)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_6_2 (a n : ℕ) : ℝ≥0 := 2 ^ (118 * (a : ℝ) ^ 3 - 100 / (202 * a) * Z * n * κ) + +/-- Lemma 7.6.2. Todo: add needed hypothesis to LaTeX -/ +lemma bound_for_tree_projection (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : + eLpNorm (approxOnCube (𝓙₆ t u₁) (‖adjointCarlesonSum (t u₂ \ 𝔖₀ t u₁ u₂) f ·‖)) 2 volume ≤ + C7_6_2 a n * + eLpNorm ((𝓘 u₁ : Set X).indicator (MB volume 𝓑 c𝓑 r𝓑 (‖f ·‖) · |>.toReal)) 2 volume := by + sorry + +/-- The constant used in `correlation_near_tree_parts`. +Has value `2 ^ (541 * a ^ 3 - Z * n / (4 * a ^ 2 + 2 * a ^ 3))` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_4_6 (a n : ℕ) : ℝ≥0 := 2 ^ (222 * (a : ℝ) ^ 3 - Z * n * 2 ^ (-10 * (a : ℝ))) + +/-- Lemma 7.4.6 -/ +lemma correlation_near_tree_parts (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) (h2u : 𝓘 u₁ ≤ 𝓘 u₂) + (hf₁ : IsBounded (range f₁)) (h2f₁ : HasCompactSupport f₁) + (hf₂ : IsBounded (range f₂)) (h2f₂ : HasCompactSupport f₂) : + ‖∫ x, adjointCarlesonSum (t u₁) g₁ x * conj (adjointCarlesonSum (t u₂ \ 𝔖₀ t u₁ u₂) g₂ x)‖₊ ≤ + C7_4_6 a n * + eLpNorm ((𝓘 u₁ : Set X).indicator (adjointBoundaryOperator t u₁ g₁) · |>.toReal) 2 volume * + eLpNorm ((𝓘 u₁ : Set X).indicator (adjointBoundaryOperator t u₂ g₂) · |>.toReal) 2 volume := by + sorry + + +/-! ## Lemmas 7.4.4 -/ + +/-- The constant used in `correlation_separated_trees`. +Has value `2 ^ (550 * a ^ 3 - 3 * n)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_4_4 (a n : ℕ) : ℝ≥0 := 2 ^ (550 * (a : ℝ) ^ 3 - 3 * n) + +lemma correlation_separated_trees_of_subset (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (h2u : 𝓘 u₁ ≤ 𝓘 u₂) + (hf₁ : IsBounded (range f₁)) (h2f₁ : HasCompactSupport f₁) + (hf₂ : IsBounded (range f₂)) (h2f₂ : HasCompactSupport f₂) : + ‖∫ x, adjointCarlesonSum (t u₁) g₁ x * conj (adjointCarlesonSum (t u₂) g₂ x)‖₊ ≤ + C7_4_4 a n * + eLpNorm + ((𝓘 u₁ ∩ 𝓘 u₂ : Set X).indicator (adjointBoundaryOperator t u₁ g₁) · |>.toReal) 2 volume * + eLpNorm + ((𝓘 u₁ ∩ 𝓘 u₂ : Set X).indicator (adjointBoundaryOperator t u₂ g₂) · |>.toReal) 2 volume := by + sorry + +/-- Lemma 7.4.4. -/ +lemma correlation_separated_trees (hu₁ : u₁ ∈ t) (hu₂ : u₂ ∈ t) (hu : u₁ ≠ u₂) + (hf₁ : IsBounded (range f₁)) (h2f₁ : HasCompactSupport f₁) + (hf₂ : IsBounded (range f₂)) (h2f₂ : HasCompactSupport f₂) : + ‖∫ x, adjointCarlesonSum (t u₁) g₁ x * conj (adjointCarlesonSum (t u₂) g₂ x)‖₊ ≤ + C7_4_4 a n * + eLpNorm + ((𝓘 u₁ ∩ 𝓘 u₂ : Set X).indicator (adjointBoundaryOperator t u₁ g₁) · |>.toReal) 2 volume * + eLpNorm + ((𝓘 u₁ ∩ 𝓘 u₂ : Set X).indicator (adjointBoundaryOperator t u₂ g₂) · |>.toReal) 2 volume := by + sorry /-! ## Section 7.7 and Proposition 2.0.4 -/ + +/-- The row-decomposition of a tree, defined in the proof of Lemma 7.7.1. +The indexing is off-by-one compared to the blueprint. -/ +def rowDecomp (t : Forest X n) (j : ℕ) : Row X n := sorry + +/-- Part of Lemma 7.7.1 -/ +@[simp] +lemma biUnion_rowDecomp : ⋃ j < 2 ^ n, t.rowDecomp j = (t : Set (𝔓 X)) := by + sorry + +/-- Part of Lemma 7.7.1 -/ +lemma pairwiseDisjoint_rowDecomp : + (Iio (2 ^ n)).PairwiseDisjoint (rowDecomp t · : ℕ → Set (𝔓 X)) := by + sorry + +@[simp] lemma rowDecomp_apply : t.rowDecomp j u = t u := by + sorry + +/-- The constant used in `row_bound`. +Has value `2 ^ (156 * a ^ 3 - n / 2)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_7_2_1 (a n : ℕ) : ℝ≥0 := 2 ^ (156 * (a : ℝ) ^ 3 - n / 2) + +/-- The constant used in `indicator_row_bound`. +Has value `2 ^ (257 * a ^ 3 - n / 2)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_7_2_2 (a n : ℕ) : ℝ≥0 := 2 ^ (257 * (a : ℝ) ^ 3 - n / 2) + +/-- Part of Lemma 7.7.2. -/ +lemma row_bound (hj : j < 2 ^ n) (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : + eLpNorm (∑ u ∈ {p | p ∈ rowDecomp t j}, adjointCarlesonSum (t u) f) 2 volume ≤ + C7_7_2_1 a n * eLpNorm f 2 volume := by + sorry + +/-- Part of Lemma 7.7.2. -/ +lemma indicator_row_bound (hj : j < 2 ^ n) (hf : IsBounded (range f)) (h2f : HasCompactSupport f) : + eLpNorm (∑ u ∈ {p | p ∈ rowDecomp t j}, F.indicator <| adjointCarlesonSum (t u) f) 2 volume ≤ + C7_7_2_2 a n * dens₂ (⋃ u ∈ t, t u) ^ (2 : ℝ)⁻¹ * eLpNorm f 2 volume := by + sorry + +/-- The constant used in `row_correlation`. +Has value `2 ^ (862 * a ^ 3 - 3 * n)` in the blueprint. -/ +-- Todo: define this recursively in terms of previous constants +def C7_7_3 (a n : ℕ) : ℝ≥0 := 2 ^ (862 * (a : ℝ) ^ 3 - 2 * n) + +/-- Lemma 7.7.3. -/ +lemma row_correlation (hjj' : j < j') (hj' : j' < 2 ^ n) + (hf₁ : IsBounded (range f₁)) (h2f₁ : HasCompactSupport f₁) + (hf₂ : IsBounded (range f₂)) (h2f₂ : HasCompactSupport f₂) : + ‖∫ x, (∑ u ∈ {p | p ∈ rowDecomp t j}, adjointCarlesonSum (t u) f₁ x) * + (∑ u ∈ {p | p ∈ rowDecomp t j'}, adjointCarlesonSum (t u) f₂ x)‖₊ ≤ + C7_7_3 a n * eLpNorm f₁ 2 volume * eLpNorm f₂ 2 volume := by + sorry + +variable (t) in +/-- The definition of `Eⱼ` defined above Lemma 7.7.4. -/ +def rowSupport (j : ℕ) : Set X := ⋃ (u ∈ rowDecomp t j) (p ∈ t u), E p + +/-- Lemma 7.7.4 -/ +lemma pairwiseDisjoint_rowSupport : + (Iio (2 ^ n)).PairwiseDisjoint (rowSupport t) := by + sorry + end TileStructure.Forest /-- The constant used in `forest_operator`. @@ -219,8 +692,7 @@ def C2_0_4 (a q : ℝ) (n : ℕ) : ℝ≥0 := 2 ^ (432 * a ^ 3 - (q - 1) / q * n theorem forest_operator {n : ℕ} (𝔉 : Forest X n) {f g : X → ℂ} (hf : Measurable f) (h2f : ∀ x, ‖f x‖ ≤ F.indicator 1 x) (hg : Measurable g) (h2g : IsBounded (support g)) : - ‖∫ x, conj (g x) * ∑ u ∈ { p | p ∈ 𝔉.𝔘 }, - ∑ p ∈ { p | p ∈ 𝔉.𝔗 u }, carlesonOn p f x‖₊ ≤ - C2_0_4 a q n * (dens₂ (X := X) (⋃ u ∈ 𝔉.𝔘, 𝔉.𝔗 u)) ^ (q⁻¹ - 2⁻¹) * + ‖∫ x, conj (g x) * ∑ u ∈ { p | p ∈ 𝔉 }, carlesonSum (𝔉 u) f x‖₊ ≤ + C2_0_4 a q n * (dens₂ (X := X) (⋃ u ∈ 𝔉, 𝔉 u)) ^ (q⁻¹ - 2⁻¹) * eLpNorm f 2 volume * eLpNorm g 2 volume := by sorry diff --git a/Carleson/HardyLittlewood.lean b/Carleson/HardyLittlewood.lean index b5df8bc5..480ae8c6 100644 --- a/Carleson/HardyLittlewood.lean +++ b/Carleson/HardyLittlewood.lean @@ -39,7 +39,7 @@ def maximalFunction (μ : Measure X) (𝓑 : Set ι) (c : ι → X) (r : ι → /-- The Hardy-Littlewood maximal function w.r.t. a collection of balls 𝓑 with exponent 1. M_𝓑 in the blueprint. -/ -abbrev MB (μ : Measure X) (𝓑 : Set ι) (c : ι → X) (r : ι → ℝ) (u : X → E) (x : X) := +abbrev MB (μ : Measure X) (𝓑 : Set ι) (c : ι → X) (r : ι → ℝ) (u : X → E) (x : X) : ℝ≥0∞ := maximalFunction μ 𝓑 c r 1 u x -- We will replace the criterion `P` used in `MeasureTheory.SublinearOn.maximalFunction` with the diff --git a/Carleson/TileStructure.lean b/Carleson/TileStructure.lean index ec608586..3ca72fd8 100644 --- a/Carleson/TileStructure.lean +++ b/Carleson/TileStructure.lean @@ -97,6 +97,12 @@ def carlesonOn (p : 𝔓 X) (f : X → ℂ) : X → ℂ := indicator (E p) fun x ↦ ∫ y, exp (I * (Q x y - Q x x)) * K x y * ψ (D ^ (- 𝔰 p) * dist x y) * f y +open Classical in +/-- The operator `T_ℭ f` defined at the bottom of Section 7.4. +We will use this in other places of the formalization as well. -/ +def carlesonSum (ℭ : Set (𝔓 X)) (f : X → ℂ) (x : X) : ℂ := + ∑ p ∈ {p | p ∈ ℭ}, carlesonOn p f x + lemma carlesonOn_def' (p : 𝔓 X) (f : X → ℂ) : carlesonOn p f = indicator (E p) fun x ↦ ∫ y, Ks (𝔰 p) x y * f y * exp (I * (Q x y - Q x x)) := by unfold carlesonOn Ks @@ -281,3 +287,41 @@ lemma isAntichain_iff_disjoint (𝔄 : Set (𝔓 X)) : IsAntichain (·≤·) (toTileLike (X := X) '' 𝔄) ↔ ∀ p p', p ∈ 𝔄 → p' ∈ 𝔄 → p ≠ p' → Disjoint (toTileLike (X := X) p).toTile (toTileLike p').toTile := sorry + +/-! ### Stack sizes -/ + +variable {C C' : Set (𝔓 X)} {x x' : X} +open scoped Classical + +/-- The number of tiles `p` in `s` whose underlying cube `𝓘 p` contains `x`. -/ +def stackSize (C : Set (𝔓 X)) (x : X) : ℕ := + ∑ p ∈ { p | p ∈ C }, (𝓘 p : Set X).indicator 1 x + +lemma stackSize_setOf_add_stackSize_setOf_not {P : 𝔓 X → Prop} : + stackSize {p ∈ C | P p} x + stackSize {p ∈ C | ¬ P p} x = stackSize C x := by + classical + simp_rw [stackSize] + conv_rhs => rw [← Finset.sum_filter_add_sum_filter_not _ P] + simp_rw [Finset.filter_filter] + congr + +lemma stackSize_congr (h : ∀ p ∈ C, x ∈ (𝓘 p : Set X) ↔ x' ∈ (𝓘 p : Set X)) : + stackSize C x = stackSize C x' := by + refine Finset.sum_congr rfl fun p hp ↦ ?_ + simp_rw [Finset.mem_filter, Finset.mem_univ, true_and] at hp + simp_rw [indicator, h p hp, Pi.one_apply] + +lemma stackSize_mono (h : C ⊆ C') : stackSize C x ≤ stackSize C' x := by + apply Finset.sum_le_sum_of_subset (fun x ↦ ?_) + simp [iff_true_intro (@h x)] + +-- Simplify the cast of `stackSize C x` from `ℕ` to `ℝ` +lemma stackSize_real (C : Set (𝔓 X)) (x : X) : (stackSize C x : ℝ) = + ∑ p ∈ { p | p ∈ C }, (𝓘 p : Set X).indicator (1 : X → ℝ) x := by + rw [stackSize, Nat.cast_sum] + refine Finset.sum_congr rfl (fun u _ ↦ ?_) + by_cases hx : x ∈ (𝓘 u : Set X) <;> simp [hx] + +lemma stackSize_measurable : Measurable fun x ↦ (stackSize C x : ℝ≥0∞) := by + simp_rw [stackSize, Nat.cast_sum, indicator, Nat.cast_ite] + refine Finset.measurable_sum _ fun _ _ ↦ Measurable.ite coeGrid_measurable ?_ ?_ <;> simp diff --git a/blueprint/src/chapter/main.tex b/blueprint/src/chapter/main.tex index 64cbaaa1..00a0aff7 100644 --- a/blueprint/src/chapter/main.tex +++ b/blueprint/src/chapter/main.tex @@ -27,7 +27,7 @@ \chapter{Introduction} -In \cite{carleson}, L. Carleson addressed a classical question regarding the convergence of Fourier series of continuous functions by proving their pointwise convergence almost everywhere. \Cref{classical-Carleson} represents a version of this result. +In \cite{carleson}, L. Carleson addressed a classical question regarding the convergence of Fourier series of continuous functions by proving their pointwise convergence almost everywhere, as stated below in \Cref{classical-carleson}. Let $f$ be a complex valued, $2\pi$-periodic bounded Borel measurable function on the real line, and for an integer $n$, define the Fourier coefficient as \begin{equation} @@ -39,25 +39,19 @@ \chapter{Introduction} \end{equation} \begin{theorem}[classical Carleson] - \label{classical-Carleson} + \label{classical-carleson} \leanok \lean{classical_carleson} - \uses{smooth-approximation, convergence-for-smooth, - control-approximation-effect} - Let $f$ be a $2\pi$-periodic complex-valued continuous function on $\mathbb{R}$. For all $\epsilon>0$, there exists a Borel set $E\subset [0,2\pi]$ with Lebesgue measure $|E|\le \epsilon$ and a positive integer $N_0$ such that for all $x\in [0,2\pi]\setminus E$ and all integers $N>N_0$, we have - \begin{equation}\label{aeconv} - |f(x)-S_N f(x)|\le \epsilon. + \uses{exceptional-set-carleson} + Let $f$ be a $2\pi$-periodic complex-valued continuous function on $\mathbb{R}$. + Then for almost all $x \in \mathbb{R}$ we have + \begin{equation}\label{eq:fourier_limit} + \lim_{n\to\infty}S_n f(x) = f(x). \end{equation} + In other words: the set of real numbers $x$ where \eqref{eq:fourier_limit} fails has Lebesgue measure $0$. \end{theorem} -%Note that mere continuity implies uniform continuity in the setting of this theorem. -By applying this theorem with a sequence of $\epsilon_n:= 2^{-n}\delta$ for $n\ge 1$ and taking the union -of corresponding exceptional sets $E_n$, we see that outside a set of measure $\delta$, the partial Fourier -sums converge pointwise for $N\to \infty$. Applying this with a sequence of $\delta$ shrinking to zero and -taking the intersection of the corresponding exceptional sets, which has measure zero, we see that the Fourier -series converges outside a set of measure zero. - -The purpose of this paper is twofold. On the one hand, it prepares computer verification of \Cref{classical-Carleson} by presenting a very detailed proof as a blueprint for coding in Lean. We pass through a bound for a generalization of the so-called Carleson operator to doubling metric measure spaces. This generalization is new, and proving these bounds constitutes the second purpose of this paper. This generalization incorporates several results from the recent literature, most prominently bounds for the polynomial Carleson operator of V. Lie \cite{lie-polynomial} as well as its generalization \cite{zk-polynomial}. A computer verification of our theorem will also entail a computer verification for the bulk of the work in these results. +The purpose of this paper is twofold. On the one hand, it prepares computer verification of \Cref{classical-carleson} by presenting a very detailed proof as a blueprint for coding in Lean. We pass through a bound for a generalization of the so-called Carleson operator to doubling metric measure spaces. This generalization is new, and proving these bounds constitutes the second purpose of this paper. This generalization incorporates several results from the recent literature, most prominently bounds for the polynomial Carleson operator of V. Lie \cite{lie-polynomial} as well as its generalization \cite{zk-polynomial}. A computer verification of our theorem will also entail a computer verification for the bulk of the work in these results. We proceed to introduce the setup for our general theorem. We carry a multi purpose parameter, a natural number @@ -2037,7 +2031,7 @@ \section{Proof of the Exceptional Sets Lemma} {\mu(F\cap B(\pc(\fp),r(\fp)))}\ge K{\mu(B(\pc(\fp),r(\fp)))}\, . $$ This ball exists by definition of $\fP_{F,G}$ - and $\dens_2$. By applying \Cref{Hardy-Littlewood}to the collection of balls + and $\dens_2$. By applying \Cref{Hardy-Littlewood} to the collection of balls $$ \mathcal{B} = \{B(\pc(\fp),r(\fp)) \ : \ \fp \in \fP_{F,G}\} $$ @@ -3513,7 +3507,7 @@ \section{The density arguments}\label{sec-TT*-T*T} = 2^{255a^3+110a + 1} {\dens_1(\mathfrak{A})^{\frac 1p}} \int |g|(y)(M_{\mathcal{B}', p'}g)(y) \, dy\,. \end{equation} -Applying Cauchy-Schwarz and using \Cref{Hardy-Littlewood}estimates the last display by +Applying Cauchy-Schwarz and using \Cref{Hardy-Littlewood} estimates the last display by \begin{equation} 2^{255a^3+110a + 1} \dens_1(\mathfrak{A})^{\frac 1p} \|g\|_2 \|M_{\mathcal{B}', p'} g\|_2 @@ -4755,6 +4749,8 @@ \section{The quantitative \texorpdfstring{$L^2$}{L2} tree estimate} \begin{lemma}[densities tree bound] \label{densities-tree-bound} + \leanok + \lean{TileStructure.Forest.density_tree_bound1, TileStructure.Forest.density_tree_bound2} \uses{tree-projection-estimate,local-dens1-tree-bound,local-dens2-tree-bound} Let $\fu \in \fU$. Then for all $f,g$ bounded with bounded support \begin{equation} @@ -4772,6 +4768,8 @@ \section{The quantitative \texorpdfstring{$L^2$}{L2} tree estimate} \begin{lemma}[local dens1 tree bound] \label{local-dens1-tree-bound} + \leanok + \lean{TileStructure.Forest.local_dens1_tree_bound} \uses{monotone-cube-metrics} Let $\fu \in \fU$ and $L \in \mathcal{L}(\fT(\fu))$. Then \begin{equation} @@ -4782,6 +4780,8 @@ \section{The quantitative \texorpdfstring{$L^2$}{L2} tree estimate} \begin{lemma}[local dens2 tree bound] \label{local-dens2-tree-bound} + \leanok + \lean{TileStructure.Forest.local_dens2_tree_bound} Let $J \in \mathcal{J}(\fT(\fu))$ be such that there exist $\fq \in \fT(\fu)$ with $J \cap \scI(\fq) \ne \emptyset$. Then $$ \mu(F \cap J) \le 2^{200a^3 + 19} \dens_2(\fT(\fu)) \mu(J)\,. @@ -4948,6 +4948,8 @@ \section{Almost orthogonality of separated trees} \begin{lemma}[adjoint tile support] \label{adjoint-tile-support} + \leanok + \lean{TileStructure.Forest.adjoint_tile_support1, TileStructure.Forest.adjoint_tile_support2} For each $\fp \in \fP$, we have $$ T_{\fp}^* g = \mathbf{1}_{B(\pc(\fp), 5D^{\ps(\fp)})} T_{\fp}^* \mathbf{1}_{\scI(\fp)} g\,. @@ -4979,6 +4981,8 @@ \section{Almost orthogonality of separated trees} \begin{lemma}[adjoint tree estimate] \label{adjoint-tree-estimate} + \leanok + \lean{TileStructure.Forest.adjoint_tree_estimate} \uses{densities-tree-bound} For all bounded $g$ with bounded support, we have that $$ @@ -5004,15 +5008,17 @@ \section{Almost orthogonality of separated trees} $$ \begin{lemma}[adjoint tree control] \label{adjoint-tree-control} + \leanok + \lean{TileStructure.Forest.adjoint_tree_control} \uses{adjoint-tree-estimate} - We have for all bounded $g$ with bounded support + We have for all $\fu \in \fU$ and bounded $g$ with bounded support $$ \|S_{2, \fu} g\|_2 \le 2^{156a^3} \|g\|_2\,. $$ \end{lemma} \begin{proof} - This follows immediately from Minkowski's inequality, \Cref{Hardy-Littlewood}and \Cref{adjoint-tree-estimate}, using that $a \ge 4$. + This follows immediately from Minkowski's inequality, \Cref{Hardy-Littlewood} and \Cref{adjoint-tree-estimate}, using that $a \ge 4$. \end{proof} @@ -5020,6 +5026,8 @@ \section{Almost orthogonality of separated trees} \begin{lemma}[correlation separated trees] \label{correlation-separated-trees} + \leanok + \lean{TileStructure.Forest.correlation_separated_trees} \uses{correlation-distant-tree-parts,correlation-near-tree-parts} For any $\fu_1 \ne \fu_2 \in \fU$ and all bounded $g_1, g_2$ with bounded support, we have \begin{equation} @@ -5046,6 +5054,8 @@ \section{Almost orthogonality of separated trees} \begin{lemma}[correlation distant tree parts] \label{correlation-distant-tree-parts} + \leanok + \lean{TileStructure.Forest.correlation_distant_tree_parts} \uses{Holder-van-der-Corput,Lipschitz-partition-unity,Holder-correlation-tree,lower-oscillation-bound} We have for all $\fu_1 \ne \fu_2 \in \fU$ with $\scI(\fu_1) \subset \scI(\fu_2)$ and all bounded $g_1, g_2$ with bounded support \begin{equation} @@ -5059,6 +5069,8 @@ \section{Almost orthogonality of separated trees} \end{lemma} \begin{lemma}[correlation near tree parts] \label{correlation-near-tree-parts} + \leanok + \lean{TileStructure.Forest.correlation_near_tree_parts} \uses{tree-projection-estimate,dyadic-partition-2,bound-for-tree-projection} We have for all $\fu_1 \ne \fu_2 \in \fU$ with $\scI(\fu_1) \subset \scI(\fu_2)$ and all bounded $g_1, g_2$ with bounded support \begin{equation} @@ -5067,7 +5079,7 @@ \section{Almost orthogonality of separated trees} \end{equation} \begin{equation} \label{eq-rhs-small-sep-tree} - \le 2^{222a^3} 2^{-Zn 2^{-10a}} \prod_{j =1}^2 \| S_{2, \fu_j} g_j\|_{L^2(\scI(\fu_1)}\,. + \le 2^{222a^3} 2^{-Zn 2^{-10a}} \prod_{j =1}^2 \| S_{2, \fu_j} g_j\|_{L^2(\scI(\fu_1))}\,. \end{equation} \end{lemma} @@ -5075,6 +5087,8 @@ \section{Almost orthogonality of separated trees} \begin{lemma}[overlap implies distance] \label{overlap-implies-distance} + \leanok + \lean{TileStructure.Forest.𝔗_subset_𝔖₀, TileStructure.Forest.overlap_implies_distance} Let $\fu_1 \ne \fu_2 \in \fU$ with $\scI(\fu_1) \subset \scI(\fu_2)$. If $\fp \in \fT(\fu_1) \cup \fT(\fu_2)$ with $\scI(\fp) \cap \scI(\fu_1) \ne \emptyset$, then $\fp \in \mathfrak{S}$. In particular, we have $\fT(\fu_1) \subset \mathfrak{S}$. \end{lemma} @@ -5113,6 +5127,8 @@ \subsection{A partition of unity} \begin{lemma}[dyadic partition 1] \label{dyadic-partition-1} + \leanok + \lean{TileStructure.Forest.union_𝓙₅, TileStructure.Forest.pairwiseDisjoint_𝓙₅} \uses{dyadic-partitions} We have that $$ @@ -5133,7 +5149,10 @@ \subsection{A partition of unity} \begin{lemma}[Lipschitz partition unity] \label{Lipschitz-partition-unity} - \uses{dyadic-partition-1,moderate-scale-change} + \leanok + \lean{TileStructure.Forest.sum_χ, TileStructure.Forest.χ_mem_Icc, + TileStructure.Forest.dist_χ_χ_le} + \uses{dyadic-partition-1,moderate-scale-change} There exists a family of functions $\chi_J$, $J \in \mathcal{J}'$ such that \begin{equation} \label{eq-pao-1} \mathbf{1}_{\scI(\fu_1)} = \sum_{J \in \mathcal{J}'} \chi_J\,, @@ -5153,6 +5172,8 @@ \subsection{A partition of unity} \begin{lemma}[moderate scale change] \label{moderate-scale-change} + \leanok + \lean{TileStructure.Forest.moderate_scale_change} If $J, J' \in \mathcal{J'}$ with $$ @@ -5250,6 +5271,8 @@ \subsection{H\"older estimates for adjoint tree operators} \begin{lemma}[Holder correlation tree] \label{Holder-correlation-tree} + \leanok + \lean{TileStructure.Forest.holder_correlation_tree} \uses{global-tree-control-2} We have for all $J \in \mathcal{J}'$ that \begin{equation} @@ -5263,6 +5286,8 @@ \subsection{H\"older estimates for adjoint tree operators} We begin with the following H\"older continuity estimate for adjoints of operators associated to tiles. \begin{lemma}[Holder correlation tile] \label{Holder-correlation-tile} + \leanok + \lean{TileStructure.Forest.holder_correlation_tile} \uses{adjoint-tile-support} Let $\fu \in \fU$ and $\fp \in \fT(\fu)$. Then for all $y, y' \in X$ and all bounded $g$ with bounded support, we have $$ @@ -5385,8 +5410,10 @@ \subsection{H\"older estimates for adjoint tree operators} \begin{lemma}[limited scale impact] \label{limited-scale-impact} + \leanok + \lean{TileStructure.Forest.limited_scale_impact} \uses{overlap-implies-distance} - Let $\fp \in \fT_2 \setminus \mathfrak{S}$, $J \in \mathcal{J}'$ and suppose that + Let $\fp \in \fT(\fu_2) \setminus \mathfrak{S}$, $J \in \mathcal{J}'$ and suppose that $$ B(\scI(\fp)) \cap B^\circ(J) \ne \emptyset\,. $$ @@ -5433,6 +5460,8 @@ \subsection{H\"older estimates for adjoint tree operators} \begin{lemma}[local tree control] \label{local-tree-control} + \leanok + \lean{TileStructure.Forest.local_tree_control} \uses{limited-scale-impact} For all $J \in \mathcal{J}'$ and all bounded $g$ with bounded support $$ @@ -5475,6 +5504,8 @@ \subsection{H\"older estimates for adjoint tree operators} \begin{lemma}[scales impacting interval] \label{scales-impacting-interval} + \leanok + \lean{TileStructure.Forest.scales_impacting_interval} \uses{overlap-implies-distance} Let $\fC = \fT(\fu_1)$ or $\fC = \fT(\fu_2) \cap \mathfrak{S}$. Then for each $J \in \mathcal{J}'$ and $\fp \in \fC$ with $B(\scI(\fp)) \cap B(J) \neq \emptyset$, we have $\ps(\fp) \ge s(J)$. \end{lemma} @@ -5485,6 +5516,8 @@ \subsection{H\"older estimates for adjoint tree operators} \begin{lemma}[global tree control 1] \label{global-tree-control-1} + \leanok + \lean{TileStructure.Forest.global_tree_control1_1, TileStructure.Forest.global_tree_control1_2} \uses{Holder-correlation-tile,scales-impacting-interval} Let $\fC = \fT(\fu_1)$ or $\fC = \fT(\fu_2) \cap \mathfrak{S}$. Then for each $J \in \mathcal{J}'$ and all bounded $g$ with bounded support, we have \begin{align} @@ -5536,10 +5569,12 @@ \subsection{H\"older estimates for adjoint tree operators} \begin{lemma}[global tree control 2] \label{global-tree-control-2} + \leanok + \lean{TileStructure.Forest.global_tree_control2} \uses{global-tree-control-1, local-tree-control} We have for all $J \in \mathcal{J}'$ and all bounded $g$ with bounded support $$ - \sup_{B(J)} |T^*_{\mathfrak{T}_2 \cap \mathfrak{S}} g| \le \inf_{B^\circ{}(J)} |T^*_{\mathfrak{T}_2} g| + 2^{155a^3} \inf_{J} M_{\mathcal{B},1}|g|\,. + \sup_{B(J)} |T^*_{\fT(\fu_2) \cap \mathfrak{S}} g| \le \inf_{B^\circ{}(J)} |T^*_{\fT(\fu_2)} g| + 2^{155a^3} \inf_{J} M_{\mathcal{B},1}|g|\,. $$ \end{lemma} @@ -5603,6 +5638,8 @@ \subsection{The van der Corput estimate} \label{subsubsec-van-der-corput} \begin{lemma}[lower oscillation bound] \label{lower-oscillation-bound} + \leanok + \lean{TileStructure.Forest.lower_oscillation_bound} \uses{overlap-implies-distance} For all $J \in \mathcal{J}'$, we have that $$ @@ -5677,6 +5714,8 @@ \section{Proof of The Remaining Tiles Lemma} note that this is different from the $\mathcal{J}'$ defined in the previous subsection. \begin{lemma}[dyadic partition 2] \label{dyadic-partition-2} + \leanok + \lean{TileStructure.Forest.union_𝓙₆, TileStructure.Forest.pairwiseDisjoint_𝓙₆} \uses{dyadic-partitions} We have $$ @@ -5692,10 +5731,12 @@ \section{Proof of The Remaining Tiles Lemma} \begin{lemma}[bound for tree projection] \label{bound-for-tree-projection} + \leanok + \lean{TileStructure.Forest.bound_for_tree_projection} \uses{adjoint-tile-support,overlap-implies-distance,dyadic-partition-2,thin-scale-impact,square-function-count} We have $$ - \|P_{\mathcal{J}'}|T_{\mathfrak{T}_2 \setminus \mathfrak{S}}^* g_2|\|_2 + \|P_{\mathcal{J}'}|T_{\fT(\fu_2) \setminus \mathfrak{S}}^* g_2|\|_2 \le 2^{118a^3} 2^{-\frac{100}{202a}Zn\kappa} \|\mathbf{1}_{\scI(\fu_1)} M_{\mathcal{B},1} |g_2|\|_2 $$ \end{lemma} @@ -5723,7 +5764,9 @@ \section{Proof of The Remaining Tiles Lemma} \begin{lemma}[thin scale impact] \label{thin-scale-impact} - If $\fp \in \fT_2 \setminus \mathfrak{S}$ and $J \in \mathcal{J'}$ with $B(\scI(\fp)) \cap B(J) \ne \emptyset$, then + % \leanok + % \lean{TileStructure.Forest.thin_scale_impact} + If $\fp \in \fT(\fu_2) \setminus \mathfrak{S}$ and $J \in \mathcal{J'}$ with $B(\scI(\fp)) \cap B(J) \ne \emptyset$, then $$ \ps(\fp) \le s(J) + 2 - \frac{Zn}{202a^3}\,. $$ @@ -5771,6 +5814,8 @@ \section{Proof of The Remaining Tiles Lemma} \begin{lemma}[square function count] \label{square-function-count} + \leanok + \lean{TileStructure.Forest.square_function_count} For each $J \in \mathcal{J}'$, we have $$ \frac{1}{\mu(J)} \int_J \Bigg(\sum_{\substack{I \in \mathcal{D}, s(I) = s(J) - s\\ I \cap \scI(\fu_1) = \emptyset\\ @@ -5778,8 +5823,7 @@ \section{Proof of The Remaining Tiles Lemma} $$ \end{lemma} - \begin{proof}[Proof of \Cref{square-function-count}] - \proves{square-function-count} + \begin{proof} Since $J \in \mathcal{J}'$ we have $J \subset \scI(\fu_1)$. Thus, if $B(I) \cap J \ne \emptyset$ then \begin{equation} \label{eq-sep-small-incl} @@ -5808,7 +5852,7 @@ \section{Proof of The Remaining Tiles Lemma} \proves{bound-for-tree-projection} Expanding the definition of $P_{\mathcal{J}'}$, we have $$ - \|P_{\mathcal{J}'}|T_{\mathfrak{T}_2 \setminus \mathfrak{S}}^* g_2|\|_2 + \|P_{\mathcal{J}'}|T_{\fT(\fu_2) \setminus \mathfrak{S}}^* g_2|\|_2 $$ $$ = \left(\sum_{J \in \mathcal{J}'} \frac{1}{\mu(J)} \left|\int_J \sum_{\fp \in \fT(\fu_2) \setminus \mathfrak{S}} T_{\fp}^* g_2 \, \mathrm{d}\mu(y) \right|^2 \right)^{1/2}\,. @@ -5902,6 +5946,9 @@ \section{Forests} \begin{lemma}[forest row decomposition] \label{forest-row-decomposition} + \leanok + \lean{TileStructure.Forest.rowDecomp, TileStructure.Forest.biUnion_rowDecomp, + TileStructure.Forest.pairwiseDisjoint_rowDecomp} Let $(\fU, \fT)$ be an $n$-forest. Then there exists a decomposition $$ \fU = \dot{\bigcup_{1 \le j \le 2^n}} \fU_j @@ -5927,6 +5974,8 @@ \section{Forests} \begin{lemma}[row bound] \label{row-bound} + \leanok + \lean{TileStructure.Forest.row_bound, TileStructure.Forest.indicator_row_bound} \uses{densities-tree-bound,adjoint-tile-support} For each $1 \le j \le 2^n$ and each bounded $g$ with bounded support, we have \begin{equation} @@ -5973,6 +6022,8 @@ \section{Forests} \begin{lemma}[row correlation] \label{row-correlation} + \leanok + \lean{TileStructure.Forest.row_correlation} \uses{adjoint-tree-control,correlation-separated-trees} For all $1 \le j < j' \le 2^n$ and for all bounded $g_1, g_2$ with bounded support, it holds that $$ @@ -6030,6 +6081,8 @@ \section{Forests} \begin{lemma}[disjoint row support] \label{disjoint-row-support} + \leanok + \lean{TileStructure.Forest.pairwiseDisjoint_rowSupport} The sets $E_j$, $1 \le j \le 2^n$ are pairwise disjoint. \end{lemma} @@ -6110,6 +6163,8 @@ \chapter{Proof of the H\"older cancellative condition} \begin{lemma}[Lipschitz Holder approximation] \label{Lipschitz-Holder-approximation} + % \leanok + % \lean{TileStructure.Forest.} \leanok \lean{support_holderApprox_subset, dist_holderApprox_le, lipschitzWith_holderApprox} Let $z\in X$ and $R>0$. Let $\varphi: X \to \mathbb{C}$ be a function supported in the ball @@ -7715,8 +7770,6 @@ \section{Calder\'on-Zygmund Decomposition} \chapter{Proof of The Classical Carleson Theorem} - - The convergence of partial Fourier sums is proved in \Cref{10classical} in two steps. In the first step, we establish convergence on a suitable dense subclass of functions. We choose smooth functions as subclass, the convergence is stated in \Cref{convergence-for-smooth} and proved in \Cref{10smooth}. @@ -7810,9 +7863,20 @@ \section{The classical Carleson theorem} \end{equation} \end{lemma} -We are now ready to prove classical Carleson: -\begin{proof} [Proof of \Cref{classical-Carleson}] -\proves{classical-Carleson} +We are now ready to prove \Cref{classical-carleson}. We first prove a version with explicit exceptional sets. + +\begin{theorem}[classical Carleson with exceptional sets] + \label{exceptional-set-carleson} + \leanok + \lean{exceptional_set_carleson} + \uses{smooth-approximation, convergence-for-smooth, + control-approximation-effect} + Let $f$ be a $2\pi$-periodic complex-valued continuous function on $\mathbb{R}$. For all $\epsilon>0$, there exists a Borel set $E\subset [0,2\pi]$ with Lebesgue measure $|E|\le \epsilon$ and a positive integer $N_0$ such that for all $x\in [0,2\pi]\setminus E$ and all integers $N>N_0$, we have + \begin{equation}\label{aeconv} + |f(x)-S_N f(x)|\le \epsilon. + \end{equation} +\end{theorem} +\begin{proof} \leanok Let $N_0$ be as in \Cref{convergence-for-smooth}. For every @@ -7833,6 +7897,18 @@ \section{The classical Carleson theorem} This shows \eqref{aeconv} for the given $E$ and $N_0$. \end{proof} +Now we turn to the proof of the statement of Carleson theorem given in the introduction. +\begin{proof}[Proof of \Cref{classical-carleson}] +\proves{classical-carleson} +\leanok +%Note that mere continuity implies uniform continuity in the setting of this theorem. +By applying \Cref{exceptional-set-carleson} with a sequence of $\epsilon_n:= 2^{-n}\delta$ for $n\ge 1$ and taking the union +of corresponding exceptional sets $E_n$, we see that outside a set of measure $\delta$, the partial Fourier +sums converge pointwise for $N\to \infty$. Applying this with a sequence of $\delta$ shrinking to zero and +taking the intersection of the corresponding exceptional sets, which has measure zero, we see that the Fourier +series converges outside a set of measure zero. +\end{proof} + Let $\kappa:\R\to \C$ be the function defined by $\kappa(0)=0$ and for $0<|x|<1$ \begin{equation}\label{eq-hilker} From 44c867567e8e69bf13a0f4dfccb566b343f35a0b Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Fri, 13 Sep 2024 20:23:01 +0200 Subject: [PATCH 3/4] refactor(RealInterpolation): golf (#126) --- Carleson/RealInterpolation.lean | 110 ++++++++++---------------------- 1 file changed, 35 insertions(+), 75 deletions(-) diff --git a/Carleson/RealInterpolation.lean b/Carleson/RealInterpolation.lean index de0a116e..416390e0 100644 --- a/Carleson/RealInterpolation.lean +++ b/Carleson/RealInterpolation.lean @@ -3075,29 +3075,18 @@ lemma weaktype_estimate_trunc_compl_top {C₀ : ℝ≥0} (hC₀ : C₀ > 0) {p p Trans.trans (distribution_mono_right (Trans.trans obs (zero_le (ENNReal.ofReal t)))) meas_eLpNormEssSup_lt · have p_pos : p > 0 := lt_trans hp₀ hp₀p - have snorm_p_pos : eLpNorm f p μ ≠ 0 := by - intro snorm_0 - apply Ne.symm (ne_of_lt snorm_pos) - apply eLpNormEssSup_eq_zero_iff.mpr - exact (eLpNorm_eq_zero_iff hf.1 p_pos.ne.symm).mp snorm_0 + have snorm_p_pos : eLpNorm f p μ ≠ 0 := + fun snorm_0 ↦ Ne.symm (ne_of_lt snorm_pos) (eLpNormEssSup_eq_zero_iff.mpr + ((eLpNorm_eq_zero_iff hf.1 p_pos.ne.symm).mp snorm_0)) have term_pos : (ENNReal.ofNNReal C₀) ^ p₀.toReal * eLpNorm f p μ ^ p.toReal > 0 := by - apply ENNReal.mul_pos - · apply ne_of_gt - refine rpow_pos_of_nonneg (by positivity) (by positivity) - · apply ne_of_gt - refine rpow_pos_of_nonneg (by positivity) (by positivity) + apply ENNReal.mul_pos <;> exact (rpow_pos_of_nonneg (by positivity) (by positivity)).ne.symm have term_ne_top : (ENNReal.ofNNReal C₀) ^ p₀.toReal * eLpNorm f p μ ^ p.toReal ≠ ⊤ - := by - apply mul_ne_top - · refine rpow_ne_top' ?_ coe_ne_top - exact ENNReal.coe_ne_zero.mpr hC₀.ne.symm - · exact rpow_ne_top' snorm_p_pos (Memℒp.eLpNorm_ne_top hf) + := mul_ne_top (rpow_ne_top' (ENNReal.coe_ne_zero.mpr hC₀.ne.symm) coe_ne_top) + (rpow_ne_top' snorm_p_pos (Memℒp.eLpNorm_ne_top hf)) have d_pos : d > 0 := by rw [hdeq] - apply Real.rpow_pos_of_pos - rw [← zero_toReal] - exact toReal_strict_mono term_ne_top term_pos + exact Real.rpow_pos_of_pos (zero_toReal ▸ toReal_strict_mono term_ne_top term_pos) _ have a_pos : a > 0 := by rw [ha]; positivity have obs : Memℒp (f - trunc f a) p₀ μ := trunc_compl_Lp_Lq_lower hp ⟨hp₀, hp₀p⟩ a_pos hf have wt_est := (h₀T (f - trunc f a) obs).2 @@ -3174,21 +3163,13 @@ lemma weaktype_estimate_trunc_top {C₁ : ℝ≥0} (hC₁ : C₁ > 0) {p p₁ q exact (eLpNorm_eq_zero_iff hf.1 hp.ne.symm).mp snorm_0 have term_pos : (ENNReal.ofNNReal C₁) ^ p₁.toReal * eLpNorm f p μ ^ p.toReal > 0 := by - apply ENNReal.mul_pos - · apply ne_of_gt - exact rpow_pos_of_nonneg (by positivity) (by positivity) - · apply ne_of_gt - exact rpow_pos_of_nonneg (by positivity) (by positivity) - have term_ne_top : (ENNReal.ofNNReal C₁) ^ p₁.toReal * eLpNorm f p μ ^ p.toReal ≠ ⊤ := by - apply mul_ne_top - · refine rpow_ne_top' ?_ coe_ne_top - exact ENNReal.coe_ne_zero.mpr hC₁.ne.symm - · exact rpow_ne_top' snorm_p_pos (Memℒp.eLpNorm_ne_top hf) + apply ENNReal.mul_pos <;> exact (rpow_pos_of_nonneg (by positivity) (by positivity)).ne.symm + have term_ne_top : (ENNReal.ofNNReal C₁) ^ p₁.toReal * eLpNorm f p μ ^ p.toReal ≠ ⊤ := + mul_ne_top (rpow_ne_top' (ENNReal.coe_ne_zero.mpr hC₁.ne.symm) coe_ne_top) + (rpow_ne_top' snorm_p_pos (Memℒp.eLpNorm_ne_top hf)) have d_pos : d > 0 := by rw [hdeq] - apply Real.rpow_pos_of_pos - rw [← zero_toReal] - exact toReal_strict_mono term_ne_top term_pos + exact Real.rpow_pos_of_pos (zero_toReal ▸ toReal_strict_mono term_ne_top term_pos) _ calc _ ≤ ↑C₁ ^ p₁.toReal * ((ENNReal.ofReal (a ^ (p₁.toReal - p.toReal))) * eLpNorm f p μ ^ p.toReal) @@ -3372,7 +3353,6 @@ end MeasureTheory end - noncomputable section open NNReal ENNReal MeasureTheory Set ComputationsChoiceExponent @@ -3444,14 +3424,11 @@ lemma estimate_norm_rpow_range_operator {q : ℝ} {f : α → E₁} (ENNReal.ofReal s) ν * ENNReal.ofReal (s^(q - 1)) + distribution (T (f - trunc f (tc.ton s))) (ENNReal.ofReal s) ν * ENNReal.ofReal (s^(q - 1)) := by rw [rewrite_norm_func hq hA hTf] - apply mul_le_mul' - · exact le_refl _ - · apply setLIntegral_mono' measurableSet_Ioi - intro s s_pos - rw [← add_mul] - apply mul_le_mul' - · exact estimate_distribution_Subadditive_trunc s_pos (tc.ran_ton s s_pos) (le_of_lt hA) ht - · exact le_refl _ + apply mul_le_mul' (le_refl _) + apply setLIntegral_mono' measurableSet_Ioi + intro s s_pos + rw [← add_mul] + exact mul_le_mul' (estimate_distribution_Subadditive_trunc s_pos (tc.ran_ton s s_pos) (le_of_lt hA) ht) (le_refl _) lemma estimate_norm_rpow_range_operator' [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] @@ -3474,7 +3451,7 @@ lemma estimate_norm_rpow_range_operator' eLpNorm (f - trunc f (tc.ton s)) (p₀) μ ^ q₀.toReal * ENNReal.ofReal (s ^ (q.toReal - q₀.toReal - 1))) := by have : ∀ q' q : ℝ, -q' + (q - 1) = q - q' - 1 := by intro q' q; group - rw [← this, ← this] + repeat rw [← this] have p_pos : p > 0 := lt_trans hp₀ hp₀p -- TODO: is there a way to use lintegral_rw₂ conveniently? rw [lintegral_rw_aux power_aux_2, lintegral_rw_aux power_aux_2] @@ -3584,31 +3561,23 @@ lemma simplify_factor₀ {D : ℝ} · congr 1 rw [mul_assoc, ← mul_add] rw [eq_exponents₂ (t := t)] <;> try assumption - rw [mul_assoc, mul_assoc, ← mul_add] - rw [neg_mul] + rw [mul_assoc, mul_assoc, ← mul_add, neg_mul] rw [eq_exponents₃ (t := t)] <;> try assumption simp only [neg_mul, neg_neg] - rw [← mul_assoc, ← add_mul] - rw [← preservation_interpolation ht hp₀.1 hp₁.1 hp] - rw [toReal_inv] + rw [← mul_assoc, ← add_mul, ← preservation_interpolation ht hp₀.1 hp₁.1 hp, toReal_inv] field_simp · exact ne_zero_of_Ioo hF · exact ne_top_of_Ioo hF · exact ne_zero_of_Ioo hF · exact ne_top_of_Ioo hF · exact coe_ne_top - · apply ENNReal.inv_ne_zero.mpr - exact d_ne_top_aux₁ hC₁ - · apply ENNReal.inv_ne_zero.mpr - exact d_ne_top_aux₂ hF + · exact ENNReal.inv_ne_zero.mpr (d_ne_top_aux₁ hC₁) + · exact ENNReal.inv_ne_zero.mpr (d_ne_top_aux₂ hF) · exact d_ne_zero_aux₁ hC₀ · exact d_ne_zero_aux₀ hF · exact d_ne_zero_aux₂ hC₀ hF - · apply mul_ne_zero - · apply ENNReal.inv_ne_zero.mpr - exact d_ne_top_aux₁ hC₁ - · apply ENNReal.inv_ne_zero.mpr - exact d_ne_top_aux₂ hF + · exact mul_ne_zero (ENNReal.inv_ne_zero.mpr (d_ne_top_aux₁ hC₁)) + (ENNReal.inv_ne_zero.mpr (d_ne_top_aux₂ hF)) · exact Or.inr (d_ne_top_aux₂ hF) · exact Or.inr (d_ne_zero_aux₀ hF) · exact d_ne_top_aux₄ hC₀ hC₁ hF @@ -3631,7 +3600,8 @@ lemma simplify_factor₁ {D : ℝ} have p₁pos : p₁ > 0 := hp₁.1 have q₁pos : q₁ > 0 := lt_of_lt_of_le hp₁.1 hp₁.2 have p₁ne_top : p₁ ≠ ⊤ := ne_top_of_le_ne_top hq₁' hp₁.2 - have hp' : p ∈ Ioo 0 ⊤ := interp_exp_in_Ioo_zero_top ht p₀pos p₁pos (by right; exact p₁ne_top) hp + have hp' : p ∈ Ioo 0 ⊤ := interp_exp_in_Ioo_zero_top ht p₀pos p₁pos + (Decidable.not_or_of_imp fun _ ↦ p₁ne_top) hp have p_toReal_pos : p.toReal > 0 := toReal_pos_of_Ioo hp' rw [hD] unfold d @@ -3649,34 +3619,24 @@ lemma simplify_factor₁ {D : ℝ} congr 3 · rw [eq_exponents₆] <;> try assumption · rw [eq_exponents₅] <;> try assumption - · rw [mul_assoc, mul_assoc, ← mul_add] - rw [eq_exponents₈] <;> try assumption - rw [neg_mul] - rw [eq_exponents₇ (ht := ht)] <;> try assumption - rw [← mul_add, ← add_mul, add_comm] - rw [← preservation_interpolation ht hp₀.1 hp₁.1 hp] - rw [toReal_inv] + · rw [mul_assoc, mul_assoc, ← mul_add, eq_exponents₈, neg_mul, + eq_exponents₇ (ht := ht)] <;> try assumption + rw [← mul_add, ← add_mul, add_comm, ← preservation_interpolation ht hp₀.1 hp₁.1 hp, + toReal_inv] field_simp · exact ne_zero_of_Ioo hF · exact ne_top_of_Ioo hF · exact ne_zero_of_Ioo hF · exact ne_top_of_Ioo hF - · apply (ENNReal.coe_pos.mpr hC₁).ne.symm - · exact coe_ne_top - · apply ENNReal.inv_ne_zero.mpr - apply rpow_ne_top' - · apply (ENNReal.coe_pos.mpr hC₁).ne.symm + · exact (ENNReal.coe_pos.mpr hC₁).ne.symm · exact coe_ne_top - · apply ENNReal.inv_ne_zero.mpr - exact d_ne_top_aux₂ hF + · exact ENNReal.inv_ne_zero.mpr (rpow_ne_top' ((ENNReal.coe_pos.mpr hC₁).ne.symm) coe_ne_top) + · exact ENNReal.inv_ne_zero.mpr (d_ne_top_aux₂ hF) · exact d_ne_zero_aux₁ hC₀ · exact d_ne_zero_aux₀ hF · exact d_ne_zero_aux₂ hC₀ hF - · apply mul_ne_zero - · apply ENNReal.inv_ne_zero.mpr - exact d_ne_top_aux₁ hC₁ - · apply ENNReal.inv_ne_zero.mpr - exact d_ne_top_aux₂ hF + · exact mul_ne_zero (ENNReal.inv_ne_zero.mpr (d_ne_top_aux₁ hC₁)) + (ENNReal.inv_ne_zero.mpr (d_ne_top_aux₂ hF)) · exact Or.inr (d_ne_top_aux₂ hF) · exact Or.inr (d_ne_zero_aux₀ hF) · exact d_ne_top_aux₄ hC₀ hC₁ hF From b3f42a752509a55c1ac2789be39331885506be9f Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Sat, 14 Sep 2024 01:38:36 +0400 Subject: [PATCH 4/4] Some real interpolation cleanup (#127) --- Carleson/Forest.lean | 1 + Carleson/RealInterpolation.lean | 926 ++++++++++++++------------------ 2 files changed, 418 insertions(+), 509 deletions(-) diff --git a/Carleson/Forest.lean b/Carleson/Forest.lean index 0f44176c..e1e5f1ae 100644 --- a/Carleson/Forest.lean +++ b/Carleson/Forest.lean @@ -28,6 +28,7 @@ variable (X) in /-- An `n`-forest -/ structure Forest (n : ℕ) where 𝔘 : Set (𝔓 X) + /-- The value of `𝔗 u` only matters when `u ∈ 𝔘`. -/ 𝔗 : 𝔓 X → Set (𝔓 X) nonempty' {u} (hu : u ∈ 𝔘) : (𝔗 u).Nonempty ordConnected' {u} (hu : u ∈ 𝔘) : OrdConnected (𝔗 u) -- (2.0.33) diff --git a/Carleson/RealInterpolation.lean b/Carleson/RealInterpolation.lean index 416390e0..168040c2 100644 --- a/Carleson/RealInterpolation.lean +++ b/Carleson/RealInterpolation.lean @@ -22,106 +22,101 @@ import Mathlib.Analysis.SpecialFunctions.ImproperIntegrals - Definitions - Proof of the real interpolation theorem -/ + +set_option linter.unusedVariables false -- contrapose! is noisy + noncomputable section open ENNReal Real Set MeasureTheory +-- Note (F): can we make `t : ℝ≥0∞` for a large part of the proof? variable {p₀ q₀ p₁ q₁ p q : ℝ≥0∞} {t : ℝ} /-! ## Convenience results for working with (interpolated) exponents -/ namespace ComputationsInterpolatedExponents -lemma preservation_positivity {p q : ℝ} (hp : p > 0) (hq : q > 0) (ht : t ∈ Ioo 0 1) : - 0 < (1 - t) * p⁻¹ + t * q⁻¹ := - add_pos (mul_pos (sub_pos.mpr ht.2) (inv_pos_of_pos hp)) - (mul_pos ht.1 (inv_pos_of_pos hq)) +-- unused +lemma ofReal_mem_Ioo_0_1 (h : t ∈ Ioo 0 1) : ENNReal.ofReal t ∈ Ioo 0 1 := + ⟨ofReal_pos.mpr h.1, ofReal_lt_one.mpr h.2⟩ -lemma ENNReal_preservation_positivity₀ {p q : ℝ≥0∞} (ht : t ∈ Ioo 0 1) (hpq : p ≠ ⊤ ∨ q ≠ ⊤) : - 0 < (1 - ENNReal.ofReal t) * p⁻¹ + (ENNReal.ofReal t) * q⁻¹ := by +lemma ENNReal_preservation_positivity₀ (ht : t ∈ Ioo 0 1) (hpq : p ≠ ⊤ ∨ q ≠ ⊤) : + 0 < (1 - ENNReal.ofReal t) * p⁻¹ + ENNReal.ofReal t * q⁻¹ := by have t_mem : ENNReal.ofReal t ∈ Ioo 0 1 := ⟨ofReal_pos.mpr ht.1, ENNReal.ofReal_one ▸ (ofReal_lt_ofReal_iff zero_lt_one).mpr ht.2⟩ - cases hpq <;> rename_i dir - · exact Left.add_pos_of_pos_of_nonneg (mul_pos ((tsub_pos_of_lt t_mem.2).ne.symm) - (ENNReal.inv_ne_zero.mpr (dir))) (zero_le _) + obtain dir|dir := hpq + · exact Left.add_pos_of_pos_of_nonneg (mul_pos ((tsub_pos_of_lt t_mem.2).ne') + (ENNReal.inv_ne_zero.mpr dir)) (zero_le _) · exact Right.add_pos_of_nonneg_of_pos (zero_le _) - (mul_pos ((ofReal_pos.mpr ht.1).ne.symm) (ENNReal.inv_ne_zero.mpr (dir))) + (mul_pos ((ofReal_pos.mpr ht.1).ne') (ENNReal.inv_ne_zero.mpr (dir))) -lemma ENNReal_preservation_positivity {p q : ℝ≥0∞} (ht : t ∈ Ioo 0 1) (hpq : p ≠ q) : - 0 < (1 - ENNReal.ofReal t) * p⁻¹ + (ENNReal.ofReal t) * q⁻¹ := by +lemma ENNReal_preservation_positivity (ht : t ∈ Ioo 0 1) (hpq : p ≠ q) : + 0 < (1 - ENNReal.ofReal t) * p⁻¹ + ENNReal.ofReal t * q⁻¹ := by apply ENNReal_preservation_positivity₀ ht cases (lt_or_gt_of_ne hpq) <;> rename_i dir <;> exact Ne.ne_or_ne ⊤ hpq -lemma ENNReal_preservation_positivity' {p p₀ p₁ : ℝ≥0∞} (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : p > 0 := - inv_lt_top.mp (hp ▸ Ne.lt_top (add_ne_top.mpr ⟨mul_ne_top (Ne.symm (ne_of_beq_false (by rfl))) - (inv_ne_top.mpr hp₀.ne.symm), mul_ne_top (coe_ne_top) (inv_ne_top.mpr hp₁.ne.symm)⟩)) +lemma ENNReal_preservation_positivity' (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p > 0 := by + rw [← inv_inv p, hp] + simp [ENNReal.mul_eq_top, sub_eq_zero, hp₀.ne', hp₁.ne'] -lemma interp_exp_ne_top {p p₀ p₁ : ℝ≥0∞} (hp₀p₁ : p₀ ≠ p₁) - (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : p ≠ ⊤ := - ENNReal.inv_ne_zero.mp (hp ▸ (ENNReal_preservation_positivity ht hp₀p₁).ne.symm) +lemma interp_exp_ne_top (hp₀p₁ : p₀ ≠ p₁) (ht : t ∈ Ioo 0 1) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p ≠ ⊤ := + ENNReal.inv_ne_zero.mp <| hp ▸ (ENNReal_preservation_positivity ht hp₀p₁).ne' -lemma interp_exp_ne_top' {p p₀ p₁ : ℝ≥0∞} (hp₀p₁ : p₀ ≠ ⊤ ∨ p₁ ≠ ⊤) - (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : p ≠ ⊤ := - ENNReal.inv_ne_zero.mp (hp ▸ (ENNReal_preservation_positivity₀ ht hp₀p₁).ne.symm) +lemma interp_exp_ne_top' (hp₀p₁ : p₀ ≠ ⊤ ∨ p₁ ≠ ⊤) (ht : t ∈ Ioo 0 1) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p ≠ ⊤ := + ENNReal.inv_ne_zero.mp (hp ▸ (ENNReal_preservation_positivity₀ ht hp₀p₁).ne') -lemma interp_exp_eq {p p₀ p₁ : ℝ≥0∞} (hp₀p₁ : p₀ = p₁) +lemma interp_exp_eq (hp₀p₁ : p₀ = p₁) (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p₀ = p := by - have t_pos := ht.1 - have : 1 - ENNReal.ofReal t + ENNReal.ofReal t = 1 := by - rw [← ENNReal.ofReal_one, ← ENNReal.ofReal_sub, ← ENNReal.ofReal_add (p := 1 - t) (q := t), - sub_add_cancel, ofReal_one] <;> try positivity - · exact le_of_lt (Ioo.one_sub_mem ht).1 - rw [← hp₀p₁, ← add_mul, this, one_mul] at hp - rw [← inv_inv (a := p₀), ← inv_inv (a := p), hp] - -lemma interp_exp_lt_top {p p₀ p₁ : ℝ≥0∞} (hp₀p₁ : p₀ ≠ p₁) + rw [← inv_inv (a := p), hp, ← hp₀p₁, ← add_mul, + tsub_add_cancel_of_le <| ofReal_lt_one.mpr ht.2 |>.le, one_mul, inv_inv] + +lemma interp_exp_lt_top (hp₀p₁ : p₀ ≠ p₁) (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : p < ⊤ := + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p < ⊤ := Ne.lt_top <| interp_exp_ne_top hp₀p₁ ht hp -lemma interp_exp_lt_top' {p p₀ p₁ : ℝ≥0∞} (hp₀p₁ : p₀ ≠ ⊤ ∨ p₁ ≠ ⊤) +lemma interp_exp_lt_top' (hp₀p₁ : p₀ ≠ ⊤ ∨ p₁ ≠ ⊤) (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : p < ⊤ := + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p < ⊤ := Ne.lt_top <| interp_exp_ne_top' hp₀p₁ ht hp -lemma interp_exp_between {p p₀ p₁ : ℝ≥0∞} (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) +lemma interp_exp_between (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ < p₁) (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : p ∈ Ioo p₀ p₁ := by + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p ∈ Ioo p₀ p₁ := by refine ⟨?_, ?_⟩ · apply ENNReal.inv_lt_inv.mp rw [hp] - have : p₀⁻¹ = (1 - (ENNReal.ofReal t)) * p₀⁻¹ + (ENNReal.ofReal t) * p₀⁻¹ := by + have : p₀⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₀⁻¹ := by rw [← add_mul, tsub_add_eq_max, max_eq_left_of_lt, one_mul] exact ofReal_lt_one.mpr ht.2 nth_rw 2 [this] gcongr - · exact mul_ne_top (sub_ne_top top_ne_one.symm) (inv_ne_top.mpr hp₀.ne.symm) - · exact (ofReal_pos.mpr ht.1).ne.symm + · exact mul_ne_top (sub_ne_top top_ne_one.symm) (inv_ne_top.mpr hp₀.ne') + · exact (ofReal_pos.mpr ht.1).ne' · exact coe_ne_top · apply ENNReal.inv_lt_inv.mp rw [hp] - have : p₁⁻¹ = (1 - (ENNReal.ofReal t)) * p₁⁻¹ + (ENNReal.ofReal t) * p₁⁻¹ := by + have : p₁⁻¹ = (1 - ENNReal.ofReal t) * p₁⁻¹ + ENNReal.ofReal t * p₁⁻¹ := by rw [← add_mul, tsub_add_eq_max, max_eq_left_of_lt, one_mul] exact ofReal_lt_one.mpr ht.2 nth_rw 1 [this] gcongr - · exact mul_ne_top coe_ne_top (inv_ne_top.mpr hp₁.ne.symm) - · exact (tsub_pos_iff_lt.mpr (ofReal_lt_one.mpr ht.2)).ne.symm + · exact mul_ne_top coe_ne_top (inv_ne_top.mpr hp₁.ne') + · exact (tsub_pos_iff_lt.mpr (ofReal_lt_one.mpr ht.2)).ne' · exact coe_ne_top -lemma one_le_interp_exp_aux {p p₀ p₁ : ℝ≥0∞} (hp₀ : 1 ≤ p₀) (hp₁ : 1 ≤ p₁) - (hp₀p₁ : p₀ < p₁) (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : 1 ≤ p := - le_of_lt (lt_of_le_of_lt hp₀ (interp_exp_between (lt_of_lt_of_le zero_lt_one hp₀) - (lt_of_lt_of_le zero_lt_one hp₁) hp₀p₁ ht hp).1) +lemma one_le_interp_exp_aux (hp₀ : 1 ≤ p₀) (hp₁ : 1 ≤ p₁) (hp₀p₁ : p₀ < p₁) (ht : t ∈ Ioo 0 1) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : 1 ≤ p := + hp₀.trans_lt + (interp_exp_between (zero_lt_one.trans_le hp₀) (zero_lt_one.trans_le hp₁) hp₀p₁ ht hp).1 |>.le -lemma switch_exponents {p p₀ p₁ : ℝ≥0∞} (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : +lemma switch_exponents (ht : t ∈ Ioo 0 1) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p⁻¹ = (1 - ENNReal.ofReal (1 - t)) * p₁⁻¹ + ENNReal.ofReal (1 - t) * p₀⁻¹ := by rw [add_comm, ← ofReal_one, ← ofReal_sub, _root_.sub_sub_cancel, ofReal_sub, ofReal_one] · exact hp @@ -129,25 +124,20 @@ lemma switch_exponents {p p₀ p₁ : ℝ≥0∞} (ht : t ∈ Ioo 0 1) · exact (Ioo.one_sub_mem ht).1.le lemma one_le_toReal {a : ℝ≥0∞} (ha₁ : 1 ≤ a) (ha₂ : a < ⊤) : 1 ≤ a.toReal := - toReal_mono (LT.lt.ne_top ha₂) ha₁ + toReal_mono ha₂.ne_top ha₁ -lemma one_le_interp {p p₀ p₁ : ℝ≥0∞} (hp₀ : 1 ≤ p₀) (hp₁ : 1 ≤ p₁) +lemma one_le_interp (hp₀ : 1 ≤ p₀) (hp₁ : 1 ≤ p₁) (hp₀p₁ : p₀ ≠ p₁) (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : 1 ≤ p := by + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : 1 ≤ p := by rcases (lt_or_gt_of_ne hp₀p₁) with p₀lt_p₁ | p₁lt_p₀ · exact one_le_interp_exp_aux hp₀ hp₁ p₀lt_p₁ ht hp · exact one_le_interp_exp_aux hp₁ hp₀ p₁lt_p₀ (Ioo.one_sub_mem ht) (switch_exponents ht hp) -lemma one_le_interp_toReal {p p₀ p₁ : ℝ≥0∞} (hp₀ : 1 ≤ p₀) (hp₁ : 1 ≤ p₁) +lemma one_le_interp_toReal (hp₀ : 1 ≤ p₀) (hp₁ : 1 ≤ p₁) (hp₀p₁ : p₀ ≠ p₁) (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : 1 ≤ p.toReal := + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : 1 ≤ p.toReal := one_le_toReal (one_le_interp hp₀ hp₁ hp₀p₁ ht hp) (Ne.lt_top (interp_exp_ne_top hp₀p₁ ht hp)) -lemma ENNReal_preservation_positivity_real {p p₀ p₁ : ℝ≥0∞} (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) - (hp₀p₁ : p₀ ≠ p₁) (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : p.toReal > 0 := - toReal_pos (ENNReal_preservation_positivity' hp₀ hp₁ hp).ne.symm (interp_exp_ne_top hp₀p₁ ht hp) - lemma coe_rpow_ne_top {a : ℝ} {q : ℝ} (hq : q ≥ 0): ENNReal.ofReal a ^ q ≠ ⊤ := rpow_ne_top_of_nonneg hq coe_ne_top @@ -165,7 +155,7 @@ lemma rpow_ne_top' {a : ℝ≥0∞} {q : ℝ} (ha : a ≠ 0) (ha' : a ≠ ⊤) · exact (ha' a_top).elim lemma exp_toReal_pos' {q : ℝ≥0∞} (hq : q ≥ 1) (hq' : q < ⊤) : q.toReal > 0 := - toReal_pos (lt_of_lt_of_le zero_lt_one hq).ne.symm (LT.lt.ne_top hq') + toReal_pos (lt_of_lt_of_le zero_lt_one hq).ne' (LT.lt.ne_top hq') lemma ne_top_of_Ico {p q r : ℝ≥0∞} (hq : q ∈ Ico p r) : q ≠ ⊤ := hq.2.ne_top @@ -176,24 +166,24 @@ lemma ne_top_of_Ioo {p q r : ℝ≥0∞} (hq : q ∈ Ioo p r) : q ≠ ⊤ := hq. lemma lt_top_of_Ioo {p q r : ℝ≥0∞} (hq : q ∈ Ioo p r) : q < ⊤ := (ne_top_of_Ioo hq).lt_top lemma ne_top_of_Ioc {p q r : ℝ≥0∞} (hq : q ∈ Ioc p r) (hr : r < ⊤) : q ≠ ⊤ := - (lt_of_le_of_lt hq.2 hr).ne_top + hq.2.trans_lt hr |>.ne_top lemma pos_of_rb_Ioc {p q r : ℝ≥0∞} (hr : q ∈ Ioc p r) : 0 < r := - Trans.trans (Trans.trans (zero_le p) hr.left) hr.right + zero_le p |>.trans_lt hr.1 |>.trans_le hr.2 lemma exp_toReal_ne_zero {q : ℝ≥0∞} (hq : q ≥ 1) (hq' : q < ⊤) : q.toReal ≠ 0 := - (exp_toReal_pos' hq hq').ne.symm + (exp_toReal_pos' hq hq').ne' -- TODO: remove the top one? lemma exp_toReal_ne_zero' {q : ℝ≥0∞} (hq : q > 0) (hq' : q ≠ ⊤) : q.toReal ≠ 0 := - (toReal_pos hq.ne.symm hq').ne.symm + (toReal_pos hq.ne' hq').ne' lemma exp_toReal_ne_zero_of_Ico {q p : ℝ≥0∞} (hq : q ∈ Ico 1 p) : q.toReal ≠ 0 := exp_toReal_ne_zero hq.1 (lt_top_of_Ico hq) lemma pos_of_Ioo {p q r : ℝ≥0∞} (hq : q ∈ Ioo p r) : q > 0 := pos_of_gt hq.1 -lemma ne_zero_of_Ioo {p q r : ℝ≥0∞} (hq : q ∈ Ioo p r) : q ≠ 0 := (pos_of_gt hq.1).ne.symm +lemma ne_zero_of_Ioo {p q r : ℝ≥0∞} (hq : q ∈ Ioo p r) : q ≠ 0 := (pos_of_gt hq.1).ne' lemma pos_of_Icc_1 {p q : ℝ≥0∞} (hp : p ∈ Icc 1 q) : p > 0 := lt_of_lt_of_le zero_lt_one hp.1 @@ -214,7 +204,7 @@ lemma eq_of_rpow_eq (a b: ℝ≥0∞) (c : ℝ) (hc : c ≠ 0) (h : a ^ c = b ^ exact congrFun (congrArg HPow.hPow h) c⁻¹ lemma le_of_rpow_le {a b: ℝ≥0∞} {c : ℝ} (hc : c > 0) (h : a ^ c ≤ b ^ c) : a ≤ b := by - rw [← ENNReal.rpow_rpow_inv hc.ne.symm a, ← ENNReal.rpow_rpow_inv hc.ne.symm b] + rw [← ENNReal.rpow_rpow_inv hc.ne' a, ← ENNReal.rpow_rpow_inv hc.ne' b] exact (ENNReal.rpow_le_rpow_iff (inv_pos_of_pos hc)).mpr h -- TODO : decide if this is wanted @@ -222,60 +212,60 @@ lemma le_of_rpow_le {a b: ℝ≥0∞} {c : ℝ} (hc : c > 0) (h : a ^ c ≤ b ^ -- coe x := ENNReal.ofReal x lemma coe_inv_exponent (hp₀ : p₀ > 0) : ENNReal.ofReal (p₀⁻¹.toReal) = p₀⁻¹ := - ofReal_toReal_eq_iff.mpr (inv_ne_top.mpr hp₀.ne.symm) + ofReal_toReal_eq_iff.mpr (inv_ne_top.mpr hp₀.ne') lemma inv_of_interpolated_pos' (hp₀p₁ : p₀ ≠ p₁) (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : 0 < p⁻¹ := + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : 0 < p⁻¹ := ENNReal.inv_pos.mpr (interp_exp_ne_top hp₀p₁ ht hp) -- TODO: remove, this is redundant, but for now mirror the development for reals... lemma interpolated_pos' (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹): 0 < p := + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹): 0 < p := ENNReal_preservation_positivity' hp₀ hp₁ hp lemma exp_toReal_pos (hp₀ : p₀ > 0) (hp₀' : p₀ ≠ ⊤) : 0 < p₀.toReal := - toReal_pos hp₀.ne.symm hp₀' + toReal_pos hp₀.ne' hp₀' lemma interp_exp_in_Ioo_zero_top (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ ⊤ ∨ p₁ ≠ ⊤) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : p ∈ Ioo 0 ⊤ := + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p ∈ Ioo 0 ⊤ := ⟨interpolated_pos' hp₀ hp₁ hp, interp_exp_lt_top' hp₀p₁ ht hp⟩ lemma inv_toReal_pos_of_ne_top (hp₀ : p₀ > 0) (hp' : p₀ ≠ ⊤) : p₀⁻¹.toReal > 0 := by rw [toReal_inv]; exact inv_pos_of_pos (exp_toReal_pos hp₀ hp') lemma inv_toReal_ne_zero_of_ne_top (hp₀ : p₀ > 0) (hp' : p₀ ≠ ⊤) : p₀⁻¹.toReal ≠ 0 := - (inv_toReal_pos_of_ne_top hp₀ hp').ne.symm + (inv_toReal_pos_of_ne_top hp₀ hp').ne' lemma interp_exp_toReal_pos (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : 0 < p.toReal := - toReal_pos (interpolated_pos' hp₀ hp₁ hp).ne.symm (interp_exp_ne_top hp₀p₁ ht hp) + toReal_pos (interpolated_pos' hp₀ hp₁ hp).ne' (interp_exp_ne_top hp₀p₁ ht hp) lemma interp_exp_toReal_pos' (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) (hp₀p₁ : p₀ ≠ ⊤ ∨ p₁ ≠ ⊤) : 0 < p.toReal := - toReal_pos (interpolated_pos' hp₀ hp₁ hp).ne.symm (interp_exp_ne_top' hp₀p₁ ht hp) + toReal_pos (interpolated_pos' hp₀ hp₁ hp).ne' (interp_exp_ne_top' hp₀p₁ ht hp) lemma interp_exp_inv_pos (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : 0 < p⁻¹.toReal := by rw [toReal_inv]; exact inv_pos_of_pos (interp_exp_toReal_pos ht hp₀ hp₁ hp₀p₁ hp) lemma interp_exp_inv_ne_zero (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : p⁻¹.toReal ≠ 0 := - (interp_exp_inv_pos ht hp₀ hp₁ hp₀p₁ hp).ne.symm + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p⁻¹.toReal ≠ 0 := + (interp_exp_inv_pos ht hp₀ hp₁ hp₀p₁ hp).ne' lemma preservation_interpolation (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) - (hp₁ : p₁ > 0) (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp₁ : p₁ > 0) (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p⁻¹.toReal = (1 - t) * (p₀⁻¹).toReal + t * (p₁⁻¹).toReal := by rw [← one_toReal, ← toReal_ofReal (le_of_lt ht.1), ← ENNReal.toReal_sub_of_le] · rw [← toReal_mul, ← toReal_mul, ← toReal_add] · exact congrArg ENNReal.toReal hp - · exact mul_ne_top (sub_ne_top (top_ne_one.symm)) (inv_ne_top.mpr hp₀.ne.symm) - · exact mul_ne_top coe_ne_top (inv_ne_top.mpr hp₁.ne.symm) + · exact mul_ne_top (sub_ne_top (top_ne_one.symm)) (inv_ne_top.mpr hp₀.ne') + · exact mul_ne_top coe_ne_top (inv_ne_top.mpr hp₁.ne') · exact ofReal_le_one.mpr ht.2.le · exact top_ne_one.symm @@ -303,27 +293,27 @@ lemma ne_inv_toReal_exponents (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : rw [← inv_inv p₀, ← inv_inv p₁] apply congrArg Inv.inv have coe_p₀ : ENNReal.ofReal (p₀⁻¹).toReal = p₀⁻¹ := - ofReal_toReal_eq_iff.mpr (inv_ne_top.mpr hp₀.ne.symm) + ofReal_toReal_eq_iff.mpr (inv_ne_top.mpr hp₀.ne') have coe_p₁ : ENNReal.ofReal (p₁⁻¹).toReal = p₁⁻¹ := - ofReal_toReal_eq_iff.mpr (inv_ne_top.mpr hp₁.ne.symm) + ofReal_toReal_eq_iff.mpr (inv_ne_top.mpr hp₁.ne') rw [← coe_p₀, ← coe_p₁] exact congrArg ENNReal.ofReal h lemma ne_inv_toReal_exp_interp_exp (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : (p₀⁻¹.toReal ≠ p⁻¹.toReal) := by rw [preservation_interpolation ht hp₀ hp₁ hp, ← sub_ne_zero, _root_.sub_mul, one_mul, add_comm_sub, sub_add_eq_sub_sub, sub_self, zero_sub, neg_sub, ← _root_.mul_sub] - exact mul_ne_zero ht.1.ne.symm (sub_ne_zero_of_ne (ne_inv_toReal_exponents hp₀ hp₁ hp₀p₁)) + exact mul_ne_zero ht.1.ne' (sub_ne_zero_of_ne (ne_inv_toReal_exponents hp₀ hp₁ hp₀p₁)) lemma ne_sub_toReal_exp (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) : p₁⁻¹.toReal - p₀⁻¹.toReal ≠ 0 := sub_ne_zero_of_ne (ne_inv_toReal_exponents hp₀ hp₁ hp₀p₁).symm lemma ne_toReal_exp_interp_exp (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p₀.toReal ≠ p.toReal := by intro h apply ne_inv_toReal_exp_interp_exp ht hp₀ hp₁ hp₀p₁ hp @@ -331,25 +321,25 @@ lemma ne_toReal_exp_interp_exp (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : exact congrArg Inv.inv h lemma ne_toReal_exp_interp_exp₁ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p.toReal ≠ p₁.toReal := (ne_toReal_exp_interp_exp (Ioo.one_sub_mem ht) hp₁ hp₀ (Ne.symm hp₀p₁) (switch_exponents ht hp)).symm lemma ofReal_inv_interp_sub_exp_pos₁ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) : ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ > 0 := ofReal_pos.mpr (inv_pos_of_pos (abs_sub_pos.mpr (ne_toReal_exp_interp_exp₁ ht hq₀ hq₁ hq₀q₁ hq))) lemma ofReal_inv_interp_sub_exp_pos₀ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) - (hq₀q₁ : q₀ ≠ q₁) (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) : + (hq₀q₁ : q₀ ≠ q₁) (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) : ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹ > 0 := ofReal_pos.mpr (inv_pos_of_pos (abs_sub_pos.mpr (Ne.symm (ne_toReal_exp_interp_exp ht hq₀ hq₁ hq₀q₁ hq)))) lemma exp_lt_iff (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p < p₀ ↔ p₁ < p₀ := by rcases lt_or_gt_of_ne hp₀p₁ with p₀lt_p₁ | p₁lt_p₀ · exact ⟨fun h ↦ (not_le_of_gt h (le_of_lt (interp_exp_between hp₀ hp₁ p₀lt_p₁ ht hp).1)).elim, @@ -358,7 +348,7 @@ lemma exp_lt_iff (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp fun _ ↦ (interp_exp_between hp₁ hp₀ p₁lt_p₀ (Ioo.one_sub_mem ht) (switch_exponents ht hp)).2⟩ lemma exp_gt_iff (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p₀ < p ↔ p₀ < p₁ := by rcases lt_or_gt_of_ne hp₀p₁ with p₀lt_p₁ | p₁lt_p₀ · exact ⟨fun _ ↦ p₀lt_p₁, fun _ ↦ (interp_exp_between hp₀ hp₁ p₀lt_p₁ ht hp).1⟩ @@ -367,25 +357,25 @@ lemma exp_gt_iff (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp fun h ↦ (not_le_of_gt h (le_of_lt p₁lt_p₀)).elim⟩ lemma exp_lt_exp_gt_iff (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p < p₀ ↔ p₁ < p := by rw [exp_lt_iff ht hp₀ hp₁ hp₀p₁ hp, ← exp_gt_iff (Ioo.one_sub_mem ht) hp₁ hp₀ (Ne.symm hp₀p₁) (switch_exponents ht hp)] lemma exp_gt_exp_lt_iff (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p₀ < p ↔ p < p₁ := by rw [exp_gt_iff ht hp₀ hp₁ hp₀p₁ hp, ← exp_lt_iff (Ioo.one_sub_mem ht) hp₁ hp₀ (Ne.symm hp₀p₁) (switch_exponents ht hp)] lemma exp_lt_iff₁ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p < p₁ ↔ p₀ < p₁ := by rw [← exp_gt_exp_lt_iff ht hp₀ hp₁ hp₀p₁ hp] exact exp_gt_iff ht hp₀ hp₁ hp₀p₁ hp lemma exp_gt_iff₁ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) : p₁ < p ↔ p₁ < p₀ := by rw [← exp_lt_exp_gt_iff ht hp₀ hp₁ hp₀p₁ hp] exact exp_lt_iff ht hp₀ hp₁ hp₀p₁ hp @@ -424,7 +414,7 @@ lemma ζ_equality₁ (ht : t ∈ Ioo 0 1) : ((1 - t) * p₀⁻¹.toReal + t * p₁⁻¹.toReal - p₀⁻¹.toReal)) := by unfold ζ have t_pos : 0 < t := ht.1 - rw [← mul_div_mul_right _ _ t_pos.ne.symm, mul_assoc _ _ t, mul_assoc _ _ t] + rw [← mul_div_mul_right _ _ t_pos.ne', mul_assoc _ _ t, mul_assoc _ _ t] congr <;> ring lemma ζ_equality₂ (ht : t ∈ Ioo 0 1) : @@ -447,8 +437,8 @@ lemma ζ_symm : lemma ζ_equality₃ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) (hq₀' : q₀ ≠ ⊤) : @ζ p₀ q₀ p₁ q₁ t = (p₀.toReal * (q₀.toReal - q.toReal)) / (q₀.toReal * (p₀.toReal - p.toReal)) := by @@ -459,7 +449,7 @@ lemma ζ_equality₃ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) mul_pos (mul_pos (mul_pos (interp_exp_toReal_pos ht hp₀ hp₁ hp₀p₁ hp) (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq)) (exp_toReal_pos hp₀ hp₀')) (exp_toReal_pos hq₀ hq₀') - rw [← mul_div_mul_right _ _ hne.ne.symm] + rw [← mul_div_mul_right _ _ hne.ne'] have eq₁ : p⁻¹.toReal * (q⁻¹.toReal - q₀⁻¹.toReal) * (p.toReal * q.toReal * p₀.toReal * q₀.toReal) = p₀.toReal * (p⁻¹.toReal * p.toReal) * ((q⁻¹.toReal * q.toReal) * q₀.toReal - @@ -491,8 +481,8 @@ lemma coe_one_sub (ht : t ∈ Ioo 0 1) : lemma ζ_equality₄ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₁' : p₁ ≠ ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₁' : p₁ ≠ ⊤) (hq₁' : q₁ ≠ ⊤) : @ζ p₀ q₀ p₁ q₁ t = (p₁.toReal * (q₁.toReal - q.toReal)) / (q₁.toReal * (p₁.toReal - p.toReal)) := by @@ -506,8 +496,8 @@ lemma ζ_equality₄ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) lemma ζ_equality₅ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) (hq₀' : q₀ ≠ ⊤) : p₀.toReal + (@ζ p₀ q₀ p₁ q₁ t)⁻¹ * (q.toReal - q₀.toReal) * (p₀.toReal / q₀.toReal) = p.toReal := by @@ -521,13 +511,13 @@ lemma ζ_equality₅ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) rw [inv_mul_cancel₀, inv_mul_cancel₀, inv_mul_cancel₀] · simp only [one_mul, mul_one, _root_.sub_sub_cancel] · exact sub_ne_zero_of_ne (ne_toReal_exp_interp_exp ht hq₀ hq₁ hq₀q₁ hq) - · exact (exp_toReal_pos hp₀ hp₀').ne.symm - · exact (exp_toReal_pos hq₀ hq₀').ne.symm + · exact (exp_toReal_pos hp₀ hp₀').ne' + · exact (exp_toReal_pos hq₀ hq₀').ne' lemma ζ_equality₆ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₁' : p₁ ≠ ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₁' : p₁ ≠ ⊤) (hq₁' : q₁ ≠ ⊤) : p₁.toReal + (@ζ p₀ q₀ p₁ q₁ t)⁻¹ * (q.toReal - q₁.toReal) * (p₁.toReal / q₁.toReal) = p.toReal := by @@ -537,17 +527,17 @@ lemma ζ_equality₆ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) lemma ζ_equality₇ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) (hq₀' : q₀ = ⊤) : @ζ p₀ q₀ p₁ q₁ t = p₀.toReal / (p₀.toReal - p.toReal) := by rw [ζ_equality₁ ht, ← preservation_interpolation ht hp₀ hp₁ hp, ← preservation_interpolation ht hq₀ hq₁ hq, hq₀'] simp only [inv_top, zero_toReal, sub_zero, mul_zero, zero_add] have obs : p₀.toReal * p.toReal * q.toReal > 0 := - mul_pos (mul_pos (toReal_pos hp₀.ne.symm hp₀') (interp_exp_toReal_pos ht hp₀ hp₁ hp₀p₁ hp)) + mul_pos (mul_pos (toReal_pos hp₀.ne' hp₀') (interp_exp_toReal_pos ht hp₀ hp₁ hp₀p₁ hp)) (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq) - rw [← mul_div_mul_right _ _ obs.ne.symm] + rw [← mul_div_mul_right _ _ obs.ne'] congr · calc _ = (p.toReal⁻¹ * p.toReal) * (q.toReal⁻¹ * q.toReal) * p₀.toReal := by @@ -555,8 +545,8 @@ lemma ζ_equality₇ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) ring _ = _ := by rw [inv_mul_cancel₀, inv_mul_cancel₀, one_mul, one_mul] - · exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne.symm - · exact (interp_exp_toReal_pos ht hp₀ hp₁ hp₀p₁ hp).ne.symm + · exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne' + · exact (interp_exp_toReal_pos ht hp₀ hp₁ hp₀p₁ hp).ne' · calc _ = (q.toReal⁻¹ * q.toReal) * (p.toReal⁻¹ * p.toReal * p₀.toReal - p₀.toReal⁻¹ * p₀.toReal * p.toReal) := by @@ -564,14 +554,14 @@ lemma ζ_equality₇ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) ring _ = _ := by repeat rw [inv_mul_cancel₀, one_mul] - · exact (toReal_pos hp₀.ne.symm hp₀').ne.symm - · exact (interp_exp_toReal_pos ht hp₀ hp₁ hp₀p₁ hp).ne.symm - · exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne.symm + · exact (toReal_pos hp₀.ne' hp₀').ne' + · exact (interp_exp_toReal_pos ht hp₀ hp₁ hp₀p₁ hp).ne' + · exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne' lemma ζ_equality₈ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₁' : p₁ ≠ ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₁' : p₁ ≠ ⊤) (hq₁' : q₁ = ⊤) : @ζ p₀ q₀ p₁ q₁ t = p₁.toReal / (p₁.toReal - p.toReal) := ζ_symm ▸ ζ_equality₇ (Ioo.one_sub_mem ht) hp₁ hq₁ hp₀ hq₀ (Ne.symm hp₀p₁) (Ne.symm hq₀q₁) @@ -579,16 +569,16 @@ lemma ζ_equality₈ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) lemma ζ_eq_top_top (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₁' : p₁ = ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₁' : p₁ = ⊤) (hq₁' : q₁ = ⊤) : @ζ p₀ q₀ p₁ q₁ t = 1 := by rw [ζ_equality₂ ht, ← preservation_interpolation ht hp₀ hp₁ hp, ← preservation_interpolation ht hq₀ hq₁ hq, hp₁', hq₁'] simp only [inv_top, zero_toReal, sub_zero] rw [mul_comm, div_eq_mul_inv, mul_inv_cancel₀] - exact ne_of_gt (mul_pos (interp_exp_inv_pos ht hq₀ hq₁ hq₀q₁ hq) - (interp_exp_inv_pos ht hp₀ hp₁ hp₀p₁ hp)) + exact (mul_pos (interp_exp_inv_pos ht hq₀ hq₁ hq₀q₁ hq) + (interp_exp_inv_pos ht hp₀ hp₁ hp₀p₁ hp)).ne' lemma ζ_pos_iff_aux (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₀' : p₀ ≠ ⊤) (hq₀' : q₀ ≠ ⊤) : ( 0 < p₀.toReal * (q₀.toReal - q.toReal) / (q₀.toReal * (p₀.toReal - p.toReal))) ↔ @@ -605,23 +595,23 @@ lemma ζ_pos_iff_aux (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₀' : p₀ ≠ · exact exp_toReal_pos hp₀ hp₀' lemma preservation_inequality (ht : t ∈ Ioo 0 1) (hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) (hp₀' : p₀ ≠ ⊤) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) (hp₀' : p₀ ≠ ⊤) : p.toReal < p₀.toReal ↔ p < p₀ := toReal_lt_toReal (interp_exp_ne_top hp₀p₁ ht hp) hp₀' lemma preservation_inequality' (ht : t ∈ Ioo 0 1)(hp₀p₁ : p₀ ≠ p₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) (hp₀' : p₀ ≠ ⊤) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) (hp₀' : p₀ ≠ ⊤) : p₀.toReal < p.toReal ↔ p₀ < p := toReal_lt_toReal hp₀' (interp_exp_ne_top hp₀p₁ ht hp) lemma preservation_inequality_of_lt₀ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hq₀q₁ : q₀ < q₁) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hq₀q₁ : q₀ < q₁) : q₀.toReal < q.toReal := (toReal_lt_toReal (LT.lt.ne_top hq₀q₁) (interp_exp_ne_top (ne_of_lt hq₀q₁) ht hq)).mpr ((exp_gt_iff ht hq₀ hq₁ (ne_of_lt hq₀q₁) hq).mpr hq₀q₁) lemma preservation_inequality_of_lt₁ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hq₀q₁ : q₀ < q₁) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hq₀q₁ : q₀ < q₁) (hq₁' : q₁ ≠ ⊤): q.toReal < q₁.toReal := (toReal_lt_toReal (interp_exp_ne_top (ne_of_lt hq₀q₁) ht hq) hq₁').mpr @@ -629,8 +619,8 @@ lemma preservation_inequality_of_lt₁ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) ( lemma ζ_pos_toReal_iff₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) (hq₀' : q₀ ≠ ⊤) : (0 < @ζ p₀ q₀ p₁ q₁ t) ↔ ((q.toReal < q₀.toReal) ∧ (p.toReal < p₀.toReal)) ∨ ((q₀.toReal < q.toReal) ∧ (p₀.toReal < p.toReal)) := by @@ -639,8 +629,8 @@ lemma ζ_pos_toReal_iff₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ lemma ζ_pos_toReal_iff₁ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₁' : p₁ ≠ ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₁' : p₁ ≠ ⊤) (hq₁' : q₁ ≠ ⊤) : (0 < @ζ p₀ q₀ p₁ q₁ t) ↔ ((q.toReal < q₁.toReal) ∧ (p.toReal < p₁.toReal)) ∨ @@ -668,18 +658,6 @@ lemma inv_toReal_iff (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) : Iff.trans (toReal_lt_toReal (ne_of_lt (inv_lt_top.mpr hp₀)) (ne_of_lt (inv_lt_top.mpr hp₁))) ENNReal.inv_lt_inv --- TODO: check where this is used, replace by ζ_pos_iff' -lemma ζ_pos_iff₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) - (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) - (hq₀' : q₀ ≠ ⊤) : - (0 < @ζ p₀ q₀ p₁ q₁ t) ↔ - ((q < q₀) ∧ (p < p₀)) ∨ ((q₀ < q) ∧ (p₀ < p)) := by - rw [ζ_pos_toReal_iff₀ ht hp₀ hq₀ hp₁ hq₁ hp₀p₁ hq₀q₁ hp hq hp₀' hq₀', - preservation_inequality ht hq₀q₁ hq hq₀', preservation_inequality ht hp₀p₁ hp hp₀', - preservation_inequality' ht hq₀q₁ hq hq₀', preservation_inequality' ht hp₀p₁ hp hp₀'] - lemma ζ_pos_iff (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) : (0 < @ζ p₀ q₀ p₁ q₁ t) ↔ @@ -690,8 +668,8 @@ lemma ζ_pos_iff (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp lemma ζ_pos_iff' (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) : (0 < @ζ p₀ q₀ p₁ q₁ t) ↔ ((q < q₀) ∧ (p < p₀)) ∨ ((q₀ < q) ∧ (p₀ < p)) := by rw [ζ_pos_iff ht hp₀ hq₀ hp₁ hq₁ hp₀p₁ hq₀q₁, @@ -700,8 +678,8 @@ lemma ζ_pos_iff' (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) lemma ζ_pos_iff_of_lt₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : (0 < @ζ p₀ q₀ p₁ q₁ t) ↔ (q₀ < q) := by rw [ζ_pos_iff' ht hp₀ hq₀ hp₁ hq₁ (ne_of_lt hp₀p₁') hq₀q₁ hp hq] @@ -711,8 +689,8 @@ lemma ζ_pos_iff_of_lt₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ lemma ζ_pos_iff_of_lt₁ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : (0 < @ζ p₀ q₀ p₁ q₁ t) ↔ (q < q₁) := by rw [← exp_gt_exp_lt_iff ht hq₀ hq₁ hq₀q₁ hq] @@ -742,8 +720,8 @@ lemma ζ_neg_iff (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp lemma ζ_neg_iff' (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) : (@ζ p₀ q₀ p₁ q₁ t) < 0 ↔ ((q < q₀) ∧ (p₀ < p)) ∨ ((q₀ < q) ∧ (p < p₀)) := by rw [ζ_neg_iff ht hp₀ hq₀ hp₁ hq₁ hp₀p₁ hq₀q₁, ← exp_lt_iff ht hp₀ hp₁ hp₀p₁ hp, @@ -752,8 +730,8 @@ lemma ζ_neg_iff' (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp lemma ζ_neg_iff_of_lt₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : (@ζ p₀ q₀ p₁ q₁ t) < 0 ↔ q < q₀ := by rw [ζ_neg_iff' ht hp₀ hq₀ hp₁ hq₁ (ne_of_lt hp₀p₁') hq₀q₁ hp hq] rw [← exp_gt_iff (p := p) ht hp₀ hp₁ (ne_of_lt hp₀p₁') hp] at hp₀p₁' @@ -762,8 +740,8 @@ lemma ζ_neg_iff_of_lt₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ lemma ζ_neg_iff_of_lt₁ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : (@ζ p₀ q₀ p₁ q₁ t) < 0 ↔ q₁ < q := by rw [← exp_lt_exp_gt_iff ht hq₀ hq₁ hq₀q₁ hq] exact ζ_neg_iff_of_lt₀ ht hp₀ hq₀ hp₁ hq₁ hq₀q₁ hp hq hp₀p₁' @@ -784,8 +762,8 @@ lemma ζ_neg_iff_aux (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₀' : p₀ ≠ lemma ζ_neg_toReal_iff₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) (hq₀' : q₀ ≠ ⊤) : (@ζ p₀ q₀ p₁ q₁ t < 0) ↔ ((q.toReal < q₀.toReal) ∧ (p₀.toReal < p.toReal)) ∨ @@ -795,8 +773,8 @@ lemma ζ_neg_toReal_iff₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ lemma ζ_neg_toReal_iff₁ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₁' : p₁ ≠ ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₁' : p₁ ≠ ⊤) (hq₁' : q₁ ≠ ⊤) : (@ζ p₀ q₀ p₁ q₁ t < 0) ↔ ((q.toReal < q₁.toReal) ∧ (p₁.toReal < p.toReal)) ∨ ((q₁.toReal < q.toReal) ∧ @@ -806,8 +784,8 @@ lemma ζ_neg_toReal_iff₁ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ lemma ζ_neg_iff₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀' : p₀ ≠ ⊤) (hq₀' : q₀ ≠ ⊤) : (@ζ p₀ q₀ p₁ q₁ t < 0) ↔ ((q < q₀) ∧ (p₀ < p)) ∨ ((q₀ < q) ∧ (p < p₀)) := by rw [ζ_neg_toReal_iff₀ ht hp₀ hq₀ hp₁ hq₁ hp₀p₁ hq₀q₁ hp hq hp₀' hq₀', @@ -818,19 +796,19 @@ lemma ζ_ne_zero (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp (hp₀p₁ : p₀ ≠ p₁) (hq₀q₁ : q₀ ≠ q₁) : (@ζ p₀ q₀ p₁ q₁ t ≠ 0) := by refine div_ne_zero ?_ ?_ - · apply mul_ne_zero (ne_of_gt (preservation_positivity_inv_toReal ht hp₀ hp₁ hp₀p₁)) + · apply mul_ne_zero (preservation_positivity_inv_toReal ht hp₀ hp₁ hp₀p₁).ne' refine sub_ne_zero_of_ne (Ne.symm fun h ↦ hq₀q₁ ?_) rw [← inv_inv q₀, ← inv_inv q₁, ← coe_inv_exponent hq₀, ← coe_inv_exponent hq₁] exact congrArg Inv.inv (congrArg ENNReal.ofReal h) - · apply mul_ne_zero (ne_of_gt (preservation_positivity_inv_toReal ht hq₀ hq₁ hq₀q₁)) + · apply mul_ne_zero (preservation_positivity_inv_toReal ht hq₀ hq₁ hq₀q₁).ne' refine sub_ne_zero_of_ne (Ne.symm fun h ↦ hp₀p₁ ?_) rw [← inv_inv p₀, ← inv_inv p₁, ← coe_inv_exponent hp₀, ← coe_inv_exponent hp₁] exact congrArg Inv.inv (congrArg ENNReal.ofReal h) lemma ζ_le_zero_iff_of_lt₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : (@ζ p₀ q₀ p₁ q₁ t ≤ 0) ↔ q < q₀ := by constructor <;> intro h · rcases (Decidable.lt_or_eq_of_le h) with ζ_lt_0 | ζ_eq_0 @@ -840,14 +818,14 @@ lemma ζ_le_zero_iff_of_lt₀ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q lemma ζ_le_zero_iff_of_lt₁ (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hp₀p₁' : p₀ < p₁) : (@ζ p₀ q₀ p₁ q₁ t) ≤ 0 ↔ q₁ < q := by rw [← exp_lt_exp_gt_iff ht hq₀ hq₁ hq₀q₁ hq] exact ζ_le_zero_iff_of_lt₀ ht hp₀ hq₀ hp₁ hq₁ hq₀q₁ hp hq hp₀p₁' lemma eq_exponents₀ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hq₀' : q₀ ≠ ⊤) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hq₀' : q₀ ≠ ⊤) : (q₀.toReal + q₁⁻¹.toReal / (q₁⁻¹.toReal - q₀⁻¹.toReal) * (q.toReal - q₀.toReal)) = (1 - t) * q.toReal := by rw [mul_comm_div, ← mul_div_assoc, add_div'] @@ -860,8 +838,8 @@ lemma eq_exponents₀ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) _ = q₁⁻¹.toReal * q.toReal - q⁻¹.toReal * q.toReal := by rw [toReal_inv, toReal_inv, toReal_inv, mul_inv_cancel₀, inv_mul_cancel₀] · ring - · exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne.symm - · exact (toReal_pos hq₀.ne.symm hq₀').ne.symm + · exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne' + · exact (toReal_pos hq₀.ne' hq₀').ne' _ = q.toReal * (q₁⁻¹.toReal - q⁻¹.toReal) := by ring _ = _ := by rw [preservation_interpolation ht hq₀ hq₁ hq] @@ -873,7 +851,7 @@ lemma eq_exponents₀ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) · exact ne_sub_toReal_exp hq₀ hq₁ hq₀q₁ lemma eq_exponents₂ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hq₀' : q₀ ≠ ⊤) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hq₀' : q₀ ≠ ⊤) : (q₀.toReal / p₀.toReal + p₀⁻¹.toReal * q₁⁻¹.toReal / (q₁⁻¹.toReal - q₀⁻¹.toReal) * (q.toReal - q₀.toReal)) = (1 - t) * p₀⁻¹.toReal * q.toReal := by rw [div_eq_inv_mul] @@ -887,8 +865,8 @@ lemma eq_exponents₂ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) _ = q₁⁻¹.toReal * q.toReal - q⁻¹.toReal * q.toReal := by rw [toReal_inv, toReal_inv, toReal_inv, mul_inv_cancel₀, inv_mul_cancel₀] · ring - · exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne.symm - · exact (toReal_pos hq₀.ne.symm hq₀').ne.symm + · exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne' + · exact (toReal_pos hq₀.ne' hq₀').ne' _ = q.toReal * (q₁⁻¹.toReal - q⁻¹.toReal) := by ring _ = _ := by rw [preservation_interpolation ht hq₀ hq₁ hq] @@ -900,7 +878,7 @@ lemma eq_exponents₂ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) · exact ne_sub_toReal_exp hq₀ hq₁ hq₀q₁ lemma eq_exponents₁ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hq₀' : q₀ ≠ ⊤) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hq₀' : q₀ ≠ ⊤) : (q₀⁻¹.toReal / (q₁⁻¹.toReal - q₀⁻¹.toReal)) * (q.toReal - q₀.toReal) = - t * q.toReal := by rw [mul_comm_div, ← mul_div_assoc] have : q₀⁻¹.toReal * (q.toReal - q₀.toReal) = - t * q.toReal * (q₁⁻¹.toReal - q₀⁻¹.toReal) := by @@ -908,10 +886,10 @@ lemma eq_exponents₁ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) _ = (q₀⁻¹.toReal * q.toReal - q₀⁻¹.toReal * q₀.toReal) := by ring _ = (q₀⁻¹.toReal * q.toReal - 1) := by rw [toReal_inv, inv_mul_cancel₀] - exact (exp_toReal_pos hq₀ hq₀').ne.symm + exact (exp_toReal_pos hq₀ hq₀').ne' _ = (q₀⁻¹.toReal * q.toReal - q⁻¹.toReal * q.toReal) := by rw [toReal_inv, toReal_inv, inv_mul_cancel₀] - exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne.symm + exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne' _ = q.toReal * (q₀⁻¹.toReal - q⁻¹.toReal) := by ring _ = _ := by rw [preservation_interpolation ht hq₀ hq₁ hq] @@ -921,7 +899,7 @@ lemma eq_exponents₁ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) -- TODO: simplify these proofs with statements above lemma eq_exponents₃ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hq₀' : q₀ ≠ ⊤) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hq₀' : q₀ ≠ ⊤) : (p₁⁻¹.toReal * q₀⁻¹.toReal / (q₁⁻¹.toReal - q₀⁻¹.toReal)) * (q.toReal - q₀.toReal) = - t * p₁⁻¹.toReal * q.toReal := by rw [mul_comm_div, ← mul_div_assoc] @@ -931,10 +909,10 @@ lemma eq_exponents₃ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) _ = p₁⁻¹.toReal * (q₀⁻¹.toReal * q.toReal - q₀⁻¹.toReal * q₀.toReal) := by ring _ = p₁⁻¹.toReal * (q₀⁻¹.toReal * q.toReal - 1) := by rw [toReal_inv, toReal_inv, inv_mul_cancel₀] - apply (exp_toReal_pos hq₀ hq₀').ne.symm + apply (exp_toReal_pos hq₀ hq₀').ne' _ = p₁⁻¹.toReal * (q₀⁻¹.toReal * q.toReal - q⁻¹.toReal * q.toReal) := by rw [toReal_inv, toReal_inv, toReal_inv, inv_mul_cancel₀] - exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne.symm + exact (interp_exp_toReal_pos ht hq₀ hq₁ hq₀q₁ hq).ne' _ = p₁⁻¹.toReal * q.toReal * (q₀⁻¹.toReal - q⁻¹.toReal) := by ring _ = _ := by rw [preservation_interpolation ht hq₀ hq₁ hq] @@ -950,7 +928,7 @@ lemma eq_exponents₄ : _ = _ := by congr; rw [neg_inv, neg_sub] lemma eq_exponents₅ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hq₁' : q₁ ≠ ⊤): + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hq₁' : q₁ ≠ ⊤): (q₁.toReal + -(q₀⁻¹.toReal / (q₁⁻¹.toReal - q₀⁻¹.toReal) * (q.toReal - q₁.toReal))) = t * q.toReal := by rw [eq_exponents₄, neg_mul, neg_neg, eq_exponents₀ (Ioo.one_sub_mem ht) hq₁ hq₀ (Ne.symm hq₀q₁) @@ -958,14 +936,14 @@ lemma eq_exponents₅ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) ring lemma eq_exponents₆ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hq₁' : q₁ ≠ ⊤) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hq₁' : q₁ ≠ ⊤) : q₁⁻¹.toReal / (q₁⁻¹.toReal - q₀⁻¹.toReal) * (q.toReal - q₁.toReal) = (1 - t) * q.toReal := by rw [← neg_neg (a := q₁⁻¹.toReal / (q₁⁻¹.toReal - q₀⁻¹.toReal)), ← eq_exponents₄, neg_mul, eq_exponents₁ (Ioo.one_sub_mem ht) hq₁ hq₀ (Ne.symm hq₀q₁) (switch_exponents ht hq) hq₁'] ring lemma eq_exponents₇ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hq₁' : q₁ ≠ ⊤) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hq₁' : q₁ ≠ ⊤) : q₁.toReal / p₁.toReal + -(p₁⁻¹.toReal * q₀⁻¹.toReal / (q₁⁻¹.toReal - q₀⁻¹.toReal) * (q.toReal - q₁.toReal)) = t * p₁⁻¹.toReal * q.toReal := by @@ -978,7 +956,7 @@ lemma eq_exponents₇ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) ring lemma eq_exponents₈ (ht : t ∈ Ioo 0 1) (hq₀ : q₀ > 0) (hq₁ : q₁ > 0) (hq₀q₁ : q₀ ≠ q₁) - (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) (hq₁' : q₁ ≠ ⊤) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hq₁' : q₁ ≠ ⊤) : p₀⁻¹.toReal * q₁⁻¹.toReal / (q₁⁻¹.toReal - q₀⁻¹.toReal) * (q.toReal - q₁.toReal) = (1 - t) * p₀⁻¹.toReal * q.toReal := calc _ = p₀⁻¹.toReal * (q₁⁻¹.toReal / (q₁⁻¹.toReal - q₀⁻¹.toReal) * (q.toReal - q₁.toReal)) := by ring @@ -1023,10 +1001,10 @@ structure ToneCouple where else ∀ s ∈ Ioi (0 : ℝ), ∀ t ∈ Ioi (0 : ℝ), (ton s < t ↔ inv t < s) ∧ (t < ton s ↔ s < inv t) /-- A scaled power function gives rise to a ToneCouple. -/ -instance spf_to_tc (spf : ScaledPowerFunction) : ToneCouple where +def spf_to_tc (spf : ScaledPowerFunction) : ToneCouple where ton := fun s : ℝ ↦ (s / spf.d) ^ spf.σ inv := fun t : ℝ ↦ spf.d * t ^ spf.σ⁻¹ - mon := if (spf.σ > 0) then true else false + mon := if spf.σ > 0 then true else false ran_ton := fun t ht ↦ rpow_pos_of_pos (div_pos ht spf.hd) _ ran_inv := fun t ht ↦ mul_pos spf.hd (rpow_pos_of_pos ht spf.σ⁻¹) ton_is_ton := by @@ -1040,7 +1018,7 @@ instance spf_to_tc (spf : ScaledPowerFunction) : ToneCouple where · simp only [Bool.false_eq_true, ↓reduceIte] intro s (hs : s > 0) t (ht : t > 0) hst rcases spf.hσ with σ_pos | σ_neg - · exact False.elim (sgn_σ σ_pos) + · exact (sgn_σ σ_pos).elim · simp only exact (Real.rpow_lt_rpow_iff_of_neg (div_pos ht spf.hd) (div_pos hs spf.hd) σ_neg).mpr (div_lt_div_of_pos_right hst spf.hd) @@ -1050,18 +1028,18 @@ instance spf_to_tc (spf : ScaledPowerFunction) : ToneCouple where intro s hs t ht constructor · rw [← Real.lt_rpow_inv_iff_of_pos (div_nonneg (le_of_lt hs) (le_of_lt spf.hd)) - (le_of_lt ht) sgn_σ, ← _root_.mul_lt_mul_left spf.hd, mul_div_cancel₀ _ spf.hd.ne.symm] + (le_of_lt ht) sgn_σ, ← _root_.mul_lt_mul_left spf.hd, mul_div_cancel₀ _ spf.hd.ne'] · rw [← Real.rpow_inv_lt_iff_of_pos (le_of_lt ht) (div_nonneg (le_of_lt hs) (le_of_lt spf.hd)) - sgn_σ, ← _root_.mul_lt_mul_left spf.hd, mul_div_cancel₀ _ spf.hd.ne.symm] + sgn_σ, ← _root_.mul_lt_mul_left spf.hd, mul_div_cancel₀ _ spf.hd.ne'] · simp only [↓reduceIte, mem_Ioi] intro s hs t ht rcases spf.hσ with σ_pos | σ_neg · contradiction · constructor · rw [← Real.rpow_inv_lt_iff_of_neg ht (div_pos hs spf.hd) σ_neg, - ← _root_.mul_lt_mul_left spf.hd, mul_div_cancel₀ _ spf.hd.ne.symm] + ← _root_.mul_lt_mul_left spf.hd, mul_div_cancel₀ _ spf.hd.ne'] · rw [← Real.lt_rpow_inv_iff_of_neg (div_pos hs spf.hd) ht σ_neg, - ← _root_.mul_lt_mul_left spf.hd, mul_div_cancel₀ _ spf.hd.ne.symm] + ← _root_.mul_lt_mul_left spf.hd, mul_div_cancel₀ _ spf.hd.ne'] end @@ -1080,8 +1058,7 @@ variable {α α' E E₁ E₂ E₃ : Type*} {m : MeasurableSpace α} {m' : Measur [NormedAddCommGroup E₃] [MeasurableSpace E] [BorelSpace E] [MeasurableSpace E₃] [BorelSpace E₃] - -- (L : E₁ →L[ℝ] E₂ →L[ℝ] E₃) - {f : α → E₁} {t : ℝ} -- {s x y : ℝ≥0∞} + {f : α → E₁} {t : ℝ} {T : (α → E₁) → (α' → E₂)} /-! ## Results about the particular choice of scale @@ -1112,24 +1089,24 @@ lemma d_ne_top_aux₀ {b : ℝ} {F : ℝ≥0∞} (hF : F ∈ Ioo 0 ⊤) : F ^ b lemma d_ne_zero_aux₀ {b : ℝ} (hF : eLpNorm f p μ ∈ Ioo 0 ⊤) : (eLpNorm f p μ ^ p.toReal) ^ b ≠ 0 := - (ENNReal.rpow_pos (d_pos_aux₀ hF) (d_ne_top_aux₀ hF)).ne.symm + (ENNReal.rpow_pos (d_pos_aux₀ hF) (d_ne_top_aux₀ hF)).ne' lemma d_ne_zero_aux₁ {C : ℝ≥0} {c : ℝ} (hC : C > 0) : (ENNReal.ofNNReal C) ^ c ≠ 0 := - (ENNReal.rpow_pos (ENNReal.coe_pos.mpr hC) coe_ne_top).ne.symm + (ENNReal.rpow_pos (ENNReal.coe_pos.mpr hC) coe_ne_top).ne' lemma d_ne_zero_aux₂ {C : ℝ≥0} {b c : ℝ} (hC : C > 0) (hF : eLpNorm f p μ ∈ Ioo 0 ⊤) : C ^ c * (eLpNorm f p μ ^ p.toReal) ^ b ≠ 0 := - (ENNReal.mul_pos (d_ne_zero_aux₁ hC) (d_ne_zero_aux₀ hF)).ne.symm + (ENNReal.mul_pos (d_ne_zero_aux₁ hC) (d_ne_zero_aux₀ hF)).ne' lemma d_ne_top_aux₁ {C : ℝ≥0} {c : ℝ} (hC : C > 0) : (ENNReal.ofNNReal C) ^ c ≠ ⊤ := - rpow_ne_top' (ENNReal.coe_pos.mpr hC).ne.symm coe_ne_top + rpow_ne_top' (ENNReal.coe_pos.mpr hC).ne' coe_ne_top lemma d_ne_top_aux₂ {b : ℝ} (hF : eLpNorm f p μ ∈ Ioo 0 ⊤) : (eLpNorm f p μ ^ p.toReal) ^ b ≠ ⊤ := - rpow_ne_top' (d_pos_aux₀ hF).ne.symm (d_ne_top_aux₀ hF) + rpow_ne_top' (d_pos_aux₀ hF).ne' (d_ne_top_aux₀ hF) lemma d_ne_top_aux₃ {C : ℝ≥0} {b c : ℝ} (hC : C > 0) (hF : eLpNorm f p μ ∈ Ioo 0 ⊤) : @@ -1166,10 +1143,10 @@ lemma d_eq_top₀ (hp₀ : p₀ > 0) (hq₁ : q₁ > 0) (hp₀' : p₀ ≠ ⊤) · rw [div_eq_mul_inv, mul_inv_cancel₀, ENNReal.rpow_one] · rw [toReal_rpow, ENNReal.mul_rpow_of_nonneg] · rw [ENNReal.rpow_rpow_inv, toReal_inv] - exact (exp_toReal_pos hp₀ hp₀').ne.symm + exact (exp_toReal_pos hp₀ hp₀').ne' · positivity - · exact (inv_toReal_pos_of_ne_top hq₁ ((hq₀' ▸ hq₀q₁).symm)).ne.symm - · exact (inv_toReal_pos_of_ne_top hq₁ ((hq₀' ▸ hq₀q₁).symm)).ne.symm + · exact (inv_toReal_pos_of_ne_top hq₁ ((hq₀' ▸ hq₀q₁).symm)).ne' + · exact (inv_toReal_pos_of_ne_top hq₁ ((hq₀' ▸ hq₀q₁).symm)).ne' lemma d_eq_top₁ (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hp₁' : p₁ ≠ ⊤) (hq₁' : q₁ = ⊤) (hq₀q₁ : q₀ ≠ q₁) (hC₁ : C₁ > 0) : @@ -1187,11 +1164,11 @@ lemma d_eq_top₁ (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hp₁' : p₁ ≠ ⊤) · rw [← ENNReal.rpow_neg_one, ← ENNReal.rpow_mul, toReal_inv, mul_neg, mul_one, neg_neg, toReal_mul, coe_toReal] · left; exact ENNReal.inv_ne_zero.mpr coe_ne_top - · left; exact inv_ne_top.mpr <| (ENNReal.coe_pos.mpr hC₁).ne.symm - · exact (exp_toReal_pos hp₁ hp₁').ne.symm + · left; exact inv_ne_top.mpr <| (ENNReal.coe_pos.mpr hC₁).ne' + · exact (exp_toReal_pos hp₁ hp₁').ne' · positivity - · exact (inv_toReal_pos_of_ne_top hq₀ (hq₁' ▸ hq₀q₁)).ne.symm - · exact (inv_toReal_pos_of_ne_top hq₀ (hq₁' ▸ hq₀q₁)).ne.symm + · exact (inv_toReal_pos_of_ne_top hq₀ (hq₁' ▸ hq₀q₁)).ne' + · exact (inv_toReal_pos_of_ne_top hq₀ (hq₁' ▸ hq₀q₁)).ne' lemma d_eq_top_of_eq (hC₁ : C₁ > 0) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hq₀' : q₀ ≠ ⊤) (hp₀': p₀ ≠ ⊤) (hp₁ : p₁ > 0) (hp₀p₁ : p₀ = p₁) (hpp₀: p = p₀) (hq₁' : q₁ = ⊤) : @@ -1212,8 +1189,8 @@ lemma d_eq_top_top (hq₀ : q₀ > 0) (hq₀q₁ : q₀ ≠ q₁) (hp₁' : p₁ zero_mul, one_div] rw [div_neg, div_eq_mul_inv, mul_inv_cancel₀] · rw [ENNReal.rpow_neg, ENNReal.rpow_one, inv_inv, coe_toReal] - · exact ne_of_gt (toReal_pos (ENNReal.inv_ne_zero.mpr (hq₁' ▸ hq₀q₁)) - (ENNReal.inv_ne_top.mpr hq₀.ne.symm)) + · exact (toReal_pos (ENNReal.inv_ne_zero.mpr (hq₁' ▸ hq₀q₁)) + (ENNReal.inv_ne_top.mpr hq₀.ne')).ne' /-- The particular choice of scaled power function that works in the proof of the real interpolation theorem. -/ @@ -1315,9 +1292,8 @@ lemma lintegral_scale_constant' {f: ℝ → ENNReal} {a : ℝ} (h : a ≠ 0): -- local convenience function lemma lintegral_rw_aux {g : ℝ → ℝ≥0∞} {f₁ f₂ : ℝ → ℝ≥0∞} {A : Set ℝ} (heq : f₁ =ᶠ[ae (volume.restrict A)] f₂) : - (∫⁻ s in A, g s * f₁ s) = - (∫⁻ s in A, g s * f₂ s) := -(lintegral_rw₂ (Filter.EventuallyEq.refl (ae (volume.restrict A)) g) heq HMul.hMul) + ∫⁻ s in A, g s * f₁ s = ∫⁻ s in A, g s * f₂ s := + (lintegral_rw₂ (Filter.EventuallyEq.refl (ae (volume.restrict A)) g) heq HMul.hMul) lemma power_aux {p q : ℝ} : (fun s ↦ ENNReal.ofReal s ^ (p + q)) =ᶠ[ae (volume.restrict (Ioi (0 : ℝ)))] @@ -1340,15 +1316,10 @@ lemma ofReal_rpow_of_pos_aux {p : ℝ} : lemma extract_constant_double_integral_rpow {f : ℝ → ℝ → ℝ≥0∞} {q : ℝ} (hq : q ≥ 0) {a : ℝ≥0∞} (ha : a ≠ ⊤): - ∫⁻ (s : ℝ) in Ioi 0, (∫⁻ (t : ℝ) in Ioi 0, a * f s t)^q = - (a ^ q) * ∫⁻ (s : ℝ) in Ioi 0, (∫⁻ (t : ℝ) in Ioi 0, f s t)^q := by - rw [← lintegral_const_mul']; swap; exact rpow_ne_top_of_nonneg hq ha - apply congr - · rfl - · ext s - rw [← ENNReal.mul_rpow_of_nonneg _ _ hq] - congr - rw [lintegral_const_mul' a (f s) ha] + ∫⁻ (s : ℝ) in Ioi 0, (∫⁻ (t : ℝ) in Ioi 0, a * f s t) ^ q = + a ^ q * ∫⁻ (s : ℝ) in Ioi 0, (∫⁻ (t : ℝ) in Ioi 0, f s t) ^ q := by + simp_rw [← lintegral_const_mul' _ _ (rpow_ne_top_of_nonneg hq ha), + ← ENNReal.mul_rpow_of_nonneg _ _ hq, lintegral_const_mul' a _ ha] lemma ofReal_rpow_rpow_aux {p : ℝ} : (fun s ↦ ENNReal.ofReal s ^ p) =ᶠ[ae (volume.restrict (Ioi (0 : ℝ)))] @@ -1362,7 +1333,7 @@ lemma lintegral_rpow_of_gt {β γ : ℝ} (hβ : β > 0) (hγ : γ > -1) : have hγ2 : γ + 1 > 0 := by linarith rw [setLIntegral_congr Ioo_ae_eq_Ioc, ← ofReal_integral_eq_lintegral_ofReal] · rw [← intervalIntegral.integral_of_le (le_of_lt hβ), integral_rpow] - · rw [Real.zero_rpow hγ2.ne.symm, sub_zero] + · rw [Real.zero_rpow hγ2.ne', sub_zero] · exact Or.inl hγ · apply (@intervalIntegral.intervalIntegrable_rpow' 0 β γ ?_).1 linarith @@ -1380,7 +1351,7 @@ open NNReal ENNReal MeasureTheory Set ComputationsInterpolatedExponents variable {α α' E E₁ E₂ E₃ : Type*} {m : MeasurableSpace α} {m' : MeasurableSpace α'} {p p' q p₀ q₀ p₁ q₁: ℝ≥0∞} {c : ℝ≥0} {a : ℝ} {μ : Measure α} {ν : Measure α'} - {f : α → E₁} {t : ℝ} -- {s x y : ℝ≥0∞} + {f : α → E₁} {t : ℝ} {T : (α → E₁) → (α' → E₂)} /-! ## Results about truncations of a function @@ -1393,16 +1364,10 @@ def trunc [NormedAddCommGroup E₁] (f : α → E₁) (t : ℝ) (x : α) : E₁ /-- The complement of a `t`-truncatoin of a function `f`. -/ def trunc_compl [NormedAddCommGroup E₁] (f : α → E₁) (t : ℝ) : α → E₁ := f - trunc f t -lemma trunc_compl_eq [NormedAddCommGroup E₁] {a : ℝ} {f : α → E₁} [NormedAddCommGroup E₁] : - (f - trunc f a) = fun x ↦ if a < ‖f x‖ then f x else 0 := by +lemma trunc_compl_eq [NormedAddCommGroup E₁] {a : ℝ} {f : α → E₁} : + f - trunc f a = fun x ↦ if a < ‖f x‖ then f x else 0 := by ext x - unfold trunc - simp only [Pi.sub_apply] - split_ifs with h h₁ - · contrapose! h; assumption - · rw [sub_self] - · rw [sub_zero] - · contrapose! h; linarith + simp_rw [Pi.sub_apply, trunc, ← not_lt, ite_not, apply_ite (f x - ·), sub_zero, sub_self] /-- A function to deal with truncations and complement of truncations in one go. -/ def trnc [NormedAddCommGroup E₁] (j : Bool) (f : α → E₁) (t : ℝ) : α → E₁ := @@ -1443,73 +1408,66 @@ lemma trunc_compl_of_nonpos {f : α → E₁} {a : ℝ} [NormedAddCommGroup E₁ /-! ## Measurability properties of truncations -/ -@[measurability, fun_prop] -lemma measurable_trunc [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] - (hf : Measurable f) : - Measurable (trunc f t) := by - apply Measurable.ite - · apply measurableSet_le <;> fun_prop - · exact hf - · exact measurable_const +-- @[measurability, fun_prop] +-- lemma _root_.Measurable.trunc [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] +-- (hf : Measurable f) : Measurable (trunc f t) := by +-- refine hf.ite (measurableSet_le ?_ ?_) measurable_const <;> fun_prop -@[measurability, fun_prop] -lemma measurable_trunc_compl [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] - (hf : Measurable f) : - Measurable (f - trunc f t) := by - rw [trunc_compl_eq] - apply Measurable.ite ?_ hf measurable_const - exact measurableSet_lt measurable_const hf.norm +-- @[measurability, fun_prop] +-- lemma _root_.Measurable.trunc_compl [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] +-- (hf : Measurable f) : +-- Measurable (f - trunc f t) := by +-- rw [trunc_compl_eq] +-- apply Measurable.ite ?_ hf measurable_const +-- exact measurableSet_lt measurable_const hf.norm @[measurability] -lemma stronglyMeasurable_trunc [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] - (hf : StronglyMeasurable f) : - StronglyMeasurable (trunc f t) := by +protected lemma StronglyMeasurable.trunc [NormedAddCommGroup E₁] + (hf : StronglyMeasurable f) : StronglyMeasurable (trunc f t) := by apply StronglyMeasurable.ite ?_ hf stronglyMeasurable_const - exact measurableSet_le hf.measurable.norm measurable_const + exact measurableSet_le hf.norm stronglyMeasurable_const @[measurability] -lemma stronglyMeasurable_trunc_compl [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] - (hf : StronglyMeasurable f) : - StronglyMeasurable (f - trunc f t) := by +protected lemma StronglyMeasurable.trunc_compl [NormedAddCommGroup E₁] + (hf : StronglyMeasurable f) : StronglyMeasurable (f - trunc f t) := by rw [trunc_compl_eq] - apply StronglyMeasurable.ite ?_ hf stronglyMeasurable_const - exact measurableSet_lt measurable_const hf.measurable.norm - -@[measurability, fun_prop] -lemma aeMeasurable_trunc [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] - (hf : AEMeasurable f μ) : - AEMeasurable (trunc f t) μ := by - rcases hf with ⟨g, ⟨wg1, wg2⟩⟩ - exists (trunc g t) - constructor - · apply wg1.indicator (s := {x | ‖g x‖ ≤ t}) (measurableSet_le wg1.norm measurable_const) - apply measure_mono_null ?_ wg2 - intro x - contrapose - simp only [mem_compl_iff, mem_setOf_eq, not_not] - intro h₂ - unfold trunc - rewrite [h₂] - rfl - -@[measurability, fun_prop] -lemma aeMeasurable_trunc_compl [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] - (hf : AEMeasurable f μ) : - AEMeasurable (trunc_compl f t) μ := by - rcases hf with ⟨g, ⟨wg1, wg2⟩⟩ - exists (trunc_compl g t) - constructor - · unfold trunc_compl - rw [trunc_compl_eq] - exact wg1.indicator (s := {x | t < ‖g x‖}) (measurableSet_lt measurable_const wg1.norm) - · apply measure_mono_null ?_ wg2 - intro x - contrapose - simp only [mem_compl_iff, mem_setOf_eq, not_not] - intro f_eq_g; unfold trunc_compl; unfold trunc; dsimp only [Pi.sub_apply]; rw [f_eq_g] + exact hf.ite (measurableSet_lt stronglyMeasurable_const hf.norm) stronglyMeasurable_const + +-- @[measurability, fun_prop] +-- lemma aemeasurable_trunc [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] +-- (hf : AEMeasurable f μ) : +-- AEMeasurable (trunc f t) μ := by +-- rcases hf with ⟨g, ⟨wg1, wg2⟩⟩ +-- exists (trunc g t) +-- constructor +-- · apply wg1.indicator (s := {x | ‖g x‖ ≤ t}) (measurableSet_le wg1.norm measurable_const) +-- apply measure_mono_null ?_ wg2 +-- intro x +-- contrapose +-- simp only [mem_compl_iff, mem_setOf_eq, not_not] +-- intro h₂ +-- unfold trunc +-- rewrite [h₂] +-- rfl + +-- @[measurability, fun_prop] +-- lemma aeMeasurable_trunc_compl [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] +-- (hf : AEMeasurable f μ) : +-- AEMeasurable (trunc_compl f t) μ := by +-- rcases hf with ⟨g, ⟨wg1, wg2⟩⟩ +-- exists (trunc_compl g t) +-- constructor +-- · unfold trunc_compl +-- rw [trunc_compl_eq] +-- exact wg1.indicator (s := {x | t < ‖g x‖}) (measurableSet_lt measurable_const wg1.norm) +-- · apply measure_mono_null ?_ wg2 +-- intro x +-- contrapose +-- simp only [mem_compl_iff, mem_setOf_eq, not_not] +-- intro f_eq_g; unfold trunc_compl; unfold trunc; dsimp only [Pi.sub_apply]; rw [f_eq_g] @[measurability] -lemma aestronglyMeasurable_trunc [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] +lemma aestronglyMeasurable_trunc [NormedAddCommGroup E₁] (hf : AEStronglyMeasurable f μ) : AEStronglyMeasurable (trunc f t) μ := by rcases hf with ⟨g, ⟨wg1, wg2⟩⟩ @@ -1527,7 +1485,7 @@ lemma aestronglyMeasurable_trunc [MeasurableSpace E₁] [NormedAddCommGroup E₁ rfl @[measurability] -lemma aestronglyMeasurable_trunc_compl [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] +lemma aestronglyMeasurable_trunc_compl [NormedAddCommGroup E₁] (hf : AEStronglyMeasurable f μ) : AEStronglyMeasurable (f - trunc f t) μ := by rcases hf with ⟨g, ⟨wg1, wg2⟩⟩ @@ -1546,8 +1504,7 @@ lemma aestronglyMeasurable_trunc_compl [MeasurableSpace E₁] [NormedAddCommGrou rw [h₂] @[measurability] -lemma aestronglyMeasurable_trnc {j : Bool} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] +lemma aestronglyMeasurable_trnc {j : Bool} [NormedAddCommGroup E₁] (hf : AEStronglyMeasurable f μ) : AEStronglyMeasurable (trnc j f t) μ := by rcases j @@ -1566,20 +1523,16 @@ lemma trunc_le {f : α → E₁} {a : ℝ} [NormedAddCommGroup E₁] (x : α) : /-- A small lemma that is helpful for rewriting -/ lemma coe_coe_eq_ofReal (a : ℝ) : ofNNReal a.toNNReal = ENNReal.ofReal a := by rfl -lemma trunc_eLpNormEssSup_le {f : α → E₁} {a : ℝ} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] : - eLpNormEssSup (trunc f a) μ ≤ - ENNReal.ofReal (max 0 a) := by +lemma trunc_eLpNormEssSup_le {f : α → E₁} {a : ℝ} [NormedAddCommGroup E₁] : + eLpNormEssSup (trunc f a) μ ≤ ENNReal.ofReal (max 0 a) := by apply essSup_le_of_ae_le apply ae_of_all intro x simp only [← norm_toNNReal, coe_coe_eq_ofReal] exact ofReal_le_ofReal (trunc_le x) -lemma trunc_mono {f : α → E₁} {a b : ℝ} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] - (hab : a ≤ b) {x : α} : - ‖trunc f a x‖ ≤ ‖trunc f b x‖ := by +lemma trunc_mono {f : α → E₁} {a b : ℝ} [NormedAddCommGroup E₁] + (hab : a ≤ b) {x : α} : ‖trunc f a x‖ ≤ ‖trunc f b x‖ := by unfold trunc split_ifs · rfl @@ -1588,28 +1541,23 @@ lemma trunc_mono {f : α → E₁} {a b : ℝ} · exact le_refl _ /-- The norm of the truncation is monotone in the truncation parameter -/ -lemma norm_trunc_mono {f : α → E₁} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] : +lemma norm_trunc_mono {f : α → E₁} [NormedAddCommGroup E₁] : Monotone fun s ↦ eLpNorm (trunc f s) p μ := by intros a b hab; apply eLpNorm_mono; apply trunc_mono; exact hab -lemma trunc_buildup_norm {f : α → E₁} {a : ℝ} {x : α} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] : +lemma trunc_buildup_norm {f : α → E₁} {a : ℝ} {x : α} [NormedAddCommGroup E₁] : ‖trunc f a x‖ + ‖(f - trunc f a) x‖ = ‖f x‖ := by - unfold trunc; simp only [Pi.sub_apply]; split_ifs with h <;> simp + simp only [trunc, Pi.sub_apply]; split_ifs with h <;> simp -lemma trunc_le_func {f : α → E₁} {a : ℝ} {x : α} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] : ‖trunc f a x‖ ≤ ‖f x‖ := by +lemma trunc_le_func {f : α → E₁} {a : ℝ} {x : α} [NormedAddCommGroup E₁] : + ‖trunc f a x‖ ≤ ‖f x‖ := by unfold trunc; split_ifs <;> simp -lemma trunc_compl_le_func {f : α → E₁} {a : ℝ} {x : α} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] : ‖(f - trunc f a) x‖ ≤ ‖f x‖ - := by +lemma trunc_compl_le_func {f : α → E₁} {a : ℝ} {x : α} [NormedAddCommGroup E₁] : + ‖(f - trunc f a) x‖ ≤ ‖f x‖ := by rw [trunc_compl_eq]; dsimp only; split_ifs <;> simp -lemma trunc_compl_anti {f : α → E₁} {a b : ℝ} {x : α} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] - (hab : a ≤ b) : +lemma trunc_compl_anti {f : α → E₁} {a b : ℝ} {x : α} [NormedAddCommGroup E₁] (hab : a ≤ b) : ‖(f - trunc f b) x‖ ≤ ‖(f - trunc f a) x‖ := by have obs : ‖trunc f a x‖ + ‖(f - trunc f a) x‖ = ‖trunc f b x‖ + ‖(f - trunc f b) x‖ := by rw [trunc_buildup_norm, trunc_buildup_norm] @@ -1617,25 +1565,24 @@ lemma trunc_compl_anti {f : α → E₁} {a b : ℝ} {x : α} linarith /-- The norm of the complement of the truncation is antitone in the truncation parameter -/ -lemma norm_trunc_compl_anti {f : α → E₁} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] : +lemma norm_trunc_compl_anti {f : α → E₁} [NormedAddCommGroup E₁] : Antitone (fun s ↦ eLpNorm (f - trunc f s) p μ) := fun _a _b hab ↦ eLpNorm_mono (fun _ ↦ trunc_compl_anti hab) /-- The norm of the truncation is meaurable in the truncation parameter -/ @[measurability, fun_prop] -lemma norm_trunc_measurable [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] : +lemma norm_trunc_measurable [NormedAddCommGroup E₁] : Measurable (fun s ↦ eLpNorm (trunc f s) p μ) := norm_trunc_mono.measurable /-- The norm of the complement of the truncation is measurable in the truncation parameter -/ @[measurability, fun_prop] -lemma norm_trunc_compl_measurable [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] : +lemma norm_trunc_compl_measurable [NormedAddCommGroup E₁] : Measurable (fun s ↦ eLpNorm (f - trunc f s) p μ) := norm_trunc_compl_anti.measurable -lemma trnc_le_func {j : Bool} {f : α → E₁} {a : ℝ} {x : α} - [NormedAddCommGroup E₁] : ‖trnc j f a x‖ ≤ ‖f x‖ := by +lemma trnc_le_func {j : Bool} {f : α → E₁} {a : ℝ} {x : α} [NormedAddCommGroup E₁] : + ‖trnc j f a x‖ ≤ ‖f x‖ := by unfold trnc trunc rcases j · dsimp only [Pi.sub_apply] @@ -1673,8 +1620,7 @@ lemma rpow_le_rpow_of_exponent_le_base_le {a b t γ : ℝ} (ht : t > 0) (htγ : have γ_pos : 0 < γ := lt_of_lt_of_le ht htγ rw [Real.rpow_sub γ_pos] refine (ENNReal.mul_le_mul_left (a := ENNReal.ofReal (γ ^ (-b) )) ?_ ?_).mp ?_ - · apply ne_of_gt - refine ofReal_pos.mpr ?_ + · refine ofReal_pos.mpr ?_ |>.ne' exact Real.rpow_pos_of_pos γ_pos (-b) · exact coe_ne_top · rw [← ofReal_mul, ← mul_assoc, ← ofReal_mul, ← mul_div_assoc, ← Real.rpow_add, neg_add_cancel, @@ -1696,7 +1642,7 @@ lemma rpow_le_rpow_of_exponent_le_base_ge {a b t γ : ℝ} (hγ : γ > 0) (htγ have t_pos : 0 < t := gt_of_ge_of_gt htγ hγ rw [Real.rpow_sub hγ] refine (ENNReal.mul_le_mul_left (a := ENNReal.ofReal (γ ^ (-a) )) ?_ coe_ne_top).mp ?_ - · exact (ofReal_pos.mpr (Real.rpow_pos_of_pos hγ (-a))).ne.symm + · exact (ofReal_pos.mpr (Real.rpow_pos_of_pos hγ (-a))).ne' · rw [← ofReal_mul, ← mul_assoc, ← ofReal_mul, ← mul_div_assoc, ← Real.rpow_add, neg_add_cancel, Real.rpow_zero, ← ofReal_mul, mul_comm] <;> try positivity nth_rw 2 [mul_comm] @@ -1709,8 +1655,7 @@ lemma rpow_le_rpow_of_exponent_le_base_ge {a b t γ : ℝ} (hγ : γ > 0) (htγ apply ofReal_le_ofReal exact Real.rpow_le_rpow_of_exponent_le ((one_le_div hγ).mpr htγ) hab -lemma trunc_preserves_Lp {p : ℝ≥0∞} {a : ℝ} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] +lemma trunc_preserves_Lp {p : ℝ≥0∞} {a : ℝ} [NormedAddCommGroup E₁] (hf : Memℒp f p μ) : Memℒp (trunc f a) p μ := by refine ⟨aestronglyMeasurable_trunc hf.1, ?_⟩ @@ -1721,15 +1666,11 @@ lemma trunc_preserves_Lp {p : ℝ≥0∞} {a : ℝ} unfold trunc split_ifs with is_fx_le_a <;> simp -lemma snorm_trunc_compl_le {p : ℝ≥0∞} {a : ℝ} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] : - eLpNorm (f - trunc f a) p μ ≤ - eLpNorm f p μ := - eLpNorm_mono (fun _ ↦ trunc_compl_le_func) +-- lemma snorm_trunc_compl_le {p : ℝ≥0∞} {a : ℝ} [NormedAddCommGroup E₁] : +-- eLpNorm (f - trunc f a) p μ ≤ eLpNorm f p μ := +-- eLpNorm_mono (fun _ ↦ trunc_compl_le_func) -lemma trunc_compl_preserves_Lp {p : ℝ≥0∞} {a : ℝ} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] - (hf : Memℒp f p μ) : +lemma trunc_compl_preserves_Lp {p : ℝ≥0∞} {a : ℝ} [NormedAddCommGroup E₁] (hf : Memℒp f p μ) : Memℒp (f - trunc f a) p μ := Memℒp.sub hf (trunc_preserves_Lp hf) @@ -1740,8 +1681,8 @@ lemma estimate_eLpNorm_trunc_compl {p q : ℝ≥0∞} [MeasurableSpace E₁] [No eLpNorm f p μ ^ p.toReal := by unfold eLpNorm eLpNorm' have q_ne_top: q ≠ ⊤ := LT.lt.ne_top hpq.2 - have p_ne_zero : p ≠ 0 := (lt_trans hpq.1 hpq.2).ne.symm - have q_ne_zero : q ≠ 0 := hpq.1.ne.symm + have p_ne_zero : p ≠ 0 := (lt_trans hpq.1 hpq.2).ne' + have q_ne_zero : q ≠ 0 := hpq.1.ne' have q_toReal_pos : q.toReal > 0 := exp_toReal_pos hpq.1 q_ne_top split_ifs · contradiction @@ -1797,8 +1738,8 @@ lemma estimate_eLpNorm_trunc {p q : ℝ≥0∞} unfold eLpNorm eLpNorm' have : q ≠ ⊤ := hq have p_ne_top : p ≠ ⊤ := LT.lt.ne_top hpq.2 - have : p ≠ 0 := hpq.1.ne.symm - have : q ≠ 0 := (lt_trans hpq.1 hpq.2).ne.symm + have : p ≠ 0 := hpq.1.ne' + have : q ≠ 0 := (lt_trans hpq.1 hpq.2).ne' have q_toReal_pos : q.toReal > 0 := toReal_pos this hq split_ifs · contradiction @@ -1855,13 +1796,13 @@ lemma trunc_Lp_Lq_higher [MeasurableSpace E₁] [NormedAddCommGroup E₁] [Borel rcases (eq_or_ne q ⊤) with q_eq_top | q_ne_top · rw [q_eq_top, eLpNorm_exponent_top] exact Trans.trans trunc_eLpNormEssSup_le coe_lt_top - · have : q.toReal > 0 := toReal_pos (lt_trans hpq.1 hpq.2).ne.symm q_ne_top + · have : q.toReal > 0 := toReal_pos (lt_trans hpq.1 hpq.2).ne' q_ne_top refine (rpow_lt_top_iff_of_pos this).mp ?_ have := (estimate_eLpNorm_trunc (a := a) q_ne_top hpq (AEStronglyMeasurable.aemeasurable hf.1)) refine lt_of_le_of_lt this ?_ apply mul_lt_top coe_lt_top refine (rpow_lt_top_iff_of_pos ?_).mpr hf.2 - exact toReal_pos hpq.1.ne.symm hpq.2.ne_top + exact toReal_pos hpq.1.ne' hpq.2.ne_top /-- If `f` is in `Lp`, the complement of the truncation is in `Lq` for `q < p`. -/ lemma trunc_compl_Lp_Lq_lower [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] @@ -1869,13 +1810,13 @@ lemma trunc_compl_Lp_Lq_lower [MeasurableSpace E₁] [NormedAddCommGroup E₁] [ (hpq : q ∈ Ioo 0 p) (ha : a > 0) (hf : Memℒp f p μ) : Memℒp (trnc ⊥ f a) q μ := by refine ⟨aestronglyMeasurable_trnc hf.1, ?_⟩ - have : q.toReal > 0 := toReal_pos hpq.left.ne.symm (LT.lt.ne_top hpq.right) + have : q.toReal > 0 := toReal_pos hpq.left.ne' (LT.lt.ne_top hpq.right) refine (rpow_lt_top_iff_of_pos this).mp ?_ refine lt_of_le_of_lt (estimate_eLpNorm_trunc_compl hp hpq hf.1.aemeasurable ha) ?_ apply mul_lt_top coe_lt_top refine (rpow_lt_top_iff_of_pos ?_).mpr hf.2 - exact toReal_pos (lt_trans hpq.left hpq.right).ne.symm hp + exact toReal_pos (lt_trans hpq.left hpq.right).ne' hp end MeasureTheory @@ -1893,8 +1834,7 @@ variable {α α' E E₁ E₂ E₃ : Type*} {m : MeasurableSpace α} {m' : Measur [NormedAddCommGroup E₁] [NormedAddCommGroup E₂] [MeasurableSpace E] [BorelSpace E] - -- (L : E₁ →L[ℝ] E₂ →L[ℝ] E₃) - {f : α → E₁} {t : ℝ} -- {s x y : ℝ≥0∞} + {f : α → E₁} {t : ℝ} {T : (α → E₁) → (α' → E₂)} /-! ## Some results about the integrals of truncations @@ -1977,8 +1917,7 @@ lemma lintegral_trunc_mul₀ {g : ℝ → ℝ≥0∞} {j : Bool} {x : α} {tc : · simp [hp] · have := (mon_pf s (lt_trans (tc.ran_inv ‖f x‖ hfx) hs) (‖f x‖) hfx).2.mpr hs contrapose! h; linarith - · simp only [bne_self_eq_false, Bool.false_eq_true, not_false_eq_true, decide_True, - ↓reduceIte] + · simp only [bne_self_eq_false, Bool.false_eq_true, not_false_eq_true, decide_True] intro hs split_ifs with h · have := (mon_pf s hs.1 (‖f x‖) hfx).1.mpr hs.2 @@ -2135,8 +2074,8 @@ variable {α α' E E₁ E₂ E₃ : Type*} {m : MeasurableSpace α} {m' : Measur {a : ℝ} {f : α → E₁} {t : ℝ} {T : (α → E₁) → (α' → E₂)} -/-! ## Minkowski's integral inequality --/ + +/-! ## Minkowski's integral inequality -/ namespace MeasureTheory lemma rpow_add_of_pos (a : ℝ≥0∞) (c d : ℝ) (hc : c > 0) (hd : d > 0): @@ -2392,7 +2331,7 @@ lemma representationLp {μ : Measure α} [SigmaFinite μ] {f : α → ℝ≥0∞ lemma aemeasurability_prod₁ {α : Type u_1} {β : Type u_3} [MeasurableSpace α] [MeasurableSpace β] {μ : Measure α} {ν : Measure β} [SFinite ν] - [SFinite μ] ⦃f : α × β → ENNReal⦄ + ⦃f : α × β → ENNReal⦄ (hf : AEMeasurable f (μ.prod ν)) : ∀ᵐ x : α ∂μ, AEMeasurable (f ∘ (Prod.mk x)) ν := by rcases hf with ⟨g, hg⟩ @@ -2414,7 +2353,7 @@ lemma aemeasurability_prod₂ {α : Type u_1} {β : Type u_3} lemma aemeasurability_integral_component {α : Type u_1} {β : Type u_3} [MeasurableSpace α] [MeasurableSpace β] {μ : Measure α} {ν : Measure β} [SFinite ν] - [SFinite μ] ⦃f : α × β → ENNReal⦄ + ⦃f : α × β → ENNReal⦄ (hf : AEMeasurable f (μ.prod ν)) : AEMeasurable (fun x ↦ ∫⁻ (y : β), f (x, y) ∂ν) μ := by rcases hf with ⟨g, hg⟩ @@ -2479,7 +2418,7 @@ lemma lintegral_lintegral_pow_swap_rpow {α : Type u_1} {β : Type u_3} {p : ℝ (∫⁻ (y : β), (∫⁻ (x : α), (f x y) ^ p ∂μ) ^ p⁻¹ ∂ν) ^ p := by have p_pos : p > 0 := lt_of_lt_of_le zero_lt_one hp refine le_of_rpow_le (inv_pos_of_pos p_pos) ?_ - rw [ENNReal.rpow_rpow_inv p_pos.ne.symm] + rw [ENNReal.rpow_rpow_inv p_pos.ne'] exact lintegral_lintegral_pow_swap hp hf /-! ## Apply Minkowski's integral inequality to truncations @@ -2540,8 +2479,7 @@ lemma truncation_compl_ton_measurable {f : α → E₁} refine (aemeasurable_indicator_iff₀ ?hs).mpr hf.restrict.snd.restrict apply indicator_ton_measurable_lt hf.restrict -lemma restrict_to_support {a : ℝ} {p : ℝ} (hp : p > 0) - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] (f : α → E₁) : +lemma restrict_to_support {a : ℝ} {p : ℝ} (hp : p > 0) [NormedAddCommGroup E₁] (f : α → E₁) : ∫⁻ x : α in Function.support f, ‖trunc f a x‖₊ ^ p ∂ μ = ∫⁻ x : α, ‖trunc f a x‖₊ ^ p ∂μ := by apply setLIntegral_eq_of_support_subset unfold Function.support @@ -2552,8 +2490,7 @@ lemma restrict_to_support {a : ℝ} {p : ℝ} (hp : p > 0) intro f_zero simp_rw [f_zero]; simp [hp] -lemma restrict_to_support_trunc_compl {a : ℝ} {p : ℝ} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] (hp : p > 0) +lemma restrict_to_support_trunc_compl {a : ℝ} {p : ℝ} [NormedAddCommGroup E₁] (hp : p > 0) (f : α → E₁) : ∫⁻ x : α in Function.support f, ‖(f - trunc f a) x‖₊ ^ p ∂μ = ∫⁻ x : α, ‖(f - trunc f a) x‖₊ ^ p ∂μ := by @@ -2568,8 +2505,7 @@ lemma restrict_to_support_trunc_compl {a : ℝ} {p : ℝ} simp_rw [f_zero] simp [hp] -lemma restrict_to_support_trnc {a : ℝ} {p : ℝ} {j : Bool} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] (hp : p > 0) +lemma restrict_to_support_trnc {a : ℝ} {p : ℝ} {j : Bool} [NormedAddCommGroup E₁] (hp : p > 0) (f : α → E₁) : ∫⁻ x : α in Function.support f, ‖trnc j f a x‖₊ ^ p ∂μ = ∫⁻ x : α, ‖trnc j f a x‖₊ ^ p ∂μ := by @@ -2782,8 +2718,8 @@ lemma estimate_trnc₁ {spf : ScaledPowerFunction} {j : Bool} [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] (ht : t ∈ Ioo 0 1) (hp₀ : p₀ > 0) (hq₀ : q₀ > 0) (hp₁ : p₁ > 0) (hq₁ : q₁ > 0) (hpq : sel j p₀ p₁ ≤ sel j q₀ q₁) (hp' : sel j p₀ p₁ ≠ ⊤) (hq' : sel j q₀ q₁ ≠ ⊤) (hp₀p₁ : p₀ < p₁) - (hq₀q₁ : q₀ ≠ q₁) (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) + (hq₀q₁ : q₀ ≠ q₁) (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hf : AEMeasurable f μ) (hf₂ : SigmaFinite (μ.restrict (Function.support f))) (hspf : spf.σ = @ζ p₀ q₀ p₁ q₁ t) : ∫⁻ s : ℝ in Ioi 0, @@ -2808,15 +2744,15 @@ lemma estimate_trnc₁ {spf : ScaledPowerFunction} {j : Bool} · cases j · unfold sel dsimp only - apply hp₀.ne.symm + apply hp₀.ne' · unfold sel dsimp only - apply hp₁.ne.symm + apply hp₁.ne' · exact hp' · apply toReal_pos · cases j - · exact hq₀.ne.symm - · exact hq₁.ne.symm + · exact hq₀.ne' + · exact hq₁.ne' · exact hq' · exact toReal_mono hq' hpq · exact hf @@ -2888,7 +2824,7 @@ lemma estimate_trnc₁ {spf : ScaledPowerFunction} {j : Bool} congr rw [← one_div] refine Eq.symm (eLpNorm_eq_lintegral_rpow_nnnorm ?_ ?_) - · exact (interpolated_pos' hp₀ hp₁ hp).ne.symm + · exact (interpolated_pos' hp₀ hp₁ hp).ne' · exact interp_exp_ne_top (ne_of_lt hp₀p₁) ht hp -- TODO: move this to Weaktype.lean? @@ -2908,7 +2844,7 @@ lemma wnorm_eq_zero_iff {f : α → E₁} {p : ℝ≥0∞} [NormedAddCommGroup E rw [essSup_eq_sInf] at this let b := (min (sInf {a : ℝ≥0∞ | μ {x | a < ↑‖f x‖₊} = 0}) 1) / 2 have b_lt_inf : b < min (sInf {a : ℝ≥0∞ | μ {x | a < ↑‖f x‖₊} = 0}) 1 := - ENNReal.half_lt_self (lt_min this zero_lt_one).ne.symm + ENNReal.half_lt_self (lt_min this zero_lt_one).ne' (ne_of_lt (lt_of_le_of_lt (min_le_right _ 1) one_lt_top)) have meas_ne_zero : μ {x | b < ↑‖f x‖₊} ≠ 0 := by intro h @@ -2917,7 +2853,7 @@ lemma wnorm_eq_zero_iff {f : α → E₁} {p : ℝ≥0∞} [NormedAddCommGroup E calc _ < _ := b_lt_inf _ ≤ _ := min_le_left _ _ - have b_ne_0 : b ≠ 0 := (ENNReal.half_pos (lt_min this zero_lt_one).ne.symm).ne.symm + have b_ne_0 : b ≠ 0 := (ENNReal.half_pos (lt_min this zero_lt_one).ne').ne' have p_toReal_inv_pos : p.toReal⁻¹ > 0 := inv_pos_of_pos (toReal_pos hp h₀) have coe_b : ENNReal.ofNNReal b.toNNReal = b := coe_toNNReal (LT.lt.ne_top b_lt_inf) have : distribution f b μ = 0 := by @@ -2952,7 +2888,7 @@ lemma weaktype_estimate {C₀ : ℝ≥0} {p : ℝ≥0∞} {q : ℝ≥0∞} {f : eLpNorm f p μ ^ q.toReal * ENNReal.ofReal (t ^ (-q.toReal)) := by have wt_est := (h₀T f hf).2 -- the weaktype estimate have q_pos : q.toReal > 0 := by - refine toReal_pos hq.ne.symm hq'.ne_top + refine toReal_pos hq.ne' hq'.ne_top have tq_pos : ENNReal.ofReal t ^ q.toReal > 0 := coe_pow_pos ht have tq_ne_top : (ENNReal.ofReal t) ^ q.toReal ≠ ⊤ := coe_rpow_ne_top' q_pos -- have hq₁ : q.toReal = q := by exact toReal_ofReal q_nonneg @@ -2975,7 +2911,6 @@ lemma weaktype_estimate_top {C : ℝ≥0} {p : ℝ≥0∞} {q : ℝ≥0∞} have wt_est := (hT f hf).2 unfold wnorm at wt_est split_ifs at wt_est - have : eLpNormEssSup (T f) ν ^ p.toReal ≤ (C * eLpNorm f p μ) ^ p.toReal := by gcongr have snorm_est : eLpNormEssSup (T f) ν ≤ ENNReal.ofReal t := le_trans wt_est ht apply nonpos_iff_eq_zero.mp calc @@ -2992,14 +2927,14 @@ lemma weaktype_aux₀ {p₀ q₀ p q : ℝ≥0∞} [NormedAddCommGroup E₁] [No (h₀T : HasWeakType T p₀ q₀ μ ν C₀) (hf : AEStronglyMeasurable f μ) (hF : eLpNorm f p μ = 0) : eLpNorm (T f) q ν = 0 := by - have f_ae_0 : f =ᵐ[μ] 0 := (eLpNorm_eq_zero_iff hf hp.ne.symm).mp hF - have hf₂ : eLpNorm f p₀ μ = 0 := (eLpNorm_eq_zero_iff hf hp₀.ne.symm).mpr f_ae_0 + have f_ae_0 : f =ᵐ[μ] 0 := (eLpNorm_eq_zero_iff hf hp.ne').mp hF + have hf₂ : eLpNorm f p₀ μ = 0 := (eLpNorm_eq_zero_iff hf hp₀.ne').mpr f_ae_0 have hf₁ : Memℒp f p₀ μ := ⟨hf, by rw [hf₂]; exact zero_lt_top⟩ have := (h₀T f hf₁).2 rw [hf₂, mul_zero] at this have wnorm_0 : wnorm (T f) q₀ ν = 0 := nonpos_iff_eq_zero.mp this - have : (T f) =ᵐ[ν] 0 := (wnorm_eq_zero_iff hq₀.ne.symm).mp wnorm_0 - exact (eLpNorm_eq_zero_iff (h₀T _ hf₁).1 hq.ne.symm).mpr this + have : (T f) =ᵐ[ν] 0 := (wnorm_eq_zero_iff hq₀.ne').mp wnorm_0 + exact (eLpNorm_eq_zero_iff (h₀T _ hf₁).1 hq.ne').mpr this lemma weaktype_estimate_trunc_compl {C₀ : ℝ≥0} {p p₀: ℝ≥0∞} {f : α → E₁} [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] [NormedAddCommGroup E₂] @@ -3076,13 +3011,13 @@ lemma weaktype_estimate_trunc_compl_top {C₀ : ℝ≥0} (hC₀ : C₀ > 0) {p p meas_eLpNormEssSup_lt · have p_pos : p > 0 := lt_trans hp₀ hp₀p have snorm_p_pos : eLpNorm f p μ ≠ 0 := - fun snorm_0 ↦ Ne.symm (ne_of_lt snorm_pos) (eLpNormEssSup_eq_zero_iff.mpr - ((eLpNorm_eq_zero_iff hf.1 p_pos.ne.symm).mp snorm_0)) + fun snorm_0 ↦ snorm_pos.ne' <| eLpNormEssSup_eq_zero_iff.mpr <| + (eLpNorm_eq_zero_iff hf.1 p_pos.ne').mp snorm_0 have term_pos : (ENNReal.ofNNReal C₀) ^ p₀.toReal * eLpNorm f p μ ^ p.toReal > 0 := by - apply ENNReal.mul_pos <;> exact (rpow_pos_of_nonneg (by positivity) (by positivity)).ne.symm + apply ENNReal.mul_pos <;> exact (rpow_pos_of_nonneg (by positivity) (by positivity)).ne' have term_ne_top : (ENNReal.ofNNReal C₀) ^ p₀.toReal * eLpNorm f p μ ^ p.toReal ≠ ⊤ - := mul_ne_top (rpow_ne_top' (ENNReal.coe_ne_zero.mpr hC₀.ne.symm) coe_ne_top) + := mul_ne_top (rpow_ne_top' (ENNReal.coe_ne_zero.mpr hC₀.ne') coe_ne_top) (rpow_ne_top' snorm_p_pos (Memℒp.eLpNorm_ne_top hf)) have d_pos : d > 0 := by rw [hdeq] @@ -3126,7 +3061,7 @@ lemma weaktype_estimate_trunc_compl_top {C₀ : ℝ≥0} (hC₀ : C₀ > 0) {p p _ = _ := meas_eLpNormEssSup_lt lemma weaktype_estimate_trunc_top {C₁ : ℝ≥0} (hC₁ : C₁ > 0) {p p₁ q₁ : ℝ≥0∞} [MeasurableSpace E₁] - [NormedAddCommGroup E₁] [NormedAddCommGroup E₁] + [NormedAddCommGroup E₁] [BorelSpace E₁] [NormedAddCommGroup E₂] (hp : 0 < p) (hp₁ : p₁ < ⊤) (hq₁ : q₁ = ⊤) (hp₁p : p < p₁) {f : α → E₁} (hf : Memℒp f p μ) (h₁T : HasWeakType T p₁ q₁ μ ν C₁) {t : ℝ} (ht : t > 0) {a : ℝ} {d : ℝ} -- (hd : d > 0) @@ -3160,12 +3095,11 @@ lemma weaktype_estimate_trunc_top {C₁ : ℝ≥0} (hC₁ : C₁ > 0) {p p₁ q intro snorm_0 apply Ne.symm (ne_of_lt snorm_pos) apply eLpNormEssSup_eq_zero_iff.mpr - exact (eLpNorm_eq_zero_iff hf.1 hp.ne.symm).mp snorm_0 - have term_pos : (ENNReal.ofNNReal C₁) ^ p₁.toReal * eLpNorm f p μ ^ p.toReal > 0 - := by - apply ENNReal.mul_pos <;> exact (rpow_pos_of_nonneg (by positivity) (by positivity)).ne.symm + exact (eLpNorm_eq_zero_iff hf.1 hp.ne').mp snorm_0 + have term_pos : (ENNReal.ofNNReal C₁) ^ p₁.toReal * eLpNorm f p μ ^ p.toReal > 0 := by + apply ENNReal.mul_pos <;> exact (rpow_pos_of_nonneg (by positivity) (by positivity)).ne' have term_ne_top : (ENNReal.ofNNReal C₁) ^ p₁.toReal * eLpNorm f p μ ^ p.toReal ≠ ⊤ := - mul_ne_top (rpow_ne_top' (ENNReal.coe_ne_zero.mpr hC₁.ne.symm) coe_ne_top) + mul_ne_top (rpow_ne_top' (ENNReal.coe_ne_zero.mpr hC₁.ne') coe_ne_top) (rpow_ne_top' snorm_p_pos (Memℒp.eLpNorm_ne_top hf)) have d_pos : d > 0 := by rw [hdeq] @@ -3185,7 +3119,7 @@ lemma weaktype_estimate_trunc_top {C₁ : ℝ≥0} (hC₁ : C₁ > 0) {p p₁ q rw [ENNReal.ofReal_div_of_pos] <;> try positivity rw [div_eq_mul_inv] ring - · exact (sub_pos.mpr (toReal_strict_mono p₁ne_top hp₁p)).ne.symm + · exact (sub_pos.mpr (toReal_strict_mono p₁ne_top hp₁p)).ne' · positivity _ = _ := by rw [ofReal_rpow_of_pos ht] @@ -3216,11 +3150,8 @@ variable {α α' E E₁ E₂ E₃ : Type*} {m : MeasurableSpace α} {m' : Measur [NormedAddCommGroup E₂] [NormedAddCommGroup E₃] [MeasurableSpace E] [BorelSpace E] - -- [MeasurableSpace E₁] [BorelSpace E₁] - -- [MeasurableSpace E₂] [BorelSpace E₂] [MeasurableSpace E₃] [BorelSpace E₃] - -- (L : E₁ →L[ℝ] E₂ →L[ℝ] E₃) - {f : α → E₁} {t : ℝ} -- {s x y : ℝ≥0∞} + {f : α → E₁} {t : ℝ} {T : (α → E₁) → (α' → E₂)} namespace MeasureTheory @@ -3242,7 +3173,7 @@ def SubadditiveOn (T : (α → E₁) → α' → E₂) (P : (α → E₁) → Pr namespace SubadditiveOn -def antitone {T : (α → E₁) → α' → E₂} {P P' : (α → E₁) → Prop} +lemma antitone {T : (α → E₁) → α' → E₂} {P P' : (α → E₁) → Prop} (h : ∀ {u : α → E₁}, P u → P' u) {A : ℝ} (sa : SubadditiveOn T P' A) : SubadditiveOn T P A := fun f g x hf hg ↦ sa f g x (h hf) (h hg) @@ -3362,8 +3293,7 @@ variable {α α' E E₁ E₂ E₃ : Type*} {m : MeasurableSpace α} {m' : Measur {p p' q p₀ q₀ p₁ q₁: ℝ≥0∞} {C₀ C₁ : ℝ≥0} {μ : Measure α} {ν : Measure α'} {a : ℝ}-- truncation parameter - -- (L : E₁ →L[ℝ] E₂ →L[ℝ] E₃) - {f : α → E₁} {t : ℝ} -- {s x y : ℝ≥0∞} + {f : α → E₁} {t : ℝ} {T : (α → E₁) → (α' → E₂)} /-! ## Proof of the real interpolation theorem @@ -3376,7 +3306,7 @@ namespace MeasureTheory /-- Proposition that expresses that the map `T` map between function spaces preserves AE strong measurability on L^p. -/ def PreservesAEStrongMeasurability - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [NormedAddCommGroup E₂] + [NormedAddCommGroup E₁] [NormedAddCommGroup E₂] (T : (α → E₁) → α' → E₂) (p : ℝ≥0∞) : Prop := ∀ ⦃f : α → E₁⦄, Memℒp f p μ → AEStronglyMeasurable (T f) ν @@ -3444,10 +3374,10 @@ lemma estimate_norm_rpow_range_operator' ENNReal.ofReal (s^(q.toReal - 1)) + distribution (T (f - trunc f (tc.ton s))) (ENNReal.ofReal s) ν * ENNReal.ofReal (s^(q.toReal - 1)) ≤ - (if (q₁ < ⊤) then 1 else 0) * (C₁ ^ q₁.toReal * (∫⁻ s in Ioi (0 : ℝ), + (if q₁ < ⊤ then 1 else 0) * (C₁ ^ q₁.toReal * (∫⁻ s in Ioi (0 : ℝ), eLpNorm (trunc f (tc.ton s)) p₁ μ ^ q₁.toReal * ENNReal.ofReal (s ^ (q.toReal - q₁.toReal - 1)))) + - (if (q₀ < ⊤) then 1 else 0) * (C₀ ^ q₀.toReal * ∫⁻ s in Ioi (0 : ℝ), + (if q₀ < ⊤ then 1 else 0) * (C₀ ^ q₀.toReal * ∫⁻ s in Ioi (0 : ℝ), eLpNorm (f - trunc f (tc.ton s)) (p₀) μ ^ q₀.toReal * ENNReal.ofReal (s ^ (q.toReal - q₀.toReal - 1))) := by have : ∀ q' q : ℝ, -q' + (q - 1) = q - q' - 1 := by intro q' q; group @@ -3477,9 +3407,9 @@ lemma estimate_norm_rpow_range_operator' have tone := tc.ton_is_ton split_ifs at tone · apply aemeasurable_restrict_of_monotoneOn measurableSet_Ioi - exact Monotone.comp_monotoneOn norm_trunc_mono (StrictMonoOn.monotoneOn tone) + exact Monotone.comp_monotoneOn norm_trunc_mono tone.monotoneOn · apply aemeasurable_restrict_of_antitoneOn measurableSet_Ioi - exact Monotone.comp_antitoneOn norm_trunc_mono (StrictAntiOn.antitoneOn tone) + exact Monotone.comp_antitoneOn norm_trunc_mono tone.antitoneOn · apply AEMeasurable.coe_nnreal_ennreal · apply AEMeasurable.real_toNNReal · apply AEMeasurable.pow_const @@ -3525,8 +3455,8 @@ lemma simplify_factor₀ {D : ℝ} [NormedAddCommGroup E₁] (hq₀' : q₀ ≠ ⊤) (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (ht : t ∈ Ioo 0 1) - (hq₀q₁ : q₀ ≠ q₁) (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) + (hq₀q₁ : q₀ ≠ q₁) (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hC₀ : C₀ > 0) (hC₁ : C₁ > 0) (hF : eLpNorm f p μ ∈ Ioo 0 ⊤) (hD : D = @d _ E₁ _ p p₀ q₀ p₁ q₁ C₀ C₁ μ _ f) : @@ -3587,8 +3517,8 @@ lemma simplify_factor₁ {D : ℝ} [NormedAddCommGroup E₁] (hq₁' : q₁ ≠ ⊤) (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (ht : t ∈ Ioo 0 1) - (hq₀q₁ : q₀ ≠ q₁) (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) + (hq₀q₁ : q₀ ≠ q₁) (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hC₀ : C₀ > 0) (hC₁ : C₁ > 0) (hF : eLpNorm f p μ ∈ Ioo 0 ⊤) (hD : D = @d _ E₁ _ p p₀ q₀ p₁ q₁ C₀ C₁ μ _ f) : @@ -3628,9 +3558,9 @@ lemma simplify_factor₁ {D : ℝ} · exact ne_top_of_Ioo hF · exact ne_zero_of_Ioo hF · exact ne_top_of_Ioo hF - · exact (ENNReal.coe_pos.mpr hC₁).ne.symm + · exact (ENNReal.coe_pos.mpr hC₁).ne' · exact coe_ne_top - · exact ENNReal.inv_ne_zero.mpr (rpow_ne_top' ((ENNReal.coe_pos.mpr hC₁).ne.symm) coe_ne_top) + · exact ENNReal.inv_ne_zero.mpr (rpow_ne_top' ((ENNReal.coe_pos.mpr hC₁).ne') coe_ne_top) · exact ENNReal.inv_ne_zero.mpr (d_ne_top_aux₂ hF) · exact d_ne_zero_aux₁ hC₀ · exact d_ne_zero_aux₀ hF @@ -3642,7 +3572,7 @@ lemma simplify_factor₁ {D : ℝ} · exact d_ne_top_aux₄ hC₀ hC₁ hF · exact d_pos hC₀ hC₁ hF -instance finite_spanning_sets_from_lintegrable {g : α → ℝ≥0∞} (hg : AEMeasurable g μ) +def finite_spanning_sets_from_lintegrable {g : α → ℝ≥0∞} (hg : AEMeasurable g μ) (hg_int : ∫⁻ x, g x ∂μ < ⊤) : Measure.FiniteSpanningSetsIn (μ.restrict (Function.support g)) (Set.univ) where set := fun n ↦ if n = 0 then {x | g x = 0} else { x | 1 / (n + 1) ≤ g x } @@ -3697,14 +3627,14 @@ instance finite_spanning_sets_from_lintegrable {g : α → ℝ≥0∞} (hg : AEM nth_rw 1 [← add_zero (n : ℝ≥0∞)] refine (ENNReal.add_lt_add_iff_left coe_ne_top).mpr zero_lt_one -instance support_sigma_finite_of_lintegrable {g : α → ℝ≥0∞} (hg : AEMeasurable g μ) +lemma support_sigma_finite_of_lintegrable {g : α → ℝ≥0∞} (hg : AEMeasurable g μ) (hg_int : ∫⁻ x, g x ∂μ < ⊤) : SigmaFinite (μ.restrict (Function.support g)) where out' := by apply nonempty_of_exists use (finite_spanning_sets_from_lintegrable hg hg_int) -instance support_sigma_finite_from_Memℒp +lemma support_sigma_finite_from_Memℒp [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] (hf : Memℒp f p μ) (hp : p ≠ ⊤) (hp' : p ≠ 0) : SigmaFinite (μ.restrict (Function.support f)) := by @@ -3736,13 +3666,13 @@ instance support_sigma_finite_from_Memℒp · contradiction · exact lintegral_rpow_nnnorm_lt_top_of_eLpNorm'_lt_top (toReal_pos hp' hp) obs -instance support_sfinite_from_Memℒp - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] (hf : Memℒp f p μ) - (hp : p ≠ ⊤) (hp' : p ≠ 0) : - SFinite (μ.restrict (Function.support f)) := by - have : SigmaFinite (μ.restrict (Function.support f)) := by - apply support_sigma_finite_from_Memℒp hf hp hp' - apply instSFiniteOfSigmaFinite +-- lemma support_sfinite_from_Memℒp +-- [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] (hf : Memℒp f p μ) +-- (hp : p ≠ ⊤) (hp' : p ≠ 0) : +-- SFinite (μ.restrict (Function.support f)) := by +-- have : SigmaFinite (μ.restrict (Function.support f)) := by +-- apply support_sigma_finite_from_Memℒp hf hp hp' +-- apply instSFiniteOfSigmaFinite lemma combine_estimates₀ {A : ℝ} (hA : A > 0) [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] @@ -3750,8 +3680,8 @@ lemma combine_estimates₀ {A : ℝ} (hA : A > 0) {spf : ScaledPowerFunction} (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (ht : t ∈ Ioo 0 1) (hp₀p₁ : p₀ < p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hf : Memℒp f p μ) (hT : Subadditive_trunc T A f ν) (hC₀ : C₀ > 0) (hC₁ : C₁ > 0) (hF : eLpNorm f p μ ∈ Ioo 0 ⊤) @@ -3762,8 +3692,8 @@ lemma combine_estimates₀ {A : ℝ} (hA : A > 0) (h₂T : PreservesAEStrongMeasurability T p (ν := ν) (μ := μ)) : ∫⁻ x , ‖T f x‖₊ ^ q.toReal ∂ν ≤ ENNReal.ofReal ((2 * A) ^ q.toReal * q.toReal) * - ((if (q₁ < ⊤) then 1 else 0) * ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ + - (if (q₀ < ⊤) then 1 else 0) * ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹) * + ((if q₁ < ⊤ then 1 else 0) * ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ + + (if q₀ < ⊤ then 1 else 0) * ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹) * C₀ ^ ((1 - t) * q.toReal) * C₁ ^ (t * q.toReal) * eLpNorm f p μ ^ q.toReal := by have one_le_p₀ := hp₀.1 have one_le_p1 := hp₁.1 @@ -3775,9 +3705,8 @@ lemma combine_estimates₀ {A : ℝ} (hA : A > 0) have : SigmaFinite (μ.restrict (Function.support f)) := by apply support_sigma_finite_from_Memℒp (p := p) hf · exact interp_exp_ne_top (ne_of_lt hp₀p₁) ht hp - · exact p_pos.ne.symm + · exact p_pos.ne' let tc := spf_to_tc spf - have := spf.hd calc ∫⁻ x , ‖T f x‖₊ ^ q.toReal ∂ν ≤ ENNReal.ofReal ((2 * A) ^ q.toReal * q.toReal) * ∫⁻ s in Ioi (0 : ℝ), @@ -3791,10 +3720,10 @@ lemma combine_estimates₀ {A : ℝ} (hA : A > 0) · exact hT · exact AEStronglyMeasurable.aemeasurable (h₂T hf) _ ≤ ENNReal.ofReal ((2 * A)^q.toReal * q.toReal) * - ((if (q₁ < ⊤) then 1 else 0) * (C₁ ^ q₁.toReal * (∫⁻ s in Ioi (0 : ℝ), + ((if q₁ < ⊤ then 1 else 0) * (C₁ ^ q₁.toReal * (∫⁻ s in Ioi (0 : ℝ), eLpNorm (trunc f (tc.ton s)) p₁ μ ^ q₁.toReal * ENNReal.ofReal (s ^ (q.toReal - q₁.toReal - 1)))) + - (if (q₀ < ⊤) then 1 else 0) * (C₀ ^ q₀.toReal * ∫⁻ s in Ioi (0 : ℝ), + (if q₀ < ⊤ then 1 else 0) * (C₀ ^ q₀.toReal * ∫⁻ s in Ioi (0 : ℝ), eLpNorm (f - trunc f (tc.ton s)) p₀ μ ^ q₀.toReal * ENNReal.ofReal (s ^ (q.toReal - q₀.toReal - 1)))) := by gcongr @@ -3845,11 +3774,11 @@ lemma combine_estimates₀ {A : ℝ} (hA : A > 0) · exact Ne.lt_top p₁ne_top · exact (interp_exp_between p₀pos p₁pos hp₀p₁ ht hp).2 _ ≤ ENNReal.ofReal ((2 * A)^q.toReal * q.toReal) * - ((if (q₁ < ⊤) then 1 else 0) * (C₁ ^ q₁.toReal * + ((if q₁ < ⊤ then 1 else 0) * (C₁ ^ q₁.toReal * (ENNReal.ofReal (spf.d ^ (q.toReal - q₁.toReal)) * ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ * ((eLpNorm f p μ) ^ p.toReal) ^ ((sel ⊤ p₀ p₁).toReal ⁻¹ * (sel ⊤ q₀ q₁).toReal))) + - (if (q₀ < ⊤) then 1 else 0) * (C₀ ^ q₀.toReal * + (if q₀ < ⊤ then 1 else 0) * (C₀ ^ q₀.toReal * (ENNReal.ofReal (spf.d ^ (q.toReal - q₀.toReal)) * ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹ * (((eLpNorm f p μ) ^ p.toReal) ^ ((sel ⊥ p₀ p₁).toReal ⁻¹ * (sel ⊥ q₀ q₁).toReal))))) := by apply mul_le_mul_left' @@ -3913,8 +3842,8 @@ lemma combine_estimates₁ {A : ℝ} [MeasurableSpace E₁] [NormedAddCommGroup {spf : ScaledPowerFunction} (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (ht : t ∈ Ioo 0 1) (hp₀p₁ : p₀ < p₁) (hq₀q₁ : q₀ ≠ q₁) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hf : Memℒp f p μ) (hT : Subadditive_trunc T A f ν) (h₁T : HasWeakType T p₁ q₁ μ ν C₁) (h₀T : HasWeakType T p₀ q₀ μ ν C₀) @@ -3925,17 +3854,17 @@ lemma combine_estimates₁ {A : ℝ} [MeasurableSpace E₁] [NormedAddCommGroup (lt_of_lt_of_le hp₁.1 hp₁.2) (ne_of_lt hp₀p₁) hC₀ hC₁ hF) : eLpNorm (T f) q ν ≤ ENNReal.ofReal (2 * A) * q ^ q⁻¹.toReal * - (((if (q₁ < ⊤) then 1 else 0) * ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ + - (if (q₀ < ⊤) then 1 else 0) * ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹)) ^ q⁻¹.toReal * + (((if q₁ < ⊤ then 1 else 0) * ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ + + (if q₀ < ⊤ then 1 else 0) * ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹)) ^ q⁻¹.toReal * C₀ ^ (1 - t) * C₁ ^ t * eLpNorm f p μ := by - have q_ne_zero : q ≠ 0 := (interpolated_pos' (lt_of_lt_of_le hp₀.1 hp₀.2) (lt_of_lt_of_le hp₁.1 hp₁.2) hq).ne.symm + have q_ne_zero : q ≠ 0 := (interpolated_pos' (lt_of_lt_of_le hp₀.1 hp₀.2) (lt_of_lt_of_le hp₁.1 hp₁.2) hq).ne' have q_ne_top : q ≠ ⊤ := interp_exp_ne_top hq₀q₁ ht hq have q'pos : q.toReal > 0 := toReal_pos q_ne_zero q_ne_top refine le_of_rpow_le q'pos ?_ calc _ = ∫⁻ x , ‖T f x‖₊ ^ q.toReal ∂ν := by unfold eLpNorm eLpNorm' - split_ifs <;> [contradiction; rw [one_div, ENNReal.rpow_inv_rpow q'pos.ne.symm]] + split_ifs <;> [contradiction; rw [one_div, ENNReal.rpow_inv_rpow q'pos.ne']] _ ≤ _ := by apply combine_estimates₀ (hT := hT) (p := p) <;> try assumption @@ -3948,13 +3877,13 @@ lemma combine_estimates₁ {A : ℝ} [MeasurableSpace E₁] [NormedAddCommGroup · rw [toReal_inv] rw [ENNReal.rpow_inv_rpow] · exact ofReal_toReal_eq_iff.mpr q_ne_top - · exact q'pos.ne.symm + · exact q'pos.ne' · rw [toReal_inv] rw [ENNReal.rpow_inv_rpow] - exact q'pos.ne.symm + exact q'pos.ne' lemma simplify_factor₃ [NormedAddCommGroup E₁] (hp₀ : p₀ > 0) (hp₀' : p₀ ≠ ⊤) (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) (hp₀p₁ : p₀ = p₁) : + (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) (hp₀p₁ : p₀ = p₁) : C₀ ^ q₀.toReal * (eLpNorm f p μ ^ p.toReal) ^ (q₀.toReal / p₀.toReal) = (↑C₀ * eLpNorm f p μ) ^ q₀.toReal := by have hp' : p = p₀ := (interp_exp_eq hp₀p₁ ht hp).symm -- TODO abstract @@ -3963,22 +3892,21 @@ lemma simplify_factor₃ [NormedAddCommGroup E₁] (hp₀ : p₀ > 0) (hp₀' : try positivity apply ne_of_gt apply toReal_pos - · apply hp₀.ne.symm + · apply hp₀.ne' · exact hp₀' lemma simplify_factor_aux₄ [NormedAddCommGroup E₁] (hq₀' : q₀ ≠ ⊤) (hp₀ : p₀ ∈ Ioc 0 q₀) (ht : t ∈ Ioo 0 1) - (hp₀p₁ : p₀ = p₁) (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) + (hp₀p₁ : p₀ = p₁) (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) (hF : eLpNorm f p μ ∈ Ioo 0 ⊤) : ↑C₀ ^ ((1 - t) * q.toReal) * (eLpNorm f p μ ^ p.toReal) ^ ((1 - t) * p₀⁻¹.toReal * q.toReal) * ↑C₁ ^ (t * q.toReal) * (eLpNorm f p μ ^ p.toReal) ^ (t * p₁⁻¹.toReal * q.toReal) = C₀ ^ ((1 - t) * q.toReal) * C₁ ^ (t * q.toReal) * eLpNorm f p μ ^ q.toReal := by have hp' : p = p₀ := Eq.symm (interp_exp_eq hp₀p₁ ht hp) - have := (Ioo.one_sub_mem ht).1 have p₀pos : p₀ > 0 := hp₀.1 have p₀ne_top : p₀ ≠ ⊤ := ne_top_of_le_ne_top hq₀' hp₀.2 - have p₀toReal_pos : p₀.toReal > 0 := toReal_pos p₀pos.ne.symm (p₀ne_top) + have p₀toReal_pos : p₀.toReal > 0 := toReal_pos p₀pos.ne' (p₀ne_top) rw [hp', ← hp₀p₁] have : ∀ a b c d : ℝ≥0∞, a * b * c * d = a * c * (b * d) := by intro a b c d; ring rw [this] @@ -3989,16 +3917,15 @@ lemma simplify_factor_aux₄ [NormedAddCommGroup E₁] (hq₀' : q₀ ≠ ⊤) ring_nf field_simp · rw [← hp'] - apply ne_of_gt - exact d_pos_aux₀ hF + exact d_pos_aux₀ hF |>.ne' · rw [← hp'] refine d_ne_top_aux₀ hF lemma simplify_factor₄ {D : ℝ} [NormedAddCommGroup E₁] (hq₀' : q₀ ≠ ⊤) (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (ht : t ∈ Ioo 0 1) (hp₀p₁ : p₀ = p₁) - (hq₀q₁ : q₀ ≠ q₁) (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) + (hq₀q₁ : q₀ ≠ q₁) (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hC₀ : C₀ > 0) (hC₁ : C₁ > 0) (hF : eLpNorm f p μ ∈ Ioo 0 ⊤) (hD : D = @d _ E₁ _ p p₀ q₀ p₁ q₁ C₀ C₁ μ _ f) : @@ -4014,8 +3941,8 @@ lemma simplify_factor₅ {D : ℝ} [NormedAddCommGroup E₁] (hq₁' : q₁ ≠ (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (ht : t ∈ Ioo 0 1) (hp₀p₁ : p₀ = p₁) - (hq₀q₁ : q₀ ≠ q₁) (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹) + (hq₀q₁ : q₀ ≠ q₁) (hp : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹) (hC₀ : C₀ > 0) (hC₁ : C₁ > 0) (hF : eLpNorm f p μ ∈ Ioo 0 ⊤) (hD : D = @d _ E₁ _ p p₀ q₀ p₁ q₁ C₀ C₁ μ _ f) : @@ -4031,11 +3958,11 @@ lemma simplify_factor₅ {D : ℝ} [NormedAddCommGroup E₁] (hq₁' : q₁ ≠ /-- The trivial case for the estimate in the real interpolation theorem when the function `Lp` norm of `f` vanishes. -/ lemma exists_hasStrongType_real_interpolation_aux₀ {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [NormedAddCommGroup E₂] + [NormedAddCommGroup E₁] [NormedAddCommGroup E₂] (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) {C₀ : ℝ≥0} - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) / p₀ + (ENNReal.ofReal t) / p₁) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) / q₀ + (ENNReal.ofReal t) / q₁) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) / p₀ + (ENNReal.ofReal t) / p₁) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) / q₀ + (ENNReal.ofReal t) / q₁) (h₀T : HasWeakType T p₀ q₀ μ ν C₀) (h₂T : PreservesAEStrongMeasurability (μ := μ) (ν := ν) T p) (hf : Memℒp f p μ) (hF : eLpNorm f p μ = 0) : @@ -4047,14 +3974,14 @@ lemma exists_hasStrongType_real_interpolation_aux₀ {p₀ p₁ q₀ q₁ p q : have q₀pos : q₀ > 0 := pos_of_rb_Ioc hp₀ have q₁pos : q₁ > 0 := pos_of_rb_Ioc hp₁ have q_pos : q > 0 := interpolated_pos' q₀pos q₁pos hq - have f_ae_0 : f =ᵐ[μ] 0 := (eLpNorm_eq_zero_iff hf.1 p_pos.ne.symm).mp hF - have hf₂ : eLpNorm f p₀ μ = 0 := (eLpNorm_eq_zero_iff hf.1 p₀pos.ne.symm).mpr f_ae_0 + have f_ae_0 : f =ᵐ[μ] 0 := (eLpNorm_eq_zero_iff hf.1 p_pos.ne').mp hF + have hf₂ : eLpNorm f p₀ μ = 0 := (eLpNorm_eq_zero_iff hf.1 p₀pos.ne').mpr f_ae_0 have hf₁ : Memℒp f p₀ μ := ⟨hf.1, by rw [hf₂]; exact zero_lt_top⟩ have := (h₀T f hf₁).2 rw [hf₂, mul_zero] at this have wnorm_0 : wnorm (T f) q₀ ν = 0 := nonpos_iff_eq_zero.mp this - have : (T f) =ᵐ[ν] 0 := (wnorm_eq_zero_iff q₀pos.ne.symm).mp wnorm_0 - exact (eLpNorm_eq_zero_iff (h₂T hf) q_pos.ne.symm).mpr this + have : (T f) =ᵐ[ν] 0 := (wnorm_eq_zero_iff q₀pos.ne').mp wnorm_0 + exact (eLpNorm_eq_zero_iff (h₂T hf) q_pos.ne').mpr this /-- The estimate for the real interpolation theorem in case `p₀ < p₁`. -/ lemma exists_hasStrongType_real_interpolation_aux {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} {A : ℝ} @@ -4062,15 +3989,15 @@ lemma exists_hasStrongType_real_interpolation_aux {p₀ p₁ q₀ q₁ p q : ℝ [MeasurableSpace E₂] [NormedAddCommGroup E₂] [BorelSpace E₂] (hA : A > 0) (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (hp₀p₁ : p₀ < p₁) (hq₀q₁ : q₀ ≠ q₁) {C₀ C₁ : ℝ≥0} (ht : t ∈ Ioo 0 1) (hC₀ : 0 < C₀) (hC₁ : 0 < C₁) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) / p₀ + (ENNReal.ofReal t) / p₁) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) / q₀ + (ENNReal.ofReal t) / q₁) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) / p₀ + (ENNReal.ofReal t) / p₁) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) / q₀ + (ENNReal.ofReal t) / q₁) (hT : Subadditive_trunc T A f ν) (h₀T : HasWeakType T p₀ q₀ μ ν C₀) (h₁T : HasWeakType T p₁ q₁ μ ν C₁) (h₂T : PreservesAEStrongMeasurability (μ := μ) (ν := ν) T p) (hf : Memℒp f p μ) : eLpNorm (T f) q ν ≤ ENNReal.ofReal (2 * A) * q ^ q⁻¹.toReal * - (((if (q₁ < ⊤) then 1 else 0) * ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ + - (if (q₀ < ⊤) then 1 else 0) * ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹)) ^ q⁻¹.toReal * + (((if q₁ < ⊤ then 1 else 0) * ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ + + (if q₀ < ⊤ then 1 else 0) * ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹)) ^ q⁻¹.toReal * C₀ ^ (1 - t) * C₁ ^ t * eLpNorm f p μ := by have hq₀ : q₀ > 0 := pos_of_rb_Ioc hp₀ have hq₁ : q₁ > 0 := pos_of_rb_Ioc hp₁ @@ -4085,12 +4012,11 @@ lemma exists_hasStrongType_real_interpolation_aux {p₀ p₁ q₀ q₁ p q : ℝ -- TODO: the below lemmas were split because otherwise the lean server would crash -- (seems to be related to the linter?) (after the merge) -lemma exists_hasStrongType_real_interpolation_aux₁ {f : α → E₁} [MeasurableSpace E₁] - [NormedAddCommGroup E₁] [BorelSpace E₁] +lemma exists_hasStrongType_real_interpolation_aux₁ {f : α → E₁} [NormedAddCommGroup E₁] (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (hp₀p₁ : p₀ = p₁) (hq₀q₁ : q₀ < q₁) {C₀ C₁ : ℝ≥0} (ht : t ∈ Ioo 0 1) (hC₀ : 0 < C₀) (hC₁ : 0 < C₁) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) / p₀ + (ENNReal.ofReal t) / p₁) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) / q₀ + (ENNReal.ofReal t) / q₁) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) / p₀ + (ENNReal.ofReal t) / p₁) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) / q₀ + (ENNReal.ofReal t) / q₁) (hF : eLpNorm f p μ ∈ Ioo 0 ⊤) : (ENNReal.ofReal q.toReal * ((C₀ * eLpNorm f p μ )^ q₀.toReal * @@ -4111,7 +4037,7 @@ lemma exists_hasStrongType_real_interpolation_aux₁ {f : α → E₁} [Measurab have q₀lt_q_toReal : q₀.toReal < q.toReal := preservation_inequality_of_lt₀ ht q₀pos q₁pos hq hq₀q₁ have q_toReal_ne_zero : q.toReal ≠ 0 := - (interp_exp_toReal_pos' ht q₀pos q₁pos hq (Or.inl q₀ne_top)).ne.symm + (interp_exp_toReal_pos' ht q₀pos q₁pos hq (Or.inl q₀ne_top)).ne' have M_pos : M > 0 := by apply d_pos <;> try assumption have coe_q : ENNReal.ofReal q.toReal = q := @@ -4167,13 +4093,12 @@ lemma exists_hasStrongType_real_interpolation_aux₁ {f : α → E₁} [Measurab /-- The main estimate in the real interpolation theorem for `p₀ = p₁`, before taking roots, for the case `q₀ < q₁`. -/ -lemma exists_hasStrongType_real_interpolation_aux₂ {f : α → E₁} [MeasurableSpace E₁] - [NormedAddCommGroup E₁] - [BorelSpace E₁] [MeasurableSpace E₂] [NormedAddCommGroup E₂] [BorelSpace E₂] +lemma exists_hasStrongType_real_interpolation_aux₂ {f : α → E₁} + [NormedAddCommGroup E₁] [MeasurableSpace E₂] [NormedAddCommGroup E₂] [BorelSpace E₂] (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (hp₀p₁ : p₀ = p₁) (hq₀q₁ : q₀ < q₁) {C₀ C₁ : ℝ≥0} (ht : t ∈ Ioo 0 1) (hC₀ : 0 < C₀) (hC₁ : 0 < C₁) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) / p₀ + (ENNReal.ofReal t) / p₁) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) / q₀ + (ENNReal.ofReal t) / q₁) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) / p₀ + (ENNReal.ofReal t) / p₁) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) / q₀ + (ENNReal.ofReal t) / q₁) (h₀T : HasWeakType T p₀ q₀ μ ν C₀) (h₁T : HasWeakType T p₁ q₁ μ ν C₁) (h₂T : PreservesAEStrongMeasurability (μ := μ) (ν := ν) T p) (hf : Memℒp f p μ) : @@ -4190,7 +4115,7 @@ lemma exists_hasStrongType_real_interpolation_aux₂ {f : α → E₁} [Measurab have q₀ne_top : q₀ ≠ ⊤ := LT.lt.ne_top hq₀q₁ have p₀ne_top : p₀ ≠ ⊤ := ne_top_of_le_ne_top q₀ne_top hp₀.2 have q_toReal_ne_zero : q.toReal ≠ 0 := - (interp_exp_toReal_pos' ht q₀pos q₁pos hq (Or.inl q₀ne_top)).ne.symm + (interp_exp_toReal_pos' ht q₀pos q₁pos hq (Or.inl q₀ne_top)).ne' have p_eq_p₀ : p = p₀ := Eq.symm (interp_exp_eq hp₀p₁ ht hp) rcases (eq_zero_or_pos (eLpNorm f p μ)) with hF | snorm_pos · refine le_of_eq_of_le ?_ (zero_le _) @@ -4305,12 +4230,11 @@ lemma exists_hasStrongType_real_interpolation_aux₂ {f : α → E₁} [Measurab /-- The main estimate for the real interpolation theorem for `p₀ = p₁`, requiring `q₀ ≠ q₁`, before taking roots. -/ lemma exists_hasStrongType_real_interpolation_aux₃ {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} - [MeasurableSpace E₁] [NormedAddCommGroup E₁] [BorelSpace E₁] [MeasurableSpace E₂] - [NormedAddCommGroup E₂] [BorelSpace E₂] + [NormedAddCommGroup E₁] [MeasurableSpace E₂] [NormedAddCommGroup E₂] [BorelSpace E₂] (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (hp₀p₁ : p₀ = p₁) (hq₀q₁ : q₀ ≠ q₁) {C₀ C₁ : ℝ≥0} (ht : t ∈ Ioo 0 1) (hC₀ : 0 < C₀) (hC₁ : 0 < C₁) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) / p₀ + (ENNReal.ofReal t) / p₁) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) / q₀ + (ENNReal.ofReal t) / q₁) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) / p₀ + (ENNReal.ofReal t) / p₁) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) / q₀ + (ENNReal.ofReal t) / q₁) (h₀T : HasWeakType T p₀ q₀ μ ν C₀) (h₁T : HasWeakType T p₁ q₁ μ ν C₁) (h₂T : PreservesAEStrongMeasurability (μ := μ) (ν := ν) T p) (hf : Memℒp f p μ) : @@ -4339,15 +4263,15 @@ lemma exists_hasStrongType_real_interpolation_aux₄ {p₀ p₁ q₀ q₁ p q : [MeasurableSpace E₂] [NormedAddCommGroup E₂] [BorelSpace E₂] (hA : A > 0) (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (hq₀q₁ : q₀ ≠ q₁) {C₀ C₁ : ℝ≥0} (ht : t ∈ Ioo 0 1) (hC₀ : 0 < C₀) (hC₁ : 0 < C₁) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) / p₀ + (ENNReal.ofReal t) / p₁) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) / q₀ + (ENNReal.ofReal t) / q₁) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) / p₀ + (ENNReal.ofReal t) / p₁) + (hq : q⁻¹ = (1 - ENNReal.ofReal t) / q₀ + (ENNReal.ofReal t) / q₁) (hT : Subadditive_trunc T A f ν) (h₀T : HasWeakType T p₀ q₀ μ ν C₀) (h₁T : HasWeakType T p₁ q₁ μ ν C₁) (h₂T : PreservesAEStrongMeasurability (μ := μ) (ν := ν) T p) (hf : Memℒp f p μ) : eLpNorm (T f) q ν ≤ (if p₀ = p₁ then 1 else ENNReal.ofReal (2 * A)) * q ^ q⁻¹.toReal * - (((if (q₁ < ⊤) then 1 else 0) * ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ + - (if (q₀ < ⊤) then 1 else 0) * ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹)) ^ q⁻¹.toReal * + (((if q₁ < ⊤ then 1 else 0) * ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ + + (if q₀ < ⊤ then 1 else 0) * ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹)) ^ q⁻¹.toReal * C₀ ^ (1 - t) * C₁ ^ t * eLpNorm f p μ := by let M := if p₀ = p₁ then 1 else ENNReal.ofReal (2 * A) have hM : M = if p₀ = p₁ then 1 else ENNReal.ofReal (2 * A) := rfl @@ -4393,18 +4317,15 @@ lemma exists_hasStrongType_real_interpolation_aux₄ {p₀ p₁ q₀ q₁ p q : an element of `ℝ≥0∞`. -/ def C_realInterpolation_ENNReal (p₀ p₁ q₀ q₁ q : ℝ≥0∞) (C₀ C₁: ℝ≥0) (A : ℝ≥0) (t : ℝ) := (if p₀ = p₁ then 1 else ENNReal.ofReal (2 * A)) * q ^ q⁻¹.toReal * - (((if (q₁ < ⊤) then 1 else 0) * ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ + - (if (q₀ < ⊤) then 1 else 0) * ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹)) ^ q⁻¹.toReal * + (((if q₁ < ⊤ then 1 else 0) * ENNReal.ofReal |q.toReal - q₁.toReal|⁻¹ + + (if q₀ < ⊤ then 1 else 0) * ENNReal.ofReal |q.toReal - q₀.toReal|⁻¹)) ^ q⁻¹.toReal * C₀ ^ (1 - t) * C₁ ^ t -lemma C_realInterpolation_ENNReal_ne_top {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} {A : ℝ≥0} +lemma C_realInterpolation_ENNReal_ne_top {p₀ p₁ q₀ q₁ q : ℝ≥0∞} {A : ℝ≥0} (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (hq₀q₁ : q₀ ≠ q₁) {C₀ C₁ : ℝ≥0} (ht : t ∈ Ioo 0 1) (hC₀ : 0 < C₀) (hC₁ : 0 < C₁) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) / p₀ + (ENNReal.ofReal t) / p₁) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) / q₀ + (ENNReal.ofReal t) / q₁) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) / q₀ + (ENNReal.ofReal t) / q₁) : C_realInterpolation_ENNReal p₀ p₁ q₀ q₁ q C₀ C₁ A t ≠ ⊤ := by - have p₀pos : p₀ > 0 := hp₀.1 - have p₁pos : p₁ > 0 := hp₁.1 have q₀pos : q₀ > 0 := pos_of_rb_Ioc hp₀ have q₁pos : q₁ > 0 := pos_of_rb_Ioc hp₁ unfold C_realInterpolation_ENNReal @@ -4416,8 +4337,7 @@ lemma C_realInterpolation_ENNReal_ne_top {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} · exact Ne.symm top_ne_one · exact coe_ne_top · apply rpow_ne_top' - · apply ne_of_gt - exact interpolated_pos' q₀pos q₁pos hq + · exact interpolated_pos' q₀pos q₁pos hq |>.ne' · exact interp_exp_ne_top hq₀q₁ ht hq · apply rpow_ne_top' · split_ifs @@ -4427,33 +4347,26 @@ lemma C_realInterpolation_ENNReal_ne_top {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} · exact ofReal_inv_interp_sub_exp_pos₁ ht q₀pos q₁pos hq₀q₁ hq · exact ofReal_inv_interp_sub_exp_pos₀ ht q₀pos q₁pos hq₀q₁ hq · rw [one_mul, zero_mul, add_zero] - apply ne_of_gt - exact ofReal_inv_interp_sub_exp_pos₁ ht q₀pos q₁pos hq₀q₁ hq + exact ofReal_inv_interp_sub_exp_pos₁ ht q₀pos q₁pos hq₀q₁ hq |>.ne' · rw [zero_mul, one_mul, zero_add] - apply ne_of_gt - exact ofReal_inv_interp_sub_exp_pos₀ ht q₀pos q₁pos hq₀q₁ hq + exact ofReal_inv_interp_sub_exp_pos₀ ht q₀pos q₁pos hq₀q₁ hq |>.ne' · have q₀top : q₀ = ⊤ := not_lt_top.mp (by assumption) have q₁top : q₁ = ⊤ := not_lt_top.mp (by assumption) rw [q₀top, q₁top] at hq₀q₁ simp only [ne_eq, not_true_eq_false] at hq₀q₁ · split_ifs <;> exact Ne.symm (ne_of_beq_false rfl) · apply rpow_ne_top' - · apply ne_of_gt - exact ENNReal.coe_pos.mpr hC₀ + · exact ENNReal.coe_pos.mpr hC₀ |>.ne' · exact coe_ne_top · apply rpow_ne_top' - · apply ne_of_gt - exact ENNReal.coe_pos.mpr hC₁ + · exact ENNReal.coe_pos.mpr hC₁ |>.ne' · exact coe_ne_top -lemma C_realInterpolation_ENNReal_pos {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} {A : ℝ≥0} (hA : A > 0) +lemma C_realInterpolation_ENNReal_pos {p₀ p₁ q₀ q₁ q : ℝ≥0∞} {A : ℝ≥0} (hA : A > 0) (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (hq₀q₁ : q₀ ≠ q₁) {C₀ C₁ : ℝ≥0} (ht : t ∈ Ioo 0 1) (hC₀ : 0 < C₀) (hC₁ : 0 < C₁) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) / p₀ + (ENNReal.ofReal t) / p₁) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) / q₀ + (ENNReal.ofReal t) / q₁) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) / q₀ + (ENNReal.ofReal t) / q₁) : C_realInterpolation_ENNReal p₀ p₁ q₀ q₁ q C₀ C₁ A t > 0 := by - have p₀pos : p₀ > 0 := hp₀.1 - have p₁pos : p₁ > 0 := hp₁.1 have q₀pos : q₀ > 0 := pos_of_rb_Ioc hp₀ have q₁pos : q₁ > 0 := pos_of_rb_Ioc hp₁ unfold C_realInterpolation_ENNReal @@ -4464,8 +4377,7 @@ lemma C_realInterpolation_ENNReal_pos {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} {A · split_ifs · exact one_ne_zero · rw [← ofReal_zero] - apply ne_of_gt - refine (ofReal_lt_ofReal_iff_of_nonneg ?_).mpr ?_ + refine (ofReal_lt_ofReal_iff_of_nonneg ?_).mpr ?_ |>.ne' · rfl · apply _root_.mul_pos · exact zero_lt_two @@ -4512,13 +4424,12 @@ lemma C_realInterpolation_ENNReal_pos {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} {A /-- The constant occurring in the real interpolation theorem. -/ -- todo: check order of arguments def C_realInterpolation (p₀ p₁ q₀ q₁ q : ℝ≥0∞) (C₀ C₁ A : ℝ≥0) (t : ℝ) : ℝ≥0 := - ENNReal.toNNReal (C_realInterpolation_ENNReal p₀ p₁ q₀ q₁ q C₀ C₁ A t) + C_realInterpolation_ENNReal p₀ p₁ q₀ q₁ q C₀ C₁ A t |>.toNNReal -lemma C_realInterpolation_pos {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} {A : ℝ≥0} (hA : A > 0) +lemma C_realInterpolation_pos {p₀ p₁ q₀ q₁ q : ℝ≥0∞} {A : ℝ≥0} (hA : A > 0) (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (hq₀q₁ : q₀ ≠ q₁) {C₀ C₁ : ℝ≥0} (ht : t ∈ Ioo 0 1) (hC₀ : 0 < C₀) (hC₁ : 0 < C₁) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) / p₀ + (ENNReal.ofReal t) / p₁) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) / q₀ + (ENNReal.ofReal t) / q₁) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) / q₀ + (ENNReal.ofReal t) / q₁) : 0 < C_realInterpolation p₀ p₁ q₀ q₁ q C₀ C₁ A t := by unfold C_realInterpolation refine toNNReal_pos ?_ ?_ @@ -4526,11 +4437,10 @@ lemma C_realInterpolation_pos {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} {A : ℝ≥ apply C_realInterpolation_ENNReal_pos <;> try assumption · apply C_realInterpolation_ENNReal_ne_top (A := A) <;> assumption -lemma coe_C_realInterpolation {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} {A : ℝ≥0} +lemma coe_C_realInterpolation {p₀ p₁ q₀ q₁ q : ℝ≥0∞} {A : ℝ≥0} (hp₀ : p₀ ∈ Ioc 0 q₀) (hp₁ : p₁ ∈ Ioc 0 q₁) (hq₀q₁ : q₀ ≠ q₁) {C₀ C₁ : ℝ≥0} (ht : t ∈ Ioo 0 1) (hC₀ : 0 < C₀) (hC₁ : 0 < C₁) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) / p₀ + (ENNReal.ofReal t) / p₁) - (hq : q⁻¹ = (1 - (ENNReal.ofReal t)) / q₀ + (ENNReal.ofReal t) / q₁) : + (hq : q⁻¹ = (1 - ENNReal.ofReal t) / q₀ + (ENNReal.ofReal t) / q₁) : ENNReal.ofNNReal (C_realInterpolation p₀ p₁ q₀ q₁ q C₀ C₁ A t) = C_realInterpolation_ENNReal p₀ p₁ q₀ q₁ q C₀ C₁ A t := by unfold C_realInterpolation @@ -4542,7 +4452,7 @@ lemma Subadditive_trunc_from_SubadditiveOn_Lp₀p₁ {p₀ p₁ p : ℝ≥0∞} [NormedAddCommGroup E₂] (hp₀ : p₀ > 0) (hp₁ : p₁ > 0) {A : ℝ≥0} (ht : t ∈ Ioo 0 1) - (hp : p⁻¹ = (1 - (ENNReal.ofReal t)) / p₀ + ENNReal.ofReal t / p₁) + (hp : p⁻¹ = (1 - ENNReal.ofReal t) / p₀ + ENNReal.ofReal t / p₁) (hT : SubadditiveOn T (fun f ↦ Memℒp f p₀ μ ∨ Memℒp f p₁ μ) A) (hf : Memℒp f p μ) : Subadditive_trunc T A f ν := by @@ -4581,8 +4491,7 @@ lemma Subadditive_trunc_from_SubadditiveOn_Lp₀p₁ {p₀ p₁ p : ℝ≥0∞} · exact interp_exp_ne_top (ne_of_lt p₁lt_p₀) (Ioo.one_sub_mem ht) (switch_exponents ht hp) · constructor · exact hp₁ - · exact (interp_exp_between hp₁ hp₀ p₁lt_p₀ (Ioo.one_sub_mem ht) - (switch_exponents ht hp)).1 + · exact (interp_exp_between hp₁ hp₀ p₁lt_p₀ (Ioo.one_sub_mem ht) (switch_exponents ht hp)).1 /-- Marcinkiewicz real interpolation theorem. -/ theorem exists_hasStrongType_real_interpolation {p₀ p₁ q₀ q₁ p q : ℝ≥0∞} @@ -4595,17 +4504,16 @@ theorem exists_hasStrongType_real_interpolation {p₀ p₁ q₀ q₁ p q : ℝ (hT : SubadditiveOn T (fun f ↦ Memℒp f p₀ μ ∨ Memℒp f p₁ μ) A) (h₀T : HasWeakType T p₀ q₀ μ ν C₀) (h₁T : HasWeakType T p₁ q₁ μ ν C₁) : HasStrongType T p q μ ν (C_realInterpolation p₀ p₁ q₀ q₁ q C₀ C₁ A t) := by - intro f - intro hf + intros f hf refine ⟨hmT f hf, ?_⟩ - have hp' : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + (ENNReal.ofReal t) * p₁⁻¹ := by + have hp' : p⁻¹ = (1 - ENNReal.ofReal t) * p₀⁻¹ + ENNReal.ofReal t * p₁⁻¹ := by rw [hp]; congr <;> exact Eq.symm Real.toNNReal_coe - have hq' : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + (ENNReal.ofReal t) * q₁⁻¹ := by + have hq' : q⁻¹ = (1 - ENNReal.ofReal t) * q₀⁻¹ + ENNReal.ofReal t * q₁⁻¹ := by rw [hq]; congr <;> exact Eq.symm Real.toNNReal_coe have obs : Subadditive_trunc T A f ν := Subadditive_trunc_from_SubadditiveOn_Lp₀p₁ hp₀.1 hp₁.1 ht hp' hT hf rw [coe_C_realInterpolation hp₀ hp₁ hq₀q₁] <;> try assumption - apply exists_hasStrongType_real_interpolation_aux₄ <;> try assumption + apply exists_hasStrongType_real_interpolation_aux₄ <;> assumption /- State and prove Remark 1.2.7 -/