diff --git a/Mathlib.lean b/Mathlib.lean index cf5613827a4b8..b907ebbf3fcd7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -188,6 +188,7 @@ import Mathlib.Algebra.Group.InjSurj import Mathlib.Algebra.Group.MinimalAxioms import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.Group.OrderSynonym +import Mathlib.Algebra.Group.PNatPowAssoc import Mathlib.Algebra.Group.Pi import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Semiconj.Basic @@ -2546,6 +2547,7 @@ import Mathlib.MeasureTheory.Constructions.Pi import Mathlib.MeasureTheory.Constructions.Polish import Mathlib.MeasureTheory.Constructions.Prod.Basic import Mathlib.MeasureTheory.Constructions.Prod.Integral +import Mathlib.MeasureTheory.Constructions.Projective import Mathlib.MeasureTheory.Covering.Besicovitch import Mathlib.MeasureTheory.Covering.BesicovitchVectorSpace import Mathlib.MeasureTheory.Covering.DensityTheorem @@ -2682,6 +2684,7 @@ import Mathlib.MeasureTheory.Measure.VectorMeasure import Mathlib.MeasureTheory.Measure.WithDensity import Mathlib.MeasureTheory.Measure.WithDensityVectorMeasure import Mathlib.MeasureTheory.PiSystem +import Mathlib.MeasureTheory.SetSemiring import Mathlib.MeasureTheory.Tactic import Mathlib.ModelTheory.Algebra.Field.Basic import Mathlib.ModelTheory.Algebra.Field.CharP diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean b/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean index 3ebe9757195df..27f58562ab426 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean @@ -99,7 +99,7 @@ open Nat theorem of_convergence_epsilon : ∀ ε > (0 : K), ∃ N : ℕ, ∀ n ≥ N, |v - (of v).convergents n| < ε := by intro ε ε_pos - -- use the archimedean property to obtian a suitable N + -- use the archimedean property to obtain a suitable N rcases (exists_nat_gt (1 / ε) : ∃ N' : ℕ, 1 / ε < N') with ⟨N', one_div_ε_lt_N'⟩ let N := max N' 5 -- set minimum to 5 to have N ≤ fib N work @@ -115,7 +115,7 @@ theorem of_convergence_epsilon : let nB := g.denominators (n + 1) have abs_v_sub_conv_le : |v - g.convergents n| ≤ 1 / (B * nB) := abs_sub_convergents_le not_terminated_at_n - suffices : 1 / (B * nB) < ε; exact lt_of_le_of_lt abs_v_sub_conv_le this + suffices 1 / (B * nB) < ε from lt_of_le_of_lt abs_v_sub_conv_le this -- show that `0 < (B * nB)` and then multiply by `B * nB` to get rid of the division have nB_ineq : (fib (n + 2) : K) ≤ nB := haveI : ¬g.TerminatedAt (n + 1 - 1) := not_terminated_at_n @@ -126,10 +126,10 @@ theorem of_convergence_epsilon : have zero_lt_B : 0 < B := B_ineq.trans_lt' $ mod_cast fib_pos.2 n.succ_pos have nB_pos : 0 < nB := nB_ineq.trans_lt' $ mod_cast fib_pos.2 $ succ_pos _ have zero_lt_mul_conts : 0 < B * nB := by positivity - suffices : 1 < ε * (B * nB); exact (div_lt_iff zero_lt_mul_conts).mpr this - -- use that `N ≥ n` was obtained from the archimedean property to show the following + suffices 1 < ε * (B * nB) from (div_lt_iff zero_lt_mul_conts).mpr this + -- use that `N' ≥ n` was obtained from the archimedean property to show the following calc 1 < ε * (N' : K) := (div_lt_iff' ε_pos).mp one_div_ε_lt_N' - _ ≤ ε * (B * nB) := ?_ + _ ≤ ε * (B * nB) := ?_ -- cancel `ε` gcongr calc diff --git a/Mathlib/Algebra/Group/PNatPowAssoc.lean b/Mathlib/Algebra/Group/PNatPowAssoc.lean new file mode 100644 index 0000000000000..56f7b7342fcbb --- /dev/null +++ b/Mathlib/Algebra/Group/PNatPowAssoc.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2023 Scott Carnahan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Carnahan +-/ + +import Mathlib.Algebra.Group.Defs +import Mathlib.Algebra.GroupPower.Basic +import Mathlib.Algebra.Group.Prod +import Mathlib.Data.PNat.Basic +import Mathlib.GroupTheory.GroupAction.Prod + +/-! +# Typeclasses for power-associative structures + +In this file we define power-associativity for algebraic structures with a multiplication operation. +The class is a Prop-valued mixin named `PNatPowAssoc`, where `PNat` means only strictly positive +powers are considered. + +## Results + +- `ppow_add` a defining property: `x ^ (k + n) = x ^ k * x ^ n` +- `ppow_one` a defining property: `x ^ 1 = x` +- `ppow_assoc` strictly positive powers of an element have associative multiplication. +- `ppow_comm` `x ^ m * x ^ n = x ^ n * x ^ m` for strictly positive `m` and `n`. +- `ppow_mul` `x ^ (m * n) = (x ^ m) ^ n` for strictly positive `m` and `n`. +- `ppow_eq_pow` monoid exponentiation coincides with semigroup exponentiation. + +## Instances + +- PNatPowAssoc for products and Pi types + +## Todo + +* `NatPowAssoc` for `MulOneClass` - more or less the same flow +* It seems unlikely that anyone will want `NatSMulAssoc` and `PNatSmulAssoc` as additive versions of + power-associativity, but we have found that it is not hard to write. + +-/ + +variable {M : Type*} + +/-- A `Prop`-valued mixin for power-associative multiplication in the non-unital setting. -/ +class PNatPowAssoc (M : Type*) [Mul M] [Pow M ℕ+] : Prop where + /-- Multiplication is power-associative. -/ + protected ppow_add : ∀ (k n : ℕ+) (x : M), x ^ (k + n) = x ^ k * x ^ n + /-- Exponent one is identity. -/ + protected ppow_one : ∀ (x : M), x ^ (1 : ℕ+) = x + +section Mul + +variable [Mul M] [Pow M ℕ+] [PNatPowAssoc M] + +theorem ppow_add (k n : ℕ+) (x : M) : x ^ (k + n) = x ^ k * x ^ n := + PNatPowAssoc.ppow_add k n x + +@[simp] +theorem ppow_one (x : M) : x ^ (1 : ℕ+) = x := + PNatPowAssoc.ppow_one x + +theorem ppow_mul_assoc (k m n : ℕ+) (x : M) : + (x ^ k * x ^ m) * x ^ n = x ^ k * (x ^ m * x ^ n) := by + simp only [← ppow_add, add_assoc] + +theorem ppow_mul_comm (m n : ℕ+) (x : M) : + x ^ m * x ^ n = x ^ n * x ^ m := by simp only [← ppow_add, add_comm] + +theorem ppow_mul (x : M) (m n : ℕ+) : x ^ (m * n) = (x ^ m) ^ n := by + refine PNat.recOn n ?_ fun k hk ↦ ?_ + rw [ppow_one, mul_one] + rw [ppow_add, ppow_one, mul_add, ppow_add, mul_one, hk] + +theorem ppow_mul' (x : M) (m n : ℕ+) : x ^ (m * n) = (x ^ n) ^ m := by + rw [mul_comm] + exact ppow_mul x n m + +end Mul + +instance Pi.instPNatPowAssoc {ι : Type*} {α : ι → Type*} [∀ i, Mul <| α i] [∀ i, Pow (α i) ℕ+] + [∀ i, PNatPowAssoc <| α i] : PNatPowAssoc (∀ i, α i) where + ppow_add _ _ _ := by ext; simp [ppow_add] + ppow_one _ := by ext; simp + +instance Prod.instPNatPowAssoc {N : Type*} [Mul M] [Pow M ℕ+] [PNatPowAssoc M] [Mul N] [Pow N ℕ+] + [PNatPowAssoc N] : PNatPowAssoc (M × N) where + ppow_add _ _ _ := by ext <;> simp [ppow_add] + ppow_one _ := by ext <;> simp + +theorem ppow_eq_pow [Monoid M] [Pow M ℕ+] [PNatPowAssoc M] (x : M) (n : ℕ+) : + x ^ n = x ^ (n : ℕ) := by + refine PNat.recOn n ?_ fun k hk ↦ ?_ + rw [ppow_one, PNat.one_coe, pow_one] + rw [ppow_add, ppow_one, PNat.add_coe, pow_add, PNat.one_coe, pow_one, ← hk] diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 4f8636d9e4510..25e7ceeafef91 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -603,6 +603,9 @@ theorem isBigO_refl (f : α → E) (l : Filter α) : f =O[l] f := (isBigOWith_refl f l).isBigO #align asymptotics.is_O_refl Asymptotics.isBigO_refl +theorem _root_.Filter.EventuallyEq.isBigO {f₁ f₂ : α → E} (hf : f₁ =ᶠ[l] f₂) : f₁ =O[l] f₂ := + hf.trans_isBigO (isBigO_refl _ _) + theorem IsBigOWith.trans_le (hfg : IsBigOWith c l f g) (hgk : ∀ x, ‖g x‖ ≤ ‖k x‖) (hc : 0 ≤ c) : IsBigOWith c l f k := (hfg.trans (isBigOWith_of_le l hgk) hc).congr_const <| mul_one c diff --git a/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean b/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean index db6ce46b89988..d89e402b272dd 100644 --- a/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean +++ b/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean @@ -245,8 +245,8 @@ theorem hasIntegral_GP_pderiv (f : (Fin (n + 1) → ℝ) → E) /- At a point `x ∉ s`, we unfold the definition of Fréchet differentiability, then use an estimate we proved earlier in this file. -/ rcases exists_pos_mul_lt ε0 (2 * c) with ⟨ε', ε'0, hlt⟩ - rcases (nhdsWithin_hasBasis nhds_basis_closedBall _).mem_iff.1 ((Hd x hx).def ε'0) with - ⟨δ, δ0, Hδ⟩ + rcases (nhdsWithin_hasBasis nhds_basis_closedBall _).mem_iff.1 ((Hd x hx).isLittleO.def ε'0) + with ⟨δ, δ0, Hδ⟩ refine' ⟨δ, δ0, fun J hle hJδ hxJ hJc => _⟩ simp only [BoxAdditiveMap.volume_apply, Box.volume_apply, dist_eq_norm] refine' (norm_volume_sub_integral_face_upper_sub_lower_smul_le _ diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index 8577837699cf2..0a3df02e596b3 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -262,7 +262,7 @@ theorem UniqueDiffWithinAt.eq_deriv (s : Set 𝕜) (H : UniqueDiffWithinAt 𝕜 theorem hasDerivAtFilter_iff_isLittleO : HasDerivAtFilter f f' x L ↔ (fun x' : 𝕜 => f x' - f x - (x' - x) • f') =o[L] fun x' => x' - x := - Iff.rfl + hasFDerivAtFilter_iff_isLittleO .. #align has_deriv_at_filter_iff_is_o hasDerivAtFilter_iff_isLittleO theorem hasDerivAtFilter_iff_tendsto : @@ -274,7 +274,7 @@ theorem hasDerivAtFilter_iff_tendsto : theorem hasDerivWithinAt_iff_isLittleO : HasDerivWithinAt f f' s x ↔ (fun x' : 𝕜 => f x' - f x - (x' - x) • f') =o[𝓝[s] x] fun x' => x' - x := - Iff.rfl + hasFDerivAtFilter_iff_isLittleO .. #align has_deriv_within_at_iff_is_o hasDerivWithinAt_iff_isLittleO theorem hasDerivWithinAt_iff_tendsto : @@ -285,7 +285,7 @@ theorem hasDerivWithinAt_iff_tendsto : theorem hasDerivAt_iff_isLittleO : HasDerivAt f f' x ↔ (fun x' : 𝕜 => f x' - f x - (x' - x) • f') =o[𝓝 x] fun x' => x' - x := - Iff.rfl + hasFDerivAtFilter_iff_isLittleO .. #align has_deriv_at_iff_is_o hasDerivAt_iff_isLittleO theorem hasDerivAt_iff_tendsto : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Add.lean b/Mathlib/Analysis/Calculus/FDeriv/Add.lean index 5f2cf65e0b266..1ba6ce06efccf 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Add.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Add.lean @@ -123,9 +123,9 @@ nonrec theorem HasStrictFDerivAt.add (hf : HasStrictFDerivAt f f' x) abel #align has_strict_fderiv_at.add HasStrictFDerivAt.add -nonrec theorem HasFDerivAtFilter.add (hf : HasFDerivAtFilter f f' x L) +theorem HasFDerivAtFilter.add (hf : HasFDerivAtFilter f f' x L) (hg : HasFDerivAtFilter g g' x L) : HasFDerivAtFilter (fun y => f y + g y) (f' + g') x L := - (hf.add hg).congr_left fun _ => by + .of_isLittleO <| (hf.isLittleO.add hg.isLittleO).congr_left fun _ => by simp only [LinearMap.sub_apply, LinearMap.add_apply, map_sub, map_add, add_apply] abel #align has_fderiv_at_filter.add HasFDerivAtFilter.add @@ -337,7 +337,7 @@ theorem HasStrictFDerivAt.sum (h : ∀ i ∈ u, HasStrictFDerivAt (A i) (A' i) x theorem HasFDerivAtFilter.sum (h : ∀ i ∈ u, HasFDerivAtFilter (A i) (A' i) x L) : HasFDerivAtFilter (fun y => ∑ i in u, A i y) (∑ i in u, A' i) x L := by - dsimp [HasFDerivAtFilter] at * + simp only [hasFDerivAtFilter_iff_isLittleO] at * convert IsLittleO.sum h simp [ContinuousLinearMap.sum_apply] #align has_fderiv_at_filter.sum HasFDerivAtFilter.sum diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index 8a546f66dd6d1..00478ed34a603 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -141,8 +141,9 @@ variable {G' : Type*} [NormedAddCommGroup G'] [NormedSpace 𝕜 G'] is designed to be specialized for `L = 𝓝 x` (in `HasFDerivAt`), giving rise to the usual notion of Fréchet derivative, and for `L = 𝓝[s] x` (in `HasFDerivWithinAt`), giving rise to the notion of Fréchet derivative along the set `s`. -/ -def HasFDerivAtFilter (f : E → F) (f' : E →L[𝕜] F) (x : E) (L : Filter E) := - (fun x' => f x' - f x - f' (x' - x)) =o[L] fun x' => x' - x +@[mk_iff hasFDerivAtFilter_iff_isLittleO] +structure HasFDerivAtFilter (f : E → F) (f' : E →L[𝕜] F) (x : E) (L : Filter E) : Prop where + of_isLittleO :: isLittleO : (fun x' => f x' - f x - f' (x' - x)) =o[L] fun x' => x' - x #align has_fderiv_at_filter HasFDerivAtFilter /-- A function `f` has the continuous linear map `f'` as derivative at `x` within a set `s` if @@ -257,7 +258,7 @@ theorem HasFDerivWithinAt.lim (h : HasFDerivWithinAt f f' s x) {α : Type*} (l : constructor · apply tendsto_const_nhds.add (tangentConeAt.lim_zero l clim cdlim) · rwa [tendsto_principal] - have : (fun y => f y - f x - f' (y - x)) =o[𝓝[s] x] fun y => y - x := h + have : (fun y => f y - f x - f' (y - x)) =o[𝓝[s] x] fun y => y - x := h.isLittleO have : (fun n => f (x + d n) - f x - f' (x + d n - x)) =o[l] fun n => x + d n - x := this.comp_tendsto tendsto_arg have : (fun n => f (x + d n) - f x - f' (d n)) =o[l] d := by simpa only [add_sub_cancel'] @@ -312,8 +313,8 @@ theorem hasFDerivAtFilter_iff_tendsto : have h : ∀ x', ‖x' - x‖ = 0 → ‖f x' - f x - f' (x' - x)‖ = 0 := fun x' hx' => by rw [sub_eq_zero.1 (norm_eq_zero.1 hx')] simp - unfold HasFDerivAtFilter - rw [← isLittleO_norm_left, ← isLittleO_norm_right, isLittleO_iff_tendsto h] + rw [hasFDerivAtFilter_iff_isLittleO, ← isLittleO_norm_left, ← isLittleO_norm_right, + isLittleO_iff_tendsto h] exact tendsto_congr fun _ => div_eq_inv_mul _ _ #align has_fderiv_at_filter_iff_tendsto hasFDerivAtFilter_iff_tendsto @@ -330,7 +331,7 @@ theorem hasFDerivAt_iff_tendsto : theorem hasFDerivAt_iff_isLittleO_nhds_zero : HasFDerivAt f f' x ↔ (fun h : E => f (x + h) - f x - f' h) =o[𝓝 0] fun h => h := by - rw [HasFDerivAt, HasFDerivAtFilter, ← map_add_left_nhds_zero x, isLittleO_map] + rw [HasFDerivAt, hasFDerivAtFilter_iff_isLittleO, ← map_add_left_nhds_zero x, isLittleO_map] simp [(· ∘ ·)] #align has_fderiv_at_iff_is_o_nhds_zero hasFDerivAt_iff_isLittleO_nhds_zero @@ -368,7 +369,7 @@ theorem HasFDerivAt.le_of_lipschitz {f : E → F} {f' : E →L[𝕜] F} {x₀ : nonrec theorem HasFDerivAtFilter.mono (h : HasFDerivAtFilter f f' x L₂) (hst : L₁ ≤ L₂) : HasFDerivAtFilter f f' x L₁ := - h.mono hst + .of_isLittleO <| h.isLittleO.mono hst #align has_fderiv_at_filter.mono HasFDerivAtFilter.mono theorem HasFDerivWithinAt.mono_of_mem (h : HasFDerivWithinAt f f' t x) (hst : t ∈ 𝓝[s] x) : @@ -420,7 +421,7 @@ lemma hasFDerivWithinAt_of_isOpen (h : IsOpen s) (hx : x ∈ s) : theorem hasFDerivWithinAt_insert {y : E} : HasFDerivWithinAt f f' (insert y s) x ↔ HasFDerivWithinAt f f' s x := by rcases eq_or_ne x y with (rfl | h) - · simp_rw [HasFDerivWithinAt, HasFDerivAtFilter] + · simp_rw [HasFDerivWithinAt, hasFDerivAtFilter_iff_isLittleO] apply Asymptotics.isLittleO_insert simp only [sub_self, map_zero] refine' ⟨fun h => h.mono <| subset_insert y s, fun hf => hf.mono_of_mem _⟩ @@ -449,13 +450,13 @@ set_option linter.uppercaseLean3 false in theorem HasFDerivAtFilter.isBigO_sub (h : HasFDerivAtFilter f f' x L) : (fun x' => f x' - f x) =O[L] fun x' => x' - x := - h.isBigO.congr_of_sub.2 (f'.isBigO_sub _ _) + h.isLittleO.isBigO.congr_of_sub.2 (f'.isBigO_sub _ _) set_option linter.uppercaseLean3 false in #align has_fderiv_at_filter.is_O_sub HasFDerivAtFilter.isBigO_sub protected theorem HasStrictFDerivAt.hasFDerivAt (hf : HasStrictFDerivAt f f' x) : HasFDerivAt f f' x := by - rw [HasFDerivAt, HasFDerivAtFilter, isLittleO_iff] + rw [HasFDerivAt, hasFDerivAtFilter_iff_isLittleO, isLittleO_iff] exact fun c hc => tendsto_id.prod_mk_nhds tendsto_const_nhds (isLittleO_iff.1 hf hc) #align has_strict_fderiv_at.has_fderiv_at HasStrictFDerivAt.hasFDerivAt @@ -513,7 +514,7 @@ theorem hasFDerivWithinAt_inter (h : t ∈ 𝓝 x) : theorem HasFDerivWithinAt.union (hs : HasFDerivWithinAt f f' s x) (ht : HasFDerivWithinAt f f' t x) : HasFDerivWithinAt f f' (s ∪ t) x := by simp only [HasFDerivWithinAt, nhdsWithin_union] - exact hs.sup ht + exact .of_isLittleO <| hs.isLittleO.sup ht.isLittleO #align has_fderiv_within_at.union HasFDerivWithinAt.union theorem HasFDerivWithinAt.hasFDerivAt (h : HasFDerivWithinAt f f' s x) (hs : s ∈ 𝓝 x) : @@ -530,7 +531,7 @@ theorem DifferentiableWithinAt.differentiableAt (h : DifferentiableWithinAt 𝕜 as this statement is empty. -/ theorem HasFDerivWithinAt.of_nhdsWithin_eq_bot (h : 𝓝[s\{x}] x = ⊥) : HasFDerivWithinAt f f' s x := by - rw [← hasFDerivWithinAt_diff_singleton x, HasFDerivWithinAt, h] + rw [← hasFDerivWithinAt_diff_singleton x, HasFDerivWithinAt, h, hasFDerivAtFilter_iff_isLittleO] apply isLittleO_bot /-- If `x` is not in the closure of `s`, then `f` has any derivative at `x` within `s`, @@ -735,7 +736,7 @@ theorem fderivWithin_mem_iff {f : E → F} {t : Set E} {s : Set (E →L[𝕜] F) theorem Asymptotics.IsBigO.hasFDerivWithinAt {s : Set E} {x₀ : E} {n : ℕ} (h : f =O[𝓝[s] x₀] fun x => ‖x - x₀‖ ^ n) (hx₀ : x₀ ∈ s) (hn : 1 < n) : HasFDerivWithinAt f (0 : E →L[𝕜] F) s x₀ := by - simp_rw [HasFDerivWithinAt, HasFDerivAtFilter, + simp_rw [HasFDerivWithinAt, hasFDerivAtFilter_iff_isLittleO, h.eq_zero_of_norm_pow_within hx₀ <| zero_lt_one.trans hn, zero_apply, sub_zero, h.trans_isLittleO ((isLittleO_pow_sub_sub x₀ hn).mono nhdsWithin_le_nhds)] set_option linter.uppercaseLean3 false in @@ -831,7 +832,7 @@ theorem HasFDerivAtFilter.isBigO_sub_rev (hf : HasFDerivAtFilter f f' x L) {C} (hf' : AntilipschitzWith C f') : (fun x' => x' - x) =O[L] fun x' => f x' - f x := have : (fun x' => x' - x) =O[L] fun x' => f' (x' - x) := isBigO_iff.2 ⟨C, eventually_of_forall fun _ => ZeroHomClass.bound_of_antilipschitz f' hf' _⟩ - (this.trans (hf.trans_isBigO this).right_isBigO_add).congr (fun _ => rfl) fun _ => + (this.trans (hf.isLittleO.trans_isBigO this).right_isBigO_add).congr (fun _ => rfl) fun _ => sub_add_cancel _ _ set_option linter.uppercaseLean3 false in #align has_fderiv_at_filter.is_O_sub_rev HasFDerivAtFilter.isBigO_sub_rev @@ -915,9 +916,9 @@ theorem HasStrictFDerivAt.congr_of_eventuallyEq (h : HasStrictFDerivAt f f' x) #align has_strict_fderiv_at.congr_of_eventually_eq HasStrictFDerivAt.congr_of_eventuallyEq theorem Filter.EventuallyEq.hasFDerivAtFilter_iff (h₀ : f₀ =ᶠ[L] f₁) (hx : f₀ x = f₁ x) - (h₁ : ∀ x, f₀' x = f₁' x) : HasFDerivAtFilter f₀ f₀' x L ↔ HasFDerivAtFilter f₁ f₁' x L := - isLittleO_congr (h₀.mono fun y hy => by simp only [hy, h₁, hx]) - (eventually_of_forall fun _ => _root_.rfl) + (h₁ : ∀ x, f₀' x = f₁' x) : HasFDerivAtFilter f₀ f₀' x L ↔ HasFDerivAtFilter f₁ f₁' x L := by + simp only [hasFDerivAtFilter_iff_isLittleO] + exact isLittleO_congr (h₀.mono fun y hy => by simp only [hy, h₁, hx]) .rfl #align filter.eventually_eq.has_fderiv_at_filter_iff Filter.EventuallyEq.hasFDerivAtFilter_iff theorem HasFDerivAtFilter.congr_of_eventuallyEq (h : HasFDerivAtFilter f f' x L) (hL : f₁ =ᶠ[L] f) @@ -1073,7 +1074,7 @@ theorem hasStrictFDerivAt_id (x : E) : HasStrictFDerivAt id (id 𝕜 E) x := #align has_strict_fderiv_at_id hasStrictFDerivAt_id theorem hasFDerivAtFilter_id (x : E) (L : Filter E) : HasFDerivAtFilter id (id 𝕜 E) x L := - (isLittleO_zero _ _).congr_left <| by simp + .of_isLittleO <| (isLittleO_zero _ _).congr_left <| by simp #align has_fderiv_at_filter_id hasFDerivAtFilter_id theorem hasFDerivWithinAt_id (x : E) (s : Set E) : HasFDerivWithinAt id (id 𝕜 E) s x := @@ -1143,7 +1144,7 @@ theorem hasStrictFDerivAt_const (c : F) (x : E) : theorem hasFDerivAtFilter_const (c : F) (x : E) (L : Filter E) : HasFDerivAtFilter (fun _ => c) (0 : E →L[𝕜] F) x L := - (isLittleO_zero _ _).congr_left fun _ => by simp only [zero_apply, sub_self] + .of_isLittleO <| (isLittleO_zero _ _).congr_left fun _ => by simp only [zero_apply, sub_self] #align has_fderiv_at_filter_const hasFDerivAtFilter_const theorem hasFDerivWithinAt_const (c : F) (x : E) (s : Set E) : @@ -1192,8 +1193,8 @@ theorem differentiableOn_const (c : F) : DifferentiableOn 𝕜 (fun _ => c) s := theorem hasFDerivWithinAt_singleton (f : E → F) (x : E) : HasFDerivWithinAt f (0 : E →L[𝕜] F) {x} x := by - simp only [HasFDerivWithinAt, nhdsWithin_singleton, HasFDerivAtFilter, isLittleO_pure, - ContinuousLinearMap.zero_apply, sub_self] + simp only [HasFDerivWithinAt, nhdsWithin_singleton, hasFDerivAtFilter_iff_isLittleO, + isLittleO_pure, ContinuousLinearMap.zero_apply, sub_self] #align has_fderiv_within_at_singleton hasFDerivWithinAt_singleton theorem hasFDerivAt_of_subsingleton [h : Subsingleton E] (f : E → F) (x : E) : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean index 2b32baaa9fb41..25fe25066b5dc 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean @@ -62,27 +62,26 @@ variable (x) theorem HasFDerivAtFilter.comp {g : F → G} {g' : F →L[𝕜] G} {L' : Filter F} (hg : HasFDerivAtFilter g g' (f x) L') (hf : HasFDerivAtFilter f f' x L) (hL : Tendsto f L L') : HasFDerivAtFilter (g ∘ f) (g'.comp f') x L := by - let eq₁ := (g'.isBigO_comp _ _).trans_isLittleO hf - let eq₂ := (hg.comp_tendsto hL).trans_isBigO hf.isBigO_sub - refine' eq₂.triangle (eq₁.congr_left fun x' => _) + let eq₁ := (g'.isBigO_comp _ _).trans_isLittleO hf.isLittleO + let eq₂ := (hg.isLittleO.comp_tendsto hL).trans_isBigO hf.isBigO_sub + refine .of_isLittleO <| eq₂.triangle <| eq₁.congr_left fun x' => ?_ simp #align has_fderiv_at_filter.comp HasFDerivAtFilter.comp /- A readable version of the previous theorem, a general form of the chain rule. -/ example {g : F → G} {g' : F →L[𝕜] G} (hg : HasFDerivAtFilter g g' (f x) (L.map f)) (hf : HasFDerivAtFilter f f' x L) : HasFDerivAtFilter (g ∘ f) (g'.comp f') x L := by - unfold HasFDerivAtFilter at hg have := calc (fun x' => g (f x') - g (f x) - g' (f x' - f x)) =o[L] fun x' => f x' - f x := - hg.comp_tendsto le_rfl + hg.isLittleO.comp_tendsto le_rfl _ =O[L] fun x' => x' - x := hf.isBigO_sub - refine' this.triangle _ + refine' .of_isLittleO <| this.triangle _ calc (fun x' : E => g' (f x' - f x) - g'.comp f' (x' - x)) _ =ᶠ[L] fun x' => g' (f x' - f x - f' (x' - x)) := eventually_of_forall fun x' => by simp _ =O[L] fun x' => f x' - f x - f' (x' - x) := (g'.isBigO_comp _ _) - _ =o[L] fun x' => x' - x := hf + _ =o[L] fun x' => x' - x := hf.isLittleO theorem HasFDerivWithinAt.comp {g : F → G} {g' : F →L[𝕜] G} {t : Set F} (hg : HasFDerivWithinAt g g' t (f x)) (hf : HasFDerivWithinAt f f' s x) (hst : MapsTo f s t) : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean index 257012e267825..5bd75dc053eed 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean @@ -393,9 +393,9 @@ theorem HasFDerivAt.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] F} {g fun x : F => f' (g x - g a) - (x - a) := by refine' ((f'.symm : F →L[𝕜] E).isBigO_comp _ _).congr (fun x => _) fun _ => rfl simp - refine' this.trans_isLittleO _ + refine HasFDerivAtFilter.of_isLittleO <| this.trans_isLittleO ?_ clear this - refine ((hf.comp_tendsto hg).symm.congr' (hfg.mono ?_) .rfl).trans_isBigO ?_ + refine ((hf.isLittleO.comp_tendsto hg).symm.congr' (hfg.mono ?_) .rfl).trans_isBigO ?_ · intro p hp simp [hp, hfg.self_of_nhds] · refine' ((hf.isBigO_sub_rev f'.antilipschitz).comp_tendsto hg).congr' @@ -432,7 +432,7 @@ theorem HasFDerivWithinAt.eventually_ne (h : HasFDerivWithinAt f f' s x) rw [nhdsWithin, diff_eq, ← inf_principal, ← inf_assoc, eventually_inf_principal] have A : (fun z => z - x) =O[𝓝[s] x] fun z => f' (z - x) := isBigO_iff.2 <| hf'.imp fun C hC => eventually_of_forall fun z => hC _ - have : (fun z => f z - f x) ~[𝓝[s] x] fun z => f' (z - x) := h.trans_isBigO A + have : (fun z => f z - f x) ~[𝓝[s] x] fun z => f' (z - x) := h.isLittleO.trans_isBigO A simpa [not_imp_not, sub_eq_zero] using (A.trans this.isBigO_symm).eq_zero_imp #align has_fderiv_within_at.eventually_ne HasFDerivWithinAt.eventually_ne diff --git a/Mathlib/Analysis/Calculus/FDeriv/Extend.lean b/Mathlib/Analysis/Calculus/FDeriv/Extend.lean index cb148441d143b..e2859c8d2b745 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Extend.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Extend.lean @@ -45,7 +45,7 @@ theorem has_fderiv_at_boundary_of_tendsto_fderiv {f : E → F} {s : Set E} {x : by_cases hx : x ∉ closure s · rw [← closure_closure] at hx; exact hasFDerivWithinAt_of_nmem_closure hx push_neg at hx - rw [HasFDerivWithinAt, HasFDerivAtFilter, Asymptotics.isLittleO_iff] + rw [HasFDerivWithinAt, hasFDerivAtFilter_iff_isLittleO, Asymptotics.isLittleO_iff] /- One needs to show that `‖f y - f x - f' (y - x)‖ ≤ ε ‖y - x‖` for `y` close to `x` in `closure s`, where `ε` is an arbitrary positive constant. By continuity of the functions, it suffices to prove this for nearby points inside `s`. In a neighborhood of `x`, the derivative diff --git a/Mathlib/Analysis/Calculus/FDeriv/Linear.lean b/Mathlib/Analysis/Calculus/FDeriv/Linear.lean index 8295330fe567b..5c3ad26b156f4 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Linear.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Linear.lean @@ -64,7 +64,7 @@ protected theorem ContinuousLinearMap.hasStrictFDerivAt {x : E} : HasStrictFDeri #align continuous_linear_map.has_strict_fderiv_at ContinuousLinearMap.hasStrictFDerivAt protected theorem ContinuousLinearMap.hasFDerivAtFilter : HasFDerivAtFilter e e x L := - (isLittleO_zero _ _).congr_left fun x => by simp only [e.map_sub, sub_self] + .of_isLittleO <| (isLittleO_zero _ _).congr_left fun x => by simp only [e.map_sub, sub_self] #align continuous_linear_map.has_fderiv_at_filter ContinuousLinearMap.hasFDerivAtFilter protected theorem ContinuousLinearMap.hasFDerivWithinAt : HasFDerivWithinAt e e s x := diff --git a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean index 1e45579362f25..f68b4dd71db0b 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean @@ -163,22 +163,21 @@ theorem le_of_mem_A {r ε : ℝ} {L : E →L[𝕜] F} {x : E} (hx : x ∈ A f L theorem mem_A_of_differentiable {ε : ℝ} (hε : 0 < ε) {x : E} (hx : DifferentiableAt 𝕜 f x) : ∃ R > 0, ∀ r ∈ Ioo (0 : ℝ) R, x ∈ A f (fderiv 𝕜 f x) r ε := by - have := hx.hasFDerivAt - simp only [HasFDerivAt, HasFDerivAtFilter, isLittleO_iff] at this let δ := (ε / 2) / 2 - have hδ : 0 < δ := by positivity - rcases eventually_nhds_iff_ball.1 (this hδ) with ⟨R, R_pos, hR⟩ + obtain ⟨R, R_pos, hR⟩ : + ∃ R > 0, ∀ y ∈ ball x R, ‖f y - f x - fderiv 𝕜 f x (y - x)‖ ≤ δ * ‖y - x‖ := + eventually_nhds_iff_ball.1 <| hx.hasFDerivAt.isLittleO.bound <| by positivity refine' ⟨R, R_pos, fun r hr => _⟩ - have : r ∈ Ioc (r / 2) r := ⟨half_lt_self hr.1, le_rfl⟩ + have : r ∈ Ioc (r / 2) r := right_mem_Ioc.2 <| half_lt_self hr.1 refine' ⟨r, this, fun y hy z hz => _⟩ calc ‖f z - f y - (fderiv 𝕜 f x) (z - y)‖ = ‖f z - f x - (fderiv 𝕜 f x) (z - x) - (f y - f x - (fderiv 𝕜 f x) (y - x))‖ := - by congr 1; simp only [ContinuousLinearMap.map_sub]; abel + by simp only [map_sub]; abel_nf _ ≤ ‖f z - f x - (fderiv 𝕜 f x) (z - x)‖ + ‖f y - f x - (fderiv 𝕜 f x) (y - x)‖ := - (norm_sub_le _ _) + norm_sub_le _ _ _ ≤ δ * ‖z - x‖ + δ * ‖y - x‖ := - (add_le_add (hR _ (lt_trans (mem_ball.1 hz) hr.2)) (hR _ (lt_trans (mem_ball.1 hy) hr.2))) + add_le_add (hR _ (ball_subset_ball hr.2.le hz)) (hR _ (ball_subset_ball hr.2.le hy)) _ ≤ δ * r + δ * r := by rw [mem_ball_iff_norm] at hz hy; gcongr _ = (ε / 2) * r := by ring _ < ε * r := by gcongr; exacts [hr.1, half_lt_self hε] diff --git a/Mathlib/Analysis/Calculus/FDeriv/Prod.lean b/Mathlib/Analysis/Calculus/FDeriv/Prod.lean index 8241c33847598..1dc744fb39f0e 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Prod.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Prod.lean @@ -67,7 +67,7 @@ protected theorem HasStrictFDerivAt.prod (hf₁ : HasStrictFDerivAt f₁ f₁' x theorem HasFDerivAtFilter.prod (hf₁ : HasFDerivAtFilter f₁ f₁' x L) (hf₂ : HasFDerivAtFilter f₂ f₂' x L) : HasFDerivAtFilter (fun x => (f₁ x, f₂ x)) (f₁'.prod f₂') x L := - hf₁.prod_left hf₂ + .of_isLittleO <| hf₁.isLittleO.prod_left hf₂.isLittleO #align has_fderiv_at_filter.prod HasFDerivAtFilter.prod nonrec theorem HasFDerivWithinAt.prod (hf₁ : HasFDerivWithinAt f₁ f₁' s x) @@ -391,7 +391,7 @@ theorem hasStrictFDerivAt_pi : theorem hasFDerivAtFilter_pi' : HasFDerivAtFilter Φ Φ' x L ↔ ∀ i, HasFDerivAtFilter (fun x => Φ x i) ((proj i).comp Φ') x L := by - simp only [HasFDerivAtFilter, ContinuousLinearMap.coe_pi] + simp only [hasFDerivAtFilter_iff_isLittleO, ContinuousLinearMap.coe_pi] exact isLittleO_pi #align has_fderiv_at_filter_pi' hasFDerivAtFilter_pi' diff --git a/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean b/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean index c40472a04f6e5..9131e5db3ac8b 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean @@ -56,17 +56,17 @@ theorem HasStrictFDerivAt.restrictScalars (h : HasStrictFDerivAt f f' x) : theorem HasFDerivAtFilter.restrictScalars {L} (h : HasFDerivAtFilter f f' x L) : HasFDerivAtFilter f (f'.restrictScalars 𝕜) x L := - h + .of_isLittleO h.1 #align has_fderiv_at_filter.restrict_scalars HasFDerivAtFilter.restrictScalars theorem HasFDerivAt.restrictScalars (h : HasFDerivAt f f' x) : HasFDerivAt f (f'.restrictScalars 𝕜) x := - h + .of_isLittleO h.1 #align has_fderiv_at.restrict_scalars HasFDerivAt.restrictScalars theorem HasFDerivWithinAt.restrictScalars (h : HasFDerivWithinAt f f' s x) : HasFDerivWithinAt f (f'.restrictScalars 𝕜) s x := - h + .of_isLittleO h.1 #align has_fderiv_within_at.restrict_scalars HasFDerivWithinAt.restrictScalars theorem DifferentiableAt.restrictScalars (h : DifferentiableAt 𝕜' f x) : DifferentiableAt 𝕜 f x := @@ -86,16 +86,16 @@ theorem Differentiable.restrictScalars (h : Differentiable 𝕜' f) : Differenti (h x).restrictScalars 𝕜 #align differentiable.restrict_scalars Differentiable.restrictScalars -theorem hasFDerivWithinAt_of_restrictScalars {g' : E →L[𝕜] F} (h : HasFDerivWithinAt f g' s x) +theorem HasFDerivWithinAt.of_restrictScalars {g' : E →L[𝕜] F} (h : HasFDerivWithinAt f g' s x) (H : f'.restrictScalars 𝕜 = g') : HasFDerivWithinAt f f' s x := by rw [← H] at h - exact h -#align has_fderiv_within_at_of_restrict_scalars hasFDerivWithinAt_of_restrictScalars + exact .of_isLittleO h.1 +#align has_fderiv_within_at_of_restrict_scalars HasFDerivWithinAt.of_restrictScalars theorem hasFDerivAt_of_restrictScalars {g' : E →L[𝕜] F} (h : HasFDerivAt f g' x) (H : f'.restrictScalars 𝕜 = g') : HasFDerivAt f f' x := by rw [← H] at h - exact h + exact .of_isLittleO h.1 #align has_fderiv_at_of_restrict_scalars hasFDerivAt_of_restrictScalars theorem DifferentiableAt.fderiv_restrictScalars (h : DifferentiableAt 𝕜' f x) : @@ -110,7 +110,7 @@ theorem differentiableWithinAt_iff_restrictScalars (hf : DifferentiableWithinAt · rintro ⟨g', hg'⟩ exact ⟨g', hs.eq (hg'.restrictScalars 𝕜) hf.hasFDerivWithinAt⟩ · rintro ⟨f', hf'⟩ - exact ⟨f', hasFDerivWithinAt_of_restrictScalars 𝕜 hf.hasFDerivWithinAt hf'⟩ + exact ⟨f', hf.hasFDerivWithinAt.of_restrictScalars 𝕜 hf'⟩ #align differentiable_within_at_iff_restrict_scalars differentiableWithinAt_iff_restrictScalars theorem differentiableAt_iff_restrictScalars (hf : DifferentiableAt 𝕜 f x) : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean index 9a9643591c762..6e17b9d117492 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean @@ -76,7 +76,7 @@ theorem Convex.taylor_approx_two_segment {v w : E} (hv : x + v ∈ interior s) (isLittleO_iff.2 fun ε εpos => _) (isBigO_const_mul_self ((‖v‖ + ‖w‖) * ‖w‖) _ _) -- consider a ball of radius `δ` around `x` in which the Taylor approximation for `f''` is -- good up to `δ`. - rw [HasFDerivWithinAt, HasFDerivAtFilter, isLittleO_iff] at hx + rw [HasFDerivWithinAt, hasFDerivAtFilter_iff_isLittleO, isLittleO_iff] at hx rcases Metric.mem_nhdsWithin_iff.1 (hx εpos) with ⟨δ, δpos, sδ⟩ have E1 : ∀ᶠ h in 𝓝[>] (0 : ℝ), h * (‖v‖ + ‖w‖) < δ := by have : Filter.Tendsto (fun h => h * (‖v‖ + ‖w‖)) (𝓝[>] (0 : ℝ)) (𝓝 (0 * (‖v‖ + ‖w‖))) := diff --git a/Mathlib/Analysis/Calculus/Gradient/Basic.lean b/Mathlib/Analysis/Calculus/Gradient/Basic.lean index d7d14e5c2e825..61c7eda849864 100644 --- a/Mathlib/Analysis/Calculus/Gradient/Basic.lean +++ b/Mathlib/Analysis/Calculus/Gradient/Basic.lean @@ -211,12 +211,12 @@ section GradientProperties theorem hasGradientAtFilter_iff_isLittleO : HasGradientAtFilter f f' x L ↔ (fun x' : F => f x' - f x - ⟪f', x' - x⟫) =o[L] fun x' => x' - x := - Iff.rfl + hasFDerivAtFilter_iff_isLittleO .. theorem hasGradientWithinAt_iff_isLittleO : HasGradientWithinAt f f' s x ↔ (fun x' : F => f x' - f x - ⟪f', x' - x⟫) =o[𝓝[s] x] fun x' => x' - x := - Iff.rfl + hasGradientAtFilter_iff_isLittleO theorem hasGradientWithinAt_iff_tendsto : HasGradientWithinAt f f' s x ↔ @@ -225,7 +225,7 @@ theorem hasGradientWithinAt_iff_tendsto : theorem hasGradientAt_iff_isLittleO : HasGradientAt f f' x ↔ (fun x' : F => f x' - f x - ⟪f', x' - x⟫) =o[𝓝 x] fun x' => x' - x := - Iff.rfl + hasGradientAtFilter_iff_isLittleO theorem hasGradientAt_iff_tendsto : HasGradientAt f f' x ↔ diff --git a/Mathlib/Analysis/Convex/Hull.lean b/Mathlib/Analysis/Convex/Hull.lean index 7815423e211aa..4e3454c157cb4 100644 --- a/Mathlib/Analysis/Convex/Hull.lean +++ b/Mathlib/Analysis/Convex/Hull.lean @@ -40,12 +40,8 @@ variable (𝕜) variable [AddCommMonoid E] [AddCommMonoid F] [Module 𝕜 E] [Module 𝕜 F] /-- The convex hull of a set `s` is the minimal convex set that includes `s`. -/ -def convexHull : ClosureOperator (Set E) := - ClosureOperator.mk₃ (fun s => ⋂ (t : Set E) (_ : s ⊆ t) (_ : Convex 𝕜 t), t) (Convex 𝕜) - (fun _ => - Set.subset_iInter fun _ => Set.subset_iInter fun hst => Set.subset_iInter fun _ => hst) - (fun _ => convex_iInter fun _ => convex_iInter fun _ => convex_iInter id) fun _ t hst ht => - Set.iInter_subset_of_subset t <| Set.iInter_subset_of_subset hst <| Set.iInter_subset _ ht +@[simps! isClosed] +def convexHull : ClosureOperator (Set E) := .ofCompletePred (Convex 𝕜) fun _ ↦ convex_sInter #align convex_hull convexHull variable (s : Set E) @@ -54,13 +50,11 @@ theorem subset_convexHull : s ⊆ convexHull 𝕜 s := (convexHull 𝕜).le_closure s #align subset_convex_hull subset_convexHull -theorem convex_convexHull : Convex 𝕜 (convexHull 𝕜 s) := - ClosureOperator.closure_mem_mk₃ s +theorem convex_convexHull : Convex 𝕜 (convexHull 𝕜 s) := (convexHull 𝕜).isClosed_closure s #align convex_convex_hull convex_convexHull -theorem convexHull_eq_iInter : convexHull 𝕜 s = - ⋂ (t : Set E) (_ : s ⊆ t) (_ : Convex 𝕜 t), t := - rfl +theorem convexHull_eq_iInter : convexHull 𝕜 s = ⋂ (t : Set E) (_ : s ⊆ t) (_ : Convex 𝕜 t), t := by + simp [convexHull, iInter_subtype, iInter_and] #align convex_hull_eq_Inter convexHull_eq_iInter variable {𝕜 s} {t : Set E} {x y : E} @@ -69,12 +63,11 @@ theorem mem_convexHull_iff : x ∈ convexHull 𝕜 s ↔ ∀ t, s ⊆ t → Conv simp_rw [convexHull_eq_iInter, mem_iInter] #align mem_convex_hull_iff mem_convexHull_iff -theorem convexHull_min (hst : s ⊆ t) (ht : Convex 𝕜 t) : convexHull 𝕜 s ⊆ t := - ClosureOperator.closure_le_mk₃_iff (show s ≤ t from hst) ht +theorem convexHull_min : s ⊆ t → Convex 𝕜 t → convexHull 𝕜 s ⊆ t := (convexHull 𝕜).closure_min #align convex_hull_min convexHull_min theorem Convex.convexHull_subset_iff (ht : Convex 𝕜 t) : convexHull 𝕜 s ⊆ t ↔ s ⊆ t := - ⟨(subset_convexHull _ _).trans, fun h => convexHull_min h ht⟩ + (show (convexHull 𝕜).IsClosed t from ht).closure_le_iff #align convex.convex_hull_subset_iff Convex.convexHull_subset_iff @[mono] @@ -82,7 +75,9 @@ theorem convexHull_mono (hst : s ⊆ t) : convexHull 𝕜 s ⊆ convexHull 𝕜 (convexHull 𝕜).monotone hst #align convex_hull_mono convexHull_mono -theorem Convex.convexHull_eq : Convex 𝕜 s → convexHull 𝕜 s = s := ClosureOperator.mem_mk₃_closed.2 +lemma convexHull_eq_self : convexHull 𝕜 s = s ↔ Convex 𝕜 s := (convexHull 𝕜).isClosed_iff.symm + +alias ⟨_, Convex.convexHull_eq⟩ := convexHull_eq_self #align convex.convex_hull_eq Convex.convexHull_eq @[simp] diff --git a/Mathlib/Analysis/MellinTransform.lean b/Mathlib/Analysis/MellinTransform.lean index 5cdb0a9d8a64a..c9d67860f041b 100644 --- a/Mathlib/Analysis/MellinTransform.lean +++ b/Mathlib/Analysis/MellinTransform.lean @@ -99,8 +99,6 @@ theorem MellinConvergent.comp_rpow {f : ℝ → E} {s : ℂ} {a : ℝ} (ha : a add_sub_assoc, sub_add_cancel] #align mellin_convergent.comp_rpow MellinConvergent.comp_rpow -variable [CompleteSpace E] - /-- The Mellin transform of a function `f` (for a complex exponent `s`), defined as the integral of `t ^ (s - 1) • f` over `Ioi 0`. -/ def mellin (f : ℝ → E) (s : ℂ) : E := @@ -123,10 +121,14 @@ theorem mellin_div_const (f : ℝ → ℂ) (s a : ℂ) : mellin (fun t => f t / simp_rw [mellin, smul_eq_mul, ← mul_div_assoc, integral_div] #align mellin_div_const mellin_div_const -theorem mellin_comp_rpow (f : ℝ → E) (s : ℂ) {a : ℝ} (ha : a ≠ 0) : +theorem mellin_comp_rpow (f : ℝ → E) (s : ℂ) (a : ℝ) : mellin (fun t => f (t ^ a)) s = |a|⁻¹ • mellin f (s / a) := by - -- note: this is also true for a = 0 (both sides are zero), but this is mathematically - -- uninteresting and rather time-consuming to check + /- This is true for `a = 0` as all sides are undefined but turn out to vanish thanks to our + convention. The interesting case is `a ≠ 0` -/ + rcases eq_or_ne a 0 with rfl|ha + · by_cases hE : CompleteSpace E + · simp [integral_smul_const, mellin, setIntegral_Ioi_zero_cpow] + · simp [integral, mellin, hE] simp_rw [mellin] conv_rhs => rw [← integral_comp_rpow_Ioi _ ha, ← integral_smul] refine' set_integral_congr measurableSet_Ioi fun t ht => _ @@ -162,7 +164,7 @@ theorem mellin_comp_mul_right (f : ℝ → E) (s : ℂ) {a : ℝ} (ha : 0 < a) : #align mellin_comp_mul_right mellin_comp_mul_right theorem mellin_comp_inv (f : ℝ → E) (s : ℂ) : mellin (fun t => f t⁻¹) s = mellin f (-s) := by - simp_rw [← rpow_neg_one, mellin_comp_rpow _ _ (neg_ne_zero.mpr one_ne_zero), abs_neg, abs_one, + simp_rw [← rpow_neg_one, mellin_comp_rpow _ _ _, abs_neg, abs_one, inv_one, one_smul, ofReal_neg, ofReal_one, div_neg, div_one] #align mellin_comp_inv mellin_comp_inv @@ -250,7 +252,7 @@ theorem mellin_convergent_zero_of_isBigO {b : ℝ} {f : ℝ → ℝ} · show HasFiniteIntegral (fun t => d * t ^ (s - b - 1)) _ refine' (Integrable.hasFiniteIntegral _).const_mul _ rw [← IntegrableOn, ← integrableOn_Ioc_iff_integrableOn_Ioo, ← - intervalIntegrable_iff_integrable_Ioc_of_le hε.le] + intervalIntegrable_iff_integrableOn_Ioc_of_le hε.le] exact intervalIntegral.intervalIntegrable_rpow' (by linarith) · refine' (ae_restrict_iff' measurableSet_Ioo).mpr (eventually_of_forall fun t ht => _) rw [mul_comm, norm_mul] @@ -471,7 +473,7 @@ theorem hasMellin_one_Ioc {s : ℂ} (hs : 0 < re s) : simp_rw [HasMellin, mellin, MellinConvergent, ← indicator_smul, IntegrableOn, integrable_indicator_iff aux3, smul_eq_mul, integral_indicator aux3, mul_one, IntegrableOn, Measure.restrict_restrict_of_subset Ioc_subset_Ioi_self] - rw [← IntegrableOn, ← intervalIntegrable_iff_integrable_Ioc_of_le zero_le_one] + rw [← IntegrableOn, ← intervalIntegrable_iff_integrableOn_Ioc_of_le zero_le_one] refine' ⟨intervalIntegral.intervalIntegrable_cpow' aux1, _⟩ rw [← intervalIntegral.integral_of_le zero_le_one, integral_cpow (Or.inl aux1), sub_add_cancel, ofReal_zero, ofReal_one, one_cpow, zero_cpow aux2, sub_zero] diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 193888790ef6c..508e3a1bdf050 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -494,8 +494,7 @@ section theorem norm_ofSubsingleton [Subsingleton ι] (i : ι) (f : G →L[𝕜] G') : ‖ofSubsingleton 𝕜 G G' i f‖ = ‖f‖ := by letI : Unique ι := uniqueOfSubsingleton i - simp only [norm_def, ContinuousLinearMap.norm_def, (Equiv.funUnique _ _).symm.surjective.forall, - Fintype.prod_subsingleton _ i]; rfl + simp [norm_def, ContinuousLinearMap.norm_def, (Equiv.funUnique _ _).symm.surjective.forall] @[simp] theorem nnnorm_ofSubsingleton [Subsingleton ι] (i : ι) (f : G →L[𝕜] G') : diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index 27c8b2afc3abf..462ed8d781a71 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -75,7 +75,7 @@ theorem GammaIntegral_convergent {s : ℝ} (h : 0 < s) : constructor · rw [← integrableOn_Icc_iff_integrableOn_Ioc] refine' IntegrableOn.continuousOn_mul continuousOn_id.neg.exp _ isCompact_Icc - refine' (intervalIntegrable_iff_integrable_Icc_of_le zero_le_one).mp _ + refine' (intervalIntegrable_iff_integrableOn_Icc_of_le zero_le_one).mp _ exact intervalIntegrable_rpow' (by linarith) · refine' integrable_of_isBigO_exp_neg one_half_pos _ (Gamma_integrand_isLittleO _).isBigO refine' continuousOn_id.neg.exp.mul (continuousOn_id.rpow_const _) @@ -169,7 +169,7 @@ theorem tendsto_partialGamma {s : ℂ} (hs : 0 < s.re) : private theorem Gamma_integrand_interval_integrable (s : ℂ) {X : ℝ} (hs : 0 < s.re) (hX : 0 ≤ X) : IntervalIntegrable (fun x => (-x).exp * x ^ (s - 1) : ℝ → ℂ) volume 0 X := by - rw [intervalIntegrable_iff_integrable_Ioc_of_le hX] + rw [intervalIntegrable_iff_integrableOn_Ioc_of_le hX] exact IntegrableOn.mono_set (GammaIntegral_convergent hs) Ioc_subset_Ioi_self private theorem Gamma_integrand_deriv_integrable_A {s : ℂ} (hs : 0 < s.re) {X : ℝ} (hX : 0 ≤ X) : @@ -182,7 +182,7 @@ private theorem Gamma_integrand_deriv_integrable_B {s : ℂ} (hs : 0 < s.re) {Y IntervalIntegrable (fun x : ℝ => (-x).exp * (s * x ^ (s - 1)) : ℝ → ℂ) volume 0 Y := by have : (fun x => (-x).exp * (s * x ^ (s - 1)) : ℝ → ℂ) = (fun x => s * ((-x).exp * x ^ (s - 1)) : ℝ → ℂ) := by ext1; ring - rw [this, intervalIntegrable_iff_integrable_Ioc_of_le hY] + rw [this, intervalIntegrable_iff_integrableOn_Ioc_of_le hY] constructor · refine' (continuousOn_const.mul _).aestronglyMeasurable measurableSet_Ioc apply (continuous_ofReal.comp continuous_neg.exp).continuousOn.mul @@ -306,7 +306,7 @@ theorem GammaAux_recurrence2 (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : /-- The `Γ` function (of a complex variable `s`). -/ -- @[pp_nodot] -- Porting note: removed -def Gamma (s : ℂ) : ℂ := +irreducible_def Gamma (s : ℂ) : ℂ := GammaAux ⌊1 - s.re⌋₊ s #align complex.Gamma Complex.Gamma @@ -377,7 +377,8 @@ theorem Gamma_neg_nat_eq_zero (n : ℕ) : Gamma (-n) = 0 := by #align complex.Gamma_neg_nat_eq_zero Complex.Gamma_neg_nat_eq_zero theorem Gamma_conj (s : ℂ) : Gamma (conj s) = conj (Gamma s) := by - suffices : ∀ (n : ℕ) (s : ℂ), GammaAux n (conj s) = conj (GammaAux n s); exact this _ _ + suffices ∀ (n : ℕ) (s : ℂ), GammaAux n (conj s) = conj (GammaAux n s) by + simp [Gamma, this] intro n induction' n with n IH · rw [GammaAux]; exact GammaIntegral_conj diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean index 2bd846666489f..f5af478858237 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean @@ -300,7 +300,7 @@ theorem approx_Gamma_integral_tendsto_Gamma_integral {s : ℂ} (hs : 0 < re s) : intro n rw [integrable_indicator_iff (measurableSet_Ioc : MeasurableSet (Ioc (_ : ℝ) _)), IntegrableOn, Measure.restrict_restrict_of_subset Ioc_subset_Ioi_self, ← IntegrableOn, ← - intervalIntegrable_iff_integrable_Ioc_of_le (by positivity : (0 : ℝ) ≤ n)] + intervalIntegrable_iff_integrableOn_Ioc_of_le (by positivity : (0 : ℝ) ≤ n)] apply IntervalIntegrable.continuousOn_mul · refine' intervalIntegral.intervalIntegrable_cpow' _ rwa [sub_re, one_re, ← zero_sub, sub_lt_sub_iff_right] diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian.lean index a592bdbab64d3..02eb3077c9828 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian.lean @@ -85,7 +85,7 @@ theorem integrableOn_rpow_mul_exp_neg_rpow {p s : ℝ} (hs : -1 < s) (hp : 1 ≤ constructor · rw [← integrableOn_Icc_iff_integrableOn_Ioc] refine IntegrableOn.mul_continuousOn ?_ ?_ isCompact_Icc - · refine (intervalIntegrable_iff_integrable_Icc_of_le zero_le_one).mp ?_ + · refine (intervalIntegrable_iff_integrableOn_Icc_of_le zero_le_one).mp ?_ exact intervalIntegral.intervalIntegrable_rpow' hs · intro x _ change ContinuousWithinAt ((fun x => exp (- x)) ∘ (fun x => x ^ p)) (Icc 0 1) x diff --git a/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean b/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean index 1356ffd72599c..3eebe992bab84 100644 --- a/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean @@ -73,6 +73,36 @@ theorem integrableOn_Ioi_rpow_of_lt {a : ℝ} (ha : a < -1) {c : ℝ} (hc : 0 < integrableOn_Ioi_deriv_of_nonneg' hd (fun t ht => rpow_nonneg_of_nonneg (hc.trans ht).le a) ht #align integrable_on_Ioi_rpow_of_lt integrableOn_Ioi_rpow_of_lt +theorem integrableOn_Ioi_rpow_iff {s t : ℝ} (ht : 0 < t) : + IntegrableOn (fun x ↦ x ^ s) (Ioi t) ↔ s < -1 := by + refine ⟨fun h ↦ ?_, fun h ↦ integrableOn_Ioi_rpow_of_lt h ht⟩ + contrapose! h + intro H + have H' : IntegrableOn (fun x ↦ x ^ s) (Ioi (max 1 t)) := + H.mono (Set.Ioi_subset_Ioi (le_max_right _ _)) le_rfl + have : IntegrableOn (fun x ↦ x⁻¹) (Ioi (max 1 t)) := by + apply H'.mono' measurable_inv.aestronglyMeasurable + filter_upwards [ae_restrict_mem measurableSet_Ioi] with x hx + have x_one : 1 ≤ x := ((le_max_left _ _).trans_lt (mem_Ioi.1 hx)).le + simp only [norm_inv, Real.norm_eq_abs, abs_of_nonneg (zero_le_one.trans x_one)] + rw [← Real.rpow_neg_one x] + exact Real.rpow_le_rpow_of_exponent_le x_one h + exact not_IntegrableOn_Ioi_inv this + +/-- The real power function with any exponent is not integrable on `(0, +∞)`. -/ +theorem not_integrableOn_Ioi_rpow (s : ℝ) : ¬ IntegrableOn (fun x ↦ x ^ s) (Ioi (0 : ℝ)) := by + intro h + rcases le_or_lt s (-1) with hs|hs + · have : IntegrableOn (fun x ↦ x ^ s) (Ioo (0 : ℝ) 1) := h.mono Ioo_subset_Ioi_self le_rfl + rw [integrableOn_Ioo_rpow_iff zero_lt_one] at this + exact hs.not_lt this + · have : IntegrableOn (fun x ↦ x ^ s) (Ioi 1) := h.mono (Ioi_subset_Ioi zero_le_one) le_rfl + rw [integrableOn_Ioi_rpow_iff zero_lt_one] at this + exact hs.not_lt this + +theorem setIntegral_Ioi_zero_rpow (s : ℝ) : ∫ x in Ioi (0 : ℝ), x ^ s = 0 := + MeasureTheory.integral_undef (not_integrableOn_Ioi_rpow s) + theorem integral_Ioi_rpow_of_lt {a : ℝ} (ha : a < -1) {c : ℝ} (hc : 0 < c) : ∫ t : ℝ in Ioi c, t ^ a = -c ^ (a + 1) / (a + 1) := by have hd : ∀ (x : ℝ) (_ : x ∈ Ici c), HasDerivAt (fun t => t ^ (a + 1) / (a + 1)) (x ^ a) x := by @@ -97,6 +127,33 @@ theorem integrableOn_Ioi_cpow_of_lt {a : ℂ} (ha : a.re < -1) {c : ℝ} (hc : 0 (Complex.continuousAt_ofReal_cpow_const _ _ (Or.inr (hc.trans ht).ne')).continuousWithinAt #align integrable_on_Ioi_cpow_of_lt integrableOn_Ioi_cpow_of_lt +theorem integrableOn_Ioi_cpow_iff {s : ℂ} {t : ℝ} (ht : 0 < t) : + IntegrableOn (fun x : ℝ ↦ (x : ℂ) ^ s) (Ioi t) ↔ s.re < -1 := by + refine ⟨fun h ↦ ?_, fun h ↦ integrableOn_Ioi_cpow_of_lt h ht⟩ + have B : IntegrableOn (fun a ↦ a ^ s.re) (Ioi t) := by + apply (integrableOn_congr_fun _ measurableSet_Ioi).1 h.norm + intro a ha + have : 0 < a := ht.trans ha + simp [Complex.abs_cpow_eq_rpow_re_of_pos this] + rwa [integrableOn_Ioi_rpow_iff ht] at B + +/-- The complex power function with any exponent is not integrable on `(0, +∞)`. -/ +theorem not_integrableOn_Ioi_cpow (s : ℂ) : + ¬ IntegrableOn (fun x : ℝ ↦ (x : ℂ) ^ s) (Ioi (0 : ℝ)) := by + intro h + rcases le_or_lt s.re (-1) with hs|hs + · have : IntegrableOn (fun x : ℝ ↦ (x : ℂ) ^ s) (Ioo (0 : ℝ) 1) := + h.mono Ioo_subset_Ioi_self le_rfl + rw [integrableOn_Ioo_cpow_iff zero_lt_one] at this + exact hs.not_lt this + · have : IntegrableOn (fun x : ℝ ↦ (x : ℂ) ^ s) (Ioi 1) := + h.mono (Ioi_subset_Ioi zero_le_one) le_rfl + rw [integrableOn_Ioi_cpow_iff zero_lt_one] at this + exact hs.not_lt this + +theorem setIntegral_Ioi_zero_cpow (s : ℂ) : ∫ x in Ioi (0 : ℝ), (x : ℂ) ^ s = 0 := + MeasureTheory.integral_undef (not_integrableOn_Ioi_cpow s) + theorem integral_Ioi_cpow_of_lt {a : ℂ} (ha : a.re < -1) {c : ℝ} (hc : 0 < c) : (∫ t : ℝ in Ioi c, (t : ℂ) ^ a) = -(c : ℂ) ^ (a + 1) / (a + 1) := by refine' diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index 1866918bd3f9e..ce8ad9de3f0fa 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -5,6 +5,7 @@ Authors: Benjamin Davidson -/ import Mathlib.MeasureTheory.Integral.FundThmCalculus import Mathlib.Analysis.SpecialFunctions.Trigonometric.ArctanDeriv +import Mathlib.Analysis.SpecialFunctions.NonIntegrable #align_import analysis.special_functions.integrals from "leanprover-community/mathlib"@"011cafb4a5bc695875d186e245d6b3df03bf6c40" @@ -93,6 +94,26 @@ theorem intervalIntegrable_rpow' {r : ℝ} (h : -1 < r) : rpow_def_of_pos hx.1, rpow_def_of_neg (by linarith [hx.1] : -x < 0)] #align interval_integral.interval_integrable_rpow' intervalIntegral.intervalIntegrable_rpow' +/-- The power function `x ↦ x^s` is integrable on `(0, t)` iff `-1 < s`. -/ +lemma integrableOn_Ioo_rpow_iff {s t : ℝ} (ht : 0 < t) : + IntegrableOn (fun x ↦ x ^ s) (Ioo (0 : ℝ) t) ↔ -1 < s := by + refine ⟨fun h ↦ ?_, fun h ↦ by simpa [intervalIntegrable_iff_integrableOn_Ioo_of_le ht.le] + using intervalIntegrable_rpow' h (a := 0) (b := t)⟩ + contrapose! h + intro H + have I : 0 < min 1 t := lt_min zero_lt_one ht + have H' : IntegrableOn (fun x ↦ x ^ s) (Ioo 0 (min 1 t)) := + H.mono (Set.Ioo_subset_Ioo le_rfl (min_le_right _ _)) le_rfl + have : IntegrableOn (fun x ↦ x⁻¹) (Ioo 0 (min 1 t)) := by + apply H'.mono' measurable_inv.aestronglyMeasurable + filter_upwards [ae_restrict_mem measurableSet_Ioo] with x hx + simp only [norm_inv, Real.norm_eq_abs, abs_of_nonneg (le_of_lt hx.1)] + rwa [← Real.rpow_neg_one x, Real.rpow_le_rpow_left_iff_of_base_lt_one hx.1] + exact lt_of_lt_of_le hx.2 (min_le_left _ _) + have : IntervalIntegrable (fun x ↦ x⁻¹) volume 0 (min 1 t) := by + rwa [intervalIntegrable_iff_integrableOn_Ioo_of_le I.le] + simp [intervalIntegrable_inv_iff, I.ne] at this + /-- See `intervalIntegrable_cpow'` for a version with a weaker hypothesis on `r`, but assuming the measure is volume. -/ theorem intervalIntegrable_cpow {r : ℂ} (h : 0 ≤ r.re ∨ (0 : ℝ) ∉ [[a, b]]) : @@ -117,13 +138,13 @@ theorem intervalIntegrable_cpow {r : ℂ} (h : 0 ≤ r.re ∨ (0 : ℝ) ∉ [[a, rcases le_or_lt 0 c with (hc | hc) · -- case `0 ≤ c`: integrand is identically 1 have : IntervalIntegrable (fun _ => 1 : ℝ → ℝ) μ 0 c := intervalIntegrable_const - rw [intervalIntegrable_iff_integrable_Ioc_of_le hc] at this ⊢ + rw [intervalIntegrable_iff_integrableOn_Ioc_of_le hc] at this ⊢ refine' IntegrableOn.congr_fun this (fun x hx => _) measurableSet_Ioc dsimp only rw [Complex.norm_eq_abs, Complex.abs_cpow_eq_rpow_re_of_pos hx.1, ← h', rpow_zero] · -- case `c < 0`: integrand is identically constant, *except* at `x = 0` if `r ≠ 0`. apply IntervalIntegrable.symm - rw [intervalIntegrable_iff_integrable_Ioc_of_le hc.le] + rw [intervalIntegrable_iff_integrableOn_Ioc_of_le hc.le] have : Ioc c 0 = Ioo c 0 ∪ {(0 : ℝ)} := by rw [← Ioo_union_Icc_eq_Ioc hc (le_refl 0), ← Icc_def] simp_rw [← le_antisymm_iff, setOf_eq_eq_singleton'] @@ -177,6 +198,17 @@ theorem intervalIntegrable_cpow' {r : ℂ} (h : -1 < r.re) : simp #align interval_integral.interval_integrable_cpow' intervalIntegral.intervalIntegrable_cpow' +/-- The complex power function `x ↦ x^s` is integrable on `(0, t)` iff `-1 < s.re`. -/ +theorem integrableOn_Ioo_cpow_iff {s : ℂ} {t : ℝ} (ht : 0 < t) : + IntegrableOn (fun x : ℝ ↦ (x : ℂ) ^ s) (Ioo (0 : ℝ) t) ↔ -1 < s.re := by + refine ⟨fun h ↦ ?_, fun h ↦ by simpa [intervalIntegrable_iff_integrableOn_Ioo_of_le ht.le] + using intervalIntegrable_cpow' h (a := 0) (b := t)⟩ + have B : IntegrableOn (fun a ↦ a ^ s.re) (Ioo 0 t) := by + apply (integrableOn_congr_fun _ measurableSet_Ioo).1 h.norm + intro a ha + simp [Complex.abs_cpow_eq_rpow_re_of_pos ha.1] + rwa [integrableOn_Ioo_rpow_iff ht] at B + @[simp] theorem intervalIntegrable_id : IntervalIntegrable (fun x => x) μ a b := continuous_id.intervalIntegrable a b diff --git a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean index ac0728f081189..282ce63bab599 100644 --- a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean +++ b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean @@ -90,7 +90,7 @@ theorem finite_integral_rpow_sub_one_pow_aux {r : ℝ} (n : ℕ) (hnr : (n : ℝ exact hr.le refine' lt_of_le_of_lt (set_lintegral_mono' measurableSet_Ioc h_int) _ refine' IntegrableOn.set_lintegral_lt_top _ - rw [← intervalIntegrable_iff_integrable_Ioc_of_le zero_le_one] + rw [← intervalIntegrable_iff_integrableOn_Ioc_of_le zero_le_one] apply intervalIntegral.intervalIntegrable_rpow' rwa [neg_lt_neg_iff, inv_mul_lt_iff' hr, one_mul] #align finite_integral_rpow_sub_one_pow_aux finite_integral_rpow_sub_one_pow_aux diff --git a/Mathlib/Analysis/SpecialFunctions/NonIntegrable.lean b/Mathlib/Analysis/SpecialFunctions/NonIntegrable.lean index 1ab1e87d8af20..0a86d7cbd9ae2 100644 --- a/Mathlib/Analysis/SpecialFunctions/NonIntegrable.lean +++ b/Mathlib/Analysis/SpecialFunctions/NonIntegrable.lean @@ -43,56 +43,93 @@ open scoped MeasureTheory Topology Interval NNReal ENNReal open MeasureTheory TopologicalSpace Set Filter Asymptotics intervalIntegral -variable {E F : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [SecondCountableTopology E] - [CompleteSpace E] [NormedAddCommGroup F] +variable {E F : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [NormedAddCommGroup F] /-- If `f` is eventually differentiable along a nontrivial filter `l : Filter ℝ` that is generated by convex sets, the norm of `f` tends to infinity along `l`, and `f' = O(g)` along `l`, where `f'` -is the derivative of `f`, then `g` is not integrable on any interval `a..b` such that -`[a, b] ∈ l`. -/ -theorem not_intervalIntegrable_of_tendsto_norm_atTop_of_deriv_isBigO_filter {f : ℝ → E} {g : ℝ → F} - {a b : ℝ} (l : Filter ℝ) [NeBot l] [TendstoIxxClass Icc l l] (hl : [[a, b]] ∈ l) - (hd : ∀ᶠ x in l, DifferentiableAt ℝ f x) (hf : Tendsto (fun x => ‖f x‖) l atTop) - (hfg : deriv f =O[l] g) : ¬IntervalIntegrable g volume a b := by +is the derivative of `f`, then `g` is not integrable on any set `k` belonging to `l`. +Auxiliary version assuming that `E` is complete. -/ +theorem not_integrableOn_of_tendsto_norm_atTop_of_deriv_isBigO_filter_aux + [CompleteSpace E] {f : ℝ → E} {g : ℝ → F} + {k : Set ℝ} (l : Filter ℝ) [NeBot l] [TendstoIxxClass Icc l l] + (hl : k ∈ l) (hd : ∀ᶠ x in l, DifferentiableAt ℝ f x) (hf : Tendsto (fun x => ‖f x‖) l atTop) + (hfg : deriv f =O[l] g) : ¬IntegrableOn g k := by intro hgi obtain ⟨C, hC₀, s, hsl, hsub, hfd, hg⟩ : - ∃ (C : ℝ) (_ : 0 ≤ C), ∃ s ∈ l, (∀ x ∈ s, ∀ y ∈ s, [[x, y]] ⊆ [[a, b]]) ∧ + ∃ (C : ℝ) (_ : 0 ≤ C), ∃ s ∈ l, (∀ x ∈ s, ∀ y ∈ s, [[x, y]] ⊆ k) ∧ (∀ x ∈ s, ∀ y ∈ s, ∀ z ∈ [[x, y]], DifferentiableAt ℝ f z) ∧ ∀ x ∈ s, ∀ y ∈ s, ∀ z ∈ [[x, y]], ‖deriv f z‖ ≤ C * ‖g z‖ := by rcases hfg.exists_nonneg with ⟨C, C₀, hC⟩ have h : ∀ᶠ x : ℝ × ℝ in l.prod l, - ∀ y ∈ [[x.1, x.2]], (DifferentiableAt ℝ f y ∧ ‖deriv f y‖ ≤ C * ‖g y‖) ∧ y ∈ [[a, b]] := + ∀ y ∈ [[x.1, x.2]], (DifferentiableAt ℝ f y ∧ ‖deriv f y‖ ≤ C * ‖g y‖) ∧ y ∈ k := (tendsto_fst.uIcc tendsto_snd).eventually ((hd.and hC.bound).and hl).smallSets rcases mem_prod_self_iff.1 h with ⟨s, hsl, hs⟩ simp only [prod_subset_iff, mem_setOf_eq] at hs exact ⟨C, C₀, s, hsl, fun x hx y hy z hz => (hs x hx y hy z hz).2, fun x hx y hy z hz => (hs x hx y hy z hz).1.1, fun x hx y hy z hz => (hs x hx y hy z hz).1.2⟩ - replace hgi : IntervalIntegrable (fun x => C * ‖g x‖) volume a b - · convert hgi.norm.smul C using 1 - obtain ⟨c, hc, d, hd, hlt⟩ : ∃ c ∈ s, ∃ d ∈ s, (‖f c‖ + ∫ y in Ι a b, C * ‖g y‖) < ‖f d‖ := by + replace hgi : IntegrableOn (fun x ↦ C * ‖g x‖) k := by exact hgi.norm.smul C + obtain ⟨c, hc, d, hd, hlt⟩ : ∃ c ∈ s, ∃ d ∈ s, (‖f c‖ + ∫ y in k, C * ‖g y‖) < ‖f d‖ := by rcases Filter.nonempty_of_mem hsl with ⟨c, hc⟩ - have : ∀ᶠ x in l, (‖f c‖ + ∫ y in Ι a b, C * ‖g y‖) < ‖f x‖ := + have : ∀ᶠ x in l, (‖f c‖ + ∫ y in k, C * ‖g y‖) < ‖f x‖ := hf.eventually (eventually_gt_atTop _) exact ⟨c, hc, (this.and hsl).exists.imp fun d hd => ⟨hd.2, hd.1⟩⟩ specialize hsub c hc d hd; specialize hfd c hc d hd - replace hg : ∀ x ∈ Ι c d, ‖deriv f x‖ ≤ C * ‖g x‖; - exact fun z hz => hg c hc d hd z ⟨hz.1.le, hz.2⟩ + replace hg : ∀ x ∈ Ι c d, ‖deriv f x‖ ≤ C * ‖g x‖ := + fun z hz => hg c hc d hd z ⟨hz.1.le, hz.2⟩ have hg_ae : ∀ᵐ x ∂volume.restrict (Ι c d), ‖deriv f x‖ ≤ C * ‖g x‖ := (ae_restrict_mem measurableSet_uIoc).mono hg - have hsub' : Ι c d ⊆ Ι a b := uIoc_subset_uIoc_of_uIcc_subset_uIcc hsub - have hfi : IntervalIntegrable (deriv f) volume c d := - (hgi.mono_set hsub).mono_fun' (aestronglyMeasurable_deriv _ _) hg_ae + have hsub' : Ι c d ⊆ k := Subset.trans Ioc_subset_Icc_self hsub + have hfi : IntervalIntegrable (deriv f) volume c d := by + rw [intervalIntegrable_iff] + have : IntegrableOn (fun x ↦ C * ‖g x‖) (Ι c d) := IntegrableOn.mono hgi hsub' le_rfl + exact Integrable.mono' this (aestronglyMeasurable_deriv _ _) hg_ae refine' hlt.not_le (sub_le_iff_le_add'.1 _) calc ‖f d‖ - ‖f c‖ ≤ ‖f d - f c‖ := norm_sub_norm_le _ _ - _ = ‖∫ x in c..d, deriv f x‖ := (congr_arg _ (integral_deriv_eq_sub hfd hfi).symm) - _ = ‖∫ x in Ι c d, deriv f x‖ := (norm_integral_eq_norm_integral_Ioc _) - _ ≤ ∫ x in Ι c d, ‖deriv f x‖ := (norm_integral_le_integral_norm _) + _ = ‖∫ x in c..d, deriv f x‖ := congr_arg _ (integral_deriv_eq_sub hfd hfi).symm + _ = ‖∫ x in Ι c d, deriv f x‖ := norm_integral_eq_norm_integral_Ioc _ + _ ≤ ∫ x in Ι c d, ‖deriv f x‖ := norm_integral_le_integral_norm _ _ ≤ ∫ x in Ι c d, C * ‖g x‖ := - (set_integral_mono_on hfi.norm.def (hgi.def.mono_set hsub') measurableSet_uIoc hg) - _ ≤ ∫ x in Ι a b, C * ‖g x‖ := - set_integral_mono_set hgi.def (ae_of_all _ fun x => mul_nonneg hC₀ (norm_nonneg _)) - hsub'.eventuallyLE + set_integral_mono_on hfi.norm.def (hgi.mono_set hsub') measurableSet_uIoc hg + _ ≤ ∫ x in k, C * ‖g x‖ := by + apply set_integral_mono_set hgi + (ae_of_all _ fun x => mul_nonneg hC₀ (norm_nonneg _)) hsub'.eventuallyLE + +theorem not_integrableOn_of_tendsto_norm_atTop_of_deriv_isBigO_filter + {f : ℝ → E} {g : ℝ → F} + {k : Set ℝ} (l : Filter ℝ) [NeBot l] [TendstoIxxClass Icc l l] + (hl : k ∈ l) (hd : ∀ᶠ x in l, DifferentiableAt ℝ f x) (hf : Tendsto (fun x => ‖f x‖) l atTop) + (hfg : deriv f =O[l] g) : ¬IntegrableOn g k := by + let a : E →ₗᵢ[ℝ] UniformSpace.Completion E := UniformSpace.Completion.toComplₗᵢ + let f' := a ∘ f + have h'd : ∀ᶠ x in l, DifferentiableAt ℝ f' x := by + filter_upwards [hd] with x hx using a.toContinuousLinearMap.differentiableAt.comp x hx + have h'f : Tendsto (fun x => ‖f' x‖) l atTop := hf.congr (fun x ↦ by simp) + have h'fg : deriv f' =O[l] g := by + apply IsBigO.trans _ hfg + rw [← isBigO_norm_norm] + suffices (fun x ↦ ‖deriv f' x‖) =ᶠ[l] (fun x ↦ ‖deriv f x‖) by exact this.isBigO + filter_upwards [hd] with x hx + have : deriv f' x = a (deriv f x) := by + rw [fderiv.comp_deriv x _ hx] + · have : fderiv ℝ a (f x) = a.toContinuousLinearMap := a.toContinuousLinearMap.fderiv + simp only [this] + rfl + · exact a.toContinuousLinearMap.differentiableAt + simp only [this] + simp + exact not_integrableOn_of_tendsto_norm_atTop_of_deriv_isBigO_filter_aux l hl h'd h'f h'fg + +/-- If `f` is eventually differentiable along a nontrivial filter `l : Filter ℝ` that is generated +by convex sets, the norm of `f` tends to infinity along `l`, and `f' = O(g)` along `l`, where `f'` +is the derivative of `f`, then `g` is not integrable on any interval `a..b` such that +`[a, b] ∈ l`. -/ +theorem not_intervalIntegrable_of_tendsto_norm_atTop_of_deriv_isBigO_filter {f : ℝ → E} {g : ℝ → F} + {a b : ℝ} (l : Filter ℝ) [NeBot l] [TendstoIxxClass Icc l l] (hl : [[a, b]] ∈ l) + (hd : ∀ᶠ x in l, DifferentiableAt ℝ f x) (hf : Tendsto (fun x => ‖f x‖) l atTop) + (hfg : deriv f =O[l] g) : ¬IntervalIntegrable g volume a b := by + rw [intervalIntegrable_iff'] + exact not_integrableOn_of_tendsto_norm_atTop_of_deriv_isBigO_filter _ hl hd hf hfg set_option linter.uppercaseLean3 false in #align not_interval_integrable_of_tendsto_norm_at_top_of_deriv_is_O_filter not_intervalIntegrable_of_tendsto_norm_atTop_of_deriv_isBigO_filter @@ -174,3 +211,19 @@ theorem intervalIntegrable_inv_iff {a b : ℝ} : IntervalIntegrable (fun x => x⁻¹) volume a b ↔ a = b ∨ (0 : ℝ) ∉ [[a, b]] := by simp only [← intervalIntegrable_sub_inv_iff, sub_zero] #align interval_integrable_inv_iff intervalIntegrable_inv_iff + +/-- The function `fun x ↦ x⁻¹` is not integrable on any interval `[a, +∞)`. -/ +theorem not_IntegrableOn_Ici_inv {a : ℝ} : + ¬ IntegrableOn (fun x => x⁻¹) (Ici a) := by + have A : ∀ᶠ x in atTop, HasDerivAt (fun x => Real.log x) x⁻¹ x := by + filter_upwards [Ioi_mem_atTop 0] with x hx using Real.hasDerivAt_log (ne_of_gt hx) + have B : Tendsto (fun x => ‖Real.log x‖) atTop atTop := + tendsto_norm_atTop_atTop.comp Real.tendsto_log_atTop + exact not_integrableOn_of_tendsto_norm_atTop_of_deriv_isBigO_filter atTop (Ici_mem_atTop a) + (A.mono (fun x hx ↦ hx.differentiableAt)) B + (Filter.EventuallyEq.isBigO (A.mono (fun x hx ↦ hx.deriv))) + +/-- The function `fun x ↦ x⁻¹` is not integrable on any interval `(a, +∞)`. -/ +theorem not_IntegrableOn_Ioi_inv {a : ℝ} : + ¬ IntegrableOn (·⁻¹) (Ioi a) := by + simpa only [IntegrableOn, restrict_Ioi_eq_restrict_Ici] using not_IntegrableOn_Ici_inv diff --git a/Mathlib/CategoryTheory/Sites/Closed.lean b/Mathlib/CategoryTheory/Sites/Closed.lean index 7cd46668661a1..82452a339ea72 100644 --- a/Mathlib/CategoryTheory/Sites/Closed.lean +++ b/Mathlib/CategoryTheory/Sites/Closed.lean @@ -104,18 +104,17 @@ theorem close_isClosed {X : C} (S : Sieve X) : J₁.IsClosed (J₁.close S) := fun _ g hg => J₁.arrow_trans g _ S hg fun _ hS => hS #align category_theory.grothendieck_topology.close_is_closed CategoryTheory.GrothendieckTopology.close_isClosed +/-- A Grothendieck topology induces a natural family of closure operators on sieves. -/ +@[simps! isClosed] +def closureOperator (X : C) : ClosureOperator (Sieve X) := + .ofPred J₁.close J₁.IsClosed J₁.le_close J₁.close_isClosed fun _ _ ↦ J₁.le_close_of_isClosed +#align category_theory.grothendieck_topology.closure_operator CategoryTheory.GrothendieckTopology.closureOperator + +#align category_theory.grothendieck_topology.closed_iff_closed CategoryTheory.GrothendieckTopology.closureOperator_isClosed + /-- The sieve `S` is closed iff its closure is equal to itself. -/ -theorem isClosed_iff_close_eq_self {X : C} (S : Sieve X) : J₁.IsClosed S ↔ J₁.close S = S := by - constructor - · intro h - apply le_antisymm - · intro Y f hf - rw [← J₁.covers_iff_mem_of_isClosed h] - apply hf - · apply J₁.le_close - · intro e - rw [← e] - apply J₁.close_isClosed +theorem isClosed_iff_close_eq_self {X : C} (S : Sieve X) : J₁.IsClosed S ↔ J₁.close S = S := + (J₁.closureOperator _).isClosed_iff #align category_theory.grothendieck_topology.is_closed_iff_close_eq_self CategoryTheory.GrothendieckTopology.isClosed_iff_close_eq_self theorem close_eq_self_of_isClosed {X : C} {S : Sieve X} (hS : J₁.IsClosed S) : J₁.close S = S := @@ -136,13 +135,12 @@ theorem pullback_close {X Y : C} (f : Y ⟶ X) (S : Sieve X) : @[mono] theorem monotone_close {X : C} : Monotone (J₁.close : Sieve X → Sieve X) := - fun _ S₂ h => J₁.le_close_of_isClosed (h.trans (J₁.le_close _)) (J₁.close_isClosed S₂) + (J₁.closureOperator _).monotone #align category_theory.grothendieck_topology.monotone_close CategoryTheory.GrothendieckTopology.monotone_close @[simp] theorem close_close {X : C} (S : Sieve X) : J₁.close (J₁.close S) = J₁.close S := - le_antisymm (J₁.le_close_of_isClosed le_rfl (J₁.close_isClosed S)) - (J₁.monotone_close (J₁.le_close _)) + (J₁.closureOperator _).idempotent _ #align category_theory.grothendieck_topology.close_close CategoryTheory.GrothendieckTopology.close_close /-- @@ -162,20 +160,6 @@ theorem close_eq_top_iff_mem {X : C} (S : Sieve X) : J₁.close S = ⊤ ↔ S apply J₁.pullback_stable _ hS #align category_theory.grothendieck_topology.close_eq_top_iff_mem CategoryTheory.GrothendieckTopology.close_eq_top_iff_mem -/-- A Grothendieck topology induces a natural family of closure operators on sieves. -/ -@[simps!] -def closureOperator (X : C) : ClosureOperator (Sieve X) := - ClosureOperator.mk' J₁.close - (fun _ S₂ h => J₁.le_close_of_isClosed (h.trans (J₁.le_close _)) (J₁.close_isClosed S₂)) - J₁.le_close fun S => J₁.le_close_of_isClosed le_rfl (J₁.close_isClosed S) -#align category_theory.grothendieck_topology.closure_operator CategoryTheory.GrothendieckTopology.closureOperator - -@[simp] -theorem closed_iff_closed {X : C} (S : Sieve X) : - S ∈ (J₁.closureOperator X).closed ↔ J₁.IsClosed S := - (J₁.isClosed_iff_close_eq_self S).symm -#align category_theory.grothendieck_topology.closed_iff_closed CategoryTheory.GrothendieckTopology.closed_iff_closed - end GrothendieckTopology /-- diff --git a/Mathlib/Condensed/Explicit.lean b/Mathlib/Condensed/Explicit.lean index e83935d583441..191d57b9867c9 100644 --- a/Mathlib/Condensed/Explicit.lean +++ b/Mathlib/Condensed/Explicit.lean @@ -21,7 +21,7 @@ We give the following three explicit descriptions of condensed sets: * `Condensed.ofSheafStonean`: A finite-product-preserving presheaf on `CompHaus`, satisfying `EqualizerCondition`. -The property `EqualizerCondition` is defined in `Mathlib/CategoryTheory/Sites/RegularExtensive` +The property `EqualizerCondition` is defined in `Mathlib/CategoryTheory/Sites/RegularExtensive.lean` and it says that for any effective epi `X ⟶ B` (in this case that is equivalent to being a continuous surjection), the presheaf `F` exhibits `F(B)` as the equalizer of the two maps `F(X) ⇉ F(X ×_B X)` diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 3e7ff55772748..46f67f53d9fc1 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -2106,31 +2106,10 @@ theorem cons_nthLe_drop_succ {l : List α} {n : ℕ} (hn : n < l.length) : #align list.cons_nth_le_drop_succ List.cons_nthLe_drop_succ #align list.drop_nil List.drop_nil - -@[simp] -theorem drop_one : ∀ l : List α, drop 1 l = tail l - | [] | _ :: _ => rfl #align list.drop_one List.drop_one - -theorem drop_add : ∀ (m n) (l : List α), drop (m + n) l = drop m (drop n l) - | _, 0, _ => rfl - | _, _ + 1, [] => drop_nil.symm - | m, n + 1, _ :: _ => drop_add m n _ #align list.drop_add List.drop_add - -@[simp] -theorem drop_left : ∀ l₁ l₂ : List α, drop (length l₁) (l₁ ++ l₂) = l₂ - | [], _ => rfl - | _ :: l₁, l₂ => drop_left l₁ l₂ #align list.drop_left List.drop_left - -theorem drop_left' {l₁ l₂ : List α} {n} (h : length l₁ = n) : drop n (l₁ ++ l₂) = l₂ := by - rw [← h]; apply drop_left #align list.drop_left' List.drop_left' - -theorem drop_eq_get_cons : ∀ {n} {l : List α} (h), drop n l = get l ⟨n, h⟩ :: drop (n + 1) l - | 0, _ :: _, _ => rfl - | n + 1, _ :: _, _ => @drop_eq_get_cons n _ _ #align list.drop_eq_nth_le_cons List.drop_eq_get_consₓ -- nth_le vs get #align list.drop_length List.drop_length diff --git a/Mathlib/Data/List/Zip.lean b/Mathlib/Data/List/Zip.lean index 09deeb716d8f5..dba949738c56a 100644 --- a/Mathlib/Data/List/Zip.lean +++ b/Mathlib/Data/List/Zip.lean @@ -34,12 +34,7 @@ variable {α : Type u} {β γ δ ε : Type*} #align list.zip_cons_cons List.zip_cons_cons #align list.zip_with_nil_left List.zipWith_nil_left #align list.zip_with_nil_right List.zipWith_nil_right - -@[simp] -theorem zipWith_eq_nil_iff {f : α → β → γ} {l l'} : zipWith f l l' = [] ↔ l = [] ∨ l' = [] := by - cases l <;> cases l' <;> simp #align list.zip_with_eq_nil_iff List.zipWith_eq_nil_iff - #align list.zip_nil_left List.zip_nil_left #align list.zip_nil_right List.zip_nil_right @@ -86,35 +81,14 @@ theorem lt_length_right_of_zip {i : ℕ} {l : List α} {l' : List β} (h : i < ( lt_length_right_of_zipWith h #align list.lt_length_right_of_zip List.lt_length_right_of_zip -theorem zip_append : - ∀ {l₁ r₁ : List α} {l₂ r₂ : List β} (_h : length l₁ = length l₂), - zip (l₁ ++ r₁) (l₂ ++ r₂) = zip l₁ l₂ ++ zip r₁ r₂ - | [], r₁, l₂, r₂, h => by simp only [eq_nil_of_length_eq_zero h.symm]; rfl - | l₁, r₁, [], r₂, h => by simp only [eq_nil_of_length_eq_zero h]; rfl - | a :: l₁, r₁, b :: l₂, r₂, h => by - simp only [cons_append, zip_cons_cons, zip_append (succ.inj h)] #align list.zip_append List.zip_append - #align list.zip_map List.zip_map #align list.zip_map_left List.zip_map_left #align list.zip_map_right List.zip_map_right #align list.zip_with_map List.zipWith_map #align list.zip_with_map_left List.zipWith_map_left #align list.zip_with_map_right List.zipWith_map_right - -theorem zip_map' (f : α → β) (g : α → γ) : - ∀ l : List α, zip (l.map f) (l.map g) = l.map fun a => (f a, g a) - | [] => rfl - | a :: l => by simp only [map, zip_cons_cons, zip_map'] #align list.zip_map' List.zip_map' - -theorem map_zipWith {δ : Type*} (f : α → β) (g : γ → δ → α) (l : List γ) (l' : List δ) : - map f (zipWith g l l') = zipWith (fun x y => f (g x y)) l l' := by - induction' l with hd tl hl generalizing l' - · simp - · cases l' - · simp - · simp [hl] #align list.map_zip_with List.map_zipWith theorem mem_zip {a b} : ∀ {l₁ : List α} {l₂ : List β}, (a, b) ∈ zip l₁ l₂ → a ∈ l₁ ∧ b ∈ l₂ @@ -125,30 +99,9 @@ theorem mem_zip {a b} : ∀ {l₁ : List α} {l₂ : List β}, (a, b) ∈ zip l exact ⟨Mem.tail _ this.1, Mem.tail _ this.2⟩ #align list.mem_zip List.mem_zip -theorem map_fst_zip : - ∀ (l₁ : List α) (l₂ : List β), l₁.length ≤ l₂.length → map Prod.fst (zip l₁ l₂) = l₁ - | [], bs, _ => rfl - | _ :: as, _ :: bs, h => by - simp? [succ_le_succ_iff] at h says simp only [length_cons, succ_le_succ_iff] at h - change _ :: map Prod.fst (zip as bs) = _ :: as - rw [map_fst_zip as bs h] - | a :: as, [], h => by simp at h #align list.map_fst_zip List.map_fst_zip - -theorem map_snd_zip : - ∀ (l₁ : List α) (l₂ : List β), l₂.length ≤ l₁.length → map Prod.snd (zip l₁ l₂) = l₂ - | _, [], _ => by - rw [zip_nil_right] - rfl - | [], b :: bs, h => by simp at h - | a :: as, b :: bs, h => by - simp? [succ_le_succ_iff] at h says simp only [length_cons, succ_le_succ_iff] at h - change _ :: map Prod.snd (zip as bs) = _ :: bs - rw [map_snd_zip as bs h] #align list.map_snd_zip List.map_snd_zip - #align list.unzip_nil List.unzip_nil - #align list.unzip_cons List.unzip_cons theorem unzip_eq_map : ∀ l : List (α × β), unzip l = (l.map Prod.fst, l.map Prod.snd) @@ -413,40 +366,9 @@ section Distrib variable (f : α → β → γ) (l : List α) (l' : List β) (n : ℕ) -theorem zipWith_distrib_take : (zipWith f l l').take n = zipWith f (l.take n) (l'.take n) := by - induction' l with hd tl hl generalizing l' n - · simp - · cases l' - · simp - · cases n - · simp - · simp [hl] #align list.zip_with_distrib_take List.zipWith_distrib_take - -theorem zipWith_distrib_drop : (zipWith f l l').drop n = zipWith f (l.drop n) (l'.drop n) := by - induction' l with hd tl hl generalizing l' n - · simp - · cases l' - · simp - · cases n - · simp - · simp [hl] #align list.zip_with_distrib_drop List.zipWith_distrib_drop - -theorem zipWith_distrib_tail : (zipWith f l l').tail = zipWith f l.tail l'.tail := by - simp_rw [← drop_one, zipWith_distrib_drop] #align list.zip_with_distrib_tail List.zipWith_distrib_tail - -theorem zipWith_append (f : α → β → γ) (l la : List α) (l' lb : List β) - (h : l.length = l'.length) : - zipWith f (l ++ la) (l' ++ lb) = zipWith f l l' ++ zipWith f la lb := by - induction' l with hd tl hl generalizing l' - · have : l' = [] := eq_nil_of_length_eq_zero (by simpa using h.symm) - simp [this] - · cases l' - · simp at h - · simp only [add_left_inj, length] at h - simp [hl _ h] #align list.zip_with_append List.zipWith_append theorem zipWith_distrib_reverse (h : l.length = l'.length) : diff --git a/Mathlib/FieldTheory/RatFunc.lean b/Mathlib/FieldTheory/RatFunc.lean index b4edc9a70cdf7..52b7364155cc5 100644 --- a/Mathlib/FieldTheory/RatFunc.lean +++ b/Mathlib/FieldTheory/RatFunc.lean @@ -1284,7 +1284,7 @@ theorem num_denom_mul (x y : RatFunc K) : #align ratfunc.num_denom_mul RatFunc.num_denom_mul theorem num_dvd {x : RatFunc K} {p : K[X]} (hp : p ≠ 0) : - num x ∣ p ↔ ∃ (q : K[X]) (_ : q ≠ 0), x = algebraMap _ _ p / algebraMap _ _ q := by + num x ∣ p ↔ ∃ q : K[X], q ≠ 0 ∧ x = algebraMap _ _ p / algebraMap _ _ q := by constructor · rintro ⟨q, rfl⟩ obtain ⟨_hx, hq⟩ := mul_ne_zero_iff.mp hp diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index 3b95deaa7c007..ce8e024089543 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -653,25 +653,31 @@ theorem exists_subgroup_card_pow_prime [Fintype G] (p : ℕ) {n : ℕ} [Fact p.P ⟨K, hK.1⟩ #align sylow.exists_subgroup_card_pow_prime Sylow.exists_subgroup_card_pow_prime -lemma exists_subgroup_card_pow_prime_of_le_card {m p : ℕ} (hp : p.Prime) (h : IsPGroup p G) - (hm : p ^ m ≤ Nat.card G) : ∃ H : Subgroup G, Nat.card H = p ^ m := by +/-- A special case of **Sylow's first theorem**. If `G` is a `p`-group of size at least `p ^ n` +then there is a subgroup of cardinality `p ^ n`. -/ +lemma exists_subgroup_card_pow_prime_of_le_card {n p : ℕ} (hp : p.Prime) (h : IsPGroup p G) + (hn : p ^ n ≤ Nat.card G) : ∃ H : Subgroup G, Nat.card H = p ^ n := by have : Fact p.Prime := ⟨hp⟩ - have : Finite G := Nat.finite_of_card_ne_zero $ by linarith [Nat.one_le_pow m p hp.pos] + have : Finite G := Nat.finite_of_card_ne_zero $ by linarith [Nat.one_le_pow n p hp.pos] cases nonempty_fintype G - obtain ⟨n, hn⟩ := h.exists_card_eq - simp_rw [Nat.card_eq_fintype_card] at hn hm ⊢ + obtain ⟨m, hm⟩ := h.exists_card_eq + simp_rw [Nat.card_eq_fintype_card] at hm hn ⊢ refine exists_subgroup_card_pow_prime _ ?_ - rw [hn] at hm ⊢ - exact pow_dvd_pow _ $ (pow_le_pow_iff_right hp.one_lt).1 hm - -lemma exists_subgroup_le_card_pow_prime_of_le_card {m p : ℕ} (hp : p.Prime) (h : IsPGroup p G) - {H : Subgroup G} (hm : p ^ m ≤ Nat.card H) : ∃ H' ≤ H, Nat.card H' = p ^ m := by - obtain ⟨H', H'card⟩ := exists_subgroup_card_pow_prime_of_le_card hp (h.to_subgroup H) hm + rw [hm] at hn ⊢ + exact pow_dvd_pow _ $ (pow_le_pow_iff_right hp.one_lt).1 hn + +/-- A special case of **Sylow's first theorem**. If `G` is a `p`-group and `H` a subgroup of size at +least `p ^ n` then there is a subgroup of `H` of cardinality `p ^ n`. -/ +lemma exists_subgroup_le_card_pow_prime_of_le_card {n p : ℕ} (hp : p.Prime) (h : IsPGroup p G) + {H : Subgroup G} (hn : p ^ n ≤ Nat.card H) : ∃ H' ≤ H, Nat.card H' = p ^ n := by + obtain ⟨H', H'card⟩ := exists_subgroup_card_pow_prime_of_le_card hp (h.to_subgroup H) hn refine ⟨H'.map H.subtype, map_subtype_le _, ?_⟩ rw [← H'card] let e : H' ≃* H'.map H.subtype := H'.equivMapOfInjective (Subgroup.subtype H) H.subtype_injective exact Nat.card_congr e.symm.toEquiv +/-- A special case of **Sylow's first theorem**. If `G` is a `p`-group and `H` a subgroup of size at +least `k` then there is a subgroup of `H` of cardinality between `k / p` and `k`. -/ lemma exists_subgroup_le_card_le {k p : ℕ} (hp : p.Prime) (h : IsPGroup p G) {H : Subgroup G} (hk : k ≤ Nat.card H) (hk₀ : k ≠ 0) : ∃ H' ≤ H, Nat.card H' ≤ k ∧ k < p * Nat.card H' := by obtain ⟨m, hmk, hkm⟩ : ∃ s, p ^ s ≤ k ∧ k < p ^ (s + 1) := diff --git a/Mathlib/LinearAlgebra/FreeModule/Finite/Rank.lean b/Mathlib/LinearAlgebra/FreeModule/Finite/Rank.lean index 046c7e91812e0..a05dfe14b478b 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Finite/Rank.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Finite/Rank.lean @@ -615,21 +615,22 @@ end DivisionRing section ZeroRank -variable [Ring R] [StrongRankCondition R] [AddCommGroup V] [Module R V] - -attribute [local instance] nontrivial_of_invariantBasisNumber +variable [Ring R] [AddCommGroup V] [Module R V] open FiniteDimensional theorem Module.finite_of_rank_eq_nat [Module.Free R V] {n : ℕ} (h : Module.rank R V = n) : Module.Finite R V := by - have := Cardinal.mk_lt_aleph0_iff.mp - (((Free.rank_eq_card_chooseBasisIndex R V).symm.trans h).trans_lt (nat_lt_aleph0 n)) - exact Module.Finite.of_basis (Free.chooseBasis R V) + nontriviality R + obtain ⟨⟨ι, b⟩⟩ := Module.Free.exists_basis (R := R) (M := V) + have := mk_lt_aleph0_iff.mp <| cardinal_le_rank_of_linearIndependent + b.linearIndependent |>.trans_eq h |>.trans_lt <| nat_lt_aleph0 n + exact Module.Finite.of_basis b theorem Module.finite_of_rank_eq_zero [NoZeroSMulDivisors R V] (h : Module.rank R V = 0) : Module.Finite R V := by + nontriviality R rw [rank_zero_iff] at h infer_instance @@ -651,12 +652,14 @@ variable {R V} theorem Submodule.bot_eq_top_of_rank_eq_zero [NoZeroSMulDivisors R V] (h : Module.rank R V = 0) : (⊥ : Submodule R V) = ⊤ := by + nontriviality R rw [rank_zero_iff] at h exact Subsingleton.elim _ _ #align bot_eq_top_of_rank_eq_zero Submodule.bot_eq_top_of_rank_eq_zero +/-- See `rank_subsingleton` for the reason that `Nontrivial R` is needed. -/ @[simp] -theorem Submodule.rank_eq_zero [NoZeroSMulDivisors R V] {S : Submodule R V} : +theorem Submodule.rank_eq_zero [Nontrivial R] [NoZeroSMulDivisors R V] {S : Submodule R V} : Module.rank R S = 0 ↔ S = ⊥ := ⟨fun h => (Submodule.eq_bot_iff _).2 fun x hx => @@ -667,7 +670,8 @@ theorem Submodule.rank_eq_zero [NoZeroSMulDivisors R V] {S : Submodule R V} : #align rank_eq_zero Submodule.rank_eq_zero @[simp] -theorem Submodule.finrank_eq_zero [NoZeroSMulDivisors R V] {S : Submodule R V} [Module.Finite R S] : +theorem Submodule.finrank_eq_zero [StrongRankCondition R] [NoZeroSMulDivisors R V] + {S : Submodule R V} [Module.Finite R S] : finrank R S = 0 ↔ S = ⊥ := by rw [← Submodule.rank_eq_zero, ← finrank_eq_rank, ← @Nat.cast_zero Cardinal, Cardinal.natCast_inj] #align finrank_eq_zero Submodule.finrank_eq_zero @@ -678,7 +682,7 @@ namespace Submodule open IsNoetherian FiniteDimensional -variable [AddCommGroup V] [Ring R] [StrongRankCondition R] [Module R V] +variable [Ring R] [AddCommGroup V] [Module R V] theorem fg_iff_finite (s : Submodule R V) : s.FG ↔ Module.Finite R s := (finite_def.trans (fg_top s)).symm @@ -714,7 +718,7 @@ end Submodule section -variable [Ring R] [StrongRankCondition R] [AddCommGroup V] [Module R V] +variable [Ring R] [AddCommGroup V] [Module R V] instance Module.Finite.finsupp {ι : Type*} [_root_.Finite ι] [Module.Finite R V] : Module.Finite R (ι →₀ V) := diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 6bf9c137dbc8a..3963650a42f8f 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -630,7 +630,7 @@ def piUnique [Unique α] (β : α → Sort*) : (∀ i, β i) ≃ β default wher right_inv x := rfl /-- If `α` has a unique term, then the type of function `α → β` is equivalent to `β`. -/ -@[simps! (config := .asFn) apply] +@[simps! (config := .asFn) apply symm_apply] def funUnique (α β) [Unique.{u} α] : (α → β) ≃ β := piUnique _ #align equiv.fun_unique Equiv.funUnique #align equiv.fun_unique_apply Equiv.funUnique_apply diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 3a0489aaaa9b7..1859d0aabf727 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -101,11 +101,11 @@ theorem TopologicalSpace.IsTopologicalBasis.borel_eq_generateFrom [TopologicalSp borel_eq_generateFrom_of_subbasis hs.eq_generateFrom #align topological_space.is_topological_basis.borel_eq_generate_from TopologicalSpace.IsTopologicalBasis.borel_eq_generateFrom -theorem isPiSystem_isOpen [TopologicalSpace α] : IsPiSystem (IsOpen : Set α → Prop) := +theorem isPiSystem_isOpen [TopologicalSpace α] : IsPiSystem ({s : Set α | IsOpen s}) := fun _s hs _t ht _ => IsOpen.inter hs ht #align is_pi_system_is_open isPiSystem_isOpen -lemma isPiSystem_isClosed [TopologicalSpace α] : IsPiSystem (IsClosed : Set α → Prop) := +lemma isPiSystem_isClosed [TopologicalSpace α] : IsPiSystem ({s : Set α | IsClosed s}) := fun _s hs _t ht _ ↦ IsClosed.inter hs ht theorem borel_eq_generateFrom_isClosed [TopologicalSpace α] : @@ -130,7 +130,9 @@ theorem borel_eq_generateFrom_Iio : borel α = .generateFrom (range Iio) := by letI : MeasurableSpace α := MeasurableSpace.generateFrom (range Iio) have H : ∀ a : α, MeasurableSet (Iio a) := fun a => GenerateMeasurable.basic _ ⟨_, rfl⟩ refine' generateFrom_le _ - rintro _ ⟨a, rfl | rfl⟩ <;> [skip; apply H] + rintro _ ⟨a, rfl | rfl⟩ + swap + · apply H by_cases h : ∃ a', ∀ b, a < b ↔ a' ≤ b · rcases h with ⟨a', ha'⟩ rw [(_ : Ioi a = (Iio a')ᶜ)] @@ -151,7 +153,6 @@ theorem borel_eq_generateFrom_Iio : borel α = .generateFrom (range Iio) := by ⟨fun ab => le_of_not_lt fun h' => h ⟨b, ab, h'⟩, lt_of_lt_of_le ax⟩⟩) with ⟨a', h₁, h₂⟩ · exact ⟨a', h₁, le_of_lt h₂⟩ rw [this] - skip apply MeasurableSet.iUnion exact fun _ => (H _).compl · rw [forall_range_iff] @@ -1285,7 +1286,7 @@ protected theorem Monotone.measurable [LinearOrder β] [OrderClosedTopology β] theorem aemeasurable_restrict_of_monotoneOn [LinearOrder β] [OrderClosedTopology β] {μ : Measure β} {s : Set β} (hs : MeasurableSet s) {f : β → α} (hf : MonotoneOn f s) : AEMeasurable f (μ.restrict s) := - have : Monotone (f ∘ (↑) : s → α) := fun ⟨x, hx⟩ ⟨y, hy⟩=> fun (hxy : x ≤ y) => hf hx hy hxy + have : Monotone (f ∘ (↑) : s → α) := fun ⟨x, hx⟩ ⟨y, hy⟩ => fun (hxy : x ≤ y) => hf hx hy hxy aemeasurable_restrict_of_measurable_subtype hs this.measurable #align ae_measurable_restrict_of_monotone_on aemeasurable_restrict_of_monotoneOn diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean index f4f2064acf9d4..72fb2c01d96f1 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean @@ -694,7 +694,7 @@ theorem prodAssoc_prod [SFinite τ] : (sFiniteSeq μ p.1.1).prod ((sFiniteSeq ν p.1.2).prod (sFiniteSeq τ p.2))) := by ext s hs rw [sum_apply _ hs, sum_apply _ hs, ← (Equiv.prodAssoc _ _ _).tsum_eq] - rfl + simp only [Equiv.prodAssoc_apply] rw [← sum_sFiniteSeq μ, ← sum_sFiniteSeq ν, ← sum_sFiniteSeq τ, prod_sum, prod_sum, map_sum MeasurableEquiv.prodAssoc.measurable.aemeasurable, prod_sum, prod_sum, this] congr @@ -820,8 +820,8 @@ theorem skew_product [SFinite μa] [SFinite μc] {f : α → β} (hf : MeasurePr ← hf.lintegral_comp (measurable_measure_prod_mk_left hs)] apply lintegral_congr_ae filter_upwards [hg] with a ha - rw [← ha, map_apply hgm.of_uncurry_left (measurable_prod_mk_left hs)] - rfl + rw [← ha, map_apply hgm.of_uncurry_left (measurable_prod_mk_left hs), preimage_preimage, + preimage_preimage] #align measure_theory.measure_preserving.skew_product MeasureTheory.MeasurePreserving.skew_product /-- If `f : α → β` sends the measure `μa` to `μb` and `g : γ → δ` sends the measure `μc` to `μd`, diff --git a/Mathlib/MeasureTheory/Constructions/Projective.lean b/Mathlib/MeasureTheory/Constructions/Projective.lean new file mode 100644 index 0000000000000..74f1c69477b4a --- /dev/null +++ b/Mathlib/MeasureTheory/Constructions/Projective.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne, Peter Pfaffelhuber +-/ +import Mathlib.MeasureTheory.Constructions.Cylinders + +/-! +# Projective measure families and projective limits + +A family of measures indexed by finite sets of `ι` is projective if, for finite sets `J ⊆ I`, +the projection from `∀ i : I, α i` to `∀ i : J, α i` maps `P I` to `P J`. +A measure `μ` is the projective limit of such a family of measures if for all `I : Finset ι`, +the projection from `∀ i, α i` to `∀ i : I, α i` maps `μ` to `P I`. + +## Main definitions + +* `MeasureTheory.IsProjectiveMeasureFamily`: `P : ∀ J : Finset ι, Measure (∀ j : J, α j)` is + projective if the projection from `∀ i : I, α i` to `∀ i : J, α i` maps `P I` to `P J` + for all `J ⊆ I`. +* `MeasureTheory.IsProjectiveLimit`: `μ` is the projective limit of the measure family `P` if for + all `I : Finset ι`, the map of `μ` by the projection to `I` is `P I`. + +## Main statements + +* `MeasureTheory.IsProjectiveLimit.unique`: the projective limit of a family of finite measures + is unique. + +-/ + +open Set + +namespace MeasureTheory + +variable {ι : Type*} {α : ι → Type*} [∀ i, MeasurableSpace (α i)] + {P : ∀ J : Finset ι, Measure (∀ j : J, α j)} + +/-- A family of measures indexed by finite sets of `ι` is projective if, for finite sets `J ⊆ I`, +the projection from `∀ i : I, α i` to `∀ i : J, α i` maps `P I` to `P J`. -/ +def IsProjectiveMeasureFamily [∀ i, MeasurableSpace (α i)] + (P : ∀ J : Finset ι, Measure (∀ j : J, α j)) : Prop := + ∀ (I J : Finset ι) (hJI : J ⊆ I), + P J = (P I).map (fun (x : ∀ i : I, α i) (j : J) ↦ x ⟨j, hJI j.2⟩) + +namespace IsProjectiveMeasureFamily + +variable {I J : Finset ι} + +/-- Auxiliary lemma for `measure_univ_eq`. -/ +lemma measure_univ_eq_of_subset (hP : IsProjectiveMeasureFamily P) (hJI : J ⊆ I) : + P I univ = P J univ := by + classical + have : (univ : Set (∀ i : I, α i)) = + (fun x : ∀ i : I, α i ↦ fun i : J ↦ x ⟨i, hJI i.2⟩) ⁻¹' (univ : Set (∀ i : J, α i)) := + by rw [preimage_univ] + rw [this, ← Measure.map_apply _ MeasurableSet.univ] + · rw [hP I J hJI] + · exact measurable_pi_lambda _ (fun _ ↦ measurable_pi_apply _) + +lemma measure_univ_eq (hP : IsProjectiveMeasureFamily P) (I J : Finset ι) : + P I univ = P J univ := by + classical + rw [← hP.measure_univ_eq_of_subset (Finset.subset_union_left I J), + ← hP.measure_univ_eq_of_subset (Finset.subset_union_right I J)] + +lemma congr_cylinder_of_subset (hP : IsProjectiveMeasureFamily P) + {S : Set (∀ i : I, α i)} {T : Set (∀ i : J, α i)} (hT : MeasurableSet T) + (h_eq : cylinder I S = cylinder J T) (hJI : J ⊆ I) : + P I S = P J T := by + cases isEmpty_or_nonempty (∀ i, α i) with + | inl h => + suffices ∀ I, P I univ = 0 by + simp only [Measure.measure_univ_eq_zero] at this + simp [this] + intro I + simp only [isEmpty_pi] at h + obtain ⟨i, hi_empty⟩ := h + rw [measure_univ_eq hP I {i}] + have : (univ : Set ((j : {x // x ∈ ({i} : Finset ι)}) → α j)) = ∅ := by simp [hi_empty] + simp [this] + | inr h => + have : S = (fun f : ∀ i : I, α i ↦ fun j : J ↦ f ⟨j, hJI j.prop⟩) ⁻¹' T := + eq_of_cylinder_eq_of_subset h_eq hJI + rw [hP I J hJI, Measure.map_apply _ hT, this] + exact measurable_pi_lambda _ (fun _ ↦ measurable_pi_apply _) + +lemma congr_cylinder (hP : IsProjectiveMeasureFamily P) + {S : Set (∀ i : I, α i)} {T : Set (∀ i : J, α i)} (hS : MeasurableSet S) (hT : MeasurableSet T) + (h_eq : cylinder I S = cylinder J T) : + P I S = P J T := by + classical + let U := (fun f : ∀ i : (I ∪ J : Finset ι), α i + ↦ fun j : I ↦ f ⟨j, Finset.mem_union_left J j.prop⟩) ⁻¹' S ∩ + (fun f ↦ fun j : J ↦ f ⟨j, Finset.mem_union_right I j.prop⟩) ⁻¹' T + suffices : P (I ∪ J) U = P I S ∧ P (I ∪ J) U = P J T + exact this.1.symm.trans this.2 + constructor + · have h_eq_union : cylinder I S = cylinder (I ∪ J) U := by + rw [← inter_cylinder, h_eq, inter_self] + exact hP.congr_cylinder_of_subset hS h_eq_union.symm (Finset.subset_union_left _ _) + · have h_eq_union : cylinder J T = cylinder (I ∪ J) U := by + rw [← inter_cylinder, h_eq, inter_self] + exact hP.congr_cylinder_of_subset hT h_eq_union.symm (Finset.subset_union_right _ _) + +end IsProjectiveMeasureFamily + +/-- A measure `μ` is the projective limit of a family of measures indexed by finite sets of `ι` if +for all `I : Finset ι`, the projection from `∀ i, α i` to `∀ i : I, α i` maps `μ` to `P I`. -/ +def IsProjectiveLimit (μ : Measure (∀ i, α i)) + (P : ∀ J : Finset ι, Measure (∀ j : J, α j)) : Prop := + ∀ I : Finset ι, (μ.map fun x : ∀ i, α i ↦ fun i : I ↦ x i) = P I + +namespace IsProjectiveLimit + +variable {μ ν : Measure (∀ i, α i)} + +lemma measure_cylinder (h : IsProjectiveLimit μ P) + (I : Finset ι) {s : Set (∀ i : I, α i)} (hs : MeasurableSet s) : + μ (cylinder I s) = P I s := by + rw [cylinder, ← Measure.map_apply _ hs, h I] + exact measurable_pi_lambda _ (fun _ ↦ measurable_pi_apply _) + +lemma measure_univ_eq (hμ : IsProjectiveLimit μ P) (I : Finset ι) : + μ univ = P I univ := by + rw [← cylinder_univ I, hμ.measure_cylinder _ MeasurableSet.univ] + +lemma isFiniteMeasure [∀ i, IsFiniteMeasure (P i)] (hμ : IsProjectiveLimit μ P) : + IsFiniteMeasure μ := by + constructor + rw [hμ.measure_univ_eq (∅ : Finset ι)] + exact measure_lt_top _ _ + +lemma isProbabilityMeasure [∀ i, IsProbabilityMeasure (P i)] (hμ : IsProjectiveLimit μ P) : + IsProbabilityMeasure μ := by + constructor + rw [hμ.measure_univ_eq (∅ : Finset ι)] + exact measure_univ + +lemma measure_univ_unique (hμ : IsProjectiveLimit μ P) (hν : IsProjectiveLimit ν P) : + μ univ = ν univ := by + rw [hμ.measure_univ_eq (∅ : Finset ι), hν.measure_univ_eq (∅ : Finset ι)] + +/-- The projective limit of a family of finite measures is unique. -/ +theorem unique [∀ i, IsFiniteMeasure (P i)] + (hμ : IsProjectiveLimit μ P) (hν : IsProjectiveLimit ν P) : + μ = ν := by + haveI : IsFiniteMeasure μ := hμ.isFiniteMeasure + refine ext_of_generate_finite (measurableCylinders α) generateFrom_measurableCylinders.symm + isPiSystem_measurableCylinders (fun s hs ↦ ?_) (hμ.measure_univ_unique hν) + obtain ⟨I, S, hS, rfl⟩ := (mem_measurableCylinders _).mp hs + rw [hμ.measure_cylinder _ hS, hν.measure_cylinder _ hS] + +end IsProjectiveLimit + +end MeasureTheory diff --git a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean index 2d7ffe0b4235d..d5d43fd1a86d3 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean @@ -130,8 +130,7 @@ theorem singularPart_mutuallySingular (s : SignedMeasure α) (μ : Measure α) : s.toJordanDecomposition.posPart.singularPart μ ⟂ₘ s.toJordanDecomposition.negPart.singularPart μ := by by_cases hl : s.HaveLebesgueDecomposition μ - · haveI := hl - obtain ⟨i, hi, hpos, hneg⟩ := s.toJordanDecomposition.mutuallySingular + · obtain ⟨i, hi, hpos, hneg⟩ := s.toJordanDecomposition.mutuallySingular rw [s.toJordanDecomposition.posPart.haveLebesgueDecomposition_add μ] at hpos rw [s.toJordanDecomposition.negPart.haveLebesgueDecomposition_add μ] at hneg rw [add_apply, add_eq_zero_iff] at hpos hneg @@ -153,16 +152,14 @@ theorem singularPart_totalVariation (s : SignedMeasure α) (μ : Measure α) : ⟨s.toJordanDecomposition.posPart.singularPart μ, s.toJordanDecomposition.negPart.singularPart μ, singularPart_mutuallySingular s μ⟩ := by refine' JordanDecomposition.toSignedMeasure_injective _ - rw [toSignedMeasure_toJordanDecomposition] - rfl + rw [toSignedMeasure_toJordanDecomposition, singularPart, JordanDecomposition.toSignedMeasure] · rw [totalVariation, this] #align measure_theory.signed_measure.singular_part_total_variation MeasureTheory.SignedMeasure.singularPart_totalVariation nonrec theorem mutuallySingular_singularPart (s : SignedMeasure α) (μ : Measure α) : singularPart s μ ⟂ᵥ μ.toENNRealVectorMeasure := by - rw [mutuallySingular_ennreal_iff, singularPart_totalVariation] - change _ ⟂ₘ VectorMeasure.equivMeasure.toFun (VectorMeasure.equivMeasure.invFun μ) - rw [VectorMeasure.equivMeasure.right_inv μ] + rw [mutuallySingular_ennreal_iff, singularPart_totalVariation, + VectorMeasure.ennrealToMeasure_toENNRealVectorMeasure] exact (mutuallySingular_singularPart _ _).add_left (mutuallySingular_singularPart _ _) #align measure_theory.signed_measure.mutually_singular_singular_part MeasureTheory.SignedMeasure.mutuallySingular_singularPart @@ -208,8 +205,8 @@ variable (s μ) /-- **The Lebesgue Decomposition theorem between a signed measure and a measure**: Given a signed measure `s` and a σ-finite measure `μ`, there exist a signed measure `t` and a measurable and integrable function `f`, such that `t` is mutually singular with respect to `μ` -and `s = t + μ.with_densityᵥ f`. In this case `t = s.singular_part μ` and -`f = s.rn_deriv μ`. -/ +and `s = t + μ.withDensityᵥ f`. In this case `t = s.singularPart μ` and +`f = s.rnDeriv μ`. -/ theorem singularPart_add_withDensity_rnDeriv_eq [s.HaveLebesgueDecomposition μ] : s.singularPart μ + μ.withDensityᵥ (s.rnDeriv μ) = s := by conv_rhs => @@ -239,11 +236,8 @@ theorem jordanDecomposition_add_withDensity_mutuallySingular {f : α → ℝ} (h (htμ : t ⟂ᵥ μ.toENNRealVectorMeasure) : (t.toJordanDecomposition.posPart + μ.withDensity fun x : α => ENNReal.ofReal (f x)) ⟂ₘ t.toJordanDecomposition.negPart + μ.withDensity fun x : α => ENNReal.ofReal (-f x) := by - rw [mutuallySingular_ennreal_iff, totalVariation_mutuallySingular_iff] at htμ - change - _ ⟂ₘ VectorMeasure.equivMeasure.toFun (VectorMeasure.equivMeasure.invFun μ) ∧ - _ ⟂ₘ VectorMeasure.equivMeasure.toFun (VectorMeasure.equivMeasure.invFun μ) at htμ - rw [VectorMeasure.equivMeasure.right_inv] at htμ + rw [mutuallySingular_ennreal_iff, totalVariation_mutuallySingular_iff, + VectorMeasure.ennrealToMeasure_toENNRealVectorMeasure] at htμ exact ((JordanDecomposition.mutuallySingular _).add_right (htμ.1.mono_ac (refl _) (withDensity_absolutelyContinuous _ _))).add_left @@ -308,20 +302,17 @@ private theorem eq_singularPart' (t : SignedMeasure α) {f : α → ℝ} (hf : M (hfi : Integrable f μ) (htμ : t ⟂ᵥ μ.toENNRealVectorMeasure) (hadd : s = t + μ.withDensityᵥ f) : t = s.singularPart μ := by have htμ' := htμ - rw [mutuallySingular_ennreal_iff, totalVariation_mutuallySingular_iff] at htμ - change - _ ⟂ₘ VectorMeasure.equivMeasure.toFun (VectorMeasure.equivMeasure.invFun μ) ∧ - _ ⟂ₘ VectorMeasure.equivMeasure.toFun (VectorMeasure.equivMeasure.invFun μ) at htμ - rw [VectorMeasure.equivMeasure.right_inv] at htμ - · rw [singularPart, ← t.toSignedMeasure_toJordanDecomposition, - JordanDecomposition.toSignedMeasure] - congr - · have hfpos : Measurable fun x => ENNReal.ofReal (f x) := by measurability - refine' eq_singularPart hfpos htμ.1 _ - rw [toJordanDecomposition_eq_of_eq_add_withDensity hf hfi htμ' hadd] - · have hfneg : Measurable fun x => ENNReal.ofReal (-f x) := by measurability - refine' eq_singularPart hfneg htμ.2 _ - rw [toJordanDecomposition_eq_of_eq_add_withDensity hf hfi htμ' hadd] + rw [mutuallySingular_ennreal_iff, totalVariation_mutuallySingular_iff, + VectorMeasure.ennrealToMeasure_toENNRealVectorMeasure] at htμ + rw [singularPart, ← t.toSignedMeasure_toJordanDecomposition, + JordanDecomposition.toSignedMeasure] + congr + · have hfpos : Measurable fun x => ENNReal.ofReal (f x) := by measurability + refine' eq_singularPart hfpos htμ.1 _ + rw [toJordanDecomposition_eq_of_eq_add_withDensity hf hfi htμ' hadd] + · have hfneg : Measurable fun x => ENNReal.ofReal (-f x) := by measurability + refine' eq_singularPart hfneg htμ.2 _ + rw [toJordanDecomposition_eq_of_eq_add_withDensity hf hfi htμ' hadd] /-- Given a measure `μ`, signed measures `s` and `t`, and a function `f` such that `t` is mutually singular with respect to `μ` and `s = t + μ.withDensityᵥ f`, we have @@ -371,21 +362,22 @@ theorem singularPart_smul_nnreal (s : SignedMeasure α) (μ : Measure α) (r : nonrec theorem singularPart_smul (s : SignedMeasure α) (μ : Measure α) (r : ℝ) : (r • s).singularPart μ = r • s.singularPart μ := by - by_cases hr : 0 ≤ r - · lift r to ℝ≥0 using hr + cases le_or_lt 0 r with + | inl hr => + lift r to ℝ≥0 using hr exact singularPart_smul_nnreal s μ r - · rw [singularPart, singularPart] + | inr hr => + rw [singularPart, singularPart] conv_lhs => congr · congr · rw [toJordanDecomposition_smul_real, - JordanDecomposition.real_smul_posPart_neg _ _ (not_le.1 hr), singularPart_smul] + JordanDecomposition.real_smul_posPart_neg _ _ hr, singularPart_smul] · congr · rw [toJordanDecomposition_smul_real, - JordanDecomposition.real_smul_negPart_neg _ _ (not_le.1 hr), singularPart_smul] - rw [toSignedMeasure_smul, toSignedMeasure_smul, ← neg_sub, ← smul_sub] - change -(((-r).toNNReal : ℝ) • (_ : SignedMeasure α)) = _ - rw [← neg_smul, Real.coe_toNNReal _ (le_of_lt (neg_pos.mpr (not_le.1 hr))), neg_neg] + JordanDecomposition.real_smul_negPart_neg _ _ hr, singularPart_smul] + rw [toSignedMeasure_smul, toSignedMeasure_smul, ← neg_sub, ← smul_sub, NNReal.smul_def, + ← neg_smul, Real.coe_toNNReal _ (le_of_lt (neg_pos.mpr hr)), neg_neg] #align measure_theory.signed_measure.singular_part_smul MeasureTheory.SignedMeasure.singularPart_smul theorem singularPart_add (s t : SignedMeasure α) (μ : Measure α) [s.HaveLebesgueDecomposition μ] @@ -395,8 +387,8 @@ theorem singularPart_add (s t : SignedMeasure α) (μ : Measure α) [s.HaveLebes (eq_singularPart _ (s.rnDeriv μ + t.rnDeriv μ) ((mutuallySingular_singularPart s μ).add_left (mutuallySingular_singularPart t μ)) _).symm - erw [withDensityᵥ_add (integrable_rnDeriv s μ) (integrable_rnDeriv t μ)] - rw [add_assoc, add_comm (t.singularPart μ), add_assoc, add_comm _ (t.singularPart μ), + rw [withDensityᵥ_add (integrable_rnDeriv s μ) (integrable_rnDeriv t μ), add_assoc, + add_comm (t.singularPart μ), add_assoc, add_comm _ (t.singularPart μ), singularPart_add_withDensity_rnDeriv_eq, ← add_assoc, singularPart_add_withDensity_rnDeriv_eq] #align measure_theory.signed_measure.singular_part_add MeasureTheory.SignedMeasure.singularPart_add @@ -417,7 +409,7 @@ theorem eq_rnDeriv (t : SignedMeasure α) (f : α → ℝ) (hfi : Integrable f have hadd' : s = t + μ.withDensityᵥ f' := by convert hadd using 2 exact WithDensityᵥEq.congr_ae hfi.1.ae_eq_mk.symm - haveI := haveLebesgueDecomposition_mk μ hfi.1.measurable_mk htμ hadd' + have := haveLebesgueDecomposition_mk μ hfi.1.measurable_mk htμ hadd' refine' (Integrable.ae_eq_of_withDensityᵥ_eq (integrable_rnDeriv _ _) hfi _).symm rw [← add_right_inj t, ← hadd, eq_singularPart _ f htμ hadd, singularPart_add_withDensity_rnDeriv_eq] @@ -437,11 +429,9 @@ theorem rnDeriv_smul (s : SignedMeasure α) (μ : Measure α) [s.HaveLebesgueDec refine' Integrable.ae_eq_of_withDensityᵥ_eq (integrable_rnDeriv _ _) ((integrable_rnDeriv _ _).smul r) _ - change _ = μ.withDensityᵥ ((r : ℝ) • s.rnDeriv μ) - rw [withDensityᵥ_smul (rnDeriv s μ) (r : ℝ), ← add_right_inj ((r • s).singularPart μ), - singularPart_add_withDensity_rnDeriv_eq, singularPart_smul] - change _ = _ + r • _ - rw [← smul_add, singularPart_add_withDensity_rnDeriv_eq] + rw [withDensityᵥ_smul (rnDeriv s μ) r, ← add_right_inj ((r • s).singularPart μ), + singularPart_add_withDensity_rnDeriv_eq, singularPart_smul, ← smul_add, + singularPart_add_withDensity_rnDeriv_eq] #align measure_theory.signed_measure.rn_deriv_smul MeasureTheory.SignedMeasure.rnDeriv_smul theorem rnDeriv_add (s t : SignedMeasure α) (μ : Measure α) [s.HaveLebesgueDecomposition μ] diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index 0afc446ce88dc..8cd27dd4d467d 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -154,7 +154,7 @@ theorem exists_closed_cover_approximatesLinearOn_of_hasFDerivWithinAt [SecondCou simpa only [sub_pos] using mem_ball_iff_norm.mp hz obtain ⟨δ, δpos, hδ⟩ : ∃ (δ : ℝ), 0 < δ ∧ ball x δ ∩ s ⊆ {y | ‖f y - f x - (f' x) (y - x)‖ ≤ ε * ‖y - x‖} := - Metric.mem_nhdsWithin_iff.1 (IsLittleO.def (hf' x xs) εpos) + Metric.mem_nhdsWithin_iff.1 ((hf' x xs).isLittleO.def εpos) obtain ⟨n, hn⟩ : ∃ n, u n < δ := ((tendsto_order.1 u_lim).2 _ δpos).exists refine' ⟨n, ⟨z, zT⟩, ⟨xs, _⟩⟩ intro y hy @@ -494,7 +494,7 @@ theorem _root_.ApproximatesLinearOn.norm_fderiv_sub_le {A : E →L[ℝ] E} {δ : (measure_closedBall_pos μ z εpos).ne' obtain ⟨ρ, ρpos, hρ⟩ : ∃ ρ > 0, ball x ρ ∩ s ⊆ {y : E | ‖f y - f x - (f' x) (y - x)‖ ≤ ε * ‖y - x‖} := - mem_nhdsWithin_iff.1 (IsLittleO.def (hf' x xs) εpos) + mem_nhdsWithin_iff.1 ((hf' x xs).isLittleO.def εpos) -- for small enough `r`, the rescaled ball `r • closedBall z ε` is included in the set where -- `f y - f x` is well approximated by `f' x (y - x)`. have B₂ : ∀ᶠ r in 𝓝[>] (0 : ℝ), {x} + r • closedBall z ε ⊆ ball x ρ := by diff --git a/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean b/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean index b1eaedb0a1fdf..b4d52eb14ae5f 100644 --- a/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean +++ b/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean @@ -401,7 +401,7 @@ theorem integral_eq_of_hasDerivWithinAt_off_countable_of_le (f f' : ℝ → E) { · exact fun x y => (OrderIso.funUnique (Fin 1) ℝ).symm.le_iff_le · exact (volume_preserving_funUnique (Fin 1) ℝ).symm _ · intro x; rw [Fin.sum_univ_one, hF', e_symm, Pi.single_eq_same, one_smul] - · rw [intervalIntegrable_iff_integrable_Ioc_of_le hle] at Hi + · rw [intervalIntegrable_iff_integrableOn_Ioc_of_le hle] at Hi exact Hi.congr_set_ae Ioc_ae_eq_Icc.symm _ = f b - f a := by simp only [Fin.sum_univ_one, e_symm] diff --git a/Mathlib/MeasureTheory/Integral/ExpDecay.lean b/Mathlib/MeasureTheory/Integral/ExpDecay.lean index fe702d2970c9e..53891a6243f24 100644 --- a/Mathlib/MeasureTheory/Integral/ExpDecay.lean +++ b/Mathlib/MeasureTheory/Integral/ExpDecay.lean @@ -46,7 +46,7 @@ theorem integrable_of_isBigO_exp_neg {f : ℝ → ℝ} {a b : ℝ} (h0 : 0 < b) let v := max a r -- show integrable on `(a, v]` from continuity have int_left : IntegrableOn f (Ioc a v) := by - rw [← intervalIntegrable_iff_integrable_Ioc_of_le (le_max_left a r)] + rw [← intervalIntegrable_iff_integrableOn_Ioc_of_le (le_max_left a r)] have u : Icc a v ⊆ Ici a := Icc_subset_Ici_self exact (h1.mono u).intervalIntegrable_of_Icc (le_max_left a r) suffices IntegrableOn f (Ioi v) by diff --git a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean index 4d1ddc8ff4110..edff5f8d486e1 100644 --- a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean +++ b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean @@ -149,7 +149,7 @@ open MeasureTheory Set Classical Filter Function open scoped Classical Topology Filter ENNReal BigOperators Interval NNReal -variable {ι 𝕜 E F A : Type*} [NormedAddCommGroup E] [CompleteSpace E] [NormedSpace ℝ E] +variable {ι 𝕜 E F A : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] namespace intervalIntegral @@ -275,6 +275,8 @@ theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae' [IsMeasurablyGenera (hv : Tendsto v lt l) : (fun t => (∫ x in u t..v t, f x ∂μ) - ∫ _ in u t..v t, c ∂μ) =o[lt] fun t => ∫ _ in u t..v t, (1 : ℝ) ∂μ := by + by_cases hE : CompleteSpace E; swap + · simp [intervalIntegral, integral, hE] have A := hf.integral_sub_linear_isLittleO_ae hfm hl (hu.Ioc hv) have B := hf.integral_sub_linear_isLittleO_ae hfm hl (hv.Ioc hu) simp_rw [integral_const', sub_smul] @@ -285,6 +287,8 @@ theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae' [IsMeasurablyGenera abel #align interval_integral.measure_integral_sub_linear_is_o_of_tendsto_ae' intervalIntegral.measure_integral_sub_linear_isLittleO_of_tendsto_ae' +variable [CompleteSpace E] + /-- **Fundamental theorem of calculus-1**, local version for any measure. Let filters `l` and `l'` be related by `TendstoIxxClass Ioc`. If `f` has a finite limit `c` at `l ⊓ μ.ae`, where `μ` is a measure @@ -484,7 +488,8 @@ at `(a, b)` provided that `f` is integrable on `a..b` and is continuous at `a` a -/ -variable {f : ℝ → E} {c ca cb : E} {l l' la la' lb lb' : Filter ℝ} {lt : Filter ι} {a b z : ℝ} +variable [CompleteSpace E] + {f : ℝ → E} {c ca cb : E} {l l' la la' lb lb' : Filter ℝ} {lt : Filter ι} {a b z : ℝ} {u v ua ub va vb : ι → ℝ} [FTCFilter a la la'] [FTCFilter b lb lb'] /-! @@ -834,7 +839,7 @@ theorem integral_hasFDerivWithinAt_of_tendsto_ae (hf : IntervalIntegrable f volu integral_sub_integral_sub_linear_isLittleO_of_tendsto_ae hf hmeas_a hmeas_b ha hb (tendsto_const_pure.mono_right FTCFilter.pure_le : Tendsto _ _ (𝓝[s] a)) tendsto_fst (tendsto_const_pure.mono_right FTCFilter.pure_le : Tendsto _ _ (𝓝[t] b)) tendsto_snd - refine' (this.congr_left _).trans_isBigO _ + refine .of_isLittleO <| (this.congr_left ?_).trans_isBigO ?_ · intro x; simp [sub_smul]; abel · exact isBigO_fst_prod.norm_left.add isBigO_snd_prod.norm_left #align interval_integral.integral_has_fderiv_within_at_of_tendsto_ae intervalIntegral.integral_hasFDerivWithinAt_of_tendsto_ae @@ -899,7 +904,7 @@ theorem integral_hasDerivWithinAt_of_tendsto_ae_right (hf : IntervalIntegrable f {s t : Set ℝ} [FTCFilter b (𝓝[s] b) (𝓝[t] b)] (hmeas : StronglyMeasurableAtFilter f (𝓝[t] b)) (hb : Tendsto f (𝓝[t] b ⊓ volume.ae) (𝓝 c)) : HasDerivWithinAt (fun u => ∫ x in a..u, f x) c s b := - integral_sub_integral_sub_linear_isLittleO_of_tendsto_ae_right hf hmeas hb + .of_isLittleO <| integral_sub_integral_sub_linear_isLittleO_of_tendsto_ae_right hf hmeas hb (tendsto_const_pure.mono_right FTCFilter.pure_le) tendsto_id #align interval_integral.integral_has_deriv_within_at_of_tendsto_ae_right intervalIntegral.integral_hasDerivWithinAt_of_tendsto_ae_right @@ -1093,10 +1098,10 @@ theorem sub_le_integral_of_hasDeriv_right_of_le_Ico (hab : a ≤ b) _ ≤ (∫ w in a..t, (G' w).toReal) + ∫ w in t..x, (G' w).toReal := (add_le_add ht.1 hx) _ = ∫ w in a..x, (G' w).toReal := by apply integral_add_adjacent_intervals - · rw [intervalIntegrable_iff_integrable_Ioc_of_le ht.2.1] + · rw [intervalIntegrable_iff_integrableOn_Ioc_of_le ht.2.1] exact IntegrableOn.mono_set G'int (Ioc_subset_Icc_self.trans (Icc_subset_Icc le_rfl ht.2.2.le)) - · rw [intervalIntegrable_iff_integrable_Ioc_of_le h'x.1.le] + · rw [intervalIntegrable_iff_integrableOn_Ioc_of_le h'x.1.le] apply IntegrableOn.mono_set G'int exact Ioc_subset_Icc_self.trans (Icc_subset_Icc ht.2.1 (h'x.2.trans (min_le_right _ _))) -- now that we know that `s` contains `[a, b]`, we get the desired result by applying this to `b`. @@ -1161,7 +1166,7 @@ theorem integral_eq_sub_of_hasDeriv_right_of_le_real (hab : a ≤ b) (sub_le_integral_of_hasDeriv_right_of_le hab hcont hderiv g'int fun _ _ => le_rfl) #align interval_integral.integral_eq_sub_of_has_deriv_right_of_le_real intervalIntegral.integral_eq_sub_of_hasDeriv_right_of_le_real -variable {f f' : ℝ → E} +variable [CompleteSpace E] {f f' : ℝ → E} /-- **Fundamental theorem of calculus-2**: If `f : ℝ → E` is continuous on `[a, b]` (where `a ≤ b`) and has a right derivative at `f' x` for all `x` in `(a, b)`, and `f'` is integrable on `[a, b]`, @@ -1173,7 +1178,7 @@ theorem integral_eq_sub_of_hasDeriv_right_of_le (hab : a ≤ b) (hcont : Continu rw [← g.intervalIntegral_comp_comm f'int, g.map_sub] exact integral_eq_sub_of_hasDeriv_right_of_le_real hab (g.continuous.comp_continuousOn hcont) (fun x hx => g.hasFDerivAt.comp_hasDerivWithinAt x (hderiv x hx)) - (g.integrable_comp ((intervalIntegrable_iff_integrable_Icc_of_le hab).1 f'int)) + (g.integrable_comp ((intervalIntegrable_iff_integrableOn_Icc_of_le hab).1 f'int)) #align interval_integral.integral_eq_sub_of_has_deriv_right_of_le intervalIntegral.integral_eq_sub_of_hasDeriv_right_of_le /-- Fundamental theorem of calculus-2: If `f : ℝ → E` is continuous on `[a, b]` and @@ -1362,19 +1367,22 @@ end Parts ### Integration by substitution / Change of variables -/ - section SMul +variable {G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] + /-- Change of variables, general form. If `f` is continuous on `[a, b]` and has right-derivative `f'` in `(a, b)`, `g` is continuous on `f '' (a, b)` and integrable on `f '' [a, b]`, and `f' x • (g ∘ f) x` is integrable on `[a, b]`, then we can substitute `u = f x` to get `∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. -/ -theorem integral_comp_smul_deriv''' {f f' : ℝ → ℝ} {g : ℝ → E} (hf : ContinuousOn f [[a, b]]) +theorem integral_comp_smul_deriv''' {f f' : ℝ → ℝ} {g : ℝ → G} (hf : ContinuousOn f [[a, b]]) (hff' : ∀ x ∈ Ioo (min a b) (max a b), HasDerivWithinAt f (f' x) (Ioi x) x) (hg_cont : ContinuousOn g (f '' Ioo (min a b) (max a b))) (hg1 : IntegrableOn g (f '' [[a, b]])) (hg2 : IntegrableOn (fun x => f' x • (g ∘ f) x) [[a, b]]) : (∫ x in a..b, f' x • (g ∘ f) x) = ∫ u in f a..f b, g u := by + by_cases hG : CompleteSpace G; swap + · simp [intervalIntegral, integral, hG] rw [hf.image_uIcc, ← intervalIntegrable_iff'] at hg1 have h_cont : ContinuousOn (fun u => ∫ t in f a..f u, g t) [[a, b]] := by refine' (continuousOn_primitive_interval' hg1 _).comp hf _ @@ -1415,7 +1423,7 @@ theorem integral_comp_smul_deriv''' {f f' : ℝ → ℝ} {g : ℝ → E} (hf : C continuous right-derivative `f'` in `(a, b)`, and `g` is continuous on `f '' [a, b]` then we can substitute `u = f x` to get `∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. -/ -theorem integral_comp_smul_deriv'' {f f' : ℝ → ℝ} {g : ℝ → E} (hf : ContinuousOn f [[a, b]]) +theorem integral_comp_smul_deriv'' {f f' : ℝ → ℝ} {g : ℝ → G} (hf : ContinuousOn f [[a, b]]) (hff' : ∀ x ∈ Ioo (min a b) (max a b), HasDerivWithinAt f (f' x) (Ioi x) x) (hf' : ContinuousOn f' [[a, b]]) (hg : ContinuousOn g (f '' [[a, b]])) : (∫ x in a..b, f' x • (g ∘ f) x) = ∫ u in f a..f b, g u := by @@ -1432,7 +1440,7 @@ and `g` is continuous on `f '' [a, b]`, then we can substitute `u = f x` to get Compared to `intervalIntegral.integral_comp_smul_deriv` we only require that `g` is continuous on `f '' [a, b]`. -/ -theorem integral_comp_smul_deriv' {f f' : ℝ → ℝ} {g : ℝ → E} +theorem integral_comp_smul_deriv' {f f' : ℝ → ℝ} {g : ℝ → G} (h : ∀ x ∈ uIcc a b, HasDerivAt f (f' x) x) (h' : ContinuousOn f' (uIcc a b)) (hg : ContinuousOn g (f '' [[a, b]])) : (∫ x in a..b, f' x • (g ∘ f) x) = ∫ x in f a..f b, g x := @@ -1444,7 +1452,7 @@ theorem integral_comp_smul_deriv' {f f' : ℝ → ℝ} {g : ℝ → E} and `g` is continuous, then we can substitute `u = f x` to get `∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. -/ -theorem integral_comp_smul_deriv {f f' : ℝ → ℝ} {g : ℝ → E} +theorem integral_comp_smul_deriv {f f' : ℝ → ℝ} {g : ℝ → G} (h : ∀ x ∈ uIcc a b, HasDerivAt f (f' x) x) (h' : ContinuousOn f' (uIcc a b)) (hg : Continuous g) : (∫ x in a..b, f' x • (g ∘ f) x) = ∫ x in f a..f b, g x := integral_comp_smul_deriv' h h' hg.continuousOn diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index 5caffbd50b781..9d483849f878b 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -683,7 +683,7 @@ theorem integral_Ioi_of_hasDerivAt_of_tendsto (hcont : ContinuousWithinAt f (Ici apply intervalIntegral.integral_eq_sub_of_hasDerivAt_of_le h'x (hcont.mono Icc_subset_Ici_self) fun y hy => hderiv y hy.1 - rw [intervalIntegrable_iff_integrable_Ioc_of_le h'x] + rw [intervalIntegrable_iff_integrableOn_Ioc_of_le h'x] exact f'int.mono (fun y hy => hy.1) le_rfl #align measure_theory.integral_Ioi_of_has_deriv_at_of_tendsto MeasureTheory.integral_Ioi_of_hasDerivAt_of_tendsto @@ -731,7 +731,7 @@ theorem integrableOn_Ioi_deriv_of_nonneg (hcont : ContinuousWithinAt g (Ici a) a symm apply intervalIntegral.integral_eq_sub_of_hasDerivAt_of_le h'x (hcont.mono Icc_subset_Ici_self) fun y hy => hderiv y hy.1 - rw [intervalIntegrable_iff_integrable_Ioc_of_le h'x] + rw [intervalIntegrable_iff_integrableOn_Ioc_of_le h'x] exact intervalIntegral.integrableOn_deriv_of_nonneg (hcont.mono Icc_subset_Ici_self) (fun y hy => hderiv y hy.1) fun y hy => g'pos y hy.1 _ = ∫ y in a..id x, ‖g' y‖ := by @@ -835,7 +835,7 @@ theorem integral_Iic_of_hasDerivAt_of_tendsto (hcont : ContinuousWithinAt f (Iic symm apply intervalIntegral.integral_eq_sub_of_hasDerivAt_of_le hx (hcont.mono Icc_subset_Iic_self) fun y hy => hderiv y hy.2 - rw [intervalIntegrable_iff_integrable_Ioc_of_le hx] + rw [intervalIntegrable_iff_integrableOn_Ioc_of_le hx] exact f'int.mono (fun y hy => hy.2) le_rfl /-- **Fundamental theorem of calculus-2**, on semi-infinite intervals `(-∞, a)`. @@ -867,7 +867,7 @@ open Real open scoped Interval -variable {E : Type*} {f : ℝ → E} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +variable {E : Type*} {f : ℝ → E} [NormedAddCommGroup E] [NormedSpace ℝ E] /-- Change-of-variables formula for `Ioi` integrals of vector-valued functions, proved by taking limits from the result for finite intervals. -/ diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 3e6ad1de93990..cbfc79abcb452 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -86,20 +86,28 @@ theorem IntervalIntegrable.def (h : IntervalIntegrable f μ a b) : IntegrableOn intervalIntegrable_iff.mp h #align interval_integrable.def IntervalIntegrable.def -theorem intervalIntegrable_iff_integrable_Ioc_of_le (hab : a ≤ b) : +theorem intervalIntegrable_iff_integrableOn_Ioc_of_le (hab : a ≤ b) : IntervalIntegrable f μ a b ↔ IntegrableOn f (Ioc a b) μ := by rw [intervalIntegrable_iff, uIoc_of_le hab] -#align interval_integrable_iff_integrable_Ioc_of_le intervalIntegrable_iff_integrable_Ioc_of_le +#align interval_integrable_iff_integrable_Ioc_of_le intervalIntegrable_iff_integrableOn_Ioc_of_le theorem intervalIntegrable_iff' [NoAtoms μ] : IntervalIntegrable f μ a b ↔ IntegrableOn f (uIcc a b) μ := by rw [intervalIntegrable_iff, ← Icc_min_max, uIoc, integrableOn_Icc_iff_integrableOn_Ioc] #align interval_integrable_iff' intervalIntegrable_iff' -theorem intervalIntegrable_iff_integrable_Icc_of_le {f : ℝ → E} {a b : ℝ} (hab : a ≤ b) +theorem intervalIntegrable_iff_integrableOn_Icc_of_le {f : ℝ → E} {a b : ℝ} (hab : a ≤ b) {μ : Measure ℝ} [NoAtoms μ] : IntervalIntegrable f μ a b ↔ IntegrableOn f (Icc a b) μ := by - rw [intervalIntegrable_iff_integrable_Ioc_of_le hab, integrableOn_Icc_iff_integrableOn_Ioc] -#align interval_integrable_iff_integrable_Icc_of_le intervalIntegrable_iff_integrable_Icc_of_le + rw [intervalIntegrable_iff_integrableOn_Ioc_of_le hab, integrableOn_Icc_iff_integrableOn_Ioc] +#align interval_integrable_iff_integrable_Icc_of_le intervalIntegrable_iff_integrableOn_Icc_of_le + +theorem intervalIntegrable_iff_integrableOn_Ico_of_le [NoAtoms μ] (hab : a ≤ b) : + IntervalIntegrable f μ a b ↔ IntegrableOn f (Ico a b) μ := by + rw [intervalIntegrable_iff_integrableOn_Icc_of_le hab, integrableOn_Icc_iff_integrableOn_Ico] + +theorem intervalIntegrable_iff_integrableOn_Ioo_of_le [NoAtoms μ] (hab : a ≤ b) : + IntervalIntegrable f μ a b ↔ IntegrableOn f (Ioo a b) μ := by + rw [intervalIntegrable_iff_integrableOn_Icc_of_le hab, integrableOn_Icc_iff_integrableOn_Ioo] /-- If a function is integrable with respect to a given measure `μ` then it is interval integrable with respect to `μ` on `uIcc a b`. -/ @@ -1197,7 +1205,8 @@ theorem continuousOn_primitive [NoAtoms μ] (h_int : IntegrableOn f (Icc a b) μ rw [continuousOn_congr this] intro x₀ _ refine' continuousWithinAt_primitive (measure_singleton x₀) _ - simp only [intervalIntegrable_iff_integrable_Ioc_of_le, min_eq_left, max_eq_right, h, min_self] + simp only [intervalIntegrable_iff_integrableOn_Ioc_of_le, min_eq_left, max_eq_right, h, + min_self] exact h_int.mono Ioc_subset_Icc_self le_rfl · rw [Icc_eq_empty h] exact continuousOn_empty _ diff --git a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean index 7ed4718d1723a..db55576954577 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean @@ -53,7 +53,7 @@ end ContinuousLinearEquiv variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [MeasurableSpace E] [BorelSpace E] [FiniteDimensional ℝ E] (μ : Measure E) [IsAddHaarMeasure μ] {F : Type*} [NormedAddCommGroup F] - [NormedSpace ℝ F] [CompleteSpace F] + [NormedSpace ℝ F] variable {s : Set E} @@ -62,6 +62,8 @@ integral of `f`. The formula we give works even when `f` is not integrable or `R thanks to the convention that a non-integrable function has integral zero. -/ theorem integral_comp_smul (f : E → F) (R : ℝ) : (∫ x, f (R • x) ∂μ) = |(R ^ finrank ℝ E)⁻¹| • ∫ x, f x ∂μ := by + by_cases hF : CompleteSpace F; swap + · simp [integral, hF] rcases eq_or_ne R 0 with (rfl | hR) · simp only [zero_smul, integral_const] rcases Nat.eq_zero_or_pos (finrank ℝ E) with (hE | hE) diff --git a/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean b/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean index 11f3a47892f38..6edc67e0db2ad 100644 --- a/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean +++ b/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean @@ -210,10 +210,11 @@ theorem measure_isClosed_eq_of_forall_lintegral_eq_of_isFiniteMeasure {Ω : Type (h : ∀ (f : Ω →ᵇ ℝ≥0), ∫⁻ x, f x ∂μ = ∫⁻ x, f x ∂ν) {F : Set Ω} (F_closed : IsClosed F) : μ F = ν F := by have ν_finite : IsFiniteMeasure ν := by + constructor have whole := h 1 simp only [BoundedContinuousFunction.coe_one, Pi.one_apply, coe_one, lintegral_const, one_mul] at whole - refine ⟨by simpa [← whole] using IsFiniteMeasure.measure_univ_lt_top⟩ + simpa [← whole] using IsFiniteMeasure.measure_univ_lt_top have obs_μ := HasOuterApproxClosed.tendsto_lintegral_apprSeq F_closed μ have obs_ν := HasOuterApproxClosed.tendsto_lintegral_apprSeq F_closed ν simp_rw [h] at obs_μ @@ -231,7 +232,6 @@ theorem ext_of_forall_lintegral_eq_of_IsFiniteMeasure {Ω : Type*} · exact fun F F_closed ↦ key F_closed · exact key isClosed_univ · rw [BorelSpace.measurable_eq (α := Ω), borel_eq_generateFrom_isClosed] - rfl end MeasureTheory -- namespace diff --git a/Mathlib/MeasureTheory/Measure/VectorMeasure.lean b/Mathlib/MeasureTheory/Measure/VectorMeasure.lean index 69aedbea53977..32477fa432965 100644 --- a/Mathlib/MeasureTheory/Measure/VectorMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/VectorMeasure.lean @@ -521,6 +521,17 @@ theorem ennrealToMeasure_apply {m : MeasurableSpace α} {v : VectorMeasure α rw [ennrealToMeasure, ofMeasurable_apply _ hs] #align measure_theory.vector_measure.ennreal_to_measure_apply MeasureTheory.VectorMeasure.ennrealToMeasure_apply +@[simp] +theorem _root_.MeasureTheory.Measure.toENNRealVectorMeasure_ennrealToMeasure + (μ : VectorMeasure α ℝ≥0∞) : + toENNRealVectorMeasure (ennrealToMeasure μ) = μ := ext fun s hs => by + rw [toENNRealVectorMeasure_apply_measurable hs, ennrealToMeasure_apply hs] + +@[simp] +theorem ennrealToMeasure_toENNRealVectorMeasure (μ : Measure α) : + ennrealToMeasure (toENNRealVectorMeasure μ) = μ := Measure.ext fun s hs => by + rw [ennrealToMeasure_apply hs, toENNRealVectorMeasure_apply_measurable hs] + /-- The equiv between `VectorMeasure α ℝ≥0∞` and `Measure α` formed by `MeasureTheory.VectorMeasure.ennrealToMeasure` and `MeasureTheory.Measure.toENNRealVectorMeasure`. -/ @@ -528,10 +539,8 @@ theorem ennrealToMeasure_apply {m : MeasurableSpace α} {v : VectorMeasure α def equivMeasure [MeasurableSpace α] : VectorMeasure α ℝ≥0∞ ≃ Measure α where toFun := ennrealToMeasure invFun := toENNRealVectorMeasure - left_inv _ := ext fun s hs => by - rw [toENNRealVectorMeasure_apply_measurable hs, ennrealToMeasure_apply hs] - right_inv _ := Measure.ext fun s hs => by - rw [ennrealToMeasure_apply hs, toENNRealVectorMeasure_apply_measurable hs] + left_inv := toENNRealVectorMeasure_ennrealToMeasure + right_inv := ennrealToMeasure_toENNRealVectorMeasure #align measure_theory.vector_measure.equiv_measure MeasureTheory.VectorMeasure.equivMeasure end diff --git a/Mathlib/MeasureTheory/SetSemiring.lean b/Mathlib/MeasureTheory/SetSemiring.lean new file mode 100644 index 0000000000000..e8f0ab12c8fe3 --- /dev/null +++ b/Mathlib/MeasureTheory/SetSemiring.lean @@ -0,0 +1,278 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne, Peter Pfaffelhuber +-/ +import Mathlib.Data.Set.Pairwise.Lattice +import Mathlib.Data.Real.ENNReal +import Mathlib.MeasureTheory.PiSystem + +/-! # Semirings of sets + +A semi-ring of sets `C` (in the sense of measure theory) is a family of sets containing `∅`, +stable by intersection and such that for all `s, t ∈ C`, `t \ s` is equal to a disjoint union of +finitely many sets in `C`. Note that a semi-ring of sets may not contain unions. + +An important example of a semi-ring of sets is intervals in `ℝ`. The intersection of two intervals +is an interval (possibly empty). The union of two intervals may not be an interval. +The set difference of two intervals may not be an interval, but it will be a disjoint union of +two intervals. + +## Main definitions + +* `MeasureTheory.IsSetSemiring C`: property of being a semi-ring of sets. +* `MeasureTheory.IsSetSemiring.diffFinset hs ht`: for `s, t` in a semi-ring `C` + (with `hC : IsSetSemiring C`) with `hs : s ∈ C`, `ht : t ∈ C`, this is a `Finset` of + pairwise disjoint sets such that `s \ t = ⋃₀ hC.diffFinset hs ht`. +* `MeasureTheory.IsSetSemiring.diffFinset₀ hs hI`: for `hs : s ∈ C` and a finset `I` of sets in `C` + (with `hI : ↑I ⊆ C`), this is a `Finset` of pairwise disjoint sets such that + `s \ ⋃₀ I = ⋃₀ hC.diffFinset₀ hs hI`. + +## Main statements + +* `MeasureTheory.IsSetSemiring.exists_disjoint_finset_diff_eq`: the existence of the `Finset` given + by the definition `IsSetSemiring.diffFinset₀` (see above). + +-/ + +open Finset Set + +open scoped ENNReal BigOperators + +namespace MeasureTheory + +variable {α : Type*} {C : Set (Set α)} {s t : Set α} + +/-- A semi-ring of sets `C` is a family of sets containing `∅`, stable by intersection and such that +for all `s, t ∈ C`, `s \ t` is equal to a disjoint union of finitely many sets in `C`. -/ +structure IsSetSemiring (C : Set (Set α)) : Prop where + empty_mem : ∅ ∈ C + inter_mem : ∀ s ∈ C, ∀ t ∈ C, s ∩ t ∈ C + diff_eq_Union' : ∀ s ∈ C, ∀ t ∈ C, + ∃ I : Finset (Set α), ↑I ⊆ C ∧ PairwiseDisjoint (I : Set (Set α)) id ∧ s \ t = ⋃₀ I + +namespace IsSetSemiring + +lemma isPiSystem (hC : IsSetSemiring C) : IsPiSystem C := fun s hs t ht _ ↦ hC.inter_mem s hs t ht + +section diffFinset + +open Classical in +/-- In a semi-ring of sets `C`, for all sets `s, t ∈ C`, `s \ t` is equal to a disjoint union of +finitely many sets in `C`. The finite set of sets in the union is not unique, but this definition +gives an arbitrary `Finset (Set α)` that satisfies the equality. + +We remove the empty set to ensure that `t ∉ hC.diffFinset hs ht` even if `t = ∅`. -/ +noncomputable def diffFinset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + Finset (Set α) := + (hC.diff_eq_Union' s hs t ht).choose \ {∅} + +lemma empty_not_mem_diffFinset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + ∅ ∉ hC.diffFinset hs ht := by + classical + simp only [diffFinset, mem_sdiff, Finset.mem_singleton, eq_self_iff_true, not_true, + and_false_iff, not_false_iff] + +lemma diffFinset_subset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + ↑(hC.diffFinset hs ht) ⊆ C := by + classical + simp only [diffFinset, coe_sdiff, coe_singleton, diff_singleton_subset_iff] + exact (hC.diff_eq_Union' s hs t ht).choose_spec.1.trans (Set.subset_insert _ _) + +lemma pairwiseDisjoint_diffFinset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + PairwiseDisjoint (hC.diffFinset hs ht : Set (Set α)) id := by + classical + simp only [diffFinset, coe_sdiff, coe_singleton] + exact Set.PairwiseDisjoint.subset (hC.diff_eq_Union' s hs t ht).choose_spec.2.1 + (Set.diff_subset _ _) + +lemma sUnion_diffFinset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + ⋃₀ hC.diffFinset hs ht = s \ t := by + classical + rw [(hC.diff_eq_Union' s hs t ht).choose_spec.2.2] + simp only [diffFinset, coe_sdiff, coe_singleton, diff_singleton_subset_iff] + rw [sUnion_diff_singleton_empty] + +lemma not_mem_diffFinset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + t ∉ hC.diffFinset hs ht := by + intro hs_mem + suffices t ⊆ s \ t by + have h := @disjoint_sdiff_self_right _ t s _ + specialize h le_rfl this + simp only [Set.bot_eq_empty, Set.le_eq_subset, subset_empty_iff] at h + refine hC.empty_not_mem_diffFinset hs ht ?_ + rwa [← h] + rw [← hC.sUnion_diffFinset hs ht] + exact subset_sUnion_of_mem hs_mem + +lemma sUnion_insert_diffFinset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) (hst : t ⊆ s) : + ⋃₀ insert t (hC.diffFinset hs ht) = s := by + conv_rhs => rw [← union_diff_cancel hst, ← hC.sUnion_diffFinset hs ht] + simp only [mem_coe, sUnion_insert] + +lemma disjoint_sUnion_diffFinset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + Disjoint t (⋃₀ hC.diffFinset hs ht) := by + rw [hC.sUnion_diffFinset] + exact disjoint_sdiff_right + +lemma pairwiseDisjoint_insert_diffFinset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + PairwiseDisjoint (insert t (hC.diffFinset hs ht) : Set (Set α)) id := by + have h := hC.pairwiseDisjoint_diffFinset hs ht + refine PairwiseDisjoint.insert_of_not_mem h (hC.not_mem_diffFinset hs ht) fun u hu ↦ ?_ + simp_rw [id.def] + refine Disjoint.mono_right ?_ (hC.disjoint_sUnion_diffFinset hs ht) + simp only [Set.le_eq_subset] + exact subset_sUnion_of_mem hu + +end diffFinset + +section diffFinset₀ + +variable {I : Finset (Set α)} + +/-- In a semiring of sets `C`, for all set `s ∈ C` and finite set of sets `I ⊆ C`, there is a +finite set of sets in `C` whose union is `s \ ⋃₀ I`. +See `IsSetSemiring.diffFinset₀` for a definition that gives such a set. -/ +lemma exists_disjoint_finset_diff_eq (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : + ∃ (J : Finset (Set α)) (_h_ss : ↑J ⊆ C) (_h_dis : PairwiseDisjoint (J : Set (Set α)) id), + s \ ⋃₀ I = ⋃₀ J := by + classical + revert hI + refine Finset.induction ?_ ?_ I + · intro + simp only [coe_empty, sUnion_empty, diff_empty, exists_prop] + refine ⟨{s}, singleton_subset_set_iff.mpr hs, ?_⟩ + simp only [coe_singleton, pairwiseDisjoint_singleton, sUnion_singleton, eq_self_iff_true, + and_self_iff] + intro t I' _ h h_insert_subset + rw [coe_insert] at h_insert_subset + have ht : t ∈ C := h_insert_subset (Set.mem_insert _ _) + obtain ⟨J, h_ss, h_dis, h_eq⟩ := h ((Set.subset_insert _ _).trans h_insert_subset) + let Ju : ∀ u ∈ C, Finset (Set α) := fun u hu ↦ hC.diffFinset hu ht + have hJu_subset : ∀ (u) (hu : u ∈ C), ↑(Ju u hu) ⊆ C := by + intro u hu x hx + exact hC.diffFinset_subset hu ht hx + have hJu_disj : ∀ (u) (hu : u ∈ C), (Ju u hu : Set (Set α)).PairwiseDisjoint id := fun u hu ↦ + hC.pairwiseDisjoint_diffFinset hu ht + have hJu_sUnion : ∀ (u) (hu : u ∈ C), ⋃₀ (Ju u hu : Set (Set α)) = u \ t := + fun u hu ↦ hC.sUnion_diffFinset hu ht + have hJu_disj' : ∀ (u) (hu : u ∈ C) (v) (hv : v ∈ C) (_h_dis : Disjoint u v), + Disjoint (⋃₀ (Ju u hu : Set (Set α))) (⋃₀ ↑(Ju v hv)) :=by + intro u hu v hv huv_disj + rw [hJu_sUnion, hJu_sUnion] + exact disjoint_of_subset (Set.diff_subset u t) (Set.diff_subset v t) huv_disj + let J' : Finset (Set α) := Finset.biUnion (Finset.univ : Finset J) fun u ↦ Ju u (h_ss u.prop) + have hJ'_subset : ↑J' ⊆ C := by + intro u + simp only [Subtype.coe_mk, univ_eq_attach, coe_biUnion, mem_coe, mem_attach, iUnion_true, + mem_iUnion, Finset.exists_coe, bex_imp] + intro v hv huvt + exact hJu_subset v (h_ss hv) huvt + refine ⟨J', hJ'_subset, ?_, ?_⟩ + · rw [Finset.coe_biUnion] + refine PairwiseDisjoint.biUnion ?_ ?_ + · simp only [univ_eq_attach, mem_coe, id.def, iSup_eq_iUnion] + simp_rw [PairwiseDisjoint, Set.Pairwise, Function.onFun] + intro x _ y _ hxy + have hxy_disj : Disjoint (x : Set α) y := by + by_contra h_contra + refine hxy ?_ + refine Subtype.ext ?_ + exact h_dis.elim x.prop y.prop h_contra + convert hJu_disj' (x : Set α) (h_ss x.prop) y (h_ss y.prop) hxy_disj + · rw [sUnion_eq_biUnion] + congr + · rw [sUnion_eq_biUnion] + congr + · exact fun u _ ↦ hJu_disj _ _ + · rw [coe_insert, sUnion_insert, Set.union_comm, ← Set.diff_diff, h_eq] + simp_rw [sUnion_eq_biUnion, Set.iUnion_diff] + simp only [Subtype.coe_mk, mem_coe, Finset.mem_biUnion, Finset.mem_univ, exists_true_left, + Finset.exists_coe, iUnion_exists, true_and] + rw [iUnion_comm] + refine iUnion_congr fun i ↦ ?_ + by_cases hi : i ∈ J + · simp only [hi, iUnion_true, exists_prop] + rw [← hJu_sUnion i (h_ss hi), sUnion_eq_biUnion] + simp only [mem_coe] + · simp only [hi, iUnion_of_empty, iUnion_empty] + +open Classical in +/-- In a semiring of sets `C`, for all set `s ∈ C` and finite set of sets `I ⊆ C`, +`diffFinset₀` is a finite set of sets in `C` such that `s \ ⋃₀ I = ⋃₀ (hC.diffFinset₀ hs I hI)`. +`diffFinset` is a special case of `diffFinset₀` where `I` is a singleton. -/ +noncomputable def diffFinset₀ (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : Finset (Set α) := + (hC.exists_disjoint_finset_diff_eq hs hI).choose \ {∅} + +lemma empty_not_mem_diffFinset₀ (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : + ∅ ∉ hC.diffFinset₀ hs hI := by + classical + simp only [diffFinset₀, mem_sdiff, Finset.mem_singleton, eq_self_iff_true, not_true, + and_false_iff, not_false_iff] + +lemma diffFinset₀_subset (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : + ↑(hC.diffFinset₀ hs hI) ⊆ C := by + classical + simp only [diffFinset₀, coe_sdiff, coe_singleton, diff_singleton_subset_iff] + exact (hC.exists_disjoint_finset_diff_eq hs hI).choose_spec.choose.trans (Set.subset_insert _ _) + +lemma pairwiseDisjoint_diffFinset₀ (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : + PairwiseDisjoint (hC.diffFinset₀ hs hI : Set (Set α)) id := by + classical + simp only [diffFinset₀, coe_sdiff, coe_singleton] + exact Set.PairwiseDisjoint.subset + (hC.exists_disjoint_finset_diff_eq hs hI).choose_spec.choose_spec.choose (Set.diff_subset _ _) + +lemma diff_sUnion_eq_sUnion_diffFinset₀ (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : + s \ ⋃₀ I = ⋃₀ hC.diffFinset₀ hs hI := by + classical + rw [(hC.exists_disjoint_finset_diff_eq hs hI).choose_spec.choose_spec.choose_spec] + simp only [diffFinset₀, coe_sdiff, coe_singleton, diff_singleton_subset_iff] + rw [sUnion_diff_singleton_empty] + +lemma sUnion_diffFinset₀_subset (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : + ⋃₀ (hC.diffFinset₀ hs hI : Set (Set α)) ⊆ s := by + rw [← hC.diff_sUnion_eq_sUnion_diffFinset₀] + exact diff_subset _ _ + +lemma disjoint_sUnion_diffFinset₀ (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : + Disjoint (⋃₀ (I : Set (Set α))) (⋃₀ hC.diffFinset₀ hs hI) := by + rw [← hC.diff_sUnion_eq_sUnion_diffFinset₀]; exact Set.disjoint_sdiff_right + +lemma disjoint_diffFinset₀ (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : + Disjoint I (hC.diffFinset₀ hs hI) := by + by_contra h + rw [Finset.not_disjoint_iff] at h + obtain ⟨u, huI, hu_diffFinset₀⟩ := h + have h_disj : u ≤ ⊥ := hC.disjoint_sUnion_diffFinset₀ hs hI (subset_sUnion_of_mem huI) + (subset_sUnion_of_mem hu_diffFinset₀) + simp only [Set.bot_eq_empty, Set.le_eq_subset, subset_empty_iff] at h_disj + refine hC.empty_not_mem_diffFinset₀ hs hI ?_ + rwa [h_disj] at hu_diffFinset₀ + +lemma pairwiseDisjoint_union_diffFinset₀ (hC : IsSetSemiring C) (hs : s ∈ C) + (hI : ↑I ⊆ C) (h_dis : PairwiseDisjoint (I : Set (Set α)) id) : + PairwiseDisjoint (I ∪ hC.diffFinset₀ hs hI : Set (Set α)) id := by + rw [pairwiseDisjoint_union] + refine ⟨h_dis, hC.pairwiseDisjoint_diffFinset₀ hs hI, fun u hu v hv _ ↦ ?_⟩ + simp_rw [id.def] + exact disjoint_of_subset (subset_sUnion_of_mem hu) (subset_sUnion_of_mem hv) + (hC.disjoint_sUnion_diffFinset₀ hs hI) + +lemma sUnion_union_sUnion_diffFinset₀_of_subset (hC : IsSetSemiring C) (hs : s ∈ C) + (hI : ↑I ⊆ C) (hI_ss : ⋃₀ ↑I ⊆ s) : + ⋃₀ I ∪ ⋃₀ hC.diffFinset₀ hs hI = s := by + conv_rhs => rw [← union_diff_cancel hI_ss, hC.diff_sUnion_eq_sUnion_diffFinset₀ hs hI] + +lemma sUnion_union_diffFinset₀_of_subset (hC : IsSetSemiring C) (hs : s ∈ C) + (hI : ↑I ⊆ C) (hI_ss : ⋃₀ ↑I ⊆ s) [DecidableEq (Set α)] : + ⋃₀ ↑(I ∪ hC.diffFinset₀ hs hI) = s := by + conv_rhs => rw [← sUnion_union_sUnion_diffFinset₀_of_subset hC hs hI hI_ss] + simp_rw [coe_union] + rw [sUnion_union] + +end diffFinset₀ + +end IsSetSemiring + +end MeasureTheory diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index d628a349a5c8e..6aaf80c8eab22 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -288,8 +288,9 @@ theorem map_div_left_divisors : n.divisors.map ⟨fun d => (n / d, d), fun p₁ p₂ => congr_arg Prod.snd⟩ = n.divisorsAntidiagonal := by apply Finset.map_injective (Equiv.prodComm _ _).toEmbedding + ext rw [map_swap_divisorsAntidiagonal, ← map_div_right_divisors, Finset.map_map] - rfl + simp #align nat.map_div_left_divisors Nat.map_div_left_divisors theorem sum_divisors_eq_sum_properDivisors_add_self : @@ -344,7 +345,7 @@ theorem divisors_injective : Function.Injective divisors := @[simp] theorem divisors_inj {a b : ℕ} : a.divisors = b.divisors ↔ a = b := - ⟨fun x => divisors_injective x, congrArg divisors⟩ + divisors_injective.eq_iff theorem eq_properDivisors_of_subset_of_sum_eq_sum {s : Finset ℕ} (hsub : s ⊆ n.properDivisors) : ((∑ x in s, x) = ∑ x in n.properDivisors, x) → s = n.properDivisors := by diff --git a/Mathlib/NumberTheory/ZetaFunction.lean b/Mathlib/NumberTheory/ZetaFunction.lean index ae85bd1b85b88..e6a04813ebfd8 100644 --- a/Mathlib/NumberTheory/ZetaFunction.lean +++ b/Mathlib/NumberTheory/ZetaFunction.lean @@ -172,8 +172,10 @@ theorem locally_integrable_zetaKernel₂ : LocallyIntegrableOn zetaKernel₂ (Io #align locally_integrable_zeta_kernel₂ locally_integrable_zetaKernel₂ /-- Functional equation for `zetaKernel₂`. -/ -theorem zetaKernel₂_one_div {t : ℝ} (ht : 0 < t) : +theorem zetaKernel₂_one_div {t : ℝ} (ht : 0 ≤ t) : zetaKernel₂ (1 / t) = sqrt t * zetaKernel₂ t := by + rcases ht.eq_or_lt with rfl|h't + · simp [zetaKernel₂, zetaKernel₁] have aux : ∀ {u : ℝ} (_ : 1 < u), zetaKernel₂ (1 / u) = sqrt u * zetaKernel₂ u := by intro u hu simp_rw [zetaKernel₂, Pi.add_apply] @@ -195,9 +197,9 @@ theorem zetaKernel₂_one_div {t : ℝ} (ht : 0 < t) : rcases lt_trichotomy 1 t with (h | h | h) · exact aux h · simp only [← h, div_self, Ne.def, one_ne_zero, not_false_iff, sqrt_one, ofReal_one, one_mul] - · have := aux (show 1 < 1 / t by rwa [lt_one_div (zero_lt_one' ℝ) ht, div_one]) + · have := aux (show 1 < 1 / t by rwa [lt_one_div (zero_lt_one' ℝ) h't, div_one]) rw [one_div_one_div] at this - rw [this, ← mul_assoc, ← ofReal_mul, ← sqrt_mul ht.le, mul_one_div_cancel ht.ne', sqrt_one, + rw [this, ← mul_assoc, ← ofReal_mul, ← sqrt_mul ht, mul_one_div_cancel h't.ne', sqrt_one, ofReal_one, one_mul] #align zeta_kernel₂_one_div zetaKernel₂_one_div @@ -240,7 +242,7 @@ theorem isBigO_zero_zetaKernel₂ : IsBigO (𝓝[>] 0) zetaKernel₂ fun t => ex simp_rw [← one_div] at h1 have h2 : zetaKernel₂ ∘ Div.div 1 =ᶠ[𝓝[>] 0] fun t => sqrt t * zetaKernel₂ t := eventually_of_mem self_mem_nhdsWithin fun t ht => by - dsimp only; rw [← zetaKernel₂_one_div ht]; rfl + dsimp only; rw [← zetaKernel₂_one_div (le_of_lt ht)]; rfl have h3 := h1.congr' h2 (EventuallyEq.refl _ _) have h4 := h3.mul (isBigO_refl (fun t : ℝ => 1 / (sqrt t : ℂ)) (𝓝[>] 0)).norm_right refine h4.congr' ?_ ?_ @@ -641,12 +643,12 @@ theorem riemannZeta_four : riemannZeta 4 = π ^ 4 / 90 := by `Λ₀(1 - s) = Λ₀ s`. -/ theorem riemannCompletedZeta₀_one_sub (s : ℂ) : riemannCompletedZeta₀ (1 - s) = riemannCompletedZeta₀ s := by - have := mellin_comp_rpow zetaKernel₂ (s / 2 - 1 / 2) neg_one_lt_zero.ne + have := mellin_comp_rpow zetaKernel₂ (s / 2 - 1 / 2) (-1) simp_rw [rpow_neg_one, ← one_div, abs_neg, abs_one, div_one, one_smul, ofReal_neg, ofReal_one, div_neg, div_one, neg_sub] at this conv_lhs => rw [riemannCompletedZeta₀, sub_div, ← this] refine set_integral_congr measurableSet_Ioi fun t ht => ?_ - simp_rw [zetaKernel₂_one_div ht, smul_eq_mul, ← mul_assoc, sqrt_eq_rpow, + simp_rw [zetaKernel₂_one_div (le_of_lt ht), smul_eq_mul, ← mul_assoc, sqrt_eq_rpow, ofReal_cpow (le_of_lt ht), ← cpow_add _ _ (ofReal_ne_zero.mpr <| ne_of_gt ht)] congr 2 push_cast @@ -797,3 +799,37 @@ theorem riemannZeta_neg_nat_eq_bernoulli (k : ℕ) : rw [inv_mul_cancel (ofReal_ne_zero.mpr <| pow_ne_zero _ pi_pos.ne'), inv_mul_cancel (ofReal_ne_zero.mpr <| pow_ne_zero _ two_ne_zero), one_mul, one_mul] #align riemann_zeta_neg_nat_eq_bernoulli riemannZeta_neg_nat_eq_bernoulli + +/-- The residue of `Λ(s)` at `s = 1` is equal to 1. -/ +lemma riemannCompletedZeta_residue_one : + Tendsto (fun s ↦ (s - 1) * riemannCompletedZeta s) (𝓝[≠] 1) (𝓝 1) := by + unfold riemannCompletedZeta + simp_rw [mul_add, mul_sub, (by simp : 𝓝 (1 : ℂ) = 𝓝 (0 - 0 + 1))] + refine ((Tendsto.sub ?_ ?_).mono_left nhdsWithin_le_nhds).add ?_ + · rw [(by simp : 𝓝 (0 : ℂ) = 𝓝 ((1 - 1) * riemannCompletedZeta₀ 1))] + apply ((continuous_sub_right _).mul differentiable_completed_zeta₀.continuous).tendsto + · rw [(by simp : 𝓝 (0 : ℂ) = 𝓝 ((1 - 1) * (1 / 1)))] + exact (((continuous_sub_right _).continuousAt).mul <| + continuousAt_const.div continuousAt_id one_ne_zero) + · refine (tendsto_const_nhds.mono_left nhdsWithin_le_nhds).congr' ?_ + refine eventually_nhdsWithin_of_forall (fun s hs ↦ ?_) + simpa using (div_self <| sub_ne_zero_of_ne <| Set.mem_compl_singleton_iff.mpr hs).symm + +/-- The residue of `ζ(s)` at `s = 1` is equal to 1. -/ +lemma riemannZeta_residue_one : Tendsto (fun s ↦ (s - 1) * riemannZeta s) (𝓝[≠] 1) (𝓝 1) := by + suffices : Tendsto (fun s => (s - 1) * + (π ^ (s / 2) * riemannCompletedZeta s / Gamma (s / 2))) (𝓝[≠] 1) (𝓝 1) + · refine this.congr' <| (eventually_ne_nhdsWithin one_ne_zero).mp (eventually_of_forall ?_) + intro x hx + simp [riemannZeta_def, Function.update_noteq hx] + have h0 : Tendsto (fun s ↦ π ^ (s / 2) : ℂ → ℂ) (𝓝[≠] 1) (𝓝 (π ^ (1 / 2 : ℂ))) + · refine ((continuousAt_id.div_const _).const_cpow ?_).tendsto.mono_left nhdsWithin_le_nhds + exact Or.inl <| ofReal_ne_zero.mpr pi_ne_zero + have h1 : Tendsto (fun s : ℂ ↦ 1 / Gamma (s / 2)) (𝓝[≠] 1) (𝓝 (1 / π ^ (1 / 2 : ℂ))) + · rw [← Complex.Gamma_one_half_eq] + refine (Continuous.tendsto ?_ _).mono_left nhdsWithin_le_nhds + simpa using differentiable_one_div_Gamma.continuous.comp (continuous_id.div_const _) + convert (riemannCompletedZeta_residue_one.mul h0).mul h1 using 2 with s + · ring + · have := fun h ↦ ofReal_ne_zero.mpr pi_ne_zero ((cpow_eq_zero_iff _ (1 / 2)).mp h).1 + field_simp diff --git a/Mathlib/Order/Closure.lean b/Mathlib/Order/Closure.lean index 2b9f73336ed62..af6f24417d78e 100644 --- a/Mathlib/Order/Closure.lean +++ b/Mathlib/Order/Closure.lean @@ -61,16 +61,22 @@ structure ClosureOperator [Preorder α] extends α →o α where le_closure' : ∀ x, x ≤ toFun x /-- Closures are idempotent -/ idempotent' : ∀ x, toFun (toFun x) = toFun x + /-- Predicate for an element to be closed. + + By default, this is defined as `c.IsClosed x := (c x = x)` (see `isClosed_iff`). + We allow an override to fix definitional equalities. -/ + IsClosed (x : α) : Prop := toFun x = x + isClosed_iff {x : α} : IsClosed x ↔ toFun x = x := by aesop #align closure_operator ClosureOperator namespace ClosureOperator instance [Preorder α] : OrderHomClass (ClosureOperator α) α α where coe c := c.1 - coe_injective' := by rintro ⟨⟩ ⟨⟩ h; congr; exact FunLike.ext' h + coe_injective' := by rintro ⟨⟩ ⟨⟩ h; obtain rfl := FunLike.ext' h; congr with x; simp [*] map_rel f _ _ h := f.mono h -initialize_simps_projections ClosureOperator (toFun → apply) +initialize_simps_projections ClosureOperator (toFun → apply, IsClosed → isClosed) section PartialOrder @@ -82,8 +88,11 @@ def id : ClosureOperator α where toOrderHom := OrderHom.id le_closure' _ := le_rfl idempotent' _ := rfl + IsClosed _ := True #align closure_operator.id ClosureOperator.id #align closure_operator.id_apply ClosureOperator.id_apply +#align closure_operator.closed ClosureOperator.IsClosed +#align closure_operator.mem_closed_iff ClosureOperator.isClosed_iff instance : Inhabited (ClosureOperator α) := ⟨id α⟩ @@ -91,9 +100,8 @@ instance : Inhabited (ClosureOperator α) := variable {α} [PartialOrder α] (c : ClosureOperator α) @[ext] -theorem ext : ∀ c₁ c₂ : ClosureOperator α, (c₁ : α → α) = (c₂ : α → α) → c₁ = c₂ - | ⟨⟨c₁, _⟩, _, _⟩, ⟨⟨c₂, _⟩, _, _⟩, h => by - congr +theorem ext : ∀ c₁ c₂ : ClosureOperator α, (c₁ : α → α) = (c₂ : α → α) → c₁ = c₂ := + FunLike.coe_injective #align closure_operator.ext ClosureOperator.ext /-- Constructor for a closure operator using the weaker idempotency axiom: `f (f x) ≤ f x`. -/ @@ -119,30 +127,20 @@ def mk₂ (f : α → α) (hf : ∀ x, x ≤ f x) (hmin : ∀ ⦃x y⦄, x ≤ f #align closure_operator.mk₂ ClosureOperator.mk₂ #align closure_operator.mk₂_apply ClosureOperator.mk₂_apply -/-- Expanded out version of `mk₂`. `p` implies being closed. This constructor should be used when -you already know a sufficient condition for being closed and using `mem_mk₃_closed` will avoid you -the (slight) hassle of having to prove it both inside and outside the constructor. -/ +/-- Construct a closure operator from an inflationary function `f` and a "closedness" predicate `p` +witnessing minimality of `f x` among closed elements greater than `x`. -/ @[simps!] -def mk₃ (f : α → α) (p : α → Prop) (hf : ∀ x, x ≤ f x) (hfp : ∀ x, p (f x)) - (hmin : ∀ ⦃x y⦄, x ≤ y → p y → f x ≤ y) : ClosureOperator α := - mk₂ f hf fun _ y hxy => hmin hxy (hfp y) -#align closure_operator.mk₃ ClosureOperator.mk₃ -#align closure_operator.mk₃_apply ClosureOperator.mk₃_apply - -/-- This lemma shows that the image of `x` of a closure operator built from the `mk₃` constructor -respects `p`, the property that was fed into it. -/ -theorem closure_mem_mk₃ {f : α → α} {p : α → Prop} {hf : ∀ x, x ≤ f x} {hfp : ∀ x, p (f x)} - {hmin : ∀ ⦃x y⦄, x ≤ y → p y → f x ≤ y} (x : α) : p (mk₃ f p hf hfp hmin x) := - hfp x -#align closure_operator.closure_mem_mk₃ ClosureOperator.closure_mem_mk₃ - -/-- Analogue of `closure_le_closed_iff_le` but with the `p` that was fed into the `mk₃` constructor. --/ -theorem closure_le_mk₃_iff {f : α → α} {p : α → Prop} {hf : ∀ x, x ≤ f x} {hfp : ∀ x, p (f x)} - {hmin : ∀ ⦃x y⦄, x ≤ y → p y → f x ≤ y} {x y : α} (hxy : x ≤ y) (hy : p y) : - mk₃ f p hf hfp hmin x ≤ y := - hmin hxy hy -#align closure_operator.closure_le_mk₃_iff ClosureOperator.closure_le_mk₃_iff +def ofPred (f : α → α) (p : α → Prop) (hf : ∀ x, x ≤ f x) (hfp : ∀ x, p (f x)) + (hmin : ∀ ⦃x y⦄, x ≤ y → p y → f x ≤ y) : ClosureOperator α where + __ := mk₂ f hf fun _ y hxy => hmin hxy (hfp y) + IsClosed := p + isClosed_iff := ⟨λ hx ↦ (hmin le_rfl hx).antisymm $ hf _, λ hx ↦ hx ▸ hfp _⟩ +#align closure_operator.mk₃ ClosureOperator.ofPred +#align closure_operator.mk₃_apply ClosureOperator.ofPred_apply +#align closure_operator.mem_mk₃_closed ClosureOperator.ofPred_isClosed + +#noalign closure_operator.closure_mem_ofPred +#noalign closure_operator.closure_le_ofPred_iff @[mono] theorem monotone : Monotone c := @@ -160,66 +158,48 @@ theorem idempotent (x : α) : c (c x) = c x := c.idempotent' x #align closure_operator.idempotent ClosureOperator.idempotent -theorem le_closure_iff (x y : α) : x ≤ c y ↔ c x ≤ c y := - ⟨fun h => c.idempotent y ▸ c.monotone h, fun h => (c.le_closure x).trans h⟩ -#align closure_operator.le_closure_iff ClosureOperator.le_closure_iff +@[simp] lemma isClosed_closure (x : α) : c.IsClosed (c x) := c.isClosed_iff.2 $ c.idempotent x +#align closure_operator.closure_is_closed ClosureOperator.isClosed_closure -/-- An element `x` is closed for the closure operator `c` if it is a fixed point for it. -/ -def closed : Set α := {x | c x = x} -#align closure_operator.closed ClosureOperator.closed +/-- The type of elements closed under a closure operator. -/ +abbrev Closeds := {x // c.IsClosed x} -theorem mem_closed_iff (x : α) : x ∈ c.closed ↔ c x = x := - Iff.rfl -#align closure_operator.mem_closed_iff ClosureOperator.mem_closed_iff +/-- Send an element to a closed element (by taking the closure). -/ +def toCloseds (x : α) : c.Closeds := ⟨c x, c.isClosed_closure x⟩ +#align closure_operator.to_closed ClosureOperator.toCloseds -theorem mem_closed_iff_closure_le (x : α) : x ∈ c.closed ↔ c x ≤ x := - ⟨le_of_eq, fun h => h.antisymm (c.le_closure x)⟩ -#align closure_operator.mem_closed_iff_closure_le ClosureOperator.mem_closed_iff_closure_le +variable {c} {x y : α} -theorem closure_eq_self_of_mem_closed {x : α} (h : x ∈ c.closed) : c x = x := - h -#align closure_operator.closure_eq_self_of_mem_closed ClosureOperator.closure_eq_self_of_mem_closed +theorem IsClosed.closure_eq : c.IsClosed x → c x = x := c.isClosed_iff.1 +#align closure_operator.closure_eq_self_of_mem_closed ClosureOperator.IsClosed.closure_eq -@[simp] -theorem closure_is_closed (x : α) : c x ∈ c.closed := - c.idempotent x -#align closure_operator.closure_is_closed ClosureOperator.closure_is_closed +theorem isClosed_iff_closure_le : c.IsClosed x ↔ c x ≤ x := + ⟨fun h ↦ h.closure_eq.le, fun h ↦ c.isClosed_iff.2 $ h.antisymm $ c.le_closure x⟩ +#align closure_operator.mem_closed_iff_closure_le ClosureOperator.isClosed_iff_closure_le /-- The set of closed elements for `c` is exactly its range. -/ -theorem closed_eq_range_close : c.closed = Set.range c := - Set.ext fun x => - ⟨fun h => ⟨x, h⟩, by - rintro ⟨y, rfl⟩ - apply c.idempotent⟩ -#align closure_operator.closed_eq_range_close ClosureOperator.closed_eq_range_close +theorem setOf_isClosed_eq_range_closure : {x | c.IsClosed x} = Set.range c := by + ext x; exact ⟨fun hx ↦ ⟨x, hx.closure_eq⟩, by rintro ⟨y, rfl⟩; exact c.isClosed_closure _⟩ +#align closure_operator.closed_eq_range_close ClosureOperator.setOf_isClosed_eq_range_closure -/-- Send an `x` to an element of the set of closed elements (by taking the closure). -/ -def toClosed (x : α) : c.closed := - ⟨c x, c.closure_is_closed x⟩ -#align closure_operator.to_closed ClosureOperator.toClosed +theorem le_closure_iff : x ≤ c y ↔ c x ≤ c y := + ⟨fun h ↦ c.idempotent y ▸ c.monotone h, (c.le_closure x).trans⟩ +#align closure_operator.le_closure_iff ClosureOperator.le_closure_iff @[simp] -theorem closure_le_closed_iff_le (x : α) {y : α} (hy : c.closed y) : c x ≤ y ↔ x ≤ y := by - rw [← c.closure_eq_self_of_mem_closed hy, ← le_closure_iff] -#align closure_operator.closure_le_closed_iff_le ClosureOperator.closure_le_closed_iff_le +theorem IsClosed.closure_le_iff (hy : c.IsClosed y) : c x ≤ y ↔ x ≤ y := by + rw [← hy.closure_eq, ← le_closure_iff] +#align closure_operator.closure_le_closed_iff_le ClosureOperator.IsClosed.closure_le_iff + +lemma closure_min (hxy : x ≤ y) (hy : c.IsClosed y) : c x ≤ y := hy.closure_le_iff.2 hxy /-- A closure operator is equal to the closure operator obtained by feeding `c.closed` into the -`mk₃` constructor. -/ -theorem eq_mk₃_closed (c : ClosureOperator α) : - c = - mk₃ c c.closed c.le_closure c.closure_is_closed fun x y hxy hy => - (c.closure_le_closed_iff_le x hy).2 hxy := by +`ofPred` constructor. -/ +theorem eq_ofPred_closed (c : ClosureOperator α) : + c = ofPred c c.IsClosed c.le_closure c.isClosed_closure fun x y ↦ closure_min := by ext rfl -#align closure_operator.eq_mk₃_closed ClosureOperator.eq_mk₃_closed - -/-- The property `p` fed into the `mk₃` constructor exactly corresponds to being closed. -/ -@[simp] theorem mem_mk₃_closed {f : α → α} {p : α → Prop} {hf hfp hmin} {x : α} : - x ∈ (mk₃ f p hf hfp hmin).closed ↔ p x := by - refine' ⟨λ hx ↦ _, λ hx ↦ (hmin le_rfl hx).antisymm (hf _)⟩ - rw [← closure_eq_self_of_mem_closed _ hx] - exact hfp _ -#align closure_operator.mem_mk₃_closed ClosureOperator.mem_mk₃_closedₓ +#align closure_operator.eq_mk₃_closed ClosureOperator.eq_ofPred_closed end PartialOrder @@ -234,9 +214,8 @@ theorem closure_top : c ⊤ = ⊤ := le_top.antisymm (c.le_closure _) #align closure_operator.closure_top ClosureOperator.closure_top -theorem top_mem_closed : ⊤ ∈ c.closed := - c.closure_top -#align closure_operator.top_mem_closed ClosureOperator.top_mem_closed +@[simp] lemma isClosed_top : c.IsClosed ⊤ := c.isClosed_iff.2 c.closure_top +#align closure_operator.top_mem_closed ClosureOperator.isClosed_top end OrderTop @@ -254,7 +233,7 @@ theorem closure_sup_closure_le (x y : α) : c x ⊔ c y ≤ c (x ⊔ y) := #align closure_operator.closure_sup_closure_le ClosureOperator.closure_sup_closure_le theorem closure_sup_closure_left (x y : α) : c (c x ⊔ y) = c (x ⊔ y) := - ((c.le_closure_iff _ _).1 + (le_closure_iff.1 (sup_le (c.monotone le_sup_left) (le_sup_right.trans (c.le_closure _)))).antisymm (c.monotone (sup_le_sup_right (c.le_closure _) _)) #align closure_operator.closure_sup_closure_left ClosureOperator.closure_sup_closure_left @@ -274,33 +253,23 @@ section CompleteLattice variable [CompleteLattice α] (c : ClosureOperator α) {p : α → Prop} /-- Define a closure operator from a predicate that's preserved under infima. -/ -def ofPred (p : α → Prop) (hsinf : ∀ s, (∀ a ∈ s, p a) → p (sInf s)) : ClosureOperator α := - ClosureOperator.mk₃ (fun a ↦ ⨅ b : {b // p b ∧ a ≤ b}, b) p +@[simps!] +def ofCompletePred (p : α → Prop) (hsinf : ∀ s, (∀ a ∈ s, p a) → p (sInf s)) : ClosureOperator α := + ofPred (fun a ↦ ⨅ b : {b // a ≤ b ∧ p b}, b) p (fun a ↦ by simp [forall_swap]) - (fun a ↦ hsinf _ $ forall_range_iff.2 $ fun b ↦ b.2.1) - (fun a b hab hb ↦ iInf_le_of_le ⟨b, hb, hab⟩ le_rfl) - -/-- This lemma shows that the image of `x` of a closure operator built from the `ofPred` constructor -respects `p`, the property that was fed into it. -/ -lemma ofPred_spec {sinf} (x : α) : p (ofPred p sinf x) := closure_mem_mk₃ _ - -/-- The property `p` fed into the `ofPred` constructor exactly corresponds to being closed. -/ -@[simp] lemma closed_ofPred {hsinf} : (ofPred p hsinf).closed = {x | p x} := by - ext; exact mem_mk₃_closed - -/-- The property `p` fed into the `ofPred` constructor exactly corresponds to being closed. -/ -lemma mem_closed_ofPred {hsinf} {x : α} : x ∈ (ofPred p hsinf).closed ↔ p x := mem_mk₃_closed + (fun a ↦ hsinf _ $ forall_range_iff.2 $ fun b ↦ b.2.2) + (fun a b hab hb ↦ iInf_le_of_le ⟨b, hab, hb⟩ le_rfl) @[simp] theorem closure_iSup_closure (f : ι → α) : c (⨆ i, c (f i)) = c (⨆ i, f i) := - le_antisymm ((c.le_closure_iff _ _).1 <| iSup_le fun i => c.monotone <| le_iSup f i) <| + le_antisymm (le_closure_iff.1 <| iSup_le fun i => c.monotone <| le_iSup f i) <| c.monotone <| iSup_mono fun _ => c.le_closure _ #align closure_operator.closure_supr_closure ClosureOperator.closure_iSup_closure @[simp] theorem closure_iSup₂_closure (f : ∀ i, κ i → α) : c (⨆ (i) (j), c (f i j)) = c (⨆ (i) (j), f i j) := - le_antisymm ((c.le_closure_iff _ _).1 <| iSup₂_le fun i j => c.monotone <| le_iSup₂ i j) <| + le_antisymm (le_closure_iff.1 <| iSup₂_le fun i j => c.monotone <| le_iSup₂ i j) <| c.monotone <| iSup₂_mono fun _ _ => c.le_closure _ #align closure_operator.closure_supr₂_closure ClosureOperator.closure_iSup₂_closure @@ -389,7 +358,7 @@ theorem idempotent (x : α) : u (l (u (l x))) = u (l x) := #align lower_adjoint.idempotent LowerAdjoint.idempotent theorem le_closure_iff (x y : α) : x ≤ u (l y) ↔ u (l x) ≤ u (l y) := - l.closureOperator.le_closure_iff _ _ + l.closureOperator.le_closure_iff #align lower_adjoint.le_closure_iff LowerAdjoint.le_closure_iff end PartialOrder @@ -417,7 +386,7 @@ section PartialOrder variable [PartialOrder α] [PartialOrder β] {u : β → α} (l : LowerAdjoint u) theorem mem_closed_iff_closure_le (x : α) : x ∈ l.closed ↔ u (l x) ≤ x := - l.closureOperator.mem_closed_iff_closure_le _ + l.closureOperator.isClosed_iff_closure_le #align lower_adjoint.mem_closed_iff_closure_le LowerAdjoint.mem_closed_iff_closure_le @[simp, nolint simpNF] -- Porting note: lemma does prove itself, seems to be a linter error @@ -427,7 +396,7 @@ theorem closure_is_closed (x : α) : u (l x) ∈ l.closed := /-- The set of closed elements for `l` is the range of `u ∘ l`. -/ theorem closed_eq_range_close : l.closed = Set.range (u ∘ l) := - l.closureOperator.closed_eq_range_close + l.closureOperator.setOf_isClosed_eq_range_closure #align lower_adjoint.closed_eq_range_close LowerAdjoint.closed_eq_range_close /-- Send an `x` to an element of the set of closed elements (by taking the closure). -/ @@ -437,7 +406,7 @@ def toClosed (x : α) : l.closed := @[simp] theorem closure_le_closed_iff_le (x : α) {y : α} (hy : l.closed y) : u (l x) ≤ y ↔ x ≤ y := - l.closureOperator.closure_le_closed_iff_le x hy + (show l.closureOperator.IsClosed y from hy).closure_le_iff #align lower_adjoint.closure_le_closed_iff_le LowerAdjoint.closure_le_closed_iff_le end PartialOrder @@ -572,10 +541,10 @@ def GaloisConnection.closureOperator [PartialOrder α] [Preorder β] {l : α → #align galois_connection.closure_operator_apply GaloisConnection.closureOperator_apply /-- The set of closed elements has a Galois insertion to the underlying type. -/ -def _root_.ClosureOperator.gi [PartialOrder α] (c : ClosureOperator α) : - GaloisInsertion c.toClosed (↑) where - choice x hx := ⟨x, hx.antisymm (c.le_closure x)⟩ - gc _ y := c.closure_le_closed_iff_le _ y.2 +def ClosureOperator.gi [PartialOrder α] (c : ClosureOperator α) : + GaloisInsertion c.toCloseds (↑) where + choice x hx := ⟨x, isClosed_iff_closure_le.2 hx⟩ + gc _ y := y.2.closure_le_iff le_l_u _ := c.le_closure _ choice_eq x hx := le_antisymm (c.le_closure x) hx #align closure_operator.gi ClosureOperator.gi diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index c96460ab769d1..3b635d1d87aa4 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -14,13 +14,14 @@ import Mathlib.Data.Set.Finite ## Main definitions * `Filter` : filters on a set; -* `Filter.Principal` : specific filters; -* `map`, `comap` : operations on filters; +* `Filter.principal` : filter of all sets containing a given set; +* `Filter.map`, `Filter.comap` : operations on filters; * `Filter.Tendsto` : limit with respect to filters; * `Filter.Eventually` : `f.eventually p` means `{x | p x} ∈ f`; * `Filter.Frequently` : `f.frequently p` means `{x | ¬p x} ∉ f`; -* `filter_upwards [h₁, ..., hₙ]` : takes a list of proofs `hᵢ : sᵢ ∈ f`, and replaces a goal `s ∈ f` - with `∀ x, x ∈ s₁ → ... → x ∈ sₙ → x ∈ s`; +* `filter_upwards [h₁, ..., hₙ]` : + a tactic that takes a list of proofs `hᵢ : sᵢ ∈ f`, + and replaces a goal `s ∈ f` with `∀ x, x ∈ s₁ → ... → x ∈ sₙ → x ∈ s`; * `Filter.NeBot f` : a utility class stating that `f` is a non-trivial filter. Filters on a type `X` are sets of sets of `X` satisfying three conditions. They are mostly used to @@ -1760,6 +1761,69 @@ theorem EventuallyLE.union {s t s' t' : Set α} {l : Filter α} (h : s ≤ᶠ[l] h'.mp <| h.mono fun _ => Or.imp #align filter.eventually_le.union Filter.EventuallyLE.union +protected lemma EventuallyLE.iUnion [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i ≤ᶠ[l] t i) : (⋃ i, s i) ≤ᶠ[l] ⋃ i, t i := + (eventually_all.2 h).mono fun _x hx hx' ↦ + let ⟨i, hi⟩ := mem_iUnion.1 hx'; mem_iUnion.2 ⟨i, hx i hi⟩ + +protected lemma EventuallyEq.iUnion [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i =ᶠ[l] t i) : (⋃ i, s i) =ᶠ[l] ⋃ i, t i := + (EventuallyLE.iUnion fun i ↦ (h i).le).antisymm <| .iUnion fun i ↦ (h i).symm.le + +protected lemma EventuallyLE.iInter [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i ≤ᶠ[l] t i) : (⋂ i, s i) ≤ᶠ[l] ⋂ i, t i := + (eventually_all.2 h).mono fun _x hx hx' ↦ mem_iInter.2 fun i ↦ hx i (mem_iInter.1 hx' i) + +protected lemma EventuallyEq.iInter [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i =ᶠ[l] t i) : (⋂ i, s i) =ᶠ[l] ⋂ i, t i := + (EventuallyLE.iInter fun i ↦ (h i).le).antisymm <| .iInter fun i ↦ (h i).symm.le + +lemma _root_.Set.Finite.eventuallyLE_iUnion {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋃ i ∈ s, f i) ≤ᶠ[l] (⋃ i ∈ s, g i) := by + have := hs.to_subtype + rw [biUnion_eq_iUnion, biUnion_eq_iUnion] + exact .iUnion fun i ↦ hle i.1 i.2 + +alias EventuallyLE.biUnion := Set.Finite.eventuallyLE_iUnion + +lemma _root_.Set.Finite.eventuallyEq_iUnion {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋃ i ∈ s, f i) =ᶠ[l] (⋃ i ∈ s, g i) := + (EventuallyLE.biUnion hs fun i hi ↦ (heq i hi).le).antisymm <| + .biUnion hs fun i hi ↦ (heq i hi).symm.le + +alias EventuallyEq.biUnion := Set.Finite.eventuallyEq_iUnion + +lemma _root_.Set.Finite.eventuallyLE_iInter {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋂ i ∈ s, f i) ≤ᶠ[l] (⋂ i ∈ s, g i) := by + have := hs.to_subtype + rw [biInter_eq_iInter, biInter_eq_iInter] + exact .iInter fun i ↦ hle i.1 i.2 + +alias EventuallyLE.biInter := Set.Finite.eventuallyLE_iInter + +lemma _root_.Set.Finite.eventuallyEq_iInter {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋂ i ∈ s, f i) =ᶠ[l] (⋂ i ∈ s, g i) := + (EventuallyLE.biInter hs fun i hi ↦ (heq i hi).le).antisymm <| + .biInter hs fun i hi ↦ (heq i hi).symm.le + +alias EventuallyEq.biInter := Set.Finite.eventuallyEq_iInter + +lemma _root_.Finset.eventuallyLE_iUnion {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋃ i ∈ s, f i) ≤ᶠ[l] (⋃ i ∈ s, g i) := + .biUnion s.finite_toSet hle + +lemma _root_.Finset.eventuallyEq_iUnion {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋃ i ∈ s, f i) =ᶠ[l] (⋃ i ∈ s, g i) := + .biUnion s.finite_toSet heq + +lemma _root_.Finset.eventuallyLE_iInter {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋂ i ∈ s, f i) ≤ᶠ[l] (⋂ i ∈ s, g i) := + .biInter s.finite_toSet hle + +lemma _root_.Finset.eventuallyEq_iInter {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋂ i ∈ s, f i) =ᶠ[l] (⋂ i ∈ s, g i) := + .biInter s.finite_toSet heq + @[mono] theorem EventuallyLE.compl {s t : Set α} {l : Filter α} (h : s ≤ᶠ[l] t) : (tᶜ : Set α) ≤ᶠ[l] (sᶜ : Set α) := diff --git a/Mathlib/Order/Filter/Cofinite.lean b/Mathlib/Order/Filter/Cofinite.lean index 5654ad5b45923..2275cbf66e058 100644 --- a/Mathlib/Order/Filter/Cofinite.lean +++ b/Mathlib/Order/Filter/Cofinite.lean @@ -65,10 +65,20 @@ theorem cofinite_eq_bot [Finite α] : @cofinite α = ⊥ := cofinite_eq_bot_iff. theorem frequently_cofinite_iff_infinite {p : α → Prop} : (∃ᶠ x in cofinite, p x) ↔ Set.Infinite { x | p x } := by - simp only [Filter.Frequently, Filter.Eventually, mem_cofinite, compl_setOf, not_not, - Set.Infinite] + simp only [Filter.Frequently, eventually_cofinite, not_not, Set.Infinite] #align filter.frequently_cofinite_iff_infinite Filter.frequently_cofinite_iff_infinite +lemma frequently_cofinite_mem_iff_infinite {s : Set α} : (∃ᶠ x in cofinite, x ∈ s) ↔ s.Infinite := + frequently_cofinite_iff_infinite + +alias ⟨_, _root_.Set.Infinite.frequently_cofinite⟩ := frequently_cofinite_mem_iff_infinite + +@[simp] +lemma cofinite_inf_principal_neBot_iff {s : Set α} : (cofinite ⊓ 𝓟 s).NeBot ↔ s.Infinite := + frequently_mem_iff_neBot.symm.trans frequently_cofinite_mem_iff_infinite + +alias ⟨_, _root_.Set.Infinite.cofinite_inf_principal_neBot⟩ := cofinite_inf_principal_neBot_iff + theorem _root_.Set.Finite.compl_mem_cofinite {s : Set α} (hs : s.Finite) : sᶜ ∈ @cofinite α := mem_cofinite.2 <| (compl_compl s).symm ▸ hs #align set.finite.compl_mem_cofinite Set.Finite.compl_mem_cofinite @@ -124,14 +134,10 @@ theorem coprodᵢ_cofinite {α : ι → Type*} [Finite ι] : set_option linter.uppercaseLean3 false in #align filter.Coprod_cofinite Filter.coprodᵢ_cofinite -@[simp] theorem disjoint_cofinite_left : Disjoint cofinite l ↔ ∃ s ∈ l, Set.Finite s := by - simp only [hasBasis_cofinite.disjoint_iff l.basis_sets, id, disjoint_compl_left_iff_subset] - exact ⟨fun ⟨s, hs, t, ht, hts⟩ => ⟨t, ht, hs.subset hts⟩, - fun ⟨s, hs, hsf⟩ => ⟨s, hsf, s, hs, Subset.rfl⟩⟩ + simp [l.basis_sets.disjoint_iff_right] #align filter.disjoint_cofinite_left Filter.disjoint_cofinite_left -@[simp] theorem disjoint_cofinite_right : Disjoint l cofinite ↔ ∃ s ∈ l, Set.Finite s := disjoint_comm.trans disjoint_cofinite_left #align filter.disjoint_cofinite_right Filter.disjoint_cofinite_right @@ -152,6 +158,14 @@ end Filter open Filter +lemma Set.Finite.cofinite_inf_principal_compl {s : Set α} (hs : s.Finite) : + cofinite ⊓ 𝓟 sᶜ = cofinite := by + simpa using hs.compl_mem_cofinite + +lemma Set.Finite.cofinite_inf_principal_diff {s t : Set α} (ht : t.Finite) : + cofinite ⊓ 𝓟 (s \ t) = cofinite ⊓ 𝓟 s := by + rw [diff_eq, ← inf_principal, ← inf_assoc, inf_right_comm, ht.cofinite_inf_principal_compl] + /-- For natural numbers the filters `Filter.cofinite` and `Filter.atTop` coincide. -/ theorem Nat.cofinite_eq_atTop : @cofinite ℕ = atTop := by refine' le_antisymm _ atTop_le_cofinite diff --git a/Mathlib/Order/Filter/Interval.lean b/Mathlib/Order/Filter/Interval.lean index 23f60e8d167af..a5b3e3166251b 100644 --- a/Mathlib/Order/Filter/Interval.lean +++ b/Mathlib/Order/Filter/Interval.lean @@ -12,10 +12,41 @@ import Mathlib.Order.Filter.AtTopBot /-! # Convergence of intervals +## Motivation + +If a function tends to infinity somewhere, then its derivative is not integrable around this place. +One should be careful about this statement: "somewhere" could mean a point, but also convergence +from the left or from the right, or it could also be infinity, and "around this place" will refer +to these directed neighborhoods. Therefore, the above theorem has many variants. Instead of stating +all these variants, one can look for the common abstraction and have a single version. One has to +be careful: if one considers convergence along a sequence, then the function may tend to infinity +but have a derivative which is small along the sequence (with big jumps inbetween), so in the end +the derivative may be integrable on a neighborhood of the sequence. What really matters for such +calculus issues in terms of derivatives is that whole intervals are included in the sets we +consider. + +The right common abstraction is provided in this file, as the `TendstoIxxClass` typeclass. +It takes as parameters a class of bounded intervals and two real filters `l₁` and `l₂`. +An instance `TendstoIxxClass Icc l₁ l₂` registers that, if `aₙ` and `bₙ` are converging towards +the filter `l₁`, then the intervals `Icc aₙ bₙ` are eventually contained in any given set +belonging to `l₂`. For instance, for `l₁ = 𝓝[>] x` and `l₂ = 𝓝[≥] x`, the strict and large right +neighborhoods of `x` respectively, then given any large right neighborhood `s ∈ 𝓝[≥] x` and any two +sequences `xₙ` and `yₙ` converging strictly to the right of `x`, +then the interval `[xₙ, yₙ]` is eventually contained in `s`. Therefore, the instance +`TendstoIxxClass Icc (𝓝[>] x) (𝓝[≥] x)` holds. Note that one could have taken as +well `l₂ = 𝓝[>] x`, but that `l₁ = 𝓝[≥] x` and `l₂ = 𝓝[>] x` wouldn't work. + +With this formalism, the above theorem would read: if `TendstoIxxClass Icc l l` and `f` tends +to infinity along `l`, then its derivative is not integrable on any element of `l`. +Beyond this simple example, this typeclass plays a prominent role in generic formulations of +the fundamental theorem of calculus. + +## Main definition + If both `a` and `b` tend to some filter `l₁`, sometimes this implies that `Ixx a b` tends to `l₂.smallSets`, i.e., for any `s ∈ l₂` eventually `Ixx a b` becomes a subset of `s`. Here and below -`Ixx` is one of `Set.Icc`, `Set.Ico`, `Set.Ioc`, and `Set.Ioo`. We define `Filter.TendstoIxxClass -Ixx l₁ l₂` to be a typeclass representing this property. +`Ixx` is one of `Set.Icc`, `Set.Ico`, `Set.Ioc`, and `Set.Ioo`. +We define `Filter.TendstoIxxClass Ixx l₁ l₂` to be a typeclass representing this property. The instances provide the best `l₂` for a given `l₁`. In many cases `l₁ = l₂` but sometimes we can drop an endpoint from an interval: e.g., we prove diff --git a/Mathlib/Order/Filter/NAry.lean b/Mathlib/Order/Filter/NAry.lean index 39b454f9a2079..a8916ef2ecd52 100644 --- a/Mathlib/Order/Filter/NAry.lean +++ b/Mathlib/Order/Filter/NAry.lean @@ -100,94 +100,59 @@ theorem le_map₂_iff {h : Filter γ} : #align filter.le_map₂_iff Filter.le_map₂_iff @[simp] -theorem map₂_bot_left : map₂ m ⊥ g = ⊥ := - empty_mem_iff_bot.1 ⟨∅, univ, trivial, univ_mem, image2_empty_left.subset⟩ -#align filter.map₂_bot_left Filter.map₂_bot_left +theorem map₂_eq_bot_iff : map₂ m f g = ⊥ ↔ f = ⊥ ∨ g = ⊥ := by simp [← map_prod_eq_map₂] +#align filter.map₂_eq_bot_iff Filter.map₂_eq_bot_iff @[simp] -theorem map₂_bot_right : map₂ m f ⊥ = ⊥ := - empty_mem_iff_bot.1 ⟨univ, ∅, univ_mem, trivial, image2_empty_right.subset⟩ -#align filter.map₂_bot_right Filter.map₂_bot_right +theorem map₂_bot_left : map₂ m ⊥ g = ⊥ := map₂_eq_bot_iff.2 <| .inl rfl +#align filter.map₂_bot_left Filter.map₂_bot_left @[simp] -theorem map₂_eq_bot_iff : map₂ m f g = ⊥ ↔ f = ⊥ ∨ g = ⊥ := by - simp only [← empty_mem_iff_bot, mem_map₂_iff, subset_empty_iff, image2_eq_empty_iff] - constructor - · rintro ⟨s, t, hs, ht, rfl | rfl⟩ - · exact Or.inl hs - · exact Or.inr ht - · rintro (h | h) - · exact ⟨_, _, h, univ_mem, Or.inl rfl⟩ - · exact ⟨_, _, univ_mem, h, Or.inr rfl⟩ -#align filter.map₂_eq_bot_iff Filter.map₂_eq_bot_iff +theorem map₂_bot_right : map₂ m f ⊥ = ⊥ := map₂_eq_bot_iff.2 <| .inr rfl +#align filter.map₂_bot_right Filter.map₂_bot_right @[simp] -theorem map₂_neBot_iff : (map₂ m f g).NeBot ↔ f.NeBot ∧ g.NeBot := by - simp_rw [neBot_iff] - exact map₂_eq_bot_iff.not.trans not_or +theorem map₂_neBot_iff : (map₂ m f g).NeBot ↔ f.NeBot ∧ g.NeBot := by simp [neBot_iff, not_or] #align filter.map₂_ne_bot_iff Filter.map₂_neBot_iff -theorem NeBot.map₂ (hf : f.NeBot) (hg : g.NeBot) : (map₂ m f g).NeBot := +protected theorem NeBot.map₂ (hf : f.NeBot) (hg : g.NeBot) : (map₂ m f g).NeBot := map₂_neBot_iff.2 ⟨hf, hg⟩ #align filter.ne_bot.map₂ Filter.NeBot.map₂ --- Porting note: Why do I have to specify the `Filter` namespace for `map₂` here? -theorem NeBot.of_map₂_left (h : (Filter.map₂ m f g).NeBot) : f.NeBot := +instance map₂.neBot [NeBot f] [NeBot g] : NeBot (map₂ m f g) := .map₂ ‹_› ‹_› + +theorem NeBot.of_map₂_left (h : (map₂ m f g).NeBot) : f.NeBot := (map₂_neBot_iff.1 h).1 #align filter.ne_bot.of_map₂_left Filter.NeBot.of_map₂_left -theorem NeBot.of_map₂_right (h : (Filter.map₂ m f g).NeBot) : g.NeBot := +theorem NeBot.of_map₂_right (h : (map₂ m f g).NeBot) : g.NeBot := (map₂_neBot_iff.1 h).2 #align filter.ne_bot.of_map₂_right Filter.NeBot.of_map₂_right theorem map₂_sup_left : map₂ m (f₁ ⊔ f₂) g = map₂ m f₁ g ⊔ map₂ m f₂ g := by - ext u - constructor - · rintro ⟨s, t, ⟨h₁, h₂⟩, ht, hu⟩ - exact ⟨mem_of_superset (image2_mem_map₂ h₁ ht) hu, mem_of_superset (image2_mem_map₂ h₂ ht) hu⟩ - · rintro ⟨⟨s₁, t₁, hs₁, ht₁, hu₁⟩, s₂, t₂, hs₂, ht₂, hu₂⟩ - refine' ⟨s₁ ∪ s₂, t₁ ∩ t₂, union_mem_sup hs₁ hs₂, inter_mem ht₁ ht₂, _⟩ - rw [image2_union_left] - exact - union_subset ((image2_subset_left <| inter_subset_left _ _).trans hu₁) - ((image2_subset_left <| inter_subset_right _ _).trans hu₂) + simp_rw [← map_prod_eq_map₂, sup_prod, map_sup] #align filter.map₂_sup_left Filter.map₂_sup_left theorem map₂_sup_right : map₂ m f (g₁ ⊔ g₂) = map₂ m f g₁ ⊔ map₂ m f g₂ := by - ext u - constructor - · rintro ⟨s, t, hs, ⟨h₁, h₂⟩, hu⟩ - exact ⟨mem_of_superset (image2_mem_map₂ hs h₁) hu, mem_of_superset (image2_mem_map₂ hs h₂) hu⟩ - · rintro ⟨⟨s₁, t₁, hs₁, ht₁, hu₁⟩, s₂, t₂, hs₂, ht₂, hu₂⟩ - refine' ⟨s₁ ∩ s₂, t₁ ∪ t₂, inter_mem hs₁ hs₂, union_mem_sup ht₁ ht₂, _⟩ - rw [image2_union_right] - exact - union_subset ((image2_subset_right <| inter_subset_left _ _).trans hu₁) - ((image2_subset_right <| inter_subset_right _ _).trans hu₂) + simp_rw [← map_prod_eq_map₂, prod_sup, map_sup] #align filter.map₂_sup_right Filter.map₂_sup_right theorem map₂_inf_subset_left : map₂ m (f₁ ⊓ f₂) g ≤ map₂ m f₁ g ⊓ map₂ m f₂ g := - le_inf (map₂_mono_right inf_le_left) (map₂_mono_right inf_le_right) + Monotone.map_inf_le (fun _ _ ↦ map₂_mono_right) f₁ f₂ #align filter.map₂_inf_subset_left Filter.map₂_inf_subset_left theorem map₂_inf_subset_right : map₂ m f (g₁ ⊓ g₂) ≤ map₂ m f g₁ ⊓ map₂ m f g₂ := - le_inf (map₂_mono_left inf_le_left) (map₂_mono_left inf_le_right) + Monotone.map_inf_le (fun _ _ ↦ map₂_mono_left) g₁ g₂ #align filter.map₂_inf_subset_right Filter.map₂_inf_subset_right @[simp] -theorem map₂_pure_left : map₂ m (pure a) g = g.map fun b => m a b := - Filter.ext fun u => - ⟨fun ⟨s, t, hs, ht, hu⟩ => - mem_of_superset (image_mem_map ht) ((image_subset_image2_right <| mem_pure.1 hs).trans hu), - fun h => ⟨{a}, _, singleton_mem_pure, h, by rw [image2_singleton_left, image_subset_iff]⟩⟩ +theorem map₂_pure_left : map₂ m (pure a) g = g.map (m a) := by + rw [← map_prod_eq_map₂, pure_prod, map_map]; rfl #align filter.map₂_pure_left Filter.map₂_pure_left @[simp] -theorem map₂_pure_right : map₂ m f (pure b) = f.map fun a => m a b := - Filter.ext fun u => - ⟨fun ⟨s, t, hs, ht, hu⟩ => - mem_of_superset (image_mem_map hs) ((image_subset_image2_left <| mem_pure.1 ht).trans hu), - fun h => ⟨_, {b}, h, singleton_mem_pure, by rw [image2_singleton_right, image_subset_iff]⟩⟩ +theorem map₂_pure_right : map₂ m f (pure b) = f.map (m · b) := by + rw [← map_prod_eq_map₂, prod_pure, map_map]; rfl #align filter.map₂_pure_right Filter.map₂_pure_right theorem map₂_pure : map₂ m (pure a) (pure b) = pure (m a b) := by rw [map₂_pure_right, map_pure] @@ -200,23 +165,18 @@ theorem map₂_swap (m : α → β → γ) (f : Filter α) (g : Filter β) : #align filter.map₂_swap Filter.map₂_swap @[simp] -theorem map₂_left (h : g.NeBot) : map₂ (fun x _ => x) f g = f := by - ext u - refine' ⟨_, fun hu => ⟨_, _, hu, univ_mem, (image2_left <| h.nonempty_of_mem univ_mem).subset⟩⟩ - rintro ⟨s, t, hs, ht, hu⟩ - rw [image2_left (h.nonempty_of_mem ht)] at hu - exact mem_of_superset hs hu +theorem map₂_left [NeBot g] : map₂ (fun x _ => x) f g = f := by + rw [← map_prod_eq_map₂, map_fst_prod] #align filter.map₂_left Filter.map₂_left @[simp] -theorem map₂_right (h : f.NeBot) : map₂ (fun _ y => y) f g = g := by rw [map₂_swap, map₂_left h] +theorem map₂_right [NeBot f] : map₂ (fun _ y => y) f g = g := by rw [map₂_swap, map₂_left] #align filter.map₂_right Filter.map₂_right /-- The image of a ternary function `m : α → β → γ → δ` as a function `Filter α → Filter β → Filter γ → Filter δ`. Mathematically this should be thought of as the image of the corresponding function `α × β × γ → δ`. -/ -def map₃ (m : α → β → γ → δ) (f : Filter α) (g : Filter β) (h : Filter γ) : Filter δ - where +def map₃ (m : α → β → γ → δ) (f : Filter α) (g : Filter β) (h : Filter γ) : Filter δ where sets := { s | ∃ u v w, u ∈ f ∧ v ∈ g ∧ w ∈ h ∧ image3 m u v w ⊆ s } univ_sets := ⟨univ, univ, univ, univ_sets _, univ_sets _, univ_sets _, subset_univ _⟩ sets_of_superset hs hst := diff --git a/Mathlib/Order/Filter/Pointwise.lean b/Mathlib/Order/Filter/Pointwise.lean index 4ff5c60ceda8a..69a10231cf860 100644 --- a/Mathlib/Order/Filter/Pointwise.lean +++ b/Mathlib/Order/Filter/Pointwise.lean @@ -102,13 +102,13 @@ theorem principal_one : 𝓟 1 = (1 : Filter α) := #align filter.principal_one Filter.principal_one #align filter.principal_zero Filter.principal_zero -@[to_additive] -- TODO: make this a scoped instance in the `Pointwise` namespace +@[to_additive] theorem one_neBot : (1 : Filter α).NeBot := Filter.pure_neBot #align filter.one_ne_bot Filter.one_neBot #align filter.zero_ne_bot Filter.zero_neBot -scoped[Pointwise] attribute [instance] Filter.one_neBot +scoped[Pointwise] attribute [instance] one_neBot zero_neBot @[to_additive (attr := simp)] protected theorem map_one' (f : α → β) : (1 : Filter α).map f = pure (f 1) := @@ -224,10 +224,15 @@ theorem neBot_inv_iff : f⁻¹.NeBot ↔ NeBot f := #align filter.ne_bot_neg_iff Filter.neBot_neg_iff @[to_additive] -theorem NeBot.inv : f.NeBot → f⁻¹.NeBot := fun h => h.map _ +protected theorem NeBot.inv : f.NeBot → f⁻¹.NeBot := fun h => h.map _ #align filter.ne_bot.inv Filter.NeBot.inv #align filter.ne_bot.neg Filter.NeBot.neg +@[to_additive neg.instNeBot] +lemma inv.instNeBot [NeBot f] : NeBot f⁻¹ := .inv ‹_› + +scoped[Pointwise] attribute [instance] inv.instNeBot neg.instNeBot + end Inv section InvolutiveInv @@ -336,7 +341,7 @@ lemma mul_neBot_iff : (f * g).NeBot ↔ f.NeBot ∧ g.NeBot := #align filter.add_ne_bot_iff Filter.add_neBot_iff @[to_additive] -theorem NeBot.mul : NeBot f → NeBot g → NeBot (f * g) := +protected theorem NeBot.mul : NeBot f → NeBot g → NeBot (f * g) := NeBot.map₂ #align filter.ne_bot.mul Filter.NeBot.mul #align filter.ne_bot.add Filter.NeBot.add @@ -353,6 +358,11 @@ theorem NeBot.of_mul_right : (f * g).NeBot → g.NeBot := #align filter.ne_bot.of_mul_right Filter.NeBot.of_mul_right #align filter.ne_bot.of_add_right Filter.NeBot.of_add_right +@[to_additive add.instNeBot] +protected lemma mul.instNeBot [NeBot f] [NeBot g] : NeBot (f * g) := .mul ‹_› ‹_› + +scoped[Pointwise] attribute [instance] mul.instNeBot add.instNeBot + @[to_additive (attr := simp)] theorem pure_mul : pure a * g = g.map (a * ·) := map₂_pure_left @@ -418,7 +428,6 @@ end Mul /-! ### Filter subtraction/division -/ - section Div variable [Div α] {f f₁ f₂ g g₁ g₂ h : Filter α} {s t : Set α} {a b : α} @@ -470,14 +479,14 @@ theorem div_eq_bot_iff : f / g = ⊥ ↔ f = ⊥ ∨ g = ⊥ := #align filter.div_eq_bot_iff Filter.div_eq_bot_iff #align filter.sub_eq_bot_iff Filter.sub_eq_bot_iff -@[to_additive (attr := simp)] -- TODO: make this a scoped instance in the `Pointwise` namespace +@[to_additive (attr := simp)] theorem div_neBot_iff : (f / g).NeBot ↔ f.NeBot ∧ g.NeBot := map₂_neBot_iff #align filter.div_ne_bot_iff Filter.div_neBot_iff #align filter.sub_ne_bot_iff Filter.sub_neBot_iff @[to_additive] -theorem NeBot.div : NeBot f → NeBot g → NeBot (f / g) := +protected theorem NeBot.div : NeBot f → NeBot g → NeBot (f / g) := NeBot.map₂ #align filter.ne_bot.div Filter.NeBot.div #align filter.ne_bot.sub Filter.NeBot.sub @@ -494,6 +503,11 @@ theorem NeBot.of_div_right : (f / g).NeBot → g.NeBot := #align filter.ne_bot.of_div_right Filter.NeBot.of_div_right #align filter.ne_bot.of_sub_right Filter.NeBot.of_sub_right +@[to_additive sub.instNeBot] +lemma div.instNeBot [NeBot f] [NeBot g] : NeBot (f / g) := .div ‹_› ‹_› + +scoped[Pointwise] attribute [instance] div.instNeBot sub.instNeBot + @[to_additive (attr := simp)] theorem pure_div : pure a / g = g.map (a / ·) := map₂_pure_left @@ -832,7 +846,6 @@ variable [MulZeroClass α] {f g : Filter α} /-! Note that `Filter` is not a `MulZeroClass` because `0 * ⊥ ≠ 0`. -/ - theorem NeBot.mul_zero_nonneg (hf : f.NeBot) : 0 ≤ f * 0 := le_mul_iff.2 fun _ h₁ _ h₂ => let ⟨_, ha⟩ := hf.nonempty_of_mem h₁ @@ -1001,7 +1014,7 @@ theorem smul_neBot_iff : (f • g).NeBot ↔ f.NeBot ∧ g.NeBot := #align filter.vadd_ne_bot_iff Filter.vadd_neBot_iff @[to_additive] -theorem NeBot.smul : NeBot f → NeBot g → NeBot (f • g) := +protected theorem NeBot.smul : NeBot f → NeBot g → NeBot (f • g) := NeBot.map₂ #align filter.ne_bot.smul Filter.NeBot.smul #align filter.ne_bot.vadd Filter.NeBot.vadd @@ -1018,6 +1031,11 @@ theorem NeBot.of_smul_right : (f • g).NeBot → g.NeBot := #align filter.ne_bot.of_smul_right Filter.NeBot.of_smul_right #align filter.ne_bot.of_vadd_right Filter.NeBot.of_vadd_right +@[to_additive vadd.instNeBot] +lemma smul.instNeBot [NeBot f] [NeBot g] : NeBot (f • g) := .smul ‹_› ‹_› + +scoped[Pointwise] attribute [instance] smul.instNeBot vadd.instNeBot + @[to_additive (attr := simp)] theorem pure_smul : (pure a : Filter α) • g = g.map (a • ·) := map₂_pure_left @@ -1117,7 +1135,7 @@ theorem vsub_neBot_iff : (f -ᵥ g : Filter α).NeBot ↔ f.NeBot ∧ g.NeBot := map₂_neBot_iff #align filter.vsub_ne_bot_iff Filter.vsub_neBot_iff -theorem NeBot.vsub : NeBot f → NeBot g → NeBot (f -ᵥ g) := +protected theorem NeBot.vsub : NeBot f → NeBot g → NeBot (f -ᵥ g) := NeBot.map₂ #align filter.ne_bot.vsub Filter.NeBot.vsub @@ -1129,6 +1147,10 @@ theorem NeBot.of_vsub_right : (f -ᵥ g : Filter α).NeBot → g.NeBot := NeBot.of_map₂_right #align filter.ne_bot.of_vsub_right Filter.NeBot.of_vsub_right +lemma vsub.instNeBot [NeBot f] [NeBot g] : NeBot (f -ᵥ g) := .vsub ‹_› ‹_› + +scoped[Pointwise] attribute [instance] vsub.instNeBot + @[simp] theorem pure_vsub : (pure a : Filter β) -ᵥ g = g.map (a -ᵥ ·) := map₂_pure_left @@ -1225,6 +1247,11 @@ theorem NeBot.of_smul_filter : (a • f).NeBot → f.NeBot := #align filter.ne_bot.of_smul_filter Filter.NeBot.of_smul_filter #align filter.ne_bot.of_vadd_filter Filter.NeBot.of_vadd_filter +@[to_additive vadd_filter.instNeBot] +lemma smul_filter.instNeBot [NeBot f] : NeBot (a • f) := .smul_filter ‹_› + +scoped[Pointwise] attribute [instance] smul_filter.instNeBot vadd_filter.instNeBot + @[to_additive] theorem smul_filter_le_smul_filter (hf : f₁ ≤ f₂) : a • f₁ ≤ a • f₂ := map_mono hf @@ -1350,7 +1377,6 @@ Note that we have neither `SMulWithZero α (Filter β)` nor `SMulWithZero (Filte because `0 * ⊥ ≠ 0`. -/ - theorem NeBot.smul_zero_nonneg (hf : f.NeBot) : 0 ≤ f • (0 : Filter β) := le_smul_iff.2 fun _ h₁ _ h₂ => let ⟨_, ha⟩ := hf.nonempty_of_mem h₁ diff --git a/Mathlib/Order/Filter/Prod.lean b/Mathlib/Order/Filter/Prod.lean index 56db461ed0e58..e7deb41fa8b80 100644 --- a/Mathlib/Order/Filter/Prod.lean +++ b/Mathlib/Order/Filter/Prod.lean @@ -413,6 +413,7 @@ theorem prod_pure_pure {a : α} {b : β} : (pure a : Filter α) ×ˢ (pure b : Filter β) = pure (a, b) := by simp #align filter.prod_pure_pure Filter.prod_pure_pure +@[simp] theorem prod_eq_bot : f ×ˢ g = ⊥ ↔ f = ⊥ ∨ g = ⊥ := by simp_rw [← empty_mem_iff_bot, mem_prod_iff, subset_empty_iff, prod_eq_empty_iff, ← exists_prop, Subtype.exists', exists_or, exists_const, Subtype.exists, exists_prop, exists_eq_right] @@ -428,11 +429,11 @@ theorem prod_neBot : NeBot (f ×ˢ g) ↔ NeBot f ∧ NeBot g := by simp only [neBot_iff, Ne, prod_eq_bot, not_or] #align filter.prod_ne_bot Filter.prod_neBot -theorem NeBot.prod (hf : NeBot f) (hg : NeBot g) : NeBot (f ×ˢ g) := prod_neBot.2 ⟨hf, hg⟩ +protected theorem NeBot.prod (hf : NeBot f) (hg : NeBot g) : NeBot (f ×ˢ g) := prod_neBot.2 ⟨hf, hg⟩ #align filter.ne_bot.prod Filter.NeBot.prod -instance prod_neBot' [hf : NeBot f] [hg : NeBot g] : NeBot (f ×ˢ g) := hf.prod hg -#align filter.prod_ne_bot' Filter.prod_neBot' +instance prod.instNeBot [hf : NeBot f] [hg : NeBot g] : NeBot (f ×ˢ g) := hf.prod hg +#align filter.prod_ne_bot' Filter.prod.instNeBot @[simp] lemma disjoint_prod {f' : Filter α} {g' : Filter β} : diff --git a/Mathlib/Order/SupClosed.lean b/Mathlib/Order/SupClosed.lean index c485fbbdd6b49..18169ee234f31 100644 --- a/Mathlib/Order/SupClosed.lean +++ b/Mathlib/Order/SupClosed.lean @@ -252,7 +252,8 @@ section SemilatticeSup variable [SemilatticeSup α] [SemilatticeSup β] {s t : Set α} {a b : α} /-- Every set in a join-semilattice generates a set closed under join. -/ -def supClosure : ClosureOperator (Set α) := ClosureOperator.mk₃ +@[simps! isClosed] +def supClosure : ClosureOperator (Set α) := .ofPred (λ s ↦ {a | ∃ (t : Finset α) (ht : t.Nonempty), ↑t ⊆ s ∧ t.sup' ht id = a}) SupClosed (λ s a ha ↦ ⟨{a}, singleton_nonempty _, by simpa⟩) @@ -266,12 +267,11 @@ def supClosure : ClosureOperator (Set α) := ClosureOperator.mk₃ @[simp] lemma subset_supClosure {s : Set α} : s ⊆ supClosure s := supClosure.le_closure _ -@[simp] lemma supClosed_supClosure {s : Set α} : SupClosed (supClosure s) := -ClosureOperator.closure_mem_mk₃ _ +@[simp] lemma supClosed_supClosure : SupClosed (supClosure s) := supClosure.isClosed_closure _ lemma supClosure_mono : Monotone (supClosure : Set α → Set α) := supClosure.monotone -@[simp] lemma supClosure_eq_self : supClosure s = s ↔ SupClosed s := ClosureOperator.mem_mk₃_closed +@[simp] lemma supClosure_eq_self : supClosure s = s ↔ SupClosed s := supClosure.isClosed_iff.symm alias ⟨_, SupClosed.supClosure_eq⟩ := supClosure_eq_self @@ -296,11 +296,7 @@ lemma finsetSup'_mem_supClosure {ι : Type*} {t : Finset ι} (ht : t.Nonempty) { (hf : ∀ i ∈ t, f i ∈ s) : t.sup' ht f ∈ supClosure s := supClosed_supClosure.finsetSup'_mem _ $ fun _i hi ↦ subset_supClosure $ hf _ hi -@[simp] lemma supClosure_closed : supClosure.closed = {s : Set α | SupClosed s} := by - ext; exact ClosureOperator.mem_mk₃_closed - -lemma supClosure_min (hst : s ⊆ t) (ht : SupClosed t) : supClosure s ⊆ t := - ClosureOperator.closure_le_mk₃_iff hst ht +lemma supClosure_min : s ⊆ t → SupClosed t → supClosure s ⊆ t := supClosure.closure_min /-- The semilatice generated by a finite set is finite. -/ protected lemma Set.Finite.supClosure (hs : s.Finite) : (supClosure s).Finite := by @@ -319,7 +315,8 @@ section SemilatticeInf variable [SemilatticeInf α] [SemilatticeInf β] {s t : Set α} {a b : α} /-- Every set in a join-semilattice generates a set closed under join. -/ -def infClosure : ClosureOperator (Set α) := ClosureOperator.mk₃ +@[simps! isClosed] +def infClosure : ClosureOperator (Set α) := ClosureOperator.ofPred (λ s ↦ {a | ∃ (t : Finset α) (ht : t.Nonempty), ↑t ⊆ s ∧ t.inf' ht id = a}) InfClosed (λ s a ha ↦ ⟨{a}, singleton_nonempty _, by simpa⟩) @@ -333,12 +330,11 @@ def infClosure : ClosureOperator (Set α) := ClosureOperator.mk₃ @[simp] lemma subset_infClosure {s : Set α} : s ⊆ infClosure s := infClosure.le_closure _ -@[simp] lemma infClosed_infClosure {s : Set α} : InfClosed (infClosure s) := -ClosureOperator.closure_mem_mk₃ _ +@[simp] lemma infClosed_infClosure : InfClosed (infClosure s) := infClosure.isClosed_closure _ lemma infClosure_mono : Monotone (infClosure : Set α → Set α) := infClosure.monotone -@[simp] lemma infClosure_eq_self : infClosure s = s ↔ InfClosed s := ClosureOperator.mem_mk₃_closed +@[simp] lemma infClosure_eq_self : infClosure s = s ↔ InfClosed s := infClosure.isClosed_iff.symm alias ⟨_, InfClosed.infClosure_eq⟩ := infClosure_eq_self @@ -364,11 +360,7 @@ lemma finsetInf'_mem_infClosure {ι : Type*} {t : Finset ι} (ht : t.Nonempty) { (hf : ∀ i ∈ t, f i ∈ s) : t.inf' ht f ∈ infClosure s := infClosed_infClosure.finsetInf'_mem _ $ fun _i hi ↦ subset_infClosure $ hf _ hi -@[simp] lemma infClosure_closed : infClosure.closed = {s : Set α | InfClosed s} := by - ext; exact ClosureOperator.mem_mk₃_closed - -lemma infClosure_min (hst : s ⊆ t) (ht : InfClosed t) : infClosure s ⊆ t := - ClosureOperator.closure_le_mk₃_iff hst ht +lemma infClosure_min : s ⊆ t → InfClosed t → infClosure s ⊆ t := infClosure.closure_min /-- The semilatice generated by a finite set is finite. -/ protected lemma Set.Finite.infClosure (hs : s.Finite) : (infClosure s).Finite := by @@ -387,25 +379,22 @@ section Lattice variable [Lattice α] {s t : Set α} /-- Every set in a join-semilattice generates a set closed under join. -/ +@[simps! isClosed] def latticeClosure : ClosureOperator (Set α) := - ClosureOperator.ofPred IsSublattice $ fun _ ↦ isSublattice_sInter + .ofCompletePred IsSublattice $ fun _ ↦ isSublattice_sInter @[simp] lemma subset_latticeClosure : s ⊆ latticeClosure s := latticeClosure.le_closure _ @[simp] lemma isSublattice_latticeClosure : IsSublattice (latticeClosure s) := - ClosureOperator.ofPred_spec _ + latticeClosure.isClosed_closure _ -lemma latticeClosure_min (hst : s ⊆ t) (ht : IsSublattice t) : latticeClosure s ⊆ t := by - rw [latticeClosure, ClosureOperator.ofPred] - exact ClosureOperator.closure_le_mk₃_iff hst ht +lemma latticeClosure_min : s ⊆ t → IsSublattice t → latticeClosure s ⊆ t := + latticeClosure.closure_min lemma latticeClosure_mono : Monotone (latticeClosure : Set α → Set α) := latticeClosure.monotone @[simp] lemma latticeClosure_eq_self : latticeClosure s = s ↔ IsSublattice s := - ClosureOperator.mem_closed_ofPred - -@[simp] lemma latticeClosure_closed : latticeClosure.closed = {s : Set α | IsSublattice s} := - ClosureOperator.closed_ofPred + latticeClosure.isClosed_iff.symm alias ⟨_, IsSublattice.latticeClosure_eq⟩ := latticeClosure_eq_self diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index 54c5af38d5118..bebe3af33a68c 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -413,7 +413,9 @@ theorem independent_ne_bot_iff_independent : refine ⟨fun h ↦ ?_, fun h ↦ h.comp Subtype.val_injective⟩ simp only [independent_def] at h ⊢ intro i - cases' eq_or_ne (t i) ⊥ with hi hi; simp [hi] + cases eq_or_ne (t i) ⊥ with + | inl hi => simp [hi] + | inr hi => ?_ convert h ⟨i, hi⟩ have : ∀ j, ⨆ (_ : t j = ⊥), t j = ⊥ := fun j ↦ by simp only [iSup_eq_bot, imp_self] rw [iSup_split _ (fun j ↦ t j = ⊥), iSup_subtype] diff --git a/Mathlib/Probability/Independence/Basic.lean b/Mathlib/Probability/Independence/Basic.lean index 502124125e079..18a7e4b341b6c 100644 --- a/Mathlib/Probability/Independence/Basic.lean +++ b/Mathlib/Probability/Independence/Basic.lean @@ -668,42 +668,113 @@ theorem IndepFun.comp {_mβ : MeasurableSpace β} {_mβ' : MeasurableSpace β'} kernel.IndepFun.comp hfg hφ hψ #align probability_theory.indep_fun.comp ProbabilityTheory.IndepFun.comp +section iIndepFun +variable {β : ι → Type*} {m : ∀ i, MeasurableSpace (β i)} {f : ∀ i, Ω → β i} + +@[nontriviality] +lemma iIndepFun.of_subsingleton [IsProbabilityMeasure μ] [Subsingleton ι] : iIndepFun m f μ := + kernel.iIndepFun.of_subsingleton + +lemma iIndepFun.isProbabilityMeasure (h : iIndepFun m f μ) : IsProbabilityMeasure μ := + ⟨by simpa using h.meas_biInter (S := ∅) (s := fun _ ↦ univ)⟩ + /-- If `f` is a family of mutually independent random variables (`iIndepFun m f μ`) and `S, T` are two disjoint finite index sets, then the tuple formed by `f i` for `i ∈ S` is independent of the tuple `(f i)_i` for `i ∈ T`. -/ -theorem iIndepFun.indepFun_finset [IsProbabilityMeasure μ] {ι : Type*} {β : ι → Type*} - {m : ∀ i, MeasurableSpace (β i)} {f : ∀ i, Ω → β i} (S T : Finset ι) (hST : Disjoint S T) - (hf_Indep : iIndepFun m f μ) (hf_meas : ∀ i, Measurable (f i)) : - IndepFun (fun a (i : S) => f i a) (fun a (i : T) => f i a) μ := +lemma iIndepFun.indepFun_finset (S T : Finset ι) (hST : Disjoint S T) (hf_Indep : iIndepFun m f μ) + (hf_meas : ∀ i, Measurable (f i)) : + IndepFun (fun a (i : S) ↦ f i a) (fun a (i : T) ↦ f i a) μ := + have := hf_Indep.isProbabilityMeasure kernel.iIndepFun.indepFun_finset S T hST hf_Indep hf_meas set_option linter.uppercaseLean3 false in #align probability_theory.Indep_fun.indep_fun_finset ProbabilityTheory.iIndepFun.indepFun_finset -theorem iIndepFun.indepFun_prod [IsProbabilityMeasure μ] {ι : Type*} {β : ι → Type*} - {m : ∀ i, MeasurableSpace (β i)} {f : ∀ i, Ω → β i} (hf_Indep : iIndepFun m f μ) - (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : +lemma iIndepFun.indepFun_prod_mk (hf_Indep : iIndepFun m f μ) (hf_meas : ∀ i, Measurable (f i)) + (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : IndepFun (fun a => (f i a, f j a)) (f k) μ := - kernel.iIndepFun.indepFun_prod hf_Indep hf_meas i j k hik hjk + have := hf_Indep.isProbabilityMeasure + kernel.iIndepFun.indepFun_prod_mk hf_Indep hf_meas i j k hik hjk set_option linter.uppercaseLean3 false in -#align probability_theory.Indep_fun.indep_fun_prod ProbabilityTheory.iIndepFun.indepFun_prod +#align probability_theory.Indep_fun.indep_fun_prod ProbabilityTheory.iIndepFun.indepFun_prod_mk + +open Finset in +lemma iIndepFun.indepFun_prod_mk_prod_mk (h_indep : iIndepFun m f μ) (hf : ∀ i, Measurable (f i)) + (i j k l : ι) (hik : i ≠ k) (hil : i ≠ l) (hjk : j ≠ k) (hjl : j ≠ l) : + IndepFun (fun a ↦ (f i a, f j a)) (fun a ↦ (f k a, f l a)) μ := by + classical + let g (i j : ι) (v : Π x : ({i, j} : Finset ι), β x) : β i × β j := + ⟨v ⟨i, mem_insert_self _ _⟩, v ⟨j, mem_insert_of_mem $ mem_singleton_self _⟩⟩ + have hg (i j : ι) : Measurable (g i j) := by measurability + exact (h_indep.indepFun_finset {i, j} {k, l} (by aesop) hf).comp (hg i j) (hg k l) + +end iIndepFun + +section Mul +variable {β : Type*} {m : MeasurableSpace β} [Mul β] [MeasurableMul₂ β] {f : ι → Ω → β} @[to_additive] -theorem iIndepFun.mul [IsProbabilityMeasure μ] {ι : Type*} {β : Type*} {m : MeasurableSpace β} - [Mul β] [MeasurableMul₂ β] {f : ι → Ω → β} (hf_Indep : iIndepFun (fun _ => m) f μ) +lemma iIndepFun.indepFun_mul_left (hf_indep : iIndepFun (fun _ ↦ m) f μ) (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : IndepFun (f i * f j) (f k) μ := - kernel.iIndepFun.mul hf_Indep hf_meas i j k hik hjk + have := hf_indep.isProbabilityMeasure + kernel.iIndepFun.indepFun_mul_left hf_indep hf_meas i j k hik hjk set_option linter.uppercaseLean3 false in -#align probability_theory.Indep_fun.mul ProbabilityTheory.iIndepFun.mul +#align probability_theory.Indep_fun.mul ProbabilityTheory.iIndepFun.indepFun_mul_left set_option linter.uppercaseLean3 false in -#align probability_theory.Indep_fun.add ProbabilityTheory.iIndepFun.add +#align probability_theory.Indep_fun.add ProbabilityTheory.iIndepFun.indepFun_add_left + +@[to_additive] +lemma iIndepFun.indepFun_mul_right (hf_indep : iIndepFun (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hij : i ≠ j) (hik : i ≠ k) : + IndepFun (f i) (f j * f k) μ := + have := hf_indep.isProbabilityMeasure + kernel.iIndepFun.indepFun_mul_right hf_indep hf_meas i j k hij hik @[to_additive] -theorem iIndepFun.indepFun_finset_prod_of_not_mem [IsProbabilityMeasure μ] {ι : Type*} {β : Type*} - {m : MeasurableSpace β} [CommMonoid β] [MeasurableMul₂ β] {f : ι → Ω → β} - (hf_Indep : iIndepFun (fun _ => m) f μ) (hf_meas : ∀ i, Measurable (f i)) {s : Finset ι} {i : ι} - (hi : i ∉ s) : +lemma iIndepFun.indepFun_mul_mul (hf_indep : iIndepFun (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) + (i j k l : ι) (hik : i ≠ k) (hil : i ≠ l) (hjk : j ≠ k) (hjl : j ≠ l) : + IndepFun (f i * f j) (f k * f l) μ := + have := hf_indep.isProbabilityMeasure + kernel.iIndepFun.indepFun_mul_mul hf_indep hf_meas i j k l hik hil hjk hjl + +end Mul + +section Div +variable {β : Type*} {m : MeasurableSpace β} [Div β] [MeasurableDiv₂ β] {f : ι → Ω → β} + +@[to_additive] +lemma iIndepFun.indepFun_div_left (hf_indep : iIndepFun (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : + IndepFun (f i / f j) (f k) μ := + have := hf_indep.isProbabilityMeasure + kernel.iIndepFun.indepFun_div_left hf_indep hf_meas i j k hik hjk + +@[to_additive] +lemma iIndepFun.indepFun_div_right (hf_indep : iIndepFun (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hij : i ≠ j) (hik : i ≠ k) : + IndepFun (f i) (f j / f k) μ := + have := hf_indep.isProbabilityMeasure + kernel.iIndepFun.indepFun_div_right hf_indep hf_meas i j k hij hik + +@[to_additive] +lemma iIndepFun.indepFun_div_div (hf_indep : iIndepFun (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) + (i j k l : ι) (hik : i ≠ k) (hil : i ≠ l) (hjk : j ≠ k) (hjl : j ≠ l) : + IndepFun (f i / f j) (f k / f l) μ := + have := hf_indep.isProbabilityMeasure + kernel.iIndepFun.indepFun_div_div hf_indep hf_meas i j k l hik hil hjk hjl + +end Div + +section CommMonoid +variable {β : Type*} {m : MeasurableSpace β} [CommMonoid β] [MeasurableMul₂ β] {f : ι → Ω → β} + +@[to_additive] +lemma iIndepFun.indepFun_finset_prod_of_not_mem (hf_Indep : iIndepFun (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) {s : Finset ι} {i : ι} (hi : i ∉ s) : IndepFun (∏ j in s, f j) (f i) μ := + have := hf_Indep.isProbabilityMeasure kernel.iIndepFun.indepFun_finset_prod_of_not_mem hf_Indep hf_meas hi set_option linter.uppercaseLean3 false in #align probability_theory.Indep_fun.indep_fun_finset_prod_of_not_mem ProbabilityTheory.iIndepFun.indepFun_finset_prod_of_not_mem @@ -711,16 +782,17 @@ set_option linter.uppercaseLean3 false in #align probability_theory.Indep_fun.indep_fun_finset_sum_of_not_mem ProbabilityTheory.iIndepFun.indepFun_finset_sum_of_not_mem @[to_additive] -theorem iIndepFun.indepFun_prod_range_succ [IsProbabilityMeasure μ] {β : Type*} - {m : MeasurableSpace β} [CommMonoid β] [MeasurableMul₂ β] {f : ℕ → Ω → β} - (hf_Indep : iIndepFun (fun _ => m) f μ) (hf_meas : ∀ i, Measurable (f i)) (n : ℕ) : - IndepFun (∏ j in Finset.range n, f j) (f n) μ := +lemma iIndepFun.indepFun_prod_range_succ {f : ℕ → Ω → β} (hf_Indep : iIndepFun (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) (n : ℕ) : IndepFun (∏ j in Finset.range n, f j) (f n) μ := + have := hf_Indep.isProbabilityMeasure kernel.iIndepFun.indepFun_prod_range_succ hf_Indep hf_meas n set_option linter.uppercaseLean3 false in #align probability_theory.Indep_fun.indep_fun_prod_range_succ ProbabilityTheory.iIndepFun.indepFun_prod_range_succ set_option linter.uppercaseLean3 false in #align probability_theory.Indep_fun.indep_fun_sum_range_succ ProbabilityTheory.iIndepFun.indepFun_sum_range_succ +end CommMonoid + theorem iIndepSet.iIndepFun_indicator [Zero β] [One β] {m : MeasurableSpace β} {s : ι → Set Ω} (hs : iIndepSet s μ) : iIndepFun (fun _n => m) (fun n => (s n).indicator fun _ω => 1) μ := diff --git a/Mathlib/Probability/Independence/Conditional.lean b/Mathlib/Probability/Independence/Conditional.lean index 537198d6ff5d3..c0c12d16890ba 100644 --- a/Mathlib/Probability/Independence/Conditional.lean +++ b/Mathlib/Probability/Independence/Conditional.lean @@ -425,7 +425,7 @@ end CondIndepSet section CondIndep @[symm] -theorem CondIndep.symm {m' m₁ m₂: MeasurableSpace Ω} [mΩ : MeasurableSpace Ω] +theorem CondIndep.symm {m' m₁ m₂ : MeasurableSpace Ω} [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] (h : CondIndep m' m₁ m₂ hm' μ) : CondIndep m' m₂ m₁ hm' μ := @@ -450,7 +450,7 @@ theorem condIndep_of_condIndep_of_le_left {m' m₁ m₂ m₃ : MeasurableSpace CondIndep m' m₃ m₂ hm' μ := kernel.indep_of_indep_of_le_left h_indep h31 -theorem condIndep_of_condIndep_of_le_right {m' m₁ m₂ m₃: MeasurableSpace Ω} +theorem condIndep_of_condIndep_of_le_right {m' m₁ m₂ m₃ : MeasurableSpace Ω} [mΩ : MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] (h_indep : CondIndep m' m₁ m₂ hm' μ) (h32 : m₃ ≤ m₂) : @@ -694,6 +694,13 @@ theorem CondIndepFun.comp {γ γ' : Type*} {_mβ : MeasurableSpace β} {_mβ' : CondIndepFun m' hm' (φ ∘ f) (ψ ∘ g) μ := kernel.IndepFun.comp hfg hφ hψ +section iCondIndepFun +variable {β : ι → Type*} {m : ∀ i, MeasurableSpace (β i)} {f : ∀ i, Ω → β i} + +@[nontriviality] +lemma iCondIndepFun.of_subsingleton [Subsingleton ι] : iCondIndepFun m' hm' m f μ := + kernel.iIndepFun.of_subsingleton + /-- If `f` is a family of mutually conditionally independent random variables (`iCondIndepFun m' hm' m f μ`) and `S, T` are two disjoint finite index sets, then the tuple formed by `f i` for `i ∈ S` is conditionally independent of the tuple `(f i)_i` for `i ∈ T`. -/ @@ -703,34 +710,91 @@ theorem iCondIndepFun.condIndepFun_finset {β : ι → Type*} CondIndepFun m' hm' (fun a (i : S) => f i a) (fun a (i : T) => f i a) μ := kernel.iIndepFun.indepFun_finset S T hST hf_Indep hf_meas -theorem iCondIndepFun.condIndepFun_prod {β : ι → Type*} +theorem iCondIndepFun.condIndepFun_prod_mk {β : ι → Type*} {m : ∀ i, MeasurableSpace (β i)} {f : ∀ i, Ω → β i} (hf_Indep : iCondIndepFun m' hm' m f μ) (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : CondIndepFun m' hm' (fun a => (f i a, f j a)) (f k) μ := - kernel.iIndepFun.indepFun_prod hf_Indep hf_meas i j k hik hjk + kernel.iIndepFun.indepFun_prod_mk hf_Indep hf_meas i j k hik hjk + +open Finset in +lemma iCondIndepFun.condIndepFun_prod_mk_prod_mk (h_indep : iCondIndepFun m' hm' m f μ) + (hf : ∀ i, Measurable (f i)) + (i j k l : ι) (hik : i ≠ k) (hil : i ≠ l) (hjk : j ≠ k) (hjl : j ≠ l) : + CondIndepFun m' hm' (fun a ↦ (f i a, f j a)) (fun a ↦ (f k a, f l a)) μ := by + classical + let g (i j : ι) (v : Π x : ({i, j} : Finset ι), β x) : β i × β j := + ⟨v ⟨i, mem_insert_self _ _⟩, v ⟨j, mem_insert_of_mem $ mem_singleton_self _⟩⟩ + have hg (i j : ι) : Measurable (g i j) := by measurability + exact (h_indep.indepFun_finset {i, j} {k, l} (by aesop) hf).comp (hg i j) (hg k l) + +end iCondIndepFun + +section Mul +variable {β : Type*} {m : MeasurableSpace β} [Mul β] [MeasurableMul₂ β] {f : ι → Ω → β} @[to_additive] -theorem iCondIndepFun.mul {m : MeasurableSpace β} - [Mul β] [MeasurableMul₂ β] {f : ι → Ω → β} (hf_Indep : iCondIndepFun m' hm' (fun _ => m) f μ) +lemma iCondIndepFun.indepFun_mul_left (hf_indep : iCondIndepFun m' hm' (fun _ ↦ m) f μ) (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : CondIndepFun m' hm' (f i * f j) (f k) μ := - kernel.iIndepFun.mul hf_Indep hf_meas i j k hik hjk + kernel.iIndepFun.indepFun_mul_left hf_indep hf_meas i j k hik hjk + +@[to_additive] +lemma iCondIndepFun.indepFun_mul_right (hf_indep : iCondIndepFun m' hm' (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hij : i ≠ j) (hik : i ≠ k) : + CondIndepFun m' hm' (f i) (f j * f k) μ := + kernel.iIndepFun.indepFun_mul_right hf_indep hf_meas i j k hij hik + +@[to_additive] +lemma iCondIndepFun.indepFun_mul_mul (hf_indep : iCondIndepFun m' hm' (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) + (i j k l : ι) (hik : i ≠ k) (hil : i ≠ l) (hjk : j ≠ k) (hjl : j ≠ l) : + CondIndepFun m' hm' (f i * f j) (f k * f l) μ := + kernel.iIndepFun.indepFun_mul_mul hf_indep hf_meas i j k l hik hil hjk hjl + +end Mul + +section Div +variable {β : Type*} {m : MeasurableSpace β} [Div β] [MeasurableDiv₂ β] {f : ι → Ω → β} + +@[to_additive] +lemma iCondIndepFun.indepFun_div_left (hf_indep : iCondIndepFun m' hm' (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : + CondIndepFun m' hm' (f i / f j) (f k) μ := + kernel.iIndepFun.indepFun_div_left hf_indep hf_meas i j k hik hjk + +@[to_additive] +lemma iCondIndepFun.indepFun_div_right (hf_indep : iCondIndepFun m' hm' (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hij : i ≠ j) (hik : i ≠ k) : + CondIndepFun m' hm' (f i) (f j / f k) μ := + kernel.iIndepFun.indepFun_div_right hf_indep hf_meas i j k hij hik + +@[to_additive] +lemma iCondIndepFun.indepFun_div_div (hf_indep : iCondIndepFun m' hm' (fun _ ↦ m) f μ) + (hf_meas : ∀ i, Measurable (f i)) + (i j k l : ι) (hik : i ≠ k) (hil : i ≠ l) (hjk : j ≠ k) (hjl : j ≠ l) : + CondIndepFun m' hm' (f i / f j) (f k / f l) μ := + kernel.iIndepFun.indepFun_div_div hf_indep hf_meas i j k l hik hil hjk hjl + +end Div + +section CommMonoid +variable {β : Type*} {m : MeasurableSpace β} [CommMonoid β] [MeasurableMul₂ β] {f : ι → Ω → β} @[to_additive] theorem iCondIndepFun.condIndepFun_finset_prod_of_not_mem - {m : MeasurableSpace β} [CommMonoid β] [MeasurableMul₂ β] {f : ι → Ω → β} (hf_Indep : iCondIndepFun m' hm' (fun _ => m) f μ) (hf_meas : ∀ i, Measurable (f i)) {s : Finset ι} {i : ι} (hi : i ∉ s) : CondIndepFun m' hm' (∏ j in s, f j) (f i) μ := kernel.iIndepFun.indepFun_finset_prod_of_not_mem hf_Indep hf_meas hi @[to_additive] -theorem iCondIndepFun.condIndepFun_prod_range_succ - {m : MeasurableSpace β} [CommMonoid β] [MeasurableMul₂ β] {f : ℕ → Ω → β} +theorem iCondIndepFun.condIndepFun_prod_range_succ {f : ℕ → Ω → β} (hf_Indep : iCondIndepFun m' hm' (fun _ => m) f μ) (hf_meas : ∀ i, Measurable (f i)) (n : ℕ) : CondIndepFun m' hm' (∏ j in Finset.range n, f j) (f n) μ := kernel.iIndepFun.indepFun_prod_range_succ hf_Indep hf_meas n +end CommMonoid + theorem iCondIndepSet.iCondIndepFun_indicator [Zero β] [One β] {m : MeasurableSpace β} {s : ι → Set Ω} (hs : iCondIndepSet m' hm' s μ) : iCondIndepFun m' hm' (fun _n => m) (fun n => (s n).indicator fun _ω => 1) μ := diff --git a/Mathlib/Probability/Independence/Kernel.lean b/Mathlib/Probability/Independence/Kernel.lean index a91c7306ecc90..89457332a0f2f 100644 --- a/Mathlib/Probability/Independence/Kernel.lean +++ b/Mathlib/Probability/Independence/Kernel.lean @@ -809,14 +809,28 @@ theorem IndepFun.comp {mβ : MeasurableSpace β} {mβ' : MeasurableSpace β'} · exact ⟨φ ⁻¹' A, hφ hA, Set.preimage_comp.symm⟩ · exact ⟨ψ ⁻¹' B, hψ hB, Set.preimage_comp.symm⟩ +section iIndepFun +variable {β : ι → Type*} {m : ∀ i, MeasurableSpace (β i)} {f : ∀ i, Ω → β i} + +@[nontriviality] +lemma iIndepFun.of_subsingleton [IsMarkovKernel κ] [Subsingleton ι] : iIndepFun m f κ μ := by + refine (iIndepFun_iff_measure_inter_preimage_eq_mul ..).2 fun s f' hf' ↦ ?_ + obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty + · simp + · have : s = {x} := by ext y; simp [Subsingleton.elim y x, hx] + simp [this] + +lemma iIndepFun.ae_isProbabilityMeasure (h : iIndepFun m f κ μ) : + ∀ᵐ a ∂μ, IsProbabilityMeasure (κ a) := by + simpa [isProbabilityMeasure_iff] using h.meas_biInter (S := ∅) (s := fun _ ↦ Set.univ) + /-- If `f` is a family of mutually independent random variables (`iIndepFun m f μ`) and `S, T` are two disjoint finite index sets, then the tuple formed by `f i` for `i ∈ S` is independent of the tuple `(f i)_i` for `i ∈ T`. -/ -theorem iIndepFun.indepFun_finset [IsMarkovKernel κ] {ι : Type*} {β : ι → Type*} - {m : ∀ i, MeasurableSpace (β i)} {f : ∀ i, Ω → β i} (S T : Finset ι) (hST : Disjoint S T) +theorem iIndepFun.indepFun_finset [IsMarkovKernel κ] (S T : Finset ι) (hST : Disjoint S T) (hf_Indep : iIndepFun m f κ μ) (hf_meas : ∀ i, Measurable (f i)) : IndepFun (fun a (i : S) => f i a) (fun a (i : T) => f i a) κ μ := by - -- We introduce π-systems, build from the π-system of boxes which generates `MeasurableSpace.pi`. + -- We introduce π-systems, built from the π-system of boxes which generates `MeasurableSpace.pi`. let πSβ := Set.pi (Set.univ : Set S) '' Set.pi (Set.univ : Set S) fun i => { s : Set (β i) | MeasurableSet[m i] s } let πS := { s : Set Ω | ∃ t ∈ πSβ, (fun a (i : S) => f i a) ⁻¹' t = s } @@ -905,8 +919,7 @@ theorem iIndepFun.indepFun_finset [IsMarkovKernel κ] {ι : Type*} {β : ι → show κ a (f i ⁻¹' (sets_s' i ∩ sets_t' i)) = κ a (f i ⁻¹' (sets_t' i)) rw [h_sets_s'_univ hi, Set.univ_inter] -theorem iIndepFun.indepFun_prod [IsMarkovKernel κ] {ι : Type*} {β : ι → Type*} - {m : ∀ i, MeasurableSpace (β i)} {f : ∀ i, Ω → β i} (hf_Indep : iIndepFun m f κ μ) +theorem iIndepFun.indepFun_prod_mk [IsMarkovKernel κ] (hf_Indep : iIndepFun m f κ μ) (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : IndepFun (fun a => (f i a, f j a)) (f k) κ μ := by classical @@ -932,21 +945,84 @@ theorem iIndepFun.indepFun_prod [IsMarkovKernel κ] {ι : Type*} {β : ι → Ty simp only [Finset.mem_insert, Finset.mem_singleton, not_or] exact ⟨hik.symm, hjk.symm⟩ +open Finset in +lemma iIndepFun.indepFun_prod_mk_prod_mk [IsMarkovKernel κ] (hf_indep : iIndepFun m f κ μ) + (hf_meas : ∀ i, Measurable (f i)) + (i j k l : ι) (hik : i ≠ k) (hil : i ≠ l) (hjk : j ≠ k) (hjl : j ≠ l) : + IndepFun (fun a ↦ (f i a, f j a)) (fun a ↦ (f k a, f l a)) κ μ := by + classical + let g (i j : ι) (v : Π x : ({i, j} : Finset ι), β x) : β i × β j := + ⟨v ⟨i, mem_insert_self _ _⟩, v ⟨j, mem_insert_of_mem $ mem_singleton_self _⟩⟩ + have hg (i j : ι) : Measurable (g i j) := by measurability + exact (hf_indep.indepFun_finset {i, j} {k, l} (by aesop) hf_meas).comp (hg i j) (hg k l) + +end iIndepFun + +section Mul +variable {β : Type*} {m : MeasurableSpace β} [Mul β] [MeasurableMul₂ β] {f : ι → Ω → β} + [IsMarkovKernel κ] + @[to_additive] -theorem iIndepFun.mul [IsMarkovKernel κ] {ι : Type*} {β : Type*} {m : MeasurableSpace β} - [Mul β] [MeasurableMul₂ β] {f : ι → Ω → β} (hf_Indep : iIndepFun (fun _ => m) f κ μ) +lemma iIndepFun.indepFun_mul_left (hf_indep : iIndepFun (fun _ ↦ m) f κ μ) (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : IndepFun (f i * f j) (f k) κ μ := by have : IndepFun (fun ω => (f i ω, f j ω)) (f k) κ μ := - hf_Indep.indepFun_prod hf_meas i j k hik hjk + hf_indep.indepFun_prod_mk hf_meas i j k hik hjk change IndepFun ((fun p : β × β => p.fst * p.snd) ∘ fun ω => (f i ω, f j ω)) (id ∘ f k) κ μ - exact IndepFun.comp this (measurable_fst.mul measurable_snd) measurable_id + exact this.comp (measurable_fst.mul measurable_snd) measurable_id + +@[to_additive] +lemma iIndepFun.indepFun_mul_right (hf_indep : iIndepFun (fun _ ↦ m) f κ μ) + (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hij : i ≠ j) (hik : i ≠ k) : + IndepFun (f i) (f j * f k) κ μ := + (hf_indep.indepFun_mul_left hf_meas _ _ _ hij.symm hik.symm).symm + +@[to_additive] +lemma iIndepFun.indepFun_mul_mul (hf_indep : iIndepFun (fun _ ↦ m) f κ μ) + (hf_meas : ∀ i, Measurable (f i)) + (i j k l : ι) (hik : i ≠ k) (hil : i ≠ l) (hjk : j ≠ k) (hjl : j ≠ l) : + IndepFun (f i * f j) (f k * f l) κ μ := + (hf_indep.indepFun_prod_mk_prod_mk hf_meas i j k l hik hil hjk hjl).comp + measurable_mul measurable_mul + +end Mul + +section Div +variable {β : Type*} {m : MeasurableSpace β} [Div β] [MeasurableDiv₂ β] {f : ι → Ω → β} + [IsMarkovKernel κ] + +@[to_additive] +lemma iIndepFun.indepFun_div_left (hf_indep : iIndepFun (fun _ ↦ m) f κ μ) + (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : + IndepFun (f i / f j) (f k) κ μ := by + have : IndepFun (fun ω => (f i ω, f j ω)) (f k) κ μ := + hf_indep.indepFun_prod_mk hf_meas i j k hik hjk + change IndepFun ((fun p : β × β => p.fst / p.snd) ∘ fun ω => (f i ω, f j ω)) (id ∘ f k) κ μ + exact this.comp (measurable_fst.div measurable_snd) measurable_id + +@[to_additive] +lemma iIndepFun.indepFun_div_right (hf_indep : iIndepFun (fun _ ↦ m) f κ μ) + (hf_meas : ∀ i, Measurable (f i)) (i j k : ι) (hij : i ≠ j) (hik : i ≠ k) : + IndepFun (f i) (f j / f k) κ μ := + (hf_indep.indepFun_div_left hf_meas _ _ _ hij.symm hik.symm).symm + +@[to_additive] +lemma iIndepFun.indepFun_div_div (hf_indep : iIndepFun (fun _ ↦ m) f κ μ) + (hf_meas : ∀ i, Measurable (f i)) + (i j k l : ι) (hik : i ≠ k) (hil : i ≠ l) (hjk : j ≠ k) (hjl : j ≠ l) : + IndepFun (f i / f j) (f k / f l) κ μ := + (hf_indep.indepFun_prod_mk_prod_mk hf_meas i j k l hik hil hjk hjl).comp + measurable_div measurable_div + +end Div + +section CommMonoid +variable {β : Type*} {m : MeasurableSpace β} [CommMonoid β] [MeasurableMul₂ β] {f : ι → Ω → β} + [IsMarkovKernel κ] @[to_additive] -theorem iIndepFun.indepFun_finset_prod_of_not_mem [IsMarkovKernel κ] {ι : Type*} {β : Type*} - {m : MeasurableSpace β} [CommMonoid β] [MeasurableMul₂ β] {f : ι → Ω → β} - (hf_Indep : iIndepFun (fun _ => m) f κ μ) (hf_meas : ∀ i, Measurable (f i)) - {s : Finset ι} {i : ι} (hi : i ∉ s) : +theorem iIndepFun.indepFun_finset_prod_of_not_mem (hf_Indep : iIndepFun (fun _ ↦ m) f κ μ) + (hf_meas : ∀ i, Measurable (f i)) {s : Finset ι} {i : ι} (hi : i ∉ s) : IndepFun (∏ j in s, f j) (f i) κ μ := by classical have h_right : f i = @@ -967,12 +1043,13 @@ theorem iIndepFun.indepFun_finset_prod_of_not_mem [IsMarkovKernel κ] {ι : Type h_meas_left h_meas_right @[to_additive] -theorem iIndepFun.indepFun_prod_range_succ [IsMarkovKernel κ] {β : Type*} - {m : MeasurableSpace β} [CommMonoid β] [MeasurableMul₂ β] {f : ℕ → Ω → β} +theorem iIndepFun.indepFun_prod_range_succ {f : ℕ → Ω → β} (hf_Indep : iIndepFun (fun _ => m) f κ μ) (hf_meas : ∀ i, Measurable (f i)) (n : ℕ) : IndepFun (∏ j in Finset.range n, f j) (f n) κ μ := hf_Indep.indepFun_finset_prod_of_not_mem hf_meas Finset.not_mem_range_self +end CommMonoid + theorem iIndepSet.iIndepFun_indicator [Zero β] [One β] {m : MeasurableSpace β} {s : ι → Set Ω} (hs : iIndepSet s κ μ) : iIndepFun (fun _n => m) (fun n => (s n).indicator fun _ω => 1) κ μ := by diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian.lean index 45ecbf69526b7..8cee2bae5f281 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian.lean @@ -498,4 +498,27 @@ instance isMaximal_of_isPrime (p : Ideal R) [p.IsPrime] : p.IsMaximal := localization_surjective (nonZeroDivisors (R ⧸ p)) (FractionRing (R ⧸ p))⟩).isField _ <| Field.toIsField <| FractionRing (R ⧸ p) +lemma isPrime_iff_isMaximal (p : Ideal R) : p.IsPrime ↔ p.IsMaximal := + ⟨fun _ ↦ isMaximal_of_isPrime p, fun h ↦ h.isPrime⟩ + +variable (R) in +lemma primeSpectrum_finite : {I : Ideal R | I.IsPrime}.Finite := by + set Spec := {I : Ideal R | I.IsPrime} + obtain ⟨_, ⟨s, rfl⟩, H⟩ := set_has_minimal + (range (Finset.inf · Subtype.val : Finset Spec → Ideal R)) ⟨⊤, ∅, by simp⟩ + refine ⟨⟨s, fun p ↦ ?_⟩⟩ + classical + obtain ⟨q, hq1, hq2⟩ := p.2.inf_le'.mp <| inf_eq_right.mp <| + inf_le_right.eq_of_not_lt (H (p ⊓ s.inf Subtype.val) ⟨insert p s, by simp⟩) + rwa [← Subtype.ext <| (@isMaximal_of_isPrime _ _ _ _ q.2).eq_of_le p.2.1 hq2] + +variable (R) in +/-- +[Stacks Lemma 00J7](https://stacks.math.columbia.edu/tag/00J7) +-/ +lemma maximal_ideals_finite : {I : Ideal R | I.IsMaximal}.Finite := by + simp_rw [← isPrime_iff_isMaximal] + apply primeSpectrum_finite R + + end IsArtinianRing diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index 38eaa127af764..1bb2b2f96eafc 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -687,10 +687,11 @@ instance Module.Finite.base_change [CommSemiring R] [Semiring A] [Algebra R A] [ [Module R M] [h : Module.Finite R M] : Module.Finite A (TensorProduct R A M) := by classical obtain ⟨s, hs⟩ := h.out - refine' ⟨⟨s.image (TensorProduct.mk R A M 1), eq_top_iff.mpr fun x _ => _⟩⟩ - apply TensorProduct.induction_on (motive := _) x - · exact zero_mem _ - · intro x y + refine ⟨⟨s.image (TensorProduct.mk R A M 1), eq_top_iff.mpr ?_⟩⟩ + rintro x - + induction x using TensorProduct.induction_on with + | zero => exact zero_mem _ + | tmul x y => -- Porting note: new TC reminder haveI : IsScalarTower R A (TensorProduct R A M) := TensorProduct.isScalarTower_left rw [Finset.coe_image, ← Submodule.span_span_of_tower R, Submodule.span_image, hs, @@ -698,7 +699,7 @@ instance Module.Finite.base_change [CommSemiring R] [Semiring A] [Algebra R A] [ change _ ∈ Submodule.span A (Set.range <| TensorProduct.mk R A M 1) rw [← mul_one x, ← smul_eq_mul, ← TensorProduct.smul_tmul'] exact Submodule.smul_mem _ x (Submodule.subset_span <| Set.mem_range_self y) - · exact fun _ _ => Submodule.add_mem _ + | add x y hx hy => exact Submodule.add_mem _ hx hy #align module.finite.base_change Module.Finite.base_change instance Module.Finite.tensorProduct [CommSemiring R] [AddCommMonoid M] [Module R M] diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index 4a43d9787bd3d..53992ef79332c 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -811,8 +811,8 @@ theorem isLocalization_iff_of_algEquiv [Algebra R P] (h : S ≃ₐ[R] P) : #align is_localization.is_localization_iff_of_alg_equiv IsLocalization.isLocalization_iff_of_algEquiv theorem isLocalization_iff_of_ringEquiv (h : S ≃+* P) : - IsLocalization M S ↔ haveI := (h.toRingHom.comp <| algebraMap R S).toAlgebra; - IsLocalization M P := + IsLocalization M S ↔ + haveI := (h.toRingHom.comp <| algebraMap R S).toAlgebra; IsLocalization M P := letI := (h.toRingHom.comp <| algebraMap R S).toAlgebra isLocalization_iff_of_algEquiv M { h with commutes' := fun _ => rfl } #align is_localization.is_localization_iff_of_ring_equiv IsLocalization.isLocalization_iff_of_ringEquiv @@ -820,8 +820,8 @@ theorem isLocalization_iff_of_ringEquiv (h : S ≃+* P) : variable (S) theorem isLocalization_of_base_ringEquiv [IsLocalization M S] (h : R ≃+* P) : - haveI := ((algebraMap R S).comp h.symm.toRingHom).toAlgebra; - IsLocalization (M.map h.toMonoidHom) S := by + haveI := ((algebraMap R S).comp h.symm.toRingHom).toAlgebra + IsLocalization (M.map h.toMonoidHom) S := by letI : Algebra P S := ((algebraMap R S).comp h.symm.toRingHom).toAlgebra constructor · rintro ⟨_, ⟨y, hy, rfl⟩⟩ @@ -843,7 +843,8 @@ theorem isLocalization_of_base_ringEquiv [IsLocalization M S] (h : R ≃+* P) : #align is_localization.is_localization_of_base_ring_equiv IsLocalization.isLocalization_of_base_ringEquiv theorem isLocalization_iff_of_base_ringEquiv (h : R ≃+* P) : - IsLocalization M S ↔ haveI := ((algebraMap R S).comp h.symm.toRingHom).toAlgebra; + IsLocalization M S ↔ + haveI := ((algebraMap R S).comp h.symm.toRingHom).toAlgebra IsLocalization (M.map h.toMonoidHom) S := by letI : Algebra P S := ((algebraMap R S).comp h.symm.toRingHom).toAlgebra refine' ⟨fun _ => isLocalization_of_base_ringEquiv M S h, _⟩ diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index cfffeeaf38935..77a50b18ef731 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -169,7 +169,7 @@ theorem mk'_eq_div {r} (s : nonZeroDivisors A) : mk' K r s = algebraMap A K r / #align is_fraction_ring.mk'_eq_div IsFractionRing.mk'_eq_div theorem div_surjective (z : K) : - ∃ (x y : A) (_ : y ∈ nonZeroDivisors A), algebraMap _ _ x / algebraMap _ _ y = z := + ∃ x y : A, y ∈ nonZeroDivisors A ∧ algebraMap _ _ x / algebraMap _ _ y = z := let ⟨x, ⟨y, hy⟩, h⟩ := mk'_surjective (nonZeroDivisors A) z ⟨x, y, hy, by rwa [mk'_eq_div] at h⟩ #align is_fraction_ring.div_surjective IsFractionRing.div_surjective diff --git a/Mathlib/Testing/SlimCheck/Functions.lean b/Mathlib/Testing/SlimCheck/Functions.lean index 9e042dd0ef2f6..0c7104b1c3bc5 100644 --- a/Mathlib/Testing/SlimCheck/Functions.lean +++ b/Mathlib/Testing/SlimCheck/Functions.lean @@ -35,7 +35,7 @@ permutation. One has to be careful when generating the domain to make it vast enough that, when generating arguments to apply `f` to, they argument should be likely to lie in the domain of `f`. This is the reason that injective functions `f : ℤ → ℤ` are generated by -fixing the domain to the range `[-2*size .. -2*size]`, with `size` +fixing the domain to the range `[-2*size .. 2*size]`, with `size` the size parameter of the `gen` monad. Much of the machinery provided in this file is applicable to generate @@ -61,11 +61,11 @@ namespace SlimCheck and a default value returned when the input is not in the domain of the partial function. -`with_default f y` encodes `x ↦ f x` when `x ∈ f` and `x ↦ y` +`withDefault f y` encodes `x ↦ f x` when `x ∈ f` and `x ↦ y` otherwise. We use `Σ` to encode mappings instead of `×` because we -rely on the association list API defined in `data.list.sigma`. +rely on the association list API defined in `Mathlib/Data/List/Sigma.lean`. -/ inductive TotalFunction (α : Type u) (β : Type v) : Type max u v | withDefault : List (Σ _ : α, β) → β → TotalFunction α β @@ -89,7 +89,7 @@ def apply [DecidableEq α] : TotalFunction α β → α → β | TotalFunction.withDefault m y, x => (m.dlookup x).getD y #align slim_check.total_function.apply SlimCheck.TotalFunction.apply -/-- Implementation of `has_repr (total_function α β)`. +/-- Implementation of `Repr (TotalFunction α β)`. Creates a string for a given `finmap` and output, `x₀ ↦ y₀, .. xₙ ↦ yₙ` for each of the entries. The brackets are provided by the calling function. @@ -168,7 +168,7 @@ def zeroDefault : TotalFunction α β → TotalFunction α β variable [DecidableEq α] [DecidableEq β] -/-- The support of a zero default `total_function`. -/ +/-- The support of a zero default `TotalFunction`. -/ @[simp] def zeroDefaultSupp : TotalFunction α β → Finset α | withDefault A _ => @@ -259,11 +259,11 @@ namespace SlimCheck and a default value returned when the input is not in the domain of the partial function. -`map_to_self f` encodes `x ↦ f x` when `x ∈ f` and `x ↦ x`, +`mapToSelf f` encodes `x ↦ f x` when `x ∈ f` and `x ↦ x`, i.e. `x` to itself, otherwise. We use `Σ` to encode mappings instead of `×` because we -rely on the association list API defined in `data.list.sigma`. +rely on the association list API defined in `Mathlib/Data/List/Sigma.lean`. -/ inductive InjectiveFunction (α : Type u) : Type u | mapToSelf (xs : List (Σ _ : α, α)) : @@ -281,9 +281,9 @@ def apply [DecidableEq α] : InjectiveFunction α → α → α | InjectiveFunction.mapToSelf m _ _, x => (m.dlookup x).getD x #align slim_check.injective_function.apply SlimCheck.InjectiveFunction.apply -/-- Produce a string for a given `total_function`. +/-- Produce a string for a given `InjectiveFunction`. The output is of the form `[x₀ ↦ f x₀, .. xₙ ↦ f xₙ, x ↦ x]`. -Unlike for `total_function`, the default value is not a constant +Unlike for `TotalFunction`, the default value is not a constant but the identity function. -/ protected def repr [Repr α] : InjectiveFunction α → String @@ -299,7 +299,6 @@ def List.applyId [DecidableEq α] (xs : List (α × α)) (x : α) : α := ((xs.map Prod.toSigma).dlookup x).getD x #align slim_check.injective_function.list.apply_id SlimCheck.InjectiveFunction.List.applyId -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] theorem List.applyId_cons [DecidableEq α] (xs : List (α × α)) (x y z : α) : List.applyId ((y, z)::xs) x = if y = x then z else List.applyId xs x := by @@ -315,9 +314,9 @@ open Nat theorem List.applyId_zip_eq [DecidableEq α] {xs ys : List α} (h₀ : List.Nodup xs) (h₁ : xs.length = ys.length) (x y : α) (i : ℕ) (h₂ : xs.get? i = some x) : List.applyId.{u} (xs.zip ys) x = y ↔ ys.get? i = some y := by - induction xs generalizing ys i - case nil => cases h₂ - case cons x' xs xs_ih => + induction xs generalizing ys i with + | nil => cases h₂ + | cons x' xs xs_ih => cases i · injection h₂ with h₀; subst h₀ cases ys @@ -360,7 +359,7 @@ theorem applyId_mem_iff [DecidableEq α] {xs ys : List α} (h₀ : List.Nodup xs have h₆ := Nat.succ.inj h₁ specialize xs_ih h₅ h₃ h₄ h₆ simp only [Ne.symm h, xs_ih, List.mem_cons, false_or_iff] - suffices : val ∈ ys; tauto + suffices val ∈ ys by tauto erw [← Option.mem_def, List.mem_dlookup_iff] at h₃ simp only [Prod.toSigma, List.mem_map, heq_iff_eq, Prod.exists] at h₃ rcases h₃ with ⟨a, b, h₃, h₄, h₅⟩ @@ -489,10 +488,9 @@ protected theorem injective [DecidableEq α] (f : InjectiveFunction α) : Inject dsimp [id] at h₁ have hxs : xs = TotalFunction.List.toFinmap' (xs₀.zip xs₁) := by rw [← h₀, ← h₁, List.toFinmap']; clear h₀ h₁ xs₀ xs₁ hperm hnodup - induction xs - case nil => simp only [List.zip_nil_right, List.map_nil] - case cons xs_hd xs_tl - xs_ih => + induction xs with + | nil => simp only [List.zip_nil_right, List.map_nil] + | cons xs_hd xs_tl xs_ih => simp only [true_and_iff, Prod.toSigma, eq_self_iff_true, Sigma.eta, List.zip_cons_cons, List.map, List.cons_inj] exact xs_ih diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean index 586b37fe8e26a..23a6b5c68250c 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean @@ -1347,6 +1347,30 @@ theorem Summable.countable_support [FirstCountableTopology G] [T1Space G] (hf : Summable f) : f.support.Countable := by simpa only [ker_nhds] using hf.tendsto_cofinite_zero.countable_compl_preimage_ker +theorem summable_const_iff [Infinite β] [T2Space G] (a : G) : + Summable (fun _ : β ↦ a) ↔ a = 0 := by + refine ⟨fun h ↦ ?_, ?_⟩ + · by_contra ha + have : {a}ᶜ ∈ 𝓝 0 := compl_singleton_mem_nhds (Ne.symm ha) + have : Finite β := by + simpa [← Set.finite_univ_iff] using h.tendsto_cofinite_zero this + exact not_finite β + · rintro rfl + exact summable_zero + +@[simp] +theorem tsum_const [T2Space G] : ∑' _ : β, (a : G) = Nat.card β • a := by + rcases finite_or_infinite β with hβ|hβ + · letI : Fintype β := Fintype.ofFinite β + rw [tsum_eq_sum (s := univ) (fun x hx ↦ (hx (mem_univ x)).elim)] + simp only [sum_const, Nat.card_eq_fintype_card] + rfl + · simp only [Nat.card_eq_zero_of_infinite, zero_smul] + rcases eq_or_ne a 0 with rfl|ha + · simp + · apply tsum_eq_zero_of_not_summable + simpa [summable_const_iff] using ha + end TopologicalGroup section ConstSMul diff --git a/Mathlib/Topology/Algebra/UniformConvergence.lean b/Mathlib/Topology/Algebra/UniformConvergence.lean index 694c1a0250a5e..5e927c0ac278e 100644 --- a/Mathlib/Topology/Algebra/UniformConvergence.lean +++ b/Mathlib/Topology/Algebra/UniformConvergence.lean @@ -50,16 +50,80 @@ uniform convergence, strong dual -/ -set_option autoImplicit true +open Filter +open scoped Topology Pointwise UniformConvergence +section AlgebraicInstances -open Filter +variable {α β ι R : Type*} {𝔖 : Set <| Set α} {x : α} -open Topology Pointwise UniformConvergence +@[to_additive] instance [One β] : One (α →ᵤ β) := Pi.instOne -section AlgebraicInstances +@[to_additive (attr := simp)] +lemma UniformFun.toFun_one [One β] : toFun (1 : α →ᵤ β) = 1 := rfl + +@[to_additive (attr := simp)] +lemma UniformFun.ofFun_one [One β] : ofFun (1 : α → β) = 1 := rfl + +@[to_additive] instance [One β] : One (α →ᵤ[𝔖] β) := Pi.instOne + +@[to_additive (attr := simp)] +lemma UniformOnFun.toFun_one [One β] : toFun 𝔖 (1 : α →ᵤ[𝔖] β) = 1 := rfl + +@[to_additive (attr := simp)] +lemma UniformOnFun.one_apply [One β] : ofFun 𝔖 (1 : α → β) = 1 := rfl + +@[to_additive] instance [Mul β] : Mul (α →ᵤ β) := Pi.instMul + +@[to_additive (attr := simp)] +lemma UniformFun.toFun_mul [Mul β] (f g : α →ᵤ β) : toFun (f * g) = toFun f * toFun g := rfl + +@[to_additive (attr := simp)] +lemma UniformFun.ofFun_mul [Mul β] (f g : α → β) : ofFun (f * g) = ofFun f * ofFun g := rfl + +@[to_additive] instance [Mul β] : Mul (α →ᵤ[𝔖] β) := Pi.instMul + +@[to_additive (attr := simp)] +lemma UniformOnFun.toFun_mul [Mul β] (f g : α →ᵤ[𝔖] β) : + toFun 𝔖 (f * g) = toFun 𝔖 f * toFun 𝔖 g := + rfl + +@[to_additive (attr := simp)] +lemma UniformOnFun.ofFun_mul [Mul β] (f g : α → β) : ofFun 𝔖 (f * g) = ofFun 𝔖 f * ofFun 𝔖 g := rfl + +@[to_additive] instance [Inv β] : Inv (α →ᵤ β) := Pi.instInv + +@[to_additive (attr := simp)] +lemma UniformFun.toFun_inv [Inv β] (f : α →ᵤ β) : toFun (f⁻¹) = (toFun f)⁻¹ := rfl + +@[to_additive (attr := simp)] +lemma UniformFun.ofFun_inv [Inv β] (f : α → β) : ofFun (f⁻¹) = (ofFun f)⁻¹ := rfl + +@[to_additive] instance [Inv β] : Inv (α →ᵤ[𝔖] β) := Pi.instInv + +@[to_additive (attr := simp)] +lemma UniformOnFun.toFun_inv [Inv β] (f : α →ᵤ[𝔖] β) : toFun 𝔖 (f⁻¹) = (toFun 𝔖 f)⁻¹ := rfl + +@[to_additive (attr := simp)] +lemma UniformOnFun.ofFun_inv [Inv β] (f : α → β) : ofFun 𝔖 (f⁻¹) = (ofFun 𝔖 f)⁻¹ := rfl -variable {α β ι R : Type*} {𝔖 : Set <| Set α} +@[to_additive] instance [Div β] : Div (α →ᵤ β) := Pi.instDiv + +@[to_additive (attr := simp)] +lemma UniformFun.toFun_div [Div β] (f g : α →ᵤ β) : toFun (f / g) = toFun f / toFun g := rfl + +@[to_additive (attr := simp)] +lemma UniformFun.ofFun_div [Div β] (f g : α → β) : ofFun (f / g) = ofFun f / ofFun g := rfl + +@[to_additive] instance [Div β] : Div (α →ᵤ[𝔖] β) := Pi.instDiv + +@[to_additive (attr := simp)] +lemma UniformOnFun.toFun_div [Div β] (f g : α →ᵤ[𝔖] β) : + toFun 𝔖 (f / g) = toFun 𝔖 f / toFun 𝔖 g := + rfl + +@[to_additive (attr := simp)] +lemma UniformOnFun.ofFun_div [Div β] (f g : α → β) : ofFun 𝔖 (f / g) = ofFun 𝔖 f / ofFun 𝔖 g := rfl @[to_additive] instance [Monoid β] : Monoid (α →ᵤ β) := @@ -93,37 +157,63 @@ instance [CommGroup β] : CommGroup (α →ᵤ β) := instance [CommGroup β] : CommGroup (α →ᵤ[𝔖] β) := Pi.commGroup -instance [Semiring R] [AddCommMonoid β] [Module R β] : Module R (α →ᵤ β) := - Pi.module _ _ _ +instance {M : Type*} [SMul M β] : SMul M (α →ᵤ β) := Pi.instSMul -instance [Semiring R] [AddCommMonoid β] [Module R β] : Module R (α →ᵤ[𝔖] β) := - Pi.module _ _ _ +@[simp] +lemma UniformFun.toFun_smul {M : Type*} [SMul M β] (c : M) (f : α →ᵤ β) : + toFun (c • f) = c • toFun f := + rfl --- Porting note: unfortunately `simp` will no longer use `Pi.one_apply` etc. --- on `α →ᵤ β` or `α →ᵤ[𝔖] β`, so we restate some of these here. More may be needed later. -@[to_additive (attr := simp)] -lemma UniformFun.one_apply [Monoid β] : (1 : α →ᵤ β) x = 1 := Pi.one_apply x +@[simp] +lemma UniformFun.ofFun_smul {M : Type*} [SMul M β] (c : M) (f : α → β) : + ofFun (c • f) = c • ofFun f := + rfl -@[to_additive (attr := simp)] -lemma UniformOnFun.one_apply [Monoid β] : (1 : α →ᵤ[𝔖] β) x = 1 := Pi.one_apply x +instance {M : Type*} [SMul M β] : SMul M (α →ᵤ[𝔖] β) := Pi.instSMul -@[to_additive (attr := simp)] -lemma UniformFun.mul_apply [Monoid β] : (f * g : α →ᵤ β) x = f x * g x := Pi.mul_apply f g x +@[simp] +lemma UniformOnFun.toFun_smul {M : Type*} [SMul M β] (c : M) (f : α →ᵤ[𝔖] β) : + toFun 𝔖 (c • f) = c • toFun 𝔖 f := + rfl -@[to_additive (attr := simp)] -lemma UniformOnFun.mul_apply [Monoid β] : (f * g : α →ᵤ[𝔖] β) x = f x * g x := Pi.mul_apply f g x +@[simp] +lemma UniformOnFun.ofFun_smul {M : Type*} [SMul M β] (c : M) (f : α → β) : + ofFun 𝔖 (c • f) = c • ofFun 𝔖 f := + rfl -@[to_additive (attr := simp)] -lemma UniformFun.inv_apply [Group β] : (f : α →ᵤ β)⁻¹ x = (f x)⁻¹ := Pi.inv_apply f x +instance {M N : Type*} [SMul M N] [SMul M β] [SMul N β] [IsScalarTower M N β] : + IsScalarTower M N (α →ᵤ β) := + Pi.isScalarTower -@[to_additive (attr := simp)] -lemma UniformOnFun.inv_apply [Group β] : (f : α →ᵤ[𝔖] β)⁻¹ x = (f x)⁻¹ := Pi.inv_apply f x +instance {M N : Type*} [SMul M N] [SMul M β] [SMul N β] [IsScalarTower M N β] : + IsScalarTower M N (α →ᵤ[𝔖] β) := + Pi.isScalarTower -@[to_additive (attr := simp)] -lemma UniformFun.div_apply [Group β] : (f / g : α →ᵤ β) x = f x / g x := Pi.div_apply f g x +instance {M N : Type*} [SMul M β] [SMul N β] [SMulCommClass M N β] : + SMulCommClass M N (α →ᵤ β) := + Pi.smulCommClass -@[to_additive (attr := simp)] -lemma UniformOnFun.div_apply [Group β] : (f / g : α →ᵤ[𝔖] β) x = f x / g x := Pi.div_apply f g x +instance {M N : Type*} [SMul M β] [SMul N β] [SMulCommClass M N β] : + SMulCommClass M N (α →ᵤ[𝔖] β) := + Pi.smulCommClass + +instance {M : Type*} [Monoid M] [MulAction M β] : MulAction M (α →ᵤ β) := Pi.mulAction _ + +instance {M : Type*} [Monoid M] [MulAction M β] : MulAction M (α →ᵤ[𝔖] β) := Pi.mulAction _ + +instance {M : Type*} [Monoid M] [AddMonoid β] [DistribMulAction M β] : + DistribMulAction M (α →ᵤ β) := + Pi.distribMulAction _ + +instance {M : Type*} [Monoid M] [AddMonoid β] [DistribMulAction M β] : + DistribMulAction M (α →ᵤ[𝔖] β) := + Pi.distribMulAction _ + +instance [Semiring R] [AddCommMonoid β] [Module R β] : Module R (α →ᵤ β) := + Pi.module _ _ _ + +instance [Semiring R] [AddCommMonoid β] [Module R β] : Module R (α →ᵤ[𝔖] β) := + Pi.module _ _ _ end AlgebraicInstances @@ -146,12 +236,12 @@ instance : UniformGroup (α →ᵤ G) := @[to_additive] protected theorem UniformFun.hasBasis_nhds_one_of_basis {p : ι → Prop} {b : ι → Set G} (h : (𝓝 1 : Filter G).HasBasis p b) : - (𝓝 1 : Filter (α →ᵤ G)).HasBasis p fun i => { f : α →ᵤ G | ∀ x, f x ∈ b i } := by + (𝓝 1 : Filter (α →ᵤ G)).HasBasis p fun i => { f : α →ᵤ G | ∀ x, toFun f x ∈ b i } := by have := h.comap fun p : G × G => p.2 / p.1 rw [← uniformity_eq_comap_nhds_one] at this convert UniformFun.hasBasis_nhds_of_basis α _ (1 : α →ᵤ G) this -- Porting note: removed `ext i f` here, as it has already been done by `convert`. - simp [UniformFun.gen] + simp #align uniform_fun.has_basis_nhds_one_of_basis UniformFun.hasBasis_nhds_one_of_basis #align uniform_fun.has_basis_nhds_zero_of_basis UniformFun.hasBasis_nhds_zero_of_basis @@ -181,7 +271,7 @@ protected theorem UniformOnFun.hasBasis_nhds_one_of_basis (𝔖 : Set <| Set α) (h𝔖₂ : DirectedOn (· ⊆ ·) 𝔖) {p : ι → Prop} {b : ι → Set G} (h : (𝓝 1 : Filter G).HasBasis p b) : (𝓝 1 : Filter (α →ᵤ[𝔖] G)).HasBasis (fun Si : Set α × ι => Si.1 ∈ 𝔖 ∧ p Si.2) fun Si => - { f : α →ᵤ[𝔖] G | ∀ x ∈ Si.1, f x ∈ b Si.2 } := by + { f : α →ᵤ[𝔖] G | ∀ x ∈ Si.1, toFun 𝔖 f x ∈ b Si.2 } := by have := h.comap fun p : G × G => p.1 / p.2 rw [← uniformity_eq_comap_nhds_one_swapped] at this convert UniformOnFun.hasBasis_nhds_of_basis α _ 𝔖 (1 : α →ᵤ[𝔖] G) h𝔖₁ h𝔖₂ this diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index 5c2b5736e149c..71109c438ce32 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -1326,6 +1326,9 @@ theorem mem_closure_iff_nhdsWithin_neBot {s : Set α} {x : α} : x ∈ closure s mem_closure_iff_clusterPt #align mem_closure_iff_nhds_within_ne_bot mem_closure_iff_nhdsWithin_neBot +lemma not_mem_closure_iff_nhdsWithin_eq_bot {s : Set α} {x : α} : x ∉ closure s ↔ 𝓝[s] x = ⊥ := by + rw [mem_closure_iff_nhdsWithin_neBot, not_neBot] + /-- If `x` is not an isolated point of a topological space, then `{x}ᶜ` is dense in the whole space. -/ theorem dense_compl_singleton (x : α) [NeBot (𝓝[≠] x)] : Dense ({x}ᶜ : Set α) := by diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 5bcdfe33ff9af..821878173534c 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -796,12 +796,29 @@ theorem finite_of_compact_of_discrete [CompactSpace X] [DiscreteTopology X] : Fi Finite.of_finite_univ <| isCompact_univ.finite_of_discrete #align finite_of_compact_of_discrete finite_of_compact_of_discrete +lemma Set.Infinite.exists_accPt_cofinite_inf_principal_of_subset_isCompact + {K : Set X} (hs : s.Infinite) (hK : IsCompact K) (hsub : s ⊆ K) : + ∃ x ∈ K, AccPt x (cofinite ⊓ 𝓟 s) := + (@hK _ hs.cofinite_inf_principal_neBot (inf_le_right.trans <| principal_mono.2 hsub)).imp + fun x hx ↦ by rwa [acc_iff_cluster, inf_comm, inf_right_comm, + (finite_singleton _).cofinite_inf_principal_compl] + +lemma Set.Infinite.exists_accPt_of_subset_isCompact {K : Set X} (hs : s.Infinite) + (hK : IsCompact K) (hsub : s ⊆ K) : ∃ x ∈ K, AccPt x (𝓟 s) := + let ⟨x, hxK, hx⟩ := hs.exists_accPt_cofinite_inf_principal_of_subset_isCompact hK hsub + ⟨x, hxK, hx.mono inf_le_right⟩ + +lemma Set.Infinite.exists_accPt_cofinite_inf_principal [CompactSpace X] (hs : s.Infinite) : + ∃ x, AccPt x (cofinite ⊓ 𝓟 s) := by + simpa only [mem_univ, true_and] + using hs.exists_accPt_cofinite_inf_principal_of_subset_isCompact isCompact_univ s.subset_univ + +lemma Set.Infinite.exists_accPt_principal [CompactSpace X] (hs : s.Infinite) : ∃ x, AccPt x (𝓟 s) := + hs.exists_accPt_cofinite_inf_principal.imp fun _x hx ↦ hx.mono inf_le_right + theorem exists_nhds_ne_neBot (X : Type*) [TopologicalSpace X] [CompactSpace X] [Infinite X] : - ∃ x : X, (𝓝[≠] x).NeBot := by - by_contra! H - simp_rw [not_neBot] at H - haveI := discreteTopology_iff_nhds_ne.2 H - exact Infinite.not_finite (finite_of_compact_of_discrete : Finite X) + ∃ z : X, (𝓝[≠] z).NeBot := by + simpa [AccPt] using (@infinite_univ X _).exists_accPt_principal #align exists_nhds_ne_ne_bot exists_nhds_ne_neBot theorem finite_cover_nhds_interior [CompactSpace X] {U : X → Set X} (hU : ∀ x, U x ∈ 𝓝 x) : @@ -937,10 +954,8 @@ theorem IsCompact.finite (hs : IsCompact s) (hs' : DiscreteTopology s) : s.Finit #align is_compact.finite IsCompact.finite theorem exists_nhds_ne_inf_principal_neBot (hs : IsCompact s) (hs' : s.Infinite) : - ∃ x ∈ s, (𝓝[≠] x ⊓ 𝓟 s).NeBot := by - by_contra! H - simp_rw [not_neBot] at H - exact hs' (hs.finite <| discreteTopology_subtype_iff.mpr H) + ∃ z ∈ s, (𝓝[≠] z ⊓ 𝓟 s).NeBot := + hs'.exists_accPt_of_subset_isCompact hs Subset.rfl #align exists_nhds_ne_inf_principal_ne_bot exists_nhds_ne_inf_principal_neBot protected theorem ClosedEmbedding.noncompactSpace [NoncompactSpace X] {f : X → Y} diff --git a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean index 31c964e4a08ad..c5f184f3f8286 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean @@ -165,32 +165,37 @@ scoped[UniformConvergence] notation:25 α " →ᵤ[" 𝔖 "] " β:0 => UniformOn open UniformConvergence -instance {α β} [Nonempty β] : Nonempty (α →ᵤ β) := - Pi.Nonempty +variable {α β : Type*} {𝔖 : Set (Set α)} -instance {α β 𝔖} [Nonempty β] : Nonempty (α →ᵤ[𝔖] β) := - Pi.Nonempty +instance [Nonempty β] : Nonempty (α →ᵤ β) := Pi.Nonempty + +instance [Nonempty β] : Nonempty (α →ᵤ[𝔖] β) := Pi.Nonempty /-- Reinterpret `f : α → β` as an element of `α →ᵤ β`. -/ -def UniformFun.ofFun {α β} : (α → β) ≃ (α →ᵤ β) := +def UniformFun.ofFun : (α → β) ≃ (α →ᵤ β) := ⟨fun x => x, fun x => x, fun _ => rfl, fun _ => rfl⟩ #align uniform_fun.of_fun UniformFun.ofFun /-- Reinterpret `f : α → β` as an element of `α →ᵤ[𝔖] β`. -/ -def UniformOnFun.ofFun {α β} (𝔖) : (α → β) ≃ (α →ᵤ[𝔖] β) := +def UniformOnFun.ofFun (𝔖) : (α → β) ≃ (α →ᵤ[𝔖] β) := ⟨fun x => x, fun x => x, fun _ => rfl, fun _ => rfl⟩ #align uniform_on_fun.of_fun UniformOnFun.ofFun /-- Reinterpret `f : α →ᵤ β` as an element of `α → β`. -/ -def UniformFun.toFun {α β} : (α →ᵤ β) ≃ (α → β) := +def UniformFun.toFun : (α →ᵤ β) ≃ (α → β) := UniformFun.ofFun.symm #align uniform_fun.to_fun UniformFun.toFun /-- Reinterpret `f : α →ᵤ[𝔖] β` as an element of `α → β`. -/ -def UniformOnFun.toFun {α β} (𝔖) : (α →ᵤ[𝔖] β) ≃ (α → β) := +def UniformOnFun.toFun (𝔖) : (α →ᵤ[𝔖] β) ≃ (α → β) := (UniformOnFun.ofFun 𝔖).symm #align uniform_on_fun.to_fun UniformOnFun.toFun +@[simp] lemma UniformFun.toFun_ofFun (f : α → β) : toFun (ofFun f) = f := rfl +@[simp] lemma UniformFun.ofFun_toFun (f : α →ᵤ β) : ofFun (toFun f) = f := rfl +@[simp] lemma UniformOnFun.toFun_ofFun (f : α → β) : toFun 𝔖 (ofFun 𝔖 f) = f := rfl +@[simp] lemma UniformOnFun.ofFun_toFun (f : α →ᵤ[𝔖] β) : ofFun 𝔖 (toFun 𝔖 f) = f := rfl + -- Note: we don't declare a `CoeFun` instance because Lean wouldn't insert it when writing -- `f x` (because of definitional equality with `α → β`). end TypeAlias @@ -206,7 +211,7 @@ variable {s s' : Set α} {x : α} {p : Filter ι} {g : ι → α} /-- Basis sets for the uniformity of uniform convergence: `gen α β V` is the set of pairs `(f, g)` of functions `α →ᵤ β` such that `∀ x, (f x, g x) ∈ V`. -/ protected def gen (V : Set (β × β)) : Set ((α →ᵤ β) × (α →ᵤ β)) := - { uv : (α →ᵤ β) × (α →ᵤ β) | ∀ x, (uv.1 x, uv.2 x) ∈ V } + { uv : (α →ᵤ β) × (α →ᵤ β) | ∀ x, (toFun uv.1 x, toFun uv.2 x) ∈ V } #align uniform_fun.gen UniformFun.gen /-- If `𝓕` is a filter on `β × β`, then the set of all `UniformFun.gen α β V` for @@ -248,7 +253,7 @@ The exact definition of the lower adjoint `l` is not interesting; we will only u (in `UniformFun.mono` and `UniformFun.iInf_eq`) and that `l (Filter.map (Prod.map f f) 𝓕) = Filter.map (Prod.map ((∘) f) ((∘) f)) (l 𝓕)` for each `𝓕 : Filter (γ × γ)` and `f : γ → α` (in `UniformFun.comap_eq`). -/ -local notation "lower_adjoint" => fun 𝓐 => map (UniformFun.phi α β) (𝓐 ×ˢ ⊤) +local notation "lowerAdjoint" => fun 𝓐 => map (UniformFun.phi α β) (𝓐 ×ˢ ⊤) /-- The function `UniformFun.filter α β : Filter (β × β) → Filter ((α →ᵤ β) × (α →ᵤ β))` has a lower adjoint `l` (in the sense of `GaloisConnection`). The exact definition of `l` is not @@ -256,7 +261,7 @@ interesting; we will only use that it exists (in `UniformFun.mono` and `UniformFun.iInf_eq`) and that `l (Filter.map (Prod.map f f) 𝓕) = Filter.map (Prod.map ((∘) f) ((∘) f)) (l 𝓕)` for each `𝓕 : Filter (γ × γ)` and `f : γ → α` (in `UniformFun.comap_eq`). -/ -protected theorem gc : GaloisConnection lower_adjoint fun 𝓕 => UniformFun.filter α β 𝓕 := by +protected theorem gc : GaloisConnection lowerAdjoint fun 𝓕 => UniformFun.filter α β 𝓕 := by intro 𝓐 𝓕 symm calc @@ -271,7 +276,7 @@ protected theorem gc : GaloisConnection lower_adjoint fun 𝓕 => UniformFun.fil { uvx : ((α →ᵤ β) × (α →ᵤ β)) × α | (uvx.1.1 uvx.2, uvx.1.2 uvx.2) ∈ U } ∈ 𝓐 ×ˢ (⊤ : Filter α) := forall₂_congr fun U _hU => mem_prod_top.symm - _ ↔ lower_adjoint 𝓐 ≤ 𝓕 := Iff.rfl + _ ↔ lowerAdjoint 𝓐 ≤ 𝓕 := Iff.rfl #align uniform_fun.gc UniformFun.gc variable [UniformSpace β] @@ -305,7 +310,7 @@ local notation "𝒰(" α ", " β ", " u ")" => @UniformFun.uniformSpace α β u /-- By definition, the uniformity of `α →ᵤ β` admits the family `{(f, g) | ∀ x, (f x, g x) ∈ V}` for `V ∈ 𝓤 β` as a filter basis. -/ protected theorem hasBasis_uniformity : - (𝓤 (α →ᵤ β)).HasBasis (fun V => V ∈ 𝓤 β) (UniformFun.gen α β) := + (𝓤 (α →ᵤ β)).HasBasis (· ∈ 𝓤 β) (UniformFun.gen α β) := (UniformFun.isBasis_gen α β (𝓤 β)).hasBasis #align uniform_fun.has_basis_uniformity UniformFun.hasBasis_uniformity @@ -349,6 +354,11 @@ theorem uniformContinuous_eval (x : α) : variable {β} +@[simp] +protected lemma mem_gen {f g : α →ᵤ β} {V : Set (β × β)} : + (f, g) ∈ UniformFun.gen α β V ↔ ∀ x, (toFun f x, toFun g x) ∈ V := + .rfl + /-- If `u₁` and `u₂` are two uniform structures on `γ` and `u₁ ≤ u₂`, then `𝒰(α, γ, u₁) ≤ 𝒰(α, γ, u₂)`. -/ protected theorem mono : Monotone (@UniformFun.uniformSpace α γ) := fun _ _ hu => @@ -428,7 +438,7 @@ uniform convergence. More precisely, for any `f : γ → α`, the function `(· ∘ f) : (α →ᵤ β) → (γ →ᵤ β)` is uniformly continuous. -/ protected theorem precomp_uniformContinuous {f : γ → α} : - UniformContinuous fun g : α →ᵤ β => ofFun (g ∘ f) := by + UniformContinuous fun g : α →ᵤ β => ofFun (toFun g ∘ f) := by -- Here we simply go back to filter bases. rw [uniformContinuous_iff] change @@ -440,18 +450,12 @@ protected theorem precomp_uniformContinuous {f : γ → α} : /-- Turn a bijection `γ ≃ α` into a uniform isomorphism `(γ →ᵤ β) ≃ᵤ (α →ᵤ β)` by pre-composing. -/ -protected def congrLeft (e : γ ≃ α) : (γ →ᵤ β) ≃ᵤ (α →ᵤ β) := - { Equiv.arrowCongr e (Equiv.refl _) with - uniformContinuous_toFun := UniformFun.precomp_uniformContinuous - uniformContinuous_invFun := UniformFun.precomp_uniformContinuous } +protected def congrLeft (e : γ ≃ α) : (γ →ᵤ β) ≃ᵤ (α →ᵤ β) where + toEquiv := e.arrowCongr (.refl _) + uniformContinuous_toFun := UniformFun.precomp_uniformContinuous + uniformContinuous_invFun := UniformFun.precomp_uniformContinuous #align uniform_fun.congr_left UniformFun.congrLeft -/-- The topology of uniform convergence is T₂. -/ -instance [T2Space β] : T2Space (α →ᵤ β) where - t2 f g h := by - obtain ⟨x, hx⟩ := not_forall.mp (mt funext h) - exact separated_by_continuous (uniformContinuous_eval β x).continuous hx - /-- The natural map `UniformFun.toFun` from `α →ᵤ β` to `α → β` is uniformly continuous. In other words, the uniform structure of uniform convergence is finer than that of pointwise @@ -463,12 +467,16 @@ protected theorem uniformContinuous_toFun : UniformContinuous (toFun : (α → exact uniformContinuous_eval β x #align uniform_fun.uniform_continuous_to_fun UniformFun.uniformContinuous_toFun +/-- The topology of uniform convergence is T₂. -/ +instance [T2Space β] : T2Space (α →ᵤ β) := + .of_injective_continuous toFun.injective UniformFun.uniformContinuous_toFun.continuous + /-- The topology of uniform convergence indeed gives the same notion of convergence as `TendstoUniformly`. -/ protected theorem tendsto_iff_tendstoUniformly {F : ι → α →ᵤ β} {f : α →ᵤ β} : - Tendsto F p (𝓝 f) ↔ TendstoUniformly F f p := by + Tendsto F p (𝓝 f) ↔ TendstoUniformly (toFun ∘ F) (toFun f) p := by rw [(UniformFun.hasBasis_nhds α β f).tendsto_right_iff, TendstoUniformly] - rfl + simp only [mem_setOf, UniformFun.gen, Function.comp_def] #align uniform_fun.tendsto_iff_tendsto_uniformly UniformFun.tendsto_iff_tendstoUniformly /-- The natural bijection between `α → β × γ` and `(α → β) × (α → γ)`, upgraded to a uniform @@ -479,7 +487,7 @@ protected def uniformEquivProdArrow [UniformSpace γ] : (α →ᵤ β × γ) ≃ -- But `uβ × uγ` is defined as `comap fst uβ ⊓ comap snd uγ`, so we just have to apply -- `UniformFun.inf_eq` and `UniformFun.comap_eq`, which leaves us to check -- that some square commutes. - Equiv.toUniformEquivOfUniformInducing (Equiv.arrowProdEquivProdArrow _ _ _) $ by + Equiv.toUniformEquivOfUniformInducing (Equiv.arrowProdEquivProdArrow _ _ _) <| by constructor change comap (Prod.map (Equiv.arrowProdEquivProdArrow _ _ _) (Equiv.arrowProdEquivProdArrow _ _ _)) @@ -537,7 +545,7 @@ local notation "𝒰(" α ", " β ", " u ")" => @UniformFun.uniformSpace α β u `∀ x ∈ S, (f x, g x) ∈ V`. Note that the family `𝔖 : Set (Set α)` is only used to specify which type alias of `α → β` to use here. -/ protected def gen (𝔖) (S : Set α) (V : Set (β × β)) : Set ((α →ᵤ[𝔖] β) × (α →ᵤ[𝔖] β)) := - { uv : (α →ᵤ[𝔖] β) × (α →ᵤ[𝔖] β) | ∀ x ∈ S, (uv.1 x, uv.2 x) ∈ V } + { uv : (α →ᵤ[𝔖] β) × (α →ᵤ[𝔖] β) | ∀ x ∈ S, (toFun 𝔖 uv.1 x, toFun 𝔖 uv.2 x) ∈ V } #align uniform_on_fun.gen UniformOnFun.gen /-- For `S : Set α` and `V : Set (β × β)`, we have @@ -596,8 +604,8 @@ of `S.restrict : (α →ᵤ[𝔖] β) → (↥S →ᵤ β)` of restriction to `S the topology of uniform convergence. -/ protected theorem topologicalSpace_eq : UniformOnFun.topologicalSpace α β 𝔖 = - ⨅ (s : Set α) (_ : s ∈ 𝔖), TopologicalSpace.induced (s.restrict ∘ UniformFun.toFun) - (UniformFun.topologicalSpace s β) := by + ⨅ (s : Set α) (_ : s ∈ 𝔖), TopologicalSpace.induced + (UniformFun.ofFun ∘ s.restrict ∘ toFun 𝔖) (UniformFun.topologicalSpace s β) := by simp only [UniformOnFun.topologicalSpace, UniformSpace.toTopologicalSpace_iInf] rfl #align uniform_on_fun.topological_space_eq UniformOnFun.topologicalSpace_eq @@ -829,7 +837,7 @@ protected theorem uniformContinuous_toFun (h : ⋃₀ 𝔖 = univ) : /-- Convergence in the topology of `𝔖`-convergence means uniform convergence on `S` (in the sense of `TendstoUniformlyOn`) for all `S ∈ 𝔖`. -/ protected theorem tendsto_iff_tendstoUniformlyOn {F : ι → α →ᵤ[𝔖] β} {f : α →ᵤ[𝔖] β} : - Tendsto F p (𝓝 f) ↔ ∀ s ∈ 𝔖, TendstoUniformlyOn F f p s := by + Tendsto F p (𝓝 f) ↔ ∀ s ∈ 𝔖, TendstoUniformlyOn (toFun 𝔖 ∘ F) (toFun 𝔖 f) p s := by rw [UniformOnFun.topologicalSpace_eq, nhds_iInf, tendsto_iInf] refine' forall_congr' fun s => _ rw [nhds_iInf, tendsto_iInf] diff --git a/lake-manifest.json b/lake-manifest.json index aaaeaf41309fc..bfa77660582b7 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -4,7 +4,7 @@ [{"url": "https://github.com/leanprover/std4", "type": "git", "subDir": null, - "rev": "a17925cb39040c461c92ce63cd152f1767b02832", + "rev": "49b36f282e56ce9a695923e2df156cd4abf33fc8", "name": "std", "manifestFile": "lake-manifest.json", "inputRev": "bump/v4.5.0", diff --git a/test/Recall.lean b/test/Recall.lean index 8809da844a2b6..3222dd5eb3f86 100644 --- a/test/Recall.lean +++ b/test/Recall.lean @@ -17,8 +17,8 @@ section variable {𝕜 : Type _} [NontriviallyNormedField 𝕜] variable {E : Type _} [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable {F : Type _} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -recall HasFDerivAtFilter (f : E → F) (f' : E →L[𝕜] F) (x : E) (L : Filter E) := - (fun x' => f x' - f x - f' (x' - x)) =o[L] fun x' => x' - x +recall HasFDerivAt (f : E → F) (f' : E →L[𝕜] F) (x : E) := + HasFDerivAtFilter f f' x (nhds x) end /--