diff --git a/.github/workflows/latest_import.yml b/.github/workflows/latest_import.yml index aa7a52cd7f9811..6e9b5b0cf2cfd9 100644 --- a/.github/workflows/latest_import.yml +++ b/.github/workflows/latest_import.yml @@ -48,8 +48,10 @@ jobs: - name: add minImports linter option run: | - # set `linter.minImports option` to true in `lakefile` - sed -i -- '/^ -- '\`'latest_import.yml'\`' uses this comment/{s=^= ⟨`linter.minImports, true⟩,\n=}' lakefile.lean + # we disable checking for backticks inside single quotes with the next line + # shellcheck disable=SC2016 + # set `linter.minImports option` to true and `Elab.async` to false in `lakefile` + sed -i -- '/^ -- '\`'latest_import.yml'\`' uses this comment/{s=^= ⟨`linter.minImports, true⟩,\n ⟨`Elab.async, false⟩,\n=}' lakefile.lean # import the `minImports` linter in `Mathlib.Init` sed -i -z 's=^=import Mathlib.Tactic.Linter.MinImports\n=' Mathlib/Init.lean diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index 10a12e5f3a439f..1f6e215d7fa1ec 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -77,13 +77,15 @@ jobs: ( steps.user_permission.outputs.require-result == 'true' || steps.merge_or_delegate.outputs.bot == 'true' ) }} id: remove_labels - name: Remove "awaiting-author" + name: Remove "awaiting-author" and "maintainer-merge" # we use curl rather than octokit/request-action so that the job won't fail # (and send an annoying email) if the labels don't exist run: | - curl --request DELETE \ - --url "https://api.github.com/repos/${{ github.repository }}/issues/${{ github.event.issue.number }}${{ github.event.pull_request.number }}/labels/awaiting-author" \ - --header 'authorization: Bearer ${{ secrets.TRIAGE_TOKEN }}' + for label in awaiting-author maintainer-merge; do + curl --request DELETE \ + --url "https://api.github.com/repos/${{ github.repository }}/issues/${{ github.event.issue.number }}${{ github.event.pull_request.number }}/labels/${label}" \ + --header 'authorization: Bearer ${{ secrets.TRIAGE_TOKEN }}' + done - name: Set up Python if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && diff --git a/Mathlib.lean b/Mathlib.lean index e5041a5fd5c1bb..f62f59976932f7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1,3 +1,4 @@ +import Std import Batteries import Mathlib.Algebra.AddConstMap.Basic import Mathlib.Algebra.AddConstMap.Equiv @@ -355,6 +356,7 @@ import Mathlib.Algebra.Group.Submonoid.Defs import Mathlib.Algebra.Group.Submonoid.DistribMulAction import Mathlib.Algebra.Group.Submonoid.Finsupp import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.Group.Submonoid.MulAction import Mathlib.Algebra.Group.Submonoid.MulOpposite import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.Group.Submonoid.Pointwise @@ -449,6 +451,7 @@ import Mathlib.Algebra.Homology.Embedding.StupidTrunc import Mathlib.Algebra.Homology.Embedding.TruncGE import Mathlib.Algebra.Homology.Embedding.TruncGEHomology import Mathlib.Algebra.Homology.Embedding.TruncLE +import Mathlib.Algebra.Homology.Embedding.TruncLEHomology import Mathlib.Algebra.Homology.ExactSequence import Mathlib.Algebra.Homology.Factorizations.Basic import Mathlib.Algebra.Homology.Functor @@ -1166,6 +1169,7 @@ import Mathlib.Analysis.BoxIntegral.Partition.Tagged import Mathlib.Analysis.BoxIntegral.UnitPartition import Mathlib.Analysis.CStarAlgebra.ApproximateUnit import Mathlib.Analysis.CStarAlgebra.Basic +import Mathlib.Analysis.CStarAlgebra.CStarMatrix import Mathlib.Analysis.CStarAlgebra.Classes import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Basic import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances @@ -1732,6 +1736,7 @@ import Mathlib.CategoryTheory.Category.ULift import Mathlib.CategoryTheory.ChosenFiniteProducts import Mathlib.CategoryTheory.ChosenFiniteProducts.Cat import Mathlib.CategoryTheory.ChosenFiniteProducts.FunctorCategory +import Mathlib.CategoryTheory.ChosenFiniteProducts.InfSemilattice import Mathlib.CategoryTheory.Closed.Cartesian import Mathlib.CategoryTheory.Closed.Enrichment import Mathlib.CategoryTheory.Closed.Functor @@ -1845,7 +1850,7 @@ import Mathlib.CategoryTheory.Galois.Prorepresentability import Mathlib.CategoryTheory.Galois.Topology import Mathlib.CategoryTheory.Generator.Abelian import Mathlib.CategoryTheory.Generator.Basic -import Mathlib.CategoryTheory.Generator.Coproduct +import Mathlib.CategoryTheory.Generator.Indization import Mathlib.CategoryTheory.Generator.Preadditive import Mathlib.CategoryTheory.Generator.Presheaf import Mathlib.CategoryTheory.Generator.Sheaf @@ -1930,6 +1935,7 @@ import Mathlib.CategoryTheory.Limits.Indization.Equalizers import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits import Mathlib.CategoryTheory.Limits.Indization.IndObject import Mathlib.CategoryTheory.Limits.Indization.LocallySmall +import Mathlib.CategoryTheory.Limits.Indization.ParallelPair import Mathlib.CategoryTheory.Limits.Indization.Products import Mathlib.CategoryTheory.Limits.IsConnected import Mathlib.CategoryTheory.Limits.IsLimit @@ -1939,6 +1945,7 @@ import Mathlib.CategoryTheory.Limits.MorphismProperty import Mathlib.CategoryTheory.Limits.Opposites import Mathlib.CategoryTheory.Limits.Over import Mathlib.CategoryTheory.Limits.Pi +import Mathlib.CategoryTheory.Limits.Preorder import Mathlib.CategoryTheory.Limits.Preserves.Basic import Mathlib.CategoryTheory.Limits.Preserves.Filtered import Mathlib.CategoryTheory.Limits.Preserves.Finite @@ -2143,6 +2150,7 @@ import Mathlib.CategoryTheory.Preadditive.Projective import Mathlib.CategoryTheory.Preadditive.ProjectiveResolution import Mathlib.CategoryTheory.Preadditive.Schur import Mathlib.CategoryTheory.Preadditive.SingleObj +import Mathlib.CategoryTheory.Preadditive.Transfer import Mathlib.CategoryTheory.Preadditive.Yoneda.Basic import Mathlib.CategoryTheory.Preadditive.Yoneda.Injective import Mathlib.CategoryTheory.Preadditive.Yoneda.Limits @@ -2263,6 +2271,7 @@ import Mathlib.CategoryTheory.Subobject.Types import Mathlib.CategoryTheory.Subobject.WellPowered import Mathlib.CategoryTheory.Subpresheaf.Basic import Mathlib.CategoryTheory.Subpresheaf.Image +import Mathlib.CategoryTheory.Subpresheaf.OfSection import Mathlib.CategoryTheory.Subpresheaf.Sieves import Mathlib.CategoryTheory.Subterminal import Mathlib.CategoryTheory.Sums.Associator @@ -2500,6 +2509,7 @@ import Mathlib.Data.Complex.FiniteDimensional import Mathlib.Data.Complex.Module import Mathlib.Data.Complex.Order import Mathlib.Data.Complex.Orientation +import Mathlib.Data.Complex.Trigonometric import Mathlib.Data.Countable.Basic import Mathlib.Data.Countable.Defs import Mathlib.Data.Countable.Small @@ -3082,7 +3092,6 @@ import Mathlib.Deprecated.Cardinal.Finite import Mathlib.Deprecated.Cardinal.PartENat import Mathlib.Deprecated.Combinator import Mathlib.Deprecated.Equiv -import Mathlib.Deprecated.Group import Mathlib.Deprecated.HashMap import Mathlib.Deprecated.LazyList import Mathlib.Deprecated.Logic @@ -3090,11 +3099,6 @@ import Mathlib.Deprecated.MinMax import Mathlib.Deprecated.NatLemmas import Mathlib.Deprecated.Order import Mathlib.Deprecated.RelClasses -import Mathlib.Deprecated.Ring -import Mathlib.Deprecated.Subfield -import Mathlib.Deprecated.Subgroup -import Mathlib.Deprecated.Submonoid -import Mathlib.Deprecated.Subring import Mathlib.Dynamics.BirkhoffSum.Average import Mathlib.Dynamics.BirkhoffSum.Basic import Mathlib.Dynamics.BirkhoffSum.NormedSpace @@ -3231,8 +3235,8 @@ import Mathlib.Geometry.Manifold.IntegralCurve.Basic import Mathlib.Geometry.Manifold.IntegralCurve.ExistUnique import Mathlib.Geometry.Manifold.IntegralCurve.Transform import Mathlib.Geometry.Manifold.IntegralCurve.UniformTime -import Mathlib.Geometry.Manifold.InteriorBoundary -import Mathlib.Geometry.Manifold.IsManifold +import Mathlib.Geometry.Manifold.IsManifold.Basic +import Mathlib.Geometry.Manifold.IsManifold.InteriorBoundary import Mathlib.Geometry.Manifold.LocalDiffeomorph import Mathlib.Geometry.Manifold.LocalInvariantProperties import Mathlib.Geometry.Manifold.MFDeriv.Atlas @@ -3742,6 +3746,7 @@ import Mathlib.MeasureTheory.Covering.OneDim import Mathlib.MeasureTheory.Covering.Vitali import Mathlib.MeasureTheory.Covering.VitaliFamily import Mathlib.MeasureTheory.Decomposition.Exhaustion +import Mathlib.MeasureTheory.Decomposition.IntegralRNDeriv import Mathlib.MeasureTheory.Decomposition.Jordan import Mathlib.MeasureTheory.Decomposition.Lebesgue import Mathlib.MeasureTheory.Decomposition.RadonNikodym @@ -3768,7 +3773,9 @@ import Mathlib.MeasureTheory.Function.EssSup import Mathlib.MeasureTheory.Function.Floor import Mathlib.MeasureTheory.Function.Intersectivity import Mathlib.MeasureTheory.Function.Jacobian -import Mathlib.MeasureTheory.Function.L1Space +import Mathlib.MeasureTheory.Function.L1Space.AEEqFun +import Mathlib.MeasureTheory.Function.L1Space.HasFiniteIntegral +import Mathlib.MeasureTheory.Function.L1Space.Integrable import Mathlib.MeasureTheory.Function.L2Space import Mathlib.MeasureTheory.Function.LocallyIntegrable import Mathlib.MeasureTheory.Function.LpOrder @@ -4019,6 +4026,7 @@ import Mathlib.NumberTheory.LSeries.Nonvanishing import Mathlib.NumberTheory.LSeries.Positivity import Mathlib.NumberTheory.LSeries.PrimesInAP import Mathlib.NumberTheory.LSeries.RiemannZeta +import Mathlib.NumberTheory.LSeries.SumCoeff import Mathlib.NumberTheory.LSeries.ZMod import Mathlib.NumberTheory.LegendreSymbol.AddCharacter import Mathlib.NumberTheory.LegendreSymbol.Basic @@ -4155,6 +4163,7 @@ import Mathlib.Order.Compare import Mathlib.Order.CompleteBooleanAlgebra import Mathlib.Order.CompleteLattice import Mathlib.Order.CompleteLattice.Finset +import Mathlib.Order.CompleteLattice.SetLike import Mathlib.Order.CompleteLatticeIntervals import Mathlib.Order.CompletePartialOrder import Mathlib.Order.CompleteSublattice @@ -4355,7 +4364,10 @@ import Mathlib.Probability.Independence.ZeroOne import Mathlib.Probability.Integration import Mathlib.Probability.Kernel.Basic import Mathlib.Probability.Kernel.Composition.Basic +import Mathlib.Probability.Kernel.Composition.CompNotation import Mathlib.Probability.Kernel.Composition.IntegralCompProd +import Mathlib.Probability.Kernel.Composition.Lemmas +import Mathlib.Probability.Kernel.Composition.MeasureComp import Mathlib.Probability.Kernel.Composition.MeasureCompProd import Mathlib.Probability.Kernel.Composition.ParallelComp import Mathlib.Probability.Kernel.CondDistrib @@ -4384,7 +4396,9 @@ import Mathlib.Probability.Martingale.OptionalSampling import Mathlib.Probability.Martingale.OptionalStopping import Mathlib.Probability.Martingale.Upcrossing import Mathlib.Probability.Moments.Basic +import Mathlib.Probability.Moments.ComplexMGF import Mathlib.Probability.Moments.IntegrableExpMul +import Mathlib.Probability.Moments.MGFAnalytic import Mathlib.Probability.Notation import Mathlib.Probability.ProbabilityMassFunction.Basic import Mathlib.Probability.ProbabilityMassFunction.Binomial @@ -4632,6 +4646,7 @@ import Mathlib.RingTheory.Localization.LocalizationLocalization import Mathlib.RingTheory.Localization.Module import Mathlib.RingTheory.Localization.NormTrace import Mathlib.RingTheory.Localization.NumDen +import Mathlib.RingTheory.Localization.Pi import Mathlib.RingTheory.Localization.Submodule import Mathlib.RingTheory.MatrixAlgebra import Mathlib.RingTheory.Multiplicity @@ -4860,6 +4875,7 @@ import Mathlib.SetTheory.Cardinal.SchroederBernstein import Mathlib.SetTheory.Cardinal.Subfield import Mathlib.SetTheory.Cardinal.ToNat import Mathlib.SetTheory.Cardinal.UnivLE +import Mathlib.SetTheory.Descriptive.Tree import Mathlib.SetTheory.Game.Basic import Mathlib.SetTheory.Game.Birthday import Mathlib.SetTheory.Game.Domineering diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index c853185785a742..7977a04a58c145 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -38,6 +38,8 @@ multiplicative group actions). -/ +open scoped Pointwise + /-- An `AddTorsor G P` gives a structure to the nonempty type `P`, acted on by an `AddGroup G` with a transitive and free action given by the `+ᵥ` operation and a corresponding subtraction given by the @@ -162,8 +164,6 @@ theorem vadd_eq_vadd_iff_neg_add_eq_vsub {v₁ v₂ : G} {p₁ p₂ : P} : namespace Set -open Pointwise - theorem singleton_vsub_self (p : P) : ({p} : Set P) -ᵥ {p} = {(0 : G)} := by rw [Set.singleton_vsub_singleton, vsub_self] @@ -231,6 +231,13 @@ theorem vadd_eq_vadd_iff_sub_eq_vsub {v₁ v₂ : G} {p₁ p₂ : P} : theorem vsub_sub_vsub_comm (p₁ p₂ p₃ p₄ : P) : p₁ -ᵥ p₂ - (p₃ -ᵥ p₄) = p₁ -ᵥ p₃ - (p₂ -ᵥ p₄) := by rw [← vsub_vadd_eq_vsub_sub, vsub_vadd_comm, vsub_vadd_eq_vsub_sub] +namespace Set + +@[simp] lemma vadd_set_vsub_vadd_set (v : G) (s t : Set P) : (v +ᵥ s) -ᵥ (v +ᵥ t) = s -ᵥ t := by + ext; simp [mem_vsub, mem_vadd_set] + +end Set + end comm namespace Prod diff --git a/Mathlib/Algebra/Algebra/Pi.lean b/Mathlib/Algebra/Algebra/Pi.lean index 188a138794019c..d04b77f4ffb485 100644 --- a/Mathlib/Algebra/Algebra/Pi.lean +++ b/Mathlib/Algebra/Algebra/Pi.lean @@ -48,6 +48,15 @@ theorem algebraMap_apply {_ : CommSemiring R} [_s : ∀ i, Semiring (f i)] [∀ (a : R) (i : I) : algebraMap R (∀ i, f i) a i = algebraMap R (f i) a := rfl +variable {I} in +instance (g : I → Type*) [∀ i, CommSemiring (f i)] [∀ i, Semiring (g i)] + [∀ i, Algebra (f i) (g i)] : Algebra (∀ i, f i) (∀ i, g i) where + algebraMap := Pi.ringHom fun _ ↦ (algebraMap _ _).comp (Pi.evalRingHom f _) + commutes' _ _ := funext fun _ ↦ Algebra.commutes _ _ + smul_def' _ _ := funext fun _ ↦ Algebra.smul_def _ _ + +example [∀ i, CommSemiring (f i)] : Pi.instAlgebraForall f f = Algebra.id _ := rfl + -- One could also build a `∀ i, R i`-algebra structure on `∀ i, A i`, -- when each `A i` is an `R i`-algebra, although I'm not sure that it's useful. variable {I} (R) @@ -70,6 +79,16 @@ def evalAlgHom {_ : CommSemiring R} [∀ i, Semiring (f i)] [∀ i, Algebra R (f toFun := fun f => f i commutes' := fun _ => rfl } +@[simp] +theorem algHom_evalAlgHom [CommSemiring R] [s : ∀ i, Semiring (f i)] [∀ i, Algebra R (f i)] : + algHom R f (evalAlgHom R f) = AlgHom.id R (Π i, f i) := rfl + +/-- `Pi.algHom` commutes with composition. -/ +theorem algHom_comp [CommSemiring R] [∀ i, Semiring (f i)] [∀ i, Algebra R (f i)] + {A B : Type*} [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] + (g : ∀ i, B →ₐ[R] f i) (h : A →ₐ[R] B) : + (algHom R f g).comp h = algHom R f (fun i ↦ (g i).comp h) := rfl + variable (A B : Type*) [CommSemiring R] [Semiring B] [Algebra R B] /-- `Function.const` as an `AlgHom`. The name matches `Pi.constRingHom`, `Pi.constMonoidHom`, diff --git a/Mathlib/Algebra/BigOperators/Fin.lean b/Mathlib/Algebra/BigOperators/Fin.lean index 2a445b2244dcc9..4d62dc8fa144aa 100644 --- a/Mathlib/Algebra/BigOperators/Fin.lean +++ b/Mathlib/Algebra/BigOperators/Fin.lean @@ -55,7 +55,7 @@ theorem prod_univ_zero [CommMonoid β] (f : Fin 0 → β) : ∏ i, f i = 1 := /-- A product of a function `f : Fin (n + 1) → β` over all `Fin (n + 1)` is the product of `f x`, for some `x : Fin (n + 1)` times the remaining product -/ @[to_additive "A sum of a function `f : Fin (n + 1) → β` over all `Fin (n + 1)` is the sum of -`f x`, for some `x : Fin (n + 1)` plus the remaining product"] +`f x`, for some `x : Fin (n + 1)` plus the remaining sum"] theorem prod_univ_succAbove [CommMonoid β] {n : ℕ} (f : Fin (n + 1) → β) (x : Fin (n + 1)) : ∏ i, f i = f x * ∏ i : Fin n, f (x.succAbove i) := by rw [univ_succAbove, prod_cons, Finset.prod_map _ x.succAboveEmb] @@ -64,7 +64,7 @@ theorem prod_univ_succAbove [CommMonoid β] {n : ℕ} (f : Fin (n + 1) → β) ( /-- A product of a function `f : Fin (n + 1) → β` over all `Fin (n + 1)` is the product of `f 0` plus the remaining product -/ @[to_additive "A sum of a function `f : Fin (n + 1) → β` over all `Fin (n + 1)` is the sum of -`f 0` plus the remaining product"] +`f 0` plus the remaining sum"] theorem prod_univ_succ [CommMonoid β] {n : ℕ} (f : Fin (n + 1) → β) : ∏ i, f i = f 0 * ∏ i : Fin n, f i.succ := prod_univ_succAbove f 0 @@ -215,17 +215,7 @@ theorem partialProd_left_inv {G : Type*} [Group G] (f : Fin (n + 1) → G) : @[to_additive] theorem partialProd_right_inv {G : Type*} [Group G] (f : Fin n → G) (i : Fin n) : (partialProd f (Fin.castSucc i))⁻¹ * partialProd f i.succ = f i := by - obtain ⟨i, hn⟩ := i - induction i with - | zero => simp [-Fin.succ_mk, partialProd_succ] - | succ i hi => - specialize hi (lt_trans (Nat.lt_succ_self i) hn) - simp only [Fin.coe_eq_castSucc, Fin.succ_mk, Fin.castSucc_mk] at hi ⊢ - rw [← Fin.succ_mk _ _ (lt_trans (Nat.lt_succ_self _) hn), ← Fin.succ_mk _ _ hn] - simp only [partialProd_succ, mul_inv_rev, Fin.castSucc_mk] - -- Porting note: was - -- assoc_rw [hi, inv_mul_cancel_left] - rw [← mul_assoc, mul_left_eq_self, mul_assoc, hi, inv_mul_cancel] + rw [partialProd_succ, inv_mul_cancel_left] /-- Let `(g₀, g₁, ..., gₙ)` be a tuple of elements in `Gⁿ⁺¹`. Then if `k < j`, this says `(g₀g₁...gₖ₋₁)⁻¹ * g₀g₁...gₖ = gₖ`. diff --git a/Mathlib/Algebra/BigOperators/Pi.lean b/Mathlib/Algebra/BigOperators/Pi.lean index 88511371f4a9ef..0db9fbd3746737 100644 --- a/Mathlib/Algebra/BigOperators/Pi.lean +++ b/Mathlib/Algebra/BigOperators/Pi.lean @@ -181,3 +181,22 @@ def Pi.monoidHomMulEquiv {ι : Type*} [Fintype ι] [DecidableEq ι] (M : ι → MonoidHom.mul_apply, mul_apply] end MulEquiv + +variable [Finite ι] [DecidableEq ι] {M : Type*} + +-- manually additivized to fix variable names +-- See https://github.com/leanprover-community/mathlib4/issues/11462 +lemma Pi.single_induction [AddCommMonoid M] (p : (ι → M) → Prop) (f : ι → M) + (zero : p 0) (add : ∀ f g, p f → p g → p (f + g)) + (single : ∀ i m, p (Pi.single i m)) : p f := by + cases nonempty_fintype ι + rw [← Finset.univ_sum_single f] + exact Finset.sum_induction _ _ add zero (by simp [single]) + +@[to_additive (attr := elab_as_elim) existing] +lemma Pi.mulSingle_induction [CommMonoid M] (p : (ι → M) → Prop) (f : ι → M) + (one : p 1) (mul : ∀ f g, p f → p g → p (f * g)) + (mulSingle : ∀ i m, p (Pi.mulSingle i m)) : p f := by + cases nonempty_fintype ι + rw [← Finset.univ_prod_mulSingle f] + exact Finset.prod_induction _ _ mul one (by simp [mulSingle]) diff --git a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean index 87dd2422804195..808f200d06da42 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean @@ -61,15 +61,38 @@ variable {R} in structure Hom (A B : AlgebraCat.{v} R) where private mk :: /-- The underlying algebra map. -/ - hom : A →ₐ[R] B + hom' : A →ₐ[R] B instance : Category (AlgebraCat.{v} R) where Hom A B := Hom A B id A := ⟨AlgHom.id R A⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ -instance {M N : AlgebraCat.{v} R} : CoeFun (M ⟶ N) (fun _ ↦ M → N) where - coe f := f.hom +instance : ConcreteCategory (AlgebraCat.{v} R) (· →ₐ[R] ·) where + hom := Hom.hom' + ofHom := Hom.mk + +variable {R} in +/-- Turn a morphism in `AlgebraCat` back into an `AlgHom`. -/ +abbrev Hom.hom {A B : AlgebraCat.{v} R} (f : Hom A B) := + ConcreteCategory.hom (C := AlgebraCat R) f + +variable {R} in +/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat`. -/ +abbrev ofHom {A B : Type v} [Ring A] [Ring B] [Algebra R A] [Algebra R B] (f : A →ₐ[R] B) : + of R A ⟶ of R B := + ConcreteCategory.ofHom (C := AlgebraCat R) f + +variable {R} in +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (A B : AlgebraCat.{v} R) (f : Hom A B) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ @[simp] lemma hom_id {A : AlgebraCat.{v} R} : (𝟙 A : A ⟶ A).hom = AlgHom.id R A := rfl @@ -90,11 +113,7 @@ lemma comp_apply {A B C : AlgebraCat.{v} R} (f : A ⟶ B) (g : B ⟶ C) (a : A) lemma hom_ext {A B : AlgebraCat.{v} R} {f g : A ⟶ B} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat R`. -/ -abbrev ofHom {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] - (f : X →ₐ[R] Y) : of R X ⟶ of R Y := - ⟨f⟩ - +@[simp] lemma hom_ofHom {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] (f : X →ₐ[R] Y) : (ofHom f).hom = f := rfl @@ -127,12 +146,6 @@ lemma hom_inv_apply {A B : AlgebraCat.{v} R} (e : A ≅ B) (x : B) : e.hom (e.in instance : Inhabited (AlgebraCat R) := ⟨of R R⟩ -instance : HasForget.{v} (AlgebraCat.{v} R) where - forget := - { obj := fun R => R - map := fun f => f.hom } - forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ - lemma forget_obj {A : AlgebraCat.{v} R} : (forget (AlgebraCat.{v} R)).obj A = A := rfl lemma forget_map {A B : AlgebraCat.{v} R} (f : A ⟶ B) : diff --git a/Mathlib/Algebra/Category/Grp/AB.lean b/Mathlib/Algebra/Category/Grp/AB.lean index dde16f6cf9ea29..3038d2cedbcace 100644 --- a/Mathlib/Algebra/Category/Grp/AB.lean +++ b/Mathlib/Algebra/Category/Grp/AB.lean @@ -32,14 +32,15 @@ noncomputable instance : intro x (hx : _ = _) dsimp at hx rcases Concrete.colimit_exists_rep S.X₂ x with ⟨j, y, rfl⟩ - rw [← comp_apply, colimMap_eq, colimit.ι_map, comp_apply, + rw [← CategoryTheory.comp_apply, colimMap_eq, colimit.ι_map, CategoryTheory.comp_apply, ← map_zero (by exact colimit.ι S.X₃ j : (S.X₃).obj j →+ ↑(colimit S.X₃))] at hx rcases Concrete.colimit_exists_of_rep_eq.{u, u, u} S.X₃ _ _ hx with ⟨k, e₁, e₂, hk : _ = S.X₃.map e₂ 0⟩ - rw [map_zero, ← comp_apply, ← NatTrans.naturality, comp_apply] at hk + rw [map_zero, ← CategoryTheory.comp_apply, ← NatTrans.naturality, CategoryTheory.comp_apply] + at hk rcases hS k hk with ⟨t, ht⟩ use colimit.ι S.X₁ k t - erw [← comp_apply, colimit.ι_map, comp_apply, ht] + erw [← CategoryTheory.comp_apply, colimit.ι_map, CategoryTheory.comp_apply, ht] exact colimit.w_apply S.X₂ e₁ y) noncomputable instance : diff --git a/Mathlib/Algebra/Category/Grp/LargeColimits.lean b/Mathlib/Algebra/Category/Grp/LargeColimits.lean index 9f9628925dad5b..9d0a40892ca36e 100644 --- a/Mathlib/Algebra/Category/Grp/LargeColimits.lean +++ b/Mathlib/Algebra/Category/Grp/LargeColimits.lean @@ -40,7 +40,7 @@ lemma isColimit_iff_bijective_desc [DecidableEq J] : apply ofHom_injective refine hc.hom_ext (fun j ↦ ?_) ext x - rw [comp_apply, comp_apply, ← Quot.ι_desc _ c j x] + rw [CategoryTheory.comp_apply, CategoryTheory.comp_apply, ← Quot.ι_desc _ c j x] exact ULift.down_injective (DFunLike.congr_fun eq (Quot.ι F j x)) · set c' : Cocone F := { pt := AddCommGrp.of (ULift (AddCircle (1 : ℚ))) diff --git a/Mathlib/Algebra/Category/ModuleCat/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Basic.lean index 4c63c2f7e9aeae..1692cf88738586 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Basic.lean @@ -89,20 +89,40 @@ variable {R} in structure Hom (M N : ModuleCat.{v} R) where private mk :: /-- The underlying linear map. -/ - hom : M →ₗ[R] N + hom' : M →ₗ[R] N instance moduleCategory : Category.{v, max (v+1) u} (ModuleCat.{v} R) where Hom M N := Hom M N id _ := ⟨LinearMap.id⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ -instance {M N : ModuleCat.{v} R} : CoeFun (M ⟶ N) (fun _ ↦ M → N) where - coe f := f.hom +instance : ConcreteCategory (ModuleCat.{v} R) (· →ₗ[R] ·) where + hom := Hom.hom' + ofHom := Hom.mk section variable {R} +/-- Turn a morphism in `ModuleCat` back into a `LinearMap`. -/ +abbrev Hom.hom {A B : ModuleCat.{v} R} (f : Hom A B) := + ConcreteCategory.hom (C := ModuleCat R) f + +/-- Typecheck a `LinearMap` as a morphism in `ModuleCat`. -/ +abbrev ofHom {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] + (f : X →ₗ[R] Y) : of R X ⟶ of R Y := + ConcreteCategory.ofHom (C := ModuleCat R) f + +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (A B : ModuleCat.{v} R) (f : Hom A B) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ + @[simp] lemma hom_id {M : ModuleCat.{v} R} : (𝟙 M : M ⟶ M).hom = LinearMap.id := rfl @@ -137,15 +157,9 @@ lemma hom_surjective {M N : ModuleCat.{v} R} : Function.Surjective (Hom.hom : (M ⟶ N) → (M →ₗ[R] N)) := hom_bijective.surjective -/-- Typecheck a `LinearMap` as a morphism in `ModuleCat R`. -/ -abbrev ofHom {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] - (f : X →ₗ[R] Y) : of R X ⟶ of R Y := - ⟨f⟩ - @[deprecated (since := "2024-10-06")] alias asHom := ModuleCat.ofHom -/- Doesn't need to be `@[simp]` since the `simp` tactic applies this rewrite automatically: -`ofHom` and `hom` are reducibly equal to the constructor and projection respectively. -/ +@[simp] lemma hom_ofHom {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) : (ofHom f).hom = f := rfl @@ -444,14 +458,17 @@ end variable (M N : ModuleCat.{v} R) -/-- `ModuleCat.Hom.hom` as an isomorphism of monoids. -/ -@[simps] -def endMulEquiv : End M ≃* (M →ₗ[R] M) where +/-- `ModuleCat.Hom.hom` as an isomorphism of rings. -/ +@[simps!] def endRingEquiv : End M ≃+* (M →ₗ[R] M) where toFun := ModuleCat.Hom.hom invFun := ModuleCat.ofHom map_mul' _ _ := rfl left_inv _ := rfl right_inv _ := rfl + map_add' _ _ := rfl + +/-- `ModuleCat.Hom.hom` as an isomorphism of monoids. -/ +@[deprecated (since := "2025-01-23")] alias endMulEquiv := endRingEquiv /-- The scalar multiplication on an object of `ModuleCat R` considered as a morphism of rings from `R` to the endomorphisms of the underlying abelian group. -/ @@ -537,9 +554,9 @@ a morphism between the underlying objects in `AddCommGrp` and the compatibility with the scalar multiplication. -/ @[simps] def homMk : M ⟶ N where - hom.toFun := φ - hom.map_add' _ _ := φ.map_add _ _ - hom.map_smul' r x := (congr_hom (hφ r) x).symm + hom'.toFun := φ + hom'.map_add' _ _ := φ.map_add _ _ + hom'.map_smul' r x := (congr_hom (hφ r) x).symm lemma forget₂_map_homMk : (forget₂ (ModuleCat R) AddCommGrp).map (homMk φ hφ) = φ := rfl @@ -568,7 +585,7 @@ variable {R : Type*} [CommRing R] namespace ModuleCat /-- Turn a bilinear map into a homomorphism. -/ -@[simps] +@[simps!] def ofHom₂ {M N P : ModuleCat.{u} R} (f : M →ₗ[R] N →ₗ[R] P) : M ⟶ of R (N ⟶ P) := ofHom <| homLinearEquiv.symm.toLinearMap ∘ₗ f diff --git a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean index 4387064834c7ea..cb8a2f9ed669a5 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean @@ -86,7 +86,7 @@ variable {J : Type w} (f : J → ModuleCat.{max w v} R) /-- The map from an arbitrary cone over an indexed family of abelian groups to the cartesian product of those groups. -/ -@[simps] +@[simps!] def lift (s : Fan f) : s.pt ⟶ ModuleCat.of R (∀ j, f j) := ofHom { toFun := fun x j => s.π.app ⟨j⟩ x diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index b922ed2155138c..fa3ae76a99db89 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -298,7 +298,7 @@ def map' {M1 M2 : ModuleCat.{v} R} (l : M1 ⟶ M2) : obj' f M1 ⟶ obj' f M2 := theorem map'_id {M : ModuleCat.{v} R} : map' f (𝟙 M) = 𝟙 _ := by ext x - simp [map'] + simp [map', obj'] theorem map'_comp {M₁ M₂ M₃ : ModuleCat.{v} R} (l₁₂ : M₁ ⟶ M₂) (l₂₃ : M₂ ⟶ M₃) : map' f (l₁₂ ≫ l₂₃) = map' f l₁₂ ≫ map' f l₂₃ := by @@ -415,7 +415,7 @@ instance : CoeFun (obj' f M) fun _ => S → M := /-- If `M, M'` are `R`-modules, then any `R`-linear map `g : M ⟶ M'` induces an `S`-linear map `(S →ₗ[R] M) ⟶ (S →ₗ[R] M')` defined by `h ↦ g ∘ h`-/ -@[simps] +@[simps!] def map' {M M' : ModuleCat R} (g : M ⟶ M') : obj' f M ⟶ obj' f M' := ofHom { toFun := fun h => g.hom.comp h @@ -610,7 +610,7 @@ Given `R`-module X and `S`-module Y and a map `g : (extendScalars f).obj X ⟶ Y map `S ⨂ X → Y`, there is a `X ⟶ (restrictScalars f).obj Y`, i.e. `R`-linear map `X ⟶ Y` by `x ↦ g (1 ⊗ x)`. -/ -@[simps hom_apply] +@[simps! hom_apply] def HomEquiv.toRestrictScalars {X Y} (g : (extendScalars f).obj X ⟶ Y) : X ⟶ (restrictScalars f).obj Y := -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. @@ -649,7 +649,7 @@ Given `R`-module X and `S`-module Y and a map `X ⟶ (restrictScalars f).obj Y`, `X ⟶ Y`, there is a map `(extend_scalars f).obj X ⟶ Y`, i.e `S`-linear map `S ⨂ X → Y` by `s ⊗ x ↦ s • g x`. -/ -@[simps hom_apply] +@[simps! hom_apply] def HomEquiv.fromExtendScalars {X Y} (g : X ⟶ (restrictScalars f).obj Y) : (extendScalars f).obj X ⟶ Y := by letI m1 : Module R S := Module.compHom S f; letI m2 : Module R Y := Module.compHom Y f @@ -734,7 +734,7 @@ def unit : 𝟭 (ModuleCat R) ⟶ extendScalars f ⋙ restrictScalars.{max v u /-- For any `S`-module Y, there is a natural `R`-linear map from `S ⨂ Y` to `Y` by `s ⊗ y ↦ s • y` -/ -@[simps hom_apply] +@[simps! hom_apply] def Counit.map {Y} : (restrictScalars f ⋙ extendScalars f).obj Y ⟶ Y := ofHom { toFun := @@ -873,8 +873,7 @@ lemma homEquiv_extendScalarsId (M : ModuleCat R) : (extendRestrictScalarsAdj (RingHom.id R)).homEquiv _ _ ((extendScalarsId R).hom.app M) = (restrictScalarsId R).inv.app M := by ext m - rw [extendRestrictScalarsAdj_homEquiv_apply, ← extendScalarsId_inv_app_apply] - erw [← comp_apply] + rw [extendRestrictScalarsAdj_homEquiv_apply, ← extendScalarsId_inv_app_apply, ← comp_apply] simp lemma extendScalarsId_hom_app_one_tmul (M : ModuleCat R) (m : M) : diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean index da141704256b25..33b8dc3cb3281e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean @@ -181,7 +181,8 @@ commutative rings. -/ noncomputable def relativeDifferentials' : PresheafOfModules.{u} (R ⋙ forget₂ _ _) where obj X := CommRingCat.KaehlerDifferential (φ'.app X) - map f := CommRingCat.KaehlerDifferential.map (φ'.naturality f) + -- Have to hint `g' := R.map f` below, or it gets unfolded weirdly. + map f := CommRingCat.KaehlerDifferential.map (g' := R.map f) (φ'.naturality f) -- Without `dsimp`, `ext` doesn't pick up the right lemmas. map_id _ := by dsimp; ext; simp map_comp _ _ := by dsimp; ext; simp @@ -192,7 +193,8 @@ attribute [simp] relativeDifferentials'_obj lemma relativeDifferentials'_map_d {X Y : Dᵒᵖ} (f : X ⟶ Y) (x : R.obj X) : DFunLike.coe (α := CommRingCat.KaehlerDifferential (φ'.app X)) (β := fun _ ↦ CommRingCat.KaehlerDifferential (φ'.app Y)) - ((relativeDifferentials' φ').map f).hom (CommRingCat.KaehlerDifferential.d x) = + (ModuleCat.Hom.hom (R := ↑(R.obj X)) ((relativeDifferentials' φ').map f)) + (CommRingCat.KaehlerDifferential.d x) = CommRingCat.KaehlerDifferential.d (R.map f x) := CommRingCat.KaehlerDifferential.map_d (φ'.naturality f) _ diff --git a/Mathlib/Algebra/Category/ModuleCat/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Limits.lean index 431ac7e1734a3d..81944fb1cb6489 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Limits.lean @@ -249,8 +249,7 @@ def directLimitIsColimit : IsColimit (directLimitCocone G f) where rfl fac s i := by ext - dsimp only [directLimitCocone, CategoryStruct.comp] - rw [LinearMap.comp_apply] + dsimp only [hom_comp, directLimitCocone, hom_ofHom, LinearMap.comp_apply] apply DirectLimit.lift_of uniq s m h := by have : diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean index 89bdba339ff9f5..1f38171bfc7b0d 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean @@ -244,7 +244,7 @@ lemma tensorLift_tmul (m : M₁) (n : M₂) : end -lemma tensor_ext {f g : M₁ ⊗ M₂ ⟶ M₃} (h : ∀ m n, f (m ⊗ₜ n) = g (m ⊗ₜ n)) : +lemma tensor_ext {f g : M₁ ⊗ M₂ ⟶ M₃} (h : ∀ m n, f.hom (m ⊗ₜ n) = g.hom (m ⊗ₜ n)) : f = g := hom_ext <| TensorProduct.ext (by ext; apply h) diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean index a619b7e79f8806..3fd18ca3d8336b 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean @@ -42,12 +42,12 @@ structure PresheafOfModules where /-- the restriction maps of a presheaf of modules -/ map {X Y : Cᵒᵖ} (f : X ⟶ Y) : obj X ⟶ (ModuleCat.restrictScalars (R.map f).hom).obj (obj Y) map_id (X : Cᵒᵖ) : - map (𝟙 X) = - (ModuleCat.restrictScalarsId' _ (congrArg RingCat.Hom.hom (R.map_id X))).inv.app _ := by + map (𝟙 X) = (ModuleCat.restrictScalarsId' (R.map (𝟙 X)).hom + (congrArg RingCat.Hom.hom (R.map_id X))).inv.app _ := by aesop_cat map_comp {X Y Z : Cᵒᵖ} (f : X ⟶ Y) (g : Y ⟶ Z) : map (f ≫ g) = map f ≫ (ModuleCat.restrictScalars _).map (map g) ≫ - (ModuleCat.restrictScalarsComp' _ _ _ + (ModuleCat.restrictScalarsComp' (R.map f).hom (R.map g).hom (R.map (f ≫ g)).hom (congrArg RingCat.Hom.hom <| R.map_comp f g)).inv.app _ := by aesop_cat namespace PresheafOfModules diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean index cd94fb824ae020..869be578d46d6b 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean @@ -53,7 +53,8 @@ noncomputable def tensorObj : PresheafOfModules (R ⋙ forget₂ _ _) where map_id X := ModuleCat.MonoidalCategory.tensor_ext (by intro m₁ m₂ dsimp [tensorObjMap] - simp) + simp + rfl) -- `ModuleCat.restrictScalarsId'App_inv_apply` doesn't get picked up due to type mismatch map_comp f g := ModuleCat.MonoidalCategory.tensor_ext (by intro m₁ m₂ dsimp [tensorObjMap] @@ -65,7 +66,8 @@ variable {M₁ M₂ M₃ M₄} lemma tensorObj_map_tmul {X Y : Cᵒᵖ} (f : X ⟶ Y) (m₁ : M₁.obj X) (m₂ : M₂.obj X) : DFunLike.coe (α := (M₁.obj X ⊗ M₂.obj X :)) (β := fun _ ↦ (ModuleCat.restrictScalars (R.map f).hom).obj (M₁.obj Y ⊗ M₂.obj Y)) - ((tensorObj M₁ M₂).map f).hom (m₁ ⊗ₜ[R.obj X] m₂) = M₁.map f m₁ ⊗ₜ[R.obj Y] M₂.map f m₂ := rfl + (ModuleCat.Hom.hom (R := ↑(R.obj X)) ((tensorObj M₁ M₂).map f)) (m₁ ⊗ₜ[R.obj X] m₂) = + M₁.map f m₁ ⊗ₜ[R.obj Y] M₂.map f m₂ := rfl /-- The tensor product of two morphisms of presheaves of modules. -/ @[simps] @@ -73,8 +75,11 @@ noncomputable def tensorHom (f : M₁ ⟶ M₂) (g : M₃ ⟶ M₄) : tensorObj app X := f.app X ⊗ g.app X naturality {X Y} φ := ModuleCat.MonoidalCategory.tensor_ext (fun m₁ m₃ ↦ by dsimp - rw [tensorObj_map_tmul, ModuleCat.MonoidalCategory.tensorHom_tmul, tensorObj_map_tmul, - naturality_apply, naturality_apply]) + rw [tensorObj_map_tmul] + -- Need `erw` because of the type mismatch in `map` and the tensor product. + erw [ModuleCat.MonoidalCategory.tensorHom_tmul, tensorObj_map_tmul] + rw [naturality_apply, naturality_apply] + simp) end Monoidal diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean index 62c75c356bf580..80d229fe13f6cd 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean @@ -28,15 +28,16 @@ namespace PresheafOfModules variable (F : C ⥤ D) +set_option maxHeartbeats 400000 in -- Seems to be a timeout in type checking the definition /-- The pushforward functor on presheaves of modules for a functor `F : C ⥤ D` and `R : Dᵒᵖ ⥤ RingCat`. On the underlying presheaves of abelian groups, it is induced by the precomposition with `F.op`. -/ def pushforward₀ (R : Dᵒᵖ ⥤ RingCat.{u}) : PresheafOfModules.{v} R ⥤ PresheafOfModules.{v} (F.op ⋙ R) where obj M := - { obj := fun X ↦ ModuleCat.of _ (M.obj (F.op.obj X)) - map := fun {X Y} f ↦ M.map (F.op.map f) - map_id := fun X ↦ by + { obj X := ModuleCat.of _ (M.obj (F.op.obj X)) + map {X Y} f := M.map (F.op.map f) + map_id X := by refine ModuleCat.hom_ext -- Work around an instance diamond for `restrictScalarsId'` (@LinearMap.ext _ _ _ _ _ _ _ _ (_) (_) _ _ _ (fun x => ?_)) @@ -46,7 +47,7 @@ def pushforward₀ (R : Dᵒᵖ ⥤ RingCat.{u}) : -- Work around an instance diamond for `restrictScalarsId'` (@LinearMap.ext _ _ _ _ _ _ _ _ (_) (_) _ _ _ (fun x => ?_)) exact (M.congr_map_apply (F.op.map_comp f g) x).trans (by simp) } - map {M₁ M₂} φ := { app := fun X ↦ φ.app _ } + map {M₁ M₂} φ := { app X := φ.app _ } /-- The pushforward of presheaves of modules commutes with the forgetful functor to presheaves of abelian groups. -/ diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean index 4a8d0a5b700861..a1d35a63c69edb 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean @@ -106,7 +106,7 @@ lemma isCompatible_map_smul : ((r₀.smul m₀).map (whiskerRight φ (forget _)) RingCat.comp_apply] have hb₀ : (φ.app (Opposite.op Z)) b₀ = (A.map (f₁.op ≫ g₁.op)) m := by dsimp [b₀] - erw [NatTrans.naturality_apply, hb₁, Functor.map_comp, comp_apply] + erw [NatTrans.naturality_apply, hb₁, Functor.map_comp, CategoryTheory.comp_apply] have ha₀' : (α.app (Opposite.op Z)) a₀ = (R.map (f₂.op ≫ g₂.op)) r := by rw [ha₀, ← op_comp, fac, op_comp] have hb₀' : (φ.app (Opposite.op Z)) b₀ = (A.map (f₂.op ≫ g₂.op)) m := by @@ -161,7 +161,7 @@ def SMulCandidate.mk' (S : Sieve X.unop) (hS : S ∈ J X.unop) apply A.isSeparated _ _ (J.pullback_stable f.unop hS) rintro Z g hg dsimp at hg - erw [← comp_apply, ← A.val.map_comp, ← NatTrans.naturality_apply, M₀.map_smul] + erw [← CategoryTheory.comp_apply, ← A.val.map_comp, ← NatTrans.naturality_apply, M₀.map_smul] refine (ha _ hg).trans (app_eq_of_isLocallyInjective α φ A.isSeparated _ _ _ _ ?_ ?_) · rw [← RingCat.comp_apply, NatTrans.naturality, RingCat.comp_apply, ha₀] apply (hr₀ _ hg).symm.trans @@ -298,9 +298,9 @@ lemma map_smul : apply A.isSeparated _ _ hS rintro Y f ⟨⟨r₀, (hr₀ : (α.app (Opposite.op Y)).hom r₀ = (R.val.map f.op).hom ((R.val.map π).hom r))⟩, ⟨m₀, hm₀⟩⟩ - rw [← comp_apply, ← Functor.map_comp, + rw [← CategoryTheory.comp_apply, ← Functor.map_comp, map_smul_eq α φ r m (π ≫ f.op) r₀ (by rw [hr₀, Functor.map_comp, RingCat.comp_apply]) m₀ - (by rw [hm₀, Functor.map_comp, comp_apply]), + (by rw [hm₀, Functor.map_comp, CategoryTheory.comp_apply]), map_smul_eq α φ (R.val.map π r) (A.val.map π m) f.op r₀ hr₀ m₀ hm₀] end Sheafify diff --git a/Mathlib/Algebra/Category/ModuleCat/Products.lean b/Mathlib/Algebra/Category/ModuleCat/Products.lean index 937317e6ea0a50..d4a7f6a1b8fcba 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Products.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Products.lean @@ -86,7 +86,7 @@ def coproductCoconeIsColimit : IsColimit (coproductCocone Z) where ext : 1 refine DirectSum.linearMap_ext _ fun i ↦ ?_ ext x - simpa only [LinearMap.coe_comp, Function.comp_apply, toModule_lof] using + simpa only [LinearMap.coe_comp, Function.comp_apply, hom_ofHom, toModule_lof] using congr($(h ⟨i⟩) x) variable [HasCoproduct Z] diff --git a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean index 5fb8394a9c465d..29c919d23199df 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean @@ -59,7 +59,7 @@ noncomputable def subobjectModule : Subobject M ≃o Submodule R M := ofHom (underlyingIso (ofHom N.subtype)).symm.toLinearEquiv.toLinearMap := by ext x rfl - rw [this, hom_comp, LinearEquiv.range_comp] + rw [this, hom_comp, hom_ofHom, LinearEquiv.range_comp] · exact (Submodule.range_subtype _).symm map_rel_iff' := fun {S T} => by refine ⟨fun h => ?_, fun h => mk_le_mk_of_comm (↟(Submodule.inclusion h)) rfl⟩ diff --git a/Mathlib/Algebra/Category/Ring/Basic.lean b/Mathlib/Algebra/Category/Ring/Basic.lean index eb6977db756324..8dd6aac24e41de 100644 --- a/Mathlib/Algebra/Category/Ring/Basic.lean +++ b/Mathlib/Algebra/Category/Ring/Basic.lean @@ -53,18 +53,37 @@ lemma of_carrier (R : SemiRingCat.{u}) : of R = R := rfl variable {R} in /-- The type of morphisms in `SemiRingCat`. -/ @[ext] -structure Hom (R S : SemiRingCat) where +structure Hom (R S : SemiRingCat.{u}) where private mk :: /-- The underlying ring hom. -/ - hom : R →+* S + hom' : R →+* S instance : Category SemiRingCat where Hom R S := Hom R S id R := ⟨RingHom.id R⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ -instance {R S : SemiRingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where - coe f := f.hom +instance : ConcreteCategory.{u} SemiRingCat (fun R S => R →+* S) where + hom := Hom.hom' + ofHom f := ⟨f⟩ + +/-- Turn a morphism in `SemiRingCat` back into a `RingHom`. -/ +abbrev Hom.hom {R S : SemiRingCat.{u}} (f : Hom R S) := + ConcreteCategory.hom (C := SemiRingCat) f + +/-- Typecheck a `RingHom` as a morphism in `SemiRingCat`. -/ +abbrev ofHom {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) : of R ⟶ of S := + ConcreteCategory.ofHom (C := SemiRingCat) f + +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (R S : SemiRingCat) (f : Hom R S) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ @[simp] lemma hom_id {R : SemiRingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl @@ -85,10 +104,7 @@ lemma comp_apply {R S T : SemiRingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : lemma hom_ext {R S : SemiRingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck a `RingHom` as a morphism in `SemiRingCat`. -/ -abbrev ofHom {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) : of R ⟶ of S := - ⟨f⟩ - +@[simp] lemma hom_ofHom {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) : (ofHom f).hom = f := rfl @[simp] @@ -120,12 +136,6 @@ lemma hom_inv_apply {R S : SemiRingCat} (e : R ≅ S) (s : S) : e.hom (e.inv s) instance : Inhabited SemiRingCat := ⟨of PUnit⟩ -instance : HasForget.{u} SemiRingCat where - forget := - { obj := fun R => R - map := fun f => f.hom } - forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ - /-- This unification hint helps with problems of the form `(forget ?C).obj R =?= carrier R'`. -/ unif_hint forget_obj_eq_coe (R R' : SemiRingCat) where R ≟ R' ⊢ @@ -154,8 +164,8 @@ instance hasForgetToAddCommMonCat : HasForget₂ SemiRingCat AddCommMonCat where @[simps] def _root_.RingEquiv.toSemiRingCatIso {R S : Type u} [Semiring R] [Semiring S] (e : R ≃+* S) : of R ≅ of S where - hom := ⟨e⟩ - inv := ⟨e.symm⟩ + hom := ofHom e + inv := ofHom e.symm instance forgetReflectIsos : (forget SemiRingCat).ReflectsIsomorphisms where reflects {X Y} f _ := by @@ -197,18 +207,37 @@ lemma of_carrier (R : RingCat.{u}) : of R = R := rfl variable {R} in /-- The type of morphisms in `RingCat`. -/ @[ext] -structure Hom (R S : RingCat) where +structure Hom (R S : RingCat.{u}) where private mk :: /-- The underlying ring hom. -/ - hom : R →+* S + hom' : R →+* S instance : Category RingCat where Hom R S := Hom R S id R := ⟨RingHom.id R⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ -instance {R S : RingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where - coe f := f.hom +instance : ConcreteCategory.{u} RingCat (fun R S => R →+* S) where + hom := Hom.hom' + ofHom f := ⟨f⟩ + +/-- Turn a morphism in `RingCat` back into a `RingHom`. -/ +abbrev Hom.hom {R S : RingCat.{u}} (f : Hom R S) := + ConcreteCategory.hom (C := RingCat) f + +/-- Typecheck a `RingHom` as a morphism in `RingCat`. -/ +abbrev ofHom {R S : Type u} [Ring R] [Ring S] (f : R →+* S) : of R ⟶ of S := + ConcreteCategory.ofHom (C := RingCat) f + +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (R S : RingCat) (f : Hom R S) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ @[simp] lemma hom_id {R : RingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl @@ -229,10 +258,7 @@ lemma comp_apply {R S T : RingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : lemma hom_ext {R S : RingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck a `RingHom` as a morphism in `RingCat`. -/ -abbrev ofHom {R S : Type u} [Ring R] [Ring S] (f : R →+* S) : of R ⟶ of S := - ⟨f⟩ - +@[simp] lemma hom_ofHom {R S : Type u} [Ring R] [Ring S] (f : R →+* S) : (ofHom f).hom = f := rfl @[simp] @@ -264,12 +290,6 @@ lemma hom_inv_apply {R S : RingCat} (e : R ≅ S) (s : S) : e.hom (e.inv s) = s instance : Inhabited RingCat := ⟨of PUnit⟩ -instance : HasForget.{u} RingCat where - forget := - { obj := fun R => R - map := fun f => f.hom } - forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ - /-- This unification hint helps with problems of the form `(forget ?C).obj R =?= carrier R'`. An example where this is needed is in applying @@ -302,8 +322,8 @@ instance hasForgetToAddCommGrp : HasForget₂ RingCat AddCommGrp where @[simps] def _root_.RingEquiv.toRingCatIso {R S : Type u} [Ring R] [Ring S] (e : R ≃+* S) : of R ≅ of S where - hom := ⟨e⟩ - inv := ⟨e.symm⟩ + hom := ofHom e + inv := ofHom e.symm instance forgetReflectIsos : (forget RingCat).ReflectsIsomorphisms where reflects {X Y} f _ := by @@ -345,18 +365,37 @@ lemma of_carrier (R : CommSemiRingCat.{u}) : of R = R := rfl variable {R} in /-- The type of morphisms in `CommSemiRingCat`. -/ @[ext] -structure Hom (R S : CommSemiRingCat) where +structure Hom (R S : CommSemiRingCat.{u}) where private mk :: /-- The underlying ring hom. -/ - hom : R →+* S + hom' : R →+* S instance : Category CommSemiRingCat where Hom R S := Hom R S id R := ⟨RingHom.id R⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ -instance {R S : CommSemiRingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where - coe f := f.hom +instance : ConcreteCategory.{u} CommSemiRingCat (fun R S => R →+* S) where + hom := Hom.hom' + ofHom f := ⟨f⟩ + +/-- Turn a morphism in `CommSemiRingCat` back into a `RingHom`. -/ +abbrev Hom.hom {R S : CommSemiRingCat.{u}} (f : Hom R S) := + ConcreteCategory.hom (C := CommSemiRingCat) f + +/-- Typecheck a `RingHom` as a morphism in `CommSemiRingCat`. -/ +abbrev ofHom {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) : of R ⟶ of S := + ConcreteCategory.ofHom (C := CommSemiRingCat) f + +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (R S : CommSemiRingCat) (f : Hom R S) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ @[simp] lemma hom_id {R : CommSemiRingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl @@ -377,10 +416,7 @@ lemma comp_apply {R S T : CommSemiRingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : lemma hom_ext {R S : CommSemiRingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck a `RingHom` as a morphism in `CommSemiRingCat`. -/ -abbrev ofHom {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) : of R ⟶ of S := - ⟨f⟩ - +@[simp] lemma hom_ofHom {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) : (ofHom f).hom = f := rfl @@ -413,12 +449,6 @@ lemma hom_inv_apply {R S : CommSemiRingCat} (e : R ≅ S) (s : S) : e.hom (e.inv instance : Inhabited CommSemiRingCat := ⟨of PUnit⟩ -instance : HasForget.{u} CommSemiRingCat where - forget := - { obj := fun R => R - map := fun f => f.hom } - forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ - /-- This unification hint helps with problems of the form `(forget ?C).obj R =?= carrier R'`. -/ unif_hint forget_obj_eq_coe (R R' : CommSemiRingCat) where R ≟ R' ⊢ @@ -449,8 +479,8 @@ instance hasForgetToCommMonCat : HasForget₂ CommSemiRingCat CommMonCat where def _root_.RingEquiv.toCommSemiRingCatIso {R S : Type u} [CommSemiring R] [CommSemiring S] (e : R ≃+* S) : of R ≅ of S where - hom := ⟨e⟩ - inv := ⟨e.symm⟩ + hom := ofHom e + inv := ofHom e.symm instance forgetReflectIsos : (forget CommSemiRingCat).ReflectsIsomorphisms where reflects {X Y} f _ := by @@ -492,18 +522,37 @@ lemma of_carrier (R : CommRingCat.{u}) : of R = R := rfl variable {R} in /-- The type of morphisms in `CommRingCat`. -/ @[ext] -structure Hom (R S : CommRingCat) where +structure Hom (R S : CommRingCat.{u}) where private mk :: /-- The underlying ring hom. -/ - hom : R →+* S + hom' : R →+* S instance : Category CommRingCat where Hom R S := Hom R S id R := ⟨RingHom.id R⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ -instance {R S : CommRingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where - coe f := f.hom +instance : ConcreteCategory.{u} CommRingCat (fun R S => R →+* S) where + hom := Hom.hom' + ofHom f := ⟨f⟩ + +/-- The underlying ring hom. -/ +abbrev Hom.hom {R S : CommRingCat.{u}} (f : Hom R S) := + ConcreteCategory.hom (C := CommRingCat) f + +/-- Typecheck a `RingHom` as a morphism in `CommRingCat`. -/ +abbrev ofHom {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) : of R ⟶ of S := + ConcreteCategory.ofHom (C := CommRingCat) f + +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (R S : CommRingCat) (f : Hom R S) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ @[simp] lemma hom_id {R : CommRingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl @@ -524,10 +573,7 @@ lemma comp_apply {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : lemma hom_ext {R S : CommRingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck a `RingHom` as a morphism in `CommRingCat`. -/ -abbrev ofHom {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) : of R ⟶ of S := - ⟨f⟩ - +@[simp] lemma hom_ofHom {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) : (ofHom f).hom = f := rfl @@ -560,12 +606,6 @@ lemma hom_inv_apply {R S : CommRingCat} (e : R ≅ S) (s : S) : e.hom (e.inv s) instance : Inhabited CommRingCat := ⟨of PUnit⟩ -instance : HasForget.{u} CommRingCat where - forget := - { obj := fun R => R - map := fun f => f.hom } - forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ - lemma forget_obj {R : CommRingCat} : (forget CommRingCat).obj R = R := rfl /-- This unification hint helps with problems of the form `(forget ?C).obj R =?= carrier R'`. @@ -606,8 +646,8 @@ instance hasForgetToAddCommMonCat : HasForget₂ CommRingCat CommSemiRingCat whe def _root_.RingEquiv.toCommRingCatIso {R S : Type u} [CommRing R] [CommRing S] (e : R ≃+* S) : of R ≅ of S where - hom := ⟨e⟩ - inv := ⟨e.symm⟩ + hom := ofHom e + inv := ofHom e.symm instance forgetReflectIsos : (forget CommRingCat).ReflectsIsomorphisms where reflects {X Y} f _ := by @@ -669,10 +709,10 @@ abbrev CommRingCatMax.{u1, u2} := CommRingCat.{max u1 u2} lemma RingCat.forget_map_apply {R S : RingCat} (f : R ⟶ S) (x : (CategoryTheory.forget RingCat).obj R) : - @DFunLike.coe _ _ _ HasForget.instFunLike f x = f x := + (forget _).map f x = f x := rfl lemma CommRingCat.forget_map_apply {R S : CommRingCat} (f : R ⟶ S) (x : (CategoryTheory.forget CommRingCat).obj R) : - @DFunLike.coe _ _ _ HasForget.instFunLike f x = f x := + (forget _).map f x = f x := rfl diff --git a/Mathlib/Algebra/Category/Ring/Epi.lean b/Mathlib/Algebra/Category/Ring/Epi.lean index ae3808ac124703..4ba86a9d7d7c35 100644 --- a/Mathlib/Algebra/Category/Ring/Epi.lean +++ b/Mathlib/Algebra/Category/Ring/Epi.lean @@ -23,8 +23,7 @@ lemma CommRingCat.epi_iff_tmul_eq_tmul {R S : Type u} [CommRing R] [CommRing S] ∀ s : S, s ⊗ₜ[R] 1 = 1 ⊗ₜ s := by constructor · intro H - #adaptation_note - /-- After https://github.com/leanprover/lean4/pull/6024 + #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 we need to add `(R := R) (A := S)` in the next line to deal with unification issues. -/ have := H.1 (CommRingCat.ofHom <| Algebra.TensorProduct.includeLeftRingHom (R := R)) (CommRingCat.ofHom <| (Algebra.TensorProduct.includeRight (R := R) (A := S)).toRingHom) diff --git a/Mathlib/Algebra/DualNumber.lean b/Mathlib/Algebra/DualNumber.lean index 6e2ef4a79c866f..2eca2386d15669 100644 --- a/Mathlib/Algebra/DualNumber.lean +++ b/Mathlib/Algebra/DualNumber.lean @@ -149,36 +149,24 @@ theorem lift_apply_apply (fe : {_fe : (A →ₐ[R] B) × B // _}) (a : A[ε]) : @[simp] theorem coe_lift_symm_apply (F : A[ε] →ₐ[R] B) : (lift.symm F).val = (F.comp (inlAlgHom _ _ _), F ε) := rfl -#adaptation_note -/-- -The new unused variable linter in -https://github.com/leanprover/lean4/pull/5338 -flags `{fe : (A →ₐ[R] B) × B // _}`. --/ +#adaptation_note /-- https://github.com/leanprover/lean4/pull/5338 +The new unused variable linter flags `{fe : (A →ₐ[R] B) × B // _}`. -/ set_option linter.unusedVariables false in /-- When applied to `inl`, `DualNumber.lift` applies the map `f : A →ₐ[R] B`. -/ @[simp] theorem lift_apply_inl (fe : {fe : (A →ₐ[R] B) × B // _}) (a : A) : lift fe (inl a : A[ε]) = fe.val.1 a := by rw [lift_apply_apply, fst_inl, snd_inl, map_zero, zero_mul, add_zero] -#adaptation_note -/-- -The new unused variable linter in -https://github.com/leanprover/lean4/pull/5338 -flags `{fe : (A →ₐ[R] B) × B // _}`. --/ +#adaptation_note /-- https://github.com/leanprover/lean4/pull/5338 +The new unused variable linter flags `{fe : (A →ₐ[R] B) × B // _}`. -/ set_option linter.unusedVariables false in /-- Scaling on the left is sent by `DualNumber.lift` to multiplication on the left -/ @[simp] theorem lift_smul (fe : {fe : (A →ₐ[R] B) × B // _}) (a : A) (ad : A[ε]) : lift fe (a • ad) = fe.val.1 a * lift fe ad := by rw [← inl_mul_eq_smul, map_mul, lift_apply_inl] -#adaptation_note -/-- -The new unused variable linter in -https://github.com/leanprover/lean4/pull/5338 -flags `{fe : (A →ₐ[R] B) × B // _}`. --/ +#adaptation_note /-- https://github.com/leanprover/lean4/pull/5338 +The new unused variable linter flags `{fe : (A →ₐ[R] B) × B // _}`. -/ set_option linter.unusedVariables false in /-- Scaling on the right is sent by `DualNumber.lift` to multiplication on the right -/ @[simp] theorem lift_op_smul (fe : {fe : (A →ₐ[R] B) × B // _}) (a : A) (ad : A[ε]) : diff --git a/Mathlib/Algebra/Free.lean b/Mathlib/Algebra/Free.lean index c0663264e177d9..0fbfaa9678903a 100644 --- a/Mathlib/Algebra/Free.lean +++ b/Mathlib/Algebra/Free.lean @@ -80,8 +80,8 @@ theorem hom_ext {β : Type v} [Mul β] {f g : FreeMagma α →ₙ* β} (h : f end FreeMagma -#adaptation_note /-- around nightly-2024-02-25, we need to write `mul x y` in the second pattern, -instead of `x * y`. --/ +#adaptation_note /-- nightly-2024-02-25 +we need to write `mul x y` in the second pattern, instead of `x * y`. --/ /-- Lifts a function `α → β` to a magma homomorphism `FreeMagma α → β` given a magma `β`. -/ def FreeMagma.liftAux {α : Type u} {β : Type v} [Mul β] (f : α → β) : FreeMagma α → β | FreeMagma.of x => f x @@ -185,8 +185,8 @@ end Category end FreeMagma -#adaptation_note /-- around nightly-2024-02-25, we need to write `mul x y` in the second pattern, - instead of `x * y`. -/ +#adaptation_note /-- nightly-2024-02-25 +we need to write `mul x y` in the second pattern, instead of `x * y`. -/ /-- `FreeMagma` is traversable. -/ protected def FreeMagma.traverse {m : Type u → Type u} [Applicative m] {α β : Type u} (F : α → m β) : FreeMagma α → m (FreeMagma β) @@ -262,8 +262,8 @@ end Category end FreeMagma -#adaptation_note /-- around nightly-2024-02-25, we need to write `mul x y` in the second pattern, -instead of `x * y`. -/ +#adaptation_note /-- nightly-2024-02-25 +we need to write `mul x y` in the second pattern, instead of `x * y`. -/ /-- Representation of an element of a free magma. -/ protected def FreeMagma.repr {α : Type u} [Repr α] : FreeMagma α → Lean.Format | FreeMagma.of x => repr x @@ -279,8 +279,8 @@ attribute [to_additive existing] FreeMagma.repr @[to_additive] instance {α : Type u} [Repr α] : Repr (FreeMagma α) := ⟨fun o _ => FreeMagma.repr o⟩ -#adaptation_note /-- around nightly-2024-02-25, we need to write `mul x y` in the second pattern, -instead of `x * y`. -/ +#adaptation_note /-- nightly-2024-02-25 +we need to write `mul x y` in the second pattern, instead of `x * y`. -/ /-- Length of an element of a free magma. -/ def FreeMagma.length {α : Type u} : FreeMagma α → ℕ | FreeMagma.of _x => 1 diff --git a/Mathlib/Algebra/FreeMonoid/Basic.lean b/Mathlib/Algebra/FreeMonoid/Basic.lean index 7340041857cfc8..596248f9831176 100644 --- a/Mathlib/Algebra/FreeMonoid/Basic.lean +++ b/Mathlib/Algebra/FreeMonoid/Basic.lean @@ -372,7 +372,7 @@ theorem map_apply_map_symm_eq {x : FreeMonoid β} (e : α ≃ β) : instance uniqueUnits : Unique (FreeMonoid α)ˣ where uniq u := Units.ext <| toList.injective <| have : toList u.val ++ toList u.inv = [] := DFunLike.congr_arg toList u.val_inv - (List.append_eq_nil.mp this).1 + (List.append_eq_nil_iff.mp this).1 @[to_additive (attr := simp)] theorem map_surjective {f : α → β} : Function.Surjective (map f) ↔ Function.Surjective f := by diff --git a/Mathlib/Algebra/GCDMonoid/Finset.lean b/Mathlib/Algebra/GCDMonoid/Finset.lean index 89d7b6ca4f3efc..245b1c3f9ec1b0 100644 --- a/Mathlib/Algebra/GCDMonoid/Finset.lean +++ b/Mathlib/Algebra/GCDMonoid/Finset.lean @@ -76,8 +76,7 @@ theorem lcm_insert [DecidableEq β] {b : β} : theorem lcm_singleton {b : β} : ({b} : Finset β).lcm f = normalize (f b) := Multiset.lcm_singleton --- Porting note: Priority changed for `simpNF` -@[simp 1100] +@[local simp] -- This will later be provable by other `simp` lemmas. theorem normalize_lcm : normalize (s.lcm f) = s.lcm f := by simp [lcm_def] theorem lcm_union [DecidableEq β] : (s₁ ∪ s₂).lcm f = GCDMonoid.lcm (s₁.lcm f) (s₂.lcm f) := @@ -149,8 +148,7 @@ theorem gcd_insert [DecidableEq β] {b : β} : theorem gcd_singleton {b : β} : ({b} : Finset β).gcd f = normalize (f b) := Multiset.gcd_singleton --- Porting note: Priority changed for `simpNF` -@[simp 1100] +@[local simp] -- This will later be provable by other `simp` lemmas. theorem normalize_gcd : normalize (s.gcd f) = s.gcd f := by simp [gcd_def] theorem gcd_union [DecidableEq β] : (s₁ ∪ s₂).gcd f = GCDMonoid.gcd (s₁.gcd f) (s₂.gcd f) := diff --git a/Mathlib/Algebra/GCDMonoid/Multiset.lean b/Mathlib/Algebra/GCDMonoid/Multiset.lean index a50a1633317558..cbe9338671d3a8 100644 --- a/Mathlib/Algebra/GCDMonoid/Multiset.lean +++ b/Mathlib/Algebra/GCDMonoid/Multiset.lean @@ -63,11 +63,7 @@ theorem dvd_lcm {s : Multiset α} {a : α} (h : a ∈ s) : a ∣ s.lcm := theorem lcm_mono {s₁ s₂ : Multiset α} (h : s₁ ⊆ s₂) : s₁.lcm ∣ s₂.lcm := lcm_dvd.2 fun _ hb ↦ dvd_lcm (h hb) -/- Porting note: Following `Algebra.GCDMonoid.Basic`'s version of `normalize_gcd`, I'm giving -this lower priority to avoid linter complaints about simp-normal form -/ -/- Porting note: Mathport seems to be replacing `Multiset.induction_on s $` with -`(Multiset.induction_on s)`, when it should be `Multiset.induction_on s <|`. -/ -@[simp 1100] +@[simp] theorem normalize_lcm (s : Multiset α) : normalize s.lcm = s.lcm := Multiset.induction_on s (by simp) fun a s _ ↦ by simp @@ -139,9 +135,7 @@ theorem gcd_dvd {s : Multiset α} {a : α} (h : a ∈ s) : s.gcd ∣ a := theorem gcd_mono {s₁ s₂ : Multiset α} (h : s₁ ⊆ s₂) : s₂.gcd ∣ s₁.gcd := dvd_gcd.2 fun _ hb ↦ gcd_dvd (h hb) -/- Porting note: Following `Algebra.GCDMonoid.Basic`'s version of `normalize_gcd`, I'm giving -this lower priority to avoid linter complaints about simp-normal form -/ -@[simp 1100] +@[simp] theorem normalize_gcd (s : Multiset α) : normalize s.gcd = s.gcd := Multiset.induction_on s (by simp) fun a s _ ↦ by simp diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index 2ffc6457f0b71d..08c5149911c7e0 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -706,6 +706,15 @@ def MonoidHom.id (M : Type*) [MulOneClass M] : M →* M where map_one' := rfl map_mul' _ _ := rfl +@[to_additive (attr := simp)] +lemma OneHom.coe_id {M : Type*} [One M] : (OneHom.id M : M → M) = _root_.id := rfl + +@[to_additive (attr := simp)] +lemma MulHom.coe_id {M : Type*} [Mul M] : (MulHom.id M : M → M) = _root_.id := rfl + +@[to_additive (attr := simp)] +lemma MonoidHom.coe_id {M : Type*} [MulOneClass M] : (MonoidHom.id M : M → M) = _root_.id := rfl + /-- Composition of `OneHom`s as a `OneHom`. -/ @[to_additive "Composition of `ZeroHom`s as a `ZeroHom`."] def OneHom.comp [One M] [One N] [One P] (hnp : OneHom N P) (hmn : OneHom M N) : OneHom M P where diff --git a/Mathlib/Algebra/Group/Pi/Lemmas.lean b/Mathlib/Algebra/Group/Pi/Lemmas.lean index 697d42405d54c7..b73a9e5ea8db5b 100644 --- a/Mathlib/Algebra/Group/Pi/Lemmas.lean +++ b/Mathlib/Algebra/Group/Pi/Lemmas.lean @@ -338,7 +338,7 @@ theorem Pi.semiconjBy_iff {x y z : ∀ i, f i} : SemiconjBy x y z ↔ ∀ i, SemiconjBy (x i) (y i) (z i) := funext_iff @[to_additive] -theorem Commute.pi {x y : ∀ i, f i} (h : ∀ i, Commute (x i) (y i)) : Commute x y := .pi h +theorem Commute.pi {x y : ∀ i, f i} (h : ∀ i, Commute (x i) (y i)) : Commute x y := SemiconjBy.pi h @[to_additive] theorem Pi.commute_iff {x y : ∀ i, f i} : Commute x y ↔ ∀ i, Commute (x i) (y i) := semiconjBy_iff diff --git a/Mathlib/Algebra/Group/Prod.lean b/Mathlib/Algebra/Group/Prod.lean index aa462854b70861..f3a94d3d34aa6e 100644 --- a/Mathlib/Algebra/Group/Prod.lean +++ b/Mathlib/Algebra/Group/Prod.lean @@ -260,7 +260,7 @@ theorem Prod.semiconjBy_iff {x y z : M × N} : @[to_additive AddCommute.prod] theorem Commute.prod {x y : M × N} (hm : Commute x.1 y.1) (hn : Commute x.2 y.2) : Commute x y := - .prod hm hn + SemiconjBy.prod hm hn @[to_additive] theorem Prod.commute_iff {x y : M × N} : diff --git a/Mathlib/Algebra/Group/Submonoid/MulAction.lean b/Mathlib/Algebra/Group/Submonoid/MulAction.lean new file mode 100644 index 00000000000000..7c37bc72c9f08b --- /dev/null +++ b/Mathlib/Algebra/Group/Submonoid/MulAction.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2021 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.Algebra.Group.Submonoid.Defs +import Mathlib.Algebra.Group.Action.Defs + +/-! +# Actions by `Submonoid`s + +These instances transfer the action by an element `m : M` of a monoid `M` written as `m • a` onto +the action by an element `s : S` of a submonoid `S : Submonoid M` such that `s • a = (s : M) • a`. + +These instances work particularly well in conjunction with `Monoid.toMulAction`, enabling +`s • m` as an alias for `↑s * m`. +-/ + +namespace Submonoid + +variable {M' : Type*} {α β : Type*} + +section MulOneClass + +variable [MulOneClass M'] + +@[to_additive] +instance smul [SMul M' α] (S : Submonoid M') : SMul S α := + SMul.comp _ S.subtype + +@[to_additive] +instance smulCommClass_left [SMul M' β] [SMul α β] [SMulCommClass M' α β] + (S : Submonoid M') : SMulCommClass S α β := + ⟨fun a _ _ => smul_comm (a : M') _ _⟩ + +@[to_additive] +instance smulCommClass_right [SMul α β] [SMul M' β] [SMulCommClass α M' β] + (S : Submonoid M') : SMulCommClass α S β := + ⟨fun a s => smul_comm a (s : M')⟩ + +/-- Note that this provides `IsScalarTower S M' M'` which is needed by `SMulMulAssoc`. -/ +instance isScalarTower [SMul α β] [SMul M' α] [SMul M' β] [IsScalarTower M' α β] + (S : Submonoid M') : + IsScalarTower S α β := + ⟨fun a => smul_assoc (a : M')⟩ + +section SMul +variable [SMul M' α] {S : Submonoid M'} + +@[to_additive] lemma smul_def (g : S) (a : α) : g • a = (g : M') • a := rfl + +@[to_additive (attr := simp)] +lemma mk_smul (g : M') (hg : g ∈ S) (a : α) : (⟨g, hg⟩ : S) • a = g • a := rfl + +end SMul +end MulOneClass + +variable [Monoid M'] + +/-- The action by a submonoid is the action by the underlying monoid. -/ +@[to_additive + "The additive action by an `AddSubmonoid` is the action by the underlying `AddMonoid`. "] +instance mulAction [MulAction M' α] (S : Submonoid M') : MulAction S α where + one_smul := one_smul M' + mul_smul m₁ m₂ := mul_smul (m₁ : M') m₂ + +example {S : Submonoid M'} : IsScalarTower S M' M' := by infer_instance + +end Submonoid diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 9037ade39062b4..a997c5650f854b 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Group.Action.Faithful import Mathlib.Algebra.Group.Nat.Defs import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Submonoid.Basic +import Mathlib.Algebra.Group.Submonoid.MulAction import Mathlib.Algebra.Group.TypeTags.Basic /-! @@ -973,75 +974,10 @@ theorem Submonoid.equivMapOfInjective_coe_mulEquiv (e : M ≃* N) : ext rfl -section Actions - -/-! ### Actions by `Submonoid`s - -These instances transfer the action by an element `m : M` of a monoid `M` written as `m • a` onto -the action by an element `s : S` of a submonoid `S : Submonoid M` such that `s • a = (s : M) • a`. - -These instances work particularly well in conjunction with `Monoid.toMulAction`, enabling -`s • m` as an alias for `↑s * m`. --/ - - -namespace Submonoid - -variable {M' : Type*} {α β : Type*} - -section MulOneClass - -variable [MulOneClass M'] - -@[to_additive] -instance smul [SMul M' α] (S : Submonoid M') : SMul S α := - SMul.comp _ S.subtype - -@[to_additive] -instance smulCommClass_left [SMul M' β] [SMul α β] [SMulCommClass M' α β] - (S : Submonoid M') : SMulCommClass S α β := - ⟨fun a _ _ => smul_comm (a : M') _ _⟩ - -@[to_additive] -instance smulCommClass_right [SMul α β] [SMul M' β] [SMulCommClass α M' β] - (S : Submonoid M') : SMulCommClass α S β := - ⟨fun a s => smul_comm a (s : M')⟩ - -/-- Note that this provides `IsScalarTower S M' M'` which is needed by `SMulMulAssoc`. -/ -instance isScalarTower [SMul α β] [SMul M' α] [SMul M' β] [IsScalarTower M' α β] - (S : Submonoid M') : - IsScalarTower S α β := - ⟨fun a => smul_assoc (a : M')⟩ - -section SMul -variable [SMul M' α] {S : Submonoid M'} - -@[to_additive] lemma smul_def (g : S) (a : α) : g • a = (g : M') • a := rfl - -@[to_additive (attr := simp)] -lemma mk_smul (g : M') (hg : g ∈ S) (a : α) : (⟨g, hg⟩ : S) • a = g • a := rfl - -instance faithfulSMul [FaithfulSMul M' α] : FaithfulSMul S α := +instance Submonoid.faithfulSMul {M' α : Type*} [MulOneClass M'] [SMul M' α] {S : Submonoid M'} + [FaithfulSMul M' α] : FaithfulSMul S α := ⟨fun h => Subtype.ext <| eq_of_smul_eq_smul h⟩ -end SMul -end MulOneClass - -variable [Monoid M'] - -/-- The action by a submonoid is the action by the underlying monoid. -/ -@[to_additive - "The additive action by an `AddSubmonoid` is the action by the underlying `AddMonoid`. "] -instance mulAction [MulAction M' α] (S : Submonoid M') : MulAction S α := - MulAction.compHom _ S.subtype - -example {S : Submonoid M'} : IsScalarTower S M' M' := by infer_instance - - -end Submonoid - -end Actions - section Units namespace Submonoid diff --git a/Mathlib/Algebra/GroupWithZero/Hom.lean b/Mathlib/Algebra/GroupWithZero/Hom.lean index f30ace0a86cbce..5f4976f69371a0 100644 --- a/Mathlib/Algebra/GroupWithZero/Hom.lean +++ b/Mathlib/Algebra/GroupWithZero/Hom.lean @@ -40,13 +40,6 @@ open Function namespace NeZero variable {F α β : Type*} [Zero α] [Zero β] [FunLike F α β] [ZeroHomClass F α β] {a : α} -#adaptation_note -/-- -We name `neZero` so it can be used as a named argument, -but since https://github.com/leanprover/lean4/pull/5338, this is considered unused, -so we need to disable the linter. --/ -set_option linter.unusedVariables false in lemma of_map (f : F) [neZero : NeZero (f a)] : NeZero a := ⟨fun h ↦ ne (f a) <| by rw [h]; exact ZeroHomClass.map_zero f⟩ diff --git a/Mathlib/Algebra/GroupWithZero/WithZero.lean b/Mathlib/Algebra/GroupWithZero/WithZero.lean index 43ea0028206c0d..9ad955134228e7 100644 --- a/Mathlib/Algebra/GroupWithZero/WithZero.lean +++ b/Mathlib/Algebra/GroupWithZero/WithZero.lean @@ -266,6 +266,10 @@ def unitsWithZeroEquiv : (WithZero α)ˣ ≃* α where right_inv _ := rfl map_mul' _ _ := coe_inj.mp <| by simp only [Units.val_mul, coe_unzero, coe_mul] +theorem coe_unitsWithZeroEquiv_eq_units_val (γ : (WithZero α)ˣ) : + ↑(unitsWithZeroEquiv γ) = γ.val := by + simp only [WithZero.unitsWithZeroEquiv, MulEquiv.coe_mk, Equiv.coe_fn_mk, WithZero.coe_unzero] + /-- Any group with zero is isomorphic to adjoining `0` to the units of itself. -/ def withZeroUnitsEquiv {G : Type*} [GroupWithZero G] [DecidablePred (fun a : G ↦ a = 0)] : diff --git a/Mathlib/Algebra/Homology/Additive.lean b/Mathlib/Algebra/Homology/Additive.lean index 86d5e417b4051e..a5c9c7078277ab 100644 --- a/Mathlib/Algebra/Homology/Additive.lean +++ b/Mathlib/Algebra/Homology/Additive.lean @@ -170,7 +170,7 @@ theorem NatTrans.mapHomologicalComplex_comp (c : ComplexShape ι) {F G H : W₁ NatTrans.mapHomologicalComplex α c ≫ NatTrans.mapHomologicalComplex β c := by aesop_cat -@[reassoc (attr := simp 1100)] +@[reassoc] theorem NatTrans.mapHomologicalComplex_naturality {c : ComplexShape ι} {F G : W₁ ⥤ W₂} [F.PreservesZeroMorphisms] [G.PreservesZeroMorphisms] (α : F ⟶ G) {C D : HomologicalComplex W₁ c} (f : C ⟶ D) : diff --git a/Mathlib/Algebra/Homology/ConcreteCategory.lean b/Mathlib/Algebra/Homology/ConcreteCategory.lean index f510e2876a6f3b..9c7f9f7bae473c 100644 --- a/Mathlib/Algebra/Homology/ConcreteCategory.lean +++ b/Mathlib/Algebra/Homology/ConcreteCategory.lean @@ -80,7 +80,7 @@ lemma δ_apply (x₃ : (forget₂ C Ab).obj (S.X₃.X i)) AddMonoidHom.zero_apply])) := by refine hS.δ_apply' i j hij _ ((forget₂ C Ab).map (S.X₂.pOpcycles i) x₂) _ ?_ ?_ · rw [← forget₂_comp_apply, ← forget₂_comp_apply, - HomologicalComplex.p_opcyclesMap, Functor.map_comp, comp_apply, + HomologicalComplex.p_opcyclesMap, Functor.map_comp, CategoryTheory.comp_apply, HomologicalComplex.homology_π_ι, forget₂_comp_apply, hx₂, HomologicalComplex.i_cyclesMk] · apply (Preadditive.mono_iff_injective (S.X₂.iCycles j)).1 inferInstance conv_lhs => diff --git a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean index 908ef5c957030b..4aae573cb738de 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean @@ -160,12 +160,12 @@ lemma mk₀_hom [HasDerivedCategory.{w'} C] (f : X ⟶ Y) : (mk₀ f).hom = ShiftedHom.mk₀ _ (by simp) ((singleFunctor C 0).map f) := by apply SmallShiftedHom.equiv_mk₀ -@[simp 1100] +@[simp] lemma mk₀_comp_mk₀ (f : X ⟶ Y) (g : Y ⟶ Z) : (mk₀ f).comp (mk₀ g) (zero_add 0) = mk₀ (f ≫ g) := by letI := HasDerivedCategory.standard C; ext; simp -@[simp 1100] +@[simp] lemma mk₀_comp_mk₀_assoc (f : X ⟶ Y) (g : Y ⟶ Z) {n : ℕ} (α : Ext Z T n) : (mk₀ f).comp ((mk₀ g).comp α (zero_add n)) (zero_add n) = (mk₀ (f ≫ g)).comp α (zero_add n) := by diff --git a/Mathlib/Algebra/Homology/Embedding/TruncGEHomology.lean b/Mathlib/Algebra/Homology/Embedding/TruncGEHomology.lean index d39f0d57c60523..33aa50866daca8 100644 --- a/Mathlib/Algebra/Homology/Embedding/TruncGEHomology.lean +++ b/Mathlib/Algebra/Homology/Embedding/TruncGEHomology.lean @@ -10,8 +10,8 @@ import Mathlib.Algebra.Homology.QuasiIso /-! # The homology of a canonical truncation -Given an embedding of complex shapes `e : Embedding c c'`, we shall -relate the homology of `K : HomologicalComplex C c'` and of +Given an embedding of complex shapes `e : Embedding c c'`, +we relate the homology of `K : HomologicalComplex C c'` and of `K.truncGE e : HomologicalComplex C c'`. The main result is that `K.πTruncGE e : K ⟶ K.truncGE e` induces a @@ -50,7 +50,8 @@ lemma hasHomology_of_not_mem_boundary (hj : ¬ e.BoundaryGE j) : (K.truncGE' e).HasHomology j := hasHomology_sc'_of_not_mem_boundary K e _ j _ rfl rfl hj -lemma quasiIsoAt_restrictionToTruncGE'_f (hj : ¬ e.BoundaryGE j) +/-- `K.restrictionToTruncGE' e` is a quasi-isomorphism in degrees that are not at the boundary. -/ +lemma quasiIsoAt_restrictionToTruncGE' (hj : ¬ e.BoundaryGE j) [(K.restriction e).HasHomology j] [(K.truncGE' e).HasHomology j] : QuasiIsoAt (K.restrictionToTruncGE' e) j := by rw [quasiIsoAt_iff] diff --git a/Mathlib/Algebra/Homology/Embedding/TruncLEHomology.lean b/Mathlib/Algebra/Homology/Embedding/TruncLEHomology.lean new file mode 100644 index 00000000000000..4e0871f9545d5d --- /dev/null +++ b/Mathlib/Algebra/Homology/Embedding/TruncLEHomology.lean @@ -0,0 +1,79 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Homology.Embedding.TruncGEHomology +import Mathlib.Algebra.Homology.Embedding.TruncLE + +/-! # The homology of a canonical truncation + +Given an embedding of complex shapes `e : Embedding c c'`, +we relate the homology of `K : HomologicalComplex C c'` and of +`K.truncLE e : HomologicalComplex C c'`. + +The main result is that `K.ιTruncLE e : K.truncLE e ⟶ K` induces a +quasi-isomorphism in degree `e.f i` for all `i`. (Note that the complex +`K.truncLE e` is exact in degrees that are not in the image of `e.f`.) + +All the results are obtained by dualising the results in the file `Embedding.TruncGEHomology`. + +-/ + +open CategoryTheory Category Limits + +namespace HomologicalComplex + +variable {ι ι' : Type*} {c : ComplexShape ι} {c' : ComplexShape ι'} + {C : Type*} [Category C] [HasZeroMorphisms C] + (K L : HomologicalComplex C c') (φ : K ⟶ L) (e : c.Embedding c') [e.IsTruncLE] + [∀ i', K.HasHomology i'] [∀ i', L.HasHomology i'] + +namespace truncLE' + +/-- `K.truncLE'ToRestriction e` is a quasi-isomorphism in degrees that are not at the boundary. -/ +lemma quasiIsoAt_truncLE'ToRestriction (j : ι) (hj : ¬ e.BoundaryLE j) + [(K.restriction e).HasHomology j] [(K.truncLE' e).HasHomology j] : + QuasiIsoAt (K.truncLE'ToRestriction e) j := by + dsimp only [truncLE'ToRestriction] + have : (K.op.restriction e.op).HasHomology j := + inferInstanceAs ((K.restriction e).op.HasHomology j) + rw [quasiIsoAt_unopFunctor_map_iff] + exact truncGE'.quasiIsoAt_restrictionToTruncGE' K.op e.op j (by simpa) + +instance truncLE'_hasHomology (i : ι) : (K.truncLE' e).HasHomology i := + inferInstanceAs ((K.op.truncGE' e.op).unop.HasHomology i) + +end truncLE' + +variable [HasZeroObject C] + +instance (i' : ι') : (K.truncLE e).HasHomology i' := + inferInstanceAs ((K.op.truncGE e.op).unop.HasHomology i') + +lemma quasiIsoAt_ιTruncLE {j : ι} {j' : ι'} (hj' : e.f j = j') : + QuasiIsoAt (K.ιTruncLE e) j' := by + have := K.op.quasiIsoAt_πTruncGE e.op hj' + exact inferInstanceAs (QuasiIsoAt ((unopFunctor _ _ ).map (K.op.πTruncGE e.op).op) j') + +instance (i : ι) : QuasiIsoAt (K.ιTruncLE e) (e.f i) := K.quasiIsoAt_ιTruncLE e rfl + +lemma quasiIso_ιTruncLE_iff_isSupported : + QuasiIso (K.ιTruncLE e) ↔ K.IsSupported e := by + rw [← quasiIso_opFunctor_map_iff, ← isSupported_op_iff] + exact K.op.quasiIso_πTruncGE_iff_isSupported e.op + +lemma acyclic_truncLE_iff_isSupportedOutside : + (K.truncLE e).Acyclic ↔ K.IsSupportedOutside e := by + rw [← acyclic_op_iff, ← isSupportedOutside_op_iff] + exact K.op.acyclic_truncGE_iff_isSupportedOutside e.op + +variable {K L} + +lemma quasiIso_truncLEMap_iff : + QuasiIso (truncLEMap φ e) ↔ ∀ (i : ι) (i' : ι') (_ : e.f i = i'), QuasiIsoAt φ i' := by + rw [← quasiIso_opFunctor_map_iff] + simp only [← quasiIsoAt_opFunctor_map_iff φ] + apply quasiIso_truncGEMap_iff + +end HomologicalComplex diff --git a/Mathlib/Algebra/Homology/HomologicalComplex.lean b/Mathlib/Algebra/Homology/HomologicalComplex.lean index 9fc057de9025d3..f09f0593e1d0f5 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplex.lean @@ -469,7 +469,7 @@ theorem dFrom_comp_xNextIso {i j : ι} (r : c.Rel i j) : theorem dFrom_comp_xNextIsoSelf {i : ι} (h : ¬c.Rel i (c.next i)) : C.dFrom i ≫ (C.xNextIsoSelf h).hom = 0 := by simp [h] -@[simp 1100] +-- This is not a simp lemma; the LHS already simplifies. theorem dTo_comp_dFrom (j : ι) : C.dTo j ≫ C.dFrom j = 0 := C.d_comp_d _ _ _ @@ -552,14 +552,12 @@ theorem next_eq (f : Hom C₁ C₂) {i j : ι} (w : c.Rel i j) : theorem comm_from (f : Hom C₁ C₂) (i : ι) : f.f i ≫ C₂.dFrom i = C₁.dFrom i ≫ f.next i := f.comm _ _ -attribute [simp 1100] comm_from_assoc attribute [simp] comm_from_apply @[reassoc, elementwise] theorem comm_to (f : Hom C₁ C₂) (j : ι) : f.prev j ≫ C₂.dTo j = C₁.dTo j ≫ f.f j := f.comm _ _ -attribute [simp 1100] comm_to_assoc attribute [simp] comm_to_apply /-- A morphism of chain complexes diff --git a/Mathlib/Algebra/Homology/Homotopy.lean b/Mathlib/Algebra/Homology/Homotopy.lean index bb13f21e29877f..8e8923988a0cf6 100644 --- a/Mathlib/Algebra/Homology/Homotopy.lean +++ b/Mathlib/Algebra/Homology/Homotopy.lean @@ -51,12 +51,12 @@ lemma dNext_eq_zero (f : ∀ i j, C.X i ⟶ D.X j) (i : ι) (hi : ¬ c.Rel i (c. dsimp [dNext] rw [shape _ _ _ hi, zero_comp] -@[simp 1100] +-- This is not a simp lemma; the LHS already simplifies. theorem dNext_comp_left (f : C ⟶ D) (g : ∀ i j, D.X i ⟶ E.X j) (i : ι) : (dNext i fun i j => f.f i ≫ g i j) = f.f i ≫ dNext i g := (f.comm_assoc _ _ _).symm -@[simp 1100] +-- This is not a simp lemma; the LHS already simplifies. theorem dNext_comp_right (f : ∀ i j, C.X i ⟶ D.X j) (g : D ⟶ E) (i : ι) : (dNext i fun i j => f i j ≫ g.f j) = dNext i f ≫ g.f i := (assoc _ _ _).symm @@ -85,12 +85,12 @@ theorem prevD_eq (f : ∀ i j, C.X i ⟶ D.X j) {j j' : ι} (w : c.Rel j' j) : obtain rfl := c.prev_eq' w rfl -@[simp 1100] +-- This is not a simp lemma; the LHS already simplifies. theorem prevD_comp_left (f : C ⟶ D) (g : ∀ i j, D.X i ⟶ E.X j) (j : ι) : (prevD j fun i j => f.f i ≫ g i j) = f.f j ≫ prevD j g := assoc _ _ _ -@[simp 1100] +-- This is not a simp lemma; the LHS already simplifies. theorem prevD_comp_right (f : ∀ i j, C.X i ⟶ D.X j) (g : D ⟶ E) (j : ι) : (prevD j fun i j => f i j ≫ g.f j) = prevD j f ≫ g.f j := by dsimp [prevD] @@ -426,21 +426,21 @@ section MkInductive variable {P Q : ChainComplex V ℕ} -@[simp 1100] +-- This is not a simp lemma; the LHS already simplifies. theorem prevD_chainComplex (f : ∀ i j, P.X i ⟶ Q.X j) (j : ℕ) : prevD j f = f j (j + 1) ≫ Q.d _ _ := by dsimp [prevD] have : (ComplexShape.down ℕ).prev j = j + 1 := ChainComplex.prev ℕ j congr 2 -@[simp 1100] +-- This is not a simp lemma; the LHS already simplifies. theorem dNext_succ_chainComplex (f : ∀ i j, P.X i ⟶ Q.X j) (i : ℕ) : dNext (i + 1) f = P.d _ _ ≫ f i (i + 1) := by dsimp [dNext] have : (ComplexShape.down ℕ).next (i + 1) = i := ChainComplex.next_nat_succ _ congr 2 -@[simp 1100] +-- This is not a simp lemma; the LHS already simplifies. theorem dNext_zero_chainComplex (f : ∀ i j, P.X i ⟶ Q.X j) : dNext 0 f = 0 := by dsimp [dNext] rw [P.shape, zero_comp] @@ -556,21 +556,21 @@ section MkCoinductive variable {P Q : CochainComplex V ℕ} -@[simp 1100] +-- This is not a simp lemma; the LHS already simplifies. theorem dNext_cochainComplex (f : ∀ i j, P.X i ⟶ Q.X j) (j : ℕ) : dNext j f = P.d _ _ ≫ f (j + 1) j := by dsimp [dNext] have : (ComplexShape.up ℕ).next j = j + 1 := CochainComplex.next ℕ j congr 2 -@[simp 1100] +-- This is not a simp lemma; the LHS already simplifies. theorem prevD_succ_cochainComplex (f : ∀ i j, P.X i ⟶ Q.X j) (i : ℕ) : prevD (i + 1) f = f (i + 1) _ ≫ Q.d i (i + 1) := by dsimp [prevD] have : (ComplexShape.up ℕ).prev (i + 1) = i := CochainComplex.prev_nat_succ i congr 2 -@[simp 1100] +-- This is not a simp lemma; the LHS already simplifies. theorem prevD_zero_cochainComplex (f : ∀ i j, P.X i ⟶ Q.X j) : prevD 0 f = 0 := by dsimp [prevD] rw [Q.shape, comp_zero] diff --git a/Mathlib/Algebra/Homology/HomotopyCategory.lean b/Mathlib/Algebra/Homology/HomotopyCategory.lean index 93400ddd35d940..b4ecba4a3831a2 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory.lean @@ -121,7 +121,6 @@ def homotopyOutMap {C D : HomologicalComplex V c} (f : C ⟶ D) : apply homotopyOfEq simp -@[simp 1100] theorem quotient_map_out_comp_out {C D E : HomotopyCategory V c} (f : C ⟶ D) (g : D ⟶ E) : (quotient V c).map (Quot.out f ≫ Quot.out g) = f ≫ g := by simp diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/MappingCone.lean b/Mathlib/Algebra/Homology/HomotopyCategory/MappingCone.lean index 57a474997c34a4..aebf5de588e5a8 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/MappingCone.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/MappingCone.lean @@ -241,10 +241,10 @@ lemma inl_v_d (i j k : ℤ) (hij : i + (-1) = j) (hik : k + (-1) = i) : rw [homotopyCofiber.inlX_d φ j i k (by dsimp; omega) (by dsimp; omega)] abel -@[reassoc (attr := simp 1100)] +@[reassoc] lemma inr_f_d (n₁ n₂ : ℤ) : (inr φ).f n₁ ≫ (mappingCone φ).d n₁ n₂ = G.d n₁ n₂ ≫ (inr φ).f n₂ := by - apply Hom.comm + simp @[reassoc] lemma d_fst_v (i j k : ℤ) (hij : i + 1 = j) (hjk : j + 1 = k) : diff --git a/Mathlib/Algebra/Homology/Opposite.lean b/Mathlib/Algebra/Homology/Opposite.lean index 5bbd8fee09f160..cb42b8f38d5d44 100644 --- a/Mathlib/Algebra/Homology/Opposite.lean +++ b/Mathlib/Algebra/Homology/Opposite.lean @@ -7,6 +7,7 @@ import Mathlib.CategoryTheory.Abelian.Opposite import Mathlib.Algebra.Homology.Additive import Mathlib.Algebra.Homology.ImageToKernel import Mathlib.Algebra.Homology.ShortComplex.HomologicalComplex +import Mathlib.Algebra.Homology.QuasiIso /-! # Opposite categories of complexes @@ -221,6 +222,86 @@ instance (K : HomologicalComplex Vᵒᵖ c) (i : ι) [K.HasHomology i] : variable {V c} +@[simp] +lemma quasiIsoAt_opFunctor_map_iff + {K L : HomologicalComplex V c} (φ : K ⟶ L) (i : ι) + [K.HasHomology i] [L.HasHomology i] : + QuasiIsoAt ((opFunctor _ _).map φ.op) i ↔ QuasiIsoAt φ i := by + simp only [quasiIsoAt_iff] + exact ShortComplex.quasiIso_opMap_iff ((shortComplexFunctor V c i).map φ) + +@[simp] +lemma quasiIsoAt_unopFunctor_map_iff + {K L : HomologicalComplex Vᵒᵖ c} (φ : K ⟶ L) (i : ι) + [K.HasHomology i] [L.HasHomology i] : + QuasiIsoAt ((unopFunctor _ _).map φ.op) i ↔ QuasiIsoAt φ i := by + rw [← quasiIsoAt_opFunctor_map_iff] + rfl + +instance {K L : HomologicalComplex V c} (φ : K ⟶ L) (i : ι) + [K.HasHomology i] [L.HasHomology i] [QuasiIsoAt φ i] : + QuasiIsoAt ((opFunctor _ _).map φ.op) i := by + rw [quasiIsoAt_opFunctor_map_iff] + infer_instance + +instance {K L : HomologicalComplex Vᵒᵖ c} (φ : K ⟶ L) (i : ι) + [K.HasHomology i] [L.HasHomology i] [QuasiIsoAt φ i] : + QuasiIsoAt ((unopFunctor _ _).map φ.op) i := by + rw [quasiIsoAt_unopFunctor_map_iff] + infer_instance + +@[simp] +lemma quasiIso_opFunctor_map_iff + {K L : HomologicalComplex V c} (φ : K ⟶ L) + [∀ i, K.HasHomology i] [∀ i, L.HasHomology i] : + QuasiIso ((opFunctor _ _).map φ.op) ↔ QuasiIso φ := by + simp only [quasiIso_iff, quasiIsoAt_opFunctor_map_iff] + +@[simp] +lemma quasiIso_unopFunctor_map_iff + {K L : HomologicalComplex Vᵒᵖ c} (φ : K ⟶ L) + [∀ i, K.HasHomology i] [∀ i, L.HasHomology i] : + QuasiIso ((unopFunctor _ _).map φ.op) ↔ QuasiIso φ := by + simp only [quasiIso_iff, quasiIsoAt_unopFunctor_map_iff] + +instance {K L : HomologicalComplex V c} (φ : K ⟶ L) + [∀ i, K.HasHomology i] [∀ i, L.HasHomology i] [QuasiIso φ] : + QuasiIso ((opFunctor _ _).map φ.op) := by + rw [quasiIso_opFunctor_map_iff] + infer_instance + +instance {K L : HomologicalComplex Vᵒᵖ c} (φ : K ⟶ L) + [∀ i, K.HasHomology i] [∀ i, L.HasHomology i] [QuasiIso φ] : + QuasiIso ((unopFunctor _ _).map φ.op) := by + rw [quasiIso_unopFunctor_map_iff] + infer_instance + +lemma ExactAt.op {K : HomologicalComplex V c} {i : ι} (h : K.ExactAt i) : + K.op.ExactAt i := + ShortComplex.Exact.op h + +lemma ExactAt.unop {K : HomologicalComplex Vᵒᵖ c} {i : ι} (h : K.ExactAt i) : + K.unop.ExactAt i := + ShortComplex.Exact.unop h + +@[simp] +lemma exactAt_op_iff (K : HomologicalComplex V c) {i : ι} : + K.op.ExactAt i ↔ K.ExactAt i := + ⟨fun h ↦ h.unop, fun h ↦ h.op⟩ + +lemma Acyclic.op {K : HomologicalComplex V c} (h : K.Acyclic) : + K.op.Acyclic := + fun i ↦ (h i).op + +lemma Acyclic.unop {K : HomologicalComplex Vᵒᵖ c} (h : K.Acyclic) : + K.unop.Acyclic := + fun i ↦ (h i).unop + +@[simp] +lemma acyclic_op_iff (K : HomologicalComplex V c) : + K.op.Acyclic ↔ K.Acyclic := + ⟨fun h ↦ h.unop, fun h ↦ h.op⟩ + /-- If `K` is a homological complex, then the homology of `K.op` identifies to the opposite of the homology of `K`. -/ def homologyOp (K : HomologicalComplex V c) (i : ι) [K.HasHomology i] : diff --git a/Mathlib/Algebra/Homology/ShortComplex/Ab.lean b/Mathlib/Algebra/Homology/ShortComplex/Ab.lean index 393b22f24e512c..f1b0f96f3b7a08 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Ab.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Ab.lean @@ -37,7 +37,7 @@ variable (S : ShortComplex Ab.{u}) @[simp] lemma ab_zero_apply (x : S.X₁) : S.g (S.f x) = 0 := by - rw [← comp_apply, S.zero] + rw [← CategoryTheory.comp_apply, S.zero] rfl /-- The canonical additive morphism `S.X₁ →+ AddMonoidHom.ker S.g` induced by `S.f`. -/ @@ -81,7 +81,7 @@ noncomputable def abCyclesIso : S.cycles ≅ AddCommGrp.of (AddMonoidHom.ker S.g lemma abCyclesIso_inv_apply_iCycles (x : AddMonoidHom.ker S.g) : S.iCycles (S.abCyclesIso.inv x) = x := by dsimp only [abCyclesIso] - rw [← comp_apply, S.abLeftHomologyData.cyclesIso_inv_comp_iCycles] + rw [← CategoryTheory.comp_apply, S.abLeftHomologyData.cyclesIso_inv_comp_iCycles] rfl /-- Given a short complex `S` of abelian groups, this is the isomorphism between @@ -129,7 +129,7 @@ lemma ab_exact_iff_range_eq_ker : S.Exact ↔ S.f.range = S.g.ker := by · intro h refine le_antisymm ?_ h rintro _ ⟨x₁, rfl⟩ - rw [AddMonoidHom.mem_ker, ← comp_apply, S.zero] + rw [AddMonoidHom.mem_ker, ← CategoryTheory.comp_apply, S.zero] rfl · intro h rw [h] diff --git a/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean b/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean index ad19544a15033f..806aabbe913404 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean @@ -50,10 +50,10 @@ lemma abelianImageToKernel_comp_kernel_ι : instance : Mono S.abelianImageToKernel := mono_of_mono_fac S.abelianImageToKernel_comp_kernel_ι -@[reassoc (attr := simp 1100)] +@[reassoc] lemma abelianImageToKernel_comp_kernel_ι_comp_cokernel_π : S.abelianImageToKernel ≫ kernel.ι S.g ≫ cokernel.π S.f = 0 := by - simp only [abelianImageToKernel_comp_kernel_ι_assoc, kernel.condition] + simp /-- `Abelian.image S.f` is the kernel of `kernel.ι S.g ≫ cokernel.π S.f` -/ noncomputable def abelianImageToKernelIsKernel : diff --git a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean index 92d14ea2564f50..610dbd89fe4145 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean @@ -32,7 +32,7 @@ lemma ShortComplex.zero_apply [Limits.HasZeroMorphisms C] [(forget₂ C Ab).PreservesZeroMorphisms] (S : ShortComplex C) (x : (forget₂ C Ab).obj S.X₁) : ((forget₂ C Ab).map S.g) (((forget₂ C Ab).map S.f) x) = 0 := by - rw [← comp_apply, ← Functor.map_comp, S.zero, Functor.map_zero] + rw [← CategoryTheory.comp_apply, ← Functor.map_comp, S.zero, Functor.map_zero] rfl section preadditive @@ -112,8 +112,8 @@ lemma i_cyclesMk [S.HasHomology] (x₂ : (forget₂ C Ab).obj S.X₂) (hx₂ : ((forget₂ C Ab).map S.g) x₂ = 0) : (forget₂ C Ab).map S.iCycles (S.cyclesMk x₂ hx₂) = x₂ := by dsimp [cyclesMk] - erw [← comp_apply, S.mapCyclesIso_hom_iCycles (forget₂ C Ab), - ← comp_apply, abCyclesIso_inv_apply_iCycles ] + erw [← CategoryTheory.comp_apply, S.mapCyclesIso_hom_iCycles (forget₂ C Ab), + ← CategoryTheory.comp_apply, abCyclesIso_inv_apply_iCycles ] end ShortComplex @@ -172,13 +172,13 @@ lemma δ_apply' (x₃ : (forget₂ C Ab).obj D.L₀.X₃) · refine ((congr_hom (e.hom.naturality D.L₁.g) x₂).symm.trans ?_).trans (congr_hom (e.hom.naturality D.v₀₁.τ₃) x₃) dsimp - rw [comp_apply, comp_apply] + rw [CategoryTheory.comp_apply, CategoryTheory.comp_apply] erw [h₂] rfl · refine ((congr_hom (e.hom.naturality D.L₂.f) x₁).symm.trans ?_).trans (congr_hom (e.hom.naturality D.v₁₂.τ₂) x₂) dsimp - rw [comp_apply, comp_apply] + rw [CategoryTheory.comp_apply, CategoryTheory.comp_apply] erw [h₁] rfl diff --git a/Mathlib/Algebra/Homology/Single.lean b/Mathlib/Algebra/Homology/Single.lean index b6dc763c87a954..16d9d5943f7329 100644 --- a/Mathlib/Algebra/Homology/Single.lean +++ b/Mathlib/Algebra/Homology/Single.lean @@ -44,8 +44,8 @@ noncomputable def single (j : ι) : V ⥤ HomologicalComplex V c where split_ifs with h · subst h simp - · #adaptation_note /-- after nightly-2024-03-07, the previous sensible proof - `rw [if_neg h]; simp` fails with "motive not type correct". + · #adaptation_note /-- nightly-2024-03-07 + the previous sensible proof `rw [if_neg h]; simp` fails with "motive not type correct". The following is horrible. -/ convert (id_zero (C := V)).symm all_goals simp [if_neg h] diff --git a/Mathlib/Algebra/Lie/CartanExists.lean b/Mathlib/Algebra/Lie/CartanExists.lean index 4e3acbb133d815..4574da2c841879 100644 --- a/Mathlib/Algebra/Lie/CartanExists.lean +++ b/Mathlib/Algebra/Lie/CartanExists.lean @@ -119,8 +119,6 @@ variable {K L : Type*} [Field K] [LieRing L] [LieAlgebra K L] [Module.Finite K L open Module LieSubalgebra LieSubmodule Polynomial Cardinal LieModule engel_isBot_of_isMin -#adaptation_note /-- otherwise there is a spurious warning on `contrapose!` below. -/ -set_option linter.unusedVariables false in /-- Let `L` be a Lie algebra of dimension `n` over a field `K` with at least `n` elements. Given a Lie subalgebra `U` of `L`, and an element `x ∈ U` such that `U ≤ engel K x`. Suppose that `engel K x` is minimal amongst the Engel subalgebras `engel K y` for `y ∈ U`. diff --git a/Mathlib/Algebra/Lie/Semisimple/Basic.lean b/Mathlib/Algebra/Lie/Semisimple/Basic.lean index 0697232c790f3f..28001afd616102 100644 --- a/Mathlib/Algebra/Lie/Semisimple/Basic.lean +++ b/Mathlib/Algebra/Lie/Semisimple/Basic.lean @@ -212,8 +212,7 @@ lemma finitelyAtomistic : ∀ s : Finset (LieIdeal R L), ↑s ⊆ {I : LieIdeal set K := s'.sup id suffices I ≤ K by obtain ⟨t, hts', htI⟩ := finitelyAtomistic s' hs'S I this - #adaptation_note - /-- Prior to https://github.com/leanprover/lean4/pull/6024 + #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 we could write `hts'.trans hs'.subset` instead of `Finset.Subset.trans hts' hs'.subset` in the next line. -/ exact ⟨t, Finset.Subset.trans hts' hs'.subset, htI⟩ diff --git a/Mathlib/Algebra/LinearRecurrence.lean b/Mathlib/Algebra/LinearRecurrence.lean index a07f5edb8399f6..b9fe57a324105f 100644 --- a/Mathlib/Algebra/LinearRecurrence.lean +++ b/Mathlib/Algebra/LinearRecurrence.lean @@ -92,11 +92,11 @@ theorem eq_mk_of_is_sol_of_eq_init {u : ℕ → α} {init : Fin E.order → α} rw [mkSol] split_ifs with h' · exact mod_cast heq ⟨n, h'⟩ - rw [← tsub_add_cancel_of_le (le_of_not_lt h'), h (n - E.order)] - congr with k - have : n - E.order + k < n := by omega - rw [eq_mk_of_is_sol_of_eq_init h heq (n - E.order + k)] - simp + · rw [← tsub_add_cancel_of_le (le_of_not_lt h'), h (n - E.order)] + congr with k + have : n - E.order + k < n := by omega + rw [eq_mk_of_is_sol_of_eq_init h heq (n - E.order + k)] + simp /-- If `u` is a solution to `E` and `init` designates its first `E.order` values, then `u = E.mkSol init`. This proves that `E.mkSol init` is the only solution diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index b0228d7304c795..ddbbd94475a751 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -482,6 +482,41 @@ theorem symm_neg : (neg R : M ≃ₗ[R] M).symm = neg R := end Neg +section Semiring + +open LinearMap + +variable {M₂₁ M₂₂ : Type*} [Semiring R] [AddCommMonoid M₁] [AddCommMonoid M₂] + [AddCommMonoid M₂₁] [AddCommMonoid M₂₂] [Module R M₁] [Module R M₂] [Module R M₂₁] [Module R M₂₂] + +/-- A linear isomorphism between the domains and codomains of two spaces of linear maps gives a +additive isomorphism between the two function spaces. + +See also `LinearEquiv.arrowCongr` for the linear version of this isomorphism. -/ +def arrowCongrAddEquiv (e₁ : M₁ ≃ₗ[R] M₂) (e₂ : M₂₁ ≃ₗ[R] M₂₂) : + (M₁ →ₗ[R] M₂₁) ≃+ (M₂ →ₗ[R] M₂₂) where + toFun f := e₂.comp (f.comp e₁.symm.toLinearMap) + invFun f := e₂.symm.comp (f.comp e₁.toLinearMap) + left_inv f := by + ext x + simp only [symm_apply_apply, Function.comp_apply, coe_comp, coe_coe] + right_inv f := by + ext x + simp only [Function.comp_apply, apply_symm_apply, coe_comp, coe_coe] + map_add' f g := by + ext x + simp only [map_add, add_apply, Function.comp_apply, coe_comp, coe_coe] + +/-- If `M` and `M₂` are linearly isomorphic then the endomorphism rings of `M` and `M₂` +are isomorphic. + +See `LinearEquiv.conj` for the linear version of this isomorphism. -/ +def conjRingEquiv (e : M₁ ≃ₗ[R] M₂) : Module.End R M₁ ≃+* Module.End R M₂ where + __ := arrowCongrAddEquiv e e + map_mul' _ _ := by ext; simp [arrowCongrAddEquiv] + +end Semiring + section CommSemiring variable [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] @@ -494,24 +529,15 @@ def smulOfUnit (a : Rˣ) : M ≃ₗ[R] M := DistribMulAction.toLinearEquiv R M a /-- A linear isomorphism between the domains and codomains of two spaces of linear maps gives a -linear isomorphism between the two function spaces. -/ +linear isomorphism between the two function spaces. + +See `LinearEquiv.arrowCongrAddEquiv` for the additive version of this isomorphism that works +over a not necessarily commutative semiring. -/ def arrowCongr {R M₁ M₂ M₂₁ M₂₂ : Sort _} [CommSemiring R] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₂₁] [AddCommMonoid M₂₂] [Module R M₁] [Module R M₂] [Module R M₂₁] [Module R M₂₂] (e₁ : M₁ ≃ₗ[R] M₂) (e₂ : M₂₁ ≃ₗ[R] M₂₂) : (M₁ →ₗ[R] M₂₁) ≃ₗ[R] M₂ →ₗ[R] M₂₂ where - toFun := fun f : M₁ →ₗ[R] M₂₁ ↦ (e₂ : M₂₁ →ₗ[R] M₂₂).comp <| f.comp (e₁.symm : M₂ →ₗ[R] M₁) - invFun f := (e₂.symm : M₂₂ →ₗ[R] M₂₁).comp <| f.comp (e₁ : M₁ →ₗ[R] M₂) - left_inv f := by - ext x - simp only [symm_apply_apply, Function.comp_apply, coe_comp, coe_coe] - right_inv f := by - ext x - simp only [Function.comp_apply, apply_symm_apply, coe_comp, coe_coe] - map_add' f g := by - ext x - simp only [map_add, add_apply, Function.comp_apply, coe_comp, coe_coe] - map_smul' c f := by - ext x - simp only [smul_apply, Function.comp_apply, coe_comp, map_smulₛₗ e₂, coe_coe] + __ := arrowCongrAddEquiv e₁ e₂ + map_smul' c f := by ext; simp [arrowCongrAddEquiv] @[simp] theorem arrowCongr_apply {R M₁ M₂ M₂₁ M₂₂ : Sort _} [CommSemiring R] [AddCommMonoid M₁] @@ -547,7 +573,11 @@ def congrRight (f : M₂ ≃ₗ[R] M₃) : (M →ₗ[R] M₂) ≃ₗ[R] M →ₗ arrowCongr (LinearEquiv.refl R M) f /-- If `M` and `M₂` are linearly isomorphic then the two spaces of linear maps from `M` and `M₂` to -themselves are linearly isomorphic. -/ +themselves are linearly isomorphic. + +See `LinearEquiv.conjRingEquiv` for the isomorphism between endomorphism rings, +which works over a not necessarily commutative semiring. -/ +-- TODO: upgrade to AlgEquiv (but this file currently cannot import AlgEquiv) def conj (e : M ≃ₗ[R] M₂) : Module.End R M ≃ₗ[R] Module.End R M₂ := arrowCongr e e @@ -582,12 +612,8 @@ isomorphism between `M₂ →ₗ[R] M` and `M₃ →ₗ[R] M`, if `M` is both an `S`-module and their actions commute. -/ def congrLeft {R} (S) [Semiring R] [Semiring S] [Module R M₂] [Module R M₃] [Module R M] [Module S M] [SMulCommClass R S M] (e : M₂ ≃ₗ[R] M₃) : (M₂ →ₗ[R] M) ≃ₗ[S] (M₃ →ₗ[R] M) where - toFun f := f.comp e.symm.toLinearMap - invFun f := f.comp e.toLinearMap - map_add' _ _ := rfl + __ := e.arrowCongrAddEquiv (.refl ..) map_smul' _ _ := rfl - left_inv f := by dsimp only; apply DFunLike.ext; exact (congr_arg f <| e.left_inv ·) - right_inv f := by dsimp only; apply DFunLike.ext; exact (congr_arg f <| e.right_inv ·) end CommSemiring diff --git a/Mathlib/Algebra/Module/LinearMap/Polynomial.lean b/Mathlib/Algebra/Module/LinearMap/Polynomial.lean index 5637e7a41bdf33..a5f9f7bb8180f3 100644 --- a/Mathlib/Algebra/Module/LinearMap/Polynomial.lean +++ b/Mathlib/Algebra/Module/LinearMap/Polynomial.lean @@ -253,11 +253,8 @@ lemma polyCharpolyAux_baseChange (A : Type*) [CommRing A] [Algebra R A] : simp only [RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, map_X, bind₁_X_right] classical rw [toMvPolynomial_comp _ (basis A (Basis.end bₘ)), ← toMvPolynomial_baseChange] - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/4119 we either need to specify the `M₂` argument, - or use `set_option maxSynthPendingDepth 2 in`. - -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 + we either need to specify the `M₂` argument, or use `set_option maxSynthPendingDepth 2 in`. -/ suffices toMvPolynomial (M₂ := (Module.End A (TensorProduct R A M))) (basis A bₘ.end) (basis A bₘ).end (tensorProduct R A M M) ij = X ij by rw [this, bind₁_X_right] diff --git a/Mathlib/Algebra/Module/Submodule/Defs.lean b/Mathlib/Algebra/Module/Submodule/Defs.lean index 55faba74532642..50ffade02ba270 100644 --- a/Mathlib/Algebra/Module/Submodule/Defs.lean +++ b/Mathlib/Algebra/Module/Submodule/Defs.lean @@ -5,6 +5,7 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Group.Subgroup.Defs import Mathlib.GroupTheory.GroupAction.SubMulAction +import Mathlib.Algebra.Group.Submonoid.Basic /-! diff --git a/Mathlib/Algebra/Order/CauSeq/Basic.lean b/Mathlib/Algebra/Order/CauSeq/Basic.lean index fb76b07c2edb27..b7366261beeeae 100644 --- a/Mathlib/Algebra/Order/CauSeq/Basic.lean +++ b/Mathlib/Algebra/Order/CauSeq/Basic.lean @@ -404,7 +404,7 @@ theorem zero_limZero : LimZero (0 : CauSeq β abv) theorem const_limZero {x : β} : LimZero (const x) ↔ x = 0 := ⟨fun H => (abv_eq_zero abv).1 <| - (eq_of_le_of_forall_le_of_dense (abv_nonneg abv _)) fun _ ε0 => + (eq_of_le_of_forall_lt_imp_le_of_dense (abv_nonneg abv _)) fun _ ε0 => let ⟨_, hi⟩ := H _ ε0 le_of_lt <| hi _ le_rfl, fun e => e.symm ▸ zero_limZero⟩ diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index d2b9446b43e012..5d80dad8dfcfd6 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -414,7 +414,7 @@ theorem le_iff_forall_one_lt_le_mul₀ {α : Type*} [LinearOrderedSemifield α] obtain rfl|hb := hb.eq_or_lt · simp_rw [zero_mul] at h exact h 2 one_lt_two - refine le_of_forall_le_of_dense fun x hbx => ?_ + refine le_of_forall_gt_imp_ge_of_dense fun x hbx => ?_ convert h (x / b) ((one_lt_div hb).mpr hbx) rw [mul_div_cancel₀ _ hb.ne'] @@ -755,14 +755,14 @@ private lemma exists_mul_right_lt₀ {a b c : α} (hc : a * b < c) : ∃ b' > b, simp_rw [mul_comm a] at hc ⊢; exact exists_mul_left_lt₀ hc lemma le_mul_of_forall_lt₀ {a b c : α} (h : ∀ a' > a, ∀ b' > b, c ≤ a' * b') : c ≤ a * b := by - refine le_of_forall_le_of_dense fun d hd ↦ ?_ + refine le_of_forall_gt_imp_ge_of_dense fun d hd ↦ ?_ obtain ⟨a', ha', hd⟩ := exists_mul_left_lt₀ hd obtain ⟨b', hb', hd⟩ := exists_mul_right_lt₀ hd exact (h a' ha' b' hb').trans hd.le lemma mul_le_of_forall_lt_of_nonneg {a b c : α} (ha : 0 ≤ a) (hc : 0 ≤ c) (h : ∀ a' ≥ 0, a' < a → ∀ b' ≥ 0, b' < b → a' * b' ≤ c) : a * b ≤ c := by - refine le_of_forall_ge_of_dense fun d d_ab ↦ ?_ + refine le_of_forall_lt_imp_le_of_dense fun d d_ab ↦ ?_ rcases lt_or_le d 0 with hd | hd · exact hd.le.trans hc obtain ⟨a', ha', d_ab⟩ := exists_lt_mul_left_of_nonneg ha hd d_ab diff --git a/Mathlib/Algebra/Order/Group/DenselyOrdered.lean b/Mathlib/Algebra/Order/Group/DenselyOrdered.lean index a019c1739800a6..a1bd1983375e5f 100644 --- a/Mathlib/Algebra/Order/Group/DenselyOrdered.lean +++ b/Mathlib/Algebra/Order/Group/DenselyOrdered.lean @@ -68,7 +68,7 @@ private lemma exists_mul_right_lt [CommGroup α] [LT α] [DenselyOrdered α] lemma le_mul_of_forall_lt [CommGroup α] [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] [DenselyOrdered α] {a b c : α} (h : ∀ a' > a, ∀ b' > b, c ≤ a' * b') : c ≤ a * b := by - refine le_of_forall_le_of_dense fun d hd ↦ ?_ + refine le_of_forall_gt_imp_ge_of_dense fun d hd ↦ ?_ obtain ⟨a', ha', hd⟩ := exists_mul_left_lt hd obtain ⟨b', hb', hd⟩ := exists_mul_right_lt hd exact (h a' ha' b' hb').trans hd.le @@ -77,7 +77,7 @@ lemma le_mul_of_forall_lt [CommGroup α] [LinearOrder α] [CovariantClass α α lemma mul_le_of_forall_lt [CommGroup α] [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] [DenselyOrdered α] {a b c : α} (h : ∀ a' < a, ∀ b' < b, a' * b' ≤ c) : a * b ≤ c := by - refine le_of_forall_ge_of_dense fun d hd ↦ ?_ + refine le_of_forall_lt_imp_le_of_dense fun d hd ↦ ?_ obtain ⟨a', ha', hd⟩ := exists_lt_mul_left hd obtain ⟨b', hb', hd⟩ := exists_lt_mul_right hd exact hd.le.trans (h a' ha' b' hb') diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean index 1bd716addf3274..c2e25592e58cd5 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean @@ -69,7 +69,7 @@ variable [LinearOrder α] [DenselyOrdered α] [Monoid α] [ExistsMulOfLE α] @[to_additive] theorem le_of_forall_one_lt_le_mul (h : ∀ ε : α, 1 < ε → a ≤ b * ε) : a ≤ b := - le_of_forall_le_of_dense fun x hxb => by + le_of_forall_gt_imp_ge_of_dense fun x hxb => by obtain ⟨ε, rfl⟩ := exists_mul_of_le hxb.le exact h _ (one_lt_of_lt_mul_right hxb) diff --git a/Mathlib/Algebra/Polynomial/AlgebraMap.lean b/Mathlib/Algebra/Polynomial/AlgebraMap.lean index 40e975a43fbfe1..2129497e1006c0 100644 --- a/Mathlib/Algebra/Polynomial/AlgebraMap.lean +++ b/Mathlib/Algebra/Polynomial/AlgebraMap.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker -/ import Mathlib.Algebra.Algebra.Pi +import Mathlib.Algebra.Algebra.Prod import Mathlib.Algebra.Algebra.Subalgebra.Basic import Mathlib.Algebra.Algebra.Tower import Mathlib.Algebra.MonoidAlgebra.Basic @@ -368,6 +369,38 @@ theorem aeval_algebraMap_apply_eq_algebraMap_eval (x : R) (p : R[X]) : aeval (algebraMap R A x) p = algebraMap R A (p.eval x) := aeval_algHom_apply (Algebra.ofId R A) x p +/-- Polynomial evaluation on a pair is a product of the evaluations on the components. -/ +theorem aeval_prod (x : A × B) : aeval (R := R) x = (aeval x.1).prod (aeval x.2) := + aeval_algHom (.fst R A B) x ▸ aeval_algHom (.snd R A B) x ▸ + (aeval x).prod_comp (.fst R A B) (.snd R A B) + +/-- Polynomial evaluation on a pair is a pair of evaluations. -/ +theorem aeval_prod_apply (x : A × B) (p : Polynomial R) : + p.aeval x = (p.aeval x.1, p.aeval x.2) := by simp [aeval_prod] + +section Pi + +variable {I : Type*} {A : I → Type*} [∀ i, Semiring (A i)] [∀ i, Algebra R (A i)] +variable (x : Π i, A i) (p : R[X]) + +/-- Polynomial evaluation on an indexed tuple is the indexed product of the evaluations +on the components. +Generalizes `Polynomial.aeval_prod` to indexed products. -/ +theorem aeval_pi (x : Π i, A i) : aeval (R := R) x = Pi.algHom R A (fun i ↦ aeval (x i)) := + (funext fun i ↦ aeval_algHom (Pi.evalAlgHom R A i) x) ▸ + (Pi.algHom_comp R A (Pi.evalAlgHom R A) (aeval x)) + +theorem aeval_pi_apply₂ (j : I) : p.aeval x j = p.aeval (x j) := + aeval_pi (R := R) x ▸ Pi.algHom_apply R A (fun i ↦ aeval (x i)) p j + +/-- Polynomial evaluation on an indexed tuple is the indexed tuple of the evaluations +on the components. +Generalizes `Polynomial.aeval_prod_apply` to indexed products. -/ +theorem aeval_pi_apply : p.aeval x = fun j ↦ p.aeval (x j) := + funext fun j ↦ aeval_pi_apply₂ x p j + +end Pi + @[simp] theorem coe_aeval_eq_eval (r : R) : (aeval r : R[X] → R) = eval r := rfl diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index 4a0d9b37f9a1d3..880597eb13c1f6 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.MonoidAlgebra.Defs import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop import Mathlib.Data.Finset.Sort import Mathlib.Tactic.FastInstance +import Mathlib.Algebra.Group.Submonoid.Operations /-! # Theory of univariate polynomials diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index eb2061427a125d..9be07ffaeff4d2 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -234,7 +234,6 @@ theorem rootMultiplicity_mul {p q : R[X]} {x : R} (hpq : p * q ≠ 0) : multiplicity_mul (prime_X_sub_C x) (finiteMultiplicity_X_sub_C _ hpq)] open Multiset in -set_option linter.unusedVariables false in theorem exists_multiset_roots [DecidableEq R] : ∀ {p : R[X]} (_ : p ≠ 0), ∃ s : Multiset R, (Multiset.card s : WithBot ℕ) ≤ degree p ∧ ∀ a, s.count a = rootMultiplicity a p @@ -245,11 +244,6 @@ theorem exists_multiset_roots [DecidableEq R] : have hpd : 0 < degree p := degree_pos_of_root hp hx have hd0 : p /ₘ (X - C x) ≠ 0 := fun h => by rw [← mul_divByMonic_eq_iff_isRoot.2 hx, h, mul_zero] at hp; exact hp rfl - #adaptation_note - /-- - Since https://github.com/leanprover/lean4/pull/5338, this is considered unused, - because it is only used in the decreasing_by clause. - -/ have wf : degree (p /ₘ (X - C x)) < degree p := degree_divByMonic_lt _ (monic_X_sub_C x) hp ((degree_X_sub_C x).symm ▸ by decide) let ⟨t, htd, htr⟩ := @exists_multiset_roots _ (p /ₘ (X - C x)) hd0 diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index 761bf9815c9242..8c113ce263a893 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -384,6 +384,15 @@ theorem center_eq_top (R) [CommRing R] : center R = ⊤ := instance : CommRing (center R) := { inferInstanceAs (CommSemiring (Subsemiring.center R)), (center R).toRing with } +/-- The center of isomorphic (not necessarily associative) rings are isomorphic. -/ +@[simps!] def centerCongr (e : R ≃+* S) : center R ≃+* center S := + NonUnitalSubsemiring.centerCongr e + +/-- The center of a (not necessarily associative) ring +is isomorphic to the center of its opposite. -/ +@[simps!] def centerToMulOpposite : center R ≃+* center Rᵐᵒᵖ := + NonUnitalSubsemiring.centerToMulOpposite + end section DivisionRing diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index e38a8720a241af..65b72af50f1f40 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -262,6 +262,17 @@ This is not an instance as it forms a non-defeq diamond with abbrev center.commSemiring' : CommSemiring (center R) := { Submonoid.center.commMonoid', (center R).toNonAssocSemiring with } +variable {R} + +/-- The center of isomorphic (not necessarily associative) semirings are isomorphic. -/ +@[simps!] def centerCongr [NonAssocSemiring S] (e : R ≃+* S) : center R ≃+* center S := + NonUnitalSubsemiring.centerCongr e + +/-- The center of a (not necessarily associative) semiring +is isomorphic to the center of its opposite. -/ +@[simps!] def centerToMulOpposite : center R ≃+* center Rᵐᵒᵖ := + NonUnitalSubsemiring.centerToMulOpposite + end NonAssocSemiring section Semiring @@ -285,7 +296,6 @@ instance decidableMemCenter {R} [Semiring R] [DecidableEq R] [Fintype R] : theorem center_eq_top (R) [CommSemiring R] : center R = ⊤ := SetLike.coe_injective (Set.center_eq_univ R) - end Semiring section Centralizer diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index a7a71a986a8cab..4772bc08cc2a91 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -630,7 +630,7 @@ lemma appLE_eq_away_map {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} (hU : IsA letI := hV.isLocalization_basicOpen (f.appLE U V e r) ext : 1 apply IsLocalization.ringHom_ext (.powers r) - rw [IsLocalization.Away.map, IsLocalization.map_comp, + rw [IsLocalization.Away.map, CommRingCat.hom_ofHom, IsLocalization.map_comp, RingHom.algebraMap_toAlgebra, RingHom.algebraMap_toAlgebra, ← CommRingCat.hom_comp, ← CommRingCat.hom_comp, Scheme.Hom.appLE_map, Scheme.Hom.map_appLE] @@ -646,7 +646,7 @@ lemma app_basicOpen_eq_away_map {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} haveI := h.isLocalization_basicOpen (f.app U r) ext : 1 apply IsLocalization.ringHom_ext (.powers r) - rw [IsLocalization.Away.map, CommRingCat.hom_comp, RingHom.comp_assoc, + rw [IsLocalization.Away.map, CommRingCat.hom_comp, RingHom.comp_assoc, CommRingCat.hom_ofHom, IsLocalization.map_comp, RingHom.algebraMap_toAlgebra, RingHom.algebraMap_toAlgebra, ← RingHom.comp_assoc, ← CommRingCat.hom_comp, ← CommRingCat.hom_comp, ← X.presheaf.map_comp] @@ -817,31 +817,35 @@ def SpecMapRestrictBasicOpenIso {R S : CommRingCat} (f : R ⟶ S) (r : R) : Arrow.mk (Spec.map <| CommRingCat.ofHom (Localization.awayMap f.hom r)) := by letI e₁ : Localization.Away r ≃ₐ[R] Γ(Spec R, basicOpen r) := IsLocalization.algEquiv (Submonoid.powers r) _ _ - letI e₂ : Localization.Away (f r) ≃ₐ[S] Γ(Spec S, basicOpen (f r)) := - IsLocalization.algEquiv (Submonoid.powers (f r)) _ _ + letI e₂ : Localization.Away (f.hom r) ≃ₐ[S] Γ(Spec S, basicOpen (f.hom r)) := + IsLocalization.algEquiv (Submonoid.powers (f.hom r)) _ _ refine Arrow.isoMk ?_ ?_ ?_ · exact (Spec (.of S)).isoOfEq (comap_basicOpen _ _) ≪≫ - (IsAffineOpen.Spec_basicOpen (f r)).isoSpec ≪≫ Scheme.Spec.mapIso e₂.toCommRingCatIso.op + (IsAffineOpen.Spec_basicOpen (f.hom r)).isoSpec ≪≫ Scheme.Spec.mapIso e₂.toCommRingCatIso.op · exact (IsAffineOpen.Spec_basicOpen r).isoSpec ≪≫ Scheme.Spec.mapIso e₁.toCommRingCatIso.op · have := AlgebraicGeometry.IsOpenImmersion.of_isLocalization (S := (Localization.Away r)) r rw [← cancel_mono (Spec.map (CommRingCat.ofHom (algebraMap R (Localization.Away r))))] simp only [Arrow.mk_left, Arrow.mk_right, Functor.id_obj, Scheme.isoOfEq_rfl, Iso.refl_trans, - Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, Iso.symm_hom, - Scheme.Spec_map, Quiver.Hom.unop_op, Arrow.mk_hom, Category.assoc, - ← Spec.map_comp] - show _ ≫ Spec.map (CommRingCat.ofHom - ((e₂.toRingHom.comp (Localization.awayMap f.hom r)).comp (algebraMap R _))) - = _ ≫ _ ≫ Spec.map (CommRingCat.ofHom (e₁.toRingHom.comp (algebraMap R _))) - rw [RingHom.comp_assoc] - conv => enter [1,2,1,1,2]; tactic => exact IsLocalization.map_comp _ - rw [← RingHom.comp_assoc] - conv => enter [1,2,1,1,1]; tactic => exact e₂.toAlgHom.comp_algebraMap - conv => enter [2,2,2,1,1]; tactic => exact e₁.toAlgHom.comp_algebraMap - show _ ≫ Spec.map (f ≫ (Scheme.ΓSpecIso S).inv ≫ - (Spec S).presheaf.map (homOfLE le_top).op) = - Spec.map f ∣_ basicOpen r ≫ _ ≫ Spec.map ((Scheme.ΓSpecIso R).inv ≫ - (Spec R).presheaf.map (homOfLE le_top).op) + Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, Scheme.Spec_map, Quiver.Hom.unop_op, + Arrow.mk_hom, Category.assoc, ← Spec.map_comp] + conv => + congr + · enter [2, 1]; tactic => + show _ = + (f ≫ (Scheme.ΓSpecIso S).inv ≫ (Spec S).presheaf.map (homOfLE le_top).op) + ext + simp only [Localization.awayMap, IsLocalization.Away.map, AlgEquiv.toRingEquiv_eq_coe, + RingEquiv.toCommRingCatIso_hom, AlgEquiv.toRingEquiv_toRingHom, CommRingCat.hom_comp, + CommRingCat.hom_ofHom, RingHom.comp_apply, IsLocalization.map_eq, RingHom.coe_coe, + AlgEquiv.commutes, IsAffineOpen.algebraMap_Spec_obj] + · enter [2, 2, 1]; tactic => + show _ = (Scheme.ΓSpecIso R).inv ≫ (Spec R).presheaf.map (homOfLE le_top).op + ext + simp only [AlgEquiv.toRingEquiv_eq_coe, RingEquiv.toCommRingCatIso_hom, + AlgEquiv.toRingEquiv_toRingHom, CommRingCat.hom_comp, CommRingCat.hom_ofHom, + RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, AlgEquiv.commutes, + IsAffineOpen.algebraMap_Spec_obj, homOfLE_leOfHom] simp only [IsAffineOpen.isoSpec_hom, homOfLE_leOfHom, Spec.map_comp, Category.assoc, Scheme.Opens.toSpecΓ_SpecMap_map_assoc, Scheme.Opens.toSpecΓ_top, Scheme.homOfLE_ι_assoc, morphismRestrict_ι_assoc] diff --git a/Mathlib/AlgebraicGeometry/AffineSpace.lean b/Mathlib/AlgebraicGeometry/AffineSpace.lean index e845edc4961402..b8779620b20642 100644 --- a/Mathlib/AlgebraicGeometry/AffineSpace.lean +++ b/Mathlib/AlgebraicGeometry/AffineSpace.lean @@ -89,7 +89,7 @@ def toSpecMvPolyIntEquiv : (X ⟶ Spec ℤ[n]) ≃ (n → Γ(X, ⊤)) where Iso.cancel_iso_inv_left, ← Iso.eq_inv_comp] apply of_mvPolynomial_int_ext intro i - rw [coe_eval₂Hom, eval₂_X] + rw [ConcreteCategory.hom_ofHom, coe_eval₂Hom, eval₂_X] rfl right_inv v := by ext i @@ -203,7 +203,7 @@ def isoOfIsAffine [IsAffine S] : ext : 1 apply ringHom_ext' · show _ = (CommRingCat.ofHom C ≫ _).hom - rw [CommRingCat.hom_comp, RingHom.comp_assoc, eval₂Hom_comp_C, + rw [CommRingCat.hom_comp, RingHom.comp_assoc, CommRingCat.hom_ofHom, eval₂Hom_comp_C, ← CommRingCat.hom_comp, ← CommRingCat.hom_ext_iff, ← cancel_mono (Scheme.ΓSpecIso _).hom] rw [← Scheme.comp_appTop, homOfVector_over, Scheme.comp_appTop] @@ -212,7 +212,7 @@ def isoOfIsAffine [IsAffine S] : rw [← Scheme.comp_appTop_assoc, Scheme.isoSpec, asIso_inv, IsIso.hom_inv_id] simp · intro i - rw [CommRingCat.comp_apply, coe_eval₂Hom] + rw [CommRingCat.comp_apply, ConcreteCategory.hom_ofHom, coe_eval₂Hom] simp only [eval₂_X] exact homOfVector_appTop_coord _ _ _ @@ -328,7 +328,7 @@ lemma map_Spec_map {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : conv_lhs => enter[2]; tactic => exact map_appTop_coord _ _ conv_rhs => enter[2]; tactic => exact SpecIso_inv_appTop_coord _ _ rw [SpecIso_inv_appTop_coord, ← CommRingCat.comp_apply, ← Scheme.ΓSpecIso_inv_naturality, - CommRingCat.comp_apply, map_X] + CommRingCat.comp_apply, ConcreteCategory.hom_ofHom, map_X] /-- The map between affine spaces over affine bases is isomorphic to the natural map between polynomial rings. -/ diff --git a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean index 87849e46136cac..042468ac880d9b 100644 --- a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean +++ b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean @@ -113,7 +113,7 @@ def Cover.bind [P.IsStableUnderComposition] (f : ∀ x : 𝒰.J, (𝒰.obj x).Co rcases (f (𝒰.f x)).covers y with ⟨z, hz⟩ change x ∈ Set.range ((f (𝒰.f x)).map ((f (𝒰.f x)).f y) ≫ 𝒰.map (𝒰.f x)).base use z - erw [comp_apply] + erw [CategoryTheory.comp_apply] rw [hz, hy] map_prop _ := P.comp_mem _ _ ((f _).map_prop _) (𝒰.map_prop _) diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index f7d458c2b79a60..2506ebe79df0c0 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -139,7 +139,7 @@ theorem toΓSpecCApp_iff -- Porting Note: Type class problem got stuck in `IsLocalization.Away.AwayMap.lift_comp` -- created instance manually. This replaces the `pick_goal` tactics have loc_inst := IsLocalization.to_basicOpen (Γ.obj (op X)) r - refine CommRingCat.hom_ext_iff.trans ?_ + refine ConcreteCategory.ext_iff.trans ?_ rw [← @IsLocalization.Away.lift_comp _ _ _ _ _ _ _ r loc_inst _ (X.isUnit_res_toΓSpecMapBasicOpen r)] --pick_goal 5; exact is_localization.to_basic_open _ r diff --git a/Mathlib/AlgebraicGeometry/Gluing.lean b/Mathlib/AlgebraicGeometry/Gluing.lean index bcdfd0adfd7723..2d70705beae896 100644 --- a/Mathlib/AlgebraicGeometry/Gluing.lean +++ b/Mathlib/AlgebraicGeometry/Gluing.lean @@ -227,7 +227,7 @@ theorem ι_eq_iff (i j : D.J) (x : (D.U i).carrier) (y : (D.U j).carrier) : (TopCat.GlueData.ι_eq_iff_rel D.toLocallyRingedSpaceGlueData.toSheafedSpaceGlueData.toPresheafedSpaceGlueData.toTopGlueData i j x y) - rw [← ((TopCat.mono_iff_injective D.isoCarrier.inv).mp _).eq_iff, ← comp_apply] + rw [← ((TopCat.mono_iff_injective D.isoCarrier.inv).mp _).eq_iff, ← CategoryTheory.comp_apply] · simp_rw [← D.ι_isoCarrier_inv] rfl -- `rfl` was not needed before https://github.com/leanprover-community/mathlib4/pull/13170 · infer_instance @@ -339,7 +339,7 @@ theorem fromGlued_injective : Function.Injective 𝒰.fromGlued.base := by intro x y h obtain ⟨i, x, rfl⟩ := 𝒰.gluedCover.ι_jointly_surjective x obtain ⟨j, y, rfl⟩ := 𝒰.gluedCover.ι_jointly_surjective y - rw [← comp_apply, ← comp_apply] at h + rw [← CategoryTheory.comp_apply, ← CategoryTheory.comp_apply] at h simp_rw [← Scheme.comp_base] at h rw [ι_fromGlued, ι_fromGlued] at h let e := @@ -389,7 +389,7 @@ instance : Epi 𝒰.fromGlued.base := by intro x obtain ⟨y, h⟩ := 𝒰.covers x use (𝒰.gluedCover.ι (𝒰.f x)).base y - rw [← comp_apply] + rw [← CategoryTheory.comp_apply] rw [← 𝒰.ι_fromGlued (𝒰.f x)] at h exact h diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean index 4fa80786dbe0c9..a941013a085e5d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean @@ -248,9 +248,8 @@ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isAffineOpen (X : Sch (H : x |_ (X.basicOpen f) = 0) : ∃ n : ℕ, f ^ n * x = 0 := by rw [← map_zero (X.presheaf.map (homOfLE <| X.basicOpen_le f : X.basicOpen f ⟶ U).op).hom] at H - #adaptation_note - /-- - Prior to nightly-2024-09-29, we could use dot notation here: + #adaptation_note /-- nightly-2024-09-29 + we could use dot notation here: `(hU.isLocalization_basicOpen f).exists_of_eq H` This is no longer possible; likely changing the signature of `IsLocalization.Away.exists_of_eq` is in order. diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 9f07373d33dce9..625901a76887ed 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -158,7 +158,7 @@ theorem sourceAffineLocally_isLocal (h₁ : RingHom.RespectsIso P) simpa using hU rw [← f.appLE_congr _ rfl this (fun f => P f.hom), IsAffineOpen.appLE_eq_away_map f (isAffineOpen_top Y) U.2 _ r] - simp only + simp only [CommRingCat.hom_ofHom] apply (config := { allowSynthFailures := true }) h₂ exact H U · introv hs hs' U @@ -167,7 +167,7 @@ theorem sourceAffineLocally_isLocal (h₁ : RingHom.RespectsIso P) simp_rw [sourceAffineLocally_morphismRestrict] at hs' have := hs' r ⟨X.basicOpen (f.appLE ⊤ U le_top r.1), U.2.basicOpen (f.appLE ⊤ U le_top r.1)⟩ (by simp [Scheme.Hom.appLE]) - rwa [IsAffineOpen.appLE_eq_away_map f (isAffineOpen_top Y) U.2, + rwa [IsAffineOpen.appLE_eq_away_map f (isAffineOpen_top Y) U.2, CommRingCat.hom_ofHom, ← h₁.is_localization_away_iff] at this variable {P} diff --git a/Mathlib/AlgebraicGeometry/Noetherian.lean b/Mathlib/AlgebraicGeometry/Noetherian.lean index 7d1c4fff13747f..73a384e72fb398 100644 --- a/Mathlib/AlgebraicGeometry/Noetherian.lean +++ b/Mathlib/AlgebraicGeometry/Noetherian.lean @@ -191,9 +191,8 @@ lemma noetherianSpace_of_isAffineOpen (U : X.Opens) (hU : IsAffineOpen U) (Scheme.restrictFunctorΓ.app (op U)).symm.commRingCatIsoToRingEquiv exact @noetherianSpace_of_isAffine _ hU _ -/-- Any open immersion `Z ⟶ X` with `X` locally Noetherian is quasi-compact. - -[Stacks: Lemma 01OX](https://stacks.math.columbia.edu/tag/01OX) -/ +/-- Any open immersion `Z ⟶ X` with `X` locally Noetherian is quasi-compact. -/ +@[stacks 01OX] instance (priority := 100) {Z : Scheme} [IsLocallyNoetherian X] {f : Z ⟶ X} [IsOpenImmersion f] : QuasiCompact f := by apply (quasiCompact_iff_forall_affine f).mpr @@ -206,9 +205,8 @@ instance (priority := 100) {Z : Scheme} [IsLocallyNoetherian X] · exact Set.inter_subset_left · exact Set.inter_subset_right -/-- A locally Noetherian scheme is quasi-separated. - -[Stacks: Lemma 01OY](https://stacks.math.columbia.edu/tag/01OY) -/ +/-- A locally Noetherian scheme is quasi-separated. -/ +@[stacks 01OY] instance (priority := 100) IsLocallyNoetherian.quasiSeparatedSpace [IsLocallyNoetherian X] : QuasiSeparatedSpace X := by apply (quasiSeparatedSpace_iff_affine X).mpr @@ -266,9 +264,8 @@ theorem isNoetherian_iff_of_finite_affine_openCover {𝒰 : Scheme.OpenCover.{v, · exact Scheme.OpenCover.compactSpace 𝒰 open CategoryTheory in -/-- A Noetherian scheme has a Noetherian underlying topological space. - -[Stacks, Lemma 01OZ](https://stacks.math.columbia.edu/tag/01OZ) -/ +/-- A Noetherian scheme has a Noetherian underlying topological space. -/ +@[stacks 01OZ] instance (priority := 100) IsNoetherian.noetherianSpace [IsNoetherian X] : NoetherianSpace X := by apply TopologicalSpace.noetherian_univ_iff.mp @@ -284,9 +281,8 @@ instance (priority := 100) IsNoetherian.noetherianSpace [IsNoetherian X] : convert noetherianSpace_of_isAffineOpen U.1 U.2 apply IsLocallyNoetherian.component_noetherian -/-- Any morphism of schemes `f : X ⟶ Y` with `X` Noetherian is quasi-compact. - -[Stacks, Lemma 01P0](https://stacks.math.columbia.edu/tag/01P0) -/ +/-- Any morphism of schemes `f : X ⟶ Y` with `X` Noetherian is quasi-compact. -/ +@[stacks 01P0] instance (priority := 100) quasiCompact_of_noetherianSpace_source {X Y : Scheme} [NoetherianSpace X] (f : X ⟶ Y) : QuasiCompact f := ⟨fun _ _ _ => NoetherianSpace.isCompact _⟩ @@ -325,9 +321,8 @@ theorem isNoetherian_Spec {R : CommRingCat} : ⟨fun _ => inferInstance, fun _ => inferInstance⟩ -/-- A Noetherian scheme has a finite number of irreducible components. - -[Stacks, Lemma 0BA8](https://stacks.math.columbia.edu/tag/0BA8) -/ +/-- A Noetherian scheme has a finite number of irreducible components. -/ +@[stacks 0BA8] theorem finite_irreducibleComponents_of_isNoetherian [IsNoetherian X] : (irreducibleComponents X).Finite := NoetherianSpace.finite_irreducibleComponents diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean index 16327672ba49a1..8a215ab44a0221 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean @@ -179,7 +179,7 @@ lemma awayMap_awayToSection : obtain ⟨⟨n, a, ⟨b, hb'⟩, i, rfl : _ = b⟩, rfl⟩ := mk_surjective a simp only [homOfLE_leOfHom, CommRingCat.hom_comp, RingHom.coe_comp, Function.comp_apply] erw [ProjectiveSpectrum.Proj.awayToSection_apply] - rw [val_awayMap_mk, Localization.mk_eq_mk', IsLocalization.map_mk', + rw [CommRingCat.hom_ofHom, val_awayMap_mk, Localization.mk_eq_mk', IsLocalization.map_mk', ← Localization.mk_eq_mk'] refine Localization.mk_eq_mk_iff.mpr ?_ rw [Localization.r_iff_exists] diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index fe628dda5c4cb5..26b9f817d1bf58 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -420,7 +420,7 @@ instance isLocalizedModule_toPushforwardStalkAlgHom : change PrimeSpectrum.basicOpen r ≤ U at hrU apply_fun (Spec.topMap (CommRingCat.ofHom (algebraMap R S)) _* (structureSheaf S).1).map (homOfLE hrU).op at e - simp only [Functor.op_map, map_zero, ← comp_apply, toOpen_res] at e + simp only [Functor.op_map, map_zero, ← CategoryTheory.comp_apply, toOpen_res] at e have : toOpen S (PrimeSpectrum.basicOpen <| algebraMap R S r) x = 0 := by refine Eq.trans ?_ e; rfl have := diff --git a/Mathlib/AlgebraicGeometry/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/StructureSheaf.lean index 55c3b14a35165f..9e79bc6c7f8b98 100644 --- a/Mathlib/AlgebraicGeometry/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/StructureSheaf.lean @@ -947,9 +947,8 @@ theorem localizationToStalk_stalkSpecializes {R : Type*} [CommRing R] {x y : Pri StructureSheaf.localizationToStalk R x := by ext : 1 apply IsLocalization.ringHom_ext (S := Localization.AtPrime y.asIdeal) y.asIdeal.primeCompl - erw [RingHom.comp_assoc] - conv_rhs => erw [RingHom.comp_assoc] - dsimp [CommRingCat.ofHom, localizationToStalk, PrimeSpectrum.localizationMapOfSpecializes] + rw [CommRingCat.hom_comp, RingHom.comp_assoc, CommRingCat.hom_comp, RingHom.comp_assoc] + dsimp [localizationToStalk, PrimeSpectrum.localizationMapOfSpecializes] rw [IsLocalization.lift_comp, IsLocalization.lift_comp, IsLocalization.lift_comp] exact CommRingCat.hom_ext_iff.mp (toStalk_stalkSpecializes h) diff --git a/Mathlib/AlgebraicTopology/SimplicialObject/Split.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Split.lean index 9061d57109b29b..8edfd646e7deed 100644 --- a/Mathlib/AlgebraicTopology/SimplicialObject/Split.lean +++ b/Mathlib/AlgebraicTopology/SimplicialObject/Split.lean @@ -376,7 +376,8 @@ theorem comp_f {S₁ S₂ S₃ : Split C} (Φ₁₂ : S₁ ⟶ S₂) (Φ₂₃ : (Φ₁₂ ≫ Φ₂₃).f n = Φ₁₂.f n ≫ Φ₂₃.f n := rfl -@[reassoc (attr := simp 1100)] +-- This is not a `@[simp]` lemma as it can later be proved by `simp`. +@[reassoc] theorem cofan_inj_naturality_symm {S₁ S₂ : Split C} (Φ : S₁ ⟶ S₂) {Δ : SimplexCategoryᵒᵖ} (A : Splitting.IndexSet Δ) : (S₁.s.cofan Δ).inj A ≫ Φ.F.app Δ = Φ.f A.1.unop.len ≫ (S₂.s.cofan Δ).inj A := by diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Boundary.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Boundary.lean index 58805409d1cebd..e9f81272b8ff26 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Boundary.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Boundary.lean @@ -41,13 +41,6 @@ def boundary (n : ℕ) : SSet.{u} where /-- The boundary `∂Δ[n]` of the `n`-th standard simplex -/ scoped[Simplicial] notation3 "∂Δ[" n "]" => SSet.boundary n -#adaptation_note -/-- -The new unused variable linter in -https://github.com/leanprover/lean4/pull/5338 -flags `{ α : Δ[n].obj m // _ }`. --/ -set_option linter.unusedVariables false in /-- The inclusion of the boundary of the `n`-th standard simplex into that standard simplex. -/ def boundaryInclusion (n : ℕ) : ∂Δ[n] ⟶ Δ[n] where app m (α : { α : Δ[n].obj m // _ }) := α diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Horn.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Horn.lean index ef8c001fd1d797..6eb839414c7599 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Horn.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Horn.lean @@ -35,13 +35,6 @@ def horn (n : ℕ) (i : Fin (n + 1)) : SSet where /-- The `i`-th horn `Λ[n, i]` of the standard `n`-simplex -/ scoped[Simplicial] notation3 "Λ[" n ", " i "]" => SSet.horn (n : ℕ) i -#adaptation_note -/-- -The new unused variable linter in -https://github.com/leanprover/lean4/pull/5338 -flags `{ α : Δ[n].obj m // _ }`. --/ -set_option linter.unusedVariables false in /-- The inclusion of the `i`-th horn of the `n`-th standard simplex into that standard simplex. -/ def hornInclusion (n : ℕ) (i : Fin (n + 1)) : Λ[n, i] ⟶ Δ[n] where app m (α : { α : Δ[n].obj m // _ }) := α diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index f607e676809c63..095e60deae80ab 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -503,13 +503,13 @@ theorem HasFPowerSeriesOnBall.comp_sub (hf : HasFPowerSeriesOnBall f p x r) (y : theorem HasFPowerSeriesWithinOnBall.hasSum_sub (hf : HasFPowerSeriesWithinOnBall f p s x r) {y : E} (hy : y ∈ (insert x s) ∩ EMetric.ball x r) : HasSum (fun n : ℕ => p n fun _ => y - x) (f y) := by - have : y - x ∈ EMetric.ball (0 : E) r := by simpa [edist_eq_coe_nnnorm_sub] using hy.2 + have : y - x ∈ EMetric.ball (0 : E) r := by simpa [edist_eq_enorm_sub] using hy.2 have := hf.hasSum (by simpa only [add_sub_cancel] using hy.1) this simpa only [add_sub_cancel] theorem HasFPowerSeriesOnBall.hasSum_sub (hf : HasFPowerSeriesOnBall f p x r) {y : E} (hy : y ∈ EMetric.ball x r) : HasSum (fun n : ℕ => p n fun _ => y - x) (f y) := by - have : y - x ∈ EMetric.ball (0 : E) r := by simpa [edist_eq_coe_nnnorm_sub] using hy + have : y - x ∈ EMetric.ball (0 : E) r := by simpa [edist_eq_enorm_sub] using hy simpa only [add_sub_cancel] using hf.hasSum this theorem HasFPowerSeriesOnBall.radius_pos (hf : HasFPowerSeriesOnBall f p x r) : 0 < p.radius := @@ -544,7 +544,7 @@ lemma HasFPowerSeriesWithinOnBall.congr {f g : E → F} {p : FormalMultilinearSe · simpa using h'' · apply h' refine ⟨hy, ?_⟩ - simpa [edist_eq_coe_nnnorm_sub] using h'y + simpa [edist_eq_enorm_sub] using h'y /-- Variant of `HasFPowerSeriesWithinOnBall.congr` in which one requests equality on `insert x s` instead of separating `x` and `s`. -/ @@ -554,7 +554,7 @@ lemma HasFPowerSeriesWithinOnBall.congr' {f g : E → F} {p : FormalMultilinearS HasFPowerSeriesWithinOnBall g p s x r := by refine ⟨h.r_le, h.r_pos, fun {y} hy h'y ↦ ?_⟩ convert h.hasSum hy h'y using 1 - exact h' ⟨hy, by simpa [edist_eq_coe_nnnorm_sub] using h'y⟩ + exact h' ⟨hy, by simpa [edist_eq_enorm_sub] using h'y⟩ lemma HasFPowerSeriesWithinAt.congr {f g : E → F} {p : FormalMultilinearSeries 𝕜 E F} {s : Set E} {x : E} (h : HasFPowerSeriesWithinAt f p s x) (h' : g =ᶠ[𝓝[s] x] f) (h'' : g x = f x) : @@ -576,7 +576,7 @@ theorem HasFPowerSeriesOnBall.congr (hf : HasFPowerSeriesOnBall f p x r) hasSum := fun {y} hy => by convert hf.hasSum hy using 1 apply hg.symm - simpa [edist_eq_coe_nnnorm_sub] using hy } + simpa [edist_eq_enorm_sub] using hy } theorem HasFPowerSeriesAt.congr (hf : HasFPowerSeriesAt f p x) (hg : f =ᶠ[𝓝 x] g) : HasFPowerSeriesAt g p x := by @@ -679,7 +679,7 @@ theorem HasFPowerSeriesWithinAt.mono_of_mem_nhdsWithin rcases hy with rfl | hy · simp apply mem_insert_of_mem _ (hr' ?_) - simp only [EMetric.mem_ball, edist_eq_coe_nnnorm_sub, sub_zero, lt_min_iff, mem_inter_iff, + simp only [EMetric.mem_ball, edist_eq_enorm_sub, sub_zero, lt_min_iff, mem_inter_iff, add_sub_cancel_left, hy, and_true] at h'y ⊢ exact h'y.2 @@ -912,7 +912,7 @@ theorem HasFPowerSeriesWithinOnBall.tendsto_partialSum_prod {y : E} (atTop ×ˢ 𝓝 y) (𝓝 0) by simpa using A.add this apply Metric.tendsto_nhds.2 (fun ε εpos ↦ ?_) obtain ⟨r', yr', r'r⟩ : ∃ (r' : ℝ≥0), ‖y‖₊ < r' ∧ r' < r := by - simp [edist_eq_coe_nnnorm] at hy + simp [edist_zero_eq_enorm] at hy simpa using ENNReal.lt_iff_exists_nnreal_btwn.1 hy have yr'_2 : ‖y‖ < r' := by simpa [← coe_nnnorm] using yr' have S : Summable fun n ↦ ‖p n‖ * ↑r' ^ n := p.summable_norm_mul_pow (r'r.trans_le hf.r_le) @@ -991,7 +991,7 @@ theorem HasFPowerSeriesWithinOnBall.uniform_geometric_approx' {r' : ℝ≥0} have hr'0 : 0 < (r' : ℝ) := (norm_nonneg _).trans_lt yr' have : y ∈ EMetric.ball (0 : E) r := by refine mem_emetric_ball_zero_iff.2 (lt_trans ?_ h) - exact mod_cast yr' + simpa [enorm] using yr' rw [norm_sub_rev, ← mul_div_right_comm] have ya : a * (‖y‖ / ↑r') ≤ a := mul_le_of_le_one_right ha.1.le (div_le_one_of_le₀ yr'.le r'.coe_nonneg) @@ -1102,7 +1102,7 @@ theorem HasFPowerSeriesWithinOnBall.isBigO_image_sub_image_sub_deriv_principal zero_add, ← Subsingleton.pi_single_eq (0 : Fin 1) (y.1 - x), Pi.single, ← Subsingleton.pi_single_eq (0 : Fin 1) (y.2 - x), Pi.single, ← (p 1).map_update_sub, ← Pi.single, Subsingleton.pi_single_eq, sub_sub_sub_cancel_right] - rw [EMetric.mem_ball, edist_eq_coe_nnnorm_sub, ENNReal.coe_lt_coe] at hy' + rw [EMetric.mem_ball, edist_eq_enorm_sub, enorm_lt_coe] at hy' set B : ℕ → ℝ := fun n => C * (a / r') ^ 2 * (‖y - (x, x)‖ * ‖y.1 - y.2‖) * ((n + 2) * a ^ n) have hAB : ∀ n, ‖A (n + 2)‖ ≤ B n := fun n => calc @@ -1282,7 +1282,7 @@ theorem HasFPowerSeriesWithinOnBall.tendstoLocallyUniformlyOn' · ext z simp · intro z - simp [edist_eq_coe_nnnorm, edist_eq_coe_nnnorm_sub] + simp [edist_zero_eq_enorm, edist_eq_enorm_sub] /-- If a function admits a power series expansion at `x`, then it is the locally uniform limit of the partial sums of this power series on the disk of convergence, i.e., `f y` diff --git a/Mathlib/Analysis/Analytic/CPolynomial.lean b/Mathlib/Analysis/Analytic/CPolynomial.lean index 8c2c58d2957d56..838c7af23b5b86 100644 --- a/Mathlib/Analysis/Analytic/CPolynomial.lean +++ b/Mathlib/Analysis/Analytic/CPolynomial.lean @@ -278,7 +278,7 @@ theorem HasFiniteFPowerSeriesOnBall.eq_partialSum' ∀ y ∈ EMetric.ball x r, ∀ m, n ≤ m → f y = p.partialSum m (y - x) := by intro y hy m hm - rw [EMetric.mem_ball, edist_eq_coe_nnnorm_sub, ← mem_emetric_ball_zero_iff] at hy + rw [EMetric.mem_ball, edist_eq_enorm_sub, ← mem_emetric_ball_zero_iff] at hy rw [← (HasFiniteFPowerSeriesOnBall.eq_partialSum hf _ hy m hm), add_sub_cancel] /-! The particular cases where `f` has a finite power series bounded by `0` or `1`. -/ @@ -300,8 +300,8 @@ theorem HasFiniteFPowerSeriesOnBall.bound_zero_of_eq_zero (hf : ∀ y ∈ EMetri rw [hf (x + y)] · convert hasSum_zero rw [hp, ContinuousMultilinearMap.zero_apply] - · rwa [EMetric.mem_ball, edist_eq_coe_nnnorm_sub, add_comm, add_sub_cancel_right, - ← edist_eq_coe_nnnorm, ← EMetric.mem_ball] + · rwa [EMetric.mem_ball, edist_eq_enorm_sub, add_comm, add_sub_cancel_right, + ← edist_zero_eq_enorm, ← EMetric.mem_ball] /-- If `f` has a formal power series at `x` bounded by `0`, then `f` is equal to `0` in a neighborhood of `x`. -/ @@ -396,8 +396,8 @@ lemma changeOriginSeriesTerm_bound (p : FormalMultilinearSeries 𝕜 E F) {n : (hn : ∀ (m : ℕ), n ≤ m → p m = 0) (k l : ℕ) {s : Finset (Fin (k + l))} (hs : s.card = l) (hkl : n ≤ k + l) : p.changeOriginSeriesTerm k l s hs = 0 := by - #adaptation_note - /-- `set_option maxSynthPendingDepth 2` required after https://github.com/leanprover/lean4/pull/4119 -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 + `set_option maxSynthPendingDepth 2` required after https://github.com/leanprover/lean4/pull/4119 -/ set_option maxSynthPendingDepth 2 in rw [changeOriginSeriesTerm, hn _ hkl, map_zero] @@ -504,8 +504,7 @@ theorem HasFiniteFPowerSeriesOnBall.changeOrigin (hf : HasFiniteFPowerSeriesOnBa FormalMultilinearSeries.sum (FormalMultilinearSeries.changeOrigin p y) z := by rw [mem_emetric_ball_zero_iff, lt_tsub_iff_right, add_comm] at hz rw [p.changeOrigin_eval_of_finite hf.finite, add_assoc, hf.sum] - refine mem_emetric_ball_zero_iff.2 (lt_of_le_of_lt ?_ hz) - exact mod_cast nnnorm_add_le y z + exact mem_emetric_ball_zero_iff.2 ((enorm_add_le _ _).trans_lt hz) rw [this] apply (p.changeOrigin y).hasSum_of_finite fun _ => p.changeOrigin_finite_of_finite hf.finite @@ -514,7 +513,7 @@ it is continuously polynomial at every point of this ball. -/ theorem HasFiniteFPowerSeriesOnBall.cPolynomialAt_of_mem (hf : HasFiniteFPowerSeriesOnBall f p x n r) (h : y ∈ EMetric.ball x r) : CPolynomialAt 𝕜 f y := by - have : (‖y - x‖₊ : ℝ≥0∞) < r := by simpa [edist_eq_coe_nnnorm_sub] using h + have : (‖y - x‖₊ : ℝ≥0∞) < r := by simpa [edist_eq_enorm_sub] using h have := hf.changeOrigin this rw [add_sub_cancel] at this exact this.cPolynomialAt diff --git a/Mathlib/Analysis/Analytic/ChangeOrigin.lean b/Mathlib/Analysis/Analytic/ChangeOrigin.lean index 0e2e48a529be8c..a9dd1bdf959940 100644 --- a/Mathlib/Analysis/Analytic/ChangeOrigin.lean +++ b/Mathlib/Analysis/Analytic/ChangeOrigin.lean @@ -306,8 +306,8 @@ variable [CompleteSpace F] {f : E → F} {p : FormalMultilinearSeries 𝕜 E F} it also admits a power series on any subball of this ball (even with a different center provided it belongs to `s`), given by `p.changeOrigin`. -/ theorem HasFPowerSeriesWithinOnBall.changeOrigin (hf : HasFPowerSeriesWithinOnBall f p s x r) - (h : (‖y‖₊ : ℝ≥0∞) < r) (hy : x + y ∈ insert x s) : - HasFPowerSeriesWithinOnBall f (p.changeOrigin y) s (x + y) (r - ‖y‖₊) where + (h : ‖y‖ₑ < r) (hy : x + y ∈ insert x s) : + HasFPowerSeriesWithinOnBall f (p.changeOrigin y) s (x + y) (r - ‖y‖ₑ) where r_le := by apply le_trans _ p.changeOrigin_radius exact tsub_le_tsub hf.r_le le_rfl @@ -322,8 +322,7 @@ theorem HasFPowerSeriesWithinOnBall.changeOrigin (hf : HasFPowerSeriesWithinOnBa rw [insert_eq_of_mem hy] at this apply this simpa [add_assoc] using h'z - refine mem_emetric_ball_zero_iff.2 (lt_of_le_of_lt ?_ hz) - exact mod_cast nnnorm_add_le y z + exact mem_emetric_ball_zero_iff.2 (lt_of_le_of_lt (enorm_add_le _ _) hz) rw [this] apply (p.changeOrigin y).hasSum refine EMetric.ball_subset_ball (le_trans ?_ p.changeOrigin_radius) hz @@ -342,7 +341,7 @@ it is analytic at every point of this ball. -/ theorem HasFPowerSeriesWithinOnBall.analyticWithinAt_of_mem (hf : HasFPowerSeriesWithinOnBall f p s x r) (h : y ∈ insert x s ∩ EMetric.ball x r) : AnalyticWithinAt 𝕜 f s y := by - have : (‖y - x‖₊ : ℝ≥0∞) < r := by simpa [edist_eq_coe_nnnorm_sub] using h.2 + have : (‖y - x‖₊ : ℝ≥0∞) < r := by simpa [edist_eq_enorm_sub] using h.2 have := hf.changeOrigin this (by simpa using h.1) rw [add_sub_cancel] at this exact this.analyticWithinAt diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index ae4386b4831c7f..6844e170f22e95 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -742,7 +742,7 @@ theorem HasFPowerSeriesWithinAt.comp {g : F → G} {f : E → F} {q : FormalMult apply hδ have : y ∈ EMetric.ball (0 : E) δ := (EMetric.ball_subset_ball (le_trans (min_le_left _ _) (min_le_right _ _))) hy - simpa [-Set.mem_insert_iff, edist_eq_coe_nnnorm_sub, h'y] + simpa [-Set.mem_insert_iff, edist_eq_enorm_sub, h'y] /- Now the proof starts. To show that the sum of `q.comp p` at `y` is `g (f (x + y))`, we will write `q.comp p` applied to `y` as a big sum over all compositions. Since the sum is summable, to get its convergence it suffices to get @@ -775,7 +775,7 @@ theorem HasFPowerSeriesWithinAt.comp {g : F → G} {f : E → F} {q : FormalMult have : Tendsto (fun (z : ℕ × F) ↦ q.partialSum z.1 z.2) (atTop ×ˢ 𝓝 (f (x + y) - f x)) (𝓝 (g (f x + (f (x + y) - f x)))) := by apply Hg.tendsto_partialSum_prod (y := f (x + y) - f x) - · simpa [edist_eq_coe_nnnorm_sub] using fy_mem.2 + · simpa [edist_eq_enorm_sub] using fy_mem.2 · simpa using fy_mem.1 simpa using this.comp A -- Third step: the sum over all compositions in `compPartialSumTarget 0 n n` converges to @@ -806,9 +806,9 @@ theorem HasFPowerSeriesWithinAt.comp {g : F → G} {f : E → F} {q : FormalMult apply mul_le_mul_of_nonneg_left _ (norm_nonneg _) rw [Finset.prod_const, Finset.card_fin] gcongr - rw [EMetric.mem_ball, edist_eq_coe_nnnorm] at hy + rw [EMetric.mem_ball, edist_zero_eq_enorm] at hy have := le_trans (le_of_lt hy) (min_le_right _ _) - rwa [ENNReal.coe_le_coe, ← NNReal.coe_le_coe, coe_nnnorm] at this + rwa [enorm_le_coe, ← NNReal.coe_le_coe, coe_nnnorm] at this tendsto_nhds_of_cauchySeq_of_subseq cau compPartialSumTarget_tendsto_atTop C -- Fifth step: the sum over `n` of `q.comp p n` can be expressed as a particular resummation of -- the sum over all compositions, by grouping together the compositions of the same diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index 6ef7fff93cce3a..bb9bba9211dc66 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -592,9 +592,9 @@ lemma HasFPowerSeriesAt.tendsto_partialSum_prod_of_comp apply mul_le_mul_of_nonneg_left _ (norm_nonneg _) rw [Finset.prod_const, Finset.card_fin] gcongr - rw [EMetric.mem_ball, edist_eq_coe_nnnorm] at hy + rw [EMetric.mem_ball, edist_zero_eq_enorm] at hy have := le_trans (le_of_lt hy) (min_le_right _ _) - rwa [ENNReal.coe_le_coe, ← NNReal.coe_le_coe, coe_nnnorm] at this + rwa [enorm_le_coe, ← NNReal.coe_le_coe, coe_nnnorm] at this apply HasSum.of_sigma (fun b ↦ hasSum_fintype _) ?_ cau simpa [FormalMultilinearSeries.comp] using h0.hasSum hy0 have B : Tendsto (fun (n : ℕ × ℕ) => ∑ i ∈ compPartialSumTarget 0 n.1 n.2, @@ -686,7 +686,7 @@ theorem PartialHomeomorph.hasFPowerSeriesAt_symm (f : PartialHomeomorph E F) {a refine ⟨min r (p.leftInv i a).radius, min_le_right _ _, lt_min r_pos (radius_leftInv_pos_of_radius_pos h.radius_pos hp), fun {y} hy ↦ ?_⟩ have : y + f a ∈ EMetric.ball (f a) r := by - simp only [EMetric.mem_ball, edist_eq_coe_nnnorm_sub, sub_zero, lt_min_iff, + simp only [EMetric.mem_ball, edist_eq_enorm_sub, sub_zero, lt_min_iff, add_sub_cancel_right] at hy ⊢ exact hy.1 simpa [add_comm] using hr this diff --git a/Mathlib/Analysis/Analytic/Meromorphic.lean b/Mathlib/Analysis/Analytic/Meromorphic.lean index 86a7ea98ddcc01..ff69888be1dc23 100644 --- a/Mathlib/Analysis/Analytic/Meromorphic.lean +++ b/Mathlib/Analysis/Analytic/Meromorphic.lean @@ -278,6 +278,30 @@ lemma _root_.AnalyticAt.meromorphicAt_order {f : 𝕜 → E} {x : 𝕜} (hf : An rcases (hf.order_eq_nat_iff _).mp hn.symm with ⟨g, h1, h2, h3⟩ exact ⟨g, h1, h2, h3.filter_mono nhdsWithin_le_nhds⟩ +/-- The order is additive when multiplying scalar-valued and vector-valued meromorphic functions. -/ +theorem order_smul {f : 𝕜 → 𝕜} {g : 𝕜 → E} {x : 𝕜} + (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : + (hf.smul hg).order = hf.order + hg.order := by + -- Trivial cases: one of the functions vanishes around z₀ + cases' h₂f : hf.order with m h₂f + · simp only [top_add, order_eq_top_iff] at h₂f ⊢ + filter_upwards [h₂f] with z hz using by simp [hz] + cases' h₂g : hg.order with n h₂f + · simp only [add_top, order_eq_top_iff] at h₂g ⊢ + filter_upwards [h₂g] with z hz using by simp [hz] + -- Non-trivial case: both functions do not vanish around z₀ + rw [← WithTop.coe_add, order_eq_int_iff] + obtain ⟨F, h₁F, h₂F, h₃F⟩ := (hf.order_eq_int_iff _).1 h₂f + obtain ⟨G, h₁G, h₂G, h₃G⟩ := (hg.order_eq_int_iff _).1 h₂g + use F • G, h₁F.smul h₁G, by simp [h₂F, h₂G] + filter_upwards [self_mem_nhdsWithin, h₃F, h₃G] with a ha hfa hga + simp [hfa, hga, smul_comm (F a), zpow_add₀ (sub_ne_zero.mpr ha), mul_smul] + +/-- The order is additive when multiplying meromorphic functions. -/ +theorem order_mul {f g : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : + (hf.mul hg).order = hf.order + hg.order := + hf.order_smul hg + lemma iff_eventuallyEq_zpow_smul_analyticAt {f : 𝕜 → E} {x : 𝕜} : MeromorphicAt f x ↔ ∃ (n : ℤ) (g : 𝕜 → E), AnalyticAt 𝕜 g x ∧ ∀ᶠ z in 𝓝[≠] x, f z = (z - x) ^ n • g z := by refine ⟨fun ⟨n, hn⟩ ↦ ⟨-n, _, ⟨hn, eventually_nhdsWithin_iff.mpr ?_⟩⟩, ?_⟩ diff --git a/Mathlib/Analysis/Analytic/OfScalars.lean b/Mathlib/Analysis/Analytic/OfScalars.lean index 5674b07141092c..555c1c81c911a8 100644 --- a/Mathlib/Analysis/Analytic/OfScalars.lean +++ b/Mathlib/Analysis/Analytic/OfScalars.lean @@ -82,6 +82,11 @@ theorem ofScalars_apply_zero (n : ℕ) : rw [ofScalars] cases n <;> simp +@[simp] +lemma coeff_ofScalars {𝕜 : Type*} [NontriviallyNormedField 𝕜] {p : ℕ → 𝕜} {n : ℕ} : + (FormalMultilinearSeries.ofScalars 𝕜 p).coeff n = p n := by + simp [FormalMultilinearSeries.coeff, FormalMultilinearSeries.ofScalars, List.prod_ofFn] + theorem ofScalars_add (c' : ℕ → 𝕜) : ofScalars E (c + c') = ofScalars E c + ofScalars E c' := by unfold ofScalars simp_rw [Pi.add_apply, Pi.add_def _ _] diff --git a/Mathlib/Analysis/Analytic/Uniqueness.lean b/Mathlib/Analysis/Analytic/Uniqueness.lean index 6c836c8afcab6f..7bd05ac5c6aea5 100644 --- a/Mathlib/Analysis/Analytic/Uniqueness.lean +++ b/Mathlib/Analysis/Analytic/Uniqueness.lean @@ -183,7 +183,7 @@ theorem eqOn_zero_of_preconnected_of_eventuallyEq_zero_aux [CompleteSpace F] {f EMetric.mem_closure_iff.1 xu (r / 2) (ENNReal.half_pos hp.r_pos.ne') let q := p.changeOrigin (y - x) have has_series : HasFPowerSeriesOnBall f q y (r / 2) := by - have A : (‖y - x‖₊ : ℝ≥0∞) < r / 2 := by rwa [edist_comm, edist_eq_coe_nnnorm_sub] at hxy + have A : (‖y - x‖₊ : ℝ≥0∞) < r / 2 := by rwa [edist_comm, edist_eq_enorm_sub] at hxy have := hp.changeOrigin (A.trans_le ENNReal.half_le_self) simp only [add_sub_cancel] at this apply this.mono (ENNReal.half_pos hp.r_pos.ne') diff --git a/Mathlib/Analysis/Analytic/Within.lean b/Mathlib/Analysis/Analytic/Within.lean index fd9bd854f032d5..9ffa3528476c86 100644 --- a/Mathlib/Analysis/Analytic/Within.lean +++ b/Mathlib/Analysis/Analytic/Within.lean @@ -124,14 +124,14 @@ lemma hasFPowerSeriesWithinOnBall_iff_exists_hasFPowerSeriesOnBall [CompleteSpac · intro h refine ⟨fun y ↦ p.sum (y - x), ?_, ?_⟩ · intro y ⟨ys,yb⟩ - simp only [EMetric.mem_ball, edist_eq_coe_nnnorm_sub] at yb + simp only [EMetric.mem_ball, edist_eq_enorm_sub] at yb have e0 := p.hasSum (x := y - x) ?_ have e1 := (h.hasSum (y := y - x) ?_ ?_) · simp only [add_sub_cancel] at e1 exact e1.unique e0 · simpa only [add_sub_cancel] - · simpa only [EMetric.mem_ball, edist_eq_coe_nnnorm] - · simp only [EMetric.mem_ball, edist_eq_coe_nnnorm] + · simpa only [EMetric.mem_ball, edist_zero_eq_enorm] + · simp only [EMetric.mem_ball, edist_zero_eq_enorm] exact lt_of_lt_of_le yb h.r_le · refine ⟨h.r_le, h.r_pos, ?_⟩ intro y lt @@ -145,7 +145,7 @@ lemma hasFPowerSeriesWithinOnBall_iff_exists_hasFPowerSeriesOnBall [CompleteSpac rw [hfg] · exact hg.hasSum lt · refine ⟨ys, ?_⟩ - simpa only [EMetric.mem_ball, edist_eq_coe_nnnorm_sub, add_sub_cancel_left, sub_zero] using lt + simpa only [EMetric.mem_ball, edist_eq_enorm_sub, add_sub_cancel_left, sub_zero] using lt /-- `f` has power series `p` at `x` iff some local extension of `f` has that series -/ lemma hasFPowerSeriesWithinAt_iff_exists_hasFPowerSeriesAt [CompleteSpace F] {f : E → F} diff --git a/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean b/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean index e1c0e57e890c53..00668c051caf31 100644 --- a/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean @@ -3,9 +3,8 @@ Copyright (c) 2021 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ -import Mathlib.Analysis.Normed.Order.Basic -import Mathlib.Analysis.Asymptotics.Lemmas -import Mathlib.Analysis.Normed.Module.Basic +import Mathlib.Analysis.Asymptotics.AsymptoticEquivalent +import Mathlib.Analysis.SpecificLimits.Basic /-! # A collection of specific asymptotic results @@ -143,3 +142,19 @@ theorem Filter.Tendsto.cesaro {u : ℕ → ℝ} {l : ℝ} (h : Tendsto u atTop ( h.cesaro_smul end Real + +section NormedLinearOrderedField + +variable {R : Type*} [NormedLinearOrderedField R] [OrderTopology R] [FloorRing R] + +theorem Asymptotics.isEquivalent_nat_floor : + (fun (x : R) ↦ ↑⌊x⌋₊) ~[atTop] (fun x ↦ x) := by + refine isEquivalent_of_tendsto_one ?_ tendsto_nat_floor_div_atTop + filter_upwards with x hx using by rw [hx, Nat.floor_zero, Nat.cast_eq_zero] + +theorem Asymptotics.isEquivalent_nat_ceil : + (fun (x : R) ↦ ↑⌈x⌉₊) ~[atTop] (fun x ↦ x) := by + refine isEquivalent_of_tendsto_one ?_ tendsto_nat_ceil_div_atTop + filter_upwards with x hx using by rw [hx, Nat.ceil_zero, Nat.cast_eq_zero] + +end NormedLinearOrderedField diff --git a/Mathlib/Analysis/Asymptotics/TVS.lean b/Mathlib/Analysis/Asymptotics/TVS.lean index 5d0db2d9253a32..59720fc3eebddc 100644 --- a/Mathlib/Analysis/Asymptotics/TVS.lean +++ b/Mathlib/Analysis/Asymptotics/TVS.lean @@ -193,6 +193,17 @@ lemma IsLittleOTVS.insert [TopologicalSpace α] {x : α} {s : Set α} lemma IsLittleOTVS.bot : f =o[𝕜;⊥] g := fun u hU => ⟨univ, by simp⟩ +theorem IsLittleOTVS.add [TopologicalAddGroup E] [ContinuousSMul 𝕜 E] + {f₁ f₂ : α → E} {g : α → F} {l : Filter α} + (h₁ : f₁ =o[𝕜; l] g) (h₂ : f₂ =o[𝕜; l] g) : (f₁ + f₂) =o[𝕜; l] g := by + rw [(nhds_basis_balanced 𝕜 E).add_self.isLittleOTVS_iff (basis_sets _)] + rintro U ⟨hU, hUb⟩ + rcases ((h₁.eventually_smallSets U hU).and (h₂.eventually_smallSets U hU)).exists_mem_of_smallSets + with ⟨V, hV, hVf₁, hVf₂⟩ + refine ⟨V, hV, fun ε hε ↦ ?_⟩ + filter_upwards [hVf₁ ε hε, hVf₂ ε hε] with x hx₁ hx₂ + exact (egauge_add_add_le hUb hUb _ _).trans (max_le hx₁ hx₂) + protected lemma IsLittleOTVS.smul_left (h : f =o[𝕜;l] g) (c : α → 𝕜) : (fun x ↦ c x • f x) =o[𝕜;l] (fun x ↦ c x • g x) := by unfold IsLittleOTVS at * diff --git a/Mathlib/Analysis/BoxIntegral/Integrability.lean b/Mathlib/Analysis/BoxIntegral/Integrability.lean index a8a89c77948867..695710a0bc2224 100644 --- a/Mathlib/Analysis/BoxIntegral/Integrability.lean +++ b/Mathlib/Analysis/BoxIntegral/Integrability.lean @@ -223,7 +223,7 @@ theorem IntegrableOn.hasBoxIntegral [CompleteSpace E] {f : (ι → ℝ) → E} { -- Choose `N` such that the integral of `‖f N x - g x‖` is less than or equal to `ε`. obtain ⟨N₀, hN₀⟩ : ∃ N : ℕ, ∫ x in I, ‖f N x - g x‖ ∂μ ≤ ε := by have : Tendsto (fun n => ∫⁻ x in I, ‖f n x - g x‖₊ ∂μ) atTop (𝓝 0) := - SimpleFunc.tendsto_approxOn_range_L1_nnnorm hg.measurable hgi + SimpleFunc.tendsto_approxOn_range_L1_enorm hg.measurable hgi refine (this.eventually (ge_mem_nhds ε0')).exists.imp fun N hN => ?_ exact integral_coe_le_of_lintegral_coe_le hN -- For each `x`, we choose `Nx x ≥ N₀` such that `dist (f Nx x) (g x) ≤ ε`. diff --git a/Mathlib/Analysis/CStarAlgebra/CStarMatrix.lean b/Mathlib/Analysis/CStarAlgebra/CStarMatrix.lean new file mode 100644 index 00000000000000..27bbb72bc3611b --- /dev/null +++ b/Mathlib/Analysis/CStarAlgebra/CStarMatrix.lean @@ -0,0 +1,684 @@ +/- +Copyright (c) 2025 Frédéric Dupuis. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Frédéric Dupuis +-/ + +import Mathlib.Analysis.CStarAlgebra.Module.Constructions +import Mathlib.Analysis.Matrix +import Mathlib.Topology.UniformSpace.Matrix + +/-! +# Matrices with entries in a C⋆-algebra + +This file creates a type copy of `Matrix m n A` called `CStarMatrix m n A` meant for matrices with +entries in a C⋆-algebra `A`. Its action on `C⋆ᵐᵒᵈ (n → A)` (via `Matrix.mulVec`) gives +it the operator norm, and this norm makes `CStarMatrix n n A` a C⋆-algebra. + +## Main declarations + ++ `CStarMatrix m n A`: the type copy ++ `CStarMatrix.instNonUnitalCStarAlgebra`: square matrices with entries in a non-unital C⋆-algebra + form a non-unital C⋆-algebra ++ `CStarMatrix.instCStarAlgebra`: square matrices with entries in a unital C⋆-algebra form a + unital C⋆-algebra + +## Implementation notes + +The norm on this type induces the product uniformity and bornology, but these are not defeq to +`Pi.uniformSpace` and `Pi.instBornology`. Hence, we prove the equality to the Pi instances and +replace the uniformity and bornology by the Pi ones when registering the +`NormedAddCommGroup (CStarMatrix m n A)` instance. See the docstring of the `TopologyAux` section +below for more details. +-/ + +open scoped ComplexOrder Topology Uniformity Bornology Matrix NNReal InnerProductSpace + WithCStarModule + +/-- Type copy `Matrix m n A` meant for matrices with entries in a C⋆-algebra. This is +a C⋆-algebra when `m = n`. -/ +def CStarMatrix (m : Type*) (n : Type*) (A : Type*) := Matrix m n A + +namespace CStarMatrix + +variable {m n A B R S : Type*} + +section basic + +variable (m n A) in +/-- The equivalence between `Matrix m n A` and `CStarMatrix m n A`. -/ +def ofMatrix {m n A : Type*} : Matrix m n A ≃ CStarMatrix m n A := Equiv.refl _ + +@[simp] +lemma ofMatrix_apply {M : Matrix m n A} {i : m} : (ofMatrix M) i = M i := rfl + +@[simp] +lemma ofMatrix_symm_apply {M : CStarMatrix m n A} {i : m} : (ofMatrix.symm M) i = M i := rfl + +theorem ext_iff {M N : CStarMatrix m n A} : (∀ i j, M i j = N i j) ↔ M = N := + ⟨fun h => funext fun i => funext <| h i, fun h => by simp [h]⟩ + +@[ext] +lemma ext {M₁ M₂ : CStarMatrix m n A} (h : ∀ i j, M₁ i j = M₂ i j) : M₁ = M₂ := ext_iff.mp h + +/-- `M.map f` is the matrix obtained by applying `f` to each entry of the matrix `M`. -/ +def map (M : CStarMatrix m n A) (f : A → B) : CStarMatrix m n B := + ofMatrix fun i j => f (M i j) + +@[simp] +theorem map_apply {M : CStarMatrix m n A} {f : A → B} {i : m} {j : n} : M.map f i j = f (M i j) := + rfl + +@[simp] +theorem map_id (M : CStarMatrix m n A) : M.map id = M := by + ext + rfl + +@[simp] +theorem map_id' (M : CStarMatrix m n A) : M.map (·) = M := map_id M + +theorem map_map {C : Type*} {M : Matrix m n A} {f : A → B} {g : B → C} : + (M.map f).map g = M.map (g ∘ f) := by ext; rfl + +theorem map_injective {f : A → B} (hf : Function.Injective f) : + Function.Injective fun M : CStarMatrix m n A => M.map f := fun _ _ h => + ext fun i j => hf <| ext_iff.mpr h i j + +/-- The transpose of a matrix. -/ +def transpose (M : CStarMatrix m n A) : CStarMatrix n m A := + ofMatrix fun x y => M y x + +@[simp] +theorem transpose_apply (M : CStarMatrix m n A) (i j) : transpose M i j = M j i := + rfl + +/-- The conjugate transpose of a matrix defined in term of `star`. -/ +def conjTranspose [Star A] (M : CStarMatrix m n A) : CStarMatrix n m A := + M.transpose.map star + +instance instStar [Star A] : Star (CStarMatrix n n A) where + star M := M.conjTranspose + +instance instInvolutiveStar [InvolutiveStar A] : InvolutiveStar (CStarMatrix n n A) where + star_involutive := star_involutive (R := Matrix n n A) + +instance instInhabited [Inhabited A] : Inhabited (CStarMatrix m n A) := + inferInstanceAs <| Inhabited <| m → n → A + +instance instDecidableEq [DecidableEq A] [Fintype m] [Fintype n] : + DecidableEq (CStarMatrix m n A) := + Fintype.decidablePiFintype + +instance {n m} [Fintype m] [DecidableEq m] [Fintype n] [DecidableEq n] (α) [Fintype α] : + Fintype (CStarMatrix m n α) := inferInstanceAs (Fintype (m → n → α)) + +instance {n m} [Finite m] [Finite n] (α) [Finite α] : + Finite (CStarMatrix m n α) := inferInstanceAs (Finite (m → n → α)) + +instance instAdd [Add A] : Add (CStarMatrix m n A) := + Pi.instAdd + +instance instAddSemigroup [AddSemigroup A] : AddSemigroup (CStarMatrix m n A) := + Pi.addSemigroup + +instance instAddCommSemigroup [AddCommSemigroup A] : AddCommSemigroup (CStarMatrix m n A) := + Pi.addCommSemigroup + +instance instZero [Zero A] : Zero (CStarMatrix m n A) := + Pi.instZero + +instance instAddZeroClass [AddZeroClass A] : AddZeroClass (CStarMatrix m n A) := + Pi.addZeroClass + +instance instAddMonoid [AddMonoid A] : AddMonoid (CStarMatrix m n A) := + Pi.addMonoid + +instance instAddCommMonoid [AddCommMonoid A] : AddCommMonoid (CStarMatrix m n A) := + Pi.addCommMonoid + +instance instNeg [Neg A] : Neg (CStarMatrix m n A) := + Pi.instNeg + +instance instSub [Sub A] : Sub (CStarMatrix m n A) := + Pi.instSub + +instance instAddGroup [AddGroup A] : AddGroup (CStarMatrix m n A) := + Pi.addGroup + +instance instAddCommGroup [AddCommGroup A] : AddCommGroup (CStarMatrix m n A) := + Pi.addCommGroup + +instance instUnique [Unique A] : Unique (CStarMatrix m n A) := + Pi.unique + +instance instSubsingleton [Subsingleton A] : Subsingleton (CStarMatrix m n A) := + inferInstanceAs <| Subsingleton <| m → n → A + +instance instNontrivial [Nonempty m] [Nonempty n] [Nontrivial A] : Nontrivial (CStarMatrix m n A) := + Function.nontrivial + +instance instSMul [SMul R A] : SMul R (CStarMatrix m n A) := + Pi.instSMul + +instance instSMulCommClass [SMul R A] [SMul S A] [SMulCommClass R S A] : + SMulCommClass R S (CStarMatrix m n A) := + Pi.smulCommClass + +instance instIsScalarTower [SMul R S] [SMul R A] [SMul S A] [IsScalarTower R S A] : + IsScalarTower R S (CStarMatrix m n A) := + Pi.isScalarTower + +instance instIsCentralScalar [SMul R A] [SMul Rᵐᵒᵖ A] [IsCentralScalar R A] : + IsCentralScalar R (CStarMatrix m n A) := + Pi.isCentralScalar + +instance instMulAction [Monoid R] [MulAction R A] : MulAction R (CStarMatrix m n A) := + Pi.mulAction _ + +instance instDistribMulAction [Monoid R] [AddMonoid A] [DistribMulAction R A] : + DistribMulAction R (CStarMatrix m n A) := + Pi.distribMulAction _ + +instance instModule [Semiring R] [AddCommMonoid A] [Module R A] : Module R (CStarMatrix m n A) := + Pi.module _ _ _ + +@[simp, nolint simpNF] +theorem zero_apply [Zero A] (i : m) (j : n) : (0 : CStarMatrix m n A) i j = 0 := rfl + +@[simp] theorem add_apply [Add A] (M N : CStarMatrix m n A) (i : m) (j : n) : + (M + N) i j = (M i j) + (N i j) := rfl + +@[simp] theorem smul_apply [SMul B A] (r : B) (M : Matrix m n A) (i : m) (j : n) : + (r • M) i j = r • (M i j) := rfl + +@[simp] theorem sub_apply [Sub A] (M N : Matrix m n A) (i : m) (j : n) : + (M - N) i j = (M i j) - (N i j) := rfl + +@[simp] theorem neg_apply [Neg A] (M : Matrix m n A) (i : m) (j : n) : + (-M) i j = -(M i j) := rfl + +/-! simp-normal form pulls `of` to the outside, to match the `Matrix` API. -/ + +@[simp] theorem of_zero [Zero A] : ofMatrix (0 : Matrix m n A) = 0 := rfl + +@[simp] theorem of_add_of [Add A] (f g : Matrix m n A) : + ofMatrix f + ofMatrix g = ofMatrix (f + g) := rfl + +@[simp] +theorem of_sub_of [Sub A] (f g : Matrix m n A) : ofMatrix f - ofMatrix g = ofMatrix (f - g) := + rfl + +@[simp] theorem neg_of [Neg A] (f : Matrix m n A) : -ofMatrix f = ofMatrix (-f) := rfl + +@[simp] theorem smul_of [SMul R A] (r : R) (f : Matrix m n A) : + r • ofMatrix f = ofMatrix (r • f) := rfl + +instance instStarAddMonoid [AddMonoid A] [StarAddMonoid A] : StarAddMonoid (CStarMatrix n n A) where + star_add := star_add (R := Matrix n n A) + +instance instStarModule [Star R] [Star A] [SMul R A] [StarModule R A] : + StarModule R (CStarMatrix n n A) where + star_smul := star_smul (A := Matrix n n A) + +/-- The equivalence to matrices, bundled as a linear equivalence. -/ +def ofMatrixₗ [AddCommMonoid A] [Semiring R] [Module R A] : + (Matrix m n A) ≃ₗ[R] CStarMatrix m n A := LinearEquiv.refl _ _ + +section zero_one + +variable [Zero A] [One A] [DecidableEq n] + +instance instOne : One (CStarMatrix n n A) := inferInstanceAs <| One (Matrix n n A) + +theorem one_apply {i j} : (1 : CStarMatrix n n A) i j = if i = j then 1 else 0 := rfl + +@[simp] +theorem one_apply_eq (i) : (1 : CStarMatrix n n A) i i = 1 := Matrix.one_apply_eq _ + +@[simp] theorem one_apply_ne {i j} : i ≠ j → (1 : CStarMatrix n n A) i j = 0 := Matrix.one_apply_ne + +theorem one_apply_ne' {i j} : j ≠ i → (1 : CStarMatrix n n A) i j = 0 := Matrix.one_apply_ne' + +instance instAddMonoidWithOne [AddMonoidWithOne A] : AddMonoidWithOne (CStarMatrix n n A) where + +instance instAddGroupWithOne [AddGroupWithOne A] : AddGroupWithOne (CStarMatrix n n A) where + __ := instAddGroup + __ := instAddMonoidWithOne + +instance instAddCommMonoidWithOne [AddCommMonoidWithOne A] : + AddCommMonoidWithOne (CStarMatrix n n A) where + __ := instAddCommMonoid + __ := instAddMonoidWithOne + +instance instAddCommGroupWithOne [AddCommGroupWithOne A] : + AddCommGroupWithOne (CStarMatrix n n A) where + __ := instAddCommGroup + __ := instAddGroupWithOne + +-- We want to be lower priority than `instHMul`, but without this we can't have operands with +-- implicit dimensions. +@[default_instance 100] +instance {l : Type*} [Fintype m] [Mul A] [AddCommMonoid A] : + HMul (CStarMatrix l m A) (CStarMatrix m n A) (CStarMatrix l n A) where + hMul M N := ofMatrix (ofMatrix.symm M * ofMatrix.symm N) + +instance [Fintype n] [Mul A] [AddCommMonoid A] : Mul (CStarMatrix n n A) where mul M N := M * N + +end zero_one + +theorem mul_apply {l : Type*} [Fintype m] [Mul A] [AddCommMonoid A] {M : CStarMatrix l m A} + {N : CStarMatrix m n A} {i k} : (M * N) i k = ∑ j, M i j * N j k := rfl + +theorem mul_apply' {l : Type*} [Fintype m] [Mul A] [AddCommMonoid A] {M : CStarMatrix l m A} + {N : CStarMatrix m n A} {i k} : (M * N) i k = (fun j => M i j) ⬝ᵥ fun j => N j k := rfl + +@[simp] +theorem smul_mul {l : Type*} [Fintype n] [Monoid R] [AddCommMonoid A] [Mul A] [DistribMulAction R A] + [IsScalarTower R A A] (a : R) (M : CStarMatrix m n A) (N : CStarMatrix n l A) : + (a • M) * N = a • (M * N) := Matrix.smul_mul a M N + +theorem mul_smul {l : Type*} [Fintype n] [Monoid R] [AddCommMonoid A] [Mul A] [DistribMulAction R A] + [SMulCommClass R A A] (M : Matrix m n A) (a : R) (N : Matrix n l A) : + M * (a • N) = a • (M * N) := Matrix.mul_smul M a N + +instance instNonUnitalNonAssocSemiring [Fintype n] [NonUnitalNonAssocSemiring A] : + NonUnitalNonAssocSemiring (CStarMatrix n n A) := + inferInstanceAs <| NonUnitalNonAssocSemiring (Matrix n n A) + +instance instNonUnitalNonAssocRing [Fintype n] [NonUnitalNonAssocRing A] : + NonUnitalNonAssocRing (CStarMatrix n n A) := + inferInstanceAs <| NonUnitalNonAssocRing (Matrix n n A) + +instance instNonUnitalSemiring [Fintype n] [NonUnitalSemiring A] : + NonUnitalSemiring (CStarMatrix n n A) := + inferInstanceAs <| NonUnitalSemiring (Matrix n n A) + +instance instNonAssocSemiring [Fintype n] [DecidableEq n] [NonAssocSemiring A] : + NonAssocSemiring (CStarMatrix n n A) := + inferInstanceAs <| NonAssocSemiring (Matrix n n A) + +instance instNonUnitalRing [Fintype n] [NonUnitalRing A] : + NonUnitalRing (CStarMatrix n n A) := + inferInstanceAs <| NonUnitalRing (Matrix n n A) + +instance instNonAssocRing [Fintype n] [DecidableEq n] [NonAssocRing A] : + NonAssocRing (CStarMatrix n n A) := + inferInstanceAs <| NonAssocRing (Matrix n n A) + +instance instSemiring [Fintype n] [DecidableEq n] [Semiring A] : + Semiring (CStarMatrix n n A) := + inferInstanceAs <| Semiring (Matrix n n A) + +instance instRing [Fintype n] [DecidableEq n] [Ring A] : Ring (CStarMatrix n n A) := + inferInstanceAs <| Ring (Matrix n n A) + +/-- `ofMatrix` bundled as a ring equivalence. -/ +def ofMatrixRingEquiv [Fintype n] [Semiring A] : + Matrix n n A ≃+* CStarMatrix n n A := + { ofMatrix with + map_mul' := fun _ _ => rfl + map_add' := fun _ _ => rfl } + +instance instStarRing [Fintype n] [NonUnitalSemiring A] [StarRing A] : + StarRing (CStarMatrix n n A) := inferInstanceAs <| StarRing (Matrix n n A) + +instance instAlgebra [Fintype n] [DecidableEq n] [CommSemiring R] [Semiring A] [Algebra R A] : + Algebra R (CStarMatrix n n A) := inferInstanceAs <| Algebra R (Matrix n n A) + +/-- `ofMatrix` bundled as a star algebra equivalence. -/ +def ofMatrixStarAlgEquiv [Fintype n] [SMul ℂ A] [Semiring A] [StarRing A] : + Matrix n n A ≃⋆ₐ[ℂ] CStarMatrix n n A := + { ofMatrixRingEquiv with + map_star' := fun _ => rfl + map_smul' := fun _ _ => rfl } + +lemma ofMatrix_eq_ofMatrixStarAlgEquiv [Fintype n] [SMul ℂ A] [Semiring A] [StarRing A] : + (ofMatrix : Matrix n n A → CStarMatrix n n A) + = (ofMatrixStarAlgEquiv : Matrix n n A → CStarMatrix n n A) := rfl + +end basic + +variable [Fintype m] [Fintype n] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] + + +/-- Interpret a `CStarMatrix m n A` as a continuous linear map acting on `C⋆ᵐᵒᵈ (n → A)`. -/ +def toCLM : CStarMatrix m n A →ₗ[ℂ] (C⋆ᵐᵒᵈ (n → A)) →L[ℂ] (C⋆ᵐᵒᵈ (m → A)) where + toFun M := { toFun := (WithCStarModule.equivL ℂ).symm ∘ M.mulVec ∘ WithCStarModule.equivL ℂ + map_add' := M.mulVec_add + map_smul' := M.mulVec_smul + cont := by + simp only [LinearMap.coe_mk, AddHom.coe_mk] + exact Continuous.comp (by fun_prop) (by fun_prop) } + map_add' M₁ M₂ := by + ext + simp only [ContinuousLinearMap.coe_mk', LinearMap.coe_mk, AddHom.coe_mk, Function.comp_apply, + WithCStarModule.equivL_apply, WithCStarModule.equivL_symm_apply, + WithCStarModule.equiv_symm_pi_apply, ContinuousLinearMap.add_apply, WithCStarModule.add_apply] + rw [Matrix.add_mulVec, Pi.add_apply] + map_smul' c M := by + ext x i + simp only [ContinuousLinearMap.coe_mk', LinearMap.coe_mk, AddHom.coe_mk, Function.comp_apply, + WithCStarModule.equivL_apply, WithCStarModule.equivL_symm_apply, + WithCStarModule.equiv_symm_pi_apply, Matrix.mulVec, dotProduct, + WithCStarModule.equiv_pi_apply, RingHom.id_apply, ContinuousLinearMap.coe_smul', + Pi.smul_apply, WithCStarModule.smul_apply, Finset.smul_sum] + congr + ext j + rw [CStarMatrix.smul_apply, smul_mul_assoc] + +lemma toCLM_apply {M : CStarMatrix m n A} {v : C⋆ᵐᵒᵈ (n → A)} : + toCLM M v = (WithCStarModule.equiv _).symm (M.mulVec v) := rfl + +lemma toCLM_apply_eq_sum {M : CStarMatrix m n A} {v : C⋆ᵐᵒᵈ (n → A)} : + toCLM M v = (WithCStarModule.equiv _).symm (fun i => ∑ j, M i j * v j) := by + ext i + simp [toCLM_apply, Matrix.mulVec, dotProduct] + +/-- Interpret a `CStarMatrix m n A` as a continuous linear map acting on `C⋆ᵐᵒᵈ (n → A)`. This +version is specialized to the case `m = n` and is bundled as a non-unital algebra homomorphism. -/ +def toCLMNonUnitalAlgHom : CStarMatrix n n A →ₙₐ[ℂ] (C⋆ᵐᵒᵈ (n → A)) →L[ℂ] (C⋆ᵐᵒᵈ (n → A)) := + { toCLM (n := n) (m := n) with + map_zero' := by ext1; simp + map_mul' := by intros; ext; simp [toCLM] } + +lemma toCLMNonUnitalAlgHom_eq_toCLM {M : CStarMatrix n n A} : + toCLMNonUnitalAlgHom (A := A) M = toCLM M := rfl + +open WithCStarModule in +@[simp high] +lemma toCLM_apply_single [DecidableEq n] {M : CStarMatrix m n A} {j : n} (a : A) : + (toCLM M) (equiv _ |>.symm <| Pi.single j a) = (equiv _).symm (fun i => M i j * a) := by + ext + simp [toCLM_apply, EmbeddingLike.apply_eq_iff_eq, equiv, Equiv.refl] + +open WithCStarModule in +lemma toCLM_apply_single_apply [DecidableEq n] {M : CStarMatrix m n A} {i : m} {j : n} (a : A) : + (toCLM M) (equiv _ |>.symm <| Pi.single j a) i = M i j * a := by simp + +open WithCStarModule in +lemma mul_entry_mul_eq_inner_toCLM [DecidableEq m] [DecidableEq n] {M : CStarMatrix m n A} + {i : m} {j : n} (a b : A) : + star a * M i j * b + = ⟪equiv _ |>.symm (Pi.single i a), toCLM M (equiv _ |>.symm <| Pi.single j b)⟫_A := by + simp [mul_assoc, inner_def] + +lemma toCLM_injective [DecidableEq n] : Function.Injective (toCLM (A := A) (m := m) (n := n)) := by + rw [injective_iff_map_eq_zero] + intro M h + ext i j + rw [Matrix.zero_apply, ← norm_eq_zero, ← sq_eq_zero_iff, sq, ← CStarRing.norm_self_mul_star, + ← toCLM_apply_single_apply] + simp [h] + +open WithCStarModule in +lemma inner_toCLM_conjTranspose_left {M : CStarMatrix m n A} {v : C⋆ᵐᵒᵈ (m → A)} + {w : C⋆ᵐᵒᵈ (n → A)} : ⟪toCLM Mᴴ v, w⟫_A = ⟪v, toCLM M w⟫_A := by + simp only [toCLM_apply_eq_sum, pi_inner, equiv_symm_pi_apply, inner_def, Finset.mul_sum, + Matrix.conjTranspose_apply, star_sum, star_mul, star_star, Finset.sum_mul] + rw [Finset.sum_comm] + simp_rw [mul_assoc] + +lemma inner_toCLM_conjTranspose_right {M : CStarMatrix m n A} {v : C⋆ᵐᵒᵈ (n → A)} + {w : C⋆ᵐᵒᵈ (m → A)} : ⟪v, toCLM Mᴴ w⟫_A = ⟪toCLM M v, w⟫_A := by + apply Eq.symm + simpa using inner_toCLM_conjTranspose_left (M := Mᴴ) (v := v) (w := w) + +/-- The operator norm on `CStarMatrix m n A`. -/ +noncomputable instance instNorm : Norm (CStarMatrix m n A) where + norm M := ‖toCLM M‖ + +lemma norm_def {M : CStarMatrix m n A} : ‖M‖ = ‖toCLM M‖ := rfl + +lemma norm_def' {M : CStarMatrix n n A} : ‖M‖ = ‖toCLMNonUnitalAlgHom (A := A) M‖ := rfl + +set_option maxSynthPendingDepth 2 in +lemma normedSpaceCore [DecidableEq n]: NormedSpace.Core ℂ (CStarMatrix m n A) where + norm_nonneg M := (toCLM M).opNorm_nonneg + norm_smul c M := by rw [norm_def, norm_def, map_smul, norm_smul _ (toCLM M)] + norm_triangle M₁ M₂ := by simpa [← map_add] using norm_add_le (toCLM M₁) (toCLM M₂) + norm_eq_zero_iff := by + simpa only [norm_def, norm_eq_zero, ← injective_iff_map_eq_zero'] using toCLM_injective + +open WithCStarModule in +lemma norm_entry_le_norm [DecidableEq n] {M : CStarMatrix m n A} {i : m} {j : n} : + ‖M i j‖ ≤ ‖M‖ := by + suffices ‖M i j‖ * ‖M i j‖ ≤ ‖M‖ * ‖M i j‖ by + obtain (h | h) := eq_zero_or_norm_pos (M i j) + · set_option maxSynthPendingDepth 2 in + simp [h, norm_def] + · exact le_of_mul_le_mul_right this h + rw [← CStarRing.norm_self_mul_star, ← toCLM_apply_single_apply] + apply norm_apply_le_norm _ _ |>.trans + apply (toCLM M).le_opNorm _ |>.trans + simp [norm_def] + +open CStarModule in +lemma norm_le_of_forall_inner_le {M : CStarMatrix m n A} {C : ℝ≥0} + (h : ∀ v w, ‖⟪w, toCLM M v⟫_A‖ ≤ C * ‖v‖ * ‖w‖) : ‖M‖ ≤ C := by + refine (toCLM M).opNorm_le_bound (by simp) fun v ↦ ?_ + obtain (h₀ | h₀) := (norm_nonneg (toCLM M v)).eq_or_lt + · rw [← h₀] + positivity + · refine le_of_mul_le_mul_right ?_ h₀ + simpa [← sq, norm_sq_eq] using h .. + +end CStarMatrix + +section TopologyAux +/- +## Replacing the uniformity and bornology + +Note that while the norm that we have defined on `CStarMatrix m n A` induces the product uniformity, +it is not defeq to `Pi.uniformSpace`. In this section, we show that the norm indeed does induce +the product topology and use this fact to properly set up the +`NormedAddCommGroup (CStarMatrix m n A)` instance such that the uniformity is still +`Pi.uniformSpace` and the bornology is `Pi.instBornology`. + +To do this, we locally register a `NormedAddCommGroup` instance on `CStarMatrix` which registers +the "bad" topology, and we also locally use the matrix norm `Matrix.normedAddCommGroup` +(which takes the norm of the biggest entry as the norm of the matrix) +in order to show that the map `ofMatrix` is bilipschitz. We then finally register the +`NormedAddCommGroup (C⋆ᵐᵒᵈ (n → A))` instance via `NormedAddCommGroup.ofCoreReplaceAll`. +-/ + +namespace CStarMatrix + +variable {m n A : Type*} [Fintype m] [Fintype n] [DecidableEq n] + [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] + +private noncomputable def normedAddCommGroupAux : NormedAddCommGroup (CStarMatrix m n A) := + .ofCore CStarMatrix.normedSpaceCore + +attribute [local instance] normedAddCommGroupAux + +private def normedSpaceAux : NormedSpace ℂ (CStarMatrix m n A) := + .ofCore CStarMatrix.normedSpaceCore + +/- In this `Aux` section, we locally activate the following instances: a norm on `CStarMatrix` + which induces a topology that is not defeq with the matrix one, and the elementwise norm on + matrices, in order to show that the two topologies are in fact equal -/ +attribute [local instance] normedSpaceAux Matrix.normedAddCommGroup Matrix.normedSpace + +private lemma nnnorm_le_of_forall_inner_le {M : CStarMatrix m n A} {C : ℝ≥0} + (h : ∀ v w, ‖⟪w, CStarMatrix.toCLM M v⟫_A‖₊ ≤ C * ‖v‖₊ * ‖w‖₊) : ‖M‖₊ ≤ C := + CStarMatrix.norm_le_of_forall_inner_le fun v w => h v w + +open Finset in +private lemma lipschitzWith_toMatrixAux [DecidableEq m] : + LipschitzWith 1 (ofMatrixₗ.symm (R := ℂ) : CStarMatrix m n A → Matrix m n A) := by + refine AddMonoidHomClass.lipschitz_of_bound_nnnorm _ _ fun M => ?_ + rw [one_mul, ← NNReal.coe_le_coe, coe_nnnorm, coe_nnnorm, Matrix.norm_le_iff (norm_nonneg _)] + exact fun _ _ ↦ CStarMatrix.norm_entry_le_norm + +open CStarMatrix WithCStarModule in +private lemma antilipschitzWith_toMatrixAux : + AntilipschitzWith (Fintype.card m * Fintype.card n) + (ofMatrixₗ.symm (R := ℂ) : CStarMatrix m n A → Matrix m n A) := by + refine AddMonoidHomClass.antilipschitz_of_bound _ fun M => ?_ + calc + ‖M‖ ≤ ∑ i, ∑ j, ‖M i j‖ := by + rw [norm_def] + refine (toCLM M).opNorm_le_bound (by positivity) fun v => ?_ + simp only [toCLM_apply_eq_sum, equiv_symm_pi_apply, Finset.sum_mul] + apply pi_norm_le_sum_norm _ |>.trans + gcongr with i _ + apply norm_sum_le _ _ |>.trans + gcongr with j _ + apply norm_mul_le _ _ |>.trans + gcongr + exact norm_apply_le_norm v j + _ ≤ ∑ _ : m, ∑ _ : n, ‖ofMatrixₗ.symm (R := ℂ) M‖ := by + gcongr with i _ j _ + exact ofMatrixₗ.symm (R := ℂ) M |>.norm_entry_le_entrywise_sup_norm + _ = _ := by simp [mul_assoc] + +private lemma uniformInducing_toMatrixAux [DecidableEq m] : + IsUniformInducing (ofMatrix.symm : CStarMatrix m n A → Matrix m n A) := + AntilipschitzWith.isUniformInducing antilipschitzWith_toMatrixAux + lipschitzWith_toMatrixAux.uniformContinuous + +private lemma uniformity_eq_aux [DecidableEq m] : + 𝓤 (CStarMatrix m n A) = (𝓤[Pi.uniformSpace _] : + Filter (CStarMatrix m n A × CStarMatrix m n A)) := by + have : + (fun x : CStarMatrix m n A × CStarMatrix m n A => ⟨ofMatrix.symm x.1, ofMatrix.symm x.2⟩) + = id := by + ext i <;> rfl + rw [← uniformInducing_toMatrixAux.comap_uniformity, this, Filter.comap_id] + rfl + +open Bornology in +private lemma cobounded_eq_aux [DecidableEq m] : + cobounded (CStarMatrix m n A) = @cobounded _ Pi.instBornology := by + have : cobounded (CStarMatrix m n A) = Filter.comap ofMatrix.symm (cobounded _) := by + refine le_antisymm ?_ ?_ + · exact antilipschitzWith_toMatrixAux.tendsto_cobounded.le_comap + · exact lipschitzWith_toMatrixAux.comap_cobounded_le + exact this.trans Filter.comap_id + +end CStarMatrix + +end TopologyAux + +namespace CStarMatrix + +section non_unital + +variable {A : Type*} [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] + +variable {m n : Type*} [Fintype m] [Fintype n] [DecidableEq m] + +instance instTopologicalSpace : TopologicalSpace (CStarMatrix m n A) := Pi.topologicalSpace +instance instUniformSpace : UniformSpace (CStarMatrix m n A) := Pi.uniformSpace _ +instance instBornology : Bornology (CStarMatrix m n A) := Pi.instBornology +instance instCompleteSpace : CompleteSpace (CStarMatrix m n A) := Pi.complete _ +instance instT2Space : T2Space (CStarMatrix m n A) := Pi.t2Space +instance instT3Space : T3Space (CStarMatrix m n A) := _root_.instT3Space + +instance instTopologicalAddGroup : TopologicalAddGroup (CStarMatrix m n A) := + Pi.topologicalAddGroup + +instance instUniformAddGroup : UniformAddGroup (CStarMatrix m n A) := + Pi.instUniformAddGroup + +instance instContinuousSMul {R : Type*} [SMul R A] [TopologicalSpace R] [ContinuousSMul R A] : + ContinuousSMul R (CStarMatrix m n A) := instContinuousSMulForall + +noncomputable instance instNormedAddCommGroup [DecidableEq n] : + NormedAddCommGroup (CStarMatrix m n A) := + .ofCoreReplaceAll CStarMatrix.normedSpaceCore + CStarMatrix.uniformity_eq_aux.symm + fun _ => Filter.ext_iff.1 CStarMatrix.cobounded_eq_aux.symm _ + +instance instNormedSpace [DecidableEq n] : NormedSpace ℂ (CStarMatrix m n A) := + .ofCore CStarMatrix.normedSpaceCore + +noncomputable instance instNonUnitalNormedRing [DecidableEq n] : + NonUnitalNormedRing (CStarMatrix n n A) where + dist_eq _ _ := rfl + norm_mul _ _ := by + set_option maxSynthPendingDepth 2 in + simpa only [norm_def', map_mul] using norm_mul_le _ _ + +open ContinuousLinearMap CStarModule in +/-- Matrices with entries in a C⋆-algebra form a C⋆-algebra. -/ +instance instCStarRing [DecidableEq n] : CStarRing (CStarMatrix n n A) where + norm_mul_self_le M := by + have hmain : ‖M‖ ≤ √‖star M * M‖ := by + change ‖toCLM M‖ ≤ √‖star M * M‖ + rw [opNorm_le_iff (by positivity)] + intro v + rw [norm_eq_sqrt_norm_inner_self, ← inner_toCLM_conjTranspose_right] + have h₁ : ‖⟪v, (toCLM Mᴴ) ((toCLM M) v)⟫_A‖ ≤ ‖star M * M‖ * ‖v‖ ^ 2 := calc + _ ≤ ‖v‖ * ‖(toCLM Mᴴ) (toCLM M v)‖ := norm_inner_le (C⋆ᵐᵒᵈ (n → A)) + _ ≤ ‖v‖ * ‖(toCLM Mᴴ).comp (toCLM M)‖ * ‖v‖ := by + rw [mul_assoc] + gcongr + rw [← ContinuousLinearMap.comp_apply] + exact le_opNorm ((toCLM Mᴴ).comp (toCLM M)) v + _ = ‖(toCLM Mᴴ).comp (toCLM M)‖ * ‖v‖ ^ 2 := by ring + _ = ‖star M * M‖ * ‖v‖ ^ 2 := by + congr + simp only [← toCLMNonUnitalAlgHom_eq_toCLM, Matrix.star_eq_conjTranspose, + map_mul] + rfl + have h₂ : ‖v‖ = √(‖v‖ ^ 2) := by simp + rw [h₂, ← Real.sqrt_mul] + gcongr + positivity + rw [← pow_two, ← Real.sqrt_le_sqrt_iff (by positivity)] + simp [hmain] + +/-- Matrices with entries in a non-unital C⋆-algebra form a non-unital C⋆-algebra. -/ +noncomputable instance instNonUnitalCStarAlgebra [DecidableEq n] : + NonUnitalCStarAlgebra (CStarMatrix n n A) where + smul_assoc x y z := by simp + smul_comm m a b := (Matrix.mul_smul _ _ _).symm + +instance instPartialOrder [DecidableEq n] : + PartialOrder (CStarMatrix n n A) := CStarAlgebra.spectralOrder _ +instance instStarOrderedRing [DecidableEq n] : + StarOrderedRing (CStarMatrix n n A) := CStarAlgebra.spectralOrderedRing _ + +end non_unital + +section unital + +variable {A : Type*} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] + +variable {n : Type*} [Fintype n] [DecidableEq n] + +noncomputable instance instNormedRing : NormedRing (CStarMatrix n n A) where + dist_eq _ _ := rfl + norm_mul := norm_mul_le + +noncomputable instance instNormedAlgebra : NormedAlgebra ℂ (CStarMatrix n n A) where + norm_smul_le r M := by simpa only [norm_def, map_smul] using (toCLM M).opNorm_smul_le r + +/-- Matrices with entries in a unital C⋆-algebra form a unital C⋆-algebra. -/ +noncomputable instance instCStarAlgebra [DecidableEq n] : CStarAlgebra (CStarMatrix n n A) where + +end unital + +section + +variable {m n A : Type*} [NonUnitalCStarAlgebra A] + +lemma uniformEmbedding_ofMatrix : + IsUniformEmbedding (ofMatrix : Matrix m n A → CStarMatrix m n A) where + comap_uniformity := Filter.comap_id' + injective := fun ⦃_ _⦄ a ↦ a + +/-- `ofMatrix` bundled as a continuous linear equivalence. -/ +def ofMatrixL : Matrix m n A ≃L[ℂ] CStarMatrix m n A := + { ofMatrixₗ with + continuous_toFun := continuous_id + continuous_invFun := continuous_id } + +lemma ofMatrix_eq_ofMatrixL : + (ofMatrix : Matrix m n A → CStarMatrix m n A) + = (ofMatrixL : Matrix m n A → CStarMatrix m n A) := rfl + +end + +end CStarMatrix diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index c093bde67133cd..62ae49c6840de3 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -291,7 +291,8 @@ lemma le_iff_norm_sqrt_mul_rpow {a b : A} (hbu : IsUnit b) (ha : 0 ≤ a) (hb : lift b to Aˣ using hbu have hbab : 0 ≤ (b : A) ^ (-(1 / 2) : ℝ) * a * (b : A) ^ (-(1 / 2) : ℝ) := conjugate_nonneg_of_nonneg ha rpow_nonneg - #adaptation_note /-- 2024-11-10 added `(R := A)` -/ + #adaptation_note /-- 2024-11-10 + added `(R := A)` -/ conv_rhs => rw [← sq_le_one_iff₀ (norm_nonneg _), sq, ← CStarRing.norm_star_mul_self, star_mul, IsSelfAdjoint.of_nonneg (R := A) sqrt_nonneg, IsSelfAdjoint.of_nonneg rpow_nonneg, diff --git a/Mathlib/Analysis/CStarAlgebra/Module/Constructions.lean b/Mathlib/Analysis/CStarAlgebra/Module/Constructions.lean index 64d2ba88bc9c29..872bf20de1001d 100644 --- a/Mathlib/Analysis/CStarAlgebra/Module/Constructions.lean +++ b/Mathlib/Analysis/CStarAlgebra/Module/Constructions.lean @@ -296,6 +296,13 @@ private lemma isBounded_pi_iff_aux (s : Set (C⋆ᵐᵒᵈ (Π i, E i))) : end Aux +noncomputable instance : PseudoMetricSpace (C⋆ᵐᵒᵈ (Π i, E i)) := + .ofSeminormedAddCommGroupCoreReplaceAll + normedSpaceCore.toCore uniformity_pi_eq_aux isBounded_pi_iff_aux + +noncomputable instance : SeminormedAddCommGroup (C⋆ᵐᵒᵈ (Π i, E i)) := + .ofCoreReplaceAll normedSpaceCore.toCore uniformity_pi_eq_aux isBounded_pi_iff_aux + noncomputable instance : NormedAddCommGroup (C⋆ᵐᵒᵈ (Π i, E i)) := .ofCoreReplaceAll normedSpaceCore uniformity_pi_eq_aux isBounded_pi_iff_aux diff --git a/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean b/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean index 95c7421f62a168..02b57dc8603b2a 100644 --- a/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean +++ b/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean @@ -6,6 +6,8 @@ Authors: Jireh Loreaux import Mathlib.RingTheory.Finiteness.Defs import Mathlib.Topology.Bornology.Constructions import Mathlib.Topology.UniformSpace.Equiv +import Mathlib.Topology.Algebra.Module.Equiv +import Mathlib.Topology.Algebra.UniformGroup.Basic /-! # Type synonym for types with a `CStarModule` structure @@ -75,11 +77,21 @@ instance instUnique [Unique E] : Unique (WithCStarModule E) := ‹Unique E› /-! ## `WithCStarModule E` inherits various module-adjacent structures from `E`. -/ +instance instZero [Zero E] : Zero (WithCStarModule E) := ‹Zero E› +instance instAdd [Add E] : Add (WithCStarModule E) := ‹Add E› +instance instSub [Sub E] : Sub (WithCStarModule E) := ‹Sub E› +instance instNeg [Neg E] : Neg (WithCStarModule E) := ‹Neg E› +instance instAddMonoid [AddMonoid E] : AddMonoid (WithCStarModule E) := ‹AddMonoid E› +instance instSubNegMonoid [SubNegMonoid E] : SubNegMonoid (WithCStarModule E) := ‹SubNegMonoid E› +instance instSubNegZeroMonoid [SubNegZeroMonoid E] : SubNegZeroMonoid (WithCStarModule E) := + ‹SubNegZeroMonoid E› + instance instAddCommGroup [AddCommGroup E] : AddCommGroup (WithCStarModule E) := ‹AddCommGroup E› + instance instSMul {R : Type*} [SMul R E] : SMul R (WithCStarModule E) := ‹SMul R E› + instance instModule {R : Type*} [Semiring R] [AddCommGroup E] [Module R E] : - Module R (WithCStarModule E) := - ‹Module R E› + Module R (WithCStarModule E) := ‹Module R E› instance instIsScalarTower [SMul R R'] [SMul R E] [SMul R' E] [IsScalarTower R R' E] : IsScalarTower R R' (WithCStarModule E) := @@ -89,11 +101,6 @@ instance instSMulCommClass [SMul R E] [SMul R' E] [SMulCommClass R R' E] : SMulCommClass R R' (WithCStarModule E) := ‹SMulCommClass R R' E› -instance instModuleFinite [Semiring R] [AddCommGroup E] [Module R E] [Module.Finite R E] : - Module.Finite R (WithCStarModule E) := - ‹Module.Finite R E› - - section Equiv variable {R E} @@ -151,6 +158,12 @@ theorem equiv_symm_smul : (equiv E).symm (c • x') = c • (equiv E).symm x' := end Equiv +/-- `WithCStarModule.equiv` as an additive equivalence. -/ +def addEquiv [AddCommGroup E] : C⋆ᵐᵒᵈ E ≃+ E := + { AddEquiv.refl _ with + toFun := equiv _ + invFun := (equiv _).symm } + /-- `WithCStarModule.equiv` as a linear equivalence. -/ @[simps (config := .asFn)] def linearEquiv [Semiring R] [AddCommGroup E] [Module R E] : C⋆ᵐᵒᵈ E ≃ₗ[R] E := @@ -158,6 +171,16 @@ def linearEquiv [Semiring R] [AddCommGroup E] [Module R E] : C⋆ᵐᵒᵈ E ≃ toFun := equiv _ invFun := (equiv _).symm } +lemma map_top_submodule {R : Type*} [Semiring R] [AddCommGroup E] [Module R E] : + (⊤ : Submodule R E).map (linearEquiv R E).symm = ⊤ := by + ext x + refine ⟨fun _ => trivial, fun _ => ?_⟩ + rw [Submodule.mem_map] + exact ⟨linearEquiv R E x, by simp⟩ + +instance instModuleFinite [Semiring R] [AddCommGroup E] [Module R E] [Module.Finite R E] : + Module.Finite R (WithCStarModule E) := inferInstanceAs (Module.Finite R E) + /-! ## `WithCStarModule E` inherits the uniformity and bornology from `E`. -/ variable {E} @@ -170,9 +193,26 @@ instance [Bornology E] : Bornology (C⋆ᵐᵒᵈ E) := Bornology.induced <| equ def uniformEquiv [UniformSpace E] : C⋆ᵐᵒᵈ E ≃ᵤ E := equiv E |>.toUniformEquivOfIsUniformInducing ⟨rfl⟩ +/-- `WithCStarModule.equiv` as a continuous linear equivalence between `C⋆ᵐᵒᵈ E` and `E`. -/ +@[simps! apply symm_apply] +def equivL [Semiring R] [AddCommGroup E] [UniformSpace E] [Module R E] : C⋆ᵐᵒᵈ E ≃L[R] E := + { linearEquiv R E with + continuous_toFun := UniformEquiv.continuous uniformEquiv + continuous_invFun := UniformEquiv.continuous uniformEquiv.symm } + instance [UniformSpace E] [CompleteSpace E] : CompleteSpace (C⋆ᵐᵒᵈ E) := uniformEquiv.completeSpace_iff.mpr inferInstance +instance [AddCommGroup E] [UniformSpace E] [ContinuousAdd E] : ContinuousAdd (C⋆ᵐᵒᵈ E) := + ContinuousAdd.induced (addEquiv E) + +instance [AddCommGroup E] [UniformSpace E] [UniformAddGroup E] : UniformAddGroup (C⋆ᵐᵒᵈ E) := + UniformAddGroup.comap (addEquiv E) + +instance [Semiring R] [TopologicalSpace R] [AddCommGroup E] [UniformSpace E] [Module R E] + [ContinuousSMul R E] : ContinuousSMul R (C⋆ᵐᵒᵈ E) := + ContinuousSMul.induced (linearEquiv R E) + end Basic /-! ## Prod diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index 7cc8a905e12eca..9ed44eadb36478 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -279,11 +279,8 @@ theorem hasFTaylorSeriesUpToOn_succ_iff_left {n : ℕ} : rw [this] exact h.2.2 -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/4119, -without `set_option maxSynthPendingDepth 2` this proof needs substantial repair. --/ +#adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 +without `set_option maxSynthPendingDepth 2` this proof needs substantial repair. -/ set_option maxSynthPendingDepth 2 in -- Porting note: this was split out from `hasFTaylorSeriesUpToOn_succ_iff_right` to avoid a timeout. theorem HasFTaylorSeriesUpToOn.shift_of_succ @@ -489,9 +486,8 @@ theorem iteratedFDerivWithin_succ_apply_right {n : ℕ} (hs : UniqueDiffOn 𝕜 rw [fderivWithin_congr A (A x hx)] _ = (I ∘ fderivWithin 𝕜 (iteratedFDerivWithin 𝕜 n (fderivWithin 𝕜 f s) s) s x : E → E[×n + 1]→L[𝕜] F) (m 0) (tail m) := by - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/4119 we need to either use + #adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 + we need to either use `set_option maxSynthPendingDepth 2 in` or fill in an explicit argument as ``` diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index 64f66f524e715c..9f7a1638661298 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -658,6 +658,9 @@ theorem deriv_const' : (deriv fun _ : 𝕜 => c) = fun _ => 0 := theorem derivWithin_const : derivWithin (fun _ => c) s = 0 := by ext; simp [derivWithin] +@[simp] +theorem derivWithin_zero : derivWithin (0 : 𝕜 → F) s = 0 := derivWithin_const _ _ + end Const section Continuous diff --git a/Mathlib/Analysis/Calculus/Deriv/Mul.lean b/Mathlib/Analysis/Calculus/Deriv/Mul.lean index b97172a75dda63..3da9a8ed349913 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Mul.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Mul.lean @@ -246,6 +246,17 @@ theorem derivWithin_mul_const (hc : DifferentiableWithinAt 𝕜 c s x) (d : 𝔸 · exact (hc.hasDerivWithinAt.mul_const d).derivWithin hxs · simp [derivWithin_zero_of_isolated hxs] +lemma derivWithin_mul_const_field (u : 𝕜') : + derivWithin (fun y => v y * u) s x = derivWithin v s x * u := by + by_cases hv : DifferentiableWithinAt 𝕜 v s x + · rw [derivWithin_mul_const hv u] + by_cases hu : u = 0 + · simp [hu] + rw [derivWithin_zero_of_not_differentiableWithinAt hv, zero_mul, + derivWithin_zero_of_not_differentiableWithinAt] + have : v = fun x ↦ (v x * u) * u⁻¹ := by ext; simp [hu] + exact fun h_diff ↦ hv <| this ▸ h_diff.mul_const _ + theorem deriv_mul_const (hc : DifferentiableAt 𝕜 c x) (d : 𝔸) : deriv (fun y => c y * d) x = deriv c x * d := (hc.hasDerivAt.mul_const d).deriv @@ -284,6 +295,11 @@ theorem derivWithin_const_mul (c : 𝔸) (hd : DifferentiableWithinAt 𝕜 d s x · exact (hd.hasDerivWithinAt.const_mul c).derivWithin hxs · simp [derivWithin_zero_of_isolated hxs] +lemma derivWithin_const_mul_field (u : 𝕜') : + derivWithin (fun y => u * v y) s x = u * derivWithin v s x := by + simp_rw [mul_comm u] + exact derivWithin_mul_const_field u + theorem deriv_const_mul (c : 𝔸) (hd : DifferentiableAt 𝕜 d x) : deriv (fun y => c * d y) x = c * deriv d x := (hd.hasDerivAt.const_mul c).deriv diff --git a/Mathlib/Analysis/Calculus/Deriv/Prod.lean b/Mathlib/Analysis/Calculus/Deriv/Prod.lean index caa6c116835431..61e10a2f7aeefa 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Prod.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Prod.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Gabriel Ebner, Yury Kudryashov +Authors: Gabriel Ebner, Yury Kudryashov, Eric Wieser -/ import Mathlib.Analysis.Calculus.Deriv.Basic import Mathlib.Analysis.Calculus.FDeriv.Prod @@ -13,7 +13,7 @@ In this file we prove lemmas about derivatives of functions `f : 𝕜 → E × F `f : 𝕜 → (Π i, E i)`. For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of -`analysis/calculus/deriv/basic`. +`Mathlib/Analysis/Calculus/Deriv/Basic.lean`. ## Keywords @@ -90,3 +90,88 @@ theorem deriv_pi (h : ∀ i, DifferentiableAt 𝕜 (fun x => φ x i) x) : (hasDerivAt_pi.2 fun i => (h i).hasDerivAt).deriv end Pi + + +/-! +### Derivatives of tuples `f : 𝕜 → Π i : Fin n.succ, F' i` + +These can be used to prove results about functions of the form `fun x ↦ ![f x, g x, h x]`, +as `Matrix.vecCons` is defeq to `Fin.cons`. +-/ +section PiFin + +variable {n : Nat} {F' : Fin n.succ → Type*} +variable [∀ i, NormedAddCommGroup (F' i)] [∀ i, NormedSpace 𝕜 (F' i)] +variable {φ : 𝕜 → F' 0} {φs : 𝕜 → ∀ i, F' (Fin.succ i)} + +theorem hasStrictDerivAt_finCons {φ' : Π i, F' i} : + HasStrictDerivAt (fun x => Fin.cons (φ x) (φs x)) φ' x ↔ + HasStrictDerivAt φ (φ' 0) x ∧ HasStrictDerivAt φs (fun i => φ' i.succ) x := + hasStrictFDerivAt_finCons + +/-- A variant of `hasStrictDerivAt_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem hasStrictDerivAt_finCons' {φ' : F' 0} {φs' : Π i, F' (Fin.succ i)} : + HasStrictDerivAt (fun x => Fin.cons (φ x) (φs x)) (Fin.cons φ' φs') x ↔ + HasStrictDerivAt φ φ' x ∧ HasStrictDerivAt φs φs' x := + hasStrictDerivAt_finCons + +theorem HasStrictDerivAt.finCons {φ' : F' 0} {φs' : Π i, F' (Fin.succ i)} + (h : HasStrictDerivAt φ φ' x) (hs : HasStrictDerivAt φs φs' x) : + HasStrictDerivAt (fun x => Fin.cons (φ x) (φs x)) (Fin.cons φ' φs') x := + hasStrictDerivAt_finCons'.mpr ⟨h, hs⟩ + +theorem hasDerivAtFilter_finCons {φ' : Π i, F' i} {l : Filter 𝕜} : + HasDerivAtFilter (fun x => Fin.cons (φ x) (φs x)) φ' x l ↔ + HasDerivAtFilter φ (φ' 0) x l ∧ HasDerivAtFilter φs (fun i => φ' i.succ) x l := + hasFDerivAtFilter_finCons + +/-- A variant of `hasDerivAtFilter_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem hasDerivAtFilter_finCons' {φ' : F' 0} {φs' : Π i, F' (Fin.succ i)} {l : Filter 𝕜} : + HasDerivAtFilter (fun x => Fin.cons (φ x) (φs x)) (Fin.cons φ' φs') x l ↔ + HasDerivAtFilter φ φ' x l ∧ HasDerivAtFilter φs φs' x l := + hasDerivAtFilter_finCons + +theorem HasDerivAtFilter.finCons {φ' : F' 0} {φs' : Π i, F' (Fin.succ i)} {l : Filter 𝕜} + (h : HasDerivAtFilter φ φ' x l) (hs : HasDerivAtFilter φs φs' x l) : + HasDerivAtFilter (fun x => Fin.cons (φ x) (φs x)) (Fin.cons φ' φs') x l := + hasDerivAtFilter_finCons'.mpr ⟨h, hs⟩ + +theorem hasDerivAt_finCons {φ' : Π i, F' i} : + HasDerivAt (fun x => Fin.cons (φ x) (φs x)) φ' x ↔ + HasDerivAt φ (φ' 0) x ∧ HasDerivAt φs (fun i => φ' i.succ) x := + hasDerivAtFilter_finCons + +/-- A variant of `hasDerivAt_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem hasDerivAt_finCons' {φ' : F' 0} {φs' : Π i, F' (Fin.succ i)} : + HasDerivAt (fun x => Fin.cons (φ x) (φs x)) (Fin.cons φ' φs') x ↔ + HasDerivAt φ φ' x ∧ HasDerivAt φs φs' x := + hasDerivAt_finCons + +theorem HasDerivAt.finCons {φ' : F' 0} {φs' : Π i, F' (Fin.succ i)} + (h : HasDerivAt φ φ' x) (hs : HasDerivAt φs φs' x) : + HasDerivAt (fun x => Fin.cons (φ x) (φs x)) (Fin.cons φ' φs') x := + hasDerivAt_finCons'.mpr ⟨h, hs⟩ + +theorem hasDerivWithinAt_finCons {φ' : Π i, F' i} : + HasDerivWithinAt (fun x => Fin.cons (φ x) (φs x)) φ' s x ↔ + HasDerivWithinAt φ (φ' 0) s x ∧ HasDerivWithinAt φs (fun i => φ' i.succ) s x := + hasDerivAtFilter_finCons + +/-- A variant of `hasDerivWithinAt_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem hasDerivWithinAt_finCons' {φ' : F' 0} {φs' : Π i, F' (Fin.succ i)} : + HasDerivWithinAt (fun x => Fin.cons (φ x) (φs x)) (Fin.cons φ' φs') s x ↔ + HasDerivWithinAt φ φ' s x ∧ HasDerivWithinAt φs φs' s x := + hasDerivAtFilter_finCons + +theorem HasDerivWithinAt.finCons {φ' : F' 0} {φs' : Π i, F' (Fin.succ i)} + (h : HasDerivWithinAt φ φ' s x) (hs : HasDerivWithinAt φs φs' s x) : + HasDerivWithinAt (fun x => Fin.cons (φ x) (φs x)) (Fin.cons φ' φs') s x := + hasDerivWithinAt_finCons'.mpr ⟨h, hs⟩ + +-- TODO: write the `Fin.cons` versions of `derivWithin_pi` and `deriv_pi` + +end PiFin diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 09a1501984d890..00cea385a3f4a6 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -205,7 +205,7 @@ protected theorem HasFPowerSeriesOnBall.fderiv [CompleteSpace F] (h.r_pos.trans_le h.r_le)).mono h.r_pos h.r_le).comp_sub x dsimp only rw [← h.fderiv_eq, add_sub_cancel] - simpa only [edist_eq_coe_nnnorm_sub, EMetric.mem_ball] using hz + simpa only [edist_eq_enorm_sub, EMetric.mem_ball] using hz /-- If a function has a power series within a set on a ball, then so does its derivative. -/ protected theorem HasFPowerSeriesWithinOnBall.fderivWithin [CompleteSpace F] @@ -220,7 +220,7 @@ protected theorem HasFPowerSeriesWithinOnBall.fderivWithin [CompleteSpace F] (h.r_pos.trans_le h.r_le)).mono h.r_pos h.r_le).comp_sub x · dsimp only rw [← h.fderivWithin_eq _ _ hu, add_sub_cancel] - · simpa only [edist_eq_coe_nnnorm_sub, EMetric.mem_ball] using hz.2 + · simpa only [edist_eq_enorm_sub, EMetric.mem_ball] using hz.2 · simpa using hz.1 /-- If a function has a power series within a set on a ball, then so does its derivative. For a @@ -315,7 +315,7 @@ theorem HasFPowerSeriesWithinOnBall.hasSum_derivSeries_of_hasFDerivWithinAt rw [← b.isEmbedding.hasSum_iff] have : HasFPowerSeriesWithinOnBall (a ∘ f) (a.compFormalMultilinearSeries p) s x r := a.comp_hasFPowerSeriesWithinOnBall h - have Z := (this.fderivWithin hu).hasSum h'y (by simpa [edist_eq_coe_nnnorm] using hy) + have Z := (this.fderivWithin hu).hasSum h'y (by simpa [edist_zero_eq_enorm] using hy) have : fderivWithin 𝕜 (a ∘ f) (insert x s) (x + y) = a ∘L f' := by apply HasFDerivWithinAt.fderivWithin _ (hu _ h'y) exact a.hasFDerivAt.comp_hasFDerivWithinAt (x + y) hf' @@ -339,7 +339,7 @@ protected theorem HasFPowerSeriesWithinOnBall.fderivWithin_of_mem_of_analyticOn (h : AnalyticOn 𝕜 f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : HasFPowerSeriesWithinOnBall (fderivWithin 𝕜 f s) p.derivSeries s x r := by refine ⟨hr.r_le.trans p.radius_le_radius_derivSeries, hr.r_pos, fun {y} hy h'y ↦ ?_⟩ - apply hr.hasSum_derivSeries_of_hasFDerivWithinAt (by simpa [edist_eq_coe_nnnorm] using h'y) hy + apply hr.hasSum_derivSeries_of_hasFDerivWithinAt (by simpa [edist_zero_eq_enorm] using h'y) hy · rw [insert_eq_of_mem hx] at hy ⊢ apply DifferentiableWithinAt.hasFDerivWithinAt exact h.differentiableOn _ hy @@ -489,7 +489,7 @@ protected theorem HasFiniteFPowerSeriesOnBall.fderiv ((p.hasFiniteFPowerSeriesOnBall_changeOrigin 1 h.finite).mono h.r_pos le_top).comp_sub x dsimp only rw [← h.fderiv_eq, add_sub_cancel] - simpa only [edist_eq_coe_nnnorm_sub, EMetric.mem_ball] using hz + simpa only [edist_eq_enorm_sub, EMetric.mem_ball] using hz /-- If a function has a finite power series on a ball, then so does its derivative. This is a variant of `HasFiniteFPowerSeriesOnBall.fderiv` where the degree of `f` is `< n` diff --git a/Mathlib/Analysis/Calculus/FDeriv/Prod.lean b/Mathlib/Analysis/Calculus/FDeriv/Prod.lean index c5e33c862226a1..892ed8c6dfa95e 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Prod.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Prod.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov +Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov, Eric Wieser -/ import Mathlib.Analysis.Calculus.FDeriv.Linear import Mathlib.Analysis.Calculus.FDeriv.Comp @@ -487,6 +487,187 @@ theorem fderiv_pi (h : ∀ i, DifferentiableAt 𝕜 (φ i) x) : end Pi +/-! +### Derivatives of tuples `f : E → Π i : Fin n.succ, F' i` + +These can be used to prove results about functions of the form `fun x ↦ ![f x, g x, h x]`, +as `Matrix.vecCons` is defeq to `Fin.cons`. +-/ +section PiFin + +variable {n : Nat} {F' : Fin n.succ → Type*} +variable [∀ i, NormedAddCommGroup (F' i)] [∀ i, NormedSpace 𝕜 (F' i)] +variable {φ : E → F' 0} {φs : E → ∀ i, F' (Fin.succ i)} + +theorem hasStrictFDerivAt_finCons {φ' : E →L[𝕜] Π i, F' i} : + HasStrictFDerivAt (fun x => Fin.cons (φ x) (φs x)) φ' x ↔ + HasStrictFDerivAt φ (.proj 0 ∘L φ') x ∧ + HasStrictFDerivAt φs (Pi.compRightL 𝕜 F' Fin.succ ∘L φ') x := by + rw [hasStrictFDerivAt_pi', Fin.forall_fin_succ, hasStrictFDerivAt_pi'] + dsimp [ContinuousLinearMap.comp, LinearMap.comp, Function.comp_def] + simp only [Fin.cons_zero, Fin.cons_succ] + +/-- A variant of `hasStrictFDerivAt_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem hasStrictFDerivAt_finCons' + {φ' : E →L[𝕜] F' 0} {φs' : E →L[𝕜] Π i, F' (Fin.succ i)} : + HasStrictFDerivAt (fun x => Fin.cons (φ x) (φs x)) (φ'.finCons φs') x ↔ + HasStrictFDerivAt φ φ' x ∧ HasStrictFDerivAt φs φs' x := + hasStrictFDerivAt_finCons + +@[fun_prop] +theorem HasStrictFDerivAt.finCons + {φ' : E →L[𝕜] F' 0} {φs' : E →L[𝕜] Π i, F' (Fin.succ i)} + (h : HasStrictFDerivAt φ φ' x) (hs : HasStrictFDerivAt φs φs' x) : + HasStrictFDerivAt (fun x => Fin.cons (φ x) (φs x)) (φ'.finCons φs') x := + hasStrictFDerivAt_finCons'.mpr ⟨h, hs⟩ + +theorem hasFDerivAtFilter_finCons + {φ' : E →L[𝕜] Π i, F' i} {l : Filter E} : + HasFDerivAtFilter (fun x => Fin.cons (φ x) (φs x)) φ' x l ↔ + HasFDerivAtFilter φ (.proj 0 ∘L φ') x l ∧ + HasFDerivAtFilter φs (Pi.compRightL 𝕜 F' Fin.succ ∘L φ') x l := by + rw [hasFDerivAtFilter_pi', Fin.forall_fin_succ, hasFDerivAtFilter_pi'] + dsimp [ContinuousLinearMap.comp, LinearMap.comp, Function.comp_def] + simp only [Fin.cons_zero, Fin.cons_succ] + +/-- A variant of `hasFDerivAtFilter_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem hasFDerivAtFilter_finCons' + {φ' : E →L[𝕜] F' 0} {φs' : E →L[𝕜] Π i, F' (Fin.succ i)} {l : Filter E} : + HasFDerivAtFilter (fun x => Fin.cons (φ x) (φs x)) (φ'.finCons φs') x l ↔ + HasFDerivAtFilter φ φ' x l ∧ HasFDerivAtFilter φs φs' x l := + hasFDerivAtFilter_finCons + +theorem HasFDerivAtFilter.finCons + {φ' : E →L[𝕜] F' 0} {φs' : E →L[𝕜] Π i, F' (Fin.succ i)} {l : Filter E} + (h : HasFDerivAtFilter φ φ' x l) (hs : HasFDerivAtFilter φs φs' x l) : + HasFDerivAtFilter (fun x => Fin.cons (φ x) (φs x)) (φ'.finCons φs') x l := + hasFDerivAtFilter_finCons'.mpr ⟨h, hs⟩ + +theorem hasFDerivAt_finCons + {φ' : E →L[𝕜] Π i, F' i} : + HasFDerivAt (fun x => Fin.cons (φ x) (φs x)) φ' x ↔ + HasFDerivAt φ (.proj 0 ∘L φ') x ∧ HasFDerivAt φs (Pi.compRightL 𝕜 F' Fin.succ ∘L φ') x := + hasFDerivAtFilter_finCons + +/-- A variant of `hasFDerivAt_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem hasFDerivAt_finCons' + {φ' : E →L[𝕜] F' 0} {φs' : E →L[𝕜] Π i, F' (Fin.succ i)} : + HasFDerivAt (fun x => Fin.cons (φ x) (φs x)) (φ'.finCons φs') x ↔ + HasFDerivAt φ φ' x ∧ HasFDerivAt φs φs' x := + hasFDerivAt_finCons + +@[fun_prop] +theorem HasFDerivAt.finCons + {φ' : E →L[𝕜] F' 0} {φs' : E →L[𝕜] Π i, F' (Fin.succ i)} + (h : HasFDerivAt φ φ' x) (hs : HasFDerivAt φs φs' x) : + HasFDerivAt (fun x => Fin.cons (φ x) (φs x)) (φ'.finCons φs') x := + hasFDerivAt_finCons'.mpr ⟨h, hs⟩ + +theorem hasFDerivWithinAt_finCons + {φ' : E →L[𝕜] Π i, F' i} : + HasFDerivWithinAt (fun x => Fin.cons (φ x) (φs x)) φ' s x ↔ + HasFDerivWithinAt φ (.proj 0 ∘L φ') s x ∧ + HasFDerivWithinAt φs (Pi.compRightL 𝕜 F' Fin.succ ∘L φ') s x := + hasFDerivAtFilter_finCons + +/-- A variant of `hasFDerivWithinAt_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem hasFDerivWithinAt_finCons' + {φ' : E →L[𝕜] F' 0} {φs' : E →L[𝕜] Π i, F' (Fin.succ i)} : + HasFDerivWithinAt (fun x => Fin.cons (φ x) (φs x)) (φ'.finCons φs') s x ↔ + HasFDerivWithinAt φ φ' s x ∧ HasFDerivWithinAt φs φs' s x := + hasFDerivAtFilter_finCons + +@[fun_prop] +theorem HasFDerivWithinAt.finCons + {φ' : E →L[𝕜] F' 0} {φs' : E →L[𝕜] Π i, F' (Fin.succ i)} + (h : HasFDerivWithinAt φ φ' s x) (hs : HasFDerivWithinAt φs φs' s x) : + HasFDerivWithinAt (fun x => Fin.cons (φ x) (φs x)) (φ'.finCons φs') s x := + hasFDerivWithinAt_finCons'.mpr ⟨h, hs⟩ + +theorem differentiableWithinAt_finCons : + DifferentiableWithinAt 𝕜 (fun x => Fin.cons (φ x) (φs x)) s x ↔ + DifferentiableWithinAt 𝕜 φ s x ∧ DifferentiableWithinAt 𝕜 φs s x := by + rw [differentiableWithinAt_pi, Fin.forall_fin_succ, differentiableWithinAt_pi] + simp only [Fin.cons_zero, Fin.cons_succ] + +/-- A variant of `differentiableWithinAt_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem differentiableWithinAt_finCons' : + DifferentiableWithinAt 𝕜 (fun x => Fin.cons (φ x) (φs x)) s x ↔ + DifferentiableWithinAt 𝕜 φ s x ∧ DifferentiableWithinAt 𝕜 φs s x := + differentiableWithinAt_finCons + +@[fun_prop] +theorem DifferentiableWithinAt.finCons + (h : DifferentiableWithinAt 𝕜 φ s x) (hs : DifferentiableWithinAt 𝕜 φs s x) : + DifferentiableWithinAt 𝕜 (fun x => Fin.cons (φ x) (φs x)) s x := + differentiableWithinAt_finCons'.mpr ⟨h, hs⟩ + +theorem differentiableAt_finCons : + DifferentiableAt 𝕜 (fun x => Fin.cons (φ x) (φs x)) x ↔ + DifferentiableAt 𝕜 φ x ∧ DifferentiableAt 𝕜 φs x := by + rw [differentiableAt_pi, Fin.forall_fin_succ, differentiableAt_pi] + simp only [Fin.cons_zero, Fin.cons_succ] + +/-- A variant of `differentiableAt_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem differentiableAt_finCons' : + DifferentiableAt 𝕜 (fun x => Fin.cons (φ x) (φs x)) x ↔ + DifferentiableAt 𝕜 φ x ∧ DifferentiableAt 𝕜 φs x := + differentiableAt_finCons + +@[fun_prop] +theorem DifferentiableAt.finCons + (h : DifferentiableAt 𝕜 φ x) (hs : DifferentiableAt 𝕜 φs x) : + DifferentiableAt 𝕜 (fun x => Fin.cons (φ x) (φs x)) x := + differentiableAt_finCons'.mpr ⟨h, hs⟩ + +theorem differentiableOn_finCons : + DifferentiableOn 𝕜 (fun x => Fin.cons (φ x) (φs x)) s ↔ + DifferentiableOn 𝕜 φ s ∧ DifferentiableOn 𝕜 φs s := by + rw [differentiableOn_pi, Fin.forall_fin_succ, differentiableOn_pi] + simp only [Fin.cons_zero, Fin.cons_succ] + +/-- A variant of `differentiableOn_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem differentiableOn_finCons' : + DifferentiableOn 𝕜 (fun x => Fin.cons (φ x) (φs x)) s ↔ + DifferentiableOn 𝕜 φ s ∧ DifferentiableOn 𝕜 φs s := + differentiableOn_finCons + +@[fun_prop] +theorem DifferentiableOn.finCons + (h : DifferentiableOn 𝕜 φ s) (hs : DifferentiableOn 𝕜 φs s) : + DifferentiableOn 𝕜 (fun x => Fin.cons (φ x) (φs x)) s := + differentiableOn_finCons'.mpr ⟨h, hs⟩ + +theorem differentiable_finCons : + Differentiable 𝕜 (fun x => Fin.cons (φ x) (φs x)) ↔ + Differentiable 𝕜 φ ∧ Differentiable 𝕜 φs := by + rw [differentiable_pi, Fin.forall_fin_succ, differentiable_pi] + simp only [Fin.cons_zero, Fin.cons_succ] + +/-- A variant of `differentiable_finCons` where the derivative variables are free on the RHS +instead. -/ +theorem differentiable_finCons' : + Differentiable 𝕜 (fun x => Fin.cons (φ x) (φs x)) ↔ + Differentiable 𝕜 φ ∧ Differentiable 𝕜 φs := + differentiable_finCons + +@[fun_prop] +theorem Differentiable.finCons + (h : Differentiable 𝕜 φ) (hs : Differentiable 𝕜 φs) : + Differentiable 𝕜 (fun x => Fin.cons (φ x) (φs x)) := + differentiable_finCons'.mpr ⟨h, hs⟩ + +-- TODO: write the `Fin.cons` versions of `fderivWithin_pi` and `fderiv_pi` + +end PiFin + end CartesianProduct end diff --git a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean index 4254d8378f1d65..78250c59bb934c 100644 --- a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean +++ b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean @@ -64,11 +64,8 @@ end Module namespace FormalMultilinearSeries -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/4481 -the `simpNF` linter incorrectly claims this lemma can't be applied by `simp`. --/ +#adaptation_note /-- https://github.com/leanprover/lean4/pull/4481 +the `simpNF` linter incorrectly claims this lemma can't be applied by `simp`. -/ @[simp, nolint simpNF] theorem zero_apply (n : ℕ) : (0 : FormalMultilinearSeries 𝕜 E F) n = 0 := rfl diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean index bd390d2a8ebcb5..d51fff9a736c09 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean @@ -290,3 +290,21 @@ theorem iteratedDeriv_eq_iterate : iteratedDeriv n f = deriv^[n] f := by derivative. -/ theorem iteratedDeriv_succ' : iteratedDeriv (n + 1) f = iteratedDeriv n (deriv f) := by rw [iteratedDeriv_eq_iterate, iteratedDeriv_eq_iterate]; rfl + +lemma AnalyticAt.hasFPowerSeriesAt {𝕜 : Type*} [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] + [CharZero 𝕜] {f : 𝕜 → 𝕜} {x : 𝕜} (h : AnalyticAt 𝕜 f x) : + HasFPowerSeriesAt f + (FormalMultilinearSeries.ofScalars 𝕜 (fun n ↦ iteratedDeriv n f x / n.factorial)) x := by + obtain ⟨p, hp⟩ := h + convert hp + obtain ⟨r, hpr⟩ := hp + ext n u + have h_fact_smul := hpr.factorial_smul 1 + simp only [FormalMultilinearSeries.apply_eq_prod_smul_coeff, Finset.prod_const, Finset.card_univ, + Fintype.card_fin, smul_eq_mul, nsmul_eq_mul, one_pow, one_mul] at h_fact_smul + simp only [FormalMultilinearSeries.apply_eq_prod_smul_coeff, + FormalMultilinearSeries.coeff_ofScalars, smul_eq_mul, mul_eq_mul_left_iff] + left + rw [div_eq_iff, mul_comm, h_fact_smul, ← iteratedDeriv_eq_iteratedFDeriv] + norm_cast + exact Nat.factorial_ne_zero _ diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean index fb639723208c38..e797df9f8f6a72 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean @@ -152,6 +152,18 @@ lemma iteratedDeriv_sub (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) : iteratedDerivWithin_sub (Set.mem_univ _) uniqueDiffOn_univ (contDiffOn_univ.mpr hf) (contDiffOn_univ.mpr hg) +theorem iteratedDeriv_const_smul {n : ℕ} {f : 𝕜 → F} (h : ContDiffAt 𝕜 n f x) (c : 𝕜) : + iteratedDeriv n (c • f) x = c • iteratedDeriv n f x := by + simpa only [iteratedDerivWithin_univ] using + iteratedDerivWithin_const_smul (Set.mem_univ x) uniqueDiffOn_univ + c (contDiffWithinAt_univ.mpr h) + +theorem iteratedDeriv_const_mul {n : ℕ} {f : 𝕜 → 𝕜} (h : ContDiffAt 𝕜 n f x) (c : 𝕜) : + iteratedDeriv n (fun z => c * f z) x = c * iteratedDeriv n f x := by + simpa only [iteratedDerivWithin_univ] using + iteratedDerivWithin_const_mul (Set.mem_univ x) uniqueDiffOn_univ + c (contDiffWithinAt_univ.mpr h) + theorem iteratedDeriv_comp_const_smul {n : ℕ} {f : 𝕜 → F} (h : ContDiff 𝕜 n f) (c : 𝕜) : iteratedDeriv n (fun x => f (c * x)) = fun x => c ^ n • iteratedDeriv n f (c * x) := by funext x @@ -159,16 +171,10 @@ theorem iteratedDeriv_comp_const_smul {n : ℕ} {f : 𝕜 → F} (h : ContDiff iteratedDerivWithin_comp_const_smul (Set.mem_univ x) uniqueDiffOn_univ (contDiffOn_univ.mpr h) c (Set.mapsTo_univ _ _) -@[deprecated (since := "2024-12-20")] -alias iteratedDeriv_const_smul := iteratedDeriv_comp_const_smul - theorem iteratedDeriv_comp_const_mul {n : ℕ} {f : 𝕜 → 𝕜} (h : ContDiff 𝕜 n f) (c : 𝕜) : iteratedDeriv n (fun x => f (c * x)) = fun x => c ^ n * iteratedDeriv n f (c * x) := by simpa only [smul_eq_mul] using iteratedDeriv_comp_const_smul h c -@[deprecated (since := "2024-12-20")] -alias iteratedDeriv_const_mul := iteratedDeriv_comp_const_mul - lemma iteratedDeriv_comp_neg (n : ℕ) (f : 𝕜 → F) (a : 𝕜) : iteratedDeriv n (fun x ↦ f (-x)) a = (-1 : 𝕜) ^ n • iteratedDeriv n f (-a) := by induction' n with n ih generalizing a diff --git a/Mathlib/Analysis/Calculus/TangentCone.lean b/Mathlib/Analysis/Calculus/TangentCone.lean index 68f5491fa0644a..e3061bb13fff2c 100644 --- a/Mathlib/Analysis/Calculus/TangentCone.lean +++ b/Mathlib/Analysis/Calculus/TangentCone.lean @@ -70,15 +70,17 @@ def UniqueDiffOn (s : Set E) : Prop := end TangentCone -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] -variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable {G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] -variable {𝕜} {x y : E} {s t : Set E} +variable {𝕜} +variable {E F G : Type*} section TangentCone -- This section is devoted to the properties of the tangent cone. + open NormedField +section TVS +variable [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] +variable {x y : E} {s t : Set E} theorem mem_tangentConeAt_of_pow_smul {r : 𝕜} (hr₀ : r ≠ 0) (hr : ‖r‖ < 1) (hs : ∀ᶠ n : ℕ in atTop, x + r ^ n • y ∈ s) : y ∈ tangentConeAt 𝕜 s x := by @@ -96,6 +98,14 @@ theorem tangentCone_mono (h : s ⊆ t) : tangentConeAt 𝕜 s x ⊆ tangentConeA rintro y ⟨c, d, ds, ctop, clim⟩ exact ⟨c, d, mem_of_superset ds fun n hn => h hn, ctop, clim⟩ +end TVS + +section Normed +variable [NormedAddCommGroup E] [NormedSpace 𝕜 E] +variable [NormedAddCommGroup F] [NormedSpace 𝕜 F] +variable [NormedAddCommGroup G] [NormedSpace ℝ G] +variable {x y : E} {s t : Set E} + /-- Auxiliary lemma ensuring that, under the assumptions defining the tangent cone, the sequence `d` tends to 0 at infinity. -/ theorem tangentConeAt.lim_zero {α : Type*} (l : Filter α) {c : α → 𝕜} {d : α → E} @@ -337,6 +347,8 @@ theorem tangentCone_eq_univ {s : Set 𝕜} {x : 𝕜} (hx : (𝓝[s \ {x}] x).Ne · convert tendsto_const_nhds (α := ℕ) (x := y) with n simp [mul_assoc, inv_mul_cancel₀ (d_ne n)] +end Normed + end TangentCone section UniqueDiff @@ -346,6 +358,9 @@ section UniqueDiff This section is devoted to properties of the predicates `UniqueDiffWithinAt` and `UniqueDiffOn`. -/ +section TVS +variable [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] +variable {x y : E} {s t : Set E} theorem UniqueDiffOn.uniqueDiffWithinAt {s : Set E} {x} (hs : UniqueDiffOn 𝕜 s) (h : x ∈ s) : UniqueDiffWithinAt 𝕜 s x := @@ -364,6 +379,13 @@ theorem uniqueDiffOn_empty : UniqueDiffOn 𝕜 (∅ : Set E) := theorem UniqueDiffWithinAt.congr_pt (h : UniqueDiffWithinAt 𝕜 s x) (hy : x = y) : UniqueDiffWithinAt 𝕜 s y := hy ▸ h +end TVS + +section Normed +variable [NormedAddCommGroup E] [NormedSpace 𝕜 E] +variable [NormedAddCommGroup F] [NormedSpace 𝕜 F] +variable {x y : E} {s t : Set E} + theorem UniqueDiffWithinAt.mono_nhds (h : UniqueDiffWithinAt 𝕜 s x) (st : 𝓝[s] x ≤ 𝓝[t] x) : UniqueDiffWithinAt 𝕜 t x := by simp only [uniqueDiffWithinAt_iff] at * @@ -457,6 +479,11 @@ theorem UniqueDiffOn.univ_pi (ι : Type*) [Finite ι] (E : ι → Type*) (h : ∀ i, UniqueDiffOn 𝕜 (s i)) : UniqueDiffOn 𝕜 (Set.pi univ s) := UniqueDiffOn.pi _ _ _ _ fun i _ => h i +end Normed + +section RealNormed +variable [NormedAddCommGroup G] [NormedSpace ℝ G] + /-- In a real vector space, a convex set with nonempty interior is a set of unique differentiability at every point of its closure. -/ theorem uniqueDiffWithinAt_convex {s : Set G} (conv : Convex ℝ s) (hs : (interior s).Nonempty) @@ -479,6 +506,10 @@ theorem uniqueDiffOn_convex {s : Set G} (conv : Convex ℝ s) (hs : (interior s) UniqueDiffOn ℝ s := fun _ xs => uniqueDiffWithinAt_convex conv hs (subset_closure xs) +end RealNormed + +section Real + theorem uniqueDiffOn_Ici (a : ℝ) : UniqueDiffOn ℝ (Ici a) := uniqueDiffOn_convex (convex_Ici a) <| by simp only [interior_Ici, nonempty_Ioi] @@ -532,4 +563,6 @@ theorem uniqueDiffWithinAt_or_nhdsWithin_eq_bot (s : Set 𝕜) (x : 𝕜) : apply neBot_of_le (hf := h) exact nhdsWithin_mono _ diff_subset +end Real + end UniqueDiff diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index 73c4f7ae5cdfa4..3d4217f58338eb 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -5,7 +5,7 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Data.Complex.Module import Mathlib.Data.Complex.Order -import Mathlib.Data.Complex.Exponential +import Mathlib.Data.Complex.Trigonometric import Mathlib.Analysis.RCLike.Basic import Mathlib.Topology.Algebra.InfiniteSum.Field import Mathlib.Topology.Algebra.InfiniteSum.Module @@ -177,11 +177,11 @@ theorem norm_int_of_nonneg {n : ℤ} (hn : 0 ≤ n) : ‖(n : ℂ)‖ = n := by lemma normSq_eq_norm_sq (z : ℂ) : normSq z = ‖z‖ ^ 2 := by rw [normSq_eq_abs, norm_eq_abs] -@[continuity] +@[continuity, fun_prop] theorem continuous_abs : Continuous abs := continuous_norm -@[continuity] +@[continuity, fun_prop] theorem continuous_normSq : Continuous normSq := by simpa [← normSq_eq_abs] using continuous_abs.pow 2 diff --git a/Mathlib/Analysis/Complex/CauchyIntegral.lean b/Mathlib/Analysis/Complex/CauchyIntegral.lean index 99b3e0d3c22dc6..41557430b9bf36 100644 --- a/Mathlib/Analysis/Complex/CauchyIntegral.lean +++ b/Mathlib/Analysis/Complex/CauchyIntegral.lean @@ -343,7 +343,7 @@ theorem circleIntegral_sub_center_inv_smul_of_differentiable_on_off_countable_of (hd : ∀ z ∈ (ball c R \ {c}) \ s, DifferentiableAt ℂ f z) (hy : Tendsto f (𝓝[{c}ᶜ] c) (𝓝 y)) : (∮ z in C(c, R), (z - c)⁻¹ • f z) = (2 * π * I : ℂ) • y := by rw [← sub_eq_zero, ← norm_le_zero_iff] - refine le_of_forall_le_of_dense fun ε ε0 => ?_ + refine le_of_forall_gt_imp_ge_of_dense fun ε ε0 => ?_ obtain ⟨δ, δ0, hδ⟩ : ∃ δ > (0 : ℝ), ∀ z ∈ closedBall c δ \ {c}, dist (f z) y < ε / (2 * π) := ((nhdsWithin_hasBasis nhds_basis_closedBall _).tendsto_iff nhds_basis_ball).1 hy _ (div_pos ε0 Real.two_pi_pos) @@ -536,7 +536,7 @@ theorem hasFPowerSeriesOnBall_of_differentiable_off_countable {R : ℝ≥0} {c : hasSum := fun {w} hw => by have hw' : c + w ∈ ball c R := by simpa only [add_mem_ball_iff_norm, ← coe_nnnorm, mem_emetric_ball_zero_iff, - NNReal.coe_lt_coe, ENNReal.coe_lt_coe] using hw + NNReal.coe_lt_coe, enorm_lt_coe] using hw rw [← two_pi_I_inv_smul_circleIntegral_sub_inv_smul_of_differentiable_on_off_countable hs hw' hc hd] exact (hasFPowerSeriesOn_cauchy_integral diff --git a/Mathlib/Analysis/Complex/Liouville.lean b/Mathlib/Analysis/Complex/Liouville.lean index aee36657dab9e8..75d46597fd3d1e 100644 --- a/Mathlib/Analysis/Complex/Liouville.lean +++ b/Mathlib/Analysis/Complex/Liouville.lean @@ -89,7 +89,7 @@ theorem liouville_theorem_aux {f : ℂ → F} (hf : Differentiable ℂ f) (hb : exact ⟨max C 1, lt_max_iff.2 (Or.inr zero_lt_one), fun z => (hC (f z) (mem_range_self _)).trans (le_max_left _ _)⟩ - refine norm_le_zero_iff.1 (le_of_forall_le_of_dense fun ε ε₀ => ?_) + refine norm_le_zero_iff.1 (le_of_forall_gt_imp_ge_of_dense fun ε ε₀ => ?_) calc ‖deriv f c‖ ≤ C / (C / ε) := norm_deriv_le_of_forall_mem_sphere_norm_le (div_pos C₀ ε₀) hf.diffContOnCl fun z _ => hC z diff --git a/Mathlib/Analysis/Complex/PhragmenLindelof.lean b/Mathlib/Analysis/Complex/PhragmenLindelof.lean index ee37b9bc62b367..1d6cb0068727fe 100644 --- a/Mathlib/Analysis/Complex/PhragmenLindelof.lean +++ b/Mathlib/Analysis/Complex/PhragmenLindelof.lean @@ -116,7 +116,7 @@ theorem horizontal_strip (hfd : DiffContOnCl ℂ f (im ⁻¹' Ioo a b)) cases' hza with hza hza; · exact hle_a _ hza.symm cases' hzb with hzb hzb; · exact hle_b _ hzb wlog hC₀ : 0 < C generalizing C - · refine le_of_forall_le_of_dense fun C' hC' => this (fun w hw => ?_) (fun w hw => ?_) ?_ + · refine le_of_forall_gt_imp_ge_of_dense fun C' hC' => this (fun w hw => ?_) (fun w hw => ?_) ?_ · exact (hle_a _ hw).trans hC'.le · exact (hle_b _ hw).trans hC'.le · refine ((norm_nonneg (f (a * I))).trans (hle_a _ ?_)).trans_lt hC' diff --git a/Mathlib/Analysis/Complex/TaylorSeries.lean b/Mathlib/Analysis/Complex/TaylorSeries.lean index dbd3037467b3ac..4c7a2f85e0047c 100644 --- a/Mathlib/Analysis/Complex/TaylorSeries.lean +++ b/Mathlib/Analysis/Complex/TaylorSeries.lean @@ -24,9 +24,8 @@ see `Complex.hasSum_taylorSeries_of_entire`, `Complex.taylorSeries_eq_of_entire` `Complex.taylorSeries_eq_of_entire'`. -/ -#adaptation_note -/-- -Due to https://github.com/leanprover/lean4/issues/5126, we've had to change all +#adaptation_note /-- https://github.com/leanprover/lean4/issues/5126 +we've had to change all the `⦃⦄` stricit implicit variable statements in this file to normal `{}` implicit variables. Once this issue is fixed, we should change them back. For now it doesn't break anything downstream. diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index c92c0c0ce665bc..eba4e36e6505b5 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -157,11 +157,8 @@ theorem cmp_dist_eq_cmp_dist_coe_center (z w : ℍ) (r : ℝ) : ((mul_neg_of_pos_of_neg w.im_pos (sinh_neg_iff.2 hr₀)).trans_le dist_nonneg).cmp_eq_gt.symm] have hr₀' : 0 ≤ w.im * Real.sinh r := by positivity have hzw₀ : 0 < 2 * z.im * w.im := by positivity - #adaptation_note - /-- - After the bug fix in https://github.com/leanprover/lean4/pull/6024, - we need to give Lean the hint `(y := w.im * Real.sinh r)`. - -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 + we need to give Lean the hint `(y := w.im * Real.sinh r)`. -/ simp only [← cosh_strictMonoOn.cmp_map_eq dist_nonneg hr₀, ← (pow_left_strictMonoOn₀ two_ne_zero).cmp_map_eq dist_nonneg (y := w.im * Real.sinh r) hr₀', dist_coe_center_sq] diff --git a/Mathlib/Analysis/Convex/Basic.lean b/Mathlib/Analysis/Convex/Basic.lean index cbf1474ae125fc..1894c691b8c59e 100644 --- a/Mathlib/Analysis/Convex/Basic.lean +++ b/Mathlib/Analysis/Convex/Basic.lean @@ -624,8 +624,8 @@ theorem single_mem_stdSimplex (i : ι) : Pi.single i 1 ∈ stdSimplex 𝕜 ι := theorem ite_eq_mem_stdSimplex (i : ι) : (if i = · then (1 : 𝕜) else 0) ∈ stdSimplex 𝕜 ι := by simpa only [@eq_comm _ i, ← Pi.single_apply] using single_mem_stdSimplex 𝕜 i -#adaptation_note /-- as of `nightly-2024-03-11`, we need a type annotation on the segment in the -following two lemmas. -/ +#adaptation_note /-- nightly-2024-03-11 +we need a type annotation on the segment in the following two lemmas. -/ /-- The edges are contained in the simplex. -/ lemma segment_single_subset_stdSimplex (i j : ι) : diff --git a/Mathlib/Analysis/Convex/Birkhoff.lean b/Mathlib/Analysis/Convex/Birkhoff.lean index d69061ff3199b0..3b8c6303e749c5 100644 --- a/Mathlib/Analysis/Convex/Birkhoff.lean +++ b/Mathlib/Analysis/Convex/Birkhoff.lean @@ -5,6 +5,7 @@ Authors: Bhavik Mehta -/ import Mathlib.Analysis.Convex.Combination +import Mathlib.Analysis.Convex.Extreme import Mathlib.Combinatorics.Hall.Basic import Mathlib.Data.Matrix.DoublyStochastic import Mathlib.Tactic.Linarith @@ -18,10 +19,11 @@ import Mathlib.Tactic.Linarith combination of permutation matrices. * `doublyStochastic_eq_convexHull_perm`: The set of doubly stochastic matrices is the convex hull of the permutation matrices. +* `extremePoints_doublyStochastic`: The set of extreme points of the doubly stochastic matrices is + the set of permutation matrices. ## TODO -* Show that the extreme points of doubly stochastic matrices are the permutation matrices. * Show that for `x y : n → R`, `x` is majorized by `y` if and only if there is a doubly stochastic matrix `M` such that `M *ᵥ y = x`. @@ -169,4 +171,28 @@ theorem doublyStochastic_eq_convexHull_permMatrix : obtain ⟨w, hw1, hw2, hw3⟩ := exists_eq_sum_perm_of_mem_doublyStochastic hM exact mem_convexHull_of_exists_fintype w (·.permMatrix R) hw1 hw2 (by simp) hw3 +/-- +The set of extreme points of the doubly stochastic matrices is the set of permutation matrices. +-/ +theorem extremePoints_doublyStochastic : + Set.extremePoints R (doublyStochastic R n) = {σ.permMatrix R | σ : Equiv.Perm n} := by + refine subset_antisymm ?_ ?_ + · rw [doublyStochastic_eq_convexHull_permMatrix] + exact extremePoints_convexHull_subset + rintro _ ⟨σ, rfl⟩ + refine ⟨permMatrix_mem_doublyStochastic, fun x₁ hx₁ x₂ hx₂ hσ ↦ ?_⟩ + suffices ∀ i j : n, x₁ i j = x₂ i j by + obtain rfl : x₁ = x₂ := by simpa [← Matrix.ext_iff] + simp_all + intro i j + have h₁ : σ.permMatrix R i j ∈ openSegment R (x₁ i j) (x₂ i j) := + image_openSegment _ (entryLinearMap R R i j).toAffineMap x₁ x₂ ▸ ⟨_, hσ, rfl⟩ + by_contra! h + have h₂ : openSegment R (x₁ i j) (x₂ i j) ⊆ Set.Ioo 0 1 := by + rw [openSegment_eq_Ioo' h] + apply Set.Ioo_subset_Ioo <;> + simp_all [nonneg_of_mem_doublyStochastic, le_one_of_mem_doublyStochastic] + specialize h₂ h₁ + aesop + end LinearOrderedField diff --git a/Mathlib/Analysis/Convex/EGauge.lean b/Mathlib/Analysis/Convex/EGauge.lean index 5083bf5351560e..6e85fb91802748 100644 --- a/Mathlib/Analysis/Convex/EGauge.lean +++ b/Mathlib/Analysis/Convex/EGauge.lean @@ -32,15 +32,15 @@ section SMul /-- The Minkowski functional for vector spaces over normed fields. Given a set `s` in a vector space over a normed field `𝕜`, `egauge s` is the functional which sends `x : E` -to the infimum of `‖c‖₊` over `c` such that `x` belongs to `s` scaled by `c`. +to the infimum of `‖c‖ₑ` over `c` such that `x` belongs to `s` scaled by `c`. -The definition only requires `𝕜` to have a `NNNorm` instance +The definition only requires `𝕜` to have a `ENorm` instance and `(· • ·) : 𝕜 → E → E` to be defined. This way the definition applies, e.g., to `𝕜 = ℝ≥0`. For `𝕜 = ℝ≥0`, the function is equal (up to conversion to `ℝ`) to the usual Minkowski functional defined in `gauge`. -/ -noncomputable def egauge (𝕜 : Type*) [NNNorm 𝕜] {E : Type*} [SMul 𝕜 E] (s : Set E) (x : E) : ℝ≥0∞ := - ⨅ (c : 𝕜) (_ : x ∈ c • s), ‖c‖₊ +noncomputable def egauge (𝕜 : Type*) [ENorm 𝕜] {E : Type*} [SMul 𝕜 E] (s : Set E) (x : E) : ℝ≥0∞ := + ⨅ (c : 𝕜) (_ : x ∈ c • s), ‖c‖ₑ variable (𝕜 : Type*) [NNNorm 𝕜] {E : Type*} [SMul 𝕜 E] {c : 𝕜} {s t : Set E} {x : E} {r : ℝ≥0∞} @@ -52,13 +52,13 @@ lemma egauge_anti (h : s ⊆ t) (x : E) : egauge 𝕜 t x ≤ egauge 𝕜 s x := variable {𝕜} -lemma egauge_le_of_mem_smul (h : x ∈ c • s) : egauge 𝕜 s x ≤ ‖c‖₊ := iInf₂_le c h +lemma egauge_le_of_mem_smul (h : x ∈ c • s) : egauge 𝕜 s x ≤ ‖c‖ₑ := iInf₂_le c h -lemma le_egauge_iff : r ≤ egauge 𝕜 s x ↔ ∀ c : 𝕜, x ∈ c • s → r ≤ ‖c‖₊ := le_iInf₂_iff +lemma le_egauge_iff : r ≤ egauge 𝕜 s x ↔ ∀ c : 𝕜, x ∈ c • s → r ≤ ‖c‖ₑ := le_iInf₂_iff lemma egauge_eq_top : egauge 𝕜 s x = ∞ ↔ ∀ c : 𝕜, x ∉ c • s := by simp [egauge] -lemma egauge_lt_iff : egauge 𝕜 s x < r ↔ ∃ c : 𝕜, x ∈ c • s ∧ ‖c‖₊ < r := by +lemma egauge_lt_iff : egauge 𝕜 s x < r ↔ ∃ c : 𝕜, x ∈ c • s ∧ ‖c‖ₑ < r := by simp [egauge, iInf_lt_iff] lemma egauge_union (s t : Set E) (x : E) : egauge 𝕜 (s ∪ t) x = egauge 𝕜 s x ⊓ egauge 𝕜 t x := by @@ -67,7 +67,7 @@ lemma egauge_union (s t : Set E) (x : E) : egauge 𝕜 (s ∪ t) x = egauge 𝕜 lemma le_egauge_inter (s t : Set E) (x : E) : egauge 𝕜 s x ⊔ egauge 𝕜 t x ≤ egauge 𝕜 (s ∩ t) x := - max_le_iff.2 ⟨egauge_anti _ inter_subset_left _, egauge_anti _ inter_subset_right _⟩ + max_le (egauge_anti _ inter_subset_left _) (egauge_anti _ inter_subset_right _) end SMul @@ -87,25 +87,24 @@ section Module variable {𝕜 : Type*} [NormedDivisionRing 𝕜] {E : Type*} [AddCommGroup E] [Module 𝕜 E] {c : 𝕜} {s : Set E} {x : E} -/-- If `c • x ∈ s` and `c ≠ 0`, then `egauge 𝕜 s x` is at most `((‖c‖₊⁻¹ : ℝ≥0) : ℝ≥0∞). +/-- If `c • x ∈ s` and `c ≠ 0`, then `egauge 𝕜 s x` is at most `(‖c‖₊⁻¹ : ℝ≥0)`. See also `egauge_le_of_smul_mem`. -/ -lemma egauge_le_of_smul_mem_of_ne (h : c • x ∈ s) (hc : c ≠ 0) : - egauge 𝕜 s x ≤ ↑(‖c‖₊⁻¹ : ℝ≥0) := by +lemma egauge_le_of_smul_mem_of_ne (h : c • x ∈ s) (hc : c ≠ 0) : egauge 𝕜 s x ≤ (‖c‖₊⁻¹ : ℝ≥0) := by rw [← nnnorm_inv] exact egauge_le_of_mem_smul <| (mem_inv_smul_set_iff₀ hc _ _).2 h -/-- If `c • x ∈ s`, then `egauge 𝕜 s x` is at most `(‖c‖₊ : ℝ≥0∞)⁻¹. +/-- If `c • x ∈ s`, then `egauge 𝕜 s x` is at most `‖c‖ₑ⁻¹`. See also `egauge_le_of_smul_mem_of_ne`. -/ -lemma egauge_le_of_smul_mem (h : c • x ∈ s) : egauge 𝕜 s x ≤ (‖c‖₊ : ℝ≥0∞)⁻¹ := by +lemma egauge_le_of_smul_mem (h : c • x ∈ s) : egauge 𝕜 s x ≤ ‖c‖ₑ⁻¹ := by rcases eq_or_ne c 0 with rfl | hc · simp · exact (egauge_le_of_smul_mem_of_ne h hc).trans ENNReal.coe_inv_le lemma mem_of_egauge_lt_one (hs : Balanced 𝕜 s) (hx : egauge 𝕜 s x < 1) : x ∈ s := let ⟨c, hxc, hc⟩ := egauge_lt_iff.1 hx - hs c (mod_cast hc.le) hxc + hs c (by simpa [enorm] using hc.le) hxc lemma egauge_eq_zero_iff : egauge 𝕜 s x = 0 ↔ ∃ᶠ c : 𝕜 in 𝓝 0, x ∈ c • s := by refine (iInf₂_eq_bot _).trans ?_ @@ -128,25 +127,25 @@ lemma egauge_le_one (h : x ∈ s) : egauge 𝕜 s x ≤ 1 := by variable {𝕜} lemma le_egauge_smul_left (c : 𝕜) (s : Set E) (x : E) : - egauge 𝕜 s x / ‖c‖₊ ≤ egauge 𝕜 (c • s) x := by + egauge 𝕜 s x / ‖c‖ₑ ≤ egauge 𝕜 (c • s) x := by simp_rw [le_egauge_iff, smul_smul] rintro a ⟨x, hx, rfl⟩ apply ENNReal.div_le_of_le_mul - rw [← ENNReal.coe_mul, ← nnnorm_mul] + rw [← enorm_mul] exact egauge_le_of_mem_smul <| smul_mem_smul_set hx lemma egauge_smul_left (hc : c ≠ 0) (s : Set E) (x : E) : - egauge 𝕜 (c • s) x = egauge 𝕜 s x / ‖c‖₊ := by + egauge 𝕜 (c • s) x = egauge 𝕜 s x / ‖c‖ₑ := by refine le_antisymm ?_ (le_egauge_smul_left _ _ _) rw [ENNReal.le_div_iff_mul_le (by simp [*]) (by simp)] calc - egauge 𝕜 (c • s) x * ‖c‖₊ = egauge 𝕜 (c • s) x / ‖c⁻¹‖₊ := by - rw [nnnorm_inv, ENNReal.coe_inv (by simpa), div_eq_mul_inv, inv_inv] + egauge 𝕜 (c • s) x * ‖c‖ₑ = egauge 𝕜 (c • s) x / ‖c⁻¹‖ₑ := by + rw [enorm_inv (by simpa), div_eq_mul_inv, inv_inv] _ ≤ egauge 𝕜 (c⁻¹ • c • s) x := le_egauge_smul_left _ _ _ _ = egauge 𝕜 s x := by rw [inv_smul_smul₀ hc] lemma le_egauge_smul_right (c : 𝕜) (s : Set E) (x : E) : - ‖c‖₊ * egauge 𝕜 s x ≤ egauge 𝕜 s (c • x) := by + ‖c‖ₑ * egauge 𝕜 s x ≤ egauge 𝕜 s (c • x) := by rw [le_egauge_iff] rintro a ⟨y, hy, hxy⟩ rcases eq_or_ne c 0 with rfl | hc @@ -157,51 +156,69 @@ lemma le_egauge_smul_right (c : 𝕜) (s : Set E) (x : E) : simp only [mul_smul, hxy, inv_smul_smul₀ hc] lemma egauge_smul_right (h : c = 0 → s.Nonempty) (x : E) : - egauge 𝕜 s (c • x) = ‖c‖₊ * egauge 𝕜 s x := by + egauge 𝕜 s (c • x) = ‖c‖ₑ * egauge 𝕜 s x := by refine le_antisymm ?_ (le_egauge_smul_right c s x) rcases eq_or_ne c 0 with rfl | hc · simp [egauge_zero_right _ (h rfl)] - · rw [mul_comm, ← ENNReal.div_le_iff_le_mul (.inl <| by simpa) (.inl ENNReal.coe_ne_top), - ENNReal.div_eq_inv_mul, ← ENNReal.coe_inv (by simpa), ← nnnorm_inv] + · rw [mul_comm, ← ENNReal.div_le_iff_le_mul (.inl <| by simpa) (.inl enorm_ne_top), + ENNReal.div_eq_inv_mul, ← enorm_inv (by simpa)] refine (le_egauge_smul_right _ _ _).trans_eq ?_ rw [inv_smul_smul₀ hc] end Module +section VectorSpace + +variable {𝕜 : Type*} [NormedField 𝕜] {E : Type*} [AddCommGroup E] [Module 𝕜 E] + +theorem egauge_add_add_le {U V : Set E} (hU : Balanced 𝕜 U) (hV : Balanced 𝕜 V) (a b : E) : + egauge 𝕜 (U + V) (a + b) ≤ max (egauge 𝕜 U a) (egauge 𝕜 V b) := by + refine le_of_forall_lt' fun c hc ↦ ?_ + simp only [max_lt_iff, egauge_lt_iff] at hc ⊢ + rcases hc with ⟨⟨x, hx, hxc⟩, ⟨y, hy, hyc⟩⟩ + wlog hxy : ‖x‖ ≤ ‖y‖ generalizing a b x y U V + · simpa only [add_comm] using this hV hU b a y hy hyc x hx hxc (le_of_not_le hxy) + refine ⟨y, ?_, hyc⟩ + rw [smul_add] + exact add_mem_add (hU.smul_mono hxy hx) hy + +end VectorSpace + section SeminormedAddCommGroup variable (𝕜 : Type*) [NormedField 𝕜] {E : Type*} [SeminormedAddCommGroup E] [NormedSpace 𝕜 E] -lemma div_le_egauge_closedBall (r : ℝ≥0) (x : E) : ‖x‖₊ / r ≤ egauge 𝕜 (closedBall 0 r) x := by +lemma div_le_egauge_closedBall (r : ℝ≥0) (x : E) : ‖x‖ₑ / r ≤ egauge 𝕜 (closedBall 0 r) x := by rw [le_egauge_iff] rintro c ⟨y, hy, rfl⟩ rw [mem_closedBall_zero_iff, ← coe_nnnorm, NNReal.coe_le_coe] at hy - simp only [nnnorm_smul, ENNReal.coe_mul] + rw [enorm_smul] apply ENNReal.div_le_of_le_mul gcongr + rwa [enorm_le_coe] -lemma le_egauge_closedBall_one (x : E) : ‖x‖₊ ≤ egauge 𝕜 (closedBall 0 1) x := by +lemma le_egauge_closedBall_one (x : E) : ‖x‖ₑ ≤ egauge 𝕜 (closedBall 0 1) x := by simpa using div_le_egauge_closedBall 𝕜 1 x -lemma div_le_egauge_ball (r : ℝ≥0) (x : E) : ‖x‖₊ / r ≤ egauge 𝕜 (ball 0 r) x := +lemma div_le_egauge_ball (r : ℝ≥0) (x : E) : ‖x‖ₑ / r ≤ egauge 𝕜 (ball 0 r) x := (div_le_egauge_closedBall 𝕜 r x).trans <| egauge_anti _ ball_subset_closedBall _ -lemma le_egauge_ball_one (x : E) : ‖x‖₊ ≤ egauge 𝕜 (ball 0 1) x := by +lemma le_egauge_ball_one (x : E) : ‖x‖ₑ ≤ egauge 𝕜 (ball 0 1) x := by simpa using div_le_egauge_ball 𝕜 1 x variable {𝕜} variable {c : 𝕜} {x : E} {r : ℝ≥0} lemma egauge_ball_le_of_one_lt_norm (hc : 1 < ‖c‖) (h₀ : r ≠ 0 ∨ ‖x‖ ≠ 0) : - egauge 𝕜 (ball 0 r) x ≤ ‖c‖₊ * ‖x‖₊ / r := by + egauge 𝕜 (ball 0 r) x ≤ ‖c‖ₑ * ‖x‖ₑ / r := by letI : NontriviallyNormedField 𝕜 := ⟨c, hc⟩ rcases (zero_le r).eq_or_lt with rfl | hr · rw [ENNReal.coe_zero, ENNReal.div_zero (mul_ne_zero _ _)] · apply le_top · simpa using one_pos.trans hc - · simpa [← NNReal.coe_eq_zero] using h₀ + · simpa [enorm, ← NNReal.coe_eq_zero] using h₀ · rcases eq_or_ne ‖x‖ 0 with hx | hx - · have hx' : ‖x‖₊ = 0 := by rwa [← coe_nnnorm, NNReal.coe_eq_zero] at hx + · have hx' : ‖x‖ₑ = 0 := by simpa [enorm, ← coe_nnnorm, NNReal.coe_eq_zero] using hx simp [egauge_eq_zero_iff, hx'] refine (frequently_iff_neBot.2 (inferInstance : NeBot (𝓝[≠] (0 : 𝕜)))).mono fun c hc ↦ ?_ simp [mem_smul_set_iff_inv_smul_mem₀ hc, norm_smul, hx, hr] @@ -210,10 +227,10 @@ lemma egauge_ball_le_of_one_lt_norm (hc : 1 < ‖c‖) (h₀ : r ≠ 0 ∨ ‖x egauge 𝕜 (ball 0 r) x ≤ ↑(‖a‖₊⁻¹) := egauge_le_of_smul_mem_of_ne (mem_ball_zero_iff.2 har) ha₀ _ ≤ ↑(‖c‖₊ * ‖x‖₊ / r) := by rwa [ENNReal.coe_le_coe, div_eq_inv_mul, ← mul_assoc] - _ ≤ ‖c‖₊ * ‖x‖₊ / r := ENNReal.coe_div_le.trans <| by rw [ENNReal.coe_mul] + _ ≤ ‖c‖ₑ * ‖x‖ₑ / r := ENNReal.coe_div_le.trans <| by simp [ENNReal.coe_mul, enorm] lemma egauge_ball_one_le_of_one_lt_norm (hc : 1 < ‖c‖) (x : E) : - egauge 𝕜 (ball 0 1) x ≤ ‖c‖₊ * ‖x‖₊ := by + egauge 𝕜 (ball 0 1) x ≤ ‖c‖ₑ * ‖x‖ₑ := by simpa using egauge_ball_le_of_one_lt_norm hc (.inl one_ne_zero) end SeminormedAddCommGroup diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 616027a97b9a72..f6b41f8197a662 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -300,7 +300,7 @@ theorem lintegral_pow_le_pow_lintegral_fderiv_aux [Fintype ι] {p : ℝ} (hp : Real.IsConjExponent #ι p) {u : (ι → ℝ) → F} (hu : ContDiff ℝ 1 u) (h2u : HasCompactSupport u) : - ∫⁻ x, (‖u x‖₊ : ℝ≥0∞) ^ p ≤ (∫⁻ x, ‖fderiv ℝ u x‖₊) ^ p := by + ∫⁻ x, ‖u x‖ₑ ^ p ≤ (∫⁻ x, ‖fderiv ℝ u x‖ₑ) ^ p := by classical /- For a function `f` in one variable and `t ∈ ℝ` we have `|f(t)| = `|∫_{-∞}^t Df(s)∂s| ≤ ∫_ℝ |Df(s)| ∂s` where we use the fundamental theorem of calculus. @@ -311,43 +311,41 @@ theorem lintegral_pow_le_pow_lintegral_fderiv_aux [Fintype ι] have : (1 : ℝ) ≤ ↑#ι - 1 := by have hι : (2 : ℝ) ≤ #ι := by exact_mod_cast hp.one_lt linarith - calc ∫⁻ x, (‖u x‖₊ : ℝ≥0∞) ^ p - = ∫⁻ x, ((‖u x‖₊ : ℝ≥0∞) ^ (1 / (#ι - 1 : ℝ))) ^ (#ι : ℝ) := by + calc ∫⁻ x, ‖u x‖ₑ ^ p + = ∫⁻ x, (‖u x‖ₑ ^ (1 / (#ι - 1 : ℝ))) ^ (#ι : ℝ) := by -- a little algebraic manipulation of the exponent congr! 2 with x rw [← ENNReal.rpow_mul, hp.conj_eq] field_simp - _ = ∫⁻ x, ∏ _i : ι, (‖u x‖₊ : ℝ≥0∞) ^ (1 / (#ι - 1 : ℝ)) := by + _ = ∫⁻ x, ∏ _i : ι, ‖u x‖ₑ ^ (1 / (#ι - 1 : ℝ)) := by -- express the left-hand integrand as a product of identical factors congr! 2 with x simp_rw [prod_const, card_univ] norm_cast - _ ≤ ∫⁻ x, ∏ i, (∫⁻ xᵢ, ‖fderiv ℝ u (update x i xᵢ)‖₊) ^ ((1 : ℝ) / (#ι - 1 : ℝ)) := ?_ - _ ≤ (∫⁻ x, ‖fderiv ℝ u x‖₊) ^ p := by + _ ≤ ∫⁻ x, ∏ i, (∫⁻ xᵢ, ‖fderiv ℝ u (update x i xᵢ)‖ₑ) ^ ((1 : ℝ) / (#ι - 1 : ℝ)) := ?_ + _ ≤ (∫⁻ x, ‖fderiv ℝ u x‖ₑ) ^ p := by -- apply the grid-lines lemma apply lintegral_prod_lintegral_pow_le _ hp have : Continuous (fderiv ℝ u) := hu.continuous_fderiv le_rfl fun_prop -- we estimate |u x| using the fundamental theorem of calculus. gcongr with x i - calc (‖u x‖₊ : ℝ≥0∞) - _ ≤ ∫⁻ xᵢ in Iic (x i), ‖deriv (u ∘ update x i) xᵢ‖₊ := by - apply le_trans (by simp) (HasCompactSupport.ennnorm_le_lintegral_Ici_deriv _ _ _) + calc ‖u x‖ₑ + _ ≤ ∫⁻ xᵢ in Iic (x i), ‖deriv (u ∘ update x i) xᵢ‖ₑ := by + apply le_trans (by simp) (HasCompactSupport.enorm_le_lintegral_Ici_deriv _ _ _) · exact hu.comp (by convert contDiff_update 1 x i) · exact h2u.comp_isClosedEmbedding (isClosedEmbedding_update x i) - _ ≤ ∫⁻ xᵢ, (‖fderiv ℝ u (update x i xᵢ)‖₊ : ℝ≥0∞) := ?_ + _ ≤ ∫⁻ xᵢ, ‖fderiv ℝ u (update x i xᵢ)‖ₑ := ?_ gcongr · exact Measure.restrict_le_self intro y dsimp - gcongr -- bound the derivative which appears - calc ‖deriv (u ∘ update x i) y‖₊ = ‖fderiv ℝ u (update x i y) (deriv (update x i) y)‖₊ := by + calc ‖deriv (u ∘ update x i) y‖ₑ = ‖fderiv ℝ u (update x i y) (deriv (update x i) y)‖ₑ := by rw [fderiv_comp_deriv _ (hu.differentiable le_rfl).differentiableAt (hasDerivAt_update x i y).differentiableAt] - _ ≤ ‖fderiv ℝ u (update x i y)‖₊ * ‖deriv (update x i) y‖₊ := - ContinuousLinearMap.le_opNNNorm .. - _ ≤ ‖fderiv ℝ u (update x i y)‖₊ := by simp [deriv_update, Pi.nnnorm_single] + _ ≤ ‖fderiv ℝ u (update x i y)‖ₑ * ‖deriv (update x i) y‖ₑ := ContinuousLinearMap.le_opENorm _ _ + _ ≤ ‖fderiv ℝ u (update x i y)‖ₑ := by simp [deriv_update, Pi.enorm_single] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [MeasurableSpace E] [BorelSpace E] [FiniteDimensional ℝ E] (μ : Measure E) [IsAddHaarMeasure μ] @@ -374,8 +372,8 @@ Lebesgue integral of the Fréchet derivative of `u`. -/ theorem lintegral_pow_le_pow_lintegral_fderiv {u : E → F} (hu : ContDiff ℝ 1 u) (h2u : HasCompactSupport u) {p : ℝ} (hp : Real.IsConjExponent (finrank ℝ E) p) : - ∫⁻ x, (‖u x‖₊ : ℝ≥0∞) ^ p ∂μ ≤ - lintegralPowLePowLIntegralFDerivConst μ p * (∫⁻ x, ‖fderiv ℝ u x‖₊ ∂μ) ^ p := by + ∫⁻ x, ‖u x‖ₑ ^ p ∂μ ≤ + lintegralPowLePowLIntegralFDerivConst μ p * (∫⁻ x, ‖fderiv ℝ u x‖ₑ ∂μ) ^ p := by /- We reduce to the case where `E` is `ℝⁿ`, for which we have already proved the result using an explicit basis in `MeasureTheory.lintegral_pow_le_pow_lintegral_fderiv_aux`. This proof is not too hard, but takes quite some steps, reasoning about the equivalence @@ -403,30 +401,28 @@ theorem lintegral_pow_le_pow_lintegral_fderiv {u : E → F} have hv : ContDiff ℝ 1 v := hu.comp e.symm.contDiff have h2v : HasCompactSupport v := h2u.comp_homeomorph e.symm.toHomeomorph have := - calc ∫⁻ x, (‖u x‖₊ : ℝ≥0∞) ^ p ∂(volume : Measure (ι → ℝ)).map e.symm - = ∫⁻ y, (‖v y‖₊ : ℝ≥0∞) ^ p := by + calc ∫⁻ x, ‖u x‖ₑ ^ p ∂(volume : Measure (ι → ℝ)).map e.symm + = ∫⁻ y, ‖v y‖ₑ ^ p := by refine lintegral_map ?_ e.symm.continuous.measurable borelize F exact hu.continuous.measurable.nnnorm.coe_nnreal_ennreal.pow_const _ - _ ≤ (∫⁻ y, ‖fderiv ℝ v y‖₊) ^ p := - lintegral_pow_le_pow_lintegral_fderiv_aux hp hv h2v - _ = (∫⁻ y, ‖(fderiv ℝ u (e.symm y)).comp (fderiv ℝ e.symm y)‖₊) ^ p := by + _ ≤ (∫⁻ y, ‖fderiv ℝ v y‖ₑ) ^ p := lintegral_pow_le_pow_lintegral_fderiv_aux hp hv h2v + _ = (∫⁻ y, ‖(fderiv ℝ u (e.symm y)).comp (fderiv ℝ e.symm y)‖ₑ) ^ p := by congr! with y apply fderiv_comp _ (hu.differentiable le_rfl _) exact e.symm.differentiableAt - _ ≤ (∫⁻ y, ‖fderiv ℝ u (e.symm y)‖₊ * ‖(e.symm : (ι → ℝ) →L[ℝ] E)‖₊) ^ p := by + _ ≤ (∫⁻ y, ‖fderiv ℝ u (e.symm y)‖ₑ * ‖(e.symm : (ι → ℝ) →L[ℝ] E)‖ₑ) ^ p := by gcongr with y - norm_cast rw [e.symm.fderiv] - apply ContinuousLinearMap.opNNNorm_comp_le - _ = (‖(e.symm : (ι → ℝ) →L[ℝ] E)‖₊ * ∫⁻ y, ‖fderiv ℝ u (e.symm y)‖₊) ^ p := by + apply ContinuousLinearMap.opENorm_comp_le + _ = (‖(e.symm : (ι → ℝ) →L[ℝ] E)‖ₑ * ∫⁻ y, ‖fderiv ℝ u (e.symm y)‖ₑ) ^ p := by rw [lintegral_mul_const, mul_comm] refine (Continuous.nnnorm ?_).measurable.coe_nnreal_ennreal exact (hu.continuous_fderiv le_rfl).comp e.symm.continuous - _ = (‖(e.symm : (ι → ℝ) →L[ℝ] E)‖₊ ^ p : ℝ≥0) * (∫⁻ y, ‖fderiv ℝ u (e.symm y)‖₊) ^ p := by - rw [ENNReal.mul_rpow_of_nonneg _ _ h0p, ← ENNReal.coe_rpow_of_nonneg _ h0p] + _ = (‖(e.symm : (ι → ℝ) →L[ℝ] E)‖₊ ^ p : ℝ≥0) * (∫⁻ y, ‖fderiv ℝ u (e.symm y)‖ₑ) ^ p := by + rw [ENNReal.mul_rpow_of_nonneg _ _ h0p, enorm_eq_nnnorm, ← ENNReal.coe_rpow_of_nonneg _ h0p] _ = (‖(e.symm : (ι → ℝ) →L[ℝ] E)‖₊ ^ p : ℝ≥0) - * (∫⁻ x, ‖fderiv ℝ u x‖₊ ∂(volume : Measure (ι → ℝ)).map e.symm) ^ p := by + * (∫⁻ x, ‖fderiv ℝ u x‖ₑ ∂(volume : Measure (ι → ℝ)).map e.symm) ^ p := by congr rw [lintegral_map _ e.symm.continuous.measurable] have : Continuous (fderiv ℝ u) := hu.continuous_fderiv le_rfl @@ -449,7 +445,7 @@ theorem eLpNorm_le_eLpNorm_fderiv_one {u : E → F} (hu : ContDiff ℝ 1 u) (h2 {p : ℝ≥0} (hp : NNReal.IsConjExponent (finrank ℝ E) p) : eLpNorm u p μ ≤ eLpNormLESNormFDerivOneConst μ p * eLpNorm (fderiv ℝ u) 1 μ := by have h0p : 0 < (p : ℝ) := hp.coe.symm.pos - rw [eLpNorm_one_eq_lintegral_nnnorm, + rw [eLpNorm_one_eq_lintegral_enorm, ← ENNReal.rpow_le_rpow_iff h0p, ENNReal.mul_rpow_of_nonneg _ _ h0p.le, ← ENNReal.coe_rpow_of_nonneg _ h0p.le, eLpNormLESNormFDerivOneConst, ← NNReal.rpow_mul, eLpNorm_nnreal_pow_eq_lintegral hp.symm.pos.ne', @@ -535,16 +531,16 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq_inner {u : E → F'} have : (p : ℝ) * (n - 1) - (n - p) = n * (p - 1) := by ring field_simp [this]; ring have h4γ : (γ : ℝ) ≠ 0 := (zero_lt_one.trans h1γ).ne' - by_cases h3u : ∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ = 0 + by_cases h3u : ∫⁻ x, ‖u x‖ₑ ^ (p' : ℝ) ∂μ = 0 · rw [eLpNorm_nnreal_eq_lintegral h0p', h3u, ENNReal.zero_rpow_of_pos] <;> positivity - have h4u : ∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ ≠ ∞ := by - refine lintegral_rpow_nnnorm_lt_top_of_eLpNorm'_lt_top + have h4u : ∫⁻ x, ‖u x‖ₑ ^ (p' : ℝ) ∂μ ≠ ∞ := by + refine lintegral_rpow_enorm_lt_top_of_eLpNorm'_lt_top ((NNReal.coe_pos.trans pos_iff_ne_zero).mpr h0p') ?_ |>.ne rw [← eLpNorm_nnreal_eq_eLpNorm' h0p'] exact hu.continuous.memℒp_of_hasCompactSupport (μ := μ) h2u |>.eLpNorm_lt_top - have h5u : (∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ) ^ (1 / q) ≠ 0 := + have h5u : (∫⁻ x, ‖u x‖ₑ ^ (p' : ℝ) ∂μ) ^ (1 / q) ≠ 0 := ENNReal.rpow_pos (pos_iff_ne_zero.mpr h3u) h4u |>.ne' - have h6u : (∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ) ^ (1 / q) ≠ ∞ := + have h6u : (∫⁻ x, ‖u x‖ₑ ^ (p' : ℝ) ∂μ) ^ (1 / q) ≠ ∞ := ENNReal.rpow_ne_top_of_nonneg (div_nonneg zero_le_one hq.symm.nonneg) h4u have h7u := hu.continuous -- for fun_prop have h8u := (hu.fderiv_right (m := 0) le_rfl).continuous -- for fun_prop @@ -553,20 +549,21 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq_inner {u : E → F'} have h2v : HasCompactSupport v := h2u.norm.rpow_const h4γ set C := eLpNormLESNormFDerivOneConst μ n' have := - calc (∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ) ^ (1 / (n' : ℝ)) = eLpNorm v n' μ := by + calc (∫⁻ x, ‖u x‖ₑ ^ (p' : ℝ) ∂μ) ^ (1 / (n' : ℝ)) = eLpNorm v n' μ := by rw [← h2γ, eLpNorm_nnreal_eq_lintegral hn.symm.pos.ne'] - simp (discharger := positivity) [v, Real.nnnorm_rpow_of_nonneg, ENNReal.rpow_mul, + simp (discharger := positivity) [v, Real.enorm_rpow_of_nonneg, ENNReal.rpow_mul, ← ENNReal.coe_rpow_of_nonneg] _ ≤ C * eLpNorm (fderiv ℝ v) 1 μ := eLpNorm_le_eLpNorm_fderiv_one μ hv h2v hn - _ = C * ∫⁻ x, ‖fderiv ℝ v x‖₊ ∂μ := by rw [eLpNorm_one_eq_lintegral_nnnorm] - _ ≤ C * γ * ∫⁻ x, ‖u x‖₊ ^ ((γ : ℝ) - 1) * ‖fderiv ℝ u x‖₊ ∂μ := by + _ = C * ∫⁻ x, ‖fderiv ℝ v x‖ₑ ∂μ := by rw [eLpNorm_one_eq_lintegral_enorm] + _ ≤ C * γ * ∫⁻ x, ‖u x‖ₑ ^ ((γ : ℝ) - 1) * ‖fderiv ℝ u x‖ₑ ∂μ := by rw [mul_assoc, ← lintegral_const_mul γ] gcongr - simp_rw [← mul_assoc, ← ENNReal.coe_rpow_of_nonneg _ (sub_nonneg.mpr h1γ.le)] - exact ENNReal.coe_le_coe.mpr <| nnnorm_fderiv_norm_rpow_le (hu.differentiable le_rfl) h1γ + simp_rw [← mul_assoc] + exact enorm_fderiv_norm_rpow_le (hu.differentiable le_rfl) h1γ + dsimp [enorm] fun_prop - _ ≤ C * γ * ((∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ) ^ (1 / q) * - (∫⁻ x, ‖fderiv ℝ u x‖₊ ^ (p : ℝ) ∂μ) ^ (1 / (p : ℝ))) := by + _ ≤ C * γ * ((∫⁻ x, ‖u x‖ₑ ^ (p' : ℝ) ∂μ) ^ (1 / q) * + (∫⁻ x, ‖fderiv ℝ u x‖ₑ ^ (p : ℝ) ∂μ) ^ (1 / (p : ℝ))) := by gcongr convert ENNReal.lintegral_mul_le_Lp_mul_Lq μ (.symm <| .conjExponent <| show 1 < (p : ℝ) from hp) ?_ ?_ using 5 @@ -574,12 +571,12 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq_inner {u : E → F'} · borelize F' fun_prop · fun_prop - _ = C * γ * (∫⁻ x, ‖fderiv ℝ u x‖₊ ^ (p : ℝ) ∂μ) ^ (1 / (p : ℝ)) * - (∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ) ^ (1 / q) := by ring + _ = C * γ * (∫⁻ x, ‖fderiv ℝ u x‖ₑ ^ (p : ℝ) ∂μ) ^ (1 / (p : ℝ)) * + (∫⁻ x, ‖u x‖ₑ ^ (p' : ℝ) ∂μ) ^ (1 / q) := by ring calc eLpNorm u p' μ - = (∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ) ^ (1 / (p' : ℝ)) := eLpNorm_nnreal_eq_lintegral h0p' - _ ≤ C * γ * (∫⁻ x, ‖fderiv ℝ u x‖₊ ^ (p : ℝ) ∂μ) ^ (1 / (p : ℝ)) := by + = (∫⁻ x, ‖u x‖ₑ ^ (p' : ℝ) ∂μ) ^ (1 / (p' : ℝ)) := eLpNorm_nnreal_eq_lintegral h0p' + _ ≤ C * γ * (∫⁻ x, ‖fderiv ℝ u x‖ₑ ^ (p : ℝ) ∂μ) ^ (1 / (p : ℝ)) := by rwa [← h2q, ENNReal.rpow_sub _ _ h3u h4u, ENNReal.div_le_iff h5u h6u] _ = eLpNormLESNormFDerivOfEqInnerConst μ p * eLpNorm (fderiv ℝ u) (↑p) μ := by suffices (C : ℝ) * γ = eLpNormLESNormFDerivOfEqInnerConst μ p by @@ -683,8 +680,8 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_le [FiniteDimensional ℝ F] have hp' : p'⁻¹ = p⁻¹ - (finrank ℝ E : ℝ)⁻¹ := by rw [inv_inv, NNReal.coe_sub] · simp - · #adaptation_note - /-- This should just be `gcongr`, but this is not working as of nightly-2024-11-20. + · #adaptation_note /-- nightly-2024-11-20 + This should just be `gcongr`, but this is not working as of nightly-2024-11-20. Possibly related to #19262 (since this proof fails at `with_reducible_and_instances`). -/ exact inv_anti₀ (by positivity) h2p.le have : (q : ℝ≥0∞) ≤ p' := by @@ -694,8 +691,8 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_le [FiniteDimensional ℝ F] · dsimp have : 0 < p⁻¹ - (finrank ℝ E : ℝ≥0)⁻¹ := by simp only [tsub_pos_iff_lt] - #adaptation_note - /-- This should just be `gcongr`, but this is not working as of nightly-2024-11-20. + #adaptation_note /-- nightly-2024-11-20 + This should just be `gcongr`, but this is not working as of nightly-2024-11-20. Possibly related to #19262 (since this proof fails at `with_reducible_and_instances`). -/ exact inv_strictAnti₀ (by positivity) h2p positivity diff --git a/Mathlib/Analysis/LocallyConvex/Polar.lean b/Mathlib/Analysis/LocallyConvex/Polar.lean index af5dc28df843c0..d49aedb9e44496 100644 --- a/Mathlib/Analysis/LocallyConvex/Polar.lean +++ b/Mathlib/Analysis/LocallyConvex/Polar.lean @@ -139,7 +139,7 @@ variable (B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜) theorem polar_univ (h : SeparatingRight B) : B.polar Set.univ = {(0 : F)} := by rw [Set.eq_singleton_iff_unique_mem] refine ⟨by simp only [zero_mem_polar], fun y hy => h _ fun x => ?_⟩ - refine norm_le_zero_iff.mp (le_of_forall_le_of_dense fun ε hε => ?_) + refine norm_le_zero_iff.mp (le_of_forall_gt_imp_ge_of_dense fun ε hε => ?_) rcases NormedField.exists_norm_lt 𝕜 hε with ⟨c, hc, hcε⟩ calc ‖B x y‖ = ‖c‖ * ‖B (c⁻¹ • x) y‖ := by diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index 4f081227992a49..649fa425cabf54 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -184,10 +184,8 @@ theorem exists_nnnorm_eq_spectralRadius_of_nonempty [ProperSpace 𝕜] {a : A} ( theorem spectralRadius_lt_of_forall_lt_of_nonempty [ProperSpace 𝕜] {a : A} (ha : (σ a).Nonempty) {r : ℝ≥0} (hr : ∀ k ∈ σ a, ‖k‖₊ < r) : spectralRadius 𝕜 a < r := - sSup_image.symm.trans_lt <| - ((spectrum.isCompact a).sSup_lt_iff_of_continuous ha - (ENNReal.continuous_coe.comp continuous_nnnorm).continuousOn (r : ℝ≥0∞)).mpr - (by dsimp only [(· ∘ ·)]; exact mod_cast hr) + sSup_image.symm.trans_lt <| ((spectrum.isCompact a).sSup_lt_iff_of_continuous ha + continuous_enorm.continuousOn (r : ℝ≥0∞)).mpr (by simpa using hr) open ENNReal Polynomial diff --git a/Mathlib/Analysis/Normed/Field/Ultra.lean b/Mathlib/Analysis/Normed/Field/Ultra.lean index 8ef392b9b1afd4..79e9e92787437d 100644 --- a/Mathlib/Analysis/Normed/Field/Ultra.lean +++ b/Mathlib/Analysis/Normed/Field/Ultra.lean @@ -76,7 +76,7 @@ lemma isUltrametricDist_of_forall_pow_norm_le_nsmul_pow_max_one_norm -- that avoids explicitly mentioning `m`-th roots. -- First note it suffices to show that `‖x + 1‖ ≤ a` for all `a : ℝ` with `max ‖x‖ 1 < a`. rw [max_comm] - refine le_of_forall_le_of_dense fun a ha ↦ ?_ + refine le_of_forall_gt_imp_ge_of_dense fun a ha ↦ ?_ have ha' : 1 < a := (max_lt_iff.mp ha).left -- `max 1 ‖x‖ < a`, so there must be some `m : ℕ` such that `m + 1 < (a / max 1 ‖x‖) ^ m` -- by the virtue of exponential growth being faster than linear growth diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 0b033f8d126185..d773e1f72c4f5c 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -686,8 +686,7 @@ alias nndist_eq_nnnorm := nndist_eq_nnnorm_sub theorem nndist_one_right (a : E) : nndist a 1 = ‖a‖₊ := by simp [nndist_eq_nnnorm_div] @[to_additive (attr := simp)] -theorem edist_one_right (a : E) : edist a 1 = ‖a‖₊ := by - rw [edist_nndist, nndist_one_right] +lemma edist_one_right (a : E) : edist a 1 = ‖a‖ₑ := by simp [edist_nndist, nndist_one_right, enorm] @[to_additive (attr := simp) nnnorm_zero] theorem nnnorm_one' : ‖(1 : E)‖₊ = 0 := NNReal.eq norm_one' @@ -800,44 +799,51 @@ lemma nnnorm_div_eq_nnnorm_right {x : E} (y : E) (h : ‖x‖₊ = 0) : ‖x / y lemma nnnorm_div_eq_nnnorm_left (x : E) {y : E} (h : ‖y‖₊ = 0) : ‖x / y‖₊ = ‖x‖₊ := NNReal.eq <| norm_div_eq_norm_left _ <| congr_arg NNReal.toReal h -@[to_additive ofReal_norm_eq_coe_nnnorm] -theorem ofReal_norm_eq_coe_nnnorm' (a : E) : ENNReal.ofReal ‖a‖ = ‖a‖₊ := - ENNReal.ofReal_eq_coe_nnreal _ - /-- The non negative norm seen as an `ENNReal` and then as a `Real` is equal to the norm. -/ @[to_additive toReal_coe_nnnorm "The non negative norm seen as an `ENNReal` and then as a `Real` is equal to the norm."] theorem toReal_coe_nnnorm' (a : E) : (‖a‖₊ : ℝ≥0∞).toReal = ‖a‖ := rfl -@[to_additive] -theorem edist_eq_coe_nnnorm_div (a b : E) : edist a b = ‖a / b‖₊ := by - rw [edist_dist, dist_eq_norm_div, ofReal_norm_eq_coe_nnnorm'] - -@[to_additive edist_eq_coe_nnnorm] -theorem edist_eq_coe_nnnorm' (x : E) : edist x 1 = (‖x‖₊ : ℝ≥0∞) := by - rw [edist_eq_coe_nnnorm_div, div_one] - open scoped symmDiff in @[to_additive] theorem edist_mulIndicator (s t : Set α) (f : α → E) (x : α) : edist (s.mulIndicator f x) (t.mulIndicator f x) = ‖(s ∆ t).mulIndicator f x‖₊ := by rw [edist_nndist, nndist_mulIndicator] -@[to_additive] -theorem mem_emetric_ball_one_iff {r : ℝ≥0∞} : a ∈ EMetric.ball (1 : E) r ↔ ↑‖a‖₊ < r := by - rw [EMetric.mem_ball, edist_eq_coe_nnnorm'] - end NNNorm section ENorm @[to_additive (attr := simp) enorm_zero] lemma enorm_one' : ‖(1 : E)‖ₑ = 0 := by simp [enorm] +@[to_additive ofReal_norm_eq_enorm] +lemma ofReal_norm_eq_enorm' (a : E) : .ofReal ‖a‖ = ‖a‖ₑ := ENNReal.ofReal_eq_coe_nnreal _ + +@[deprecated (since := "2025-01-17")] alias ofReal_norm_eq_coe_nnnorm := ofReal_norm_eq_enorm +@[deprecated (since := "2025-01-17")] alias ofReal_norm_eq_coe_nnnorm' := ofReal_norm_eq_enorm' + instance : ENorm ℝ≥0∞ where enorm x := x @[simp] lemma enorm_eq_self (x : ℝ≥0∞) : ‖x‖ₑ = x := rfl +@[to_additive] +theorem edist_eq_enorm_div (a b : E) : edist a b = ‖a / b‖ₑ := by + rw [edist_dist, dist_eq_norm_div, ofReal_norm_eq_enorm'] + +@[deprecated (since := "2025-01-17")] alias edist_eq_coe_nnnorm_sub := edist_eq_enorm_sub +@[deprecated (since := "2025-01-17")] alias edist_eq_coe_nnnorm_div := edist_eq_enorm_div + +@[to_additive] +theorem edist_one_eq_enorm (x : E) : edist x 1 = ‖x‖ₑ := by rw [edist_eq_enorm_div, div_one] + +@[deprecated (since := "2025-01-17")] alias edist_eq_coe_nnnorm := edist_zero_eq_enorm +@[deprecated (since := "2025-01-17")] alias edist_eq_coe_nnnorm' := edist_one_eq_enorm + +@[to_additive] +theorem mem_emetric_ball_one_iff {r : ℝ≥0∞} : a ∈ EMetric.ball 1 r ↔ ‖a‖ₑ < r := by + rw [EMetric.mem_ball, edist_one_eq_enorm] + end ENorm @[to_additive] @@ -1319,11 +1325,15 @@ lemma enorm_of_nonneg (hr : 0 ≤ r) : ‖r‖ₑ = .ofReal r := by @[simp] lemma nnnorm_abs (r : ℝ) : ‖|r|‖₊ = ‖r‖₊ := by simp [nnnorm] @[simp] lemma enorm_abs (r : ℝ) : ‖|r|‖ₑ = ‖r‖ₑ := by simp [enorm] -theorem ennnorm_eq_ofReal (hr : 0 ≤ r) : (‖r‖₊ : ℝ≥0∞) = ENNReal.ofReal r := by - rw [← ofReal_norm_eq_coe_nnnorm, norm_of_nonneg hr] +theorem enorm_eq_ofReal (hr : 0 ≤ r) : ‖r‖ₑ = .ofReal r := by + rw [← ofReal_norm_eq_enorm, norm_of_nonneg hr] + +@[deprecated (since := "2025-01-17")] alias ennnorm_eq_ofReal := enorm_eq_ofReal -theorem ennnorm_eq_ofReal_abs (r : ℝ) : (‖r‖₊ : ℝ≥0∞) = ENNReal.ofReal |r| := by - rw [← Real.nnnorm_abs r, Real.ennnorm_eq_ofReal (abs_nonneg _)] +theorem enorm_eq_ofReal_abs (r : ℝ) : ‖r‖ₑ = ENNReal.ofReal |r| := by + rw [← enorm_eq_ofReal (abs_nonneg _), enorm_abs] + +@[deprecated (since := "2025-01-17")] alias ennnorm_eq_ofReal_abs := enorm_eq_ofReal_abs theorem toNNReal_eq_nnnorm_of_nonneg (hr : 0 ≤ r) : r.toNNReal = ‖r‖₊ := by rw [Real.toNNReal_of_nonneg hr] @@ -1331,12 +1341,10 @@ theorem toNNReal_eq_nnnorm_of_nonneg (hr : 0 ≤ r) : r.toNNReal = ‖r‖₊ := rw [coe_mk, coe_nnnorm r, Real.norm_eq_abs r, abs_of_nonneg hr] -- Porting note: this is due to the change from `Subtype.val` to `NNReal.toReal` for the coercion -theorem ofReal_le_ennnorm (r : ℝ) : ENNReal.ofReal r ≤ ‖r‖₊ := by - obtain hr | hr := le_total 0 r - · exact (Real.ennnorm_eq_ofReal hr).ge - · rw [ENNReal.ofReal_eq_zero.2 hr] - exact bot_le --- Porting note: should this be renamed to `Real.ofReal_le_nnnorm`? +theorem ofReal_le_enorm (r : ℝ) : ENNReal.ofReal r ≤ ‖r‖ₑ := by + rw [enorm_eq_ofReal_abs]; gcongr; exact le_abs_self _ + +@[deprecated (since := "2025-01-17")] alias ofReal_le_ennnorm := ofReal_le_enorm end Real diff --git a/Mathlib/Analysis/Normed/Lp/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean index 0c311a281bb35c..dfd02e7678382b 100644 --- a/Mathlib/Analysis/Normed/Lp/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -95,13 +95,10 @@ variable [Semiring 𝕜] [∀ i, SeminormedAddCommGroup (β i)] variable [∀ i, Module 𝕜 (β i)] (c : 𝕜) variable (x y : PiLp p β) (i : ι) -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/4481 +#adaptation_note /-- https://github.com/leanprover/lean4/pull/4481 the `simpNF` linter incorrectly claims this lemma can't be applied by `simp`. -(It appears to also be unused in Mathlib.) --/ +(It appears to also be unused in Mathlib.) -/ @[simp, nolint simpNF] theorem zero_apply : (0 : PiLp p β) i = 0 := rfl diff --git a/Mathlib/Analysis/Normed/Module/Dual.lean b/Mathlib/Analysis/Normed/Module/Dual.lean index 8a947e649a0b2e..279e098d12e2c2 100644 --- a/Mathlib/Analysis/Normed/Module/Dual.lean +++ b/Mathlib/Analysis/Normed/Module/Dual.lean @@ -240,7 +240,7 @@ theorem polar_ball {𝕜 E : Type*} [RCLike 𝕜] [NormedAddCommGroup E] [Normed apply le_antisymm · intro x hx rw [mem_closedBall_zero_iff] - apply le_of_forall_le_of_dense + apply le_of_forall_gt_imp_ge_of_dense intro a ha rw [← mem_closedBall_zero_iff, ← (mul_div_cancel_left₀ a (Ne.symm (ne_of_lt hr)))] rw [← RCLike.norm_of_nonneg (K := 𝕜) (le_trans zero_le_one diff --git a/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean b/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean index 9ad44b68f8091d..72c464afef4f11 100644 --- a/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean +++ b/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean @@ -193,12 +193,9 @@ theorem isBoundedLinearMap_prod_multilinear {E : ι → Type*} [∀ i, Seminorme (ContinuousMultilinearMap.prodL 𝕜 E F G).toContinuousLinearEquiv |>.toContinuousLinearMap.isBoundedLinearMap -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/6024 +#adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 we needed to add the named arguments `(ι := ι) (G := F)` -to `ContinuousMultilinearMap.compContinuousLinearMapL`. --/ +to `ContinuousMultilinearMap.compContinuousLinearMapL`. -/ /-- Given a fixed continuous linear map `g`, associating to a continuous multilinear map `f` the continuous multilinear map `f (g m₁, ..., g mₙ)` is a bounded linear operation. -/ theorem isBoundedLinearMap_continuousMultilinearMap_comp_linear (g : G →L[𝕜] E) : diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean index 59589fe9246181..4ce0cf8cdda550 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean @@ -145,8 +145,7 @@ variable {E : Type u} [NormedAddCommGroup E] [NormedSpace 𝕜 E] open ContinuousLinearEquiv Submodule theorem coord_norm' {x : E} (h : x ≠ 0) : ‖(‖x‖ : 𝕜) • coord 𝕜 x h‖ = 1 := by - #adaptation_note - /-- + #adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 `set_option maxSynthPendingDepth 2` required after https://github.com/leanprover/lean4/pull/4119 Alternatively, we can add: ``` diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean index 8193941231cfa9..471a7a72a27a8a 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean @@ -146,11 +146,8 @@ def flip (f : E →SL[σ₁₃] F →SL[σ₂₃] G) : F →SL[σ₂₃] E →SL ‖f‖ fun y x => (f.le_opNorm₂ x y).trans_eq <| by simp only [mul_right_comm] private theorem le_norm_flip (f : E →SL[σ₁₃] F →SL[σ₂₃] G) : ‖f‖ ≤ ‖flip f‖ := - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/4119 we either need - to specify the `f.flip` argument, or use `set_option maxSynthPendingDepth 2 in`. - -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 + we either need to specify the `f.flip` argument, or use `set_option maxSynthPendingDepth 2 in`. -/ f.opNorm_le_bound₂ (norm_nonneg f.flip) fun x y => by rw [mul_right_comm] exact (flip f).le_opNorm₂ y x diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean index 45424cd36fd7de..aec0cfaa5eab13 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean @@ -73,3 +73,47 @@ theorem AnalyticOn.cpow (fs : AnalyticOn ℂ f s) (gs : AnalyticOn ℂ g s) theorem AnalyticOnNhd.cpow (fs : AnalyticOnNhd ℂ f s) (gs : AnalyticOnNhd ℂ g s) (m : ∀ z ∈ s, f z ∈ slitPlane) : AnalyticOnNhd ℂ (fun z ↦ f z ^ g z) s := fun z n ↦ (fs z n).cpow (gs z n) (m z n) + +section ReOfReal + +variable {f : ℂ → ℂ} {s : Set ℝ} {x : ℝ} + +lemma AnalyticAt.re_ofReal (hf : AnalyticAt ℂ f x) : + AnalyticAt ℝ (fun x : ℝ ↦ (f x).re) x := + (Complex.reCLM.analyticAt _).comp (hf.restrictScalars.comp (Complex.ofRealCLM.analyticAt _)) + +lemma AnalyticAt.im_ofReal (hf : AnalyticAt ℂ f x) : + AnalyticAt ℝ (fun x : ℝ ↦ (f x).im) x := + (Complex.imCLM.analyticAt _).comp (hf.restrictScalars.comp (Complex.ofRealCLM.analyticAt _)) + +lemma AnalyticWithinAt.re_ofReal (hf : AnalyticWithinAt ℂ f (ofReal '' s) x) : + AnalyticWithinAt ℝ (fun x : ℝ ↦ (f x).re) s x := + ((Complex.reCLM.analyticWithinAt _ _).comp hf.restrictScalars (mapsTo_image f _)).comp + (Complex.ofRealCLM.analyticWithinAt _ _) (mapsTo_image ofReal s) + +lemma AnalyticWithinAt.im_ofReal (hf : AnalyticWithinAt ℂ f (ofReal '' s) x) : + AnalyticWithinAt ℝ (fun x : ℝ ↦ (f x).im) s x := + ((Complex.imCLM.analyticWithinAt _ _).comp hf.restrictScalars (mapsTo_image f _)).comp + (Complex.ofRealCLM.analyticWithinAt _ _) (mapsTo_image ofReal s) + +lemma AnalyticOn.re_ofReal (hf : AnalyticOn ℂ f (ofReal '' s)) : + AnalyticOn ℝ (fun x : ℝ ↦ (f x).re) s := + ((Complex.reCLM.analyticOn _).comp hf.restrictScalars (mapsTo_image f _)).comp + (Complex.ofRealCLM.analyticOn _) (mapsTo_image ofReal s) + +lemma AnalyticOn.im_ofReal (hf : AnalyticOn ℂ f (ofReal '' s)) : + AnalyticOn ℝ (fun x : ℝ ↦ (f x).im) s := + ((Complex.imCLM.analyticOn _).comp hf.restrictScalars (mapsTo_image f _)).comp + (Complex.ofRealCLM.analyticOn _) (mapsTo_image ofReal s) + +lemma AnalyticOnNhd.re_ofReal (hf : AnalyticOnNhd ℂ f (ofReal '' s)) : + AnalyticOnNhd ℝ (fun x : ℝ ↦ (f x).re) s := + ((Complex.reCLM.analyticOnNhd _).comp hf.restrictScalars (mapsTo_image f _)).comp + (Complex.ofRealCLM.analyticOnNhd _) (mapsTo_image ofReal s) + +lemma AnalyticOnNhd.im_ofReal (hf : AnalyticOnNhd ℂ f (ofReal '' s)) : + AnalyticOnNhd ℝ (fun x : ℝ ↦ (f x).im) s := + ((Complex.imCLM.analyticOnNhd _).comp hf.restrictScalars (mapsTo_image f _)).comp + (Complex.ofRealCLM.analyticOnNhd _) (mapsTo_image ofReal s) + +end ReOfReal diff --git a/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean b/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean index 3ef729e02e9f10..4e8ea696d800af 100644 --- a/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean @@ -81,6 +81,14 @@ theorem integrableOn_Ioi_rpow_iff {s t : ℝ} (ht : 0 < t) : exact Real.rpow_le_rpow_of_exponent_le x_one h exact not_IntegrableOn_Ioi_inv this +theorem integrableAtFilter_rpow_atTop_iff {s : ℝ} : + IntegrableAtFilter (fun x : ℝ ↦ x ^ s) atTop ↔ s < -1 := by + refine ⟨fun ⟨t, ht, hint⟩ ↦ ?_, fun h ↦ + ⟨Set.Ioi 1, Ioi_mem_atTop 1, (integrableOn_Ioi_rpow_iff zero_lt_one).mpr h⟩⟩ + obtain ⟨a, ha⟩ := mem_atTop_sets.mp ht + refine (integrableOn_Ioi_rpow_iff (zero_lt_one.trans_le (le_max_right a 1))).mp ?_ + exact hint.mono_set <| fun x hx ↦ ha _ <| (le_max_left a 1).trans hx.le + /-- 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 @@ -127,6 +135,24 @@ theorem integrableOn_Ioi_cpow_iff {s : ℂ} {t : ℝ} (ht : 0 < t) : simp [Complex.abs_cpow_eq_rpow_re_of_pos this] rwa [integrableOn_Ioi_rpow_iff ht] at B +theorem integrableOn_Ioi_deriv_ofReal_cpow {s : ℂ} {t : ℝ} (ht : 0 < t) (hs : s.re < 0) : + IntegrableOn (deriv fun x : ℝ ↦ (x : ℂ) ^ s) (Set.Ioi t) := by + have h : IntegrableOn (fun x : ℝ ↦ s * x ^ (s - 1)) (Set.Ioi t) := by + refine (integrableOn_Ioi_cpow_of_lt ?_ ht).const_mul _ + rwa [Complex.sub_re, Complex.one_re, sub_lt_iff_lt_add, neg_add_cancel] + refine h.congr_fun (fun x hx ↦ ?_) measurableSet_Ioi + rw [Complex.deriv_ofReal_cpow_const (ht.trans hx).ne' (fun h ↦ (Complex.zero_re ▸ h ▸ hs).false)] + +theorem integrableOn_Ioi_deriv_norm_ofReal_cpow {s : ℂ} {t : ℝ} (ht : 0 < t) (hs : s.re ≤ 0): + IntegrableOn (deriv fun x : ℝ ↦ ‖(x : ℂ) ^ s‖) (Set.Ioi t) := by + rw [integrableOn_congr_fun (fun x hx ↦ by + rw [deriv_norm_ofReal_cpow _ (ht.trans hx)]) measurableSet_Ioi] + obtain hs | hs := eq_or_lt_of_le hs + · simp_rw [hs, zero_mul] + exact integrableOn_zero + · replace hs : s.re - 1 < - 1 := by rwa [sub_lt_iff_lt_add, neg_add_cancel] + exact (integrableOn_Ioi_rpow_of_lt hs ht).const_mul s.re + /-- 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 diff --git a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean index 595b8aa9b7f348..a773a828ece6ab 100644 --- a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean +++ b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean @@ -132,9 +132,9 @@ theorem integrable_one_add_norm {r : ℝ} (hnr : (finrank ℝ E : ℝ) < r) : constructor · apply Measurable.aestronglyMeasurable (by fun_prop) -- Lower Lebesgue integral - have : (∫⁻ a : E, ‖(1 + ‖a‖) ^ (-r)‖₊ ∂μ) = ∫⁻ a : E, ENNReal.ofReal ((1 + ‖a‖) ^ (-r)) ∂μ := - lintegral_nnnorm_eq_of_nonneg fun _ => rpow_nonneg (by positivity) _ - rw [hasFiniteIntegral_iff_nnnorm, this] + have : (∫⁻ a : E, ‖(1 + ‖a‖) ^ (-r)‖ₑ ∂μ) = ∫⁻ a : E, ENNReal.ofReal ((1 + ‖a‖) ^ (-r)) ∂μ := + lintegral_enorm_of_nonneg fun _ => rpow_nonneg (by positivity) _ + rw [hasFiniteIntegral_iff_enorm, this] exact finite_integral_one_add_norm hnr theorem integrable_rpow_neg_one_add_norm_sq {r : ℝ} (hnr : (finrank ℝ E : ℝ) < r) : diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Summable.lean b/Mathlib/Analysis/SpecialFunctions/Log/Summable.lean index 3c1969101ea969..0bb757f1e28ba1 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Summable.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Summable.lean @@ -22,7 +22,7 @@ products. open Filter Function Complex Real -open scoped Interval Topology BigOperators Nat Classical Complex +open scoped Interval Topology BigOperators Nat Complex variable {ι : Type*} diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean index 0407d95d6e2ccc..30cb5251b049e8 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean @@ -42,6 +42,13 @@ theorem cpow_eq_zero_iff (x y : ℂ) : x ^ y = 0 ↔ x = 0 ∧ y ≠ 0 := by simp only [cpow_def] split_ifs <;> simp [*, exp_ne_zero] +theorem cpow_ne_zero_iff {x y : ℂ} : + x ^ y ≠ 0 ↔ x ≠ 0 ∨ y = 0 := by + rw [ne_eq, cpow_eq_zero_iff, not_and_or, ne_eq, not_not] + +theorem cpow_ne_zero_iff_of_exponent_ne_zero {x y : ℂ} (hy : y ≠ 0) : + x ^ y ≠ 0 ↔ x ≠ 0 := by simp [hy] + @[simp] theorem zero_cpow {x : ℂ} (h : x ≠ 0) : (0 : ℂ) ^ x = 0 := by simp [cpow_def, *] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index 901b5b2038c028..88657855528d15 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -623,6 +623,13 @@ theorem deriv_rpow_const (hf : DifferentiableAt ℝ f x) (hx : f x ≠ 0 ∨ 1 deriv (fun x => f x ^ p) x = deriv f x * p * f x ^ (p - 1) := (hf.hasDerivAt.rpow_const hx).deriv +theorem deriv_norm_ofReal_cpow (c : ℂ) {t : ℝ} (ht : 0 < t) : + (deriv fun x : ℝ ↦ ‖(x : ℂ) ^ c‖) t = c.re * t ^ (c.re - 1) := by + rw [EventuallyEq.deriv_eq (f := fun x ↦ x ^ c.re)] + · rw [Real.deriv_rpow_const (Or.inl ht.ne')] + · filter_upwards [eventually_gt_nhds ht] with x hx + rw [Complex.norm_eq_abs, Complex.abs_cpow_eq_rpow_re_of_pos hx] + lemma isTheta_deriv_rpow_const_atTop {p : ℝ} (hp : p ≠ 0) : deriv (fun (x : ℝ) => x ^ p) =Θ[atTop] fun x => x ^ (p-1) := by calc deriv (fun (x : ℝ) => x ^ p) =ᶠ[atTop] fun x => p * x ^ (p - 1) := by diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index a49168a3cabb83..0a8b78fdc039e3 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -318,6 +318,12 @@ theorem abs_cpow_eq_rpow_re_of_nonneg {x : ℝ} (hx : 0 ≤ x) {y : ℂ} (hy : r abs (x ^ y) = x ^ re y := by rw [abs_cpow_of_imp] <;> simp [*, arg_ofReal_of_nonneg, _root_.abs_of_nonneg] +open Filter in +lemma norm_ofReal_cpow_eventually_eq_atTop (c : ℂ) : + (fun t : ℝ ↦ ‖(t : ℂ) ^ c‖) =ᶠ[atTop] fun t ↦ t ^ c.re := by + filter_upwards [eventually_gt_atTop 0] with t ht + rw [Complex.norm_eq_abs, abs_cpow_eq_rpow_re_of_pos ht] + lemma norm_natCast_cpow_of_re_ne_zero (n : ℕ) {s : ℂ} (hs : s.re ≠ 0) : ‖(n : ℂ) ^ s‖ = (n : ℝ) ^ (s.re) := by rw [norm_eq_abs, ← ofReal_natCast, abs_cpow_eq_rpow_re_of_nonneg n.cast_nonneg hs] @@ -962,6 +968,35 @@ lemma norm_log_natCast_le_rpow_div (n : ℕ) {ε : ℝ} (hε : 0 < ε) : ‖log end Complex +namespace Asymptotics + +open Filter + +variable {E : Type*} [SeminormedRing E] (a b c : ℝ) + +theorem IsBigO.mul_atTop_rpow_of_isBigO_rpow {f g : ℝ → E} + (hf : f =O[atTop] fun t ↦ (t : ℝ) ^ a) (hg : g =O[atTop] fun t ↦ (t : ℝ) ^ b) + (h : a + b ≤ c) : + (f * g) =O[atTop] fun t ↦ (t : ℝ) ^ c := by + refine (hf.mul hg).trans (Eventually.isBigO ?_) + filter_upwards [eventually_ge_atTop 1] with t ht + rw [← Real.rpow_add (zero_lt_one.trans_le ht), Real.norm_of_nonneg (Real.rpow_nonneg + (zero_le_one.trans ht) (a + b))] + exact Real.rpow_le_rpow_of_exponent_le ht h + +theorem IsBigO.mul_atTop_rpow_natCast_of_isBigO_rpow {f g : ℕ → E} + (hf : f =O[atTop] fun n ↦ (n : ℝ) ^ a) (hg : g =O[atTop] fun n ↦ (n : ℝ) ^ b) + (h : a + b ≤ c) : + (f * g) =O[atTop] fun n ↦ (n : ℝ) ^ c := by + refine (hf.mul hg).trans (Eventually.isBigO ?_) + filter_upwards [eventually_ge_atTop 1] with t ht + replace ht : 1 ≤ (t : ℝ) := Nat.one_le_cast.mpr ht + rw [← Real.rpow_add (zero_lt_one.trans_le ht), Real.norm_of_nonneg (Real.rpow_nonneg + (zero_le_one.trans ht) (a + b))] + exact Real.rpow_le_rpow_of_exponent_le ht h + +end Asymptotics + /-! ## Square roots of reals -/ @@ -1068,14 +1103,6 @@ theorem isRat_rpow_neg {a b : ℝ} {nb : ℕ} IsRat (a ^ b) num den := by rwa [pb.out, Real.rpow_intCast] -#adaptation_note -/-- -Since https://github.com/leanprover/lean4/pull/5338, -the unused variable linter can not see usages of variables in -`haveI' : ⋯ =Q ⋯ := ⟨⟩` clauses, so generates many false positives. --/ -set_option linter.unusedVariables false - /-- Evaluates expressions of the form `a ^ b` when `a` and `b` are both reals. -/ @[norm_num (_ : ℝ) ^ (_ : ℝ)] def evalRPow : NormNumExt where eval {u α} e := do diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean index 049879b8369028..e56f20010f8cf1 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import Mathlib.Data.Complex.Exponential +import Mathlib.Data.Complex.Trigonometric import Mathlib.Data.Complex.Module import Mathlib.RingTheory.Polynomial.Chebyshev diff --git a/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean b/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean index 5812a21f568584..b81f5bddd383a9 100644 --- a/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean +++ b/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean @@ -128,12 +128,6 @@ section Five variable {R₁ R₂ : ComposableArrows C 4} (hR₁ : R₁.Exact) (hR₂ : R₂.Exact) (φ : R₁ ⟶ R₂) include hR₁ hR₂ -#adaptation_note /-- nightly-2024-03-11 -We turn off simprocs here. -Ideally someone will investigate whether `simp` lemmas can be rearranged -so that this works without the `set_option`, -*or* come up with a proposal regarding finer control of disabling simprocs. -/ -set_option simprocs false in /-- The five lemma. -/ theorem isIso_of_epi_of_isIso_of_isIso_of_mono (h₀ : Epi (app' φ 0)) (h₁ : IsIso (app' φ 1)) (h₂ : IsIso (app' φ 3)) (h₃ : Mono (app' φ 4)) : IsIso (app' φ 2) := by diff --git a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean index 9110373e9c67a4..dac1bafa0aea28 100644 --- a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean @@ -298,7 +298,6 @@ lemma ofCocomplex_d_0_1 : (ofCocomplex Z).d 0 1 = d (Injective.ι Z) := by simp [ofCocomplex] -#adaptation_note /-- Since nightly-2024-03-11, this takes forever now -/ lemma ofCocomplex_exactAt_succ (n : ℕ) : (ofCocomplex Z).ExactAt (n + 1) := by rw [HomologicalComplex.exactAt_iff' _ n (n + 1) (n + 1 + 1) (by simp) (by simp)] diff --git a/Mathlib/CategoryTheory/Abelian/Transfer.lean b/Mathlib/CategoryTheory/Abelian/Transfer.lean index 20ca075bae7579..114d4d2a53902d 100644 --- a/Mathlib/CategoryTheory/Abelian/Transfer.lean +++ b/Mathlib/CategoryTheory/Abelian/Transfer.lean @@ -8,6 +8,7 @@ import Mathlib.CategoryTheory.Abelian.Basic import Mathlib.CategoryTheory.Adjunction.Limits import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Kernels import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor +import Mathlib.CategoryTheory.Preadditive.Transfer /-! # Transferring "abelian-ness" across a functor @@ -202,21 +203,14 @@ lemma inverse_map_add {P Q : ShrinkHoms C} (f g : P ⟶ Q) : variable (C) -noncomputable instance preadditive : - Preadditive.{w} (ShrinkHoms C) where - homGroup := homGroup - add_comp _ _ _ _ _ _ := by - apply (inverse C).map_injective - simp only [inverse_map_add, Functor.map_comp, Preadditive.add_comp] - comp_add _ _ _ _ _ _ := by - apply (inverse C).map_injective - simp only [inverse_map_add, Functor.map_comp, Preadditive.comp_add] - -instance : (inverse C).Additive where - map_add := by apply inverse_map_add - -instance : (functor C).Additive where - map_add := by apply functor_map_add +instance preadditive : Preadditive.{w} (ShrinkHoms C) := + .ofFullyFaithful (equivalence C).fullyFaithfulInverse + +instance : (inverse C).Additive := + (equivalence C).symm.fullyFaithfulFunctor.additive_ofFullyFaithful + +instance : (functor C).Additive := + (equivalence C).symm.additive_inverse_of_FullyFaithful instance hasLimitsOfShape (J : Type*) [Category J] [HasLimitsOfShape J C] : HasLimitsOfShape.{_, _, w} J (ShrinkHoms C) := diff --git a/Mathlib/CategoryTheory/Action/Basic.lean b/Mathlib/CategoryTheory/Action/Basic.lean index 1c61d3fcf4c62b..f68dccaad46f25 100644 --- a/Mathlib/CategoryTheory/Action/Basic.lean +++ b/Mathlib/CategoryTheory/Action/Basic.lean @@ -42,7 +42,7 @@ namespace Action variable {V} -@[simp 1100] +@[simp] theorem ρ_one {G : MonCat.{u}} (A : Action V G) : A.ρ 1 = 𝟙 A.V := by rw [MonoidHom.map_one]; rfl /-- When a group acts, we can lift the action to the group of automorphisms. -/ diff --git a/Mathlib/CategoryTheory/Adjunction/Lifting/Left.lean b/Mathlib/CategoryTheory/Adjunction/Lifting/Left.lean index 432970c912ce8e..508de191a1af9d 100644 --- a/Mathlib/CategoryTheory/Adjunction/Lifting/Left.lean +++ b/Mathlib/CategoryTheory/Adjunction/Lifting/Left.lean @@ -122,13 +122,6 @@ variable [HasReflexiveCoequalizers A] noncomputable def constructLeftAdjointObj (Y : B) : A := coequalizer (F'.map (U.map (adj₁.counit.app Y))) (otherMap _ _ adj₁ adj₂ Y) -#adaptation_note -/-- -The new unused variable linter in -https://github.com/leanprover/lean4/pull/5338 -flags `{ z : F.obj (U.obj X) ⟶ R.obj Y // _ }`. --/ -set_option linter.unusedVariables false in /-- The homset equivalence which helps show that `R` is a right adjoint. -/ @[simps!] noncomputable def constructLeftAdjointEquiv [∀ X : B, RegularEpi (adj₁.counit.app X)] (Y : A) diff --git a/Mathlib/CategoryTheory/Category/Basic.lean b/Mathlib/CategoryTheory/Category/Basic.lean index e47e8bcebd744b..247f3da889691e 100644 --- a/Mathlib/CategoryTheory/Category/Basic.lean +++ b/Mathlib/CategoryTheory/Category/Basic.lean @@ -23,7 +23,7 @@ Introduces notations in the `CategoryTheory` scope Users may like to add `g ⊚ f` for composition in the standard convention, using ```lean -local notation g ` ⊚ `:80 f:80 := category.comp f g -- type as \oo +local notation:80 g " ⊚ " f:80 => CategoryTheory.CategoryStruct.comp f g -- type as \oo ``` ## Porting note diff --git a/Mathlib/CategoryTheory/ChosenFiniteProducts/InfSemilattice.lean b/Mathlib/CategoryTheory/ChosenFiniteProducts/InfSemilattice.lean new file mode 100644 index 00000000000000..ae74a6c833fadf --- /dev/null +++ b/Mathlib/CategoryTheory/ChosenFiniteProducts/InfSemilattice.lean @@ -0,0 +1,42 @@ +/- +Copyright (c) 2025 Sina Hazratpour. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sina Hazratpour +-/ + +import Mathlib.CategoryTheory.ChosenFiniteProducts +import Mathlib.CategoryTheory.Limits.Preorder + +/-! +# Chosen finite products for the preorder category of a meet-semilattice with a greatest element + +The preorder category of a meet-semilattice `C` with a greatest element has chosen finite products. + +A symmetric monoidal structure on the preorder category is automatically provided by the +instances `monoidalOfChosenFiniteProducts` and `ChosenFiniteProducts.instSymmetricCategory`. + +-/ + +namespace CategoryTheory + +open Category MonoidalCategory + +universe u + +variable (C : Type u) [SemilatticeInf C] [OrderTop C] + +namespace SemilatticeInf + +/-- Chosen finite products for the preorder category of a meet-semilattice with a greatest element-/ +noncomputable scoped instance chosenFiniteProducts : ChosenFiniteProducts C where + terminal := ⟨_, Preorder.isTerminalTop C⟩ + product X Y := ⟨_, Preorder.isLimitBinaryFan X Y⟩ + +lemma tensorObj {C : Type u} [SemilatticeInf C] [OrderTop C] {X Y : C} : X ⊗ Y = X ⊓ Y := rfl + +lemma tensorUnit {C : Type u} [SemilatticeInf C] [OrderTop C] : + 𝟙_ C = ⊤ := rfl + +end SemilatticeInf + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/CofilteredSystem.lean b/Mathlib/CategoryTheory/CofilteredSystem.lean index a3088857c8cf8e..103f619ddd6525 100644 --- a/Mathlib/CategoryTheory/CofilteredSystem.lean +++ b/Mathlib/CategoryTheory/CofilteredSystem.lean @@ -170,15 +170,11 @@ def toPreimages : J ⥤ Type v where rw [← mem_preimage, preimage_preimage, mem_preimage] convert h (g ≫ f); rw [F.map_comp]; rfl map_id j := by - #adaptation_note /-- nightly-2024-03-16: simp was - simp (config := { unfoldPartialApp := true }) only [MapsTo.restrict, Subtype.map, F.map_id] -/ - simp only [MapsTo.restrict, Subtype.map_def, F.map_id] + simp +unfoldPartialApp only [MapsTo.restrict, Subtype.map, F.map_id] ext rfl map_comp f g := by - #adaptation_note /-- nightly-2024-03-16: simp was - simp (config := { unfoldPartialApp := true }) only [MapsTo.restrict, Subtype.map, F.map_comp] -/ - simp only [MapsTo.restrict, Subtype.map_def, F.map_comp] + simp +unfoldPartialApp only [MapsTo.restrict, Subtype.map, F.map_comp] rfl instance toPreimages_finite [∀ j, Finite (F.obj j)] : ∀ j, Finite ((F.toPreimages s).obj j) := @@ -253,15 +249,11 @@ def toEventualRanges : J ⥤ Type v where obj j := F.eventualRange j map f := (F.eventualRange_mapsTo f).restrict _ _ _ map_id i := by - #adaptation_note /--- nightly-2024-03-16: simp was - simp (config := { unfoldPartialApp := true }) only [MapsTo.restrict, Subtype.map, F.map_id] -/ - simp only [MapsTo.restrict, Subtype.map_def, F.map_id] + simp +unfoldPartialApp only [MapsTo.restrict, Subtype.map, F.map_id] ext rfl map_comp _ _ := by - #adaptation_note /-- nightly-2024-03-16: simp was - simp (config := { unfoldPartialApp := true }) only [MapsTo.restrict, Subtype.map, F.map_comp] -/ - simp only [MapsTo.restrict, Subtype.map_def, F.map_comp] + simp +unfoldPartialApp only [MapsTo.restrict, Subtype.map, F.map_comp] rfl instance toEventualRanges_finite [∀ j, Finite (F.obj j)] : ∀ j, Finite (F.toEventualRanges.obj j) := diff --git a/Mathlib/CategoryTheory/Comma/Final.lean b/Mathlib/CategoryTheory/Comma/Final.lean index 6c77ce36c3afbc..23e5a766393055 100644 --- a/Mathlib/CategoryTheory/Comma/Final.lean +++ b/Mathlib/CategoryTheory/Comma/Final.lean @@ -5,7 +5,9 @@ Authors: Jakob von Raumer -/ import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction import Mathlib.CategoryTheory.Limits.IsConnected +import Mathlib.CategoryTheory.Limits.Sifted import Mathlib.CategoryTheory.Filtered.Final +import Mathlib.CategoryTheory.Filtered.Flat import Mathlib.CategoryTheory.Grothendieck import Mathlib.CategoryTheory.Comma.StructuredArrow.CommaMap @@ -19,9 +21,12 @@ final and `A` is connected. We then use this in a proof that derives finality of `map` between two comma categories on a quasi-commutative diagram of functors, some of which need to be final. +Finally we prove filteredness of a `Comma L R` and finality of `snd L R`, given that `R` is final +and `A` and `B` are filtered. + ## References -* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Lemma 3.4.3 & 3.4.4 +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Lemma 3.4.3 -- 3.4.5 -/ universe v₁ v₂ v₃ v₄ v₅ v₆ u₁ u₂ u₃ u₄ u₅ u₆ @@ -80,9 +85,9 @@ instance final_fst [R.Final] : (fst L R).Final := by exact (Functor.associator _ _ _).symm.trans (Iso.compInverseIso (mapFst _ _)) instance initial_snd [L.Initial] : (snd L R).Initial := by - haveI : ((opFunctor L R).leftOp ⋙ fst R.op L.op).Final := + have : ((opFunctor L R).leftOp ⋙ fst R.op L.op).Final := final_equivalence_comp (opEquiv L R).functor.leftOp (fst R.op L.op) - haveI : (snd L R).op.Final := final_of_natIso (opFunctorCompFst _ _) + have : (snd L R).op.Final := final_of_natIso (opFunctorCompFst _ _) apply initial_of_final_op /-- `Comma L R` with `L : A ⥤ T` and `R : B ⥤ T` is connected if `R` is final and `A` is @@ -109,9 +114,9 @@ end NonSmall Let `F`, `G`, `R` and `R'` be final and `B` be filtered. Then, the induced functor between the comma categories of the first and second row of the diagram is final. -/ -lemma map_final {A : Type u₁} [Category.{v₁} A] {B : Type u₂} [Category.{u₂} B] {T : Type u₃} +lemma map_final {A : Type u₁} [Category.{v₁} A] {B : Type u₂} [Category.{v₂} B] {T : Type u₃} [Category.{v₃} T] {L : A ⥤ T} {R : B ⥤ T} {A' : Type u₄} [Category.{v₄} A'] {B' : Type u₅} - [Category.{v₅} B'] {T' : Type u₂} [Category.{u₂} T'] {L' : A' ⥤ T'} {R' : B' ⥤ T'} {F : A ⥤ A'} + [Category.{v₅} B'] {T' : Type u₆} [Category.{v₆} T'] {L' : A' ⥤ T'} {R' : B' ⥤ T'} {F : A ⥤ A'} {G : B ⥤ B'} {H : T ⥤ T'} (iL : F ⋙ L' ≅ L ⋙ H) (iR : G ⋙ R' ≅ R ⋙ H) [IsFiltered B] [R.Final] [R'.Final] [F.Final] [G.Final] : (Comma.map iL.hom iR.inv).Final := ⟨fun ⟨i₂, j₂, u₂⟩ => by @@ -132,6 +137,57 @@ lemma map_final {A : Type u₁} [Category.{v₁} A] {B : Type u₂} [Category.{u rw [IsIso.Iso.inv_inv] infer_instance⟩ +section Filtered + +variable {A : Type u₁} [Category.{v₁} A] +variable {B : Type u₂} [Category.{v₂} B] +variable {T : Type u₃} [Category.{v₃} T] +variable (L : A ⥤ T) (R : B ⥤ T) + +attribute [local instance] map_final in +/-- Let `A` and `B` be filtered categories, `R : B ⥤ T` be final and `L : A ⥤ T`. Then, the +comma category `Comma L R` is filtered. -/ +instance isFiltered_of_final [IsFiltered A] [IsFiltered B] [R.Final] : IsFiltered (Comma L R) := by + haveI (a : A) : IsFiltered (Comma (fromPUnit (L.obj a)) R) := + R.final_iff_isFiltered_structuredArrow.mp inferInstance (L.obj a) + have (a : A) : (fromPUnit (Over.mk (𝟙 a))).Final := final_const_of_isTerminal Over.mkIdTerminal + let η (a : A) : fromPUnit (Over.mk (𝟙 a)) ⋙ Over.forget a ⋙ L ≅ fromPUnit (L.obj a) := + NatIso.ofComponents (fun _ => Iso.refl _) + have (a : A) := IsFiltered.of_final (map (L := fromPUnit (L.obj a)) (F := 𝟭 T) (η a).hom + ((Iso.refl (𝟭 B ⋙ R)).inv)) + have : RepresentablyCoflat (fst L R) := + ⟨fun a => IsFiltered.of_equivalence (CostructuredArrow.ofCommaFstEquivalence L R a).symm⟩ + apply isFiltered_of_representablyCoflat (fst L R) + +attribute [local instance] isFiltered_of_final in +/-- Let `A` and `B` be cofiltered categories, `L : A ⥤ T` be initial and `R : B ⥤ T`. Then, the +comma category `Comma L R` is cofiltered. -/ +lemma isCofiltered_of_initial [IsCofiltered A] [IsCofiltered B] [L.Initial] : + IsCofiltered (Comma L R) := + IsCofiltered.of_equivalence (Comma.opEquiv _ _).symm + +attribute [local instance] final_of_isFiltered_of_pUnit in +/-- Let `A` and `B` be filtered categories, `R : B ⥤ T` be final and `R : A ⥤ T`. Then, the +projection `snd L R : Comma L R ⥤ B` is final. -/ +instance final_snd [IsFiltered A] [IsFiltered B] [R.Final] : (snd L R).Final := by + let iL : star.{1} A ⋙ 𝟭 _ ≅ L ⋙ star _ := Iso.refl _ + let iR : 𝟭 B ⋙ star.{1} B ≅ R ⋙ star _ := Iso.refl _ + have := map_final iL iR + let s := (equivProd (𝟭 _) (star B)).trans <| prod.leftUnitorEquivalence B + let iS : map iL.hom iR.inv ⋙ s.functor ≅ snd L R := + NatIso.ofComponents (fun _ => Iso.refl _) (fun f => by simp [iL, iR, s]) + apply final_of_natIso iS + +/-- Let `A` and `B` be cofiltered categories, `L : A ⥤ T` be initial and `R : B ⥤ T`. Then, the +projection `fst L R : Comma L R ⥤ A` is initial. -/ +instance initial_fst [IsCofiltered A] [IsCofiltered B] [L.Initial] : (fst L R).Initial := by + have : ((opFunctor L R).leftOp ⋙ snd R.op L.op).Final := + final_equivalence_comp (opEquiv L R).functor.leftOp _ + have : (fst L R).op.Final := final_of_natIso <| opFunctorCompSnd _ _ + apply initial_of_final_op + +end Filtered + end Comma end CategoryTheory diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow/CommaMap.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/CommaMap.lean index 5a2d0d2c797989..9a0d5b34cc35b6 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow/CommaMap.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/CommaMap.lean @@ -23,7 +23,7 @@ noncomputable section variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] {T : Type u₃} [Category.{v₃} T] {L : C ⥤ T} {R : D ⥤ T} {C' : Type u₄} [Category.{v₄} C'] {D' : Type u₅} [Category.{v₅} D'] {T' : Type u₆} - [Category.{u₆} T'] {L' : C' ⥤ T'} {R' : D' ⥤ T'} {F₁ : C ⥤ C'} {F₂ : D ⥤ D'} {F : T ⥤ T'} + [Category.{v₆} T'] {L' : C' ⥤ T'} {R' : D' ⥤ T'} {F₁ : C ⥤ C'} {F₂ : D ⥤ D'} {F : T ⥤ T'} (α : F₁ ⋙ L' ⟶ L ⋙ F) (β : R ⋙ F ⟶ F₂ ⋙ R') /-- The functor establishing the equivalence `StructuredArrow.commaMapEquivalence`. -/ diff --git a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean index fe740cc36ff77b..aaf7e3bd093855 100644 --- a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean +++ b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean @@ -17,7 +17,6 @@ coalgebras over `G`. ## TODO -* Prove the dual result about the structure map of the terminal coalgebra of an endofunctor. * Prove that if the countable infinite product over the powers of the endofunctor exists, then algebras over the endofunctor coincide with algebras over the free monad on the endofunctor. -/ @@ -382,6 +381,35 @@ def equivOfNatIso {F G : C ⥤ C} (α : F ≅ G) : Coalgebra F ≌ Coalgebra G w counitIso := (functorOfNatTransComp _ _).symm ≪≫ functorOfNatTransEq (by simp) ≪≫ functorOfNatTransId +namespace Terminal + +variable {A} (h : @Limits.IsTerminal (Coalgebra F) _ A) + +/-- The inverse of the structure map of an terminal coalgebra -/ +@[simp] +def strInv : F.obj A.1 ⟶ A.1 := + (h.from ⟨F.obj A.V, F.map A.str⟩).f + +theorem right_inv' : + ⟨A.str ≫ strInv h, by rw [Category.assoc, F.map_comp, strInv, ← Hom.h] ⟩ = 𝟙 A := + Limits.IsTerminal.hom_ext h _ (𝟙 A) + +theorem right_inv : A.str ≫ strInv h = 𝟙 _ := + congr_arg Hom.f (right_inv' h) + +theorem left_inv : strInv h ≫ A.str = 𝟙 _ := by + rw [strInv, ← (h.from ⟨F.obj A.V, F.map A.str⟩).h, ← F.map_id, ← F.map_comp] + congr + exact right_inv h + +/-- The structure map of the terminal coalgebra is an isomorphism, +hence endofunctors preserve their terminal coalgebras +-/ +theorem str_isIso (h : Limits.IsTerminal A) : IsIso A.str := + { out := ⟨strInv h, right_inv _, left_inv _⟩ } + +end Terminal + end Coalgebra namespace Adjunction diff --git a/Mathlib/CategoryTheory/Galois/Topology.lean b/Mathlib/CategoryTheory/Galois/Topology.lean index 17d6329fb58cdc..7020a117b13ee3 100644 --- a/Mathlib/CategoryTheory/Galois/Topology.lean +++ b/Mathlib/CategoryTheory/Galois/Topology.lean @@ -17,7 +17,7 @@ embedding of `Aut F` into `∀ X, Aut (F.obj X)` where ## References -- Stacks Project: Tag 0BMQ +- [Stacks 0BMQ](https://stacks.math.columbia.edu/tag/0BMQ) -/ diff --git a/Mathlib/CategoryTheory/Generator/Basic.lean b/Mathlib/CategoryTheory/Generator/Basic.lean index 049514c52cef97..70168c6d7626de 100644 --- a/Mathlib/CategoryTheory/Generator/Basic.lean +++ b/Mathlib/CategoryTheory/Generator/Basic.lean @@ -527,15 +527,21 @@ theorem isSeparator_coprod_of_isSeparator_right (G H : C) [HasBinaryCoproduct G (hH : IsSeparator H) : IsSeparator (G ⨿ H) := (isSeparator_coprod _ _).2 <| IsSeparating.mono hH <| by simp +lemma isSeparator_of_isColimit_cofan {β : Type w} {f : β → C} + (hf : IsSeparating (Set.range f)) {c : Cofan f} (hc : IsColimit c) : IsSeparator c.pt := by + refine (isSeparator_def _).2 fun X Y u v huv => hf _ _ fun Z hZ g => ?_ + obtain ⟨b, rfl⟩ := Set.mem_range.1 hZ + classical simpa using c.ι.app ⟨b⟩ ≫= huv (hc.desc (Cofan.mk _ (Pi.single b g))) + theorem isSeparator_sigma {β : Type w} (f : β → C) [HasCoproduct f] : IsSeparator (∐ f) ↔ IsSeparating (Set.range f) := by - refine - ⟨fun h X Y u v huv => ?_, fun h => - (isSeparator_def _).2 fun X Y u v huv => h _ _ fun Z hZ g => ?_⟩ - · refine h.def _ _ fun g => colimit.hom_ext fun b => ?_ - simpa using huv (f b.as) (by simp) (colimit.ι (Discrete.functor f) _ ≫ g) - · obtain ⟨b, rfl⟩ := Set.mem_range.1 hZ - classical simpa using Sigma.ι f b ≫= huv (Sigma.desc (Pi.single b g)) + refine ⟨fun h X Y u v huv => ?_, fun h => isSeparator_of_isColimit_cofan h (colimit.isColimit _)⟩ + refine h.def _ _ fun g => colimit.hom_ext fun b => ?_ + simpa using huv (f b.as) (by simp) (colimit.ι (Discrete.functor f) _ ≫ g) + +theorem IsSeparating.isSeparator_coproduct {β : Type w} {f : β → C} [HasCoproduct f] + (hS : IsSeparating (Set.range f)) : IsSeparator (∐ f) := + (isSeparator_sigma _).2 hS theorem isSeparator_sigma_of_isSeparator {β : Type w} (f : β → C) [HasCoproduct f] (b : β) (hb : IsSeparator (f b)) : IsSeparator (∐ f) := diff --git a/Mathlib/CategoryTheory/Generator/Coproduct.lean b/Mathlib/CategoryTheory/Generator/Coproduct.lean deleted file mode 100644 index 0e948301d48193..00000000000000 --- a/Mathlib/CategoryTheory/Generator/Coproduct.lean +++ /dev/null @@ -1,43 +0,0 @@ -/- -Copyright (c) 2024 Joël Riou. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Joël Riou --/ - -import Mathlib.CategoryTheory.Generator.Basic -import Mathlib.CategoryTheory.Limits.Shapes.Products - -/-! -# The coproduct of a separating family of objects is separating - -If a family of objects `S : ι → C` in a category with zero morphisms -is separating, then the coproduct of `S` is a separator in `C`. - --/ - -universe w v u - -namespace CategoryTheory.IsSeparating - -open Limits - -variable {C : Type u} [Category.{v} C] [HasZeroMorphisms C] - {ι : Type w} {S : ι → C} - -open Classical in -lemma isSeparator_of_isColimit_cofan - (hS : IsSeparating (Set.range S)) {c : Cofan S} (hc : IsColimit c) : - IsSeparator c.pt := by - intro X Y f g h - apply hS - rintro _ ⟨i, rfl⟩ α - let β : c.pt ⟶ X := Cofan.IsColimit.desc hc - (fun j ↦ if hij : i = j then eqToHom (by rw [hij]) ≫ α else 0) - have hβ : c.inj i ≫ β = α := by simp [β] - simp only [← hβ, Category.assoc, h c.pt (by simp) β] - -lemma isSeparator_coproduct (hS : IsSeparating (Set.range S)) [HasCoproduct S] : - IsSeparator (∐ S) := - isSeparator_of_isColimit_cofan hS (colimit.isColimit _) - -end CategoryTheory.IsSeparating diff --git a/Mathlib/CategoryTheory/Generator/Indization.lean b/Mathlib/CategoryTheory/Generator/Indization.lean new file mode 100644 index 00000000000000..4c79e7d781802e --- /dev/null +++ b/Mathlib/CategoryTheory/Generator/Indization.lean @@ -0,0 +1,37 @@ +/- +Copyright (c) 2025 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Generator.Basic +import Mathlib.CategoryTheory.Limits.Indization.Category + +/-! +# Separating set in the category of ind-objects + +We construct a separating set in the category of ind-objects. + +## Future work + +Once we have constructed zero morphisms in the category of ind-objects, we will be able to show +that under sufficient conditions, the category of ind-objects has a separating object. + +-/ + +universe v u + +namespace CategoryTheory + +open Limits + +section + +variable {C : Type u} [Category.{v} C] + +theorem Ind.isSeparating_range_yoneda : IsSeparating (Set.range (Ind.yoneda : C ⥤ _).obj) := by + refine fun X Y f g h => (cancel_epi (Ind.colimitPresentationCompYoneda X).hom).1 ?_ + exact colimit.hom_ext (fun i => by simp [← Category.assoc, h]) + +end + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Generator/Presheaf.lean b/Mathlib/CategoryTheory/Generator/Presheaf.lean index d12ab50d4882a9..61042589d44a2e 100644 --- a/Mathlib/CategoryTheory/Generator/Presheaf.lean +++ b/Mathlib/CategoryTheory/Generator/Presheaf.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.CategoryTheory.Generator.Coproduct +import Mathlib.CategoryTheory.Generator.Basic import Mathlib.CategoryTheory.Limits.FunctorCategory.Basic /-! diff --git a/Mathlib/CategoryTheory/GradedObject.lean b/Mathlib/CategoryTheory/GradedObject.lean index 60536139bb7171..0b144c74283155 100644 --- a/Mathlib/CategoryTheory/GradedObject.lean +++ b/Mathlib/CategoryTheory/GradedObject.lean @@ -216,11 +216,8 @@ theorem shiftFunctor_map_apply {β : Type*} [AddCommGroup β] (s : β) instance [HasZeroMorphisms C] (β : Type w) (X Y : GradedObject β C) : Zero (X ⟶ Y) := ⟨fun _ => 0⟩ -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/4481 -the `simpNF` linter incorrectly claims this lemma can't be applied by `simp`. --/ +#adaptation_note /-- https://github.com/leanprover/lean4/pull/4481 +the `simpNF` linter incorrectly claims this lemma can't be applied by `simp`. -/ @[simp, nolint simpNF] theorem zero_apply [HasZeroMorphisms C] (β : Type w) (X Y : GradedObject β C) (b : β) : (0 : X ⟶ Y) b = 0 := diff --git a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean index b10fab11e47d85..3467d00d9731fd 100644 --- a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean +++ b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean @@ -312,9 +312,8 @@ lemma associator_naturality (f₁ : X₁ ⟶ Y₁) (f₂ : X₂ ⟶ Y₂) (f₃ [HasGoodTensor₁₂Tensor Y₁ Y₂ Y₃] [HasGoodTensorTensor₂₃ Y₁ Y₂ Y₃] : tensorHom (tensorHom f₁ f₂) f₃ ≫ (associator Y₁ Y₂ Y₃).hom = (associator X₁ X₂ X₃).hom ≫ tensorHom f₁ (tensorHom f₂ f₃) := by - #adaptation_note - /-- this used to be aesop_cat, but that broke with - https://github.com/leanprover/lean4/pull/4154 -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/4154 + this used to be aesop_cat -/ ext x i₁ i₂ i₃ h : 2 simp only [categoryOfGradedObjects_comp, ιTensorObj₃'_tensorHom_assoc, associator_conjugation, ιTensorObj₃'_associator_hom, assoc, Iso.inv_hom_id_assoc, diff --git a/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean index d5e68a50703d7c..1b1231d4bd3c37 100644 --- a/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean +++ b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean @@ -74,7 +74,7 @@ lemma colimit_no_zero_smul_divisor rw [← this, ← map_smul (colimit.ι F j).hom] at hx obtain ⟨j', i, h⟩ := Concrete.colimit_rep_eq_zero (hx := hx) obtain ⟨j'', H⟩ := H - simpa [elementwise_of% (colimit.w F), this, map_zero] using congr(colimit.ι F _ + simpa [elementwise_of% (colimit.w F), ← this, map_zero] using congr(colimit.ι F _ $(H (IsFiltered.sup {j, j', j''} { ⟨j, j', by simp, by simp, i⟩ }) (IsFiltered.toSup _ _ <| by simp) (F.map (IsFiltered.toSup _ _ <| by simp) x) diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index 57bf33aa6b3ab0..c5a23c859f29df 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -169,8 +169,6 @@ theorem Cocone.w {F : J ⥤ C} (c : Cocone F) {j j' : J} (f : j ⟶ j') : rw [c.ι.naturality f] apply comp_id -attribute [simp 1001] Cocone.w_assoc - end variable {F : J ⥤ C} diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 93331eeffeb40a..3431f905c7954d 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -335,11 +335,8 @@ section variable (G) /-- When `F : C ⥤ D` is final, and `G : D ⥤ E` has a colimit, then `F ⋙ G` has a colimit also and -`colimit (F ⋙ G) ≅ colimit G` - -https://stacks.math.columbia.edu/tag/04E7 --/ -@[simps! (config := .lemmasOnly)] +`colimit (F ⋙ G) ≅ colimit G`. -/ +@[simps! (config := .lemmasOnly), stacks 04E7] def colimitIso [HasColimit G] : colimit (F ⋙ G) ≅ colimit G := asIso (colimit.pre G F) @@ -684,11 +681,8 @@ section variable (G) /-- When `F : C ⥤ D` is initial, and `G : D ⥤ E` has a limit, then `F ⋙ G` has a limit also and -`limit (F ⋙ G) ≅ limit G` - -https://stacks.math.columbia.edu/tag/04E7 --/ -@[simps! (config := .lemmasOnly)] +`limit (F ⋙ G) ≅ limit G`. -/ +@[simps! (config := .lemmasOnly), stacks 04E7] def limitIso [HasLimit G] : limit (F ⋙ G) ≅ limit G := (asIso (limit.pre G F)).symm diff --git a/Mathlib/CategoryTheory/Limits/Indization/Category.lean b/Mathlib/CategoryTheory/Limits/Indization/Category.lean index 7b2b555d9cb66e..c1cf76d2a04fe0 100644 --- a/Mathlib/CategoryTheory/Limits/Indization/Category.lean +++ b/Mathlib/CategoryTheory/Limits/Indization/Category.lean @@ -7,6 +7,7 @@ import Mathlib.CategoryTheory.Functor.Flat import Mathlib.CategoryTheory.Limits.Constructions.Filtered import Mathlib.CategoryTheory.Limits.FullSubcategory import Mathlib.CategoryTheory.Limits.Indization.LocallySmall +import Mathlib.CategoryTheory.Limits.Indization.ParallelPair import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits import Mathlib.CategoryTheory.Limits.Indization.Products @@ -30,7 +31,8 @@ Adopting the theorem numbering of [Kashiwara2006], we show that * if `C` has finite coproducts, then `Ind C` has small coproducts (6.1.18(ii)), * if `C` has products indexed by `α`, then `Ind C` has products indexed by `α`, and the functor `Ind C ⥤ Cᵒᵖ ⥤ Type v` creates such products (6.1.17) -* the functor `C ⥤ Ind C` preserves small limits. +* the functor `C ⥤ Ind C` preserves small limits, +* if `C` has coequalizers, then `Ind C` has coequalizers (6.1.18(i)). More limit-colimit properties will follow. @@ -43,7 +45,7 @@ Note that: * [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Chapter 6 -/ -universe v u +universe w v u namespace CategoryTheory @@ -180,14 +182,14 @@ noncomputable def Ind.limCompInclusion {I : Type v} [SmallCategory I] [IsFiltere _ ≅ (whiskeringRight _ _ _).obj yoneda ⋙ colim := isoWhiskerRight ((whiskeringRight _ _ _).mapIso (Ind.yonedaCompInclusion)) colim -instance {α : Type v} [SmallCategory α] [FinCategory α] [HasLimitsOfShape α C] {I : Type v} +instance {α : Type w} [SmallCategory α] [FinCategory α] [HasLimitsOfShape α C] {I : Type v} [SmallCategory I] [IsFiltered I] : PreservesLimitsOfShape α (Ind.lim I : (I ⥤ C) ⥤ _) := haveI : PreservesLimitsOfShape α (Ind.lim I ⋙ Ind.inclusion C) := preservesLimitsOfShape_of_natIso Ind.limCompInclusion.symm preservesLimitsOfShape_of_reflects_of_preserves _ (Ind.inclusion C) -instance {α : Type v} [SmallCategory α] [FinCategory α] [HasColimitsOfShape α C] {I : Type v} +instance {α : Type w} [SmallCategory α] [FinCategory α] [HasColimitsOfShape α C] {I : Type v} [SmallCategory I] [IsFiltered I] : PreservesColimitsOfShape α (Ind.lim I : (I ⥤ C) ⥤ _) := inferInstanceAs (PreservesColimitsOfShape α (_ ⋙ colim)) @@ -209,11 +211,43 @@ instance {α : Type v} [Finite α] [HasColimitsOfShape (Discrete α) C] : -- ``` -- from the fact that finite limits commute with filtered colimits and from the fact that -- `Ind.yoneda` preserves finite colimits. - apply hasColimitOfIso iso.symm + exact hasColimitOfIso iso.symm instance [HasFiniteCoproducts C] : HasCoproducts.{v} (Ind C) := have : HasFiniteCoproducts (Ind C) := ⟨fun _ => hasColimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift)⟩ hasCoproducts_of_finite_and_filtered +/-- Given an `IndParallelPairPresentation f g`, we can understand the parallel pair `(f, g)` as +the colimit of `(P.φ, P.ψ)` in `Ind C`. -/ +noncomputable def IndParallelPairPresentation.parallelPairIsoParallelPairCompIndYoneda + {A B : Ind C} {f g : A ⟶ B} + (P : IndParallelPairPresentation ((Ind.inclusion _).map f) ((Ind.inclusion _).map g)) : + parallelPair f g ≅ parallelPair P.φ P.ψ ⋙ Ind.lim P.I := + ((whiskeringRight WalkingParallelPair _ _).obj (Ind.inclusion C)).preimageIso <| + diagramIsoParallelPair _ ≪≫ + P.parallelPairIsoParallelPairCompYoneda ≪≫ + isoWhiskerLeft (parallelPair _ _) Ind.limCompInclusion.symm + +instance [HasColimitsOfShape WalkingParallelPair C] : + HasColimitsOfShape WalkingParallelPair (Ind C) := by + refine ⟨fun F => ?_⟩ + obtain ⟨P⟩ := nonempty_indParallelPairPresentation (F.obj WalkingParallelPair.zero).2 + (F.obj WalkingParallelPair.one).2 (Ind.inclusion _ |>.map <| F.map WalkingParallelPairHom.left) + (Ind.inclusion _ |>.map <| F.map WalkingParallelPairHom.right) + exact hasColimitOfIso (diagramIsoParallelPair _ ≪≫ P.parallelPairIsoParallelPairCompIndYoneda) + +/-- A way to understand morphisms in `Ind C`: every morphism is induced by a natural transformation +of diagrams. -/ +theorem Ind.exists_nonempty_arrow_mk_iso_ind_lim {A B : Ind C} {f : A ⟶ B} : + ∃ (I : Type v) (_ : SmallCategory I) (_ : IsFiltered I) (F G : I ⥤ C) (φ : F ⟶ G), + Nonempty (Arrow.mk f ≅ Arrow.mk ((Ind.lim _).map φ)) := by + obtain ⟨P⟩ := nonempty_indParallelPairPresentation A.2 B.2 + (Ind.inclusion _ |>.map f) (Ind.inclusion _ |>.map f) + refine ⟨P.I, inferInstance, inferInstance, P.F₁, P.F₂, P.φ, ⟨Arrow.isoMk ?_ ?_ ?_⟩⟩ + · exact P.parallelPairIsoParallelPairCompIndYoneda.app WalkingParallelPair.zero + · exact P.parallelPairIsoParallelPairCompIndYoneda.app WalkingParallelPair.one + · simpa using + (P.parallelPairIsoParallelPairCompIndYoneda.hom.naturality WalkingParallelPairHom.left).symm + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean index ef9fd1afa688dc..880fad04850627 100644 --- a/Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean @@ -4,16 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits +import Mathlib.CategoryTheory.Limits.FullSubcategory +import Mathlib.CategoryTheory.Limits.Indization.ParallelPair /-! # Equalizers of ind-objects -The eventual goal of this file is to show that if a category `C` has equalizers, then ind-objects -are closed under equalizers. For now, it contains one of the main prerequisites, namely we show -that under sufficient conditions the limit of a diagram in `I ⥤ C` taken in `Cᵒᵖ ⥤ Type v` is an -ind-object. +We show that if a category `C` has equalizers, then ind-objects are closed under equalizers. -## References +# References * [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Section 6.1 -/ @@ -60,4 +59,21 @@ theorem isIndObject_limit_comp_yoneda_comp_colim end +/-- If `C` has equalizers. then ind-objects are closed under equalizers. + +This is Proposition 6.1.17(i) in [Kashiwara2006]. +-/ +theorem closedUnderLimitsOfShape_walkingParallelPair_isIndObject [HasEqualizers C] : + ClosedUnderLimitsOfShape WalkingParallelPair (IsIndObject (C := C)) := by + apply closedUnderLimitsOfShape_of_limit + intro F hF h + obtain ⟨P⟩ := nonempty_indParallelPairPresentation (h WalkingParallelPair.zero) + (h WalkingParallelPair.one) (F.map WalkingParallelPairHom.left) + (F.map WalkingParallelPairHom.right) + exact IsIndObject.map + (HasLimit.isoOfNatIso (P.parallelPairIsoParallelPairCompYoneda.symm ≪≫ + (diagramIsoParallelPair _).symm)).hom + (isIndObject_limit_comp_yoneda_comp_colim (parallelPair P.φ P.ψ) + (fun i => isIndObject_limit_comp_yoneda _)) + end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Indization/ParallelPair.lean b/Mathlib/CategoryTheory/Limits/Indization/ParallelPair.lean new file mode 100644 index 00000000000000..68c898bc326347 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Indization/ParallelPair.lean @@ -0,0 +1,171 @@ +/- +Copyright (c) 2025 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Comma.Final +import Mathlib.CategoryTheory.Limits.Indization.IndObject + +/-! +# Parallel pairs of natural transformations between ind-objects + +We show that if `A` and `B` are ind-objects and `f` and `g` are natural transformations between +`A` and `B`, then there is a small filtered category `I` such that `A`, `B`, `f` and `g` are +commonly presented by diagrams and natural transformations in `I ⥤ C`. + + +## References +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Proposition 6.1.15 (though + our proof is more direct). +-/ + +universe v₁ v₂ v₃ u₁ u₂ u₃ + +namespace CategoryTheory + +open Limits + +variable {C : Type u₁} [Category.{v₁} C] + +/-- Structure containing data exhibiting two parallel natural transformations `f` and `g` between +presheaves `A` and `B` as induced by a natural transformation in a functor category exhibiting +`A` and `B` as ind-objects. -/ +structure IndParallelPairPresentation {A B : Cᵒᵖ ⥤ Type v₁} (f g : A ⟶ B) where + /-- The indexing category. -/ + I : Type v₁ + /-- Category instance on the indexing category. -/ + [ℐ : SmallCategory I] + [hI : IsFiltered I] + /-- The diagram presenting `A`. -/ + F₁ : I ⥤ C + /-- The diagram presenting `B`. -/ + F₂ : I ⥤ C + /-- The cocone on `F₁` with apex `A`. -/ + ι₁ : F₁ ⋙ yoneda ⟶ (Functor.const I).obj A + /-- The cocone on `F₁` with apex `A` is a colimit cocone. -/ + isColimit₁ : IsColimit (Cocone.mk A ι₁) + /-- The cocone on `F₂` with apex `B`. -/ + ι₂ : F₂ ⋙ yoneda ⟶ (Functor.const I).obj B + /-- The cocone on `F₂` with apex `B` is a colimit cocone. -/ + isColimit₂ : IsColimit (Cocone.mk B ι₂) + /-- The natural transformation presenting `f`. -/ + φ : F₁ ⟶ F₂ + /-- The natural transformation presenting `g`. -/ + ψ : F₁ ⟶ F₂ + /-- `f` is in fact presented by `φ`. -/ + hf : f = IsColimit.map isColimit₁ (Cocone.mk B ι₂) (whiskerRight φ yoneda) + /-- `g` is in fact presented by `ψ`. -/ + hg : g = IsColimit.map isColimit₁ (Cocone.mk B ι₂) (whiskerRight ψ yoneda) + +instance {A B : Cᵒᵖ ⥤ Type v₁} {f g : A ⟶ B} (P : IndParallelPairPresentation f g) : + SmallCategory P.I := P.ℐ +instance {A B : Cᵒᵖ ⥤ Type v₁} {f g : A ⟶ B} (P : IndParallelPairPresentation f g) : + IsFiltered P.I := P.hI + +namespace NonemptyParallelPairPresentationAux + +variable {A B : Cᵒᵖ ⥤ Type v₁} (f g : A ⟶ B) (P₁ : IndObjectPresentation A) + (P₂ : IndObjectPresentation B) + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +abbrev K : Type v₁ := + Comma ((P₁.toCostructuredArrow ⋙ CostructuredArrow.map f).prod' + (P₁.toCostructuredArrow ⋙ CostructuredArrow.map g)) + (P₂.toCostructuredArrow.prod' P₂.toCostructuredArrow) + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +abbrev F₁ : K f g P₁ P₂ ⥤ C := Comma.fst _ _ ⋙ P₁.F +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +abbrev F₂ : K f g P₁ P₂ ⥤ C := Comma.snd _ _ ⋙ P₂.F + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +abbrev ι₁ : F₁ f g P₁ P₂ ⋙ yoneda ⟶ (Functor.const (K f g P₁ P₂)).obj A := + whiskerLeft (Comma.fst _ _) P₁.ι + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +noncomputable abbrev isColimit₁ : IsColimit (Cocone.mk A (ι₁ f g P₁ P₂)) := + (Functor.Final.isColimitWhiskerEquiv _ _).symm P₁.isColimit + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +abbrev ι₂ : F₂ f g P₁ P₂ ⋙ yoneda ⟶ (Functor.const (K f g P₁ P₂)).obj B := + whiskerLeft (Comma.snd _ _) P₂.ι + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +noncomputable abbrev isColimit₂ : IsColimit (Cocone.mk B (ι₂ f g P₁ P₂)) := + (Functor.Final.isColimitWhiskerEquiv _ _).symm P₂.isColimit + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +def ϕ : F₁ f g P₁ P₂ ⟶ F₂ f g P₁ P₂ where + app h := h.hom.1.left + naturality _ _ h := by + have := h.w + simp only [Functor.prod'_obj, Functor.comp_obj, prod_Hom, Functor.prod'_map, Functor.comp_map, + prod_comp, Prod.mk.injEq, CostructuredArrow.hom_eq_iff, CostructuredArrow.map_obj_left, + IndObjectPresentation.toCostructuredArrow_obj_left, CostructuredArrow.comp_left, + CostructuredArrow.map_map_left, IndObjectPresentation.toCostructuredArrow_map_left] at this + exact this.1 + +theorem hf : f = IsColimit.map (isColimit₁ f g P₁ P₂) + (Cocone.mk B (ι₂ f g P₁ P₂)) (whiskerRight (ϕ f g P₁ P₂) yoneda) := by + refine (isColimit₁ f g P₁ P₂).hom_ext (fun i => ?_) + rw [IsColimit.ι_map] + simpa using i.hom.1.w.symm + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +def ψ : F₁ f g P₁ P₂ ⟶ F₂ f g P₁ P₂ where + app h := h.hom.2.left + naturality _ _ h := by + have := h.w + simp only [Functor.prod'_obj, Functor.comp_obj, prod_Hom, Functor.prod'_map, Functor.comp_map, + prod_comp, Prod.mk.injEq, CostructuredArrow.hom_eq_iff, CostructuredArrow.map_obj_left, + IndObjectPresentation.toCostructuredArrow_obj_left, CostructuredArrow.comp_left, + CostructuredArrow.map_map_left, IndObjectPresentation.toCostructuredArrow_map_left] at this + exact this.2 + +theorem hg : g = IsColimit.map (isColimit₁ f g P₁ P₂) + (Cocone.mk B (ι₂ f g P₁ P₂)) (whiskerRight (ψ f g P₁ P₂) yoneda) := by + refine (isColimit₁ f g P₁ P₂).hom_ext (fun i => ?_) + rw [IsColimit.ι_map] + simpa using i.hom.2.w.symm + +attribute [local instance] Comma.isFiltered_of_final in +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +noncomputable def presentation : IndParallelPairPresentation f g where + I := K f g P₁ P₂ + F₁ := F₁ f g P₁ P₂ + F₂ := F₂ f g P₁ P₂ + ι₁ := ι₁ f g P₁ P₂ + isColimit₁ := isColimit₁ f g P₁ P₂ + ι₂ := ι₂ f g P₁ P₂ + isColimit₂ := isColimit₂ f g P₁ P₂ + φ := ϕ f g P₁ P₂ + ψ := ψ f g P₁ P₂ + hf := hf f g P₁ P₂ + hg := hg f g P₁ P₂ + +end NonemptyParallelPairPresentationAux + +theorem nonempty_indParallelPairPresentation {A B : Cᵒᵖ ⥤ Type v₁} (hA : IsIndObject A) + (hB : IsIndObject B) (f g : A ⟶ B) : Nonempty (IndParallelPairPresentation f g) := + ⟨NonemptyParallelPairPresentationAux.presentation f g hA.presentation hB.presentation⟩ + +namespace IndParallelPairPresentation + +/-- Given an `IndParallelPairPresentation f g`, we can understand the parallel pair `(f, g)` +as the colimit of `(P.φ, P.ψ)` in `Cᵒᵖ ⥤ Type v`. -/ +noncomputable def parallelPairIsoParallelPairCompYoneda {A B : Cᵒᵖ ⥤ Type v₁} {f g : A ⟶ B} + (P : IndParallelPairPresentation f g) : + parallelPair f g ≅ parallelPair P.φ P.ψ ⋙ (whiskeringRight _ _ _).obj yoneda ⋙ colim := + parallelPair.ext + (P.isColimit₁.coconePointUniqueUpToIso (colimit.isColimit _)) + (P.isColimit₂.coconePointUniqueUpToIso (colimit.isColimit _)) + (P.isColimit₁.hom_ext (fun j => by + simp [P.hf, P.isColimit₁.ι_map_assoc, P.isColimit₁.comp_coconePointUniqueUpToIso_hom_assoc, + P.isColimit₂.comp_coconePointUniqueUpToIso_hom])) + (P.isColimit₁.hom_ext (fun j => by + simp [P.hg, P.isColimit₁.ι_map_assoc, P.isColimit₁.comp_coconePointUniqueUpToIso_hom_assoc, + P.isColimit₂.comp_coconePointUniqueUpToIso_hom])) + +end IndParallelPairPresentation + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Preorder.lean b/Mathlib/CategoryTheory/Limits/Preorder.lean new file mode 100644 index 00000000000000..605ebbea32e56b --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Preorder.lean @@ -0,0 +1,75 @@ +/- +Copyright (c) 2025 Sina Hazratpour. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sina Hazratpour, Joël Riou +-/ + +import Mathlib.CategoryTheory.Limits.Shapes.BinaryProducts + +/-! +# (Co)limits in a preorder category + +We provide basic results about the nullary and binary (co)products in the associated category of +a preordered type. + +-/ + +universe u + +open CategoryTheory Limits + +namespace Preorder + +variable (C : Type u) + +section OrderBot + +variable [Preorder C] [OrderBot C] + +/-- The least element in a preordered type is initial in the category +associated to this preorder. -/ +def isInitialBot : IsInitial (⊥ : C) := IsInitial.ofUnique _ + +instance : HasInitial C := hasInitial_of_unique ⊥ + +end OrderBot + +section OrderTop + +variable [Preorder C] [OrderTop C] + +/-- The greatest element of a preordered type is terminal in the category +associated to this preorder. -/ +def isTerminalTop : IsTerminal (⊤ : C) := IsTerminal.ofUnique _ + +instance : HasTerminal C := hasTerminal_of_unique ⊤ + +end OrderTop + +section SemilatticeInf + +variable {C} [SemilatticeInf C] + +/-- The infimum of two elements in a preordered type is a binary product in +the category associated to this preorder. -/ +def isLimitBinaryFan (X Y : C) : + IsLimit (BinaryFan.mk (P := X ⊓ Y) (homOfLE inf_le_left) (homOfLE inf_le_right)) := + BinaryFan.isLimitMk (fun s ↦ homOfLE (le_inf (leOfHom s.fst) (leOfHom s.snd))) + (by intros; rfl) (by intros; rfl) (by intros; rfl) + +end SemilatticeInf + +section SemilatticeSup + +variable {C} [SemilatticeSup C] + +/-- The supremum of two elements in a preordered type is a binary coproduct +in the category associated to this preorder. -/ +def isColimitBinaryCofan (X Y : C) : + IsColimit (BinaryCofan.mk (P := X ⊔ Y) (homOfLE le_sup_left) (homOfLE le_sup_right)) := + BinaryCofan.isColimitMk (fun s ↦ homOfLE (sup_le (leOfHom s.inl) (leOfHom s.inr))) + (by intros; rfl) (by intros; rfl) (by intros; rfl) + +end SemilatticeSup + +end Preorder diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean b/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean index 6f0db0a360ddbb..9064ae11e1cb93 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean @@ -219,9 +219,7 @@ instance sequentialFunctor_initial : (sequentialFunctor J).Initial where · right exact ⟨CostructuredArrow.homMk (homOfLE h).op rfl⟩ -/-- -This is proved in https://stacks.math.columbia.edu/tag/0032 --/ +@[stacks 0032] proof_wanted preorder_of_cofiltered (J : Type*) [Category J] [IsCofiltered J] : ∃ (I : Type*) (_ : Preorder I) (_ : IsCofiltered I) (F : I ⥤ J), F.Initial diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean index 6e11689c3fc8bb..6dc28a06e73c91 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean @@ -420,7 +420,8 @@ def kernelCompMono {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [HasKernel f] [Mono g simp) inv := kernel.lift _ (kernel.ι _) (by simp) -#adaptation_note /-- nightly-2024-04-01 The `symm` wasn't previously necessary. -/ +#adaptation_note /-- nightly-2024-04-01 +The `symm` wasn't previously necessary. -/ instance hasKernel_iso_comp {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] [HasKernel g] : HasKernel (f ≫ g) where exists_limit := @@ -874,7 +875,8 @@ theorem cokernel_not_mono_of_nonzero (w : f ≠ 0) : ¬Mono (cokernel.π f) := f theorem cokernel_not_iso_of_nonzero (w : f ≠ 0) : IsIso (cokernel.π f) → False := fun _ => cokernel_not_mono_of_nonzero w inferInstance -#adaptation_note /-- nightly-2024-04-01 The `symm` wasn't previously necessary. -/ +#adaptation_note /-- nightly-2024-04-01 +The `symm` wasn't previously necessary. -/ -- TODO the remainder of this section has obvious generalizations to `HasCoequalizer f g`. instance hasCokernel_comp_iso {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [HasCokernel f] [IsIso g] : HasCokernel (f ≫ g) where diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Preorder/Basic.lean b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/Basic.lean index fd2ca7692b49c4..cf21ab0f2fbac8 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Preorder/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.CategoryTheory.Limits.Shapes.Terminal +import Mathlib.CategoryTheory.Limits.Preorder /-! # Limits and colimits indexed by preorders @@ -27,12 +27,6 @@ section OrderBot variable [OrderBot J] -/-- The least element in a preordered type is initial in the category -associated to this preorder. -/ -def isInitialBot : IsInitial (⊥ : J) := IsInitial.ofUnique _ - -instance : HasInitial J := hasInitial_of_unique ⊥ - instance : HasLimitsOfShape J C := ⟨fun _ ↦ by infer_instance⟩ end OrderBot @@ -41,12 +35,6 @@ section OrderTop variable [OrderTop J] -/-- The greatest element of a preordered type is terminal in the category -associated to this preorder. -/ -def isTerminalBot : IsTerminal (⊤ : J) := IsTerminal.ofUnique _ - -instance : HasTerminal J := hasTerminal_of_unique ⊤ - instance : HasColimitsOfShape J C := ⟨fun _ ↦ by infer_instance⟩ end OrderTop diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean index 315be3a49b8d54..9574c460133413 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean @@ -160,11 +160,7 @@ instance hasPullback_of_right_factors_mono : HasPullback i (f ≫ i) := by instance pullback_snd_iso_of_right_factors_mono : IsIso (pullback.snd i (f ≫ i)) := by - #adaptation_note /-- nightly-testing 2024-04-01 - this could not be placed directly in the `show from` without `dsimp` -/ have := limit.isoLimitCone_hom_π ⟨_, pullbackIsPullbackOfCompMono (𝟙 _) f i⟩ WalkingCospan.right - dsimp only [cospan_right, id_eq, eq_mpr_eq_cast, PullbackCone.mk_pt, PullbackCone.mk_π_app, - Functor.const_obj_obj, cospan_one] at this convert (congrArg IsIso (show _ ≫ pullback.snd (𝟙 Z) f = _ from this)).mp inferInstance · exact (Category.id_comp _).symm · exact (Category.id_comp _).symm @@ -176,11 +172,7 @@ instance hasPullback_of_left_factors_mono : HasPullback (f ≫ i) i := by instance pullback_snd_iso_of_left_factors_mono : IsIso (pullback.fst (f ≫ i) i) := by - #adaptation_note /-- nightly-testing 2024-04-01 - this could not be placed directly in the `show from` without `dsimp` -/ have := limit.isoLimitCone_hom_π ⟨_, pullbackIsPullbackOfCompMono f (𝟙 _) i⟩ WalkingCospan.left - dsimp only [cospan_left, id_eq, eq_mpr_eq_cast, PullbackCone.mk_pt, PullbackCone.mk_π_app, - Functor.const_obj_obj, cospan_one] at this convert (congrArg IsIso (show _ ≫ pullback.fst f (𝟙 Z) = _ from this)).mp inferInstance · exact (Category.id_comp _).symm · exact (Category.id_comp _).symm diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean index 5d87af2751e5f3..99b3d3e5f71a0b 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean @@ -55,6 +55,8 @@ example [UnivLE.{v, u}] : HasProducts.{v} (Type u) := inferInstance instance : HasProducts.{v} (Type v) := inferInstance /-- A restatement of `Types.Limit.lift_π_apply` that uses `Pi.π` and `Pi.lift`. -/ +-- The increased `@[simp]` priority here results in a minor speed up in +-- `Mathlib.CategoryTheory.Sites.EqualizerSheafCondition`. @[simp 1001] theorem pi_lift_π_apply {β : Type v} [Small.{u} β] (f : β → Type u) {P : Type u} (s : ∀ b, P ⟶ f b) (b : β) (x : P) : @@ -69,7 +71,7 @@ theorem pi_lift_π_apply' {β : Type v} (f : β → Type v) {P : Type v} simp /-- A restatement of `Types.Limit.map_π_apply` that uses `Pi.π` and `Pi.map`. -/ -@[simp 1001] +@[simp] theorem pi_map_π_apply {β : Type v} [Small.{u} β] {f g : β → Type u} (α : ∀ j, f j ⟶ g j) (b : β) (x) : (Pi.π g b : ∏ᶜ g → g b) (Pi.map α x) = α b ((Pi.π f b : ∏ᶜ f → f b) x) := diff --git a/Mathlib/CategoryTheory/Limits/Shapes/WideEqualizers.lean b/Mathlib/CategoryTheory/Limits/Shapes/WideEqualizers.lean index 349e66c0e30f53..04dfbf1805cdec 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/WideEqualizers.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/WideEqualizers.lean @@ -517,11 +517,9 @@ abbrev wideEqualizer.ι : wideEqualizer f ⟶ X := abbrev wideEqualizer.trident : Trident f := limit.cone (parallelFamily f) -@[simp] theorem wideEqualizer.trident_ι : (wideEqualizer.trident f).ι = wideEqualizer.ι f := rfl -@[simp 1100] theorem wideEqualizer.trident_π_app_zero : (wideEqualizer.trident f).π.app zero = wideEqualizer.ι f := rfl @@ -543,11 +541,11 @@ abbrev wideEqualizer.lift [Nonempty J] {W : C} (k : W ⟶ X) (h : ∀ j₁ j₂, W ⟶ wideEqualizer f := limit.lift (parallelFamily f) (Trident.ofι k h) -@[reassoc (attr := simp 1100)] +@[reassoc] theorem wideEqualizer.lift_ι [Nonempty J] {W : C} (k : W ⟶ X) (h : ∀ j₁ j₂, k ≫ f j₁ = k ≫ f j₂) : - wideEqualizer.lift k h ≫ wideEqualizer.ι f = k := - limit.lift_π _ _ + wideEqualizer.lift k h ≫ wideEqualizer.ι f = k := by + simp /-- A morphism `k : W ⟶ X` satisfying `∀ j₁ j₂, k ≫ f j₁ = k ≫ f j₂` induces a morphism `l : W ⟶ wideEqualizer f` satisfying `l ≫ wideEqualizer.ι f = k`. -/ @@ -604,11 +602,9 @@ abbrev wideCoequalizer.π : Y ⟶ wideCoequalizer f := abbrev wideCoequalizer.cotrident : Cotrident f := colimit.cocone (parallelFamily f) -@[simp] theorem wideCoequalizer.cotrident_π : (wideCoequalizer.cotrident f).π = wideCoequalizer.π f := rfl -@[simp 1100] theorem wideCoequalizer.cotrident_ι_app_one : (wideCoequalizer.cotrident f).ι.app one = wideCoequalizer.π f := rfl @@ -631,11 +627,11 @@ abbrev wideCoequalizer.desc [Nonempty J] {W : C} (k : Y ⟶ W) (h : ∀ j₁ j wideCoequalizer f ⟶ W := colimit.desc (parallelFamily f) (Cotrident.ofπ k h) -@[reassoc (attr := simp 1100)] +@[reassoc] theorem wideCoequalizer.π_desc [Nonempty J] {W : C} (k : Y ⟶ W) (h : ∀ j₁ j₂, f j₁ ≫ k = f j₂ ≫ k) : - wideCoequalizer.π f ≫ wideCoequalizer.desc k h = k := - colimit.ι_desc _ _ + wideCoequalizer.π f ≫ wideCoequalizer.desc k h = k := by + simp /-- Any morphism `k : Y ⟶ W` satisfying `∀ j₁ j₂, f j₁ ≫ k = f j₂ ≫ k` induces a morphism `l : wideCoequalizer f ⟶ W` satisfying `wideCoequalizer.π ≫ g = l`. -/ diff --git a/Mathlib/CategoryTheory/Monad/Monadicity.lean b/Mathlib/CategoryTheory/Monad/Monadicity.lean index 591201fa317258..b18ae5ed47930c 100644 --- a/Mathlib/CategoryTheory/Monad/Monadicity.lean +++ b/Mathlib/CategoryTheory/Monad/Monadicity.lean @@ -79,13 +79,6 @@ def comparisonLeftAdjointObj (A : adj.toMonad.Algebra) [HasCoequalizer (F.map A.a) (adj.counit.app _)] : D := coequalizer (F.map A.a) (adj.counit.app _) -#adaptation_note -/-- -The new unused variable linter in -https://github.com/leanprover/lean4/pull/5338 -flags `{ f : F.obj A.A ⟶ B // _ }`. --/ -set_option linter.unusedVariables false in /-- We have a bijection of homsets which will be used to construct the left adjoint to the comparison functor. @@ -337,8 +330,8 @@ instance [CreatesColimitOfIsSplitPair G] : ∀ (A : Algebra adj.toMonad), fun _ => CreatesColimitOfIsSplitPair.out _ _ /-- -Beck's monadicity theorem. If `G` has a left adjoint and creates coequalizers of `G`-split pairs, -then it is monadic. +**Beck's monadicity theorem**: if `G` has a left adjoint and creates coequalizers of `G`-split +pairs, then it is monadic. This is the converse of `createsGSplitCoequalizersOfMonadic`. -/ def monadicOfCreatesGSplitCoequalizers [CreatesColimitOfIsSplitPair G] : @@ -351,7 +344,7 @@ def monadicOfCreatesGSplitCoequalizers [CreatesColimitOfIsSplitPair G] : have : ReflectsColimitOfIsSplitPair G := ⟨by intros; infer_instance⟩ exact monadicOfHasPreservesReflectsGSplitCoequalizers adj -/-- An alternate version of Beck's monadicity theorem. If `G` reflects isomorphisms, preserves +/-- An alternate version of **Beck's monadicity theorem**: if `G` reflects isomorphisms, preserves coequalizers of `G`-split pairs and `C` has coequalizers of `G`-split pairs, then it is monadic. -/ def monadicOfHasPreservesGSplitCoequalizersOfReflectsIsomorphisms [G.ReflectsIsomorphisms] diff --git a/Mathlib/CategoryTheory/Monoidal/Mod_.lean b/Mathlib/CategoryTheory/Monoidal/Mod_.lean index cfb9163cb552cb..570447a4586ad9 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mod_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mod_.lean @@ -88,14 +88,11 @@ def forget : Mod_ A ⥤ C where open CategoryTheory.MonoidalCategory -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/6053 +#adaptation_note /-- https://github.com/leanprover/lean4/pull/6053 we needed to increase the `maxHeartbeats` limit if we didn't write an explicit proof for `map_id` and `map_comp`. -This may indicate a configuration problem in Aesop. --/ +This may indicate a configuration problem in Aesop. -/ /-- A morphism of monoid objects induces a "restriction" or "comap" functor between the categories of module objects. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index f87726f0ec8203..b03e0f9b350c37 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -246,14 +246,11 @@ namespace CategoryTheory.Functor variable {C} {D : Type u₂} [Category.{v₂} D] [MonoidalCategory.{v₂} D] -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/6053 +#adaptation_note /-- https://github.com/leanprover/lean4/pull/6053 we needed to increase the `maxHeartbeats` limit if we didn't write an explicit proof for `map_id` and `map_comp`. -This may indicate a configuration problem in Aesop. --/ +This may indicate a configuration problem in Aesop. -/ -- TODO: mapMod F A : Mod A ⥤ Mod (F.mapMon A) /-- A lax monoidal functor takes monoid objects to monoid objects. diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean index 09afedcc3b9aa8..86dba8755a920d 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean @@ -148,12 +148,10 @@ attribute [instance] HasLeftDual.exact open ExactPairing HasRightDual HasLeftDual MonoidalCategory -#adaptation_note -/-- +#adaptation_note /-- https://github.com/leanprover/lean4/pull/4596 The overlapping notation for `leftDual` and `leftAdjointMate` become more problematic in after https://github.com/leanprover/lean4/pull/4596, and we sometimes have to disambiguate with -e.g. `(ᘁX : C)` where previously just `ᘁX` was enough. --/ +e.g. `(ᘁX : C)` where previously just `ᘁX` was enough. -/ @[inherit_doc] prefix:1024 "ᘁ" => leftDual @[inherit_doc] postfix:1024 "ᘁ" => rightDual diff --git a/Mathlib/CategoryTheory/Preadditive/Transfer.lean b/Mathlib/CategoryTheory/Preadditive/Transfer.lean new file mode 100644 index 00000000000000..273c431bbe7b02 --- /dev/null +++ b/Mathlib/CategoryTheory/Preadditive/Transfer.lean @@ -0,0 +1,63 @@ +/- +Copyright (c) 2025 Jakob von Raumer. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jakob von Raumer +-/ +import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor +import Mathlib.Algebra.Equiv.TransferInstance + +/-! +# Pulling back a preadditive structure along a fully faithful functor + +A preadditive structure on a category `D` transfers to a preadditive structure on `C` for a given +fully faithful functor `F : C ⥤ D`. +-/ +namespace CategoryTheory + +open Limits + +universe v₁ v₂ u₁ u₂ + +variable {C : Type u₁} [Category.{v₁} C] +variable {D : Type u₂} [Category.{v₂} D] [Preadditive D] +variable {F : C ⥤ D} (hF : F.FullyFaithful) + +namespace Preadditive + + +/-- If `D` is a preadditive category, any fully faithful functor `F : C ⥤ D` induces a preadditive +structure on `C`. -/ +def ofFullyFaithful : Preadditive C where + homGroup P Q := hF.homEquiv.addCommGroup + add_comp P Q R f f' g := hF.map_injective (by simp [Equiv.add_def]) + comp_add P Q R f g g' := hF.map_injective (by simp [Equiv.add_def]) + +end Preadditive + +open Preadditive +namespace Functor.FullyFaithful + +/-- The preadditive structure on `C` induced by a fully faithful functor `F : C ⥤ D` makes `F` an +additive functor. -/ +lemma additive_ofFullyFaithful : + letI : Preadditive C := Preadditive.ofFullyFaithful hF + F.Additive := + letI : Preadditive C := Preadditive.ofFullyFaithful hF + { map_add := by simp [Equiv.add_def] } + +end Functor.FullyFaithful + +namespace Equivalence + +/-- The preadditive structure on `C` induced by an equivalence `e : C ≌ D` makes `e.inverse` an +additive functor. -/ +lemma additive_inverse_of_FullyFaithful (e : C ≌ D) : + letI : Preadditive C := ofFullyFaithful e.fullyFaithfulFunctor + e.inverse.Additive := + letI : Preadditive C := ofFullyFaithful e.fullyFaithfulFunctor + letI : e.functor.Additive := e.fullyFaithfulFunctor.additive_ofFullyFaithful + e.inverse_additive + +end Equivalence + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Shift/ShiftSequence.lean b/Mathlib/CategoryTheory/Shift/ShiftSequence.lean index 78359306417b7b..fce6ef65bd5b59 100644 --- a/Mathlib/CategoryTheory/Shift/ShiftSequence.lean +++ b/Mathlib/CategoryTheory/Shift/ShiftSequence.lean @@ -85,17 +85,17 @@ def shiftIso (n a a' : M) (ha' : n + a = a') : shiftFunctor C n ⋙ F.shift a ≅ F.shift a' := ShiftSequence.shiftIso n a a' ha' -@[reassoc (attr := simp 1100)] +@[reassoc (attr := simp)] lemma shiftIso_hom_naturality {X Y : C} (n a a' : M) (ha' : n + a = a') (f : X ⟶ Y) : (shift F a).map (f⟦n⟧') ≫ (shiftIso F n a a' ha').hom.app Y = (shiftIso F n a a' ha').hom.app X ≫ (shift F a').map f := (F.shiftIso n a a' ha').hom.naturality f -@[reassoc (attr := simp 1100)] +@[reassoc] lemma shiftIso_inv_naturality {X Y : C} (n a a' : M) (ha' : n + a = a') (f : X ⟶ Y) : (shift F a').map f ≫ (shiftIso F n a a' ha').inv.app Y = - (shiftIso F n a a' ha').inv.app X ≫ (shift F a).map (f⟦n⟧') := - (F.shiftIso n a a' ha').inv.naturality f + (shiftIso F n a a' ha').inv.app X ≫ (shift F a).map (f⟦n⟧') := by + simp variable (M) diff --git a/Mathlib/CategoryTheory/Shift/ShiftedHom.lean b/Mathlib/CategoryTheory/Shift/ShiftedHom.lean index 09bea544345c60..0f182e8d965ba7 100644 --- a/Mathlib/CategoryTheory/Shift/ShiftedHom.lean +++ b/Mathlib/CategoryTheory/Shift/ShiftedHom.lean @@ -92,7 +92,7 @@ lemma comp_mk₀_id {a : M} (f : ShiftedHom X Y a) (m₀ : M) (hm₀ : m₀ = 0) f.comp (mk₀ m₀ hm₀ (𝟙 Y)) (by rw [hm₀, zero_add]) = f := by simp [comp_mk₀] -@[simp 1100] +@[simp] lemma mk₀_comp_mk₀ (f : X ⟶ Y) (g : Y ⟶ Z) {a b c : M} (h : b + a = c) (ha : a = 0) (hb : b = 0) : (mk₀ a ha f).comp (mk₀ b hb g) h = mk₀ c (by rw [← h, ha, hb, add_zero]) (f ≫ g) := by diff --git a/Mathlib/CategoryTheory/Sites/Canonical.lean b/Mathlib/CategoryTheory/Sites/Canonical.lean index 8dd61b381382af..be99c646d4406d 100644 --- a/Mathlib/CategoryTheory/Sites/Canonical.lean +++ b/Mathlib/CategoryTheory/Sites/Canonical.lean @@ -139,11 +139,9 @@ theorem isSheafFor_trans (P : Cᵒᵖ ⥤ Type v) (R S : Sieve X) rw [this] apply hR' hf -/-- Construct the finest (largest) Grothendieck topology for which the given presheaf is a sheaf. - -This is a special case of https://stacks.math.columbia.edu/tag/00Z9, but following a different -proof (see the comments there). --/ +/-- Construct the finest (largest) Grothendieck topology for which the given presheaf is a sheaf. -/ +@[stacks 00Z9 "This is a special case of the Stacks entry, but following a different +proof (see the Stacks comments)."] def finestTopologySingle (P : Cᵒᵖ ⥤ Type v) : GrothendieckTopology C where sieves X S := ∀ (Y) (f : Y ⟶ X), Presieve.IsSheafFor P (S.pullback f : Presieve Y) top_mem' X Y f := by @@ -164,11 +162,9 @@ def finestTopologySingle (P : Cᵒᵖ ⥤ Type v) : GrothendieckTopology C where rw [pullback_id, pullback_comp] at this apply this -/-- -Construct the finest (largest) Grothendieck topology for which all the given presheaves are sheaves. - -This is equal to the construction of . --/ +/-- Construct the finest (largest) Grothendieck topology for which all the given presheaves are +sheaves. -/ +@[stacks 00Z9 "Equal to that Stacks construction"] def finestTopology (Ps : Set (Cᵒᵖ ⥤ Type v)) : GrothendieckTopology C := sInf (finestTopologySingle '' Ps) diff --git a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean index 5137ac167f17ea..820d2cdb384bee 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean @@ -87,13 +87,7 @@ instance : haveI := F.reflects_precoherent; F.IsDenseSubsite (coherentTopology C) (coherentTopology D) where functorPushforward_mem_iff := by rw [eq_induced F] - #adaptation_note - /-- - This proof used to be `rfl`, - but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. - It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 - -/ - exact Iff.rfl + rfl lemma coverPreserving : haveI := F.reflects_precoherent CoverPreserving (coherentTopology _) (coherentTopology _) F := @@ -191,13 +185,7 @@ instance : haveI := F.reflects_preregular; F.IsDenseSubsite (regularTopology C) (regularTopology D) where functorPushforward_mem_iff := by rw [eq_induced F] - #adaptation_note - /-- - This proof used to be `rfl`, - but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. - It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 - -/ - exact Iff.rfl + rfl lemma coverPreserving : haveI := F.reflects_preregular CoverPreserving (regularTopology _) (regularTopology _) F := diff --git a/Mathlib/CategoryTheory/Sites/CoverLifting.lean b/Mathlib/CategoryTheory/Sites/CoverLifting.lean index b6098d47f159a5..db34551ea03238 100644 --- a/Mathlib/CategoryTheory/Sites/CoverLifting.lean +++ b/Mathlib/CategoryTheory/Sites/CoverLifting.lean @@ -211,11 +211,9 @@ variable [∀ (F : Cᵒᵖ ⥤ A), G.op.HasPointwiseRightKanExtension F] /-- If `G` is cocontinuous, then `G.op.ran` pushes sheaves to sheaves. -This is SGA 4 III 2.2. An alternative reference is -https://stacks.math.columbia.edu/tag/00XK (where results -are obtained under the additional assumption that -`C` and `D` have pullbacks). --/ +This is SGA 4 III 2.2. -/ +@[stacks 00XK "Alternative reference. There, results are obtained under the additional assumption +that `C` and `D` have pullbacks."] theorem ran_isSheaf_of_isCocontinuous (ℱ : Sheaf J A) : Presheaf.IsSheaf K (G.op.ran.obj ℱ.val) := by rw [Presheaf.isSheaf_iff_multifork] diff --git a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean index fce2b79404c900..f7d3acf8d8072d 100644 --- a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean +++ b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean @@ -154,11 +154,8 @@ theorem compatiblePreservingOfDownwardsClosed (F : C ⥤ D) [F.Full] [F.Faithful variable {F J K} -/-- If `F` is cover-preserving and compatible-preserving, -then `F` is a continuous functor. - -This result is basically . --/ +/-- If `F` is cover-preserving and compatible-preserving, then `F` is a continuous functor. -/ +@[stacks 00WW "This is basically this Stacks entry."] lemma Functor.isContinuous_of_coverPreserving (hF₁ : CompatiblePreserving.{w} K F) (hF₂ : CoverPreserving J K F) : Functor.IsContinuous.{w} F J K where op_comp_isSheaf_of_types G X S hS x hx := by diff --git a/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean b/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean index 406251e3f3f1bf..c52235dba062f2 100644 --- a/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean +++ b/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean @@ -45,8 +45,9 @@ noncomputable section /-- The middle object of the fork diagram given in Equation (3) of [MM92], as well as the fork diagram -of . +of the Stacks entry. -/ +@[stacks 00VM "This is the middle object of the fork diagram there."] def FirstObj : Type max v u := ∏ᶜ fun f : ΣY, { f : Y ⟶ X // R f } => P.obj (op f.1) @@ -78,8 +79,9 @@ instance : Inhabited (FirstObj P ((⊥ : Sieve X) : Presieve X)) := /-- The left morphism of the fork diagram given in Equation (3) of [MM92], as well as the fork diagram -of . +of the Stacks entry. -/ +@[stacks 00VM "This is the left morphism of the fork diagram there."] def forkMap : P.obj (op X) ⟶ FirstObj P R := Pi.lift fun f => P.map f.2.1.op @@ -176,15 +178,17 @@ namespace Presieve variable [R.hasPullbacks] /-- -The rightmost object of the fork diagram of https://stacks.math.columbia.edu/tag/00VM, which +The rightmost object of the fork diagram of the Stacks entry, which contains the data used to check a family of elements for a presieve is compatible. -/ -@[simp] def SecondObj : Type max v u := +@[simp, stacks 00VM "This is the rightmost object of the fork diagram there."] +def SecondObj : Type max v u := ∏ᶜ fun fg : (ΣY, { f : Y ⟶ X // R f }) × ΣZ, { g : Z ⟶ X // R g } => haveI := Presieve.hasPullbacks.has_pullbacks fg.1.2.2 fg.2.2.2 P.obj (op (pullback fg.1.2.1 fg.2.2.1)) -/-- The map `pr₀*` of . -/ +/-- The map `pr₀*` of the Stacks entry. -/ +@[stacks 00VM "This is the map `pr₀*` there."] def firstMap : FirstObj P R ⟶ SecondObj P R := Pi.lift fun fg => haveI := Presieve.hasPullbacks.has_pullbacks fg.1.2.2 fg.2.2.2 @@ -193,7 +197,8 @@ def firstMap : FirstObj P R ⟶ SecondObj P R := instance [HasPullbacks C] : Inhabited (SecondObj P (⊥ : Presieve X)) := ⟨firstMap _ _ default⟩ -/-- The map `pr₁*` of . -/ +/-- The map `pr₁*` of the Stacks entry. -/ +@[stacks 00VM "This is the map `pr₁*` there."] def secondMap : FirstObj P R ⟶ SecondObj P R := Pi.lift fun fg => haveI := Presieve.hasPullbacks.has_pullbacks fg.1.2.2 fg.2.2.2 @@ -254,10 +259,11 @@ variable {B : C} {I : Type} (X : I → C) (π : (i : I) → X i ⟶ B) -- TODO: allow `I : Type w` /-- -The middle object of the fork diagram of . +The middle object of the fork diagram of the Stacks entry. The difference between this and `Equalizer.FirstObj P (ofArrows X π)` arises if the family of arrows `π` contains duplicates. The `Presieve.ofArrows` doesn't see those. -/ +@[stacks 00VM "The middle object of the fork diagram there."] def FirstObj : Type w := ∏ᶜ (fun i ↦ P.obj (op (X i))) @[ext] @@ -268,10 +274,11 @@ lemma FirstObj.ext (z₁ z₂ : FirstObj P X) (h : ∀ i, (Pi.π _ i : FirstObj exact h i /-- -The rightmost object of the fork diagram of https://stacks.math.columbia.edu/tag/00VM. +The rightmost object of the fork diagram of the Stacks entry. The difference between this and `Equalizer.Presieve.SecondObj P (ofArrows X π)` arises if the family of arrows `π` contains duplicates. The `Presieve.ofArrows` doesn't see those. -/ +@[stacks 00VM "The rightmost object of the fork diagram there."] def SecondObj : Type w := ∏ᶜ (fun (ij : I × I) ↦ P.obj (op (pullback (π ij.1) (π ij.2)))) diff --git a/Mathlib/CategoryTheory/Sites/Grothendieck.lean b/Mathlib/CategoryTheory/Sites/Grothendieck.lean index 816a9f82f51ce0..d581e6a006f6af 100644 --- a/Mathlib/CategoryTheory/Sites/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Sites/Grothendieck.lean @@ -265,13 +265,7 @@ lemma mem_sInf (s : Set (GrothendieckTopology C)) {X : C} (S : Sieve X) : @[stacks 00Z7] theorem isGLB_sInf (s : Set (GrothendieckTopology C)) : IsGLB s (sInf s) := by refine @IsGLB.of_image _ _ _ _ sieves ?_ _ _ ?_ - · #adaptation_note - /-- - This proof used to be `rfl`, - but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. - It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 - -/ - exact Iff.rfl + · rfl · exact _root_.isGLB_sInf _ /-- Construct a complete lattice from the `Inf`, but make the trivial and discrete topologies diff --git a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean index 0d7effb4b905f3..e229cdcd97a5f1 100644 --- a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean +++ b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean @@ -83,9 +83,10 @@ A presheaf is a sheaf (resp, separated) if every *compatible* family of elements (resp, at most one) amalgamation. This data is referred to as a `family` in [MM92], Chapter III, Section 4. It is also a concrete -version of the elements of the middle object in https://stacks.math.columbia.edu/tag/00VM which is +version of the elements of the middle object in the Stacks entry which is more useful for direct calculations. It is also used implicitly in Definition C2.1.2 in [Elephant]. -/ +@[stacks 00VM "This is a concrete version of the elements of the middle object there."] def FamilyOfElements (P : Cᵒᵖ ⥤ Type w) (R : Presieve X) := ∀ ⦃Y : C⦄ (f : Y ⟶ X), R f → P.obj (op Y) @@ -416,8 +417,8 @@ This version is also useful to establish that being a sheaf is preserved under i presheaves. See the discussion before Equation (3) of [MM92], Chapter III, Section 4. See also C2.1.4 of -[Elephant]. This is also a direct reformulation of . --/ +[Elephant]. -/ +@[stacks 00Z8 "Direct reformulation"] def YonedaSheafCondition (P : Cᵒᵖ ⥤ Type v₁) (S : Sieve X) : Prop := ∀ f : S.functor ⟶ P, ∃! g, S.functorInclusion ≫ g = f diff --git a/Mathlib/CategoryTheory/Sites/LocallyInjective.lean b/Mathlib/CategoryTheory/Sites/LocallyInjective.lean index 64b1d4140a669e..9f60ac73c6f1a7 100644 --- a/Mathlib/CategoryTheory/Sites/LocallyInjective.lean +++ b/Mathlib/CategoryTheory/Sites/LocallyInjective.lean @@ -243,8 +243,8 @@ lemma mono_of_isLocallyInjective [IsLocallyInjective φ] : Mono φ := by infer_instance instance {F G : Sheaf J (Type w)} (f : F ⟶ G) : - IsLocallyInjective (imageSheafι f) := by - dsimp [imageSheafι] + IsLocallyInjective (Sheaf.imageι f) := by + dsimp [Sheaf.imageι] infer_instance end Sheaf diff --git a/Mathlib/CategoryTheory/Sites/LocallySurjective.lean b/Mathlib/CategoryTheory/Sites/LocallySurjective.lean index e3c5a1393050fd..390f9e08622c43 100644 --- a/Mathlib/CategoryTheory/Sites/LocallySurjective.lean +++ b/Mathlib/CategoryTheory/Sites/LocallySurjective.lean @@ -50,7 +50,7 @@ def imageSieve {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s : G.obj (op U)) : rw [op_comp, G.map_comp, comp_apply, ← ht, elementwise_of% f.naturality] theorem imageSieve_eq_sieveOfSection {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s : G.obj (op U)) : - imageSieve f s = (imagePresheaf (whiskerRight f (forget A))).sieveOfSection s := + imageSieve f s = (Subpresheaf.range (whiskerRight f (forget A))).sieveOfSection s := rfl theorem imageSieve_whisker_forget {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s : G.obj (op U)) : @@ -92,13 +92,13 @@ instance {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) [IsLocallySurjective J f] : imageSieve_mem s := imageSieve_mem J f s theorem isLocallySurjective_iff_imagePresheaf_sheafify_eq_top {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) : - IsLocallySurjective J f ↔ (imagePresheaf (whiskerRight f (forget A))).sheafify J = ⊤ := by - simp only [Subpresheaf.ext_iff, funext_iff, Set.ext_iff, top_subpresheaf_obj, + IsLocallySurjective J f ↔ (Subpresheaf.range (whiskerRight f (forget A))).sheafify J = ⊤ := by + simp only [Subpresheaf.ext_iff, funext_iff, Set.ext_iff, Subpresheaf.top_obj, Set.top_eq_univ, Set.mem_univ, iff_true] exact ⟨fun H _ => H.imageSieve_mem, fun H => ⟨H _⟩⟩ theorem isLocallySurjective_iff_imagePresheaf_sheafify_eq_top' {F G : Cᵒᵖ ⥤ Type w} (f : F ⟶ G) : - IsLocallySurjective J f ↔ (imagePresheaf f).sheafify J = ⊤ := by + IsLocallySurjective J f ↔ (Subpresheaf.range f).sheafify J = ⊤ := by apply isLocallySurjective_iff_imagePresheaf_sheafify_eq_top theorem isLocallySurjective_iff_whisker_forget {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) : @@ -248,7 +248,7 @@ lemma isLocallySurjective_comp_iff infer_instance instance {F₁ F₂ : Cᵒᵖ ⥤ Type w} (f : F₁ ⟶ F₂) : - IsLocallySurjective J (toImagePresheafSheafify J f) where + IsLocallySurjective J (Subpresheaf.toRangeSheafify J f) where imageSieve_mem {X} := by rintro ⟨s, hs⟩ refine J.superset_covering ?_ hs @@ -257,20 +257,20 @@ instance {F₁ F₂ : Cᵒᵖ ⥤ Type w} (f : F₁ ⟶ F₂) : /-- The image of `F` in `J.sheafify F` is isomorphic to the sheafification. -/ noncomputable def sheafificationIsoImagePresheaf (F : Cᵒᵖ ⥤ Type max u v) : - J.sheafify F ≅ ((imagePresheaf (J.toSheafify F)).sheafify J).toPresheaf where + J.sheafify F ≅ ((Subpresheaf.range (J.toSheafify F)).sheafify J).toPresheaf where hom := - J.sheafifyLift (toImagePresheafSheafify J _) + J.sheafifyLift (Subpresheaf.toRangeSheafify J _) ((isSheaf_iff_isSheaf_of_type J _).mpr <| Subpresheaf.sheafify_isSheaf _ <| (isSheaf_iff_isSheaf_of_type J _).mp <| GrothendieckTopology.sheafify_isSheaf J _) inv := Subpresheaf.ι _ hom_inv_id := - J.sheafify_hom_ext _ _ (J.sheafify_isSheaf _) (by simp [toImagePresheafSheafify]) + J.sheafify_hom_ext _ _ (J.sheafify_isSheaf _) (by simp [Subpresheaf.toRangeSheafify]) inv_hom_id := by rw [← cancel_mono (Subpresheaf.ι _), Category.id_comp, Category.assoc] refine Eq.trans ?_ (Category.comp_id _) congr 1 - exact J.sheafify_hom_ext _ _ (J.sheafify_isSheaf _) (by simp [toImagePresheafSheafify]) + exact J.sheafify_hom_ext _ _ (J.sheafify_isSheaf _) (by simp [Subpresheaf.toRangeSheafify]) section @@ -328,8 +328,8 @@ instance isLocallySurjective_of_iso [IsIso φ] : IsLocallySurjective φ := by infer_instance instance {F G : Sheaf J (Type w)} (f : F ⟶ G) : - IsLocallySurjective (toImageSheaf f) := by - dsimp [toImageSheaf] + IsLocallySurjective (Sheaf.toImage f) := by + dsimp [Sheaf.toImage] infer_instance variable [J.HasSheafCompose (forget A)] @@ -339,11 +339,11 @@ instance [IsLocallySurjective φ] : (Presheaf.isLocallySurjective_iff_whisker_forget J φ.val).1 inferInstance theorem isLocallySurjective_iff_isIso {F G : Sheaf J (Type w)} (f : F ⟶ G) : - IsLocallySurjective f ↔ IsIso (imageSheafι f) := by + IsLocallySurjective f ↔ IsIso (Sheaf.imageι f) := by dsimp only [IsLocallySurjective] - rw [imageSheafι, Presheaf.isLocallySurjective_iff_imagePresheaf_sheafify_eq_top', + rw [Sheaf.imageι, Presheaf.isLocallySurjective_iff_imagePresheaf_sheafify_eq_top', Subpresheaf.eq_top_iff_isIso] - exact isIso_iff_of_reflects_iso (f := imageSheafι f) (F := sheafToPresheaf J (Type w)) + exact isIso_iff_of_reflects_iso (f := Sheaf.imageι f) (F := sheafToPresheaf J (Type w)) instance epi_of_isLocallySurjective' {F₁ F₂ : Sheaf J (Type w)} (φ : F₁ ⟶ F₂) [IsLocallySurjective φ] : Epi φ where @@ -369,7 +369,7 @@ lemma isLocallySurjective_iff_epi {F G : Sheaf J (Type w)} (φ : F ⟶ G) · intro infer_instance · intro - have := epi_of_epi_fac (toImageSheaf_ι φ) + have := epi_of_epi_fac (Sheaf.toImage_ι φ) rw [isLocallySurjective_iff_isIso φ] apply isIso_of_mono_of_epi diff --git a/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean b/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean index b15c8b4f9117fa..4452691a4e8ba5 100644 --- a/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean +++ b/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean @@ -57,11 +57,8 @@ instance : Group (ZeroCochain G U) := Pi.group namespace Cochain₀ -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/4481 -the `simpNF` linter incorrectly claims this lemma can't be applied by `simp`. --/ +#adaptation_note /-- https://github.com/leanprover/lean4/pull/4481 +the `simpNF` linter incorrectly claims this lemma can't be applied by `simp`. -/ @[simp, nolint simpNF] lemma one_apply (i : I) : (1 : ZeroCochain G U) i = 1 := rfl diff --git a/Mathlib/CategoryTheory/Sites/Pretopology.lean b/Mathlib/CategoryTheory/Sites/Pretopology.lean index 6285e9b9a6bde9..1732a166ae6f3a 100644 --- a/Mathlib/CategoryTheory/Sites/Pretopology.lean +++ b/Mathlib/CategoryTheory/Sites/Pretopology.lean @@ -52,12 +52,10 @@ three axioms: In some sense, a pretopology can be seen as Grothendieck topology with weaker saturation conditions, in that each covering is not necessarily downward closed. -See: https://ncatlab.org/nlab/show/Grothendieck+pretopology, or -https://stacks.math.columbia.edu/tag/00VH, or [MM92] Chapter III, Section 2, Definition 2. -Note that Stacks calls a category together with a pretopology a site, and [MM92] calls this -a basis for a topology. --/ -@[ext] +See: https://ncatlab.org/nlab/show/Grothendieck+pretopology or [MM92] Chapter III, +Section 2, Definition 2. -/ +@[ext, stacks 00VH "Note that Stacks calls a category together with a pretopology a site, +and [MM92] calls this a basis for a topology."] structure Pretopology where coverings : ∀ X : C, Set (Presieve X) has_isos : ∀ ⦃X Y⦄ (f : Y ⟶ X) [IsIso f], Presieve.singleton f ∈ coverings X diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index 9b7c318a4906ac..b213f408e9c076 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -68,10 +68,8 @@ variable (J : GrothendieckTopology C) -- We follow https://stacks.math.columbia.edu/tag/00VL definition 00VR /-- A sheaf of A is a presheaf P : Cᵒᵖ => A such that for every E : A, the -presheaf of types given by sending U : C to Hom_{A}(E, P U) is a sheaf of types. - -https://stacks.math.columbia.edu/tag/00VR --/ +presheaf of types given by sending U : C to Hom_{A}(E, P U) is a sheaf of types. -/ +@[stacks 00VR] def IsSheaf (P : Cᵒᵖ ⥤ A) : Prop := ∀ E : A, Presieve.IsSheaf J (P ⋙ coyoneda.obj (op E)) @@ -586,34 +584,35 @@ section variable [HasProducts.{max u₁ v₁} A] variable [HasProducts.{max u₁ v₁} A'] -/-- -The middle object of the fork diagram given in Equation (3) of [MM92], as well as the fork diagram -of . --/ +/-- The middle object of the fork diagram given in Equation (3) of [MM92], as well as the fork +diagram of the Stacks entry. -/ +@[stacks 00VM "The middle object of the fork diagram there."] def firstObj : A := ∏ᶜ fun f : ΣV, { f : V ⟶ U // R f } => P.obj (op f.1) -/-- -The left morphism of the fork diagram given in Equation (3) of [MM92], as well as the fork diagram -of . --/ +/-- The left morphism of the fork diagram given in Equation (3) of [MM92], as well as the fork +diagram of the Stacks entry. -/ +@[stacks 00VM "The left morphism the fork diagram there."] def forkMap : P.obj (op U) ⟶ firstObj R P := Pi.lift fun f => P.map f.2.1.op variable [HasPullbacks C] -/-- The rightmost object of the fork diagram of https://stacks.math.columbia.edu/tag/00VM, which +/-- The rightmost object of the fork diagram of the Stacks entry, which contains the data used to check a family of elements for a presieve is compatible. -/ +@[stacks 00VM "The rightmost object of the fork diagram there."] def secondObj : A := ∏ᶜ fun fg : (ΣV, { f : V ⟶ U // R f }) × ΣW, { g : W ⟶ U // R g } => P.obj (op (pullback fg.1.2.1 fg.2.2.1)) -/-- The map `pr₀*` of . -/ +/-- The map `pr₀*` of the Stacks entry. -/ +@[stacks 00VM "The map `pr₀*` there."] def firstMap : firstObj R P ⟶ secondObj R P := Pi.lift fun _ => Pi.π _ _ ≫ P.map (pullback.fst _ _).op -/-- The map `pr₁*` of . -/ +/-- The map `pr₁*` of the Stacks entry. -/ +@[stacks 00VM "The map `pr₁*` there."] def secondMap : firstObj R P ⟶ secondObj R P := Pi.lift fun _ => Pi.π _ _ ≫ P.map (pullback.snd _ _).op diff --git a/Mathlib/CategoryTheory/Sites/Subsheaf.lean b/Mathlib/CategoryTheory/Sites/Subsheaf.lean index 3be19de81341bc..03d5262b3fe05d 100644 --- a/Mathlib/CategoryTheory/Sites/Subsheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Subsheaf.lean @@ -194,16 +194,17 @@ variable (J) /-- A morphism factors through the sheafification of the image presheaf. -/ @[simps!] -def toImagePresheafSheafify (f : F' ⟶ F) : F' ⟶ ((imagePresheaf f).sheafify J).toPresheaf := - toImagePresheaf f ≫ Subpresheaf.homOfLe ((imagePresheaf f).le_sheafify J) +def Subpresheaf.toRangeSheafify (f : F' ⟶ F) : F' ⟶ ((Subpresheaf.range f).sheafify J).toPresheaf := + toRange f ≫ Subpresheaf.homOfLe ((range f).le_sheafify J) + variable {J} /-- The image sheaf of a morphism between sheaves, defined to be the sheafification of `image_presheaf`. -/ @[simps] -def imageSheaf {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Sheaf J (Type w) := - ⟨((imagePresheaf f.1).sheafify J).toPresheaf, by +def Sheaf.image {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Sheaf J (Type w) := + ⟨((Subpresheaf.range f.1).sheafify J).toPresheaf, by rw [isSheaf_iff_isSheaf_of_type] apply Subpresheaf.sheafify_isSheaf rw [← isSheaf_iff_isSheaf_of_type] @@ -211,34 +212,35 @@ def imageSheaf {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Sheaf J (Type w) := /-- A morphism factors through the image sheaf. -/ @[simps] -def toImageSheaf {F F' : Sheaf J (Type w)} (f : F ⟶ F') : F ⟶ imageSheaf f := - ⟨toImagePresheafSheafify J f.1⟩ +def Sheaf.toImage {F F' : Sheaf J (Type w)} (f : F ⟶ F') : F ⟶ Sheaf.image f := + ⟨Subpresheaf.toRangeSheafify J f.1⟩ /-- The inclusion of the image sheaf to the target. -/ @[simps] -def imageSheafι {F F' : Sheaf J (Type w)} (f : F ⟶ F') : imageSheaf f ⟶ F' := +def Sheaf.imageι {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Sheaf.image f ⟶ F' := ⟨Subpresheaf.ι _⟩ + @[reassoc (attr := simp)] -theorem toImageSheaf_ι {F F' : Sheaf J (Type w)} (f : F ⟶ F') : - toImageSheaf f ≫ imageSheafι f = f := by +theorem Sheaf.toImage_ι {F F' : Sheaf J (Type w)} (f : F ⟶ F') : + toImage f ≫ imageι f = f := by ext1 - simp [toImagePresheafSheafify] + simp [Subpresheaf.toRangeSheafify] -instance {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Mono (imageSheafι f) := +instance {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Mono (Sheaf.imageι f) := (sheafToPresheaf J _).mono_of_mono_map (by dsimp infer_instance) -instance {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Epi (toImageSheaf f) := by +instance {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Epi (Sheaf.toImage f) := by refine ⟨@fun G' g₁ g₂ e => ?_⟩ ext U ⟨s, hx⟩ apply ((isSheaf_iff_isSheaf_of_type J _).mp G'.2 _ hx).isSeparatedFor.ext rintro V i ⟨y, e'⟩ change (g₁.val.app _ ≫ G'.val.map _) _ = (g₂.val.app _ ≫ G'.val.map _) _ rw [← NatTrans.naturality, ← NatTrans.naturality] - have E : (toImageSheaf f).val.app (op V) y = (imageSheaf f).val.map i.op ⟨s, hx⟩ := + have E : (Sheaf.toImage f).val.app (op V) y = (Sheaf.image f).val.map i.op ⟨s, hx⟩ := Subtype.ext e' have := congr_arg (fun f : F ⟶ G' => (Sheaf.Hom.val f).app _ y) e dsimp at this ⊢ @@ -246,9 +248,9 @@ instance {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Epi (toImageSheaf f) := by /-- The mono factorization given by `image_sheaf` for a morphism. -/ def imageMonoFactorization {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Limits.MonoFactorisation f where - I := imageSheaf f - m := imageSheafι f - e := toImageSheaf f + I := Sheaf.image f + m := Sheaf.imageι f + e := Sheaf.toImage f /-- The mono factorization given by `image_sheaf` for a morphism is an image. -/ noncomputable def imageFactorization {F F' : Sheaf J (Type (max v u))} (f : F ⟶ F') : @@ -257,14 +259,13 @@ noncomputable def imageFactorization {F F' : Sheaf J (Type (max v u))} (f : F isImage := { lift := fun I => by haveI M := (Sheaf.Hom.mono_iff_presheaf_mono J (Type (max v u)) _).mp I.m_mono - haveI := isIso_toImagePresheaf I.m.1 - refine ⟨Subpresheaf.homOfLe ?_ ≫ inv (toImagePresheaf I.m.1)⟩ + refine ⟨Subpresheaf.homOfLe ?_ ≫ inv (Subpresheaf.toRange I.m.1)⟩ apply Subpresheaf.sheafify_le · conv_lhs => rw [← I.fac] - apply imagePresheaf_comp_le + apply Subpresheaf.range_comp_le · rw [← isSheaf_iff_isSheaf_of_type] exact F'.2 - · apply Presieve.isSheaf_iso J (asIso <| toImagePresheaf I.m.1) + · apply Presieve.isSheaf_iso J (asIso <| Subpresheaf.toRange I.m.1) rw [← isSheaf_iff_isSheaf_of_type] exact I.I.2 lift_fac := fun I => by @@ -273,10 +274,17 @@ noncomputable def imageFactorization {F F' : Sheaf J (Type (max v u))} (f : F generalize_proofs h rw [← Subpresheaf.homOfLe_ι h, Category.assoc] congr 1 - rw [IsIso.inv_comp_eq, toImagePresheaf_ι] } + rw [IsIso.inv_comp_eq, Subpresheaf.toRange_ι] } instance : Limits.HasImages (Sheaf J (Type max v u)) := - ⟨@fun _ _ f => ⟨⟨imageFactorization f⟩⟩⟩ + ⟨fun f => ⟨⟨imageFactorization f⟩⟩⟩ + +@[deprecated (since := "2025-01-25")] alias toImagePresheafSheafify := + Subpresheaf.toRangeSheafify +@[deprecated (since := "2025-01-25")] alias imageSheaf := Sheaf.image +@[deprecated (since := "2025-01-25")] alias toImageSheaf := Sheaf.toImage +@[deprecated (since := "2025-01-25")] alias imageSheafι := Sheaf.imageι +@[deprecated (since := "2025-01-25")] alias toImageSheaf_ι := Sheaf.toImage_ι end Image diff --git a/Mathlib/CategoryTheory/Subpresheaf/Basic.lean b/Mathlib/CategoryTheory/Subpresheaf/Basic.lean index 080283e8b18806..890d33e462b17e 100644 --- a/Mathlib/CategoryTheory/Subpresheaf/Basic.lean +++ b/Mathlib/CategoryTheory/Subpresheaf/Basic.lean @@ -1,11 +1,11 @@ /- Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Andrew Yang +Authors: Andrew Yang, Joël Riou -/ import Mathlib.CategoryTheory.Elementwise -import Mathlib.Data.Set.Basic - +import Mathlib.Data.Set.Lattice +import Mathlib.Order.CompleteLattice /-! @@ -47,15 +47,99 @@ variable {F F' F'' : Cᵒᵖ ⥤ Type w} (G G' : Subpresheaf F) instance : PartialOrder (Subpresheaf F) := PartialOrder.lift Subpresheaf.obj (fun _ _ => Subpresheaf.ext) -instance : Top (Subpresheaf F) := - ⟨⟨fun _ => ⊤, @fun U _ _ x _ => by simp⟩⟩ +instance : CompleteLattice (Subpresheaf F) where + sup F G := + { obj U := F.obj U ⊔ G.obj U + map _ _ := by + rintro (h|h) + · exact Or.inl (F.map _ h) + · exact Or.inr (G.map _ h) } + le_sup_left _ _ _ := by simp + le_sup_right _ _ _ := by simp + sup_le F G H h₁ h₂ U := by + rintro x (h|h) + · exact h₁ _ h + · exact h₂ _ h + inf S T := + { obj U := S.obj U ⊓ T.obj U + map _ _ h := ⟨S.map _ h.1, T.map _ h.2⟩} + inf_le_left _ _ _ _ h := h.1 + inf_le_right _ _ _ _ h := h.2 + le_inf _ _ _ h₁ h₂ _ _ h := ⟨h₁ _ h, h₂ _ h⟩ + sSup S := + { obj U := sSup (Set.image (fun T ↦ T.obj U) S) + map f x hx := by + obtain ⟨_, ⟨F, h, rfl⟩, h'⟩ := hx + simp only [Set.sSup_eq_sUnion, Set.sUnion_image, Set.preimage_iUnion, + Set.mem_iUnion, Set.mem_preimage, exists_prop] + exact ⟨_, h, F.map f h'⟩ } + le_sSup _ _ _ _ _ := by aesop + sSup_le _ _ _ _ _ := by aesop + sInf S := + { obj U := sInf (Set.image (fun T ↦ T.obj U) S) + map f x hx := by + rintro _ ⟨F, h, rfl⟩ + exact F.map f (hx _ ⟨_, h, rfl⟩) } + sInf_le _ _ _ _ _ := by aesop + le_sInf _ _ _ _ _ := by aesop + bot := + { obj U := ⊥ + map := by simp } + bot_le _ _ := bot_le + top := + { obj U := ⊤ + map := by simp } + le_top _ _ := le_top + +namespace Subpresheaf + +lemma le_def (S T : Subpresheaf F) : S ≤ T ↔ ∀ U, S.obj U ≤ T.obj U := Iff.rfl + +variable (F) + +@[simp] lemma top_obj (i : Cᵒᵖ) : (⊤ : Subpresheaf F).obj i = ⊤ := rfl +@[simp] lemma bot_obj (i : Cᵒᵖ) : (⊥ : Subpresheaf F).obj i = ⊥ := rfl + +variable {F} + +lemma sSup_obj (S : Set (Subpresheaf F)) (U : Cᵒᵖ) : + (sSup S).obj U = sSup (Set.image (fun T ↦ T.obj U) S) := rfl + +lemma sInf_obj (S : Set (Subpresheaf F)) (U : Cᵒᵖ) : + (sInf S).obj U = sInf (Set.image (fun T ↦ T.obj U) S) := rfl + +@[simp] +lemma iSup_obj {ι : Type*} (S : ι → Subpresheaf F) (U : Cᵒᵖ) : + (⨆ i, S i).obj U = ⋃ i, (S i).obj U := by + simp [iSup, sSup_obj] + +@[simp] +lemma iInf_obj {ι : Type*} (S : ι → Subpresheaf F) (U : Cᵒᵖ) : + (⨅ i, S i).obj U = ⋂ i, (S i).obj U := by + simp [iInf, sInf_obj] + +@[simp] +lemma max_obj (S T : Subpresheaf F) (i : Cᵒᵖ) : + (S ⊔ T).obj i = S.obj i ∪ T.obj i := rfl + +@[simp] +lemma min_obj (S T : Subpresheaf F) (i : Cᵒᵖ) : + (S ⊓ T).obj i = S.obj i ∩ T.obj i := rfl + +lemma max_min (S₁ S₂ T : Subpresheaf F) : + (S₁ ⊔ S₂) ⊓ T = (S₁ ⊓ T) ⊔ (S₂ ⊓ T) := by + aesop + +lemma iSup_min {ι : Type*} (S : ι → Subpresheaf F) (T : Subpresheaf F) : + (⨆ i, S i) ⊓ T = ⨆ i, S i ⊓ T := by + aesop instance : Nonempty (Subpresheaf F) := inferInstance /-- The subpresheaf as a presheaf. -/ @[simps!] -def Subpresheaf.toPresheaf : Cᵒᵖ ⥤ Type w where +def toPresheaf : Cᵒᵖ ⥤ Type w where obj U := G.obj U map := @fun _ _ i x => ⟨F.map i x, G.map i x.prop⟩ map_id X := by @@ -72,7 +156,7 @@ instance {U} : CoeHead (G.toPresheaf.obj U) (F.obj U) where /-- The inclusion of a subpresheaf to the original presheaf. -/ @[simps] -def Subpresheaf.ι : G.toPresheaf ⟶ F where app _ x := x +def ι : G.toPresheaf ⟶ F where app _ x := x instance : Mono G.ι := ⟨@fun _ _ _ e => @@ -81,7 +165,7 @@ instance : Mono G.ι := /-- The inclusion of a subpresheaf to a larger subpresheaf -/ @[simps] -def Subpresheaf.homOfLe {G G' : Subpresheaf F} (h : G ≤ G') : G.toPresheaf ⟶ G'.toPresheaf where +def homOfLe {G G' : Subpresheaf F} (h : G ≤ G') : G.toPresheaf ⟶ G'.toPresheaf where app U x := ⟨x, h U x.prop⟩ instance {G G' : Subpresheaf F} (h : G ≤ G') : Mono (Subpresheaf.homOfLe h) := @@ -92,7 +176,7 @@ instance {G G' : Subpresheaf F} (h : G ≤ G') : Mono (Subpresheaf.homOfLe h) := Subtype.ext <| (congr_arg Subtype.val <| (congr_fun (congr_app e U) x :) :)⟩ @[reassoc (attr := simp)] -theorem Subpresheaf.homOfLe_ι {G G' : Subpresheaf F} (h : G ≤ G') : +theorem homOfLe_ι {G G' : Subpresheaf F} (h : G ≤ G') : Subpresheaf.homOfLe h ≫ G'.ι = G.ι := by ext rfl @@ -103,7 +187,7 @@ instance : IsIso (Subpresheaf.ι (⊤ : Subpresheaf F)) := by rw [isIso_iff_bijective] exact ⟨Subtype.coe_injective, fun x => ⟨⟨x, _root_.trivial⟩, rfl⟩⟩ -theorem Subpresheaf.eq_top_iff_isIso : G = ⊤ ↔ IsIso G.ι := by +theorem eq_top_iff_isIso : G = ⊤ ↔ IsIso G.ι := by constructor · rintro rfl infer_instance @@ -113,12 +197,12 @@ theorem Subpresheaf.eq_top_iff_isIso : G = ⊤ ↔ IsIso G.ι := by rw [← IsIso.inv_hom_id_apply (G.ι.app U) x] exact ((inv (G.ι.app U)) x).2 -theorem Subpresheaf.nat_trans_naturality (f : F' ⟶ G.toPresheaf) {U V : Cᵒᵖ} (i : U ⟶ V) +theorem nat_trans_naturality (f : F' ⟶ G.toPresheaf) {U V : Cᵒᵖ} (i : U ⟶ V) (x : F'.obj U) : (f.app V (F'.map i x)).1 = F.map i (f.app U x).1 := congr_arg Subtype.val (FunctorToTypes.naturality _ _ f i x) -@[simp] -theorem top_subpresheaf_obj (U) : (⊤ : Subpresheaf F).obj U = ⊤ := - rfl +end Subpresheaf + +@[deprecated (since := "2025-01-23")] alias top_subpresheaf_obj := Subpresheaf.top_obj end CategoryTheory diff --git a/Mathlib/CategoryTheory/Subpresheaf/Image.lean b/Mathlib/CategoryTheory/Subpresheaf/Image.lean index 367029af657f05..4247b888fb3af6 100644 --- a/Mathlib/CategoryTheory/Subpresheaf/Image.lean +++ b/Mathlib/CategoryTheory/Subpresheaf/Image.lean @@ -1,20 +1,18 @@ /- Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Andrew Yang +Authors: Andrew Yang, Joël Riou -/ - import Mathlib.CategoryTheory.Subpresheaf.Basic import Mathlib.CategoryTheory.Limits.FunctorCategory.EpiMono /-! # The image of a subpresheaf -Given a morphism of presheaves of types `f : F' ⟶ F`, we define `imagePresheaf f` -as a subpresheaf of `F`. - -## TODO (@joelriou) -* introduce `Subpresheaf.image` and `Subpresheaf.preimage` +Given a morphism of presheaves of types `p : F' ⟶ F`, we define its range +`Subpresheaf.range p`. More generally, if `G' : Subpresheaf F'`, we +define `G'.image p : Subpresheaf F` as the image of `G'` by `f`, and +if `G : Subpresheaf F`, we define its preimage `G.preimage f : Subpresheaf F'`. -/ @@ -22,63 +20,171 @@ universe w v u namespace CategoryTheory -variable {C : Type u} [Category.{v} C] {F F' F'' : Cᵒᵖ ⥤ Type w} {G : Subpresheaf F} +variable {C : Type u} [Category.{v} C] {F F' F'' : Cᵒᵖ ⥤ Type w} -/-- If the image of a morphism falls in a subpresheaf, then the morphism factors through it. -/ -@[simps!] -def Subpresheaf.lift (f : F' ⟶ F) (hf : ∀ U x, f.app U x ∈ G.obj U) : F' ⟶ G.toPresheaf where - app U x := ⟨f.app U x, hf U x⟩ - naturality := by - have := elementwise_of% f.naturality - intros - refine funext fun x => Subtype.ext ?_ - simp only [toPresheaf_obj, types_comp_apply] - exact this _ _ +namespace Subpresheaf -@[reassoc (attr := simp)] -theorem Subpresheaf.lift_ι (f : F' ⟶ F) (hf : ∀ U x, f.app U x ∈ G.obj U) : - G.lift f hf ≫ G.ι = f := by - ext - rfl +section range -/-- The image presheaf of a morphism, whose components are the set-theoretic images. -/ +/-- The range of a morphism of presheaves of types, as a subpresheaf of the target. -/ @[simps] -def imagePresheaf (f : F' ⟶ F) : Subpresheaf F where - obj U := Set.range (f.app U) +def range (p : F' ⟶ F) : Subpresheaf F where + obj U := Set.range (p.app U) map := by rintro U V i _ ⟨x, rfl⟩ - have := elementwise_of% f.naturality - exact ⟨_, this i x⟩ + exact ⟨_, FunctorToTypes.naturality _ _ p i x⟩ + +variable (F) in +lemma range_id : range (𝟙 F) = ⊤ := by aesop @[simp] -theorem imagePresheaf_id : imagePresheaf (𝟙 F) = ⊤ := by - ext - simp +lemma range_ι (G : Subpresheaf F) : range G.ι = G := by aesop + +end range + +section lift -/-- A morphism factors through the image presheaf. -/ +variable (f : F' ⟶ F) {G : Subpresheaf F} (hf : range f ≤ G) + +/-- If the image of a morphism falls in a subpresheaf, then the morphism factors through it. -/ @[simps!] -def toImagePresheaf (f : F' ⟶ F) : F' ⟶ (imagePresheaf f).toPresheaf := - (imagePresheaf f).lift f fun _ _ => Set.mem_range_self _ +def lift : F' ⟶ G.toPresheaf where + app U x := ⟨f.app U x, hf U (by simp)⟩ + naturality _ _ g := by + ext x + simpa [Subtype.ext_iff] using FunctorToTypes.naturality _ _ f g x @[reassoc (attr := simp)] -theorem toImagePresheaf_ι (f : F' ⟶ F) : toImagePresheaf f ≫ (imagePresheaf f).ι = f := - (imagePresheaf f).lift_ι _ _ - -theorem imagePresheaf_comp_le (f₁ : F ⟶ F') (f₂ : F' ⟶ F'') : - imagePresheaf (f₁ ≫ f₂) ≤ imagePresheaf f₂ := fun U _ hx => ⟨f₁.app U hx.choose, hx.choose_spec⟩ - -instance isIso_toImagePresheaf {F F' : Cᵒᵖ ⥤ Type w} (f : F ⟶ F') [hf : Mono f] : - IsIso (toImagePresheaf f) := by - have : ∀ (X : Cᵒᵖ), IsIso ((toImagePresheaf f).app X) := by - intro X - rw [isIso_iff_bijective] - constructor - · intro x y e - have := (NatTrans.mono_iff_mono_app f).mp hf X - rw [mono_iff_injective] at this - exact this (congr_arg Subtype.val e :) - · rintro ⟨_, ⟨x, rfl⟩⟩ - exact ⟨x, rfl⟩ - apply NatIso.isIso_of_isIso_app +theorem lift_ι : lift f hf ≫ G.ι = f := rfl + +end lift + +section range + +variable (p : F' ⟶ F) + +/-- Given a morphism `p : F' ⟶ F` of presheaves of types, this is the morphism +from `F'` to its range. -/ +def toRange : + F' ⟶ (range p).toPresheaf := + lift p (by rfl) + +@[reassoc (attr := simp)] +lemma toRange_ι : toRange p ≫ (range p).ι = p := rfl + +lemma toRange_app_val {i : Cᵒᵖ} (x : F'.obj i) : + ((toRange p).app i x).val = p.app i x := by + simp [toRange] + +@[simp] +lemma range_toRange : range (toRange p) = ⊤ := by + ext i ⟨x, hx⟩ + dsimp at hx ⊢ + simp only [Set.mem_range, Set.mem_univ, iff_true] + simp only [Set.mem_range] at hx + obtain ⟨y, rfl⟩ := hx + exact ⟨y, rfl⟩ + +lemma epi_iff_range_eq_top : + Epi p ↔ range p = ⊤ := by + simp [NatTrans.epi_iff_epi_app, epi_iff_surjective, Subpresheaf.ext_iff, funext_iff, + Set.range_eq_univ] + +instance : Epi (toRange p) := by simp [epi_iff_range_eq_top] + +instance [Mono p] : IsIso (toRange p) := by + have := mono_of_mono_fac (toRange_ι p) + rw [NatTrans.isIso_iff_isIso_app] + intro i + rw [isIso_iff_bijective] + constructor + · rw [← mono_iff_injective] + infer_instance + · rw [← epi_iff_surjective] + infer_instance + +lemma range_comp_le (f : F ⟶ F') (g : F' ⟶ F'') : + range (f ≫ g) ≤ range g := fun _ _ _ ↦ by aesop + +end range + +section image + +variable (G : Subpresheaf F) (f : F ⟶ F') + +/-- The image of a subpresheaf by a morphism of presheaves of types. -/ +@[simps] +def image : Subpresheaf F' where + obj i := (f.app i) '' (G.obj i) + map := by + rintro Δ Δ' φ _ ⟨x, hx, rfl⟩ + exact ⟨F.map φ x, G.map φ hx, by apply FunctorToTypes.naturality⟩ + +lemma image_top : (⊤ : Subpresheaf F).image f = range f := by aesop + +lemma image_comp (g : F' ⟶ F'') : + G.image (f ≫ g) = (G.image f).image g := by aesop + +lemma range_comp (g : F' ⟶ F'') : + range (f ≫ g) = (range f).image g := by aesop + +end image + +section preimage + +/-- The preimage of a subpresheaf by a morphism of presheaves of types. -/ +@[simps] +def preimage (G : Subpresheaf F) (p : F' ⟶ F) : Subpresheaf F' where + obj n := p.app n ⁻¹' (G.obj n) + map f := (Set.preimage_mono (G.map f)).trans (by + simp only [Set.preimage_preimage, FunctorToTypes.naturality _ _ p f] + rfl) + +@[simp] +lemma preimage_id (G : Subpresheaf F) : + G.preimage (𝟙 F) = G := by aesop + +lemma preimage_comp (G : Subpresheaf F) (f : F'' ⟶ F') (g : F' ⟶ F) : + G.preimage (f ≫ g) = (G.preimage g).preimage f := by aesop + +lemma image_le_iff (G : Subpresheaf F) (f : F ⟶ F') (G' : Subpresheaf F') : + G.image f ≤ G' ↔ G ≤ G'.preimage f := by + simp [Subpresheaf.le_def] + +/-- Given a morphism `p : F' ⟶ F` of presheaves of types and `G : Subpresheaf F`, +this is the morphism from the preimage of `G` by `p` to `G`. -/ +def fromPreimage (G : Subpresheaf F) (p : F' ⟶ F) : + (G.preimage p).toPresheaf ⟶ G.toPresheaf := + lift ((G.preimage p).ι ≫ p) (by + rw [range_comp, range_ι, image_le_iff]) + +@[reassoc] +lemma fromPreimage_ι (G : Subpresheaf F) (p : F' ⟶ F) : + G.fromPreimage p ≫ G.ι = (G.preimage p).ι ≫ p := rfl + +lemma preimage_eq_top_iff (G : Subpresheaf F) (p : F' ⟶ F) : + G.preimage p = ⊤ ↔ range p ≤ G := by + rw [← image_top, image_le_iff] + aesop + +@[simp] +lemma preimage_image_of_epi (G : Subpresheaf F) (p : F' ⟶ F) [hp : Epi p] : + (G.preimage p).image p = G := by + apply le_antisymm + · rw [image_le_iff] + · intro i x hx + simp only [NatTrans.epi_iff_epi_app, epi_iff_surjective] at hp + obtain ⟨y, rfl⟩ := hp _ x + exact ⟨y, hx, rfl⟩ + +end preimage + +end Subpresheaf + +@[deprecated (since := "2025-01-25")] alias imagePresheaf := Subpresheaf.range +@[deprecated (since := "2025-01-25")] alias imagePresheaf_id := Subpresheaf.range_id +@[deprecated (since := "2025-01-25")] alias toImagePresheaf := Subpresheaf.toRange +@[deprecated (since := "2025-01-25")] alias toImagePresheaf_ι := Subpresheaf.toRange_ι +@[deprecated (since := "2025-01-25")] alias imagePresheaf_comp_le := Subpresheaf.range_comp_le end CategoryTheory diff --git a/Mathlib/CategoryTheory/Subpresheaf/OfSection.lean b/Mathlib/CategoryTheory/Subpresheaf/OfSection.lean new file mode 100644 index 00000000000000..3bb6facb0c17fc --- /dev/null +++ b/Mathlib/CategoryTheory/Subpresheaf/OfSection.lean @@ -0,0 +1,105 @@ +/- +Copyright (c) 2025 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Subpresheaf.Image +import Mathlib.CategoryTheory.Yoneda + +/-! +# The subpresheaf generated by a section + +Given a presheaf of types `F : Cᵒᵖ ⥤ Type w` and a section `x : F.obj X`, +we define `Subpresheaf.ofSection x : Subpresheaf F` as the subpresheaf +of `F` generated by `x`. + +-/ + +universe w v u + +namespace CategoryTheory + +variable {C : Type u} [Category.{v} C] + +namespace Subpresheaf + +section + +variable {F : Cᵒᵖ ⥤ Type w} {X : Cᵒᵖ} (x : F.obj X) + +/-- The subpresheaf of `F : Cᵒᵖ ⥤ Type w` that is generated +by a section `x : F.obj X`. -/ +@[simps (config := .lemmasOnly)] +def ofSection : Subpresheaf F where + obj U := setOf (fun u ↦ ∃ (f : X ⟶ U), F.map f x = u) + map {U V} g := by + rintro _ ⟨f, rfl⟩ + exact ⟨f ≫ g, by simp⟩ + +lemma mem_ofSection_obj : x ∈ (ofSection x).obj X := ⟨𝟙 _, by simp⟩ + +@[simp] +lemma ofSection_le_iff (G : Subpresheaf F) : + ofSection x ≤ G ↔ x ∈ G.obj X := by + constructor + · intro hx + exact hx _ (mem_ofSection_obj x) + · rintro hx U _ ⟨f, rfl⟩ + exact G.map f hx + +@[simp] +lemma ofSection_image {F' : Cᵒᵖ ⥤ Type w} (f : F ⟶ F') : + (ofSection x).image f = ofSection (f.app _ x) := by + apply le_antisymm + · rw [image_le_iff, ofSection_le_iff, preimage_obj, Set.mem_preimage] + exact ⟨𝟙 X, by simp⟩ + · simp only [ofSection_le_iff, image_obj, Set.mem_image] + exact ⟨x, mem_ofSection_obj x, rfl⟩ + +end + +section + +variable {F : Cᵒᵖ ⥤ Type v} + +lemma ofSection_eq_range {X : Cᵒᵖ} (x : F.obj X) : + ofSection x = range (yonedaEquiv.symm x) := by + ext U y + simp only [ofSection_obj, Set.mem_setOf_eq, Opposite.op_unop, range_obj, yoneda_obj_obj, + Set.mem_range] + constructor + · rintro ⟨f, rfl⟩ + exact ⟨f.unop, rfl⟩ + · rintro ⟨f, rfl⟩ + exact ⟨f.op, rfl⟩ + +lemma range_eq_ofSection {X : C} (f : yoneda.obj X ⟶ F) : + range f = ofSection (yonedaEquiv f) := by + rw [ofSection_eq_range, Equiv.symm_apply_apply] + +end + +section + +variable {F : Cᵒᵖ ⥤ Type max v w} + +lemma ofSection_eq_range' {X : Cᵒᵖ} (x : F.obj X) : + ofSection x = range ((yonedaCompUliftFunctorEquiv F X.unop).symm x) := by + ext U y + simp only [Opposite.op_unop, range_obj, Functor.comp_obj, yoneda_obj_obj, uliftFunctor_obj, + Set.mem_range, ULift.exists] + constructor + · rintro ⟨f, rfl⟩ + exact ⟨f.unop, rfl⟩ + · rintro ⟨f, rfl⟩ + exact ⟨f.op, rfl⟩ + +lemma range_eq_ofSection' {X : C} (f : yoneda.obj X ⋙ uliftFunctor.{w} ⟶ F) : + range f = ofSection ((yonedaCompUliftFunctorEquiv F X) f) := by + rw [ofSection_eq_range', Equiv.symm_apply_apply] + +end + +end Subpresheaf + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Subpresheaf/Sieves.lean b/Mathlib/CategoryTheory/Subpresheaf/Sieves.lean index 12ffd8506ee789..7c84f551edcd18 100644 --- a/Mathlib/CategoryTheory/Subpresheaf/Sieves.lean +++ b/Mathlib/CategoryTheory/Subpresheaf/Sieves.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ - import Mathlib.CategoryTheory.Subpresheaf.Basic import Mathlib.CategoryTheory.Sites.IsSheafFor diff --git a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean index 99b4e12ea64112..add552a2ee67aa 100644 --- a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean +++ b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean @@ -182,7 +182,8 @@ theorem ruzsa_triangle_inequality_mul_mul_mul (A B C : Finset G) : refine le_trans ?_ (mul_le_mul (hUA _ hB') (cast_le.2 <| card_le_card <| mul_subset_mul_right hU.2) (zero_le _) (zero_le _)) - #adaptation_note /-- 2024-11-01 `le_div_iff₀` is synthesizing wrong `GroupWithZero` without `@` -/ + #adaptation_note /-- 2024-11-01 + `le_div_iff₀` is synthesizing wrong `GroupWithZero` without `@` -/ rw [← mul_div_right_comm, ← mul_assoc, @le_div_iff₀ _ (_) _ _ _ _ _ _ _ (cast_pos.2 hU.1.card_pos)] exact mod_cast pluennecke_petridis_inequality_mul C (mul_aux hU.1 hU.2 hUA) diff --git a/Mathlib/Combinatorics/Enumerative/DyckWord.lean b/Mathlib/Combinatorics/Enumerative/DyckWord.lean index 5a8e966e55c59c..a48532e748a046 100644 --- a/Mathlib/Combinatorics/Enumerative/DyckWord.lean +++ b/Mathlib/Combinatorics/Enumerative/DyckWord.lean @@ -97,7 +97,7 @@ lemma toList_ne_nil : p.toList ≠ [] ↔ p ≠ 0 := toList_eq_nil.ne instance : Unique (AddUnits DyckWord) where uniq p := by obtain ⟨a, b, h, -⟩ := p - obtain ⟨ha, hb⟩ := append_eq_nil.mp (toList_eq_nil.mpr h) + obtain ⟨ha, hb⟩ := append_eq_nil_iff.mp (toList_eq_nil.mpr h) congr · exact toList_eq_nil.mp ha · exact toList_eq_nil.mp hb diff --git a/Mathlib/Combinatorics/SimpleGraph/Clique.lean b/Mathlib/Combinatorics/SimpleGraph/Clique.lean index 7738b767a8d3a8..433d517a5cba77 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Clique.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Clique.lean @@ -367,6 +367,10 @@ protected theorem CliqueFree.replaceVertex [DecidableEq α] (h : G.CliqueFree n) conv at hφ => enter [a, b]; rw [G.adj_replaceVertex_iff_of_ne _ (mt a) (mt b)] exact hφ +@[simp] +lemma cliqueFree_one : G.CliqueFree 1 ↔ IsEmpty α := by + simp [CliqueFree, isEmpty_iff] + @[simp] theorem cliqueFree_two : G.CliqueFree 2 ↔ G = ⊥ := by classical diff --git a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean index 54483becc09063..c0fa0618bf9ad3 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean @@ -137,6 +137,10 @@ theorem isEmpty_of_colorable_zero (h : G.Colorable 0) : IsEmpty V := by obtain ⟨i, hi⟩ := h.some v exact Nat.not_lt_zero _ hi +@[simp] +lemma colorable_zero_iff : G.Colorable 0 ↔ IsEmpty V := + ⟨G.isEmpty_of_colorable_zero, fun _ ↦ G.colorable_of_isEmpty 0⟩ + /-- The "tautological" coloring of a graph, using the vertices of the graph as colors. -/ def selfColoring : G.Coloring V := Coloring.mk id fun {_ _} => G.ne_of_adj diff --git a/Mathlib/Combinatorics/SimpleGraph/Matching.lean b/Mathlib/Combinatorics/SimpleGraph/Matching.lean index e98fe155b060bd..310d9aa562bd12 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Matching.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Matching.lean @@ -243,7 +243,7 @@ lemma IsPerfectMatching.induce_connectedComponent_isMatching (h : M.IsPerfectMat simpa [h.2.verts_eq_univ] using h.1.induce_connectedComponent c @[simp] -lemma IsPerfectMatching.toSubgraph_spanningCoe_iff (h : M.spanningCoe ≤ G') : +lemma IsPerfectMatching.toSubgraph_iff (h : M.spanningCoe ≤ G') : (G'.toSubgraph M.spanningCoe h).IsPerfectMatching ↔ M.IsPerfectMatching := by simp only [isPerfectMatching_iff, toSubgraph_adj, spanningCoe_adj] @@ -256,13 +256,11 @@ section Finite lemma even_card_of_isPerfectMatching [Fintype V] [DecidableEq V] [DecidableRel G.Adj] (c : ConnectedComponent G) (hM : M.IsPerfectMatching) : Even (Fintype.card c.supp) := by - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/5020, some instances that use the chain of coercions + #adaptation_note /-- https://github.com/leanprover/lean4/pull/5020 + some instances that use the chain of coercions `[SetLike X], X → Set α → Sort _` are blocked by the discrimination tree. This can be fixed by redeclaring the instance for `X` - using the double coercion but the proper fix seems to avoid the double coercion. - -/ + using the double coercion but the proper fix seems to avoid the double coercion. -/ letI : DecidablePred fun x ↦ x ∈ (M.induce c.supp).verts := fun a ↦ G.instDecidableMemSupp c a simpa using (hM.induce_connectedComponent_isMatching c).even_card @@ -339,9 +337,18 @@ lemma IsCycles.other_adj_of_adj (h : G.IsCycles) (hadj : G.Adj v w) : obtain ⟨w', hww'⟩ := (G.neighborSet v).exists_ne_of_one_lt_ncard (by omega) w exact ⟨w', ⟨hww'.2.symm, hww'.1⟩⟩ +lemma IsCycles.induce_supp (c : G.ConnectedComponent) (h : G.IsCycles) : + (G.induce c.supp).spanningCoe.IsCycles := by + intro v ⟨w, hw⟩ + rw [mem_neighborSet, c.adj_spanningCoe_induce_supp] at hw + rw [← h ⟨w, hw.2⟩] + congr + ext w' + simp only [mem_neighborSet, c.adj_spanningCoe_induce_supp, hw, true_and] + open scoped symmDiff -lemma Subgraph.IsPerfectMatching.symmDiff_spanningCoe_IsCycles +lemma Subgraph.IsPerfectMatching.symmDiff_isCycles {M : Subgraph G} {M' : Subgraph G'} (hM : M.IsPerfectMatching) (hM' : M'.IsPerfectMatching) : (M.spanningCoe ∆ M'.spanningCoe).IsCycles := by intro v @@ -368,10 +375,12 @@ for `SimpleGraph` rather than `SimpleGraph.Subgraph`. def IsAlternating (G G' : SimpleGraph V) := ∀ ⦃v w w': V⦄, w ≠ w' → G.Adj v w → G.Adj v w' → (G'.Adj v w ↔ ¬ G'.Adj v w') -lemma IsPerfectMatching.symmDiff_spanningCoe_of_isAlternating {M : Subgraph G} - (hM : M.IsPerfectMatching) (hG' : G'.IsAlternating M.spanningCoe) (hG'cyc : G'.IsCycles) : - (SimpleGraph.toSubgraph (M.spanningCoe ∆ G') - (by rfl)).IsPerfectMatching := by +lemma IsAlternating.mono {G'' : SimpleGraph V} (halt : G.IsAlternating G') (h : G'' ≤ G) : + G''.IsAlternating G' := fun _ _ _ hww' hvw hvw' ↦ halt hww' (h hvw) (h hvw') + +lemma IsPerfectMatching.symmDiff_of_isAlternating (hM : M.IsPerfectMatching) + (hG' : G'.IsAlternating M.spanningCoe) (hG'cyc : G'.IsCycles) : + (⊤ : Subgraph (M.spanningCoe ∆ G')).IsPerfectMatching := by rw [Subgraph.isPerfectMatching_iff] intro v simp only [toSubgraph_adj, symmDiff_def, sup_adj, sdiff_adj, Subgraph.spanningCoe_adj] @@ -380,20 +389,33 @@ lemma IsPerfectMatching.symmDiff_spanningCoe_of_isAlternating {M : Subgraph G} · obtain ⟨w', hw'⟩ := hG'cyc.other_adj_of_adj h have hmadj : M.Adj v w ↔ ¬M.Adj v w' := by simpa using hG' hw'.1 h hw'.2 use w' - simp only [hmadj.mp hw.1, hw'.2, not_true_eq_false, and_self, not_false_eq_true, or_true, - true_and] + simp only [Subgraph.top_adj, sup_adj, sdiff_adj, Subgraph.spanningCoe_adj, hmadj.mp hw.1, hw'.2, + not_true_eq_false, and_self, not_false_eq_true, or_true, true_and] rintro y (hl | hr) · aesop · obtain ⟨w'', hw''⟩ := hG'cyc.other_adj_of_adj hr.1 by_contra! hc - simp_all only [show M.Adj v y ↔ ¬M.Adj v w' from by simpa using hG' hc hr.1 hw'.2, - not_false_eq_true, ne_eq, iff_true, not_true_eq_false, and_false] + simp_all [show M.Adj v y ↔ ¬M.Adj v w' from by simpa using hG' hc hr.1 hw'.2] · use w - simp only [hw.1, h, not_false_eq_true, and_self, not_true_eq_false, or_false, true_and] + simp only [Subgraph.top_adj, sup_adj, sdiff_adj, Subgraph.spanningCoe_adj, hw.1, h, + not_false_eq_true, and_self, not_true_eq_false, or_false, true_and] rintro y (hl | hr) · exact hw.2 _ hl.1 · have ⟨w', hw'⟩ := hG'cyc.other_adj_of_adj hr.1 - simp_all only [show M.Adj v y ↔ ¬M.Adj v w' from by simpa using hG' hw'.1 hr.1 hw'.2, not_not, - ne_eq, and_false] + simp_all [show M.Adj v y ↔ ¬M.Adj v w' from by simpa using hG' hw'.1 hr.1 hw'.2] + +lemma IsPerfectMatching.isAlternating_symmDiff_left {M' : Subgraph G'} + (hM : M.IsPerfectMatching) (hM' : M'.IsPerfectMatching) : + (M.spanningCoe ∆ M'.spanningCoe).IsAlternating M.spanningCoe := by + intro v w w' hww' hvw hvw' + obtain ⟨v1, hm1, hv1⟩ := hM.1 (hM.2 v) + obtain ⟨v2, hm2, hv2⟩ := hM'.1 (hM'.2 v) + simp only [symmDiff_def] at * + aesop + +lemma IsPerfectMatching.isAlternating_symmDiff_right + {M' : Subgraph G'} (hM : M.IsPerfectMatching) (hM' : M'.IsPerfectMatching) : + (M.spanningCoe ∆ M'.spanningCoe).IsAlternating M'.spanningCoe := by + simpa [symmDiff_comm] using isAlternating_symmDiff_left hM' hM end SimpleGraph diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index 3211d0a9774fee..d016e6d94ad996 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -1071,6 +1071,22 @@ theorem mem_supp_iff (C : G.ConnectedComponent) (v : V) : v ∈ C.supp ↔ G.connectedComponentMk v = C := Iff.rfl +lemma mem_supp_congr_adj {v w : V} (c : G.ConnectedComponent) (hadj : G.Adj v w) : + v ∈ c.supp ↔ w ∈ c.supp := by + simp only [ConnectedComponent.mem_supp_iff] at * + constructor <;> intro h <;> simp only [← h] <;> apply connectedComponentMk_eq_of_adj + · exact hadj.symm + · exact hadj + +lemma adj_spanningCoe_induce_supp {v w : V} (c : G.ConnectedComponent) : + (G.induce c.supp).spanningCoe.Adj v w ↔ v ∈ c.supp ∧ G.Adj v w := by + by_cases h : v ∈ c.supp + · refine ⟨by aesop, ?_⟩ + intro h' + have : w ∈ c.supp := by rwa [c.mem_supp_congr_adj h'.2] at h + aesop + · aesop + theorem connectedComponentMk_mem {v : V} : v ∈ G.connectedComponentMk v := rfl diff --git a/Mathlib/Combinatorics/SimpleGraph/Sum.lean b/Mathlib/Combinatorics/SimpleGraph/Sum.lean index 25de075d43605d..aaf460a2f7c93e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Sum.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Sum.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Iván Renison -/ import Mathlib.Combinatorics.SimpleGraph.Basic +import Mathlib.Combinatorics.SimpleGraph.Coloring import Mathlib.Combinatorics.SimpleGraph.Maps /-! @@ -46,18 +47,12 @@ variable {G : SimpleGraph α} {H : SimpleGraph β} /-- The disjoint sum is commutative up to isomorphism. `Iso.sumComm` as a graph isomorphism. -/ @[simps!] def Iso.sumComm : G ⊕g H ≃g H ⊕g G := ⟨Equiv.sumComm α β, by - intro u v - cases u <;> cases v <;> simp⟩ + rintro (u | u) (v | v) <;> simp⟩ /-- The disjoint sum is associative up to isomorphism. `Iso.sumAssoc` as a graph isomorphism. -/ @[simps!] def Iso.sumAssoc {I : SimpleGraph γ} : (G ⊕g H) ⊕g I ≃g G ⊕g (H ⊕g I) := ⟨Equiv.sumAssoc α β γ, by - intro u v - cases u <;> cases v <;> rename_i u v - · cases u <;> cases v <;> simp - · cases u <;> simp - · cases v <;> simp - · simp⟩ + rintro ((u | u) | u) ((v | v) | v) <;> simp⟩ /-- The embedding of `G` into `G ⊕g H`. -/ @[simps] @@ -73,4 +68,71 @@ def Embedding.sumInr : H ↪g G ⊕g H where inj' u v := by simp map_rel_iff' := by simp +/-- Color `G ⊕g H` with colorings of `G` and `H` -/ +def Coloring.sum (cG : G.Coloring γ) (cH : H.Coloring γ) : (G ⊕g H).Coloring γ where + toFun := Sum.elim cG cH + map_rel' {u v} huv := by cases u <;> cases v <;> simp_all [cG.valid, cH.valid] + +/-- Get coloring of `G` from coloring of `G ⊕g H` -/ +def Coloring.sumLeft (c : (G ⊕g H).Coloring γ) : G.Coloring γ := c.comp Embedding.sumInl.toHom + +/-- Get coloring of `H` from coloring of `G ⊕g H` -/ +def Coloring.sumRight (c : (G ⊕g H).Coloring γ) : H.Coloring γ := c.comp Embedding.sumInr.toHom + +@[simp] +theorem Coloring.sumLeft_sum (cG : G.Coloring γ) (cH : H.Coloring γ) : (cG.sum cH).sumLeft = cG := + rfl + +@[simp] +theorem Coloring.sumRight_sum (cG : G.Coloring γ) (cH : H.Coloring γ) : (cG.sum cH).sumRight = cH := + rfl + +@[simp] +theorem Coloring.sum_sumLeft_sumRight (c : (G ⊕g H).Coloring γ) : c.sumLeft.sum c.sumRight = c := by + ext (u | u) <;> rfl + +/-- Bijection between `(G ⊕g H).Coloring γ` and `G.Coloring γ × H.Coloring γ` -/ +def Coloring.sumEquiv : (G ⊕g H).Coloring γ ≃ G.Coloring γ × H.Coloring γ where + toFun c := ⟨c.sumLeft, c.sumRight⟩ + invFun p := p.1.sum p.2 + left_inv c := by simp [sum_sumLeft_sumRight c] + right_inv p := rfl + +/-- Color `G ⊕g H` with `Fin (n + m)` given a coloring of `G` with `Fin n` and a coloring of `H` +with `Fin m` -/ +def Coloring.sumFin {n m : ℕ} (cG : G.Coloring (Fin n)) (cH : H.Coloring (Fin m)) : + (G ⊕g H).Coloring (Fin (max n m)) := sum + (G.recolorOfEmbedding (Fin.castLEEmb (n.le_max_left m)) cG) + (H.recolorOfEmbedding (Fin.castLEEmb (n.le_max_right m)) cH) + +theorem Colorable.sum_max {n m : ℕ} (hG : G.Colorable n) (hH : H.Colorable m) : + (G ⊕g H).Colorable (max n m) := Nonempty.intro (hG.some.sumFin hH.some) + +theorem Colorable.of_sum_left {n : ℕ} (h : (G ⊕g H).Colorable n) : G.Colorable n := + Nonempty.intro (h.some.sumLeft) + +theorem Colorable.of_sum_right {n : ℕ} (h : (G ⊕g H).Colorable n) : H.Colorable n := + Nonempty.intro (h.some.sumRight) + +@[simp] +theorem colorable_sum {n : ℕ} : (G ⊕g H).Colorable n ↔ G.Colorable n ∧ H.Colorable n := + ⟨fun cGH => ⟨cGH.of_sum_left, cGH.of_sum_right⟩, + fun ⟨cG, cH⟩ => by rw [← n.max_self]; exact cG.sum_max cH⟩ + +theorem chromaticNumber_le_sum_left : G.chromaticNumber ≤ (G ⊕g H).chromaticNumber := + chromaticNumber_le_of_forall_imp (fun _ h ↦ h.of_sum_left) + +theorem chromaticNumber_le_sum_right : H.chromaticNumber ≤ (G ⊕g H).chromaticNumber := + chromaticNumber_le_of_forall_imp (fun _ h ↦ h.of_sum_right) + +@[simp] +theorem chromaticNumber_sum : + (G ⊕g H).chromaticNumber = max G.chromaticNumber H.chromaticNumber := by + refine eq_max chromaticNumber_le_sum_left chromaticNumber_le_sum_right ?_ + rintro (n | n) hG hH + · simp [show (none : ℕ∞) = (⊤ : ℕ∞) from rfl] + · let cG : G.Coloring (Fin n) := (chromaticNumber_le_iff_colorable.mp hG).some + let cH : H.Coloring (Fin n) := (chromaticNumber_le_iff_colorable.mp hH).some + exact chromaticNumber_le_iff_colorable.mpr (Nonempty.intro (cG.sum cH)) + end SimpleGraph diff --git a/Mathlib/Computability/DFA.lean b/Mathlib/Computability/DFA.lean index e55cb83d9c17c5..dd5c1c6b0348b8 100644 --- a/Mathlib/Computability/DFA.lean +++ b/Mathlib/Computability/DFA.lean @@ -6,7 +6,6 @@ Authors: Fox Thomson, Chris Wong import Mathlib.Computability.Language import Mathlib.Data.Countable.Small import Mathlib.Data.Fintype.Card -import Mathlib.Data.List.Indexes import Mathlib.Tactic.NormNum /-! @@ -185,9 +184,9 @@ theorem comap_id : M.comap id = M := rfl @[simp] theorem evalFrom_comap (f : α' → α) (s : σ) (x : List α') : (M.comap f).evalFrom s x = M.evalFrom s (x.map f) := by - induction x using List.list_reverse_induction with - | base => simp - | ind x a ih => simp [ih] + induction x using List.reverseRecOn with + | nil => simp + | append_singleton x a ih => simp [ih] @[simp] theorem eval_comap (f : α' → α) (x : List α') : (M.comap f).eval x = M.eval (x.map f) := by @@ -226,9 +225,9 @@ theorem symm_reindex (g : σ ≃ σ') : (reindex (α := α) g).symm = reindex g. @[simp] theorem evalFrom_reindex (g : σ ≃ σ') (s : σ') (x : List α) : (reindex g M).evalFrom s x = g (M.evalFrom (g.symm s) x) := by - induction x using List.list_reverse_induction with - | base => simp - | ind x a ih => simp [ih] + induction x using List.reverseRecOn with + | nil => simp + | append_singleton x a ih => simp [ih] @[simp] theorem eval_reindex (g : σ ≃ σ') (x : List α) : (reindex g M).eval x = g (M.eval x) := by diff --git a/Mathlib/Computability/MyhillNerode.lean b/Mathlib/Computability/MyhillNerode.lean index 0ab520eb652381..9a9fc6bd140eee 100644 --- a/Mathlib/Computability/MyhillNerode.lean +++ b/Mathlib/Computability/MyhillNerode.lean @@ -85,9 +85,9 @@ theorem accepts_toDFA : L.toDFA.accepts = L := by ext x rw [DFA.mem_accepts] suffices L.toDFA.eval x = L.leftQuotient x by simp [this] - induction x using List.list_reverse_induction with - | base => simp - | ind x a ih => simp [ih, leftQuotient_append] + induction x using List.reverseRecOn with + | nil => simp + | append_singleton x a ih => simp [ih, leftQuotient_append] theorem IsRegular.of_finite_range_leftQuotient (h : Set.Finite (Set.range L.leftQuotient)) : L.IsRegular := diff --git a/Mathlib/Computability/PartrecCode.lean b/Mathlib/Computability/PartrecCode.lean index 4a04dbd781f9d2..fc3cd1f4049019 100644 --- a/Mathlib/Computability/PartrecCode.lean +++ b/Mathlib/Computability/PartrecCode.lean @@ -32,6 +32,7 @@ of some code. * `Nat.Partrec.Code.exists_code`: Partial recursiveness is equivalent to being the eval of a code. * `Nat.Partrec.Code.evaln_prim`: `evaln` is primitive recursive. * `Nat.Partrec.Code.fixed_point`: Roger's fixed point theorem. +* `Nat.Partrec.Code.fixed_point₂`: Kleene's second recursion theorem. ## References @@ -979,9 +980,9 @@ theorem eval_part : Partrec₂ eval := (evaln_prim.to_comp.comp ((Computable.snd.pair (fst.comp fst)).pair (snd.comp fst))).to₂).of_eq fun a => by simp [eval_eq_rfindOpt] -/-- Roger's fixed-point theorem: Any total, computable `f` has a fixed point: That is, under the -interpretation given by `Nat.Partrec.Code.eval`, there is a code `c` such that `c` and `f c` have -the same evaluation. +/-- **Roger's fixed-point theorem**: any total, computable `f` has a fixed point. +That is, under the interpretation given by `Nat.Partrec.Code.eval`, there is a code `c` +such that `c` and `f c` have the same evaluation. -/ theorem fixed_point {f : Code → Code} (hf : Computable f) : ∃ c : Code, eval (f c) = eval c := let g (x y : ℕ) : Part ℕ := eval (ofNat Code x) x >>= fun b => eval (ofNat Code b) y @@ -1000,6 +1001,7 @@ theorem fixed_point {f : Code → Code} (hf : Computable f) : ∃ c : Code, eval show eval (f (curry cg (encode cF))) n = eval (curry cg (encode cF)) n by simp [F, g, eg', eF', Part.map_id']⟩ +/-- **Kleene's second recursion theorem** -/ theorem fixed_point₂ {f : Code → ℕ →. ℕ} (hf : Partrec₂ f) : ∃ c : Code, eval c = f c := let ⟨cf, ef⟩ := exists_code.1 hf (fixed_point (curry_prim.comp (_root_.Primrec.const cf) Primrec.encode).to_comp).imp fun c e => diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index 20b710c3ba8449..a4c9369593bb28 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -333,7 +333,7 @@ theorem const (x : σ) : Primrec₂ fun (_ : α) (_ : β) => x := Primrec.const _ protected theorem pair : Primrec₂ (@Prod.mk α β) := - .pair .fst .snd + Primrec.pair .fst .snd theorem left : Primrec₂ fun (a : α) (_ : β) => a := .fst diff --git a/Mathlib/Computability/RegularExpressions.lean b/Mathlib/Computability/RegularExpressions.lean index 4e10a5bf7e78c1..a1922965cec30b 100644 --- a/Mathlib/Computability/RegularExpressions.lean +++ b/Mathlib/Computability/RegularExpressions.lean @@ -92,7 +92,7 @@ theorem comp_def (P Q : RegularExpression α) : comp P Q = P * Q := rfl -- This was renamed to `matches'` during the port of Lean 4 as `matches` is a reserved word. -#adaptation_note /-- around nightly-2024-02-25, +#adaptation_note /-- nightly-2024-02-25 we need to write `comp x y` in the pattern `comp P Q`, instead of `x * y`. -/ /-- `matches' P` provides a language which contains all strings that `P` matches -/ -- Porting note: was '@[simp] but removed based on @@ -136,7 +136,7 @@ theorem matches'_pow (P : RegularExpression α) : ∀ n : ℕ, (P ^ n).matches' theorem matches'_star (P : RegularExpression α) : P.star.matches' = P.matches'∗ := rfl -#adaptation_note /-- around nightly-2024-02-25, +#adaptation_note /-- nightly-2024-02-25 we need to write `comp x y` in the pattern `comp P Q`, instead of `x * y`. -/ /-- `matchEpsilon P` is true if and only if `P` matches the empty string -/ def matchEpsilon : RegularExpression α → Bool @@ -150,7 +150,7 @@ def matchEpsilon : RegularExpression α → Bool section DecidableEq variable [DecidableEq α] -#adaptation_note /-- around nightly-2024-02-25, +#adaptation_note /-- nightly-2024-02-25 we need to write `comp x y` in the pattern `comp P Q`, instead of `x * y`. -/ /-- `P.deriv a` matches `x` if `P` matches `a :: x`, the Brzozowski derivative of `P` with respect to `a` -/ @@ -231,7 +231,7 @@ theorem mul_rmatch_iff (P Q : RegularExpression α) (x : List α) : rw [rmatch, rmatch] rwa [Bool.and_eq_true_iff] at h · rintro ⟨t, u, h₁, h₂⟩ - cases' List.append_eq_nil.1 h₁.symm with ht hu + cases' List.append_eq_nil_iff.1 h₁.symm with ht hu subst ht subst hu repeat rw [rmatch] at h₂ @@ -343,7 +343,7 @@ instance (P : RegularExpression α) : DecidablePred (· ∈ P.matches') := fun _ end DecidableEq -#adaptation_note /-- around nightly-2024-02-25, +#adaptation_note /-- nightly-2024-02-25 we need to write `comp x y` in the pattern `comp P Q`, instead of `x * y`. -/ /-- Map the alphabet of a regular expression. -/ @[simp] @@ -361,7 +361,7 @@ protected theorem map_pow (f : α → β) (P : RegularExpression α) : | 0 => by unfold map; rfl | n + 1 => (congr_arg (· * map f P) (RegularExpression.map_pow f P n) :) -#adaptation_note /-- around nightly-2024-02-25, +#adaptation_note /-- nightly-2024-02-25 we need to write `comp x y` in the pattern `comp P Q`, instead of `x * y`. -/ @[simp] theorem map_id : ∀ P : RegularExpression α, P.map id = P @@ -372,7 +372,7 @@ theorem map_id : ∀ P : RegularExpression α, P.map id = P | comp R S => by simp_rw [map, map_id]; rfl | star R => by simp_rw [map, map_id] -#adaptation_note /-- around nightly-2024-02-25, +#adaptation_note /-- nightly-2024-02-25 we need to write `comp x y` in the pattern `comp P Q`, instead of `x * y`. -/ @[simp] theorem map_map (g : β → γ) (f : α → β) : ∀ P : RegularExpression α, (P.map f).map g = P.map (g ∘ f) @@ -383,7 +383,7 @@ theorem map_map (g : β → γ) (f : α → β) : ∀ P : RegularExpression α, | comp R S => by simp only [map, Function.comp_apply, map_map] | star R => by simp only [map, Function.comp_apply, map_map] -#adaptation_note /-- around nightly-2024-02-25, +#adaptation_note /-- nightly-2024-02-25 we need to write `comp x y` in the pattern `comp R S`, instead of `x * y` (and the `erw` was just `rw`). -/ /-- The language of the map is the map of the language. -/ diff --git a/Mathlib/Computability/Tape.lean b/Mathlib/Computability/Tape.lean index cf293a33770e00..920c5c776721a1 100644 --- a/Mathlib/Computability/Tape.lean +++ b/Mathlib/Computability/Tape.lean @@ -377,7 +377,7 @@ theorem ListBlank.append_assoc {Γ} [Inhabited Γ] (l₁ l₂ : List Γ) (l₃ : element is sent to a sequence of default elements. -/ def ListBlank.flatMap {Γ Γ'} [Inhabited Γ] [Inhabited Γ'] (l : ListBlank Γ) (f : Γ → List Γ') (hf : ∃ n, f default = List.replicate n default) : ListBlank Γ' := by - apply l.liftOn (fun l ↦ ListBlank.mk (List.flatMap l f)) + apply l.liftOn (fun l ↦ ListBlank.mk (l.flatMap f)) rintro l _ ⟨i, rfl⟩; cases' hf with n e; refine Quotient.sound' (Or.inl ⟨i * n, ?_⟩) rw [List.flatMap_append, mul_comm]; congr induction' i with i IH diff --git a/Mathlib/Control/Fold.lean b/Mathlib/Control/Fold.lean index 62e4e80dddadbb..fd6adaeacbf13e 100644 --- a/Mathlib/Control/Fold.lean +++ b/Mathlib/Control/Fold.lean @@ -118,10 +118,7 @@ def Foldl.ofFreeMonoid (f : β → α → β) : FreeMonoid α →* Monoid.Foldl map_one' := rfl map_mul' := by intros - #adaptation_note /-- nightly-2024-03-16: simp was - simp only [FreeMonoid.toList_mul, flip, unop_op, List.foldl_append, op_inj] -/ - simp only [FreeMonoid.toList_mul, unop_op, List.foldl_append, op_inj, Function.flip_def, - List.foldl_append] + simp only [FreeMonoid.toList_mul, unop_op, List.foldl_append, op_inj, Function.flip_def] rfl abbrev Foldr (α : Type u) : Type u := @@ -280,9 +277,6 @@ theorem foldr.ofFreeMonoid_comp_of (f : β → α → α) : theorem foldlm.ofFreeMonoid_comp_of {m} [Monad m] [LawfulMonad m] (f : α → β → m α) : foldlM.ofFreeMonoid f ∘ FreeMonoid.of = foldlM.mk ∘ flip f := by ext1 x - #adaptation_note /-- nightly-2024-03-16: simp was - simp only [foldlM.ofFreeMonoid, flip, MonoidHom.coe_mk, OneHom.coe_mk, Function.comp_apply, - FreeMonoid.toList_of, List.foldlM_cons, List.foldlM_nil, bind_pure, foldlM.mk, op_inj] -/ simp only [foldlM.ofFreeMonoid, Function.flip_def, MonoidHom.coe_mk, OneHom.coe_mk, Function.comp_apply, FreeMonoid.toList_of, List.foldlM_cons, List.foldlM_nil, bind_pure, foldlM.mk, op_inj] @@ -292,8 +286,6 @@ theorem foldlm.ofFreeMonoid_comp_of {m} [Monad m] [LawfulMonad m] (f : α → β theorem foldrm.ofFreeMonoid_comp_of {m} [Monad m] [LawfulMonad m] (f : β → α → m α) : foldrM.ofFreeMonoid f ∘ FreeMonoid.of = foldrM.mk ∘ f := by ext - #adaptation_note /-- nightly-2024-03-16: simp was - simp [(· ∘ ·), foldrM.ofFreeMonoid, foldrM.mk, flip] -/ simp [(· ∘ ·), foldrM.ofFreeMonoid, foldrM.mk, Function.flip_def] theorem toList_spec (xs : t α) : toList xs = FreeMonoid.toList (foldMap FreeMonoid.of xs) := @@ -306,8 +298,6 @@ theorem toList_spec (xs : t α) : toList xs = FreeMonoid.toList (foldMap FreeMon simp only [FreeMonoid.reverse_reverse, List.foldr_reverse, List.foldl_flip_cons_eq_append, List.append_nil, List.reverse_reverse] _ = (unop (Foldl.ofFreeMonoid (flip cons) (foldMap FreeMonoid.of xs)) []).reverse := by - #adaptation_note /-- nightly-2024-03-16: simp was - simp [flip, List.foldr_reverse, Foldl.ofFreeMonoid, unop_op] -/ simp [Function.flip_def, List.foldr_reverse, Foldl.ofFreeMonoid, unop_op] _ = toList xs := by rw [foldMap_hom_free (Foldl.ofFreeMonoid (flip <| @cons α))] @@ -337,8 +327,6 @@ theorem toList_map (f : α → β) (xs : t α) : toList (f <$> xs) = f <$> toLis @[simp] theorem foldl_map (g : β → γ) (f : α → γ → α) (a : α) (l : t β) : foldl f a (g <$> l) = foldl (fun x y => f x (g y)) a l := by - #adaptation_note /-- nightly-2024-03-16: simp was - simp only [foldl, foldMap_map, (· ∘ ·), flip] -/ simp only [foldl, foldMap_map, Function.comp_def, Function.flip_def] @[simp] @@ -382,8 +370,6 @@ theorem foldrm_toList (f : α → β → m β) (x : β) (xs : t α) : @[simp] theorem foldlm_map (g : β → γ) (f : α → γ → m α) (a : α) (l : t β) : foldlm f a (g <$> l) = foldlm (fun x y => f x (g y)) a l := by - #adaptation_note /-- nightly-2024-03-16: simp was - simp only [foldlm, foldMap_map, (· ∘ ·), flip] -/ simp only [foldlm, foldMap_map, Function.comp_def, Function.flip_def] @[simp] diff --git a/Mathlib/Data/Bool/Basic.lean b/Mathlib/Data/Bool/Basic.lean index 5f5bfc9ba28c85..0857104a5b71e5 100644 --- a/Mathlib/Data/Bool/Basic.lean +++ b/Mathlib/Data/Bool/Basic.lean @@ -49,8 +49,8 @@ theorem or_eq_true_eq_eq_true_or_eq_true (a b : Bool) : theorem not_eq_true_eq_eq_false (a : Bool) : (not a = true) = (a = false) := by cases a <;> simp -#adaptation_note /-- this is no longer a simp lemma, - as after nightly-2024-03-05 the LHS simplifies. -/ +#adaptation_note /-- nightly-2024-03-05 +this is no longer a simp lemma, as the LHS simplifies. -/ theorem and_eq_false_eq_eq_false_or_eq_false (a b : Bool) : ((a && b) = false) = (a = false ∨ b = false) := by cases a <;> cases b <;> simp diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 7531c38a54ff39..92b673fb0b86ab 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -291,10 +291,16 @@ theorem lim_abs (f : CauSeq ℂ Complex.abs) : lim (cauSeqAbs f) = Complex.abs ( let ⟨i, hi⟩ := equiv_lim f ε ε0 ⟨i, fun j hj => lt_of_le_of_lt (Complex.abs.abs_abv_sub_le_abv_sub _ _) (hi j hj)⟩ +lemma ne_zero_of_re_pos {s : ℂ} (hs : 0 < s.re) : s ≠ 0 := + fun h ↦ (zero_re ▸ h ▸ hs).false + lemma ne_zero_of_one_lt_re {s : ℂ} (hs : 1 < s.re) : s ≠ 0 := - fun h ↦ ((zero_re ▸ h ▸ hs).trans zero_lt_one).false + ne_zero_of_re_pos <| zero_lt_one.trans hs + +lemma re_neg_ne_zero_of_re_pos {s : ℂ} (hs : 0 < s.re) : (-s).re ≠ 0 := + ne_iff_lt_or_gt.mpr <| Or.inl <| neg_re s ▸ (neg_lt_zero.mpr hs) lemma re_neg_ne_zero_of_one_lt_re {s : ℂ} (hs : 1 < s.re) : (-s).re ≠ 0 := - ne_iff_lt_or_gt.mpr <| Or.inl <| neg_re s ▸ by linarith + re_neg_ne_zero_of_re_pos <| zero_lt_one.trans hs end Complex diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 8972434892eb7d..fd4dfde3db5b64 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -12,10 +12,9 @@ import Mathlib.Tactic.Bound.Attribute import Mathlib.Algebra.CharP.Defs /-! -# Exponential, trigonometric and hyperbolic trigonometric functions +# Exponential Function -This file contains the definitions of the real and complex exponential, sine, cosine, tangent, -hyperbolic sine, hyperbolic cosine, and hyperbolic tangent functions. +This file contains the definitions of the real and complex exponential function. -/ @@ -51,40 +50,6 @@ def exp' (z : ℂ) : CauSeq ℂ Complex.abs := def exp (z : ℂ) : ℂ := CauSeq.lim (exp' z) -/-- The complex sine function, defined via `exp` -/ -@[pp_nodot] -def sin (z : ℂ) : ℂ := - (exp (-z * I) - exp (z * I)) * I / 2 - -/-- The complex cosine function, defined via `exp` -/ -@[pp_nodot] -def cos (z : ℂ) : ℂ := - (exp (z * I) + exp (-z * I)) / 2 - -/-- The complex tangent function, defined as `sin z / cos z` -/ -@[pp_nodot] -def tan (z : ℂ) : ℂ := - sin z / cos z - -/-- The complex cotangent function, defined as `cos z / sin z` -/ -def cot (z : ℂ) : ℂ := - cos z / sin z - -/-- The complex hyperbolic sine function, defined via `exp` -/ -@[pp_nodot] -def sinh (z : ℂ) : ℂ := - (exp z - exp (-z)) / 2 - -/-- The complex hyperbolic cosine function, defined via `exp` -/ -@[pp_nodot] -def cosh (z : ℂ) : ℂ := - (exp z + exp (-z)) / 2 - -/-- The complex hyperbolic tangent function, defined as `sinh z / cosh z` -/ -@[pp_nodot] -def tanh (z : ℂ) : ℂ := - sinh z / cosh z - /-- scoped notation for the complex exponential function -/ scoped notation "cexp" => Complex.exp @@ -103,41 +68,6 @@ noncomputable section nonrec def exp (x : ℝ) : ℝ := (exp x).re -/-- The real sine function, defined as the real part of the complex sine -/ -@[pp_nodot] -nonrec def sin (x : ℝ) : ℝ := - (sin x).re - -/-- The real cosine function, defined as the real part of the complex cosine -/ -@[pp_nodot] -nonrec def cos (x : ℝ) : ℝ := - (cos x).re - -/-- The real tangent function, defined as the real part of the complex tangent -/ -@[pp_nodot] -nonrec def tan (x : ℝ) : ℝ := - (tan x).re - -/-- The real cotangent function, defined as the real part of the complex cotangent -/ -nonrec def cot (x : ℝ) : ℝ := - (cot x).re - -/-- The real hypebolic sine function, defined as the real part of the complex hyperbolic sine -/ -@[pp_nodot] -nonrec def sinh (x : ℝ) : ℝ := - (sinh x).re - -/-- The real hypebolic cosine function, defined as the real part of the complex hyperbolic cosine -/ -@[pp_nodot] -nonrec def cosh (x : ℝ) : ℝ := - (cosh x).re - -/-- The real hypebolic tangent function, defined as the real part of -the complex hyperbolic tangent -/ -@[pp_nodot] -nonrec def tanh (x : ℝ) : ℝ := - (tanh x).re - /-- scoped notation for the real exponential function -/ scoped notation "rexp" => Real.exp @@ -248,427 +178,6 @@ theorem exp_ofReal_im (x : ℝ) : (exp x).im = 0 := by rw [← ofReal_exp_ofReal theorem exp_ofReal_re (x : ℝ) : (exp x).re = Real.exp x := rfl -theorem two_sinh : 2 * sinh x = exp x - exp (-x) := - mul_div_cancel₀ _ two_ne_zero - -theorem two_cosh : 2 * cosh x = exp x + exp (-x) := - mul_div_cancel₀ _ two_ne_zero - -@[simp] -theorem sinh_zero : sinh 0 = 0 := by simp [sinh] - -@[simp] -theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul] - -private theorem sinh_add_aux {a b c d : ℂ} : - (a - b) * (c + d) + (a + b) * (c - d) = 2 * (a * c - b * d) := by ring - -theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by - rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← - mul_assoc, two_sinh, mul_left_comm, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, - mul_left_comm, two_cosh, ← mul_assoc, two_cosh] - exact sinh_add_aux - -@[simp] -theorem cosh_zero : cosh 0 = 1 := by simp [cosh] - -@[simp] -theorem cosh_neg : cosh (-x) = cosh x := by simp [add_comm, cosh, exp_neg] - -private theorem cosh_add_aux {a b c d : ℂ} : - (a + b) * (c + d) + (a - b) * (c - d) = 2 * (a * c + b * d) := by ring - -theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by - rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← - mul_assoc, two_cosh, ← mul_assoc, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, - mul_left_comm, two_cosh, mul_left_comm, two_sinh] - exact cosh_add_aux - -theorem sinh_sub : sinh (x - y) = sinh x * cosh y - cosh x * sinh y := by - simp [sub_eq_add_neg, sinh_add, sinh_neg, cosh_neg] - -theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by - simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg] - -theorem sinh_conj : sinh (conj x) = conj (sinh x) := by - rw [sinh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_sub, sinh, map_div₀] - -- Porting note: not nice - simp [← one_add_one_eq_two] - -@[simp] -theorem ofReal_sinh_ofReal_re (x : ℝ) : ((sinh x).re : ℂ) = sinh x := - conj_eq_iff_re.1 <| by rw [← sinh_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_sinh (x : ℝ) : (Real.sinh x : ℂ) = sinh x := - ofReal_sinh_ofReal_re _ - -@[simp] -theorem sinh_ofReal_im (x : ℝ) : (sinh x).im = 0 := by rw [← ofReal_sinh_ofReal_re, ofReal_im] - -theorem sinh_ofReal_re (x : ℝ) : (sinh x).re = Real.sinh x := - rfl - -theorem cosh_conj : cosh (conj x) = conj (cosh x) := by - rw [cosh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_add, cosh, map_div₀] - -- Porting note: not nice - simp [← one_add_one_eq_two] - -theorem ofReal_cosh_ofReal_re (x : ℝ) : ((cosh x).re : ℂ) = cosh x := - conj_eq_iff_re.1 <| by rw [← cosh_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_cosh (x : ℝ) : (Real.cosh x : ℂ) = cosh x := - ofReal_cosh_ofReal_re _ - -@[simp] -theorem cosh_ofReal_im (x : ℝ) : (cosh x).im = 0 := by rw [← ofReal_cosh_ofReal_re, ofReal_im] - -@[simp] -theorem cosh_ofReal_re (x : ℝ) : (cosh x).re = Real.cosh x := - rfl - -theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := - rfl - -@[simp] -theorem tanh_zero : tanh 0 = 0 := by simp [tanh] - -@[simp] -theorem tanh_neg : tanh (-x) = -tanh x := by simp [tanh, neg_div] - -theorem tanh_conj : tanh (conj x) = conj (tanh x) := by - rw [tanh, sinh_conj, cosh_conj, ← map_div₀, tanh] - -@[simp] -theorem ofReal_tanh_ofReal_re (x : ℝ) : ((tanh x).re : ℂ) = tanh x := - conj_eq_iff_re.1 <| by rw [← tanh_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_tanh (x : ℝ) : (Real.tanh x : ℂ) = tanh x := - ofReal_tanh_ofReal_re _ - -@[simp] -theorem tanh_ofReal_im (x : ℝ) : (tanh x).im = 0 := by rw [← ofReal_tanh_ofReal_re, ofReal_im] - -theorem tanh_ofReal_re (x : ℝ) : (tanh x).re = Real.tanh x := - rfl - -@[simp] -theorem cosh_add_sinh : cosh x + sinh x = exp x := by - rw [← mul_right_inj' (two_ne_zero' ℂ), mul_add, two_cosh, two_sinh, add_add_sub_cancel, two_mul] - -@[simp] -theorem sinh_add_cosh : sinh x + cosh x = exp x := by rw [add_comm, cosh_add_sinh] - -@[simp] -theorem exp_sub_cosh : exp x - cosh x = sinh x := - sub_eq_iff_eq_add.2 (sinh_add_cosh x).symm - -@[simp] -theorem exp_sub_sinh : exp x - sinh x = cosh x := - sub_eq_iff_eq_add.2 (cosh_add_sinh x).symm - -@[simp] -theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := by - rw [← mul_right_inj' (two_ne_zero' ℂ), mul_sub, two_cosh, two_sinh, add_sub_sub_cancel, two_mul] - -@[simp] -theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_sub_sinh] - -@[simp] -theorem cosh_sq_sub_sinh_sq : cosh x ^ 2 - sinh x ^ 2 = 1 := by - rw [sq_sub_sq, cosh_add_sinh, cosh_sub_sinh, ← exp_add, add_neg_cancel, exp_zero] - -theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by - rw [← cosh_sq_sub_sinh_sq x] - ring - -theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by - rw [← cosh_sq_sub_sinh_sq x] - ring - -theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by rw [two_mul, cosh_add, sq, sq] - -theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by - rw [two_mul, sinh_add] - ring - -theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by - have h1 : x + 2 * x = 3 * x := by ring - rw [← h1, cosh_add x (2 * x)] - simp only [cosh_two_mul, sinh_two_mul] - have h2 : sinh x * (2 * sinh x * cosh x) = 2 * cosh x * sinh x ^ 2 := by ring - rw [h2, sinh_sq] - ring - -theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by - have h1 : x + 2 * x = 3 * x := by ring - rw [← h1, sinh_add x (2 * x)] - simp only [cosh_two_mul, sinh_two_mul] - have h2 : cosh x * (2 * sinh x * cosh x) = 2 * sinh x * cosh x ^ 2 := by ring - rw [h2, cosh_sq] - ring - -@[simp] -theorem sin_zero : sin 0 = 0 := by simp [sin] - -@[simp] -theorem sin_neg : sin (-x) = -sin x := by - simp [sin, sub_eq_add_neg, exp_neg, (neg_div _ _).symm, add_mul] - -theorem two_sin : 2 * sin x = (exp (-x * I) - exp (x * I)) * I := - mul_div_cancel₀ _ two_ne_zero - -theorem two_cos : 2 * cos x = exp (x * I) + exp (-x * I) := - mul_div_cancel₀ _ two_ne_zero - -theorem sinh_mul_I : sinh (x * I) = sin x * I := by - rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, ← mul_assoc, two_sin, mul_assoc, I_mul_I, - mul_neg_one, neg_sub, neg_mul_eq_neg_mul] - -theorem cosh_mul_I : cosh (x * I) = cos x := by - rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, two_cos, neg_mul_eq_neg_mul] - -theorem tanh_mul_I : tanh (x * I) = tan x * I := by - rw [tanh_eq_sinh_div_cosh, cosh_mul_I, sinh_mul_I, mul_div_right_comm, tan] - -theorem cos_mul_I : cos (x * I) = cosh x := by rw [← cosh_mul_I]; ring_nf; simp - -theorem sin_mul_I : sin (x * I) = sinh x * I := by - have h : I * sin (x * I) = -sinh x := by - rw [mul_comm, ← sinh_mul_I] - ring_nf - simp - rw [← neg_neg (sinh x), ← h] - apply Complex.ext <;> simp - -theorem tan_mul_I : tan (x * I) = tanh x * I := by - rw [tan, sin_mul_I, cos_mul_I, mul_div_right_comm, tanh_eq_sinh_div_cosh] - -theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := by - rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, add_mul, add_mul, mul_right_comm, ← sinh_mul_I, - mul_assoc, ← sinh_mul_I, ← cosh_mul_I, ← cosh_mul_I, sinh_add] - -@[simp] -theorem cos_zero : cos 0 = 1 := by simp [cos] - -@[simp] -theorem cos_neg : cos (-x) = cos x := by simp [cos, sub_eq_add_neg, exp_neg, add_comm] - -theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := by - rw [← cosh_mul_I, add_mul, cosh_add, cosh_mul_I, cosh_mul_I, sinh_mul_I, sinh_mul_I, - mul_mul_mul_comm, I_mul_I, mul_neg_one, sub_eq_add_neg] - -theorem sin_sub : sin (x - y) = sin x * cos y - cos x * sin y := by - simp [sub_eq_add_neg, sin_add, sin_neg, cos_neg] - -theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by - simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] - -theorem sin_add_mul_I (x y : ℂ) : sin (x + y * I) = sin x * cosh y + cos x * sinh y * I := by - rw [sin_add, cos_mul_I, sin_mul_I, mul_assoc] - -theorem sin_eq (z : ℂ) : sin z = sin z.re * cosh z.im + cos z.re * sinh z.im * I := by - convert sin_add_mul_I z.re z.im; exact (re_add_im z).symm - -theorem cos_add_mul_I (x y : ℂ) : cos (x + y * I) = cos x * cosh y - sin x * sinh y * I := by - rw [cos_add, cos_mul_I, sin_mul_I, mul_assoc] - -theorem cos_eq (z : ℂ) : cos z = cos z.re * cosh z.im - sin z.re * sinh z.im * I := by - convert cos_add_mul_I z.re z.im; exact (re_add_im z).symm - -theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := by - have s1 := sin_add ((x + y) / 2) ((x - y) / 2) - have s2 := sin_sub ((x + y) / 2) ((x - y) / 2) - rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel_right, add_self_div_two] at s1 - rw [div_sub_div_same, ← sub_add, add_sub_cancel_left, add_self_div_two] at s2 - rw [s1, s2] - ring - -theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := by - have s1 := cos_add ((x + y) / 2) ((x - y) / 2) - have s2 := cos_sub ((x + y) / 2) ((x - y) / 2) - rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel_right, add_self_div_two] at s1 - rw [div_sub_div_same, ← sub_add, add_sub_cancel_left, add_self_div_two] at s2 - rw [s1, s2] - ring - -theorem sin_add_sin : sin x + sin y = 2 * sin ((x + y) / 2) * cos ((x - y) / 2) := by - simpa using sin_sub_sin x (-y) - -theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := by - calc - cos x + cos y = cos ((x + y) / 2 + (x - y) / 2) + cos ((x + y) / 2 - (x - y) / 2) := ?_ - _ = - cos ((x + y) / 2) * cos ((x - y) / 2) - sin ((x + y) / 2) * sin ((x - y) / 2) + - (cos ((x + y) / 2) * cos ((x - y) / 2) + sin ((x + y) / 2) * sin ((x - y) / 2)) := - ?_ - _ = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := ?_ - - · congr <;> field_simp - · rw [cos_add, cos_sub] - ring - -theorem sin_conj : sin (conj x) = conj (sin x) := by - rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, ← conj_neg_I, ← RingHom.map_mul, ← RingHom.map_mul, - sinh_conj, mul_neg, sinh_neg, sinh_mul_I, mul_neg] - -@[simp] -theorem ofReal_sin_ofReal_re (x : ℝ) : ((sin x).re : ℂ) = sin x := - conj_eq_iff_re.1 <| by rw [← sin_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_sin (x : ℝ) : (Real.sin x : ℂ) = sin x := - ofReal_sin_ofReal_re _ - -@[simp] -theorem sin_ofReal_im (x : ℝ) : (sin x).im = 0 := by rw [← ofReal_sin_ofReal_re, ofReal_im] - -theorem sin_ofReal_re (x : ℝ) : (sin x).re = Real.sin x := - rfl - -theorem cos_conj : cos (conj x) = conj (cos x) := by - rw [← cosh_mul_I, ← conj_neg_I, ← RingHom.map_mul, ← cosh_mul_I, cosh_conj, mul_neg, cosh_neg] - -@[simp] -theorem ofReal_cos_ofReal_re (x : ℝ) : ((cos x).re : ℂ) = cos x := - conj_eq_iff_re.1 <| by rw [← cos_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_cos (x : ℝ) : (Real.cos x : ℂ) = cos x := - ofReal_cos_ofReal_re _ - -@[simp] -theorem cos_ofReal_im (x : ℝ) : (cos x).im = 0 := by rw [← ofReal_cos_ofReal_re, ofReal_im] - -theorem cos_ofReal_re (x : ℝ) : (cos x).re = Real.cos x := - rfl - -@[simp] -theorem tan_zero : tan 0 = 0 := by simp [tan] - -theorem tan_eq_sin_div_cos : tan x = sin x / cos x := - rfl - -theorem cot_eq_cos_div_sin : cot x = cos x / sin x := - rfl - -theorem tan_mul_cos {x : ℂ} (hx : cos x ≠ 0) : tan x * cos x = sin x := by - rw [tan_eq_sin_div_cos, div_mul_cancel₀ _ hx] - -@[simp] -theorem tan_neg : tan (-x) = -tan x := by simp [tan, neg_div] - -theorem tan_conj : tan (conj x) = conj (tan x) := by rw [tan, sin_conj, cos_conj, ← map_div₀, tan] - -theorem cot_conj : cot (conj x) = conj (cot x) := by rw [cot, sin_conj, cos_conj, ← map_div₀, cot] - -@[simp] -theorem ofReal_tan_ofReal_re (x : ℝ) : ((tan x).re : ℂ) = tan x := - conj_eq_iff_re.1 <| by rw [← tan_conj, conj_ofReal] - -@[simp] -theorem ofReal_cot_ofReal_re (x : ℝ) : ((cot x).re : ℂ) = cot x := - conj_eq_iff_re.1 <| by rw [← cot_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_tan (x : ℝ) : (Real.tan x : ℂ) = tan x := - ofReal_tan_ofReal_re _ - -@[simp, norm_cast] -theorem ofReal_cot (x : ℝ) : (Real.cot x : ℂ) = cot x := - ofReal_cot_ofReal_re _ - -@[simp] -theorem tan_ofReal_im (x : ℝ) : (tan x).im = 0 := by rw [← ofReal_tan_ofReal_re, ofReal_im] - -theorem tan_ofReal_re (x : ℝ) : (tan x).re = Real.tan x := - rfl - -theorem cos_add_sin_I : cos x + sin x * I = exp (x * I) := by - rw [← cosh_add_sinh, sinh_mul_I, cosh_mul_I] - -theorem cos_sub_sin_I : cos x - sin x * I = exp (-x * I) := by - rw [neg_mul, ← cosh_sub_sinh, sinh_mul_I, cosh_mul_I] - -@[simp] -theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := - Eq.trans (by rw [cosh_mul_I, sinh_mul_I, mul_pow, I_sq, mul_neg_one, sub_neg_eq_add, add_comm]) - (cosh_sq_sub_sinh_sq (x * I)) - -@[simp] -theorem cos_sq_add_sin_sq : cos x ^ 2 + sin x ^ 2 = 1 := by rw [add_comm, sin_sq_add_cos_sq] - -theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := by rw [two_mul, cos_add, ← sq, ← sq] - -theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := by - rw [cos_two_mul', eq_sub_iff_add_eq.2 (sin_sq_add_cos_sq x), ← sub_add, sub_add_eq_add_sub, - two_mul] - -theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := by - rw [two_mul, sin_add, two_mul, add_mul, mul_comm] - -theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := by - simp [cos_two_mul, div_add_div_same, mul_div_cancel_left₀, two_ne_zero, -one_div] - -theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_left] - -theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_right] - -theorem inv_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := by - rw [tan_eq_sin_div_cos, div_pow] - field_simp - -theorem tan_sq_div_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : - tan x ^ 2 / (1 + tan x ^ 2) = sin x ^ 2 := by - simp only [← tan_mul_cos hx, mul_pow, ← inv_one_add_tan_sq hx, div_eq_mul_inv, one_mul] - -theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by - have h1 : x + 2 * x = 3 * x := by ring - rw [← h1, cos_add x (2 * x)] - simp only [cos_two_mul, sin_two_mul, mul_add, mul_sub, mul_one, sq] - have h2 : 4 * cos x ^ 3 = 2 * cos x * cos x * cos x + 2 * cos x * cos x ^ 2 := by ring - rw [h2, cos_sq'] - ring - -theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by - have h1 : x + 2 * x = 3 * x := by ring - rw [← h1, sin_add x (2 * x)] - simp only [cos_two_mul, sin_two_mul, cos_sq'] - have h2 : cos x * (2 * sin x * cos x) = 2 * sin x * cos x ^ 2 := by ring - rw [h2, cos_sq'] - ring - -theorem exp_mul_I : exp (x * I) = cos x + sin x * I := - (cos_add_sin_I _).symm - -theorem exp_add_mul_I : exp (x + y * I) = exp x * (cos y + sin y * I) := by rw [exp_add, exp_mul_I] - -theorem exp_eq_exp_re_mul_sin_add_cos : exp x = exp x.re * (cos x.im + sin x.im * I) := by - rw [← exp_add_mul_I, re_add_im] - -theorem exp_re : (exp x).re = Real.exp x.re * Real.cos x.im := by - rw [exp_eq_exp_re_mul_sin_add_cos] - simp [exp_ofReal_re, cos_ofReal_re] - -theorem exp_im : (exp x).im = Real.exp x.re * Real.sin x.im := by - rw [exp_eq_exp_re_mul_sin_add_cos] - simp [exp_ofReal_re, sin_ofReal_re] - -@[simp] -theorem exp_ofReal_mul_I_re (x : ℝ) : (exp (x * I)).re = Real.cos x := by - simp [exp_mul_I, cos_ofReal_re] - -@[simp] -theorem exp_ofReal_mul_I_im (x : ℝ) : (exp (x * I)).im = Real.sin x := by - simp [exp_mul_I, sin_ofReal_re] - -/-- **De Moivre's formula** -/ -theorem cos_add_sin_mul_I_pow (n : ℕ) (z : ℂ) : - (cos z + sin z * I) ^ n = cos (↑n * z) + sin (↑n * z) * I := by - rw [← exp_mul_I, ← exp_mul_I, ← exp_nat_mul, mul_assoc] - end Complex namespace Real @@ -716,241 +225,6 @@ nonrec theorem exp_neg : exp (-x) = (exp x)⁻¹ := theorem exp_sub : exp (x - y) = exp x / exp y := by simp [sub_eq_add_neg, exp_add, exp_neg, div_eq_mul_inv] -@[simp] -theorem sin_zero : sin 0 = 0 := by simp [sin] - -@[simp] -theorem sin_neg : sin (-x) = -sin x := by simp [sin, exp_neg, (neg_div _ _).symm, add_mul] - -nonrec theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := - ofReal_injective <| by simp [sin_add] - -@[simp] -theorem cos_zero : cos 0 = 1 := by simp [cos] - -@[simp] -theorem cos_neg : cos (-x) = cos x := by simp [cos, exp_neg] - -@[simp] -theorem cos_abs : cos |x| = cos x := by - cases le_total x 0 <;> simp only [*, _root_.abs_of_nonneg, abs_of_nonpos, cos_neg] - -nonrec theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := - ofReal_injective <| by simp [cos_add] - -theorem sin_sub : sin (x - y) = sin x * cos y - cos x * sin y := by - simp [sub_eq_add_neg, sin_add, sin_neg, cos_neg] - -theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by - simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] - -nonrec theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := - ofReal_injective <| by simp [sin_sub_sin] - -nonrec theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := - ofReal_injective <| by simp [cos_sub_cos] - -nonrec theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := - ofReal_injective <| by simp [cos_add_cos] - -theorem two_mul_sin_mul_sin (x y : ℝ) : 2 * sin x * sin y = cos (x - y) - cos (x + y) := by - simp [cos_add, cos_sub] - ring - -theorem two_mul_cos_mul_cos (x y : ℝ) : 2 * cos x * cos y = cos (x - y) + cos (x + y) := by - simp [cos_add, cos_sub] - ring - -theorem two_mul_sin_mul_cos (x y : ℝ) : 2 * sin x * cos y = sin (x - y) + sin (x + y) := by - simp [sin_add, sin_sub] - ring - -nonrec theorem tan_eq_sin_div_cos : tan x = sin x / cos x := - ofReal_injective <| by simp only [ofReal_tan, tan_eq_sin_div_cos, ofReal_div, ofReal_sin, - ofReal_cos] - -nonrec theorem cot_eq_cos_div_sin : cot x = cos x / sin x := - ofReal_injective <| by simp [cot_eq_cos_div_sin] - -theorem tan_mul_cos {x : ℝ} (hx : cos x ≠ 0) : tan x * cos x = sin x := by - rw [tan_eq_sin_div_cos, div_mul_cancel₀ _ hx] - -@[simp] -theorem tan_zero : tan 0 = 0 := by simp [tan] - -@[simp] -theorem tan_neg : tan (-x) = -tan x := by simp [tan, neg_div] - -@[simp] -nonrec theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := - ofReal_injective (by simp [sin_sq_add_cos_sq]) - -@[simp] -theorem cos_sq_add_sin_sq : cos x ^ 2 + sin x ^ 2 = 1 := by rw [add_comm, sin_sq_add_cos_sq] - -theorem sin_sq_le_one : sin x ^ 2 ≤ 1 := by - rw [← sin_sq_add_cos_sq x]; exact le_add_of_nonneg_right (sq_nonneg _) - -theorem cos_sq_le_one : cos x ^ 2 ≤ 1 := by - rw [← sin_sq_add_cos_sq x]; exact le_add_of_nonneg_left (sq_nonneg _) - -theorem abs_sin_le_one : |sin x| ≤ 1 := - abs_le_one_iff_mul_self_le_one.2 <| by simp only [← sq, sin_sq_le_one] - -theorem abs_cos_le_one : |cos x| ≤ 1 := - abs_le_one_iff_mul_self_le_one.2 <| by simp only [← sq, cos_sq_le_one] - -theorem sin_le_one : sin x ≤ 1 := - (abs_le.1 (abs_sin_le_one _)).2 - -theorem cos_le_one : cos x ≤ 1 := - (abs_le.1 (abs_cos_le_one _)).2 - -theorem neg_one_le_sin : -1 ≤ sin x := - (abs_le.1 (abs_sin_le_one _)).1 - -theorem neg_one_le_cos : -1 ≤ cos x := - (abs_le.1 (abs_cos_le_one _)).1 - -nonrec theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := - ofReal_injective <| by simp [cos_two_mul] - -nonrec theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := - ofReal_injective <| by simp [cos_two_mul'] - -nonrec theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := - ofReal_injective <| by simp [sin_two_mul] - -nonrec theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := - ofReal_injective <| by simp [cos_sq] - -theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_left] - -theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := - eq_sub_iff_add_eq.2 <| sin_sq_add_cos_sq _ - -lemma sin_sq_eq_half_sub : sin x ^ 2 = 1 / 2 - cos (2 * x) / 2 := by - rw [sin_sq, cos_sq, ← sub_sub, sub_half] - -theorem abs_sin_eq_sqrt_one_sub_cos_sq (x : ℝ) : |sin x| = √(1 - cos x ^ 2) := by - rw [← sin_sq, sqrt_sq_eq_abs] - -theorem abs_cos_eq_sqrt_one_sub_sin_sq (x : ℝ) : |cos x| = √(1 - sin x ^ 2) := by - rw [← cos_sq', sqrt_sq_eq_abs] - -theorem inv_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := - have : Complex.cos x ≠ 0 := mt (congr_arg re) hx - ofReal_inj.1 <| by simpa using Complex.inv_one_add_tan_sq this - -theorem tan_sq_div_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : - tan x ^ 2 / (1 + tan x ^ 2) = sin x ^ 2 := by - simp only [← tan_mul_cos hx, mul_pow, ← inv_one_add_tan_sq hx, div_eq_mul_inv, one_mul] - -theorem inv_sqrt_one_add_tan_sq {x : ℝ} (hx : 0 < cos x) : (√(1 + tan x ^ 2))⁻¹ = cos x := by - rw [← sqrt_sq hx.le, ← sqrt_inv, inv_one_add_tan_sq hx.ne'] - -theorem tan_div_sqrt_one_add_tan_sq {x : ℝ} (hx : 0 < cos x) : - tan x / √(1 + tan x ^ 2) = sin x := by - rw [← tan_mul_cos hx.ne', ← inv_sqrt_one_add_tan_sq hx, div_eq_mul_inv] - -nonrec theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by - rw [← ofReal_inj]; simp [cos_three_mul] - -nonrec theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by - rw [← ofReal_inj]; simp [sin_three_mul] - -/-- The definition of `sinh` in terms of `exp`. -/ -nonrec theorem sinh_eq (x : ℝ) : sinh x = (exp x - exp (-x)) / 2 := - ofReal_injective <| by simp [Complex.sinh] - -@[simp] -theorem sinh_zero : sinh 0 = 0 := by simp [sinh] - -@[simp] -theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul] - -nonrec theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by - rw [← ofReal_inj]; simp [sinh_add] - -/-- The definition of `cosh` in terms of `exp`. -/ -theorem cosh_eq (x : ℝ) : cosh x = (exp x + exp (-x)) / 2 := - eq_div_of_mul_eq two_ne_zero <| by - rw [cosh, exp, exp, Complex.ofReal_neg, Complex.cosh, mul_two, ← Complex.add_re, ← mul_two, - div_mul_cancel₀ _ (two_ne_zero' ℂ), Complex.add_re] - -@[simp] -theorem cosh_zero : cosh 0 = 1 := by simp [cosh] - -@[simp] -theorem cosh_neg : cosh (-x) = cosh x := - ofReal_inj.1 <| by simp - -@[simp] -theorem cosh_abs : cosh |x| = cosh x := by - cases le_total x 0 <;> simp [*, _root_.abs_of_nonneg, abs_of_nonpos] - -nonrec theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by - rw [← ofReal_inj]; simp [cosh_add] - -theorem sinh_sub : sinh (x - y) = sinh x * cosh y - cosh x * sinh y := by - simp [sub_eq_add_neg, sinh_add, sinh_neg, cosh_neg] - -theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by - simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg] - -nonrec theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := - ofReal_inj.1 <| by simp [tanh_eq_sinh_div_cosh] - -@[simp] -theorem tanh_zero : tanh 0 = 0 := by simp [tanh] - -@[simp] -theorem tanh_neg : tanh (-x) = -tanh x := by simp [tanh, neg_div] - -@[simp] -theorem cosh_add_sinh : cosh x + sinh x = exp x := by rw [← ofReal_inj]; simp - -@[simp] -theorem sinh_add_cosh : sinh x + cosh x = exp x := by rw [add_comm, cosh_add_sinh] - -@[simp] -theorem exp_sub_cosh : exp x - cosh x = sinh x := - sub_eq_iff_eq_add.2 (sinh_add_cosh x).symm - -@[simp] -theorem exp_sub_sinh : exp x - sinh x = cosh x := - sub_eq_iff_eq_add.2 (cosh_add_sinh x).symm - -@[simp] -theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := by - rw [← ofReal_inj] - simp - -@[simp] -theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_sub_sinh] - -@[simp] -theorem cosh_sq_sub_sinh_sq (x : ℝ) : cosh x ^ 2 - sinh x ^ 2 = 1 := by rw [← ofReal_inj]; simp - -nonrec theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by rw [← ofReal_inj]; simp [cosh_sq] - -theorem cosh_sq' : cosh x ^ 2 = 1 + sinh x ^ 2 := - (cosh_sq x).trans (add_comm _ _) - -nonrec theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by rw [← ofReal_inj]; simp [sinh_sq] - -nonrec theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by - rw [← ofReal_inj]; simp [cosh_two_mul] - -nonrec theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by - rw [← ofReal_inj]; simp [sinh_two_mul] - -nonrec theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by - rw [← ofReal_inj]; simp [cosh_three_mul] - -nonrec theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by - rw [← ofReal_inj]; simp [sinh_three_mul] - open IsAbsoluteValue Nat theorem sum_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) (n : ℕ) : ∑ i ∈ range n, x ^ i / i ! ≤ exp x := @@ -1055,13 +329,6 @@ theorem exp_le_one_iff {x : ℝ} : exp x ≤ 1 ↔ x ≤ 0 := theorem one_le_exp_iff {x : ℝ} : 1 ≤ exp x ↔ 0 ≤ x := exp_zero ▸ exp_le_exp -/-- `Real.cosh` is always positive -/ -theorem cosh_pos (x : ℝ) : 0 < Real.cosh x := - (cosh_eq x).symm ▸ half_pos (add_pos (exp_pos x) (exp_pos (-x))) - -theorem sinh_lt_cosh : sinh x < cosh x := - lt_of_pow_lt_pow_left₀ 2 (cosh_pos _).le <| (cosh_sq x).symm ▸ lt_add_one _ - end Real namespace Complex @@ -1334,125 +601,6 @@ theorem exp_1_approx_succ_eq {n} {a₁ b₁ : ℝ} {m : ℕ} (en : n + 1 = m) {r theorem exp_approx_start (x a b : ℝ) (h : |exp x - expNear 0 x a| ≤ |x| ^ 0 / Nat.factorial 0 * b) : |exp x - a| ≤ b := by simpa using h -theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| ^ 4 * (5 / 96) := - calc - |cos x - (1 - x ^ 2 / 2)| = Complex.abs (Complex.cos x - (1 - (x : ℂ) ^ 2 / 2)) := by - rw [← abs_ofReal]; simp - _ = Complex.abs ((Complex.exp (x * I) + Complex.exp (-x * I) - (2 - (x : ℂ) ^ 2)) / 2) := by - simp [Complex.cos, sub_div, add_div, neg_div, div_self (two_ne_zero' ℂ)] - _ = abs - (((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) + - (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial)) / 2) := - (congr_arg Complex.abs - (congr_arg (fun x : ℂ => x / 2) - (by - simp only [sum_range_succ, neg_mul, pow_succ, pow_zero, mul_one, range_zero, sum_empty, - Nat.factorial, Nat.cast_one, ne_eq, one_ne_zero, not_false_eq_true, div_self, - zero_add, div_one, Nat.mul_one, Nat.cast_succ, Nat.cast_mul, Nat.cast_ofNat, mul_neg, - neg_neg] - apply Complex.ext <;> simp [div_eq_mul_inv, normSq] <;> ring_nf - ))) - _ ≤ abs ((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2) + - abs ((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2) := by - rw [add_div]; exact Complex.abs.add_le _ _ - _ = abs (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2 + - abs (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2 := by - simp [map_div₀] - _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ) : ℝ)⁻¹) / 2 + - Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ) : ℝ)⁻¹) / 2 := by - gcongr - · exact Complex.exp_bound (by simpa) (by decide) - · exact Complex.exp_bound (by simpa) (by decide) - _ ≤ |x| ^ 4 * (5 / 96) := by norm_num [Nat.factorial] - -theorem sin_bound {x : ℝ} (hx : |x| ≤ 1) : |sin x - (x - x ^ 3 / 6)| ≤ |x| ^ 4 * (5 / 96) := - calc - |sin x - (x - x ^ 3 / 6)| = Complex.abs (Complex.sin x - (x - x ^ 3 / 6 : ℝ)) := by - rw [← abs_ofReal]; simp - _ = Complex.abs (((Complex.exp (-x * I) - Complex.exp (x * I)) * I - - (2 * x - x ^ 3 / 3 : ℝ)) / 2) := by - simp [Complex.sin, sub_div, add_div, neg_div, mul_div_cancel_left₀ _ (two_ne_zero' ℂ), - div_div, show (3 : ℂ) * 2 = 6 by norm_num] - _ = Complex.abs (((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) - - (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial)) * I / 2) := - (congr_arg Complex.abs - (congr_arg (fun x : ℂ => x / 2) - (by - simp only [sum_range_succ, neg_mul, pow_succ, pow_zero, mul_one, ofReal_sub, ofReal_mul, - ofReal_ofNat, ofReal_div, range_zero, sum_empty, Nat.factorial, Nat.cast_one, ne_eq, - one_ne_zero, not_false_eq_true, div_self, zero_add, div_one, mul_neg, neg_neg, - Nat.mul_one, Nat.cast_succ, Nat.cast_mul, Nat.cast_ofNat] - apply Complex.ext <;> simp [div_eq_mul_inv, normSq]; ring))) - _ ≤ abs ((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) * I / 2) + - abs (-((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) * I) / 2) := by - rw [sub_mul, sub_eq_add_neg, add_div]; exact Complex.abs.add_le _ _ - _ = abs (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2 + - abs (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2 := by - simp [add_comm, map_div₀] - _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * (Nat.factorial 4 * (4 : ℕ) : ℝ)⁻¹) / 2 + - Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * (Nat.factorial 4 * (4 : ℕ) : ℝ)⁻¹) / 2 := by - gcongr - · exact Complex.exp_bound (by simpa) (by decide) - · exact Complex.exp_bound (by simpa) (by decide) - _ ≤ |x| ^ 4 * (5 / 96) := by norm_num [Nat.factorial] - -theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := - calc 0 < 1 - x ^ 2 / 2 - |x| ^ 4 * (5 / 96) := - sub_pos.2 <| - lt_sub_iff_add_lt.2 - (calc - |x| ^ 4 * (5 / 96) + x ^ 2 / 2 ≤ 1 * (5 / 96) + 1 / 2 := by - gcongr - · exact pow_le_one₀ (abs_nonneg _) hx - · rw [sq, ← abs_mul_self, abs_mul] - exact mul_le_one₀ hx (abs_nonneg _) hx - _ < 1 := by norm_num) - _ ≤ cos x := sub_le_comm.1 (abs_sub_le_iff.1 (cos_bound hx)).2 - -theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < sin x := - calc 0 < x - x ^ 3 / 6 - |x| ^ 4 * (5 / 96) := - sub_pos.2 <| lt_sub_iff_add_lt.2 - (calc - |x| ^ 4 * (5 / 96) + x ^ 3 / 6 ≤ x * (5 / 96) + x / 6 := by - gcongr - · calc - |x| ^ 4 ≤ |x| ^ 1 := - pow_le_pow_of_le_one (abs_nonneg _) - (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]) (by decide) - _ = x := by simp [_root_.abs_of_nonneg (le_of_lt hx0)] - · calc - x ^ 3 ≤ x ^ 1 := pow_le_pow_of_le_one (le_of_lt hx0) hx (by decide) - _ = x := pow_one _ - _ < x := by linarith) - _ ≤ sin x := - sub_le_comm.1 (abs_sub_le_iff.1 (sin_bound (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]))).2 - -theorem sin_pos_of_pos_of_le_two {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 2) : 0 < sin x := - have : x / 2 ≤ 1 := (div_le_iff₀ (by norm_num)).mpr (by simpa) - calc - 0 < 2 * sin (x / 2) * cos (x / 2) := - mul_pos (mul_pos (by norm_num) (sin_pos_of_pos_of_le_one (half_pos hx0) this)) - (cos_pos_of_le_one (by rwa [_root_.abs_of_nonneg (le_of_lt (half_pos hx0))])) - _ = sin x := by rw [← sin_two_mul, two_mul, add_halves] - -theorem cos_one_le : cos 1 ≤ 2 / 3 := - calc - cos 1 ≤ |(1 : ℝ)| ^ 4 * (5 / 96) + (1 - 1 ^ 2 / 2) := - sub_le_iff_le_add.1 (abs_sub_le_iff.1 (cos_bound (by simp))).1 - _ ≤ 2 / 3 := by norm_num - -theorem cos_one_pos : 0 < cos 1 := - cos_pos_of_le_one (le_of_eq abs_one) - -theorem cos_two_neg : cos 2 < 0 := - calc cos 2 = cos (2 * 1) := congr_arg cos (mul_one _).symm - _ = _ := Real.cos_two_mul 1 - _ ≤ 2 * (2 / 3) ^ 2 - 1 := by - gcongr - · exact cos_one_pos.le - · apply cos_one_le - _ < 0 := by norm_num - theorem exp_bound_div_one_sub_of_interval' {x : ℝ} (h1 : 0 < x) (h2 : x < 1) : Real.exp x < 1 / (1 - x) := by have H : 0 < 1 - (1 + x + x ^ 2) * (1 - x) := calc @@ -1527,41 +675,13 @@ def evalExp : PositivityExt where eval {u α} _ _ e := do pure (.positive q(Real.exp_pos $a)) | _, _, _ => throwError "not Real.exp" -/-- Extension for the `positivity` tactic: `Real.cosh` is always positive. -/ -@[positivity Real.cosh _] -def evalCosh : PositivityExt where eval {u α} _ _ e := do - match u, α, e with - | 0, ~q(ℝ), ~q(Real.cosh $a) => - assertInstancesCommute - return .positive q(Real.cosh_pos $a) - | _, _, _ => throwError "not Real.cosh" - -example (x : ℝ) : 0 < x.cosh := by positivity - end Mathlib.Meta.Positivity namespace Complex -@[simp] -theorem abs_cos_add_sin_mul_I (x : ℝ) : abs (cos x + sin x * I) = 1 := by - have := Real.sin_sq_add_cos_sq x - simp_all [add_comm, abs, normSq, sq, sin_ofReal_re, cos_ofReal_re, mul_re] - @[simp] theorem abs_exp_ofReal (x : ℝ) : abs (exp x) = Real.exp x := by rw [← ofReal_exp] exact abs_of_nonneg (le_of_lt (Real.exp_pos _)) -@[simp] -theorem abs_exp_ofReal_mul_I (x : ℝ) : abs (exp (x * I)) = 1 := by - rw [exp_mul_I, abs_cos_add_sin_mul_I] - -theorem abs_exp (z : ℂ) : abs (exp z) = Real.exp z.re := by - rw [exp_eq_exp_re_mul_sin_add_cos, map_mul, abs_exp_ofReal, abs_cos_add_sin_mul_I, mul_one] - -theorem abs_exp_eq_iff_re_eq {x y : ℂ} : abs (exp x) = abs (exp y) ↔ x.re = y.re := by - rw [abs_exp, abs_exp, Real.exp_eq_exp] - end Complex - -set_option linter.style.longFile 1700 diff --git a/Mathlib/Data/Complex/Trigonometric.lean b/Mathlib/Data/Complex/Trigonometric.lean new file mode 100644 index 00000000000000..ea6b53a5110384 --- /dev/null +++ b/Mathlib/Data/Complex/Trigonometric.lean @@ -0,0 +1,951 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Abhimanyu Pallavi Sudhir +-/ +import Mathlib.Data.Complex.Exponential + +/-! +# Trigonometric and hyperbolic trigonometric functions + +This file contains the definitions of the sine, cosine, tangent, +hyperbolic sine, hyperbolic cosine, and hyperbolic tangent functions. + +-/ + +open CauSeq Finset IsAbsoluteValue +open scoped ComplexConjugate + +namespace Complex + +noncomputable section + +/-- The complex sine function, defined via `exp` -/ +@[pp_nodot] +def sin (z : ℂ) : ℂ := + (exp (-z * I) - exp (z * I)) * I / 2 + +/-- The complex cosine function, defined via `exp` -/ +@[pp_nodot] +def cos (z : ℂ) : ℂ := + (exp (z * I) + exp (-z * I)) / 2 + +/-- The complex tangent function, defined as `sin z / cos z` -/ +@[pp_nodot] +def tan (z : ℂ) : ℂ := + sin z / cos z + +/-- The complex cotangent function, defined as `cos z / sin z` -/ +def cot (z : ℂ) : ℂ := + cos z / sin z + +/-- The complex hyperbolic sine function, defined via `exp` -/ +@[pp_nodot] +def sinh (z : ℂ) : ℂ := + (exp z - exp (-z)) / 2 + +/-- The complex hyperbolic cosine function, defined via `exp` -/ +@[pp_nodot] +def cosh (z : ℂ) : ℂ := + (exp z + exp (-z)) / 2 + +/-- The complex hyperbolic tangent function, defined as `sinh z / cosh z` -/ +@[pp_nodot] +def tanh (z : ℂ) : ℂ := + sinh z / cosh z + +end + +end Complex + +namespace Real + +open Complex + +noncomputable section + +/-- The real sine function, defined as the real part of the complex sine -/ +@[pp_nodot] +nonrec def sin (x : ℝ) : ℝ := + (sin x).re + +/-- The real cosine function, defined as the real part of the complex cosine -/ +@[pp_nodot] +nonrec def cos (x : ℝ) : ℝ := + (cos x).re + +/-- The real tangent function, defined as the real part of the complex tangent -/ +@[pp_nodot] +nonrec def tan (x : ℝ) : ℝ := + (tan x).re + +/-- The real cotangent function, defined as the real part of the complex cotangent -/ +nonrec def cot (x : ℝ) : ℝ := + (cot x).re + +/-- The real hypebolic sine function, defined as the real part of the complex hyperbolic sine -/ +@[pp_nodot] +nonrec def sinh (x : ℝ) : ℝ := + (sinh x).re + +/-- The real hypebolic cosine function, defined as the real part of the complex hyperbolic cosine -/ +@[pp_nodot] +nonrec def cosh (x : ℝ) : ℝ := + (cosh x).re + +/-- The real hypebolic tangent function, defined as the real part of +the complex hyperbolic tangent -/ +@[pp_nodot] +nonrec def tanh (x : ℝ) : ℝ := + (tanh x).re + +end + +end Real + +namespace Complex + +variable (x y : ℂ) + +theorem two_sinh : 2 * sinh x = exp x - exp (-x) := + mul_div_cancel₀ _ two_ne_zero + +theorem two_cosh : 2 * cosh x = exp x + exp (-x) := + mul_div_cancel₀ _ two_ne_zero + +@[simp] +theorem sinh_zero : sinh 0 = 0 := by simp [sinh] + +@[simp] +theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul] + +private theorem sinh_add_aux {a b c d : ℂ} : + (a - b) * (c + d) + (a + b) * (c - d) = 2 * (a * c - b * d) := by ring + +theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← + mul_assoc, two_sinh, mul_left_comm, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, + mul_left_comm, two_cosh, ← mul_assoc, two_cosh] + exact sinh_add_aux + +@[simp] +theorem cosh_zero : cosh 0 = 1 := by simp [cosh] + +@[simp] +theorem cosh_neg : cosh (-x) = cosh x := by simp [add_comm, cosh, exp_neg] + +private theorem cosh_add_aux {a b c d : ℂ} : + (a + b) * (c + d) + (a - b) * (c - d) = 2 * (a * c + b * d) := by ring + +theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← + mul_assoc, two_cosh, ← mul_assoc, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, + mul_left_comm, two_cosh, mul_left_comm, two_sinh] + exact cosh_add_aux + +theorem sinh_sub : sinh (x - y) = sinh x * cosh y - cosh x * sinh y := by + simp [sub_eq_add_neg, sinh_add, sinh_neg, cosh_neg] + +theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by + simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg] + +theorem sinh_conj : sinh (conj x) = conj (sinh x) := by + rw [sinh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_sub, sinh, map_div₀] + -- Porting note: not nice + simp [← one_add_one_eq_two] + +@[simp] +theorem ofReal_sinh_ofReal_re (x : ℝ) : ((sinh x).re : ℂ) = sinh x := + conj_eq_iff_re.1 <| by rw [← sinh_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_sinh (x : ℝ) : (Real.sinh x : ℂ) = sinh x := + ofReal_sinh_ofReal_re _ + +@[simp] +theorem sinh_ofReal_im (x : ℝ) : (sinh x).im = 0 := by rw [← ofReal_sinh_ofReal_re, ofReal_im] + +theorem sinh_ofReal_re (x : ℝ) : (sinh x).re = Real.sinh x := + rfl + +theorem cosh_conj : cosh (conj x) = conj (cosh x) := by + rw [cosh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_add, cosh, map_div₀] + -- Porting note: not nice + simp [← one_add_one_eq_two] + +theorem ofReal_cosh_ofReal_re (x : ℝ) : ((cosh x).re : ℂ) = cosh x := + conj_eq_iff_re.1 <| by rw [← cosh_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_cosh (x : ℝ) : (Real.cosh x : ℂ) = cosh x := + ofReal_cosh_ofReal_re _ + +@[simp] +theorem cosh_ofReal_im (x : ℝ) : (cosh x).im = 0 := by rw [← ofReal_cosh_ofReal_re, ofReal_im] + +@[simp] +theorem cosh_ofReal_re (x : ℝ) : (cosh x).re = Real.cosh x := + rfl + +theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := + rfl + +@[simp] +theorem tanh_zero : tanh 0 = 0 := by simp [tanh] + +@[simp] +theorem tanh_neg : tanh (-x) = -tanh x := by simp [tanh, neg_div] + +theorem tanh_conj : tanh (conj x) = conj (tanh x) := by + rw [tanh, sinh_conj, cosh_conj, ← map_div₀, tanh] + +@[simp] +theorem ofReal_tanh_ofReal_re (x : ℝ) : ((tanh x).re : ℂ) = tanh x := + conj_eq_iff_re.1 <| by rw [← tanh_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_tanh (x : ℝ) : (Real.tanh x : ℂ) = tanh x := + ofReal_tanh_ofReal_re _ + +@[simp] +theorem tanh_ofReal_im (x : ℝ) : (tanh x).im = 0 := by rw [← ofReal_tanh_ofReal_re, ofReal_im] + +theorem tanh_ofReal_re (x : ℝ) : (tanh x).re = Real.tanh x := + rfl + +@[simp] +theorem cosh_add_sinh : cosh x + sinh x = exp x := by + rw [← mul_right_inj' (two_ne_zero' ℂ), mul_add, two_cosh, two_sinh, add_add_sub_cancel, two_mul] + +@[simp] +theorem sinh_add_cosh : sinh x + cosh x = exp x := by rw [add_comm, cosh_add_sinh] + +@[simp] +theorem exp_sub_cosh : exp x - cosh x = sinh x := + sub_eq_iff_eq_add.2 (sinh_add_cosh x).symm + +@[simp] +theorem exp_sub_sinh : exp x - sinh x = cosh x := + sub_eq_iff_eq_add.2 (cosh_add_sinh x).symm + +@[simp] +theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := by + rw [← mul_right_inj' (two_ne_zero' ℂ), mul_sub, two_cosh, two_sinh, add_sub_sub_cancel, two_mul] + +@[simp] +theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_sub_sinh] + +@[simp] +theorem cosh_sq_sub_sinh_sq : cosh x ^ 2 - sinh x ^ 2 = 1 := by + rw [sq_sub_sq, cosh_add_sinh, cosh_sub_sinh, ← exp_add, add_neg_cancel, exp_zero] + +theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by + rw [← cosh_sq_sub_sinh_sq x] + ring + +theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by + rw [← cosh_sq_sub_sinh_sq x] + ring + +theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by rw [two_mul, cosh_add, sq, sq] + +theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by + rw [two_mul, sinh_add] + ring + +theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, cosh_add x (2 * x)] + simp only [cosh_two_mul, sinh_two_mul] + have h2 : sinh x * (2 * sinh x * cosh x) = 2 * cosh x * sinh x ^ 2 := by ring + rw [h2, sinh_sq] + ring + +theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, sinh_add x (2 * x)] + simp only [cosh_two_mul, sinh_two_mul] + have h2 : cosh x * (2 * sinh x * cosh x) = 2 * sinh x * cosh x ^ 2 := by ring + rw [h2, cosh_sq] + ring + +@[simp] +theorem sin_zero : sin 0 = 0 := by simp [sin] + +@[simp] +theorem sin_neg : sin (-x) = -sin x := by + simp [sin, sub_eq_add_neg, exp_neg, (neg_div _ _).symm, add_mul] + +theorem two_sin : 2 * sin x = (exp (-x * I) - exp (x * I)) * I := + mul_div_cancel₀ _ two_ne_zero + +theorem two_cos : 2 * cos x = exp (x * I) + exp (-x * I) := + mul_div_cancel₀ _ two_ne_zero + +theorem sinh_mul_I : sinh (x * I) = sin x * I := by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, ← mul_assoc, two_sin, mul_assoc, I_mul_I, + mul_neg_one, neg_sub, neg_mul_eq_neg_mul] + +theorem cosh_mul_I : cosh (x * I) = cos x := by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, two_cos, neg_mul_eq_neg_mul] + +theorem tanh_mul_I : tanh (x * I) = tan x * I := by + rw [tanh_eq_sinh_div_cosh, cosh_mul_I, sinh_mul_I, mul_div_right_comm, tan] + +theorem cos_mul_I : cos (x * I) = cosh x := by rw [← cosh_mul_I]; ring_nf; simp + +theorem sin_mul_I : sin (x * I) = sinh x * I := by + have h : I * sin (x * I) = -sinh x := by + rw [mul_comm, ← sinh_mul_I] + ring_nf + simp + rw [← neg_neg (sinh x), ← h] + apply Complex.ext <;> simp + +theorem tan_mul_I : tan (x * I) = tanh x * I := by + rw [tan, sin_mul_I, cos_mul_I, mul_div_right_comm, tanh_eq_sinh_div_cosh] + +theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := by + rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, add_mul, add_mul, mul_right_comm, ← sinh_mul_I, + mul_assoc, ← sinh_mul_I, ← cosh_mul_I, ← cosh_mul_I, sinh_add] + +@[simp] +theorem cos_zero : cos 0 = 1 := by simp [cos] + +@[simp] +theorem cos_neg : cos (-x) = cos x := by simp [cos, sub_eq_add_neg, exp_neg, add_comm] + +theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := by + rw [← cosh_mul_I, add_mul, cosh_add, cosh_mul_I, cosh_mul_I, sinh_mul_I, sinh_mul_I, + mul_mul_mul_comm, I_mul_I, mul_neg_one, sub_eq_add_neg] + +theorem sin_sub : sin (x - y) = sin x * cos y - cos x * sin y := by + simp [sub_eq_add_neg, sin_add, sin_neg, cos_neg] + +theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by + simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] + +theorem sin_add_mul_I (x y : ℂ) : sin (x + y * I) = sin x * cosh y + cos x * sinh y * I := by + rw [sin_add, cos_mul_I, sin_mul_I, mul_assoc] + +theorem sin_eq (z : ℂ) : sin z = sin z.re * cosh z.im + cos z.re * sinh z.im * I := by + convert sin_add_mul_I z.re z.im; exact (re_add_im z).symm + +theorem cos_add_mul_I (x y : ℂ) : cos (x + y * I) = cos x * cosh y - sin x * sinh y * I := by + rw [cos_add, cos_mul_I, sin_mul_I, mul_assoc] + +theorem cos_eq (z : ℂ) : cos z = cos z.re * cosh z.im - sin z.re * sinh z.im * I := by + convert cos_add_mul_I z.re z.im; exact (re_add_im z).symm + +theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := by + have s1 := sin_add ((x + y) / 2) ((x - y) / 2) + have s2 := sin_sub ((x + y) / 2) ((x - y) / 2) + rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel_right, add_self_div_two] at s1 + rw [div_sub_div_same, ← sub_add, add_sub_cancel_left, add_self_div_two] at s2 + rw [s1, s2] + ring + +theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := by + have s1 := cos_add ((x + y) / 2) ((x - y) / 2) + have s2 := cos_sub ((x + y) / 2) ((x - y) / 2) + rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel_right, add_self_div_two] at s1 + rw [div_sub_div_same, ← sub_add, add_sub_cancel_left, add_self_div_two] at s2 + rw [s1, s2] + ring + +theorem sin_add_sin : sin x + sin y = 2 * sin ((x + y) / 2) * cos ((x - y) / 2) := by + simpa using sin_sub_sin x (-y) + +theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := by + calc + cos x + cos y = cos ((x + y) / 2 + (x - y) / 2) + cos ((x + y) / 2 - (x - y) / 2) := ?_ + _ = + cos ((x + y) / 2) * cos ((x - y) / 2) - sin ((x + y) / 2) * sin ((x - y) / 2) + + (cos ((x + y) / 2) * cos ((x - y) / 2) + sin ((x + y) / 2) * sin ((x - y) / 2)) := + ?_ + _ = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := ?_ + + · congr <;> field_simp + · rw [cos_add, cos_sub] + ring + +theorem sin_conj : sin (conj x) = conj (sin x) := by + rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, ← conj_neg_I, ← RingHom.map_mul, ← RingHom.map_mul, + sinh_conj, mul_neg, sinh_neg, sinh_mul_I, mul_neg] + +@[simp] +theorem ofReal_sin_ofReal_re (x : ℝ) : ((sin x).re : ℂ) = sin x := + conj_eq_iff_re.1 <| by rw [← sin_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_sin (x : ℝ) : (Real.sin x : ℂ) = sin x := + ofReal_sin_ofReal_re _ + +@[simp] +theorem sin_ofReal_im (x : ℝ) : (sin x).im = 0 := by rw [← ofReal_sin_ofReal_re, ofReal_im] + +theorem sin_ofReal_re (x : ℝ) : (sin x).re = Real.sin x := + rfl + +theorem cos_conj : cos (conj x) = conj (cos x) := by + rw [← cosh_mul_I, ← conj_neg_I, ← RingHom.map_mul, ← cosh_mul_I, cosh_conj, mul_neg, cosh_neg] + +@[simp] +theorem ofReal_cos_ofReal_re (x : ℝ) : ((cos x).re : ℂ) = cos x := + conj_eq_iff_re.1 <| by rw [← cos_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_cos (x : ℝ) : (Real.cos x : ℂ) = cos x := + ofReal_cos_ofReal_re _ + +@[simp] +theorem cos_ofReal_im (x : ℝ) : (cos x).im = 0 := by rw [← ofReal_cos_ofReal_re, ofReal_im] + +theorem cos_ofReal_re (x : ℝ) : (cos x).re = Real.cos x := + rfl + +@[simp] +theorem tan_zero : tan 0 = 0 := by simp [tan] + +theorem tan_eq_sin_div_cos : tan x = sin x / cos x := + rfl + +theorem cot_eq_cos_div_sin : cot x = cos x / sin x := + rfl + +theorem tan_mul_cos {x : ℂ} (hx : cos x ≠ 0) : tan x * cos x = sin x := by + rw [tan_eq_sin_div_cos, div_mul_cancel₀ _ hx] + +@[simp] +theorem tan_neg : tan (-x) = -tan x := by simp [tan, neg_div] + +theorem tan_conj : tan (conj x) = conj (tan x) := by rw [tan, sin_conj, cos_conj, ← map_div₀, tan] + +theorem cot_conj : cot (conj x) = conj (cot x) := by rw [cot, sin_conj, cos_conj, ← map_div₀, cot] + +@[simp] +theorem ofReal_tan_ofReal_re (x : ℝ) : ((tan x).re : ℂ) = tan x := + conj_eq_iff_re.1 <| by rw [← tan_conj, conj_ofReal] + +@[simp] +theorem ofReal_cot_ofReal_re (x : ℝ) : ((cot x).re : ℂ) = cot x := + conj_eq_iff_re.1 <| by rw [← cot_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_tan (x : ℝ) : (Real.tan x : ℂ) = tan x := + ofReal_tan_ofReal_re _ + +@[simp, norm_cast] +theorem ofReal_cot (x : ℝ) : (Real.cot x : ℂ) = cot x := + ofReal_cot_ofReal_re _ + +@[simp] +theorem tan_ofReal_im (x : ℝ) : (tan x).im = 0 := by rw [← ofReal_tan_ofReal_re, ofReal_im] + +theorem tan_ofReal_re (x : ℝ) : (tan x).re = Real.tan x := + rfl + +theorem cos_add_sin_I : cos x + sin x * I = exp (x * I) := by + rw [← cosh_add_sinh, sinh_mul_I, cosh_mul_I] + +theorem cos_sub_sin_I : cos x - sin x * I = exp (-x * I) := by + rw [neg_mul, ← cosh_sub_sinh, sinh_mul_I, cosh_mul_I] + +@[simp] +theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := + Eq.trans (by rw [cosh_mul_I, sinh_mul_I, mul_pow, I_sq, mul_neg_one, sub_neg_eq_add, add_comm]) + (cosh_sq_sub_sinh_sq (x * I)) + +@[simp] +theorem cos_sq_add_sin_sq : cos x ^ 2 + sin x ^ 2 = 1 := by rw [add_comm, sin_sq_add_cos_sq] + +theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := by rw [two_mul, cos_add, ← sq, ← sq] + +theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := by + rw [cos_two_mul', eq_sub_iff_add_eq.2 (sin_sq_add_cos_sq x), ← sub_add, sub_add_eq_add_sub, + two_mul] + +theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := by + rw [two_mul, sin_add, two_mul, add_mul, mul_comm] + +theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := by + simp [cos_two_mul, div_add_div_same, mul_div_cancel_left₀, two_ne_zero, -one_div] + +theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_left] + +theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_right] + +theorem inv_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := by + rw [tan_eq_sin_div_cos, div_pow] + field_simp + +theorem tan_sq_div_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : + tan x ^ 2 / (1 + tan x ^ 2) = sin x ^ 2 := by + simp only [← tan_mul_cos hx, mul_pow, ← inv_one_add_tan_sq hx, div_eq_mul_inv, one_mul] + +theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, cos_add x (2 * x)] + simp only [cos_two_mul, sin_two_mul, mul_add, mul_sub, mul_one, sq] + have h2 : 4 * cos x ^ 3 = 2 * cos x * cos x * cos x + 2 * cos x * cos x ^ 2 := by ring + rw [h2, cos_sq'] + ring + +theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, sin_add x (2 * x)] + simp only [cos_two_mul, sin_two_mul, cos_sq'] + have h2 : cos x * (2 * sin x * cos x) = 2 * sin x * cos x ^ 2 := by ring + rw [h2, cos_sq'] + ring + +theorem exp_mul_I : exp (x * I) = cos x + sin x * I := + (cos_add_sin_I _).symm + +theorem exp_add_mul_I : exp (x + y * I) = exp x * (cos y + sin y * I) := by rw [exp_add, exp_mul_I] + +theorem exp_eq_exp_re_mul_sin_add_cos : exp x = exp x.re * (cos x.im + sin x.im * I) := by + rw [← exp_add_mul_I, re_add_im] + +theorem exp_re : (exp x).re = Real.exp x.re * Real.cos x.im := by + rw [exp_eq_exp_re_mul_sin_add_cos] + simp [exp_ofReal_re, cos_ofReal_re] + +theorem exp_im : (exp x).im = Real.exp x.re * Real.sin x.im := by + rw [exp_eq_exp_re_mul_sin_add_cos] + simp [exp_ofReal_re, sin_ofReal_re] + +@[simp] +theorem exp_ofReal_mul_I_re (x : ℝ) : (exp (x * I)).re = Real.cos x := by + simp [exp_mul_I, cos_ofReal_re] + +@[simp] +theorem exp_ofReal_mul_I_im (x : ℝ) : (exp (x * I)).im = Real.sin x := by + simp [exp_mul_I, sin_ofReal_re] + +/-- **De Moivre's formula** -/ +theorem cos_add_sin_mul_I_pow (n : ℕ) (z : ℂ) : + (cos z + sin z * I) ^ n = cos (↑n * z) + sin (↑n * z) * I := by + rw [← exp_mul_I, ← exp_mul_I, ← exp_nat_mul, mul_assoc] + +end Complex + +namespace Real + +open Complex + +variable (x y : ℝ) + +@[simp] +theorem sin_zero : sin 0 = 0 := by simp [sin] + +@[simp] +theorem sin_neg : sin (-x) = -sin x := by simp [sin, exp_neg, (neg_div _ _).symm, add_mul] + +nonrec theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := + ofReal_injective <| by simp [sin_add] + +@[simp] +theorem cos_zero : cos 0 = 1 := by simp [cos] + +@[simp] +theorem cos_neg : cos (-x) = cos x := by simp [cos, exp_neg] + +@[simp] +theorem cos_abs : cos |x| = cos x := by + cases le_total x 0 <;> simp only [*, _root_.abs_of_nonneg, abs_of_nonpos, cos_neg] + +nonrec theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := + ofReal_injective <| by simp [cos_add] + +theorem sin_sub : sin (x - y) = sin x * cos y - cos x * sin y := by + simp [sub_eq_add_neg, sin_add, sin_neg, cos_neg] + +theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by + simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] + +nonrec theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := + ofReal_injective <| by simp [sin_sub_sin] + +nonrec theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := + ofReal_injective <| by simp [cos_sub_cos] + +nonrec theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := + ofReal_injective <| by simp [cos_add_cos] + +theorem two_mul_sin_mul_sin (x y : ℝ) : 2 * sin x * sin y = cos (x - y) - cos (x + y) := by + simp [cos_add, cos_sub] + ring + +theorem two_mul_cos_mul_cos (x y : ℝ) : 2 * cos x * cos y = cos (x - y) + cos (x + y) := by + simp [cos_add, cos_sub] + ring + +theorem two_mul_sin_mul_cos (x y : ℝ) : 2 * sin x * cos y = sin (x - y) + sin (x + y) := by + simp [sin_add, sin_sub] + ring + +nonrec theorem tan_eq_sin_div_cos : tan x = sin x / cos x := + ofReal_injective <| by simp only [ofReal_tan, tan_eq_sin_div_cos, ofReal_div, ofReal_sin, + ofReal_cos] + +nonrec theorem cot_eq_cos_div_sin : cot x = cos x / sin x := + ofReal_injective <| by simp [cot_eq_cos_div_sin] + +theorem tan_mul_cos {x : ℝ} (hx : cos x ≠ 0) : tan x * cos x = sin x := by + rw [tan_eq_sin_div_cos, div_mul_cancel₀ _ hx] + +@[simp] +theorem tan_zero : tan 0 = 0 := by simp [tan] + +@[simp] +theorem tan_neg : tan (-x) = -tan x := by simp [tan, neg_div] + +@[simp] +nonrec theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := + ofReal_injective (by simp [sin_sq_add_cos_sq]) + +@[simp] +theorem cos_sq_add_sin_sq : cos x ^ 2 + sin x ^ 2 = 1 := by rw [add_comm, sin_sq_add_cos_sq] + +theorem sin_sq_le_one : sin x ^ 2 ≤ 1 := by + rw [← sin_sq_add_cos_sq x]; exact le_add_of_nonneg_right (sq_nonneg _) + +theorem cos_sq_le_one : cos x ^ 2 ≤ 1 := by + rw [← sin_sq_add_cos_sq x]; exact le_add_of_nonneg_left (sq_nonneg _) + +theorem abs_sin_le_one : |sin x| ≤ 1 := + abs_le_one_iff_mul_self_le_one.2 <| by simp only [← sq, sin_sq_le_one] + +theorem abs_cos_le_one : |cos x| ≤ 1 := + abs_le_one_iff_mul_self_le_one.2 <| by simp only [← sq, cos_sq_le_one] + +theorem sin_le_one : sin x ≤ 1 := + (abs_le.1 (abs_sin_le_one _)).2 + +theorem cos_le_one : cos x ≤ 1 := + (abs_le.1 (abs_cos_le_one _)).2 + +theorem neg_one_le_sin : -1 ≤ sin x := + (abs_le.1 (abs_sin_le_one _)).1 + +theorem neg_one_le_cos : -1 ≤ cos x := + (abs_le.1 (abs_cos_le_one _)).1 + +nonrec theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := + ofReal_injective <| by simp [cos_two_mul] + +nonrec theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := + ofReal_injective <| by simp [cos_two_mul'] + +nonrec theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := + ofReal_injective <| by simp [sin_two_mul] + +nonrec theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := + ofReal_injective <| by simp [cos_sq] + +theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_left] + +theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := + eq_sub_iff_add_eq.2 <| sin_sq_add_cos_sq _ + +lemma sin_sq_eq_half_sub : sin x ^ 2 = 1 / 2 - cos (2 * x) / 2 := by + rw [sin_sq, cos_sq, ← sub_sub, sub_half] + +theorem abs_sin_eq_sqrt_one_sub_cos_sq (x : ℝ) : |sin x| = √(1 - cos x ^ 2) := by + rw [← sin_sq, sqrt_sq_eq_abs] + +theorem abs_cos_eq_sqrt_one_sub_sin_sq (x : ℝ) : |cos x| = √(1 - sin x ^ 2) := by + rw [← cos_sq', sqrt_sq_eq_abs] + +theorem inv_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := + have : Complex.cos x ≠ 0 := mt (congr_arg re) hx + ofReal_inj.1 <| by simpa using Complex.inv_one_add_tan_sq this + +theorem tan_sq_div_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : + tan x ^ 2 / (1 + tan x ^ 2) = sin x ^ 2 := by + simp only [← tan_mul_cos hx, mul_pow, ← inv_one_add_tan_sq hx, div_eq_mul_inv, one_mul] + +theorem inv_sqrt_one_add_tan_sq {x : ℝ} (hx : 0 < cos x) : (√(1 + tan x ^ 2))⁻¹ = cos x := by + rw [← sqrt_sq hx.le, ← sqrt_inv, inv_one_add_tan_sq hx.ne'] + +theorem tan_div_sqrt_one_add_tan_sq {x : ℝ} (hx : 0 < cos x) : + tan x / √(1 + tan x ^ 2) = sin x := by + rw [← tan_mul_cos hx.ne', ← inv_sqrt_one_add_tan_sq hx, div_eq_mul_inv] + +nonrec theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by + rw [← ofReal_inj]; simp [cos_three_mul] + +nonrec theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by + rw [← ofReal_inj]; simp [sin_three_mul] + +/-- The definition of `sinh` in terms of `exp`. -/ +nonrec theorem sinh_eq (x : ℝ) : sinh x = (exp x - exp (-x)) / 2 := + ofReal_injective <| by simp [Complex.sinh] + +@[simp] +theorem sinh_zero : sinh 0 = 0 := by simp [sinh] + +@[simp] +theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul] + +nonrec theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by + rw [← ofReal_inj]; simp [sinh_add] + +/-- The definition of `cosh` in terms of `exp`. -/ +theorem cosh_eq (x : ℝ) : cosh x = (exp x + exp (-x)) / 2 := + eq_div_of_mul_eq two_ne_zero <| by + rw [cosh, exp, exp, Complex.ofReal_neg, Complex.cosh, mul_two, ← Complex.add_re, ← mul_two, + div_mul_cancel₀ _ (two_ne_zero' ℂ), Complex.add_re] + +@[simp] +theorem cosh_zero : cosh 0 = 1 := by simp [cosh] + +@[simp] +theorem cosh_neg : cosh (-x) = cosh x := + ofReal_inj.1 <| by simp + +@[simp] +theorem cosh_abs : cosh |x| = cosh x := by + cases le_total x 0 <;> simp [*, _root_.abs_of_nonneg, abs_of_nonpos] + +nonrec theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by + rw [← ofReal_inj]; simp [cosh_add] + +theorem sinh_sub : sinh (x - y) = sinh x * cosh y - cosh x * sinh y := by + simp [sub_eq_add_neg, sinh_add, sinh_neg, cosh_neg] + +theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by + simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg] + +nonrec theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := + ofReal_inj.1 <| by simp [tanh_eq_sinh_div_cosh] + +@[simp] +theorem tanh_zero : tanh 0 = 0 := by simp [tanh] + +@[simp] +theorem tanh_neg : tanh (-x) = -tanh x := by simp [tanh, neg_div] + +@[simp] +theorem cosh_add_sinh : cosh x + sinh x = exp x := by rw [← ofReal_inj]; simp + +@[simp] +theorem sinh_add_cosh : sinh x + cosh x = exp x := by rw [add_comm, cosh_add_sinh] + +@[simp] +theorem exp_sub_cosh : exp x - cosh x = sinh x := + sub_eq_iff_eq_add.2 (sinh_add_cosh x).symm + +@[simp] +theorem exp_sub_sinh : exp x - sinh x = cosh x := + sub_eq_iff_eq_add.2 (cosh_add_sinh x).symm + +@[simp] +theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := by + rw [← ofReal_inj] + simp + +@[simp] +theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_sub_sinh] + +@[simp] +theorem cosh_sq_sub_sinh_sq (x : ℝ) : cosh x ^ 2 - sinh x ^ 2 = 1 := by rw [← ofReal_inj]; simp + +nonrec theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by rw [← ofReal_inj]; simp [cosh_sq] + +theorem cosh_sq' : cosh x ^ 2 = 1 + sinh x ^ 2 := + (cosh_sq x).trans (add_comm _ _) + +nonrec theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by rw [← ofReal_inj]; simp [sinh_sq] + +nonrec theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by + rw [← ofReal_inj]; simp [cosh_two_mul] + +nonrec theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by + rw [← ofReal_inj]; simp [sinh_two_mul] + +nonrec theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by + rw [← ofReal_inj]; simp [cosh_three_mul] + +nonrec theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by + rw [← ofReal_inj]; simp [sinh_three_mul] + +open IsAbsoluteValue Nat + +private theorem add_one_lt_exp_of_pos {x : ℝ} (hx : 0 < x) : x + 1 < exp x := + (by nlinarith : x + 1 < 1 + x + x ^ 2 / 2).trans_le (quadratic_le_exp_of_nonneg hx.le) + +private theorem add_one_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : x + 1 ≤ exp x := by + rcases eq_or_lt_of_le hx with (rfl | h) + · simp + exact (add_one_lt_exp_of_pos h).le + +/-- `Real.cosh` is always positive -/ +theorem cosh_pos (x : ℝ) : 0 < Real.cosh x := + (cosh_eq x).symm ▸ half_pos (add_pos (exp_pos x) (exp_pos (-x))) + +theorem sinh_lt_cosh : sinh x < cosh x := + lt_of_pow_lt_pow_left₀ 2 (cosh_pos _).le <| (cosh_sq x).symm ▸ lt_add_one _ + +end Real + +namespace Real + +open Complex Finset + +theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| ^ 4 * (5 / 96) := + calc + |cos x - (1 - x ^ 2 / 2)| = Complex.abs (Complex.cos x - (1 - (x : ℂ) ^ 2 / 2)) := by + rw [← abs_ofReal]; simp + _ = Complex.abs ((Complex.exp (x * I) + Complex.exp (-x * I) - (2 - (x : ℂ) ^ 2)) / 2) := by + simp [Complex.cos, sub_div, add_div, neg_div, div_self (two_ne_zero' ℂ)] + _ = abs + (((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) + + (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial)) / 2) := + (congr_arg Complex.abs + (congr_arg (fun x : ℂ => x / 2) + (by + simp only [sum_range_succ, neg_mul, pow_succ, pow_zero, mul_one, range_zero, sum_empty, + Nat.factorial, Nat.cast_one, ne_eq, one_ne_zero, not_false_eq_true, div_self, + zero_add, div_one, Nat.mul_one, Nat.cast_succ, Nat.cast_mul, Nat.cast_ofNat, mul_neg, + neg_neg] + apply Complex.ext <;> simp [div_eq_mul_inv, normSq] <;> ring_nf + ))) + _ ≤ abs ((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2) + + abs ((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2) := by + rw [add_div]; exact Complex.abs.add_le _ _ + _ = abs (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2 + + abs (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2 := by + simp [map_div₀] + _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ) : ℝ)⁻¹) / 2 + + Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ) : ℝ)⁻¹) / 2 := by + gcongr + · exact Complex.exp_bound (by simpa) (by decide) + · exact Complex.exp_bound (by simpa) (by decide) + _ ≤ |x| ^ 4 * (5 / 96) := by norm_num [Nat.factorial] + +theorem sin_bound {x : ℝ} (hx : |x| ≤ 1) : |sin x - (x - x ^ 3 / 6)| ≤ |x| ^ 4 * (5 / 96) := + calc + |sin x - (x - x ^ 3 / 6)| = Complex.abs (Complex.sin x - (x - x ^ 3 / 6 : ℝ)) := by + rw [← abs_ofReal]; simp + _ = Complex.abs (((Complex.exp (-x * I) - Complex.exp (x * I)) * I - + (2 * x - x ^ 3 / 3 : ℝ)) / 2) := by + simp [Complex.sin, sub_div, add_div, neg_div, mul_div_cancel_left₀ _ (two_ne_zero' ℂ), + div_div, show (3 : ℂ) * 2 = 6 by norm_num] + _ = Complex.abs (((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) - + (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial)) * I / 2) := + (congr_arg Complex.abs + (congr_arg (fun x : ℂ => x / 2) + (by + simp only [sum_range_succ, neg_mul, pow_succ, pow_zero, mul_one, ofReal_sub, ofReal_mul, + ofReal_ofNat, ofReal_div, range_zero, sum_empty, Nat.factorial, Nat.cast_one, ne_eq, + one_ne_zero, not_false_eq_true, div_self, zero_add, div_one, mul_neg, neg_neg, + Nat.mul_one, Nat.cast_succ, Nat.cast_mul, Nat.cast_ofNat] + apply Complex.ext <;> simp [div_eq_mul_inv, normSq]; ring))) + _ ≤ abs ((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) * I / 2) + + abs (-((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) * I) / 2) := by + rw [sub_mul, sub_eq_add_neg, add_div]; exact Complex.abs.add_le _ _ + _ = abs (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2 + + abs (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2 := by + simp [add_comm, map_div₀] + _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * (Nat.factorial 4 * (4 : ℕ) : ℝ)⁻¹) / 2 + + Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * (Nat.factorial 4 * (4 : ℕ) : ℝ)⁻¹) / 2 := by + gcongr + · exact Complex.exp_bound (by simpa) (by decide) + · exact Complex.exp_bound (by simpa) (by decide) + _ ≤ |x| ^ 4 * (5 / 96) := by norm_num [Nat.factorial] + +theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := + calc 0 < 1 - x ^ 2 / 2 - |x| ^ 4 * (5 / 96) := + sub_pos.2 <| + lt_sub_iff_add_lt.2 + (calc + |x| ^ 4 * (5 / 96) + x ^ 2 / 2 ≤ 1 * (5 / 96) + 1 / 2 := by + gcongr + · exact pow_le_one₀ (abs_nonneg _) hx + · rw [sq, ← abs_mul_self, abs_mul] + exact mul_le_one₀ hx (abs_nonneg _) hx + _ < 1 := by norm_num) + _ ≤ cos x := sub_le_comm.1 (abs_sub_le_iff.1 (cos_bound hx)).2 + +theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < sin x := + calc 0 < x - x ^ 3 / 6 - |x| ^ 4 * (5 / 96) := + sub_pos.2 <| lt_sub_iff_add_lt.2 + (calc + |x| ^ 4 * (5 / 96) + x ^ 3 / 6 ≤ x * (5 / 96) + x / 6 := by + gcongr + · calc + |x| ^ 4 ≤ |x| ^ 1 := + pow_le_pow_of_le_one (abs_nonneg _) + (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]) (by decide) + _ = x := by simp [_root_.abs_of_nonneg (le_of_lt hx0)] + · calc + x ^ 3 ≤ x ^ 1 := pow_le_pow_of_le_one (le_of_lt hx0) hx (by decide) + _ = x := pow_one _ + _ < x := by linarith) + _ ≤ sin x := + sub_le_comm.1 (abs_sub_le_iff.1 (sin_bound (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]))).2 + +theorem sin_pos_of_pos_of_le_two {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 2) : 0 < sin x := + have : x / 2 ≤ 1 := (div_le_iff₀ (by norm_num)).mpr (by simpa) + calc + 0 < 2 * sin (x / 2) * cos (x / 2) := + mul_pos (mul_pos (by norm_num) (sin_pos_of_pos_of_le_one (half_pos hx0) this)) + (cos_pos_of_le_one (by rwa [_root_.abs_of_nonneg (le_of_lt (half_pos hx0))])) + _ = sin x := by rw [← sin_two_mul, two_mul, add_halves] + +theorem cos_one_le : cos 1 ≤ 2 / 3 := + calc + cos 1 ≤ |(1 : ℝ)| ^ 4 * (5 / 96) + (1 - 1 ^ 2 / 2) := + sub_le_iff_le_add.1 (abs_sub_le_iff.1 (cos_bound (by simp))).1 + _ ≤ 2 / 3 := by norm_num + +theorem cos_one_pos : 0 < cos 1 := + cos_pos_of_le_one (le_of_eq abs_one) + +theorem cos_two_neg : cos 2 < 0 := + calc cos 2 = cos (2 * 1) := congr_arg cos (mul_one _).symm + _ = _ := Real.cos_two_mul 1 + _ ≤ 2 * (2 / 3) ^ 2 - 1 := by + gcongr + · exact cos_one_pos.le + · apply cos_one_le + _ < 0 := by norm_num + +end Real + +namespace Mathlib.Meta.Positivity +open Lean.Meta Qq + +/-- Extension for the `positivity` tactic: `Real.cosh` is always positive. -/ +@[positivity Real.cosh _] +def evalCosh : PositivityExt where eval {u α} _ _ e := do + match u, α, e with + | 0, ~q(ℝ), ~q(Real.cosh $a) => + assertInstancesCommute + return .positive q(Real.cosh_pos $a) + | _, _, _ => throwError "not Real.cosh" + +example (x : ℝ) : 0 < x.cosh := by positivity + +end Mathlib.Meta.Positivity + +namespace Complex + +@[simp] +theorem abs_cos_add_sin_mul_I (x : ℝ) : abs (cos x + sin x * I) = 1 := by + have := Real.sin_sq_add_cos_sq x + simp_all [add_comm, abs, normSq, sq, sin_ofReal_re, cos_ofReal_re, mul_re] + +@[simp] +theorem abs_exp_ofReal_mul_I (x : ℝ) : abs (exp (x * I)) = 1 := by + rw [exp_mul_I, abs_cos_add_sin_mul_I] + +theorem abs_exp (z : ℂ) : abs (exp z) = Real.exp z.re := by + rw [exp_eq_exp_re_mul_sin_add_cos, map_mul, abs_exp_ofReal, abs_cos_add_sin_mul_I, mul_one] + +theorem abs_exp_eq_iff_re_eq {x y : ℂ} : abs (exp x) = abs (exp y) ↔ x.re = y.re := by + rw [abs_exp, abs_exp, Real.exp_eq_exp] + +end Complex diff --git a/Mathlib/Data/DFinsupp/Order.lean b/Mathlib/Data/DFinsupp/Order.lean index 9201ee22c9791b..0ea38f0ada1209 100644 --- a/Mathlib/Data/DFinsupp/Order.lean +++ b/Mathlib/Data/DFinsupp/Order.lean @@ -44,14 +44,7 @@ lemma le_def : f ≤ g ↔ ∀ i, f i ≤ g i := Iff.rfl def orderEmbeddingToFun : (Π₀ i, α i) ↪o ∀ i, α i where toFun := DFunLike.coe inj' := DFunLike.coe_injective - map_rel_iff' := - #adaptation_note - /-- - This proof used to be `rfl`, - but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. - It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 - -/ - Iff.rfl + map_rel_iff' := Iff.rfl @[simp, norm_cast] lemma coe_orderEmbeddingToFun : ⇑(orderEmbeddingToFun (α := α)) = DFunLike.coe := rfl diff --git a/Mathlib/Data/DFinsupp/Sigma.lean b/Mathlib/Data/DFinsupp/Sigma.lean index 10d83e6df6193c..c0fd6d64a85c2f 100644 --- a/Mathlib/Data/DFinsupp/Sigma.lean +++ b/Mathlib/Data/DFinsupp/Sigma.lean @@ -63,8 +63,7 @@ theorem sigmaCurry_zero [∀ i j, Zero (δ i j)] : @[simp] theorem sigmaCurry_add [∀ i j, AddZeroClass (δ i j)] (f g : Π₀ (i : Σ _, _), δ i.1 i.2) : - #adaptation_note - /-- After https://github.com/leanprover/lean4/pull/6024 + #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 we needed to add the `(_ : Π₀ (i) (j), δ i j)` type annotation. -/ sigmaCurry (f + g) = (sigmaCurry f + sigmaCurry g : Π₀ (i) (j), δ i j) := by ext (i j) @@ -73,8 +72,7 @@ theorem sigmaCurry_add [∀ i j, AddZeroClass (δ i j)] (f g : Π₀ (i : Σ _, @[simp] theorem sigmaCurry_smul [Monoid γ] [∀ i j, AddMonoid (δ i j)] [∀ i j, DistribMulAction γ (δ i j)] (r : γ) (f : Π₀ (i : Σ _, _), δ i.1 i.2) : - #adaptation_note - /-- After https://github.com/leanprover/lean4/pull/6024 + #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 we needed to add the `(_ : Π₀ (i) (j), δ i j)` type annotation. -/ sigmaCurry (r • f) = (r • sigmaCurry f : Π₀ (i) (j), δ i j) := by ext (i j) diff --git a/Mathlib/Data/DFinsupp/WellFounded.lean b/Mathlib/Data/DFinsupp/WellFounded.lean index 4b312a0571ec77..ac37b3707bd633 100644 --- a/Mathlib/Data/DFinsupp/WellFounded.lean +++ b/Mathlib/Data/DFinsupp/WellFounded.lean @@ -212,9 +212,7 @@ protected theorem DFinsupp.wellFoundedLT [∀ i, Zero (α i)] [∀ i, Preorder ( refine Lex.wellFounded' ?_ (fun i ↦ IsWellFounded.wf) ?_ · rintro i ⟨a⟩ apply hbot - · #adaptation_note /-- nightly-2024-03-16: simp was - simp (config := { unfoldPartialApp := true }) only [Function.swap] -/ - simp only [Function.swap_def] + · simp +unfoldPartialApp only [Function.swap] exact IsWellFounded.wf refine Subrelation.wf (fun h => ?_) <| InvImage.wf (mapRange e fun _ ↦ rfl) this have := IsStrictOrder.swap (@WellOrderingRel ι) diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 9b9e842c4e348f..63249a86934640 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -44,7 +44,8 @@ protected theorem div_eq_inv_mul : a / b = b⁻¹ * a := by rw [div_eq_mul_inv, show sInf { b : ℝ≥0∞ | 1 ≤ 0 * b } = ∞ by simp @[simp] theorem inv_top : ∞⁻¹ = 0 := - bot_unique <| le_of_forall_le_of_dense fun a (h : 0 < a) => sInf_le <| by simp [*, h.ne', top_mul] + bot_unique <| le_of_forall_gt_imp_ge_of_dense fun a (h : 0 < a) => sInf_le <| by + simp [*, h.ne', top_mul] theorem coe_inv_le : (↑r⁻¹ : ℝ≥0∞) ≤ (↑r)⁻¹ := le_sInf fun b (hb : 1 ≤ ↑r * b) => @@ -444,7 +445,7 @@ instance : SMulPosMono ℝ≥0 ℝ≥0∞ where elim _r _ _a _b hab := mul_le_mul_right' (coe_le_coe.2 hab) _ theorem le_of_forall_nnreal_lt {x y : ℝ≥0∞} (h : ∀ r : ℝ≥0, ↑r < x → ↑r ≤ y) : x ≤ y := by - refine le_of_forall_ge_of_dense fun r hr => ?_ + refine le_of_forall_lt_imp_le_of_dense fun r hr => ?_ lift r to ℝ≥0 using ne_top_of_lt hr exact h r hr @@ -533,7 +534,7 @@ private lemma exists_lt_mul_right {a b c : ℝ≥0∞} (hc : c < a * b) : ∃ b' simp_rw [mul_comm a] at hc ⊢; exact exists_lt_mul_left hc lemma mul_le_of_forall_lt {a b c : ℝ≥0∞} (h : ∀ a' < a, ∀ b' < b, a' * b' ≤ c) : a * b ≤ c := by - refine le_of_forall_ge_of_dense fun d hd ↦ ?_ + refine le_of_forall_lt_imp_le_of_dense fun d hd ↦ ?_ obtain ⟨a', ha', hd⟩ := exists_lt_mul_left hd obtain ⟨b', hb', hd⟩ := exists_lt_mul_right hd exact le_trans hd.le <| h _ ha' _ hb' diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 7dc978fddb1020..331cbae4643f70 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -240,6 +240,12 @@ This one instead uses a `NeZero n` typeclass hypothesis. theorem val_zero' (n : ℕ) [NeZero n] : ((0 : Fin n) : ℕ) = 0 := rfl +/-- Fin.mk_zero` in `Lean` only applies in `Fin (n + 1)`. +This one instead uses a `NeZero n` typeclass hypothesis. +-/ +@[simp] +theorem mk_zero' (n : ℕ) [NeZero n] : (⟨0, pos_of_neZero n⟩ : Fin n) = 0 := rfl + /-- The `Fin.zero_le` in `Lean` only applies in `Fin (n+1)`. This one instead uses a `NeZero n` typeclass hypothesis. @@ -565,9 +571,15 @@ theorem leftInverse_cast (eq : n = m) : LeftInverse (Fin.cast eq.symm) (Fin.cast theorem rightInverse_cast (eq : n = m) : RightInverse (Fin.cast eq.symm) (Fin.cast eq) := fun _ => rfl +@[simp] +theorem cast_inj (eq : n = m) {a b : Fin n} : a.cast eq = b.cast eq ↔ a = b := by + simp [← val_inj] + +@[simp] theorem cast_lt_cast (eq : n = m) {a b : Fin n} : a.cast eq < b.cast eq ↔ a < b := Iff.rfl +@[simp] theorem cast_le_cast (eq : n = m) {a b : Fin n} : a.cast eq ≤ b.cast eq ↔ a ≤ b := Iff.rfl diff --git a/Mathlib/Data/Finite/Sum.lean b/Mathlib/Data/Finite/Sum.lean index 177f8e038134a3..4c3ecc2999ef54 100644 --- a/Mathlib/Data/Finite/Sum.lean +++ b/Mathlib/Data/Finite/Sum.lean @@ -24,4 +24,13 @@ theorem sum_left (β) [Finite (α ⊕ β)] : Finite α := theorem sum_right (α) [Finite (α ⊕ β)] : Finite β := of_injective (Sum.inr : β → α ⊕ β) Sum.inr_injective +instance {α β : Sort*} [Finite α] [Finite β] : Finite (α ⊕' β) := + of_equiv _ ((Equiv.psumEquivSum _ _).symm.trans (Equiv.plift.psumCongr Equiv.plift)) + +theorem psum_left {α β : Sort*} [Finite (α ⊕' β)] : Finite α := + of_injective (PSum.inl : α → α ⊕' β) PSum.inl_injective + +theorem psum_right {α β : Sort*} [Finite (α ⊕' β)] : Finite β := + of_injective (PSum.inr : β → α ⊕' β) PSum.inr_injective + end Finite diff --git a/Mathlib/Data/Finset/Lattice/Fold.lean b/Mathlib/Data/Finset/Lattice/Fold.lean index 0a04809d8a4e61..c5ae9f5e4dcdb9 100644 --- a/Mathlib/Data/Finset/Lattice/Fold.lean +++ b/Mathlib/Data/Finset/Lattice/Fold.lean @@ -250,6 +250,11 @@ theorem sup_mem (s : Set α) (w₁ : ⊥ ∈ s) (w₂ : ∀ᵉ (x ∈ s) (y ∈ protected theorem sup_eq_bot_iff (f : β → α) (S : Finset β) : S.sup f = ⊥ ↔ ∀ s ∈ S, f s = ⊥ := by classical induction' S using Finset.induction with a S _ hi <;> simp [*] +@[simp] +theorem sup_eq_bot_of_isEmpty [IsEmpty β] (f : β → α) (S : Finset β) : S.sup f = ⊥ := by + rw [Finset.sup_eq_bot_iff] + exact fun x _ => False.elim <| IsEmpty.false x + end Sup theorem sup_eq_iSup [CompleteLattice β] (s : Finset α) (f : α → β) : s.sup f = ⨆ a ∈ s, f a := diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 6abe2e3b9c254f..8edd359b3d1c99 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -1466,8 +1466,8 @@ theorem extendDomain_subtypeDomain (f : α →₀ M) (hf : ∀ a ∈ f.support, ext a by_cases h : P a · exact dif_pos h - · #adaptation_note - /-- Prior to nightly-2024-06-18, this `rw` was done by `dsimp`. -/ + · #adaptation_note /-- nightly-2024-06-18 + this `rw` was done by `dsimp`. -/ rw [extendDomain_toFun] dsimp rw [if_neg h, eq_comm, ← not_mem_support_iff] @@ -1478,8 +1478,8 @@ theorem extendDomain_subtypeDomain (f : α →₀ M) (hf : ∀ a ∈ f.support, theorem extendDomain_single (a : Subtype P) (m : M) : (single a m).extendDomain = single a.val m := by ext a' - #adaptation_note - /-- Prior to nightly-2024-06-18, this `rw` was instead `dsimp only`. -/ + #adaptation_note /-- nightly-2024-06-18 + this `rw` was instead `dsimp only`. -/ rw [extendDomain_toFun] obtain rfl | ha := eq_or_ne a.val a' · simp_rw [single_eq_same, dif_pos a.prop] diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index ff4691dc547893..3207bf60506dd0 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -1060,11 +1060,7 @@ variable [Fintype α] [DecidableEq β] {f : α → β} /-- `bijInv f` is the unique inverse to a bijection `f`. This acts as a computable alternative to `Function.invFun`. -/ def bijInv (f_bij : Bijective f) (b : β) : α := - Fintype.choose (fun a => f a = b) - (by - rcases f_bij.right b with ⟨a', fa_eq_b⟩ - rw [← fa_eq_b] - exact ⟨a', ⟨rfl, fun a h => f_bij.left h⟩⟩) + Fintype.choose (fun a => f a = b) (f_bij.existsUnique b) theorem leftInverse_bijInv (f_bij : Bijective f) : LeftInverse (bijInv f_bij) f := fun a => f_bij.left (choose_spec (fun a' => f a' = f a) _) diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index d51295c269ce17..187b3b0cb21bcc 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -29,6 +29,10 @@ open Nat hiding one_pos namespace List +-- Renamed in lean core; to be removed with the version bump. +alias replicate_append_replicate := append_replicate_replicate +alias append_eq_nil_iff := append_eq_nil + universe u v w variable {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {l₁ l₂ : List α} @@ -208,7 +212,7 @@ theorem eq_replicate_length {a : α} : ∀ {l : List α}, l = replicate l.length | (b :: l) => by simp [eq_replicate_length, replicate_succ] theorem replicate_add (m n) (a : α) : replicate (m + n) a = replicate m a ++ replicate n a := by - rw [append_replicate_replicate] + rw [replicate_append_replicate] theorem replicate_subset_singleton (n) (a : α) : replicate n a ⊆ [a] := fun _ h => mem_singleton.2 (eq_of_mem_replicate h) @@ -379,7 +383,8 @@ theorem getLastI_eq_getLast? [Inhabited α] : ∀ l : List α, l.getLastI = l.ge | [_, _, _] => rfl | _ :: _ :: c :: l => by simp [getLastI, getLastI_eq_getLast? (c :: l)] -#adaptation_note /-- 2024-07-10: removed `@[simp]` since the LHS simplifies using the simp set. -/ +#adaptation_note /-- 2024-07-10 +removed `@[simp]` since the LHS simplifies using the simp set. -/ theorem getLast?_append_cons : ∀ (l₁ : List α) (a : α) (l₂ : List α), getLast? (l₁ ++ a :: l₂) = getLast? (a :: l₂) | [], _, _ => rfl @@ -670,19 +675,19 @@ variable [DecidableEq α] -- indexOf_cons_eq _ rfl @[simp] -theorem indexOf_cons_self (a : α) (l : List α) : indexOf a (a :: l) = 0 := by +theorem indexOf_cons_self {a : α} {l : List α} : indexOf a (a :: l) = 0 := by rw [indexOf, findIdx_cons, beq_self_eq_true, cond] -- fun e => if_pos e theorem indexOf_cons_eq {a b : α} (l : List α) : b = a → indexOf a (b :: l) = 0 - | e => by rw [← e]; exact indexOf_cons_self b l + | e => by rw [← e]; exact indexOf_cons_self -- fun n => if_neg n @[simp] theorem indexOf_cons_ne {a b : α} (l : List α) : b ≠ a → indexOf a (b :: l) = succ (indexOf a l) | h => by simp only [indexOf, findIdx_cons, Bool.cond_eq_ite, beq_iff_eq, h, ite_false] -theorem indexOf_eq_length {a : α} {l : List α} : indexOf a l = length l ↔ a ∉ l := by +theorem indexOf_eq_length_iff {a : α} {l : List α} : indexOf a l = length l ↔ a ∉ l := by induction' l with b l ih · exact iff_of_true rfl (not_mem_nil _) simp only [length, mem_cons, indexOf_cons, eq_comm] @@ -693,9 +698,12 @@ theorem indexOf_eq_length {a : α} {l : List α} : indexOf a l = length l ↔ a rw [← ih] exact succ_inj' +@[deprecated (since := "2025-01-28")] +alias indexOf_eq_length := indexOf_eq_length_iff + @[simp] theorem indexOf_of_not_mem {l : List α} {a : α} : a ∉ l → indexOf a l = length l := - indexOf_eq_length.2 + indexOf_eq_length_iff.2 theorem indexOf_le_length {a : α} {l : List α} : indexOf a l ≤ length l := by induction' l with b l ih; · rfl @@ -705,8 +713,8 @@ theorem indexOf_le_length {a : α} {l : List α} : indexOf a l ≤ length l := b · rw [if_neg h]; exact succ_le_succ ih theorem indexOf_lt_length_iff {a} {l : List α} : indexOf a l < length l ↔ a ∈ l := - ⟨fun h => Decidable.byContradiction fun al => Nat.ne_of_lt h <| indexOf_eq_length.2 al, - fun al => (lt_of_le_of_ne indexOf_le_length) fun h => indexOf_eq_length.1 h al⟩ + ⟨fun h => Decidable.byContradiction fun al => Nat.ne_of_lt h <| indexOf_eq_length_iff.2 al, + fun al => (lt_of_le_of_ne indexOf_le_length) fun h => indexOf_eq_length_iff.1 h al⟩ @[deprecated (since := "2025-01-22")] alias indexOf_lt_length := indexOf_lt_length_iff @@ -875,7 +883,7 @@ theorem flatMap_pure_eq_map (f : α → β) (l : List α) : l.flatMap (pure ∘ @[deprecated (since := "2024-10-16")] alias bind_pure_eq_map := flatMap_pure_eq_map theorem flatMap_congr {l : List α} {f g : α → List β} (h : ∀ x ∈ l, f x = g x) : - List.flatMap l f = List.flatMap l g := + l.flatMap f = l.flatMap g := (congr_arg List.flatten <| map_congr_left h :) @[deprecated (since := "2024-10-16")] alias bind_congr := flatMap_congr @@ -1670,11 +1678,9 @@ theorem filter_eq_foldr (p : α → Bool) (l : List α) : filter p l = foldr (fun a out => bif p a then a :: out else out) [] l := by induction l <;> simp [*, filter]; rfl -#adaptation_note -/-- +#adaptation_note /-- nightly-2024-07-27 This has to be temporarily renamed to avoid an unintentional collision. -The prime should be removed at nightly-2024-07-27. --/ +The prime should be removed at nightly-2024-07-27. -/ @[simp] theorem filter_subset' (l : List α) : filter p l ⊆ l := (filter_sublist l).subset diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index 9718d8cc7eb2c8..a0351827a88e0a 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -488,10 +488,6 @@ theorem length_mapAccumr₂ : end MapAccumr -/- #adaptation_note: this attribute should be removed after Mathlib moves to v4.15.0-rc1. -/ -set_option allowUnsafeReducibility true in -attribute [semireducible] Fin.foldr.loop - section Deprecated @[deprecated List.mem_cons (since := "2024-08-10")] diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index e9680556a8109d..e307495cd9f2cd 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -9,6 +9,14 @@ import Mathlib.Data.List.Basic # Lemmas about List.*Idx functions. Some specification lemmas for `List.mapIdx`, `List.mapIdxM`, `List.foldlIdx` and `List.foldrIdx`. + +As of 2025-01-29, these are not used anywhere in Mathlib. Moreover, with +`List.enum` and `List.enumFrom` being replaced by `List.zipIdx` +in Lean's `nightly-2025-01-29` release, they now use deprecated functions and theorems. +Rather than updating this unused material, we are deprecating it. +Anyone wanting to restore this material is welcome to do so, but will need to update uses of +`List.enum` and `List.enumFrom` to use `List.zipIdx` instead. +However, note that this material will later be implemented in the Lean standard library. -/ assert_not_exists MonoidWithZero @@ -23,17 +31,10 @@ variable {α : Type u} {β : Type v} section MapIdx +@[deprecated reverseRecOn (since := "2025-01-28")] theorem list_reverse_induction (p : List α → Prop) (base : p []) - (ind : ∀ (l : List α) (e : α), p l → p (l ++ [e])) : (∀ (l : List α), p l) := by - let q := fun l ↦ p (reverse l) - have pq : ∀ l, p (reverse l) → q l := by simp only [q, reverse_reverse]; intro; exact id - have qp : ∀ l, q (reverse l) → p l := by simp only [q, reverse_reverse]; intro; exact id - intro l - apply qp - generalize (reverse l) = l - induction' l with head tail ih - · apply pq; simp only [reverse_nil, base] - · apply pq; simp only [reverse_cons]; apply ind; apply qp; rw [reverse_reverse]; exact ih + (ind : ∀ (l : List α) (e : α), p l → p (l ++ [e])) : (∀ (l : List α), p l) := + fun l => l.reverseRecOn base ind @[deprecated (since := "2024-10-15")] alias mapIdxGo_length := mapIdx_go_length @@ -41,7 +42,7 @@ theorem mapIdx_append_one : ∀ {f : ℕ → α → β} {l : List α} {e : α}, mapIdx f (l ++ [e]) = mapIdx f l ++ [f l.length e] := mapIdx_concat -@[local simp] +@[deprecated "Deprecated without replacement." (since := "2025-01-29"), local simp] theorem map_enumFrom_eq_zipWith : ∀ (l : List α) (n : ℕ) (f : ℕ → α → β), map (uncurry f) (enumFrom n l) = zipWith (fun i ↦ f (i + n)) (range (length l)) l := by intro l @@ -67,6 +68,7 @@ theorem map_enumFrom_eq_zipWith : ∀ (l : List α) (n : ℕ) (f : ℕ → α @[deprecated (since := "2024-10-15")] alias mapIdx_eq_nil := mapIdx_eq_nil_iff +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem get_mapIdx (l : List α) (f : ℕ → α → β) (i : ℕ) (h : i < l.length) (h' : i < (l.mapIdx f).length := h.trans_le length_mapIdx.ge) : (l.mapIdx f).get ⟨i, h'⟩ = f i (l.get ⟨i, h⟩) := by @@ -159,30 +161,41 @@ section FoldrIdx -- Porting note: Changed argument order of `foldrIdxSpec` to align better with `foldrIdx`. /-- Specification of `foldrIdx`. -/ +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] def foldrIdxSpec (f : ℕ → α → β → β) (b : β) (as : List α) (start : ℕ) : β := foldr (uncurry f) b <| enumFrom start as +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem foldrIdxSpec_cons (f : ℕ → α → β → β) (b a as start) : foldrIdxSpec f b (a :: as) start = f start a (foldrIdxSpec f b as (start + 1)) := rfl +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem foldrIdx_eq_foldrIdxSpec (f : ℕ → α → β → β) (b as start) : foldrIdx f b as start = foldrIdxSpec f b as start := by induction as generalizing start · rfl · simp only [foldrIdx, foldrIdxSpec_cons, *] +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem foldrIdx_eq_foldr_enum (f : ℕ → α → β → β) (b : β) (as : List α) : foldrIdx f b as = foldr (uncurry f) b (enum as) := by simp only [foldrIdx, foldrIdxSpec, foldrIdx_eq_foldrIdxSpec, enum] end FoldrIdx +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem indexesValues_eq_filter_enum (p : α → Prop) [DecidablePred p] (as : List α) : indexesValues p as = filter (p ∘ Prod.snd) (enum as) := by simp (config := { unfoldPartialApp := true }) [indexesValues, foldrIdx_eq_foldr_enum, uncurry, filter_eq_foldr, cond_eq_if] +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem findIdxs_eq_map_indexesValues (p : α → Prop) [DecidablePred p] (as : List α) : findIdxs p as = map Prod.fst (indexesValues p as) := by simp (config := { unfoldPartialApp := true }) only [indexesValues_eq_filter_enum, @@ -193,19 +206,26 @@ section FoldlIdx -- Porting note: Changed argument order of `foldlIdxSpec` to align better with `foldlIdx`. /-- Specification of `foldlIdx`. -/ +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] def foldlIdxSpec (f : ℕ → α → β → α) (a : α) (bs : List β) (start : ℕ) : α := foldl (fun a p ↦ f p.fst a p.snd) a <| enumFrom start bs +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem foldlIdxSpec_cons (f : ℕ → α → β → α) (a b bs start) : foldlIdxSpec f a (b :: bs) start = foldlIdxSpec f (f start a b) bs (start + 1) := rfl +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem foldlIdx_eq_foldlIdxSpec (f : ℕ → α → β → α) (a bs start) : foldlIdx f a bs start = foldlIdxSpec f a bs start := by induction bs generalizing start a · rfl · simp [foldlIdxSpec, *] +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem foldlIdx_eq_foldl_enum (f : ℕ → α → β → α) (a : α) (bs : List β) : foldlIdx f a bs = foldl (fun a p ↦ f p.fst a p.snd) a (enum bs) := by simp only [foldlIdx, foldlIdxSpec, foldlIdx_eq_foldlIdxSpec, enum] @@ -217,11 +237,15 @@ section FoldIdxM -- Porting note: `foldrM_eq_foldr` now depends on `[LawfulMonad m]` variable {m : Type u → Type v} [Monad m] +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem foldrIdxM_eq_foldrM_enum {β} (f : ℕ → α → β → m β) (b : β) (as : List α) [LawfulMonad m] : foldrIdxM f b as = foldrM (uncurry f) b (enum as) := by simp (config := { unfoldPartialApp := true }) only [foldrIdxM, foldrM_eq_foldr, foldrIdx_eq_foldr_enum, uncurry] +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem foldlIdxM_eq_foldlM_enum [LawfulMonad m] {β} (f : ℕ → β → α → m β) (b : β) (as : List α) : foldlIdxM f b as = List.foldlM (fun b p ↦ f p.fst b p.snd) b (enum as) := by rw [foldlIdxM, foldlM_eq_foldl, foldlIdx_eq_foldl_enum] @@ -234,15 +258,20 @@ section MapIdxM variable {m : Type u → Type v} [Monad m] /-- Specification of `mapIdxMAux`. -/ +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] def mapIdxMAuxSpec {β} (f : ℕ → α → m β) (start : ℕ) (as : List α) : m (List β) := List.traverse (uncurry f) <| enumFrom start as -- Note: `traverse` the class method would require a less universe-polymorphic -- `m : Type u → Type u`. +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem mapIdxMAuxSpec_cons {β} (f : ℕ → α → m β) (start : ℕ) (a : α) (as : List α) : mapIdxMAuxSpec f start (a :: as) = cons <$> f start a <*> mapIdxMAuxSpec f (start + 1) as := rfl +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem mapIdxMGo_eq_mapIdxMAuxSpec [LawfulMonad m] {β} (f : ℕ → α → m β) (arr : Array β) (as : List α) : mapIdxM.go f as arr = (arr.toList ++ ·) <$> mapIdxMAuxSpec f arr.size as := by @@ -266,6 +295,8 @@ theorem mapIdxMGo_eq_mapIdxMAuxSpec simp only [Array.push_toList, append_assoc, singleton_append, Array.size_push, map_eq_pure_bind] +set_option linter.deprecated false in +@[deprecated "Deprecated without replacement." (since := "2025-01-29")] theorem mapIdxM_eq_mmap_enum [LawfulMonad m] {β} (f : ℕ → α → m β) (as : List α) : as.mapIdxM f = List.traverse (uncurry f) (enum as) := by simp only [mapIdxM, mapIdxMGo_eq_mapIdxMAuxSpec, Array.toList_toArray, diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index 71564ef6cb8486..241a87b598d882 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -38,7 +38,7 @@ def argAux (a : Option α) (b : α) : Option α := @[simp] theorem foldl_argAux_eq_none : l.foldl (argAux r) o = none ↔ l = [] ∧ o = none := List.reverseRecOn l (by simp) fun tl hd => by - simp only [foldl_append, foldl_cons, argAux, foldl_nil, append_eq_nil, and_false, false_and, + simp only [foldl_append, foldl_cons, argAux, foldl_nil, append_eq_nil_iff, and_false, false_and, iff_false] cases foldl (argAux r) o tl · simp diff --git a/Mathlib/Data/List/Monad.lean b/Mathlib/Data/List/Monad.lean index 3132c8376d9229..0c420b1774ff69 100644 --- a/Mathlib/Data/List/Monad.lean +++ b/Mathlib/Data/List/Monad.lean @@ -16,9 +16,9 @@ namespace List variable {α : Type u} instance instMonad : Monad List.{u} where - pure := @List.singleton - bind := @List.flatMap - map := @List.map + pure x := [x] + bind l f := l.flatMap f + map f l := l.map f @[simp] theorem pure_def (a : α) : pure a = [a] := rfl diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index bdd3205fa2f21a..d598a0ba42d67f 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -326,7 +326,7 @@ theorem perm_permutations'Aux_comm (a b : α) (l : List α) : map (cons c) ((permutations'Aux a l).flatMap (permutations'Aux b)) := by intros a' b' simp only [flatMap_map, permutations'Aux] - show List.flatMap (permutations'Aux _ l) (fun a => ([b' :: c :: a] ++ + show (permutations'Aux _ l).flatMap (fun a => ([b' :: c :: a] ++ map (cons c) (permutations'Aux _ a))) ~ _ refine (flatMap_append_perm _ (fun x => [b' :: c :: x]) _).symm.trans ?_ rw [← map_eq_flatMap, ← map_flatMap] diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index 69638ec2666c4c..d7eeeed30a9772 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -49,6 +49,17 @@ theorem stdBasisMatrix_zero (i : m) (j : n) : stdBasisMatrix i j (0 : α) = 0 := ext simp +@[simp] +lemma transpose_stdBasisMatrix (i : m) (j : n) (a : α) : + (stdBasisMatrix i j a)ᵀ = stdBasisMatrix j i a := by + aesop (add unsafe unfold stdBasisMatrix) + +@[simp] +lemma map_stdBasisMatrix (i : m) (j : n) (a : α) {β : Type*} [Zero β] + {F : Type*} [FunLike F α β] [ZeroHomClass F α β] (f : F) : + (stdBasisMatrix i j a).map f = stdBasisMatrix i j (f a) := by + aesop (add unsafe unfold stdBasisMatrix) + end Zero theorem stdBasisMatrix_add [AddZeroClass α] (i : m) (j : n) (a b : α) : diff --git a/Mathlib/Data/Matrix/ConjTranspose.lean b/Mathlib/Data/Matrix/ConjTranspose.lean index 4b856340799ca6..f43539add8aae8 100644 --- a/Mathlib/Data/Matrix/ConjTranspose.lean +++ b/Mathlib/Data/Matrix/ConjTranspose.lean @@ -11,6 +11,7 @@ import Mathlib.Algebra.Star.BigOperators import Mathlib.Algebra.Star.Module import Mathlib.Algebra.Star.Pi import Mathlib.Data.Fintype.BigOperators +import Mathlib.Data.Matrix.Basis import Mathlib.Data.Matrix.Mul /-! @@ -40,6 +41,13 @@ def conjTranspose [Star α] (M : Matrix m n α) : Matrix n m α := @[inherit_doc] scoped postfix:1024 "ᴴ" => Matrix.conjTranspose +@[simp] +lemma conjTranspose_stdBasisMatrix [DecidableEq n] [DecidableEq m] [AddMonoid α] + [StarAddMonoid α] (i : m) (j : n) (a : α) : + (stdBasisMatrix i j a)ᴴ = stdBasisMatrix j i (star a) := by + show (stdBasisMatrix i j a).transpose.map starAddEquiv = stdBasisMatrix j i (star a) + simp + section Diagonal variable [DecidableEq n] diff --git a/Mathlib/Data/Matrix/Mul.lean b/Mathlib/Data/Matrix/Mul.lean index 4f71ea9008fa3a..c074c0ef47fffb 100644 --- a/Mathlib/Data/Matrix/Mul.lean +++ b/Mathlib/Data/Matrix/Mul.lean @@ -722,12 +722,12 @@ theorem mulVec_smul [Fintype n] [Monoid R] [NonUnitalNonAssocSemiring S] [Distri @[simp] theorem mulVec_single [Fintype n] [DecidableEq n] [NonUnitalNonAssocSemiring R] (M : Matrix m n R) - (j : n) (x : R) : M *ᵥ Pi.single j x = fun i => M i j * x := + (j : n) (x : R) : M *ᵥ Pi.single j x = MulOpposite.op x • Mᵀ j := funext fun _ => dotProduct_single _ _ _ @[simp] theorem single_vecMul [Fintype m] [DecidableEq m] [NonUnitalNonAssocSemiring R] (M : Matrix m n R) - (i : m) (x : R) : Pi.single i x ᵥ* M = fun j => x * M i j := + (i : m) (x : R) : Pi.single i x ᵥ* M = x • M i := funext fun _ => single_dotProduct _ _ _ theorem mulVec_single_one [Fintype n] [DecidableEq n] [NonAssocSemiring R] @@ -736,7 +736,7 @@ theorem mulVec_single_one [Fintype n] [DecidableEq n] [NonAssocSemiring R] theorem single_one_vecMul [Fintype m] [DecidableEq m] [NonAssocSemiring R] (i : m) (M : Matrix m n R) : - Pi.single i 1 ᵥ* M = M i := by simp + Pi.single i 1 ᵥ* M = M i := by ext; simp theorem diagonal_mulVec_single [Fintype n] [DecidableEq n] [NonUnitalNonAssocSemiring R] (v : n → R) (j : n) (x : R) : diagonal v *ᵥ Pi.single j x = Pi.single j (v j * x) := by @@ -781,12 +781,28 @@ section NonAssocSemiring variable [NonAssocSemiring α] -theorem mulVec_one [Fintype n] (A : Matrix m n α) : A *ᵥ 1 = fun i => ∑ j, A i j := by +theorem mulVec_one [Fintype n] (A : Matrix m n α) : A *ᵥ 1 = ∑ j, Aᵀ j := by ext; simp [mulVec, dotProduct] -theorem vec_one_mul [Fintype m] (A : Matrix m n α) : 1 ᵥ* A = fun j => ∑ i, A i j := by +theorem one_vecMul [Fintype m] (A : Matrix m n α) : 1 ᵥ* A = ∑ i, A i := by ext; simp [vecMul, dotProduct] +@[deprecated (since := "2025-01-26")] alias vec_one_mul := one_vecMul + +lemma ext_of_mulVec_single [DecidableEq n] [Fintype n] {M N : Matrix m n α} + (h : ∀ i, M *ᵥ Pi.single i 1 = N *ᵥ Pi.single i 1) : + M = N := by + ext i j + simp_rw [mulVec_single_one] at h + exact congrFun (h j) i + +lemma ext_of_single_vecMul [DecidableEq m] [Fintype m] {M N : Matrix m n α} + (h : ∀ i, Pi.single i 1 ᵥ* M = Pi.single i 1 ᵥ* N) : + M = N := by + ext i j + simp_rw [single_one_vecMul] at h + exact congrFun (h i) j + variable [Fintype m] [Fintype n] [DecidableEq m] @[simp] diff --git a/Mathlib/Data/Matrix/PEquiv.lean b/Mathlib/Data/Matrix/PEquiv.lean index eb70fede2430c4..98cb355cc674a8 100644 --- a/Mathlib/Data/Matrix/PEquiv.lean +++ b/Mathlib/Data/Matrix/PEquiv.lean @@ -40,7 +40,7 @@ open Matrix universe u v variable {k l m n : Type*} -variable {α : Type v} +variable {α β : Type*} open Matrix @@ -55,13 +55,28 @@ theorem toMatrix_apply [DecidableEq n] [Zero α] [One α] (f : m ≃. n) (i j) : toMatrix f i j = if j ∈ f i then (1 : α) else 0 := rfl -theorem mul_matrix_apply [Fintype m] [DecidableEq m] [Semiring α] (f : l ≃. m) (M : Matrix m n α) +theorem toMatrix_mul_apply [Fintype m] [DecidableEq m] [Semiring α] (f : l ≃. m) (M : Matrix m n α) (i j) : (f.toMatrix * M :) i j = Option.casesOn (f i) 0 fun fi => M fi j := by dsimp [toMatrix, Matrix.mul_apply] cases' h : f i with fi · simp [h] · rw [Finset.sum_eq_single fi] <;> simp +contextual [h, eq_comm] +@[deprecated (since := "2025-01-27")] alias mul_matrix_apply := toMatrix_mul_apply + +theorem mul_toMatrix_apply [Fintype m] [Semiring α] [DecidableEq n] (M : Matrix l m α) (f : m ≃. n) + (i j) : (M * f.toMatrix :) i j = Option.casesOn (f.symm j) 0 (M i) := by + dsimp [Matrix.mul_apply, toMatrix_apply] + cases' h : f.symm j with fj + · simp [h, ← f.eq_some_iff] + · rw [Finset.sum_eq_single fj] + · simp [h, ← f.eq_some_iff] + · rintro b - n + simp [h, ← f.eq_some_iff, n.symm] + · simp + +@[deprecated (since := "2025-01-27")] alias matrix_mul_apply := mul_toMatrix_apply + theorem toMatrix_symm [DecidableEq m] [DecidableEq n] [Zero α] [One α] (f : m ≃. n) : (f.symm.toMatrix : Matrix n m α) = f.toMatrixᵀ := by ext @@ -74,32 +89,53 @@ theorem toMatrix_refl [DecidableEq n] [Zero α] [One α] : ext simp [toMatrix_apply, one_apply] -theorem matrix_mul_apply [Fintype m] [Semiring α] [DecidableEq n] (M : Matrix l m α) (f : m ≃. n) - (i j) : (M * f.toMatrix :) i j = Option.casesOn (f.symm j) 0 fun fj => M i fj := by - dsimp [toMatrix, Matrix.mul_apply] - cases' h : f.symm j with fj - · simp [h, ← f.eq_some_iff] - · rw [Finset.sum_eq_single fj] - · simp [h, ← f.eq_some_iff] - · rintro b - n - simp [h, ← f.eq_some_iff, n.symm] - · simp +@[simp] +theorem toMatrix_toPEquiv_apply [DecidableEq n] [Zero α] [One α] (f : m ≃ n) (i) : + f.toPEquiv.toMatrix i = Pi.single (f i) (1 : α) := by + ext + simp [toMatrix_apply, Pi.single_apply, eq_comm] + +@[simp] +theorem transpose_toMatrix_toPEquiv_apply + [DecidableEq m] [DecidableEq n] [Zero α] [One α] (f : m ≃ n) (j) : + f.toPEquiv.toMatrixᵀ j = Pi.single (f.symm j) (1 : α) := by + ext + simp [toMatrix_apply, Pi.single_apply, eq_comm, ← Equiv.apply_eq_iff_eq_symm_apply] -theorem toPEquiv_mul_matrix [Fintype m] [DecidableEq m] [Semiring α] (f : m ≃ m) - (M : Matrix m n α) : f.toPEquiv.toMatrix * M = M.submatrix f id := by +theorem toMatrix_toPEquiv_mul [Fintype m] [DecidableEq m] + [Semiring α] (f : l ≃ m) (M : Matrix m n α) : + f.toPEquiv.toMatrix * M = M.submatrix f id := by ext i j - rw [mul_matrix_apply, Equiv.toPEquiv_apply, submatrix_apply, id] + rw [toMatrix_mul_apply, Equiv.toPEquiv_apply, submatrix_apply, id] -theorem mul_toPEquiv_toMatrix {m n α : Type*} [Fintype n] [DecidableEq n] [Semiring α] (f : n ≃ n) - (M : Matrix m n α) : M * f.toPEquiv.toMatrix = M.submatrix id f.symm := +@[deprecated (since := "2025-01-27")] alias toPEquiv_mul_matrix := toMatrix_toPEquiv_mul + +theorem mul_toMatrix_toPEquiv [Fintype m] [DecidableEq n] + [Semiring α] (M : Matrix l m α) (f : m ≃ n) : + (M * f.toPEquiv.toMatrix) = M.submatrix id f.symm := Matrix.ext fun i j => by - rw [PEquiv.matrix_mul_apply, ← Equiv.toPEquiv_symm, Equiv.toPEquiv_apply, + rw [PEquiv.mul_toMatrix_apply, ← Equiv.toPEquiv_symm, Equiv.toPEquiv_apply, Matrix.submatrix_apply, id] +@[deprecated (since := "2025-01-27")] alias mul_toPEquiv_toMatrix := mul_toMatrix_toPEquiv + +lemma toMatrix_toPEquiv_mulVec [DecidableEq n] [Fintype n] + [NonAssocSemiring α] (σ : m ≃ n) (a : n → α) : + σ.toPEquiv.toMatrix *ᵥ a = a ∘ σ := by + ext j + simp [toMatrix, mulVec, dotProduct] + +lemma vecMul_toMatrix_toPEquiv [DecidableEq n] [Fintype m] + [NonAssocSemiring α] (σ : m ≃ n) (a : m → α) : + a ᵥ* σ.toPEquiv.toMatrix = a ∘ σ.symm := by + classical + ext j + simp [toMatrix, σ.apply_eq_iff_eq_symm_apply, vecMul, dotProduct] + theorem toMatrix_trans [Fintype m] [DecidableEq m] [DecidableEq n] [Semiring α] (f : l ≃. m) (g : m ≃. n) : ((f.trans g).toMatrix : Matrix l n α) = f.toMatrix * g.toMatrix := by ext i j - rw [mul_matrix_apply] + rw [toMatrix_mul_apply] dsimp [toMatrix, PEquiv.trans] cases f i <;> simp @@ -110,19 +146,18 @@ theorem toMatrix_bot [DecidableEq n] [Zero α] [One α] : theorem toMatrix_injective [DecidableEq n] [MonoidWithZero α] [Nontrivial α] : Function.Injective (@toMatrix m n α _ _ _) := by - classical - intro f g - refine not_imp_not.1 ?_ - simp only [Matrix.ext_iff.symm, toMatrix_apply, PEquiv.ext_iff, not_forall, exists_imp] - intro i hi - use i - cases' hf : f i with fi - · cases' hg : g i with gi - · rw [hf, hg] at hi; exact (hi rfl).elim - · use gi - simp - · use fi - simp [hf.symm, Ne.symm hi] + intro f g + refine not_imp_not.1 ?_ + simp only [Matrix.ext_iff.symm, toMatrix_apply, PEquiv.ext_iff, not_forall, exists_imp] + intro i hi + use i + cases' hf : f i with fi + · cases' hg : g i with gi + · rw [hf, hg] at hi; exact (hi rfl).elim + · use gi + simp + · use fi + simp [hf.symm, Ne.symm hi] theorem toMatrix_swap [DecidableEq n] [Ring α] (i j : n) : (Equiv.swap i j).toPEquiv.toMatrix = @@ -152,8 +187,16 @@ theorem single_mul_single_right [Fintype n] [Fintype k] [DecidableEq n] [Decidab rw [← Matrix.mul_assoc, single_mul_single] /-- We can also define permutation matrices by permuting the rows of the identity matrix. -/ -theorem equiv_toPEquiv_toMatrix [DecidableEq n] [Zero α] [One α] (σ : Equiv n n) (i j : n) : - σ.toPEquiv.toMatrix i j = (1 : Matrix n n α) (σ i) j := - if_congr Option.some_inj rfl rfl +theorem toMatrix_toPEquiv_eq [DecidableEq n] [Zero α] [One α] (σ : Equiv.Perm n) : + σ.toPEquiv.toMatrix = (1 : Matrix n n α).submatrix σ id := + Matrix.ext fun _ _ => if_congr Option.some_inj rfl rfl + +@[deprecated (since := "2025-01-27")] alias equiv_toPEquiv_toMatrix := toMatrix_toPEquiv_eq + +@[simp] +lemma map_toMatrix [DecidableEq n] [NonAssocSemiring α] [NonAssocSemiring β] + (f : α →+* β) (σ : m ≃. n) : σ.toMatrix.map f = σ.toMatrix := by + ext i j + simp end PEquiv diff --git a/Mathlib/Data/Matroid/Basic.lean b/Mathlib/Data/Matroid/Basic.lean index 60a143ac60a47d..4bef987a38c329 100644 --- a/Mathlib/Data/Matroid/Basic.lean +++ b/Mathlib/Data/Matroid/Basic.lean @@ -217,6 +217,9 @@ namespace Matroid variable {α : Type*} {M : Matroid α} +instance (M : Matroid α) : Nonempty {B // M.Base B} := + nonempty_subtype.2 M.exists_base + /-- Typeclass for a matroid having finite ground set. Just a wrapper for `M.E.Finite`-/ @[mk_iff] protected class Finite (M : Matroid α) : Prop where /-- The ground set is finite -/ diff --git a/Mathlib/Data/Matroid/Circuit.lean b/Mathlib/Data/Matroid/Circuit.lean index d8f568c93c09e2..d074210e882958 100644 --- a/Mathlib/Data/Matroid/Circuit.lean +++ b/Mathlib/Data/Matroid/Circuit.lean @@ -16,6 +16,24 @@ In matroids arising from graphs, circuits correspond to graphical cycles. # Main Declarations * `Matroid.Circuit M C` means that `C` is minimally dependent in `M`. +* For an `Indep`endent set `I` whose closure contains an element `e ∉ I`, + `Matroid.fundCircuit M e I` is the unique circuit contained in `insert e I`. +* `Matroid.Indep.fundCircuit_circuit` states that `Matroid.fundCircuit M e I` is indeed a circuit. +* `Circuit.eq_fundCircuit_of_subset` states that `Matroid.fundCircuit M e I` is the + unique circuit contained in `insert e I`. +* `Matroid.dep_iff_superset_circuit` states that the dependent subsets of the ground set + are precisely those that contain a circuit. +* `Matroid.ext_circuit` : a matroid is determined by its collection of circuits. + +# Implementation Details + +Since `Matroid.fundCircuit M e I` is only sensible if `I` is independent and `e ∈ M.closure I \ I`, +to avoid hypotheses being explicitly included in the definition, +junk values need to be chosen if either hypothesis fails. +The definition is chosen so that the junk values satisfy +`M.fundCircuit e I = {e}` for `e ∈ I` or `e ∉ M.E` and +`M.fundCircuit e I = insert e I` if `e ∈ M.E \ M.closure I`. +These make the useful statement `e ∈ M.fundCircuit e I ⊆ insert e I` true unconditionally. -/ variable {α : Type*} {M : Matroid α} {C C' I X Y R : Set α} {e f x y : α} @@ -87,6 +105,10 @@ lemma Circuit.not_ssubset (hC : M.Circuit C) (hC' : M.Circuit C') : ¬C' ⊂ C : lemma Circuit.eq_of_subset_circuit (hC : M.Circuit C) (hC' : M.Circuit C') (h : C ⊆ C') : C = C' := hC'.eq_of_dep_subset hC.dep h +lemma Circuit.eq_of_superset_circuit (hC : M.Circuit C) (hC' : M.Circuit C') (h : C' ⊆ C) : + C = C' := + (hC'.eq_of_subset_circuit hC h).symm + lemma circuit_iff_dep_forall_diff_singleton_indep : M.Circuit C ↔ M.Dep C ∧ ∀ e ∈ C, M.Indep (C \ {e}) := by wlog hCE : C ⊆ M.E @@ -155,8 +177,7 @@ lemma Circuit.mem_closure_diff_singleton_of_mem (hC : M.Circuit C) (heC : e ∈ /-! ### Restriction -/ -lemma Circuit.circuit_restrict_of_subset (hC : M.Circuit C) (hCR : C ⊆ R) : - (M ↾ R).Circuit C := by +lemma Circuit.circuit_restrict_of_subset (hC : M.Circuit C) (hCR : C ⊆ R) : (M ↾ R).Circuit C := by simp_rw [circuit_iff, restrict_dep_iff, dep_iff, and_imp] at * exact ⟨⟨hC.1.1, hCR⟩, fun I hI _ hIC ↦ hC.2 hI (hIC.trans hC.1.2) hIC⟩ @@ -166,4 +187,167 @@ lemma restrict_circuit_iff (hR : R ⊆ M.E := by aesop_mat) : simp_rw [circuit_iff, restrict_dep_iff, and_imp, dep_iff] exact fun hC hCR h ↦ ⟨⟨⟨hC,hCR.trans hR⟩,fun I hI hIC ↦ h hI.1 (hIC.trans hCR) hIC⟩,hCR⟩ +/-! ### Fundamental Circuits -/ + +/-- For an independent set `I` and some `e ∈ M.closure I \ I`, +`M.fundCircuit e I` is the unique circuit contained in `insert e I`. +For the fact that this is a circuit, see `Matroid.Indep.fundCircuit_circuit`, +and the fact that it is unique, see `Matroid.Circuit.eq_fundCircuit_of_subset`. +Has the junk value `{e}` if `e ∈ I` or `e ∉ M.E`, and `insert e I` if `e ∈ M.E \ M.closure I`. -/ +def fundCircuit (M : Matroid α) (e : α) (I : Set α) : Set α := + insert e (I ∩ ⋂₀ {J | J ⊆ I ∧ M.closure {e} ⊆ M.closure J}) + +lemma fundCircuit_eq_sInter (he : e ∈ M.closure I) : + M.fundCircuit e I = insert e (⋂₀ {J | J ⊆ I ∧ e ∈ M.closure J}) := by + rw [fundCircuit] + simp_rw [closure_subset_closure_iff_subset_closure + (show {e} ⊆ M.E by simpa using mem_ground_of_mem_closure he), singleton_subset_iff] + rw [inter_eq_self_of_subset_right (sInter_subset_of_mem (by simpa))] + +lemma fundCircuit_subset_insert (M : Matroid α) (e : α) (I : Set α) : + M.fundCircuit e I ⊆ insert e I := + insert_subset_insert inter_subset_left + +lemma fundCircuit_subset_ground (he : e ∈ M.E) (hI : I ⊆ M.E := by aesop_mat) : + M.fundCircuit e I ⊆ M.E := + (M.fundCircuit_subset_insert e I).trans (insert_subset he hI) + +lemma mem_fundCircuit (M : Matroid α) (e : α) (I : Set α) : e ∈ fundCircuit M e I := + mem_insert .. + +lemma fundCircuit_diff_eq_inter (M : Matroid α) (heI : e ∉ I) : + (M.fundCircuit e I) \ {e} = (M.fundCircuit e I) ∩ I := + (subset_inter diff_subset (by simp [fundCircuit_subset_insert])).antisymm + (subset_diff_singleton inter_subset_left (by simp [heI])) + +/-- The fundamental circuit of `e` and `X` has the junk value `{e}` if `e ∈ X` -/ +lemma fundCircuit_eq_of_mem (heX : e ∈ X) : M.fundCircuit e X = {e} := by + suffices h : ∀ a ∈ X, (∀ t ⊆ X, M.closure {e} ⊆ M.closure t → a ∈ t) → a = e by + simpa [subset_antisymm_iff, fundCircuit] + exact fun b hbX h ↦ h _ (singleton_subset_iff.2 heX) Subset.rfl + +lemma fundCircuit_eq_of_not_mem_ground (heX : e ∉ M.E) : M.fundCircuit e X = {e} := by + suffices h : ∀ a ∈ X, (∀ t ⊆ X, M.closure {e} ⊆ M.closure t → a ∈ t) → a = e by + simpa [subset_antisymm_iff, fundCircuit] + simp_rw [← M.closure_inter_ground {e}, singleton_inter_eq_empty.2 heX] + exact fun a haX h ↦ by simpa using h ∅ (empty_subset X) rfl.subset + +lemma Indep.fundCircuit_circuit (hI : M.Indep I) (hecl : e ∈ M.closure I) (heI : e ∉ I) : + M.Circuit (M.fundCircuit e I) := by + have aux : ⋂₀ {J | J ⊆ I ∧ e ∈ M.closure J} ⊆ I := sInter_subset_of_mem (by simpa) + rw [fundCircuit_eq_sInter hecl] + refine (hI.subset aux).insert_circuit_of_forall ?_ ?_ ?_ + · simp [show ∃ x ⊆ I, e ∈ M.closure x ∧ e ∉ x from ⟨I, by simp [hecl, heI]⟩] + · rw [hI.closure_sInter_eq_biInter_closure_of_forall_subset ⟨I, by simpa⟩ (by simp +contextual)] + simp + simp only [mem_sInter, mem_setOf_eq, and_imp] + exact fun f hf hecl ↦ (hf _ (diff_subset.trans aux) hecl).2 rfl + +lemma Indep.mem_fundCircuit_iff (hI : M.Indep I) (hecl : e ∈ M.closure I) (heI : e ∉ I) : + x ∈ M.fundCircuit e I ↔ M.Indep (insert e I \ {x}) := by + obtain rfl | hne := eq_or_ne x e + · simp [hI.diff, mem_fundCircuit] + suffices (∀ t ⊆ I, e ∈ M.closure t → x ∈ t) ↔ e ∉ M.closure (I \ {x}) by + simpa [fundCircuit_eq_sInter hecl, hne, ← insert_diff_singleton_comm hne.symm, + (hI.diff _).insert_indep_iff, mem_ground_of_mem_closure hecl, heI] + refine ⟨fun h hecl ↦ (h _ diff_subset hecl).2 rfl, fun h J hJ heJ ↦ by_contra fun hxJ ↦ h ?_⟩ + exact M.closure_subset_closure (subset_diff_singleton hJ hxJ) heJ + +lemma Base.fundCircuit_circuit {B : Set α} (hB : M.Base B) (hxE : x ∈ M.E) (hxB : x ∉ B) : + M.Circuit (M.fundCircuit x B) := + hB.indep.fundCircuit_circuit (by rwa [hB.closure_eq]) hxB + +/-- For `I` independent, `M.fundCircuit e I` is the only circuit contained in `insert e I`. -/ +lemma Circuit.eq_fundCircuit_of_subset (hC : M.Circuit C) (hI : M.Indep I) (hCss : C ⊆ insert e I) : + C = M.fundCircuit e I := by + obtain hCI | ⟨heC, hCeI⟩ := subset_insert_iff.1 hCss + · exact (hC.not_indep (hI.subset hCI)).elim + suffices hss : M.fundCircuit e I ⊆ C by + refine hC.eq_of_superset_circuit (hI.fundCircuit_circuit ?_ fun heI ↦ ?_) hss + · rw [hI.mem_closure_iff] + exact .inl (hC.dep.superset hCss (insert_subset (hC.subset_ground heC) hI.subset_ground)) + exact hC.not_indep (hI.subset (hCss.trans (by simp [heI]))) + rw [fundCircuit_eq_sInter <| + M.closure_subset_closure hCeI <| hC.mem_closure_diff_singleton_of_mem heC] + refine insert_subset heC <| (sInter_subset_of_mem (t := C \ {e}) ?_).trans diff_subset + simp [hCss, hC.mem_closure_diff_singleton_of_mem heC] + +lemma fundCircuit_restrict {R : Set α} (hIR : I ⊆ R) (heR : e ∈ R) (hR : R ⊆ M.E) : + (M ↾ R).fundCircuit e I = M.fundCircuit e I := by + simp_rw [fundCircuit, M.restrict_closure_eq (R := R) (X := {e}) (by simpa)] + refine subset_antisymm (insert_subset_insert (inter_subset_inter_right _ ?_)) + (insert_subset_insert (inter_subset_inter_right _ ?_)) + · refine subset_sInter fun J ⟨hJI, heJ⟩ ↦ sInter_subset_of_mem ⟨hJI, ?_⟩ + simp only [restrict_closure_eq', union_subset_iff, subset_union_right, and_true] + refine (inter_subset_inter_left _ ?_).trans subset_union_left + rwa [inter_eq_self_of_subset_left (hJI.trans hIR)] + refine subset_sInter fun J ⟨hJI, heJ⟩ ↦ sInter_subset_of_mem + ⟨hJI, M.closure_subset_closure_of_subset_closure ?_⟩ + rw [restrict_closure_eq _ (hJI.trans hIR) hR] at heJ + simp only [subset_inter_iff, inter_subset_right, and_true] at heJ + exact subset_trans (by simpa [M.mem_closure_of_mem' (mem_singleton e) (hR heR)]) heJ + +@[simp] lemma fundCircuit_restrict_univ (M : Matroid α) : + (M ↾ univ).fundCircuit e I = M.fundCircuit e I := by + have aux (A B) : M.closure A ⊆ B ∪ univ \ M.E ↔ M.closure A ⊆ B := by + refine ⟨fun h ↦ ?_, fun h ↦ h.trans subset_union_left⟩ + refine (subset_inter h (M.closure_subset_ground A)).trans ?_ + simp [union_inter_distrib_right] + simp [fundCircuit, aux] + +/-! ### Dependence -/ + +lemma Dep.exists_circuit_subset (hX : M.Dep X) : ∃ C, C ⊆ X ∧ M.Circuit C := by + obtain ⟨I, hI⟩ := M.exists_basis X + obtain ⟨e, heX, heI⟩ := exists_of_ssubset + (hI.subset.ssubset_of_ne (by rintro rfl; exact hI.indep.not_dep hX)) + exact ⟨M.fundCircuit e I, (M.fundCircuit_subset_insert e I).trans (insert_subset heX hI.subset), + hI.indep.fundCircuit_circuit (hI.subset_closure heX) heI⟩ + +lemma dep_iff_superset_circuit (hX : X ⊆ M.E := by aesop_mat) : + M.Dep X ↔ ∃ C, C ⊆ X ∧ M.Circuit C := + ⟨Dep.exists_circuit_subset, fun ⟨C, hCX, hC⟩ ↦ hC.dep.superset hCX⟩ + +/-- A version of `Matroid.dep_iff_superset_circuit` that has the supportedness hypothesis +as part of the equivalence, rather than a hypothesis. -/ +lemma dep_iff_superset_circuit' : M.Dep X ↔ (∃ C, C ⊆ X ∧ M.Circuit C) ∧ X ⊆ M.E := + ⟨fun h ↦ ⟨h.exists_circuit_subset, h.subset_ground⟩, fun ⟨⟨C, hCX, hC⟩, h⟩ ↦ hC.dep.superset hCX⟩ + +/-- A version of `Matroid.indep_iff_forall_subset_not_circuit` that has the supportedness hypothesis +as part of the equivalence, rather than a hypothesis. -/ +lemma indep_iff_forall_subset_not_circuit' : + M.Indep I ↔ (∀ C, C ⊆ I → ¬M.Circuit C) ∧ I ⊆ M.E := by + simp_rw [indep_iff_not_dep, dep_iff_superset_circuit'] + aesop + +lemma indep_iff_forall_subset_not_circuit (hI : I ⊆ M.E := by aesop_mat) : + M.Indep I ↔ ∀ C, C ⊆ I → ¬M.Circuit C := by + rw [indep_iff_forall_subset_not_circuit', and_iff_left hI] + +/-! ### Extensionality -/ + +lemma ext_circuit {M₁ M₂ : Matroid α} (hE : M₁.E = M₂.E) + (h : ∀ ⦃C⦄, C ⊆ M₁.E → (M₁.Circuit C ↔ M₂.Circuit C)) : M₁ = M₂ := by + have h' {C} : M₁.Circuit C ↔ M₂.Circuit C := + (em (C ⊆ M₁.E)).elim (h (C := C)) (fun hC ↦ iff_of_false (mt Circuit.subset_ground hC) + (mt Circuit.subset_ground fun hss ↦ hC (hss.trans_eq hE.symm))) + refine ext_indep hE fun I hI ↦ ?_ + simp_rw [indep_iff_forall_subset_not_circuit hI, h', + indep_iff_forall_subset_not_circuit (hI.trans_eq hE)] + +/-- A stronger version of `Matroid.ext_circuit`: +two matroids on the same ground set are equal if no circuit of one is independent in the other. -/ +lemma ext_circuit_not_indep {M₁ M₂ : Matroid α} (hE : M₁.E = M₂.E) + (h₁ : ∀ C, M₁.Circuit C → ¬ M₂.Indep C) (h₂ : ∀ C, M₂.Circuit C → ¬ M₁.Indep C) : + M₁ = M₂ := by + refine ext_circuit hE fun C hCE ↦ ⟨fun hC ↦ ?_, fun hC ↦ ?_⟩ + · obtain ⟨C', hC'C, hC'⟩ := ((not_indep_iff (by rwa [← hE])).1 (h₁ C hC)).exists_circuit_subset + rwa [← hC.eq_of_not_indep_subset (h₂ C' hC') hC'C] + obtain ⟨C', hC'C, hC'⟩ := ((not_indep_iff hCE).1 (h₂ C hC)).exists_circuit_subset + rwa [← hC.eq_of_not_indep_subset (h₁ C' hC') hC'C] + +lemma ext_iff_circuit {M₁ M₂ : Matroid α} : + M₁ = M₂ ↔ M₁.E = M₂.E ∧ ∀ C, M₁.Circuit C ↔ M₂.Circuit C := + ⟨fun h ↦ by simp [h], fun h ↦ ext_circuit h.1 fun C hC ↦ h.2 (C := C)⟩ + end Matroid diff --git a/Mathlib/Data/Matroid/Rank/Cardinal.lean b/Mathlib/Data/Matroid/Rank/Cardinal.lean index b12a9f56efadbb..b08d04f6201229 100644 --- a/Mathlib/Data/Matroid/Rank/Cardinal.lean +++ b/Mathlib/Data/Matroid/Rank/Cardinal.lean @@ -3,46 +3,294 @@ Copyright (c) 2025 Peter Nelson and Junyan Xu. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Peter Nelson, Junyan Xu -/ -import Mathlib.Data.Matroid.Restrict +import Mathlib.Data.Matroid.Closure +import Mathlib.Data.Matroid.Map import Mathlib.SetTheory.Cardinal.Arithmetic /-! -# Invariance of cardinality of bases of a finitary matroid +# Cardinal-valued rank In a finitary matroid, all bases have the same cardinality. -In fact, something stronger holds: if `B` and `B'` are bases, then `#(B \ B') = #(B' \ B)`. -This file provides proofs of these facts, -as well as variants when each of `B` and `B'` is a `Basis` or `Basis'` of some common set `X`. +In fact, something stronger holds: if `I` and `J` are both bases for a set `X`, +then `#(I \ J) = #(J \ I)` and (consequently) `#I = #J`. +This file introduces a typeclass `InvariantCardinalRank` that applies to any matroid +such that this property holds for all `I`, `J` and `X`. -Some extra assumption like `Finitary` is necessary for these to be true, +A matroid satisfying this condition has a well-defined cardinality-valued rank function, +both for itself and all its minors. + +# Main Declarations + +* `Matroid.InvariantCardinalRank` : a typeclass capturing the idea that a matroid and all its minors + have a well-behaved cardinal-valued rank function. +* `Matroid.cRank M` is the supremum of the cardinalities of the bases of matroid `M`. +* `Matroid.cRk M X` is the supremum of the cardinalities of the bases of a set `X` in a matroid `M`. +* `invariantCardinalRank_of_finitary` is the instance + showing that `Finitary` matroids are `InvariantCardinalRank`. +* `cRk_inter_add_cRk_union_le` states that cardinal rank is submodular. + +# Notes + +It is not the case that all matroids are `InvariantCardinalRank`, since the equicardinality of bases in general matroids is independent of ZFC -(see the docstring of `Data.Matroid.Basic`). -Lemmas like `Matroid.Base.cardinalMk_diff_comm_of_finitary` -become true for all matroids if they are weakened by replacing `Cardinal.mk` -with the cruder `ℕ∞`-valued `encard`; see, for example, `Matroid.Base.encard_diff_comm`. +(see the module docstring of `Mathlib.Data.Matroid.Basic`). +Lemmas like `Matroid.Base.cardinalMk_diff_comm` become true for all matroids +only if they are weakened by replacing `Cardinal.mk` +with the cruder `ℕ∞`-valued `Set.encard`; see, for example, `Matroid.Base.encard_diff_comm`. + +# Implementation Details + +Since the functions `cRank` and `cRk` are defined as suprema, +independently of the `Matroid.InvariantCardinalRank` typeclass, +they are well-defined for all matroids. +However, for matroids that do not satisfy `InvariantCardinalRank`, they are badly behaved. +For instance, in general `cRk` is not submodular, +and its value may differ on a set `X` and the closure of `X`. +We state and prove theorems without `InvariantCardinalRank` whenever possible, +which sometime makes their proofs longer than they would be with the instance. -TODO +# TODO + +* Higgs' theorem : if the generalized continuum hypothesis holds, + then every matroid is `InvariantCardinalRank`. -* Higg's theorem that, if the generalized continuum hypothesis holds, - all bases of any matroid are equicardinal. -* API for a `Cardinal`-valued rank function. -/ -variable {α : Type*} {M : Matroid α} {I J B B' X : Set α} [M.Finitary] +universe u v + +variable {α : Type u} {β : Type v} {f : α → β} {M : Matroid α} {I J B B' X Y : Set α} open Cardinal Set namespace Matroid -theorem Base.cardinalMk_diff_comm_of_finitary (hB : M.Base B) (hB' : M.Base B') : - #(B \ B' : Set α) = #(B' \ B : Set α) := by - wlog hge : #(B' \ B : Set α) ≤ #(B \ B' : Set α) with aux - · exact (aux hB' hB (not_le.1 hge).le).symm +section Rank + +variable {κ : Cardinal} + +/-- The rank (supremum of the cardinalities of bases) of a matroid `M` as a `Cardinal`. -/ +noncomputable def cRank (M : Matroid α) := ⨆ (B : {B // M.Base B}), #B + +/-- The rank (supremum of the cardinalities of bases) of a set `X` in a matroid `M`, +as a `Cardinal`. -/ +noncomputable def cRk (M : Matroid α) (X : Set α) := (M ↾ X).cRank + +theorem Base.cardinalMk_le_cRank (hB : M.Base B) : #B ≤ M.cRank := + le_ciSup (f := fun (B : {B // M.Base B}) ↦ #(B.1)) (bddAbove_range _) ⟨B, hB⟩ + +theorem Basis'.cardinalMk_le_cRk (hIX : M.Basis' I X) : #I ≤ M.cRk X := + (base_restrict_iff'.2 hIX).cardinalMk_le_cRank + +theorem Basis.cardinalMk_le_cRk (hIX : M.Basis I X) : #I ≤ M.cRk X := + hIX.basis'.cardinalMk_le_cRk + +theorem cRank_le_iff : M.cRank ≤ κ ↔ ∀ ⦃B⦄, M.Base B → #B ≤ κ := + ⟨fun h _ hB ↦ (hB.cardinalMk_le_cRank.trans h), fun h ↦ ciSup_le fun ⟨_, hB⟩ ↦ h hB⟩ + +theorem cRk_le_iff : M.cRk X ≤ κ ↔ ∀ ⦃I⦄, M.Basis' I X → #I ≤ κ := by + simp_rw [cRk, cRank_le_iff, base_restrict_iff'] + +theorem Indep.cardinalMk_le_cRk_of_subset (hI : M.Indep I) (hIX : I ⊆ X) : #I ≤ M.cRk X := + let ⟨_, hJ, hIJ⟩ := hI.subset_basis'_of_subset hIX + (mk_le_mk_of_subset hIJ).trans hJ.cardinalMk_le_cRk + +theorem cRk_le_cardinalMk (M : Matroid α) (X : Set α) : M.cRk X ≤ #X := + ciSup_le fun ⟨_, hI⟩ ↦ mk_le_mk_of_subset hI.subset_ground + +@[simp] theorem cRk_ground (M : Matroid α) : M.cRk M.E = M.cRank := by + rw [cRk, restrict_ground_eq_self] + +@[simp] theorem cRank_restrict (M : Matroid α) (X : Set α) : (M ↾ X).cRank = M.cRk X := rfl + +theorem cRk_mono (M : Matroid α) : Monotone M.cRk := by + simp only [Monotone, le_eq_subset, cRk_le_iff] + intro X Y hXY I hIX + obtain ⟨J, hJ, hIJ⟩ := hIX.indep.subset_basis'_of_subset (hIX.subset.trans hXY) + exact (mk_le_mk_of_subset hIJ).trans hJ.cardinalMk_le_cRk + +theorem cRk_le_of_subset (M : Matroid α) (hXY : X ⊆ Y) : M.cRk X ≤ M.cRk Y := + M.cRk_mono hXY + +@[simp] theorem cRk_inter_ground (M : Matroid α) (X : Set α) : M.cRk (X ∩ M.E) = M.cRk X := + (M.cRk_le_of_subset inter_subset_left).antisymm <| cRk_le_iff.2 + fun _ h ↦ h.basis_inter_ground.cardinalMk_le_cRk + +theorem cRk_restrict_subset (M : Matroid α) (hYX : Y ⊆ X) : (M ↾ X).cRk Y = M.cRk Y := by + have aux : ∀ ⦃I⦄, M.Basis' I Y ↔ (M ↾ X).Basis' I Y := by + simp_rw [basis'_restrict_iff, inter_eq_self_of_subset_left hYX, iff_self_and] + exact fun I h ↦ h.subset.trans hYX + simp_rw [le_antisymm_iff, cRk_le_iff] + exact ⟨fun I hI ↦ (aux.2 hI).cardinalMk_le_cRk, fun I hI ↦ (aux.1 hI).cardinalMk_le_cRk⟩ + +theorem cRk_restrict (M : Matroid α) (X Y : Set α) : (M ↾ X).cRk Y = M.cRk (X ∩ Y) := by + rw [← cRk_inter_ground, restrict_ground_eq, cRk_restrict_subset _ inter_subset_right, + inter_comm] + +theorem Indep.cRk_eq_cardinalMk (hI : M.Indep I) : #I = M.cRk I := + (M.cRk_le_cardinalMk I).antisymm' (hI.basis_self.cardinalMk_le_cRk) + +@[simp] theorem cRk_map_image_lift (M : Matroid α) (hf : InjOn f M.E) (X : Set α) + (hX : X ⊆ M.E := by aesop_mat) : lift.{u,v} ((M.map f hf).cRk (f '' X)) = lift (M.cRk X) := by + nth_rw 1 [cRk, cRank, le_antisymm_iff, lift_iSup (bddAbove_range _), cRk, cRank, cRk, cRank] + nth_rw 2 [lift_iSup (bddAbove_range _)] + simp only [ciSup_le_iff (bddAbove_range _), ge_iff_le, Subtype.forall, base_restrict_iff', + basis'_iff_basis hX, basis'_iff_basis (show f '' X ⊆ (M.map f hf).E from image_mono hX)] + refine ⟨fun I hI ↦ ?_, fun I hI ↦ ?_⟩ + · obtain ⟨I, X', hIX, rfl, hXX'⟩ := map_basis_iff'.1 hI + rw [mk_image_eq_of_injOn_lift _ _ (hf.mono hIX.indep.subset_ground), lift_le] + obtain rfl : X = X' := by rwa [hf.image_eq_image_iff hX hIX.subset_ground] at hXX' + exact hIX.cardinalMk_le_cRk + rw [← mk_image_eq_of_injOn_lift _ _ (hf.mono hI.indep.subset_ground), lift_le] + exact (hI.map hf).cardinalMk_le_cRk + +@[simp] theorem cRk_map_image {β : Type u} {f : α → β} (M : Matroid α) (hf : InjOn f M.E) + (X : Set α) (hX : X ⊆ M.E := by aesop_mat) : (M.map f hf).cRk (f '' X) = M.cRk X := + lift_inj.1 <| M.cRk_map_image_lift .. + +theorem cRk_map_eq {β : Type u} {f : α → β} {X : Set β} (M : Matroid α) (hf : InjOn f M.E) : + (M.map f hf).cRk X = M.cRk (f ⁻¹' X) := by + rw [← M.cRk_inter_ground, ← M.cRk_map_image hf _, image_preimage_inter, ← map_ground _ _ hf, + cRk_inter_ground] + +@[simp] theorem cRk_comap_lift (M : Matroid β) (f : α → β) (X : Set α) : + lift.{v,u} ((M.comap f).cRk X) = lift (M.cRk (f '' X)) := by + nth_rw 1 [cRk, cRank, le_antisymm_iff, lift_iSup (bddAbove_range _), cRk, cRank, cRk, cRank] + nth_rw 2 [lift_iSup (bddAbove_range _)] + simp only [ciSup_le_iff (bddAbove_range _), ge_iff_le, Subtype.forall, base_restrict_iff', + comap_basis'_iff, and_imp] + refine ⟨fun I hI hfI hIX ↦ ?_, fun I hIX ↦ ?_⟩ + · rw [← mk_image_eq_of_injOn_lift _ _ hfI, lift_le] + exact hI.cardinalMk_le_cRk + obtain ⟨I₀, hI₀X, rfl, hfI₀⟩ := show ∃ I₀ ⊆ X, f '' I₀ = I ∧ InjOn f I₀ by + obtain ⟨I₀, hI₀ss, hbij⟩ := exists_subset_bijOn (f ⁻¹' I ∩ X) f + refine ⟨I₀, hI₀ss.trans inter_subset_right, ?_, hbij.injOn⟩ + rw [hbij.image_eq, image_preimage_inter, inter_eq_self_of_subset_left hIX.subset] + rw [mk_image_eq_of_injOn_lift _ _ hfI₀, lift_le] + exact Basis'.cardinalMk_le_cRk <| comap_basis'_iff.2 ⟨hIX, hfI₀, hI₀X⟩ + +@[simp] theorem cRk_comap {β : Type u} (M : Matroid β) (f : α → β) (X : Set α) : + (M.comap f).cRk X = M.cRk (f '' X) := + lift_inj.1 <| M.cRk_comap_lift .. + +end Rank + +section Invariant + +/-- A class stating that cardinality-valued rank is well-defined +(i.e. all bases are equicardinal) for a matroid `M` and its minors. +Notably, this holds for `Finitary` matroids; see `Matroid.invariantCardinalRank_of_finitary`. -/ +@[mk_iff] +class InvariantCardinalRank (M : Matroid α) : Prop where + forall_card_basis_diff : + ∀ ⦃I J X⦄, M.Basis I X → M.Basis J X → #(I \ J : Set α) = #(J \ I : Set α) + +variable [InvariantCardinalRank M] + +theorem Basis.cardinalMk_diff_comm (hIX : M.Basis I X) (hJX : M.Basis J X) : + #(I \ J : Set α) = #(J \ I : Set α) := + InvariantCardinalRank.forall_card_basis_diff hIX hJX + +theorem Basis'.cardinalMk_diff_comm (hIX : M.Basis' I X) (hJX : M.Basis' J X) : + #(I \ J : Set α) = #(J \ I : Set α) := + hIX.basis_inter_ground.cardinalMk_diff_comm hJX.basis_inter_ground + +theorem Base.cardinalMk_diff_comm (hB : M.Base B) (hB' : M.Base B') : + #(B \ B' : Set α) = #(B' \ B : Set α) := + hB.basis_ground.cardinalMk_diff_comm hB'.basis_ground + +theorem Basis.cardinalMk_eq (hIX : M.Basis I X) (hJX : M.Basis J X) : #I = #J := by + rw [← diff_union_inter I J, + mk_union_of_disjoint (disjoint_sdiff_left.mono_right inter_subset_right), + hIX.cardinalMk_diff_comm hJX, + ← mk_union_of_disjoint (disjoint_sdiff_left.mono_right inter_subset_left), + inter_comm, diff_union_inter] + +theorem Basis'.cardinalMk_eq (hIX : M.Basis' I X) (hJX : M.Basis' J X) : #I = #J := + hIX.basis_inter_ground.cardinalMk_eq hJX.basis_inter_ground + +theorem Base.cardinalMk_eq (hB : M.Base B) (hB' : M.Base B') : #B = #B' := + hB.basis_ground.cardinalMk_eq hB'.basis_ground + +theorem Indep.cardinalMk_le_base (hI : M.Indep I) (hB : M.Base B) : #I ≤ #B := + have ⟨_B', hB', hIB'⟩ := hI.exists_base_superset + hB'.cardinalMk_eq hB ▸ mk_le_mk_of_subset hIB' + +theorem Indep.cardinalMk_le_basis' (hI : M.Indep I) (hJ : M.Basis' J X) (hIX : I ⊆ X) : + #I ≤ #J := + have ⟨_J', hJ', hIJ'⟩ := hI.subset_basis'_of_subset hIX + hJ'.cardinalMk_eq hJ ▸ mk_le_mk_of_subset hIJ' + +theorem Indep.cardinalMk_le_basis (hI : M.Indep I) (hJ : M.Basis J X) (hIX : I ⊆ X) : + #I ≤ #J := + hI.cardinalMk_le_basis' hJ.basis' hIX + +theorem Base.cardinalMk_eq_cRank (hB : M.Base B) : #B = M.cRank := by + have hrw : ∀ B' : {B : Set α // M.Base B}, #B' = #B := fun B' ↦ B'.2.cardinalMk_eq hB + simp [cRank, hrw] + +/-- Restrictions of matroids with cardinal rank functions have cardinal rank functions- -/ +instance invariantCardinalRank_restrict [InvariantCardinalRank M] : + InvariantCardinalRank (M ↾ X) := by + refine ⟨fun I J Y hI hJ ↦ ?_⟩ + rw [basis_restrict_iff'] at hI hJ + exact hI.1.cardinalMk_diff_comm hJ.1 + +theorem Basis'.cardinalMk_eq_cRk (hIX : M.Basis' I X) : #I = M.cRk X := by + rw [cRk, (base_restrict_iff'.2 hIX).cardinalMk_eq_cRank] + +theorem Basis.cardinalMk_eq_cRk (hIX : M.Basis I X) : #I = M.cRk X := + hIX.basis'.cardinalMk_eq_cRk + +@[simp] theorem cRk_closure (M : Matroid α) [InvariantCardinalRank M] (X : Set α) : + M.cRk (M.closure X) = M.cRk X := by + obtain ⟨I, hI⟩ := M.exists_basis' X + rw [← hI.basis_closure_right.cardinalMk_eq_cRk, ← hI.cardinalMk_eq_cRk] + +theorem cRk_closure_congr (hXY : M.closure X = M.closure Y) : M.cRk X = M.cRk Y := by + rw [← cRk_closure, hXY, cRk_closure] + +variable (M : Matroid α) [InvariantCardinalRank M] (e : α) (X Y : Set α) + +@[simp] theorem cRk_union_closure_right_eq : M.cRk (X ∪ M.closure Y) = M.cRk (X ∪ Y) := + M.cRk_closure_congr (M.closure_union_closure_right_eq _ _) + +@[simp] theorem cRk_union_closure_left_eq : M.cRk (M.closure X ∪ Y) = M.cRk (X ∪ Y) := + M.cRk_closure_congr (M.closure_union_closure_left_eq _ _) + +@[simp] theorem cRk_insert_closure_eq : M.cRk (insert e (M.closure X)) = M.cRk (insert e X) := by + rw [← union_singleton, cRk_union_closure_left_eq, union_singleton] + +theorem cRk_union_closure_eq : M.cRk (M.closure X ∪ M.closure Y) = M.cRk (X ∪ Y) := by + simp + +/-- The `Cardinal` rank function is submodular. -/ +theorem cRk_inter_add_cRk_union_le : M.cRk (X ∩ Y) + M.cRk (X ∪ Y) ≤ M.cRk X + M.cRk Y := by + obtain ⟨Ii, hIi⟩ := M.exists_basis' (X ∩ Y) + obtain ⟨IX, hIX, hIX'⟩ := + hIi.indep.subset_basis'_of_subset (hIi.subset.trans inter_subset_left) + obtain ⟨IY, hIY, hIY'⟩ := + hIi.indep.subset_basis'_of_subset (hIi.subset.trans inter_subset_right) + rw [← cRk_union_closure_eq, ← hIX.closure_eq_closure, ← hIY.closure_eq_closure, + cRk_union_closure_eq, ← hIi.cardinalMk_eq_cRk, ← hIX.cardinalMk_eq_cRk, + ← hIY.cardinalMk_eq_cRk, ← mk_union_add_mk_inter, add_comm] + exact add_le_add (M.cRk_le_cardinalMk _) (mk_le_mk_of_subset (subset_inter hIX' hIY')) + +end Invariant + +section Instances + +/-- `Finitary` matroids have a cardinality-valued rank function. -/ +instance invariantCardinalRank_of_finitary [Finitary M] : InvariantCardinalRank M := by + suffices aux : ∀ ⦃B B'⦄ ⦃N : Matroid α⦄, Finitary N → N.Base B → N.Base B' → + #(B \ B' : Set α) ≤ #(B' \ B : Set α) from + ⟨fun I J X hI hJ ↦ (aux (restrict_finitary X) hI.base_restrict hJ.base_restrict).antisymm + (aux (restrict_finitary X) hJ.base_restrict hI.base_restrict)⟩ + intro B B' N hfin hB hB' by_cases h : (B' \ B).Finite · rw [← cast_ncard h, ← cast_ncard, hB.ncard_diff_comm hB'] - exact (diff_finite_comm hB' hB).mp h + exact (hB'.diff_finite_comm hB).mp h rw [← Set.Infinite, ← infinite_coe_iff] at h - have (a : α) (ha : a ∈ B' \ B) : ∃ S : Set α, Finite S ∧ S ⊆ B ∧ ¬ M.Indep (insert a S) := by + have (a : α) (ha : a ∈ B' \ B) : ∃ S : Set α, Finite S ∧ S ⊆ B ∧ ¬ N.Indep (insert a S) := by have := (hB.insert_dep ⟨hB'.subset_ground ha.1, ha.2⟩).1 contrapose! this exact Finitary.indep_of_forall_finite _ fun J hJ fin ↦ (this (J \ {a}) fin.diff.to_subtype <| @@ -50,7 +298,7 @@ theorem Base.cardinalMk_diff_comm_of_finitary (hB : M.Base B) (hB' : M.Base B') choose S S_fin hSB dep using this let U := ⋃ a : ↥(B' \ B), S a a.2 suffices B \ B' ⊆ U by - refine hge.antisymm' <| (mk_le_mk_of_subset this).trans <| (mk_iUnion_le ..).trans + refine (mk_le_mk_of_subset this).trans <| (mk_iUnion_le ..).trans <| (mul_le_max_of_aleph0_le_left (by simp)).trans ?_ simp only [sup_le_iff, le_refl, true_and] exact ciSup_le' fun e ↦ (lt_aleph0_of_finite _).le.trans <| by simp @@ -63,40 +311,35 @@ theorem Base.cardinalMk_diff_comm_of_finitary (hB : M.Base B) (hB' : M.Base B') refine dep a this (ind.subset <| insert_subset_insert <| .trans ?_ subset_union_right) exact subset_iUnion_of_subset ⟨a, this⟩ subset_rfl -theorem Base.cardinalMk_eq_of_finitary (hB : M.Base B) (hB' : M.Base B') : #B = #B' := by - rw [← diff_union_inter B B', - mk_union_of_disjoint (disjoint_sdiff_left.mono_right inter_subset_right), - hB.cardinalMk_diff_comm_of_finitary hB', - ← mk_union_of_disjoint (disjoint_sdiff_left.mono_right inter_subset_left), - inter_comm, diff_union_inter] - -theorem Basis'.cardinalMk_diff_comm_of_finitary (hIX : M.Basis' I X) (hJX : M.Basis' J X) : - #(I \ J : Set α) = #(J \ I : Set α) := by - rw [← base_restrict_iff'] at hIX hJX - rw [hIX.cardinalMk_diff_comm_of_finitary hJX] - -theorem Basis.cardinalMk_diff_comm_of_finitary (hIX : M.Basis I X) (hJX : M.Basis J X) : - #(I \ J : Set α) = #(J \ I : Set α) := - hIX.basis'.cardinalMk_diff_comm_of_finitary hJX.basis' - -theorem Basis'.cardinalMk_eq_of_finitary (hIX : M.Basis' I X) (hJX : M.Basis' J X) : #I = #J := by - rw [← base_restrict_iff'] at hIX hJX - rw [hIX.cardinalMk_eq_of_finitary hJX] - -theorem Basis.cardinalMk_eq_of_finitary (hIX : M.Basis I X) (hJX : M.Basis J X) : #I = #J := - hIX.basis'.cardinalMk_eq_of_finitary hJX.basis' +instance invariantCardinalRank_map (M : Matroid α) [InvariantCardinalRank M] (hf : InjOn f M.E) : + InvariantCardinalRank (M.map f hf) := by + refine ⟨fun I J X hI hJ ↦ ?_⟩ + obtain ⟨I, X, hIX, rfl, rfl⟩ := map_basis_iff'.1 hI + obtain ⟨J, X', hJX, rfl, h'⟩ := map_basis_iff'.1 hJ + obtain rfl : X = X' := by + rwa [InjOn.image_eq_image_iff hf hIX.subset_ground hJX.subset_ground] at h' + have hcard := hIX.cardinalMk_diff_comm hJX + rwa [← lift_inj.{u,v}, + ← mk_image_eq_of_injOn_lift _ _ (hf.mono ((hIX.indep.diff _).subset_ground)), + ← mk_image_eq_of_injOn_lift _ _ (hf.mono ((hJX.indep.diff _).subset_ground)), + lift_inj, (hf.mono hIX.indep.subset_ground).image_diff, + (hf.mono hJX.indep.subset_ground).image_diff, inter_comm, + hf.image_inter hJX.indep.subset_ground hIX.indep.subset_ground, + diff_inter_self_eq_diff, diff_self_inter] at hcard -theorem Indep.cardinalMk_le_base_of_finitary (hI : M.Indep I) (hB : M.Base B) : #I ≤ #B := - have ⟨_B', hB', hIB'⟩ := hI.exists_base_superset - hB'.cardinalMk_eq_of_finitary hB ▸ mk_le_mk_of_subset hIB' +instance invariantCardinalRank_comap (M : Matroid β) [InvariantCardinalRank M] (f : α → β) : + InvariantCardinalRank (M.comap f) := by + refine ⟨fun I J X hI hJ ↦ ?_⟩ + obtain ⟨hI, hfI, hIX⟩ := comap_basis_iff.1 hI + obtain ⟨hJ, hfJ, hJX⟩ := comap_basis_iff.1 hJ + rw [← lift_inj.{u,v}, ← mk_image_eq_of_injOn_lift _ _ (hfI.mono diff_subset), + ← mk_image_eq_of_injOn_lift _ _ (hfJ.mono diff_subset), lift_inj, hfI.image_diff, + hfJ.image_diff, ← diff_union_diff_cancel inter_subset_left (image_inter_subset f I J), + inter_comm, diff_inter_self_eq_diff, mk_union_of_disjoint, hI.cardinalMk_diff_comm hJ, + ← diff_union_diff_cancel inter_subset_left (image_inter_subset f J I), inter_comm, + diff_inter_self_eq_diff, mk_union_of_disjoint, inter_comm J I] <;> + exact disjoint_sdiff_left.mono_right (diff_subset.trans inter_subset_left) -theorem Indep.cardinalMk_le_basis'_of_finitary (hI : M.Indep I) (hJ : M.Basis' J X) (hIX : I ⊆ X) : - #I ≤ #J := - have ⟨_J', hJ', hIJ'⟩ := hI.subset_basis'_of_subset hIX - hJ'.cardinalMk_eq_of_finitary hJ ▸ mk_le_mk_of_subset hIJ' - -theorem Indep.cardinalMk_le_basis_of_finitary (hI : M.Indep I) (hJ : M.Basis J X) (hIX : I ⊆ X) : - #I ≤ #J := - hI.cardinalMk_le_basis'_of_finitary hJ.basis' hIX +end Instances end Matroid diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index 95519a608d3833..be4890f48ea145 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -800,7 +800,7 @@ nonrec theorem div_le_div_left {a b c : ℝ≥0} (a0 : 0 < a) (b0 : 0 < b) (c0 : div_le_div_iff_of_pos_left a0 b0 c0 theorem le_of_forall_lt_one_mul_le {x y : ℝ≥0} (h : ∀ a < 1, a * x ≤ y) : x ≤ y := - le_of_forall_ge_of_dense fun a ha => by + le_of_forall_lt_imp_le_of_dense fun a ha => by have hx : x ≠ 0 := pos_iff_ne_zero.1 (lt_of_le_of_lt (zero_le _) ha) have hx' : x⁻¹ ≠ 0 := by rwa [Ne, inv_eq_zero] have : a * x⁻¹ < 1 := by rwa [← lt_inv_iff_mul_lt hx', inv_inv] diff --git a/Mathlib/Data/Nat/Choose/Bounds.lean b/Mathlib/Data/Nat/Choose/Bounds.lean index d4b21ddf606d9a..e89df1b02a1add 100644 --- a/Mathlib/Data/Nat/Choose/Bounds.lean +++ b/Mathlib/Data/Nat/Choose/Bounds.lean @@ -50,4 +50,20 @@ theorem pow_le_choose (r n : ℕ) : ((n + 1 - r : ℕ) ^ r : α) / r ! ≤ n.cho exact n.pow_sub_le_descFactorial r exact mod_cast r.factorial_pos +theorem choose_succ_le_two_pow (n k : ℕ) : (n + 1).choose k ≤ 2 ^ n := by + by_cases lt : n + 1 < k + · simp [choose_eq_zero_of_lt lt] + · cases' n with n + · cases k <;> simp_all + · cases' k with k + · rw [choose_zero_right] + exact Nat.one_le_two_pow + · rw [choose_succ_succ', two_pow_succ] + exact Nat.add_le_add (choose_succ_le_two_pow n k) (choose_succ_le_two_pow n (k + 1)) + +theorem choose_le_two_pow (n k : ℕ) (p : 0 < n) : n.choose k < 2 ^ n := by + refine lt_of_le_of_lt ?_ (Nat.two_pow_pred_lt_two_pow p) + rw [← Nat.sub_add_cancel p] + exact choose_succ_le_two_pow (n - 1) k + end Nat diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index 2d4196b2fe33ab..f576bc2d52e16d 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -24,14 +24,6 @@ assert_not_exists OrderTop namespace Nat -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/5338 we just unused argument warnings, -but these are used in the decreasing by blocks. -If instead we inline the `have` blocks, the unusedHavesSuffices linter triggers. --/ -set_option linter.unusedVariables false - /-! ### Floor logarithm -/ diff --git a/Mathlib/Data/PEquiv.lean b/Mathlib/Data/PEquiv.lean index 007830364ceb2b..fa6cbb13d14d73 100644 --- a/Mathlib/Data/PEquiv.lean +++ b/Mathlib/Data/PEquiv.lean @@ -432,6 +432,7 @@ theorem toPEquiv_trans (f : α ≃ β) (g : β ≃ γ) : theorem toPEquiv_symm (f : α ≃ β) : f.symm.toPEquiv = f.toPEquiv.symm := rfl +@[simp] theorem toPEquiv_apply (f : α ≃ β) (x : α) : f.toPEquiv x = some (f x) := rfl diff --git a/Mathlib/Data/PFunctor/Univariate/M.lean b/Mathlib/Data/PFunctor/Univariate/M.lean index 70b6e183f4a7e6..0d37382ac25ee4 100644 --- a/Mathlib/Data/PFunctor/Univariate/M.lean +++ b/Mathlib/Data/PFunctor/Univariate/M.lean @@ -161,7 +161,7 @@ end Approx open Approx /-- Internal definition for `M`. It is needed to avoid name clashes -between `M.mk` and `M.cases_on` and the declarations generated for +between `M.mk` and `M.casesOn` and the declarations generated for the structure -/ structure MIntl where /-- An `n`-th level approximation, for each depth `n` -/ @@ -378,8 +378,7 @@ theorem casesOn_mk' {r : M F → Sort*} {a} (x : F.B a → M F) /-- `IsPath p x` tells us if `p` is a valid path through `x` -/ inductive IsPath : Path F → M F → Prop | nil (x : M F) : IsPath [] x - | - cons (xs : Path F) {a} (x : M F) (f : F.B a → M F) (i : F.B a) : + | cons (xs : Path F) {a} (x : M F) (f : F.B a → M F) (i : F.B a) : x = M.mk ⟨a, f⟩ → IsPath xs (f i) → IsPath (⟨a, i⟩ :: xs) x theorem isPath_cons {xs : Path F} {a a'} {f : F.B a → M F} {i : F.B a'} : diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index 039acaddf3405f..3e9bf433eb681a 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -42,7 +42,7 @@ theorem of_near (f : ℕ → ℚ) (x : ℝ) (h : ∀ ε > 0, ∃ i, ∀ j ≥ i, ⟨isCauSeq_iff_lift.2 (CauSeq.of_near _ (const abs x) h), sub_eq_zero.1 <| abs_eq_zero.1 <| - (eq_of_le_of_forall_le_of_dense (abs_nonneg _)) fun _ε ε0 => + (eq_of_le_of_forall_lt_imp_le_of_dense (abs_nonneg _)) fun _ε ε0 => mk_near_of_forall_near <| (h _ ε0).imp fun _i h j ij => le_of_lt (h j ij)⟩ theorem exists_floor (x : ℝ) : ∃ ub : ℤ, (ub : ℝ) ≤ x ∧ ∀ z : ℤ, (z : ℝ) ≤ x → z ≤ ub := @@ -88,7 +88,7 @@ theorem exists_isLUB (hne : s.Nonempty) (hbdd : BddAbove s) : ∃ x, IsLUB s x : simpa using sub_lt_iff_lt_add'.2 (lt_of_le_of_lt hy <| sub_lt_iff_lt_add.1 <| hf₂ _ k0 _ yS) let g : CauSeq ℚ abs := ⟨fun n => f n / n, hg⟩ refine ⟨mk g, ⟨fun x xS => ?_, fun y h => ?_⟩⟩ - · refine le_of_forall_ge_of_dense fun z xz => ?_ + · refine le_of_forall_lt_imp_le_of_dense fun z xz => ?_ cases' exists_nat_gt (x - z)⁻¹ with K hK refine le_mk_of_forall_le ⟨K, fun n nK => ?_⟩ replace xz := sub_pos.2 xz diff --git a/Mathlib/Data/Real/EReal.lean b/Mathlib/Data/Real/EReal.lean index 17c70dbd11726e..8e308ba21e9aa2 100644 --- a/Mathlib/Data/Real/EReal.lean +++ b/Mathlib/Data/Real/EReal.lean @@ -1291,7 +1291,7 @@ private lemma exists_lt_add_right {a b c : EReal} (hc : c < a + b) : ∃ b' < b, simp_rw [add_comm a] at hc ⊢; exact exists_lt_add_left hc lemma add_le_of_forall_lt {a b c : EReal} (h : ∀ a' < a, ∀ b' < b, a' + b' ≤ c) : a + b ≤ c := by - refine le_of_forall_ge_of_dense fun d hd ↦ ?_ + refine le_of_forall_lt_imp_le_of_dense fun d hd ↦ ?_ obtain ⟨a', ha', hd⟩ := exists_lt_add_left hd obtain ⟨b', hb', hd⟩ := exists_lt_add_right hd exact hd.le.trans (h _ ha' _ hb') diff --git a/Mathlib/Data/Rel.lean b/Mathlib/Data/Rel.lean index ced1fb359df561..cae0b3a83e60e8 100644 --- a/Mathlib/Data/Rel.lean +++ b/Mathlib/Data/Rel.lean @@ -136,12 +136,10 @@ theorem inv_comp (r : Rel α β) (s : Rel β γ) : inv (r • s) = inv s • inv @[simp] theorem inv_bot : (⊥ : Rel α β).inv = (⊥ : Rel β α) := by - #adaptation_note /-- nightly-2024-03-16: simp was `simp [Bot.bot, inv, flip]` -/ simp [Bot.bot, inv, Function.flip_def] @[simp] theorem inv_top : (⊤ : Rel α β).inv = (⊤ : Rel β α) := by - #adaptation_note /-- nightly-2024-03-16: simp was `simp [Top.top, inv, flip]` -/ simp [Top.top, inv, Function.flip_def] /-- Image of a set under a relation -/ diff --git a/Mathlib/Data/SetLike/Basic.lean b/Mathlib/Data/SetLike/Basic.lean index 44719c1ca0a00c..9f99e745a2ce21 100644 --- a/Mathlib/Data/SetLike/Basic.lean +++ b/Mathlib/Data/SetLike/Basic.lean @@ -214,4 +214,14 @@ theorem exists_of_lt : p < q → ∃ x ∈ q, x ∉ p := theorem lt_iff_le_and_exists : p < q ↔ p ≤ q ∧ ∃ x ∈ q, x ∉ p := by rw [lt_iff_le_not_le, not_le_iff_exists] +/-- membership is inherited from `Set X` -/ +abbrev instSubtypeSet {X} {p : Set X → Prop} : SetLike {s // p s} X where + coe := (↑) + coe_injective' := Subtype.val_injective + +/-- membership is inherited from `S` -/ +abbrev instSubtype {X S} [SetLike S X] {p : S → Prop} : SetLike {s // p s} X where + coe := (↑) + coe_injective' := SetLike.coe_injective.comp Subtype.val_injective + end SetLike diff --git a/Mathlib/Data/Sign.lean b/Mathlib/Data/Sign.lean index 4607e88ca0d3ec..ffb6b9cb726dbb 100644 --- a/Mathlib/Data/Sign.lean +++ b/Mathlib/Data/Sign.lean @@ -115,11 +115,8 @@ instance : BoundedOrder SignType where le_top := LE.of_pos bot := -1 bot_le := - #adaptation_note - /-- - Added `by exact` after https://github.com/leanprover/lean4/pull/6053, - but don't understand why it was needed. - -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/6053 + Added `by exact`, but don't understand why it was needed. -/ by exact LE.of_neg instance : HasDistribNeg SignType := diff --git a/Mathlib/Data/Subtype.lean b/Mathlib/Data/Subtype.lean index ea0164bca8d418..ebde31a801bf53 100644 --- a/Mathlib/Data/Subtype.lean +++ b/Mathlib/Data/Subtype.lean @@ -155,7 +155,6 @@ def map {p : α → Prop} {q : β → Prop} (f : α → β) (h : ∀ a, p a → Subtype p → Subtype q := fun x ↦ ⟨f x, h x x.prop⟩ -#adaptation_note /-- nightly-2024-03-16: added to replace simp [Subtype.map] -/ theorem map_def {p : α → Prop} {q : β → Prop} (f : α → β) (h : ∀ a, p a → q (f a)) : map f h = fun x ↦ ⟨f x, h x x.prop⟩ := rfl diff --git a/Mathlib/Data/Sum/Basic.lean b/Mathlib/Data/Sum/Basic.lean index 30d502873b1397..cf572757b919fa 100644 --- a/Mathlib/Data/Sum/Basic.lean +++ b/Mathlib/Data/Sum/Basic.lean @@ -269,3 +269,17 @@ def in₂ (c : γ) : α ⊕ (β ⊕ γ) := inr <| inr c end Sum3 + +/-! +### PSum +-/ + +namespace PSum + +variable {α β : Sort*} + +theorem inl_injective : Function.Injective (PSum.inl : α → α ⊕' β) := fun _ _ ↦ inl.inj + +theorem inr_injective : Function.Injective (PSum.inr : β → α ⊕' β) := fun _ _ ↦ inr.inj + +end PSum diff --git a/Mathlib/Deprecated/Group.lean b/Mathlib/Deprecated/Group.lean deleted file mode 100644 index dd913b42c9bec0..00000000000000 --- a/Mathlib/Deprecated/Group.lean +++ /dev/null @@ -1,391 +0,0 @@ -/- -Copyright (c) 2019 Yury Kudryashov. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury Kudryashov --/ -import Mathlib.Algebra.Group.Units.Hom -import Mathlib.Algebra.Ring.Hom.Defs -import Mathlib.Algebra.Group.TypeTags.Basic - -/-! -# Unbundled monoid and group homomorphisms - -This file is deprecated, and is no longer imported by anything in mathlib other than other -deprecated files, and test files. You should not need to import it. - -This file defines predicates for unbundled monoid and group homomorphisms. Instead of using -this file, please use `MonoidHom`, defined in `Algebra.Hom.Group`, with notation `→*`, for -morphisms between monoids or groups. For example use `φ : G →* H` to represent a group -homomorphism between multiplicative groups, and `ψ : A →+ B` to represent a group homomorphism -between additive groups. - -## Main Definitions - -`IsMonoidHom` (deprecated), `IsGroupHom` (deprecated) - -## Tags - -IsGroupHom, IsMonoidHom - --/ - - -universe u v - -variable {α : Type u} {β : Type v} - -/-- Predicate for maps which preserve an addition. -/ -structure IsAddHom {α β : Type*} [Add α] [Add β] (f : α → β) : Prop where - /-- The proposition that `f` preserves addition. -/ - map_add : ∀ x y, f (x + y) = f x + f y - -/-- Predicate for maps which preserve a multiplication. -/ -@[to_additive] -structure IsMulHom {α β : Type*} [Mul α] [Mul β] (f : α → β) : Prop where - /-- The proposition that `f` preserves multiplication. -/ - map_mul : ∀ x y, f (x * y) = f x * f y - -namespace IsMulHom - -variable [Mul α] [Mul β] {γ : Type*} [Mul γ] - -/-- The identity map preserves multiplication. -/ -@[to_additive "The identity map preserves addition"] -theorem id : IsMulHom (id : α → α) := - { map_mul := fun _ _ => rfl } - -/-- The composition of maps which preserve multiplication, also preserves multiplication. -/ -@[to_additive "The composition of addition preserving maps also preserves addition"] -theorem comp {f : α → β} {g : β → γ} (hf : IsMulHom f) (hg : IsMulHom g) : IsMulHom (g ∘ f) := - { map_mul := fun x y => by simp only [Function.comp, hf.map_mul, hg.map_mul] } - -/-- A product of maps which preserve multiplication, -preserves multiplication when the target is commutative. -/ -@[to_additive - "A sum of maps which preserves addition, preserves addition when the target - is commutative."] -theorem mul {α β} [Semigroup α] [CommSemigroup β] {f g : α → β} (hf : IsMulHom f) - (hg : IsMulHom g) : IsMulHom fun a => f a * g a := - { map_mul := fun a b => by - simp only [hf.map_mul, hg.map_mul, mul_comm, mul_assoc, mul_left_comm] } - -/-- The inverse of a map which preserves multiplication, -preserves multiplication when the target is commutative. -/ -@[to_additive - "The negation of a map which preserves addition, preserves addition when - the target is commutative."] -theorem inv {α β} [Mul α] [CommGroup β] {f : α → β} (hf : IsMulHom f) : IsMulHom fun a => (f a)⁻¹ := - { map_mul := fun a b => (hf.map_mul a b).symm ▸ mul_inv _ _ } - -end IsMulHom - -/-- Predicate for additive monoid homomorphisms -(deprecated -- use the bundled `MonoidHom` version). -/ -structure IsAddMonoidHom [AddZeroClass α] [AddZeroClass β] (f : α → β) extends IsAddHom f : - Prop where - /-- The proposition that `f` preserves the additive identity. -/ - map_zero : f 0 = 0 - -/-- Predicate for monoid homomorphisms (deprecated -- use the bundled `MonoidHom` version). -/ -@[to_additive] -structure IsMonoidHom [MulOneClass α] [MulOneClass β] (f : α → β) extends IsMulHom f : Prop where - /-- The proposition that `f` preserves the multiplicative identity. -/ - map_one : f 1 = 1 - -namespace MonoidHom - -variable {M : Type*} {N : Type*} {mM : MulOneClass M} {mN : MulOneClass N} - -/-- Interpret a map `f : M → N` as a homomorphism `M →* N`. -/ -@[to_additive "Interpret a map `f : M → N` as a homomorphism `M →+ N`."] -def of {f : M → N} (h : IsMonoidHom f) : M →* N where - toFun := f - map_one' := h.2 - map_mul' := h.1.1 - -@[to_additive (attr := simp)] -theorem coe_of {f : M → N} (hf : IsMonoidHom f) : ⇑(MonoidHom.of hf) = f := - rfl - -@[to_additive] -theorem isMonoidHom_coe (f : M →* N) : IsMonoidHom (f : M → N) := - { map_mul := f.map_mul - map_one := f.map_one } - -end MonoidHom - -namespace MulEquiv - -variable {M : Type*} {N : Type*} [MulOneClass M] [MulOneClass N] - -/-- A multiplicative isomorphism preserves multiplication (deprecated). -/ -@[to_additive "An additive isomorphism preserves addition (deprecated)."] -theorem isMulHom (h : M ≃* N) : IsMulHom h := - ⟨map_mul h⟩ - -/-- A multiplicative bijection between two monoids is a monoid hom - (deprecated -- use `MulEquiv.toMonoidHom`). -/ -@[to_additive - "An additive bijection between two additive monoids is an additive - monoid hom (deprecated). "] -theorem isMonoidHom (h : M ≃* N) : IsMonoidHom h := - { map_mul := map_mul h - map_one := h.map_one } - -end MulEquiv - -namespace IsMonoidHom - -/-- A monoid homomorphism preserves multiplication. -/ -@[to_additive "An additive monoid homomorphism preserves addition."] -theorem map_mul' [MulOneClass α] [MulOneClass β] {f : α → β} (hf : IsMonoidHom f) (x y) : - f (x * y) = f x * f y := - hf.map_mul x y - -/-- The inverse of a map which preserves multiplication, -preserves multiplication when the target is commutative. -/ -@[to_additive - "The negation of a map which preserves addition, preserves addition - when the target is commutative."] -theorem inv {α β} [MulOneClass α] [CommGroup β] {f : α → β} (hf : IsMonoidHom f) : - IsMonoidHom fun a => (f a)⁻¹ := - { map_one := hf.map_one.symm ▸ inv_one - map_mul := fun a b => (hf.map_mul a b).symm ▸ mul_inv _ _ } - -end IsMonoidHom - -/-- A map to a group preserving multiplication is a monoid homomorphism. -/ -@[to_additive "A map to an additive group preserving addition is an additive monoid -homomorphism."] -theorem IsMulHom.to_isMonoidHom [MulOneClass α] [Group β] {f : α → β} (hf : IsMulHom f) : - IsMonoidHom f := - { map_one := (mul_right_eq_self (a := f 1)).1 <| by rw [← hf.map_mul, one_mul] - map_mul := hf.map_mul } - -namespace IsMonoidHom - -variable [MulOneClass α] [MulOneClass β] {f : α → β} - -/-- The identity map is a monoid homomorphism. -/ -@[to_additive "The identity map is an additive monoid homomorphism."] -theorem id : IsMonoidHom (@id α) := - { map_one := rfl - map_mul := fun _ _ => rfl } - -/-- The composite of two monoid homomorphisms is a monoid homomorphism. -/ -@[to_additive - "The composite of two additive monoid homomorphisms is an additive monoid - homomorphism."] -theorem comp (hf : IsMonoidHom f) {γ} [MulOneClass γ] {g : β → γ} (hg : IsMonoidHom g) : - IsMonoidHom (g ∘ f) := - { IsMulHom.comp hf.toIsMulHom hg.toIsMulHom with - map_one := show g _ = 1 by rw [hf.map_one, hg.map_one] } - -end IsMonoidHom - -namespace IsAddMonoidHom - -/-- Left multiplication in a ring is an additive monoid morphism. -/ -theorem isAddMonoidHom_mul_left {γ : Type*} [NonUnitalNonAssocSemiring γ] (x : γ) : - IsAddMonoidHom fun y : γ => x * y := - { map_zero := mul_zero x - map_add := fun y z => mul_add x y z } - -/-- Right multiplication in a ring is an additive monoid morphism. -/ -theorem isAddMonoidHom_mul_right {γ : Type*} [NonUnitalNonAssocSemiring γ] (x : γ) : - IsAddMonoidHom fun y : γ => y * x := - { map_zero := zero_mul x - map_add := fun y z => add_mul y z x } - -end IsAddMonoidHom - -/-- Predicate for additive group homomorphism (deprecated -- use bundled `MonoidHom`). -/ -structure IsAddGroupHom [AddGroup α] [AddGroup β] (f : α → β) extends IsAddHom f : Prop - -/-- Predicate for group homomorphisms (deprecated -- use bundled `MonoidHom`). -/ -@[to_additive] -structure IsGroupHom [Group α] [Group β] (f : α → β) extends IsMulHom f : Prop - -@[to_additive] -theorem MonoidHom.isGroupHom {G H : Type*} {_ : Group G} {_ : Group H} (f : G →* H) : - IsGroupHom (f : G → H) := - { map_mul := f.map_mul } - -@[to_additive] -theorem MulEquiv.isGroupHom {G H : Type*} {_ : Group G} {_ : Group H} (h : G ≃* H) : - IsGroupHom h := - { map_mul := map_mul h } - -/-- Construct `IsGroupHom` from its only hypothesis. -/ -@[to_additive "Construct `IsAddGroupHom` from its only hypothesis."] -theorem IsGroupHom.mk' [Group α] [Group β] {f : α → β} (hf : ∀ x y, f (x * y) = f x * f y) : - IsGroupHom f := - { map_mul := hf } - -namespace IsGroupHom - -variable [Group α] [Group β] {f : α → β} (hf : IsGroupHom f) - -/-- The identity is a group homomorphism. -/ -@[to_additive "The identity is an additive group homomorphism."] -theorem id : IsGroupHom (@id α) := - { map_mul := fun _ _ => rfl } - -section -include hf - -open IsMulHom (map_mul) - -theorem map_mul' : ∀ x y, f (x * y) = f x * f y := - hf.toIsMulHom.map_mul - -/-- A group homomorphism is a monoid homomorphism. -/ -@[to_additive "An additive group homomorphism is an additive monoid homomorphism."] -theorem to_isMonoidHom : IsMonoidHom f := - hf.toIsMulHom.to_isMonoidHom - -/-- A group homomorphism sends 1 to 1. -/ -@[to_additive "An additive group homomorphism sends 0 to 0."] -theorem map_one : f 1 = 1 := - hf.to_isMonoidHom.map_one - -/-- A group homomorphism sends inverses to inverses. -/ -@[to_additive "An additive group homomorphism sends negations to negations."] -theorem map_inv (a : α) : f a⁻¹ = (f a)⁻¹ := - eq_inv_of_mul_eq_one_left <| by rw [← hf.map_mul, inv_mul_cancel, hf.map_one] - -@[to_additive] -theorem map_div (a b : α) : f (a / b) = f a / f b := by - simp_rw [div_eq_mul_inv, hf.map_mul, hf.map_inv] - -/-- The composition of two group homomorphisms is a group homomorphism. -/ -@[to_additive - "The composition of two additive group homomorphisms is an additive - group homomorphism."] -theorem comp {γ} [Group γ] {g : β → γ} (hg : IsGroupHom g) : - IsGroupHom (g ∘ f) := - { IsMulHom.comp hf.toIsMulHom hg.toIsMulHom with } - -/-- A group homomorphism is injective iff its kernel is trivial. -/ -@[to_additive "An additive group homomorphism is injective if its kernel is trivial."] -theorem injective_iff : Function.Injective f ↔ ∀ a, f a = 1 → a = 1 := - ⟨fun h _ => by rw [← hf.map_one]; exact @h _ _, fun h x y hxy => - eq_of_div_eq_one <| h _ <| by rwa [hf.map_div, div_eq_one]⟩ - -end - -/-- The product of group homomorphisms is a group homomorphism if the target is commutative. -/ -@[to_additive - "The sum of two additive group homomorphisms is an additive group homomorphism - if the target is commutative."] -theorem mul {α β} [Group α] [CommGroup β] {f g : α → β} (hf : IsGroupHom f) (hg : IsGroupHom g) : - IsGroupHom fun a => f a * g a := - { map_mul := (hf.toIsMulHom.mul hg.toIsMulHom).map_mul } - -/-- The inverse of a group homomorphism is a group homomorphism if the target is commutative. -/ -@[to_additive - "The negation of an additive group homomorphism is an additive group homomorphism - if the target is commutative."] -theorem inv {α β} [Group α] [CommGroup β] {f : α → β} (hf : IsGroupHom f) : - IsGroupHom fun a => (f a)⁻¹ := - { map_mul := hf.toIsMulHom.inv.map_mul } - -end IsGroupHom - -namespace RingHom - -/-! -These instances look redundant, because `Deprecated.Ring` provides `IsRingHom` for a `→+*`. -Nevertheless these are harmless, and helpful for stripping out dependencies on `Deprecated.Ring`. --/ - - -variable {R : Type*} {S : Type*} - -section - -variable [NonAssocSemiring R] [NonAssocSemiring S] - -theorem to_isMonoidHom (f : R →+* S) : IsMonoidHom f := - { map_one := f.map_one - map_mul := f.map_mul } - -theorem to_isAddMonoidHom (f : R →+* S) : IsAddMonoidHom f := - { map_zero := f.map_zero - map_add := f.map_add } - -end - -section - -variable [Ring R] [Ring S] - -theorem to_isAddGroupHom (f : R →+* S) : IsAddGroupHom f := - { map_add := f.map_add } - -end - -end RingHom - -/-- Inversion is a group homomorphism if the group is commutative. -/ -@[to_additive - "Negation is an `AddGroup` homomorphism if the `AddGroup` is commutative."] -theorem Inv.isGroupHom [CommGroup α] : IsGroupHom (Inv.inv : α → α) := - { map_mul := mul_inv } - -/-- The difference of two additive group homomorphisms is an additive group -homomorphism if the target is commutative. -/ -theorem IsAddGroupHom.sub {α β} [AddGroup α] [AddCommGroup β] {f g : α → β} (hf : IsAddGroupHom f) - (hg : IsAddGroupHom g) : IsAddGroupHom fun a => f a - g a := by - simpa only [sub_eq_add_neg] using hf.add hg.neg - -namespace Units - -variable {M : Type*} {N : Type*} [Monoid M] [Monoid N] - -/-- The group homomorphism on units induced by a multiplicative morphism. -/ -abbrev map' {f : M → N} (hf : IsMonoidHom f) : Mˣ →* Nˣ := - map (MonoidHom.of hf) - -@[simp] -theorem coe_map' {f : M → N} (hf : IsMonoidHom f) (x : Mˣ) : ↑((map' hf : Mˣ → Nˣ) x) = f x := - rfl - -theorem coe_isMonoidHom : IsMonoidHom (↑· : Mˣ → M) := - (coeHom M).isMonoidHom_coe - -end Units - -namespace IsUnit - -variable {M : Type*} {N : Type*} [Monoid M] [Monoid N] - -theorem map' {f : M → N} (hf : IsMonoidHom f) {x : M} (h : IsUnit x) : IsUnit (f x) := - h.map (MonoidHom.of hf) - -end IsUnit - -theorem Additive.isAddHom [Mul α] [Mul β] {f : α → β} (hf : IsMulHom f) : - @IsAddHom (Additive α) (Additive β) _ _ f := - { map_add := hf.map_mul } - -theorem Multiplicative.isMulHom [Add α] [Add β] {f : α → β} (hf : IsAddHom f) : - @IsMulHom (Multiplicative α) (Multiplicative β) _ _ f := - { map_mul := hf.map_add } - --- defeq abuse -theorem Additive.isAddMonoidHom [MulOneClass α] [MulOneClass β] {f : α → β} - (hf : IsMonoidHom f) : @IsAddMonoidHom (Additive α) (Additive β) _ _ f := - { Additive.isAddHom hf.toIsMulHom with map_zero := hf.map_one } - -theorem Multiplicative.isMonoidHom [AddZeroClass α] [AddZeroClass β] {f : α → β} - (hf : IsAddMonoidHom f) : @IsMonoidHom (Multiplicative α) (Multiplicative β) _ _ f := - { Multiplicative.isMulHom hf.toIsAddHom with map_one := hf.map_zero } - -theorem Additive.isAddGroupHom [Group α] [Group β] {f : α → β} (hf : IsGroupHom f) : - @IsAddGroupHom (Additive α) (Additive β) _ _ f := - { map_add := hf.toIsMulHom.map_mul } - -theorem Multiplicative.isGroupHom [AddGroup α] [AddGroup β] {f : α → β} (hf : IsAddGroupHom f) : - @IsGroupHom (Multiplicative α) (Multiplicative β) _ _ f := - { map_mul := hf.toIsAddHom.map_add } diff --git a/Mathlib/Deprecated/Ring.lean b/Mathlib/Deprecated/Ring.lean deleted file mode 100644 index fc34ab12c07eeb..00000000000000 --- a/Mathlib/Deprecated/Ring.lean +++ /dev/null @@ -1,150 +0,0 @@ -/- -Copyright (c) 2020 Mario Carneiro. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro --/ -import Mathlib.Deprecated.Group - -/-! -# Unbundled semiring and ring homomorphisms (deprecated) - -This file is deprecated, and is no longer imported by anything in mathlib other than other -deprecated files, and test files. You should not need to import it. - -This file defines predicates for unbundled semiring and ring homomorphisms. Instead of using -this file, please use `RingHom`, defined in `Algebra.Hom.Ring`, with notation `→+*`, for -morphisms between semirings or rings. For example use `φ : A →+* B` to represent a -ring homomorphism. - -## Main Definitions - -`IsSemiringHom` (deprecated), `IsRingHom` (deprecated) - -## Tags - -IsSemiringHom, IsRingHom - --/ - - -universe u v - -variable {α : Type u} - -/-- Predicate for semiring homomorphisms (deprecated -- use the bundled `RingHom` version). -/ -structure IsSemiringHom {α : Type u} {β : Type v} [Semiring α] [Semiring β] (f : α → β) : Prop where - /-- The proposition that `f` preserves the additive identity. -/ - map_zero : f 0 = 0 - /-- The proposition that `f` preserves the multiplicative identity. -/ - map_one : f 1 = 1 - /-- The proposition that `f` preserves addition. -/ - map_add : ∀ x y, f (x + y) = f x + f y - /-- The proposition that `f` preserves multiplication. -/ - map_mul : ∀ x y, f (x * y) = f x * f y - -namespace IsSemiringHom - -variable {β : Type v} [Semiring α] [Semiring β] -variable {f : α → β} - -/-- The identity map is a semiring homomorphism. -/ -theorem id : IsSemiringHom (@id α) := by constructor <;> intros <;> rfl - -/-- The composition of two semiring homomorphisms is a semiring homomorphism. -/ -theorem comp (hf : IsSemiringHom f) {γ} [Semiring γ] {g : β → γ} (hg : IsSemiringHom g) : - IsSemiringHom (g ∘ f) := - { map_zero := by simpa [map_zero hf] using map_zero hg - map_one := by simpa [map_one hf] using map_one hg - map_add := fun {x y} => by simp [map_add hf, map_add hg] - map_mul := fun {x y} => by simp [map_mul hf, map_mul hg] } - -/-- A semiring homomorphism is an additive monoid homomorphism. -/ -theorem to_isAddMonoidHom (hf : IsSemiringHom f) : IsAddMonoidHom f := - { ‹IsSemiringHom f› with map_add := by apply @‹IsSemiringHom f›.map_add } - -/-- A semiring homomorphism is a monoid homomorphism. -/ -theorem to_isMonoidHom (hf : IsSemiringHom f) : IsMonoidHom f := - { ‹IsSemiringHom f› with } - -end IsSemiringHom - -/-- Predicate for ring homomorphisms (deprecated -- use the bundled `RingHom` version). -/ -structure IsRingHom {α : Type u} {β : Type v} [Ring α] [Ring β] (f : α → β) : Prop where - /-- The proposition that `f` preserves the multiplicative identity. -/ - map_one : f 1 = 1 - /-- The proposition that `f` preserves multiplication. -/ - map_mul : ∀ x y, f (x * y) = f x * f y - /-- The proposition that `f` preserves addition. -/ - map_add : ∀ x y, f (x + y) = f x + f y - -namespace IsRingHom - -variable {β : Type v} [Ring α] [Ring β] - -/-- A map of rings that is a semiring homomorphism is also a ring homomorphism. -/ -theorem of_semiring {f : α → β} (H : IsSemiringHom f) : IsRingHom f := - { H with } - -variable {f : α → β} {x y : α} - -/-- Ring homomorphisms map zero to zero. -/ -theorem map_zero (hf : IsRingHom f) : f 0 = 0 := - calc - f 0 = f (0 + 0) - f 0 := by rw [hf.map_add]; simp - _ = 0 := by simp - -/-- Ring homomorphisms preserve additive inverses. -/ -theorem map_neg (hf : IsRingHom f) : f (-x) = -f x := - calc - f (-x) = f (-x + x) - f x := by rw [hf.map_add]; simp - _ = -f x := by simp [hf.map_zero] - -/-- Ring homomorphisms preserve subtraction. -/ -theorem map_sub (hf : IsRingHom f) : f (x - y) = f x - f y := by - simp [sub_eq_add_neg, hf.map_add, hf.map_neg] - -/-- The identity map is a ring homomorphism. -/ -theorem id : IsRingHom (@id α) := by constructor <;> intros <;> rfl - --- see Note [no instance on morphisms] -/-- The composition of two ring homomorphisms is a ring homomorphism. -/ -theorem comp (hf : IsRingHom f) {γ} [Ring γ] {g : β → γ} (hg : IsRingHom g) : IsRingHom (g ∘ f) := - { map_add := fun x y => by simp only [Function.comp_apply, map_add hf, map_add hg] - map_mul := fun x y => by simp only [Function.comp_apply, map_mul hf, map_mul hg] - map_one := by simp only [Function.comp_apply, map_one hf, map_one hg] } - -/-- A ring homomorphism is also a semiring homomorphism. -/ -theorem to_isSemiringHom (hf : IsRingHom f) : IsSemiringHom f := - { ‹IsRingHom f› with map_zero := map_zero hf } - -theorem to_isAddGroupHom (hf : IsRingHom f) : IsAddGroupHom f := - { map_add := hf.map_add } - -end IsRingHom - -variable {β : Type v} {rα : Semiring α} {rβ : Semiring β} - -namespace RingHom - -section - -/-- Interpret `f : α → β` with `IsSemiringHom f` as a ring homomorphism. -/ -def of {f : α → β} (hf : IsSemiringHom f) : α →+* β := - { MonoidHom.of hf.to_isMonoidHom, AddMonoidHom.of hf.to_isAddMonoidHom with toFun := f } - -@[simp] -theorem coe_of {f : α → β} (hf : IsSemiringHom f) : ⇑(of hf) = f := - rfl - -theorem to_isSemiringHom (f : α →+* β) : IsSemiringHom f := - { map_zero := f.map_zero - map_one := f.map_one - map_add := f.map_add - map_mul := f.map_mul } - -end - -theorem to_isRingHom {α γ} [Ring α] [Ring γ] (g : α →+* γ) : IsRingHom g := - IsRingHom.of_semiring g.to_isSemiringHom - -end RingHom diff --git a/Mathlib/Deprecated/Subfield.lean b/Mathlib/Deprecated/Subfield.lean deleted file mode 100644 index 3ce93a3a26acc4..00000000000000 --- a/Mathlib/Deprecated/Subfield.lean +++ /dev/null @@ -1,150 +0,0 @@ -/- -Copyright (c) 2018 Andreas Swerdlow. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Andreas Swerdlow --/ -import Mathlib.Algebra.Field.Basic -import Mathlib.Deprecated.Subring -import Mathlib.Algebra.GroupWithZero.Units.Lemmas - -/-! -# Unbundled subfields (deprecated) - -This file is deprecated, and is no longer imported by anything in mathlib other than other -deprecated files, and test files. You should not need to import it. - -This file defines predicates for unbundled subfields. Instead of using this file, please use -`Subfield`, defined in `FieldTheory.Subfield`, for subfields of fields. - -## Main definitions - -`IsSubfield (S : Set F) : Prop` : the predicate that `S` is the underlying set of a subfield -of the field `F`. The bundled variant `Subfield F` should be used in preference to this. - -## Tags - -IsSubfield, subfield --/ - - -variable {F : Type*} [Field F] (S : Set F) - -/-- `IsSubfield (S : Set F)` is the predicate saying that a given subset of a field is -the set underlying a subfield. This structure is deprecated; use the bundled variant -`Subfield F` to model subfields of a field. -/ -structure IsSubfield extends IsSubring S : Prop where - inv_mem : ∀ {x : F}, x ∈ S → x⁻¹ ∈ S - -theorem IsSubfield.div_mem {S : Set F} (hS : IsSubfield S) {x y : F} (hx : x ∈ S) (hy : y ∈ S) : - x / y ∈ S := by - rw [div_eq_mul_inv] - exact hS.toIsSubring.toIsSubmonoid.mul_mem hx (hS.inv_mem hy) - -theorem IsSubfield.pow_mem {a : F} {n : ℤ} {s : Set F} (hs : IsSubfield s) (h : a ∈ s) : - a ^ n ∈ s := by - cases' n with n n - · suffices a ^ (n : ℤ) ∈ s by exact this - rw [zpow_natCast] - exact hs.toIsSubring.toIsSubmonoid.pow_mem h - · rw [zpow_negSucc] - exact hs.inv_mem (hs.toIsSubring.toIsSubmonoid.pow_mem h) - -theorem Univ.isSubfield : IsSubfield (@Set.univ F) := - { Univ.isSubmonoid, IsAddSubgroup.univ_addSubgroup with - inv_mem := fun _ ↦ trivial } - -theorem Preimage.isSubfield {K : Type*} [Field K] (f : F →+* K) {s : Set K} (hs : IsSubfield s) : - IsSubfield (f ⁻¹' s) := - { f.isSubring_preimage hs.toIsSubring with - inv_mem := fun {a} (ha : f a ∈ s) ↦ show f a⁻¹ ∈ s by - rw [map_inv₀] - exact hs.inv_mem ha } - -theorem Image.isSubfield {K : Type*} [Field K] (f : F →+* K) {s : Set F} (hs : IsSubfield s) : - IsSubfield (f '' s) := - { f.isSubring_image hs.toIsSubring with - inv_mem := fun ⟨x, xmem, ha⟩ ↦ ⟨x⁻¹, hs.inv_mem xmem, ha ▸ map_inv₀ f x⟩ } - -theorem Range.isSubfield {K : Type*} [Field K] (f : F →+* K) : IsSubfield (Set.range f) := by - rw [← Set.image_univ] - apply Image.isSubfield _ Univ.isSubfield - -namespace Field - -/-- `Field.closure s` is the minimal subfield that includes `s`. -/ -def closure : Set F := - { x | ∃ y ∈ Ring.closure S, ∃ z ∈ Ring.closure S, y / z = x } - -variable {S} - -theorem ring_closure_subset : Ring.closure S ⊆ closure S := - fun x hx ↦ ⟨x, hx, 1, Ring.closure.isSubring.toIsSubmonoid.one_mem, div_one x⟩ - -theorem closure.isSubmonoid : IsSubmonoid (closure S) := - { mul_mem := by - rintro _ _ ⟨p, hp, q, hq, hq0, rfl⟩ ⟨r, hr, s, hs, hs0, rfl⟩ - exact ⟨p * r, IsSubmonoid.mul_mem Ring.closure.isSubring.toIsSubmonoid hp hr, q * s, - IsSubmonoid.mul_mem Ring.closure.isSubring.toIsSubmonoid hq hs, - (div_mul_div_comm _ _ _ _).symm⟩ - one_mem := ring_closure_subset <| IsSubmonoid.one_mem Ring.closure.isSubring.toIsSubmonoid } - -theorem closure.isSubfield : IsSubfield (closure S) := - { closure.isSubmonoid with - add_mem := by - intro a b ha hb - rcases id ha with ⟨p, hp, q, hq, rfl⟩ - rcases id hb with ⟨r, hr, s, hs, rfl⟩ - by_cases hq0 : q = 0 - · rwa [hq0, div_zero, zero_add] - by_cases hs0 : s = 0 - · rwa [hs0, div_zero, add_zero] - exact ⟨p * s + q * r, - IsAddSubmonoid.add_mem Ring.closure.isSubring.toIsAddSubgroup.toIsAddSubmonoid - (Ring.closure.isSubring.toIsSubmonoid.mul_mem hp hs) - (Ring.closure.isSubring.toIsSubmonoid.mul_mem hq hr), - q * s, Ring.closure.isSubring.toIsSubmonoid.mul_mem hq hs, (div_add_div p r hq0 hs0).symm⟩ - zero_mem := ring_closure_subset Ring.closure.isSubring.toIsAddSubgroup.toIsAddSubmonoid.zero_mem - neg_mem := by - rintro _ ⟨p, hp, q, hq, rfl⟩ - exact ⟨-p, Ring.closure.isSubring.toIsAddSubgroup.neg_mem hp, q, hq, neg_div q p⟩ - inv_mem := by - rintro _ ⟨p, hp, q, hq, rfl⟩ - exact ⟨q, hq, p, hp, (inv_div _ _).symm⟩ } - -theorem mem_closure {a : F} (ha : a ∈ S) : a ∈ closure S := - ring_closure_subset <| Ring.mem_closure ha - -theorem subset_closure : S ⊆ closure S := - fun _ ↦ mem_closure - -theorem closure_subset {T : Set F} (hT : IsSubfield T) (H : S ⊆ T) : closure S ⊆ T := by - rintro _ ⟨p, hp, q, hq, hq0, rfl⟩ - exact hT.div_mem (Ring.closure_subset hT.toIsSubring H hp) - (Ring.closure_subset hT.toIsSubring H hq) - -theorem closure_subset_iff {s t : Set F} (ht : IsSubfield t) : closure s ⊆ t ↔ s ⊆ t := - ⟨Set.Subset.trans subset_closure, closure_subset ht⟩ - -@[gcongr] -theorem closure_mono {s t : Set F} (H : s ⊆ t) : closure s ⊆ closure t := - closure_subset closure.isSubfield <| Set.Subset.trans H subset_closure - -end Field - -theorem isSubfield_iUnion_of_directed {ι : Type*} [Nonempty ι] {s : ι → Set F} - (hs : ∀ i, IsSubfield (s i)) (directed : ∀ i j, ∃ k, s i ⊆ s k ∧ s j ⊆ s k) : - IsSubfield (⋃ i, s i) := - { inv_mem := fun hx ↦ - let ⟨i, hi⟩ := Set.mem_iUnion.1 hx - Set.mem_iUnion.2 ⟨i, (hs i).inv_mem hi⟩ - toIsSubring := isSubring_iUnion_of_directed (fun i ↦ (hs i).toIsSubring) directed } - -theorem IsSubfield.inter {S₁ S₂ : Set F} (hS₁ : IsSubfield S₁) (hS₂ : IsSubfield S₂) : - IsSubfield (S₁ ∩ S₂) := - { IsSubring.inter hS₁.toIsSubring hS₂.toIsSubring with - inv_mem := fun hx ↦ ⟨hS₁.inv_mem hx.1, hS₂.inv_mem hx.2⟩ } - -theorem IsSubfield.iInter {ι : Sort*} {S : ι → Set F} (h : ∀ y : ι, IsSubfield (S y)) : - IsSubfield (Set.iInter S) := - { IsSubring.iInter fun y ↦ (h y).toIsSubring with - inv_mem := fun hx ↦ Set.mem_iInter.2 fun y ↦ (h y).inv_mem <| Set.mem_iInter.1 hx y } diff --git a/Mathlib/Deprecated/Subgroup.lean b/Mathlib/Deprecated/Subgroup.lean deleted file mode 100644 index 3842d160c48acc..00000000000000 --- a/Mathlib/Deprecated/Subgroup.lean +++ /dev/null @@ -1,627 +0,0 @@ -/- -Copyright (c) 2018 Johannes Hölzl. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johannes Hölzl, Mitchell Rowett, Kim Morrison, Johan Commelin, Mario Carneiro, - Michael Howes --/ -import Mathlib.Algebra.Group.Subgroup.Basic -import Mathlib.Deprecated.Submonoid - -/-! -# Unbundled subgroups (deprecated) - -This file is deprecated, and is no longer imported by anything in mathlib other than other -deprecated files, and test files. You should not need to import it. - -This file defines unbundled multiplicative and additive subgroups. Instead of using this file, -please use `Subgroup G` and `AddSubgroup A`, defined in `Mathlib.Algebra.Group.Subgroup.Basic`. - -## Main definitions - -`IsAddSubgroup (S : Set A)` : the predicate that `S` is the underlying subset of an additive -subgroup of `A`. The bundled variant `AddSubgroup A` should be used in preference to this. - -`IsSubgroup (S : Set G)` : the predicate that `S` is the underlying subset of a subgroup -of `G`. The bundled variant `Subgroup G` should be used in preference to this. - -## Tags - -subgroup, subgroups, IsSubgroup --/ - - -open Set Function - -variable {G : Type*} {H : Type*} {A : Type*} {a b : G} - -section Group - -variable [Group G] [AddGroup A] - -/-- `s` is an additive subgroup: a set containing 0 and closed under addition and negation. -/ -structure IsAddSubgroup (s : Set A) extends IsAddSubmonoid s : Prop where - /-- The proposition that `s` is closed under negation. -/ - neg_mem {a} : a ∈ s → -a ∈ s - -/-- `s` is a subgroup: a set containing 1 and closed under multiplication and inverse. -/ -@[to_additive] -structure IsSubgroup (s : Set G) extends IsSubmonoid s : Prop where - /-- The proposition that `s` is closed under inverse. -/ - inv_mem {a} : a ∈ s → a⁻¹ ∈ s - -@[to_additive] -theorem IsSubgroup.div_mem {s : Set G} (hs : IsSubgroup s) {x y : G} (hx : x ∈ s) (hy : y ∈ s) : - x / y ∈ s := by simpa only [div_eq_mul_inv] using hs.mul_mem hx (hs.inv_mem hy) - -theorem Additive.isAddSubgroup {s : Set G} (hs : IsSubgroup s) : @IsAddSubgroup (Additive G) _ s := - @IsAddSubgroup.mk (Additive G) _ _ (Additive.isAddSubmonoid hs.toIsSubmonoid) hs.inv_mem - -theorem Additive.isAddSubgroup_iff {s : Set G} : @IsAddSubgroup (Additive G) _ s ↔ IsSubgroup s := - ⟨by rintro ⟨⟨h₁, h₂⟩, h₃⟩; exact @IsSubgroup.mk G _ _ ⟨h₁, @h₂⟩ @h₃, fun h => - Additive.isAddSubgroup h⟩ - -theorem Multiplicative.isSubgroup {s : Set A} (hs : IsAddSubgroup s) : - @IsSubgroup (Multiplicative A) _ s := - @IsSubgroup.mk (Multiplicative A) _ _ (Multiplicative.isSubmonoid hs.toIsAddSubmonoid) hs.neg_mem - -theorem Multiplicative.isSubgroup_iff {s : Set A} : - @IsSubgroup (Multiplicative A) _ s ↔ IsAddSubgroup s := - ⟨by rintro ⟨⟨h₁, h₂⟩, h₃⟩; exact @IsAddSubgroup.mk A _ _ ⟨h₁, @h₂⟩ @h₃, fun h => - Multiplicative.isSubgroup h⟩ - -@[to_additive of_add_neg] -theorem IsSubgroup.of_div (s : Set G) (one_mem : (1 : G) ∈ s) - (div_mem : ∀ {a b : G}, a ∈ s → b ∈ s → a * b⁻¹ ∈ s) : IsSubgroup s := - have inv_mem : ∀ a, a ∈ s → a⁻¹ ∈ s := fun a ha => by - have : 1 * a⁻¹ ∈ s := div_mem one_mem ha - convert this using 1 - rw [one_mul] - { inv_mem := inv_mem _ - mul_mem := fun {a b} ha hb => by - have : a * b⁻¹⁻¹ ∈ s := div_mem ha (inv_mem b hb) - convert this - rw [inv_inv] - one_mem } - -theorem IsAddSubgroup.of_sub (s : Set A) (zero_mem : (0 : A) ∈ s) - (sub_mem : ∀ {a b : A}, a ∈ s → b ∈ s → a - b ∈ s) : IsAddSubgroup s := - IsAddSubgroup.of_add_neg s zero_mem fun {x y} hx hy => by - simpa only [sub_eq_add_neg] using sub_mem hx hy - -@[to_additive] -theorem IsSubgroup.inter {s₁ s₂ : Set G} (hs₁ : IsSubgroup s₁) (hs₂ : IsSubgroup s₂) : - IsSubgroup (s₁ ∩ s₂) := - { IsSubmonoid.inter hs₁.toIsSubmonoid hs₂.toIsSubmonoid with - inv_mem := fun hx => ⟨hs₁.inv_mem hx.1, hs₂.inv_mem hx.2⟩ } - -@[to_additive] -theorem IsSubgroup.iInter {ι : Sort*} {s : ι → Set G} (hs : ∀ y : ι, IsSubgroup (s y)) : - IsSubgroup (Set.iInter s) := - { IsSubmonoid.iInter fun y => (hs y).toIsSubmonoid with - inv_mem := fun h => - Set.mem_iInter.2 fun y => IsSubgroup.inv_mem (hs _) (Set.mem_iInter.1 h y) } - -@[to_additive] -theorem isSubgroup_iUnion_of_directed {ι : Type*} [Nonempty ι] {s : ι → Set G} - (hs : ∀ i, IsSubgroup (s i)) (directed : ∀ i j, ∃ k, s i ⊆ s k ∧ s j ⊆ s k) : - IsSubgroup (⋃ i, s i) := - { inv_mem := fun ha => - let ⟨i, hi⟩ := Set.mem_iUnion.1 ha - Set.mem_iUnion.2 ⟨i, (hs i).inv_mem hi⟩ - toIsSubmonoid := isSubmonoid_iUnion_of_directed (fun i => (hs i).toIsSubmonoid) directed } - -end Group - -namespace IsSubgroup - -open IsSubmonoid - -variable [Group G] {s : Set G} (hs : IsSubgroup s) -include hs - -@[to_additive] -theorem inv_mem_iff : a⁻¹ ∈ s ↔ a ∈ s := - ⟨fun h => by simpa using hs.inv_mem h, inv_mem hs⟩ - -@[to_additive] -theorem mul_mem_cancel_right (h : a ∈ s) : b * a ∈ s ↔ b ∈ s := - ⟨fun hba => by simpa using hs.mul_mem hba (hs.inv_mem h), fun hb => hs.mul_mem hb h⟩ - -@[to_additive] -theorem mul_mem_cancel_left (h : a ∈ s) : a * b ∈ s ↔ b ∈ s := - ⟨fun hab => by simpa using hs.mul_mem (hs.inv_mem h) hab, hs.mul_mem h⟩ - -end IsSubgroup - -/-- `IsNormalAddSubgroup (s : Set A)` expresses the fact that `s` is a normal additive subgroup -of the additive group `A`. Important: the preferred way to say this in Lean is via bundled -subgroups `S : AddSubgroup A` and `hs : S.normal`, and not via this structure. -/ -structure IsNormalAddSubgroup [AddGroup A] (s : Set A) extends IsAddSubgroup s : Prop where - /-- The proposition that `s` is closed under (additive) conjugation. -/ - normal : ∀ n ∈ s, ∀ g : A, g + n + -g ∈ s - -/-- `IsNormalSubgroup (s : Set G)` expresses the fact that `s` is a normal subgroup -of the group `G`. Important: the preferred way to say this in Lean is via bundled -subgroups `S : Subgroup G` and not via this structure. -/ -@[to_additive] -structure IsNormalSubgroup [Group G] (s : Set G) extends IsSubgroup s : Prop where - /-- The proposition that `s` is closed under conjugation. -/ - normal : ∀ n ∈ s, ∀ g : G, g * n * g⁻¹ ∈ s - -@[to_additive] -theorem isNormalSubgroup_of_commGroup [CommGroup G] {s : Set G} (hs : IsSubgroup s) : - IsNormalSubgroup s := - { hs with normal := fun n hn g => by rwa [mul_right_comm, mul_inv_cancel, one_mul] } - -theorem Additive.isNormalAddSubgroup [Group G] {s : Set G} (hs : IsNormalSubgroup s) : - @IsNormalAddSubgroup (Additive G) _ s := - @IsNormalAddSubgroup.mk (Additive G) _ _ (Additive.isAddSubgroup hs.toIsSubgroup) - (@IsNormalSubgroup.normal _ ‹Group (Additive G)› _ hs) - -- Porting note: Lean needs help synthesising - -theorem Additive.isNormalAddSubgroup_iff [Group G] {s : Set G} : - @IsNormalAddSubgroup (Additive G) _ s ↔ IsNormalSubgroup s := - ⟨by rintro ⟨h₁, h₂⟩; exact @IsNormalSubgroup.mk G _ _ (Additive.isAddSubgroup_iff.1 h₁) @h₂, - fun h => Additive.isNormalAddSubgroup h⟩ - -theorem Multiplicative.isNormalSubgroup [AddGroup A] {s : Set A} (hs : IsNormalAddSubgroup s) : - @IsNormalSubgroup (Multiplicative A) _ s := - @IsNormalSubgroup.mk (Multiplicative A) _ _ (Multiplicative.isSubgroup hs.toIsAddSubgroup) - (@IsNormalAddSubgroup.normal _ ‹AddGroup (Multiplicative A)› _ hs) - -theorem Multiplicative.isNormalSubgroup_iff [AddGroup A] {s : Set A} : - @IsNormalSubgroup (Multiplicative A) _ s ↔ IsNormalAddSubgroup s := - ⟨by - rintro ⟨h₁, h₂⟩ - exact @IsNormalAddSubgroup.mk A _ _ (Multiplicative.isSubgroup_iff.1 h₁) @h₂, - fun h => Multiplicative.isNormalSubgroup h⟩ - -namespace IsSubgroup - -variable [Group G] - --- Normal subgroup properties -@[to_additive] -theorem mem_norm_comm {s : Set G} (hs : IsNormalSubgroup s) {a b : G} (hab : a * b ∈ s) : - b * a ∈ s := by - simpa using (hs.normal (a * b) hab a⁻¹) - -@[to_additive] -theorem mem_norm_comm_iff {s : Set G} (hs : IsNormalSubgroup s) {a b : G} : a * b ∈ s ↔ b * a ∈ s := - ⟨mem_norm_comm hs, mem_norm_comm hs⟩ - -/-- The trivial subgroup -/ -@[to_additive "the trivial additive subgroup"] -def trivial (G : Type*) [Group G] : Set G := - {1} - -@[to_additive (attr := simp)] -theorem mem_trivial {g : G} : g ∈ trivial G ↔ g = 1 := - mem_singleton_iff - -@[to_additive] -theorem trivial_normal : IsNormalSubgroup (trivial G) := by refine ⟨⟨⟨?_, ?_⟩, ?_⟩, ?_⟩ <;> simp - -@[to_additive] -theorem eq_trivial_iff {s : Set G} (hs : IsSubgroup s) : s = trivial G ↔ ∀ x ∈ s, x = (1 : G) := by - simp only [Set.ext_iff, IsSubgroup.mem_trivial] - exact ⟨fun h x => (h x).1, fun h x => ⟨h x, fun hx => hx.symm ▸ hs.toIsSubmonoid.one_mem⟩⟩ - -@[to_additive] -theorem univ_subgroup : IsNormalSubgroup (@univ G) := by refine ⟨⟨⟨?_, ?_⟩, ?_⟩, ?_⟩ <;> simp - -/-- The underlying set of the center of a group. -/ -@[to_additive addCenter "The underlying set of the center of an additive group."] -def center (G : Type*) [Group G] : Set G := - { z | ∀ g, g * z = z * g } - -@[to_additive mem_add_center] -theorem mem_center {a : G} : a ∈ center G ↔ ∀ g, g * a = a * g := - Iff.rfl - -@[to_additive add_center_normal] -theorem center_normal : IsNormalSubgroup (center G) := - { one_mem := by simp [center] - mul_mem := fun ha hb g => by - rw [← mul_assoc, mem_center.2 ha g, mul_assoc, mem_center.2 hb g, ← mul_assoc] - inv_mem := fun {a} ha g => - calc - g * a⁻¹ = a⁻¹ * (g * a) * a⁻¹ := by simp [ha g] - _ = a⁻¹ * g := by rw [← mul_assoc, mul_assoc]; simp - normal := fun n ha g h => - calc - h * (g * n * g⁻¹) = h * n := by simp [ha g, mul_assoc] - _ = g * g⁻¹ * n * h := by rw [ha h]; simp - _ = g * n * g⁻¹ * h := by rw [mul_assoc g, ha g⁻¹, ← mul_assoc] - } - -/-- The underlying set of the normalizer of a subset `S : Set G` of a group `G`. That is, - the elements `g : G` such that `g * S * g⁻¹ = S`. -/ -@[to_additive addNormalizer - "The underlying set of the normalizer of a subset `S : Set A` of an - additive group `A`. That is, the elements `a : A` such that `a + S - a = S`."] -def normalizer (s : Set G) : Set G := - { g : G | ∀ n, n ∈ s ↔ g * n * g⁻¹ ∈ s } - -@[to_additive] -theorem normalizer_isSubgroup (s : Set G) : IsSubgroup (normalizer s) := - { one_mem := by simp [normalizer] - mul_mem := fun {a b} - (ha : ∀ n, n ∈ s ↔ a * n * a⁻¹ ∈ s) (hb : ∀ n, n ∈ s ↔ b * n * b⁻¹ ∈ s) n => - by rw [mul_inv_rev, ← mul_assoc, mul_assoc a, mul_assoc a, ← ha, ← hb] - inv_mem := fun {a} (ha : ∀ n, n ∈ s ↔ a * n * a⁻¹ ∈ s) n => by - rw [ha (a⁻¹ * n * a⁻¹⁻¹)]; simp [mul_assoc] } - -@[to_additive subset_add_normalizer] -theorem subset_normalizer {s : Set G} (hs : IsSubgroup s) : s ⊆ normalizer s := fun g hg n => by - rw [IsSubgroup.mul_mem_cancel_right hs ((IsSubgroup.inv_mem_iff hs).2 hg), - IsSubgroup.mul_mem_cancel_left hs hg] - -end IsSubgroup - --- Homomorphism subgroups -namespace IsGroupHom - -open IsSubmonoid IsSubgroup - -/-- `ker f : Set G` is the underlying subset of the kernel of a map `G → H`. -/ -@[to_additive "`ker f : Set A` is the underlying subset of the kernel of a map `A → B`"] -def ker [Group H] (f : G → H) : Set G := - preimage f (trivial H) - -@[to_additive] -theorem mem_ker [Group H] (f : G → H) {x : G} : x ∈ ker f ↔ f x = 1 := - mem_trivial - -variable [Group G] [Group H] - -@[to_additive] -theorem one_ker_inv {f : G → H} (hf : IsGroupHom f) {a b : G} (h : f (a * b⁻¹) = 1) : - f a = f b := by - rw [hf.map_mul, hf.map_inv] at h - rw [← inv_inv (f b), eq_inv_of_mul_eq_one_left h] - -@[to_additive] -theorem one_ker_inv' {f : G → H} (hf : IsGroupHom f) {a b : G} (h : f (a⁻¹ * b) = 1) : - f a = f b := by - rw [hf.map_mul, hf.map_inv] at h - apply inv_injective - rw [eq_inv_of_mul_eq_one_left h] - -@[to_additive] -theorem inv_ker_one {f : G → H} (hf : IsGroupHom f) {a b : G} (h : f a = f b) : - f (a * b⁻¹) = 1 := by - have : f a * (f b)⁻¹ = 1 := by rw [h, mul_inv_cancel] - rwa [← hf.map_inv, ← hf.map_mul] at this - -@[to_additive] -theorem inv_ker_one' {f : G → H} (hf : IsGroupHom f) {a b : G} (h : f a = f b) : - f (a⁻¹ * b) = 1 := by - have : (f a)⁻¹ * f b = 1 := by rw [h, inv_mul_cancel] - rwa [← hf.map_inv, ← hf.map_mul] at this - -@[to_additive] -theorem one_iff_ker_inv {f : G → H} (hf : IsGroupHom f) (a b : G) : f a = f b ↔ f (a * b⁻¹) = 1 := - ⟨hf.inv_ker_one, hf.one_ker_inv⟩ - -@[to_additive] -theorem one_iff_ker_inv' {f : G → H} (hf : IsGroupHom f) (a b : G) : f a = f b ↔ f (a⁻¹ * b) = 1 := - ⟨hf.inv_ker_one', hf.one_ker_inv'⟩ - -@[to_additive] -theorem inv_iff_ker {f : G → H} (hf : IsGroupHom f) (a b : G) : f a = f b ↔ a * b⁻¹ ∈ ker f := by - rw [mem_ker]; exact one_iff_ker_inv hf _ _ - -@[to_additive] -theorem inv_iff_ker' {f : G → H} (hf : IsGroupHom f) (a b : G) : f a = f b ↔ a⁻¹ * b ∈ ker f := by - rw [mem_ker]; exact one_iff_ker_inv' hf _ _ - -@[to_additive] -theorem image_subgroup {f : G → H} (hf : IsGroupHom f) {s : Set G} (hs : IsSubgroup s) : - IsSubgroup (f '' s) := - { mul_mem := fun {a₁ a₂} ⟨b₁, hb₁, eq₁⟩ ⟨b₂, hb₂, eq₂⟩ => - ⟨b₁ * b₂, hs.mul_mem hb₁ hb₂, by simp [eq₁, eq₂, hf.map_mul]⟩ - one_mem := ⟨1, hs.toIsSubmonoid.one_mem, hf.map_one⟩ - inv_mem := fun {a} ⟨b, hb, Eq⟩ => - ⟨b⁻¹, hs.inv_mem hb, by - rw [hf.map_inv] - simp [*]⟩ } - -@[to_additive] -theorem range_subgroup {f : G → H} (hf : IsGroupHom f) : IsSubgroup (Set.range f) := - @Set.image_univ _ _ f ▸ hf.image_subgroup univ_subgroup.toIsSubgroup - -attribute [local simp] IsSubmonoid.one_mem IsSubgroup.inv_mem - IsSubmonoid.mul_mem IsNormalSubgroup.normal - -@[to_additive] -theorem preimage {f : G → H} (hf : IsGroupHom f) {s : Set H} (hs : IsSubgroup s) : - IsSubgroup (f ⁻¹' s) where - one_mem := by simp [hf.map_one, hs.one_mem] - mul_mem := by simp_all [hf.map_mul, hs.mul_mem] - inv_mem := by simp_all [hf.map_inv] - -@[to_additive] -theorem preimage_normal {f : G → H} (hf : IsGroupHom f) {s : Set H} (hs : IsNormalSubgroup s) : - IsNormalSubgroup (f ⁻¹' s) := - { one_mem := by simp [hf.map_one, hs.toIsSubgroup.one_mem] - mul_mem := by simp +contextual [hf.map_mul, hs.toIsSubgroup.mul_mem] - inv_mem := by simp +contextual [hf.map_inv, hs.toIsSubgroup.inv_mem] - normal := by simp +contextual [hs.normal, hf.map_mul, hf.map_inv] } - -@[to_additive] -theorem isNormalSubgroup_ker {f : G → H} (hf : IsGroupHom f) : IsNormalSubgroup (ker f) := - hf.preimage_normal trivial_normal - -@[to_additive] -theorem injective_of_trivial_ker {f : G → H} (hf : IsGroupHom f) (h : ker f = trivial G) : - Function.Injective f := by - intro a₁ a₂ hfa - simp [Set.ext_iff, ker, IsSubgroup.trivial] at h - have ha : a₁ * a₂⁻¹ = 1 := by rw [← h]; exact hf.inv_ker_one hfa - rw [eq_inv_of_mul_eq_one_left ha, inv_inv a₂] - -@[to_additive] -theorem trivial_ker_of_injective {f : G → H} (hf : IsGroupHom f) (h : Function.Injective f) : - ker f = trivial G := - Set.ext fun x => - Iff.intro - (fun hx => by - suffices f x = f 1 by simpa using h this - simp [hf.map_one]; rwa [mem_ker] at hx) - (by simp +contextual [mem_ker, hf.map_one]) - -@[to_additive] -theorem injective_iff_trivial_ker {f : G → H} (hf : IsGroupHom f) : - Function.Injective f ↔ ker f = trivial G := - ⟨hf.trivial_ker_of_injective, hf.injective_of_trivial_ker⟩ - -@[to_additive] -theorem trivial_ker_iff_eq_one {f : G → H} (hf : IsGroupHom f) : - ker f = trivial G ↔ ∀ x, f x = 1 → x = 1 := by - rw [Set.ext_iff] - simpa [ker] using ⟨fun h x hx => (h x).1 hx, fun h x => ⟨h x, fun hx => by rw [hx, hf.map_one]⟩⟩ - -end IsGroupHom - -namespace AddGroup - -variable [AddGroup A] - -/-- If `A` is an additive group and `s : Set A`, then `InClosure s : Set A` is the underlying -subset of the subgroup generated by `s`. -/ -inductive InClosure (s : Set A) : A → Prop - | basic {a : A} : a ∈ s → InClosure s a - | zero : InClosure s 0 - | neg {a : A} : InClosure s a → InClosure s (-a) - | add {a b : A} : InClosure s a → InClosure s b → InClosure s (a + b) - -end AddGroup - -namespace Group - -open IsSubmonoid IsSubgroup - -variable [Group G] {s : Set G} - -/-- If `G` is a group and `s : Set G`, then `InClosure s : Set G` is the underlying -subset of the subgroup generated by `s`. -/ -@[to_additive] -inductive InClosure (s : Set G) : G → Prop - | basic {a : G} : a ∈ s → InClosure s a - | one : InClosure s 1 - | inv {a : G} : InClosure s a → InClosure s a⁻¹ - | mul {a b : G} : InClosure s a → InClosure s b → InClosure s (a * b) - -/-- `Group.closure s` is the subgroup generated by `s`, i.e. the smallest subgroup containing `s`. --/ -@[to_additive - "`AddGroup.closure s` is the additive subgroup generated by `s`, i.e., the - smallest additive subgroup containing `s`."] -def closure (s : Set G) : Set G := - { a | InClosure s a } - -@[to_additive] -theorem mem_closure {a : G} : a ∈ s → a ∈ closure s := - InClosure.basic - -@[to_additive] -theorem closure.isSubgroup (s : Set G) : IsSubgroup (closure s) := - { one_mem := InClosure.one - mul_mem := InClosure.mul - inv_mem := InClosure.inv } - -@[to_additive] -theorem subset_closure {s : Set G} : s ⊆ closure s := fun _ => mem_closure - -@[to_additive] -theorem closure_subset {s t : Set G} (ht : IsSubgroup t) (h : s ⊆ t) : closure s ⊆ t := fun a ha => - by induction ha <;> simp [h _, *, ht.one_mem, ht.mul_mem, IsSubgroup.inv_mem_iff] - -@[to_additive] -theorem closure_subset_iff {s t : Set G} (ht : IsSubgroup t) : closure s ⊆ t ↔ s ⊆ t := - ⟨fun h _ ha => h (mem_closure ha), fun h _ ha => closure_subset ht h ha⟩ - -@[to_additive (attr := gcongr)] -theorem closure_mono {s t : Set G} (h : s ⊆ t) : closure s ⊆ closure t := - closure_subset (closure.isSubgroup _) <| Set.Subset.trans h subset_closure - -@[to_additive (attr := simp)] -theorem closure_subgroup {s : Set G} (hs : IsSubgroup s) : closure s = s := - Set.Subset.antisymm (closure_subset hs <| Set.Subset.refl s) subset_closure - -@[to_additive] -theorem exists_list_of_mem_closure {s : Set G} {a : G} (h : a ∈ closure s) : - ∃ l : List G, (∀ x ∈ l, x ∈ s ∨ x⁻¹ ∈ s) ∧ l.prod = a := - InClosure.recOn h - (fun {x} hxs => ⟨[x], List.forall_mem_singleton.2 <| Or.inl hxs, List.prod_singleton⟩) - ⟨[], List.forall_mem_nil _, rfl⟩ - (fun {_} _ ⟨L, HL1, HL2⟩ => - ⟨L.reverse.map Inv.inv, fun _ hx => - let ⟨y, hy1, hy2⟩ := List.exists_of_mem_map hx - hy2 ▸ Or.imp id (by rw [inv_inv]; exact id) (HL1 _ <| List.mem_reverse.1 hy1).symm, - HL2 ▸ - List.recOn L inv_one.symm fun hd tl ih => by - rw [List.reverse_cons, List.map_append, List.prod_append, ih, List.map_singleton, - List.prod_cons, List.prod_nil, mul_one, List.prod_cons, mul_inv_rev]⟩) - fun {x y} _ _ ⟨L1, HL1, HL2⟩ ⟨L2, HL3, HL4⟩ => - ⟨L1 ++ L2, List.forall_mem_append.2 ⟨HL1, HL3⟩, by rw [List.prod_append, HL2, HL4]⟩ - -@[to_additive] -theorem image_closure [Group H] {f : G → H} (hf : IsGroupHom f) (s : Set G) : - f '' closure s = closure (f '' s) := - le_antisymm - (by - rintro _ ⟨x, hx, rfl⟩ - exact InClosure.recOn hx - (by intros _ ha; exact subset_closure (mem_image_of_mem f ha)) - (by - rw [hf.map_one] - apply IsSubmonoid.one_mem (closure.isSubgroup _).toIsSubmonoid) - (by - intros _ _ - rw [hf.map_inv] - apply IsSubgroup.inv_mem (closure.isSubgroup _)) - (by - intros _ _ _ _ ha hb - rw [hf.map_mul] - exact (closure.isSubgroup (f '' s)).toIsSubmonoid.mul_mem ha hb)) - (closure_subset (hf.image_subgroup <| closure.isSubgroup _) <| - Set.image_subset _ subset_closure) - -@[to_additive] -theorem mclosure_subset {s : Set G} : Monoid.Closure s ⊆ closure s := - Monoid.closure_subset (closure.isSubgroup _).toIsSubmonoid <| subset_closure - -@[to_additive] -theorem mclosure_inv_subset {s : Set G} : Monoid.Closure (Inv.inv ⁻¹' s) ⊆ closure s := - Monoid.closure_subset (closure.isSubgroup _).toIsSubmonoid fun x hx => - inv_inv x ▸ ((closure.isSubgroup _).inv_mem <| subset_closure hx) - -@[to_additive] -theorem closure_eq_mclosure {s : Set G} : closure s = Monoid.Closure (s ∪ Inv.inv ⁻¹' s) := - Set.Subset.antisymm - (@closure_subset _ _ _ (Monoid.Closure (s ∪ Inv.inv ⁻¹' s)) - { one_mem := (Monoid.closure.isSubmonoid _).one_mem - mul_mem := (Monoid.closure.isSubmonoid _).mul_mem - inv_mem := fun hx => - Monoid.InClosure.recOn hx - (fun {x} hx => - Or.casesOn hx - (fun hx => - Monoid.subset_closure <| Or.inr <| show x⁻¹⁻¹ ∈ s from (inv_inv x).symm ▸ hx) - fun hx => Monoid.subset_closure <| Or.inl hx) - ((@inv_one G _).symm ▸ IsSubmonoid.one_mem (Monoid.closure.isSubmonoid _)) - fun {x y} _ _ ihx ihy => - (mul_inv_rev x y).symm ▸ IsSubmonoid.mul_mem (Monoid.closure.isSubmonoid _) ihy ihx } - (Set.Subset.trans Set.subset_union_left Monoid.subset_closure)) - (Monoid.closure_subset (closure.isSubgroup _).toIsSubmonoid <| - Set.union_subset subset_closure fun x hx => - inv_inv x ▸ (IsSubgroup.inv_mem (closure.isSubgroup _) <| subset_closure hx)) - -@[to_additive] -theorem mem_closure_union_iff {G : Type*} [CommGroup G] {s t : Set G} {x : G} : - x ∈ closure (s ∪ t) ↔ ∃ y ∈ closure s, ∃ z ∈ closure t, y * z = x := by - simp only [closure_eq_mclosure, Monoid.mem_closure_union_iff, exists_prop, preimage_union] - constructor - · rintro ⟨_, ⟨ys, hys, yt, hyt, rfl⟩, _, ⟨zs, hzs, zt, hzt, rfl⟩, rfl⟩ - refine ⟨_, ⟨_, hys, _, hzs, rfl⟩, _, ⟨_, hyt, _, hzt, rfl⟩, ?_⟩ - rw [mul_assoc, mul_assoc, mul_left_comm zs] - · rintro ⟨_, ⟨ys, hys, zs, hzs, rfl⟩, _, ⟨yt, hyt, zt, hzt, rfl⟩, rfl⟩ - refine ⟨_, ⟨ys, hys, yt, hyt, rfl⟩, _, ⟨zs, hzs, zt, hzt, rfl⟩, ?_⟩ - rw [mul_assoc, mul_assoc, mul_left_comm yt] - -end Group - -namespace IsSubgroup - -variable [Group G] - -@[to_additive] -theorem trivial_eq_closure : trivial G = Group.closure ∅ := - Subset.antisymm (by simp [Set.subset_def, (Group.closure.isSubgroup _).one_mem]) - (Group.closure_subset trivial_normal.toIsSubgroup <| by simp) - -end IsSubgroup - -/- The normal closure of a set s is the subgroup closure of all the conjugates of -elements of s. It is the smallest normal subgroup containing s. -/ -namespace Group - -variable {s : Set G} [Group G] - -theorem conjugatesOf_subset {t : Set G} (ht : IsNormalSubgroup t) {a : G} (h : a ∈ t) : - conjugatesOf a ⊆ t := fun x hc => by - obtain ⟨c, w⟩ := isConj_iff.1 hc - have H := IsNormalSubgroup.normal ht a h c - rwa [← w] - -theorem conjugatesOfSet_subset' {s t : Set G} (ht : IsNormalSubgroup t) (h : s ⊆ t) : - conjugatesOfSet s ⊆ t := - Set.iUnion₂_subset fun _ H => conjugatesOf_subset ht (h H) - -/-- The normal closure of a set s is the subgroup closure of all the conjugates of -elements of s. It is the smallest normal subgroup containing s. -/ -def normalClosure (s : Set G) : Set G := - closure (conjugatesOfSet s) - -theorem conjugatesOfSet_subset_normalClosure : conjugatesOfSet s ⊆ normalClosure s := - subset_closure - -theorem subset_normalClosure : s ⊆ normalClosure s := - Set.Subset.trans subset_conjugatesOfSet conjugatesOfSet_subset_normalClosure - -/-- The normal closure of a set is a subgroup. -/ -theorem normalClosure.isSubgroup (s : Set G) : IsSubgroup (normalClosure s) := - closure.isSubgroup (conjugatesOfSet s) - -/-- The normal closure of s is a normal subgroup. -/ -theorem normalClosure.is_normal : IsNormalSubgroup (normalClosure s) := - { normalClosure.isSubgroup _ with - normal := fun n h g => by - induction' h with x hx x hx ihx x y hx hy ihx ihy - · exact conjugatesOfSet_subset_normalClosure (conj_mem_conjugatesOfSet hx) - · simpa using (normalClosure.isSubgroup s).one_mem - · rw [← conj_inv] - exact (normalClosure.isSubgroup _).inv_mem ihx - · rw [← conj_mul] - exact (normalClosure.isSubgroup _).toIsSubmonoid.mul_mem ihx ihy } - -/-- The normal closure of s is the smallest normal subgroup containing s. -/ -theorem normalClosure_subset {s t : Set G} (ht : IsNormalSubgroup t) (h : s ⊆ t) : - normalClosure s ⊆ t := fun a w => by - induction w with - | basic hx => exact conjugatesOfSet_subset' ht h <| hx - | one => exact ht.toIsSubgroup.toIsSubmonoid.one_mem - | inv _ ihx => exact ht.toIsSubgroup.inv_mem ihx - | mul _ _ ihx ihy => exact ht.toIsSubgroup.toIsSubmonoid.mul_mem ihx ihy - -theorem normalClosure_subset_iff {s t : Set G} (ht : IsNormalSubgroup t) : - s ⊆ t ↔ normalClosure s ⊆ t := - ⟨normalClosure_subset ht, Set.Subset.trans subset_normalClosure⟩ - -@[gcongr] -theorem normalClosure_mono {s t : Set G} : s ⊆ t → normalClosure s ⊆ normalClosure t := fun h => - normalClosure_subset normalClosure.is_normal (Set.Subset.trans h subset_normalClosure) - -end Group - -/-- Create a bundled subgroup from a set `s` and `[IsSubgroup s]`. -/ -@[to_additive "Create a bundled additive subgroup from a set `s` and `[IsAddSubgroup s]`."] -def Subgroup.of [Group G] {s : Set G} (h : IsSubgroup s) : Subgroup G where - carrier := s - one_mem' := h.1.1 - mul_mem' := h.1.2 - inv_mem' := h.2 - -@[to_additive] -theorem Subgroup.isSubgroup [Group G] (K : Subgroup G) : IsSubgroup (K : Set G) := - { one_mem := K.one_mem' - mul_mem := K.mul_mem' - inv_mem := K.inv_mem' } - --- this will never fire if it's an instance -@[to_additive] -theorem Subgroup.of_normal [Group G] (s : Set G) (h : IsSubgroup s) (n : IsNormalSubgroup s) : - Subgroup.Normal (Subgroup.of h) := - { conj_mem := n.normal } diff --git a/Mathlib/Deprecated/Submonoid.lean b/Mathlib/Deprecated/Submonoid.lean deleted file mode 100644 index 25e8af4f8368fb..00000000000000 --- a/Mathlib/Deprecated/Submonoid.lean +++ /dev/null @@ -1,363 +0,0 @@ -/- -Copyright (c) 2018 Johannes Hölzl. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johannes Hölzl, Kenny Lau, Johan Commelin, Mario Carneiro, Kevin Buzzard --/ -import Mathlib.Algebra.Group.Submonoid.Defs -import Mathlib.Deprecated.Group -import Mathlib.Data.Set.Lattice -import Mathlib.Algebra.BigOperators.Group.Finset.Defs - -/-! -# Unbundled submonoids (deprecated) - -This file is deprecated, and is no longer imported by anything in mathlib other than other -deprecated files, and test files. You should not need to import it. - -This file defines unbundled multiplicative and additive submonoids. Instead of using this file, -please use `Submonoid G` and `AddSubmonoid A`, defined in `GroupTheory.Submonoid.Basic`. - -## Main definitions - -`IsAddSubmonoid (S : Set M)` : the predicate that `S` is the underlying subset of an additive -submonoid of `M`. The bundled variant `AddSubmonoid M` should be used in preference to this. - -`IsSubmonoid (S : Set M)` : the predicate that `S` is the underlying subset of a submonoid -of `M`. The bundled variant `Submonoid M` should be used in preference to this. - -## Tags -Submonoid, Submonoids, IsSubmonoid --/ - - -variable {M : Type*} [Monoid M] {s : Set M} {A : Type*} [AddMonoid A] - -/-- `s` is an additive submonoid: a set containing 0 and closed under addition. -Note that this structure is deprecated, and the bundled variant `AddSubmonoid A` should be -preferred. -/ -structure IsAddSubmonoid (s : Set A) : Prop where - /-- The proposition that s contains 0. -/ - zero_mem : (0 : A) ∈ s - /-- The proposition that s is closed under addition. -/ - add_mem {a b} : a ∈ s → b ∈ s → a + b ∈ s - -/-- `s` is a submonoid: a set containing 1 and closed under multiplication. -Note that this structure is deprecated, and the bundled variant `Submonoid M` should be -preferred. -/ -@[to_additive] -structure IsSubmonoid (s : Set M) : Prop where - /-- The proposition that s contains 1. -/ - one_mem : (1 : M) ∈ s - /-- The proposition that s is closed under multiplication. -/ - mul_mem {a b} : a ∈ s → b ∈ s → a * b ∈ s - -theorem Additive.isAddSubmonoid {s : Set M} : - IsSubmonoid s → @IsAddSubmonoid (Additive M) _ s - | ⟨h₁, h₂⟩ => ⟨h₁, @h₂⟩ - -theorem Additive.isAddSubmonoid_iff {s : Set M} : - @IsAddSubmonoid (Additive M) _ s ↔ IsSubmonoid s := - ⟨fun ⟨h₁, h₂⟩ => ⟨h₁, @h₂⟩, Additive.isAddSubmonoid⟩ - -theorem Multiplicative.isSubmonoid {s : Set A} : - IsAddSubmonoid s → @IsSubmonoid (Multiplicative A) _ s - | ⟨h₁, h₂⟩ => ⟨h₁, @h₂⟩ - -theorem Multiplicative.isSubmonoid_iff {s : Set A} : - @IsSubmonoid (Multiplicative A) _ s ↔ IsAddSubmonoid s := - ⟨fun ⟨h₁, h₂⟩ => ⟨h₁, @h₂⟩, Multiplicative.isSubmonoid⟩ - -/-- The intersection of two submonoids of a monoid `M` is a submonoid of `M`. -/ -@[to_additive - "The intersection of two `AddSubmonoid`s of an `AddMonoid` `M` is an `AddSubmonoid` of M."] -theorem IsSubmonoid.inter {s₁ s₂ : Set M} (is₁ : IsSubmonoid s₁) (is₂ : IsSubmonoid s₂) : - IsSubmonoid (s₁ ∩ s₂) := - { one_mem := ⟨is₁.one_mem, is₂.one_mem⟩ - mul_mem := @fun _ _ hx hy => ⟨is₁.mul_mem hx.1 hy.1, is₂.mul_mem hx.2 hy.2⟩ } - -/-- The intersection of an indexed set of submonoids of a monoid `M` is a submonoid of `M`. -/ -@[to_additive - "The intersection of an indexed set of `AddSubmonoid`s of an `AddMonoid` `M` is - an `AddSubmonoid` of `M`."] -theorem IsSubmonoid.iInter {ι : Sort*} {s : ι → Set M} (h : ∀ y : ι, IsSubmonoid (s y)) : - IsSubmonoid (Set.iInter s) := - { one_mem := Set.mem_iInter.2 fun y => (h y).one_mem - mul_mem := fun h₁ h₂ => - Set.mem_iInter.2 fun y => (h y).mul_mem (Set.mem_iInter.1 h₁ y) (Set.mem_iInter.1 h₂ y) } - -/-- The union of an indexed, directed, nonempty set of submonoids of a monoid `M` is a submonoid - of `M`. -/ -@[to_additive - "The union of an indexed, directed, nonempty set of `AddSubmonoid`s of an `AddMonoid` `M` - is an `AddSubmonoid` of `M`. "] -theorem isSubmonoid_iUnion_of_directed {ι : Type*} [hι : Nonempty ι] {s : ι → Set M} - (hs : ∀ i, IsSubmonoid (s i)) (Directed : ∀ i j, ∃ k, s i ⊆ s k ∧ s j ⊆ s k) : - IsSubmonoid (⋃ i, s i) := - { one_mem := - let ⟨i⟩ := hι - Set.mem_iUnion.2 ⟨i, (hs i).one_mem⟩ - mul_mem := fun ha hb => - let ⟨i, hi⟩ := Set.mem_iUnion.1 ha - let ⟨j, hj⟩ := Set.mem_iUnion.1 hb - let ⟨k, hk⟩ := Directed i j - Set.mem_iUnion.2 ⟨k, (hs k).mul_mem (hk.1 hi) (hk.2 hj)⟩ } - -section powers - -/-- The set of natural number powers `1, x, x², ...` of an element `x` of a monoid. -/ -@[to_additive - "The set of natural number multiples `0, x, 2x, ...` of an element `x` of an `AddMonoid`."] -def powers (x : M) : Set M := - { y | ∃ n : ℕ, x ^ n = y } - -/-- 1 is in the set of natural number powers of an element of a monoid. -/ -@[to_additive "0 is in the set of natural number multiples of an element of an `AddMonoid`."] -theorem powers.one_mem {x : M} : (1 : M) ∈ powers x := - ⟨0, pow_zero _⟩ - -/-- An element of a monoid is in the set of that element's natural number powers. -/ -@[to_additive - "An element of an `AddMonoid` is in the set of that element's natural number multiples."] -theorem powers.self_mem {x : M} : x ∈ powers x := - ⟨1, pow_one _⟩ - -/-- The set of natural number powers of an element of a monoid is closed under multiplication. -/ -@[to_additive - "The set of natural number multiples of an element of an `AddMonoid` is closed under - addition."] -theorem powers.mul_mem {x y z : M} : y ∈ powers x → z ∈ powers x → y * z ∈ powers x := - fun ⟨n₁, h₁⟩ ⟨n₂, h₂⟩ => ⟨n₁ + n₂, by simp only [pow_add, *]⟩ - -/-- The set of natural number powers of an element of a monoid `M` is a submonoid of `M`. -/ -@[to_additive - "The set of natural number multiples of an element of an `AddMonoid` `M` is - an `AddSubmonoid` of `M`."] -theorem powers.isSubmonoid (x : M) : IsSubmonoid (powers x) := - { one_mem := powers.one_mem - mul_mem := powers.mul_mem } - -/-- A monoid is a submonoid of itself. -/ -@[to_additive "An `AddMonoid` is an `AddSubmonoid` of itself."] -theorem Univ.isSubmonoid : IsSubmonoid (@Set.univ M) := by constructor <;> simp - -/-- The preimage of a submonoid under a monoid hom is a submonoid of the domain. -/ -@[to_additive - "The preimage of an `AddSubmonoid` under an `AddMonoid` hom is - an `AddSubmonoid` of the domain."] -theorem IsSubmonoid.preimage {N : Type*} [Monoid N] {f : M → N} (hf : IsMonoidHom f) {s : Set N} - (hs : IsSubmonoid s) : IsSubmonoid (f ⁻¹' s) := - { one_mem := show f 1 ∈ s by (rw [IsMonoidHom.map_one hf]; exact hs.one_mem) - mul_mem := fun {a b} (ha : f a ∈ s) (hb : f b ∈ s) => - show f (a * b) ∈ s by (rw [IsMonoidHom.map_mul' hf]; exact hs.mul_mem ha hb) } - -/-- The image of a submonoid under a monoid hom is a submonoid of the codomain. -/ -@[to_additive - "The image of an `AddSubmonoid` under an `AddMonoid` hom is an `AddSubmonoid` of the - codomain."] -theorem IsSubmonoid.image {γ : Type*} [Monoid γ] {f : M → γ} (hf : IsMonoidHom f) {s : Set M} - (hs : IsSubmonoid s) : IsSubmonoid (f '' s) := - { one_mem := ⟨1, hs.one_mem, hf.map_one⟩ - mul_mem := @fun a b ⟨x, hx⟩ ⟨y, hy⟩ => - ⟨x * y, hs.mul_mem hx.1 hy.1, by rw [hf.map_mul, hx.2, hy.2]⟩ } - -/-- The image of a monoid hom is a submonoid of the codomain. -/ -@[to_additive "The image of an `AddMonoid` hom is an `AddSubmonoid` of the codomain."] -theorem Range.isSubmonoid {γ : Type*} [Monoid γ] {f : M → γ} (hf : IsMonoidHom f) : - IsSubmonoid (Set.range f) := by - rw [← Set.image_univ] - exact Univ.isSubmonoid.image hf - -/-- Submonoids are closed under natural powers. -/ -@[to_additive - "An `AddSubmonoid` is closed under multiplication by naturals."] -theorem IsSubmonoid.pow_mem {a : M} (hs : IsSubmonoid s) (h : a ∈ s) : ∀ {n : ℕ}, a ^ n ∈ s - | 0 => by - rw [pow_zero] - exact hs.one_mem - | n + 1 => by - rw [pow_succ] - exact hs.mul_mem (IsSubmonoid.pow_mem hs h) h - -/-- The set of natural number powers of an element of a `Submonoid` is a subset of the -`Submonoid`. -/ -@[to_additive - "The set of natural number multiples of an element of an `AddSubmonoid` is a subset of - the `AddSubmonoid`."] -theorem IsSubmonoid.powers_subset {a : M} (hs : IsSubmonoid s) (h : a ∈ s) : powers a ⊆ s := - fun _ ⟨_, hx⟩ => hx ▸ hs.pow_mem h - -end powers - -namespace IsSubmonoid - -/-- The product of a list of elements of a submonoid is an element of the submonoid. -/ -@[to_additive - "The sum of a list of elements of an `AddSubmonoid` is an element of the `AddSubmonoid`."] -theorem list_prod_mem (hs : IsSubmonoid s) : ∀ {l : List M}, (∀ x ∈ l, x ∈ s) → l.prod ∈ s - | [], _ => hs.one_mem - | a :: l, h => - suffices a * l.prod ∈ s by simpa - have : a ∈ s ∧ ∀ x ∈ l, x ∈ s := by simpa using h - hs.mul_mem this.1 (list_prod_mem hs this.2) - -/-- The product of a multiset of elements of a submonoid of a `CommMonoid` is an element of -the submonoid. -/ -@[to_additive - "The sum of a multiset of elements of an `AddSubmonoid` of an `AddCommMonoid` - is an element of the `AddSubmonoid`. "] -theorem multiset_prod_mem {M} [CommMonoid M] {s : Set M} (hs : IsSubmonoid s) (m : Multiset M) : - (∀ a ∈ m, a ∈ s) → m.prod ∈ s := by - refine Quotient.inductionOn m fun l hl => ?_ - rw [Multiset.quot_mk_to_coe, Multiset.prod_coe] - exact list_prod_mem hs hl - -/-- The product of elements of a submonoid of a `CommMonoid` indexed by a `Finset` is an element -of the submonoid. -/ -@[to_additive - "The sum of elements of an `AddSubmonoid` of an `AddCommMonoid` indexed by - a `Finset` is an element of the `AddSubmonoid`."] -theorem finset_prod_mem {M A} [CommMonoid M] {s : Set M} (hs : IsSubmonoid s) (f : A → M) : - ∀ t : Finset A, (∀ b ∈ t, f b ∈ s) → (∏ b ∈ t, f b) ∈ s - | ⟨m, hm⟩, _ => multiset_prod_mem hs _ (by simpa) - -end IsSubmonoid - -namespace AddMonoid - -/-- The inductively defined membership predicate for the submonoid generated by a subset of a - monoid. -/ -inductive InClosure (s : Set A) : A → Prop - | basic {a : A} : a ∈ s → InClosure _ a - | zero : InClosure _ 0 - | add {a b : A} : InClosure _ a → InClosure _ b → InClosure _ (a + b) - -end AddMonoid - -namespace Monoid - -/-- The inductively defined membership predicate for the `Submonoid` generated by a subset of an - monoid. -/ -@[to_additive] -inductive InClosure (s : Set M) : M → Prop - | basic {a : M} : a ∈ s → InClosure _ a - | one : InClosure _ 1 - | mul {a b : M} : InClosure _ a → InClosure _ b → InClosure _ (a * b) - -/-- The inductively defined submonoid generated by a subset of a monoid. -/ -@[to_additive - "The inductively defined `AddSubmonoid` generated by a subset of an `AddMonoid`."] -def Closure (s : Set M) : Set M := - { a | InClosure s a } - -@[to_additive] -theorem closure.isSubmonoid (s : Set M) : IsSubmonoid (Closure s) := - { one_mem := InClosure.one - mul_mem := InClosure.mul } - -/-- A subset of a monoid is contained in the submonoid it generates. -/ -@[to_additive - "A subset of an `AddMonoid` is contained in the `AddSubmonoid` it generates."] -theorem subset_closure {s : Set M} : s ⊆ Closure s := fun _ => InClosure.basic - -/-- The submonoid generated by a set is contained in any submonoid that contains the set. -/ -@[to_additive - "The `AddSubmonoid` generated by a set is contained in any `AddSubmonoid` that - contains the set."] -theorem closure_subset {s t : Set M} (ht : IsSubmonoid t) (h : s ⊆ t) : Closure s ⊆ t := fun a ha => - by induction ha <;> simp [h _, *, IsSubmonoid.one_mem, IsSubmonoid.mul_mem] - -/-- Given subsets `t` and `s` of a monoid `M`, if `s ⊆ t`, the submonoid of `M` generated by `s` is - contained in the submonoid generated by `t`. -/ -@[to_additive (attr := gcongr) - "Given subsets `t` and `s` of an `AddMonoid M`, if `s ⊆ t`, the `AddSubmonoid` - of `M` generated by `s` is contained in the `AddSubmonoid` generated by `t`."] -theorem closure_mono {s t : Set M} (h : s ⊆ t) : Closure s ⊆ Closure t := - closure_subset (closure.isSubmonoid t) <| Set.Subset.trans h subset_closure - -/-- The submonoid generated by an element of a monoid equals the set of natural number powers of - the element. -/ -@[to_additive - "The `AddSubmonoid` generated by an element of an `AddMonoid` equals the set of - natural number multiples of the element."] -theorem closure_singleton {x : M} : Closure ({x} : Set M) = powers x := - Set.eq_of_subset_of_subset - (closure_subset (powers.isSubmonoid x) <| Set.singleton_subset_iff.2 <| powers.self_mem) <| - IsSubmonoid.powers_subset (closure.isSubmonoid _) <| - Set.singleton_subset_iff.1 <| subset_closure - -/-- The image under a monoid hom of the submonoid generated by a set equals the submonoid generated - by the image of the set under the monoid hom. -/ -@[to_additive - "The image under an `AddMonoid` hom of the `AddSubmonoid` generated by a set equals - the `AddSubmonoid` generated by the image of the set under the `AddMonoid` hom."] -theorem image_closure {A : Type*} [Monoid A] {f : M → A} (hf : IsMonoidHom f) (s : Set M) : - f '' Closure s = Closure (f '' s) := - le_antisymm - (by - rintro _ ⟨x, hx, rfl⟩ - induction' hx with z hz - · solve_by_elim [subset_closure, Set.mem_image_of_mem] - · rw [hf.map_one] - apply IsSubmonoid.one_mem (closure.isSubmonoid (f '' s)) - · rw [hf.map_mul] - solve_by_elim [(closure.isSubmonoid _).mul_mem] ) - (closure_subset (IsSubmonoid.image hf (closure.isSubmonoid _)) <| - Set.image_subset _ subset_closure) - -/-- Given an element `a` of the submonoid of a monoid `M` generated by a set `s`, there exists -a list of elements of `s` whose product is `a`. -/ -@[to_additive - "Given an element `a` of the `AddSubmonoid` of an `AddMonoid M` generated by - a set `s`, there exists a list of elements of `s` whose sum is `a`."] -theorem exists_list_of_mem_closure {s : Set M} {a : M} (h : a ∈ Closure s) : - ∃ l : List M, (∀ x ∈ l, x ∈ s) ∧ l.prod = a := by - induction h with - | @basic a ha => exists [a]; simp [ha] - | one => exists []; simp - | mul _ _ ha hb => - rcases ha with ⟨la, ha, eqa⟩ - rcases hb with ⟨lb, hb, eqb⟩ - exists la ++ lb - simp only [List.mem_append, or_imp, List.prod_append, eqa.symm, eqb.symm, and_true] - exact fun a => ⟨ha a, hb a⟩ - -/-- Given sets `s, t` of a commutative monoid `M`, `x ∈ M` is in the submonoid of `M` generated by - `s ∪ t` iff there exists an element of the submonoid generated by `s` and an element of the - submonoid generated by `t` whose product is `x`. -/ -@[to_additive - "Given sets `s, t` of a commutative `AddMonoid M`, `x ∈ M` is in the `AddSubmonoid` - of `M` generated by `s ∪ t` iff there exists an element of the `AddSubmonoid` generated by `s` - and an element of the `AddSubmonoid` generated by `t` whose sum is `x`."] -theorem mem_closure_union_iff {M : Type*} [CommMonoid M] {s t : Set M} {x : M} : - x ∈ Closure (s ∪ t) ↔ ∃ y ∈ Closure s, ∃ z ∈ Closure t, y * z = x := - ⟨fun hx => - let ⟨L, HL1, HL2⟩ := exists_list_of_mem_closure hx - HL2 ▸ - List.recOn L - (fun _ => - ⟨1, (closure.isSubmonoid _).one_mem, 1, (closure.isSubmonoid _).one_mem, mul_one _⟩) - (fun hd tl ih HL1 => - let ⟨y, hy, z, hz, hyzx⟩ := ih (List.forall_mem_of_forall_mem_cons HL1) - Or.casesOn (HL1 hd <| List.mem_cons_self _ _) - (fun hs => - ⟨hd * y, (closure.isSubmonoid _).mul_mem (subset_closure hs) hy, z, hz, by - rw [mul_assoc, List.prod_cons, ← hyzx]⟩) - fun ht => - ⟨y, hy, z * hd, (closure.isSubmonoid _).mul_mem hz (subset_closure ht), by - rw [← mul_assoc, List.prod_cons, ← hyzx, mul_comm hd]⟩) - HL1, - fun ⟨_, hy, _, hz, hyzx⟩ => - hyzx ▸ - (closure.isSubmonoid _).mul_mem (closure_mono Set.subset_union_left hy) - (closure_mono Set.subset_union_right hz)⟩ - -end Monoid - -/-- Create a bundled submonoid from a set `s` and `[IsSubmonoid s]`. -/ -@[to_additive "Create a bundled additive submonoid from a set `s` and `[IsAddSubmonoid s]`."] -def Submonoid.of {s : Set M} (h : IsSubmonoid s) : Submonoid M := - ⟨⟨s, @fun _ _ => h.2⟩, h.1⟩ - -@[to_additive] -theorem Submonoid.isSubmonoid (S : Submonoid M) : IsSubmonoid (S : Set M) := by - exact ⟨S.2, S.1.2⟩ diff --git a/Mathlib/Deprecated/Subring.lean b/Mathlib/Deprecated/Subring.lean deleted file mode 100644 index 05bea1fc07bb06..00000000000000 --- a/Mathlib/Deprecated/Subring.lean +++ /dev/null @@ -1,205 +0,0 @@ -/- -Copyright (c) 2018 Johan Commelin. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johan Commelin --/ -import Mathlib.Algebra.Ring.Subring.Defs -import Mathlib.Deprecated.Group -import Mathlib.Deprecated.Subgroup - -/-! -# Unbundled subrings (deprecated) - -This file is deprecated, and is no longer imported by anything in mathlib other than other -deprecated files, and test files. You should not need to import it. - -This file defines predicates for unbundled subrings. Instead of using this file, please use -`Subring`, defined in `Mathlib.Algebra.Ring.Subring.Basic`, for subrings of rings. - -## Main definitions - -`IsSubring (S : Set R) : Prop` : the predicate that `S` is the underlying set of a subring -of the ring `R`. The bundled variant `Subring R` should be used in preference to this. - -## Tags - -IsSubring --/ - - -universe u v - -open Group - -variable {R : Type u} [Ring R] - -/-- `S` is a subring: a set containing 1 and closed under multiplication, addition and additive -inverse. -/ -structure IsSubring (S : Set R) extends IsAddSubgroup S, IsSubmonoid S : Prop - -/-- Construct a `Subring` from a set satisfying `IsSubring`. -/ -def IsSubring.subring {S : Set R} (hs : IsSubring S) : Subring R where - carrier := S - one_mem' := hs.one_mem - mul_mem' := hs.mul_mem - zero_mem' := hs.zero_mem - add_mem' := hs.add_mem - neg_mem' := hs.neg_mem - -namespace RingHom - -theorem isSubring_preimage {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) {s : Set S} - (hs : IsSubring s) : IsSubring (f ⁻¹' s) := - { IsAddGroupHom.preimage f.to_isAddGroupHom hs.toIsAddSubgroup, - IsSubmonoid.preimage f.to_isMonoidHom hs.toIsSubmonoid with } - -theorem isSubring_image {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) {s : Set R} - (hs : IsSubring s) : IsSubring (f '' s) := - { IsAddGroupHom.image_addSubgroup f.to_isAddGroupHom hs.toIsAddSubgroup, - IsSubmonoid.image f.to_isMonoidHom hs.toIsSubmonoid with } - -theorem isSubring_set_range {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) : - IsSubring (Set.range f) := - { IsAddGroupHom.range_addSubgroup f.to_isAddGroupHom, Range.isSubmonoid f.to_isMonoidHom with } - -end RingHom - -theorem IsSubring.inter {S₁ S₂ : Set R} (hS₁ : IsSubring S₁) (hS₂ : IsSubring S₂) : - IsSubring (S₁ ∩ S₂) := - { IsAddSubgroup.inter hS₁.toIsAddSubgroup hS₂.toIsAddSubgroup, - IsSubmonoid.inter hS₁.toIsSubmonoid hS₂.toIsSubmonoid with } - -theorem IsSubring.iInter {ι : Sort*} {S : ι → Set R} (h : ∀ y : ι, IsSubring (S y)) : - IsSubring (Set.iInter S) := - { IsAddSubgroup.iInter fun i ↦ (h i).toIsAddSubgroup, - IsSubmonoid.iInter fun i ↦ (h i).toIsSubmonoid with } - -theorem isSubring_iUnion_of_directed {ι : Type*} [Nonempty ι] {s : ι → Set R} - (h : ∀ i, IsSubring (s i)) (directed : ∀ i j, ∃ k, s i ⊆ s k ∧ s j ⊆ s k) : - IsSubring (⋃ i, s i) := - { toIsAddSubgroup := isAddSubgroup_iUnion_of_directed (fun i ↦ (h i).toIsAddSubgroup) directed - toIsSubmonoid := isSubmonoid_iUnion_of_directed (fun i ↦ (h i).toIsSubmonoid) directed } - -namespace Ring - -/-- The smallest subring containing a given subset of a ring, considered as a set. This function -is deprecated; use `Subring.closure`. -/ -def closure (s : Set R) := - AddGroup.closure (Monoid.Closure s) - -variable {s : Set R} - --- attribute [local reducible] closure -- Porting note: not available in Lean4 - -theorem exists_list_of_mem_closure {a : R} (h : a ∈ closure s) : - ∃ L : List (List R), (∀ l ∈ L, ∀ x ∈ l, x ∈ s ∨ x = (-1 : R)) ∧ (L.map List.prod).sum = a := - AddGroup.InClosure.recOn h - fun {x} hx ↦ match x, Monoid.exists_list_of_mem_closure hx with - | _, ⟨L, h1, rfl⟩ => - ⟨[L], List.forall_mem_singleton.2 fun r hr ↦ Or.inl (h1 r hr), List.sum_singleton⟩ - ⟨[], List.forall_mem_nil _, rfl⟩ - fun {b} _ ih ↦ match b, ih with - | _, ⟨L1, h1, rfl⟩ => - ⟨L1.map (List.cons (-1)), - fun L2 h2 ↦ match L2, List.mem_map.1 h2 with - | _, ⟨L3, h3, rfl⟩ => List.forall_mem_cons.2 ⟨Or.inr rfl, h1 L3 h3⟩, by - simp only [List.map_map, Function.comp_def, List.prod_cons, neg_one_mul] - refine List.recOn L1 neg_zero.symm fun hd tl ih ↦ ?_ - rw [List.map_cons, List.sum_cons, ih, List.map_cons, List.sum_cons, neg_add]⟩ - fun {r1 r2} _ _ ih1 ih2 ↦ match r1, r2, ih1, ih2 with - | _, _, ⟨L1, h1, rfl⟩, ⟨L2, h2, rfl⟩ => - ⟨L1 ++ L2, List.forall_mem_append.2 ⟨h1, h2⟩, by rw [List.map_append, List.sum_append]⟩ - -@[elab_as_elim] -protected theorem InClosure.recOn {C : R → Prop} {x : R} (hx : x ∈ closure s) (h1 : C 1) - (hneg1 : C (-1)) (hs : ∀ z ∈ s, ∀ n, C n → C (z * n)) (ha : ∀ {x y}, C x → C y → C (x + y)) : - C x := by - have h0 : C 0 := add_neg_cancel (1 : R) ▸ ha h1 hneg1 - rcases exists_list_of_mem_closure hx with ⟨L, HL, rfl⟩ - clear hx - induction' L with hd tl ih - · exact h0 - rw [List.forall_mem_cons] at HL - suffices C (List.prod hd) by - rw [List.map_cons, List.sum_cons] - exact ha this (ih HL.2) - replace HL := HL.1 - clear ih tl - -- Porting note: Expanded `rsuffices` - suffices ∃ L, (∀ x ∈ L, x ∈ s) ∧ (List.prod hd = List.prod L ∨ List.prod hd = -List.prod L) by - rcases this with ⟨L, HL', HP | HP⟩ <;> rw [HP] <;> clear HP HL - · induction' L with hd tl ih - · exact h1 - rw [List.forall_mem_cons] at HL' - rw [List.prod_cons] - exact hs _ HL'.1 _ (ih HL'.2) - · induction' L with hd tl ih - · exact hneg1 - rw [List.prod_cons, neg_mul_eq_mul_neg] - rw [List.forall_mem_cons] at HL' - exact hs _ HL'.1 _ (ih HL'.2) - induction' hd with hd tl ih - · exact ⟨[], List.forall_mem_nil _, Or.inl rfl⟩ - rw [List.forall_mem_cons] at HL - rcases ih HL.2 with ⟨L, HL', HP | HP⟩ <;> cases' HL.1 with hhd hhd - · exact - ⟨hd::L, List.forall_mem_cons.2 ⟨hhd, HL'⟩, - Or.inl <| by rw [List.prod_cons, List.prod_cons, HP]⟩ - · exact ⟨L, HL', Or.inr <| by rw [List.prod_cons, hhd, neg_one_mul, HP]⟩ - · exact - ⟨hd::L, List.forall_mem_cons.2 ⟨hhd, HL'⟩, - Or.inr <| by rw [List.prod_cons, List.prod_cons, HP, neg_mul_eq_mul_neg]⟩ - · exact ⟨L, HL', Or.inl <| by rw [List.prod_cons, hhd, HP, neg_one_mul, neg_neg]⟩ - -theorem closure.isSubring : IsSubring (closure s) := - { AddGroup.closure.isAddSubgroup _ with - one_mem := AddGroup.mem_closure <| IsSubmonoid.one_mem <| Monoid.closure.isSubmonoid _ - mul_mem := fun {a _} ha hb ↦ AddGroup.InClosure.recOn hb - (fun {c} hc ↦ AddGroup.InClosure.recOn ha - (fun hd ↦ AddGroup.subset_closure ((Monoid.closure.isSubmonoid _).mul_mem hd hc)) - ((zero_mul c).symm ▸ (AddGroup.closure.isAddSubgroup _).zero_mem) - (fun {d} _ hdc ↦ neg_mul_eq_neg_mul d c ▸ (AddGroup.closure.isAddSubgroup _).neg_mem hdc) - fun {d e} _ _ hdc hec ↦ - (add_mul d e c).symm ▸ (AddGroup.closure.isAddSubgroup _).add_mem hdc hec) - ((mul_zero a).symm ▸ (AddGroup.closure.isAddSubgroup _).zero_mem) - (fun {c} _ hac ↦ neg_mul_eq_mul_neg a c ▸ (AddGroup.closure.isAddSubgroup _).neg_mem hac) - fun {c d} _ _ hac had ↦ - (mul_add a c d).symm ▸ (AddGroup.closure.isAddSubgroup _).add_mem hac had } - -theorem mem_closure {a : R} : a ∈ s → a ∈ closure s := - AddGroup.mem_closure ∘ @Monoid.subset_closure _ _ _ _ - -theorem subset_closure : s ⊆ closure s := - fun _ ↦ mem_closure - -theorem closure_subset {t : Set R} (ht : IsSubring t) : s ⊆ t → closure s ⊆ t := - AddGroup.closure_subset ht.toIsAddSubgroup ∘ Monoid.closure_subset ht.toIsSubmonoid - -theorem closure_subset_iff {s t : Set R} (ht : IsSubring t) : closure s ⊆ t ↔ s ⊆ t := - (AddGroup.closure_subset_iff ht.toIsAddSubgroup).trans - ⟨Set.Subset.trans Monoid.subset_closure, Monoid.closure_subset ht.toIsSubmonoid⟩ - -@[gcongr] -theorem closure_mono {s t : Set R} (H : s ⊆ t) : closure s ⊆ closure t := - closure_subset closure.isSubring <| Set.Subset.trans H subset_closure - -theorem image_closure {S : Type*} [Ring S] (f : R →+* S) (s : Set R) : - f '' closure s = closure (f '' s) := by - refine le_antisymm ?_ (closure_subset (RingHom.isSubring_image _ closure.isSubring) <| - Set.image_subset _ subset_closure) - rintro _ ⟨x, hx, rfl⟩ - apply AddGroup.InClosure.recOn (motive := fun {x} _ ↦ f x ∈ closure (f '' s)) hx _ <;> intros - · rw [f.map_zero] - apply closure.isSubring.zero_mem - · rw [f.map_neg] - apply closure.isSubring.neg_mem - assumption - · rw [f.map_add] - apply closure.isSubring.add_mem - assumption' - · apply AddGroup.mem_closure - rw [← Monoid.image_closure f.to_isMonoidHom] - apply Set.mem_image_of_mem - assumption - -end Ring diff --git a/Mathlib/FieldTheory/KummerExtension.lean b/Mathlib/FieldTheory/KummerExtension.lean index 50462a0390dace..352ba6623e838b 100644 --- a/Mathlib/FieldTheory/KummerExtension.lean +++ b/Mathlib/FieldTheory/KummerExtension.lean @@ -175,10 +175,8 @@ theorem Polynomial.separable_X_pow_sub_C_of_irreducible : (X ^ n - C a).Separabl AdjoinRoot.algebraMap_eq, X_pow_sub_C_eq_prod (hζ.map_of_injective (algebraMap K _).injective) hn (root_X_pow_sub_C_pow n a), separable_prod_X_sub_C_iff'] - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/5376 we need to provide this helper instance. - -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/5376 + we need to provide this helper instance. -/ have : MonoidHomClass (K →+* K[n√a]) K K[n√a] := inferInstance exact (hζ.map_of_injective (algebraMap K K[n√a]).injective).injOn_pow_mul (root_X_pow_sub_C_ne_zero (lt_of_le_of_ne (show 1 ≤ n from hn) (Ne.symm hn')) _) diff --git a/Mathlib/FieldTheory/Minpoly/Basic.lean b/Mathlib/FieldTheory/Minpoly/Basic.lean index 0b88000978e5fe..8c19cd2535cbdd 100644 --- a/Mathlib/FieldTheory/Minpoly/Basic.lean +++ b/Mathlib/FieldTheory/Minpoly/Basic.lean @@ -15,7 +15,6 @@ such as irreducibility under the assumption `B` is a domain. -/ -open scoped Classical open Polynomial Set Function variable {A B B' : Type*} @@ -24,6 +23,7 @@ section MinPolyDef variable (A) [CommRing A] [Ring B] [Algebra A B] +open scoped Classical in /-- Suppose `x : B`, where `B` is an `A`-algebra. The minimal polynomial `minpoly A x` of `x` @@ -64,6 +64,7 @@ theorem ne_zero_iff [Nontrivial A] : minpoly A x ≠ 0 ↔ IsIntegral A x := theorem algHom_eq (f : B →ₐ[A] B') (hf : Function.Injective f) (x : B) : minpoly A (f x) = minpoly A x := by + classical simp_rw [minpoly, isIntegral_algHom_iff _ hf, ← Polynomial.aeval_def, aeval_algHom, AlgHom.comp_apply, _root_.map_eq_zero_iff f hf] diff --git a/Mathlib/FieldTheory/Minpoly/Field.lean b/Mathlib/FieldTheory/Minpoly/Field.lean index 0652b50888c6b4..b3f90cac4e69d1 100644 --- a/Mathlib/FieldTheory/Minpoly/Field.lean +++ b/Mathlib/FieldTheory/Minpoly/Field.lean @@ -18,7 +18,6 @@ are irreducible, and uniquely determined by their defining property. -/ -open scoped Classical open Polynomial Set Function minpoly namespace minpoly @@ -181,6 +180,7 @@ theorem neg {B : Type*} [CommRing B] [Algebra A B] (x : B) : section AlgHomFintype +open scoped Classical in /-- A technical finiteness result. -/ noncomputable def Fintype.subtypeProd {E : Type*} {X : Set E} (hX : X.Finite) {L : Type*} (F : E → Multiset L) : Fintype (∀ x : X, { l : L // l ∈ F x }) := diff --git a/Mathlib/FieldTheory/RatFunc/AsPolynomial.lean b/Mathlib/FieldTheory/RatFunc/AsPolynomial.lean index 14a02db0ca5883..5c4ceecc798cce 100644 --- a/Mathlib/FieldTheory/RatFunc/AsPolynomial.lean +++ b/Mathlib/FieldTheory/RatFunc/AsPolynomial.lean @@ -32,8 +32,6 @@ namespace RatFunc section Eval -open scoped Classical - open scoped nonZeroDivisors Polynomial open RatFunc diff --git a/Mathlib/FieldTheory/RatFunc/Basic.lean b/Mathlib/FieldTheory/RatFunc/Basic.lean index c5ca7de7845f57..9a1cde7fbd08bb 100644 --- a/Mathlib/FieldTheory/RatFunc/Basic.lean +++ b/Mathlib/FieldTheory/RatFunc/Basic.lean @@ -48,8 +48,6 @@ universe u v noncomputable section -open scoped Classical - open scoped nonZeroDivisors Polynomial variable {K : Type u} @@ -320,6 +318,8 @@ open RatFunc variable {G₀ L R S F : Type*} [CommGroupWithZero G₀] [Field L] [CommRing R] [CommRing S] variable [FunLike F R[X] S[X]] + +open scoped Classical in /-- Lift a monoid homomorphism that maps polynomials `φ : R[X] →* S[X]` to a `RatFunc R →* RatFunc S`, on the condition that `φ` maps non zero divisors to non zero divisors, @@ -741,6 +741,7 @@ open GCDMonoid Polynomial variable [Field K] +open scoped Classical in /-- `RatFunc.numDenom` are numerator and denominator of a rational function over a field, normalized such that the denominator is monic. -/ def numDenom (x : RatFunc K) : K[X] × K[X] := @@ -775,6 +776,7 @@ def numDenom (x : RatFunc K) : K[X] × K[X] := rw [inv_div, mul_comm, mul_div_assoc, ← mul_assoc, inv_inv, mul_inv_cancel₀ ha', one_mul, inv_div]) +open scoped Classical in @[simp] theorem numDenom_div (p : K[X]) {q : K[X]} (hq : q ≠ 0) : numDenom (algebraMap _ _ p / algebraMap _ _ q) = @@ -790,6 +792,7 @@ normalized such that the denominator is monic. -/ def num (x : RatFunc K) : K[X] := x.numDenom.1 +open scoped Classical in private theorem num_div' (p : K[X]) {q : K[X]} (hq : q ≠ 0) : num (algebraMap _ _ p / algebraMap _ _ q) = Polynomial.C (q / gcd p q).leadingCoeff⁻¹ * (p / gcd p q) := by @@ -798,6 +801,7 @@ private theorem num_div' (p : K[X]) {q : K[X]} (hq : q ≠ 0) : @[simp] theorem num_zero : num (0 : RatFunc K) = 0 := by convert num_div' (0 : K[X]) one_ne_zero <;> simp +open scoped Classical in @[simp] theorem num_div (p q : K[X]) : num (algebraMap _ _ p / algebraMap _ _ q) = @@ -814,10 +818,12 @@ theorem num_algebraMap (p : K[X]) : num (algebraMap _ _ p) = p := by convert num theorem num_div_dvd (p : K[X]) {q : K[X]} (hq : q ≠ 0) : num (algebraMap _ _ p / algebraMap _ _ q) ∣ p := by + classical rw [num_div _ q, C_mul_dvd] · exact EuclideanDomain.div_dvd_of_dvd (gcd_dvd_left p q) · simpa only [Ne, inv_eq_zero, Polynomial.leadingCoeff_eq_zero] using right_div_gcd_ne_zero hq +open scoped Classical in /-- A version of `num_div_dvd` with the LHS in simp normal form -/ @[simp] theorem num_div_dvd' (p : K[X]) {q : K[X]} (hq : q ≠ 0) : @@ -828,6 +834,7 @@ normalized such that it is monic. -/ def denom (x : RatFunc K) : K[X] := x.numDenom.2 +open scoped Classical in @[simp] theorem denom_div (p : K[X]) {q : K[X]} (hq : q ≠ 0) : denom (algebraMap _ _ p / algebraMap _ _ q) = @@ -835,6 +842,7 @@ theorem denom_div (p : K[X]) {q : K[X]} (hq : q ≠ 0) : rw [denom, numDenom_div _ hq] theorem monic_denom (x : RatFunc K) : (denom x).Monic := by + classical induction x using RatFunc.induction_on with | f p q hq => rw [denom_div p hq, mul_comm] @@ -857,6 +865,7 @@ theorem denom_algebraMap (p : K[X]) : denom (algebraMap _ (RatFunc K) p) = 1 := @[simp] theorem denom_div_dvd (p q : K[X]) : denom (algebraMap _ _ p / algebraMap _ _ q) ∣ q := by + classical by_cases hq : q = 0 · simp [hq] rw [denom_div _ hq, C_mul_dvd] @@ -865,6 +874,7 @@ theorem denom_div_dvd (p q : K[X]) : denom (algebraMap _ _ p / algebraMap _ _ q) @[simp] theorem num_div_denom (x : RatFunc K) : algebraMap _ _ (num x) / algebraMap _ _ (denom x) = x := by + classical induction' x using RatFunc.induction_on with p q hq have q_div_ne_zero : q / gcd p q ≠ 0 := right_div_gcd_ne_zero hq rw [num_div p q, denom_div p hq, RingHom.map_mul, RingHom.map_mul, mul_div_mul_left, @@ -878,6 +888,7 @@ theorem num_div_denom (x : RatFunc K) : algebraMap _ _ (num x) / algebraMap _ _ exact inv_ne_zero (Polynomial.leadingCoeff_ne_zero.mpr q_div_ne_zero) theorem isCoprime_num_denom (x : RatFunc K) : IsCoprime x.num x.denom := by + classical induction' x using RatFunc.induction_on with p q hq rw [num_div, denom_div _ hq] exact (isCoprime_mul_unit_left diff --git a/Mathlib/FieldTheory/RatFunc/Defs.lean b/Mathlib/FieldTheory/RatFunc/Defs.lean index a0e7e1e83de5be..c3fc63ae2deaa7 100644 --- a/Mathlib/FieldTheory/RatFunc/Defs.lean +++ b/Mathlib/FieldTheory/RatFunc/Defs.lean @@ -48,8 +48,6 @@ the codomain is not a field or even an integral domain. noncomputable section -open scoped Classical - open scoped nonZeroDivisors Polynomial universe u v diff --git a/Mathlib/FieldTheory/Relrank.lean b/Mathlib/FieldTheory/Relrank.lean index 2ed666bbdae0ab..aa621bbab5393f 100644 --- a/Mathlib/FieldTheory/Relrank.lean +++ b/Mathlib/FieldTheory/Relrank.lean @@ -35,12 +35,10 @@ variable {E : Type v} [Field E] {L : Type w} [Field L] variable (A B C : Subfield E) -#adaptation_note -/-- +#adaptation_note /-- nightly-2024-11-14 This `synthInstance.maxHeartbeats` (and below) was required after nightly-2024-11-14; it's not exactly clear why, but we were very close to the limit previously, -so probably we should not particularly blame changes in Lean, and instead optimize in Mathlib. --/ +so probably we should not particularly blame changes in Lean, and instead optimize in Mathlib. -/ set_option synthInstance.maxHeartbeats 400000 in /-- `Subfield.relrank A B` is defined to be `[B : A ⊓ B]` as a `Cardinal`, in particular, when `A ≤ B` it is `[B : A]`, the degree of the field extension `B / A`. diff --git a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean index ef7c492a1d056c..7fc9dd70fae57d 100644 --- a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean @@ -40,7 +40,6 @@ sphere. -/ def inversion (c : P) (R : ℝ) (x : P) : P := (R / dist x c) ^ 2 • (x -ᵥ c) +ᵥ c -#adaptation_note /-- nightly-2024-03-16: added to replace simp [inversion] -/ theorem inversion_def : inversion = fun (c : P) (R : ℝ) (x : P) => (R / dist x c) ^ 2 • (x -ᵥ c) +ᵥ c := rfl diff --git a/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean b/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean index 877daf129a72af..f948ab3a60200a 100644 --- a/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean +++ b/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean @@ -89,9 +89,7 @@ theorem hasFDerivAt_inversion (hx : x ≠ c) : ((R / dist x c) ^ 2 • (reflection (ℝ ∙ (x - c))ᗮ : F →L[ℝ] F)) x := by rcases add_left_surjective c x with ⟨x, rfl⟩ have : HasFDerivAt (inversion c R) (?_ : F →L[ℝ] F) (c + x) := by - #adaptation_note /-- nightly-2024-03-16: simp was - simp (config := { unfoldPartialApp := true }) only [inversion] -/ - simp only [inversion_def] + simp +unfoldPartialApp only [inversion] simp_rw [dist_eq_norm, div_pow, div_eq_mul_inv] have A := (hasFDerivAt_id (𝕜 := ℝ) (c + x)).sub_const c have B := ((hasDerivAt_inv <| by simpa using hx).comp_hasFDerivAt _ A.norm_sq).const_mul diff --git a/Mathlib/Geometry/Manifold/AnalyticManifold.lean b/Mathlib/Geometry/Manifold/AnalyticManifold.lean index 9b6c9a3025046a..700a30607806f2 100644 --- a/Mathlib/Geometry/Manifold/AnalyticManifold.lean +++ b/Mathlib/Geometry/Manifold/AnalyticManifold.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Lee, Geoffrey Irving -/ import Mathlib.Analysis.Analytic.Constructions -import Mathlib.Geometry.Manifold.IsManifold +import Mathlib.Geometry.Manifold.IsManifold.Basic /-! # Analytic manifolds (possibly with boundary or corners) diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 7bcb7a3778950f..c5f523f226b655 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -851,7 +851,7 @@ theorem piChartedSpace_chartAt {ι : Type*} [Finite ι] (H : ι → Type*) section sum variable [TopologicalSpace H] [TopologicalSpace M] [TopologicalSpace M'] - [cm : ChartedSpace H M] [cm' : ChartedSpace H M'] [Nonempty H] [Nonempty M] [Nonempty M'] + [cm : ChartedSpace H M] [cm' : ChartedSpace H M'] [Nonempty H] /-- The disjoint union of two charted spaces on `H` is a charted space over `H`. -/ instance ChartedSpace.sum : ChartedSpace H (M ⊕ M') where diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index 4bb05df82f4ff0..ee537cfccc4d7c 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn -/ -import Mathlib.Geometry.Manifold.IsManifold +import Mathlib.Geometry.Manifold.IsManifold.Basic import Mathlib.Geometry.Manifold.LocalInvariantProperties /-! diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 78db01d88e9e28..f18b90045c1e6a 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -3,8 +3,8 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Geometry.Manifold.InteriorBoundary -import Mathlib.Geometry.Manifold.IsManifold +import Mathlib.Geometry.Manifold.IsManifold.Basic +import Mathlib.Geometry.Manifold.IsManifold.InteriorBoundary import Mathlib.Analysis.InnerProductSpace.PiL2 /-! diff --git a/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean b/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean index 9a921471207e6b..3f35ee20c69bf7 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean @@ -6,7 +6,7 @@ Authors: Winston Yin import Mathlib.Analysis.ODE.Gronwall import Mathlib.Analysis.ODE.PicardLindelof import Mathlib.Geometry.Manifold.IntegralCurve.Transform -import Mathlib.Geometry.Manifold.InteriorBoundary +import Mathlib.Geometry.Manifold.IsManifold.InteriorBoundary /-! # Existence and uniqueness of integral curves diff --git a/Mathlib/Geometry/Manifold/IsManifold.lean b/Mathlib/Geometry/Manifold/IsManifold/Basic.lean similarity index 99% rename from Mathlib/Geometry/Manifold/IsManifold.lean rename to Mathlib/Geometry/Manifold/IsManifold/Basic.lean index 69d14ee8012e61..069cea5c3f202d 100644 --- a/Mathlib/Geometry/Manifold/IsManifold.lean +++ b/Mathlib/Geometry/Manifold/IsManifold/Basic.lean @@ -829,7 +829,7 @@ variable {M' : Type*} [TopologicalSpace M'] [ChartedSpace H M'] /-- The disjoint union of two `C^n` manifolds modelled on `(E, H)` is a `C^n` manifold modeled on `(E, H)`. -/ -instance disjointUnion [Nonempty M] [Nonempty M'] [Nonempty H] : IsManifold I n (M ⊕ M') where +instance disjointUnion [Nonempty H] : IsManifold I n (M ⊕ M') where compatible {e} e' he he' := by obtain (⟨f, hf, hef⟩ | ⟨f, hf, hef⟩) := ChartedSpace.mem_atlas_sum he · obtain (⟨f', hf', he'f'⟩ | ⟨f', hf', he'f'⟩) := ChartedSpace.mem_atlas_sum he' diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/IsManifold/InteriorBoundary.lean similarity index 98% rename from Mathlib/Geometry/Manifold/InteriorBoundary.lean rename to Mathlib/Geometry/Manifold/IsManifold/InteriorBoundary.lean index 7f6c2a06f8852a..9cc24f55b41587 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/IsManifold/InteriorBoundary.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Rothgang -/ -import Mathlib.Geometry.Manifold.IsManifold +import Mathlib.Geometry.Manifold.IsManifold.Basic /-! # Interior and boundary of a manifold @@ -250,12 +250,11 @@ end prod section disjointUnion variable {M' : Type*} [TopologicalSpace M'] [ChartedSpace H M'] {n : WithTop ℕ∞} - [hM : IsManifold I n M] [hM' : IsManifold I n M'] - [Nonempty M] [Nonempty M'] [Nonempty H] + [hM : IsManifold I n M] [hM' : IsManifold I n M'] [Nonempty H] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {J : Type*} {J : ModelWithCorners 𝕜 E' H'} {N N' : Type*} [TopologicalSpace N] [TopologicalSpace N'] [ChartedSpace H' N] [ChartedSpace H' N'] - [IsManifold J n N] [IsManifold J n N'] [Nonempty N] [Nonempty N'] [Nonempty H'] + [IsManifold J n N] [IsManifold J n N'] [Nonempty H'] open Topology diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean index c269a880e9c260..8fae4d69ee6bdd 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn -/ -import Mathlib.Geometry.Manifold.IsManifold +import Mathlib.Geometry.Manifold.IsManifold.Basic import Mathlib.Geometry.Manifold.LocalInvariantProperties /-! diff --git a/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean b/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean index d920ed00741905..158ebd2729aa13 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean @@ -81,12 +81,9 @@ end Module /-! ### Linear maps between normed spaces are differentiable -/ -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/6024 +#adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 we needed to add the named arguments `𝕜 := 𝕜` and `F := ((F₂ →L[𝕜] F₃) →L[𝕜] F₁ →L[𝕜] F₃))` -to `ContinuousLinearMap.differentiable`. --/ +to `ContinuousLinearMap.differentiable`. -/ theorem MDifferentiableWithinAt.clm_precomp {f : M → F₁ →L[𝕜] F₂} {s : Set M} {x : M} (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) f s x) : MDifferentiableWithinAt I 𝓘(𝕜, (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) @@ -95,12 +92,9 @@ theorem MDifferentiableWithinAt.clm_precomp {f : M → F₁ →L[𝕜] F₂} {s (ContinuousLinearMap.differentiable (𝕜 := 𝕜) (F := ((F₂ →L[𝕜] F₃) →L[𝕜] F₁ →L[𝕜] F₃)) (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).flip) hf -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/6024 +#adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 we needed to add the named arguments `𝕜 := 𝕜` and `F := ((F₂ →L[𝕜] F₃) →L[𝕜] F₁ →L[𝕜] F₃))` -to `ContinuousLinearMap.differentiable`. --/ +to `ContinuousLinearMap.differentiable`. -/ nonrec theorem MDifferentiableAt.clm_precomp {f : M → F₁ →L[𝕜] F₂} {x : M} (hf : MDifferentiableAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) f x) : MDifferentiableAt I 𝓘(𝕜, (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) @@ -121,12 +115,9 @@ theorem MDifferentiable.clm_precomp (fun y ↦ (f y).precomp F₃ : M → (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) := fun x ↦ (hf x).clm_precomp -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/6024 +#adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 we needed to add the named arguments `𝕜 := 𝕜` and `F := ((F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃))` -to `ContinuousLinearMap.differentiable`. --/ +to `ContinuousLinearMap.differentiable`. -/ theorem MDifferentiableWithinAt.clm_postcomp {f : M → F₂ →L[𝕜] F₃} {s : Set M} {x : M} (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) f s x) : MDifferentiableWithinAt I 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) @@ -136,12 +127,9 @@ theorem MDifferentiableWithinAt.clm_postcomp {f : M → F₂ →L[𝕜] F₃} {s (ContinuousLinearMap.differentiable (𝕜 := 𝕜) (F := ((F₁ →L[𝕜] F₂) →L[𝕜] F₁ →L[𝕜] F₃)) (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃)) hf -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/6024 +#adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 we needed to add the named arguments `𝕜 := 𝕜` and `F := ((F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃))` -to `ContinuousLinearMap.differentiable`. --/ +to `ContinuousLinearMap.differentiable`. -/ theorem MDifferentiableAt.clm_postcomp {f : M → F₂ →L[𝕜] F₃} {x : M} (hf : MDifferentiableAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) f x) : MDifferentiableAt I 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) diff --git a/Mathlib/Geometry/Manifold/Metrizable.lean b/Mathlib/Geometry/Manifold/Metrizable.lean index 5427ae259b734b..fa2cb6d5dd8920 100644 --- a/Mathlib/Geometry/Manifold/Metrizable.lean +++ b/Mathlib/Geometry/Manifold/Metrizable.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Geometry.Manifold.IsManifold +import Mathlib.Geometry.Manifold.IsManifold.Basic import Mathlib.Topology.Compactness.Paracompact import Mathlib.Topology.Metrizable.Urysohn diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index e05dffc7df37db..a4c26869dcdfeb 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -227,9 +227,10 @@ theorem coequalizer_π_stalk_isLocalHom (x : Y) : constructor rintro a ha rcases TopCat.Presheaf.germ_exist _ _ a with ⟨U, hU, s, rfl⟩ - rw [← CommRingCat.forget_map_apply, PresheafedSpace.stalkMap_germ_apply - (coequalizer.π (C := SheafedSpace _) f.toShHom g.toShHom) U _ hU] at ha - rw [CommRingCat.forget_map_apply] + rw [← CommRingCat.forget_map_apply, forget_map_eq_coe, + PresheafedSpace.stalkMap_germ_apply + (coequalizer.π (C := SheafedSpace _) f.toShHom g.toShHom) U _ hU] at ha + rw [coe_toHasForget_instFunLike] let V := imageBasicOpen f g U s have hV : (coequalizer.π f.toShHom g.toShHom).base ⁻¹' ((coequalizer.π f.toShHom g.toShHom).base '' V.1) = V.1 := diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean index e96b86ffdeceb6..5120aff056674d 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean @@ -113,7 +113,7 @@ lemma residueFieldMap_comp {Z : LocallyRingedSpace.{u}} (g : Y ⟶ Z) (x : X) : simp only [comp_toShHom, SheafedSpace.comp_base, Function.comp_apply, residueFieldMap, CommRingCat.hom_comp, TopCat.comp_app] simp_rw [stalkMap_comp] - apply IsLocalRing.ResidueField.map_comp + apply IsLocalRing.ResidueField.map_comp (Hom.stalkMap g (f.base x)).hom (Hom.stalkMap f x).hom @[reassoc] lemma evaluation_naturality {V : Opens Y} (x : (Opens.map f.base).obj V) : diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index c3fefa900b4902..912cb224dd41a2 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -192,14 +192,16 @@ theorem snd_invApp_t_app' (i j k : D.J) (U : Opens (pullback (D.f i j) (D.f i k) rw [this, Set.image_comp, Set.image_comp, Set.preimage_image_eq] swap · refine Function.HasLeftInverse.injective ⟨(D.t i k).base, fun x => ?_⟩ - rw [← comp_apply, ← comp_base, D.t_inv, id_base, id_apply] + rw [← CategoryTheory.comp_apply, ← comp_base, D.t_inv, id_base, CategoryTheory.id_apply] refine congr_arg (_ '' ·) ?_ refine congr_fun ?_ _ refine Set.image_eq_preimage_of_inverse ?_ ?_ · intro x - rw [← comp_apply, ← comp_base, IsIso.inv_hom_id, id_base, id_apply] + rw [← CategoryTheory.comp_apply, ← comp_base, IsIso.inv_hom_id, id_base, + CategoryTheory.id_apply] · intro x - rw [← comp_apply, ← comp_base, IsIso.hom_inv_id, id_base, id_apply] + rw [← CategoryTheory.comp_apply, ← comp_base, IsIso.hom_inv_id, id_base, + CategoryTheory.id_apply] · rw [← IsIso.eq_inv_comp, IsOpenImmersion.inv_invApp, Category.assoc, (D.t' k i j).c.naturality_assoc] simp_rw [← Category.assoc] diff --git a/Mathlib/GroupTheory/CommutingProbability.lean b/Mathlib/GroupTheory/CommutingProbability.lean index 606360fa0cce66..815d7acb04b4e7 100644 --- a/Mathlib/GroupTheory/CommutingProbability.lean +++ b/Mathlib/GroupTheory/CommutingProbability.lean @@ -26,8 +26,6 @@ assert_not_exists Ideal TwoSidedIdeal noncomputable section -open scoped Classical - open Fintype variable (M : Type*) [Mul M] @@ -79,6 +77,7 @@ variable {M} theorem commProb_eq_one_iff [h : Nonempty M] : commProb M = 1 ↔ Std.Commutative ((· * ·) : M → M → M) := by + classical haveI := Fintype.ofFinite M rw [commProb, ← Set.coe_setOf, Nat.card_eq_fintype_card, Nat.card_eq_fintype_card] rw [div_eq_one_iff_eq, ← Nat.cast_pow, Nat.cast_inj, sq, ← card_prod, diff --git a/Mathlib/GroupTheory/Coxeter/Length.lean b/Mathlib/GroupTheory/Coxeter/Length.lean index aff28c1d3b6853..c987abb55aa891 100644 --- a/Mathlib/GroupTheory/Coxeter/Length.lean +++ b/Mathlib/GroupTheory/Coxeter/Length.lean @@ -50,7 +50,7 @@ assert_not_exists TwoSidedIdeal namespace CoxeterSystem -open List Matrix Function Classical +open List Matrix Function variable {B : Type*} variable {W : Type*} [Group W] @@ -65,6 +65,7 @@ private theorem exists_word_with_prod (w : W) : ∃ n ω, ω.length = n ∧ π rcases cs.wordProd_surjective w with ⟨ω, rfl⟩ use ω.length, ω +open scoped Classical in /-- The length of `w`; i.e., the minimum number of simple reflections that must be multiplied to form `w`. -/ noncomputable def length (w : W) : ℕ := Nat.find (cs.exists_word_with_prod w) @@ -72,9 +73,11 @@ noncomputable def length (w : W) : ℕ := Nat.find (cs.exists_word_with_prod w) local prefix:100 "ℓ" => cs.length theorem exists_reduced_word (w : W) : ∃ ω, ω.length = ℓ w ∧ w = π ω := by + classical have := Nat.find_spec (cs.exists_word_with_prod w) tauto +open scoped Classical in theorem length_wordProd_le (ω : List B) : ℓ (π ω) ≤ ω.length := Nat.find_min' (cs.exists_word_with_prod (π ω)) ⟨ω, by tauto⟩ diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index 8adffc6d68bd07..0aa68664c8997c 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -47,8 +47,6 @@ universe u variable {G : Type u} -open scoped Classical - namespace Monoid section Monoid @@ -63,6 +61,7 @@ variable (G) [Monoid G] def ExponentExists := ∃ n, 0 < n ∧ ∀ g : G, g ^ n = 1 +open scoped Classical in /-- The exponent of a group is the smallest positive integer `n` such that `g ^ n = 1` for all `g ∈ G` if it exists, otherwise it is zero by convention. -/ @[to_additive @@ -132,6 +131,7 @@ theorem exponent_eq_zero_iff_forall : exponent G = 0 ↔ ∀ n > 0, ∃ g : G, g @[to_additive exponent_nsmul_eq_zero] theorem pow_exponent_eq_one (g : G) : g ^ exponent G = 1 := by + classical by_cases h : ExponentExists G · simp_rw [exponent, dif_pos h] exact (Nat.find_spec h).2 g @@ -150,6 +150,7 @@ theorem exponent_pos_of_exists (n : ℕ) (hpos : 0 < n) (hG : ∀ g : G, g ^ n = @[to_additive] theorem exponent_min' (n : ℕ) (hpos : 0 < n) (hG : ∀ g : G, g ^ n = 1) : exponent G ≤ n := by + classical rw [exponent, dif_pos] · apply Nat.find_min' exact ⟨hpos, hG⟩ @@ -467,6 +468,7 @@ theorem exponent_eq_iSup_orderOf (h : ∀ g : G, 0 < orderOf g) : simp_rw [Set.mem_range, exists_exists_eq_and] exact ⟨g, hx⟩ +open scoped Classical in @[to_additive] theorem exponent_eq_iSup_orderOf' : exponent G = if ∃ g : G, orderOf g = 0 then 0 else ⨆ g : G, orderOf g := by @@ -541,6 +543,7 @@ open Finset Monoid @[to_additive] theorem Monoid.exponent_pi_eq_zero {ι : Type*} {M : ι → Type*} [∀ i, Monoid (M i)] {j : ι} (hj : exponent (M j) = 0) : exponent ((i : ι) → M i) = 0 := by + classical rw [@exponent_eq_zero_iff, ExponentExists] at hj ⊢ push_neg at hj ⊢ peel hj with n hn _ diff --git a/Mathlib/GroupTheory/FreeAbelianGroup.lean b/Mathlib/GroupTheory/FreeAbelianGroup.lean index 030548d59fea55..60054af51dd65b 100644 --- a/Mathlib/GroupTheory/FreeAbelianGroup.lean +++ b/Mathlib/GroupTheory/FreeAbelianGroup.lean @@ -129,8 +129,7 @@ end lift section -open scoped Classical - +open scoped Classical in theorem of_injective : Function.Injective (of : α → FreeAbelianGroup α) := fun x y hoxy ↦ Classical.by_contradiction fun hxy : x ≠ y ↦ let f : FreeAbelianGroup α →+ ℤ := lift fun z ↦ if x = z then (1 : ℤ) else 0 diff --git a/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean b/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean index 477433e03185a5..95d573ec602a6f 100644 --- a/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean +++ b/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean @@ -161,8 +161,7 @@ theorem support_nsmul (k : ℕ) (h : k ≠ 0) (a : FreeAbelianGroup X) : apply support_zsmul k _ a exact mod_cast h -open scoped Classical - +open scoped Classical in theorem support_add (a b : FreeAbelianGroup X) : support (a + b) ⊆ a.support ∪ b.support := by simp only [support, AddMonoidHom.map_add] apply Finsupp.support_add diff --git a/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean b/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean index 2343d928c711c1..ce18362bd2a60f 100644 --- a/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean +++ b/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean @@ -50,8 +50,6 @@ free group, free groupoid, Nielsen-Schreier noncomputable section -open scoped Classical - universe v u /- Porting note: ./././Mathport/Syntax/Translate/Command.lean:229:11:unsupported: @@ -95,12 +93,8 @@ theorem ext_functor {G} [Groupoid.{v} G] [IsFreeGroupoid G] {X : Type v} [Group let ⟨_, _, u⟩ := @unique_lift G _ _ X _ fun (a b : Generators G) (e : a ⟶ b) => g.map (of e) _root_.trans (u _ h) (u _ fun _ _ _ => rfl).symm -#adaptation_note -/-- -The new unused variable linter in -https://github.com/leanprover/lean4/pull/5338 -flags `{ e // _ }`. --/ +#adaptation_note /-- https://github.com/leanprover/lean4/pull/5338 +The new unused variable linter flags `{ e // _ }`. -/ set_option linter.unusedVariables false in /-- An action groupoid over a free group is free. More generally, one could show that the groupoid of elements over a free groupoid is free, but this version is easier to prove and suffices for our @@ -210,6 +204,7 @@ def functorOfMonoidHom {X} [Monoid X] (f : End (root' T) →* X) : rw [comp_as_mul, ← f.map_mul] simp only [IsIso.inv_hom_id_assoc, loopOfHom, End.mul_def, Category.assoc] +open scoped Classical in /-- Given a free groupoid and an arborescence of its generating quiver, the vertex group at the root is freely generated by loops coming from generating arrows in the complement of the tree. -/ diff --git a/Mathlib/GroupTheory/GroupAction/Defs.lean b/Mathlib/GroupTheory/GroupAction/Defs.lean index d2cba75e644f68..fe04a1f0932dfb 100644 --- a/Mathlib/GroupTheory/GroupAction/Defs.lean +++ b/Mathlib/GroupTheory/GroupAction/Defs.lean @@ -5,8 +5,8 @@ Authors: Chris Hughes -/ import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Group.Subgroup.Defs -import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.Group.Submonoid.MulAction /-! # Definition of `orbit`, `fixedPoints` and `stabilizer` diff --git a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean index 69ac2dcd677060..c1d5b153c65e4b 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean @@ -6,6 +6,7 @@ Authors: Amelia Livingston import Mathlib.Algebra.Group.Submonoid.Defs import Mathlib.GroupTheory.Congruence.Hom import Mathlib.GroupTheory.OreLocalization.Basic +import Mathlib.Algebra.Group.Submonoid.Operations /-! # Localizations of commutative monoids diff --git a/Mathlib/GroupTheory/Nilpotent.lean b/Mathlib/GroupTheory/Nilpotent.lean index 2b1d6262d8668a..8c188f330ea885 100644 --- a/Mathlib/GroupTheory/Nilpotent.lean +++ b/Mathlib/GroupTheory/Nilpotent.lean @@ -355,23 +355,24 @@ theorem nilpotent_iff_lowerCentralSeries : IsNilpotent G ↔ ∃ n, lowerCentral section Classical -open scoped Classical - variable [hG : IsNilpotent G] variable (G) +open scoped Classical in /-- The nilpotency class of a nilpotent group is the smallest natural `n` such that the `n`'th term of the upper central series is `G`. -/ noncomputable def Group.nilpotencyClass : ℕ := Nat.find (IsNilpotent.nilpotent G) variable {G} +open scoped Classical in @[simp] theorem upperCentralSeries_nilpotencyClass : upperCentralSeries G (Group.nilpotencyClass G) = ⊤ := Nat.find_spec (IsNilpotent.nilpotent G) theorem upperCentralSeries_eq_top_iff_nilpotencyClass_le {n : ℕ} : upperCentralSeries G n = ⊤ ↔ Group.nilpotencyClass G ≤ n := by + classical constructor · intro h exact Nat.find_le h @@ -379,6 +380,7 @@ theorem upperCentralSeries_eq_top_iff_nilpotencyClass_le {n : ℕ} : rw [eq_top_iff, ← upperCentralSeries_nilpotencyClass] exact upperCentralSeries_mono _ h +open scoped Classical in /-- The nilpotency class of a nilpotent `G` is equal to the smallest `n` for which an ascending central series reaches `G` in its `n`'th term. -/ theorem least_ascending_central_series_length_eq_nilpotencyClass : @@ -391,6 +393,7 @@ theorem least_ascending_central_series_length_eq_nilpotencyClass : rw [← top_le_iff, ← hn] exact ascending_central_series_le_upper H hH n +open scoped Classical in /-- The nilpotency class of a nilpotent `G` is equal to the smallest `n` for which the descending central series reaches `⊥` in its `n`'th term. -/ theorem least_descending_central_series_length_eq_nilpotencyClass : @@ -409,6 +412,7 @@ theorem least_descending_central_series_length_eq_nilpotencyClass : rw [tsub_self] exact hH.1 +open scoped Classical in /-- The nilpotency class of a nilpotent `G` is equal to the length of the lower central series. -/ theorem lowerCentralSeries_length_eq_nilpotencyClass : Nat.find (nilpotent_iff_lowerCentralSeries.mp hG) = Group.nilpotencyClass (G := G) := by @@ -423,11 +427,13 @@ theorem lowerCentralSeries_length_eq_nilpotencyClass : @[simp] theorem lowerCentralSeries_nilpotencyClass : lowerCentralSeries G (Group.nilpotencyClass G) = ⊥ := by + classical rw [← lowerCentralSeries_length_eq_nilpotencyClass] exact Nat.find_spec (nilpotent_iff_lowerCentralSeries.mp hG) theorem lowerCentralSeries_eq_bot_iff_nilpotencyClass_le {n : ℕ} : lowerCentralSeries G n = ⊥ ↔ Group.nilpotencyClass G ≤ n := by + classical constructor · intro h rw [← lowerCentralSeries_length_eq_nilpotencyClass] diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index a53461a680fb88..4a207ce12d1ef9 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -10,6 +10,7 @@ import Mathlib.Algebra.Order.Group.Action import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Dynamics.PeriodicPts.Lemmas import Mathlib.GroupTheory.Index +import Mathlib.NumberTheory.Divisors import Mathlib.Order.Interval.Set.Infinite import Mathlib.Tactic.Positivity @@ -689,28 +690,12 @@ variable [Monoid G] {x : G} {n : ℕ} @[to_additive] theorem sum_card_orderOf_eq_card_pow_eq_one [Fintype G] [DecidableEq G] (hn : n ≠ 0) : - (∑ m ∈ (Finset.range n.succ).filter (· ∣ n), + (∑ m ∈ divisors n, (Finset.univ.filter fun x : G => orderOf x = m).card) = - (Finset.univ.filter fun x : G => x ^ n = 1).card := - calc - (∑ m ∈ (Finset.range n.succ).filter (· ∣ n), - (Finset.univ.filter fun x : G => orderOf x = m).card) = _ := - (Finset.card_biUnion - (by - intros - apply Finset.disjoint_filter.2 - rintro _ _ rfl; assumption)).symm - _ = _ := - congr_arg Finset.card - (Finset.ext - (by - intro x - suffices orderOf x ≤ n ∧ orderOf x ∣ n ↔ x ^ n = 1 by simpa [Nat.lt_succ_iff] - exact - ⟨fun h => by - let ⟨m, hm⟩ := h.2 - rw [hm, pow_mul, pow_orderOf_eq_one, one_pow], fun h => - ⟨orderOf_le_of_pow_eq_one hn.bot_lt h, orderOf_dvd_of_pow_eq_one h⟩⟩)) + (Finset.univ.filter fun x : G => x ^ n = 1).card := by + refine (Finset.card_biUnion ?_).symm.trans ?_ + · simp +contextual [disjoint_iff, Finset.ext_iff] + · congr; ext; simp [hn, orderOf_dvd_iff_pow_eq_one] @[to_additive] theorem orderOf_le_card_univ [Fintype G] : orderOf x ≤ Fintype.card G := diff --git a/Mathlib/GroupTheory/OreLocalization/Basic.lean b/Mathlib/GroupTheory/OreLocalization/Basic.lean index 24b43521f550bb..356e809eb79d9b 100644 --- a/Mathlib/GroupTheory/OreLocalization/Basic.lean +++ b/Mathlib/GroupTheory/OreLocalization/Basic.lean @@ -3,9 +3,10 @@ Copyright (c) 2022 Jakob von Raumer. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jakob von Raumer, Kevin Klinge, Andrew Yang -/ -import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.GroupTheory.OreLocalization.OreSet import Mathlib.Tactic.Common +import Mathlib.Algebra.Group.Submonoid.MulAction +import Mathlib.Algebra.Group.Units.Defs /-! diff --git a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean index 4f238742343a89..233d79ecb80e8a 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean @@ -39,8 +39,6 @@ The forward direction of `Equiv.Perm.isoCycle'` uses `Fintype.choose` of the uni result, relying on the `Fintype` instance of a `Cycle.nodup` subtype. It is unclear if this works faster than the `Equiv.Perm.toCycle`, which relies on recursion over `Finset.univ`. -Running `#eval` on even a simple noncyclic permutation `c[(1 : Fin 7), 2, 3] * c[0, 5]` -to show it takes a long time. TODO: is this because computing the cycle factors is slow? -/ @@ -469,9 +467,35 @@ set_option linter.unusedTactic false in notation3 (prettyPrint := false) "c["(l", "* => foldr (h t => List.cons h t) List.nil)"]" => Cycle.formPerm (Cycle.ofList l) (Iff.mpr Cycle.nodup_coe_iff (by decide)) -unsafe instance repr_perm [Repr α] : Repr (Perm α) := - ⟨fun f _ => repr (Multiset.pmap (fun (g : Perm α) (hg : g.IsCycle) => isoCycle ⟨g, hg⟩) - (Perm.cycleFactorsFinset f).val -- toCycle is faster? - fun _ hg => (mem_cycleFactorsFinset_iff.mp (Finset.mem_def.mpr hg)).left)⟩ +/-- Represents a permutation as product of disjoint cycles: +``` +#eval (c[0, 1, 2, 3] : Perm (Fin 4)) +-- c[0, 1, 2, 3] + +#eval (c[3, 1] * c[0, 2] : Perm (Fin 4)) +-- c[0, 2] * c[1, 3] + +#eval (c[1, 2, 3] * c[0, 1, 2] : Perm (Fin 4)) +-- c[0, 2] * c[1, 3] + +#eval (c[1, 2, 3] * c[0, 1, 2] * c[3, 1] * c[0, 2] : Perm (Fin 4)) +-- 1 +``` +-/ +unsafe instance instRepr [Repr α] : Repr (Perm α) where + reprPrec f prec := + -- Obtain a list of formats which represents disjoint cycles. + letI l := Quot.unquot <| Multiset.map repr <| Multiset.pmap toCycle + (Perm.cycleFactorsFinset f).val + fun _ hg => (mem_cycleFactorsFinset_iff.mp (Finset.mem_def.mpr hg)).left + -- And intercalate `*`s. + match l with + | [] => "1" + | [f] => f + | l => + -- multiple terms, use `*` precedence + (if prec ≥ 70 then Lean.Format.paren else id) + (Lean.Format.fill + (Lean.Format.joinSep l (" *" ++ Lean.Format.line))) end Equiv.Perm diff --git a/Mathlib/GroupTheory/Perm/Cycle/Type.lean b/Mathlib/GroupTheory/Perm/Cycle/Type.lean index e9a7dd8bc94db6..ebbd77c65b3711 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Type.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Type.lean @@ -406,8 +406,8 @@ theorem zero_eq : vectorsProdEqOne G 0 = {Vector.nil} := Set.eq_singleton_iff_unique_mem.mpr ⟨Eq.refl (1 : G), fun v _ => v.eq_nil⟩ theorem one_eq : vectorsProdEqOne G 1 = {Vector.nil.cons 1} := by - simp_rw [Set.eq_singleton_iff_unique_mem, mem_iff, Vector.toList_singleton, List.prod_singleton, - Vector.head_cons, true_and] + simp_rw [Set.eq_singleton_iff_unique_mem, mem_iff, List.Vector.toList_singleton, + List.prod_singleton, List.Vector.head_cons, true_and] exact fun v hv => v.cons_head_tail.symm.trans (congr_arg₂ Vector.cons hv v.tail.eq_nil) instance zeroUnique : Unique (vectorsProdEqOne G 0) := by diff --git a/Mathlib/GroupTheory/Perm/Fin.lean b/Mathlib/GroupTheory/Perm/Fin.lean index 6d0fa972e68ea1..ce47c8d3ed178c 100644 --- a/Mathlib/GroupTheory/Perm/Fin.lean +++ b/Mathlib/GroupTheory/Perm/Fin.lean @@ -142,13 +142,13 @@ def cycleRange {n : ℕ} (i : Fin n) : Perm (Fin n) := ext simp)) -theorem cycleRange_of_gt {n : ℕ} {i j : Fin n.succ} (h : i < j) : cycleRange i j = j := by +theorem cycleRange_of_gt {n : ℕ} {i j : Fin n} (h : i < j) : cycleRange i j = j := by rw [cycleRange, ofLeftInverse'_eq_ofInjective, ← Function.Embedding.toEquivRange_eq_ofInjective, ← viaFintypeEmbedding, viaFintypeEmbedding_apply_not_mem_range] simpa -theorem cycleRange_of_le {n : ℕ} {i j : Fin n.succ} (h : j ≤ i) : +theorem cycleRange_of_le {n : ℕ} [NeZero n] {i j : Fin n} (h : j ≤ i) : cycleRange i j = if j = i then 0 else j + 1 := by cases n · subsingleton @@ -165,8 +165,10 @@ theorem cycleRange_of_le {n : ℕ} {i j : Fin n.succ} (h : j ≤ i) : · rw [Fin.val_add_one_of_lt] exact lt_of_lt_of_le (lt_of_le_of_ne h (mt (congr_arg _) heq)) (le_last i) -theorem coe_cycleRange_of_le {n : ℕ} {i j : Fin n.succ} (h : j ≤ i) : +theorem coe_cycleRange_of_le {n : ℕ} {i j : Fin n} (h : j ≤ i) : (cycleRange i j : ℕ) = if j = i then 0 else (j : ℕ) + 1 := by + cases' n with n + · exact absurd le_rfl i.pos.not_le rw [cycleRange_of_le h] split_ifs with h' · rfl @@ -176,20 +178,20 @@ theorem coe_cycleRange_of_le {n : ℕ} {i j : Fin n.succ} (h : j ≤ i) : (j : ℕ) < i := Fin.lt_iff_val_lt_val.mp (lt_of_le_of_ne h h') _ ≤ n := Nat.lt_succ_iff.mp i.2) -theorem cycleRange_of_lt {n : ℕ} {i j : Fin n.succ} (h : j < i) : cycleRange i j = j + 1 := by +theorem cycleRange_of_lt {n : ℕ} [NeZero n] {i j : Fin n} (h : j < i) : cycleRange i j = j + 1 := by rw [cycleRange_of_le h.le, if_neg h.ne] -theorem coe_cycleRange_of_lt {n : ℕ} {i j : Fin n.succ} (h : j < i) : +theorem coe_cycleRange_of_lt {n : ℕ} {i j : Fin n} (h : j < i) : (cycleRange i j : ℕ) = j + 1 := by rw [coe_cycleRange_of_le h.le, if_neg h.ne] -theorem cycleRange_of_eq {n : ℕ} {i j : Fin n.succ} (h : j = i) : cycleRange i j = 0 := by +theorem cycleRange_of_eq {n : ℕ} [NeZero n] {i j : Fin n} (h : j = i) : cycleRange i j = 0 := by rw [cycleRange_of_le h.le, if_pos h] @[simp] -theorem cycleRange_self {n : ℕ} (i : Fin n.succ) : cycleRange i i = 0 := +theorem cycleRange_self {n : ℕ} [NeZero n] (i : Fin n) : cycleRange i i = 0 := cycleRange_of_eq rfl -theorem cycleRange_apply {n : ℕ} (i j : Fin n.succ) : +theorem cycleRange_apply {n : ℕ} [NeZero n] (i j : Fin n) : cycleRange i j = if j < i then j + 1 else if j = i then 0 else j := by split_ifs with h₁ h₂ · exact cycleRange_of_lt h₁ @@ -197,11 +199,11 @@ theorem cycleRange_apply {n : ℕ} (i j : Fin n.succ) : · exact cycleRange_of_gt (lt_of_le_of_ne (le_of_not_gt h₁) (Ne.symm h₂)) @[simp] -theorem cycleRange_zero (n : ℕ) : cycleRange (0 : Fin n.succ) = 1 := by +theorem cycleRange_zero (n : ℕ) [NeZero n] : cycleRange (0 : Fin n) = 1 := by ext j - refine Fin.cases ?_ (fun j => ?_) j + rcases (Fin.zero_le' j).eq_or_lt with rfl | hj · simp - · rw [cycleRange_of_gt (Fin.succ_pos j), one_apply] + · rw [cycleRange_of_gt hj, one_apply] @[simp] theorem cycleRange_last (n : ℕ) : cycleRange (last n) = finRotate (n + 1) := by @@ -209,10 +211,12 @@ theorem cycleRange_last (n : ℕ) : cycleRange (last n) = finRotate (n + 1) := b rw [coe_cycleRange_of_le (le_last _), coe_finRotate] @[simp] -theorem cycleRange_zero' {n : ℕ} (h : 0 < n) : cycleRange ⟨0, h⟩ = 1 := by - cases' n with n - · cases h - exact cycleRange_zero n +theorem cycleRange_mk_zero {n : ℕ} (h : 0 < n) : cycleRange ⟨0, h⟩ = 1 := + have : NeZero n := .of_pos h + cycleRange_zero n + +@[deprecated (since := "2025-01-28")] +alias cycleRange_zero' := cycleRange_mk_zero @[simp] theorem sign_cycleRange {n : ℕ} (i : Fin n) : Perm.sign (cycleRange i) = (-1) ^ (i : ℕ) := by @@ -248,7 +252,7 @@ theorem cycleRange_succAbove {n : ℕ} (i : Fin (n + 1)) (j : Fin n) : · rw [Fin.succAbove_of_le_castSucc _ _ h, Fin.cycleRange_of_gt (Fin.le_castSucc_iff.mp h)] @[simp] -theorem cycleRange_symm_zero {n : ℕ} (i : Fin (n + 1)) : i.cycleRange.symm 0 = i := +theorem cycleRange_symm_zero {n : ℕ} [NeZero n] (i : Fin n) : i.cycleRange.symm 0 = i := i.cycleRange.injective (by simp) @[simp] @@ -256,14 +260,15 @@ theorem cycleRange_symm_succ {n : ℕ} (i : Fin (n + 1)) (j : Fin n) : i.cycleRange.symm j.succ = i.succAbove j := i.cycleRange.injective (by simp) -theorem isCycle_cycleRange {n : ℕ} {i : Fin (n + 1)} (h0 : i ≠ 0) : IsCycle (cycleRange i) := by +theorem isCycle_cycleRange {n : ℕ} [NeZero n] {i : Fin n} (h0 : i ≠ 0) : + IsCycle (cycleRange i) := by cases' i with i hi cases i · exact (h0 rfl).elim exact isCycle_finRotate.extendDomain _ @[simp] -theorem cycleType_cycleRange {n : ℕ} {i : Fin (n + 1)} (h0 : i ≠ 0) : +theorem cycleType_cycleRange {n : ℕ} [NeZero n] {i : Fin n} (h0 : i ≠ 0) : cycleType (cycleRange i) = {(i + 1 : ℕ)} := by cases' i with i hi cases i diff --git a/Mathlib/GroupTheory/Perm/ViaEmbedding.lean b/Mathlib/GroupTheory/Perm/ViaEmbedding.lean index 4d0a88429acabe..7193c769d6b76b 100644 --- a/Mathlib/GroupTheory/Perm/ViaEmbedding.lean +++ b/Mathlib/GroupTheory/Perm/ViaEmbedding.lean @@ -19,18 +19,20 @@ namespace Perm variable (e : Perm α) (ι : α ↪ β) -open scoped Classical - +open scoped Classical in /-- Noncomputable version of `Equiv.Perm.viaFintypeEmbedding` that does not assume `Fintype` -/ noncomputable def viaEmbedding : Perm β := extendDomain e (ofInjective ι.1 ι.2) +open scoped Classical in theorem viaEmbedding_apply (x : α) : e.viaEmbedding ι (ι x) = ι (e x) := extendDomain_apply_image e (ofInjective ι.1 ι.2) x +open scoped Classical in theorem viaEmbedding_apply_of_not_mem (x : β) (hx : x ∉ Set.range ι) : e.viaEmbedding ι x = x := extendDomain_apply_not_subtype e (ofInjective ι.1 ι.2) hx +open scoped Classical in /-- `viaEmbedding` as a group homomorphism -/ noncomputable def viaEmbeddingHom : Perm α →* Perm β := extendDomainHom (ofInjective ι.1 ι.2) @@ -38,6 +40,7 @@ noncomputable def viaEmbeddingHom : Perm α →* Perm β := theorem viaEmbeddingHom_apply : viaEmbeddingHom ι e = viaEmbedding e ι := rfl +open scoped Classical in theorem viaEmbeddingHom_injective : Function.Injective (viaEmbeddingHom ι) := extendDomainHom_injective (ofInjective ι.1 ι.2) diff --git a/Mathlib/GroupTheory/QuotientGroup/Finite.lean b/Mathlib/GroupTheory/QuotientGroup/Finite.lean index d21f2e376e3738..c27b649ddef251 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Finite.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Finite.lean @@ -20,13 +20,12 @@ universe u v w x namespace Group -open scoped Classical - open QuotientGroup Subgroup variable {F G H : Type u} [Group F] [Group G] [Group H] [Fintype F] [Fintype H] variable (f : F →* G) (g : G →* H) +open scoped Classical in /-- If `F` and `H` are finite such that `ker(G →* H) ≤ im(F →* G)`, then `G` is finite. -/ @[to_additive "If `F` and `H` are finite such that `ker(G →+ H) ≤ im(F →+ G)`, then `G` is finite."] noncomputable def fintypeOfKerLeRange (h : g.ker ≤ f.range) : Fintype G := diff --git a/Mathlib/GroupTheory/Schreier.lean b/Mathlib/GroupTheory/Schreier.lean index d7bbfc3d608794..b0d77367ab543b 100644 --- a/Mathlib/GroupTheory/Schreier.lean +++ b/Mathlib/GroupTheory/Schreier.lean @@ -29,12 +29,11 @@ section CommGroup open Subgroup -open scoped Classical - variable (G : Type*) [CommGroup G] [Group.FG G] @[to_additive] theorem card_dvd_exponent_pow_rank : Nat.card G ∣ Monoid.exponent G ^ Group.rank G := by + classical obtain ⟨S, hS1, hS2⟩ := Group.rank_spec G rw [← hS1, ← Fintype.card_coe, ← Finset.card_univ, ← Finset.prod_const] let f : (∀ g : S, zpowers (g : G)) →* G := noncommPiCoprod fun s t _ x y _ _ => mul_comm x _ diff --git a/Mathlib/GroupTheory/SchurZassenhaus.lean b/Mathlib/GroupTheory/SchurZassenhaus.lean index b36b24b2c9fa80..0611210466d315 100644 --- a/Mathlib/GroupTheory/SchurZassenhaus.lean +++ b/Mathlib/GroupTheory/SchurZassenhaus.lean @@ -117,8 +117,6 @@ private theorem exists_right_complement'_of_coprime_aux (hH : Nat.Coprime (Nat.c end SchurZassenhausAbelian -open scoped Classical - universe u namespace SchurZassenhausInduction diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 44e9501b1c31a5..d000af6ee0f64a 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -456,8 +456,8 @@ private theorem card_orderOf_eq_totient_aux₁ {d : ℕ} (hd : d ∣ Fintype.car have h2 : (∑ m ∈ d.divisors, #{a : α | orderOf a = m}) = ∑ m ∈ d.divisors, φ m := by - rw [← filter_dvd_eq_divisors hd0, sum_card_orderOf_eq_card_pow_eq_one hd0, - filter_dvd_eq_divisors hd0, sum_totient, ← ha, card_pow_eq_one_eq_orderOf_aux hn a] + rw [sum_card_orderOf_eq_card_pow_eq_one hd0, sum_totient, + ← ha, card_pow_eq_one_eq_orderOf_aux hn a] simpa [← cons_self_properDivisors hd0, ← h1] using h2 @[to_additive] @@ -472,7 +472,7 @@ theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : apply lt_irrefl c calc c = ∑ m ∈ c.divisors, #{a : α | orderOf a = m} := by - simp only [← filter_dvd_eq_divisors hc0.ne', sum_card_orderOf_eq_card_pow_eq_one hc0.ne'] + simp only [sum_card_orderOf_eq_card_pow_eq_one hc0.ne'] apply congr_arg card simp [c] _ = ∑ m ∈ c.divisors.erase d, #{a : α | orderOf a = m} := by diff --git a/Mathlib/GroupTheory/Subgroup/Center.lean b/Mathlib/GroupTheory/Subgroup/Center.lean index 6dd4dbfa414472..b8dc94177cf81a 100644 --- a/Mathlib/GroupTheory/Subgroup/Center.lean +++ b/Mathlib/GroupTheory/Subgroup/Center.lean @@ -40,6 +40,16 @@ theorem center_toSubmonoid : (center G).toSubmonoid = Submonoid.center G := instance center.isCommutative : (center G).IsCommutative := ⟨⟨fun a b => Subtype.ext (b.2.comm a).symm⟩⟩ +variable {G} in +/-- The center of isomorphic groups are isomorphic. -/ +@[to_additive (attr := simps!) "The center of isomorphic additive groups are isomorphic."] +def centerCongr {H} [Group H] (e : G ≃* H) : center G ≃* center H := Submonoid.centerCongr e + +/-- The center of a group is isomorphic to the center of its opposite. -/ +@[to_additive (attr := simps!) +"The center of an additive group is isomorphic to the center of its opposite."] +def centerToMulOpposite : center G ≃* center Gᵐᵒᵖ := Submonoid.centerToMulOpposite + /-- For a group with zero, the center of the units is the same as the units of the center. -/ @[simps! apply_val_coe symm_apply_coe_val] def centerUnitsEquivUnitsCenter (G₀ : Type*) [GroupWithZero G₀] : diff --git a/Mathlib/GroupTheory/Submonoid/Center.lean b/Mathlib/GroupTheory/Submonoid/Center.lean index 44028df37b8689..addceff4248907 100644 --- a/Mathlib/GroupTheory/Submonoid/Center.lean +++ b/Mathlib/GroupTheory/Submonoid/Center.lean @@ -126,3 +126,61 @@ def unitsCenterToCenterUnits [Monoid M] : (Submonoid.center M)ˣ →* Submonoid. theorem unitsCenterToCenterUnits_injective [Monoid M] : Function.Injective (unitsCenterToCenterUnits M) := fun _a _b h => Units.ext <| Subtype.ext <| congr_arg (Units.val ∘ Subtype.val) h + +section congr + +variable {M} {N : Type*} + +@[to_additive] theorem _root_.MulEquivClass.apply_mem_center {F} [EquivLike F M N] [Mul M] [Mul N] + [MulEquivClass F M N] (e : F) {x : M} (hx : x ∈ Set.center M) : e x ∈ Set.center N := by + let e := MulEquivClass.toMulEquiv e + show e x ∈ Set.center N + constructor <;> + (intros; apply e.symm.injective; + simp only [map_mul, e.symm_apply_apply, (isMulCentral_iff _).mp hx]) + +@[to_additive] theorem _root_.MulEquivClass.apply_mem_center_iff {F} [EquivLike F M N] + [Mul M] [Mul N] [MulEquivClass F M N] (e : F) {x : M} : + e x ∈ Set.center N ↔ x ∈ Set.center M := + ⟨(by simpa using MulEquivClass.apply_mem_center (MulEquivClass.toMulEquiv e).symm ·), + MulEquivClass.apply_mem_center e⟩ + +/-- The center of isomorphic magmas are isomorphic. -/ +@[to_additive (attr := simps) "The center of isomorphic additive magmas are isomorphic."] +def Subsemigroup.centerCongr [Mul M] [Mul N] (e : M ≃* N) : center M ≃* center N where + toFun r := ⟨e r, MulEquivClass.apply_mem_center e r.2⟩ + invFun s := ⟨e.symm s, MulEquivClass.apply_mem_center e.symm s.2⟩ + left_inv _ := Subtype.ext (e.left_inv _) + right_inv _ := Subtype.ext (e.right_inv _) + map_mul' _ _ := Subtype.ext (map_mul ..) + +/-- The center of isomorphic monoids are isomorphic. -/ +@[to_additive (attr := simps!) "The center of isomorphic additive monoids are isomorphic."] +def Submonoid.centerCongr [MulOneClass M] [MulOneClass N] (e : M ≃* N) : center M ≃* center N := + Subsemigroup.centerCongr e + +@[to_additive] theorem MulOpposite.op_mem_center_iff [Mul M] {x : M} : + op x ∈ Set.center Mᵐᵒᵖ ↔ x ∈ Set.center M := by + simp_rw [Set.mem_center_iff, isMulCentral_iff, MulOpposite.forall, ← op_mul, op_inj]; aesop + +@[to_additive] theorem MulOpposite.unop_mem_center_iff [Mul M] {x : Mᵐᵒᵖ} : + unop x ∈ Set.center M ↔ x ∈ Set.center Mᵐᵒᵖ := + op_mem_center_iff.symm + +/-- The center of a magma is isomorphic to the center of its opposite. -/ +@[to_additive (attr := simps) +"The center of an additive magma is isomorphic to the center of its opposite."] +def Subsemigroup.centerToMulOpposite [Mul M] : center M ≃* center Mᵐᵒᵖ where + toFun r := ⟨_, MulOpposite.op_mem_center_iff.mpr r.2⟩ + invFun r := ⟨_, MulOpposite.unop_mem_center_iff.mpr r.2⟩ + left_inv _ := rfl + right_inv _ := rfl + map_mul' r _ := Subtype.ext (congr_arg MulOpposite.op <| r.2.1 _) + +/-- The center of a monoid is isomorphic to the center of its opposite. -/ +@[to_additive (attr := simps!) +"The center of an additive monoid is isomorphic to the center of its opposite. "] +def Submonoid.centerToMulOpposite [MulOneClass M] : center M ≃* center Mᵐᵒᵖ := + Subsemigroup.centerToMulOpposite + +end congr diff --git a/Mathlib/InformationTheory/Hamming.lean b/Mathlib/InformationTheory/Hamming.lean index 8ad9b91993a9b9..47a11f3d846157 100644 --- a/Mathlib/InformationTheory/Hamming.lean +++ b/Mathlib/InformationTheory/Hamming.lean @@ -244,6 +244,8 @@ instance [Zero α] [∀ i, Zero (β i)] [∀ i, SMulWithZero α (β i)] : SMulWi instance [∀ i, AddMonoid (β i)] : AddMonoid (Hamming β) := Pi.addMonoid +instance [∀ i, AddGroup (β i)] : AddGroup (Hamming β) := Pi.addGroup + instance [∀ i, AddCommMonoid (β i)] : AddCommMonoid (Hamming β) := Pi.addCommMonoid @@ -396,15 +398,15 @@ instance [∀ i, Zero (β i)] : Norm (Hamming β) := theorem norm_eq_hammingNorm [∀ i, Zero (β i)] (x : Hamming β) : ‖x‖ = hammingNorm (ofHamming x) := rfl --- Porting note: merged `SeminormedAddCommGroup` and `NormedAddCommGroup` instances +instance [∀ i, AddGroup (β i)] : NormedAddGroup (Hamming β) where + dist_eq := by push_cast; exact mod_cast hammingDist_eq_hammingNorm instance [∀ i, AddCommGroup (β i)] : NormedAddCommGroup (Hamming β) where dist_eq := by push_cast; exact mod_cast hammingDist_eq_hammingNorm @[simp, push_cast] -theorem nnnorm_eq_hammingNorm [∀ i, AddCommGroup (β i)] (x : Hamming β) : - ‖x‖₊ = hammingNorm (ofHamming x) := - rfl +theorem nnnorm_eq_hammingNorm [∀ i, AddGroup (β i)] (x : Hamming β) : + ‖x‖₊ = hammingNorm (ofHamming x) := rfl end diff --git a/Mathlib/Lean/Meta/CongrTheorems.lean b/Mathlib/Lean/Meta/CongrTheorems.lean index a8dafe249c20f2..fa7daf4c35b9d0 100644 --- a/Mathlib/Lean/Meta/CongrTheorems.lean +++ b/Mathlib/Lean/Meta/CongrTheorems.lean @@ -105,6 +105,11 @@ class FastIsEmpty (α : Sort u) : Prop where protected theorem FastSubsingleton.elim {α : Sort u} [h : FastSubsingleton α] : (a b : α) → a = b := h.inst.allEq +protected theorem FastSubsingleton.helim {α β : Sort u} [FastSubsingleton α] + (h₂ : α = β) (a : α) (b : β) : HEq a b := by + have : Subsingleton α := FastSubsingleton.inst + exact Subsingleton.helim h₂ a b + instance (priority := 100) {α : Type u} [inst : FastIsEmpty α] : FastSubsingleton α where inst := have := inst.inst; inferInstance diff --git a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean index 1c1eeeec108bc9..8ae2bfd5616d2b 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean @@ -275,8 +275,8 @@ noncomputable def dualBasis (B : BilinForm K V) (hB : B.Nondegenerate) (b : Basi theorem dualBasis_repr_apply (B : BilinForm K V) (hB : B.Nondegenerate) (b : Basis ι K V) (x i) : (B.dualBasis hB b).repr x i = B x (b i) := by - #adaptation_note - /-- Before https://github.com/leanprover/lean4/pull/4814, we did not need the `@` in front of `toDual_def` in the `rw`. + #adaptation_note /-- https://github.com/leanprover/lean4/pull/4814 + we did not need the `@` in front of `toDual_def` in the `rw`. I'm confused! -/ rw [dualBasis, Basis.map_repr, LinearEquiv.symm_symm, LinearEquiv.trans_apply, Basis.dualBasis_repr, @toDual_def] diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index a6c69a3ba7efa0..295a03c7af209f 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -1712,9 +1712,7 @@ theorem finiteDimensional_quot_dualCoannihilator_iff {W : Submodule K (Dual K V) FiniteDimensional K (V ⧸ W.dualCoannihilator) ↔ FiniteDimensional K W := ⟨fun _ ↦ FiniteDimensional.of_injective _ W.flip_quotDualCoannihilatorToDual_injective, fun _ ↦ by - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/4119 + #adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 the `Free K W` instance isn't found unless we use `set_option maxSynthPendingDepth 2`, or add explicit instances: ``` @@ -1896,4 +1894,4 @@ noncomputable def dualDistribEquiv : Dual R M ⊗[R] Dual R N ≃ₗ[R] Dual R ( end TensorProduct -set_option linter.style.longFile 2100 +set_option linter.style.longFile 2000 diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index fcb07a9d441a26..aa94d3765c6685 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -219,7 +219,7 @@ theorem inf_iSup_genEigenspace [FiniteDimensional K V] (h : ∀ x ∈ p, f x ∈ rw [Algebra.algebraMap_eq_smul_one, ← End.genEigenspace_nat] apply aux μ' hμ' · have := this.supIndep' (m.support.erase μ) - apply Finset.supIndep_antimono_fun _ this + apply this.antitone_fun intro μ' hμ' rw [Algebra.algebraMap_eq_smul_one, ← End.genEigenspace_nat] apply aux μ' hμ' diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean index 8befb3f75b370b..8ae8358bfca550 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean @@ -39,7 +39,7 @@ theorem PiToModule.fromMatrix_apply [DecidableEq ι] (A : Matrix ι ι R) (w : theorem PiToModule.fromMatrix_apply_single_one [DecidableEq ι] (A : Matrix ι ι R) (j : ι) : PiToModule.fromMatrix R b A (Pi.single j 1) = ∑ i : ι, A i j • b i := by rw [PiToModule.fromMatrix_apply, Fintype.linearCombination_apply, Matrix.mulVec_single] - simp_rw [mul_one] + simp_rw [MulOpposite.op_one, one_smul, transpose_apply] /-- The endomorphisms of `M` acts on `(ι → R) →ₗ[R] M`, and takes the projection to a `(ι → R) →ₗ[R] M`. -/ diff --git a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean index 6a59c9e32abf27..e61efb9e17587f 100644 --- a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean +++ b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean @@ -130,6 +130,7 @@ end CoeLemmas variable {S T : Type*} [CommRing S] [CommRing T] /-- A ring homomorphism ``f : R →+* S`` induces a homomorphism ``GLₙ(f) : GLₙ(R) →* GLₙ(S)``. -/ +@[simps! apply_val] def map (f : R →+* S) : GL n R →* GL n S := Units.map <| (RingHom.mapMatrix f).toMonoidHom @[simp] diff --git a/Mathlib/LinearAlgebra/Matrix/Hermitian.lean b/Mathlib/LinearAlgebra/Matrix/Hermitian.lean index 89a34b681a4a6a..c7e39b7cc0e1d0 100644 --- a/Mathlib/LinearAlgebra/Matrix/Hermitian.lean +++ b/Mathlib/LinearAlgebra/Matrix/Hermitian.lean @@ -279,7 +279,7 @@ theorem isHermitian_iff_isSymmetric [Fintype n] [DecidableEq n] {A : Matrix n n · intro h ext i j simpa only [(Pi.single_star i 1).symm, ← star_mulVec, mul_one, dotProduct_single, - single_vecMul, star_one, one_mul] using h (Pi.single i 1) (Pi.single j 1) + single_one_vecMul, star_one] using h (Pi.single i 1) (Pi.single j 1) end RCLike diff --git a/Mathlib/LinearAlgebra/Matrix/Permutation.lean b/Mathlib/LinearAlgebra/Matrix/Permutation.lean index 96c0d8f5fc7712..990b92bc1ee055 100644 --- a/Mathlib/LinearAlgebra/Matrix/Permutation.lean +++ b/Mathlib/LinearAlgebra/Matrix/Permutation.lean @@ -39,7 +39,7 @@ namespace Matrix /-- The determinant of a permutation matrix equals its sign. -/ @[simp] theorem det_permutation [CommRing R] : det (σ.permMatrix R) = Perm.sign σ := by - rw [← Matrix.mul_one (σ.permMatrix R), PEquiv.toPEquiv_mul_matrix, + rw [← Matrix.mul_one (σ.permMatrix R), PEquiv.toMatrix_toPEquiv_mul, det_permute, det_one, mul_one] /-- The trace of a permutation matrix equals the number of fixed points. -/ diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index 5e46c1956c392a..7776201ccffd13 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -363,8 +363,8 @@ theorem _root_.Matrix.posDef_diagonal_iff PosDef (diagonal d) ↔ ∀ i, 0 < d i := by refine ⟨fun h i => ?_, .diagonal⟩ have := h.2 (Pi.single i 1) - simp only [mulVec_single, mul_one, dotProduct_diagonal', Pi.star_apply, Pi.single_eq_same, - star_one, one_mul, Function.ne_iff, Pi.zero_apply] at this + simp_rw [mulVec_single_one, ← Pi.single_star, star_one, single_dotProduct, one_mul, + transpose_apply, diagonal_apply_eq, Function.ne_iff] at this exact this ⟨i, by simp⟩ protected theorem one [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] : diff --git a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean index 3367f16ef2d733..ef6287892c69eb 100644 --- a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean +++ b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean @@ -74,6 +74,11 @@ lemma eigenvectorUnitary_coe {𝕜 : Type*} [RCLike 𝕜] {n : Type*} [Fintype n (EuclideanSpace.basisFun n 𝕜).toBasis.toMatrix (hA.eigenvectorBasis).toBasis := rfl +@[simp] +theorem eigenvectorUnitary_transpose_apply (j : n) : + (eigenvectorUnitary hA)ᵀ j = ⇑(hA.eigenvectorBasis j) := + rfl + @[simp] theorem eigenvectorUnitary_apply (i j : n) : eigenvectorUnitary hA i j = ⇑(hA.eigenvectorBasis j) i := @@ -81,7 +86,7 @@ theorem eigenvectorUnitary_apply (i j : n) : theorem eigenvectorUnitary_mulVec (j : n) : eigenvectorUnitary hA *ᵥ Pi.single j 1 = ⇑(hA.eigenvectorBasis j) := by - simp only [mulVec_single, eigenvectorUnitary_apply, mul_one] + simp_rw [mulVec_single_one, eigenvectorUnitary_transpose_apply] theorem star_eigenvectorUnitary_mulVec (j : n) : (star (eigenvectorUnitary hA : Matrix n n 𝕜)) *ᵥ ⇑(hA.eigenvectorBasis j) = Pi.single j 1 := by diff --git a/Mathlib/LinearAlgebra/Matrix/ToLin.lean b/Mathlib/LinearAlgebra/Matrix/ToLin.lean index 1bce52c084cc20..506f38bc302e36 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLin.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLin.lean @@ -249,7 +249,7 @@ variable [Fintype n] @[simp] theorem Matrix.mulVecLin_one [DecidableEq n] : Matrix.mulVecLin (1 : Matrix n n R) = LinearMap.id := by - ext; simp [Matrix.one_apply, Pi.single_apply] + ext; simp [Matrix.one_apply, Pi.single_apply, eq_comm] @[simp] theorem Matrix.mulVecLin_mul [Fintype m] (M : Matrix l m R) (N : Matrix m n R) : diff --git a/Mathlib/LinearAlgebra/Orientation.lean b/Mathlib/LinearAlgebra/Orientation.lean index 549ff52629c7bb..004d248871f558 100644 --- a/Mathlib/LinearAlgebra/Orientation.lean +++ b/Mathlib/LinearAlgebra/Orientation.lean @@ -337,8 +337,8 @@ theorem eq_or_eq_neg [FiniteDimensional R M] (x₁ x₂ : Orientation R M ι) -- Porting note: this needs to be made explicit for the simp below have orientation_neg_neg : ∀ f : Basis ι R M, - -Basis.orientation f = Basis.orientation f := by - #adaptation_note - /-- `set_option maxSynthPendingDepth 2` required after https://github.com/leanprover/lean4/pull/4119 -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 + `set_option maxSynthPendingDepth 2` required -/ set_option maxSynthPendingDepth 2 in simp rcases e.orientation_eq_or_eq_neg x₁ with (h₁ | h₁) <;> rcases e.orientation_eq_or_eq_neg x₂ with (h₂ | h₂) <;> simp [h₁, h₂, orientation_neg_neg] diff --git a/Mathlib/LinearAlgebra/Pi.lean b/Mathlib/LinearAlgebra/Pi.lean index 38d4beb411b38d..c638022cfd82a2 100644 --- a/Mathlib/LinearAlgebra/Pi.lean +++ b/Mathlib/LinearAlgebra/Pi.lean @@ -11,6 +11,7 @@ import Mathlib.Algebra.Module.Submodule.Ker import Mathlib.Algebra.Module.Submodule.Range import Mathlib.Algebra.Module.Equiv.Basic import Mathlib.Logic.Equiv.Fin +import Mathlib.LinearAlgebra.Prod /-! # Pi types of modules @@ -571,6 +572,17 @@ noncomputable def Function.ExtendByZero.linearMap : (ι → R) →ₗ[R] η → end Extend +variable (R) in +/-- `Fin.consEquiv` as a continuous linear equivalence. -/ +@[simps] +def Fin.consLinearEquiv + {n : ℕ} (M : Fin n.succ → Type*) [Semiring R] [∀ i, AddCommMonoid (M i)] [∀ i, Module R (M i)] : + (M 0 × Π i, M (Fin.succ i)) ≃ₗ[R] (Π i, M i) where + __ := Fin.consEquiv M + map_add' x y := funext <| Fin.cases rfl (by simp) + map_smul' c x := funext <| Fin.cases rfl (by simp) + + /-! ### Bundled versions of `Matrix.vecCons` and `Matrix.vecEmpty` The idea of these definitions is to be able to define a map as `x ↦ ![f₁ x, f₂ x, f₃ x]`, where @@ -604,14 +616,8 @@ theorem LinearMap.vecEmpty_apply (m : M) : (LinearMap.vecEmpty : M →ₗ[R] Fin /-- A linear map into `Fin n.succ → M₃` can be built out of a map into `M₃` and a map into `Fin n → M₃`. -/ -def LinearMap.vecCons {n} (f : M →ₗ[R] M₂) (g : M →ₗ[R] Fin n → M₂) : M →ₗ[R] Fin n.succ → M₂ where - toFun m := Matrix.vecCons (f m) (g m) - map_add' x y := by - simp only [] - rw [f.map_add, g.map_add, Matrix.cons_add_cons (f x)] - map_smul' c x := by - simp only [] - rw [f.map_smul, g.map_smul, RingHom.id_apply, Matrix.smul_cons c (f x)] +def LinearMap.vecCons {n} (f : M →ₗ[R] M₂) (g : M →ₗ[R] Fin n → M₂) : M →ₗ[R] Fin n.succ → M₂ := + Fin.consLinearEquiv R (fun _ : Fin n.succ => M₂) ∘ₗ f.prod g @[simp] theorem LinearMap.vecCons_apply {n} (f : M →ₗ[R] M₂) (g : M →ₗ[R] Fin n → M₂) (m : M) : diff --git a/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean b/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean index 7a26465a9f79c0..6854ebb704c2fe 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean @@ -27,7 +27,7 @@ section CommRing variable [CommRing R] [CommRing A] variable [AddCommGroup M₁] [AddCommGroup M₂] [AddCommGroup N₁] [AddCommGroup N₂] variable [Algebra R A] [Module R M₁] [Module A M₁] [Module R N₁] [Module A N₁] -variable [SMulCommClass R A M₁] [IsScalarTower R A M₁] [SMulCommClass R A N₁] [IsScalarTower R A N₁] +variable [SMulCommClass R A M₁] [IsScalarTower R A M₁] [IsScalarTower R A N₁] variable [Module R M₂] [Module R N₂] section InvertibleTwo @@ -144,13 +144,13 @@ end QuadraticForm end InvertibleTwo -/-- If two quadratic forms from `A ⊗[R] M₂` agree on elements of the form `1 ⊗ m`, they are equal. +/-- If two quadratic maps from `A ⊗[R] M₂` agree on elements of the form `1 ⊗ m`, they are equal. -In other words, if a base change exists for a quadratic form, it is unique. +In other words, if a base change exists for a quadratic map, it is unique. Note that unlike `QuadraticForm.baseChange`, this does not need `Invertible (2 : R)`. -/ @[ext] -theorem baseChange_ext ⦃Q₁ Q₂ : QuadraticForm A (A ⊗[R] M₂)⦄ +theorem baseChange_ext ⦃Q₁ Q₂ : QuadraticMap A (A ⊗[R] M₂) N₁⦄ (h : ∀ m, Q₁ (1 ⊗ₜ m) = Q₂ (1 ⊗ₜ m)) : Q₁ = Q₂ := by replace h (a m) : Q₁ (a ⊗ₜ m) = Q₂ (a ⊗ₜ m) := by @@ -166,6 +166,6 @@ theorem baseChange_ext ⦃Q₁ Q₂ : QuadraticForm A (A ⊗[R] M₂)⦄ rw [← TensorProduct.tmul_add, h, h, h] replace := congr($this x y) dsimp [polar] at this - linear_combination this + hx + hy + linear_combination (norm := module) this + hx + hy end CommRing diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index feb4bde80dc93f..4135e45fe83c2e 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -274,13 +274,10 @@ theorem IsSemisimple.of_mem_adjoin_pair {a : End K M} (ha : a ∈ Algebra.adjoin (AdjoinRoot.powerBasis' <| minpoly.monic <| Algebra.IsIntegral.isIntegral f).finite have : Module.Finite R S := (AdjoinRoot.powerBasis' <| (minpoly.monic <| Algebra.IsIntegral.isIntegral g).map _).finite - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/4119 we either need - to specify the `(S := R)` argument, or use `set_option maxSynthPendingDepth 2 in`. + #adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 + we either need to specify the `(S := R)` argument, or use `set_option maxSynthPendingDepth 2 in`. - In either case this step is too slow! - -/ + In either case this step is too slow! -/ set_option maxSynthPendingDepth 2 in have : IsScalarTower K R S := .of_algebraMap_eq fun _ ↦ rfl have : Module.Finite K S := .trans R S diff --git a/Mathlib/Logic/Function/Defs.lean b/Mathlib/Logic/Function/Defs.lean index 0caf5f335d524b..8196c2377e17e6 100644 --- a/Mathlib/Logic/Function/Defs.lean +++ b/Mathlib/Logic/Function/Defs.lean @@ -48,7 +48,6 @@ scoped infixl:2 " on " => onFun abbrev swap {φ : α → β → Sort u₃} (f : ∀ x y, φ x y) : ∀ y x, φ x y := fun y x => f x y -#adaptation_note /-- nightly-2024-03-16: added to replace simp [Function.swap] -/ theorem swap_def {φ : α → β → Sort u₃} (f : ∀ x y, φ x y) : swap f = fun y x => f x y := rfl -- Porting note: removed, it was never used diff --git a/Mathlib/Logic/IsEmpty.lean b/Mathlib/Logic/IsEmpty.lean index fa48bd9fc85c3f..d6c0c753066073 100644 --- a/Mathlib/Logic/IsEmpty.lean +++ b/Mathlib/Logic/IsEmpty.lean @@ -213,10 +213,22 @@ variable {α β : Type*} (R : α → β → Prop) theorem leftTotal_empty [IsEmpty α] : LeftTotal R := by simp only [LeftTotal, IsEmpty.forall_iff] +theorem leftTotal_iff_isEmpty_left [IsEmpty β] : LeftTotal R ↔ IsEmpty α := by + simp only [LeftTotal, IsEmpty.exists_iff, isEmpty_iff] + @[simp] theorem rightTotal_empty [IsEmpty β] : RightTotal R := by simp only [RightTotal, IsEmpty.forall_iff] +theorem rightTotal_iff_isEmpty_right [IsEmpty α] : RightTotal R ↔ IsEmpty β := by + simp only [RightTotal, IsEmpty.exists_iff, isEmpty_iff, imp_self] + @[simp] theorem biTotal_empty [IsEmpty α] [IsEmpty β] : BiTotal R := ⟨leftTotal_empty R, rightTotal_empty R⟩ + +theorem biTotal_iff_isEmpty_right [IsEmpty α] : BiTotal R ↔ IsEmpty β := by + simp only [BiTotal, leftTotal_empty, rightTotal_iff_isEmpty_right, true_and] + +theorem biTotal_iff_isEmpty_left [IsEmpty β] : BiTotal R ↔ IsEmpty α := by + simp only [BiTotal, leftTotal_iff_isEmpty_left, rightTotal_empty, and_true] diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Metric.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Metric.lean index 07c0bcd00e44e1..c3611c1f706ca8 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Metric.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Metric.lean @@ -232,25 +232,28 @@ theorem measurable_nnnorm : Measurable (nnnorm : α → ℝ≥0) := continuous_nnnorm.measurable @[measurability, fun_prop] -theorem Measurable.nnnorm {f : β → α} (hf : Measurable f) : Measurable fun a => ‖f a‖₊ := +protected theorem Measurable.nnnorm {f : β → α} (hf : Measurable f) : Measurable fun a => ‖f a‖₊ := measurable_nnnorm.comp hf @[measurability, fun_prop] -theorem AEMeasurable.nnnorm {f : β → α} {μ : Measure β} (hf : AEMeasurable f μ) : +protected lemma AEMeasurable.nnnorm {f : β → α} {μ : Measure β} (hf : AEMeasurable f μ) : AEMeasurable (fun a => ‖f a‖₊) μ := measurable_nnnorm.comp_aemeasurable hf @[measurability] -theorem measurable_ennnorm : Measurable fun x : α => (‖x‖₊ : ℝ≥0∞) := - measurable_nnnorm.coe_nnreal_ennreal +lemma measurable_enorm : Measurable (enorm : α → ℝ≥0∞) := continuous_enorm.measurable @[measurability, fun_prop] -theorem Measurable.ennnorm {f : β → α} (hf : Measurable f) : Measurable fun a => (‖f a‖₊ : ℝ≥0∞) := +protected lemma Measurable.enorm {f : β → α} (hf : Measurable f) : Measurable (‖f ·‖ₑ) := hf.nnnorm.coe_nnreal_ennreal @[measurability, fun_prop] -theorem AEMeasurable.ennnorm {f : β → α} {μ : Measure β} (hf : AEMeasurable f μ) : - AEMeasurable (fun a => (‖f a‖₊ : ℝ≥0∞)) μ := - measurable_ennnorm.comp_aemeasurable hf +protected lemma AEMeasurable.enorm {f : β → α} {μ : Measure β} (hf : AEMeasurable f μ) : + AEMeasurable (‖f ·‖ₑ) μ := + measurable_enorm.comp_aemeasurable hf + +@[deprecated (since := "2025-01-21")] alias measurable_ennnorm := measurable_enorm +@[deprecated (since := "2025-01-21")] alias Measurable.ennnorm := Measurable.enorm +@[deprecated (since := "2025-01-21")] alias AEMeasurable.ennnorm := AEMeasurable.enorm end NormedAddCommGroup diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index d2ad209b30457b..e81c5ea04ae2fd 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -54,7 +54,7 @@ noncomputable section open Function Set MeasureTheory.OuterMeasure Filter MeasurableSpace Encodable -open scoped Classical Topology ENNReal +open scoped Topology ENNReal universe u v @@ -168,6 +168,7 @@ open List MeasurableEquiv variable [Encodable ι] +open scoped Classical in /-- The product measure on an encodable finite type, defined by mapping `Measure.tprod` along the equivalence `MeasurableEquiv.piMeasurableEquivTProd`. The definition `MeasureTheory.Measure.pi` should be used instead of this one. -/ @@ -176,6 +177,7 @@ def pi' : Measure (∀ i, α i) := theorem pi'_pi [∀ i, SigmaFinite (μ i)] (s : ∀ i, Set (α i)) : pi' μ (pi univ s) = ∏ i, μ i (s i) := by + classical rw [pi'] rw [← MeasurableEquiv.piMeasurableEquivTProd_symm_apply, MeasurableEquiv.map_apply, MeasurableEquiv.piMeasurableEquivTProd_symm_apply, elim_preimage_pi, tprod_tprod _ μ, ← @@ -342,6 +344,7 @@ theorem pi_empty_univ {α : Type*} [Fintype α] [IsEmpty α] {β : α → Type*} theorem pi_eval_preimage_null {i : ι} {s : Set (α i)} (hs : μ i s = 0) : Measure.pi μ (eval i ⁻¹' s) = 0 := by + classical -- WLOG, `s` is measurable rcases exists_measurable_superset_of_null hs with ⟨t, hst, _, hμt⟩ suffices Measure.pi μ (eval i ⁻¹' t) = 0 from measure_mono_null (preimage_mono hst) this diff --git a/Mathlib/MeasureTheory/Covering/Besicovitch.lean b/Mathlib/MeasureTheory/Covering/Besicovitch.lean index 9f8b24ead46f20..b2868a890cc06d 100644 --- a/Mathlib/MeasureTheory/Covering/Besicovitch.lean +++ b/Mathlib/MeasureTheory/Covering/Besicovitch.lean @@ -102,7 +102,7 @@ universe u open Metric Set Filter Fin MeasureTheory TopologicalSpace -open scoped Topology Classical ENNReal MeasureTheory NNReal +open scoped Topology ENNReal MeasureTheory NNReal /-! ### Satellite configurations @@ -528,6 +528,7 @@ theorem exist_finset_disjoint_balls_large_measure (μ : Measure α) [IsFiniteMea (rpos : ∀ x ∈ s, 0 < r x) (rle : ∀ x ∈ s, r x ≤ 1) : ∃ t : Finset α, ↑t ⊆ s ∧ μ (s \ ⋃ x ∈ t, closedBall x (r x)) ≤ N / (N + 1) * μ s ∧ (t : Set α).PairwiseDisjoint fun x => closedBall x (r x) := by + classical -- exclude the trivial case where `μ s = 0`. rcases le_or_lt (μ s) 0 with (hμs | hμs) · have : μ s = 0 := le_bot_iff.1 hμs @@ -667,6 +668,7 @@ theorem exists_disjoint_closedBall_covering_ae_of_finiteMeasure_aux (μ : Measur ∃ t : Set (α × ℝ), t.Countable ∧ (∀ p ∈ t, p.1 ∈ s) ∧ (∀ p ∈ t, p.2 ∈ f p.1) ∧ μ (s \ ⋃ (p : α × ℝ) (_ : p ∈ t), closedBall p.1 p.2) = 0 ∧ t.PairwiseDisjoint fun p => closedBall p.1 p.2 := by + classical rcases HasBesicovitchCovering.no_satelliteConfig (α := α) with ⟨N, τ, hτ, hN⟩ /- Introduce a property `P` on finsets saying that we have a nice disjoint covering of a subset of `s` by admissible balls. -/ @@ -862,6 +864,7 @@ theorem exists_closedBall_covering_tsum_measure_le (μ : Measure α) [SFinite μ disjoint subfamilies. Making sure that they are all included in a neighborhood `v` of `s'` of measure at most `ε / (2 N)`, the sum of their measures is at most `ε / 2`, completing the proof. -/ + classical obtain ⟨u, su, u_open, μu⟩ : ∃ U, U ⊇ s ∧ IsOpen U ∧ μ U ≤ μ s + ε / 2 := Set.exists_isOpen_le_add _ _ (by diff --git a/Mathlib/MeasureTheory/Covering/Differentiation.lean b/Mathlib/MeasureTheory/Covering/Differentiation.lean index 672bfe1a4f090c..f54de11ae1ded1 100644 --- a/Mathlib/MeasureTheory/Covering/Differentiation.lean +++ b/Mathlib/MeasureTheory/Covering/Differentiation.lean @@ -770,10 +770,10 @@ theorem ae_tendsto_lintegral_div {f : α → ℝ≥0∞} (hf : AEMeasurable f μ apply lintegral_congr_ae exact ae_restrict_of_ae hf.ae_eq_mk -theorem ae_tendsto_lintegral_nnnorm_sub_div'_of_integrable {f : α → E} (hf : Integrable f μ) +theorem ae_tendsto_lintegral_enorm_sub_div'_of_integrable {f : α → E} (hf : Integrable f μ) (h'f : StronglyMeasurable f) : - ∀ᵐ x ∂μ, Tendsto (fun a => (∫⁻ y in a, ‖f y - f x‖₊ ∂μ) / μ a) (v.filterAt x) (𝓝 0) := by - /- For every `c`, then `(∫⁻ y in a, ‖f y - c‖₊ ∂μ) / μ a` tends almost everywhere to `‖f x - c‖`. + ∀ᵐ x ∂μ, Tendsto (fun a => (∫⁻ y in a, ‖f y - f x‖ₑ ∂μ) / μ a) (v.filterAt x) (𝓝 0) := by + /- For every `c`, then `(∫⁻ y in a, ‖f y - c‖ₑ ∂μ) / μ a` tends almost everywhere to `‖f x - c‖`. We apply this to a countable set of `c` which is dense in the range of `f`, to deduce the desired convergence. A minor technical inconvenience is that constants are not integrable, so to apply previous @@ -784,37 +784,31 @@ theorem ae_tendsto_lintegral_nnnorm_sub_div'_of_integrable {f : α → E} (hf : have main : ∀ᵐ x ∂μ, ∀ᵉ (n : ℕ) (c ∈ t), - Tendsto (fun a => (∫⁻ y in a, ‖f y - (A.set n).indicator (fun _ => c) y‖₊ ∂μ) / μ a) - (v.filterAt x) (𝓝 ‖f x - (A.set n).indicator (fun _ => c) x‖₊) := by + Tendsto (fun a => (∫⁻ y in a, ‖f y - (A.set n).indicator (fun _ => c) y‖ₑ ∂μ) / μ a) + (v.filterAt x) (𝓝 ‖f x - (A.set n).indicator (fun _ => c) x‖ₑ) := by #adaptation_note /-- 2024-04-23 The next two lines were previously just `simp_rw [ae_all_iff, ae_ball_iff t_count]`. -/ simp_rw [ae_all_iff] intro x; rw [ae_ball_iff t_count]; revert x intro n c _ apply ae_tendsto_lintegral_div' - · refine (h'f.sub ?_).ennnorm + · refine (h'f.sub ?_).enorm exact stronglyMeasurable_const.indicator (IsOpen.measurableSet (A.set_mem n)) · apply ne_of_lt calc - (∫⁻ y, ↑‖f y - (A.set n).indicator (fun _ : α => c) y‖₊ ∂μ) ≤ - ∫⁻ y, ‖f y‖₊ + ‖(A.set n).indicator (fun _ : α => c) y‖₊ ∂μ := by - apply lintegral_mono - intro x - dsimp - rw [← ENNReal.coe_add] - exact ENNReal.coe_le_coe.2 (nnnorm_sub_le _ _) - _ = (∫⁻ y, ‖f y‖₊ ∂μ) + ∫⁻ y, ‖(A.set n).indicator (fun _ : α => c) y‖₊ ∂μ := - (lintegral_add_left h'f.ennnorm _) + ∫⁻ y, ‖f y - (A.set n).indicator (fun _ : α => c) y‖ₑ ∂μ + ≤ ∫⁻ y, ‖f y‖ₑ + ‖(A.set n).indicator (fun _ : α => c) y‖ₑ ∂μ := + lintegral_mono fun x ↦ enorm_sub_le + _ = ∫⁻ y, ‖f y‖ₑ ∂μ + ∫⁻ y, ‖(A.set n).indicator (fun _ : α => c) y‖ₑ ∂μ := + lintegral_add_left h'f.enorm _ _ < ∞ + ∞ := haveI I : Integrable ((A.set n).indicator fun _ : α => c) μ := by simp only [integrable_indicator_iff (IsOpen.measurableSet (A.set_mem n)), integrableOn_const, A.finite n, or_true] ENNReal.add_lt_add hf.2 I.2 filter_upwards [main, v.ae_eventually_measure_pos] with x hx h'x - have M : - ∀ c ∈ t, Tendsto (fun a => (∫⁻ y in a, ‖f y - c‖₊ ∂μ) / μ a) - (v.filterAt x) (𝓝 ‖f x - c‖₊) := by - intro c hc + have M c (hc : c ∈ t) : + Tendsto (fun a => (∫⁻ y in a, ‖f y - c‖ₑ ∂μ) / μ a) (v.filterAt x) (𝓝 ‖f x - c‖ₑ) := by obtain ⟨n, xn⟩ : ∃ n, x ∈ A.set n := by simpa [← A.spanning] using mem_univ x specialize hx n c hc simp only [xn, indicator_of_mem] at hx @@ -825,18 +819,18 @@ theorem ae_tendsto_lintegral_nnnorm_sub_div'_of_integrable {f : α → E} (hf : apply setLIntegral_congr_fun h'a filter_upwards with y hy using (by simp only [ha hy, indicator_of_mem]) apply ENNReal.tendsto_nhds_zero.2 fun ε εpos => ?_ - obtain ⟨c, ct, xc⟩ : ∃ c ∈ t, (‖f x - c‖₊ : ℝ≥0∞) < ε / 2 := by - simp_rw [← edist_eq_coe_nnnorm_sub] + obtain ⟨c, ct, xc⟩ : ∃ c ∈ t, ‖f x - c‖ₑ < ε / 2 := by + simp_rw [← edist_eq_enorm_sub] have : f x ∈ closure t := ht (mem_range_self _) exact EMetric.mem_closure_iff.1 this (ε / 2) (ENNReal.half_pos (ne_of_gt εpos)) filter_upwards [(tendsto_order.1 (M c ct)).2 (ε / 2) xc, h'x, v.eventually_measure_lt_top x] with a ha h'a h''a apply ENNReal.div_le_of_le_mul calc - (∫⁻ y in a, ‖f y - f x‖₊ ∂μ) ≤ ∫⁻ y in a, ‖f y - c‖₊ + ‖f x - c‖₊ ∂μ := by + (∫⁻ y in a, ‖f y - f x‖ₑ ∂μ) ≤ ∫⁻ y in a, ‖f y - c‖ₑ + ‖f x - c‖ₑ ∂μ := by apply lintegral_mono fun x => ?_ - simpa only [← edist_eq_coe_nnnorm_sub] using edist_triangle_right _ _ _ - _ = (∫⁻ y in a, ‖f y - c‖₊ ∂μ) + ∫⁻ _ in a, ‖f x - c‖₊ ∂μ := + simpa only [← edist_eq_enorm_sub] using edist_triangle_right _ _ _ + _ = (∫⁻ y in a, ‖f y - c‖ₑ ∂μ) + ∫⁻ _ in a, ‖f x - c‖ₑ ∂μ := (lintegral_add_right _ measurable_const) _ ≤ ε / 2 * μ a + ε / 2 * μ a := by gcongr @@ -846,10 +840,14 @@ theorem ae_tendsto_lintegral_nnnorm_sub_div'_of_integrable {f : α → E} (hf : gcongr _ = ε * μ a := by rw [← add_mul, ENNReal.add_halves] -theorem ae_tendsto_lintegral_nnnorm_sub_div_of_integrable {f : α → E} (hf : Integrable f μ) : - ∀ᵐ x ∂μ, Tendsto (fun a => (∫⁻ y in a, ‖f y - f x‖₊ ∂μ) / μ a) (v.filterAt x) (𝓝 0) := by +@[deprecated (since := "2025-01-22")] +alias ae_tendsto_lintegral_nnnorm_sub_div'_of_integrable := + ae_tendsto_lintegral_enorm_sub_div'_of_integrable + +theorem ae_tendsto_lintegral_enorm_sub_div_of_integrable {f : α → E} (hf : Integrable f μ) : + ∀ᵐ x ∂μ, Tendsto (fun a => (∫⁻ y in a, ‖f y - f x‖ₑ ∂μ) / μ a) (v.filterAt x) (𝓝 0) := by have I : Integrable (hf.1.mk f) μ := hf.congr hf.1.ae_eq_mk - filter_upwards [v.ae_tendsto_lintegral_nnnorm_sub_div'_of_integrable I hf.1.stronglyMeasurable_mk, + filter_upwards [v.ae_tendsto_lintegral_enorm_sub_div'_of_integrable I hf.1.stronglyMeasurable_mk, hf.1.ae_eq_mk] with x hx h'x apply hx.congr _ intro a @@ -859,14 +857,18 @@ theorem ae_tendsto_lintegral_nnnorm_sub_div_of_integrable {f : α → E} (hf : I filter_upwards [hf.1.ae_eq_mk] with y hy rw [hy, h'x] -theorem ae_tendsto_lintegral_nnnorm_sub_div {f : α → E} (hf : LocallyIntegrable f μ) : - ∀ᵐ x ∂μ, Tendsto (fun a => (∫⁻ y in a, ‖f y - f x‖₊ ∂μ) / μ a) (v.filterAt x) (𝓝 0) := by +@[deprecated (since := "2025-01-22")] +alias ae_tendsto_lintegral_nnnorm_sub_div_of_integrable := + ae_tendsto_lintegral_enorm_sub_div_of_integrable + +theorem ae_tendsto_lintegral_enorm_sub_div {f : α → E} (hf : LocallyIntegrable f μ) : + ∀ᵐ x ∂μ, Tendsto (fun a => (∫⁻ y in a, ‖f y - f x‖ₑ ∂μ) / μ a) (v.filterAt x) (𝓝 0) := by rcases hf.exists_nat_integrableOn with ⟨u, u_open, u_univ, hu⟩ have : ∀ n, ∀ᵐ x ∂μ, - Tendsto (fun a => (∫⁻ y in a, ‖(u n).indicator f y - (u n).indicator f x‖₊ ∂μ) / μ a) + Tendsto (fun a => (∫⁻ y in a, ‖(u n).indicator f y - (u n).indicator f x‖ₑ ∂μ) / μ a) (v.filterAt x) (𝓝 0) := by intro n - apply ae_tendsto_lintegral_nnnorm_sub_div_of_integrable + apply ae_tendsto_lintegral_enorm_sub_div_of_integrable exact (integrable_indicator_iff (u_open n).measurableSet).2 (hu n) filter_upwards [ae_all_iff.2 this] with x hx obtain ⟨n, hn⟩ : ∃ n, x ∈ u n := by simpa only [← u_univ, mem_iUnion] using mem_univ x @@ -877,11 +879,14 @@ theorem ae_tendsto_lintegral_nnnorm_sub_div {f : α → E} (hf : LocallyIntegrab refine setLIntegral_congr_fun h'a (Eventually.of_forall (fun y hy ↦ ?_)) rw [indicator_of_mem (ha hy) f, indicator_of_mem hn f] +@[deprecated (since := "2025-01-22")] +alias ae_tendsto_lintegral_nnnorm_sub_div := ae_tendsto_lintegral_enorm_sub_div + /-- *Lebesgue differentiation theorem*: for almost every point `x`, the average of `‖f y - f x‖` on `a` tends to `0` as `a` shrinks to `x` along a Vitali family. -/ theorem ae_tendsto_average_norm_sub {f : α → E} (hf : LocallyIntegrable f μ) : ∀ᵐ x ∂μ, Tendsto (fun a => ⨍ y in a, ‖f y - f x‖ ∂μ) (v.filterAt x) (𝓝 0) := by - filter_upwards [v.ae_tendsto_lintegral_nnnorm_sub_div hf] with x hx + filter_upwards [v.ae_tendsto_lintegral_enorm_sub_div hf] with x hx have := (ENNReal.tendsto_toReal ENNReal.zero_ne_top).comp hx simp only [ENNReal.zero_toReal] at this apply Tendsto.congr' _ this @@ -891,6 +896,7 @@ theorem ae_tendsto_average_norm_sub {f : α → E} (hf : LocallyIntegrable f μ) have A : IntegrableOn (fun y => (‖f y - f x‖₊ : ℝ)) a μ := by simp_rw [coe_nnnorm] exact (h''a.sub (integrableOn_const.2 (Or.inr h'a))).norm + dsimp [enorm] rw [lintegral_coe_eq_integral _ A, ENNReal.toReal_ofReal (by positivity)] simp only [coe_nnnorm, smul_eq_mul] diff --git a/Mathlib/MeasureTheory/Covering/Vitali.lean b/Mathlib/MeasureTheory/Covering/Vitali.lean index 5b2934b6db6d82..61f42f43718a68 100644 --- a/Mathlib/MeasureTheory/Covering/Vitali.lean +++ b/Mathlib/MeasureTheory/Covering/Vitali.lean @@ -310,7 +310,7 @@ theorem exists_disjoint_covering_ae linarith [Idist_v b ⟨bu, hbx⟩] -- we will show that, in `ball x (R x)`, almost all `s` is covered by the family `u`. refine ⟨_ ∩ ball x (R x), inter_mem_nhdsWithin _ (ball_mem_nhds _ (hR0 _)), - nonpos_iff_eq_zero.mp (le_of_forall_le_of_dense fun ε εpos => ?_)⟩ + nonpos_iff_eq_zero.mp (le_of_forall_gt_imp_ge_of_dense fun ε εpos => ?_)⟩ -- the elements of `v` are disjoint and all contained in a finite volume ball, hence the sum -- of their measures is finite. have I : (∑' a : v, μ (B a)) < ∞ := by diff --git a/Mathlib/MeasureTheory/Decomposition/IntegralRNDeriv.lean b/Mathlib/MeasureTheory/Decomposition/IntegralRNDeriv.lean new file mode 100644 index 00000000000000..e2ddf8630cf97b --- /dev/null +++ b/Mathlib/MeasureTheory/Decomposition/IntegralRNDeriv.lean @@ -0,0 +1,99 @@ +/- +Copyright (c) 2025 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne, Lorenzo Luccioli +-/ +import Mathlib.Analysis.Convex.Continuous +import Mathlib.Analysis.Convex.Integral +import Mathlib.MeasureTheory.Decomposition.RadonNikodym + +/-! +# Integrals of functions of Radon-Nikodym derivatives + +## Main statements + +* `mul_le_integral_rnDeriv_of_ac`: for a convex continuous function `f` on `[0, ∞)`, if `μ` + is absolutely continuous with respect to `ν`, then + `(ν univ).toReal * f ((μ univ).toReal / (ν univ).toReal) ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν`. + +-/ + + +open Set + +namespace MeasureTheory + +variable {α : Type*} {mα : MeasurableSpace α} {μ ν : Measure α} {f : ℝ → ℝ} + +/-- For a convex continuous function `f` on `[0, ∞)`, if `μ` is absolutely continuous +with respect to a probability measure `ν`, then +`f (μ univ).toReal ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν`. -/ +lemma le_integral_rnDeriv_of_ac [IsFiniteMeasure μ] [IsProbabilityMeasure ν] + (hf_cvx : ConvexOn ℝ (Ici 0) f) (hf_cont : ContinuousWithinAt f (Ici 0) 0) + (hf_int : Integrable (fun x ↦ f (μ.rnDeriv ν x).toReal) ν) (hμν : μ ≪ ν) : + f (μ univ).toReal ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν := by + have hf_cont' : ContinuousOn f (Ici 0) := by + intro x hx + rcases eq_or_lt_of_le (α := ℝ) (hx : 0 ≤ x) with rfl | hx_pos + · exact hf_cont + · have h := hf_cvx.continuousOn_interior x + simp only [nonempty_Iio, interior_Ici', mem_Ioi] at h + rw [continuousWithinAt_iff_continuousAt (Ioi_mem_nhds hx_pos)] at h + exact (h hx_pos).continuousWithinAt + calc f (μ univ).toReal + = f (∫ x, (μ.rnDeriv ν x).toReal ∂ν) := by rw [Measure.integral_toReal_rnDeriv hμν] + _ ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν := by + rw [← average_eq_integral, ← average_eq_integral] + exact ConvexOn.map_average_le hf_cvx hf_cont' isClosed_Ici (by simp) + Measure.integrable_toReal_rnDeriv hf_int + +/-- For a convex continuous function `f` on `[0, ∞)`, if `μ` is absolutely continuous +with respect to `ν`, then +`(ν univ).toReal * f ((μ univ).toReal / (ν univ).toReal) ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν`. -/ +lemma mul_le_integral_rnDeriv_of_ac [IsFiniteMeasure μ] [IsFiniteMeasure ν] + (hf_cvx : ConvexOn ℝ (Ici 0) f) (hf_cont : ContinuousWithinAt f (Ici 0) 0) + (hf_int : Integrable (fun x ↦ f (μ.rnDeriv ν x).toReal) ν) (hμν : μ ≪ ν) : + (ν univ).toReal * f ((μ univ).toReal / (ν univ).toReal) + ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν := by + by_cases hν : ν = 0 + · simp [hν] + have : NeZero ν := ⟨hν⟩ + let μ' := (ν univ)⁻¹ • μ + let ν' := (ν univ)⁻¹ • ν + have : IsFiniteMeasure μ' := μ.smul_finite (by simp [hν]) + have hμν' : μ' ≪ ν' := hμν.smul _ + have h_rnDeriv_eq : μ'.rnDeriv ν' =ᵐ[ν] μ.rnDeriv ν := by + have h1' : μ'.rnDeriv ν' =ᵐ[ν'] (ν univ)⁻¹ • μ.rnDeriv ν' := + Measure.rnDeriv_smul_left_of_ne_top' (μ := ν') (ν := μ) (by simp [hν]) + have h1 : μ'.rnDeriv ν' =ᵐ[ν] (ν univ)⁻¹ • μ.rnDeriv ν' := by + rwa [Measure.ae_smul_measure_eq] at h1' + simp + have h2 : μ.rnDeriv ν' =ᵐ[ν] (ν univ)⁻¹⁻¹ • μ.rnDeriv ν := + Measure.rnDeriv_smul_right_of_ne_top' (μ := ν) (ν := μ) (by simp) (by simp [hν]) + filter_upwards [h1, h2] with x h1 h2 + rw [h1, Pi.smul_apply, smul_eq_mul, h2] + simp only [inv_inv, Pi.smul_apply, smul_eq_mul] + rw [← mul_assoc, ENNReal.inv_mul_cancel, one_mul] + · simp [hν] + · simp + have h_eq : ∫ x, f (μ'.rnDeriv ν' x).toReal ∂ν' + = (ν univ).toReal⁻¹ * ∫ x, f ((μ.rnDeriv ν x).toReal) ∂ν := by + rw [integral_smul_measure, smul_eq_mul, ENNReal.toReal_inv] + congr 1 + refine integral_congr_ae ?_ + filter_upwards [h_rnDeriv_eq] with x hx + rw [hx] + have h : f (μ' univ).toReal ≤ ∫ x, f (μ'.rnDeriv ν' x).toReal ∂ν' := + le_integral_rnDeriv_of_ac hf_cvx hf_cont ?_ hμν' + swap + · refine Integrable.smul_measure ?_ (by simp [hν]) + refine (integrable_congr ?_).mpr hf_int + filter_upwards [h_rnDeriv_eq] with x hx + rw [hx] + rw [h_eq, mul_comm, ← div_le_iff₀, div_eq_inv_mul, inv_inv] at h + · convert h + · simp only [div_eq_inv_mul, Measure.smul_apply, smul_eq_mul, ENNReal.toReal_mul, + ENNReal.toReal_inv, μ'] + · simp [ENNReal.toReal_pos_iff, hν] + +end MeasureTheory diff --git a/Mathlib/MeasureTheory/Decomposition/Jordan.lean b/Mathlib/MeasureTheory/Decomposition/Jordan.lean index ce9122ea87126f..e4c8c69908a4c9 100644 --- a/Mathlib/MeasureTheory/Decomposition/Jordan.lean +++ b/Mathlib/MeasureTheory/Decomposition/Jordan.lean @@ -43,7 +43,7 @@ Jordan decomposition theorem noncomputable section -open scoped Classical MeasureTheory ENNReal NNReal +open scoped MeasureTheory ENNReal NNReal variable {α : Type*} [MeasurableSpace α] @@ -185,7 +185,6 @@ end JordanDecomposition namespace SignedMeasure -open scoped Classical open JordanDecomposition Measure Set VectorMeasure variable {s : SignedMeasure α} diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index f1cececd23a181..ca9899a429ecca 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -5,7 +5,7 @@ Authors: Kexing Ying -/ import Mathlib.MeasureTheory.Decomposition.UnsignedHahn import Mathlib.MeasureTheory.Function.AEEqOfLIntegral -import Mathlib.MeasureTheory.Function.L1Space +import Mathlib.MeasureTheory.Function.L1Space.Integrable import Mathlib.MeasureTheory.Measure.Sub /-! diff --git a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean index d0c1bf8f1d9423..f8f7b82e720718 100644 --- a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean +++ b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean @@ -43,7 +43,7 @@ Radon-Nikodym theorem noncomputable section -open scoped Classical MeasureTheory NNReal ENNReal +open scoped MeasureTheory NNReal ENNReal variable {α β : Type*} {m : MeasurableSpace α} diff --git a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean index 0a690c953d507f..003de465b8ccb2 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean @@ -37,7 +37,7 @@ Hahn decomposition theorem noncomputable section -open scoped Classical NNReal ENNReal MeasureTheory +open scoped NNReal ENNReal MeasureTheory variable {α β : Type*} [MeasurableSpace α] variable {M : Type*} [AddCommMonoid M] [TopologicalSpace M] [OrderedAddCommMonoid M] @@ -99,6 +99,7 @@ private theorem existsNatOneDivLTMeasure_of_not_negative (hi : ¬s ≤[i] 0) : let ⟨n, hn⟩ := exists_nat_one_div_lt hj ⟨n, k, hj₂, hj₁, hn⟩ +open scoped Classical in /-- Given the set `i`, if `i` is not negative, `findExistsOneDivLT s i` is the least natural number `n` such that `ExistsOneDivLT s i n`, otherwise, it returns 0. -/ private def findExistsOneDivLT (s : SignedMeasure α) (i : Set α) : ℕ := @@ -111,9 +112,11 @@ private theorem findExistsOneDivLT_spec (hi : ¬s ≤[i] 0) : private theorem findExistsOneDivLT_min (hi : ¬s ≤[i] 0) {m : ℕ} (hm : m < findExistsOneDivLT s i) : ¬ExistsOneDivLT s i m := by + classical rw [findExistsOneDivLT, dif_pos hi] at hm exact Nat.find_min _ hm +open scoped Classical in /-- Given the set `i`, if `i` is not negative, `someExistsOneDivLT` chooses the set `k` from `ExistsOneDivLT s i (findExistsOneDivLT s i)`, otherwise, it returns the empty set. -/ @@ -222,6 +225,7 @@ private theorem restrictNonposSeq_disjoint : Pairwise (Disjoint on restrictNonpo private theorem exists_subset_restrict_nonpos' (hi₁ : MeasurableSet i) (hi₂ : s i < 0) (hn : ¬∀ n : ℕ, ¬s ≤[i \ ⋃ l < n, restrictNonposSeq s i l] 0) : ∃ j : Set α, MeasurableSet j ∧ j ⊆ i ∧ s ≤[j] 0 ∧ s j < 0 := by + classical by_cases h : s ≤[i] 0 · exact ⟨i, hi₁, Set.Subset.refl _, h, hi₂⟩ push_neg at hn diff --git a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean index 354c4e16c02103..c20637659cb297 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean @@ -42,7 +42,7 @@ Lebesgue decomposition theorem noncomputable section -open scoped Classical MeasureTheory NNReal ENNReal +open scoped MeasureTheory NNReal ENNReal open Set diff --git a/Mathlib/MeasureTheory/Function/AEMeasurableOrder.lean b/Mathlib/MeasureTheory/Function/AEMeasurableOrder.lean index 0f4dba7c8de330..a26a30a2fad26d 100644 --- a/Mathlib/MeasureTheory/Function/AEMeasurableOrder.lean +++ b/Mathlib/MeasureTheory/Function/AEMeasurableOrder.lean @@ -23,7 +23,6 @@ as possible. open MeasureTheory Set TopologicalSpace -open scoped Classical open ENNReal NNReal /-- If a function `f : α → β` is such that the level sets `{f < p}` and `{q < f}` have measurable @@ -36,6 +35,7 @@ theorem MeasureTheory.aemeasurable_of_exist_almost_disjoint_supersets {α : Type (h : ∀ p ∈ s, ∀ q ∈ s, p < q → ∃ u v, MeasurableSet u ∧ MeasurableSet v ∧ { x | f x < p } ⊆ u ∧ { x | q < f x } ⊆ v ∧ μ (u ∩ v) = 0) : AEMeasurable f μ := by + classical haveI : Encodable s := s_count.toEncodable have h' : ∀ p q, ∃ u v, MeasurableSet u ∧ MeasurableSet v ∧ { x | f x < p } ⊆ u ∧ { x | q < f x } ⊆ v ∧ (p ∈ s → q ∈ s → p < q → μ (u ∩ v) = 0) := by diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean index e5fb77b18b62e5..98df7b831fb53b 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean @@ -69,7 +69,7 @@ conditional expectation, conditional expected value -/ open TopologicalSpace MeasureTheory.Lp Filter -open scoped Classical ENNReal Topology MeasureTheory +open scoped ENNReal Topology MeasureTheory namespace MeasureTheory -- 𝕜 for ℝ or ℂ @@ -80,6 +80,7 @@ variable {α β E 𝕜 : Type*} [RCLike 𝕜] {m m₀ : MeasurableSpace α} {μ section NormedAddCommGroup variable [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +open scoped Classical in variable (m) in /-- Conditional expectation of a function. It is defined as 0 if any one of the following conditions is true: @@ -108,6 +109,7 @@ theorem condExp_of_not_sigmaFinite (hm : m ≤ m₀) (hμm_not : ¬SigmaFinite ( @[deprecated (since := "2025-01-21")] alias condexp_of_not_sigmaFinite := condExp_of_not_sigmaFinite +open scoped Classical in theorem condExp_of_sigmaFinite (hm : m ≤ m₀) [hμm : SigmaFinite (μ.trim hm)] : μ[f|m] = if Integrable f μ then @@ -320,6 +322,7 @@ theorem condExp_add (hf : Integrable f μ) (hg : Integrable g μ) (m : Measurabl theorem condExp_finset_sum {ι : Type*} {s : Finset ι} {f : ι → α → E} (hf : ∀ i ∈ s, Integrable (f i) μ) (m : MeasurableSpace α) : μ[∑ i ∈ s, f i|m] =ᵐ[μ] ∑ i ∈ s, μ[f i|m] := by + classical induction' s using Finset.induction_on with i s his heq hf · rw [Finset.sum_empty, Finset.sum_empty, condExp_zero] · rw [Finset.sum_insert his, Finset.sum_insert his] diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean index 1871e8db5807f7..5bb82053bb1764 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean @@ -132,15 +132,15 @@ theorem norm_condExpIndL1Fin_le (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x ‖condExpIndL1Fin hm hs hμs x‖ ≤ (μ s).toReal * ‖x‖ := by rw [L1.norm_eq_integral_norm, ← ENNReal.toReal_ofReal (norm_nonneg x), ← ENNReal.toReal_mul, ← ENNReal.ofReal_le_iff_le_toReal (ENNReal.mul_ne_top hμs ENNReal.ofReal_ne_top), - ofReal_integral_norm_eq_lintegral_nnnorm] + ofReal_integral_norm_eq_lintegral_enorm] swap; · rw [← memℒp_one_iff_integrable]; exact Lp.memℒp _ have h_eq : - ∫⁻ a, ‖condExpIndL1Fin hm hs hμs x a‖₊ ∂μ = ∫⁻ a, ‖condExpIndSMul hm hs hμs x a‖₊ ∂μ := by + ∫⁻ a, ‖condExpIndL1Fin hm hs hμs x a‖ₑ ∂μ = ∫⁻ a, ‖condExpIndSMul hm hs hμs x a‖ₑ ∂μ := by refine lintegral_congr_ae ?_ refine (condExpIndL1Fin_ae_eq_condExpIndSMul hm hs hμs x).mono fun z hz => ?_ dsimp only rw [hz] - rw [h_eq, ofReal_norm_eq_coe_nnnorm] + rw [h_eq, ofReal_norm_eq_enorm] exact lintegral_nnnorm_condExpIndSMul_le hm hs hμs x @[deprecated (since := "2025-01-21")] alias norm_condexpIndL1Fin_le := norm_condExpIndL1Fin_le @@ -170,11 +170,10 @@ alias condexpIndL1Fin_disjoint_union := condExpIndL1Fin_disjoint_union end CondexpIndL1Fin -open scoped Classical - section CondexpIndL1 +open scoped Classical in /-- Conditional expectation of the indicator of a set, as a function in L1. Its value for sets which are not both measurable and of finite measure is not used: we set it to 0. -/ def condExpIndL1 {m m0 : MeasurableSpace α} (hm : m ≤ m0) (μ : Measure α) (s : Set α) diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean index b6524435ba77e5..c51e17e6232436 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean @@ -194,7 +194,7 @@ theorem lintegral_nnnorm_condExpL2_le (hs : MeasurableSet[m] s) (hμs : μ s ≠ dsimp only simp_rw [hx] rw [lintegral_congr_ae hg_nnnorm_eq.symm] - refine lintegral_nnnorm_le_of_forall_fin_meas_integral_eq + refine lintegral_enorm_le_of_forall_fin_meas_integral_eq hm (Lp.stronglyMeasurable f) ?_ ?_ ?_ ?_ hs hμs · exact integrableOn_Lp_of_measure_ne_top f fact_one_le_two_ennreal.elim hμs · exact hg_meas @@ -225,7 +225,7 @@ theorem condExpL2_ae_eq_zero_of_ae_eq_zero (hs : MeasurableSet[m] s) (hμs : μ dsimp only rw [hx] simp - · exact (Lp.stronglyMeasurable _).ennnorm + · exact (Lp.stronglyMeasurable _).enorm @[deprecated (since := "2025-01-21")] alias condexpL2_ae_eq_zero_of_ae_eq_zero := condExpL2_ae_eq_zero_of_ae_eq_zero @@ -383,7 +383,7 @@ theorem setLIntegral_nnnorm_condExpL2_indicator_le (hm : m ≤ m0) (hs : Measura _ = (∫⁻ a in t, ‖(condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a‖₊ ∂μ) * ‖x‖₊ := by simp_rw [nnnorm_smul, ENNReal.coe_mul] rw [lintegral_mul_const, lpMeas_coe] - exact (Lp.stronglyMeasurable _).ennnorm + exact (Lp.stronglyMeasurable _).enorm _ ≤ μ (s ∩ t) * ‖x‖₊ := mul_le_mul_right' (lintegral_nnnorm_condExpL2_indicator_le_real hs hμs ht hμt) _ @@ -488,7 +488,7 @@ theorem setLIntegral_nnnorm_condExpIndSMul_le (hm : m ≤ m0) (hs : MeasurableSe _ = (∫⁻ a in t, ‖(condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a‖₊ ∂μ) * ‖x‖₊ := by simp_rw [nnnorm_smul, ENNReal.coe_mul] rw [lintegral_mul_const, lpMeas_coe] - exact (Lp.stronglyMeasurable _).ennnorm + exact (Lp.stronglyMeasurable _).enorm _ ≤ μ (s ∩ t) * ‖x‖₊ := mul_le_mul_right' (lintegral_nnnorm_condExpL2_indicator_le_real hs hμs ht hμt) _ diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean index 41c370cc5b4b63..3f4e72c29f95a1 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean @@ -71,12 +71,12 @@ theorem eLpNorm_one_condExp_le_eLpNorm (f : α → ℝ) : eLpNorm (μ[f|m]) 1 μ (ae_of_all μ (fun x => neg_le_abs (f x) : ∀ x, -f x ≤ |f x|)))] with x hx₁ hx₂ exact abs_le_abs hx₁ hx₂ _ = eLpNorm f 1 μ := by - rw [eLpNorm_one_eq_lintegral_nnnorm, eLpNorm_one_eq_lintegral_nnnorm, - ← ENNReal.toReal_eq_toReal (hasFiniteIntegral_iff_nnnorm.mp integrable_condExp.2).ne - (hasFiniteIntegral_iff_nnnorm.mp hf.2).ne, - ← integral_norm_eq_lintegral_nnnorm + rw [eLpNorm_one_eq_lintegral_enorm, eLpNorm_one_eq_lintegral_enorm, + ← ENNReal.toReal_eq_toReal (hasFiniteIntegral_iff_enorm.mp integrable_condExp.2).ne + (hasFiniteIntegral_iff_enorm.mp hf.2).ne, + ← integral_norm_eq_lintegral_enorm (stronglyMeasurable_condExp.mono hm).aestronglyMeasurable, - ← integral_norm_eq_lintegral_nnnorm hf.1] + ← integral_norm_eq_lintegral_enorm hf.1] simp_rw [Real.norm_eq_abs] rw (config := {occs := .pos [2]}) [← integral_condExp hm] refine integral_congr_ae ?_ @@ -106,9 +106,9 @@ theorem integral_abs_condExp_le (f : α → ℝ) : ∫ x, |(μ[f|m]) x| ∂μ Algebra.id.smul_eq_mul, mul_zero] positivity rw [integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae] - · apply ENNReal.toReal_mono <;> simp_rw [← Real.norm_eq_abs, ofReal_norm_eq_coe_nnnorm] + · apply ENNReal.toReal_mono <;> simp_rw [← Real.norm_eq_abs, ofReal_norm_eq_enorm] · exact hfint.2.ne - · rw [← eLpNorm_one_eq_lintegral_nnnorm, ← eLpNorm_one_eq_lintegral_nnnorm] + · rw [← eLpNorm_one_eq_lintegral_enorm, ← eLpNorm_one_eq_lintegral_enorm] exact eLpNorm_one_condExp_le_eLpNorm _ · filter_upwards with x using abs_nonneg _ · simp_rw [← Real.norm_eq_abs] @@ -237,7 +237,7 @@ alias Integrable.uniformIntegrable_condexp := Integrable.uniformIntegrable_condE section PullOut -- TODO: this section could be generalized beyond multiplication, to any bounded bilinear map. -/-- Auxiliary lemma for `condexp_mul_of_stronglyMeasurable_left`. -/ +/-- Auxiliary lemma for `condExp_mul_of_stronglyMeasurable_left`. -/ theorem condExp_stronglyMeasurable_simpleFunc_mul (hm : m ≤ m0) (f : @SimpleFunc α m ℝ) {g : α → ℝ} (hg : Integrable g μ) : μ[(f * g : α → ℝ)|m] =ᵐ[μ] f * μ[g|m] := by have : ∀ (s c) (f : α → ℝ), Set.indicator s (Function.const α c) * f = s.indicator (c • f) := by diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean index 033ce021ee4bed..490cbe24cc8175 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean @@ -204,18 +204,22 @@ theorem integral_norm_le_of_forall_fin_meas_integral_eq (hm : m ≤ m0) {f g : /-- Let `m` be a sub-σ-algebra of `m0`, `f` an `m0`-measurable function and `g` an `m`-measurable function, such that their integrals coincide on `m`-measurable sets with finite measure. -Then `∫⁻ x in s, ‖g x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ` on all `m`-measurable sets with finite +Then `∫⁻ x in s, ‖g x‖ₑ ∂μ ≤ ∫⁻ x in s, ‖f x‖ₑ ∂μ` on all `m`-measurable sets with finite measure. -/ -theorem lintegral_nnnorm_le_of_forall_fin_meas_integral_eq (hm : m ≤ m0) {f g : α → ℝ} +theorem lintegral_enorm_le_of_forall_fin_meas_integral_eq (hm : m ≤ m0) {f g : α → ℝ} (hf : StronglyMeasurable f) (hfi : IntegrableOn f s μ) (hg : StronglyMeasurable[m] g) (hgi : IntegrableOn g s μ) (hgf : ∀ t, MeasurableSet[m] t → μ t < ∞ → ∫ x in t, g x ∂μ = ∫ x in t, f x ∂μ) - (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) : (∫⁻ x in s, ‖g x‖₊ ∂μ) ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ := by - rw [← ofReal_integral_norm_eq_lintegral_nnnorm hfi, ← - ofReal_integral_norm_eq_lintegral_nnnorm hgi, ENNReal.ofReal_le_ofReal_iff] + (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) : (∫⁻ x in s, ‖g x‖ₑ ∂μ) ≤ ∫⁻ x in s, ‖f x‖ₑ ∂μ := by + rw [← ofReal_integral_norm_eq_lintegral_enorm hfi, ← + ofReal_integral_norm_eq_lintegral_enorm hgi, ENNReal.ofReal_le_ofReal_iff] · exact integral_norm_le_of_forall_fin_meas_integral_eq hm hf hfi hg hgi hgf hs hμs · positivity +@[deprecated (since := "2025-01-21")] +alias lintegral_nnnorm_le_of_forall_fin_meas_integral_eq := + lintegral_enorm_le_of_forall_fin_meas_integral_eq + end IntegralNormLE end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean index 633e66b2ab991a..8110b019c55bb8 100644 --- a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean +++ b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean @@ -117,11 +117,8 @@ theorem exists_continuous_eLpNorm_sub_le_of_closed [μ.OuterRegular] (hp : p ≠ have gc_mem : Memℒp (fun x => g x • c) p μ := by refine Memℒp.smul_of_top_left (memℒp_top_const _) ?_ refine ⟨g.continuous.aestronglyMeasurable, ?_⟩ - have : eLpNorm (v.indicator fun _x => (1 : ℝ)) p μ < ⊤ := by - refine (eLpNorm_indicator_const_le _ _).trans_lt ?_ - simp only [lt_top_iff_ne_top, hμv.ne, nnnorm_one, ENNReal.coe_one, one_div, one_mul, Ne, - ENNReal.rpow_eq_top_iff, inv_lt_zero, false_and, or_false, not_and, not_lt, - ENNReal.toReal_nonneg, imp_true_iff] + have : eLpNorm (v.indicator fun _x => (1 : ℝ)) p μ < ⊤ := + (eLpNorm_indicator_const_le _ _).trans_lt <| by simp [lt_top_iff_ne_top, hμv.ne] refine (eLpNorm_mono fun x => ?_).trans_lt this by_cases hx : x ∈ v · simp only [hx, abs_of_nonneg (hg_range x).1, (hg_range x).2, Real.norm_eq_abs, @@ -219,8 +216,8 @@ theorem Integrable.exists_hasCompactSupport_lintegral_sub_le [R1Space α] [WeaklyLocallyCompactSpace α] [μ.Regular] {f : α → E} (hf : Integrable f μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : ∃ g : α → E, - HasCompactSupport g ∧ (∫⁻ x, ‖f x - g x‖₊ ∂μ) ≤ ε ∧ Continuous g ∧ Integrable g μ := by - simp only [← memℒp_one_iff_integrable, ← eLpNorm_one_eq_lintegral_nnnorm] at hf ⊢ + HasCompactSupport g ∧ ∫⁻ x, ‖f x - g x‖ₑ ∂μ ≤ ε ∧ Continuous g ∧ Integrable g μ := by + simp only [← memℒp_one_iff_integrable, ← eLpNorm_one_eq_lintegral_enorm] at hf ⊢ exact hf.exists_hasCompactSupport_eLpNorm_sub_le ENNReal.one_ne_top hε /-- In a locally compact space, any integrable function can be approximated by compactly supported @@ -230,7 +227,7 @@ theorem Integrable.exists_hasCompactSupport_integral_sub_le {f : α → E} (hf : Integrable f μ) {ε : ℝ} (hε : 0 < ε) : ∃ g : α → E, HasCompactSupport g ∧ (∫ x, ‖f x - g x‖ ∂μ) ≤ ε ∧ Continuous g ∧ Integrable g μ := by - simp only [← memℒp_one_iff_integrable, ← eLpNorm_one_eq_lintegral_nnnorm, ← ENNReal.ofReal_one] + simp only [← memℒp_one_iff_integrable, ← eLpNorm_one_eq_lintegral_enorm, ← ENNReal.ofReal_one] at hf ⊢ simpa using hf.exists_hasCompactSupport_integral_rpow_sub_le zero_lt_one hε @@ -309,8 +306,8 @@ theorem Memℒp.exists_boundedContinuous_integral_rpow_sub_le [μ.WeaklyRegular] version in terms of `∫⁻`. -/ theorem Integrable.exists_boundedContinuous_lintegral_sub_le [μ.WeaklyRegular] {f : α → E} (hf : Integrable f μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : - ∃ g : α →ᵇ E, (∫⁻ x, ‖f x - g x‖₊ ∂μ) ≤ ε ∧ Integrable g μ := by - simp only [← memℒp_one_iff_integrable, ← eLpNorm_one_eq_lintegral_nnnorm] at hf ⊢ + ∃ g : α →ᵇ E, ∫⁻ x, ‖f x - g x‖ₑ ∂μ ≤ ε ∧ Integrable g μ := by + simp only [← memℒp_one_iff_integrable, ← eLpNorm_one_eq_lintegral_enorm] at hf ⊢ exact hf.exists_boundedContinuous_eLpNorm_sub_le ENNReal.one_ne_top hε /-- Any integrable function can be approximated by bounded continuous functions, @@ -318,7 +315,7 @@ version in terms of `∫`. -/ theorem Integrable.exists_boundedContinuous_integral_sub_le [μ.WeaklyRegular] {f : α → E} (hf : Integrable f μ) {ε : ℝ} (hε : 0 < ε) : ∃ g : α →ᵇ E, (∫ x, ‖f x - g x‖ ∂μ) ≤ ε ∧ Integrable g μ := by - simp only [← memℒp_one_iff_integrable, ← eLpNorm_one_eq_lintegral_nnnorm, ← ENNReal.ofReal_one] + simp only [← memℒp_one_iff_integrable, ← eLpNorm_one_eq_lintegral_enorm, ← ENNReal.ofReal_one] at hf ⊢ simpa using hf.exists_boundedContinuous_integral_rpow_sub_le zero_lt_one hε diff --git a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean index f9a7daf12fa481..e8eaafe2a7ed05 100644 --- a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean +++ b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean @@ -270,7 +270,7 @@ theorem tendstoInMeasure_of_tendsto_eLpNorm_of_stronglyMeasurable (hp_ne_zero : · rw [← ENNReal.ofReal_rpow_of_pos hε] convert mul_meas_ge_le_pow_eLpNorm' μ hp_ne_zero hp_ne_top ((hf n).sub hg).aestronglyMeasurable (ENNReal.ofReal ε) - rw [dist_eq_norm, ← ENNReal.ofReal_le_ofReal_iff (norm_nonneg _), ofReal_norm_eq_coe_nnnorm] + rw [dist_eq_norm, ← ENNReal.ofReal_le_ofReal_iff (norm_nonneg _), ofReal_norm_eq_enorm] exact Iff.rfl · rw [Ne, ENNReal.ofReal_eq_zero, not_le] exact Or.inl (Real.rpow_pos_of_pos hε _) diff --git a/Mathlib/MeasureTheory/Function/Egorov.lean b/Mathlib/MeasureTheory/Function/Egorov.lean index 9e40665728a348..104398ecddfc9d 100644 --- a/Mathlib/MeasureTheory/Function/Egorov.lean +++ b/Mathlib/MeasureTheory/Function/Egorov.lean @@ -23,7 +23,6 @@ convergence in measure. noncomputable section -open scoped Classical open MeasureTheory NNReal ENNReal Topology namespace MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/L1Space/AEEqFun.lean b/Mathlib/MeasureTheory/Function/L1Space/AEEqFun.lean new file mode 100644 index 00000000000000..b63aa09f634b42 --- /dev/null +++ b/Mathlib/MeasureTheory/Function/L1Space/AEEqFun.lean @@ -0,0 +1,231 @@ +/- +Copyright (c) 2019 Zhouhang Zhou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Zhouhang Zhou +-/ +import Mathlib.MeasureTheory.Function.L1Space.Integrable + +/-! +# `L¹` space + +In this file we establish an API between `Integrable` and the space `L¹` of equivalence +classes of integrable functions, already defined as a special case of `L^p` spaces for `p = 1`. + +## Notation + +* `α →₁[μ] β` is the type of `L¹` space, where `α` is a `MeasureSpace` and `β` is a + `NormedAddCommGroup`. `f : α →ₘ β` is a "function" in `L¹`. + In comments, `[f]` is also used to denote an `L¹` function. + + `₁` can be typed as `\1`. + +## Tags + +function space, l1 + +-/ + +noncomputable section + +open EMetric ENNReal Filter MeasureTheory NNReal Set + +variable {α β : Type*} {m : MeasurableSpace α} {μ ν : Measure α} +variable [NormedAddCommGroup β] + +namespace MeasureTheory + +namespace AEEqFun + +section + +/-- A class of almost everywhere equal functions is `Integrable` if its function representative +is integrable. -/ +def Integrable (f : α →ₘ[μ] β) : Prop := + MeasureTheory.Integrable f μ + +theorem integrable_mk {f : α → β} (hf : AEStronglyMeasurable f μ) : + Integrable (mk f hf : α →ₘ[μ] β) ↔ MeasureTheory.Integrable f μ := by + simp only [Integrable] + apply integrable_congr + exact coeFn_mk f hf + +theorem integrable_coeFn {f : α →ₘ[μ] β} : MeasureTheory.Integrable f μ ↔ Integrable f := by + rw [← integrable_mk, mk_coeFn] + +theorem integrable_zero : Integrable (0 : α →ₘ[μ] β) := + (MeasureTheory.integrable_zero α β μ).congr (coeFn_mk _ _).symm + +end + +section + +theorem Integrable.neg {f : α →ₘ[μ] β} : Integrable f → Integrable (-f) := + induction_on f fun _f hfm hfi => (integrable_mk _).2 ((integrable_mk hfm).1 hfi).neg + +section + +theorem integrable_iff_mem_L1 {f : α →ₘ[μ] β} : Integrable f ↔ f ∈ (α →₁[μ] β) := by + rw [← integrable_coeFn, ← memℒp_one_iff_integrable, Lp.mem_Lp_iff_memℒp] + +theorem Integrable.add {f g : α →ₘ[μ] β} : Integrable f → Integrable g → Integrable (f + g) := by + refine induction_on₂ f g fun f hf g hg hfi hgi => ?_ + simp only [integrable_mk, mk_add_mk] at hfi hgi ⊢ + exact hfi.add hgi + +theorem Integrable.sub {f g : α →ₘ[μ] β} (hf : Integrable f) (hg : Integrable g) : + Integrable (f - g) := + (sub_eq_add_neg f g).symm ▸ hf.add hg.neg + +end + +section BoundedSMul + +variable {𝕜 : Type*} [NormedRing 𝕜] [Module 𝕜 β] [BoundedSMul 𝕜 β] + +theorem Integrable.smul {c : 𝕜} {f : α →ₘ[μ] β} : Integrable f → Integrable (c • f) := + induction_on f fun _f hfm hfi => (integrable_mk _).2 <| + by simpa using ((integrable_mk hfm).1 hfi).smul c + +end BoundedSMul + +end + +end AEEqFun + +namespace L1 + + +theorem integrable_coeFn (f : α →₁[μ] β) : Integrable f μ := by + rw [← memℒp_one_iff_integrable] + exact Lp.memℒp f + +theorem hasFiniteIntegral_coeFn (f : α →₁[μ] β) : HasFiniteIntegral f μ := + (integrable_coeFn f).hasFiniteIntegral + +theorem stronglyMeasurable_coeFn (f : α →₁[μ] β) : StronglyMeasurable f := + Lp.stronglyMeasurable f + +theorem measurable_coeFn [MeasurableSpace β] [BorelSpace β] (f : α →₁[μ] β) : Measurable f := + (Lp.stronglyMeasurable f).measurable + +theorem aestronglyMeasurable_coeFn (f : α →₁[μ] β) : AEStronglyMeasurable f μ := + Lp.aestronglyMeasurable f + +theorem aemeasurable_coeFn [MeasurableSpace β] [BorelSpace β] (f : α →₁[μ] β) : AEMeasurable f μ := + (Lp.stronglyMeasurable f).measurable.aemeasurable + +theorem edist_def (f g : α →₁[μ] β) : edist f g = ∫⁻ a, edist (f a) (g a) ∂μ := by + simp only [Lp.edist_def, eLpNorm, one_ne_zero, eLpNorm'_eq_lintegral_enorm, Pi.sub_apply, + one_toReal, ENNReal.rpow_one, ne_eq, not_false_eq_true, div_self, ite_false] + simp [edist_eq_enorm_sub] + +theorem dist_def (f g : α →₁[μ] β) : dist f g = (∫⁻ a, edist (f a) (g a) ∂μ).toReal := by + simp_rw [dist_edist, edist_def] + +theorem norm_def (f : α →₁[μ] β) : ‖f‖ = (∫⁻ a, ‖f a‖ₑ ∂μ).toReal := by + simp [Lp.norm_def, eLpNorm, eLpNorm'_eq_lintegral_enorm] + +/-- Computing the norm of a difference between two L¹-functions. Note that this is not a + special case of `norm_def` since `(f - g) x` and `f x - g x` are not equal + (but only a.e.-equal). -/ +theorem norm_sub_eq_lintegral (f g : α →₁[μ] β) : ‖f - g‖ = (∫⁻ x, ‖f x - g x‖ₑ ∂μ).toReal := by + rw [norm_def] + congr 1 + rw [lintegral_congr_ae] + filter_upwards [Lp.coeFn_sub f g] with _ ha + simp only [ha, Pi.sub_apply] + +theorem ofReal_norm_eq_lintegral (f : α →₁[μ] β) : ENNReal.ofReal ‖f‖ = ∫⁻ x, ‖f x‖ₑ ∂μ := by + rw [norm_def, ENNReal.ofReal_toReal] + exact ne_of_lt (hasFiniteIntegral_coeFn f) + +/-- Computing the norm of a difference between two L¹-functions. Note that this is not a + special case of `ofReal_norm_eq_lintegral` since `(f - g) x` and `f x - g x` are not equal + (but only a.e.-equal). -/ +theorem ofReal_norm_sub_eq_lintegral (f g : α →₁[μ] β) : + ENNReal.ofReal ‖f - g‖ = ∫⁻ x, ‖f x - g x‖ₑ ∂μ := by + simp_rw [ofReal_norm_eq_lintegral, ← edist_zero_eq_enorm] + apply lintegral_congr_ae + filter_upwards [Lp.coeFn_sub f g] with _ ha + simp only [ha, Pi.sub_apply] + +end L1 + +namespace Integrable + + +/-- Construct the equivalence class `[f]` of an integrable function `f`, as a member of the +space `Lp β 1 μ`. -/ +def toL1 (f : α → β) (hf : Integrable f μ) : α →₁[μ] β := + (memℒp_one_iff_integrable.2 hf).toLp f + +@[simp] +theorem toL1_coeFn (f : α →₁[μ] β) (hf : Integrable f μ) : hf.toL1 f = f := by + simp [Integrable.toL1] + +theorem coeFn_toL1 {f : α → β} (hf : Integrable f μ) : hf.toL1 f =ᵐ[μ] f := + AEEqFun.coeFn_mk _ _ + +@[simp] +theorem toL1_zero (h : Integrable (0 : α → β) μ) : h.toL1 0 = 0 := + rfl + +@[simp] +theorem toL1_eq_mk (f : α → β) (hf : Integrable f μ) : + (hf.toL1 f : α →ₘ[μ] β) = AEEqFun.mk f hf.aestronglyMeasurable := + rfl + +@[simp] +theorem toL1_eq_toL1_iff (f g : α → β) (hf : Integrable f μ) (hg : Integrable g μ) : + toL1 f hf = toL1 g hg ↔ f =ᵐ[μ] g := + Memℒp.toLp_eq_toLp_iff _ _ + +theorem toL1_add (f g : α → β) (hf : Integrable f μ) (hg : Integrable g μ) : + toL1 (f + g) (hf.add hg) = toL1 f hf + toL1 g hg := + rfl + +theorem toL1_neg (f : α → β) (hf : Integrable f μ) : toL1 (-f) (Integrable.neg hf) = -toL1 f hf := + rfl + +theorem toL1_sub (f g : α → β) (hf : Integrable f μ) (hg : Integrable g μ) : + toL1 (f - g) (hf.sub hg) = toL1 f hf - toL1 g hg := + rfl + +theorem norm_toL1 (f : α → β) (hf : Integrable f μ) : + ‖hf.toL1 f‖ = (∫⁻ a, edist (f a) 0 ∂μ).toReal := by + simp [toL1, Lp.norm_toLp, eLpNorm, eLpNorm'_eq_lintegral_enorm] + +theorem enorm_toL1 {f : α → β} (hf : Integrable f μ) : ‖hf.toL1 f‖ₑ = ∫⁻ a, ‖f a‖ₑ ∂μ := by + simpa [Integrable.toL1, eLpNorm, eLpNorm', enorm] using ENNReal.coe_toNNReal hf.2.ne + +@[deprecated (since := "2025-01-20")] alias nnnorm_toL1 := enorm_toL1 + +theorem norm_toL1_eq_lintegral_norm (f : α → β) (hf : Integrable f μ) : + ‖hf.toL1 f‖ = ENNReal.toReal (∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ) := by + rw [norm_toL1, lintegral_norm_eq_lintegral_edist] + +@[simp] +theorem edist_toL1_toL1 (f g : α → β) (hf : Integrable f μ) (hg : Integrable g μ) : + edist (hf.toL1 f) (hg.toL1 g) = ∫⁻ a, edist (f a) (g a) ∂μ := by + simp only [toL1, Lp.edist_toLp_toLp, eLpNorm, one_ne_zero, eLpNorm'_eq_lintegral_enorm, + Pi.sub_apply, one_toReal, ENNReal.rpow_one, ne_eq, not_false_eq_true, div_self, ite_false] + simp [edist_eq_enorm_sub] + +theorem edist_toL1_zero (f : α → β) (hf : Integrable f μ) : + edist (hf.toL1 f) 0 = ∫⁻ a, edist (f a) 0 ∂μ := by + simp only [edist_zero_right, Lp.enorm_def, toL1_eq_mk, eLpNorm_aeeqFun] + apply eLpNorm_one_eq_lintegral_enorm + +variable {𝕜 : Type*} [NormedRing 𝕜] [Module 𝕜 β] [BoundedSMul 𝕜 β] + +theorem toL1_smul (f : α → β) (hf : Integrable f μ) (k : 𝕜) : + toL1 (fun a => k • f a) (hf.smul k) = k • toL1 f hf := + rfl + +theorem toL1_smul' (f : α → β) (hf : Integrable f μ) (k : 𝕜) : + toL1 (k • f) (hf.smul k) = k • toL1 f hf := + rfl + +end Integrable + +end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/L1Space/HasFiniteIntegral.lean b/Mathlib/MeasureTheory/Function/L1Space/HasFiniteIntegral.lean new file mode 100644 index 00000000000000..8aee3621a62b03 --- /dev/null +++ b/Mathlib/MeasureTheory/Function/L1Space/HasFiniteIntegral.lean @@ -0,0 +1,398 @@ +/- +Copyright (c) 2019 Zhouhang Zhou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Zhouhang Zhou +-/ +import Mathlib.MeasureTheory.Function.StronglyMeasurable.Basic +import Mathlib.MeasureTheory.Measure.WithDensity + +/-! +# Function with finite integral + +In this file we define the predicate `HasFiniteIntegral`, which is then used to define the +predicate `Integrable` in the corresponding file. + +## Main definition + +* Let `f : α → β` be a function, where `α` is a `MeasureSpace` and `β` a `NormedAddCommGroup`. + Then `HasFiniteIntegral f` means `∫⁻ a, ‖f a‖ₑ < ∞`. + +## Tags + +finite integral + +-/ + + +noncomputable section + +open Topology ENNReal MeasureTheory NNReal + +open Set Filter TopologicalSpace ENNReal EMetric MeasureTheory + +variable {α β γ ε : Type*} {m : MeasurableSpace α} {μ ν : Measure α} +variable [NormedAddCommGroup β] [NormedAddCommGroup γ] [ENorm ε] + +namespace MeasureTheory + +/-! ### Some results about the Lebesgue integral involving a normed group -/ + +lemma lintegral_enorm_eq_lintegral_edist (f : α → β) : + ∫⁻ a, ‖f a‖ₑ ∂μ = ∫⁻ a, edist (f a) 0 ∂μ := by simp only [edist_zero_eq_enorm] + +@[deprecated (since := "2025-01-20")] +alias lintegral_nnnorm_eq_lintegral_edist := lintegral_enorm_eq_lintegral_edist + +theorem lintegral_norm_eq_lintegral_edist (f : α → β) : + ∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ = ∫⁻ a, edist (f a) 0 ∂μ := by + simp only [ofReal_norm_eq_enorm, edist_zero_eq_enorm] + +theorem lintegral_edist_triangle {f g h : α → β} (hf : AEStronglyMeasurable f μ) + (hh : AEStronglyMeasurable h μ) : + (∫⁻ a, edist (f a) (g a) ∂μ) ≤ (∫⁻ a, edist (f a) (h a) ∂μ) + ∫⁻ a, edist (g a) (h a) ∂μ := by + rw [← lintegral_add_left' (hf.edist hh)] + refine lintegral_mono fun a => ?_ + apply edist_triangle_right + +-- Yaël: Why do the following four lemmas even exist? +theorem lintegral_enorm_zero : ∫⁻ _ : α, ‖(0 : β)‖ₑ ∂μ = 0 := by simp + +theorem lintegral_enorm_add_left {f : α → β} (hf : AEStronglyMeasurable f μ) (g : α → γ) : + ∫⁻ a, ‖f a‖ₑ + ‖g a‖ₑ ∂μ = ∫⁻ a, ‖f a‖ₑ ∂μ + ∫⁻ a, ‖g a‖ₑ ∂μ := + lintegral_add_left' hf.enorm _ + +theorem lintegral_enorm_add_right (f : α → β) {g : α → γ} (hg : AEStronglyMeasurable g μ) : + ∫⁻ a, ‖f a‖ₑ + ‖g a‖ₑ ∂μ = ∫⁻ a, ‖f a‖ₑ ∂μ + ∫⁻ a, ‖g a‖ₑ ∂μ := + lintegral_add_right' _ hg.enorm + +theorem lintegral_enorm_neg {f : α → β} : ∫⁻ a, ‖(-f) a‖ₑ ∂μ = ∫⁻ a, ‖f a‖ₑ ∂μ := by simp + +@[deprecated (since := "2025-01-21")] alias lintegral_nnnorm_zero := lintegral_enorm_zero +@[deprecated (since := "2025-01-21")] alias lintegral_nnnorm_add_left := lintegral_enorm_add_left +@[deprecated (since := "2025-01-21")] alias lintegral_nnnorm_add_right := lintegral_enorm_add_right +@[deprecated (since := "2025-01-21")] alias lintegral_nnnorm_neg := lintegral_enorm_neg + +/-! ### The predicate `HasFiniteIntegral` -/ + + +/-- `HasFiniteIntegral f μ` means that the integral `∫⁻ a, ‖f a‖ ∂μ` is finite. + `HasFiniteIntegral f` means `HasFiniteIntegral f volume`. -/ +def HasFiniteIntegral {_ : MeasurableSpace α} (f : α → ε) + (μ : Measure α := by volume_tac) : Prop := + ∫⁻ a, ‖f a‖ₑ ∂μ < ∞ + +theorem hasFiniteIntegral_def {_ : MeasurableSpace α} (f : α → ε) (μ : Measure α) : + HasFiniteIntegral f μ ↔ (∫⁻ a, ‖f a‖ₑ ∂μ < ∞) := + Iff.rfl + +theorem hasFiniteIntegral_iff_enorm {f : α → β} : HasFiniteIntegral f μ ↔ ∫⁻ a, ‖f a‖ₑ ∂μ < ∞ := by + simp only [HasFiniteIntegral, ofReal_norm_eq_enorm, enorm_eq_nnnorm] + +@[deprecated (since := "2025-01-20")] +alias hasFiniteIntegral_iff_nnnorm := hasFiniteIntegral_iff_enorm + +theorem hasFiniteIntegral_iff_norm (f : α → β) : + HasFiniteIntegral f μ ↔ (∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ) < ∞ := by + simp only [hasFiniteIntegral_iff_enorm, ofReal_norm_eq_enorm] + +theorem hasFiniteIntegral_iff_edist (f : α → β) : + HasFiniteIntegral f μ ↔ (∫⁻ a, edist (f a) 0 ∂μ) < ∞ := by + simp only [hasFiniteIntegral_iff_norm, edist_dist, dist_zero_right] + +theorem hasFiniteIntegral_iff_ofReal {f : α → ℝ} (h : 0 ≤ᵐ[μ] f) : + HasFiniteIntegral f μ ↔ (∫⁻ a, ENNReal.ofReal (f a) ∂μ) < ∞ := by + rw [hasFiniteIntegral_iff_enorm, lintegral_enorm_of_ae_nonneg h] + +theorem hasFiniteIntegral_iff_ofNNReal {f : α → ℝ≥0} : + HasFiniteIntegral (fun x => (f x : ℝ)) μ ↔ (∫⁻ a, f a ∂μ) < ∞ := by + simp [hasFiniteIntegral_iff_norm] + +theorem HasFiniteIntegral.mono {f : α → β} {g : α → γ} (hg : HasFiniteIntegral g μ) + (h : ∀ᵐ a ∂μ, ‖f a‖ ≤ ‖g a‖) : HasFiniteIntegral f μ := by + simp only [hasFiniteIntegral_iff_norm] at * + calc + (∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ) ≤ ∫⁻ a : α, ENNReal.ofReal ‖g a‖ ∂μ := + lintegral_mono_ae (h.mono fun a h => ofReal_le_ofReal h) + _ < ∞ := hg + +theorem HasFiniteIntegral.mono' {f : α → β} {g : α → ℝ} (hg : HasFiniteIntegral g μ) + (h : ∀ᵐ a ∂μ, ‖f a‖ ≤ g a) : HasFiniteIntegral f μ := + hg.mono <| h.mono fun _x hx => le_trans hx (le_abs_self _) + +theorem HasFiniteIntegral.congr' {f : α → β} {g : α → γ} (hf : HasFiniteIntegral f μ) + (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : HasFiniteIntegral g μ := + hf.mono <| EventuallyEq.le <| EventuallyEq.symm h + +theorem hasFiniteIntegral_congr' {f : α → β} {g : α → γ} (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : + HasFiniteIntegral f μ ↔ HasFiniteIntegral g μ := + ⟨fun hf => hf.congr' h, fun hg => hg.congr' <| EventuallyEq.symm h⟩ + +theorem HasFiniteIntegral.congr {f g : α → β} (hf : HasFiniteIntegral f μ) (h : f =ᵐ[μ] g) : + HasFiniteIntegral g μ := + hf.congr' <| h.fun_comp norm + +theorem hasFiniteIntegral_congr {f g : α → β} (h : f =ᵐ[μ] g) : + HasFiniteIntegral f μ ↔ HasFiniteIntegral g μ := + hasFiniteIntegral_congr' <| h.fun_comp norm + +theorem hasFiniteIntegral_const_iff {c : β} : + HasFiniteIntegral (fun _ : α => c) μ ↔ c = 0 ∨ IsFiniteMeasure μ := by + simp [hasFiniteIntegral_iff_enorm, lintegral_const, lt_top_iff_ne_top, ENNReal.mul_eq_top, + or_iff_not_imp_left, isFiniteMeasure_iff] + +lemma hasFiniteIntegral_const_iff_isFiniteMeasure {c : β} (hc : c ≠ 0) : + HasFiniteIntegral (fun _ ↦ c) μ ↔ IsFiniteMeasure μ := by + simp [hasFiniteIntegral_const_iff, hc, isFiniteMeasure_iff] + +theorem hasFiniteIntegral_const [IsFiniteMeasure μ] (c : β) : + HasFiniteIntegral (fun _ : α => c) μ := + hasFiniteIntegral_const_iff.2 <| .inr ‹_› + +theorem HasFiniteIntegral.of_mem_Icc [IsFiniteMeasure μ] (a b : ℝ) {X : α → ℝ} + (h : ∀ᵐ ω ∂μ, X ω ∈ Set.Icc a b) : + HasFiniteIntegral X μ := by + apply (hasFiniteIntegral_const (max ‖a‖ ‖b‖)).mono' + filter_upwards [h.mono fun ω h ↦ h.1, h.mono fun ω h ↦ h.2] with ω using abs_le_max_abs_abs + +theorem hasFiniteIntegral_of_bounded [IsFiniteMeasure μ] {f : α → β} {C : ℝ} + (hC : ∀ᵐ a ∂μ, ‖f a‖ ≤ C) : HasFiniteIntegral f μ := + (hasFiniteIntegral_const C).mono' hC + +theorem HasFiniteIntegral.of_finite [Finite α] [IsFiniteMeasure μ] {f : α → β} : + HasFiniteIntegral f μ := + let ⟨_⟩ := nonempty_fintype α + hasFiniteIntegral_of_bounded <| ae_of_all μ <| norm_le_pi_norm f + +theorem HasFiniteIntegral.mono_measure {f : α → β} (h : HasFiniteIntegral f ν) (hμ : μ ≤ ν) : + HasFiniteIntegral f μ := + lt_of_le_of_lt (lintegral_mono' hμ le_rfl) h + +theorem HasFiniteIntegral.add_measure {f : α → β} (hμ : HasFiniteIntegral f μ) + (hν : HasFiniteIntegral f ν) : HasFiniteIntegral f (μ + ν) := by + simp only [HasFiniteIntegral, lintegral_add_measure] at * + exact add_lt_top.2 ⟨hμ, hν⟩ + +theorem HasFiniteIntegral.left_of_add_measure {f : α → β} (h : HasFiniteIntegral f (μ + ν)) : + HasFiniteIntegral f μ := + h.mono_measure <| Measure.le_add_right <| le_rfl + +theorem HasFiniteIntegral.right_of_add_measure {f : α → β} (h : HasFiniteIntegral f (μ + ν)) : + HasFiniteIntegral f ν := + h.mono_measure <| Measure.le_add_left <| le_rfl + +@[simp] +theorem hasFiniteIntegral_add_measure {f : α → β} : + HasFiniteIntegral f (μ + ν) ↔ HasFiniteIntegral f μ ∧ HasFiniteIntegral f ν := + ⟨fun h => ⟨h.left_of_add_measure, h.right_of_add_measure⟩, fun h => h.1.add_measure h.2⟩ + +theorem HasFiniteIntegral.smul_measure {f : α → β} (h : HasFiniteIntegral f μ) {c : ℝ≥0∞} + (hc : c ≠ ∞) : HasFiniteIntegral f (c • μ) := by + simp only [HasFiniteIntegral, lintegral_smul_measure] at * + exact mul_lt_top hc.lt_top h + +@[simp] +theorem hasFiniteIntegral_zero_measure {m : MeasurableSpace α} (f : α → β) : + HasFiniteIntegral f (0 : Measure α) := by + simp only [HasFiniteIntegral, lintegral_zero_measure, zero_lt_top] + +variable (α β μ) + +@[simp] +theorem hasFiniteIntegral_zero : HasFiniteIntegral (fun _ : α => (0 : β)) μ := by + simp [hasFiniteIntegral_iff_enorm] + +variable {α β μ} + +theorem HasFiniteIntegral.neg {f : α → β} (hfi : HasFiniteIntegral f μ) : + HasFiniteIntegral (-f) μ := by simpa [hasFiniteIntegral_iff_enorm] using hfi + +@[simp] +theorem hasFiniteIntegral_neg_iff {f : α → β} : HasFiniteIntegral (-f) μ ↔ HasFiniteIntegral f μ := + ⟨fun h => neg_neg f ▸ h.neg, HasFiniteIntegral.neg⟩ + +theorem HasFiniteIntegral.norm {f : α → β} (hfi : HasFiniteIntegral f μ) : + HasFiniteIntegral (fun a => ‖f a‖) μ := by simpa [hasFiniteIntegral_iff_enorm] using hfi + +theorem hasFiniteIntegral_norm_iff (f : α → β) : + HasFiniteIntegral (fun a => ‖f a‖) μ ↔ HasFiniteIntegral f μ := + hasFiniteIntegral_congr' <| Eventually.of_forall fun x => norm_norm (f x) + +theorem hasFiniteIntegral_toReal_of_lintegral_ne_top {f : α → ℝ≥0∞} (hf : ∫⁻ x, f x ∂μ ≠ ∞) : + HasFiniteIntegral (fun x ↦ (f x).toReal) μ := by + have h x : ‖(f x).toReal‖ₑ = .ofReal (f x).toReal := by + rw [Real.enorm_of_nonneg ENNReal.toReal_nonneg] + simp_rw [hasFiniteIntegral_iff_enorm, h] + refine lt_of_le_of_lt (lintegral_mono fun x => ?_) (lt_top_iff_ne_top.2 hf) + by_cases hfx : f x = ∞ + · simp [hfx] + · lift f x to ℝ≥0 using hfx with fx h + simp [← h, ← NNReal.coe_le_coe] + +lemma hasFiniteIntegral_toReal_iff {f : α → ℝ≥0∞} (hf : ∀ᵐ x ∂μ, f x ≠ ∞) : + HasFiniteIntegral (fun x ↦ (f x).toReal) μ ↔ ∫⁻ x, f x ∂μ ≠ ∞ := by + have : ∀ᵐ x ∂μ, .ofReal (f x).toReal = f x := by filter_upwards [hf] with x hx; simp [hx] + simp [hasFiniteIntegral_iff_enorm, Real.enorm_of_nonneg ENNReal.toReal_nonneg, + lintegral_congr_ae this, lt_top_iff_ne_top] + +theorem isFiniteMeasure_withDensity_ofReal {f : α → ℝ} (hfi : HasFiniteIntegral f μ) : + IsFiniteMeasure (μ.withDensity fun x => ENNReal.ofReal <| f x) := by + refine isFiniteMeasure_withDensity ((lintegral_mono fun x => ?_).trans_lt hfi).ne + exact Real.ofReal_le_enorm (f x) + +section DominatedConvergence + +variable {F : ℕ → α → β} {f : α → β} {bound : α → ℝ} + +theorem all_ae_ofReal_F_le_bound (h : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) : + ∀ n, ∀ᵐ a ∂μ, ENNReal.ofReal ‖F n a‖ ≤ ENNReal.ofReal (bound a) := fun n => + (h n).mono fun _ h => ENNReal.ofReal_le_ofReal h + +theorem all_ae_tendsto_ofReal_norm (h : ∀ᵐ a ∂μ, Tendsto (fun n => F n a) atTop <| 𝓝 <| f a) : + ∀ᵐ a ∂μ, Tendsto (fun n => ENNReal.ofReal ‖F n a‖) atTop <| 𝓝 <| ENNReal.ofReal ‖f a‖ := + h.mono fun _ h => tendsto_ofReal <| Tendsto.comp (Continuous.tendsto continuous_norm _) h + +theorem all_ae_ofReal_f_le_bound (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) + (h_lim : ∀ᵐ a ∂μ, Tendsto (fun n => F n a) atTop (𝓝 (f a))) : + ∀ᵐ a ∂μ, ENNReal.ofReal ‖f a‖ ≤ ENNReal.ofReal (bound a) := by + have F_le_bound := all_ae_ofReal_F_le_bound h_bound + rw [← ae_all_iff] at F_le_bound + apply F_le_bound.mp ((all_ae_tendsto_ofReal_norm h_lim).mono _) + intro a tendsto_norm F_le_bound + exact le_of_tendsto' tendsto_norm F_le_bound + +theorem hasFiniteIntegral_of_dominated_convergence {F : ℕ → α → β} {f : α → β} {bound : α → ℝ} + (bound_hasFiniteIntegral : HasFiniteIntegral bound μ) + (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) + (h_lim : ∀ᵐ a ∂μ, Tendsto (fun n => F n a) atTop (𝓝 (f a))) : HasFiniteIntegral f μ := by + /- `‖F n a‖ ≤ bound a` and `‖F n a‖ --> ‖f a‖` implies `‖f a‖ ≤ bound a`, + and so `∫ ‖f‖ ≤ ∫ bound < ∞` since `bound` is has_finite_integral -/ + rw [hasFiniteIntegral_iff_norm] + calc + (∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ) ≤ ∫⁻ a, ENNReal.ofReal (bound a) ∂μ := + lintegral_mono_ae <| all_ae_ofReal_f_le_bound h_bound h_lim + _ < ∞ := by + rw [← hasFiniteIntegral_iff_ofReal] + · exact bound_hasFiniteIntegral + exact (h_bound 0).mono fun a h => le_trans (norm_nonneg _) h + +theorem tendsto_lintegral_norm_of_dominated_convergence {F : ℕ → α → β} {f : α → β} {bound : α → ℝ} + (F_measurable : ∀ n, AEStronglyMeasurable (F n) μ) + (bound_hasFiniteIntegral : HasFiniteIntegral bound μ) + (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) + (h_lim : ∀ᵐ a ∂μ, Tendsto (fun n => F n a) atTop (𝓝 (f a))) : + Tendsto (fun n => ∫⁻ a, ENNReal.ofReal ‖F n a - f a‖ ∂μ) atTop (𝓝 0) := by + have f_measurable : AEStronglyMeasurable f μ := + aestronglyMeasurable_of_tendsto_ae _ F_measurable h_lim + let b a := 2 * ENNReal.ofReal (bound a) + /- `‖F n a‖ ≤ bound a` and `F n a --> f a` implies `‖f a‖ ≤ bound a`, and thus by the + triangle inequality, have `‖F n a - f a‖ ≤ 2 * (bound a)`. -/ + have hb : ∀ n, ∀ᵐ a ∂μ, ENNReal.ofReal ‖F n a - f a‖ ≤ b a := by + intro n + filter_upwards [all_ae_ofReal_F_le_bound h_bound n, + all_ae_ofReal_f_le_bound h_bound h_lim] with a h₁ h₂ + calc + ENNReal.ofReal ‖F n a - f a‖ ≤ ENNReal.ofReal ‖F n a‖ + ENNReal.ofReal ‖f a‖ := by + rw [← ENNReal.ofReal_add] + · apply ofReal_le_ofReal + apply norm_sub_le + · exact norm_nonneg _ + · exact norm_nonneg _ + _ ≤ ENNReal.ofReal (bound a) + ENNReal.ofReal (bound a) := add_le_add h₁ h₂ + _ = b a := by rw [← two_mul] + -- On the other hand, `F n a --> f a` implies that `‖F n a - f a‖ --> 0` + have h : ∀ᵐ a ∂μ, Tendsto (fun n => ENNReal.ofReal ‖F n a - f a‖) atTop (𝓝 0) := by + rw [← ENNReal.ofReal_zero] + refine h_lim.mono fun a h => (continuous_ofReal.tendsto _).comp ?_ + rwa [← tendsto_iff_norm_sub_tendsto_zero] + /- Therefore, by the dominated convergence theorem for nonnegative integration, have + ` ∫ ‖f a - F n a‖ --> 0 ` -/ + suffices Tendsto (fun n => ∫⁻ a, ENNReal.ofReal ‖F n a - f a‖ ∂μ) atTop (𝓝 (∫⁻ _ : α, 0 ∂μ)) by + rwa [lintegral_zero] at this + -- Using the dominated convergence theorem. + refine tendsto_lintegral_of_dominated_convergence' _ ?_ hb ?_ ?_ + -- Show `fun a => ‖f a - F n a‖` is almost everywhere measurable for all `n` + · exact fun n => + measurable_ofReal.comp_aemeasurable ((F_measurable n).sub f_measurable).norm.aemeasurable + -- Show `2 * bound` `HasFiniteIntegral` + · rw [hasFiniteIntegral_iff_ofReal] at bound_hasFiniteIntegral + · calc + ∫⁻ a, b a ∂μ = 2 * ∫⁻ a, ENNReal.ofReal (bound a) ∂μ := by + rw [lintegral_const_mul'] + exact coe_ne_top + _ ≠ ∞ := mul_ne_top coe_ne_top bound_hasFiniteIntegral.ne + filter_upwards [h_bound 0] with _ h using le_trans (norm_nonneg _) h + -- Show `‖f a - F n a‖ --> 0` + · exact h + +end DominatedConvergence + +section PosPart + +/-! Lemmas used for defining the positive part of an `L¹` function -/ + + +theorem HasFiniteIntegral.max_zero {f : α → ℝ} (hf : HasFiniteIntegral f μ) : + HasFiniteIntegral (fun a => max (f a) 0) μ := + hf.mono <| Eventually.of_forall fun x => by simp [abs_le, le_abs_self] + +theorem HasFiniteIntegral.min_zero {f : α → ℝ} (hf : HasFiniteIntegral f μ) : + HasFiniteIntegral (fun a => min (f a) 0) μ := + hf.mono <| Eventually.of_forall fun x => by simpa [abs_le] using neg_abs_le _ + +end PosPart + +section NormedSpace + +variable {𝕜 : Type*} + +theorem HasFiniteIntegral.smul [NormedAddCommGroup 𝕜] [SMulZeroClass 𝕜 β] [BoundedSMul 𝕜 β] (c : 𝕜) + {f : α → β} : HasFiniteIntegral f μ → HasFiniteIntegral (c • f) μ := by + simp only [HasFiniteIntegral]; intro hfi + calc + ∫⁻ a : α, ‖c • f a‖ₑ ∂μ ≤ ∫⁻ a : α, ‖c‖ₑ * ‖f a‖ₑ ∂μ := lintegral_mono fun i ↦ enorm_smul_le + _ < ∞ := by + rw [lintegral_const_mul'] + exacts [mul_lt_top coe_lt_top hfi, coe_ne_top] + +theorem hasFiniteIntegral_smul_iff [NormedRing 𝕜] [MulActionWithZero 𝕜 β] [BoundedSMul 𝕜 β] {c : 𝕜} + (hc : IsUnit c) (f : α → β) : HasFiniteIntegral (c • f) μ ↔ HasFiniteIntegral f μ := by + obtain ⟨c, rfl⟩ := hc + constructor + · intro h + simpa only [smul_smul, Units.inv_mul, one_smul] using h.smul ((c⁻¹ : 𝕜ˣ) : 𝕜) + exact HasFiniteIntegral.smul _ + +theorem HasFiniteIntegral.const_mul [NormedRing 𝕜] {f : α → 𝕜} (h : HasFiniteIntegral f μ) (c : 𝕜) : + HasFiniteIntegral (fun x => c * f x) μ := + h.smul c + +theorem HasFiniteIntegral.mul_const [NormedRing 𝕜] {f : α → 𝕜} (h : HasFiniteIntegral f μ) (c : 𝕜) : + HasFiniteIntegral (fun x => f x * c) μ := + h.smul (MulOpposite.op c) + +section count + +variable [MeasurableSingletonClass α] {f : α → β} + +/-- A function has finite integral for the counting measure iff its norm is summable. -/ +lemma hasFiniteIntegral_count_iff : + HasFiniteIntegral f Measure.count ↔ Summable (‖f ·‖) := by + simp only [hasFiniteIntegral_iff_enorm, enorm, lintegral_count, lt_top_iff_ne_top, + tsum_coe_ne_top_iff_summable, ← summable_coe, coe_nnnorm] + +end count + +section restrict + +variable {E : Type*} [NormedAddCommGroup E] {f : α → E} + +lemma HasFiniteIntegral.restrict (h : HasFiniteIntegral f μ) {s : Set α} : + HasFiniteIntegral f (μ.restrict s) := by + refine lt_of_le_of_lt ?_ h + simpa [Measure.restrict_univ] using lintegral_mono_set (subset_univ s) + +end restrict + +end NormedSpace + +end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space/Integrable.lean similarity index 61% rename from Mathlib/MeasureTheory/Function/L1Space.lean rename to Mathlib/MeasureTheory/Function/L1Space/Integrable.lean index 2624f72f454f23..37cfdc0b9cc29f 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space/Integrable.lean @@ -3,35 +3,23 @@ Copyright (c) 2019 Zhouhang Zhou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Zhouhang Zhou -/ +import Mathlib.MeasureTheory.Function.L1Space.HasFiniteIntegral import Mathlib.MeasureTheory.Function.LpOrder import Mathlib.MeasureTheory.Function.StronglyMeasurable.Lemmas /-! -# Integrable functions and `L¹` space +# Integrable functions -In the first part of this file, the predicate `Integrable` is defined and basic properties of +In this file, the predicate `Integrable` is defined and basic properties of integrable functions are proved. Such a predicate is already available under the name `Memℒp 1`. We give a direct definition which -is easier to use, and show that it is equivalent to `Memℒp 1` +is easier to use, and show that it is equivalent to `Memℒp 1`. -In the second part, we establish an API between `Integrable` and the space `L¹` of equivalence -classes of integrable functions, already defined as a special case of `L^p` spaces for `p = 1`. +## Main definition -## Notation - -* `α →₁[μ] β` is the type of `L¹` space, where `α` is a `MeasureSpace` and `β` is a - `NormedAddCommGroup` with a `SecondCountableTopology`. `f : α →ₘ β` is a "function" in `L¹`. - In comments, `[f]` is also used to denote an `L¹` function. - - `₁` can be typed as `\1`. - -## Main definitions - -* Let `f : α → β` be a function, where `α` is a `MeasureSpace` and `β` a `NormedAddCommGroup`. - Then `HasFiniteIntegral f` means `(∫⁻ a, ‖f a‖₊) < ∞`. - -* If `β` is moreover a `MeasurableSpace` then `f` is called `Integrable` if +* Let `f : α → β` be a function, where `α` is a `MeasureSpace` and `β` a `NormedAddCommGroup` + which also a `MeasurableSpace`. Then `f` is called `Integrable` if `f` is `Measurable` and `HasFiniteIntegral f` holds. ## Implementation notes @@ -41,367 +29,20 @@ To prove something for an arbitrary integrable function, a useful theorem is ## Tags -integrable, function space, l1 +integrable -/ noncomputable section -open scoped Classical -open Topology ENNReal MeasureTheory NNReal - -open Set Filter TopologicalSpace ENNReal EMetric MeasureTheory +open EMetric ENNReal Filter MeasureTheory NNReal Set variable {α β γ δ ε : Type*} {m : MeasurableSpace α} {μ ν : Measure α} [MeasurableSpace δ] variable [NormedAddCommGroup β] [NormedAddCommGroup γ] [ENorm ε] namespace MeasureTheory -/-! ### Some results about the Lebesgue integral involving a normed group -/ - - -theorem lintegral_nnnorm_eq_lintegral_edist (f : α → β) : - ∫⁻ a, ‖f a‖₊ ∂μ = ∫⁻ a, edist (f a) 0 ∂μ := by simp only [edist_eq_coe_nnnorm] - -theorem lintegral_norm_eq_lintegral_edist (f : α → β) : - ∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ = ∫⁻ a, edist (f a) 0 ∂μ := by - simp only [ofReal_norm_eq_coe_nnnorm, edist_eq_coe_nnnorm] - -theorem lintegral_edist_triangle {f g h : α → β} (hf : AEStronglyMeasurable f μ) - (hh : AEStronglyMeasurable h μ) : - (∫⁻ a, edist (f a) (g a) ∂μ) ≤ (∫⁻ a, edist (f a) (h a) ∂μ) + ∫⁻ a, edist (g a) (h a) ∂μ := by - rw [← lintegral_add_left' (hf.edist hh)] - refine lintegral_mono fun a => ?_ - apply edist_triangle_right - -theorem lintegral_nnnorm_zero : (∫⁻ _ : α, ‖(0 : β)‖₊ ∂μ) = 0 := by simp - -theorem lintegral_nnnorm_add_left {f : α → β} (hf : AEStronglyMeasurable f μ) (g : α → γ) : - ∫⁻ a, ‖f a‖₊ + ‖g a‖₊ ∂μ = (∫⁻ a, ‖f a‖₊ ∂μ) + ∫⁻ a, ‖g a‖₊ ∂μ := - lintegral_add_left' hf.ennnorm _ - -theorem lintegral_nnnorm_add_right (f : α → β) {g : α → γ} (hg : AEStronglyMeasurable g μ) : - ∫⁻ a, ‖f a‖₊ + ‖g a‖₊ ∂μ = (∫⁻ a, ‖f a‖₊ ∂μ) + ∫⁻ a, ‖g a‖₊ ∂μ := - lintegral_add_right' _ hg.ennnorm - -theorem lintegral_nnnorm_neg {f : α → β} : (∫⁻ a, ‖(-f) a‖₊ ∂μ) = ∫⁻ a, ‖f a‖₊ ∂μ := by - simp only [Pi.neg_apply, nnnorm_neg] - -/-! ### The predicate `HasFiniteIntegral` -/ - - -/-- `HasFiniteIntegral f μ` means that the integral `∫⁻ a, ‖f a‖ ∂μ` is finite. - `HasFiniteIntegral f` means `HasFiniteIntegral f volume`. -/ -def HasFiniteIntegral {_ : MeasurableSpace α} (f : α → ε) - (μ : Measure α := by volume_tac) : Prop := - (∫⁻ a, ‖f a‖ₑ ∂μ) < ∞ - -theorem hasFiniteIntegral_def {_ : MeasurableSpace α} (f : α → ε) (μ : Measure α) : - HasFiniteIntegral f μ ↔ ((∫⁻ a, ‖f a‖ₑ ∂μ) < ∞) := - Iff.rfl - -theorem hasFiniteIntegral_iff_nnnorm {f : α → β} : - HasFiniteIntegral f μ ↔ (∫⁻ a, ‖f a‖₊ ∂μ) < ∞ := by - simp only [HasFiniteIntegral, ofReal_norm_eq_coe_nnnorm, enorm_eq_nnnorm] - -theorem hasFiniteIntegral_iff_norm (f : α → β) : - HasFiniteIntegral f μ ↔ (∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ) < ∞ := by - simp only [hasFiniteIntegral_iff_nnnorm, ofReal_norm_eq_coe_nnnorm] - -theorem hasFiniteIntegral_iff_edist (f : α → β) : - HasFiniteIntegral f μ ↔ (∫⁻ a, edist (f a) 0 ∂μ) < ∞ := by - simp only [hasFiniteIntegral_iff_norm, edist_dist, dist_zero_right] - -theorem hasFiniteIntegral_iff_ofReal {f : α → ℝ} (h : 0 ≤ᵐ[μ] f) : - HasFiniteIntegral f μ ↔ (∫⁻ a, ENNReal.ofReal (f a) ∂μ) < ∞ := by - rw [hasFiniteIntegral_iff_nnnorm, lintegral_nnnorm_eq_of_ae_nonneg h] - -theorem hasFiniteIntegral_iff_ofNNReal {f : α → ℝ≥0} : - HasFiniteIntegral (fun x => (f x : ℝ)) μ ↔ (∫⁻ a, f a ∂μ) < ∞ := by - simp [hasFiniteIntegral_iff_norm] - -theorem HasFiniteIntegral.mono {f : α → β} {g : α → γ} (hg : HasFiniteIntegral g μ) - (h : ∀ᵐ a ∂μ, ‖f a‖ ≤ ‖g a‖) : HasFiniteIntegral f μ := by - simp only [hasFiniteIntegral_iff_norm] at * - calc - (∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ) ≤ ∫⁻ a : α, ENNReal.ofReal ‖g a‖ ∂μ := - lintegral_mono_ae (h.mono fun a h => ofReal_le_ofReal h) - _ < ∞ := hg - -theorem HasFiniteIntegral.mono' {f : α → β} {g : α → ℝ} (hg : HasFiniteIntegral g μ) - (h : ∀ᵐ a ∂μ, ‖f a‖ ≤ g a) : HasFiniteIntegral f μ := - hg.mono <| h.mono fun _x hx => le_trans hx (le_abs_self _) - -theorem HasFiniteIntegral.congr' {f : α → β} {g : α → γ} (hf : HasFiniteIntegral f μ) - (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : HasFiniteIntegral g μ := - hf.mono <| EventuallyEq.le <| EventuallyEq.symm h - -theorem hasFiniteIntegral_congr' {f : α → β} {g : α → γ} (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : - HasFiniteIntegral f μ ↔ HasFiniteIntegral g μ := - ⟨fun hf => hf.congr' h, fun hg => hg.congr' <| EventuallyEq.symm h⟩ - -theorem HasFiniteIntegral.congr {f g : α → β} (hf : HasFiniteIntegral f μ) (h : f =ᵐ[μ] g) : - HasFiniteIntegral g μ := - hf.congr' <| h.fun_comp norm - -theorem hasFiniteIntegral_congr {f g : α → β} (h : f =ᵐ[μ] g) : - HasFiniteIntegral f μ ↔ HasFiniteIntegral g μ := - hasFiniteIntegral_congr' <| h.fun_comp norm - -theorem hasFiniteIntegral_const_iff {c : β} : - HasFiniteIntegral (fun _ : α => c) μ ↔ c = 0 ∨ IsFiniteMeasure μ := by - simp [hasFiniteIntegral_iff_nnnorm, lintegral_const, lt_top_iff_ne_top, ENNReal.mul_eq_top, - or_iff_not_imp_left, isFiniteMeasure_iff] - -lemma hasFiniteIntegral_const_iff_isFiniteMeasure {c : β} (hc : c ≠ 0) : - HasFiniteIntegral (fun _ ↦ c) μ ↔ IsFiniteMeasure μ := by - simp [hasFiniteIntegral_const_iff, hc, isFiniteMeasure_iff] - -theorem hasFiniteIntegral_const [IsFiniteMeasure μ] (c : β) : - HasFiniteIntegral (fun _ : α => c) μ := - hasFiniteIntegral_const_iff.2 <| .inr ‹_› - -theorem HasFiniteIntegral.of_mem_Icc [IsFiniteMeasure μ] (a b : ℝ) {X : α → ℝ} - (h : ∀ᵐ ω ∂μ, X ω ∈ Set.Icc a b) : - HasFiniteIntegral X μ := by - apply (hasFiniteIntegral_const (max ‖a‖ ‖b‖)).mono' - filter_upwards [h.mono fun ω h ↦ h.1, h.mono fun ω h ↦ h.2] with ω using abs_le_max_abs_abs - -theorem hasFiniteIntegral_of_bounded [IsFiniteMeasure μ] {f : α → β} {C : ℝ} - (hC : ∀ᵐ a ∂μ, ‖f a‖ ≤ C) : HasFiniteIntegral f μ := - (hasFiniteIntegral_const C).mono' hC - -theorem HasFiniteIntegral.of_finite [Finite α] [IsFiniteMeasure μ] {f : α → β} : - HasFiniteIntegral f μ := - let ⟨_⟩ := nonempty_fintype α - hasFiniteIntegral_of_bounded <| ae_of_all μ <| norm_le_pi_norm f - -theorem HasFiniteIntegral.mono_measure {f : α → β} (h : HasFiniteIntegral f ν) (hμ : μ ≤ ν) : - HasFiniteIntegral f μ := - lt_of_le_of_lt (lintegral_mono' hμ le_rfl) h - -theorem HasFiniteIntegral.add_measure {f : α → β} (hμ : HasFiniteIntegral f μ) - (hν : HasFiniteIntegral f ν) : HasFiniteIntegral f (μ + ν) := by - simp only [HasFiniteIntegral, lintegral_add_measure] at * - exact add_lt_top.2 ⟨hμ, hν⟩ - -theorem HasFiniteIntegral.left_of_add_measure {f : α → β} (h : HasFiniteIntegral f (μ + ν)) : - HasFiniteIntegral f μ := - h.mono_measure <| Measure.le_add_right <| le_rfl - -theorem HasFiniteIntegral.right_of_add_measure {f : α → β} (h : HasFiniteIntegral f (μ + ν)) : - HasFiniteIntegral f ν := - h.mono_measure <| Measure.le_add_left <| le_rfl - -@[simp] -theorem hasFiniteIntegral_add_measure {f : α → β} : - HasFiniteIntegral f (μ + ν) ↔ HasFiniteIntegral f μ ∧ HasFiniteIntegral f ν := - ⟨fun h => ⟨h.left_of_add_measure, h.right_of_add_measure⟩, fun h => h.1.add_measure h.2⟩ - -theorem HasFiniteIntegral.smul_measure {f : α → β} (h : HasFiniteIntegral f μ) {c : ℝ≥0∞} - (hc : c ≠ ∞) : HasFiniteIntegral f (c • μ) := by - simp only [HasFiniteIntegral, lintegral_smul_measure] at * - exact mul_lt_top hc.lt_top h - -@[simp] -theorem hasFiniteIntegral_zero_measure {m : MeasurableSpace α} (f : α → β) : - HasFiniteIntegral f (0 : Measure α) := by - simp only [HasFiniteIntegral, lintegral_zero_measure, zero_lt_top] - -variable (α β μ) - -@[simp] -theorem hasFiniteIntegral_zero : HasFiniteIntegral (fun _ : α => (0 : β)) μ := by - simp [hasFiniteIntegral_iff_nnnorm] - -variable {α β μ} - -theorem HasFiniteIntegral.neg {f : α → β} (hfi : HasFiniteIntegral f μ) : - HasFiniteIntegral (-f) μ := by simpa [hasFiniteIntegral_iff_nnnorm] using hfi - -@[simp] -theorem hasFiniteIntegral_neg_iff {f : α → β} : HasFiniteIntegral (-f) μ ↔ HasFiniteIntegral f μ := - ⟨fun h => neg_neg f ▸ h.neg, HasFiniteIntegral.neg⟩ - -theorem HasFiniteIntegral.norm {f : α → β} (hfi : HasFiniteIntegral f μ) : - HasFiniteIntegral (fun a => ‖f a‖) μ := by - have eq : (fun a => (nnnorm ‖f a‖ : ℝ≥0∞)) = fun a => (‖f a‖₊ : ℝ≥0∞) := by - funext - rw [nnnorm_norm] - rwa [hasFiniteIntegral_iff_nnnorm, eq] - -theorem hasFiniteIntegral_norm_iff (f : α → β) : - HasFiniteIntegral (fun a => ‖f a‖) μ ↔ HasFiniteIntegral f μ := - hasFiniteIntegral_congr' <| Eventually.of_forall fun x => norm_norm (f x) - -theorem hasFiniteIntegral_toReal_of_lintegral_ne_top {f : α → ℝ≥0∞} (hf : ∫⁻ x, f x ∂μ ≠ ∞) : - HasFiniteIntegral (fun x ↦ (f x).toReal) μ := by - have h x : (‖(f x).toReal‖₊ : ℝ≥0∞) = ENNReal.ofNNReal ⟨(f x).toReal, ENNReal.toReal_nonneg⟩ := by - rw [Real.nnnorm_of_nonneg] - simp_rw [hasFiniteIntegral_iff_nnnorm, h] - refine lt_of_le_of_lt (lintegral_mono fun x => ?_) (lt_top_iff_ne_top.2 hf) - by_cases hfx : f x = ∞ - · simp [hfx] - · lift f x to ℝ≥0 using hfx with fx h - simp [← h, ← NNReal.coe_le_coe] - -lemma hasFiniteIntegral_toReal_iff {f : α → ℝ≥0∞} (hf_ne_top : ∀ᵐ x ∂μ, f x ≠ ∞) : - HasFiniteIntegral (fun x ↦ (f x).toReal) μ ↔ ∫⁻ x, f x ∂μ ≠ ∞ := by - have h_eq x : (‖(f x).toReal‖₊ : ℝ≥0∞) - = ENNReal.ofNNReal ⟨(f x).toReal, ENNReal.toReal_nonneg⟩ := by - rw [Real.nnnorm_of_nonneg] - simp_rw [hasFiniteIntegral_iff_nnnorm, h_eq, lt_top_iff_ne_top] - convert Iff.rfl using 2 - refine lintegral_congr_ae ?_ - filter_upwards [hf_ne_top] with x hfx - lift f x to ℝ≥0 using hfx with fx h - simp only [coe_toReal, ENNReal.coe_inj] - rfl - -theorem isFiniteMeasure_withDensity_ofReal {f : α → ℝ} (hfi : HasFiniteIntegral f μ) : - IsFiniteMeasure (μ.withDensity fun x => ENNReal.ofReal <| f x) := by - refine isFiniteMeasure_withDensity ((lintegral_mono fun x => ?_).trans_lt hfi).ne - exact Real.ofReal_le_ennnorm (f x) - -section DominatedConvergence - -variable {F : ℕ → α → β} {f : α → β} {bound : α → ℝ} - -theorem all_ae_ofReal_F_le_bound (h : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) : - ∀ n, ∀ᵐ a ∂μ, ENNReal.ofReal ‖F n a‖ ≤ ENNReal.ofReal (bound a) := fun n => - (h n).mono fun _ h => ENNReal.ofReal_le_ofReal h - -theorem all_ae_tendsto_ofReal_norm (h : ∀ᵐ a ∂μ, Tendsto (fun n => F n a) atTop <| 𝓝 <| f a) : - ∀ᵐ a ∂μ, Tendsto (fun n => ENNReal.ofReal ‖F n a‖) atTop <| 𝓝 <| ENNReal.ofReal ‖f a‖ := - h.mono fun _ h => tendsto_ofReal <| Tendsto.comp (Continuous.tendsto continuous_norm _) h - -theorem all_ae_ofReal_f_le_bound (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) - (h_lim : ∀ᵐ a ∂μ, Tendsto (fun n => F n a) atTop (𝓝 (f a))) : - ∀ᵐ a ∂μ, ENNReal.ofReal ‖f a‖ ≤ ENNReal.ofReal (bound a) := by - have F_le_bound := all_ae_ofReal_F_le_bound h_bound - rw [← ae_all_iff] at F_le_bound - apply F_le_bound.mp ((all_ae_tendsto_ofReal_norm h_lim).mono _) - intro a tendsto_norm F_le_bound - exact le_of_tendsto' tendsto_norm F_le_bound - -theorem hasFiniteIntegral_of_dominated_convergence {F : ℕ → α → β} {f : α → β} {bound : α → ℝ} - (bound_hasFiniteIntegral : HasFiniteIntegral bound μ) - (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) - (h_lim : ∀ᵐ a ∂μ, Tendsto (fun n => F n a) atTop (𝓝 (f a))) : HasFiniteIntegral f μ := by - /- `‖F n a‖ ≤ bound a` and `‖F n a‖ --> ‖f a‖` implies `‖f a‖ ≤ bound a`, - and so `∫ ‖f‖ ≤ ∫ bound < ∞` since `bound` is has_finite_integral -/ - rw [hasFiniteIntegral_iff_norm] - calc - (∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ) ≤ ∫⁻ a, ENNReal.ofReal (bound a) ∂μ := - lintegral_mono_ae <| all_ae_ofReal_f_le_bound h_bound h_lim - _ < ∞ := by - rw [← hasFiniteIntegral_iff_ofReal] - · exact bound_hasFiniteIntegral - exact (h_bound 0).mono fun a h => le_trans (norm_nonneg _) h - -theorem tendsto_lintegral_norm_of_dominated_convergence {F : ℕ → α → β} {f : α → β} {bound : α → ℝ} - (F_measurable : ∀ n, AEStronglyMeasurable (F n) μ) - (bound_hasFiniteIntegral : HasFiniteIntegral bound μ) - (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) - (h_lim : ∀ᵐ a ∂μ, Tendsto (fun n => F n a) atTop (𝓝 (f a))) : - Tendsto (fun n => ∫⁻ a, ENNReal.ofReal ‖F n a - f a‖ ∂μ) atTop (𝓝 0) := by - have f_measurable : AEStronglyMeasurable f μ := - aestronglyMeasurable_of_tendsto_ae _ F_measurable h_lim - let b a := 2 * ENNReal.ofReal (bound a) - /- `‖F n a‖ ≤ bound a` and `F n a --> f a` implies `‖f a‖ ≤ bound a`, and thus by the - triangle inequality, have `‖F n a - f a‖ ≤ 2 * (bound a)`. -/ - have hb : ∀ n, ∀ᵐ a ∂μ, ENNReal.ofReal ‖F n a - f a‖ ≤ b a := by - intro n - filter_upwards [all_ae_ofReal_F_le_bound h_bound n, - all_ae_ofReal_f_le_bound h_bound h_lim] with a h₁ h₂ - calc - ENNReal.ofReal ‖F n a - f a‖ ≤ ENNReal.ofReal ‖F n a‖ + ENNReal.ofReal ‖f a‖ := by - rw [← ENNReal.ofReal_add] - · apply ofReal_le_ofReal - apply norm_sub_le - · exact norm_nonneg _ - · exact norm_nonneg _ - _ ≤ ENNReal.ofReal (bound a) + ENNReal.ofReal (bound a) := add_le_add h₁ h₂ - _ = b a := by rw [← two_mul] - -- On the other hand, `F n a --> f a` implies that `‖F n a - f a‖ --> 0` - have h : ∀ᵐ a ∂μ, Tendsto (fun n => ENNReal.ofReal ‖F n a - f a‖) atTop (𝓝 0) := by - rw [← ENNReal.ofReal_zero] - refine h_lim.mono fun a h => (continuous_ofReal.tendsto _).comp ?_ - rwa [← tendsto_iff_norm_sub_tendsto_zero] - /- Therefore, by the dominated convergence theorem for nonnegative integration, have - ` ∫ ‖f a - F n a‖ --> 0 ` -/ - suffices Tendsto (fun n => ∫⁻ a, ENNReal.ofReal ‖F n a - f a‖ ∂μ) atTop (𝓝 (∫⁻ _ : α, 0 ∂μ)) by - rwa [lintegral_zero] at this - -- Using the dominated convergence theorem. - refine tendsto_lintegral_of_dominated_convergence' _ ?_ hb ?_ ?_ - -- Show `fun a => ‖f a - F n a‖` is almost everywhere measurable for all `n` - · exact fun n => - measurable_ofReal.comp_aemeasurable ((F_measurable n).sub f_measurable).norm.aemeasurable - -- Show `2 * bound` `HasFiniteIntegral` - · rw [hasFiniteIntegral_iff_ofReal] at bound_hasFiniteIntegral - · calc - ∫⁻ a, b a ∂μ = 2 * ∫⁻ a, ENNReal.ofReal (bound a) ∂μ := by - rw [lintegral_const_mul'] - exact coe_ne_top - _ ≠ ∞ := mul_ne_top coe_ne_top bound_hasFiniteIntegral.ne - filter_upwards [h_bound 0] with _ h using le_trans (norm_nonneg _) h - -- Show `‖f a - F n a‖ --> 0` - · exact h - -end DominatedConvergence - -section PosPart - -/-! Lemmas used for defining the positive part of an `L¹` function -/ - - -theorem HasFiniteIntegral.max_zero {f : α → ℝ} (hf : HasFiniteIntegral f μ) : - HasFiniteIntegral (fun a => max (f a) 0) μ := - hf.mono <| Eventually.of_forall fun x => by simp [abs_le, le_abs_self] - -theorem HasFiniteIntegral.min_zero {f : α → ℝ} (hf : HasFiniteIntegral f μ) : - HasFiniteIntegral (fun a => min (f a) 0) μ := - hf.mono <| Eventually.of_forall fun x => by simpa [abs_le] using neg_abs_le _ - -end PosPart - -section NormedSpace - -variable {𝕜 : Type*} - -theorem HasFiniteIntegral.smul [NormedAddCommGroup 𝕜] [SMulZeroClass 𝕜 β] [BoundedSMul 𝕜 β] (c : 𝕜) - {f : α → β} : HasFiniteIntegral f μ → HasFiniteIntegral (c • f) μ := by - simp only [HasFiniteIntegral]; intro hfi - calc - (∫⁻ a : α, ‖c • f a‖₊ ∂μ) ≤ ∫⁻ a : α, ‖c‖₊ * ‖f a‖₊ ∂μ := by - refine lintegral_mono ?_ - intro i - -- After https://github.com/leanprover/lean4/pull/2734, we need to do beta reduction `exact mod_cast` - beta_reduce - exact mod_cast (nnnorm_smul_le c (f i)) - _ < ∞ := by - rw [lintegral_const_mul'] - exacts [mul_lt_top coe_lt_top hfi, coe_ne_top] - -theorem hasFiniteIntegral_smul_iff [NormedRing 𝕜] [MulActionWithZero 𝕜 β] [BoundedSMul 𝕜 β] {c : 𝕜} - (hc : IsUnit c) (f : α → β) : HasFiniteIntegral (c • f) μ ↔ HasFiniteIntegral f μ := by - obtain ⟨c, rfl⟩ := hc - constructor - · intro h - simpa only [smul_smul, Units.inv_mul, one_smul] using h.smul ((c⁻¹ : 𝕜ˣ) : 𝕜) - exact HasFiniteIntegral.smul _ - -theorem HasFiniteIntegral.const_mul [NormedRing 𝕜] {f : α → 𝕜} (h : HasFiniteIntegral f μ) (c : 𝕜) : - HasFiniteIntegral (fun x => c * f x) μ := - h.smul c - -theorem HasFiniteIntegral.mul_const [NormedRing 𝕜] {f : α → 𝕜} (h : HasFiniteIntegral f μ) (c : 𝕜) : - HasFiniteIntegral (fun x => f x * c) μ := - h.smul (MulOpposite.op c) - -end NormedSpace - /-! ### The predicate `Integrable` -/ variable [TopologicalSpace ε] @@ -418,7 +59,7 @@ def Integrable {α} {_ : MeasurableSpace α} (f : α → ε) scoped notation "Integrable[" mα "]" => @Integrable _ _ _ _ mα theorem memℒp_one_iff_integrable {f : α → β} : Memℒp f 1 μ ↔ Integrable f μ := by - simp_rw [Integrable, hasFiniteIntegral_iff_nnnorm, Memℒp, eLpNorm_one_eq_lintegral_nnnorm] + simp_rw [Integrable, hasFiniteIntegral_iff_enorm, Memℒp, eLpNorm_one_eq_lintegral_enorm] theorem Integrable.aestronglyMeasurable {f : α → β} (hf : Integrable f μ) : AEStronglyMeasurable f μ := @@ -480,7 +121,6 @@ lemma Integrable.of_finite [Finite α] [MeasurableSingletonClass α] [IsFiniteMe @[deprecated Integrable.of_finite (since := "2024-10-05"), nolint deprecatedNoSince] lemma Integrable.of_isEmpty [IsEmpty α] {f : α → β} : Integrable f μ := .of_finite - theorem Memℒp.integrable_norm_rpow {f : α → β} {p : ℝ≥0∞} (hf : Memℒp f p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : Integrable (fun x : α => ‖f x‖ ^ p.toReal) μ := by rw [← memℒp_one_iff_integrable] @@ -537,6 +177,7 @@ theorem integrable_zero_measure {_ : MeasurableSpace α} {f : α → β} : theorem integrable_finset_sum_measure {ι} {m : MeasurableSpace α} {f : α → β} {μ : ι → Measure α} {s : Finset ι} : Integrable f (∑ i ∈ s, μ i) ↔ ∀ i ∈ s, Integrable f (μ i) := by + classical induction s using Finset.induction_on <;> simp [*] theorem Integrable.smul_measure {f : α → β} (h : Integrable f μ) {c : ℝ≥0∞} (hc : c ≠ ∞) : @@ -567,6 +208,7 @@ theorem Integrable.to_average {f : α → β} (h : Integrable f μ) : Integrable · apply h.smul_measure simpa +open scoped Classical in theorem integrable_average [IsFiniteMeasure μ] {f : α → β} : Integrable f ((μ univ)⁻¹ • μ) ↔ Integrable f μ := (eq_or_ne μ 0).by_cases (fun h => by simp [h]) fun h => @@ -625,12 +267,8 @@ variable {α β μ} theorem Integrable.add' {f g : α → β} (hf : Integrable f μ) (hg : Integrable g μ) : HasFiniteIntegral (f + g) μ := calc - (∫⁻ a, ‖f a + g a‖₊ ∂μ) ≤ ∫⁻ a, ‖f a‖₊ + ‖g a‖₊ ∂μ := - lintegral_mono fun a => by - -- After https://github.com/leanprover/lean4/pull/2734, we need to do beta reduction before `exact mod_cast` - beta_reduce - exact mod_cast nnnorm_add_le _ _ - _ = _ := lintegral_nnnorm_add_left hf.aestronglyMeasurable _ + ∫⁻ a, ‖f a + g a‖ₑ ∂μ ≤ ∫⁻ a, ‖f a‖ₑ + ‖g a‖ₑ ∂μ := lintegral_mono fun _ ↦ enorm_add_le _ _ + _ = _ := lintegral_enorm_add_left hf.aestronglyMeasurable _ _ < ∞ := add_lt_top.2 ⟨hf.hasFiniteIntegral, hg.hasFiniteIntegral⟩ @[fun_prop] @@ -789,7 +427,7 @@ theorem Integrable.bdd_mul {F : Type*} [NormedDivisionRing F] {f g : α → F} ( vector-valued function by a scalar function with finite essential supremum is integrable. -/ theorem Integrable.essSup_smul {𝕜 : Type*} [NormedField 𝕜] [NormedSpace 𝕜 β] {f : α → β} (hf : Integrable f μ) {g : α → 𝕜} (g_aestronglyMeasurable : AEStronglyMeasurable g μ) - (ess_sup_g : essSup (fun x => (‖g x‖₊ : ℝ≥0∞)) μ ≠ ∞) : + (ess_sup_g : essSup (‖g ·‖ₑ) μ ≠ ∞) : Integrable (fun x : α => g x • f x) μ := by rw [← memℒp_one_iff_integrable] at * refine ⟨g_aestronglyMeasurable.smul hf.1, ?_⟩ @@ -804,8 +442,7 @@ theorem Integrable.essSup_smul {𝕜 : Type*} [NormedField 𝕜] [NormedSpace scalar-valued function by a vector-value function with finite essential supremum is integrable. -/ theorem Integrable.smul_essSup {𝕜 : Type*} [NormedRing 𝕜] [Module 𝕜 β] [BoundedSMul 𝕜 β] {f : α → 𝕜} (hf : Integrable f μ) {g : α → β} - (g_aestronglyMeasurable : AEStronglyMeasurable g μ) - (ess_sup_g : essSup (fun x => (‖g x‖₊ : ℝ≥0∞)) μ ≠ ∞) : + (g_aestronglyMeasurable : AEStronglyMeasurable g μ) (ess_sup_g : essSup (‖g ·‖ₑ) μ ≠ ∞) : Integrable (fun x : α => f x • g x) μ := by rw [← memℒp_one_iff_integrable] at * refine ⟨hf.1.smul g_aestronglyMeasurable, ?_⟩ @@ -864,8 +501,8 @@ theorem Memℒp.integrable {q : ℝ≥0∞} (hq1 : 1 ≤ q) {f : α → β} [IsF where `‖f x‖ ≥ ε` is finite for all positive `ε`. -/ theorem Integrable.measure_norm_ge_lt_top {f : α → β} (hf : Integrable f μ) {ε : ℝ} (hε : 0 < ε) : μ { x | ε ≤ ‖f x‖ } < ∞ := by - rw [show { x | ε ≤ ‖f x‖ } = { x | ENNReal.ofReal ε ≤ ‖f x‖₊ } by - simp only [ENNReal.ofReal, Real.toNNReal_le_iff_le_coe, ENNReal.coe_le_coe, coe_nnnorm]] + rw [show { x | ε ≤ ‖f x‖ } = { x | .ofReal ε ≤ ‖f x‖ₑ } by + simp [ENNReal.ofReal, Real.toNNReal_le_iff_le_coe]] refine (meas_ge_le_mul_pow_eLpNorm μ one_ne_zero ENNReal.one_ne_top hf.1 ?_).trans_lt ?_ · simpa only [Ne, ENNReal.ofReal_eq_zero, not_le] using hε apply ENNReal.mul_lt_top @@ -944,12 +581,6 @@ section count variable [MeasurableSingletonClass α] {f : α → β} -/-- A function has finite integral for the counting measure iff its norm is summable. -/ -lemma hasFiniteIntegral_count_iff : - HasFiniteIntegral f Measure.count ↔ Summable (‖f ·‖) := by - simp only [hasFiniteIntegral_iff_nnnorm, lintegral_count, lt_top_iff_ne_top, - ENNReal.tsum_coe_ne_top_iff_summable, ← NNReal.summable_coe, coe_nnnorm] - /-- A function is integrable for the counting measure iff its norm is summable. -/ lemma integrable_count_iff : Integrable f Measure.count ↔ Summable (‖f ·‖) := by @@ -985,16 +616,11 @@ theorem integrable_withDensity_iff_integrable_coe_smul {f : α → ℝ≥0} (hf {g : α → E} : Integrable g (μ.withDensity fun x => f x) ↔ Integrable (fun x => (f x : ℝ) • g x) μ := by by_cases H : AEStronglyMeasurable (fun x : α => (f x : ℝ) • g x) μ - · simp only [Integrable, aestronglyMeasurable_withDensity_iff hf, hasFiniteIntegral_iff_nnnorm, H] + · simp only [Integrable, aestronglyMeasurable_withDensity_iff hf, hasFiniteIntegral_iff_enorm, H, + true_and] rw [lintegral_withDensity_eq_lintegral_mul₀' hf.coe_nnreal_ennreal.aemeasurable] - · rw [iff_iff_eq] - congr - ext1 x - simp only [nnnorm_smul, NNReal.nnnorm_eq, coe_mul, Pi.mul_apply] - · rw [aemeasurable_withDensity_ennreal_iff hf] - convert H.ennnorm using 1 - ext1 x - simp only [nnnorm_smul, NNReal.nnnorm_eq, coe_mul] + · simp [enorm_smul] + · simpa [aemeasurable_withDensity_ennreal_iff hf, enorm_smul] using H.enorm · simp only [Integrable, aestronglyMeasurable_withDensity_iff hf, H, false_and] theorem integrable_withDensity_iff_integrable_smul {f : α → ℝ≥0} (hf : Measurable f) {g : α → E} : @@ -1093,9 +719,8 @@ noncomputable def withDensitySMulLI {f : α → ℝ≥0} (f_meas : Measurable f) congr 1 apply lintegral_congr_ae filter_upwards [(memℒ1_smul_of_L1_withDensity f_meas u).coeFn_toLp] with x hx - rw [hx, Pi.mul_apply] - change (‖(f x : ℝ) • u x‖₊ : ℝ≥0∞) = (f x : ℝ≥0∞) * (‖u x‖₊ : ℝ≥0∞) - simp only [nnnorm_smul, NNReal.nnnorm_eq, ENNReal.coe_mul] + rw [hx] + simp [NNReal.smul_def, enorm_smul] @[simp] theorem withDensitySMulLI_apply {f : α → ℝ≥0} (f_meas : Measurable f) @@ -1110,7 +735,7 @@ section ENNReal theorem mem_ℒ1_toReal_of_lintegral_ne_top {f : α → ℝ≥0∞} (hfm : AEMeasurable f μ) (hfi : ∫⁻ x, f x ∂μ ≠ ∞) : Memℒp (fun x ↦ (f x).toReal) 1 μ := by - rw [Memℒp, eLpNorm_one_eq_lintegral_nnnorm] + rw [Memℒp, eLpNorm_one_eq_lintegral_enorm] exact ⟨(AEMeasurable.ennreal_toReal hfm).aestronglyMeasurable, hasFiniteIntegral_toReal_of_lintegral_ne_top hfi⟩ @@ -1193,8 +818,8 @@ variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] theorem integrable_smul_const {f : α → 𝕜} {c : E} (hc : c ≠ 0) : Integrable (fun x => f x • c) μ ↔ Integrable f μ := by simp_rw [Integrable, aestronglyMeasurable_smul_const_iff (f := f) hc, and_congr_right_iff, - hasFiniteIntegral_iff_nnnorm, nnnorm_smul, ENNReal.coe_mul] - intro _; rw [lintegral_mul_const' _ _ ENNReal.coe_ne_top, ENNReal.mul_lt_top_iff] + hasFiniteIntegral_iff_enorm, enorm_smul] + intro _; rw [lintegral_mul_const' _ _ enorm_ne_top, ENNReal.mul_lt_top_iff] have : ∀ x : ℝ≥0∞, x = 0 → x < ∞ := by simp simp [hc, or_iff_left_of_imp (this _)] @@ -1295,15 +920,13 @@ theorem Integrable.trim (hm : m ≤ m0) (hf_int : Integrable f μ') (hf : Strong refine ⟨hf.aestronglyMeasurable, ?_⟩ rw [HasFiniteIntegral, lintegral_trim hm _] · exact hf_int.2 - · exact @StronglyMeasurable.ennnorm _ m _ _ f hf + · exact @StronglyMeasurable.enorm _ m _ _ f hf theorem integrable_of_integrable_trim (hm : m ≤ m0) (hf_int : Integrable f (μ'.trim hm)) : Integrable f μ' := by obtain ⟨hf_meas_ae, hf⟩ := hf_int refine ⟨aestronglyMeasurable_of_aestronglyMeasurable_trim hm hf_meas_ae, ?_⟩ - rw [HasFiniteIntegral] at hf ⊢ - rwa [lintegral_trim_ae hm _] at hf - exact AEStronglyMeasurable.ennnorm hf_meas_ae + simpa [HasFiniteIntegral, lintegral_trim_ae hm hf_meas_ae.enorm] using hf end Trim @@ -1313,228 +936,22 @@ variable {E : Type*} {m0 : MeasurableSpace α} [NormedAddCommGroup E] theorem integrable_of_forall_fin_meas_le' {μ : Measure α} (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (C : ℝ≥0∞) (hC : C < ∞) {f : α → E} (hf_meas : AEStronglyMeasurable f μ) - (hf : ∀ s, MeasurableSet[m] s → μ s ≠ ∞ → (∫⁻ x in s, ‖f x‖₊ ∂μ) ≤ C) : Integrable f μ := + (hf : ∀ s, MeasurableSet[m] s → μ s ≠ ∞ → ∫⁻ x in s, ‖f x‖ₑ ∂μ ≤ C) : Integrable f μ := ⟨hf_meas, (lintegral_le_of_forall_fin_meas_trim_le hm C hf).trans_lt hC⟩ theorem integrable_of_forall_fin_meas_le [SigmaFinite μ] (C : ℝ≥0∞) (hC : C < ∞) {f : α → E} (hf_meas : AEStronglyMeasurable[m] f μ) - (hf : ∀ s : Set α, MeasurableSet[m] s → μ s ≠ ∞ → (∫⁻ x in s, ‖f x‖₊ ∂μ) ≤ C) : + (hf : ∀ s : Set α, MeasurableSet[m] s → μ s ≠ ∞ → ∫⁻ x in s, ‖f x‖ₑ ∂μ ≤ C) : Integrable f μ := have : SigmaFinite (μ.trim le_rfl) := by rwa [@trim_eq_self _ m] integrable_of_forall_fin_meas_le' le_rfl C hC hf_meas hf end SigmaFinite -/-! ### The predicate `Integrable` on measurable functions modulo a.e.-equality -/ - - -namespace AEEqFun - -section - -/-- A class of almost everywhere equal functions is `Integrable` if its function representative -is integrable. -/ -def Integrable (f : α →ₘ[μ] β) : Prop := - MeasureTheory.Integrable f μ - -theorem integrable_mk {f : α → β} (hf : AEStronglyMeasurable f μ) : - Integrable (mk f hf : α →ₘ[μ] β) ↔ MeasureTheory.Integrable f μ := by - simp only [Integrable] - apply integrable_congr - exact coeFn_mk f hf - -theorem integrable_coeFn {f : α →ₘ[μ] β} : MeasureTheory.Integrable f μ ↔ Integrable f := by - rw [← integrable_mk, mk_coeFn] - -theorem integrable_zero : Integrable (0 : α →ₘ[μ] β) := - (MeasureTheory.integrable_zero α β μ).congr (coeFn_mk _ _).symm - -end - -section - -theorem Integrable.neg {f : α →ₘ[μ] β} : Integrable f → Integrable (-f) := - induction_on f fun _f hfm hfi => (integrable_mk _).2 ((integrable_mk hfm).1 hfi).neg - -section - -theorem integrable_iff_mem_L1 {f : α →ₘ[μ] β} : Integrable f ↔ f ∈ (α →₁[μ] β) := by - rw [← integrable_coeFn, ← memℒp_one_iff_integrable, Lp.mem_Lp_iff_memℒp] - -theorem Integrable.add {f g : α →ₘ[μ] β} : Integrable f → Integrable g → Integrable (f + g) := by - refine induction_on₂ f g fun f hf g hg hfi hgi => ?_ - simp only [integrable_mk, mk_add_mk] at hfi hgi ⊢ - exact hfi.add hgi - -theorem Integrable.sub {f g : α →ₘ[μ] β} (hf : Integrable f) (hg : Integrable g) : - Integrable (f - g) := - (sub_eq_add_neg f g).symm ▸ hf.add hg.neg - -end - -section BoundedSMul - -variable {𝕜 : Type*} [NormedRing 𝕜] [Module 𝕜 β] [BoundedSMul 𝕜 β] - -theorem Integrable.smul {c : 𝕜} {f : α →ₘ[μ] β} : Integrable f → Integrable (c • f) := - induction_on f fun _f hfm hfi => (integrable_mk _).2 <| - by simpa using ((integrable_mk hfm).1 hfi).smul c - -end BoundedSMul - -end - -end AEEqFun - -namespace L1 - - -theorem integrable_coeFn (f : α →₁[μ] β) : Integrable f μ := by - rw [← memℒp_one_iff_integrable] - exact Lp.memℒp f - -theorem hasFiniteIntegral_coeFn (f : α →₁[μ] β) : HasFiniteIntegral f μ := - (integrable_coeFn f).hasFiniteIntegral - -theorem stronglyMeasurable_coeFn (f : α →₁[μ] β) : StronglyMeasurable f := - Lp.stronglyMeasurable f - -theorem measurable_coeFn [MeasurableSpace β] [BorelSpace β] (f : α →₁[μ] β) : Measurable f := - (Lp.stronglyMeasurable f).measurable - -theorem aestronglyMeasurable_coeFn (f : α →₁[μ] β) : AEStronglyMeasurable f μ := - Lp.aestronglyMeasurable f - -theorem aemeasurable_coeFn [MeasurableSpace β] [BorelSpace β] (f : α →₁[μ] β) : AEMeasurable f μ := - (Lp.stronglyMeasurable f).measurable.aemeasurable - -theorem edist_def (f g : α →₁[μ] β) : edist f g = ∫⁻ a, edist (f a) (g a) ∂μ := by - simp only [Lp.edist_def, eLpNorm, one_ne_zero, eLpNorm'_eq_lintegral_nnnorm, Pi.sub_apply, - one_toReal, ENNReal.rpow_one, ne_eq, not_false_eq_true, div_self, ite_false] - simp [edist_eq_coe_nnnorm_sub] - -theorem dist_def (f g : α →₁[μ] β) : dist f g = (∫⁻ a, edist (f a) (g a) ∂μ).toReal := by - simp_rw [dist_edist, edist_def] - -theorem norm_def (f : α →₁[μ] β) : ‖f‖ = (∫⁻ a, ‖f a‖₊ ∂μ).toReal := by - simp [Lp.norm_def, eLpNorm, eLpNorm'_eq_lintegral_nnnorm] - -/-- Computing the norm of a difference between two L¹-functions. Note that this is not a - special case of `norm_def` since `(f - g) x` and `f x - g x` are not equal - (but only a.e.-equal). -/ -theorem norm_sub_eq_lintegral (f g : α →₁[μ] β) : - ‖f - g‖ = (∫⁻ x, (‖f x - g x‖₊ : ℝ≥0∞) ∂μ).toReal := by - rw [norm_def] - congr 1 - rw [lintegral_congr_ae] - filter_upwards [Lp.coeFn_sub f g] with _ ha - simp only [ha, Pi.sub_apply] - -theorem ofReal_norm_eq_lintegral (f : α →₁[μ] β) : - ENNReal.ofReal ‖f‖ = ∫⁻ x, (‖f x‖₊ : ℝ≥0∞) ∂μ := by - rw [norm_def, ENNReal.ofReal_toReal] - exact ne_of_lt (hasFiniteIntegral_coeFn f) - -/-- Computing the norm of a difference between two L¹-functions. Note that this is not a - special case of `ofReal_norm_eq_lintegral` since `(f - g) x` and `f x - g x` are not equal - (but only a.e.-equal). -/ -theorem ofReal_norm_sub_eq_lintegral (f g : α →₁[μ] β) : - ENNReal.ofReal ‖f - g‖ = ∫⁻ x, (‖f x - g x‖₊ : ℝ≥0∞) ∂μ := by - simp_rw [ofReal_norm_eq_lintegral, ← edist_eq_coe_nnnorm] - apply lintegral_congr_ae - filter_upwards [Lp.coeFn_sub f g] with _ ha - simp only [ha, Pi.sub_apply] - -end L1 - -namespace Integrable - - -/-- Construct the equivalence class `[f]` of an integrable function `f`, as a member of the -space `L1 β 1 μ`. -/ -def toL1 (f : α → β) (hf : Integrable f μ) : α →₁[μ] β := - (memℒp_one_iff_integrable.2 hf).toLp f - -@[simp] -theorem toL1_coeFn (f : α →₁[μ] β) (hf : Integrable f μ) : hf.toL1 f = f := by - simp [Integrable.toL1] - -theorem coeFn_toL1 {f : α → β} (hf : Integrable f μ) : hf.toL1 f =ᵐ[μ] f := - AEEqFun.coeFn_mk _ _ - -@[simp] -theorem toL1_zero (h : Integrable (0 : α → β) μ) : h.toL1 0 = 0 := - rfl - -@[simp] -theorem toL1_eq_mk (f : α → β) (hf : Integrable f μ) : - (hf.toL1 f : α →ₘ[μ] β) = AEEqFun.mk f hf.aestronglyMeasurable := - rfl - -@[simp] -theorem toL1_eq_toL1_iff (f g : α → β) (hf : Integrable f μ) (hg : Integrable g μ) : - toL1 f hf = toL1 g hg ↔ f =ᵐ[μ] g := - Memℒp.toLp_eq_toLp_iff _ _ - -theorem toL1_add (f g : α → β) (hf : Integrable f μ) (hg : Integrable g μ) : - toL1 (f + g) (hf.add hg) = toL1 f hf + toL1 g hg := - rfl - -theorem toL1_neg (f : α → β) (hf : Integrable f μ) : toL1 (-f) (Integrable.neg hf) = -toL1 f hf := - rfl - -theorem toL1_sub (f g : α → β) (hf : Integrable f μ) (hg : Integrable g μ) : - toL1 (f - g) (hf.sub hg) = toL1 f hf - toL1 g hg := - rfl - -theorem norm_toL1 (f : α → β) (hf : Integrable f μ) : - ‖hf.toL1 f‖ = ENNReal.toReal (∫⁻ a, edist (f a) 0 ∂μ) := by - simp only [toL1, Lp.norm_toLp, eLpNorm, one_ne_zero, eLpNorm'_eq_lintegral_nnnorm, one_toReal, - ENNReal.rpow_one, ne_eq, not_false_eq_true, div_self, ite_false] - simp [edist_eq_coe_nnnorm] - -theorem nnnorm_toL1 {f : α → β} (hf : Integrable f μ) : - (‖hf.toL1 f‖₊ : ℝ≥0∞) = ∫⁻ a, ‖f a‖₊ ∂μ := by - simpa [Integrable.toL1, eLpNorm, eLpNorm'] using ENNReal.coe_toNNReal hf.2.ne - -theorem norm_toL1_eq_lintegral_norm (f : α → β) (hf : Integrable f μ) : - ‖hf.toL1 f‖ = ENNReal.toReal (∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ) := by - rw [norm_toL1, lintegral_norm_eq_lintegral_edist] - -@[simp] -theorem edist_toL1_toL1 (f g : α → β) (hf : Integrable f μ) (hg : Integrable g μ) : - edist (hf.toL1 f) (hg.toL1 g) = ∫⁻ a, edist (f a) (g a) ∂μ := by - simp only [toL1, Lp.edist_toLp_toLp, eLpNorm, one_ne_zero, eLpNorm'_eq_lintegral_nnnorm, - Pi.sub_apply, one_toReal, ENNReal.rpow_one, ne_eq, not_false_eq_true, div_self, ite_false] - simp [edist_eq_coe_nnnorm_sub] - -theorem edist_toL1_zero (f : α → β) (hf : Integrable f μ) : - edist (hf.toL1 f) 0 = ∫⁻ a, edist (f a) 0 ∂μ := by - simp only [edist_zero_right, Lp.nnnorm_coe_ennreal, toL1_eq_mk, eLpNorm_aeeqFun] - apply eLpNorm_one_eq_lintegral_nnnorm - -variable {𝕜 : Type*} [NormedRing 𝕜] [Module 𝕜 β] [BoundedSMul 𝕜 β] - -theorem toL1_smul (f : α → β) (hf : Integrable f μ) (k : 𝕜) : - toL1 (fun a => k • f a) (hf.smul k) = k • toL1 f hf := - rfl - -theorem toL1_smul' (f : α → β) (hf : Integrable f μ) (k : 𝕜) : - toL1 (k • f) (hf.smul k) = k • toL1 f hf := - rfl - -end Integrable - section restrict variable {E : Type*} [NormedAddCommGroup E] {f : α → E} -lemma HasFiniteIntegral.restrict (h : HasFiniteIntegral f μ) {s : Set α} : - HasFiniteIntegral f (μ.restrict s) := by - refine lt_of_le_of_lt ?_ h - convert lintegral_mono_set (μ := μ) (s := s) (t := univ) (f := fun x ↦ ↑‖f x‖₊) (subset_univ s) - exact Measure.restrict_univ.symm - /-- One should usually use `MeasureTheory.Integrable.IntegrableOn` instead. -/ lemma Integrable.restrict (hf : Integrable f μ) {s : Set α} : Integrable f (μ.restrict s) := hf.mono_measure Measure.restrict_le_self @@ -1592,5 +1009,3 @@ lemma integrable_prod {f : α → E × F} : ⟨fun h ↦ ⟨h.fst, h.snd⟩, fun h ↦ h.1.prod_mk h.2⟩ end MeasureTheory - -set_option linter.style.longFile 1700 diff --git a/Mathlib/MeasureTheory/Function/L2Space.lean b/Mathlib/MeasureTheory/Function/L2Space.lean index 3594c9e9b9d387..878c21c3d7fa32 100644 --- a/Mathlib/MeasureTheory/Function/L2Space.lean +++ b/Mathlib/MeasureTheory/Function/L2Space.lean @@ -156,7 +156,7 @@ theorem integral_inner_eq_sq_eLpNorm (f : α →₂[μ] E) : ext1 x have h_two : (2 : ℝ) = ((2 : ℕ) : ℝ) := by simp rw [← Real.rpow_natCast _ 2, ← h_two, ← - ENNReal.ofReal_rpow_of_nonneg (norm_nonneg _) zero_le_two, ofReal_norm_eq_coe_nnnorm] + ENNReal.ofReal_rpow_of_nonneg (norm_nonneg _) zero_le_two, ofReal_norm_eq_enorm] norm_cast @[deprecated (since := "2024-07-27")] @@ -169,7 +169,7 @@ private theorem norm_sq_eq_inner' (f : α →₂[μ] E) : ‖f‖ ^ 2 = RCLike.r · rw [← ENNReal.rpow_natCast, eLpNorm_eq_eLpNorm' two_ne_zero ENNReal.two_ne_top, eLpNorm', ← ENNReal.rpow_mul, one_div, h_two] simp [enorm_eq_nnnorm] - · refine (lintegral_rpow_nnnorm_lt_top_of_eLpNorm'_lt_top zero_lt_two ?_).ne + · refine (lintegral_rpow_enorm_lt_top_of_eLpNorm'_lt_top zero_lt_two ?_).ne rw [← h_two, ← eLpNorm_eq_eLpNorm' two_ne_zero ENNReal.two_ne_top] exact Lp.eLpNorm_lt_top f diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index a2019826e26f5a..34188e7455aaaf 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -70,18 +70,22 @@ this quantity is finite -/ def eLpNorm' {_ : MeasurableSpace α} (f : α → ε) (q : ℝ) (μ : Measure α) : ℝ≥0∞ := (∫⁻ a, ‖f a‖ₑ ^ q ∂μ) ^ (1 / q) -lemma eLpNorm'_eq_lintegral_nnnorm {_ : MeasurableSpace α} (f : α → F) (q : ℝ) (μ : Measure α) : - eLpNorm' f q μ = (∫⁻ a, ‖f a‖₊ ^ q ∂μ) ^ (1 / q) := +lemma eLpNorm'_eq_lintegral_enorm {_ : MeasurableSpace α} (f : α → F) (q : ℝ) (μ : Measure α) : + eLpNorm' f q μ = (∫⁻ a, ‖f a‖ₑ ^ q ∂μ) ^ (1 / q) := rfl +@[deprecated (since := "2025-01-17")] +alias eLpNorm'_eq_lintegral_nnnorm := eLpNorm'_eq_lintegral_enorm + /-- seminorm for `ℒ∞`, equal to the essential supremum of `‖f‖`. -/ def eLpNormEssSup {_ : MeasurableSpace α} (f : α → ε) (μ : Measure α) := essSup (fun x => ‖f x‖ₑ) μ -lemma eLpNormEssSup_eq_essSup_nnnorm {_ : MeasurableSpace α} (f : α → F) (μ : Measure α) : - eLpNormEssSup f μ = essSup (fun x => (‖f x‖₊ : ℝ≥0∞)) μ := - rfl +lemma eLpNormEssSup_eq_essSup_enorm {_ : MeasurableSpace α} (f : α → F) (μ : Measure α) : + eLpNormEssSup f μ = essSup (‖f ·‖ₑ) μ := rfl +@[deprecated (since := "2025-01-17")] +alias eLpNormEssSup_eq_essSup_nnnorm := eLpNormEssSup_eq_essSup_enorm /-- `ℒp` seminorm, equal to `0` for `p=0`, to `(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `0 < p < ∞` and to `essSup ‖f‖ μ` for `p = ∞`. -/ @@ -102,25 +106,31 @@ lemma eLpNorm_nnreal_eq_eLpNorm' {f : α → F} {p : ℝ≥0} (hp : p ≠ 0) : @[deprecated (since := "2024-07-27")] alias snorm_nnreal_eq_snorm' := eLpNorm_nnreal_eq_eLpNorm' -theorem eLpNorm_eq_lintegral_rpow_nnnorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {f : α → F} : - eLpNorm f p μ = (∫⁻ x, (‖f x‖₊ : ℝ≥0∞) ^ p.toReal ∂μ) ^ (1 / p.toReal) := by - rw [eLpNorm_eq_eLpNorm' hp_ne_zero hp_ne_top, eLpNorm'_eq_lintegral_nnnorm] +theorem eLpNorm_eq_lintegral_rpow_enorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {f : α → F} : + eLpNorm f p μ = (∫⁻ x, ‖f x‖ₑ ^ p.toReal ∂μ) ^ (1 / p.toReal) := by + rw [eLpNorm_eq_eLpNorm' hp_ne_zero hp_ne_top, eLpNorm'_eq_lintegral_enorm] @[deprecated (since := "2024-07-27")] -alias snorm_eq_lintegral_rpow_nnnorm := eLpNorm_eq_lintegral_rpow_nnnorm +alias snorm_eq_lintegral_rpow_nnnorm := eLpNorm_eq_lintegral_rpow_enorm + +@[deprecated (since := "2025-01-17")] +alias eLpNorm_eq_lintegral_rpow_nnnorm := eLpNorm_eq_lintegral_rpow_enorm lemma eLpNorm_nnreal_eq_lintegral {f : α → F} {p : ℝ≥0} (hp : p ≠ 0) : - eLpNorm f p μ = (∫⁻ x, ‖f x‖₊ ^ (p : ℝ) ∂μ) ^ (1 / (p : ℝ)) := + eLpNorm f p μ = (∫⁻ x, ‖f x‖ₑ ^ (p : ℝ) ∂μ) ^ (1 / (p : ℝ)) := eLpNorm_nnreal_eq_eLpNorm' hp @[deprecated (since := "2024-07-27")] alias snorm_nnreal_eq_lintegral := eLpNorm_nnreal_eq_lintegral -theorem eLpNorm_one_eq_lintegral_nnnorm {f : α → F} : eLpNorm f 1 μ = ∫⁻ x, ‖f x‖₊ ∂μ := by - simp_rw [eLpNorm_eq_lintegral_rpow_nnnorm one_ne_zero ENNReal.coe_ne_top, ENNReal.one_toReal, +theorem eLpNorm_one_eq_lintegral_enorm {f : α → F} : eLpNorm f 1 μ = ∫⁻ x, ‖f x‖ₑ ∂μ := by + simp_rw [eLpNorm_eq_lintegral_rpow_enorm one_ne_zero ENNReal.coe_ne_top, ENNReal.one_toReal, one_div_one, ENNReal.rpow_one] +@[deprecated (since := "2025-01-17")] +alias eLpNorm_one_eq_lintegral_nnnorm := eLpNorm_one_eq_lintegral_enorm + @[deprecated (since := "2024-07-27")] -alias snorm_one_eq_lintegral_nnnorm := eLpNorm_one_eq_lintegral_nnnorm +alias snorm_one_eq_lintegral_nnnorm := eLpNorm_one_eq_lintegral_enorm @[simp] theorem eLpNorm_exponent_top {f : α → F} : eLpNorm f ∞ μ = eLpNormEssSup f μ := by simp [eLpNorm] @@ -138,18 +148,21 @@ theorem Memℒp.aestronglyMeasurable {f : α → E} {p : ℝ≥0∞} (h : Memℒ AEStronglyMeasurable f μ := h.1 -theorem lintegral_rpow_nnnorm_eq_rpow_eLpNorm' {f : α → F} (hq0_lt : 0 < q) : - ∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ q ∂μ = eLpNorm' f q μ ^ q := by - rw [eLpNorm'_eq_lintegral_nnnorm, ← ENNReal.rpow_mul, one_div, inv_mul_cancel₀, ENNReal.rpow_one] - exact (ne_of_lt hq0_lt).symm +theorem lintegral_rpow_enorm_eq_rpow_eLpNorm' {f : α → F} (hq0_lt : 0 < q) : + ∫⁻ a, ‖f a‖ₑ ^ q ∂μ = eLpNorm' f q μ ^ q := by + rw [eLpNorm'_eq_lintegral_enorm, ← ENNReal.rpow_mul, one_div, inv_mul_cancel₀, ENNReal.rpow_one] + exact hq0_lt.ne' + +@[deprecated (since := "2025-01-17")] +alias lintegral_rpow_nnnorm_eq_rpow_eLpNorm' := lintegral_rpow_enorm_eq_rpow_eLpNorm' @[deprecated (since := "2024-07-27")] -alias lintegral_rpow_nnnorm_eq_rpow_snorm' := lintegral_rpow_nnnorm_eq_rpow_eLpNorm' +alias lintegral_rpow_nnnorm_eq_rpow_snorm' := lintegral_rpow_enorm_eq_rpow_eLpNorm' lemma eLpNorm_nnreal_pow_eq_lintegral {f : α → F} {p : ℝ≥0} (hp : p ≠ 0) : - eLpNorm f p μ ^ (p : ℝ) = ∫⁻ x, ‖f x‖₊ ^ (p : ℝ) ∂μ := by + eLpNorm f p μ ^ (p : ℝ) = ∫⁻ x, ‖f x‖ₑ ^ (p : ℝ) ∂μ := by simp [eLpNorm_eq_eLpNorm' (by exact_mod_cast hp) ENNReal.coe_ne_top, - lintegral_rpow_nnnorm_eq_rpow_eLpNorm' ((NNReal.coe_pos.trans pos_iff_ne_zero).mpr hp)] + lintegral_rpow_enorm_eq_rpow_eLpNorm' ((NNReal.coe_pos.trans pos_iff_ne_zero).mpr hp)] @[deprecated (since := "2024-07-27")] alias snorm_nnreal_pow_eq_lintegral := eLpNorm_nnreal_pow_eq_lintegral @@ -170,31 +183,39 @@ theorem Memℒp.eLpNorm_ne_top {f : α → E} (hfp : Memℒp f p μ) : eLpNorm f @[deprecated (since := "2024-07-27")] alias Memℒp.snorm_ne_top := Memℒp.eLpNorm_ne_top -theorem lintegral_rpow_nnnorm_lt_top_of_eLpNorm'_lt_top {f : α → F} (hq0_lt : 0 < q) - (hfq : eLpNorm' f q μ < ∞) : (∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ q ∂μ) < ∞ := by - rw [lintegral_rpow_nnnorm_eq_rpow_eLpNorm' hq0_lt] +theorem lintegral_rpow_enorm_lt_top_of_eLpNorm'_lt_top {f : α → F} (hq0_lt : 0 < q) + (hfq : eLpNorm' f q μ < ∞) : ∫⁻ a, ‖f a‖ₑ ^ q ∂μ < ∞ := by + rw [lintegral_rpow_enorm_eq_rpow_eLpNorm' hq0_lt] exact ENNReal.rpow_lt_top_of_nonneg (le_of_lt hq0_lt) (ne_of_lt hfq) @[deprecated (since := "2024-07-27")] alias lintegral_rpow_nnnorm_lt_top_of_snorm'_lt_top := - lintegral_rpow_nnnorm_lt_top_of_eLpNorm'_lt_top + lintegral_rpow_enorm_lt_top_of_eLpNorm'_lt_top + +@[deprecated (since := "2025-01-17")] +alias lintegral_rpow_nnnorm_lt_top_of_eLpNorm'_lt_top' := + lintegral_rpow_enorm_lt_top_of_eLpNorm'_lt_top -theorem lintegral_rpow_nnnorm_lt_top_of_eLpNorm_lt_top {f : α → F} (hp_ne_zero : p ≠ 0) - (hp_ne_top : p ≠ ∞) (hfp : eLpNorm f p μ < ∞) : (∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ p.toReal ∂μ) < ∞ := by - apply lintegral_rpow_nnnorm_lt_top_of_eLpNorm'_lt_top +theorem lintegral_rpow_enorm_lt_top_of_eLpNorm_lt_top {f : α → F} (hp_ne_zero : p ≠ 0) + (hp_ne_top : p ≠ ∞) (hfp : eLpNorm f p μ < ∞) : ∫⁻ a, ‖f a‖ₑ ^ p.toReal ∂μ < ∞ := by + apply lintegral_rpow_enorm_lt_top_of_eLpNorm'_lt_top · exact ENNReal.toReal_pos hp_ne_zero hp_ne_top · simpa [eLpNorm_eq_eLpNorm' hp_ne_zero hp_ne_top] using hfp @[deprecated (since := "2024-07-27")] -alias lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top := lintegral_rpow_nnnorm_lt_top_of_eLpNorm_lt_top +alias lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top := lintegral_rpow_enorm_lt_top_of_eLpNorm_lt_top + +@[deprecated (since := "2025-01-17")] +alias lintegral_rpow_nnnorm_lt_top_of_eLpNorm_lt_top := + lintegral_rpow_enorm_lt_top_of_eLpNorm_lt_top theorem eLpNorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top {f : α → F} (hp_ne_zero : p ≠ 0) - (hp_ne_top : p ≠ ∞) : eLpNorm f p μ < ∞ ↔ (∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ p.toReal ∂μ) < ∞ := - ⟨lintegral_rpow_nnnorm_lt_top_of_eLpNorm_lt_top hp_ne_zero hp_ne_top, by + (hp_ne_top : p ≠ ∞) : eLpNorm f p μ < ∞ ↔ ∫⁻ a, (‖f a‖ₑ) ^ p.toReal ∂μ < ∞ := + ⟨lintegral_rpow_enorm_lt_top_of_eLpNorm_lt_top hp_ne_zero hp_ne_top, by intro h have hp' := ENNReal.toReal_pos hp_ne_zero hp_ne_top have : 0 < 1 / p.toReal := div_pos zero_lt_one hp' - simpa [eLpNorm_eq_lintegral_rpow_nnnorm hp_ne_zero hp_ne_top] using + simpa [eLpNorm_eq_lintegral_rpow_enorm hp_ne_zero hp_ne_top] using ENNReal.rpow_lt_top_of_nonneg (le_of_lt this) (ne_of_lt h)⟩ @[deprecated (since := "2024-07-27")] @@ -224,7 +245,7 @@ theorem memℒp_zero_iff_aestronglyMeasurable {f : α → E} : @[simp] theorem eLpNorm'_zero (hp0_lt : 0 < q) : eLpNorm' (0 : α → F) q μ = 0 := by - simp [eLpNorm'_eq_lintegral_nnnorm, hp0_lt] + simp [eLpNorm'_eq_lintegral_enorm, hp0_lt] @[deprecated (since := "2024-07-27")] alias snorm'_zero := eLpNorm'_zero @@ -233,16 +254,14 @@ alias snorm'_zero := eLpNorm'_zero theorem eLpNorm'_zero' (hq0_ne : q ≠ 0) (hμ : μ ≠ 0) : eLpNorm' (0 : α → F) q μ = 0 := by rcases le_or_lt 0 q with hq0 | hq_neg · exact eLpNorm'_zero (lt_of_le_of_ne hq0 hq0_ne.symm) - · simp [eLpNorm'_eq_lintegral_nnnorm, ENNReal.rpow_eq_zero_iff, hμ, hq_neg] + · simp [eLpNorm'_eq_lintegral_enorm, ENNReal.rpow_eq_zero_iff, hμ, hq_neg] @[deprecated (since := "2024-07-27")] alias snorm'_zero' := eLpNorm'_zero' @[simp] theorem eLpNormEssSup_zero : eLpNormEssSup (0 : α → F) μ = 0 := by - simp_rw [eLpNormEssSup_eq_essSup_nnnorm, Pi.zero_apply, nnnorm_zero, ENNReal.coe_zero, - ← ENNReal.bot_eq_zero] - exact essSup_const_bot + simp [eLpNormEssSup, ← bot_eq_zero', essSup_const_bot] @[deprecated (since := "2024-07-27")] alias snormEssSup_zero := eLpNormEssSup_zero @@ -320,7 +339,7 @@ section Neg @[simp] theorem eLpNorm'_neg (f : α → F) (q : ℝ) (μ : Measure α) : eLpNorm' (-f) q μ = eLpNorm' f q μ := by - simp [eLpNorm'_eq_lintegral_nnnorm] + simp [eLpNorm'_eq_lintegral_enorm] @[deprecated (since := "2024-07-27")] alias snorm'_neg := eLpNorm'_neg @@ -330,7 +349,7 @@ theorem eLpNorm_neg (f : α → F) (p : ℝ≥0∞) (μ : Measure α) : eLpNorm by_cases h0 : p = 0 · simp [h0] by_cases h_top : p = ∞ - · simp [h_top, eLpNormEssSup_eq_essSup_nnnorm] + · simp [h_top, eLpNormEssSup_eq_essSup_enorm] simp [eLpNorm_eq_eLpNorm' h0 h_top] lemma eLpNorm_sub_comm (f g : α → E) (p : ℝ≥0∞) (μ : Measure α) : @@ -350,8 +369,8 @@ end Neg section Const theorem eLpNorm'_const (c : F) (hq_pos : 0 < q) : - eLpNorm' (fun _ : α => c) q μ = (‖c‖₊ : ℝ≥0∞) * μ Set.univ ^ (1 / q) := by - rw [eLpNorm'_eq_lintegral_nnnorm, lintegral_const, + eLpNorm' (fun _ : α => c) q μ = ‖c‖ₑ * μ Set.univ ^ (1 / q) := by + rw [eLpNorm'_eq_lintegral_enorm, lintegral_const, ENNReal.mul_rpow_of_nonneg _ _ (by simp [hq_pos.le] : 0 ≤ 1 / q)] congr rw [← ENNReal.rpow_mul] @@ -362,37 +381,33 @@ theorem eLpNorm'_const (c : F) (hq_pos : 0 < q) : alias snorm'_const := eLpNorm'_const theorem eLpNorm'_const' [IsFiniteMeasure μ] (c : F) (hc_ne_zero : c ≠ 0) (hq_ne_zero : q ≠ 0) : - eLpNorm' (fun _ : α => c) q μ = (‖c‖₊ : ℝ≥0∞) * μ Set.univ ^ (1 / q) := by - rw [eLpNorm'_eq_lintegral_nnnorm, lintegral_const, + eLpNorm' (fun _ : α => c) q μ = ‖c‖ₑ * μ Set.univ ^ (1 / q) := by + rw [eLpNorm'_eq_lintegral_enorm, lintegral_const, ENNReal.mul_rpow_of_ne_top _ (measure_ne_top μ Set.univ)] · congr rw [← ENNReal.rpow_mul] suffices hp_cancel : q * (1 / q) = 1 by rw [hp_cancel, ENNReal.rpow_one] rw [one_div, mul_inv_cancel₀ hq_ne_zero] · rw [Ne, ENNReal.rpow_eq_top_iff, not_or, not_and_or, not_and_or] - constructor - · left - rwa [ENNReal.coe_eq_zero, nnnorm_eq_zero] - · exact Or.inl ENNReal.coe_ne_top + simp [hc_ne_zero] @[deprecated (since := "2024-07-27")] alias snorm'_const' := eLpNorm'_const' -theorem eLpNormEssSup_const (c : F) (hμ : μ ≠ 0) : - eLpNormEssSup (fun _ : α => c) μ = (‖c‖₊ : ℝ≥0∞) := by - rw [eLpNormEssSup_eq_essSup_nnnorm, essSup_const _ hμ] +theorem eLpNormEssSup_const (c : F) (hμ : μ ≠ 0) : eLpNormEssSup (fun _ : α => c) μ = ‖c‖ₑ := by + rw [eLpNormEssSup_eq_essSup_enorm, essSup_const _ hμ] @[deprecated (since := "2024-07-27")] alias snormEssSup_const := eLpNormEssSup_const theorem eLpNorm'_const_of_isProbabilityMeasure (c : F) (hq_pos : 0 < q) [IsProbabilityMeasure μ] : - eLpNorm' (fun _ : α => c) q μ = (‖c‖₊ : ℝ≥0∞) := by simp [eLpNorm'_const c hq_pos, measure_univ] + eLpNorm' (fun _ : α => c) q μ = ‖c‖ₑ := by simp [eLpNorm'_const c hq_pos, measure_univ] @[deprecated (since := "2024-07-27")] alias snorm'_const_of_isProbabilityMeasure := eLpNorm'_const_of_isProbabilityMeasure theorem eLpNorm_const (c : F) (h0 : p ≠ 0) (hμ : μ ≠ 0) : - eLpNorm (fun _ : α => c) p μ = (‖c‖₊ : ℝ≥0∞) * μ Set.univ ^ (1 / ENNReal.toReal p) := by + eLpNorm (fun _ : α => c) p μ = ‖c‖ₑ * μ Set.univ ^ (1 / ENNReal.toReal p) := by by_cases h_top : p = ∞ · simp [h_top, eLpNormEssSup_const c hμ] simp [eLpNorm_eq_eLpNorm' h0 h_top, eLpNorm'_const, ENNReal.toReal_pos h0 h_top] @@ -401,7 +416,7 @@ theorem eLpNorm_const (c : F) (h0 : p ≠ 0) (hμ : μ ≠ 0) : alias snorm_const := eLpNorm_const theorem eLpNorm_const' (c : F) (h0 : p ≠ 0) (h_top : p ≠ ∞) : - eLpNorm (fun _ : α => c) p μ = (‖c‖₊ : ℝ≥0∞) * μ Set.univ ^ (1 / ENNReal.toReal p) := by + eLpNorm (fun _ : α => c) p μ = ‖c‖ₑ * μ Set.univ ^ (1 / ENNReal.toReal p) := by simp [eLpNorm_eq_eLpNorm' h0 h_top, eLpNorm'_const, ENNReal.toReal_pos h0 h_top] @[deprecated (since := "2024-07-27")] @@ -416,14 +431,11 @@ theorem eLpNorm_const_lt_top_iff {p : ℝ≥0∞} {c : F} (hp_ne_zero : p ≠ 0) by_cases hc : c = 0 · simp only [hc, true_or, eq_self_iff_true, ENNReal.zero_lt_top, eLpNorm_zero'] rw [eLpNorm_const' c hp_ne_zero hp_ne_top] - by_cases hμ_top : μ Set.univ = ∞ + obtain hμ_top | hμ_top := eq_or_ne (μ .univ) ∞ · simp [hc, hμ_top, hp] rw [ENNReal.mul_lt_top_iff] - simp only [true_and, one_div, ENNReal.rpow_eq_zero_iff, hμ, false_or, or_false, - ENNReal.coe_lt_top, nnnorm_eq_zero, ENNReal.coe_eq_zero, - MeasureTheory.Measure.measure_univ_eq_zero, hp, inv_lt_zero, hc, false_and, - inv_pos, or_self_iff, hμ_top, Ne.lt_top hμ_top, iff_true] - exact ENNReal.rpow_lt_top_of_nonneg (inv_nonneg.mpr hp.le) hμ_top + simpa [hμ, hc, hμ_top, hμ_top.lt_top] using + ENNReal.rpow_lt_top_of_nonneg (inv_nonneg.mpr hp.le) hμ_top @[deprecated (since := "2024-07-27")] alias snorm_const_lt_top_iff := eLpNorm_const_lt_top_iff @@ -439,12 +451,8 @@ theorem memℒp_const (c : E) [IsFiniteMeasure μ] : Memℒp (fun _ : α => c) p refine ENNReal.rpow_lt_top_of_nonneg ?_ (measure_ne_top μ Set.univ) simp -theorem memℒp_top_const (c : E) : Memℒp (fun _ : α => c) ∞ μ := by - refine ⟨aestronglyMeasurable_const, ?_⟩ - by_cases h : μ = 0 - · simp only [h, eLpNorm_measure_zero, ENNReal.zero_lt_top] - · rw [eLpNorm_const _ ENNReal.top_ne_zero h] - simp only [ENNReal.top_toReal, div_zero, ENNReal.rpow_zero, mul_one, ENNReal.coe_lt_top] +theorem memℒp_top_const (c : E) : Memℒp (fun _ : α => c) ∞ μ := + ⟨aestronglyMeasurable_const, by by_cases h : μ = 0 <;> simp [eLpNorm_const _, h]⟩ theorem memℒp_const_iff {p : ℝ≥0∞} {c : E} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : Memℒp (fun _ : α => c) p μ ↔ c = 0 ∨ μ Set.univ < ∞ := by @@ -457,9 +465,10 @@ variable {f : α → F} lemma eLpNorm'_mono_nnnorm_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : eLpNorm' f q μ ≤ eLpNorm' g q μ := by - simp only [eLpNorm'_eq_lintegral_nnnorm] + simp only [eLpNorm'_eq_lintegral_enorm] gcongr ?_ ^ (1/q) refine lintegral_mono_ae (h.mono fun x hx => ?_) + dsimp [enorm] gcongr @[deprecated (since := "2024-07-27")] @@ -474,9 +483,8 @@ alias snorm'_mono_ae := eLpNorm'_mono_ae theorem eLpNorm'_congr_nnnorm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ = ‖g x‖₊) : eLpNorm' f q μ = eLpNorm' g q μ := by - have : (fun x => (‖f x‖₊ : ℝ≥0∞) ^ q) =ᵐ[μ] fun x => (‖g x‖₊ : ℝ≥0∞) ^ q := - hfg.mono fun x hx => by simp_rw [hx] - simp only [eLpNorm'_eq_lintegral_nnnorm, lintegral_congr_ae this] + have : (‖f ·‖ₑ ^ q) =ᵐ[μ] (‖g ·‖ₑ ^ q) := hfg.mono fun x hx ↦ by simp [enorm, hx] + simp only [eLpNorm'_eq_lintegral_enorm, lintegral_congr_ae this] @[deprecated (since := "2024-07-27")] alias snorm'_congr_nnnorm_ae := eLpNorm'_congr_nnnorm_ae @@ -591,7 +599,7 @@ theorem eLpNorm_le_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ · simp [hp] have : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖(C : ℝ)‖₊ := hfC.mono fun x hx => hx.trans_eq C.nnnorm_eq.symm refine (eLpNorm_mono_ae this).trans_eq ?_ - rw [eLpNorm_const _ hp (NeZero.ne μ), C.nnnorm_eq, one_div, ENNReal.smul_def, smul_eq_mul] + rw [eLpNorm_const _ hp (NeZero.ne μ), C.enorm_eq, one_div, ENNReal.smul_def, smul_eq_mul] @[deprecated (since := "2024-07-27")] alias snorm_le_of_ae_nnnorm_bound := eLpNorm_le_of_ae_nnnorm_bound @@ -630,7 +638,7 @@ alias snorm_indicator_sub_indicator := eLpNorm_indicator_sub_indicator @[simp] theorem eLpNorm'_norm {f : α → F} : - eLpNorm' (fun a => ‖f a‖) q μ = eLpNorm' f q μ := by simp [eLpNorm'_eq_lintegral_nnnorm] + eLpNorm' (fun a => ‖f a‖) q μ = eLpNorm' f q μ := by simp [eLpNorm'_eq_lintegral_enorm] @[deprecated (since := "2024-07-27")] alias snorm'_norm := eLpNorm'_norm @@ -644,8 +652,8 @@ alias snorm_norm := eLpNorm_norm theorem eLpNorm'_norm_rpow (f : α → F) (p q : ℝ) (hq_pos : 0 < q) : eLpNorm' (fun x => ‖f x‖ ^ q) p μ = eLpNorm' f (p * q) μ ^ q := by - simp_rw [eLpNorm'_eq_lintegral_nnnorm, ← ENNReal.rpow_mul, ← one_div_mul_one_div, one_div, - mul_assoc, inv_mul_cancel₀ hq_pos.ne.symm, mul_one, ← ofReal_norm_eq_coe_nnnorm, + simp_rw [eLpNorm', ← ENNReal.rpow_mul, ← one_div_mul_one_div, one_div, + mul_assoc, inv_mul_cancel₀ hq_pos.ne.symm, mul_one, ← ofReal_norm_eq_enorm, Real.norm_eq_abs, abs_eq_self.mpr (Real.rpow_nonneg (norm_nonneg _) _), mul_comm p, ← ENNReal.ofReal_rpow_of_nonneg (norm_nonneg _) hq_pos.le, ENNReal.rpow_mul] @@ -658,22 +666,17 @@ theorem eLpNorm_norm_rpow (f : α → F) (hq_pos : 0 < q) : · simp [h0, ENNReal.zero_rpow_of_pos hq_pos] by_cases hp_top : p = ∞ · simp only [hp_top, eLpNorm_exponent_top, ENNReal.top_mul', hq_pos.not_le, - ENNReal.ofReal_eq_zero, if_false, eLpNorm_exponent_top, eLpNormEssSup_eq_essSup_nnnorm] - have h_rpow : - essSup (fun x : α => (‖‖f x‖ ^ q‖₊ : ℝ≥0∞)) μ = - essSup (fun x : α => (‖f x‖₊ : ℝ≥0∞) ^ q) μ := by + ENNReal.ofReal_eq_zero, if_false, eLpNorm_exponent_top, eLpNormEssSup_eq_essSup_enorm] + have h_rpow : essSup (‖‖f ·‖ ^ q‖ₑ) μ = essSup (‖f ·‖ₑ ^ q) μ := by congr ext1 x - conv_rhs => rw [← nnnorm_norm] - rw [← ENNReal.coe_rpow_of_nonneg _ hq_pos.le, ENNReal.coe_inj] - ext - push_cast - rw [Real.norm_rpow_of_nonneg (norm_nonneg _)] + conv_rhs => rw [← enorm_norm] + rw [← Real.enorm_rpow_of_nonneg (norm_nonneg _) hq_pos.le] rw [h_rpow] have h_rpow_mono := ENNReal.strictMono_rpow_of_pos hq_pos have h_rpow_surj := (ENNReal.rpow_left_bijective hq_pos.ne.symm).2 let iso := h_rpow_mono.orderIsoOfSurjective _ h_rpow_surj - exact (iso.essSup_apply (fun x => (‖f x‖₊ : ℝ≥0∞)) μ).symm + exact (iso.essSup_apply (fun x => ‖f x‖ₑ) μ).symm rw [eLpNorm_eq_eLpNorm' h0 hp_top, eLpNorm_eq_eLpNorm' _ _] swap · refine mul_ne_zero h0 ?_ @@ -774,15 +777,14 @@ lemma eLpNorm_indicator_eq_eLpNorm_restrict (hs : MeasurableSet s) : by_cases hp_zero : p = 0 · simp only [hp_zero, eLpNorm_exponent_zero] by_cases hp_top : p = ∞ - · simp_rw [hp_top, eLpNorm_exponent_top, eLpNormEssSup_eq_essSup_nnnorm, - nnnorm_indicator_eq_indicator_nnnorm, ENNReal.coe_indicator, - ENNReal.essSup_indicator_eq_essSup_restrict hs] - simp_rw [eLpNorm_eq_lintegral_rpow_nnnorm hp_zero hp_top] - suffices (∫⁻ x, (‖s.indicator f x‖₊ : ℝ≥0∞) ^ p.toReal ∂μ) = - ∫⁻ x in s, (‖f x‖₊ : ℝ≥0∞) ^ p.toReal ∂μ by rw [this] + · simp_rw [hp_top, eLpNorm_exponent_top, eLpNormEssSup_eq_essSup_enorm, + enorm_indicator_eq_indicator_enorm, ENNReal.essSup_indicator_eq_essSup_restrict hs] + simp_rw [eLpNorm_eq_lintegral_rpow_enorm hp_zero hp_top] + suffices (∫⁻ x, (‖s.indicator f x‖ₑ) ^ p.toReal ∂μ) = + ∫⁻ x in s, ‖f x‖ₑ ^ p.toReal ∂μ by rw [this] rw [← lintegral_indicator hs] congr - simp_rw [nnnorm_indicator_eq_indicator_nnnorm, ENNReal.coe_indicator] + simp_rw [enorm_indicator_eq_indicator_enorm] have h_zero : (fun x => x ^ p.toReal) (0 : ℝ≥0∞) = 0 := by simp [ENNReal.toReal_pos hp_zero hp_top] -- Porting note: The implicit argument should be specified because the elaborator can't deal with @@ -819,14 +821,14 @@ lemma eLpNormEssSup_indicator_le (s : Set α) (f : α → G) : exact Set.indicator_le_self s _ x lemma eLpNormEssSup_indicator_const_le (s : Set α) (c : G) : - eLpNormEssSup (s.indicator fun _ : α => c) μ ≤ ‖c‖₊ := by + eLpNormEssSup (s.indicator fun _ : α => c) μ ≤ ‖c‖ₑ := by by_cases hμ0 : μ = 0 · rw [hμ0, eLpNormEssSup_measure_zero] exact zero_le _ · exact (eLpNormEssSup_indicator_le s fun _ => c).trans (eLpNormEssSup_const c hμ0).le lemma eLpNormEssSup_indicator_const_eq (s : Set α) (c : G) (hμs : μ s ≠ 0) : - eLpNormEssSup (s.indicator fun _ : α => c) μ = ‖c‖₊ := by + eLpNormEssSup (s.indicator fun _ : α => c) μ = ‖c‖ₑ := by refine le_antisymm (eLpNormEssSup_indicator_const_le s c) ?_ by_contra! h have h' := ae_iff.mp (ae_lt_of_essSup_lt h) @@ -835,33 +837,33 @@ lemma eLpNormEssSup_indicator_const_eq (s : Set α) (c : G) (hμs : μ s ≠ 0) rw [Set.mem_setOf_eq, Set.indicator_of_mem hx_mem, enorm_eq_nnnorm] lemma eLpNorm_indicator_const₀ (hs : NullMeasurableSet s μ) (hp : p ≠ 0) (hp_top : p ≠ ∞) : - eLpNorm (s.indicator fun _ => c) p μ = ‖c‖₊ * μ s ^ (1 / p.toReal) := + eLpNorm (s.indicator fun _ => c) p μ = ‖c‖ₑ * μ s ^ (1 / p.toReal) := have hp_pos : 0 < p.toReal := ENNReal.toReal_pos hp hp_top calc eLpNorm (s.indicator fun _ => c) p μ - = (∫⁻ x, ((‖(s.indicator fun _ ↦ c) x‖₊ : ℝ≥0∞) ^ p.toReal) ∂μ) ^ (1 / p.toReal) := - eLpNorm_eq_lintegral_rpow_nnnorm hp hp_top - _ = (∫⁻ x, (s.indicator fun _ ↦ (‖c‖₊ : ℝ≥0∞) ^ p.toReal) x ∂μ) ^ (1 / p.toReal) := by + = (∫⁻ x, (‖(s.indicator fun _ ↦ c) x‖ₑ ^ p.toReal) ∂μ) ^ (1 / p.toReal) := + eLpNorm_eq_lintegral_rpow_enorm hp hp_top + _ = (∫⁻ x, (s.indicator fun _ ↦ ‖c‖ₑ ^ p.toReal) x ∂μ) ^ (1 / p.toReal) := by congr 2 - refine (Set.comp_indicator_const c (fun x ↦ (‖x‖₊ : ℝ≥0∞) ^ p.toReal) ?_) + refine (Set.comp_indicator_const c (fun x ↦ (‖x‖ₑ) ^ p.toReal) ?_) simp [hp_pos] - _ = ‖c‖₊ * μ s ^ (1 / p.toReal) := by + _ = ‖c‖ₑ * μ s ^ (1 / p.toReal) := by rw [lintegral_indicator_const₀ hs, ENNReal.mul_rpow_of_nonneg, ← ENNReal.rpow_mul, mul_one_div_cancel hp_pos.ne', ENNReal.rpow_one] positivity lemma eLpNorm_indicator_const (hs : MeasurableSet s) (hp : p ≠ 0) (hp_top : p ≠ ∞) : - eLpNorm (s.indicator fun _ => c) p μ = ‖c‖₊ * μ s ^ (1 / p.toReal) := + eLpNorm (s.indicator fun _ => c) p μ = ‖c‖ₑ * μ s ^ (1 / p.toReal) := eLpNorm_indicator_const₀ hs.nullMeasurableSet hp hp_top lemma eLpNorm_indicator_const' (hs : MeasurableSet s) (hμs : μ s ≠ 0) (hp : p ≠ 0) : - eLpNorm (s.indicator fun _ => c) p μ = ‖c‖₊ * μ s ^ (1 / p.toReal) := by + eLpNorm (s.indicator fun _ => c) p μ = ‖c‖ₑ * μ s ^ (1 / p.toReal) := by by_cases hp_top : p = ∞ · simp [hp_top, eLpNormEssSup_indicator_const_eq s c hμs] · exact eLpNorm_indicator_const hs hp hp_top lemma eLpNorm_indicator_const_le (c : G) (p : ℝ≥0∞) : - eLpNorm (s.indicator fun _ => c) p μ ≤ ‖c‖₊ * μ s ^ (1 / p.toReal) := by + eLpNorm (s.indicator fun _ => c) p μ ≤ ‖c‖ₑ * μ s ^ (1 / p.toReal) := by obtain rfl | hp := eq_or_ne p 0 · simp only [eLpNorm_exponent_zero, zero_le'] obtain rfl | h'p := eq_or_ne p ∞ @@ -872,9 +874,9 @@ lemma eLpNorm_indicator_const_le (c : G) (p : ℝ≥0∞) : calc eLpNorm (s.indicator fun _ => c) p μ ≤ eLpNorm (t.indicator fun _ => c) p μ := eLpNorm_mono (norm_indicator_le_of_subset (subset_toMeasurable _ _) _) - _ = ‖c‖₊ * μ t ^ (1 / p.toReal) := + _ = ‖c‖ₑ * μ t ^ (1 / p.toReal) := eLpNorm_indicator_const (measurableSet_toMeasurable ..) hp h'p - _ = ‖c‖₊ * μ s ^ (1 / p.toReal) := by rw [measure_toMeasurable] + _ = ‖c‖ₑ * μ s ^ (1 / p.toReal) := by rw [measure_toMeasurable] lemma Memℒp.indicator (hs : MeasurableSet s) (hf : Memℒp f p μ) : Memℒp (s.indicator f) p μ := ⟨hf.aestronglyMeasurable.indicator hs, lt_of_le_of_lt (eLpNorm_indicator_le f) hf.eLpNorm_lt_top⟩ @@ -916,18 +918,16 @@ protected lemma Memℒp.piecewise [DecidablePred (· ∈ s)] {g} (hs : Measurabl rw [eLpNorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top hp_zero hp_top, ← lintegral_add_compl _ hs, ENNReal.add_lt_top] constructor - · have h : ∀ᵐ x ∂μ, x ∈ s → - (‖Set.piecewise s f g x‖₊ : ℝ≥0∞) ^ p.toReal = (‖f x‖₊ : ℝ≥0∞) ^ p.toReal := by + · have h : ∀ᵐ x ∂μ, x ∈ s → ‖Set.piecewise s f g x‖ₑ ^ p.toReal = ‖f x‖ₑ ^ p.toReal := by filter_upwards with a ha using by simp [ha] rw [setLIntegral_congr_fun hs h] - exact lintegral_rpow_nnnorm_lt_top_of_eLpNorm_lt_top hp_zero hp_top hf.2 - · have h : ∀ᵐ x ∂μ, x ∈ sᶜ → - (‖Set.piecewise s f g x‖₊ : ℝ≥0∞) ^ p.toReal = (‖g x‖₊ : ℝ≥0∞) ^ p.toReal := by + exact lintegral_rpow_enorm_lt_top_of_eLpNorm_lt_top hp_zero hp_top hf.2 + · have h : ∀ᵐ x ∂μ, x ∈ sᶜ → ‖Set.piecewise s f g x‖ₑ ^ p.toReal = ‖g x‖ₑ ^ p.toReal := by filter_upwards with a ha have ha' : a ∉ s := ha simp [ha'] rw [setLIntegral_congr_fun hs.compl h] - exact lintegral_rpow_nnnorm_lt_top_of_eLpNorm_lt_top hp_zero hp_top hg.2 + exact lintegral_rpow_enorm_lt_top_of_eLpNorm_lt_top hp_zero hp_top hg.2 end Indicator @@ -938,12 +938,9 @@ theorem eLpNorm_restrict_eq_of_support_subset {s : Set α} {f : α → F} (hsf : by_cases hp0 : p = 0 · simp [hp0] by_cases hp_top : p = ∞ - · simp only [hp_top, eLpNorm_exponent_top, eLpNormEssSup_eq_essSup_nnnorm] - apply ENNReal.essSup_restrict_eq_of_support_subset - apply Function.support_subset_iff.2 (fun x hx ↦ ?_) - simp only [ne_eq, ENNReal.coe_eq_zero, nnnorm_eq_zero] at hx - exact Function.support_subset_iff.1 hsf x hx - · simp_rw [eLpNorm_eq_eLpNorm' hp0 hp_top, eLpNorm'_eq_lintegral_nnnorm] + · simp only [hp_top, eLpNorm_exponent_top, eLpNormEssSup_eq_essSup_enorm] + exact ENNReal.essSup_restrict_eq_of_support_subset fun x hx ↦ hsf <| enorm_ne_zero.1 hx + · simp_rw [eLpNorm_eq_eLpNorm' hp0 hp_top, eLpNorm'_eq_lintegral_enorm] congr 1 apply setLIntegral_eq_of_support_subset have : ¬(p.toReal ≤ 0) := by simpa only [not_le] using ENNReal.toReal_pos hp0 hp_top @@ -1047,7 +1044,7 @@ theorem Memℒp.smul_measure {f : α → E} {c : ℝ≥0∞} (hf : Memℒp f p theorem eLpNorm_one_add_measure (f : α → F) (μ ν : Measure α) : eLpNorm f 1 (μ + ν) = eLpNorm f 1 μ + eLpNorm f 1 ν := by - simp_rw [eLpNorm_one_eq_lintegral_nnnorm] + simp_rw [eLpNorm_one_eq_lintegral_enorm] rw [lintegral_add_measure _ μ ν] @[deprecated (since := "2024-07-27")] @@ -1067,11 +1064,11 @@ theorem eLpNorm_le_add_measure_left (f : α → F) (μ ν : Measure α) {p : ℝ @[deprecated (since := "2024-07-27")] alias snorm_le_add_measure_left := eLpNorm_le_add_measure_left -lemma eLpNormEssSup_eq_iSup (hμ : ∀ a, μ {a} ≠ 0) (f : α → E) : eLpNormEssSup f μ = ⨆ a, ↑‖f a‖₊ := +lemma eLpNormEssSup_eq_iSup (hμ : ∀ a, μ {a} ≠ 0) (f : α → E) : eLpNormEssSup f μ = ⨆ a, ‖f a‖ₑ := essSup_eq_iSup hμ _ @[simp] lemma eLpNormEssSup_count [MeasurableSingletonClass α] (f : α → E) : - eLpNormEssSup f .count = ⨆ a, ↑‖f a‖₊ := essSup_count _ + eLpNormEssSup f .count = ⨆ a, ‖f a‖ₑ := essSup_count _ theorem Memℒp.left_of_add_measure {f : α → E} (h : Memℒp f p (μ + ν)) : Memℒp f p μ := h.mono_measure <| Measure.le_add_right <| le_refl _ @@ -1101,22 +1098,12 @@ alias snorm'_eq_zero_of_ae_zero' := eLpNorm'_eq_zero_of_ae_zero' theorem ae_eq_zero_of_eLpNorm'_eq_zero {f : α → E} (hq0 : 0 ≤ q) (hf : AEStronglyMeasurable f μ) (h : eLpNorm' f q μ = 0) : f =ᵐ[μ] 0 := by - rw [eLpNorm'_eq_lintegral_nnnorm, ENNReal.rpow_eq_zero_iff] at h - cases h with - | inl h => - rw [lintegral_eq_zero_iff' (hf.ennnorm.pow_const q)] at h - refine h.left.mono fun x hx => ?_ - rw [Pi.zero_apply, ENNReal.rpow_eq_zero_iff] at hx - cases hx with - | inl hx => - cases' hx with hx _ - rwa [← ENNReal.coe_zero, ENNReal.coe_inj, nnnorm_eq_zero] at hx - | inr hx => - exact absurd hx.left ENNReal.coe_ne_top - | inr h => - exfalso - rw [one_div, inv_lt_zero] at h - exact hq0.not_lt h.right + simp only [eLpNorm'_eq_lintegral_enorm, lintegral_eq_zero_iff' (hf.enorm.pow_const q), one_div, + ENNReal.rpow_eq_zero_iff, inv_pos, inv_neg'', hq0.not_lt, and_false, or_false] at h + refine h.left.mono fun x hx ↦ ?_ + simp only [Pi.zero_apply, ENNReal.rpow_eq_zero_iff, enorm_eq_zero, enorm_ne_top, false_and, + or_false] at hx + exact hx.1 @[deprecated (since := "2024-07-27")] alias ae_eq_zero_of_snorm'_eq_zero := ae_eq_zero_of_eLpNorm'_eq_zero @@ -1129,15 +1116,15 @@ theorem eLpNorm'_eq_zero_iff (hq0_lt : 0 < q) {f : α → E} (hf : AEStronglyMea alias snorm'_eq_zero_iff := eLpNorm'_eq_zero_iff theorem coe_nnnorm_ae_le_eLpNormEssSup {_ : MeasurableSpace α} (f : α → F) (μ : Measure α) : - ∀ᵐ x ∂μ, (‖f x‖₊ : ℝ≥0∞) ≤ eLpNormEssSup f μ := - ENNReal.ae_le_essSup fun x => (‖f x‖₊ : ℝ≥0∞) + ∀ᵐ x ∂μ, ‖f x‖ₑ ≤ eLpNormEssSup f μ := + ENNReal.ae_le_essSup fun x => ‖f x‖ₑ @[deprecated (since := "2024-07-27")] alias coe_nnnorm_ae_le_snormEssSup := coe_nnnorm_ae_le_eLpNormEssSup @[simp] theorem eLpNormEssSup_eq_zero_iff {f : α → F} : eLpNormEssSup f μ = 0 ↔ f =ᵐ[μ] 0 := by - simp [EventuallyEq, eLpNormEssSup_eq_essSup_nnnorm] + simp [EventuallyEq, eLpNormEssSup_eq_essSup_enorm] @[deprecated (since := "2024-07-27")] alias snormEssSup_eq_zero_iff := eLpNormEssSup_eq_zero_iff @@ -1156,7 +1143,7 @@ theorem eLpNorm_eq_zero_of_ae_zero {f : α → E} (hf : f =ᵐ[μ] 0) : eLpNorm rw [← eLpNorm_zero (p := p) (μ := μ) (α := α) (F := E)] exact eLpNorm_congr_ae hf -theorem ae_le_eLpNormEssSup {f : α → F} : ∀ᵐ y ∂μ, ‖f y‖₊ ≤ eLpNormEssSup f μ := +theorem ae_le_eLpNormEssSup {f : α → F} : ∀ᵐ y ∂μ, ‖f y‖ₑ ≤ eLpNormEssSup f μ := ae_le_essSup @[deprecated (since := "2024-07-27")] @@ -1168,7 +1155,7 @@ lemma eLpNormEssSup_lt_top_iff_isBoundedUnder : simp_rw [← ENNReal.coe_le_coe, ENNReal.coe_toNNReal h.ne]; exact ae_le_eLpNormEssSup⟩ mpr := by rintro ⟨C, hC⟩; exact eLpNormEssSup_lt_top_of_ae_nnnorm_bound (C := C) hC -theorem meas_eLpNormEssSup_lt {f : α → F} : μ { y | eLpNormEssSup f μ < ‖f y‖₊ } = 0 := +theorem meas_eLpNormEssSup_lt {f : α → F} : μ { y | eLpNormEssSup f μ < ‖f y‖ₑ } = 0 := meas_essSup_lt @[deprecated (since := "2024-07-27")] @@ -1182,7 +1169,7 @@ lemma eLpNorm_lt_top_of_finite [Finite α] [IsFiniteMeasure μ] : eLpNorm f p μ exact .le_of_finite rw [eLpNorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top hp₀ hp] refine IsFiniteMeasure.lintegral_lt_top_of_bounded_to_ennreal μ ?_ - simp_rw [← ENNReal.coe_rpow_of_nonneg _ ENNReal.toReal_nonneg] + simp_rw [enorm, ← ENNReal.coe_rpow_of_nonneg _ ENNReal.toReal_nonneg] norm_cast exact Finite.exists_le _ @@ -1205,7 +1192,7 @@ variable {β : Type*} {mβ : MeasurableSpace β} {f : α → β} {g : β → E} theorem eLpNormEssSup_map_measure (hg : AEStronglyMeasurable g (Measure.map f μ)) (hf : AEMeasurable f μ) : eLpNormEssSup g (Measure.map f μ) = eLpNormEssSup (g ∘ f) μ := - essSup_map_measure hg.ennnorm hf + essSup_map_measure hg.enorm hf @[deprecated (since := "2024-07-27")] alias snormEssSup_map_measure := eLpNormEssSup_map_measure @@ -1217,8 +1204,8 @@ theorem eLpNorm_map_measure (hg : AEStronglyMeasurable g (Measure.map f μ)) by_cases hp_top : p = ∞ · simp_rw [hp_top, eLpNorm_exponent_top] exact eLpNormEssSup_map_measure hg hf - simp_rw [eLpNorm_eq_lintegral_rpow_nnnorm hp_zero hp_top] - rw [lintegral_map' (hg.ennnorm.pow_const p.toReal) hf] + simp_rw [eLpNorm_eq_lintegral_rpow_enorm hp_zero hp_top] + rw [lintegral_map' (hg.enorm.pow_const p.toReal) hf] rfl @[deprecated (since := "2024-07-27")] @@ -1267,7 +1254,7 @@ theorem _root_.MeasurableEmbedding.eLpNorm_map_measure {g : β → F} (hf : Meas by_cases hp : p = ∞ · simp_rw [hp, eLpNorm_exponent_top] exact hf.essSup_map_measure - · simp_rw [eLpNorm_eq_lintegral_rpow_nnnorm hp_zero hp] + · simp_rw [eLpNorm_eq_lintegral_rpow_enorm hp_zero hp] rw [hf.lintegral_map] rfl @@ -1289,10 +1276,10 @@ section Monotonicity theorem eLpNorm'_le_nnreal_smul_eLpNorm'_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) {p : ℝ} (hp : 0 < p) : eLpNorm' f p μ ≤ c • eLpNorm' g p μ := by - simp_rw [eLpNorm'_eq_lintegral_nnnorm] + simp_rw [eLpNorm'_eq_lintegral_enorm] rw [← ENNReal.rpow_le_rpow_iff hp, ENNReal.smul_def, smul_eq_mul, ENNReal.mul_rpow_of_nonneg _ _ hp.le] - simp_rw [← ENNReal.rpow_mul, one_div, inv_mul_cancel₀ hp.ne.symm, ENNReal.rpow_one, + simp_rw [← ENNReal.rpow_mul, one_div, inv_mul_cancel₀ hp.ne.symm, ENNReal.rpow_one, enorm, ← ENNReal.coe_rpow_of_nonneg _ hp.le, ← lintegral_const_mul' _ _ ENNReal.coe_ne_top, ← ENNReal.coe_mul] apply lintegral_mono_ae @@ -1305,10 +1292,10 @@ alias snorm'_le_nnreal_smul_snorm'_of_ae_le_mul := eLpNorm'_le_nnreal_smul_eLpNo theorem eLpNormEssSup_le_nnreal_smul_eLpNormEssSup_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : eLpNormEssSup f μ ≤ c • eLpNormEssSup g μ := calc - essSup (fun x => (‖f x‖₊ : ℝ≥0∞)) μ ≤ essSup (fun x => (↑(c * ‖g x‖₊) : ℝ≥0∞)) μ := + essSup (‖f ·‖ₑ) μ ≤ essSup (fun x => (↑(c * ‖g x‖₊) : ℝ≥0∞)) μ := essSup_mono_ae <| h.mono fun _ hx => ENNReal.coe_le_coe.mpr hx - _ = essSup (fun x => (c * ‖g x‖₊ : ℝ≥0∞)) μ := by simp_rw [ENNReal.coe_mul] - _ = c • essSup (fun x => (‖g x‖₊ : ℝ≥0∞)) μ := ENNReal.essSup_const_mul + _ = essSup (c * ‖g ·‖ₑ) μ := by simp_rw [ENNReal.coe_mul, enorm] + _ = c • essSup (‖g ·‖ₑ) μ := ENNReal.essSup_const_mul @[deprecated (since := "2024-07-27")] alias snormEssSup_le_nnreal_smul_snormEssSup_of_ae_le_mul := @@ -1384,20 +1371,20 @@ section BoundedSMul variable {𝕜 : Type*} [NormedRing 𝕜] [MulActionWithZero 𝕜 E] [MulActionWithZero 𝕜 F] variable [BoundedSMul 𝕜 E] [BoundedSMul 𝕜 F] {c : 𝕜} {f : α → F} -theorem eLpNorm'_const_smul_le (hq : 0 < q) : eLpNorm' (c • f) q μ ≤ ‖c‖₊ • eLpNorm' f q μ := +theorem eLpNorm'_const_smul_le (hq : 0 < q) : eLpNorm' (c • f) q μ ≤ ‖c‖ₑ * eLpNorm' f q μ := eLpNorm'_le_nnreal_smul_eLpNorm'_of_ae_le_mul (Eventually.of_forall fun _ => nnnorm_smul_le ..) hq @[deprecated (since := "2024-07-27")] alias snorm'_const_smul_le := eLpNorm'_const_smul_le -theorem eLpNormEssSup_const_smul_le : eLpNormEssSup (c • f) μ ≤ ‖c‖₊ • eLpNormEssSup f μ := +theorem eLpNormEssSup_const_smul_le : eLpNormEssSup (c • f) μ ≤ ‖c‖ₑ * eLpNormEssSup f μ := eLpNormEssSup_le_nnreal_smul_eLpNormEssSup_of_ae_le_mul (Eventually.of_forall fun _ => by simp [nnnorm_smul_le]) @[deprecated (since := "2024-07-27")] alias snormEssSup_const_smul_le := eLpNormEssSup_const_smul_le -theorem eLpNorm_const_smul_le : eLpNorm (c • f) p μ ≤ ‖c‖₊ • eLpNorm f p μ := +theorem eLpNorm_const_smul_le : eLpNorm (c • f) p μ ≤ ‖c‖ₑ * eLpNorm f p μ := eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul (Eventually.of_forall fun _ => by simp [nnnorm_smul_le]) _ @@ -1425,29 +1412,31 @@ variable {𝕜 : Type*} [NormedDivisionRing 𝕜] [MulActionWithZero 𝕜 E] [Mo variable [BoundedSMul 𝕜 E] [BoundedSMul 𝕜 F] theorem eLpNorm'_const_smul {f : α → F} (c : 𝕜) (hq_pos : 0 < q) : - eLpNorm' (c • f) q μ = ‖c‖₊ • eLpNorm' f q μ := by + eLpNorm' (c • f) q μ = ‖c‖ₑ * eLpNorm' f q μ := by obtain rfl | hc := eq_or_ne c 0 - · simp [eLpNorm'_eq_lintegral_nnnorm, hq_pos] - refine le_antisymm (eLpNorm'_const_smul_le hq_pos) ?_ - simpa [hc, le_inv_smul_iff_of_pos] using eLpNorm'_const_smul_le (c := c⁻¹) (f := c • f) hq_pos + · simp [eLpNorm'_eq_lintegral_enorm, hq_pos] + refine le_antisymm (eLpNorm'_const_smul_le hq_pos) <| ENNReal.mul_le_of_le_div' ?_ + simpa [enorm_inv, hc, ENNReal.div_eq_inv_mul] + using eLpNorm'_const_smul_le (c := c⁻¹) (f := c • f) hq_pos @[deprecated (since := "2024-07-27")] alias snorm'_const_smul := eLpNorm'_const_smul theorem eLpNormEssSup_const_smul (c : 𝕜) (f : α → F) : - eLpNormEssSup (c • f) μ = (‖c‖₊ : ℝ≥0∞) * eLpNormEssSup f μ := by - simp_rw [eLpNormEssSup_eq_essSup_nnnorm, Pi.smul_apply, nnnorm_smul, ENNReal.coe_mul, + eLpNormEssSup (c • f) μ = ‖c‖ₑ * eLpNormEssSup f μ := by + simp_rw [eLpNormEssSup_eq_essSup_enorm, Pi.smul_apply, enorm_smul, ENNReal.essSup_const_mul] @[deprecated (since := "2024-07-27")] alias snormEssSup_const_smul := eLpNormEssSup_const_smul theorem eLpNorm_const_smul (c : 𝕜) (f : α → F) (p : ℝ≥0∞) (μ : Measure α): - eLpNorm (c • f) p μ = (‖c‖₊ : ℝ≥0∞) * eLpNorm f p μ := by + eLpNorm (c • f) p μ = ‖c‖ₑ * eLpNorm f p μ := by obtain rfl | hc := eq_or_ne c 0 · simp - refine le_antisymm eLpNorm_const_smul_le ?_ - simpa [hc, le_inv_smul_iff_of_pos] using eLpNorm_const_smul_le (c := c⁻¹) (f := c • f) + refine le_antisymm eLpNorm_const_smul_le <| ENNReal.mul_le_of_le_div' ?_ + simpa [enorm_inv, hc, ENNReal.div_eq_inv_mul] + using eLpNorm_const_smul_le (c := c⁻¹) (f := c • f) @[deprecated (since := "2024-07-27")] alias snorm_const_smul := eLpNorm_const_smul @@ -1461,7 +1450,7 @@ end NormedSpace theorem le_eLpNorm_of_bddBelow (hp : p ≠ 0) (hp' : p ≠ ∞) {f : α → F} (C : ℝ≥0) {s : Set α} (hs : MeasurableSet s) (hf : ∀ᵐ x ∂μ, x ∈ s → C ≤ ‖f x‖₊) : C • μ s ^ (1 / p.toReal) ≤ eLpNorm f p μ := by - rw [ENNReal.smul_def, smul_eq_mul, eLpNorm_eq_lintegral_rpow_nnnorm hp hp', + rw [ENNReal.smul_def, smul_eq_mul, eLpNorm_eq_lintegral_rpow_enorm hp hp', one_div, ENNReal.le_rpow_inv_iff (ENNReal.toReal_pos hp hp'), ENNReal.mul_rpow_of_nonneg _ _ ENNReal.toReal_nonneg, ← ENNReal.rpow_mul, inv_mul_cancel₀ (ENNReal.toReal_pos hp hp').ne.symm, ENNReal.rpow_one, ← setLIntegral_const, @@ -1469,9 +1458,9 @@ theorem le_eLpNorm_of_bddBelow (hp : p ≠ 0) (hp' : p ≠ ∞) {f : α → F} ( refine lintegral_mono_ae ?_ filter_upwards [hf] with x hx by_cases hxs : x ∈ s - · simp only [Set.indicator_of_mem hxs] at hx ⊢ + · simp only [Set.indicator_of_mem, hxs, true_implies] at hx ⊢ gcongr - exact hx hxs + rwa [coe_le_enorm] · simp [Set.indicator_of_not_mem hxs] @[deprecated (since := "2024-07-27")] @@ -1511,7 +1500,7 @@ variable [MeasurableSpace E] [OpensMeasurableSpace E] {R : ℝ≥0} theorem ae_bdd_liminf_atTop_rpow_of_eLpNorm_bdd {p : ℝ≥0∞} {f : ℕ → α → E} (hfmeas : ∀ n, Measurable (f n)) (hbdd : ∀ n, eLpNorm (f n) p μ ≤ R) : - ∀ᵐ x ∂μ, liminf (fun n => ((‖f n x‖₊ : ℝ≥0∞) ^ p.toReal : ℝ≥0∞)) atTop < ∞ := by + ∀ᵐ x ∂μ, liminf (fun n => ((‖f n x‖ₑ) ^ p.toReal : ℝ≥0∞)) atTop < ∞ := by by_cases hp0 : p.toReal = 0 · simp only [hp0, ENNReal.rpow_zero] filter_upwards with _ @@ -1526,7 +1515,7 @@ theorem ae_bdd_liminf_atTop_rpow_of_eLpNorm_bdd {p : ℝ≥0∞} {f : ℕ → α (lt_of_le_of_lt ?_ (ENNReal.rpow_lt_top_of_nonneg ENNReal.toReal_nonneg ENNReal.coe_ne_top : (R : ℝ≥0∞) ^ p.toReal < ∞))).ne - simp_rw [eLpNorm_eq_lintegral_rpow_nnnorm hp hp', one_div] at hbdd + simp_rw [eLpNorm_eq_lintegral_rpow_enorm hp hp', one_div] at hbdd simp_rw [liminf_eq, eventually_atTop] exact sSup_le fun b ⟨a, ha⟩ => @@ -1537,11 +1526,11 @@ alias ae_bdd_liminf_atTop_rpow_of_snorm_bdd := ae_bdd_liminf_atTop_rpow_of_eLpNo theorem ae_bdd_liminf_atTop_of_eLpNorm_bdd {p : ℝ≥0∞} (hp : p ≠ 0) {f : ℕ → α → E} (hfmeas : ∀ n, Measurable (f n)) (hbdd : ∀ n, eLpNorm (f n) p μ ≤ R) : - ∀ᵐ x ∂μ, liminf (fun n => (‖f n x‖₊ : ℝ≥0∞)) atTop < ∞ := by + ∀ᵐ x ∂μ, liminf (fun n => (‖f n x‖ₑ)) atTop < ∞ := by by_cases hp' : p = ∞ · subst hp' simp_rw [eLpNorm_exponent_top] at hbdd - have : ∀ n, ∀ᵐ x ∂μ, (‖f n x‖₊ : ℝ≥0∞) < R + 1 := fun n => + have : ∀ n, ∀ᵐ x ∂μ, (‖f n x‖ₑ) < R + 1 := fun n => ae_lt_of_essSup_lt (lt_of_le_of_lt (hbdd n) <| ENNReal.lt_add_right ENNReal.coe_ne_top one_ne_zero) rw [← ae_all_iff] at this @@ -1551,15 +1540,15 @@ theorem ae_bdd_liminf_atTop_of_eLpNorm_bdd {p : ℝ≥0∞} (hp : p ≠ 0) {f : filter_upwards [ae_bdd_liminf_atTop_rpow_of_eLpNorm_bdd hfmeas hbdd] with x hx have hppos : 0 < p.toReal := ENNReal.toReal_pos hp hp' have : - liminf (fun n => (‖f n x‖₊ : ℝ≥0∞) ^ p.toReal) atTop = - liminf (fun n => (‖f n x‖₊ : ℝ≥0∞)) atTop ^ p.toReal := by + liminf (fun n => (‖f n x‖ₑ) ^ p.toReal) atTop = + liminf (fun n => (‖f n x‖ₑ)) atTop ^ p.toReal := by change - liminf (fun n => ENNReal.orderIsoRpow p.toReal hppos (‖f n x‖₊ : ℝ≥0∞)) atTop = - ENNReal.orderIsoRpow p.toReal hppos (liminf (fun n => (‖f n x‖₊ : ℝ≥0∞)) atTop) + liminf (fun n => ENNReal.orderIsoRpow p.toReal hppos (‖f n x‖ₑ)) atTop = + ENNReal.orderIsoRpow p.toReal hppos (liminf (fun n => (‖f n x‖ₑ)) atTop) refine (OrderIso.liminf_apply (ENNReal.orderIsoRpow p.toReal _) ?_ ?_ ?_ ?_).symm <;> isBoundedDefault rw [this] at hx - rw [← ENNReal.rpow_one (liminf (fun n => ‖f n x‖₊) atTop), ← mul_inv_cancel₀ hppos.ne.symm, + rw [← ENNReal.rpow_one (liminf (‖f · x‖ₑ) atTop), ← mul_inv_cancel₀ hppos.ne.symm, ENNReal.rpow_mul] exact ENNReal.rpow_lt_top_of_nonneg (inv_nonneg.2 hppos.le) hx.ne @@ -1587,13 +1576,13 @@ theorem Memℒp.exists_eLpNorm_indicator_compl_lt {β : Type*} [NormedAddCommGro rcases eq_or_ne p 0 with rfl | hp₀ · use ∅; simp [pos_iff_ne_zero.2 hε] -- first take care of `p = 0` · obtain ⟨s, hsm, hs, hε⟩ : - ∃ s, MeasurableSet s ∧ μ s < ∞ ∧ ∫⁻ a in sᶜ, (‖f a‖₊) ^ p.toReal ∂μ < ε ^ p.toReal := by + ∃ s, MeasurableSet s ∧ μ s < ∞ ∧ ∫⁻ a in sᶜ, (‖f a‖ₑ) ^ p.toReal ∂μ < ε ^ p.toReal := by apply exists_setLintegral_compl_lt · exact ((eLpNorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top hp₀ hp_top).1 hf.2).ne · simp [*] refine ⟨s, hsm, hs, ?_⟩ rwa [eLpNorm_indicator_eq_eLpNorm_restrict hsm.compl, - eLpNorm_eq_lintegral_rpow_nnnorm hp₀ hp_top, one_div, ENNReal.rpow_inv_lt_iff] + eLpNorm_eq_lintegral_rpow_enorm hp₀ hp_top, one_div, ENNReal.rpow_inv_lt_iff] simp [ENNReal.toReal_pos, *] @[deprecated (since := "2024-07-27")] diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean index 6248e4a2c59258..968d77d53489ec 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean @@ -21,9 +21,9 @@ variable {α E : Type*} {m0 : MeasurableSpace α} [NormedAddCommGroup E] theorem pow_mul_meas_ge_le_eLpNorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : AEStronglyMeasurable f μ) (ε : ℝ≥0∞) : (ε * μ { x | ε ≤ (‖f x‖₊ : ℝ≥0∞) ^ p.toReal }) ^ (1 / p.toReal) ≤ eLpNorm f p μ := by - rw [eLpNorm_eq_lintegral_rpow_nnnorm hp_ne_zero hp_ne_top] + rw [eLpNorm_eq_lintegral_rpow_enorm hp_ne_zero hp_ne_top] gcongr - exact mul_meas_ge_le_lintegral₀ (hf.ennnorm.pow_const _) ε + exact mul_meas_ge_le_lintegral₀ (hf.enorm.pow_const _) ε @[deprecated (since := "2024-07-27")] alias pow_mul_meas_ge_le_snorm := pow_mul_meas_ge_le_eLpNorm diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean index 649bc7c74f6ed3..e56b6c50caf6d7 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean @@ -30,16 +30,16 @@ theorem eLpNorm'_le_eLpNorm'_mul_rpow_measure_univ {p q : ℝ} (hp0_lt : 0 < p) · rw [hpq_eq, sub_self, ENNReal.rpow_zero, mul_one] have hpq : p < q := lt_of_le_of_ne hpq hpq_eq let g := fun _ : α => (1 : ℝ≥0∞) - have h_rw : (∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ p ∂μ) = ∫⁻ a, ((‖f a‖₊ : ℝ≥0∞) * g a) ^ p ∂μ := + have h_rw : (∫⁻ a, ‖f a‖ₑ ^ p ∂μ) = ∫⁻ a, (‖f a‖ₑ * g a) ^ p ∂μ := lintegral_congr fun a => by simp [g] - repeat' rw [eLpNorm'_eq_lintegral_nnnorm] + repeat' rw [eLpNorm'_eq_lintegral_enorm] rw [h_rw] let r := p * q / (q - p) have hpqr : 1 / p = 1 / q + 1 / r := by field_simp [r, hp0_lt.ne', hq0_lt.ne'] calc (∫⁻ a : α, (↑‖f a‖₊ * g a) ^ p ∂μ) ^ (1 / p) ≤ (∫⁻ a : α, ↑‖f a‖₊ ^ q ∂μ) ^ (1 / q) * (∫⁻ a : α, g a ^ r ∂μ) ^ (1 / r) := - ENNReal.lintegral_Lp_mul_le_Lq_mul_Lr hp0_lt hpq hpqr μ hf.ennnorm aemeasurable_const + ENNReal.lintegral_Lp_mul_le_Lq_mul_Lr hp0_lt hpq hpqr μ hf.enorm aemeasurable_const _ = (∫⁻ a : α, ↑‖f a‖₊ ^ q ∂μ) ^ (1 / q) * μ Set.univ ^ (1 / p - 1 / q) := by rw [hpqr]; simp [r, g] @@ -48,7 +48,7 @@ alias snorm'_le_snorm'_mul_rpow_measure_univ := eLpNorm'_le_eLpNorm'_mul_rpow_me theorem eLpNorm'_le_eLpNormEssSup_mul_rpow_measure_univ {q : ℝ} (hq_pos : 0 < q) : eLpNorm' f q μ ≤ eLpNormEssSup f μ * μ Set.univ ^ (1 / q) := by - have h_le : (∫⁻ a : α, (‖f a‖₊ : ℝ≥0∞) ^ q ∂μ) ≤ ∫⁻ _ : α, eLpNormEssSup f μ ^ q ∂μ := by + have h_le : (∫⁻ a : α, ‖f a‖ₑ ^ q ∂μ) ≤ ∫⁻ _ : α, eLpNormEssSup f μ ^ q ∂μ := by refine lintegral_mono_ae ?_ have h_nnnorm_le_eLpNorm_ess_sup := coe_nnnorm_ae_le_eLpNormEssSup f μ exact h_nnnorm_le_eLpNorm_ess_sup.mono fun x hx => by gcongr @@ -198,7 +198,7 @@ theorem eLpNorm_le_eLpNorm_top_mul_eLpNorm (p : ℝ≥0∞) (f : α → E) {g : exact ha by_cases hp_zero : p = 0 · simp only [hp_zero, eLpNorm_exponent_zero, mul_zero, le_zero_iff] - simp_rw [eLpNorm_eq_lintegral_rpow_nnnorm hp_zero hp_top, eLpNorm_exponent_top, eLpNormEssSup] + simp_rw [eLpNorm_eq_lintegral_rpow_enorm hp_zero hp_top, eLpNorm_exponent_top, eLpNormEssSup] calc (∫⁻ x, (‖b (f x) (g x)‖₊ : ℝ≥0∞) ^ p.toReal ∂μ) ^ (1 / p.toReal) ≤ (∫⁻ x, (‖f x‖₊ : ℝ≥0∞) ^ p.toReal * (‖g x‖₊ : ℝ≥0∞) ^ p.toReal ∂μ) ^ (1 / p.toReal) := by @@ -255,7 +255,7 @@ theorem eLpNorm'_le_eLpNorm'_mul_eLpNorm' {p q r : ℝ} (hf : AEStronglyMeasurab h.mono fun a ha => (ENNReal.rpow_le_rpow_iff hp0_lt).mpr <| ENNReal.coe_le_coe.mpr <| ha _ ≤ _ := ?_ simp_rw [eLpNorm', ENNReal.coe_mul] - exact ENNReal.lintegral_Lp_mul_le_Lq_mul_Lr hp0_lt hpq hpqr μ hf.ennnorm hg.ennnorm + exact ENNReal.lintegral_Lp_mul_le_Lq_mul_Lr hp0_lt hpq hpqr μ hf.enorm hg.enorm @[deprecated (since := "2024-07-27")] alias snorm'_le_snorm'_mul_snorm' := eLpNorm'_le_eLpNorm'_mul_eLpNorm' diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean index bde7165f033fba..74b789bb4680a2 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean @@ -28,7 +28,7 @@ theorem eLpNorm'_add_le (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasura (∫⁻ a, ((fun a => (‖f a‖₊ : ℝ≥0∞)) + fun a => (‖g a‖₊ : ℝ≥0∞)) a ^ q ∂μ) ^ (1 / q) := by gcongr with a simp only [Pi.add_apply, ← ENNReal.coe_add, ENNReal.coe_le_coe, nnnorm_add_le] - _ ≤ eLpNorm' f q μ + eLpNorm' g q μ := ENNReal.lintegral_Lp_add_le hf.ennnorm hg.ennnorm hq1 + _ ≤ eLpNorm' f q μ + eLpNorm' g q μ := ENNReal.lintegral_Lp_add_le hf.enorm hg.enorm hq1 @[deprecated (since := "2024-07-27")] alias snorm'_add_le := eLpNorm'_add_le @@ -41,7 +41,7 @@ theorem eLpNorm'_add_le_of_le_one (hf : AEStronglyMeasurable f μ) (hq0 : 0 ≤ gcongr with a simp only [Pi.add_apply, ← ENNReal.coe_add, ENNReal.coe_le_coe, nnnorm_add_le] _ ≤ (2 : ℝ≥0∞) ^ (1 / q - 1) * (eLpNorm' f q μ + eLpNorm' g q μ) := - ENNReal.lintegral_Lp_add_le_of_le_one hf.ennnorm hq0 hq1 + ENNReal.lintegral_Lp_add_le_of_le_one hf.enorm hq0 hq1 @[deprecated (since := "2024-07-27")] alias snorm'_add_le_of_le_one := eLpNorm'_add_le_of_le_one diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean index b6776ff95e87ac..0cb5cd60fe951f 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean @@ -51,7 +51,7 @@ theorem essSup_trim (hm : m ≤ m0) {f : α → ℝ≥0∞} (hf : Measurable[m] theorem eLpNormEssSup_trim (hm : m ≤ m0) {f : α → E} (hf : StronglyMeasurable[m] f) : eLpNormEssSup f (μ.trim hm) = eLpNormEssSup f μ := - essSup_trim _ (@StronglyMeasurable.ennnorm _ m _ _ _ hf) + essSup_trim _ (@StronglyMeasurable.enorm _ m _ _ _ hf) @[deprecated (since := "2024-07-27")] alias snormEssSup_trim := eLpNormEssSup_trim diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 7ff9c55c9cd7d8..64caa8bea97f1f 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -248,10 +248,12 @@ theorem nnnorm_def (f : Lp E p μ) : ‖f‖₊ = ENNReal.toNNReal (eLpNorm f p protected theorem coe_nnnorm (f : Lp E p μ) : (‖f‖₊ : ℝ) = ‖f‖ := rfl -@[simp, norm_cast] -theorem nnnorm_coe_ennreal (f : Lp E p μ) : (‖f‖₊ : ℝ≥0∞) = eLpNorm f p μ := +@[simp] +theorem enorm_def (f : Lp E p μ) : ‖f‖ₑ = eLpNorm f p μ := ENNReal.coe_toNNReal <| Lp.eLpNorm_ne_top f +@[deprecated (since := "2025-01-20")] alias nnnorm_coe_ennreal := enorm_def + @[simp] lemma norm_toLp (f : α → E) (hf : Memℒp f p μ) : ‖hf.toLp f‖ = ENNReal.toReal (eLpNorm f p μ) := by rw [norm_def, eLpNorm_congr_ae (Memℒp.coeFn_toLp hf)] @@ -261,8 +263,10 @@ theorem nnnorm_toLp (f : α → E) (hf : Memℒp f p μ) : ‖hf.toLp f‖₊ = ENNReal.toNNReal (eLpNorm f p μ) := NNReal.eq <| norm_toLp f hf -theorem coe_nnnorm_toLp {f : α → E} (hf : Memℒp f p μ) : (‖hf.toLp f‖₊ : ℝ≥0∞) = eLpNorm f p μ := by - rw [nnnorm_toLp f hf, ENNReal.coe_toNNReal hf.2.ne] +lemma enorm_toLp {f : α → E} (hf : Memℒp f p μ) : ‖hf.toLp f‖ₑ = eLpNorm f p μ := by + simp [enorm, nnnorm_toLp f hf, ENNReal.coe_toNNReal hf.2.ne] + +@[deprecated (since := "2025-01-20")] alias coe_nnnorm_toLp := enorm_toLp theorem dist_def (f g : Lp E p μ) : dist f g = (eLpNorm (⇑f - ⇑g) p μ).toReal := by simp_rw [dist, norm_def] @@ -402,8 +406,8 @@ instance instNormedAddCommGroup [hp : Fact (1 ≤ p)] : NormedAddCommGroup (Lp E map_zero' := norm_zero neg' := by simp add_le' := fun f g => by - suffices (‖f + g‖₊ : ℝ≥0∞) ≤ ‖f‖₊ + ‖g‖₊ from mod_cast this - simp only [Lp.nnnorm_coe_ennreal] + suffices ‖f + g‖ₑ ≤ ‖f‖ₑ + ‖g‖ₑ by simpa [← ENNReal.coe_add, enorm] using this + simp only [Lp.enorm_def] exact (eLpNorm_congr_ae (AEEqFun.coeFn_add _ _)).trans_le (eLpNorm_add_le (Lp.aestronglyMeasurable _) (Lp.aestronglyMeasurable _) hp.out) eq_zero_of_map_eq_zero' := fun _ => @@ -426,9 +430,7 @@ variable [BoundedSMul 𝕜 E] [BoundedSMul 𝕜' E] theorem const_smul_mem_Lp (c : 𝕜) (f : Lp E p μ) : c • (f : α →ₘ[μ] E) ∈ Lp E p μ := by rw [mem_Lp_iff_eLpNorm_lt_top, eLpNorm_congr_ae (AEEqFun.coeFn_smul _ _)] - refine eLpNorm_const_smul_le.trans_lt ?_ - rw [ENNReal.smul_def, smul_eq_mul, ENNReal.mul_lt_top_iff] - exact Or.inl ⟨ENNReal.coe_lt_top, f.prop⟩ + exact eLpNorm_const_smul_le.trans_lt <| ENNReal.mul_lt_top ENNReal.coe_lt_top f.prop variable (E p μ 𝕜) @@ -461,10 +463,9 @@ instance instIsScalarTower [SMul 𝕜 𝕜'] [IsScalarTower 𝕜 𝕜' E] : IsSc instance instBoundedSMul [Fact (1 ≤ p)] : BoundedSMul 𝕜 (Lp E p μ) := -- TODO: add `BoundedSMul.of_nnnorm_smul_le` BoundedSMul.of_norm_smul_le fun r f => by - suffices (‖r • f‖₊ : ℝ≥0∞) ≤ ‖r‖₊ * ‖f‖₊ from mod_cast this - rw [nnnorm_def, nnnorm_def, ENNReal.coe_toNNReal (Lp.eLpNorm_ne_top _), - eLpNorm_congr_ae (coeFn_smul _ _), ENNReal.coe_toNNReal (Lp.eLpNorm_ne_top _)] - exact eLpNorm_const_smul_le + suffices ‖r • f‖ₑ ≤ ‖r‖ₑ * ‖f‖ₑ by simpa [← ENNReal.coe_mul, enorm] using this + simpa only [eLpNorm_congr_ae (coeFn_smul _ _), enorm_def] + using eLpNorm_const_smul_le (c := r) (f := f) (p := p) end BoundedSMul @@ -536,8 +537,7 @@ theorem exists_eLpNorm_indicator_le (hp : p ≠ ∞) (c : E) {ε : ℝ≥0∞} ( have hp₀ : 0 < p := bot_lt_iff_ne_bot.2 h'p have hp₀' : 0 ≤ 1 / p.toReal := div_nonneg zero_le_one ENNReal.toReal_nonneg have hp₀'' : 0 < p.toReal := ENNReal.toReal_pos hp₀.ne' hp - obtain ⟨η, hη_pos, hη_le⟩ : - ∃ η : ℝ≥0, 0 < η ∧ (‖c‖₊ : ℝ≥0∞) * (η : ℝ≥0∞) ^ (1 / p.toReal) ≤ ε := by + obtain ⟨η, hη_pos, hη_le⟩ : ∃ η : ℝ≥0, 0 < η ∧ ‖c‖ₑ * (η : ℝ≥0∞) ^ (1 / p.toReal) ≤ ε := by have : Filter.Tendsto (fun x : ℝ≥0 => ((‖c‖₊ * x ^ (1 / p.toReal) : ℝ≥0) : ℝ≥0∞)) (𝓝 0) (𝓝 (0 : ℝ≥0)) := by @@ -548,8 +548,7 @@ theorem exists_eLpNorm_indicator_le (hp : p ≠ ∞) (c : E) {ε : ℝ≥0∞} ( obtain ⟨δ, hδ, hδε'⟩ := NNReal.nhds_zero_basis.eventually_iff.mp (this.eventually_le_const hε') obtain ⟨η, hη, hηδ⟩ := exists_between hδ refine ⟨η, hη, ?_⟩ - rw [← ENNReal.coe_rpow_of_nonneg _ hp₀', ← ENNReal.coe_mul] - exact hδε' hηδ + simpa only [← ENNReal.coe_rpow_of_nonneg _ hp₀', enorm, ← ENNReal.coe_mul] using hδε' hηδ refine ⟨η, hη_pos, fun s hs => ?_⟩ refine (eLpNorm_indicator_const_le _ _).trans (le_trans ?_ hη_le) exact mul_le_mul_left' (ENNReal.rpow_le_rpow hs hp₀') _ @@ -614,13 +613,13 @@ theorem norm_indicatorConstLp (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : ‖indicatorConstLp p hs hμs c‖ = ‖c‖ * (μ s).toReal ^ (1 / p.toReal) := by rw [Lp.norm_def, eLpNorm_congr_ae indicatorConstLp_coeFn, eLpNorm_indicator_const hs hp_ne_zero hp_ne_top, ENNReal.toReal_mul, ENNReal.toReal_rpow, - ENNReal.coe_toReal, coe_nnnorm] + toReal_enorm] theorem norm_indicatorConstLp_top (hμs_ne_zero : μ s ≠ 0) : ‖indicatorConstLp ∞ hs hμs c‖ = ‖c‖ := by rw [Lp.norm_def, eLpNorm_congr_ae indicatorConstLp_coeFn, eLpNorm_indicator_const' hs hμs_ne_zero ENNReal.top_ne_zero, ENNReal.top_toReal, - _root_.div_zero, ENNReal.rpow_zero, mul_one, ENNReal.coe_toReal, coe_nnnorm] + _root_.div_zero, ENNReal.rpow_zero, mul_one, toReal_enorm] theorem norm_indicatorConstLp' (hp_pos : p ≠ 0) (hμs_pos : μ s ≠ 0) : ‖indicatorConstLp p hs hμs c‖ = ‖c‖ * (μ s).toReal ^ (1 / p.toReal) := by @@ -634,29 +633,33 @@ theorem norm_indicatorConstLp_le : rw [indicatorConstLp, Lp.norm_toLp] refine ENNReal.toReal_le_of_le_ofReal (by positivity) ?_ refine (eLpNorm_indicator_const_le _ _).trans_eq ?_ - rw [← coe_nnnorm, ENNReal.ofReal_mul (NNReal.coe_nonneg _), ENNReal.ofReal_coe_nnreal, - ENNReal.toReal_rpow, ENNReal.ofReal_toReal] + rw [ENNReal.ofReal_mul (norm_nonneg _), ofReal_norm, ENNReal.toReal_rpow, ENNReal.ofReal_toReal] exact ENNReal.rpow_ne_top_of_nonneg (by positivity) hμs theorem nnnorm_indicatorConstLp_le : ‖indicatorConstLp p hs hμs c‖₊ ≤ ‖c‖₊ * (μ s).toNNReal ^ (1 / p.toReal) := norm_indicatorConstLp_le -theorem ennnorm_indicatorConstLp_le : - (‖indicatorConstLp p hs hμs c‖₊ : ℝ≥0∞) ≤ ‖c‖₊ * (μ s) ^ (1 / p.toReal) := by - refine (ENNReal.coe_le_coe.mpr nnnorm_indicatorConstLp_le).trans_eq ?_ - simp [ENNReal.coe_rpow_of_nonneg, ENNReal.coe_toNNReal hμs] +theorem enorm_indicatorConstLp_le : + ‖indicatorConstLp p hs hμs c‖ₑ ≤ ‖c‖ₑ * μ s ^ (1 / p.toReal) := by + simpa [ENNReal.coe_rpow_of_nonneg, ENNReal.coe_toNNReal hμs, Lp.enorm_def, ← enorm_eq_nnnorm] + using ENNReal.coe_le_coe.2 <| nnnorm_indicatorConstLp_le (c := c) (hμs := hμs) + +@[deprecated (since := "2025-01-20")] alias ennnorm_indicatorConstLp_le := enorm_indicatorConstLp_le -theorem edist_indicatorConstLp_eq_nnnorm {t : Set α} {ht : MeasurableSet t} {hμt : μ t ≠ ∞} : +theorem edist_indicatorConstLp_eq_enorm {t : Set α} {ht : MeasurableSet t} {hμt : μ t ≠ ∞} : edist (indicatorConstLp p hs hμs c) (indicatorConstLp p ht hμt c) = - ‖indicatorConstLp p (hs.symmDiff ht) (measure_symmDiff_ne_top hμs hμt) c‖₊ := by + ‖indicatorConstLp p (hs.symmDiff ht) (measure_symmDiff_ne_top hμs hμt) c‖ₑ := by unfold indicatorConstLp - rw [Lp.edist_toLp_toLp, eLpNorm_indicator_sub_indicator, Lp.coe_nnnorm_toLp] + rw [Lp.edist_toLp_toLp, eLpNorm_indicator_sub_indicator, Lp.enorm_toLp] + +@[deprecated (since := "2025-01-20")] +alias edist_indicatorConstLp_eq_nnnorm := edist_indicatorConstLp_eq_enorm theorem dist_indicatorConstLp_eq_norm {t : Set α} {ht : MeasurableSet t} {hμt : μ t ≠ ∞} : dist (indicatorConstLp p hs hμs c) (indicatorConstLp p ht hμt c) = ‖indicatorConstLp p (hs.symmDiff ht) (measure_symmDiff_ne_top hμs hμt) c‖ := by - rw [Lp.dist_edist, edist_indicatorConstLp_eq_nnnorm, ENNReal.coe_toReal, Lp.coe_nnnorm] + simp [Lp.dist_edist, edist_indicatorConstLp_eq_enorm, enorm, ENNReal.coe_toReal, Lp.coe_nnnorm] /-- A family of `indicatorConstLp` functions tends to an `indicatorConstLp`, if the underlying sets tend to the set in the sense of the measure of the symmetric difference. -/ @@ -741,12 +744,12 @@ theorem Lp.norm_const [NeZero μ] (hp_zero : p ≠ 0) : ‖Lp.const p μ c‖ = ‖c‖ * (μ Set.univ).toReal ^ (1 / p.toReal) := by have := NeZero.ne μ rw [← Memℒp.toLp_const, Lp.norm_toLp, eLpNorm_const] <;> try assumption - rw [ENNReal.toReal_mul, ENNReal.coe_toReal, ← ENNReal.toReal_rpow, coe_nnnorm] + rw [ENNReal.toReal_mul, toReal_enorm, ← ENNReal.toReal_rpow] theorem Lp.norm_const' (hp_zero : p ≠ 0) (hp_top : p ≠ ∞) : ‖Lp.const p μ c‖ = ‖c‖ * (μ Set.univ).toReal ^ (1 / p.toReal) := by rw [← Memℒp.toLp_const, Lp.norm_toLp, eLpNorm_const'] <;> try assumption - rw [ENNReal.toReal_mul, ENNReal.coe_toReal, ← ENNReal.toReal_rpow, coe_nnnorm] + rw [ENNReal.toReal_mul, toReal_enorm, ← ENNReal.toReal_rpow] theorem Lp.norm_const_le : ‖Lp.const p μ c‖ ≤ ‖c‖ * (μ Set.univ).toReal ^ (1 / p.toReal) := by rw [← indicatorConstLp_univ] @@ -1150,18 +1153,15 @@ namespace MeasureTheory namespace Lp theorem eLpNorm'_lim_eq_lintegral_liminf {ι} [Nonempty ι] [LinearOrder ι] {f : ι → α → G} {p : ℝ} - (hp_nonneg : 0 ≤ p) {f_lim : α → G} - (h_lim : ∀ᵐ x : α ∂μ, Tendsto (fun n => f n x) atTop (𝓝 (f_lim x))) : - eLpNorm' f_lim p μ = (∫⁻ a, atTop.liminf fun m => (‖f m a‖₊ : ℝ≥0∞) ^ p ∂μ) ^ (1 / p) := by - suffices h_no_pow : - (∫⁻ a, (‖f_lim a‖₊ : ℝ≥0∞) ^ p ∂μ) = ∫⁻ a, atTop.liminf fun m => (‖f m a‖₊ : ℝ≥0∞) ^ p ∂μ by - rw [eLpNorm'_eq_lintegral_nnnorm, h_no_pow] + {f_lim : α → G} (h_lim : ∀ᵐ x : α ∂μ, Tendsto (fun n => f n x) atTop (𝓝 (f_lim x))) : + eLpNorm' f_lim p μ = (∫⁻ a, atTop.liminf (‖f · a‖ₑ ^ p) ∂μ) ^ (1 / p) := by + suffices h_no_pow : (∫⁻ a, ‖f_lim a‖ₑ ^ p ∂μ) = ∫⁻ a, atTop.liminf fun m => ‖f m a‖ₑ ^ p ∂μ by + rw [eLpNorm'_eq_lintegral_enorm, h_no_pow] refine lintegral_congr_ae (h_lim.mono fun a ha => ?_) dsimp only rw [Tendsto.liminf_eq] - simp_rw [← ENNReal.coe_rpow_of_nonneg _ hp_nonneg, ENNReal.tendsto_coe] - refine ((NNReal.continuous_rpow_const hp_nonneg).tendsto ‖f_lim a‖₊).comp ?_ - exact (continuous_nnnorm.tendsto (f_lim a)).comp ha + refine (ENNReal.continuous_rpow_const.tendsto ‖f_lim a‖₊).comp ?_ + exact (continuous_enorm.tendsto (f_lim a)).comp ha @[deprecated (since := "2024-07-27")] alias snorm'_lim_eq_lintegral_liminf := eLpNorm'_lim_eq_lintegral_liminf @@ -1170,9 +1170,9 @@ theorem eLpNorm'_lim_le_liminf_eLpNorm' {E} [NormedAddCommGroup E] {f : ℕ → (hp_pos : 0 < p) (hf : ∀ n, AEStronglyMeasurable (f n) μ) {f_lim : α → E} (h_lim : ∀ᵐ x : α ∂μ, Tendsto (fun n => f n x) atTop (𝓝 (f_lim x))) : eLpNorm' f_lim p μ ≤ atTop.liminf fun n => eLpNorm' (f n) p μ := by - rw [eLpNorm'_lim_eq_lintegral_liminf hp_pos.le h_lim] + rw [eLpNorm'_lim_eq_lintegral_liminf h_lim] rw [one_div, ← ENNReal.le_rpow_inv_iff (by simp [hp_pos] : 0 < p⁻¹), inv_inv] - refine (lintegral_liminf_le' fun m => (hf m).ennnorm.pow_const _).trans_eq ?_ + refine (lintegral_liminf_le' fun m => (hf m).enorm.pow_const _).trans_eq ?_ have h_pow_liminf : atTop.liminf (fun n ↦ eLpNorm' (f n) p μ) ^ p = atTop.liminf fun n ↦ eLpNorm' (f n) p μ ^ p := by @@ -1181,7 +1181,7 @@ theorem eLpNorm'_lim_le_liminf_eLpNorm' {E} [NormedAddCommGroup E] {f : ℕ → refine (h_rpow_mono.orderIsoOfSurjective _ h_rpow_surj).liminf_apply ?_ ?_ ?_ ?_ all_goals isBoundedDefault rw [h_pow_liminf] - simp_rw [eLpNorm'_eq_lintegral_nnnorm, ← ENNReal.rpow_mul, one_div, + simp_rw [eLpNorm'_eq_lintegral_enorm, ← ENNReal.rpow_mul, one_div, inv_mul_cancel₀ hp_pos.ne.symm, ENNReal.rpow_one] @[deprecated (since := "2024-07-27")] @@ -1189,13 +1189,12 @@ alias snorm'_lim_le_liminf_snorm' := eLpNorm'_lim_le_liminf_eLpNorm' theorem eLpNorm_exponent_top_lim_eq_essSup_liminf {ι} [Nonempty ι] [LinearOrder ι] {f : ι → α → G} {f_lim : α → G} (h_lim : ∀ᵐ x : α ∂μ, Tendsto (fun n => f n x) atTop (𝓝 (f_lim x))) : - eLpNorm f_lim ∞ μ = essSup (fun x => atTop.liminf fun m => (‖f m x‖₊ : ℝ≥0∞)) μ := by - rw [eLpNorm_exponent_top, eLpNormEssSup_eq_essSup_nnnorm] + eLpNorm f_lim ∞ μ = essSup (fun x => atTop.liminf fun m => ‖f m x‖ₑ) μ := by + rw [eLpNorm_exponent_top, eLpNormEssSup_eq_essSup_enorm] refine essSup_congr_ae (h_lim.mono fun x hx => ?_) dsimp only apply (Tendsto.liminf_eq ..).symm - rw [ENNReal.tendsto_coe] - exact (continuous_nnnorm.tendsto (f_lim x)).comp hx + exact (continuous_enorm.tendsto (f_lim x)).comp hx @[deprecated (since := "2024-07-27")] alias snorm_exponent_top_lim_eq_essSup_liminf := eLpNorm_exponent_top_lim_eq_essSup_liminf @@ -1206,7 +1205,7 @@ theorem eLpNorm_exponent_top_lim_le_liminf_eLpNorm_exponent_top {ι} [Nonempty eLpNorm f_lim ∞ μ ≤ atTop.liminf fun n => eLpNorm (f n) ∞ μ := by rw [eLpNorm_exponent_top_lim_eq_essSup_liminf h_lim] simp_rw [eLpNorm_exponent_top, eLpNormEssSup] - exact ENNReal.essSup_liminf_le fun n => fun x => (‖f n x‖₊ : ℝ≥0∞) + exact ENNReal.essSup_liminf_le _ @[deprecated (since := "2024-07-27")] alias snorm_exponent_top_lim_le_liminf_snorm_exponent_top := @@ -1335,35 +1334,31 @@ private theorem eLpNorm'_sum_norm_sub_le_tsum_of_cauchy_eLpNorm' {f : ℕ → α alias snorm'_sum_norm_sub_le_tsum_of_cauchy_snorm' := eLpNorm'_sum_norm_sub_le_tsum_of_cauchy_eLpNorm' -private theorem lintegral_rpow_sum_coe_nnnorm_sub_le_rpow_tsum +private theorem lintegral_rpow_sum_enorm_sub_le_rpow_tsum {f : ℕ → α → E} {p : ℝ} (hp1 : 1 ≤ p) {B : ℕ → ℝ≥0∞} (n : ℕ) (hn : eLpNorm' (fun x => ∑ i ∈ Finset.range (n + 1), ‖f (i + 1) x - f i x‖) p μ ≤ ∑' i, B i) : - (∫⁻ a, (∑ i ∈ Finset.range (n + 1), ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞) ^ p ∂μ) ≤ - (∑' i, B i) ^ p := by + (∫⁻ a, (∑ i ∈ Finset.range (n + 1), ‖f (i + 1) a - f i a‖ₑ) ^ p ∂μ) ≤ (∑' i, B i) ^ p := by have hp_pos : 0 < p := zero_lt_one.trans_le hp1 rw [← inv_inv p, @ENNReal.le_rpow_inv_iff _ _ p⁻¹ (by simp [hp_pos]), inv_inv p] - simp_rw [eLpNorm'_eq_lintegral_nnnorm, one_div] at hn + simp_rw [eLpNorm'_eq_lintegral_enorm, one_div] at hn have h_nnnorm_nonneg : - (fun a => (‖∑ i ∈ Finset.range (n + 1), ‖f (i + 1) a - f i a‖‖₊ : ℝ≥0∞) ^ p) = fun a => - (∑ i ∈ Finset.range (n + 1), (‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)) ^ p := by + (fun a => ‖∑ i ∈ Finset.range (n + 1), ‖f (i + 1) a - f i a‖‖ₑ ^ p) = fun a => + (∑ i ∈ Finset.range (n + 1), ‖f (i + 1) a - f i a‖ₑ) ^ p := by ext1 a congr - simp_rw [← ofReal_norm_eq_coe_nnnorm] + simp_rw [← ofReal_norm_eq_enorm] rw [← ENNReal.ofReal_sum_of_nonneg] · rw [Real.norm_of_nonneg _] exact Finset.sum_nonneg fun x _ => norm_nonneg _ · exact fun x _ => norm_nonneg _ rwa [h_nnnorm_nonneg] at hn -private theorem lintegral_rpow_tsum_coe_nnnorm_sub_le_tsum {f : ℕ → α → E} +private theorem lintegral_rpow_tsum_coe_enorm_sub_le_tsum {f : ℕ → α → E} (hf : ∀ n, AEStronglyMeasurable (f n) μ) {p : ℝ} (hp1 : 1 ≤ p) {B : ℕ → ℝ≥0∞} - (h : - ∀ n, - (∫⁻ a, (∑ i ∈ Finset.range (n + 1), ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞) ^ p ∂μ) ≤ - (∑' i, B i) ^ p) : - (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞) ^ p ∂μ) ^ (1 / p) ≤ ∑' i, B i := by + (h : ∀ n, ∫⁻ a, (∑ i ∈ Finset.range (n + 1), ‖f (i + 1) a - f i a‖ₑ) ^ p ∂μ ≤ (∑' i, B i) ^ p) : + (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖ₑ) ^ p ∂μ) ^ (1 / p) ≤ ∑' i, B i := by have hp_pos : 0 < p := zero_lt_one.trans_le hp1 - suffices h_pow : (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞) ^ p ∂μ) ≤ (∑' i, B i) ^ p by + suffices h_pow : (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖ₑ) ^ p ∂μ) ≤ (∑' i, B i) ^ p by rwa [one_div, ← ENNReal.le_rpow_inv_iff (by simp [hp_pos] : 0 < p⁻¹), inv_inv] have h_tsum_1 : ∀ g : ℕ → ℝ≥0∞, ∑' i, g i = atTop.liminf fun n => ∑ i ∈ Finset.range (n + 1), g i := by @@ -1372,35 +1367,29 @@ private theorem lintegral_rpow_tsum_coe_nnnorm_sub_le_tsum {f : ℕ → α → E simp_rw [h_tsum_1 _] rw [← h_tsum_1] have h_liminf_pow : - (∫⁻ a, (atTop.liminf - fun n => ∑ i ∈ Finset.range (n + 1), (‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)) ^ p ∂μ) = - ∫⁻ a, atTop.liminf - fun n => (∑ i ∈ Finset.range (n + 1), (‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)) ^ p ∂μ := by + ∫⁻ a, (atTop.liminf fun n => ∑ i ∈ Finset.range (n + 1), ‖f (i + 1) a - f i a‖ₑ) ^ p ∂μ = + ∫⁻ a, atTop.liminf fun n => (∑ i ∈ Finset.range (n + 1), ‖f (i + 1) a - f i a‖ₑ) ^ p ∂μ := by refine lintegral_congr fun x => ?_ have h_rpow_mono := ENNReal.strictMono_rpow_of_pos (zero_lt_one.trans_le hp1) have h_rpow_surj := (ENNReal.rpow_left_bijective hp_pos.ne.symm).2 refine (h_rpow_mono.orderIsoOfSurjective _ h_rpow_surj).liminf_apply ?_ ?_ ?_ ?_ all_goals isBoundedDefault rw [h_liminf_pow] - refine (lintegral_liminf_le' ?_).trans ?_ - · exact fun n => - (Finset.aemeasurable_sum (Finset.range (n + 1)) fun i _ => - ((hf (i + 1)).sub (hf i)).ennnorm).pow_const - _ - · exact liminf_le_of_frequently_le' (Frequently.of_forall h) - -private theorem tsum_nnnorm_sub_ae_lt_top {f : ℕ → α → E} (hf : ∀ n, AEStronglyMeasurable (f n) μ) + refine (lintegral_liminf_le' fun n ↦ ?_).trans <| liminf_le_of_frequently_le' <| .of_forall h + exact ((Finset.range _).aemeasurable_sum fun i _ ↦ ((hf _).sub (hf i)).enorm).pow_const _ + +private theorem tsum_enorm_sub_ae_lt_top {f : ℕ → α → E} (hf : ∀ n, AEStronglyMeasurable (f n) μ) {p : ℝ} (hp1 : 1 ≤ p) {B : ℕ → ℝ≥0∞} (hB : ∑' i, B i ≠ ∞) - (h : (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞) ^ p ∂μ) ^ (1 / p) ≤ ∑' i, B i) : - ∀ᵐ x ∂μ, (∑' i, ‖f (i + 1) x - f i x‖₊ : ℝ≥0∞) < ∞ := by + (h : (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖ₑ) ^ p ∂μ) ^ (1 / p) ≤ ∑' i, B i) : + ∀ᵐ x ∂μ, ∑' i, ‖f (i + 1) x - f i x‖ₑ < ∞ := by have hp_pos : 0 < p := zero_lt_one.trans_le hp1 - have h_integral : (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞) ^ p ∂μ) < ∞ := by + have h_integral : ∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖ₑ) ^ p ∂μ < ∞ := by have h_tsum_lt_top : (∑' i, B i) ^ p < ∞ := ENNReal.rpow_lt_top_of_nonneg hp_pos.le hB refine lt_of_le_of_lt ?_ h_tsum_lt_top rwa [one_div, ← ENNReal.le_rpow_inv_iff (by simp [hp_pos] : 0 < p⁻¹), inv_inv] at h - have rpow_ae_lt_top : ∀ᵐ x ∂μ, (∑' i, ‖f (i + 1) x - f i x‖₊ : ℝ≥0∞) ^ p < ∞ := by + have rpow_ae_lt_top : ∀ᵐ x ∂μ, (∑' i, ‖f (i + 1) x - f i x‖ₑ) ^ p < ∞ := by refine ae_lt_top' (AEMeasurable.pow_const ?_ _) h_integral.ne - exact AEMeasurable.ennreal_tsum fun n => ((hf (n + 1)).sub (hf n)).ennnorm + exact AEMeasurable.ennreal_tsum fun n => ((hf (n + 1)).sub (hf n)).enorm refine rpow_ae_lt_top.mono fun x hx => ?_ rwa [← ENNReal.lt_rpow_inv_iff hp_pos, ENNReal.top_rpow_of_pos (by simp [hp_pos] : 0 < p⁻¹)] at hx @@ -1413,15 +1402,13 @@ theorem ae_tendsto_of_cauchy_eLpNorm' [CompleteSpace E] {f : ℕ → α → E} { have h1 : ∀ n, eLpNorm' (fun x => ∑ i ∈ Finset.range (n + 1), ‖f (i + 1) x - f i x‖) p μ ≤ ∑' i, B i := eLpNorm'_sum_norm_sub_le_tsum_of_cauchy_eLpNorm' hf hp1 h_cau - have h2 : - ∀ n, - (∫⁻ a, (∑ i ∈ Finset.range (n + 1), ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞) ^ p ∂μ) ≤ - (∑' i, B i) ^ p := - fun n => lintegral_rpow_sum_coe_nnnorm_sub_le_rpow_tsum hp1 n (h1 n) - have h3 : (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞) ^ p ∂μ) ^ (1 / p) ≤ ∑' i, B i := - lintegral_rpow_tsum_coe_nnnorm_sub_le_tsum hf hp1 h2 - have h4 : ∀ᵐ x ∂μ, (∑' i, ‖f (i + 1) x - f i x‖₊ : ℝ≥0∞) < ∞ := - tsum_nnnorm_sub_ae_lt_top hf hp1 hB h3 + have h2 n : + ∫⁻ a, (∑ i ∈ Finset.range (n + 1), ‖f (i + 1) a - f i a‖ₑ) ^ p ∂μ ≤ (∑' i, B i) ^ p := + lintegral_rpow_sum_enorm_sub_le_rpow_tsum hp1 n (h1 n) + have h3 : (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖ₑ) ^ p ∂μ) ^ (1 / p) ≤ ∑' i, B i := + lintegral_rpow_tsum_coe_enorm_sub_le_tsum hf hp1 h2 + have h4 : ∀ᵐ x ∂μ, ∑' i, ‖f (i + 1) x - f i x‖ₑ < ∞ := + tsum_enorm_sub_ae_lt_top hf hp1 hB h3 exact h4.mono fun x hx => .of_nnnorm <| ENNReal.tsum_coe_ne_top_iff_summable.mp hx.ne have h : ∀ᵐ x ∂μ, ∃ l : E, @@ -1453,7 +1440,7 @@ theorem ae_tendsto_of_cauchy_eLpNorm [CompleteSpace E] {f : ℕ → α → E} ∀ᵐ x ∂μ, ∃ l : E, atTop.Tendsto (fun n => f n x) (𝓝 l) := by by_cases hp_top : p = ∞ · simp_rw [hp_top] at * - have h_cau_ae : ∀ᵐ x ∂μ, ∀ N n m, N ≤ n → N ≤ m → (‖(f n - f m) x‖₊ : ℝ≥0∞) < B N := by + have h_cau_ae : ∀ᵐ x ∂μ, ∀ N n m, N ≤ n → N ≤ m → ‖(f n - f m) x‖ₑ < B N := by simp_rw [ae_all_iff] exact fun N n m hnN hmN => ae_lt_of_essSup_lt (h_cau N n m hnN hmN) simp_rw [eLpNorm_exponent_top, eLpNormEssSup] at h_cau @@ -1463,7 +1450,7 @@ theorem ae_tendsto_of_cauchy_eLpNorm [CompleteSpace E] {f : ℕ → α → E} specialize hx N n m hnN hmN rw [_root_.dist_eq_norm, ← ENNReal.ofReal_le_iff_le_toReal (ENNReal.ne_top_of_tsum_ne_top hB N), - ofReal_norm_eq_coe_nnnorm] + ofReal_norm_eq_enorm] exact hx.le · rw [← ENNReal.zero_toReal] exact @@ -1745,27 +1732,35 @@ namespace MeasureTheory namespace Lp -theorem pow_mul_meas_ge_le_norm (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (ε : ℝ≥0∞) : - (ε * μ { x | ε ≤ (‖f x‖₊ : ℝ≥0∞) ^ p.toReal }) ^ (1 / p.toReal) ≤ ENNReal.ofReal ‖f‖ := +/-- A version of **Markov's inequality** with elements of Lp. -/ +lemma pow_mul_meas_ge_le_enorm (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (ε : ℝ≥0∞) : + (ε * μ {x | ε ≤ ‖f x‖ₑ ^ p.toReal}) ^ (1 / p.toReal) ≤ ENNReal.ofReal ‖f‖ := (ENNReal.ofReal_toReal (eLpNorm_ne_top f)).symm ▸ pow_mul_meas_ge_le_eLpNorm μ hp_ne_zero hp_ne_top (Lp.aestronglyMeasurable f) ε -theorem mul_meas_ge_le_pow_norm (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (ε : ℝ≥0∞) : - ε * μ { x | ε ≤ (‖f x‖₊ : ℝ≥0∞) ^ p.toReal } ≤ ENNReal.ofReal ‖f‖ ^ p.toReal := +/-- A version of **Markov's inequality** with elements of Lp. -/ +lemma mul_meas_ge_le_pow_enorm (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (ε : ℝ≥0∞) : + ε * μ {x | ε ≤ ‖f x‖ₑ ^ p.toReal} ≤ ENNReal.ofReal ‖f‖ ^ p.toReal := (ENNReal.ofReal_toReal (eLpNorm_ne_top f)).symm ▸ mul_meas_ge_le_pow_eLpNorm μ hp_ne_zero hp_ne_top (Lp.aestronglyMeasurable f) ε -/-- A version of Markov's inequality with elements of Lp. -/ -theorem mul_meas_ge_le_pow_norm' (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) - (ε : ℝ≥0∞) : ε ^ p.toReal * μ { x | ε ≤ ‖f x‖₊ } ≤ ENNReal.ofReal ‖f‖ ^ p.toReal := +/-- A version of **Markov's inequality** with elements of Lp. -/ +theorem mul_meas_ge_le_pow_enorm' (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) + (ε : ℝ≥0∞) : ε ^ p.toReal * μ {x | ε ≤ ‖f x‖₊ } ≤ ENNReal.ofReal ‖f‖ ^ p.toReal := (ENNReal.ofReal_toReal (eLpNorm_ne_top f)).symm ▸ mul_meas_ge_le_pow_eLpNorm' μ hp_ne_zero hp_ne_top (Lp.aestronglyMeasurable f) ε -theorem meas_ge_le_mul_pow_norm (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {ε : ℝ≥0∞} - (hε : ε ≠ 0) : μ { x | ε ≤ ‖f x‖₊ } ≤ ε⁻¹ ^ p.toReal * ENNReal.ofReal ‖f‖ ^ p.toReal := +/-- A version of **Markov's inequality** with elements of Lp. -/ +theorem meas_ge_le_mul_pow_enorm (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {ε : ℝ≥0∞} + (hε : ε ≠ 0) : μ {x | ε ≤ ‖f x‖₊} ≤ ε⁻¹ ^ p.toReal * ENNReal.ofReal ‖f‖ ^ p.toReal := (ENNReal.ofReal_toReal (eLpNorm_ne_top f)).symm ▸ meas_ge_le_mul_pow_eLpNorm μ hp_ne_zero hp_ne_top (Lp.aestronglyMeasurable f) hε +@[deprecated (since := "2025-01-20")] alias pow_mul_meas_ge_le_norm := pow_mul_meas_ge_le_enorm +@[deprecated (since := "2025-01-20")] alias mul_meas_ge_le_pow_norm := mul_meas_ge_le_pow_enorm +@[deprecated (since := "2025-01-20")] alias mul_meas_ge_le_pow_norm' := mul_meas_ge_le_pow_enorm' +@[deprecated (since := "2025-01-20")] alias meas_ge_le_mul_pow_norm := meas_ge_le_mul_pow_enorm + end Lp end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 3d030c368a939f..80239aa069a9a6 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -28,7 +28,6 @@ open Filter ENNReal open Function (support) -open scoped Classical open Topology NNReal ENNReal MeasureTheory namespace MeasureTheory @@ -177,6 +176,7 @@ theorem sum_range_measure_preimage_singleton (f : α →ₛ β) (μ : Measure α (∑ y ∈ f.range, μ (f ⁻¹' {y})) = μ univ := by rw [f.sum_measure_preimage_singleton, coe_range, preimage_range] +open scoped Classical in /-- If-then-else as a `SimpleFunc`. -/ def piecewise (s : Set α) (hs : MeasurableSet s) (f g : α →ₛ β) : α →ₛ β := ⟨s.piecewise f g, fun _ => @@ -184,15 +184,18 @@ def piecewise (s : Set α) (hs : MeasurableSet s) (f g : α →ₛ β) : α → f.measurable.piecewise hs g.measurable trivial, (f.finite_range.union g.finite_range).subset range_ite_subset⟩ +open scoped Classical in @[simp] theorem coe_piecewise {s : Set α} (hs : MeasurableSet s) (f g : α →ₛ β) : ⇑(piecewise s hs f g) = s.piecewise f g := rfl +open scoped Classical in theorem piecewise_apply {s : Set α} (hs : MeasurableSet s) (f g : α →ₛ β) (a) : piecewise s hs f g a = if a ∈ s then f a else g a := rfl +open scoped Classical in @[simp] theorem piecewise_compl {s : Set α} (hs : MeasurableSet sᶜ) (f g : α →ₛ β) : piecewise sᶜ hs f g = piecewise s hs.of_compl g f := @@ -206,6 +209,7 @@ theorem piecewise_univ (f g : α →ₛ β) : piecewise univ MeasurableSet.univ theorem piecewise_empty (f g : α →ₛ β) : piecewise ∅ MeasurableSet.empty f g = g := coe_injective <| by simp +open scoped Classical in @[simp] theorem piecewise_same (f : α →ₛ β) {s : Set α} (hs : MeasurableSet s) : piecewise s hs f f = f := @@ -215,6 +219,7 @@ theorem support_indicator [Zero β] {s : Set α} (hs : MeasurableSet s) (f : α Function.support (f.piecewise s hs (SimpleFunc.const α 0)) = s ∩ Function.support f := Set.support_indicator +open scoped Classical in theorem range_indicator {s : Set α} (hs : MeasurableSet s) (hs_nonempty : s.Nonempty) (hs_ne_univ : s ≠ univ) (x y : β) : (piecewise s hs (const α x) (const α y)).range = {x, y} := by @@ -261,12 +266,14 @@ theorem range_map [DecidableEq γ] (g : β → γ) (f : α →ₛ β) : (f.map g theorem map_const (g : β → γ) (b : β) : (const α b).map g = const α (g b) := rfl +open scoped Classical in theorem map_preimage (f : α →ₛ β) (g : β → γ) (s : Set γ) : f.map g ⁻¹' s = f ⁻¹' ↑{b ∈ f.range | g b ∈ s} := by simp only [coe_range, sep_mem_eq, coe_map, Finset.coe_filter, ← mem_preimage, inter_comm, preimage_inter_range, ← Finset.mem_coe] exact preimage_comp +open scoped Classical in theorem map_preimage_singleton (f : α →ₛ β) (g : β → γ) (c : γ) : f.map g ⁻¹' {c} = f ⁻¹' ↑{b ∈ f.range | g b = c} := map_preimage _ _ _ @@ -560,6 +567,7 @@ instance instPreorder : Preorder (α →ₛ β) := Preorder.lift (⇑) @[gcongr] protected alias ⟨_, GCongr.coe_le_coe⟩ := coe_le_coe @[gcongr] protected alias ⟨_, GCongr.coe_lt_coe⟩ := coe_lt_coe +open scoped Classical in @[gcongr] lemma piecewise_mono (hf : ∀ a ∈ s, f₁ a ≤ f₂ a) (hg : ∀ a ∉ s, g₁ a ≤ g₂ a) : piecewise s hs f₁ g₁ ≤ piecewise s hs f₂ g₂ := Set.piecewise_mono hf hg @@ -600,6 +608,7 @@ instance instBoundedOrder [LE β] [BoundedOrder β] : BoundedOrder (α →ₛ β theorem finset_sup_apply [SemilatticeSup β] [OrderBot β] {f : γ → α →ₛ β} (s : Finset γ) (a : α) : s.sup f a = s.sup fun c => f c a := by + classical refine Finset.induction_on s rfl ?_ intro a s _ ih rw [Finset.sup_insert, Finset.sup_insert, sup_apply, ih] @@ -608,6 +617,7 @@ section Restrict variable [Zero β] +open scoped Classical in /-- Restrict a simple function `f : α →ₛ β` to a set `s`. If `s` is measurable, then `f.restrict s a = if a ∈ s then f a else 0`, otherwise `f.restrict s = const α 0`. -/ def restrict (f : α →ₛ β) (s : Set α) : α →ₛ β := @@ -620,6 +630,7 @@ theorem restrict_of_not_measurable {f : α →ₛ β} {s : Set α} (hs : ¬Measu @[simp] theorem coe_restrict (f : α →ₛ β) {s : Set α} (hs : MeasurableSet s) : ⇑(restrict f s) = indicator s f := by + classical rw [restrict, dif_pos hs, coe_piecewise, coe_zero, piecewise_eq_indicator] @[simp] @@ -628,6 +639,7 @@ theorem restrict_univ (f : α →ₛ β) : restrict f univ = f := by simp [restr @[simp] theorem restrict_empty (f : α →ₛ β) : restrict f ∅ = 0 := by simp [restrict] +open scoped Classical in theorem map_restrict_of_zero [Zero γ] {g : β → γ} (hg : g 0 = 0) (f : α →ₛ β) (s : Set α) : (f.restrict s).map g = (f.map g).restrict s := ext fun x => @@ -657,6 +669,7 @@ theorem mem_restrict_range {r : β} {s : Set α} {f : α →ₛ β} (hs : Measur r ∈ (restrict f s).range ↔ r = 0 ∧ s ≠ univ ∨ r ∈ f '' s := by rw [← Finset.mem_coe, coe_range, coe_restrict _ hs, mem_range_indicator] +open scoped Classical in theorem mem_image_of_mem_range_restrict {r : β} {s : Set α} {f : α →ₛ β} (hr : r ∈ (restrict f s).range) (h0 : r ≠ 0) : r ∈ f '' s := if hs : MeasurableSet s then by simpa [mem_restrict_range hs, h0, -mem_range] using hr @@ -664,6 +677,7 @@ theorem mem_image_of_mem_range_restrict {r : β} {s : Set α} {f : α →ₛ β} rw [restrict_of_not_measurable hs] at hr exact (h0 <| eq_zero_of_mem_range_zero hr).elim +open scoped Classical in @[gcongr, mono] theorem restrict_mono [Preorder β] (s : Set α) {f g : α →ₛ β} (H : f ≤ g) : f.restrict s ≤ g.restrict s := @@ -685,6 +699,7 @@ of the set `{i k | k ≤ n ∧ i k ≤ f a}`, see `approx_apply` and `iSup_appro def approx (i : ℕ → β) (f : α → β) (n : ℕ) : α →ₛ β := (Finset.range n).sup fun k => restrict (const α (i k)) { a : α | i k ≤ f a } +open scoped Classical in theorem approx_apply [TopologicalSpace β] [OrderClosedTopology β] [MeasurableSpace β] [OpensMeasurableSpace β] {i : ℕ → β} {f : α → β} {n : ℕ} (a : α) (hf : Measurable f) : (approx i f n : α →ₛ β) a = (Finset.range n).sup fun k => if i k ≤ f a then i k else 0 := by @@ -904,6 +919,7 @@ theorem lintegral_sum {m : MeasurableSpace α} {ι} (f : α →ₛ ℝ≥0∞) ( ENNReal.tsum_mul_left] apply ENNReal.tsum_comm +open scoped Classical in theorem restrict_lintegral (f : α →ₛ ℝ≥0∞) {s : Set α} (hs : MeasurableSet s) : (restrict f s).lintegral μ = ∑ r ∈ f.range, r * μ (f ⁻¹' {r} ∩ s) := calc @@ -1006,6 +1022,7 @@ section FinMeasSupp open Finset Function +open scoped Classical in theorem support_eq [MeasurableSpace α] [Zero β] (f : α →ₛ β) : support f = ⋃ y ∈ {y ∈ f.range | y ≠ 0}, f ⁻¹' {y} := Set.ext fun x => by @@ -1020,6 +1037,7 @@ theorem measurableSet_support [MeasurableSpace α] (f : α →ₛ β) : Measurab lemma measure_support_lt_top (f : α →ₛ β) (hf : ∀ y, y ≠ 0 → μ (f ⁻¹' {y}) < ∞) : μ (support f) < ∞ := by + classical rw [support_eq] refine (measure_biUnion_finset_le _ _).trans_lt (ENNReal.sum_lt_top.mpr fun y hy => ?_) rw [Finset.mem_filter] at hy @@ -1034,6 +1052,7 @@ theorem finMeasSupp_iff_support : f.FinMeasSupp μ ↔ μ (support f) < ∞ := Iff.rfl theorem finMeasSupp_iff : f.FinMeasSupp μ ↔ ∀ y, y ≠ 0 → μ (f ⁻¹' {y}) < ∞ := by + classical constructor · refine fun h y hy => lt_of_le_of_lt (measure_mono ?_) h exact fun x hx (H : f x = 0) => hy <| H ▸ Eq.symm hx @@ -1124,6 +1143,7 @@ protected theorem induction {α γ} [MeasurableSpace α] [AddMonoid γ] {P : Sim P (SimpleFunc.piecewise s hs (SimpleFunc.const _ c) (SimpleFunc.const _ 0))) (h_add : ∀ ⦃f g : SimpleFunc α γ⦄, Disjoint (support f) (support g) → P f → P g → P (f + g)) (f : SimpleFunc α γ) : P f := by + classical generalize h : f.range \ {0} = s rw [← Finset.coe_inj, Finset.coe_sdiff, Finset.coe_singleton, SimpleFunc.coe_range] at h induction s using Finset.induction generalizing f with diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean index d8993de25b2648..a764aa7f10af51 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean @@ -34,7 +34,6 @@ by a sequence of simple functions. open Set Function Filter TopologicalSpace ENNReal EMetric Finset -open scoped Classical open Topology ENNReal MeasureTheory variable {α β ι E F 𝕜 : Type*} diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index 7ad4405faa9c1a..a7d019b6e5ebea 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Zhouhang Zhou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Zhouhang Zhou, Yury Kudryashov, Heather Macbeth -/ -import Mathlib.MeasureTheory.Function.L1Space +import Mathlib.MeasureTheory.Function.L1Space.AEEqFun import Mathlib.MeasureTheory.Function.SimpleFuncDense /-! @@ -45,7 +45,7 @@ noncomputable section open Set Function Filter TopologicalSpace ENNReal EMetric Finset -open scoped Classical Topology ENNReal MeasureTheory +open scoped Topology ENNReal MeasureTheory variable {α β ι E F 𝕜 : Type*} @@ -73,16 +73,14 @@ theorem nnnorm_approxOn_le [OpensMeasurableSpace E] {f : β → E} (hf : Measura theorem norm_approxOn_y₀_le [OpensMeasurableSpace E] {f : β → E} (hf : Measurable f) {s : Set E} {y₀ : E} (h₀ : y₀ ∈ s) [SeparableSpace s] (x : β) (n : ℕ) : ‖approxOn f hf s y₀ h₀ n x - y₀‖ ≤ ‖f x - y₀‖ + ‖f x - y₀‖ := by - have := edist_approxOn_y0_le hf h₀ x n - repeat rw [edist_comm y₀, edist_eq_coe_nnnorm_sub] at this - exact mod_cast this + simpa [enorm, edist_eq_enorm_sub, ← ENNReal.coe_add, norm_sub_rev] + using edist_approxOn_y0_le hf h₀ x n theorem norm_approxOn_zero_le [OpensMeasurableSpace E] {f : β → E} (hf : Measurable f) {s : Set E} (h₀ : (0 : E) ∈ s) [SeparableSpace s] (x : β) (n : ℕ) : ‖approxOn f hf s 0 h₀ n x‖ ≤ ‖f x‖ + ‖f x‖ := by - have := edist_approxOn_y0_le hf h₀ x n - simp only [edist_comm (0 : E), edist_eq_coe_nnnorm] at this - exact mod_cast this + simpa [enorm, edist_eq_enorm_sub, ← ENNReal.coe_add, norm_sub_rev] + using edist_approxOn_y0_le hf h₀ x n theorem tendsto_approxOn_Lp_eLpNorm [OpensMeasurableSpace E] {f : β → E} (hf : Measurable f) {s : Set E} {y₀ : E} (h₀ : y₀ ∈ s) [SeparableSpace s] (hp_ne_top : p ≠ ∞) {μ : Measure β} @@ -91,35 +89,28 @@ theorem tendsto_approxOn_Lp_eLpNorm [OpensMeasurableSpace E] {f : β → E} (hf by_cases hp_zero : p = 0 · simpa only [hp_zero, eLpNorm_exponent_zero] using tendsto_const_nhds have hp : 0 < p.toReal := toReal_pos hp_zero hp_ne_top - suffices - Tendsto (fun n => ∫⁻ x, (‖approxOn f hf s y₀ h₀ n x - f x‖₊ : ℝ≥0∞) ^ p.toReal ∂μ) atTop - (𝓝 0) by - simp only [eLpNorm_eq_lintegral_rpow_nnnorm hp_zero hp_ne_top] + suffices Tendsto (fun n => ∫⁻ x, ‖approxOn f hf s y₀ h₀ n x - f x‖ₑ ^ p.toReal ∂μ) atTop (𝓝 0) by + simp only [eLpNorm_eq_lintegral_rpow_enorm hp_zero hp_ne_top] convert continuous_rpow_const.continuousAt.tendsto.comp this simp [zero_rpow_of_pos (_root_.inv_pos.mpr hp)] -- We simply check the conditions of the Dominated Convergence Theorem: -- (1) The function "`p`-th power of distance between `f` and the approximation" is measurable - have hF_meas : - ∀ n, Measurable fun x => (‖approxOn f hf s y₀ h₀ n x - f x‖₊ : ℝ≥0∞) ^ p.toReal := by - simpa only [← edist_eq_coe_nnnorm_sub] using fun n => + have hF_meas n : Measurable fun x => ‖approxOn f hf s y₀ h₀ n x - f x‖ₑ ^ p.toReal := by + simpa only [← edist_eq_enorm_sub] using (approxOn f hf s y₀ h₀ n).measurable_bind (fun y x => edist y (f x) ^ p.toReal) fun y => (measurable_edist_right.comp hf).pow_const p.toReal -- (2) The functions "`p`-th power of distance between `f` and the approximation" are uniformly -- bounded, at any given point, by `fun x => ‖f x - y₀‖ ^ p.toReal` - have h_bound : - ∀ n, (fun x => (‖approxOn f hf s y₀ h₀ n x - f x‖₊ : ℝ≥0∞) ^ p.toReal) ≤ᵐ[μ] fun x => - (‖f x - y₀‖₊ : ℝ≥0∞) ^ p.toReal := - fun n => - Eventually.of_forall fun x => - rpow_le_rpow (coe_mono (nnnorm_approxOn_le hf h₀ x n)) toReal_nonneg + have h_bound n : + (fun x ↦ ‖approxOn f hf s y₀ h₀ n x - f x‖ₑ ^ p.toReal) ≤ᵐ[μ] (‖f · - y₀‖ₑ ^ p.toReal) := + .of_forall fun x => rpow_le_rpow (coe_mono (nnnorm_approxOn_le hf h₀ x n)) toReal_nonneg -- (3) The bounding function `fun x => ‖f x - y₀‖ ^ p.toReal` has finite integral - have h_fin : (∫⁻ a : β, (‖f a - y₀‖₊ : ℝ≥0∞) ^ p.toReal ∂μ) ≠ ⊤ := - (lintegral_rpow_nnnorm_lt_top_of_eLpNorm_lt_top hp_zero hp_ne_top hi).ne + have h_fin : (∫⁻ a : β, ‖f a - y₀‖ₑ ^ p.toReal ∂μ) ≠ ⊤ := + (lintegral_rpow_enorm_lt_top_of_eLpNorm_lt_top hp_zero hp_ne_top hi).ne -- (4) The functions "`p`-th power of distance between `f` and the approximation" tend pointwise -- to zero have h_lim : - ∀ᵐ a : β ∂μ, - Tendsto (fun n => (‖approxOn f hf s y₀ h₀ n a - f a‖₊ : ℝ≥0∞) ^ p.toReal) atTop (𝓝 0) := by + ∀ᵐ a : β ∂μ, Tendsto (‖approxOn f hf s y₀ h₀ · a - f a‖ₑ ^ p.toReal) atTop (𝓝 0) := by filter_upwards [hμ] with a ha have : Tendsto (fun n => (approxOn f hf s y₀ h₀ n) a - f a) atTop (𝓝 (f a - f a)) := (tendsto_approxOn hf h₀ ha).sub tendsto_const_nhds @@ -222,13 +213,15 @@ section Integrable variable [MeasurableSpace β] variable [MeasurableSpace E] [NormedAddCommGroup E] -theorem tendsto_approxOn_L1_nnnorm [OpensMeasurableSpace E] {f : β → E} (hf : Measurable f) +theorem tendsto_approxOn_L1_enorm [OpensMeasurableSpace E] {f : β → E} (hf : Measurable f) {s : Set E} {y₀ : E} (h₀ : y₀ ∈ s) [SeparableSpace s] {μ : Measure β} (hμ : ∀ᵐ x ∂μ, f x ∈ closure s) (hi : HasFiniteIntegral (fun x => f x - y₀) μ) : - Tendsto (fun n => ∫⁻ x, ‖approxOn f hf s y₀ h₀ n x - f x‖₊ ∂μ) atTop (𝓝 0) := by - simpa [eLpNorm_one_eq_lintegral_nnnorm] using + Tendsto (fun n => ∫⁻ x, ‖approxOn f hf s y₀ h₀ n x - f x‖ₑ ∂μ) atTop (𝓝 0) := by + simpa [eLpNorm_one_eq_lintegral_enorm] using tendsto_approxOn_Lp_eLpNorm hf h₀ one_ne_top hμ - (by simpa [eLpNorm_one_eq_lintegral_nnnorm] using hi) + (by simpa [eLpNorm_one_eq_lintegral_enorm] using hi) + +@[deprecated (since := "2025-01-21")] alias tendsto_approxOn_L1_nnnorm := tendsto_approxOn_L1_enorm theorem integrable_approxOn [BorelSpace E] {f : β → E} {μ : Measure β} (fmeas : Measurable f) (hf : Integrable f μ) {s : Set E} {y₀ : E} (h₀ : y₀ ∈ s) [SeparableSpace s] @@ -236,14 +229,17 @@ theorem integrable_approxOn [BorelSpace E] {f : β → E} {μ : Measure β} (fme rw [← memℒp_one_iff_integrable] at hf hi₀ ⊢ exact memℒp_approxOn fmeas hf h₀ hi₀ n -theorem tendsto_approxOn_range_L1_nnnorm [OpensMeasurableSpace E] {f : β → E} {μ : Measure β} +theorem tendsto_approxOn_range_L1_enorm [OpensMeasurableSpace E] {f : β → E} {μ : Measure β} [SeparableSpace (range f ∪ {0} : Set E)] (fmeas : Measurable f) (hf : Integrable f μ) : - Tendsto (fun n => ∫⁻ x, ‖approxOn f fmeas (range f ∪ {0}) 0 (by simp) n x - f x‖₊ ∂μ) atTop + Tendsto (fun n => ∫⁻ x, ‖approxOn f fmeas (range f ∪ {0}) 0 (by simp) n x - f x‖ₑ ∂μ) atTop (𝓝 0) := by - apply tendsto_approxOn_L1_nnnorm fmeas + apply tendsto_approxOn_L1_enorm fmeas · filter_upwards with x using subset_closure (by simp) · simpa using hf.2 +@[deprecated (since := "2025-01-21")] +alias tendsto_approxOn_range_L1_nnnorm := tendsto_approxOn_range_L1_enorm + theorem integrable_approxOn_range [BorelSpace E] {f : β → E} {μ : Measure β} (fmeas : Measurable f) [SeparableSpace (range f ∪ {0} : Set E)] (hf : Integrable f μ) (n : ℕ) : Integrable (approxOn f fmeas (range f ∪ {0}) 0 (by simp) n) μ := @@ -278,10 +274,9 @@ theorem memℒp_top (f : α →ₛ E) (μ : Measure α) : Memℒp f ∞ μ := memℒp_top_of_bound f.aestronglyMeasurable C <| Eventually.of_forall hfC protected theorem eLpNorm'_eq {p : ℝ} (f : α →ₛ F) (μ : Measure α) : - eLpNorm' f p μ = (∑ y ∈ f.range, (‖y‖₊ : ℝ≥0∞) ^ p * μ (f ⁻¹' {y})) ^ (1 / p) := by - have h_map : (fun a => (‖f a‖₊ : ℝ≥0∞) ^ p) = f.map fun a : F => (‖a‖₊ : ℝ≥0∞) ^ p := by - simp; rfl - rw [eLpNorm'_eq_lintegral_nnnorm, h_map, lintegral_eq_lintegral, map_lintegral] + eLpNorm' f p μ = (∑ y ∈ f.range, ‖y‖ₑ ^ p * μ (f ⁻¹' {y})) ^ (1 / p) := by + have h_map : (‖f ·‖ₑ ^ p) = f.map (‖·‖ₑ ^ p) := by simp; rfl + rw [eLpNorm'_eq_lintegral_enorm, h_map, lintegral_eq_lintegral, map_lintegral] @[deprecated (since := "2024-07-27")] protected alias snorm'_eq := SimpleFunc.eLpNorm'_eq diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 3b54bf1d8f7dba..3ea95e6147099a 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -812,10 +812,12 @@ protected theorem nnnorm {_ : MeasurableSpace α} {β : Type*} [SeminormedAddCom continuous_nnnorm.comp_stronglyMeasurable hf @[measurability] -protected theorem ennnorm {_ : MeasurableSpace α} {β : Type*} [SeminormedAddCommGroup β] - {f : α → β} (hf : StronglyMeasurable f) : Measurable fun a => (‖f a‖₊ : ℝ≥0∞) := +protected theorem enorm {_ : MeasurableSpace α} {β : Type*} [SeminormedAddCommGroup β] + {f : α → β} (hf : StronglyMeasurable f) : Measurable (‖f ·‖ₑ) := (ENNReal.continuous_coe.comp_stronglyMeasurable hf.nnnorm).measurable +@[deprecated (since := "2025-01-21")] alias ennnorm := StronglyMeasurable.enorm + @[measurability] protected theorem real_toNNReal {_ : MeasurableSpace α} {f : α → ℝ} (hf : StronglyMeasurable f) : StronglyMeasurable fun x => (f x).toNNReal := @@ -1438,10 +1440,12 @@ protected theorem nnnorm {β : Type*} [SeminormedAddCommGroup β] {f : α → β continuous_nnnorm.comp_aestronglyMeasurable hf @[measurability] -protected theorem ennnorm {β : Type*} [SeminormedAddCommGroup β] {f : α → β} - (hf : AEStronglyMeasurable f μ) : AEMeasurable (fun a => (‖f a‖₊ : ℝ≥0∞)) μ := +protected theorem enorm {β : Type*} [SeminormedAddCommGroup β] {f : α → β} + (hf : AEStronglyMeasurable f μ) : AEMeasurable (‖f ·‖ₑ) μ := (ENNReal.continuous_coe.comp_aestronglyMeasurable hf.nnnorm).aemeasurable +@[deprecated (since := "2025-01-20")] alias ennnorm := AEStronglyMeasurable.enorm + @[aesop safe 20 apply (rule_sets := [Measurable])] protected theorem edist {β : Type*} [SeminormedAddCommGroup β] {f g : α → β} (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasurable g μ) : diff --git a/Mathlib/MeasureTheory/Function/UnifTight.lean b/Mathlib/MeasureTheory/Function/UnifTight.lean index 6c36077b3a716f..1b98552847b15d 100644 --- a/Mathlib/MeasureTheory/Function/UnifTight.lean +++ b/Mathlib/MeasureTheory/Function/UnifTight.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Igor Khavkine -/ import Mathlib.MeasureTheory.Function.ConvergenceInMeasure -import Mathlib.MeasureTheory.Function.L1Space import Mathlib.MeasureTheory.Function.UniformIntegrable /-! diff --git a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean index 172756a4d9dbd6..8ba93267f9ee78 100644 --- a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ import Mathlib.MeasureTheory.Function.ConvergenceInMeasure -import Mathlib.MeasureTheory.Function.L1Space +import Mathlib.MeasureTheory.Function.L1Space.Integrable /-! # Uniform integrability @@ -45,7 +45,7 @@ uniform integrable, uniformly absolutely continuous integral, Vitali convergence noncomputable section -open scoped Classical MeasureTheory NNReal ENNReal Topology +open scoped MeasureTheory NNReal ENNReal Topology namespace MeasureTheory @@ -126,6 +126,7 @@ protected theorem sub (hf : UnifIntegrable f p μ) (hg : UnifIntegrable g p μ) protected theorem ae_eq (hf : UnifIntegrable f p μ) (hfg : ∀ n, f n =ᵐ[μ] g n) : UnifIntegrable g p μ := by + classical intro ε hε obtain ⟨δ, hδ_pos, hfδ⟩ := hf hε refine ⟨δ, hδ_pos, fun n s hs hμs => (le_of_eq <| eLpNorm_congr_ae ?_).trans (hfδ n s hs hμs)⟩ @@ -224,12 +225,12 @@ theorem Memℒp.integral_indicator_norm_ge_le (hf : Memℒp f 1 μ) (hmeas : Str which does not require measurability. -/ theorem Memℒp.integral_indicator_norm_ge_nonneg_le_of_meas (hf : Memℒp f 1 μ) (hmeas : StronglyMeasurable f) {ε : ℝ} (hε : 0 < ε) : - ∃ M : ℝ, 0 ≤ M ∧ (∫⁻ x, ‖{ x | M ≤ ‖f x‖₊ }.indicator f x‖₊ ∂μ) ≤ ENNReal.ofReal ε := + ∃ M : ℝ, 0 ≤ M ∧ (∫⁻ x, ‖{ x | M ≤ ‖f x‖₊ }.indicator f x‖ₑ ∂μ) ≤ ENNReal.ofReal ε := let ⟨M, hM⟩ := hf.integral_indicator_norm_ge_le hmeas hε ⟨max M 0, le_max_right _ _, by simpa⟩ theorem Memℒp.integral_indicator_norm_ge_nonneg_le (hf : Memℒp f 1 μ) {ε : ℝ} (hε : 0 < ε) : - ∃ M : ℝ, 0 ≤ M ∧ (∫⁻ x, ‖{ x | M ≤ ‖f x‖₊ }.indicator f x‖₊ ∂μ) ≤ ENNReal.ofReal ε := by + ∃ M : ℝ, 0 ≤ M ∧ (∫⁻ x, ‖{ x | M ≤ ‖f x‖₊ }.indicator f x‖ₑ ∂μ) ≤ ENNReal.ofReal ε := by have hf_mk : Memℒp (hf.1.mk f) 1 μ := (memℒp_congr_ae hf.1.ae_eq_mk).mp hf obtain ⟨M, hM_pos, hfM⟩ := hf_mk.integral_indicator_norm_ge_nonneg_le_of_meas hf.1.stronglyMeasurable_mk hε @@ -281,27 +282,25 @@ theorem Memℒp.eLpNorm_indicator_norm_ge_le (hf : Memℒp f p μ) (hmeas : Stro obtain ⟨M, hM', hM⟩ := Memℒp.integral_indicator_norm_ge_nonneg_le (μ := μ) (hf.norm_rpow hp_ne_zero hp_ne_top) (Real.rpow_pos_of_pos hε p.toReal) refine ⟨M ^ (1 / p.toReal), ?_⟩ - rw [eLpNorm_eq_lintegral_rpow_nnnorm hp_ne_zero hp_ne_top, ← ENNReal.rpow_one (ENNReal.ofReal ε)] + rw [eLpNorm_eq_lintegral_rpow_enorm hp_ne_zero hp_ne_top, ← ENNReal.rpow_one (ENNReal.ofReal ε)] conv_rhs => rw [← mul_one_div_cancel (ENNReal.toReal_pos hp_ne_zero hp_ne_top).ne.symm] rw [ENNReal.rpow_mul, ENNReal.rpow_le_rpow_iff (one_div_pos.2 <| ENNReal.toReal_pos hp_ne_zero hp_ne_top), ENNReal.ofReal_rpow_of_pos hε] - convert hM - rename_i x - rw [← ENNReal.coe_rpow_of_nonneg _ ENNReal.toReal_nonneg, nnnorm_indicator_eq_indicator_nnnorm, - nnnorm_indicator_eq_indicator_nnnorm] + convert hM using 3 with x + rw [enorm_indicator_eq_indicator_enorm, enorm_indicator_eq_indicator_enorm] have hiff : M ^ (1 / p.toReal) ≤ ‖f x‖₊ ↔ M ≤ ‖‖f x‖ ^ p.toReal‖₊ := by rw [coe_nnnorm, coe_nnnorm, Real.norm_rpow_of_nonneg (norm_nonneg _), norm_norm, ← Real.rpow_le_rpow_iff hM' (Real.rpow_nonneg (norm_nonneg _) _) (one_div_pos.2 <| ENNReal.toReal_pos hp_ne_zero hp_ne_top), ← Real.rpow_mul (norm_nonneg _), mul_one_div_cancel (ENNReal.toReal_pos hp_ne_zero hp_ne_top).ne.symm, Real.rpow_one] by_cases hx : x ∈ { x : α | M ^ (1 / p.toReal) ≤ ‖f x‖₊ } - · rw [Set.indicator_of_mem hx, Set.indicator_of_mem, Real.nnnorm_of_nonneg] - · rfl + · rw [Set.indicator_of_mem hx, Set.indicator_of_mem, Real.enorm_of_nonneg (by positivity), + ← ENNReal.ofReal_rpow_of_nonneg (norm_nonneg _) ENNReal.toReal_nonneg, ofReal_norm] rw [Set.mem_setOf_eq] rwa [← hiff] · rw [Set.indicator_of_not_mem hx, Set.indicator_of_not_mem] - · simp [(ENNReal.toReal_pos hp_ne_zero hp_ne_top).ne.symm] + · simp [ENNReal.toReal_pos hp_ne_zero hp_ne_top] · rw [Set.mem_setOf_eq] rwa [← hiff] @@ -499,7 +498,7 @@ theorem eLpNorm_sub_le_of_dist_bdd (μ : Measure α) refine le_trans (eLpNorm_mono this) ?_ rw [eLpNorm_indicator_const hs hp hp'] refine mul_le_mul_right' (le_of_eq ?_) _ - rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, abs_of_nonneg hc] + rw [← ofReal_norm_eq_enorm, Real.norm_eq_abs, abs_of_nonneg hc] @[deprecated (since := "2024-07-27")] alias snorm_sub_le_of_dist_bdd := eLpNorm_sub_le_of_dist_bdd @@ -920,33 +919,23 @@ theorem uniformIntegrable_average refine ⟨δ, hδ₁, fun n s hs hle => ?_⟩ simp_rw [Finset.smul_sum, Finset.indicator_sum] refine le_trans (eLpNorm_sum_le (fun i _ => ((hf₁ i).const_smul _).indicator hs) hp) ?_ - have : ∀ i, s.indicator ((n : ℝ) ⁻¹ • f i) = (↑n : ℝ)⁻¹ • s.indicator (f i) := - fun i ↦ indicator_const_smul _ _ _ - simp_rw [this, eLpNorm_const_smul, ← Finset.mul_sum, nnnorm_inv, Real.nnnorm_natCast] - by_cases hn : (↑(↑n : ℝ≥0)⁻¹ : ℝ≥0∞) = 0 - · simp only [hn, zero_mul, zero_le] - refine le_trans ?_ (?_ : ↑(↑n : ℝ≥0)⁻¹ * n • ENNReal.ofReal ε ≤ ENNReal.ofReal ε) - · refine (ENNReal.mul_le_mul_left hn ENNReal.coe_ne_top).2 ?_ - conv_rhs => rw [← Finset.card_range n] - exact Finset.sum_le_card_nsmul _ _ _ fun i _ => hδ₂ _ _ hs hle - · simp only [ENNReal.coe_eq_zero, inv_eq_zero, Nat.cast_eq_zero] at hn - rw [nsmul_eq_mul, ← mul_assoc, ENNReal.coe_inv, ENNReal.coe_natCast, - ENNReal.inv_mul_cancel _ (ENNReal.natCast_ne_top _), one_mul] - all_goals simpa only [Ne, Nat.cast_eq_zero] + have this i : s.indicator ((n : ℝ) ⁻¹ • f i) = (↑n : ℝ)⁻¹ • s.indicator (f i) := + indicator_const_smul _ _ _ + obtain rfl | hn := eq_or_ne n 0 + · simp + simp_rw [this, eLpNorm_const_smul, ← Finset.mul_sum] + rw [enorm_inv (by positivity), Real.enorm_natCast, ← ENNReal.div_eq_inv_mul] + refine ENNReal.div_le_of_le_mul' ?_ + simpa using Finset.sum_le_card_nsmul (.range n) _ _ fun i _ => hδ₂ _ _ hs hle · obtain ⟨C, hC⟩ := hf₃ simp_rw [Finset.smul_sum] refine ⟨C, fun n => (eLpNorm_sum_le (fun i _ => (hf₁ i).const_smul _) hp).trans ?_⟩ - simp_rw [eLpNorm_const_smul, ← Finset.mul_sum, nnnorm_inv, Real.nnnorm_natCast] - by_cases hn : (↑(↑n : ℝ≥0)⁻¹ : ℝ≥0∞) = 0 - · simp only [hn, zero_mul, zero_le] - refine le_trans ?_ (?_ : ↑(↑n : ℝ≥0)⁻¹ * (n • C : ℝ≥0∞) ≤ C) - · refine (ENNReal.mul_le_mul_left hn ENNReal.coe_ne_top).2 ?_ - conv_rhs => rw [← Finset.card_range n] - exact Finset.sum_le_card_nsmul _ _ _ fun i _ => hC i - · simp only [ENNReal.coe_eq_zero, inv_eq_zero, Nat.cast_eq_zero] at hn - rw [nsmul_eq_mul, ← mul_assoc, ENNReal.coe_inv, ENNReal.coe_natCast, - ENNReal.inv_mul_cancel _ (ENNReal.natCast_ne_top _), one_mul] - all_goals simpa only [Ne, Nat.cast_eq_zero] + obtain rfl | hn := eq_or_ne n 0 + · simp + simp_rw [eLpNorm_const_smul, ← Finset.mul_sum] + rw [enorm_inv (by positivity), Real.enorm_natCast, ← ENNReal.div_eq_inv_mul] + refine ENNReal.div_le_of_le_mul' ?_ + simpa using Finset.sum_le_card_nsmul (.range n) _ _ fun i _ => hC i /-- The averaging of a uniformly integrable real-valued sequence is also uniformly integrable. -/ theorem uniformIntegrable_average_real (hp : 1 ≤ p) {f : ℕ → α → ℝ} (hf : UniformIntegrable f p μ) : diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 0dc9999c175a6f..62c46156f922a0 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -764,9 +764,7 @@ nonrec theorem _root_.MulEquiv.isHaarMeasure_map [BorelSpace G] [TopologicalGrou [TopologicalGroup H] (e : G ≃* H) (he : Continuous e) (hesymm : Continuous e.symm) : IsHaarMeasure (Measure.map e μ) := let f : G ≃ₜ H := .mk e - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/6024 + #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 we needed to write `e.toMonoidHom` instead of just `e`, to avoid unification issues. -/ isHaarMeasure_map μ e.toMonoidHom he e.surjective f.isClosedEmbedding.tendsto_cocompact diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 51627a057494c3..a9ff72df2468ee 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -868,12 +868,14 @@ theorem norm_integral_le_lintegral_norm (f : α → G) : · rw [integral_undef hf, norm_zero]; exact toReal_nonneg · simp [integral, hG] -theorem ennnorm_integral_le_lintegral_ennnorm (f : α → G) : - (‖∫ a, f a ∂μ‖₊ : ℝ≥0∞) ≤ ∫⁻ a, ‖f a‖₊ ∂μ := by - simp_rw [← ofReal_norm_eq_coe_nnnorm] +theorem enorm_integral_le_lintegral_enorm (f : α → G) : ‖∫ a, f a ∂μ‖ₑ ≤ ∫⁻ a, ‖f a‖ₑ ∂μ := by + simp_rw [← ofReal_norm_eq_enorm] apply ENNReal.ofReal_le_of_le_toReal exact norm_integral_le_lintegral_norm f +@[deprecated (since := "2025-01-21")] +alias ennnorm_integral_le_lintegral_ennnorm := enorm_integral_le_lintegral_enorm + theorem integral_eq_zero_of_ae {f : α → G} (hf : f =ᵐ[μ] 0) : ∫ a, f a ∂μ = 0 := by simp [integral_congr_ae hf, integral_zero] @@ -887,7 +889,7 @@ theorem HasFiniteIntegral.tendsto_setIntegral_nhds_zero {ι} {f : α → G} ENNReal.coe_zero] exact tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds (tendsto_setLIntegral_zero (ne_of_lt hf) hs) (fun i => zero_le _) - fun i => ennnorm_integral_le_lintegral_ennnorm _ + fun i => enorm_integral_le_lintegral_enorm _ @[deprecated (since := "2024-04-17")] alias HasFiniteIntegral.tendsto_set_integral_nhds_zero := @@ -907,7 +909,7 @@ alias Integrable.tendsto_set_integral_nhds_zero := /-- If `F i → f` in `L1`, then `∫ x, F i x ∂μ → ∫ x, f x ∂μ`. -/ theorem tendsto_integral_of_L1 {ι} (f : α → G) (hfi : Integrable f μ) {F : ι → α → G} {l : Filter ι} (hFi : ∀ᶠ i in l, Integrable (F i) μ) - (hF : Tendsto (fun i => ∫⁻ x, ‖F i x - f x‖₊ ∂μ) l (𝓝 0)) : + (hF : Tendsto (fun i => ∫⁻ x, ‖F i x - f x‖ₑ ∂μ) l (𝓝 0)) : Tendsto (fun i => ∫ x, F i x ∂μ) l (𝓝 <| ∫ x, f x ∂μ) := by by_cases hG : CompleteSpace G · simp only [integral, hG, L1.integral] @@ -919,18 +921,18 @@ lemma tendsto_integral_of_L1' {ι} (f : α → G) (hfi : Integrable f μ) {F : (hFi : ∀ᶠ i in l, Integrable (F i) μ) (hF : Tendsto (fun i ↦ eLpNorm (F i - f) 1 μ) l (𝓝 0)) : Tendsto (fun i ↦ ∫ x, F i x ∂μ) l (𝓝 (∫ x, f x ∂μ)) := by refine tendsto_integral_of_L1 f hfi hFi ?_ - simp_rw [eLpNorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF + simp_rw [eLpNorm_one_eq_lintegral_enorm, Pi.sub_apply] at hF exact hF /-- If `F i → f` in `L1`, then `∫ x in s, F i x ∂μ → ∫ x in s, f x ∂μ`. -/ lemma tendsto_setIntegral_of_L1 {ι} (f : α → G) (hfi : Integrable f μ) {F : ι → α → G} {l : Filter ι} - (hFi : ∀ᶠ i in l, Integrable (F i) μ) (hF : Tendsto (fun i ↦ ∫⁻ x, ‖F i x - f x‖₊ ∂μ) l (𝓝 0)) + (hFi : ∀ᶠ i in l, Integrable (F i) μ) (hF : Tendsto (fun i ↦ ∫⁻ x, ‖F i x - f x‖ₑ ∂μ) l (𝓝 0)) (s : Set α) : Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by refine tendsto_integral_of_L1 f hfi.restrict ?_ ?_ · filter_upwards [hFi] with i hi using hi.restrict - · simp_rw [← eLpNorm_one_eq_lintegral_nnnorm] at hF ⊢ + · simp_rw [← eLpNorm_one_eq_lintegral_enorm] at hF ⊢ exact tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds hF (fun _ ↦ zero_le') (fun _ ↦ eLpNorm_mono_measure _ Measure.restrict_le_self) @@ -944,7 +946,7 @@ lemma tendsto_setIntegral_of_L1' {ι} (f : α → G) (hfi : Integrable f μ) {F (s : Set α) : Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by refine tendsto_setIntegral_of_L1 f hfi hFi ?_ s - simp_rw [eLpNorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF + simp_rw [eLpNorm_one_eq_lintegral_enorm, Pi.sub_apply] at hF exact hF @[deprecated (since := "2024-04-17")] @@ -1048,16 +1050,22 @@ theorem integral_eq_lintegral_of_nonneg_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ] f rw [Real.norm_eq_abs, abs_of_nonneg h] rw [this, hfi]; rfl -theorem integral_norm_eq_lintegral_nnnorm {P : Type*} [NormedAddCommGroup P] {f : α → P} - (hf : AEStronglyMeasurable f μ) : ∫ x, ‖f x‖ ∂μ = ENNReal.toReal (∫⁻ x, ‖f x‖₊ ∂μ) := by +theorem integral_norm_eq_lintegral_enorm {P : Type*} [NormedAddCommGroup P] {f : α → P} + (hf : AEStronglyMeasurable f μ) : ∫ x, ‖f x‖ ∂μ = (∫⁻ x, ‖f x‖ₑ ∂μ).toReal := by rw [integral_eq_lintegral_of_nonneg_ae _ hf.norm] - · simp_rw [ofReal_norm_eq_coe_nnnorm] + · simp_rw [ofReal_norm_eq_enorm] · filter_upwards; simp_rw [Pi.zero_apply, norm_nonneg, imp_true_iff] -theorem ofReal_integral_norm_eq_lintegral_nnnorm {P : Type*} [NormedAddCommGroup P] {f : α → P} - (hf : Integrable f μ) : ENNReal.ofReal (∫ x, ‖f x‖ ∂μ) = ∫⁻ x, ‖f x‖₊ ∂μ := by - rw [integral_norm_eq_lintegral_nnnorm hf.aestronglyMeasurable, ENNReal.ofReal_toReal] - exact lt_top_iff_ne_top.mp (hasFiniteIntegral_iff_nnnorm.mpr hf.2) +@[deprecated (since := "2025-01-21")] +alias integral_norm_eq_lintegral_nnnorm := integral_norm_eq_lintegral_enorm + +theorem ofReal_integral_norm_eq_lintegral_enorm {P : Type*} [NormedAddCommGroup P] {f : α → P} + (hf : Integrable f μ) : ENNReal.ofReal (∫ x, ‖f x‖ ∂μ) = ∫⁻ x, ‖f x‖ₑ ∂μ := by + rw [integral_norm_eq_lintegral_enorm hf.aestronglyMeasurable, ENNReal.ofReal_toReal] + exact lt_top_iff_ne_top.mp (hasFiniteIntegral_iff_enorm.mpr hf.2) + +@[deprecated (since := "2025-01-21")] +alias ofReal_integral_norm_eq_lintegral_nnnorm := ofReal_integral_norm_eq_lintegral_enorm theorem integral_eq_integral_pos_part_sub_integral_neg_part {f : α → ℝ} (hf : Integrable f μ) : ∫ a, f a ∂μ = ∫ a, (Real.toNNReal (f a) : ℝ) ∂μ - ∫ a, (Real.toNNReal (-f a) : ℝ) ∂μ := by @@ -1079,13 +1087,13 @@ theorem lintegral_coe_eq_integral (f : α → ℝ≥0) (hfi : Integrable (fun x rw [← lt_top_iff_ne_top] convert hfi.hasFiniteIntegral -- Porting note: `convert` no longer unfolds `HasFiniteIntegral` - simp_rw [hasFiniteIntegral_iff_nnnorm, NNReal.nnnorm_eq] + simp_rw [hasFiniteIntegral_iff_enorm, NNReal.enorm_eq] theorem ofReal_integral_eq_lintegral_ofReal {f : α → ℝ} (hfi : Integrable f μ) (f_nn : 0 ≤ᵐ[μ] f) : ENNReal.ofReal (∫ x, f x ∂μ) = ∫⁻ x, ENNReal.ofReal (f x) ∂μ := by have : f =ᵐ[μ] (‖f ·‖) := f_nn.mono fun _x hx ↦ (abs_of_nonneg hx).symm - simp_rw [integral_congr_ae this, ofReal_integral_norm_eq_lintegral_nnnorm hfi, - ← ofReal_norm_eq_coe_nnnorm] + simp_rw [integral_congr_ae this, ofReal_integral_norm_eq_lintegral_enorm hfi, + ← ofReal_norm_eq_enorm] exact lintegral_congr_ae (this.symm.fun_comp ENNReal.ofReal) theorem integral_toReal {f : α → ℝ≥0∞} (hfm : AEMeasurable f μ) (hf : ∀ᵐ x ∂μ, f x < ∞) : @@ -1249,7 +1257,7 @@ lemma tendsto_of_integral_tendsto_of_monotone {μ : Measure α} {f : ℕ → α have h := tendsto_of_lintegral_tendsto_of_monotone ?_ h_tendsto h_mono h_bound ?_ rotate_left · exact (hF_int.1.aemeasurable.sub (hf_int 0).1.aemeasurable).ennreal_ofReal - · exact ((lintegral_ofReal_le_lintegral_nnnorm _).trans_lt (hF_int.sub (hf_int 0)).2).ne + · exact ((lintegral_ofReal_le_lintegral_enorm _).trans_lt (hF_int.sub (hf_int 0)).2).ne filter_upwards [h, hf_mono, hf_bound] with a ha ha_mono ha_bound have h1 : (fun i ↦ f i a) = fun i ↦ (f' i a).toReal + f 0 a := by unfold f' @@ -1296,11 +1304,11 @@ section NormedAddCommGroup variable {H : Type*} [NormedAddCommGroup H] theorem L1.norm_eq_integral_norm (f : α →₁[μ] H) : ‖f‖ = ∫ a, ‖f a‖ ∂μ := by - simp only [eLpNorm, eLpNorm'_eq_lintegral_nnnorm, ENNReal.one_toReal, ENNReal.rpow_one, + simp only [eLpNorm, eLpNorm'_eq_lintegral_enorm, ENNReal.one_toReal, ENNReal.rpow_one, Lp.norm_def, if_false, ENNReal.one_ne_top, one_ne_zero, _root_.div_one] rw [integral_eq_lintegral_of_nonneg_ae (Eventually.of_forall (by simp [norm_nonneg])) (Lp.aestronglyMeasurable f).norm] - simp [ofReal_norm_eq_coe_nnnorm] + simp [ofReal_norm_eq_enorm] theorem L1.dist_eq_integral_dist (f g : α →₁[μ] H) : dist f g = ∫ a, dist (f a) (g a) ∂μ := by simp only [dist_eq_norm, L1.norm_eq_integral_norm] @@ -1314,14 +1322,14 @@ theorem L1.norm_of_fun_eq_integral_norm {f : α → H} (hf : Integrable f μ) : theorem Memℒp.eLpNorm_eq_integral_rpow_norm {f : α → H} {p : ℝ≥0∞} (hp1 : p ≠ 0) (hp2 : p ≠ ∞) (hf : Memℒp f p μ) : eLpNorm f p μ = ENNReal.ofReal ((∫ a, ‖f a‖ ^ p.toReal ∂μ) ^ p.toReal⁻¹) := by - have A : ∫⁻ a : α, ENNReal.ofReal (‖f a‖ ^ p.toReal) ∂μ = ∫⁻ a : α, ‖f a‖₊ ^ p.toReal ∂μ := by - simp_rw [← ofReal_rpow_of_nonneg (norm_nonneg _) toReal_nonneg, ofReal_norm_eq_coe_nnnorm] - simp only [eLpNorm_eq_lintegral_rpow_nnnorm hp1 hp2, one_div] + have A : ∫⁻ a : α, ENNReal.ofReal (‖f a‖ ^ p.toReal) ∂μ = ∫⁻ a : α, ‖f a‖ₑ ^ p.toReal ∂μ := by + simp_rw [← ofReal_rpow_of_nonneg (norm_nonneg _) toReal_nonneg, ofReal_norm_eq_enorm] + simp only [eLpNorm_eq_lintegral_rpow_enorm hp1 hp2, one_div] rw [integral_eq_lintegral_of_nonneg_ae]; rotate_left · exact ae_of_all _ fun x => by positivity · exact (hf.aestronglyMeasurable.norm.aemeasurable.pow_const _).aestronglyMeasurable rw [A, ← ofReal_rpow_of_nonneg toReal_nonneg (inv_nonneg.2 toReal_nonneg), ofReal_toReal] - exact (lintegral_rpow_nnnorm_lt_top_of_eLpNorm_lt_top hp1 hp2 hf.2).ne + exact (lintegral_rpow_enorm_lt_top_of_eLpNorm_lt_top hp1 hp2 hf.2).ne @[deprecated (since := "2024-07-27")] alias Memℒp.snorm_eq_integral_rpow_norm := Memℒp.eLpNorm_eq_integral_rpow_norm @@ -1428,8 +1436,8 @@ theorem tendsto_integral_norm_approxOn_sub (fmeas : Measurable f) (hf : Integrable f μ) [SeparableSpace (range f ∪ {0} : Set E)] : Tendsto (fun n ↦ ∫ x, ‖SimpleFunc.approxOn f fmeas (range f ∪ {0}) 0 (by simp) n x - f x‖ ∂μ) atTop (𝓝 0) := by - convert (tendsto_toReal zero_ne_top).comp (tendsto_approxOn_range_L1_nnnorm fmeas hf) with n - rw [integral_norm_eq_lintegral_nnnorm] + convert (tendsto_toReal zero_ne_top).comp (tendsto_approxOn_range_L1_enorm fmeas hf) with n + rw [integral_norm_eq_lintegral_enorm] · simp · apply (SimpleFunc.aestronglyMeasurable _).sub apply (stronglyMeasurable_iff_measurable_separable.2 ⟨fmeas, ?_⟩ ).aestronglyMeasurable @@ -1484,9 +1492,9 @@ theorem integral_finset_sum_measure {ι} {m : MeasurableSpace α} {f : α → G} theorem nndist_integral_add_measure_le_lintegral {f : α → G} (h₁ : Integrable f μ) (h₂ : Integrable f ν) : - (nndist (∫ x, f x ∂μ) (∫ x, f x ∂(μ + ν)) : ℝ≥0∞) ≤ ∫⁻ x, ‖f x‖₊ ∂ν := by + (nndist (∫ x, f x ∂μ) (∫ x, f x ∂(μ + ν)) : ℝ≥0∞) ≤ ∫⁻ x, ‖f x‖ₑ ∂ν := by rw [integral_add_measure h₁ h₂, nndist_comm, nndist_eq_nnnorm, add_sub_cancel_left] - exact ennnorm_integral_le_lintegral_ennnorm _ + exact enorm_integral_le_lintegral_enorm _ theorem hasSum_integral_measure {ι} {m : MeasurableSpace α} {f : α → G} {μ : ι → Measure α} (hf : Integrable f (Measure.sum μ)) : @@ -1495,11 +1503,11 @@ theorem hasSum_integral_measure {ι} {m : MeasurableSpace α} {f : α → G} {μ simp only [HasSum, ← integral_finset_sum_measure fun i _ => hfi i] refine Metric.nhds_basis_ball.tendsto_right_iff.mpr fun ε ε0 => ?_ lift ε to ℝ≥0 using ε0.le - have hf_lt : (∫⁻ x, ‖f x‖₊ ∂Measure.sum μ) < ∞ := hf.2 - have hmem : ∀ᶠ y in 𝓝 (∫⁻ x, ‖f x‖₊ ∂Measure.sum μ), (∫⁻ x, ‖f x‖₊ ∂Measure.sum μ) < y + ε := by + have hf_lt : (∫⁻ x, ‖f x‖ₑ ∂Measure.sum μ) < ∞ := hf.2 + have hmem : ∀ᶠ y in 𝓝 (∫⁻ x, ‖f x‖ₑ ∂Measure.sum μ), (∫⁻ x, ‖f x‖ₑ ∂Measure.sum μ) < y + ε := by refine tendsto_id.add tendsto_const_nhds (lt_mem_nhds (α := ℝ≥0∞) <| ENNReal.lt_add_right ?_ ?_) exacts [hf_lt.ne, ENNReal.coe_ne_zero.2 (NNReal.coe_ne_zero.1 ε0.ne')] - refine ((hasSum_lintegral_measure (fun x => ‖f x‖₊) μ).eventually hmem).mono fun s hs => ?_ + refine ((hasSum_lintegral_measure (fun x => ‖f x‖ₑ) μ).eventually hmem).mono fun s hs => ?_ obtain ⟨ν, hν⟩ : ∃ ν, (∑ i ∈ s, μ i) + ν = Measure.sum μ := by refine ⟨Measure.sum fun i : ↥(sᶜ : Set ι) => μ i, ?_⟩ simpa only [← Measure.sum_coe_finset] using Measure.sum_add_sum_compl (s : Set ι) μ @@ -1677,26 +1685,26 @@ theorem integral_mul_norm_le_Lp_mul_Lq {E} [NormedAddCommGroup E] {f g : α → rw [ENNReal.toReal_rpow, ENNReal.toReal_rpow, ← ENNReal.toReal_mul] -- replace norms by nnnorm have h_left : ∫⁻ a, ENNReal.ofReal (‖f a‖ * ‖g a‖) ∂μ = - ∫⁻ a, ((fun x => (‖f x‖₊ : ℝ≥0∞)) * fun x => (‖g x‖₊ : ℝ≥0∞)) a ∂μ := by - simp_rw [Pi.mul_apply, ← ofReal_norm_eq_coe_nnnorm, ENNReal.ofReal_mul (norm_nonneg _)] - have h_right_f : ∫⁻ a, ENNReal.ofReal (‖f a‖ ^ p) ∂μ = ∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ p ∂μ := by + ∫⁻ a, ((‖f ·‖ₑ) * (‖g ·‖ₑ)) a ∂μ := by + simp_rw [Pi.mul_apply, ← ofReal_norm_eq_enorm, ENNReal.ofReal_mul (norm_nonneg _)] + have h_right_f : ∫⁻ a, .ofReal (‖f a‖ ^ p) ∂μ = ∫⁻ a, ‖f a‖ₑ ^ p ∂μ := by refine lintegral_congr fun x => ?_ - rw [← ofReal_norm_eq_coe_nnnorm, ENNReal.ofReal_rpow_of_nonneg (norm_nonneg _) hpq.nonneg] - have h_right_g : ∫⁻ a, ENNReal.ofReal (‖g a‖ ^ q) ∂μ = ∫⁻ a, (‖g a‖₊ : ℝ≥0∞) ^ q ∂μ := by + rw [← ofReal_norm_eq_enorm, ENNReal.ofReal_rpow_of_nonneg (norm_nonneg _) hpq.nonneg] + have h_right_g : ∫⁻ a, .ofReal (‖g a‖ ^ q) ∂μ = ∫⁻ a, ‖g a‖ₑ ^ q ∂μ := by refine lintegral_congr fun x => ?_ - rw [← ofReal_norm_eq_coe_nnnorm, ENNReal.ofReal_rpow_of_nonneg (norm_nonneg _) hpq.symm.nonneg] + rw [← ofReal_norm_eq_enorm, ENNReal.ofReal_rpow_of_nonneg (norm_nonneg _) hpq.symm.nonneg] rw [h_left, h_right_f, h_right_g] -- we can now apply `ENNReal.lintegral_mul_le_Lp_mul_Lq` (up to the `toReal` application) refine ENNReal.toReal_mono ?_ ?_ · refine ENNReal.mul_ne_top ?_ ?_ · convert hf.eLpNorm_ne_top - rw [eLpNorm_eq_lintegral_rpow_nnnorm] + rw [eLpNorm_eq_lintegral_rpow_enorm] · rw [ENNReal.toReal_ofReal hpq.nonneg] · rw [Ne, ENNReal.ofReal_eq_zero, not_le] exact hpq.pos · exact ENNReal.coe_ne_top · convert hg.eLpNorm_ne_top - rw [eLpNorm_eq_lintegral_rpow_nnnorm] + rw [eLpNorm_eq_lintegral_rpow_enorm] · rw [ENNReal.toReal_ofReal hpq.symm.nonneg] · rw [Ne, ENNReal.ofReal_eq_zero, not_le] exact hpq.symm.pos @@ -1838,12 +1846,12 @@ theorem integral_trim (hm : m ≤ m0) {f : β → G} (hf : StronglyMeasurable[m] integral_trim_simpleFunc hm (f_seq n) (hf_seq_int n) have h_lim_1 : atTop.Tendsto (fun n => ∫ x, f_seq n x ∂μ) (𝓝 (∫ x, f x ∂μ)) := by refine tendsto_integral_of_L1 f hf_int (Eventually.of_forall hf_seq_int) ?_ - exact SimpleFunc.tendsto_approxOn_range_L1_nnnorm (hf.mono hm).measurable hf_int + exact SimpleFunc.tendsto_approxOn_range_L1_enorm (hf.mono hm).measurable hf_int have h_lim_2 : atTop.Tendsto (fun n => ∫ x, f_seq n x ∂μ) (𝓝 (∫ x, f x ∂μ.trim hm)) := by simp_rw [hf_seq_eq] refine @tendsto_integral_of_L1 β G _ _ m (μ.trim hm) _ f (hf_int.trim hm hf) _ _ (Eventually.of_forall hf_seq_int_m) ?_ - exact @SimpleFunc.tendsto_approxOn_range_L1_nnnorm β G m _ _ _ f _ _ hf.measurable + exact @SimpleFunc.tendsto_approxOn_range_L1_enorm β G m _ _ _ f _ _ hf.measurable (hf_int.trim hm hf) exact tendsto_nhds_unique h_lim_1 h_lim_2 diff --git a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean index d76ba1fe92b5b3..6eb67e7491eb74 100644 --- a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean +++ b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean @@ -47,7 +47,7 @@ theorem lintegral_lt_top_of_nnreal (f : X →ᵇ ℝ≥0) : ∫⁻ x, f x ∂μ theorem integrable_of_nnreal [OpensMeasurableSpace X] (f : X →ᵇ ℝ≥0) : Integrable (((↑) : ℝ≥0 → ℝ) ∘ ⇑f) μ := by refine ⟨(NNReal.continuous_coe.comp f.continuous).measurable.aestronglyMeasurable, ?_⟩ - simp only [hasFiniteIntegral_iff_nnnorm, Function.comp_apply, NNReal.nnnorm_eq] + simp only [hasFiniteIntegral_iff_enorm, Function.comp_apply, NNReal.enorm_eq] exact lintegral_lt_top_of_nnreal _ f theorem integral_eq_integral_nnrealPart_sub [OpensMeasurableSpace X] (f : X →ᵇ ℝ) : diff --git a/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean b/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean index 827108c784d82d..95aead3dd63b1f 100644 --- a/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean +++ b/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean @@ -56,7 +56,7 @@ divergence theorem, Bochner integral open Set Finset TopologicalSpace Function BoxIntegral MeasureTheory Filter -open scoped Classical Topology Interval +open scoped Topology Interval universe u diff --git a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean index 3a7713df35ddae..063908dfd80a56 100644 --- a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean +++ b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean @@ -102,11 +102,11 @@ theorem hasSum_integral_of_dominated_convergence {ι} [Countable ι] {F : ι → _ ≤ ∑' n, bound n a := sum_le_tsum _ (fun n _ => ha0 n) has theorem integral_tsum {ι} [Countable ι] {f : ι → α → G} (hf : ∀ i, AEStronglyMeasurable (f i) μ) - (hf' : ∑' i, ∫⁻ a : α, ‖f i a‖₊ ∂μ ≠ ∞) : + (hf' : ∑' i, ∫⁻ a : α, ‖f i a‖ₑ ∂μ ≠ ∞) : ∫ a : α, ∑' i, f i a ∂μ = ∑' i, ∫ a : α, f i a ∂μ := by by_cases hG : CompleteSpace G; swap · simp [integral, hG] - have hf'' : ∀ i, AEMeasurable (fun x => (‖f i x‖₊ : ℝ≥0∞)) μ := fun i => (hf i).ennnorm + have hf'' i : AEMeasurable (‖f i ·‖ₑ) μ := (hf i).enorm have hhh : ∀ᵐ a : α ∂μ, Summable fun n => (‖f n a‖₊ : ℝ) := by rw [← lintegral_tsum hf''] at hf' refine (ae_lt_top' (AEMeasurable.ennreal_tsum hf'') hf').mono ?_ @@ -124,7 +124,7 @@ theorem integral_tsum {ι} [Countable ι] {f : ι → α → G} (hf : ∀ i, AES apply AEMeasurable.nnreal_tsum exact fun i => (hf i).nnnorm.aemeasurable · dsimp [HasFiniteIntegral] - have : ∫⁻ a, ∑' n, ‖f n a‖₊ ∂μ < ⊤ := by rwa [lintegral_tsum hf'', lt_top_iff_ne_top] + have : ∫⁻ a, ∑' n, ‖f n a‖ₑ ∂μ < ⊤ := by rwa [lintegral_tsum hf'', lt_top_iff_ne_top] convert this using 1 apply lintegral_congr_ae simp_rw [← coe_nnnorm, ← NNReal.coe_tsum, enorm_eq_nnnorm, NNReal.nnnorm_eq] @@ -140,14 +140,13 @@ lemma hasSum_integral_of_summable_integral_norm {ι} [Countable ι] {F : ι → · simp [integral, hE, hasSum_zero] rw [integral_tsum (fun i ↦ (hF_int i).1)] · exact (hF_sum.of_norm_bounded _ fun i ↦ norm_integral_le_integral_norm _).hasSum - have (i : ι) : ∫⁻ (a : α), ‖F i a‖₊ ∂μ = ‖(∫ a : α, ‖F i a‖ ∂μ)‖₊ := by + have (i : ι) : ∫⁻ a, ‖F i a‖ₑ ∂μ = ‖∫ a, ‖F i a‖ ∂μ‖ₑ := by + dsimp [enorm] rw [lintegral_coe_eq_integral _ (hF_int i).norm, coe_nnreal_eq, coe_nnnorm, Real.norm_of_nonneg (integral_nonneg (fun a ↦ norm_nonneg (F i a)))] simp only [coe_nnnorm] - rw [funext this, ← ENNReal.coe_tsum] - · apply coe_ne_top - · simp_rw [← NNReal.summable_coe, coe_nnnorm] - exact hF_sum.abs + rw [funext this] + exact ENNReal.tsum_coe_ne_top_iff_summable.2 <| NNReal.summable_coe.1 hF_sum.abs lemma integral_tsum_of_summable_integral_norm {ι} [Countable ι] {F : ι → α → E} (hF_int : ∀ i : ι, Integrable (F i) μ) (hF_sum : Summable fun i ↦ ∫ a, ‖F i a‖ ∂μ) : diff --git a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean index a88490002c8be2..788e73736ccf9f 100644 --- a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean +++ b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean @@ -141,10 +141,9 @@ integral, fundamental theorem of calculus, FTC-1, FTC-2, change of variables in noncomputable section -open scoped Classical open MeasureTheory Set Filter Function -open scoped Classical Topology Filter ENNReal Interval NNReal +open scoped Topology Filter ENNReal Interval NNReal variable {ι 𝕜 E A : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] @@ -1223,7 +1222,7 @@ theorem integrableOn_deriv_right_of_nonneg (hcont : ContinuousOn g (Icc a b)) let F : ℝ → ℝ := (↑) ∘ f have intF : IntegrableOn F (Ioo a b) := by refine ⟨f.measurable.coe_nnreal_real.aestronglyMeasurable, ?_⟩ - simpa only [F, hasFiniteIntegral_iff_nnnorm, comp_apply, NNReal.nnnorm_eq] using fint + simpa only [F, hasFiniteIntegral_iff_enorm, comp_apply, NNReal.enorm_eq] using fint have A : ∫⁻ x : ℝ in Ioo a b, f x = ENNReal.ofReal (∫ x in Ioo a b, F x) := lintegral_coe_eq_integral _ intF rw [A] at hf diff --git a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean index 9093ea80babe21..d5c7a91ba8e6ae 100644 --- a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean +++ b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Zhouhang Zhou, Yury Kudryashov -/ -import Mathlib.MeasureTheory.Function.L1Space +import Mathlib.MeasureTheory.Function.L1Space.Integrable import Mathlib.Analysis.NormedSpace.IndicatorFunction /-! # Functions integrable on a set and at a filter @@ -232,8 +232,8 @@ theorem MeasurePreserving.integrableOn_image [MeasurableSpace β] {e : α → β theorem integrable_indicator_iff (hs : MeasurableSet s) : Integrable (indicator s f) μ ↔ IntegrableOn f s μ := by - simp_rw [IntegrableOn, Integrable, hasFiniteIntegral_iff_nnnorm, - nnnorm_indicator_eq_indicator_nnnorm, ENNReal.coe_indicator, lintegral_indicator hs, + simp_rw [IntegrableOn, Integrable, hasFiniteIntegral_iff_enorm, + enorm_indicator_eq_indicator_enorm, lintegral_indicator hs, aestronglyMeasurable_indicator_iff hs] theorem IntegrableOn.integrable_indicator (h : IntegrableOn f s μ) (hs : MeasurableSet s) : @@ -335,7 +335,7 @@ theorem integrableOn_Lp_of_measure_ne_top {E} [NormedAddCommGroup E] {p : ℝ≥ theorem Integrable.lintegral_lt_top {f : α → ℝ} (hf : Integrable f μ) : (∫⁻ x, ENNReal.ofReal (f x) ∂μ) < ∞ := calc - (∫⁻ x, ENNReal.ofReal (f x) ∂μ) ≤ ∫⁻ x, ↑‖f x‖₊ ∂μ := lintegral_ofReal_le_lintegral_nnnorm f + (∫⁻ x, ENNReal.ofReal (f x) ∂μ) ≤ ∫⁻ x, ↑‖f x‖₊ ∂μ := lintegral_ofReal_le_lintegral_enorm f _ < ∞ := hf.2 theorem IntegrableOn.setLIntegral_lt_top {f : α → ℝ} {s : Set α} (hf : IntegrableOn f s μ) : diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index 99c22b52d22c31..e0363619f99bbb 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -396,46 +396,62 @@ section Integrable variable {α ι E : Type*} [MeasurableSpace α] {μ : Measure α} {l : Filter ι} [NormedAddCommGroup E] -theorem AECover.integrable_of_lintegral_nnnorm_bounded [l.NeBot] [l.IsCountablyGenerated] +theorem AECover.integrable_of_lintegral_enorm_bounded [l.NeBot] [l.IsCountablyGenerated] {φ : ι → Set α} (hφ : AECover μ l φ) {f : α → E} (I : ℝ) (hfm : AEStronglyMeasurable f μ) - (hbounded : ∀ᶠ i in l, (∫⁻ x in φ i, ‖f x‖₊ ∂μ) ≤ ENNReal.ofReal I) : Integrable f μ := by + (hbounded : ∀ᶠ i in l, ∫⁻ x in φ i, ‖f x‖ₑ ∂μ ≤ ENNReal.ofReal I) : Integrable f μ := by refine ⟨hfm, (le_of_tendsto ?_ hbounded).trans_lt ENNReal.ofReal_lt_top⟩ - exact hφ.lintegral_tendsto_of_countably_generated hfm.ennnorm + exact hφ.lintegral_tendsto_of_countably_generated hfm.enorm -theorem AECover.integrable_of_lintegral_nnnorm_tendsto [l.NeBot] [l.IsCountablyGenerated] +@[deprecated (since := "2025-01-22")] +alias AECover.integrable_of_lintegral_nnnorm_bounded := + AECover.integrable_of_lintegral_enorm_bounded + +theorem AECover.integrable_of_lintegral_enorm_tendsto [l.NeBot] [l.IsCountablyGenerated] {φ : ι → Set α} (hφ : AECover μ l φ) {f : α → E} (I : ℝ) (hfm : AEStronglyMeasurable f μ) - (htendsto : Tendsto (fun i => ∫⁻ x in φ i, ‖f x‖₊ ∂μ) l (𝓝 <| ENNReal.ofReal I)) : + (htendsto : Tendsto (fun i => ∫⁻ x in φ i, ‖f x‖ₑ ∂μ) l (𝓝 <| .ofReal I)) : Integrable f μ := by - refine hφ.integrable_of_lintegral_nnnorm_bounded (max 1 (I + 1)) hfm ?_ + refine hφ.integrable_of_lintegral_enorm_bounded (max 1 (I + 1)) hfm ?_ refine htendsto.eventually (ge_mem_nhds ?_) refine (ENNReal.ofReal_lt_ofReal_iff (lt_max_of_lt_left zero_lt_one)).2 ?_ exact lt_max_of_lt_right (lt_add_one I) -theorem AECover.integrable_of_lintegral_nnnorm_bounded' [l.NeBot] [l.IsCountablyGenerated] +@[deprecated (since := "2025-01-22")] +alias AECover.integrable_of_lintegral_nnnorm_tendsto := + AECover.integrable_of_lintegral_enorm_tendsto + +theorem AECover.integrable_of_lintegral_enorm_bounded' [l.NeBot] [l.IsCountablyGenerated] {φ : ι → Set α} (hφ : AECover μ l φ) {f : α → E} (I : ℝ≥0) (hfm : AEStronglyMeasurable f μ) - (hbounded : ∀ᶠ i in l, (∫⁻ x in φ i, ‖f x‖₊ ∂μ) ≤ I) : Integrable f μ := - hφ.integrable_of_lintegral_nnnorm_bounded I hfm + (hbounded : ∀ᶠ i in l, ∫⁻ x in φ i, ‖f x‖ₑ ∂μ ≤ I) : Integrable f μ := + hφ.integrable_of_lintegral_enorm_bounded I hfm (by simpa only [ENNReal.ofReal_coe_nnreal] using hbounded) -theorem AECover.integrable_of_lintegral_nnnorm_tendsto' [l.NeBot] [l.IsCountablyGenerated] +@[deprecated (since := "2025-01-22")] +alias AECover.integrable_of_lintegral_nnnorm_bounded' := + AECover.integrable_of_lintegral_enorm_bounded' + +theorem AECover.integrable_of_lintegral_enorm_tendsto' [l.NeBot] [l.IsCountablyGenerated] {φ : ι → Set α} (hφ : AECover μ l φ) {f : α → E} (I : ℝ≥0) (hfm : AEStronglyMeasurable f μ) - (htendsto : Tendsto (fun i => ∫⁻ x in φ i, ‖f x‖₊ ∂μ) l (𝓝 I)) : Integrable f μ := - hφ.integrable_of_lintegral_nnnorm_tendsto I hfm + (htendsto : Tendsto (fun i => ∫⁻ x in φ i, ‖f x‖ₑ ∂μ) l (𝓝 I)) : Integrable f μ := + hφ.integrable_of_lintegral_enorm_tendsto I hfm (by simpa only [ENNReal.ofReal_coe_nnreal] using htendsto) +@[deprecated (since := "2025-01-22")] +alias AECover.integrable_of_lintegral_nnnorm_tendsto' := + AECover.integrable_of_lintegral_enorm_tendsto' + theorem AECover.integrable_of_integral_norm_bounded [l.NeBot] [l.IsCountablyGenerated] {φ : ι → Set α} (hφ : AECover μ l φ) {f : α → E} (I : ℝ) (hfi : ∀ i, IntegrableOn f (φ i) μ) (hbounded : ∀ᶠ i in l, (∫ x in φ i, ‖f x‖ ∂μ) ≤ I) : Integrable f μ := by have hfm : AEStronglyMeasurable f μ := hφ.aestronglyMeasurable fun i => (hfi i).aestronglyMeasurable - refine hφ.integrable_of_lintegral_nnnorm_bounded I hfm ?_ + refine hφ.integrable_of_lintegral_enorm_bounded I hfm ?_ conv at hbounded in integral _ _ => rw [integral_eq_lintegral_of_nonneg_ae (ae_of_all _ fun x => @norm_nonneg E _ (f x)) hfm.norm.restrict] conv at hbounded in ENNReal.ofReal _ => rw [← coe_nnnorm, ENNReal.ofReal_coe_nnreal] refine hbounded.mono fun i hi => ?_ - rw [← ENNReal.ofReal_toReal <| ne_top_of_lt <| hasFiniteIntegral_iff_nnnorm.mp (hfi i).2] + rw [← ENNReal.ofReal_toReal <| ne_top_of_lt <| hasFiniteIntegral_iff_enorm.mp (hfi i).2] apply ENNReal.ofReal_le_ofReal hi theorem AECover.integrable_of_integral_norm_tendsto [l.NeBot] [l.IsCountablyGenerated] @@ -958,23 +974,27 @@ theorem _root_.HasCompactSupport.integral_Iic_deriv_eq (hf : ContDiff ℝ 1 f) exact h2f.filter_mono _root_.atBot_le_cocompact |>.tendsto open UniformSpace in -lemma _root_.HasCompactSupport.ennnorm_le_lintegral_Ici_deriv +lemma _root_.HasCompactSupport.enorm_le_lintegral_Ici_deriv {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] {f : ℝ → F} (hf : ContDiff ℝ 1 f) (h'f : HasCompactSupport f) (x : ℝ) : - (‖f x‖₊ : ℝ≥0∞) ≤ ∫⁻ y in Iic x, ‖deriv f y‖₊ := by + ‖f x‖ₑ ≤ ∫⁻ y in Iic x, ‖deriv f y‖ₑ := by let I : F →L[ℝ] Completion F := Completion.toComplL let f' : ℝ → Completion F := I ∘ f have hf' : ContDiff ℝ 1 f' := hf.continuousLinearMap_comp I have h'f' : HasCompactSupport f' := h'f.comp_left rfl - have : (‖f' x‖₊ : ℝ≥0∞) ≤ ∫⁻ y in Iic x, ‖deriv f' y‖₊ := by + have : ‖f' x‖ₑ ≤ ∫⁻ y in Iic x, ‖deriv f' y‖ₑ := by rw [← HasCompactSupport.integral_Iic_deriv_eq hf' h'f' x] - exact ennnorm_integral_le_lintegral_ennnorm _ + exact enorm_integral_le_lintegral_enorm _ convert this with y - · simp [f', I, Completion.nnnorm_coe] + · simp [f', I, Completion.enorm_coe] · rw [fderiv_comp_deriv _ I.differentiableAt (hf.differentiable le_rfl _)] simp only [ContinuousLinearMap.fderiv] simp [I] +@[deprecated (since := "2025-01-22")] +alias _root_.HasCompactSupport.ennnorm_le_lintegral_Ici_deriv := + HasCompactSupport.enorm_le_lintegral_Ici_deriv + end IicFTC section UnivFTC diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index bfb8d81c0bd4a7..338687b657efa8 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -49,10 +49,9 @@ integral noncomputable section -open scoped Classical open MeasureTheory Set Filter Function -open scoped Classical Topology Filter ENNReal Interval NNReal +open scoped Topology Filter ENNReal Interval NNReal variable {ι 𝕜 E F A : Type*} [NormedAddCommGroup E] diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index b0cbd7bec0648f..171fca1f4b81a8 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -323,22 +323,31 @@ theorem setLIntegral_congr_fun {f g : α → ℝ≥0∞} {s : Set α} (hs : Meas @[deprecated (since := "2024-06-29")] alias set_lintegral_congr_fun := setLIntegral_congr_fun -theorem lintegral_ofReal_le_lintegral_nnnorm (f : α → ℝ) : - ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ ∫⁻ x, ‖f x‖₊ ∂μ := by - simp_rw [← ofReal_norm_eq_coe_nnnorm] +theorem lintegral_ofReal_le_lintegral_enorm (f : α → ℝ) : + ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ ∫⁻ x, ‖f x‖ₑ ∂μ := by + simp_rw [← ofReal_norm_eq_enorm] refine lintegral_mono fun x => ENNReal.ofReal_le_ofReal ?_ rw [Real.norm_eq_abs] exact le_abs_self (f x) -theorem lintegral_nnnorm_eq_of_ae_nonneg {f : α → ℝ} (h_nonneg : 0 ≤ᵐ[μ] f) : - ∫⁻ x, ‖f x‖₊ ∂μ = ∫⁻ x, ENNReal.ofReal (f x) ∂μ := by +@[deprecated (since := "2025-01-17")] +alias lintegral_ofReal_le_lintegral_nnnorm := lintegral_ofReal_le_lintegral_enorm + +theorem lintegral_enorm_of_ae_nonneg {f : α → ℝ} (h_nonneg : 0 ≤ᵐ[μ] f) : + ∫⁻ x, ‖f x‖ₑ ∂μ = ∫⁻ x, .ofReal (f x) ∂μ := by apply lintegral_congr_ae filter_upwards [h_nonneg] with x hx - rw [Real.nnnorm_of_nonneg hx, ENNReal.ofReal_eq_coe_nnreal hx] + rw [Real.enorm_eq_ofReal hx] + +@[deprecated (since := "2025-01-17")] +alias lintegral_nnnorm_eq_of_ae_nonneg := lintegral_enorm_of_ae_nonneg + +theorem lintegral_enorm_of_nonneg {f : α → ℝ} (h_nonneg : 0 ≤ f) : + ∫⁻ x, ‖f x‖ₑ ∂μ = ∫⁻ x, .ofReal (f x) ∂μ := + lintegral_enorm_of_ae_nonneg <| .of_forall h_nonneg -theorem lintegral_nnnorm_eq_of_nonneg {f : α → ℝ} (h_nonneg : 0 ≤ f) : - ∫⁻ x, ‖f x‖₊ ∂μ = ∫⁻ x, ENNReal.ofReal (f x) ∂μ := - lintegral_nnnorm_eq_of_ae_nonneg (Filter.Eventually.of_forall h_nonneg) +@[deprecated (since := "2025-01-17")] +alias lintegral_nnnorm_eq_of_nonneg := lintegral_enorm_of_nonneg /-- **Monotone convergence theorem** -- sometimes called **Beppo-Levi convergence**. See `lintegral_iSup_directed` for a more general form. -/ diff --git a/Mathlib/MeasureTheory/Integral/Prod.lean b/Mathlib/MeasureTheory/Integral/Prod.lean index 886d5b00bc9104..193e5df6d8fb36 100644 --- a/Mathlib/MeasureTheory/Integral/Prod.lean +++ b/Mathlib/MeasureTheory/Integral/Prod.lean @@ -35,7 +35,7 @@ product measure, Fubini's theorem, Fubini-Tonelli theorem noncomputable section -open scoped Classical Topology ENNReal MeasureTheory +open scoped Topology ENNReal MeasureTheory open Set Function Real ENNReal @@ -59,7 +59,7 @@ along one of the variables (using either the Lebesgue or Bochner integral) is me theorem measurableSet_integrable [SFinite ν] ⦃f : α → β → E⦄ (hf : StronglyMeasurable (uncurry f)) : MeasurableSet {x | Integrable (f x) ν} := by simp_rw [Integrable, hf.of_uncurry_left.aestronglyMeasurable, true_and] - exact measurableSet_lt (Measurable.lintegral_prod_right hf.ennnorm) measurable_const + exact measurableSet_lt (Measurable.lintegral_prod_right hf.enorm) measurable_const section @@ -70,6 +70,7 @@ variable [NormedSpace ℝ E] This version has `f` in curried form. -/ theorem MeasureTheory.StronglyMeasurable.integral_prod_right [SFinite ν] ⦃f : α → β → E⦄ (hf : StronglyMeasurable (uncurry f)) : StronglyMeasurable fun x => ∫ y, f x y ∂ν := by + classical by_cases hE : CompleteSpace E; swap; · simp [integral, hE, stronglyMeasurable_const] borelize E haveI : SeparableSpace (range (uncurry f) ∪ {0} : Set E) := @@ -148,7 +149,7 @@ variable [SFinite ν] theorem integrable_measure_prod_mk_left {s : Set (α × β)} (hs : MeasurableSet s) (h2s : (μ.prod ν) s ≠ ∞) : Integrable (fun x => (ν (Prod.mk x ⁻¹' s)).toReal) μ := by refine ⟨(measurable_measure_prod_mk_left hs).ennreal_toReal.aemeasurable.aestronglyMeasurable, ?_⟩ - simp_rw [hasFiniteIntegral_iff_nnnorm, ennnorm_eq_ofReal toReal_nonneg] + simp_rw [hasFiniteIntegral_iff_enorm, enorm_eq_ofReal toReal_nonneg] convert h2s.lt_top using 1 rw [prod_apply hs] apply lintegral_congr_ae @@ -216,11 +217,11 @@ theorem hasFiniteIntegral_prod_iff ⦃f : α × β → E⦄ (h1f : StronglyMeasu HasFiniteIntegral f (μ.prod ν) ↔ (∀ᵐ x ∂μ, HasFiniteIntegral (fun y => f (x, y)) ν) ∧ HasFiniteIntegral (fun x => ∫ y, ‖f (x, y)‖ ∂ν) μ := by - simp only [hasFiniteIntegral_iff_nnnorm, lintegral_prod_of_measurable _ h1f.ennnorm] + simp only [hasFiniteIntegral_iff_enorm, lintegral_prod_of_measurable _ h1f.enorm] have (x) : ∀ᵐ y ∂ν, 0 ≤ ‖f (x, y)‖ := by filter_upwards with y using norm_nonneg _ simp_rw [integral_eq_lintegral_of_nonneg_ae (this _) (h1f.norm.comp_measurable measurable_prod_mk_left).aestronglyMeasurable, - ennnorm_eq_ofReal toReal_nonneg, ofReal_norm_eq_coe_nnnorm] + enorm_eq_ofReal toReal_nonneg, ofReal_norm_eq_enorm] -- this fact is probably too specialized to be its own lemma have : ∀ {p q r : Prop} (_ : r → p), (r ↔ p ∧ q) ↔ p → (r ↔ q) := fun {p q r} h1 => by rw [← and_congr_right_iff, and_iff_right_of_imp h1] @@ -228,7 +229,7 @@ theorem hasFiniteIntegral_prod_iff ⦃f : α × β → E⦄ (h1f : StronglyMeasu · intro h2f; rw [lintegral_congr_ae] filter_upwards [h2f] with x hx rw [ofReal_toReal]; rw [← lt_top_iff_ne_top]; exact hx - · intro h2f; refine ae_lt_top ?_ h2f.ne; exact h1f.ennnorm.lintegral_prod_right' + · intro h2f; refine ae_lt_top ?_ h2f.ne; exact h1f.enorm.lintegral_prod_right' theorem hasFiniteIntegral_prod_iff' ⦃f : α × β → E⦄ (h1f : AEStronglyMeasurable f (μ.prod ν)) : HasFiniteIntegral f (μ.prod ν) ↔ @@ -385,17 +386,15 @@ theorem continuous_integral_integral : refine tendsto_integral_of_L1 _ (L1.integrable_coeFn g).integral_prod_left (Eventually.of_forall fun h => (L1.integrable_coeFn h).integral_prod_left) ?_ - simp_rw [← - lintegral_fn_integral_sub (fun x => (‖x‖₊ : ℝ≥0∞)) (L1.integrable_coeFn _) - (L1.integrable_coeFn g)] + simp_rw [← lintegral_fn_integral_sub _ (L1.integrable_coeFn _) (L1.integrable_coeFn g)] apply tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds _ (fun i => zero_le _) _ - · exact fun i => ∫⁻ x, ∫⁻ y, ‖i (x, y) - g (x, y)‖₊ ∂ν ∂μ - swap; · exact fun i => lintegral_mono fun x => ennnorm_integral_le_lintegral_ennnorm _ + · exact fun i => ∫⁻ x, ∫⁻ y, ‖i (x, y) - g (x, y)‖ₑ ∂ν ∂μ + swap; · exact fun i => lintegral_mono fun x => enorm_integral_le_lintegral_enorm _ show - Tendsto (fun i : α × β →₁[μ.prod ν] E => ∫⁻ x, ∫⁻ y : β, ‖i (x, y) - g (x, y)‖₊ ∂ν ∂μ) (𝓝 g) + Tendsto (fun i : α × β →₁[μ.prod ν] E => ∫⁻ x, ∫⁻ y : β, ‖i (x, y) - g (x, y)‖ₑ ∂ν ∂μ) (𝓝 g) (𝓝 0) - have : ∀ i : α × β →₁[μ.prod ν] E, Measurable fun z => (‖i z - g z‖₊ : ℝ≥0∞) := fun i => - ((Lp.stronglyMeasurable i).sub (Lp.stronglyMeasurable g)).ennnorm + have this (i : α × β →₁[μ.prod ν] E) : Measurable fun z => ‖i z - g z‖ₑ := + ((Lp.stronglyMeasurable i).sub (Lp.stronglyMeasurable g)).enorm -- Porting note: was -- simp_rw [← lintegral_prod_of_measurable _ (this _), ← L1.ofReal_norm_sub_eq_lintegral, ← -- ofReal_zero] diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index f97aa2469f9e89..05f685c4df4b5b 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -54,7 +54,7 @@ assert_not_exists InnerProductSpace noncomputable section open Filter Function MeasureTheory RCLike Set TopologicalSpace Topology -open scoped Classical ENNReal NNReal +open scoped ENNReal NNReal variable {X Y E F : Type*} @@ -141,6 +141,7 @@ theorem integral_finset_biUnion {ι : Type*} (t : Finset ι) {s : ι → Set X} (hs : ∀ i ∈ t, MeasurableSet (s i)) (h's : Set.Pairwise (↑t) (Disjoint on s)) (hf : ∀ i ∈ t, IntegrableOn f (s i) μ) : ∫ x in ⋃ i ∈ t, s i, f x ∂μ = ∑ i ∈ t, ∫ x in s i, f x ∂μ := by + classical induction' t using Finset.induction_on with a t hat IH hs h's · simp · simp only [Finset.coe_insert, Finset.forall_mem_insert, Set.pairwise_insert, @@ -211,8 +212,7 @@ theorem ofReal_setIntegral_one_of_measure_ne_top {X : Type*} {m : MeasurableSpac ENNReal.ofReal (∫ _ in s, (1 : ℝ) ∂μ) = ENNReal.ofReal (∫ _ in s, ‖(1 : ℝ)‖ ∂μ) := by simp only [norm_one] _ = ∫⁻ _ in s, 1 ∂μ := by - rw [ofReal_integral_norm_eq_lintegral_nnnorm (integrableOn_const.2 (Or.inr hs.lt_top))] - simp only [nnnorm_one, ENNReal.coe_one] + simpa [ofReal_integral_norm_eq_lintegral_enorm (integrableOn_const.2 (.inr hs.lt_top))] _ = μ s := setLIntegral_one _ @[deprecated (since := "2024-04-17")] @@ -244,7 +244,7 @@ theorem tendsto_setIntegral_of_monotone have hSm : MeasurableSet S := MeasurableSet.iUnion_of_monotone h_mono hsm have hsub {i} : s i ⊆ S := subset_iUnion s i rw [← withDensity_apply _ hSm] at hfi' - set ν := μ.withDensity fun x => ‖f x‖₊ with hν + set ν := μ.withDensity (‖f ·‖ₑ) with hν refine Metric.nhds_basis_closedBall.tendsto_right_iff.2 fun ε ε0 => ?_ lift ε to ℝ≥0 using ε0.le have : ∀ᶠ i in atTop, ν (s i) ∈ Icc (ν S - ε) (ν S + ε) := @@ -252,7 +252,7 @@ theorem tendsto_setIntegral_of_monotone filter_upwards [this] with i hi rw [mem_closedBall_iff_norm', ← integral_diff (hsm i) hfi hsub, ← coe_nnnorm, NNReal.coe_le_coe, ← ENNReal.coe_le_coe] - refine (ennnorm_integral_le_lintegral_ennnorm _).trans ?_ + refine (enorm_integral_le_lintegral_enorm _).trans ?_ rw [← withDensity_apply _ (hSm.diff (hsm _)), ← hν, measure_diff hsub (hsm _).nullMeasurableSet] exacts [tsub_le_iff_tsub_le.mp hi.1, (hi.2.trans_lt <| ENNReal.add_lt_top.2 ⟨hfi', ENNReal.coe_lt_top⟩).ne] @@ -661,7 +661,7 @@ theorem setIntegral_gt_gt {R : ℝ} {f : X → ℝ} (hR : 0 ≤ R) (μ {x | ↑R < f x}).toReal * R < ∫ x in {x | ↑R < f x}, f x ∂μ := by have : IntegrableOn (fun _ => R) {x | ↑R < f x} μ := by refine ⟨aestronglyMeasurable_const, lt_of_le_of_lt ?_ hfint.2⟩ - refine setLIntegral_mono_ae hfint.1.ennnorm <| ae_of_all _ fun x hx => ?_ + refine setLIntegral_mono_ae hfint.1.enorm <| ae_of_all _ fun x hx => ?_ simp only [ENNReal.coe_le_coe, Real.nnnorm_of_nonneg hR, enorm_eq_nnnorm, Real.nnnorm_of_nonneg (hR.trans <| le_of_lt hx), Subtype.mk_le_mk] exact le_of_lt hx @@ -1333,6 +1333,8 @@ end ContinuousMap theorem integral_ofReal {f : X → ℝ} : ∫ x, (f x : 𝕜) ∂μ = ↑(∫ x, f x ∂μ) := (@RCLike.ofRealLI 𝕜 _).integral_comp_comm f +theorem integral_complex_ofReal {f : X → ℝ} : ∫ x, (f x : ℂ) ∂μ = ∫ x, f x ∂μ := integral_ofReal + theorem integral_re {f : X → 𝕜} (hf : Integrable f μ) : ∫ x, RCLike.re (f x) ∂μ = RCLike.re (∫ x, f x ∂μ) := (@RCLike.reCLM 𝕜 _).integral_comp_comm hf @@ -1416,10 +1418,7 @@ theorem integral_withDensity_eq_integral_smul {f : X → ℝ≥0} (f_meas : Meas · rfl · exact integral_nonneg fun x => NNReal.coe_nonneg _ · refine ⟨f_meas.coe_nnreal_real.aemeasurable.aestronglyMeasurable, ?_⟩ - rw [withDensity_apply _ s_meas] at hs - rw [hasFiniteIntegral_iff_nnnorm] - convert hs with x - simp only [NNReal.nnnorm_eq] + simpa [withDensity_apply _ s_meas, hasFiniteIntegral_iff_enorm] using hs · intro u u' _ u_int u'_int h h' change (∫ x : X, u x + u' x ∂μ.withDensity fun x : X => ↑(f x)) = ∫ x : X, f x • (u x + u' x) ∂μ diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index 9c3a4587da7c6d..2fdacefa4379d0 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -607,13 +607,13 @@ namespace SimpleFunc theorem norm_eq_sum_mul (f : α →₁ₛ[μ] G) : ‖f‖ = ∑ x ∈ (toSimpleFunc f).range, (μ (toSimpleFunc f ⁻¹' {x})).toReal * ‖x‖ := by - rw [norm_toSimpleFunc, eLpNorm_one_eq_lintegral_nnnorm] - have h_eq := SimpleFunc.map_apply (fun x => (‖x‖₊ : ℝ≥0∞)) (toSimpleFunc f) + rw [norm_toSimpleFunc, eLpNorm_one_eq_lintegral_enorm] + have h_eq := SimpleFunc.map_apply (‖·‖ₑ) (toSimpleFunc f) simp_rw [← h_eq] rw [SimpleFunc.lintegral_eq_lintegral, SimpleFunc.map_lintegral, ENNReal.toReal_sum] · congr ext1 x - rw [ENNReal.toReal_mul, mul_comm, ← ofReal_norm_eq_coe_nnnorm, + rw [ENNReal.toReal_mul, mul_comm, ← ofReal_norm_eq_enorm, ENNReal.toReal_ofReal (norm_nonneg _)] · intro x _ by_cases hx0 : x = 0 @@ -1350,14 +1350,14 @@ theorem continuous_setToFun (hT : DominatedFinMeasAdditive μ T C) : /-- If `F i → f` in `L1`, then `setToFun μ T hT (F i) → setToFun μ T hT f`. -/ theorem tendsto_setToFun_of_L1 (hT : DominatedFinMeasAdditive μ T C) {ι} (f : α → E) (hfi : Integrable f μ) {fs : ι → α → E} {l : Filter ι} (hfsi : ∀ᶠ i in l, Integrable (fs i) μ) - (hfs : Tendsto (fun i => ∫⁻ x, ‖fs i x - f x‖₊ ∂μ) l (𝓝 0)) : + (hfs : Tendsto (fun i => ∫⁻ x, ‖fs i x - f x‖ₑ ∂μ) l (𝓝 0)) : Tendsto (fun i => setToFun μ T hT (fs i)) l (𝓝 <| setToFun μ T hT f) := by classical let f_lp := hfi.toL1 f let F_lp i := if hFi : Integrable (fs i) μ then hFi.toL1 (fs i) else 0 have tendsto_L1 : Tendsto F_lp l (𝓝 f_lp) := by rw [Lp.tendsto_Lp_iff_tendsto_ℒp'] - simp_rw [eLpNorm_one_eq_lintegral_nnnorm, Pi.sub_apply] + simp_rw [eLpNorm_one_eq_lintegral_enorm, Pi.sub_apply] refine (tendsto_congr' ?_).mp hfs filter_upwards [hfsi] with i hi refine lintegral_congr_ae ?_ @@ -1380,7 +1380,7 @@ theorem tendsto_setToFun_approxOn_of_measurable (hT : DominatedFinMeasAdditive (𝓝 <| setToFun μ T hT f) := tendsto_setToFun_of_L1 hT _ hfi (Eventually.of_forall (SimpleFunc.integrable_approxOn hfm hfi h₀ h₀i)) - (SimpleFunc.tendsto_approxOn_L1_nnnorm hfm _ hs (hfi.sub h₀i).2) + (SimpleFunc.tendsto_approxOn_L1_enorm hfm _ hs (hfi.sub h₀i).2) theorem tendsto_setToFun_approxOn_of_measurable_of_range_subset (hT : DominatedFinMeasAdditive μ T C) [MeasurableSpace E] [BorelSpace E] {f : α → E} @@ -1569,7 +1569,7 @@ theorem tendsto_setToFun_of_dominated_convergence (hT : DominatedFinMeasAdditive rw [← Integrable.toL1_sub] refine ((fs_int n).sub f_int).coeFn_toL1.mono fun x hx => ?_ dsimp only - rw [hx, ofReal_norm_eq_coe_nnnorm, Pi.sub_apply] + rw [hx, ofReal_norm_eq_enorm, Pi.sub_apply] /-- Lebesgue dominated convergence theorem for filters with a countable basis -/ theorem tendsto_setToFun_filter_of_dominated_convergence (hT : DominatedFinMeasAdditive μ T C) {ι} diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean index b7600589ddd7c8..14e6554bcb7bdf 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean @@ -504,9 +504,7 @@ def arrowProdEquivProdArrow (α β γ : Type*) [MeasurableSpace α] [MeasurableS __ := Equiv.arrowProdEquivProdArrow α β γ measurable_toFun _ h := by simp_rw [Equiv.arrowProdEquivProdArrow, coe_fn_mk] - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/6024 + #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 we need provide the type hints `(a : γ → α × β)`, to avoid unification issues. -/ exact MeasurableSet.preimage h (Measurable.prod_mk diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Pi.lean b/Mathlib/MeasureTheory/MeasurableSpace/Pi.lean index 82e1418f467fcf..fc311ca2b168c0 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Pi.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Pi.lean @@ -20,7 +20,6 @@ generate the σ-algebra on the indexed product of `α i`s. noncomputable section open Function Set Filter MeasurableSpace Encodable -open scoped Classical variable {ι : Type*} {α : ι → Type*} @@ -63,6 +62,7 @@ theorem IsCountablySpanning.pi {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsCoun theorem generateFrom_pi_eq {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsCountablySpanning (C i)) : (@MeasurableSpace.pi _ _ fun i => generateFrom (C i)) = generateFrom (pi univ '' pi univ C) := by + classical cases nonempty_encodable ι apply le_antisymm · refine iSup_le ?_; intro i; rw [comap_generateFrom] diff --git a/Mathlib/MeasureTheory/Measure/AddContent.lean b/Mathlib/MeasureTheory/Measure/AddContent.lean index 2332cc948c41d4..3475421c5788ae 100644 --- a/Mathlib/MeasureTheory/Measure/AddContent.lean +++ b/Mathlib/MeasureTheory/Measure/AddContent.lean @@ -109,27 +109,27 @@ lemma addContent_union' (hs : s ∈ C) (ht : t ∈ C) (hst : s ∪ t ∈ C) (h_d section IsSetSemiring -lemma addContent_eq_add_diffFinset₀_of_subset (hC : IsSetSemiring C) +lemma addContent_eq_add_disjointOfDiffUnion_of_subset (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) (hI_ss : ∀ t ∈ I, t ⊆ s) (h_dis : PairwiseDisjoint (I : Set (Set α)) id) : - m s = ∑ i ∈ I, m i + ∑ i ∈ hC.diffFinset₀ hs hI, m i := by + m s = ∑ i ∈ I, m i + ∑ i ∈ hC.disjointOfDiffUnion hs hI, m i := by classical - conv_lhs => rw [← hC.sUnion_union_diffFinset₀_of_subset hs hI hI_ss] + conv_lhs => rw [← hC.sUnion_union_disjointOfDiffUnion_of_subset hs hI hI_ss] rw [addContent_sUnion] · rw [sum_union] - exact hC.disjoint_diffFinset₀ hs hI + exact hC.disjoint_disjointOfDiffUnion hs hI · rw [coe_union] - exact Set.union_subset hI (hC.diffFinset₀_subset hs hI) + exact Set.union_subset hI (hC.disjointOfDiffUnion_subset hs hI) · rw [coe_union] - exact hC.pairwiseDisjoint_union_diffFinset₀ hs hI h_dis - · rwa [hC.sUnion_union_diffFinset₀_of_subset hs hI hI_ss] + exact hC.pairwiseDisjoint_union_disjointOfDiffUnion hs hI h_dis + · rwa [hC.sUnion_union_disjointOfDiffUnion_of_subset hs hI hI_ss] lemma sum_addContent_le_of_subset (hC : IsSetSemiring C) (h_ss : ↑I ⊆ C) (h_dis : PairwiseDisjoint (I : Set (Set α)) id) (ht : t ∈ C) (hJt : ∀ s ∈ I, s ⊆ t) : ∑ u ∈ I, m u ≤ m t := by classical - rw [addContent_eq_add_diffFinset₀_of_subset hC ht h_ss hJt h_dis] + rw [addContent_eq_add_disjointOfDiffUnion_of_subset hC ht h_ss hJt h_dis] exact le_add_right le_rfl lemma addContent_mono (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) diff --git a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean index a782f7e429aeba..7915e996c9a658 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean @@ -393,7 +393,7 @@ lemma QuotientGroup.integral_eq_integral_automorphize {E : Type*} [NormedAddComm rw [integral_tsum] · exact fun i ↦ (hf₁.1.comp_quasiMeasurePreserving (measurePreserving_smul i μ).quasiMeasurePreserving).restrict - · rw [← h𝓕.lintegral_eq_tsum'' (fun x ↦ ‖f x‖₊)] + · rw [← h𝓕.lintegral_eq_tsum'' (‖f ·‖ₑ)] exact ne_of_lt hf₁.2 /-- This is the **Unfolding Trick**: Given a subgroup `Γ` of a group `G`, the integral of a @@ -403,7 +403,7 @@ lemma QuotientGroup.integral_eq_integral_automorphize {E : Type*} [NormedAddComm lemma QuotientGroup.integral_mul_eq_integral_automorphize_mul {K : Type*} [NormedField K] [NormedSpace ℝ K] [μ.IsMulRightInvariant] {f : G → K} (f_ℒ_1 : Integrable f μ) {g : G ⧸ Γ → K} (hg : AEStronglyMeasurable g μ_𝓕) - (g_ℒ_infinity : essSup (fun x ↦ ↑‖g x‖₊) μ_𝓕 ≠ ∞) + (g_ℒ_infinity : essSup (fun x ↦ ↑‖g x‖ₑ) μ_𝓕 ≠ ∞) (F_ae_measurable : AEStronglyMeasurable (QuotientGroup.automorphize f) μ_𝓕) : ∫ x : G, g (x : G ⧸ Γ) * (f x) ∂μ = ∫ x : G ⧸ Γ, g x * (QuotientGroup.automorphize f x) ∂μ_𝓕 := by @@ -418,8 +418,7 @@ lemma QuotientGroup.integral_mul_eq_integral_automorphize_mul {K : Type*} [Norme have : AEStronglyMeasurable (fun (x : G) ↦ g (x : (G ⧸ Γ))) μ := (hg.mono_ac h𝓕.absolutelyContinuous_map).comp_measurable meas_π refine Integrable.essSup_smul f_ℒ_1 this ?_ - have hg' : AEStronglyMeasurable (fun x ↦ (‖g x‖₊ : ℝ≥0∞)) μ_𝓕 := - (ENNReal.continuous_coe.comp continuous_nnnorm).comp_aestronglyMeasurable hg + have hg' : AEStronglyMeasurable (‖g ·‖ₑ) μ_𝓕 := continuous_enorm.comp_aestronglyMeasurable hg rw [← essSup_comp_quotientGroup_mk h𝓕 hg'.aemeasurable] exact g_ℒ_infinity have H₂ : AEStronglyMeasurable (QuotientGroup.automorphize ((g ∘ π) * f)) μ_𝓕 := by @@ -447,7 +446,7 @@ local notation "μ_𝓕" => Measure.map (@QuotientAddGroup.mk G' _ Γ') (μ'.res lemma QuotientAddGroup.integral_mul_eq_integral_automorphize_mul {K : Type*} [NormedField K] [NormedSpace ℝ K] [μ'.IsAddRightInvariant] {f : G' → K} (f_ℒ_1 : Integrable f μ') {g : G' ⧸ Γ' → K} (hg : AEStronglyMeasurable g μ_𝓕) - (g_ℒ_infinity : essSup (fun x ↦ (‖g x‖₊ : ℝ≥0∞)) μ_𝓕 ≠ ∞) + (g_ℒ_infinity : essSup (‖g ·‖ₑ) μ_𝓕 ≠ ∞) (F_ae_measurable : AEStronglyMeasurable (QuotientAddGroup.automorphize f) μ_𝓕) (h𝓕 : IsAddFundamentalDomain Γ'.op 𝓕' μ') : ∫ x : G', g (x : G' ⧸ Γ') * (f x) ∂μ' @@ -463,8 +462,7 @@ lemma QuotientAddGroup.integral_mul_eq_integral_automorphize_mul {K : Type*} [No have : AEStronglyMeasurable (fun (x : G') ↦ g (x : (G' ⧸ Γ'))) μ' := (hg.mono_ac h𝓕.absolutelyContinuous_map).comp_measurable meas_π refine Integrable.essSup_smul f_ℒ_1 this ?_ - have hg' : AEStronglyMeasurable (fun x ↦ (‖g x‖₊ : ℝ≥0∞)) μ_𝓕 := - (ENNReal.continuous_coe.comp continuous_nnnorm).comp_aestronglyMeasurable hg + have hg' : AEStronglyMeasurable (‖g ·‖ₑ) μ_𝓕 := continuous_enorm.comp_aestronglyMeasurable hg rw [← essSup_comp_quotientAddGroup_mk h𝓕 hg'.aemeasurable] exact g_ℒ_infinity have H₂ : AEStronglyMeasurable (QuotientAddGroup.automorphize ((g ∘ π) * f)) μ_𝓕 := by diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index 2f8b288b9314bd..5bd12b8bc56832 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -501,7 +501,7 @@ theorem mkMetric_le_liminf_tsum {β : Type*} {ι : β → Type*} [∀ n, Countab haveI : ∀ n, Encodable (ι n) := fun n => Encodable.ofCountable _ simp only [mkMetric_apply] refine iSup₂_le fun ε hε => ?_ - refine le_of_forall_le_of_dense fun c hc => ?_ + refine le_of_forall_gt_imp_ge_of_dense fun c hc => ?_ rcases ((frequently_lt_of_liminf_lt (by isBoundedDefault) hc).and_eventually ((hr.eventually (gt_mem_nhds hε)).and (ht.and hst))).exists with ⟨n, hn, hrn, htn, hstn⟩ diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index aa7a1f50934429..2027a22e7dfebb 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -563,7 +563,7 @@ theorem _root_.AlternatingMap.measure_parallelepiped (ω : G [⋀^Fin n]→ₗ[ conv_rhs => rw [ω.eq_smul_basis_det (finBasisOfFinrankEq ℝ G _i.out)] simp only [addHaar_parallelepiped, AlternatingMap.measure, coe_nnreal_smul_apply, AlternatingMap.smul_apply, Algebra.id.smul_eq_mul, abs_mul, ENNReal.ofReal_mul (abs_nonneg _), - Real.ennnorm_eq_ofReal_abs] + ← Real.enorm_eq_ofReal_abs, enorm] instance (ω : G [⋀^Fin n]→ₗ[ℝ] ℝ) : IsAddLeftInvariant ω.measure := by rw [AlternatingMap.measure]; infer_instance diff --git a/Mathlib/MeasureTheory/Measure/Prod.lean b/Mathlib/MeasureTheory/Measure/Prod.lean index 6248de656ceb81..e20bbe27b66ac1 100644 --- a/Mathlib/MeasureTheory/Measure/Prod.lean +++ b/Mathlib/MeasureTheory/Measure/Prod.lean @@ -851,6 +851,13 @@ instance fst.instIsProbabilityMeasure [IsProbabilityMeasure ρ] : IsProbabilityM rw [fst_univ] exact measure_univ +instance fst.instIsZeroOrProbabilityMeasure [IsZeroOrProbabilityMeasure ρ] : + IsZeroOrProbabilityMeasure ρ.fst := by + rcases eq_zero_or_isProbabilityMeasure ρ with h | h + · simp only [h, fst_zero] + infer_instance + · infer_instance + @[simp] lemma fst_prod [IsProbabilityMeasure ν] : (μ.prod ν).fst = μ := by ext1 s hs @@ -909,6 +916,13 @@ instance snd.instIsProbabilityMeasure [IsProbabilityMeasure ρ] : IsProbabilityM rw [snd_univ] exact measure_univ +instance snd.instIsZeroOrProbabilityMeasure [IsZeroOrProbabilityMeasure ρ] : + IsZeroOrProbabilityMeasure ρ.snd := by + rcases eq_zero_or_isProbabilityMeasure ρ with h | h + · simp only [h, snd_zero] + infer_instance + · infer_instance + @[simp] lemma snd_prod [IsProbabilityMeasure μ] : (μ.prod ν).snd = ν := by ext1 s hs diff --git a/Mathlib/MeasureTheory/SetSemiring.lean b/Mathlib/MeasureTheory/SetSemiring.lean index 503f9ca6b3e8d5..976dcb2c87dc3e 100644 --- a/Mathlib/MeasureTheory/SetSemiring.lean +++ b/Mathlib/MeasureTheory/SetSemiring.lean @@ -6,6 +6,7 @@ Authors: Rémy Degenne, Peter Pfaffelhuber import Mathlib.Data.Nat.Lattice import Mathlib.Data.Set.Accumulate import Mathlib.Data.Set.Pairwise.Lattice +import Mathlib.Order.CompleteLattice import Mathlib.MeasureTheory.PiSystem /-! # Semirings and rings of sets @@ -24,19 +25,27 @@ A ring of sets is a set of sets containing `∅`, stable by union, set differenc ## 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` +* `MeasureTheory.IsSetSemiring.disjointOfDiff 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`. + pairwise disjoint sets such that `s \ t = ⋃₀ hC.disjointOfDiff hs ht`. +* `MeasureTheory.IsSetSemiring.disjointOfDiffUnion 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.disjointOfDiffUnion hs hI`. +* `MeasureTheory.IsSetSemiring.disjointOfUnion hJ`: for `hJ ⊆ C`, this is a + `Finset` of pairwise disjoint sets such that `⋃₀ J = ⋃₀ hC.disjointOfUnion hJ`. * `MeasureTheory.IsSetRing`: property of being a ring of sets. ## Main statements * `MeasureTheory.IsSetSemiring.exists_disjoint_finset_diff_eq`: the existence of the `Finset` given - by the definition `IsSetSemiring.diffFinset₀` (see above). + by the definition `IsSetSemiring.disjointOfDiffUnion` (see above). +* `MeasureTheory.IsSetSemiring.disjointOfUnion_props`: In a `hC : IsSetSemiring C`, + for a `J : Finset (Set α)` with `J ⊆ C`, there is + for every `x in J` some `K x ⊆ C` finite, such that + * `⋃ x ∈ J, K x` are pairwise disjoint and do not contain ∅, + * `⋃ s ∈ K x, s ⊆ x`, + * `⋃ x ∈ J, x = ⋃ x ∈ J, ⋃ s ∈ K x, s`. -/ @@ -58,84 +67,86 @@ namespace IsSetSemiring lemma isPiSystem (hC : IsSetSemiring C) : IsPiSystem C := fun s hs t ht _ ↦ hC.inter_mem s hs t ht -section diffFinset +section disjointOfDiff open scoped 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) : +We remove the empty set to ensure that `t ∉ hC.disjointOfDiff hs ht` even if `t = ∅`. -/ +noncomputable def disjointOfDiff (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : Finset (Set α) := (hC.diff_eq_sUnion' s hs t ht).choose \ {∅} -lemma empty_not_mem_diffFinset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : - ∅ ∉ hC.diffFinset hs ht := by +lemma empty_nmem_disjointOfDiff (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + ∅ ∉ hC.disjointOfDiff hs ht := by classical - simp only [diffFinset, mem_sdiff, Finset.mem_singleton, eq_self_iff_true, not_true, - and_false, not_false_iff] + simp only [disjointOfDiff, mem_sdiff, Finset.mem_singleton, eq_self_iff_true, + not_true, and_false, not_false_iff] -lemma diffFinset_subset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : - ↑(hC.diffFinset hs ht) ⊆ C := by +lemma subset_disjointOfDiff (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + ↑(hC.disjointOfDiff hs ht) ⊆ C := by classical - simp only [diffFinset, coe_sdiff, coe_singleton, diff_singleton_subset_iff] + simp only [disjointOfDiff, coe_sdiff, coe_singleton, diff_singleton_subset_iff] exact (hC.diff_eq_sUnion' 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 +lemma pairwiseDisjoint_disjointOfDiff (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + PairwiseDisjoint (hC.disjointOfDiff hs ht : Set (Set α)) id := by classical - simp only [diffFinset, coe_sdiff, coe_singleton] + simp only [disjointOfDiff, coe_sdiff, coe_singleton] exact Set.PairwiseDisjoint.subset (hC.diff_eq_sUnion' s hs t ht).choose_spec.2.1 diff_subset -lemma sUnion_diffFinset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : - ⋃₀ hC.diffFinset hs ht = s \ t := by +lemma sUnion_disjointOfDiff (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + ⋃₀ hC.disjointOfDiff hs ht = s \ t := by classical rw [(hC.diff_eq_sUnion' s hs t ht).choose_spec.2.2] - simp only [diffFinset, coe_sdiff, coe_singleton, diff_singleton_subset_iff] + simp only [disjointOfDiff, 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 +lemma nmem_disjointOfDiff (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + t ∉ hC.disjointOfDiff 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 ?_ + refine hC.empty_nmem_disjointOfDiff hs ht ?_ rwa [← h] - rw [← hC.sUnion_diffFinset hs ht] + rw [← hC.sUnion_disjointOfDiff 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] +lemma sUnion_insert_disjointOfDiff (hC : IsSetSemiring C) (hs : s ∈ C) + (ht : t ∈ C) (hst : t ⊆ s) : + ⋃₀ insert t (hC.disjointOfDiff hs ht) = s := by + conv_rhs => rw [← union_diff_cancel hst, ← hC.sUnion_disjointOfDiff 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] +lemma disjoint_sUnion_disjointOfDiff (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : + Disjoint t (⋃₀ hC.disjointOfDiff hs ht) := by + rw [hC.sUnion_disjointOfDiff] 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 ↦ ?_ +lemma pairwiseDisjoint_insert_disjointOfDiff (hC : IsSetSemiring C) (hs : s ∈ C) + (ht : t ∈ C) : + PairwiseDisjoint (insert t (hC.disjointOfDiff hs ht) : Set (Set α)) id := by + have h := hC.pairwiseDisjoint_disjointOfDiff hs ht + refine PairwiseDisjoint.insert_of_not_mem h (hC.nmem_disjointOfDiff hs ht) fun u hu ↦ ?_ simp_rw [id] - refine Disjoint.mono_right ?_ (hC.disjoint_sUnion_diffFinset hs ht) + refine Disjoint.mono_right ?_ (hC.disjoint_sUnion_disjointOfDiff hs ht) simp only [Set.le_eq_subset] exact subset_sUnion_of_mem hu -end diffFinset +end disjointOfDiff -section diffFinset₀ +section disjointOfDiffUnion 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. -/ +See `IsSetSemiring.disjointOfDiffUnion` 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 α), ↑J ⊆ C ∧ PairwiseDisjoint (J : Set (Set α)) id ∧ s \ ⋃₀ I = ⋃₀ J := by @@ -151,14 +162,14 @@ lemma exists_disjoint_finset_diff_eq (hC : IsSetSemiring C) (hs : s ∈ C) (hI : rw [coe_insert] at hI have ht : t ∈ C := hI (Set.mem_insert _ _) obtain ⟨J, h_ss, h_dis, h_eq⟩ := h ((Set.subset_insert _ _).trans hI) - let Ju : ∀ u ∈ C, Finset (Set α) := fun u hu ↦ hC.diffFinset hu ht + let Ju : ∀ u ∈ C, Finset (Set α) := fun u hu ↦ hC.disjointOfDiff 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 + exact hC.subset_disjointOfDiff 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 + hC.pairwiseDisjoint_disjointOfDiff hu ht have hJu_sUnion : ∀ (u) (hu : u ∈ C), ⋃₀ (Ju u hu : Set (Set α)) = u \ t := - fun u hu ↦ hC.sUnion_diffFinset hu ht + fun u hu ↦ hC.sUnion_disjointOfDiff 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 @@ -202,80 +213,228 @@ lemma exists_disjoint_finset_diff_eq (hC : IsSetSemiring C) (hs : s ∈ C) (hI : open scoped 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 α) := +`disjointOfDiffUnion` is a finite set of sets in `C` such that +`s \ ⋃₀ I = ⋃₀ (hC.disjointOfDiffUnion hs I hI)`. +`disjointOfDiff` is a special case of `disjointOfDiffUnion` where `I` is a +singleton. -/ +noncomputable def disjointOfDiffUnion (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 +lemma empty_nmem_disjointOfDiffUnion (hC : IsSetSemiring C) (hs : s ∈ C) + (hI : ↑I ⊆ C) : + ∅ ∉ hC.disjointOfDiffUnion hs hI := by classical - simp only [diffFinset₀, mem_sdiff, Finset.mem_singleton, eq_self_iff_true, not_true, - and_false, not_false_iff] + simp only [disjointOfDiffUnion, mem_sdiff, Finset.mem_singleton, eq_self_iff_true, + not_true, and_false, not_false_iff] -lemma diffFinset₀_subset (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : - ↑(hC.diffFinset₀ hs hI) ⊆ C := by +lemma disjointOfDiffUnion_subset (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : + ↑(hC.disjointOfDiffUnion hs hI) ⊆ C := by classical - simp only [diffFinset₀, coe_sdiff, coe_singleton, diff_singleton_subset_iff] + simp only [disjointOfDiffUnion, coe_sdiff, coe_singleton, diff_singleton_subset_iff] exact (hC.exists_disjoint_finset_diff_eq hs hI).choose_spec.1.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 +lemma pairwiseDisjoint_disjointOfDiffUnion (hC : IsSetSemiring C) (hs : s ∈ C) + (hI : ↑I ⊆ C) : PairwiseDisjoint (hC.disjointOfDiffUnion hs hI : Set (Set α)) id := by classical - simp only [diffFinset₀, coe_sdiff, coe_singleton] + simp only [disjointOfDiffUnion, coe_sdiff, coe_singleton] exact Set.PairwiseDisjoint.subset (hC.exists_disjoint_finset_diff_eq hs hI).choose_spec.2.1 diff_subset -lemma diff_sUnion_eq_sUnion_diffFinset₀ (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : - s \ ⋃₀ I = ⋃₀ hC.diffFinset₀ hs hI := by +lemma diff_sUnion_eq_sUnion_disjointOfDiffUnion (hC : IsSetSemiring C) (hs : s ∈ C) + (hI : ↑I ⊆ C) : s \ ⋃₀ I = ⋃₀ hC.disjointOfDiffUnion hs hI := by classical rw [(hC.exists_disjoint_finset_diff_eq hs hI).choose_spec.2.2] - simp only [diffFinset₀, coe_sdiff, coe_singleton, diff_singleton_subset_iff] + simp only [disjointOfDiffUnion, 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₀] +lemma sUnion_disjointOfDiffUnion_subset (hC : IsSetSemiring C) (hs : s ∈ C) + (hI : ↑I ⊆ C) : ⋃₀ (hC.disjointOfDiffUnion hs hI : Set (Set α)) ⊆ s := by + rw [← hC.diff_sUnion_eq_sUnion_disjointOfDiffUnion] 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 +lemma subset_of_diffUnion_disjointOfDiffUnion (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) + (t : Set α) (ht : t ∈ (hC.disjointOfDiffUnion hs hI : Set (Set α))) : + t ⊆ s \ ⋃₀ I := by + revert t ht + rw [← sUnion_subset_iff, hC.diff_sUnion_eq_sUnion_disjointOfDiffUnion hs hI] + +lemma subset_of_mem_disjointOfDiffUnion (hC : IsSetSemiring C) {I : Finset (Set α)} + (hs : s ∈ C) (hI : ↑I ⊆ C) (t : Set α) + (ht : t ∈ (hC.disjointOfDiffUnion hs hI : Set (Set α))) : + t ⊆ s := by + apply le_trans <| hC.subset_of_diffUnion_disjointOfDiffUnion hs hI t ht + exact sdiff_le (a := s) (b := ⋃₀ I) + +lemma disjoint_sUnion_disjointOfDiffUnion (hC : IsSetSemiring C) (hs : s ∈ C) + (hI : ↑I ⊆ C) : + Disjoint (⋃₀ (I : Set (Set α))) (⋃₀ hC.disjointOfDiffUnion hs hI) := by + rw [← hC.diff_sUnion_eq_sUnion_disjointOfDiffUnion]; exact Set.disjoint_sdiff_right + +lemma disjoint_disjointOfDiffUnion (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : + Disjoint I (hC.disjointOfDiffUnion 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₀) + obtain ⟨u, huI, hu_disjointOfDiffUnion⟩ := h + have h_disj : u ≤ ⊥ := + hC.disjoint_sUnion_disjointOfDiffUnion hs hI (subset_sUnion_of_mem huI) + (subset_sUnion_of_mem hu_disjointOfDiffUnion) 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₀ + refine hC.empty_nmem_disjointOfDiffUnion hs hI ?_ + rwa [h_disj] at hu_disjointOfDiffUnion -lemma pairwiseDisjoint_union_diffFinset₀ (hC : IsSetSemiring C) (hs : s ∈ C) +lemma pairwiseDisjoint_union_disjointOfDiffUnion (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 + PairwiseDisjoint (I ∪ hC.disjointOfDiffUnion hs hI : Set (Set α)) id := by rw [pairwiseDisjoint_union] - refine ⟨h_dis, hC.pairwiseDisjoint_diffFinset₀ hs hI, fun u hu v hv _ ↦ ?_⟩ + refine ⟨h_dis, hC.pairwiseDisjoint_disjointOfDiffUnion hs hI, fun u hu v hv _ ↦ ?_⟩ simp_rw [id] exact disjoint_of_subset (subset_sUnion_of_mem hu) (subset_sUnion_of_mem hv) - (hC.disjoint_sUnion_diffFinset₀ hs hI) + (hC.disjoint_sUnion_disjointOfDiffUnion hs hI) -lemma sUnion_union_sUnion_diffFinset₀_of_subset (hC : IsSetSemiring C) (hs : s ∈ C) - (hI : ↑I ⊆ C) (hI_ss : ∀ t ∈ I, t ⊆ s) : - ⋃₀ I ∪ ⋃₀ hC.diffFinset₀ hs hI = s := by +lemma sUnion_union_sUnion_disjointOfDiffUnion_of_subset (hC : IsSetSemiring C) + (hs : s ∈ C) (hI : ↑I ⊆ C) (hI_ss : ∀ t ∈ I, t ⊆ s) : + ⋃₀ I ∪ ⋃₀ hC.disjointOfDiffUnion hs hI = s := by conv_rhs => rw [← union_diff_cancel (Set.sUnion_subset hI_ss : ⋃₀ ↑I ⊆ s), - hC.diff_sUnion_eq_sUnion_diffFinset₀ hs hI] + hC.diff_sUnion_eq_sUnion_disjointOfDiffUnion hs hI] -lemma sUnion_union_diffFinset₀_of_subset (hC : IsSetSemiring C) (hs : s ∈ C) +lemma sUnion_union_disjointOfDiffUnion_of_subset (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) (hI_ss : ∀ t ∈ I, t ⊆ s) [DecidableEq (Set α)] : - ⋃₀ ↑(I ∪ hC.diffFinset₀ hs hI) = s := by - conv_rhs => rw [← sUnion_union_sUnion_diffFinset₀_of_subset hC hs hI hI_ss] + ⋃₀ ↑(I ∪ hC.disjointOfDiffUnion hs hI) = s := by + conv_rhs => rw [← sUnion_union_sUnion_disjointOfDiffUnion_of_subset hC hs hI hI_ss] simp_rw [coe_union] rw [sUnion_union] -end diffFinset₀ +end disjointOfDiffUnion + +section disjointOfUnion + + +variable {j : Set α} {J : Finset (Set α)} + +open Set MeasureTheory Order + +theorem disjointOfUnion_props (hC : IsSetSemiring C) (h1 : ↑J ⊆ C) : + ∃ K : Set α → Finset (Set α), + PairwiseDisjoint J K + ∧ (∀ i ∈ J, ↑(K i) ⊆ C) + ∧ PairwiseDisjoint (⋃ x ∈ J, (K x : Set (Set α))) id + ∧ (∀ j ∈ J, ⋃₀ K j ⊆ j) + ∧ (∀ j ∈ J, ∅ ∉ K j) + ∧ ⋃₀ J = ⋃₀ (⋃ x ∈ J, (K x : Set (Set α))) := by + classical + induction J using Finset.cons_induction with + | empty => simp + | cons s J hJ hind => + rw [cons_eq_insert, coe_insert, Set.insert_subset_iff] at h1 + obtain ⟨K, hK0, ⟨hK1, hK2, hK3, hK4, hK5⟩⟩ := hind h1.2 + let K1 : Set α → Finset (Set α) := fun (t : Set α) ↦ + if t = s then (hC.disjointOfDiffUnion h1.1 h1.2) else K t + have hK1s : K1 s = hC.disjointOfDiffUnion h1.1 h1.2 := by simp [K1] + have hK1_of_ne t (ht : t ≠ s) : K1 t = K t := by simp [K1, ht] + use K1 + simp only [cons_eq_insert, disjiUnion_eq_biUnion, Finset.biUnion_insert, coe_union, coe_biUnion, + mem_coe, Set.union_subset_iff, iUnion_subset_iff, Finset.mem_insert, sUnion_subset_iff, + forall_eq_or_imp, coe_insert, sUnion_insert, exists_and_left, exists_prop] + -- two simplification rules for induction hypothesis + have ht1' : ∀ x ∈ J, K1 x = K x := fun x hx ↦ hK1_of_ne _ (fun h_eq ↦ hJ (h_eq ▸ hx)) + have ht2 : (⋃ x ∈ J, (K1 x : Set (Set α))) = ⋃ x ∈ J, ((K x : Set (Set α))) := by + apply iUnion₂_congr + intros x hx + exact_mod_cast hK1_of_ne _ (ne_of_mem_of_not_mem hx hJ) + simp only [hK1s] + refine ⟨?_, ⟨hC.disjointOfDiffUnion_subset h1.1 h1.2, ?_⟩, ?_, + ⟨hC.subset_of_mem_disjointOfDiffUnion h1.1 h1.2, ?_⟩, ?_, ?_⟩ + · apply Set.Pairwise.insert + · intro j hj i hi hij + rw [Function.onFun, ht1' j hj, ht1' i hi] + exact hK0 hj hi hij + · intro i hi _ + have h7 : Disjoint ↑(hC.disjointOfDiffUnion h1.1 h1.2) (K i : Set (Set α)) := by + refine disjoint_of_sSup_disjoint_of_le_of_le + (hC.subset_of_diffUnion_disjointOfDiffUnion h1.1 h1.2) ?_ + (@disjoint_sdiff_left _ (⋃₀ J) s) (Or.inl + (hC.empty_nmem_disjointOfDiffUnion h1.1 h1.2)) + simp only [mem_coe, Set.le_eq_subset] + apply sUnion_subset_iff.mp + exact (hK3 i hi).trans (subset_sUnion_of_mem hi) + have h8 : Function.onFun Disjoint K1 s i := by + refine Finset.disjoint_iff_inter_eq_empty.mpr ?_ + rw [ht1' i hi, hK1s] + rw [Set.disjoint_iff_inter_eq_empty] at h7 + exact_mod_cast h7 + exact ⟨h8, Disjoint.symm h8⟩ + · intros i hi + rw [ht1' i hi] + exact hK1 i hi + · simp only [iUnion_iUnion_eq_or_left] + refine pairwiseDisjoint_union.mpr ⟨?_, ?_, ?_⟩ + · rw [hK1s] + exact hC.pairwiseDisjoint_disjointOfDiffUnion h1.1 h1.2 + · simpa [ht2] + · simp only [mem_coe, mem_iUnion, exists_prop, ne_eq, id_eq, forall_exists_index, and_imp] + intros i hi j x hx h3 h4 + obtain ki : i ⊆ s \ ⋃₀ J := hC.subset_of_diffUnion_disjointOfDiffUnion h1.1 h1.2 _ + (hK1s ▸ hi) + obtain hx2 : j ⊆ x := subset_trans (subset_sUnion_of_mem (ht1' x hx ▸ h3)) (hK3 x hx) + obtain kj : j ⊆ ⋃₀ J := hx2.trans <| subset_sUnion_of_mem hx + exact disjoint_of_subset ki kj disjoint_sdiff_left + · intros a ha + simp_rw [hK1_of_ne _ (ne_of_mem_of_not_mem ha hJ)] + change ∀ t' ∈ (K a : Set (Set α)), t' ⊆ a + rw [← sUnion_subset_iff] + exact hK3 a ha + · refine ⟨hC.empty_nmem_disjointOfDiffUnion h1.1 h1.2, ?_⟩ + intros a ha + rw [ht1' a ha] + exact hK4 a ha + · simp only [iUnion_iUnion_eq_or_left, ht2, sUnion_union, ht2, K1] + simp_rw [apply_ite, hK5, + ← hC.diff_sUnion_eq_sUnion_disjointOfDiffUnion h1.1 h1.2, hK5] + simp only [↓reduceIte, diff_union_self] + +/-- For some `hJ : J ⊆ C` and `j : Set α`, where `hC : IsSetSemiring C`, this is +a `Finset (Set α)` such that `K j := hC.disjointOfUnion hJ` are disjoint +and `⋃₀ K j ⊆ j`, for `j ∈ J`. +Using these we write `⋃₀ J` as a disjoint union `⋃₀ J = ⋃₀ ⋃ x ∈ J, (K x)`. +See `MeasureTheory.IsSetSemiring.disjointOfUnion_props`.-/ +noncomputable def disjointOfUnion (hC : IsSetSemiring C) (hJ : ↑J ⊆ C) (j : Set α) := + (hC.disjointOfUnion_props hJ).choose j + +lemma pairwiseDisjoint_disjointOfUnion (hC : IsSetSemiring C) (hJ : ↑J ⊆ C) : + PairwiseDisjoint J (hC.disjointOfUnion hJ) := + (Exists.choose_spec (hC.disjointOfUnion_props hJ)).1 + +lemma disjointOfUnion_subset (hC : IsSetSemiring C) (hJ : ↑J ⊆ C) (hj : j ∈ J) : + (disjointOfUnion hC hJ j : Set (Set α)) ⊆ C := + (Exists.choose_spec (hC.disjointOfUnion_props hJ)).2.1 _ hj + +lemma pairwiseDisjoint_biUnion_disjointOfUnion (hC : IsSetSemiring C) (hJ : ↑J ⊆ C) : + PairwiseDisjoint (⋃ x ∈ J, (hC.disjointOfUnion hJ x : Set (Set α))) id := + (Exists.choose_spec (hC.disjointOfUnion_props hJ)).2.2.1 + +lemma pairwiseDisjoint_disjointOfUnion_of_mem (hC : IsSetSemiring C) (hJ : ↑J ⊆ C) (hj : j ∈ J) : + PairwiseDisjoint (hC.disjointOfUnion hJ j : Set (Set α)) id := by + apply PairwiseDisjoint.subset (hC.pairwiseDisjoint_biUnion_disjointOfUnion hJ) + exact subset_iUnion₂_of_subset j hj fun ⦃a⦄ a ↦ a + +lemma disjointOfUnion_subset_of_mem (hC : IsSetSemiring C) (hJ : ↑J ⊆ C) (hj : j ∈ J) : + ⋃₀ hC.disjointOfUnion hJ j ⊆ j := + (Exists.choose_spec (hC.disjointOfUnion_props hJ)).2.2.2.1 j hj + +lemma subset_of_mem_disjointOfUnion (hC : IsSetSemiring C) (hJ : ↑J ⊆ C) (hj : j ∈ J) {x : Set α} + (hx : x ∈ (hC.disjointOfUnion hJ) j) : x ⊆ j := + sUnion_subset_iff.mp (hC.disjointOfUnion_subset_of_mem hJ hj) x hx + +lemma empty_nmem_disjointOfUnion (hC : IsSetSemiring C) (hJ : ↑J ⊆ C) (hj : j ∈ J) : + ∅ ∉ hC.disjointOfUnion hJ j := + (Exists.choose_spec (hC.disjointOfUnion_props hJ)).2.2.2.2.1 j hj + +lemma sUnion_disjointOfUnion (hC : IsSetSemiring C) (hJ : ↑J ⊆ C) : + ⋃₀ ⋃ x ∈ J, (hC.disjointOfUnion hJ x : Set (Set α)) = ⋃₀ J + := (Exists.choose_spec (hC.disjointOfUnion_props hJ)).2.2.2.2.2.symm + +end disjointOfUnion end IsSetSemiring diff --git a/Mathlib/ModelTheory/PartialEquiv.lean b/Mathlib/ModelTheory/PartialEquiv.lean index 6a0442f49b37a2..2c595a9111371a 100644 --- a/Mathlib/ModelTheory/PartialEquiv.lean +++ b/Mathlib/ModelTheory/PartialEquiv.lean @@ -347,7 +347,8 @@ lemma partialEquivLimit_comp_inclusion {i : ι} : theorem le_partialEquivLimit (i : ι) : S i ≤ partialEquivLimit S := ⟨le_iSup (f := fun i ↦ (S i).dom) _, by - #adaptation_note /-- After https://github.com/leanprover/lean4/pull/5020, these two `simp` calls cannot be combined. -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/5020 + these two `simp` calls cannot be combined. -/ simp only [partialEquivLimit_comp_inclusion] simp only [cod_partialEquivLimit, dom_partialEquivLimit, ← Embedding.comp_assoc, subtype_comp_inclusion]⟩ diff --git a/Mathlib/NumberTheory/AbelSummation.lean b/Mathlib/NumberTheory/AbelSummation.lean index 202136f5c40add..3594e9a1f4abea 100644 --- a/Mathlib/NumberTheory/AbelSummation.lean +++ b/Mathlib/NumberTheory/AbelSummation.lean @@ -51,8 +51,8 @@ namespace abelSummationProof open intervalIntegral IntervalIntegrable -private theorem sumlocc (n : ℕ) : - ∀ᵐ t, t ∈ Set.Icc (n : ℝ) (n + 1) → ∑ k ∈ Icc 0 ⌊t⌋₊, c k = ∑ k ∈ Icc 0 n, c k := by +private theorem sumlocc {m : ℕ} (n : ℕ) : + ∀ᵐ t, t ∈ Set.Icc (n : ℝ) (n + 1) → ∑ k ∈ Icc m ⌊t⌋₊, c k = ∑ k ∈ Icc m n, c k := by filter_upwards [Ico_ae_eq_Icc] with t h ht rw [Nat.floor_eq_on_Ico _ _ (h.mpr ht)] @@ -83,11 +83,12 @@ private theorem ineqofmemIco' {k : ℕ} (hk : k ∈ Ico (⌊a⌋₊ + 1) ⌊b⌋ a ≤ k ∧ k + 1 ≤ b := ineqofmemIco (by rwa [← Finset.coe_Ico]) -private theorem integrablemulsum (ha : 0 ≤ a) {g : ℝ → 𝕜} (hg_int : IntegrableOn g (Set.Icc a b)) : - IntegrableOn (fun t ↦ g t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) (Set.Icc a b) := by +theorem _root_.integrableOn_mul_sum_Icc {m : ℕ} (ha : 0 ≤ a) {g : ℝ → 𝕜} + (hg_int : IntegrableOn g (Set.Icc a b)) : + IntegrableOn (fun t ↦ g t * ∑ k ∈ Icc m ⌊t⌋₊, c k) (Set.Icc a b) := by obtain hab | hab := le_or_gt a b · obtain hb | hb := eq_or_lt_of_le (Nat.floor_le_floor hab) - · have : ∀ᵐ t, t ∈ Set.Icc a b → ∑ k ∈ Icc 0 ⌊a⌋₊, c k = ∑ k ∈ Icc 0 ⌊t⌋₊, c k := by + · have : ∀ᵐ t, t ∈ Set.Icc a b → ∑ k ∈ Icc m ⌊a⌋₊, c k = ∑ k ∈ Icc m ⌊t⌋₊, c k := by filter_upwards [sumlocc c ⌊a⌋₊] with t ht₁ ht₂ rw [ht₁ ⟨(Nat.floor_le ha).trans ht₂.1, hb ▸ ht₂.2.trans (Nat.lt_floor_add_one b).le⟩] rw [← ae_restrict_iff' measurableSet_Icc] at this @@ -95,7 +96,7 @@ private theorem integrablemulsum (ha : 0 ≤ a) {g : ℝ → 𝕜} (hg_int : Int (hg_int.mul_const _) ((Filter.EventuallyEq.refl _ g).mul this) · have h_locint {t₁ t₂ : ℝ} {n : ℕ} (h : t₁ ≤ t₂) (h₁ : n ≤ t₁) (h₂ : t₂ ≤ n + 1) (h₃ : a ≤ t₁) (h₄ : t₂ ≤ b) : - IntervalIntegrable (fun t ↦ g t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) volume t₁ t₂ := by + IntervalIntegrable (fun t ↦ g t * ∑ k ∈ Icc m ⌊t⌋₊, c k) volume t₁ t₂ := by rw [intervalIntegrable_iff_integrableOn_Icc_of_le h] exact (IntegrableOn.mono_set (hg_int.mul_const _) (Set.Icc_subset_Icc h₃ h₄)).congr <| ae_restrict_of_ae_restrict_of_subset (Set.Icc_subset_Icc h₁ h₂) @@ -154,15 +155,15 @@ theorem _root_.sum_mul_eq_sub_sub_integral_mul (ha : 0 ≤ a) (hab : a ≤ b) -- (Note we have 5 goals, but the 1st and 3rd are identical. TODO: find a non-hacky way of dealing -- with both at once.) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux6] - exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux5) + exact (integrableOn_mul_sum_Icc c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux5) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux5] - exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_left aux6) + exact (integrableOn_mul_sum_Icc c ha hf_int).mono_set (Set.Icc_subset_Icc_left aux6) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux6] - exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux5) + exact (integrableOn_mul_sum_Icc c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux5) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux3] - exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux4) + exact (integrableOn_mul_sum_Icc c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux4) · exact fun k hk ↦ (intervalIntegrable_iff_integrableOn_Icc_of_le (mod_cast k.le_succ)).mpr - <| (integrablemulsum c ha hf_int).mono_set + <| (integrableOn_mul_sum_Icc c ha hf_int).mono_set <| (Set.Icc_subset_Icc_iff (mod_cast k.le_succ)).mpr <| mod_cast (ineqofmemIco hk) /-- A version of `sum_mul_eq_sub_sub_integral_mul` where the endpoints are `Nat`. -/ @@ -209,7 +210,7 @@ theorem sum_mul_eq_sub_integral_mul₀ (hc : c 0 = 0) (b : ℝ) f b * (∑ k ∈ Icc 0 ⌊b⌋₊, c k) - ∫ t in Set.Ioc 1 b, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k := by obtain hb | hb := le_or_gt 1 b · have : 1 ≤ ⌊b⌋₊ := (Nat.one_le_floor_iff _).mpr hb - nth_rewrite 1 [Icc_eq_cons_Ioc (by omega), sum_cons, ← Nat.Icc_succ_left, + nth_rewrite 1 [Icc_eq_cons_Ioc (Nat.zero_le _), sum_cons, ← Nat.Icc_succ_left, Icc_eq_cons_Ioc (by omega), sum_cons] rw [Nat.succ_eq_add_one, zero_add, ← Nat.floor_one (α := ℝ), sum_mul_eq_sub_sub_integral_mul c zero_le_one hb hf_diff hf_int, Nat.floor_one, Nat.cast_one, @@ -235,20 +236,22 @@ section limit open Filter Topology abelSummationProof intervalIntegral -private theorem locallyintegrablemulsum (ha : 0 ≤ a) {g : ℝ → 𝕜} (hg : IntegrableOn g (Set.Ici a)) : - LocallyIntegrableOn (fun t ↦ g t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) (Set.Ici a) := by +theorem locallyIntegrableOn_mul_sum_Icc {m : ℕ} (ha : 0 ≤ a) {g : ℝ → 𝕜} + (hg : LocallyIntegrableOn g (Set.Ici a)) : + LocallyIntegrableOn (fun t ↦ g t * ∑ k ∈ Icc m ⌊t⌋₊, c k) (Set.Ici a) := by refine (locallyIntegrableOn_iff isLocallyClosed_Ici).mpr fun K hK₁ hK₂ ↦ ?_ by_cases hK₃ : K.Nonempty · have h_inf : a ≤ sInf K := (hK₁ (hK₂.sInf_mem hK₃)) refine IntegrableOn.mono_set ?_ (Bornology.IsBounded.subset_Icc_sInf_sSup hK₂.isBounded) - refine integrablemulsum _ (ha.trans h_inf) (hg.mono_set ?_) + refine integrableOn_mul_sum_Icc _ (ha.trans h_inf) ?_ + refine hg.integrableOn_compact_subset ?_ isCompact_Icc exact (Set.Icc_subset_Ici_iff (Real.sInf_le_sSup _ hK₂.bddBelow hK₂.bddAbove)).mpr h_inf · rw [Set.not_nonempty_iff_eq_empty.mp hK₃] exact integrableOn_empty theorem tendsto_sum_mul_atTop_nhds_one_sub_integral (hf_diff : ∀ t ∈ Set.Ici 0, DifferentiableAt ℝ f t) - (hf_int : IntegrableOn (deriv f) (Set.Ici 0)) {l : 𝕜} + (hf_int : LocallyIntegrableOn (deriv f) (Set.Ici 0)) {l : 𝕜} (h_lim : Tendsto (fun n : ℕ ↦ f n * ∑ k ∈ Icc 0 n, c k) atTop (𝓝 l)) {g : ℝ → 𝕜} (hg_dom : (fun t ↦ deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) =O[atTop] g) (hg_int : IntegrableAtFilter g atTop) : @@ -259,14 +262,15 @@ theorem tendsto_sum_mul_atTop_nhds_one_sub_integral refine Tendsto.congr (fun _ ↦ by rw [← integral_of_le (Nat.cast_nonneg _)]) ?_ refine intervalIntegral_tendsto_integral_Ioi _ ?_ tendsto_natCast_atTop_atTop exact integrableOn_Ici_iff_integrableOn_Ioi.mp - <| (locallyintegrablemulsum c le_rfl hf_int).integrableOn_of_isBigO_atTop hg_dom hg_int + <| (locallyIntegrableOn_mul_sum_Icc c le_rfl hf_int).integrableOn_of_isBigO_atTop + hg_dom hg_int refine (h_lim.sub h_lim').congr (fun _ ↦ ?_) - rw [sum_mul_eq_sub_integral_mul' _ _ (fun t ht ↦ hf_diff _ ht.1) - (hf_int.mono_set Set.Icc_subset_Ici_self)] + rw [sum_mul_eq_sub_integral_mul' _ _ (fun t ht ↦ hf_diff _ ht.1)] + exact hf_int.integrableOn_compact_subset Set.Icc_subset_Ici_self isCompact_Icc theorem tendsto_sum_mul_atTop_nhds_one_sub_integral₀ (hc : c 0 = 0) (hf_diff : ∀ t ∈ Set.Ici 1, DifferentiableAt ℝ f t) - (hf_int : IntegrableOn (deriv f) (Set.Ici 1)) {l : 𝕜} + (hf_int : LocallyIntegrableOn (deriv f) (Set.Ici 1)) {l : 𝕜} (h_lim: Tendsto (fun n : ℕ ↦ f n * ∑ k ∈ Icc 0 n, c k) atTop (𝓝 l)) {g : ℝ → ℝ} (hg_dom : (fun t ↦ deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) =O[atTop] g) (hg_int : IntegrableAtFilter g atTop) : @@ -280,10 +284,11 @@ theorem tendsto_sum_mul_atTop_nhds_one_sub_integral₀ (hc : c 0 = 0) atTop (𝓝 (∫ t in Set.Ioi 1, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k)) := by refine Tendsto.congr' h (intervalIntegral_tendsto_integral_Ioi _ ?_ tendsto_natCast_atTop_atTop) exact integrableOn_Ici_iff_integrableOn_Ioi.mp - <| (locallyintegrablemulsum c zero_le_one hf_int).integrableOn_of_isBigO_atTop hg_dom hg_int + <| (locallyIntegrableOn_mul_sum_Icc c zero_le_one hf_int).integrableOn_of_isBigO_atTop + hg_dom hg_int refine (h_lim.sub h_lim').congr (fun _ ↦ ?_) - rw [sum_mul_eq_sub_integral_mul₀' _ hc _ (fun t ht ↦ hf_diff _ ht.1) - (hf_int.mono_set Set.Icc_subset_Ici_self)] + rw [sum_mul_eq_sub_integral_mul₀' _ hc _ (fun t ht ↦ hf_diff _ ht.1)] + exact hf_int.integrableOn_compact_subset Set.Icc_subset_Ici_self isCompact_Icc end limit @@ -293,7 +298,7 @@ open Filter abelSummationProof private theorem summable_mul_of_bigO_atTop_aux (m : ℕ) (h_bdd : (fun n : ℕ ↦ ‖f n‖ * ∑ k ∈ Icc 0 n, ‖c k‖) =O[atTop] fun _ ↦ (1 : ℝ)) - (hf_int : IntegrableOn (deriv (fun t ↦ ‖f t‖)) (Set.Ici (m : ℝ))) + (hf_int : LocallyIntegrableOn (deriv (fun t ↦ ‖f t‖)) (Set.Ici (m : ℝ))) (hf : ∀ n : ℕ, ∑ k ∈ Icc 0 n, ‖f k‖ * ‖c k‖ = ‖f n‖ * ∑ k ∈ Icc 0 n, ‖c k‖ - ∫ (t : ℝ) in Set.Ioc ↑m ↑n, deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 0 ⌊t⌋₊, ‖c k‖) @@ -325,26 +330,27 @@ private theorem summable_mul_of_bigO_atTop_aux (m : ℕ) · exact add_le_add_left (le_trans (neg_le_abs _) (Real.norm_eq_abs _ ▸ norm_integral_le_integral_norm _)) _ · refine add_le_add_left (setIntegral_mono_set ?_ ?_ Set.Ioc_subset_Ioi_self.eventuallyLE) C₁ - · refine integrableOn_Ici_iff_integrableOn_Ioi.mp <| + · exact integrableOn_Ici_iff_integrableOn_Ioi.mp <| (integrable_norm_iff h_mes.aestronglyMeasurable).mpr <| - (locallyintegrablemulsum _ m.cast_nonneg hf_int).integrableOn_of_isBigO_atTop hg₁ hg₂ + (locallyIntegrableOn_mul_sum_Icc _ m.cast_nonneg hf_int).integrableOn_of_isBigO_atTop + hg₁ hg₂ · filter_upwards with t using norm_nonneg _ theorem summable_mul_of_bigO_atTop (hf_diff : ∀ t ∈ Set.Ici 0, DifferentiableAt ℝ (fun x ↦ ‖f x‖) t) - (hf_int : IntegrableOn (deriv (fun t ↦ ‖f t‖)) (Set.Ici 0)) + (hf_int : LocallyIntegrableOn (deriv (fun t ↦ ‖f t‖)) (Set.Ici 0)) (h_bdd : (fun n : ℕ ↦ ‖f n‖ * ∑ k ∈ Icc 0 n, ‖c k‖) =O[atTop] fun _ ↦ (1 : ℝ)) {g : ℝ → ℝ} (hg₁ : (fun t ↦ deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 0 ⌊t⌋₊, ‖c k‖) =O[atTop] g) (hg₂ : IntegrableAtFilter g atTop) : Summable (fun n : ℕ ↦ f n * c n) := by refine summable_mul_of_bigO_atTop_aux c 0 h_bdd (by rwa [Nat.cast_zero]) (fun n ↦ ?_) hg₁ hg₂ exact_mod_cast sum_mul_eq_sub_integral_mul' _ _ (fun _ ht ↦ hf_diff _ ht.1) - (hf_int.mono_set Set.Icc_subset_Ici_self) + (hf_int.integrableOn_compact_subset Set.Icc_subset_Ici_self isCompact_Icc) /-- A version of `summable_mul_of_bigO_atTop` that can be useful to avoid difficulties near zero. -/ theorem summable_mul_of_bigO_atTop' (hf_diff : ∀ t ∈ Set.Ici 1, DifferentiableAt ℝ (fun x ↦ ‖f x‖) t) - (hf_int : IntegrableOn (deriv (fun t ↦ ‖f t‖)) (Set.Ici 1)) + (hf_int : LocallyIntegrableOn (deriv (fun t ↦ ‖f t‖)) (Set.Ici 1)) (h_bdd : (fun n : ℕ ↦ ‖f n‖ * ∑ k ∈ Icc 1 n, ‖c k‖) =O[atTop] fun _ ↦ (1 : ℝ)) {g : ℝ → ℝ} (hg₁ : (fun t ↦ deriv (fun t ↦ ‖f t‖) t * ∑ k ∈ Icc 1 ⌊t⌋₊, ‖c k‖) =O[atTop] g) (hg₂ : IntegrableAtFilter g atTop) : @@ -358,7 +364,8 @@ theorem summable_mul_of_bigO_atTop' refine Summable.congr_atTop (summable_mul_of_bigO_atTop_aux (fun n ↦ if n = 0 then 0 else c n) 1 h_bdd (by rwa [Nat.cast_one]) (fun n ↦ ?_) hg₁ hg₂) ?_ · exact_mod_cast sum_mul_eq_sub_integral_mul₀' _ (by simp only [reduceIte, norm_zero]) n - (fun _ ht ↦ hf_diff _ ht.1) (hf_int.mono_set Set.Icc_subset_Ici_self) + (fun _ ht ↦ hf_diff _ ht.1) + (hf_int.integrableOn_compact_subset Set.Icc_subset_Ici_self isCompact_Icc) · filter_upwards [eventually_ne_atTop 0] with k hk simp_rw [if_neg hk] diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean index 7e288de6b9c304..bff8147059e585 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean @@ -98,8 +98,8 @@ theorem exists_approx_aux (n : ℕ) (h : abv.IsAdmissible) : · rwa [Finset.length_toList] · ext simpa [(Finset.nodup_toList _).getElem_inj_iff] using h - · #adaptation_note - /-- This proof was nicer prior to https://github.com/leanprover/lean4/pull/4400. + · #adaptation_note /-- https://github.com/leanprover/lean4/pull/4400 + This proof was nicer before. Please feel welcome to improve it, by avoiding use of `List.get` in favour of `GetElem`. -/ have : ∀ i h, t ((Finset.univ.filter fun x ↦ t x = s).toList.get ⟨i, h⟩) = s := fun i h ↦ (Finset.mem_filter.mp (Finset.mem_toList.mp (List.get_mem _ ⟨i, h⟩))).2 diff --git a/Mathlib/NumberTheory/ClassNumber/FunctionField.lean b/Mathlib/NumberTheory/ClassNumber/FunctionField.lean index 04ead5e6912634..6c6aecd30fe2d0 100644 --- a/Mathlib/NumberTheory/ClassNumber/FunctionField.lean +++ b/Mathlib/NumberTheory/ClassNumber/FunctionField.lean @@ -29,12 +29,11 @@ variable [Algebra Fq[X] F] [Algebra (RatFunc Fq) F] variable [IsScalarTower Fq[X] (RatFunc Fq) F] variable [FunctionField Fq F] [Algebra.IsSeparable (RatFunc Fq) F] -open scoped Classical - namespace RingOfIntegers open FunctionField +open scoped Classical in noncomputable instance : Fintype (ClassGroup (ringOfIntegers Fq F)) := ClassGroup.fintypeOfAdmissibleOfFinite (RatFunc Fq) F (Polynomial.cardPowDegreeIsAdmissible : diff --git a/Mathlib/NumberTheory/Cyclotomic/Rat.lean b/Mathlib/NumberTheory/Cyclotomic/Rat.lean index dad9eca2e69658..6f56aadeff7a6b 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Rat.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Rat.lean @@ -214,12 +214,10 @@ theorem integralPowerBasis_gen [hcycl : IsCyclotomicExtension {p ^ k} ℚ K] simp only [adjoinEquivRingOfIntegers_apply, IsIntegralClosure.algebraMap_lift] rfl -#adaptation_note -/-- +#adaptation_note /-- https://github.com/leanprover/lean4/pull/5338 We name `hcycl` so it can be used as a named argument, but since https://github.com/leanprover/lean4/pull/5338, this is considered unused, -so we need to disable the linter. --/ +so we need to disable the linter. -/ set_option linter.unusedVariables false in @[simp] theorem integralPowerBasis_dim [hcycl : IsCyclotomicExtension {p ^ k} ℚ K] diff --git a/Mathlib/NumberTheory/Dioph.lean b/Mathlib/NumberTheory/Dioph.lean index 6c7c378d55eb41..2b065a6f7dd2e5 100644 --- a/Mathlib/NumberTheory/Dioph.lean +++ b/Mathlib/NumberTheory/Dioph.lean @@ -26,7 +26,7 @@ there exists `t : ℕ^β` with `p (v, t) = 0`. * `Dioph`: a predicate stating that a set is Diophantine, i.e. a set `S ⊆ ℕ^α` is Diophantine if there exists a polynomial on `α ⊕ β` such that `v ∈ S` iff there exists `t : ℕ^β` with `p (v, t) = 0`. -* `dioph_fn`: a predicate on a function stating that it is Diophantine in the sense that its graph +* `DiophFn`: a predicate on a function stating that it is Diophantine in the sense that its graph is Diophantine as a set. ## Main statements diff --git a/Mathlib/NumberTheory/LSeries/Basic.lean b/Mathlib/NumberTheory/LSeries/Basic.lean index 3faf700efb4030..3f0a95887411db 100644 --- a/Mathlib/NumberTheory/LSeries/Basic.lean +++ b/Mathlib/NumberTheory/LSeries/Basic.lean @@ -68,6 +68,12 @@ lemma term_def (f : ℕ → ℂ) (s : ℂ) (n : ℕ) : term f s n = if n = 0 then 0 else f n / n ^ s := rfl +/-- An alternate spelling of `term_def` for the case `f 0 = 0`. -/ +lemma term_def₀ {f : ℕ → ℂ} (hf : f 0 = 0) (s : ℂ) (n : ℕ) : + LSeries.term f s n = f n * (n : ℂ) ^ (- s) := by + rw [LSeries.term] + split_ifs with h <;> simp [h, hf, cpow_neg, div_eq_inv_mul, mul_comm] + @[simp] lemma term_zero (f : ℕ → ℂ) (s : ℂ) : term f s 0 = 0 := rfl diff --git a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean index 4b59c880997f0a..46c73ec9805a35 100644 --- a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean +++ b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean @@ -38,7 +38,7 @@ All definitions and theorems are in the `DirichletCharacter` namespace. `completedLFunction χ s = N ^ (s - 1 / 2) * rootNumber χ * completedLFunction χ⁻¹ s`. -/ -open HurwitzZeta Complex Finset Classical ZMod Filter +open HurwitzZeta Complex Finset ZMod Filter open scoped Real Topology @@ -198,6 +198,7 @@ section gammaFactor omit [NeZero N] -- not required for these declarations +open scoped Classical in /-- The Archimedean Gamma factor: `Gammaℝ s` if `χ` is even, and `Gammaℝ (s + 1)` otherwise. -/ noncomputable def gammaFactor (χ : DirichletCharacter ℂ N) (s : ℂ) := if χ.Even then Gammaℝ s else Gammaℝ (s + 1) @@ -258,6 +259,7 @@ lemma LFunction_eq_completed_div_gammaFactor (χ : DirichletCharacter ℂ N) (s · exact LFunction_eq_completed_div_gammaFactor_even hχ.to_fun _ (h.imp_right χ.map_zero') · apply LFunction_eq_completed_div_gammaFactor_odd hχ.to_fun +open scoped Classical in /-- Global root number of `χ` (for `χ` primitive; junk otherwise). Defined as `gaussSum χ stdAddChar / I ^ a / N ^ (1 / 2)`, where `a = 0` if even, `a = 1` if odd. (The factor @@ -277,6 +279,7 @@ namespace IsPrimitive /-- **Functional equation** for primitive Dirichlet L-functions. -/ theorem completedLFunction_one_sub {χ : DirichletCharacter ℂ N} (hχ : IsPrimitive χ) (s : ℂ) : completedLFunction χ (1 - s) = N ^ (s - 1 / 2) * rootNumber χ * completedLFunction χ⁻¹ s := by + classical -- First handle special case of Riemann zeta rcases eq_or_ne N 1 with rfl | hN · simp only [completedLFunction_modOne_eq, completedRiemannZeta_one_sub, Nat.cast_one, one_cpow, diff --git a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean index 9782d336636d6d..ceccbcc2a6eab6 100644 --- a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean +++ b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean @@ -46,7 +46,7 @@ multiples of `1 / s` and `1 / (1 - s)`. -/ noncomputable section -open Complex Filter Topology Asymptotics Real Set Classical MeasureTheory +open Complex Filter Topology Asymptotics Real Set MeasureTheory namespace HurwitzZeta diff --git a/Mathlib/NumberTheory/LSeries/Nonvanishing.lean b/Mathlib/NumberTheory/LSeries/Nonvanishing.lean index 24e61f8fb24586..c9d59f72fadfd9 100644 --- a/Mathlib/NumberTheory/LSeries/Nonvanishing.lean +++ b/Mathlib/NumberTheory/LSeries/Nonvanishing.lean @@ -374,11 +374,8 @@ private lemma LFunction_ne_zero_of_not_quadratic_or_ne_one {t : ℝ} (h : χ ^ 2 -- go via absolute value to translate into a statement over `ℝ` replace H := (H₀.trans H).norm_right simp only [norm_eq_abs, abs_ofReal] at H - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/6024 - we needed to add `(F' := ℝ)` to `H.of_norm_right`. - -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 + we needed to add `(F' := ℝ)` to `H.of_norm_right`. -/ exact isLittleO_irrefl (.of_forall (fun _ ↦ one_ne_zero)) <| (H.of_norm_right (F' := ℝ)).trans_isLittleO <| isLittleO_id_one.mono nhdsWithin_le_nhds diff --git a/Mathlib/NumberTheory/LSeries/SumCoeff.lean b/Mathlib/NumberTheory/LSeries/SumCoeff.lean new file mode 100644 index 00000000000000..3c70a12320fde4 --- /dev/null +++ b/Mathlib/NumberTheory/LSeries/SumCoeff.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2025 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import Mathlib.Analysis.Asymptotics.SpecificAsymptotics +import Mathlib.Analysis.InnerProductSpace.Calculus +import Mathlib.Analysis.SpecialFunctions.ImproperIntegrals +import Mathlib.NumberTheory.AbelSummation +import Mathlib.NumberTheory.LSeries.Basic + +/-! +# Partial sums of coefficients of L-series + +We prove several results involving partial sums of coefficients (or norm of coefficients) of +L-series. + +## Main results + +* `LSeriesSummable_of_sum_norm_bigO`: for `f : ℕ → ℂ`, if the partial sums + `∑ k ∈ Icc 1 n, ‖f k‖` are `O(n ^ r)` for some real `0 ≤ r`, then the L-series `LSeries f` + converges at `s : ℂ` for all `s` such that `r < s.re`. + +* `LSeries_eq_mul_integral` : for `f : ℕ → ℂ`, if the partial sums `∑ k ∈ Icc 1 n, f k` are + `O(n ^ r)` for some real `0 ≤ r` and the L-series `LSeries f` converges at `s : ℂ` with + `r < s.re`, then `LSeries f s = s * ∫ t in Set.Ioi 1, (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (-(s + 1))`. + +-/ + +open Finset Filter MeasureTheory Topology Complex Asymptotics + +section summable + +variable {f : ℕ → ℂ} {r : ℝ} {s : ℂ} + +private theorem LSeriesSummable_of_sum_norm_bigO_aux (hf : f 0 = 0) + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, ‖f k‖) =O[atTop] fun n ↦ (n : ℝ) ^ r) + (hr : 0 ≤ r) (hs : r < s.re) : + LSeriesSummable f s := by + have h₁ : -s ≠ 0 := neg_ne_zero.mpr <| ne_zero_of_re_pos (hr.trans_lt hs) + have h₂ : (-s).re + r ≤ 0 := by + rw [neg_re, neg_add_nonpos_iff] + exact hs.le + have h₃ (t : ℝ) (ht : t ∈ Set.Ici 1) : DifferentiableAt ℝ (fun x : ℝ ↦ ‖(x : ℂ) ^ (-s)‖) t := + have ht' : t ≠ 0 := (zero_lt_one.trans_le ht).ne' + (differentiableAt_id.ofReal_cpow_const ht' h₁).norm ℝ <| + (cpow_ne_zero_iff_of_exponent_ne_zero h₁).mpr <| ofReal_ne_zero.mpr ht' + have h₄ : (deriv fun t : ℝ ↦ ‖(t : ℂ) ^ (-s)‖) =ᶠ[atTop] fun t ↦ -s.re * t ^ (-(s.re +1)) := by + filter_upwards [eventually_gt_atTop 0] with t ht + rw [deriv_norm_ofReal_cpow _ ht, neg_re, neg_add'] + simp_rw [LSeriesSummable, funext (LSeries.term_def₀ hf s), mul_comm (f _)] + refine summable_mul_of_bigO_atTop' (f := fun t ↦ (t : ℂ) ^ (-s)) + (g := fun t ↦ t ^ (-(s.re + 1) + r)) _ h₃ ?_ ?_ ?_ ?_ + · refine (integrableOn_Ici_iff_integrableOn_Ioi.mpr + (integrableOn_Ioi_deriv_norm_ofReal_cpow zero_lt_one ?_)).locallyIntegrableOn + exact neg_re _ ▸ neg_nonpos.mpr <| hr.trans hs.le + · refine (IsBigO.mul_atTop_rpow_natCast_of_isBigO_rpow _ _ _ ?_ hO h₂).congr_right (by simp) + exact (norm_ofReal_cpow_eventually_eq_atTop _).isBigO.natCast_atTop + · refine h₄.isBigO.of_const_mul_right.mul_atTop_rpow_of_isBigO_rpow _ r _ ?_ le_rfl + exact (hO.comp_tendsto tendsto_nat_floor_atTop).trans <| + isEquivalent_nat_floor.isBigO.rpow hr (eventually_ge_atTop 0) + · rwa [integrableAtFilter_rpow_atTop_iff, neg_add_lt_iff_lt_add, add_neg_cancel_right] + +/-- If the partial sums `∑ k ∈ Icc 1 n, ‖f k‖` are `O(n ^ r)` for some real `0 ≤ r`, then the +L-series `LSeries f` converges at `s : ℂ` for all `s` such that `r < s.re`. -/ +theorem LSeriesSummable_of_sum_norm_bigO + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, ‖f k‖) =O[atTop] fun n ↦ (n : ℝ) ^ r) + (hr : 0 ≤ r) (hs : r < s.re) : + LSeriesSummable f s := by + have h₁ : (fun n ↦ if n = 0 then 0 else f n) =ᶠ[atTop] f := by + filter_upwards [eventually_ne_atTop 0] with n hn using by simp_rw [if_neg hn] + refine (LSeriesSummable_of_sum_norm_bigO_aux (if_pos rfl) ?_ hr hs).congr' _ h₁ + refine hO.congr' (Eventually.of_forall fun _ ↦ Finset.sum_congr rfl fun _ h ↦ ?_) EventuallyEq.rfl + rw [if_neg (zero_lt_one.trans_le (mem_Icc.mp h).1).ne'] + +/-- If `f` takes nonnegative real values and the partial sums `∑ k ∈ Icc 1 n, f k` are `O(n ^ r)` +for some real `0 ≤ r`, then the L-series `LSeries f` converges at `s : ℂ` for all `s` +such that `r < s.re`. -/ +theorem LSeriesSummable_of_sum_norm_bigO_and_nonneg + {f : ℕ → ℝ} (hO : (fun n ↦ ∑ k ∈ Icc 1 n, f k) =O[atTop] fun n ↦ (n : ℝ) ^ r) + (hf : ∀ n, 0 ≤ f n) (hr : 0 ≤ r) (hs : r < s.re) : + LSeriesSummable (fun n ↦ f n) s := + LSeriesSummable_of_sum_norm_bigO (by simpa [_root_.abs_of_nonneg (hf _)]) hr hs + +end summable + +section integralrepresentation + +private theorem LSeries_eq_mul_integral_aux {f : ℕ → ℂ} (hf : f 0 = 0) {r : ℝ} (hr : 0 ≤ r) {s : ℂ} + (hs : r < s.re) (hS : LSeriesSummable f s) + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, f k) =O[atTop] fun n ↦ (n : ℝ) ^ r) : + LSeries f s = s * ∫ t in Set.Ioi (1 : ℝ), (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (- (s + 1)) := by + have h₁ : (-s - 1).re + r < -1 := by + rwa [sub_re, one_re, neg_re, neg_sub_left, neg_add_lt_iff_lt_add, add_neg_cancel_comm] + have h₂ : s ≠ 0 := ne_zero_of_re_pos (hr.trans_lt hs) + have h₃ (t : ℝ) (ht : t ∈ Set.Ici 1) : DifferentiableAt ℝ (fun x : ℝ ↦ (x : ℂ) ^ (-s)) t := + differentiableAt_id.ofReal_cpow_const (zero_lt_one.trans_le ht).ne' (neg_ne_zero.mpr h₂) + have h₄ : ∀ n, ∑ k ∈ Icc 0 n, f k = ∑ k ∈ Icc 1 n, f k := fun n ↦ by + rw [← Nat.Icc_insert_succ_left n.zero_le, sum_insert (by aesop), hf, zero_add, zero_add] + simp_rw [← h₄] at hO + rw [← integral_mul_left] + refine tendsto_nhds_unique ((tendsto_add_atTop_iff_nat 1).mpr hS.hasSum.tendsto_sum_nat) ?_ + simp_rw [Nat.range_succ_eq_Icc_zero, LSeries.term_def₀ hf, mul_comm (f _)] + convert tendsto_sum_mul_atTop_nhds_one_sub_integral₀ (f := fun x ↦ (x : ℂ) ^ (-s)) (l := 0) + ?_ hf h₃ ?_ ?_ ?_ (integrableAtFilter_rpow_atTop_iff.mpr h₁) + · rw [zero_sub, ← integral_neg] + refine setIntegral_congr_fun measurableSet_Ioi fun t ht ↦ ?_ + rw [deriv_ofReal_cpow_const (zero_lt_one.trans ht).ne', h₄] + · ring_nf + · exact neg_ne_zero.mpr <| ne_zero_of_re_pos (hr.trans_lt hs) + · refine (integrableOn_Ici_iff_integrableOn_Ioi.mpr <| + integrableOn_Ioi_deriv_ofReal_cpow zero_lt_one + (by simpa using hr.trans_lt hs)).locallyIntegrableOn + · have hlim : Tendsto (fun n : ℕ ↦ (n : ℝ) ^ (-(s.re - r))) atTop (𝓝 0) := + (tendsto_rpow_neg_atTop (by rwa [sub_pos])).comp tendsto_natCast_atTop_atTop + refine (IsBigO.mul_atTop_rpow_natCast_of_isBigO_rpow (-s.re) _ _ ?_ hO ?_).trans_tendsto hlim + · exact isBigO_norm_left.mp <| (norm_ofReal_cpow_eventually_eq_atTop _).isBigO.natCast_atTop + · linarith + · refine .mul_atTop_rpow_of_isBigO_rpow (-(s + 1).re) r _ ?_ ?_ (by rw [← neg_re, neg_add']) + · simpa [- neg_add_rev, neg_add'] using isBigO_deriv_ofReal_cpow_const_atTop _ + · exact (hO.comp_tendsto tendsto_nat_floor_atTop).trans <| + isEquivalent_nat_floor.isBigO.rpow hr (eventually_ge_atTop 0) + +/-- If the partial sums `∑ k ∈ Icc 1 n, f k` are `O(n ^ r)` for some real `0 ≤ r` and the +L-series `LSeries f` converges at `s : ℂ` with `r < s.re`, then +`LSeries f s = s * ∫ t in Set.Ioi 1, (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (-(s + 1))`. -/ +theorem LSeries_eq_mul_integral (f : ℕ → ℂ) {r : ℝ} (hr : 0 ≤ r) {s : ℂ} (hs : r < s.re) + (hS : LSeriesSummable f s) + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, f k) =O[atTop] fun n ↦ (n : ℝ) ^ r) : + LSeries f s = s * ∫ t in Set.Ioi (1 : ℝ), (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (-(s + 1)) := by + rw [← LSeriesSummable_congr' s (f := fun n ↦ if n = 0 then 0 else f n) + (by filter_upwards [eventually_ne_atTop 0] with n h using if_neg h)] at hS + have (n) : ∑ k ∈ Icc 1 n, (if k = 0 then 0 else f k) = ∑ k ∈ Icc 1 n, f k := + Finset.sum_congr rfl fun k hk ↦ by rw [if_neg (zero_lt_one.trans_le (mem_Icc.mp hk).1).ne'] + rw [← LSeries_congr _ (fun _ ↦ if_neg _), LSeries_eq_mul_integral_aux (if_pos rfl) hr hs hS] <;> + simp_all + +/-- A version of `LSeries_eq_mul_integral` where we use the stronger condition that the partial sums +`∑ k ∈ Icc 1 n, ‖f k‖` are `O(n ^ r)` to deduce the integral representation. -/ +theorem LSeries_eq_mul_integral' (f : ℕ → ℂ) {r : ℝ} (hr : 0 ≤ r) {s : ℂ} (hs : r < s.re) + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, ‖f k‖) =O[atTop] fun n ↦ (n : ℝ) ^ r) : + LSeries f s = s * ∫ t in Set.Ioi (1 : ℝ), (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (-(s + 1)) := + LSeries_eq_mul_integral _ hr hs (LSeriesSummable_of_sum_norm_bigO hO hr hs) <| + (isBigO_of_le _ fun _ ↦ (norm_sum_le _ _).trans <| Real.le_norm_self _).trans hO + +/-- If `f` takes nonnegative real values and the partial sums `∑ k ∈ Icc 1 n, f k` are `O(n ^ r)` +for some real `0 ≤ r`, then for `s : ℂ` with `r < s.re`, we have +`LSeries f s = s * ∫ t in Set.Ioi 1, (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (-(s + 1))`. -/ +theorem LSeries_eq_mul_integral_of_nonneg (f : ℕ → ℝ) {r : ℝ} (hr : 0 ≤ r) {s : ℂ} (hs : r < s.re) + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, f k) =O[atTop] fun n ↦ (n : ℝ) ^ r) (hf : ∀ n, 0 ≤ f n) : + LSeries (fun n ↦ f n) s = + s * ∫ t in Set.Ioi (1 : ℝ), (∑ k ∈ Icc 1 ⌊t⌋₊, (f k : ℂ)) * t ^ (-(s + 1)) := + LSeries_eq_mul_integral' _ hr hs <| hO.congr_left fun _ ↦ by simp [_root_.abs_of_nonneg (hf _)] + +end integralrepresentation diff --git a/Mathlib/NumberTheory/LSeries/ZMod.lean b/Mathlib/NumberTheory/LSeries/ZMod.lean index e25eb91d295544..67ee56461069a2 100644 --- a/Mathlib/NumberTheory/LSeries/ZMod.lean +++ b/Mathlib/NumberTheory/LSeries/ZMod.lean @@ -55,7 +55,7 @@ Results for completed L-functions: the functional equation relating `completedLFunction Φ (1 - s)` to `completedLFunction (𝓕 Φ) s`. -/ -open HurwitzZeta Complex ZMod Finset Classical Topology Filter Set +open HurwitzZeta Complex ZMod Finset Topology Filter Set open scoped Real diff --git a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean index f408526e3fb67b..d609a2dcc3395b 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean @@ -508,10 +508,10 @@ open NumberTheorySymbols jacobiSym private def fastJacobiSymAux (a b : ℕ) (flip : Bool) (ha0 : a > 0) : ℤ := if ha4 : a % 4 = 0 then fastJacobiSymAux (a / 4) b flip - (a.div_pos (Nat.le_of_dvd ha0 (Nat.dvd_of_mod_eq_zero ha4)) (by decide)) + (Nat.div_pos (Nat.le_of_dvd ha0 (Nat.dvd_of_mod_eq_zero ha4)) (by decide)) else if ha2 : a % 2 = 0 then fastJacobiSymAux (a / 2) b (xor (b % 8 = 3 ∨ b % 8 = 5) flip) - (a.div_pos (Nat.le_of_dvd ha0 (Nat.dvd_of_mod_eq_zero ha2)) (by decide)) + (Nat.div_pos (Nat.le_of_dvd ha0 (Nat.dvd_of_mod_eq_zero ha2)) (by decide)) else if ha1 : a = 1 then if flip then -1 else 1 else if hba : b % a = 0 then diff --git a/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean b/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean index b106e2e0d0de90..321736dea55cba 100644 --- a/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean +++ b/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean @@ -34,7 +34,7 @@ hence Dirichlet L-functions, etc). `∞`. -/ -open Set Filter Topology Asymptotics Real Classical +open Set Filter Topology Asymptotics Real noncomputable section diff --git a/Mathlib/NumberTheory/ModularForms/QExpansion.lean b/Mathlib/NumberTheory/ModularForms/QExpansion.lean index 27d4e3389e4bb1..beb1790478637b 100644 --- a/Mathlib/NumberTheory/ModularForms/QExpansion.lean +++ b/Mathlib/NumberTheory/ModularForms/QExpansion.lean @@ -149,7 +149,7 @@ lemma qExpansionFormalMultilinearSeries_apply_norm (m : ℕ) : lemma qExpansionFormalMultilinearSeries_radius [NeZero n] [ModularFormClass F Γ(n) k] : 1 ≤ (qExpansionFormalMultilinearSeries n f).radius := by - refine le_of_forall_ge_of_dense fun r hr ↦ ?_ + refine le_of_forall_lt_imp_le_of_dense fun r hr ↦ ?_ lift r to NNReal using hr.ne_top apply FormalMultilinearSeries.le_radius_of_summable simp only [qExpansionFormalMultilinearSeries_apply_norm] @@ -161,8 +161,8 @@ lemma qExpansionFormalMultilinearSeries_radius [NeZero n] [ModularFormClass F Γ lemma hasFPowerSeries_cuspFunction [NeZero n] [ModularFormClass F Γ(n) k] : HasFPowerSeriesOnBall (cuspFunction n f) (qExpansionFormalMultilinearSeries n f) 0 1 := by refine ⟨qExpansionFormalMultilinearSeries_radius n f, zero_lt_one, fun hy ↦ ?_⟩ - rw [EMetric.mem_ball, edist_zero_right, ENNReal.coe_lt_one_iff, ← NNReal.coe_lt_one, - coe_nnnorm, norm_eq_abs] at hy + rw [EMetric.mem_ball, edist_zero_right, enorm_eq_nnnorm, ENNReal.coe_lt_one_iff, + ← NNReal.coe_lt_one, coe_nnnorm, norm_eq_abs] at hy simpa [qExpansionFormalMultilinearSeries] using hasSum_qExpansion_of_abs_lt n f hy end ModularFormClass diff --git a/Mathlib/NumberTheory/NumberField/AdeleRing.lean b/Mathlib/NumberTheory/NumberField/AdeleRing.lean index 927332bf8bdebf..fcbefa15560a6b 100644 --- a/Mathlib/NumberTheory/NumberField/AdeleRing.lean +++ b/Mathlib/NumberTheory/NumberField/AdeleRing.lean @@ -39,8 +39,6 @@ namespace NumberField open InfinitePlace AbsoluteValue.Completion InfinitePlace.Completion DedekindDomain IsDedekindDomain -open scoped Classical - /-! ## The infinite adele ring The infinite adele ring is the finite product of completions of a number field over its @@ -76,6 +74,7 @@ theorem algebraMap_apply (x : K) (v : InfinitePlace K) : instance locallyCompactSpace [NumberField K] : LocallyCompactSpace (InfiniteAdeleRing K) := Pi.locallyCompactSpace_of_finite +open scoped Classical in /-- The ring isomorphism between the infinite adele ring of a number field and the space `ℝ ^ r₁ × ℂ ^ r₂`, where `(r₁, r₂)` is the signature of the number field. -/ abbrev ringEquiv_mixedSpace : diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index 46dd10425cc974..8d446e73f8ab60 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -309,10 +309,9 @@ end commMap noncomputable section norm -open scoped Classical - variable {K} +open scoped Classical in /-- The norm at the infinite place `w` of an element of the mixed space. --/ def normAtPlace (w : InfinitePlace K) : (mixedSpace K) →*₀ ℝ where toFun x := if hw : IsReal w then ‖x.1 ⟨w, hw⟩‖ else ‖x.2 ⟨w, not_isReal_iff_isComplex.mp hw⟩‖ @@ -380,6 +379,7 @@ theorem exists_normAtPlace_ne_zero_iff {x : mixedSpace K} : variable [NumberField K] +open scoped Classical in theorem nnnorm_eq_sup_normAtPlace (x : mixedSpace K) : ‖x‖₊ = univ.sup fun w ↦ ⟨normAtPlace w x, normAtPlace_nonneg w x⟩ := by have : @@ -395,6 +395,7 @@ theorem nnnorm_eq_sup_normAtPlace (x : mixedSpace K) : · ext w simp [normAtPlace_apply_isComplex w.prop] +open scoped Classical in theorem norm_eq_sup'_normAtPlace (x : mixedSpace K) : ‖x‖ = univ.sup' univ_nonempty fun w ↦ normAtPlace w x := by rw [← coe_nnnorm, nnnorm_eq_sup_normAtPlace, ← sup'_eq_sup univ_nonempty, ← NNReal.val_eq_coe, @@ -460,8 +461,6 @@ end norm noncomputable section stdBasis -open scoped Classical - open Complex MeasureTheory MeasureTheory.Measure ZSpan Matrix ComplexConjugate variable [NumberField K] @@ -469,6 +468,7 @@ variable [NumberField K] /-- The type indexing the basis `stdBasis`. -/ abbrev index := {w : InfinitePlace K // IsReal w} ⊕ ({w : InfinitePlace K // IsComplex w}) × (Fin 2) +open scoped Classical in /-- The `ℝ`-basis of the mixed space of `K` formed by the vector equal to `1` at `w` and `0` elsewhere for `IsReal w` and by the couple of vectors equal to `1` (resp. `I`) at `w` and `0` elsewhere for `IsComplex w`. -/ @@ -494,6 +494,7 @@ theorem stdBasis_apply_ofIsComplex_snd (x : mixedSpace K) variable (K) +open scoped Classical in theorem fundamentalDomain_stdBasis : fundamentalDomain (stdBasis K) = (Set.univ.pi fun _ => Set.Ico 0 1) ×ˢ @@ -501,6 +502,7 @@ theorem fundamentalDomain_stdBasis : ext simp [stdBasis, mem_fundamentalDomain, Complex.measurableEquivPi] +open scoped Classical in theorem volume_fundamentalDomain_stdBasis : volume (fundamentalDomain (stdBasis K)) = 1 := by rw [fundamentalDomain_stdBasis, volume_eq_prod, prod_prod, volume_pi, volume_pi, pi_pi, pi_pi, @@ -508,6 +510,7 @@ theorem volume_fundamentalDomain_stdBasis : sub_zero, ENNReal.ofReal_one, prod_const_one, prod_const_one, prod_const_one, one_mul] exact (MeasurableSet.pi Set.countable_univ (fun _ _ => measurableSet_Ico)).nullMeasurableSet +open scoped Classical in /-- The `Equiv` between `index K` and `K →+* ℂ` defined by sending a real infinite place `w` to the unique corresponding embedding `w.embedding`, and the pair `⟨w, 0⟩` (resp. `⟨w, 1⟩`) for a complex infinite place `w` to `w.embedding` (resp. `conjugate w.embedding`). -/ @@ -544,6 +547,7 @@ theorem indexEquiv_apply_ofIsComplex_snd (w : {w : InfinitePlace K // IsComplex variable (K) +open scoped Classical in /-- The matrix that gives the representation on `stdBasis` of the image by `commMap` of an element `x` of `(K →+* ℂ) → ℂ` fixed by the map `x_φ ↦ conj x_(conjugate φ)`, see `stdBasis_repr_eq_matrixToStdBasis_mul`. -/ @@ -551,6 +555,7 @@ def matrixToStdBasis : Matrix (index K) (index K) ℂ := fromBlocks (diagonal fun _ => 1) 0 0 <| reindex (Equiv.prodComm _ _) (Equiv.prodComm _ _) (blockDiagonal (fun _ => (2 : ℂ)⁻¹ • !![1, 1; - I, I])) +open scoped Classical in theorem det_matrixToStdBasis : (matrixToStdBasis K).det = (2⁻¹ * I) ^ nrComplexPlaces K := calc @@ -563,6 +568,7 @@ theorem det_matrixToStdBasis : _ = (2⁻¹ * Complex.I) ^ Fintype.card {w : InfinitePlace K // IsComplex w} := by rw [prod_const, Fintype.card] +open scoped Classical in /-- Let `x : (K →+* ℂ) → ℂ` such that `x_φ = conj x_(conj φ)` for all `φ : K →+* ℂ`, then the representation of `commMap K x` on `stdBasis` is given (up to reindexing) by the product of `matrixToStdBasis` by `x`. -/ diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index 0eb57e9a67069c..fe204bdeade808 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -74,8 +74,6 @@ theorem convexBodyLT_convex : Convex ℝ (convexBodyLT K f) := open Fintype MeasureTheory MeasureTheory.Measure ENNReal -open scoped Classical - variable [NumberField K] /-- The fudge factor that appears in the formula for the volume of `convexBodyLT`. -/ @@ -88,6 +86,7 @@ theorem convexBodyLTFactor_ne_zero : convexBodyLTFactor K ≠ 0 := theorem one_le_convexBodyLTFactor : 1 ≤ convexBodyLTFactor K := one_le_mul (one_le_pow₀ one_le_two) (one_le_pow₀ (one_le_two.trans Real.two_le_pi)) +open scoped Classical in /-- The volume of `(ConvexBodyLt K f)` where `convexBodyLT K f` is the set of points `x` such that `‖x w‖ < f w` for all infinite places `w`. -/ theorem convexBodyLT_volume : @@ -122,6 +121,7 @@ but one and choose the value at the remaining place so that we can apply `exists_ne_zero_mem_ringOfIntegers_lt`. -/ theorem adjust_f {w₁ : InfinitePlace K} (B : ℝ≥0) (hf : ∀ w, w ≠ w₁ → f w ≠ 0) : ∃ g : InfinitePlace K → ℝ≥0, (∀ w, w ≠ w₁ → g w = f w) ∧ ∏ w, (g w) ^ mult w = B := by + classical let S := ∏ w ∈ Finset.univ.erase w₁, (f w) ^ mult w refine ⟨Function.update f w₁ ((B * S⁻¹) ^ (mult w₁ : ℝ)⁻¹), ?_, ?_⟩ · exact fun w hw => Function.update_of_ne hw _ f @@ -139,10 +139,9 @@ section convexBodyLT' open Metric ENNReal NNReal -open scoped Classical - variable (f : InfinitePlace K → ℝ≥0) (w₀ : {w : InfinitePlace K // IsComplex w}) +open scoped Classical in /-- A version of `convexBodyLT` with an additional condition at a fixed complex place. This is needed to ensure the element constructed is not real, see for example `exists_primitive_element_lt_of_isComplex`. @@ -191,8 +190,6 @@ theorem convexBodyLT'_convex : Convex ℝ (convexBodyLT' K f w₀) := by open MeasureTheory MeasureTheory.Measure -open scoped Classical - variable [NumberField K] /-- The fudge factor that appears in the formula for the volume of `convexBodyLT'`. -/ @@ -205,6 +202,7 @@ theorem convexBodyLT'Factor_ne_zero : convexBodyLT'Factor K ≠ 0 := theorem one_le_convexBodyLT'Factor : 1 ≤ convexBodyLT'Factor K := one_le_mul (one_le_pow₀ one_le_two) (one_le_pow₀ (one_le_two.trans Real.two_le_pi)) +open scoped Classical in theorem convexBodyLT'_volume : volume (convexBodyLT' K f w₀) = convexBodyLT'Factor K * ∏ w, (f w) ^ (mult w) := by have vol_box : ∀ B : ℝ≥0, volume {x : ℂ | |x.re| < 1 ∧ |x.im| < B^2} = 4*B^2 := by @@ -258,7 +256,7 @@ section convexBodySum open ENNReal MeasureTheory Fintype -open scoped Real Classical NNReal +open scoped Real NNReal variable [NumberField K] (B : ℝ) variable {K} @@ -270,6 +268,7 @@ noncomputable abbrev convexBodySumFun (x : mixedSpace K) : ℝ := ∑ w, mult w theorem convexBodySumFun_apply (x : mixedSpace K) : convexBodySumFun x = ∑ w, mult w * normAtPlace w x := rfl +open scoped Classical in theorem convexBodySumFun_apply' (x : mixedSpace K) : convexBodySumFun x = ∑ w, ‖x.1 w‖ + 2 * ∑ w, ‖x.2 w‖ := by simp_rw [convexBodySumFun_apply, ← Finset.sum_add_sum_compl {w | IsReal w}.toFinset, @@ -310,6 +309,7 @@ theorem convexBodySumFun_eq_zero_iff (x : mixedSpace K) : (mem_nonZeroDivisors_iff_ne_zero.mpr <| Nat.cast_ne_zero.mpr mult_ne_zero)] simp_rw [Finset.mem_univ, true_implies] +open scoped Classical in theorem norm_le_convexBodySumFun (x : mixedSpace K) : ‖x‖ ≤ convexBodySumFun x := by rw [norm_eq_sup'_normAtPlace] refine (Finset.sup'_le_iff _ _).mpr fun w _ ↦ ?_ @@ -333,6 +333,7 @@ theorem convexBodySumFun_continuous : `∑ w real, ‖x w‖ + 2 * ∑ w complex, ‖x w‖ ≤ B`. -/ abbrev convexBodySum : Set (mixedSpace K) := { x | convexBodySumFun x ≤ B } +open scoped Classical in theorem convexBodySum_volume_eq_zero_of_le_zero {B} (hB : B ≤ 0) : volume (convexBodySum K B) = 0 := by obtain hB | hB := lt_or_eq_of_le hB @@ -363,12 +364,14 @@ theorem convexBodySum_convex : Convex ℝ (convexBodySum K B) := by exact (abs_eq_self.mpr h).symm theorem convexBodySum_isBounded : Bornology.IsBounded (convexBodySum K B) := by + classical refine Metric.isBounded_iff.mpr ⟨B + B, fun x hx y hy => ?_⟩ refine le_trans (norm_sub_le x y) (add_le_add ?_ ?_) · exact le_trans (norm_le_convexBodySumFun x) hx · exact le_trans (norm_le_convexBodySumFun y) hy theorem convexBodySum_compact : IsCompact (convexBodySum K B) := by + classical rw [Metric.isCompact_iff_isClosed_bounded] refine ⟨?_, convexBodySum_isBounded K B⟩ convert IsClosed.preimage (convexBodySumFun_continuous K) (isClosed_Icc : IsClosed (Set.Icc 0 B)) @@ -385,6 +388,7 @@ theorem convexBodySumFactor_ne_zero : convexBodySumFactor K ≠ 0 := by (pow_ne_zero _ (div_ne_zero NNReal.pi_ne_zero two_ne_zero)) open MeasureTheory MeasureTheory.Measure Real in +open scoped Classical in theorem convexBodySum_volume : volume (convexBodySum K B) = (convexBodySumFactor K) * (.ofReal B) ^ (finrank ℚ K) := by obtain hB | hB := le_or_lt B 0 @@ -440,13 +444,13 @@ end convexBodySum section minkowski -open scoped Classical open MeasureTheory MeasureTheory.Measure Module ZSpan Real Submodule open scoped ENNReal NNReal nonZeroDivisors IntermediateField variable [NumberField K] (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) +open scoped Classical in /-- The bound that appears in **Minkowski Convex Body theorem**, see `MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure`. See `NumberField.mixedEmbedding.volume_fundamentalDomain_idealLatticeBasis_eq` and @@ -456,6 +460,7 @@ noncomputable def minkowskiBound : ℝ≥0∞ := volume (fundamentalDomain (fractionalIdealLatticeBasis K I)) * (2 : ℝ≥0∞) ^ (finrank ℝ (mixedSpace K)) +open scoped Classical in theorem volume_fundamentalDomain_fractionalIdealLatticeBasis : volume (fundamentalDomain (fractionalIdealLatticeBasis K I)) = .ofReal (FractionalIdeal.absNorm I.1) * volume (fundamentalDomain (latticeBasis K)) := by @@ -471,17 +476,20 @@ theorem volume_fundamentalDomain_fractionalIdealLatticeBasis : rw [mixedEmbedding.det_basisOfFractionalIdeal_eq_norm] theorem minkowskiBound_lt_top : minkowskiBound K I < ⊤ := by + classical refine ENNReal.mul_lt_top ?_ ?_ · exact (fundamentalDomain_isBounded _).measure_lt_top · exact ENNReal.pow_lt_top (lt_top_iff_ne_top.mpr ENNReal.two_ne_top) _ theorem minkowskiBound_pos : 0 < minkowskiBound K I := by + classical refine zero_lt_iff.mpr (mul_ne_zero ?_ ?_) · exact ZSpan.measure_fundamentalDomain_ne_zero _ · exact ENNReal.pow_ne_zero two_ne_zero _ variable {f : InfinitePlace K → ℝ≥0} (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) +open scoped Classical in /-- Let `I` be a fractional ideal of `K`. Assume that `f : InfinitePlace K → ℝ≥0` is such that `minkowskiBound K I < volume (convexBodyLT K f)` where `convexBodyLT K f` is the set of points `x` such that `‖x w‖ < f w` for all infinite places `w` (see `convexBodyLT_volume` for @@ -499,6 +507,7 @@ theorem exists_ne_zero_mem_ideal_lt (h : minkowskiBound K I < volume (convexBody obtain ⟨a, ha, rfl⟩ := hx exact ⟨a, ha, by simpa using h_nz, (convexBodyLT_mem K f).mp h_mem⟩ +open scoped Classical in /-- A version of `exists_ne_zero_mem_ideal_lt` where the absolute value of the real part of `a` is smaller than `1` at some fixed complex place. This is useful to ensure that `a` is not real. -/ theorem exists_ne_zero_mem_ideal_lt' (w₀ : {w : InfinitePlace K // IsComplex w}) @@ -515,6 +524,7 @@ theorem exists_ne_zero_mem_ideal_lt' (w₀ : {w : InfinitePlace K // IsComplex w obtain ⟨a, ha, rfl⟩ := hx exact ⟨a, ha, by simpa using h_nz, (convexBodyLT'_mem K f w₀).mp h_mem⟩ +open scoped Classical in /-- A version of `exists_ne_zero_mem_ideal_lt` for the ring of integers of `K`. -/ theorem exists_ne_zero_mem_ringOfIntegers_lt (h : minkowskiBound K ↑1 < volume (convexBodyLT K f)) : ∃ a : 𝓞 K, a ≠ 0 ∧ ∀ w : InfinitePlace K, w a < f w := by @@ -522,6 +532,7 @@ theorem exists_ne_zero_mem_ringOfIntegers_lt (h : minkowskiBound K ↑1 < volume obtain ⟨a, rfl⟩ := (FractionalIdeal.mem_one_iff _).mp h_mem exact ⟨a, RingOfIntegers.coe_ne_zero_iff.mp h_nz, h_bd⟩ +open scoped Classical in /-- A version of `exists_ne_zero_mem_ideal_lt'` for the ring of integers of `K`. -/ theorem exists_ne_zero_mem_ringOfIntegers_lt' (w₀ : {w : InfinitePlace K // IsComplex w}) (h : minkowskiBound K ↑1 < volume (convexBodyLT' K f w₀)) : @@ -535,6 +546,7 @@ theorem exists_primitive_element_lt_of_isReal {w₀ : InfinitePlace K} (hw₀ : (hB : minkowskiBound K ↑1 < convexBodyLTFactor K * B) : ∃ a : 𝓞 K, ℚ⟮(a : K)⟯ = ⊤ ∧ ∀ w : InfinitePlace K, w a < max B 1 := by + classical have : minkowskiBound K ↑1 < volume (convexBodyLT K (fun w ↦ if w = w₀ then B else 1)) := by rw [convexBodyLT_volume, ← Finset.prod_erase_mul _ _ (Finset.mem_univ w₀)] simp_rw [ite_pow, one_pow] @@ -551,6 +563,7 @@ theorem exists_primitive_element_lt_of_isComplex {w₀ : InfinitePlace K} (hw₀ {B : ℝ≥0} (hB : minkowskiBound K ↑1 < convexBodyLT'Factor K * B) : ∃ a : 𝓞 K, ℚ⟮(a : K)⟯ = ⊤ ∧ ∀ w : InfinitePlace K, w a < Real.sqrt (1 + B ^ 2) := by + classical have : minkowskiBound K ↑1 < volume (convexBodyLT' K (fun w ↦ if w = w₀ then NNReal.sqrt B else 1) ⟨w₀, hw₀⟩) := by rw [convexBodyLT'_volume, ← Finset.prod_erase_mul _ _ (Finset.mem_univ w₀)] @@ -577,6 +590,7 @@ theorem exists_primitive_element_lt_of_isComplex {w₀ : InfinitePlace K} (hw₀ rw [NNReal.coe_one, Real.le_sqrt' zero_lt_one, one_pow] norm_num +open scoped Classical in /-- Let `I` be a fractional ideal of `K`. Assume that `B : ℝ` is such that `minkowskiBound K I < volume (convexBodySum K B)` where `convexBodySum K B` is the set of points `x` such that `∑ w real, ‖x w‖ + 2 * ∑ w complex, ‖x w‖ ≤ B` (see `convexBodySum_volume` for @@ -615,6 +629,7 @@ theorem exists_ne_zero_mem_ideal_of_norm_le {B : ℝ} · rw [← Nat.cast_sum, sum_mult_eq, Nat.cast_pos] exact finrank_pos +open scoped Classical in theorem exists_ne_zero_mem_ringOfIntegers_of_norm_le {B : ℝ} (h : (minkowskiBound K ↑1) ≤ volume (convexBodySum K B)) : ∃ a : 𝓞 K, a ≠ 0 ∧ |Algebra.norm ℚ (a : K)| ≤ (B / finrank ℚ K) ^ finrank ℚ K := by diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index b14ba38413d3d6..d92da2e2f64b27 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -30,7 +30,7 @@ for `x ∈ K`, we have `Π_w ‖x‖_w = |norm(x)|` where the product is over th number field, embeddings, places, infinite places -/ -open scoped Classical Finset +open scoped Finset namespace NumberField.Embeddings @@ -97,6 +97,7 @@ variable (K A) /-- Let `B` be a real number. The set of algebraic integers in `K` whose conjugates are all smaller in norm than `B` is finite. -/ theorem finite_of_norm_le (B : ℝ) : {x : K | IsIntegral ℤ x ∧ ∀ φ : K →+* A, ‖φ x‖ ≤ B}.Finite := by + classical let C := Nat.ceil (max B 1 ^ finrank ℚ K * (finrank ℚ K).choose (finrank ℚ K / 2)) have := bUnion_roots_finite (algebraMap ℤ K) (finrank ℚ K) (finite_Icc (-C : ℤ) C) refine this.subset fun x hx => ?_; simp_rw [mem_iUnion] @@ -422,6 +423,7 @@ lemma isComplex_mk_iff {φ : K →+* ℂ} : theorem not_isReal_of_mk_isComplex {φ : K →+* ℂ} (h : IsComplex (mk φ)) : ¬ ComplexEmbedding.IsReal φ := by rwa [← isComplex_mk_iff] +open scoped Classical in /-- The multiplicity of an infinite place, that is the number of distinct complex embeddings that define it, see `card_filter_mk_eq`. -/ noncomputable def mult (w : InfinitePlace K) : ℕ := if (IsReal w) then 1 else 2 @@ -437,6 +439,7 @@ theorem one_le_mult {w : InfinitePlace K} : (1 : ℝ) ≤ mult w := by rw [← Nat.cast_one, Nat.cast_le] exact mult_pos +open scoped Classical in theorem card_filter_mk_eq [NumberField K] (w : InfinitePlace K) : #{φ | mk φ = w} = mult w := by conv_lhs => congr; congr; ext @@ -450,11 +453,13 @@ theorem card_filter_mk_eq [NumberField K] (w : InfinitePlace K) : #{φ | mk φ = · refine Finset.card_pair ?_ rwa [Ne, eq_comm, ← ComplexEmbedding.isReal_iff, ← isReal_iff] +open scoped Classical in noncomputable instance NumberField.InfinitePlace.fintype [NumberField K] : Fintype (InfinitePlace K) := Set.fintypeRange _ theorem sum_mult_eq [NumberField K] : ∑ w : InfinitePlace K, mult w = Module.finrank ℚ K := by + classical rw [← Embeddings.card K ℂ, Fintype.card, Finset.card_eq_sum_ones, ← Finset.univ.sum_fiberwise (fun φ => InfinitePlace.mk φ)] exact Finset.sum_congr rfl @@ -490,6 +495,7 @@ variable [NumberField K] `‖·‖_w` is the normalized absolute value for `w`. -/ theorem prod_eq_abs_norm (x : K) : ∏ w : InfinitePlace K, w x ^ mult w = abs (Algebra.norm ℚ x) := by + classical convert (congr_arg Complex.abs (@Algebra.norm_eq_prod_embeddings ℚ _ _ _ _ ℂ _ _ _ _ _ x)).symm · rw [map_prod, ← Fintype.prod_equiv RingHom.equivRatAlgHom (fun f => Complex.abs (f x)) (fun φ => Complex.abs (φ x)) fun _ => by simp [RingHom.equivRatAlgHom_apply]; rfl] @@ -555,25 +561,30 @@ section NumberField variable [NumberField K] +open scoped Classical in /-- The number of infinite real places of the number field `K`. -/ noncomputable abbrev nrRealPlaces := card { w : InfinitePlace K // IsReal w } @[deprecated (since := "2024-10-24")] alias NrRealPlaces := nrRealPlaces +open scoped Classical in /-- The number of infinite complex places of the number field `K`. -/ noncomputable abbrev nrComplexPlaces := card { w : InfinitePlace K // IsComplex w } @[deprecated (since := "2024-10-24")] alias NrComplexPlaces := nrComplexPlaces +open scoped Classical in theorem card_real_embeddings : card { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } = nrRealPlaces K := Fintype.card_congr mkReal theorem card_eq_nrRealPlaces_add_nrComplexPlaces : Fintype.card (InfinitePlace K) = nrRealPlaces K + nrComplexPlaces K := by + classical convert Fintype.card_subtype_or_disjoint (IsReal (K := K)) (IsComplex (K := K)) (disjoint_isReal_isComplex K) using 1 exact (Fintype.card_of_subtype _ (fun w ↦ ⟨fun _ ↦ isReal_or_isComplex w, fun _ ↦ by simp⟩)).symm +open scoped Classical in theorem card_complex_embeddings : card { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ } = 2 * nrComplexPlaces K := by suffices ∀ w : { w : InfinitePlace K // IsComplex w }, @@ -593,6 +604,7 @@ theorem card_complex_embeddings : theorem card_add_two_mul_card_eq_rank : nrRealPlaces K + 2 * nrComplexPlaces K = finrank ℚ K := by + classical rw [← card_real_embeddings, ← card_complex_embeddings, Fintype.card_subtype_compl, ← Embeddings.card K ℂ, Nat.add_sub_of_le] exact Fintype.card_subtype_le _ @@ -838,6 +850,7 @@ variable (k w) lemma nat_card_stabilizer_eq_one_or_two : Nat.card (Stab w) = 1 ∨ Nat.card (Stab w) = 2 := by + classical rw [← SetLike.coe_sort_coe, ← mk_embedding w] by_cases h : ∃ σ, ComplexEmbedding.IsConj (k := k) (embedding w) σ · obtain ⟨σ, hσ⟩ := h @@ -869,6 +882,7 @@ lemma not_isUnramified_iff_card_stabilizer_eq_two [IsGalois k K] : rw [isUnramified_iff_card_stabilizer_eq_one] obtain (e|e) := nat_card_stabilizer_eq_one_or_two k w <;> rw [e] <;> decide +open scoped Classical in lemma card_stabilizer [IsGalois k K] : Nat.card (Stab w) = if IsUnramified k w then 1 else 2 := by split @@ -934,6 +948,7 @@ variable (k K) variable [NumberField K] open Finset in +open scoped Classical in lemma card_isUnramified [NumberField k] [IsGalois k K] : #{w : InfinitePlace K | w.IsUnramified k} = #{w : InfinitePlace k | w.IsUnramifiedIn K} * finrank k K := by @@ -956,6 +971,7 @@ lemma card_isUnramified [NumberField k] [IsGalois k K] : · simp [isUnramifiedIn_comap] open Finset in +open scoped Classical in lemma card_isUnramified_compl [NumberField k] [IsGalois k K] : #({w : InfinitePlace K | w.IsUnramified k} : Finset _)ᶜ = #({w : InfinitePlace k | w.IsUnramifiedIn K} : Finset _)ᶜ * (finrank k K / 2) := by @@ -977,6 +993,7 @@ lemma card_isUnramified_compl [NumberField k] [IsGalois k K] : rwa [← isUnramifiedIn_comap] · simp [isUnramifiedIn_comap] +open scoped Classical in lemma card_eq_card_isUnramifiedIn [NumberField k] [IsGalois k K] : Fintype.card (InfinitePlace K) = #{w : InfinitePlace k | w.IsUnramifiedIn K} * finrank k K + @@ -1038,6 +1055,7 @@ open Module in lemma IsUnramifiedAtInfinitePlaces.card_infinitePlace [NumberField k] [NumberField K] [IsGalois k K] [IsUnramifiedAtInfinitePlaces k K] : Fintype.card (InfinitePlace K) = Fintype.card (InfinitePlace k) * finrank k K := by + classical rw [InfinitePlace.card_eq_card_isUnramifiedIn (k := k) (K := K), Finset.filter_true_of_mem, Finset.card_univ, Finset.card_eq_zero.mpr, zero_mul, add_zero] · exact Finset.compl_univ diff --git a/Mathlib/NumberTheory/NumberField/FinitePlaces.lean b/Mathlib/NumberTheory/NumberField/FinitePlaces.lean index d030b4900a54c4..5d5cf658791b0c 100644 --- a/Mathlib/NumberTheory/NumberField/FinitePlaces.lean +++ b/Mathlib/NumberTheory/NumberField/FinitePlaces.lean @@ -25,7 +25,7 @@ into a completion of `K` associated to a non-zero prime ideal of `𝓞 K`. number field, places, finite places -/ -open Ideal IsDedekindDomain HeightOneSpectrum NumberField WithZeroMulInt +open Ideal IsDedekindDomain HeightOneSpectrum WithZeroMulInt namespace NumberField @@ -33,9 +33,8 @@ section absoluteValue variable {K : Type*} [Field K] [NumberField K] (v : HeightOneSpectrum (𝓞 K)) -/-- The norm of a maximal ideal as an element of `ℝ≥0` is `> 1` -/ -lemma one_lt_norm : 1 < (absNorm v.asIdeal : NNReal) := by - norm_cast +/-- The norm of a maximal ideal is `> 1` -/ +lemma one_lt_norm : 1 < absNorm v.asIdeal := by by_contra! h apply IsPrime.ne_top v.isPrime rw [← absNorm_eq_one_iff] @@ -44,7 +43,12 @@ lemma one_lt_norm : 1 < (absNorm v.asIdeal : NNReal) := by exact (v.asIdeal.fintypeQuotientOfFreeOfNeBot v.ne_bot).finite omega -private lemma norm_ne_zero : (absNorm v.asIdeal : NNReal) ≠ 0 := ne_zero_of_lt (one_lt_norm v) +/-- The norm of a maximal ideal as an element of `ℝ≥0` is `> 1` -/ +lemma one_lt_norm_nnreal : 1 < (absNorm v.asIdeal : NNReal) := mod_cast one_lt_norm v + +/-- The norm of a maximal ideal as an element of `ℝ≥0` is `≠ 0` -/ +lemma norm_ne_zero : (absNorm v.asIdeal : NNReal) ≠ 0 := + ne_zero_of_lt (one_lt_norm_nnreal v) /-- The `v`-adic absolute value on `K` defined as the norm of `v` raised to negative `v`-adic valuation.-/ @@ -57,7 +61,7 @@ noncomputable def vadicAbv : AbsoluteValue K ℝ where -- the triangle inequality is implied by the ultrametric one apply le_trans _ <| max_le_add_of_nonneg (zero_le ((toNNReal (norm_ne_zero v)) (v.valuation x))) (zero_le ((toNNReal (norm_ne_zero v)) (v.valuation y))) - have h_mono := (toNNReal_strictMono (one_lt_norm v)).monotone + have h_mono := (toNNReal_strictMono (one_lt_norm_nnreal v)).monotone rw [← h_mono.map_max] --max goes inside withZeroMultIntToNNReal exact h_mono (v.valuation.map_add x y) @@ -80,7 +84,7 @@ noncomputable instance instRankOneValuedAdicCompletion : map_one' := rfl map_mul' := MonoidWithZeroHom.map_mul (toNNReal (norm_ne_zero v)) } - strictMono' := toNNReal_strictMono (one_lt_norm v) + strictMono' := toNNReal_strictMono (one_lt_norm_nnreal v) nontrivial' := by rcases Submodule.exists_mem_ne_zero_of_ne_bot v.ne_bot with ⟨x, hx1, hx2⟩ use (x : K) @@ -133,19 +137,19 @@ open FinitePlace /-- The `v`-adic norm of an integer is at most 1. -/ theorem norm_le_one (x : 𝓞 K) : ‖embedding v x‖ ≤ 1 := by - rw [norm_def', NNReal.coe_le_one, toNNReal_le_one_iff (one_lt_norm v)] + rw [norm_def', NNReal.coe_le_one, toNNReal_le_one_iff (one_lt_norm_nnreal v)] exact valuation_le_one v x /-- The `v`-adic norm of an integer is 1 if and only if it is not in the ideal. -/ theorem norm_eq_one_iff_not_mem (x : 𝓞 K) : ‖(embedding v) x‖ = 1 ↔ x ∉ v.asIdeal := by rw [norm_def_int, NNReal.coe_eq_one, toNNReal_eq_one_iff (v.intValuationDef x) - (norm_ne_zero v) (one_lt_norm v).ne', ← dvd_span_singleton, + (norm_ne_zero v) (one_lt_norm_nnreal v).ne', ← dvd_span_singleton, ← intValuation_lt_one_iff_dvd, not_lt] exact (intValuation_le_one v x).ge_iff_eq.symm /-- The `v`-adic norm of an integer is less than 1 if and only if it is in the ideal. -/ theorem norm_lt_one_iff_mem (x : 𝓞 K) : ‖embedding v x‖ < 1 ↔ x ∈ v.asIdeal := by - rw [norm_def_int, NNReal.coe_lt_one, toNNReal_lt_one_iff (one_lt_norm v), + rw [norm_def_int, NNReal.coe_lt_one, toNNReal_lt_one_iff (one_lt_norm_nnreal v), intValuation_lt_one_iff_dvd] exact dvd_span_singleton @@ -265,9 +269,9 @@ lemma embedding_mul_absNorm (v : HeightOneSpectrum (𝓞 K)) {x : 𝓞 K} (h_x_n WithZeroMulInt.toNNReal_neg_apply _ (v.valuation.ne_zero_iff.mpr (RingOfIntegers.coe_ne_zero_iff.mpr h_x_nezero))] push_cast - rw [← zpow_natCast, ← zpow_add₀ <| mod_cast (zero_lt_one.trans (one_lt_norm v)).ne'] + rw [← zpow_natCast, ← zpow_add₀ <| mod_cast (zero_lt_one.trans (one_lt_norm_nnreal v)).ne'] norm_cast - rw [zpow_eq_one_iff_right₀ (Nat.cast_nonneg' _) (mod_cast (one_lt_norm v).ne')] + rw [zpow_eq_one_iff_right₀ (Nat.cast_nonneg' _) (mod_cast (one_lt_norm_nnreal v).ne')] simp [valuation_eq_intValuationDef, intValuationDef_if_neg, h_x_nezero] end IsDedekindDomain.HeightOneSpectrum diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 1d17123834ee72..88da50e57f12cc 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -400,9 +400,7 @@ theorem logEmbeddingQuot_injective : Function.comp_apply, EmbeddingLike.apply_eq_iff_eq] at h exact (EmbeddingLike.apply_eq_iff_eq _).mp <| (QuotientGroup.kerLift_injective _).eq_iff.mp h -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/4119 +#adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 the `Module ℤ (Additive ((𝓞 K)ˣ ⧸ NumberField.Units.torsion K))` instance required below isn't found unless we use `set_option maxSynthPendingDepth 2`, or add explicit instances: diff --git a/Mathlib/NumberTheory/NumberField/Units/Regulator.lean b/Mathlib/NumberTheory/NumberField/Units/Regulator.lean index cd7cb76b5bf075..31347cb3073d0e 100644 --- a/Mathlib/NumberTheory/NumberField/Units/Regulator.lean +++ b/Mathlib/NumberTheory/NumberField/Units/Regulator.lean @@ -32,21 +32,22 @@ namespace NumberField.Units variable (K : Type*) [Field K] -open MeasureTheory Classical NumberField.InfinitePlace +open MeasureTheory NumberField.InfinitePlace NumberField NumberField.Units.dirichletUnitTheorem variable [NumberField K] +open scoped Classical in /-- The regulator of a number field `K`. -/ def regulator : ℝ := ZLattice.covolume (unitLattice K) +open scoped Classical in theorem regulator_ne_zero : regulator K ≠ 0 := ZLattice.covolume_ne_zero (unitLattice K) volume +open scoped Classical in theorem regulator_pos : 0 < regulator K := ZLattice.covolume_pos (unitLattice K) volume -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/4119 +#adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 the `Module ℤ (Additive ((𝓞 K)ˣ ⧸ NumberField.Units.torsion K))` instance required below isn't found unless we use `set_option maxSynthPendingDepth 2`, or add explicit instances: @@ -56,6 +57,7 @@ local instance : CommGroup (𝓞 K)ˣ := inferInstance -/ set_option maxSynthPendingDepth 2 -- Note this is active for the remainder of the file. +open scoped Classical in theorem regulator_eq_det' (e : {w : InfinitePlace K // w ≠ w₀} ≃ Fin (rank K)) : regulator K = |(Matrix.of fun i ↦ logEmbedding K (Additive.ofMul (fundSystem K (e i)))).det| := by @@ -63,6 +65,7 @@ theorem regulator_eq_det' (e : {w : InfinitePlace K // w ≠ w₀} ≃ Fin (rank (((basisModTorsion K).map (logEmbeddingEquiv K)).reindex e.symm), Basis.coe_reindex, Function.comp_def, Basis.map_apply, ← fundSystem_mk, Equiv.symm_symm, logEmbeddingEquiv_apply] +open scoped Classical in /-- Let `u : Fin (rank K) → (𝓞 K)ˣ` be a family of units and let `w₁` and `w₂` be two infinite places. Then, the two square matrices with entries `(mult w * log w (u i))_i, {w ≠ w_i}`, `i = 1,2`, have the same determinant in absolute value. -/ @@ -98,6 +101,7 @@ theorem abs_det_eq_abs_det (u : Fin (rank K) → (𝓞 K)ˣ) Units.norm, Rat.cast_one, Real.log_one] exact fun _ _ ↦ pow_ne_zero _ <| (map_ne_zero _).mpr (coe_ne_zero _) +open scoped Classical in /-- For any infinite place `w'`, the regulator is equal to the absolute value of the determinant of the matrix `(mult w * log w (fundSystem K i)))_i, {w ≠ w'}`. -/ theorem regulator_eq_det (w' : InfinitePlace K) (e : {w // w ≠ w'} ≃ Fin (rank K)) : diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index f16855f13878b8..f6b9067b3e3c91 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -41,7 +41,6 @@ The constructions of the ring homomorphisms go through an auxiliary constructor noncomputable section -open scoped Classical open Nat IsLocalRing Padic namespace PadicInt @@ -293,6 +292,7 @@ def residueField : IsLocalRing.ResidueField ℤ_[p] ≃+* ZMod p := by exact_mod_cast (@PadicInt.ker_toZMod p _) ▸ RingHom.quotientKerEquivOfSurjective (ZMod.ringHom_surjective PadicInt.toZMod) +open scoped Classical in /-- `appr n x` gives a value `v : ℕ` such that `x` and `↑v : ℤ_p` are congruent mod `p^n`. See `appr_spec`. -/ -- Porting note: removing irreducible solves a lot of problems diff --git a/Mathlib/NumberTheory/RamificationInertia/Basic.lean b/Mathlib/NumberTheory/RamificationInertia/Basic.lean index 24f96fd9081fbb..137fee8752dbb4 100644 --- a/Mathlib/NumberTheory/RamificationInertia/Basic.lean +++ b/Mathlib/NumberTheory/RamificationInertia/Basic.lean @@ -675,25 +675,27 @@ end FactLeComap section FactorsMap -open scoped Classical - /-! ## Properties of the factors of `p.map (algebraMap R S)` -/ variable [IsDedekindDomain S] [Algebra R S] +open scoped Classical in theorem Factors.ne_bot (P : (factors (map (algebraMap R S) p)).toFinset) : (P : Ideal S) ≠ ⊥ := (prime_of_factor _ (Multiset.mem_toFinset.mp P.2)).ne_zero +open scoped Classical in instance Factors.isPrime (P : (factors (map (algebraMap R S) p)).toFinset) : IsPrime (P : Ideal S) := Ideal.isPrime_of_prime (prime_of_factor _ (Multiset.mem_toFinset.mp P.2)) +open scoped Classical in theorem Factors.ramificationIdx_ne_zero (P : (factors (map (algebraMap R S) p)).toFinset) : ramificationIdx (algebraMap R S) p P ≠ 0 := IsDedekindDomain.ramificationIdx_ne_zero (ne_zero_of_mem_factors (Multiset.mem_toFinset.mp P.2)) (Factors.isPrime p P) (Ideal.le_of_dvd (dvd_of_mem_factors (Multiset.mem_toFinset.mp P.2))) +open scoped Classical in instance Factors.fact_ramificationIdx_neZero (P : (factors (map (algebraMap R S) p)).toFinset) : NeZero (ramificationIdx (algebraMap R S) p P) := ⟨Factors.ramificationIdx_ne_zero p P⟩ @@ -702,14 +704,17 @@ set_option synthInstance.checkSynthOrder false -- Porting note: this is okay since, as noted above, in this file the value of `f` can be inferred attribute [local instance] Quotient.algebraQuotientOfRamificationIdxNeZero +open scoped Classical in instance Factors.isScalarTower (P : (factors (map (algebraMap R S) p)).toFinset) : IsScalarTower R (R ⧸ p) (S ⧸ (P : Ideal S)) := IsScalarTower.of_algebraMap_eq' rfl +open scoped Classical in instance Factors.liesOver [p.IsMaximal] (P : (factors (map (algebraMap R S) p)).toFinset) : P.1.LiesOver p := ⟨(comap_eq_of_scalar_tower_quotient (algebraMap (R ⧸ p) (S ⧸ P.1)).injective).symm⟩ +open scoped Classical in theorem Factors.finrank_pow_ramificationIdx [p.IsMaximal] (P : (factors (map (algebraMap R S) p)).toFinset) : finrank (R ⧸ p) (S ⧸ (P : Ideal S) ^ ramificationIdx (algebraMap R S) p P) = @@ -717,6 +722,7 @@ theorem Factors.finrank_pow_ramificationIdx [p.IsMaximal] rw [finrank_prime_pow_ramificationIdx, inertiaDeg_algebraMap] exacts [Factors.ne_bot p P, NeZero.ne _] +open scoped Classical in instance Factors.finiteDimensional_quotient_pow [Module.Finite R S] [p.IsMaximal] (P : (factors (map (algebraMap R S) p)).toFinset) : FiniteDimensional (R ⧸ p) (S ⧸ (P : Ideal S) ^ ramificationIdx (algebraMap R S) p P) := by @@ -726,6 +732,7 @@ instance Factors.finiteDimensional_quotient_pow [Module.Finite R S] [p.IsMaximal universe w +open scoped Classical in /-- **Chinese remainder theorem** for a ring of integers: if the prime ideal `p : Ideal R` factors in `S` as `∏ i, P i ^ e i`, then `S ⧸ I` factors as `Π i, R ⧸ (P i ^ e i)`. -/ noncomputable def Factors.piQuotientEquiv (p : Ideal R) (hp : map (algebraMap R S) p ≠ ⊥) : @@ -752,6 +759,7 @@ theorem Factors.piQuotientEquiv_map (p : Ideal R) (hp : map (algebraMap R S) p variable (S) +open scoped Classical in /-- **Chinese remainder theorem** for a ring of integers: if the prime ideal `p : Ideal R` factors in `S` as `∏ i, P i ^ e i`, then `S ⧸ I` factors `R ⧸ I`-linearly as `Π i, R ⧸ (P i ^ e i)`. -/ @@ -767,6 +775,7 @@ noncomputable def Factors.piQuotientLinearEquiv (p : Ideal R) (hp : map (algebra RingHomCompTriple.comp_apply, Pi.mul_apply, Pi.algebraMap_apply] congr } +open scoped Classical in /-- The **fundamental identity** of ramification index `e` and inertia degree `f`: for `P` ranging over the primes lying over `p`, `∑ P, e P * f P = [Frac(S) : Frac(R)]`; here `S` is a finite `R`-module (and thus `Frac(S) : Frac(R)` is a finite extension) and `p` diff --git a/Mathlib/Order/Antisymmetrization.lean b/Mathlib/Order/Antisymmetrization.lean index a64b7b3da17a81..1ec2f0a7747118 100644 --- a/Mathlib/Order/Antisymmetrization.lean +++ b/Mathlib/Order/Antisymmetrization.lean @@ -40,7 +40,7 @@ def AntisymmRel (a b : α) : Prop := r a b ∧ r b a theorem antisymmRel_swap : AntisymmRel (swap r) = AntisymmRel r := - funext fun _ => funext fun _ => propext and_comm + funext₂ fun _ _ ↦ propext and_comm @[refl] theorem antisymmRel_refl [IsRefl α r] (a : α) : AntisymmRel r a a := diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 6b4dfb21914633..11f4a94f61cf10 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -1263,32 +1263,47 @@ instance [∀ i, Preorder (π i)] [∀ i, DenselyOrdered (π i)] : ⟨le_update_iff.2 ⟨ha.le, fun _ _ ↦ le_rfl⟩, i, by rwa [update_self]⟩, update_le_iff.2 ⟨hb.le, fun _ _ ↦ hab _⟩, i, by rwa [update_self]⟩⟩ -theorem le_of_forall_le_of_dense [LinearOrder α] [DenselyOrdered α] {a₁ a₂ : α} - (h : ∀ a, a₂ < a → a₁ ≤ a) : a₁ ≤ a₂ := +section LinearOrder +variable [LinearOrder α] [DenselyOrdered α] {a₁ a₂ : α} + +theorem le_of_forall_gt_imp_ge_of_dense (h : ∀ a, a₂ < a → a₁ ≤ a) : a₁ ≤ a₂ := le_of_not_gt fun ha ↦ let ⟨a, ha₁, ha₂⟩ := exists_between ha lt_irrefl a <| lt_of_lt_of_le ‹a < a₁› (h _ ‹a₂ < a›) -theorem eq_of_le_of_forall_le_of_dense [LinearOrder α] [DenselyOrdered α] {a₁ a₂ : α} (h₁ : a₂ ≤ a₁) - (h₂ : ∀ a, a₂ < a → a₁ ≤ a) : a₁ = a₂ := - le_antisymm (le_of_forall_le_of_dense h₂) h₁ +lemma forall_gt_imp_ge_iff_le_of_dense : (∀ a, a₂ < a → a₁ ≤ a) ↔ a₁ ≤ a₂ := + ⟨le_of_forall_gt_imp_ge_of_dense, fun ha _a ha₂ ↦ ha.trans ha₂.le⟩ + +lemma eq_of_le_of_forall_lt_imp_le_of_dense (h₁ : a₂ ≤ a₁) (h₂ : ∀ a, a₂ < a → a₁ ≤ a) : a₁ = a₂ := + le_antisymm (le_of_forall_gt_imp_ge_of_dense h₂) h₁ -theorem le_of_forall_ge_of_dense [LinearOrder α] [DenselyOrdered α] {a₁ a₂ : α} - (h : ∀ a₃ < a₁, a₃ ≤ a₂) : a₁ ≤ a₂ := +theorem le_of_forall_lt_imp_le_of_dense (h : ∀ a < a₁, a ≤ a₂) : a₁ ≤ a₂ := le_of_not_gt fun ha ↦ let ⟨a, ha₁, ha₂⟩ := exists_between ha lt_irrefl a <| lt_of_le_of_lt (h _ ‹a < a₁›) ‹a₂ < a› -theorem eq_of_le_of_forall_ge_of_dense [LinearOrder α] [DenselyOrdered α] {a₁ a₂ : α} (h₁ : a₂ ≤ a₁) - (h₂ : ∀ a₃ < a₁, a₃ ≤ a₂) : a₁ = a₂ := - (le_of_forall_ge_of_dense h₂).antisymm h₁ +lemma forall_lt_imp_le_iff_le_of_dense : (∀ a < a₁, a ≤ a₂) ↔ a₁ ≤ a₂ := + ⟨le_of_forall_lt_imp_le_of_dense, fun ha _a ha₁ ↦ ha₁.le.trans ha⟩ + +theorem eq_of_le_of_forall_gt_imp_ge_of_dense (h₁ : a₂ ≤ a₁) (h₂ : ∀ a < a₁, a ≤ a₂) : a₁ = a₂ := + (le_of_forall_lt_imp_le_of_dense h₂).antisymm h₁ + +@[deprecated (since := "2025-01-21")] +alias le_of_forall_le_of_dense := le_of_forall_gt_imp_ge_of_dense + +@[deprecated (since := "2025-01-21")] +alias le_of_forall_ge_of_dense := le_of_forall_lt_imp_le_of_dense + +@[deprecated (since := "2025-01-21")] alias forall_lt_le_iff := forall_lt_imp_le_iff_le_of_dense +@[deprecated (since := "2025-01-21")] alias forall_gt_ge_iff := forall_gt_imp_ge_iff_le_of_dense + +@[deprecated (since := "2025-01-21")] +alias eq_of_le_of_forall_le_of_dense := eq_of_le_of_forall_lt_imp_le_of_dense -theorem forall_lt_le_iff [LinearOrder α] [DenselyOrdered α] {a b : α} : (∀ c < a, c ≤ b) ↔ a ≤ b := - ⟨le_of_forall_ge_of_dense, fun hab _c hca ↦ hca.le.trans hab⟩ +@[deprecated (since := "2025-01-21")] +alias eq_of_le_of_forall_ge_of_dense := eq_of_le_of_forall_gt_imp_ge_of_dense -theorem forall_gt_ge_iff [LinearOrder α] [DenselyOrdered α] {a b : α} : - (∀ c, a < c → b ≤ c) ↔ b ≤ a := - forall_lt_le_iff (α := αᵒᵈ) +end LinearOrder theorem dense_or_discrete [LinearOrder α] (a₁ a₂ : α) : (∃ a, a₁ < a ∧ a < a₂) ∨ (∀ a, a₁ < a → a₂ ≤ a) ∧ ∀ a < a₂, a ≤ a₁ := diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index ff8f494fcda960..a161ef7fab20c1 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -458,7 +458,7 @@ theorem exists_glb_Ioi (i : γ) : ∃ j, IsGLB (Ioi i) j := variable [DenselyOrdered γ] theorem isLUB_Iio {a : γ} : IsLUB (Iio a) a := - ⟨fun _ hx => le_of_lt hx, fun _ hy => le_of_forall_ge_of_dense hy⟩ + ⟨fun _ hx => le_of_lt hx, fun _ hy => le_of_forall_lt_imp_le_of_dense hy⟩ theorem isGLB_Ioi {a : γ} : IsGLB (Ioi a) a := @isLUB_Iio γᵒᵈ _ _ a diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index 6dfb80d1a72f2e..5e469fe3ff642b 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -1739,6 +1739,18 @@ theorem disjoint_sSup_right {a : Set α} {b : α} (d : Disjoint b (sSup a)) {i} Disjoint b i := disjoint_iff_inf_le.mpr (iSup₂_le_iff.mp (iSup_inf_le_inf_sSup.trans d.le_bot) i hi :) +lemma disjoint_of_sSup_disjoint_of_le_of_le {a b : α} {c d : Set α} (hs : ∀ e ∈ c, e ≤ a) + (ht : ∀ e ∈ d, e ≤ b) (hd : Disjoint a b) (he : ⊥ ∉ c ∨ ⊥ ∉ d) : Disjoint c d := by + rw [disjoint_iff_forall_ne] + intros x hx y hy + rw [Disjoint.ne_iff] + · aesop + · exact Disjoint.mono (hs x hx) (ht y hy) hd + +lemma disjoint_of_sSup_disjoint {a b : Set α} (hd : Disjoint (sSup a) (sSup b)) + (he : ⊥ ∉ a ∨ ⊥ ∉ b) : Disjoint a b := + disjoint_of_sSup_disjoint_of_le_of_le (fun _ hc ↦ le_sSup hc) (fun _ hc ↦ le_sSup hc) hd he + end CompleteLattice -- See note [reducible non-instances] diff --git a/Mathlib/Order/CompleteLattice/SetLike.lean b/Mathlib/Order/CompleteLattice/SetLike.lean new file mode 100644 index 00000000000000..2e5426d136ae59 --- /dev/null +++ b/Mathlib/Order/CompleteLattice/SetLike.lean @@ -0,0 +1,56 @@ +/- +Copyright (c) 2024 Sven Manthe. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sven Manthe +-/ +import Mathlib.Order.CompleteSublattice + +/-! +# `SetLike` instance for elements of `CompleteSublattice (Set X)` + +This file provides lemmas for the `SetLike` instance for elements of `CompleteSublattice (Set X)` +-/ + +attribute [local instance] SetLike.instSubtypeSet + +namespace Sublattice + +variable {X : Type*} {L : Sublattice (Set X)} + +variable {S T : L} {x : X} + +@[ext] lemma ext_mem (h : ∀ x, x ∈ S ↔ x ∈ T) : S = T := SetLike.ext h + +lemma mem_subtype : x ∈ L.subtype T ↔ x ∈ T := Iff.rfl + +@[simp] lemma setLike_mem_inf : x ∈ S ⊓ T ↔ x ∈ S ∧ x ∈ T := by simp [← mem_subtype] +@[simp] lemma setLike_mem_sup : x ∈ S ⊔ T ↔ x ∈ S ∨ x ∈ T := by simp [← mem_subtype] + +@[simp] lemma setLike_mem_coe : x ∈ T.val ↔ x ∈ T := Iff.rfl +@[simp] lemma setLike_mem_mk (U : Set X) (h : U ∈ L) : x ∈ (⟨U, h⟩ : L) ↔ x ∈ U := Iff.rfl + +end Sublattice + +namespace CompleteSublattice + +variable {X : Type*} {L : CompleteSublattice (Set X)} + +variable {S T : L} {𝒮 : Set L} {I : Sort*} {f : I → L} {x : X} + +@[ext] lemma ext (h : ∀ x, x ∈ S ↔ x ∈ T) : S = T := SetLike.ext h + +lemma mem_subtype : x ∈ L.subtype T ↔ x ∈ T := Iff.rfl + +@[simp] lemma mem_inf : x ∈ S ⊓ T ↔ x ∈ S ∧ x ∈ T := by simp [← mem_subtype] +@[simp] lemma mem_sInf : x ∈ sInf 𝒮 ↔ ∀ T ∈ 𝒮, x ∈ T := by simp [← mem_subtype] +@[simp] lemma mem_iInf : x ∈ ⨅ i : I, f i ↔ ∀ i : I, x ∈ f i := by simp [← mem_subtype] +@[simp] lemma mem_top : x ∈ (⊤ : L) := by simp [← mem_subtype] + +@[simp] lemma mem_sup : x ∈ S ⊔ T ↔ x ∈ S ∨ x ∈ T := by simp [← mem_subtype] +@[simp] lemma mem_sSup : x ∈ sSup 𝒮 ↔ ∃ T ∈ 𝒮, x ∈ T := by simp [← mem_subtype] +@[simp] lemma mem_iSup : x ∈ ⨆ i : I, f i ↔ ∃ i : I, x ∈ f i := by simp [← mem_subtype] +@[simp] lemma not_mem_bot : ¬ x ∈ (⊥ : L) := by simp [← mem_subtype] + +@[simp] lemma mem_mk (U : Set X) (h : U ∈ L) : x ∈ (⟨U, h⟩ : L) ↔ x ∈ U := Iff.rfl + +end CompleteSublattice diff --git a/Mathlib/Order/DirectedInverseSystem.lean b/Mathlib/Order/DirectedInverseSystem.lean index 943b33a61c2285..fe714d8d388630 100644 --- a/Mathlib/Order/DirectedInverseSystem.lean +++ b/Mathlib/Order/DirectedInverseSystem.lean @@ -144,15 +144,13 @@ theorem mk_injective (h : ∀ i j hij, Function.Injective (f i j hij)) (i) : section map₀ -open Classical (arbitrary) - variable [Nonempty ι] (ih : ∀ i, F i) /-- "Nullary map" to construct an element in the direct limit. -/ -noncomputable def map₀ : DirectLimit F f := ⟦⟨arbitrary ι, ih _⟩⟧ +noncomputable def map₀ : DirectLimit F f := ⟦⟨Classical.arbitrary ι, ih _⟩⟧ theorem map₀_def (compat : ∀ i j h, f i j h (ih i) = ih j) (i) : map₀ f ih = ⟦⟨i, ih i⟩⟧ := - have ⟨j, hcj, hij⟩ := exists_ge_ge (arbitrary ι) i + have ⟨j, hcj, hij⟩ := exists_ge_ge (Classical.arbitrary ι) i Quotient.sound ⟨j, hcj, hij, (compat ..).trans (compat ..).symm⟩ end map₀ diff --git a/Mathlib/Order/Filter/SmallSets.lean b/Mathlib/Order/Filter/SmallSets.lean index 1e40a2da092eb3..28c02d2e8dc150 100644 --- a/Mathlib/Order/Filter/SmallSets.lean +++ b/Mathlib/Order/Filter/SmallSets.lean @@ -54,6 +54,21 @@ theorem hasBasis_smallSets (l : Filter α) : HasBasis l.smallSets (fun t : Set α => t ∈ l) powerset := l.basis_sets.smallSets +theorem Eventually.exists_mem_basis_of_smallSets {p : ι → Prop} {s : ι → Set α} {P : Set α → Prop} + (h₁ : ∀ᶠ t in l.smallSets, P t) (h₂ : HasBasis l p s) : ∃ i, p i ∧ P (s i) := + (h₂.smallSets.eventually_iff.mp h₁).imp fun _i ⟨hpi, hi⟩ ↦ ⟨hpi, hi Subset.rfl⟩ + +theorem Frequently.smallSets_of_forall_mem_basis {p : ι → Prop} {s : ι → Set α} {P : Set α → Prop} + (h₁ : ∀ i, p i → P (s i)) (h₂ : HasBasis l p s) : ∃ᶠ t in l.smallSets, P t := + h₂.smallSets.frequently_iff.mpr fun _ hi => ⟨_, Subset.rfl, h₁ _ hi⟩ + +theorem Eventually.exists_mem_of_smallSets {p : Set α → Prop} + (h : ∀ᶠ t in l.smallSets, p t) : ∃ s ∈ l, p s := + h.exists_mem_basis_of_smallSets l.basis_sets + +/-! No `Frequently.smallSets_of_forall_mem (h : ∀ s ∈ l, p s) : ∃ᶠ t in l.smallSets, p t` as +`Filter.frequently_smallSets_mem : ∃ᶠ t in l.smallSets, t ∈ l` is preferred. -/ + /-- `g` converges to `f.smallSets` if for all `s ∈ f`, eventually we have `g x ⊆ s`. -/ theorem tendsto_smallSets_iff {f : α → Set β} : Tendsto f la lb.smallSets ↔ ∀ t ∈ lb, ∀ᶠ x in la, f x ⊆ t := diff --git a/Mathlib/Order/Fin/Basic.lean b/Mathlib/Order/Fin/Basic.lean index 174adbcde047ab..8f534eba16e5a9 100644 --- a/Mathlib/Order/Fin/Basic.lean +++ b/Mathlib/Order/Fin/Basic.lean @@ -71,7 +71,7 @@ instance instLattice : Lattice (Fin n) := inferInstance lemma top_eq_last (n : ℕ) : ⊤ = Fin.last n := rfl -lemma bot_eq_zero (n : ℕ) : ⊥ = (0 : Fin (n + 1)) := rfl +lemma bot_eq_zero (n : ℕ) [NeZero n] : ⊥ = (0 : Fin n) := rfl @[simp] theorem rev_bot [NeZero n] : rev (⊥ : Fin n) = ⊤ := rfl @[simp] theorem rev_top [NeZero n] : rev (⊤ : Fin n) = ⊥ := rev_rev _ @@ -79,6 +79,26 @@ lemma bot_eq_zero (n : ℕ) : ⊥ = (0 : Fin (n + 1)) := rfl theorem rev_zero_eq_top (n : ℕ) [NeZero n] : rev (0 : Fin n) = ⊤ := rfl theorem rev_last_eq_bot (n : ℕ) : rev (last n) = ⊥ := by rw [rev_last, bot_eq_zero] +@[simp] +theorem succ_top (n : ℕ) [NeZero n] : (⊤ : Fin n).succ = ⊤ := by + rw [← rev_zero_eq_top, ← rev_zero_eq_top, ← rev_castSucc, castSucc_zero'] + +@[simp] +theorem val_top (n : ℕ) [NeZero n] : ((⊤ : Fin n) : ℕ) = n - 1 := rfl + +@[simp] +theorem zero_eq_top {n : ℕ} [NeZero n] : (0 : Fin n) = ⊤ ↔ n = 1 := by + rw [← bot_eq_zero, subsingleton_iff_bot_eq_top, subsingleton_iff_le_one, LE.le.le_iff_eq] + exact pos_of_neZero n + +@[simp] +theorem top_eq_zero {n : ℕ} [NeZero n] : (⊤ : Fin n) = 0 ↔ n = 1 := + eq_comm.trans zero_eq_top + +@[simp] +theorem cast_top {m n : ℕ} [NeZero m] [NeZero n] (h : m = n) : (⊤ : Fin m).cast h = ⊤ := by + simp [← val_inj, h] + section ToFin variable {α : Type*} [Preorder α] {f : α → Fin (n + 1)} @@ -117,6 +137,16 @@ lemma strictAnti_iff_succ_lt : StrictAnti f ↔ ∀ i : Fin n, f i.succ < f (cas lemma antitone_iff_succ_le : Antitone f ↔ ∀ i : Fin n, f i.succ ≤ f (castSucc i) := antitone_iff_forall_lt.trans <| liftFun_iff_succ (· ≥ ·) +lemma orderHom_injective_iff {α : Type*} [PartialOrder α] {n : ℕ} (f : Fin (n + 1) →o α) : + Function.Injective f ↔ ∀ (i : Fin n), f i.castSucc ≠ f i.succ := by + constructor + · intro hf i hi + have := hf hi + simp [Fin.ext_iff] at this + · intro hf + refine (strictMono_iff_lt_succ (f := f).2 fun i ↦ ?_).injective + exact lt_of_le_of_ne (f.monotone (Fin.castSucc_le_succ i)) (hf i) + end FromFin /-! #### Monotonicity -/ @@ -225,6 +255,8 @@ def succOrderEmb (n : ℕ) : Fin n ↪o Fin (n + 1) := .ofStrictMono succ strict @[simp, norm_cast] lemma coe_succOrderEmb : ⇑(succOrderEmb n) = Fin.succ := rfl +@[simp] lemma succOrderEmb_toEmbedding : (succOrderEmb n).toEmbedding = succEmb n := rfl + /-- `Fin.castLE` as an `OrderEmbedding`. `castLEEmb h i` embeds `i` into a larger `Fin` type. -/ diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index d446eb5f357437..56f30d50f82363 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -65,6 +65,9 @@ lemma range_eq_Icc_zero_sub_one (n : ℕ) (hn : n ≠ 0) : range n = Icc 0 (n - theorem _root_.Finset.range_eq_Ico : range = Ico 0 := Ico_zero_eq_range.symm +theorem range_succ_eq_Icc_zero (n : ℕ) : range (n + 1) = Icc 0 n := by + rw [range_eq_Icc_zero_sub_one _ (Nat.add_one_ne_zero _), Nat.add_sub_cancel_right] + @[simp] lemma card_Icc : #(Icc a b) = b + 1 - a := List.length_range' .. @[simp] lemma card_Ico : #(Ico a b) = b - a := List.length_range' .. @[simp] lemma card_Ioc : #(Ioc a b) = b - a := List.length_range' .. diff --git a/Mathlib/Order/RelIso/Set.lean b/Mathlib/Order/RelIso/Set.lean index d52b7837ed95bc..90cb227e3edb90 100644 --- a/Mathlib/Order/RelIso/Set.lean +++ b/Mathlib/Order/RelIso/Set.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Order.Directed import Mathlib.Order.RelIso.Basic import Mathlib.Logic.Embedding.Set import Mathlib.Logic.Equiv.Set @@ -35,6 +36,14 @@ theorem map_sup [SemilatticeSup α] [LinearOrder β] [FunLike F β α] a (m ⊔ n) = a m ⊔ a n := map_inf (α := αᵒᵈ) (β := βᵒᵈ) _ _ _ +theorem directed [FunLike F α β] [RelHomClass F r s] {ι : Sort*} {a : ι → α} {f : F} + (ha : Directed r a) : Directed s (f ∘ a) := + ha.mono_comp _ fun _ _ h ↦ map_rel f h + +theorem directedOn [FunLike F α β] [RelHomClass F r s] {f : F} + {t : Set α} (hs : DirectedOn r t) : DirectedOn s (f '' t) := + hs.mono_comp fun _ _ h ↦ map_rel f h + end RelHomClass namespace RelIso diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index 68a3a6a1f2487f..4895b49790087a 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -5,7 +5,6 @@ Authors: Aaron Anderson, Kevin Buzzard, Yaël Dillies, Eric Wieser -/ import Mathlib.Data.Finset.Sigma import Mathlib.Data.Finset.Pairwise -import Mathlib.Data.Finset.Powerset import Mathlib.Data.Fintype.Basic import Mathlib.Order.CompleteLatticeIntervals @@ -58,13 +57,27 @@ def SupIndep (s : Finset ι) (f : ι → α) : Prop := variable {s t : Finset ι} {f : ι → α} {i : ι} -instance [DecidableEq ι] [DecidableEq α] : Decidable (SupIndep s f) := by - refine @Finset.decidableForallOfDecidableSubsets _ _ _ (?_) - rintro t - - refine @Finset.decidableDforallFinset _ _ _ (?_) - rintro i - - have : Decidable (Disjoint (f i) (sup t f)) := decidable_of_iff' (_ = ⊥) disjoint_iff - infer_instance +/-- The RHS looks like the definition of `iSupIndep`. -/ +theorem supIndep_iff_disjoint_erase [DecidableEq ι] : + s.SupIndep f ↔ ∀ i ∈ s, Disjoint (f i) ((s.erase i).sup f) := + ⟨fun hs _ hi => hs (erase_subset _ _) hi (not_mem_erase _ _), fun hs _ ht i hi hit => + (hs i hi).mono_right (sup_mono fun _ hj => mem_erase.2 ⟨ne_of_mem_of_not_mem hj hit, ht hj⟩)⟩ + +/-- If both the index type and the lattice have decidable equality, +then the `SupIndep` predicate is decidable. + +TODO: speedup the definition and drop the `[DecidableEq ι]` assumption +by iterating over the pairs `(a, t)` such that `s = Finset.cons a t _` +using something like `List.eraseIdx` +or by generating both `f i` and `(s.erase i).sup f` in one loop over `s`. +Yet another possible optimization is to precompute partial suprema of `f` +over the inits and tails of the list representing `s`, +store them in 2 `Array`s, +then compute each `sup` in 1 operation. -/ +instance [DecidableEq ι] [DecidableEq α] : Decidable (SupIndep s f) := + have : ∀ i, Decidable (Disjoint (f i) ((s.erase i).sup f)) := fun _ ↦ + decidable_of_iff _ disjoint_iff.symm + decidable_of_iff _ supIndep_iff_disjoint_erase.symm theorem SupIndep.subset (ht : t.SupIndep f) (h : s ⊆ t) : s.SupIndep f := fun _ hu _ hi => ht (hu.trans h) (h hi) @@ -73,6 +86,7 @@ theorem SupIndep.subset (ht : t.SupIndep f) (h : s ⊆ t) : s.SupIndep f := fun theorem supIndep_empty (f : ι → α) : (∅ : Finset ι).SupIndep f := fun _ _ a ha => (not_mem_empty a ha).elim +@[simp] theorem supIndep_singleton (i : ι) (f : ι → α) : ({i} : Finset ι).SupIndep f := fun s hs j hji hj => by rw [eq_empty_of_ssubset_singleton ⟨hs, fun h => hj (h hji)⟩, sup_empty] @@ -82,53 +96,28 @@ theorem SupIndep.pairwiseDisjoint (hs : s.SupIndep f) : (s : Set ι).PairwiseDis fun _ ha _ hb hab => sup_singleton.subst <| hs (singleton_subset_iff.2 hb) ha <| not_mem_singleton.2 hab +@[deprecated (since := "2025-01-17")] alias sup_indep.pairwise_disjoint := SupIndep.pairwiseDisjoint + theorem SupIndep.le_sup_iff (hs : s.SupIndep f) (hts : t ⊆ s) (hi : i ∈ s) (hf : ∀ i, f i ≠ ⊥) : f i ≤ t.sup f ↔ i ∈ t := by refine ⟨fun h => ?_, le_sup⟩ by_contra hit exact hf i (disjoint_self.1 <| (hs hts hi hit).mono_right h) -/-- The RHS looks like the definition of `iSupIndep`. -/ -theorem supIndep_iff_disjoint_erase [DecidableEq ι] : - s.SupIndep f ↔ ∀ i ∈ s, Disjoint (f i) ((s.erase i).sup f) := - ⟨fun hs _ hi => hs (erase_subset _ _) hi (not_mem_erase _ _), fun hs _ ht i hi hit => - (hs i hi).mono_right (sup_mono fun _ hj => mem_erase.2 ⟨ne_of_mem_of_not_mem hj hit, ht hj⟩)⟩ +theorem SupIndep.antitone_fun {g : ι → α} (hle : ∀ x ∈ s, f x ≤ g x) (h : s.SupIndep g) : + s.SupIndep f := fun _t hts i his hit ↦ + (h hts his hit).mono (hle i his) <| Finset.sup_mono_fun fun x hx ↦ hle x <| hts hx -theorem supIndep_antimono_fun {g : ι → α} (h : ∀ x ∈ s, f x ≤ g x) (h : s.SupIndep g) : - s.SupIndep f := by - classical - induction s using Finset.induction_on with - | empty => apply Finset.supIndep_empty - | @insert i s his IH => - rename_i hle - rw [Finset.supIndep_iff_disjoint_erase] at h ⊢ - intro j hj - simp_all only [Finset.mem_insert, or_true, implies_true, true_implies, forall_eq_or_imp, - Finset.erase_insert_eq_erase, not_false_eq_true, Finset.erase_eq_of_not_mem] - obtain rfl | hj := hj - · simp only [Finset.erase_insert_eq_erase] - apply h.left.mono hle.left - apply (Finset.sup_mono _).trans (Finset.sup_mono_fun hle.right) - exact Finset.erase_subset _ _ - · apply (h.right j hj).mono (hle.right j hj) (Finset.sup_mono_fun _) - intro k hk - simp only [Finset.mem_erase, ne_eq, Finset.mem_insert] at hk - obtain ⟨-, rfl | hk⟩ := hk - · exact hle.left - · exact hle.right k hk - -theorem SupIndep.image [DecidableEq ι] {s : Finset ι'} {g : ι' → ι} (hs : s.SupIndep (f ∘ g)) : - (s.image g).SupIndep f := by +@[deprecated (since := "2025-01-17")] +alias supIndep_antimono_fun := SupIndep.antitone_fun + +protected theorem SupIndep.image [DecidableEq ι] {s : Finset ι'} {g : ι' → ι} + (hs : s.SupIndep (f ∘ g)) : (s.image g).SupIndep f := by intro t ht i hi hit - rw [mem_image] at hi - obtain ⟨i, hi, rfl⟩ := hi - haveI : DecidableEq ι' := Classical.decEq _ - suffices hts : t ⊆ (s.erase i).image g by - refine (supIndep_iff_disjoint_erase.1 hs i hi).mono_right ((sup_mono hts).trans ?_) - rw [sup_image] - rintro j hjt - obtain ⟨j, hj, rfl⟩ := mem_image.1 (ht hjt) - exact mem_image_of_mem _ (mem_erase.2 ⟨ne_of_apply_ne g (ne_of_mem_of_not_mem hjt hit), hj⟩) + rcases subset_image_iff.mp ht with ⟨t, hts, rfl⟩ + rcases mem_image.mp hi with ⟨i, his, rfl⟩ + rw [sup_image] + exact hs hts his (hit <| mem_image_of_mem _ ·) theorem supIndep_map {s : Finset ι'} {g : ι' ↪ ι} : (s.map g).SupIndep f ↔ s.SupIndep (f ∘ g) := by refine ⟨fun hs t ht i hi hit => ?_, fun hs => ?_⟩ @@ -140,24 +129,11 @@ theorem supIndep_map {s : Finset ι'} {g : ι' ↪ ι} : (s.map g).SupIndep f @[simp] theorem supIndep_pair [DecidableEq ι] {i j : ι} (hij : i ≠ j) : - ({i, j} : Finset ι).SupIndep f ↔ Disjoint (f i) (f j) := - ⟨fun h => h.pairwiseDisjoint (by simp) (by simp) hij, - fun h => by - rw [supIndep_iff_disjoint_erase] - intro k hk - rw [Finset.mem_insert, Finset.mem_singleton] at hk - obtain rfl | rfl := hk - · convert h using 1 - rw [Finset.erase_insert, Finset.sup_singleton] - simpa using hij - · convert h.symm using 1 - have : ({i, k} : Finset ι).erase k = {i} := by - ext - rw [mem_erase, mem_insert, mem_singleton, mem_singleton, and_or_left, Ne, - not_and_self_iff, or_false, and_iff_right_of_imp] - rintro rfl - exact hij - rw [this, Finset.sup_singleton]⟩ + ({i, j} : Finset ι).SupIndep f ↔ Disjoint (f i) (f j) := by + suffices Disjoint (f i) (f j) → Disjoint (f j) ((Finset.erase {i, j} j).sup f) by + simpa [supIndep_iff_disjoint_erase, hij] + rw [pair_comm] + simp [hij.symm, disjoint_comm] theorem supIndep_univ_bool (f : Bool → α) : (Finset.univ : Finset Bool).SupIndep f ↔ Disjoint (f false) (f true) := @@ -167,29 +143,14 @@ theorem supIndep_univ_bool (f : Bool → α) : @[simp] theorem supIndep_univ_fin_two (f : Fin 2 → α) : (Finset.univ : Finset (Fin 2)).SupIndep f ↔ Disjoint (f 0) (f 1) := - haveI : (0 : Fin 2) ≠ 1 := by simp + have : (0 : Fin 2) ≠ 1 := by simp supIndep_pair this -theorem SupIndep.attach (hs : s.SupIndep f) : s.attach.SupIndep fun a => f a := by - intro t _ i _ hi - classical - have : (fun (a : { x // x ∈ s }) => f ↑a) = f ∘ (fun a : { x // x ∈ s } => ↑a) := rfl - rw [this, ← Finset.sup_image] - refine hs (image_subset_iff.2 fun (j : { x // x ∈ s }) _ => j.2) i.2 fun hi' => hi ?_ - rw [mem_image] at hi' - obtain ⟨j, hj, hji⟩ := hi' - rwa [Subtype.ext hji] at hj - @[simp] theorem supIndep_attach : (s.attach.SupIndep fun a => f a) ↔ s.SupIndep f := by - refine ⟨fun h t ht i his hit => ?_, SupIndep.attach⟩ - classical - convert h (filter_subset (fun (i : { x // x ∈ s }) => (i : ι) ∈ t) _) (mem_attach _ ⟨i, ‹_›⟩) - fun hi => hit <| by simpa using hi using 1 - refine eq_of_forall_ge_iff ?_ - simp only [Finset.sup_le_iff, mem_filter, mem_attach, true_and, Function.comp_apply, - Subtype.forall, Subtype.coe_mk] - exact fun a => forall_congr' fun j => ⟨fun h _ => h, fun h hj => h (ht hj) hj⟩ + simpa [Finset.attach_map_val] using (supIndep_map (s := s.attach) (g := .subtype _)).symm + +alias ⟨_, SupIndep.attach⟩ := supIndep_attach end Lattice @@ -201,11 +162,10 @@ theorem supIndep_iff_pairwiseDisjoint : s.SupIndep f ↔ (s : Set ι).PairwiseDi ⟨SupIndep.pairwiseDisjoint, fun hs _ ht _ hi hit => Finset.disjoint_sup_right.2 fun _ hj => hs hi (ht hj) (ne_of_mem_of_not_mem hj hit).symm⟩ -alias ⟨sup_indep.pairwise_disjoint, _root_.Set.PairwiseDisjoint.supIndep⟩ := - supIndep_iff_pairwiseDisjoint +alias ⟨_, _root_.Set.PairwiseDisjoint.supIndep⟩ := supIndep_iff_pairwiseDisjoint /-- Bind operation for `SupIndep`. -/ -theorem SupIndep.sup [DecidableEq ι] {s : Finset ι'} {g : ι' → Finset ι} {f : ι → α} +protected theorem SupIndep.sup [DecidableEq ι] {s : Finset ι'} {g : ι' → Finset ι} {f : ι → α} (hs : s.SupIndep fun i => (g i).sup f) (hg : ∀ i' ∈ s, (g i').SupIndep f) : (s.sup g).SupIndep f := by simp_rw [supIndep_iff_pairwiseDisjoint] at hs hg ⊢ @@ -213,15 +173,15 @@ theorem SupIndep.sup [DecidableEq ι] {s : Finset ι'} {g : ι' → Finset ι} { exact hs.biUnion_finset hg /-- Bind operation for `SupIndep`. -/ -theorem SupIndep.biUnion [DecidableEq ι] {s : Finset ι'} {g : ι' → Finset ι} {f : ι → α} +protected theorem SupIndep.biUnion [DecidableEq ι] {s : Finset ι'} {g : ι' → Finset ι} {f : ι → α} (hs : s.SupIndep fun i => (g i).sup f) (hg : ∀ i' ∈ s, (g i').SupIndep f) : (s.biUnion g).SupIndep f := by rw [← sup_eq_biUnion] exact hs.sup hg /-- Bind operation for `SupIndep`. -/ -theorem SupIndep.sigma {β : ι → Type*} {s : Finset ι} {g : ∀ i, Finset (β i)} {f : Sigma β → α} - (hs : s.SupIndep fun i => (g i).sup fun b => f ⟨i, b⟩) +protected theorem SupIndep.sigma {β : ι → Type*} {s : Finset ι} {g : ∀ i, Finset (β i)} + {f : Sigma β → α} (hs : s.SupIndep fun i => (g i).sup fun b => f ⟨i, b⟩) (hg : ∀ i ∈ s, (g i).SupIndep fun b => f ⟨i, b⟩) : (s.sigma g).SupIndep f := by rintro t ht ⟨i, b⟩ hi hit rw [Finset.disjoint_sup_right] @@ -235,7 +195,7 @@ theorem SupIndep.sigma {β : ι → Type*} {s : Finset ι} {g : ∀ i, Finset ( · convert le_sup (α := α) hi.2; simp · convert le_sup (α := α) hj.2; simp -theorem SupIndep.product {s : Finset ι} {t : Finset ι'} {f : ι × ι' → α} +protected theorem SupIndep.product {s : Finset ι} {t : Finset ι'} {f : ι × ι' → α} (hs : s.SupIndep fun i => t.sup fun i' => f (i, i')) (ht : t.SupIndep fun i' => s.sup fun i => f (i, i')) : (s ×ˢ t).SupIndep f := by rintro u hu ⟨i, i'⟩ hi hiu diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index 63a1594ad82a06..8cbe3710d04299 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -956,7 +956,7 @@ lemma le_of_forall_lt_iff_le [LinearOrder α] [DenselyOrdered α] [NoMinOrder α | coe x => rw [le_coe_iff] rintro y rfl - exact le_of_forall_le_of_dense (by exact_mod_cast h) + exact le_of_forall_gt_imp_ge_of_dense (by exact_mod_cast h) lemma ge_of_forall_gt_iff_ge [LinearOrder α] [DenselyOrdered α] [NoMinOrder α] {x y : WithBot α} : (∀ z : α, z < x → z ≤ y) ↔ x ≤ y := by @@ -966,7 +966,7 @@ lemma ge_of_forall_gt_iff_ge [LinearOrder α] [DenselyOrdered α] [NoMinOrder α | coe y => rw [le_coe_iff] rintro h x rfl - exact le_of_forall_ge_of_dense (by exact_mod_cast h) + exact le_of_forall_lt_imp_le_of_dense (by exact_mod_cast h) section LE diff --git a/Mathlib/Probability/Density.lean b/Mathlib/Probability/Density.lean index a89d87c31a7920..3a90248fdc4734 100644 --- a/Mathlib/Probability/Density.lean +++ b/Mathlib/Probability/Density.lean @@ -289,15 +289,14 @@ theorem integral_mul_eq_integral [HasPDF X ℙ] : ∫ x, x * (pdf X ℙ volume x _ = _ := integral_pdf_smul measurable_id.aestronglyMeasurable theorem hasFiniteIntegral_mul {f : ℝ → ℝ} {g : ℝ → ℝ≥0∞} (hg : pdf X ℙ =ᵐ[volume] g) - (hgi : ∫⁻ x, ‖f x‖₊ * g x ≠ ∞) : + (hgi : ∫⁻ x, ‖f x‖ₑ * g x ≠ ∞) : HasFiniteIntegral fun x => f x * (pdf X ℙ volume x).toReal := by - rw [hasFiniteIntegral_iff_nnnorm] - have : (fun x => ↑‖f x‖₊ * g x) =ᵐ[volume] fun x => ‖f x * (pdf X ℙ volume x).toReal‖₊ := by - refine ae_eq_trans (Filter.EventuallyEq.mul (ae_eq_refl fun x => (‖f x‖₊ : ℝ≥0∞)) - (ae_eq_trans hg.symm ofReal_toReal_ae_eq.symm)) ?_ - simp_rw [← smul_eq_mul, nnnorm_smul, ENNReal.coe_mul, smul_eq_mul] - refine Filter.EventuallyEq.mul (ae_eq_refl _) ?_ - simp only [Real.ennnorm_eq_ofReal ENNReal.toReal_nonneg, ae_eq_refl] + rw [hasFiniteIntegral_iff_enorm] + have : (fun x => ‖f x‖ₑ * g x) =ᵐ[volume] fun x => ‖f x * (pdf X ℙ volume x).toReal‖ₑ := by + refine ae_eq_trans ((ae_eq_refl _).mul (ae_eq_trans hg.symm ofReal_toReal_ae_eq.symm)) ?_ + simp_rw [← smul_eq_mul, enorm_smul, smul_eq_mul] + refine .mul (ae_eq_refl _) ?_ + simp only [Real.enorm_eq_ofReal ENNReal.toReal_nonneg, ae_eq_refl] rwa [lt_top_iff_ne_top, ← lintegral_congr_ae this] end Real diff --git a/Mathlib/Probability/Distributions/Uniform.lean b/Mathlib/Probability/Distributions/Uniform.lean index 3efd4efda464d5..4c50fb841441c2 100644 --- a/Mathlib/Probability/Distributions/Uniform.lean +++ b/Mathlib/Probability/Distributions/Uniform.lean @@ -153,11 +153,11 @@ theorem mul_pdf_integrable (hcs : IsCompact s) (huX : IsUniform X s ℙ) : (measurable_pdf X ℙ).aemeasurable.ennreal_toReal.aestronglyMeasurable refine hasFiniteIntegral_mul (pdf_eq hcs.measurableSet huX) ?_ set ind := (volume s)⁻¹ • (1 : ℝ → ℝ≥0∞) - have : ∀ x, ↑‖x‖₊ * s.indicator ind x = s.indicator (fun x => ‖x‖₊ * ind x) x := fun x => + have : ∀ x, ‖x‖ₑ * s.indicator ind x = s.indicator (fun x => ‖x‖ₑ * ind x) x := fun x => (s.indicator_mul_right (fun x => ↑‖x‖₊) ind).symm simp only [ind, this, lintegral_indicator hcs.measurableSet, mul_one, Algebra.id.smul_eq_mul, Pi.one_apply, Pi.smul_apply] - rw [lintegral_mul_const _ measurable_nnnorm.coe_nnreal_ennreal] + rw [lintegral_mul_const _ measurable_enorm] exact ENNReal.mul_ne_top (setLIntegral_lt_top_of_isCompact hnt.2 hcs continuous_nnnorm).ne (ENNReal.inv_lt_top.2 (pos_iff_ne_zero.mpr hnt.1)).ne @@ -209,12 +209,13 @@ namespace PMF variable {α : Type*} -open scoped Classical NNReal ENNReal +open scoped NNReal ENNReal section UniformOfFinset /-- Uniform distribution taking the same non-zero probability on the nonempty finset `s` -/ def uniformOfFinset (s : Finset α) (hs : s.Nonempty) : PMF α := by + classical refine ofFinset (fun a => if a ∈ s then s.card⁻¹ else 0) s ?_ ?_ · simp only [Finset.sum_ite_mem, Finset.inter_self, Finset.sum_const, nsmul_eq_mul] have : (s.card : ℝ≥0∞) ≠ 0 := by @@ -225,6 +226,7 @@ def uniformOfFinset (s : Finset α) (hs : s.Nonempty) : PMF α := by variable {s : Finset α} (hs : s.Nonempty) {a : α} +open scoped Classical in @[simp] theorem uniformOfFinset_apply (a : α) : uniformOfFinset s hs a = if a ∈ s then (s.card : ℝ≥0∞)⁻¹ else 0 := @@ -249,6 +251,7 @@ section Measure variable (t : Set α) +open scoped Classical in @[simp] theorem toOuterMeasure_uniformOfFinset_apply : (uniformOfFinset s hs).toOuterMeasure t = (s.filter (· ∈ t)).card / s.card := @@ -266,6 +269,7 @@ theorem toOuterMeasure_uniformOfFinset_apply : _ = (s.filter (· ∈ t)).card / s.card := by simp only [div_eq_mul_inv, Finset.sum_const, nsmul_eq_mul] +open scoped Classical in @[simp] theorem toMeasure_uniformOfFinset_apply [MeasurableSpace α] (ht : MeasurableSet t) : (uniformOfFinset s hs).toMeasure t = (s.filter (· ∈ t)).card / s.card := @@ -298,11 +302,13 @@ section Measure variable (s : Set α) +open scoped Classical in theorem toOuterMeasure_uniformOfFintype_apply : (uniformOfFintype α).toOuterMeasure s = Fintype.card s / Fintype.card α := by rw [uniformOfFintype, toOuterMeasure_uniformOfFinset_apply,Fintype.card_ofFinset] rfl +open scoped Classical in theorem toMeasure_uniformOfFintype_apply [MeasurableSpace α] (hs : MeasurableSet s) : (uniformOfFintype α).toMeasure s = Fintype.card s / Fintype.card α := by simp [uniformOfFintype, hs] @@ -313,6 +319,7 @@ end UniformOfFintype section OfMultiset +open scoped Classical in /-- Given a non-empty multiset `s` we construct the `PMF` which sends `a` to the fraction of elements in `s` that are `a`. -/ def ofMultiset (s : Multiset α) (hs : s ≠ 0) : PMF α := @@ -334,14 +341,17 @@ def ofMultiset (s : Multiset α) (hs : s ≠ 0) : PMF α := variable {s : Multiset α} (hs : s ≠ 0) +open scoped Classical in @[simp] theorem ofMultiset_apply (a : α) : ofMultiset s hs a = s.count a / (Multiset.card s) := rfl +open scoped Classical in @[simp] theorem support_ofMultiset : (ofMultiset s hs).support = s.toFinset := Set.ext (by simp [mem_support_iff, hs]) +open scoped Classical in theorem mem_support_ofMultiset_iff (a : α) : a ∈ (ofMultiset s hs).support ↔ a ∈ s.toFinset := by simp @@ -353,6 +363,7 @@ section Measure variable (t : Set α) +open scoped Classical in @[simp] theorem toOuterMeasure_ofMultiset_apply : (ofMultiset s hs).toOuterMeasure t = @@ -361,6 +372,7 @@ theorem toOuterMeasure_ofMultiset_apply : refine tsum_congr fun x => ?_ by_cases hx : x ∈ t <;> simp [Set.indicator, hx, div_eq_mul_inv] +open scoped Classical in @[simp] theorem toMeasure_ofMultiset_apply [MeasurableSpace α] (ht : MeasurableSet t) : (ofMultiset s hs).toMeasure t = (∑' x, (s.filter (· ∈ t)).count x : ℝ≥0∞) / (Multiset.card s) := diff --git a/Mathlib/Probability/Independence/Integrable.lean b/Mathlib/Probability/Independence/Integrable.lean index c33b372d816a2c..0d06f7f91d98f2 100644 --- a/Mathlib/Probability/Independence/Integrable.lean +++ b/Mathlib/Probability/Independence/Integrable.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.MeasureTheory.Function.L1Space +import Mathlib.MeasureTheory.Function.L1Space.Integrable import Mathlib.Probability.Independence.Basic /-! diff --git a/Mathlib/Probability/Integration.lean b/Mathlib/Probability/Integration.lean index 77a6aa46b89e1e..7d987ea9c09355 100644 --- a/Mathlib/Probability/Integration.lean +++ b/Mathlib/Probability/Integration.lean @@ -153,19 +153,16 @@ theorem lintegral_prod_eq_prod_lintegral_of_indepFun {ι : Type*} [DecidableEq theorem IndepFun.integrable_mul {β : Type*} [MeasurableSpace β] {X Y : Ω → β} [NormedDivisionRing β] [BorelSpace β] (hXY : IndepFun X Y μ) (hX : Integrable X μ) (hY : Integrable Y μ) : Integrable (X * Y) μ := by - let nX : Ω → ENNReal := fun a => ‖X a‖₊ - let nY : Ω → ENNReal := fun a => ‖Y a‖₊ - have hXY' : IndepFun (fun a => ‖X a‖₊) (fun a => ‖Y a‖₊) μ := - hXY.comp measurable_nnnorm measurable_nnnorm - have hXY'' : IndepFun nX nY μ := - hXY'.comp measurable_coe_nnreal_ennreal measurable_coe_nnreal_ennreal - have hnX : AEMeasurable nX μ := hX.1.aemeasurable.nnnorm.coe_nnreal_ennreal - have hnY : AEMeasurable nY μ := hY.1.aemeasurable.nnnorm.coe_nnreal_ennreal + let nX : Ω → ℝ≥0∞ := fun a => ‖X a‖ₑ + let nY : Ω → ℝ≥0∞ := fun a => ‖Y a‖ₑ + have hXY' : IndepFun nX nY μ := hXY.comp measurable_enorm measurable_enorm + have hnX : AEMeasurable nX μ := hX.1.aemeasurable.enorm + have hnY : AEMeasurable nY μ := hY.1.aemeasurable.enorm have hmul : ∫⁻ a, nX a * nY a ∂μ = (∫⁻ a, nX a ∂μ) * ∫⁻ a, nY a ∂μ := - lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun' hnX hnY hXY'' + lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun' hnX hnY hXY' refine ⟨hX.1.mul hY.1, ?_⟩ simp only [nX, nY] at hmul - simp_rw [hasFiniteIntegral_iff_nnnorm, Pi.mul_apply, nnnorm_mul, ENNReal.coe_mul, hmul] + simp_rw [hasFiniteIntegral_iff_enorm, Pi.mul_apply, enorm_mul, hmul] exact ENNReal.mul_lt_top hX.2 hY.2 /-- If the product of two independent real-valued random variables is integrable and @@ -175,18 +172,16 @@ theorem IndepFun.integrable_left_of_integrable_mul {β : Type*} [MeasurableSpace (hX : AEStronglyMeasurable X μ) (hY : AEStronglyMeasurable Y μ) (h'Y : ¬Y =ᵐ[μ] 0) : Integrable X μ := by refine ⟨hX, ?_⟩ - have I : (∫⁻ ω, ‖Y ω‖₊ ∂μ) ≠ 0 := fun H ↦ by - have I : (fun ω => ‖Y ω‖₊ : Ω → ℝ≥0∞) =ᵐ[μ] 0 := (lintegral_eq_zero_iff' hY.ennnorm).1 H + have I : (∫⁻ ω, ‖Y ω‖ₑ ∂μ) ≠ 0 := fun H ↦ by + have I : (fun ω => ‖Y ω‖ₑ : Ω → ℝ≥0∞) =ᵐ[μ] 0 := (lintegral_eq_zero_iff' hY.enorm).1 H apply h'Y filter_upwards [I] with ω hω simpa using hω - refine hasFiniteIntegral_iff_nnnorm.mpr <| lt_top_iff_ne_top.2 fun H => ?_ - have J : IndepFun (fun ω => ‖X ω‖₊ : Ω → ℝ≥0∞) (fun ω => ‖Y ω‖₊ : Ω → ℝ≥0∞) μ := by - have M : Measurable fun x : β => (‖x‖₊ : ℝ≥0∞) := measurable_nnnorm.coe_nnreal_ennreal - apply IndepFun.comp hXY M M - have A : (∫⁻ ω, ‖X ω * Y ω‖₊ ∂μ) < ∞ := h'XY.2 - simp only [nnnorm_mul, ENNReal.coe_mul] at A - rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun'' hX.ennnorm hY.ennnorm J, H] at A + refine hasFiniteIntegral_iff_enorm.mpr <| lt_top_iff_ne_top.2 fun H => ?_ + have J : IndepFun (‖X ·‖ₑ) (‖Y ·‖ₑ) μ := hXY.comp measurable_enorm measurable_enorm + have A : ∫⁻ ω, ‖X ω * Y ω‖ₑ ∂μ < ∞ := h'XY.2 + simp only [enorm_mul, ENNReal.coe_mul] at A + rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun'' hX.enorm hY.enorm J, H] at A simp only [ENNReal.top_mul I, lt_self_iff_false] at A /-- If the product of two independent real-valued random variables is integrable and the @@ -196,19 +191,17 @@ theorem IndepFun.integrable_right_of_integrable_mul {β : Type*} [MeasurableSpac (hX : AEStronglyMeasurable X μ) (hY : AEStronglyMeasurable Y μ) (h'X : ¬X =ᵐ[μ] 0) : Integrable Y μ := by refine ⟨hY, ?_⟩ - have I : (∫⁻ ω, ‖X ω‖₊ ∂μ) ≠ 0 := fun H ↦ by - have I : (fun ω => ‖X ω‖₊ : Ω → ℝ≥0∞) =ᵐ[μ] 0 := (lintegral_eq_zero_iff' hX.ennnorm).1 H + have I : ∫⁻ ω, ‖X ω‖ₑ ∂μ ≠ 0 := fun H ↦ by + have I : ((‖X ·‖ₑ) : Ω → ℝ≥0∞) =ᵐ[μ] 0 := (lintegral_eq_zero_iff' hX.enorm).1 H apply h'X filter_upwards [I] with ω hω simpa using hω refine lt_top_iff_ne_top.2 fun H => ?_ - have J : IndepFun (fun ω => ‖X ω‖₊ : Ω → ℝ≥0∞) (fun ω => ‖Y ω‖₊ : Ω → ℝ≥0∞) μ := by - have M : Measurable fun x : β => (‖x‖₊ : ℝ≥0∞) := measurable_nnnorm.coe_nnreal_ennreal - apply IndepFun.comp hXY M M - have A : (∫⁻ ω, ‖X ω * Y ω‖₊ ∂μ) < ∞ := h'XY.2 - simp only [nnnorm_mul, ENNReal.coe_mul] at A - simp_rw [enorm_eq_nnnorm] at H - rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun'' hX.ennnorm hY.ennnorm J, H] at A + have J : IndepFun (fun ω => ‖X ω‖ₑ : Ω → ℝ≥0∞) (fun ω => ‖Y ω‖ₑ : Ω → ℝ≥0∞) μ := + IndepFun.comp hXY measurable_enorm measurable_enorm + have A : ∫⁻ ω, ‖X ω * Y ω‖ₑ ∂μ < ∞ := h'XY.2 + simp only [enorm_mul, ENNReal.coe_mul] at A + rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun'' hX.enorm hY.enorm J, H] at A simp only [ENNReal.mul_top I, lt_self_iff_false] at A /-- The (Bochner) integral of the product of two independent, nonnegative random diff --git a/Mathlib/Probability/Kernel/Composition/Basic.lean b/Mathlib/Probability/Kernel/Composition/Basic.lean index c5cdd7bb8284a8..3a6ccc33841e1d 100644 --- a/Mathlib/Probability/Kernel/Composition/Basic.lean +++ b/Mathlib/Probability/Kernel/Composition/Basic.lean @@ -158,8 +158,7 @@ theorem measurable_compProdFun (κ : Kernel α β) [IsSFiniteKernel κ] (η : Ke @[deprecated (since := "2024-08-30")] alias measurable_compProdFun_of_finite := measurable_compProdFun -open scoped Classical - +open scoped Classical in /-- Composition-Product of kernels. For s-finite kernels, it satisfies `∫⁻ bc, f bc ∂(compProd κ η a) = ∫⁻ b, ∫⁻ c, f (b, c) ∂(η (a, b)) ∂(κ a)` (see `ProbabilityTheory.Kernel.lintegral_compProd`). @@ -259,6 +258,7 @@ lemma compProd_preimage_fst {s : Set β} (hs : MeasurableSet s) (κ : Kernel α lemma compProd_deterministic_apply [MeasurableSingletonClass γ] {f : α × β → γ} (hf : Measurable f) {s : Set (β × γ)} (hs : MeasurableSet s) (κ : Kernel α β) [IsSFiniteKernel κ] (x : α) : (κ ⊗ₖ deterministic f hf) x s = κ x {b | (b, f (x, b)) ∈ s} := by + classical simp only [deterministic_apply, measurableSet_setOf, Set.mem_setOf_eq, Measure.dirac_apply, Set.mem_setOf_eq, Set.indicator_apply, Pi.one_apply, compProd_apply hs] let t := {b | (b, f (x, b)) ∈ s} diff --git a/Mathlib/Probability/Kernel/Composition/CompNotation.lean b/Mathlib/Probability/Kernel/Composition/CompNotation.lean new file mode 100644 index 00000000000000..a206f6de862820 --- /dev/null +++ b/Mathlib/Probability/Kernel/Composition/CompNotation.lean @@ -0,0 +1,56 @@ +/- +Copyright (c) 2025 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne, Lorenzo Luccioli +-/ +import Mathlib.Probability.Kernel.Basic + +/-! +# Notation for the compostition of a measure and a kernel + +This operation, for which we introduce the notation `∘ₘ`, takes `μ : Measure α` and +`κ : Kernel α β` and creates `κ ∘ₘ μ : Measure β`. The integral of a function against `κ ∘ₘ μ` is +`∫⁻ x, f x ∂(κ ∘ₘ μ) = ∫⁻ a, ∫⁻ b, f b ∂(κ a) ∂μ`. + +This file does not define composition but only introduces notation for +`MeasureTheory.Measure.bind μ κ`. + +## Notations + +* `κ ∘ₘ μ = MeasureTheory.Measure.bind μ κ`, for `κ` a kernel. +-/ + +/- This file is only for lemmas that are direct specializations of `Measure.bind` to kernels, +anything more involved shoud go elsewhere (for example the `MeasureComp` file). -/ +assert_not_exists ProbabilityTheory.Kernel.compProd + +open ProbabilityTheory + +namespace MeasureTheory.Measure + +variable {α β : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + {μ : Measure α} {κ : Kernel α β} + +/-- Composition of a measure and a kernel. + +Notation for `MeasureTheory.Measure.bind` -/ +scoped[ProbabilityTheory] notation3 κ " ∘ₘ " μ:100 => MeasureTheory.Measure.bind μ κ + +@[simp] +lemma comp_apply_univ [IsMarkovKernel κ] : (κ ∘ₘ μ) Set.univ = μ Set.univ := by + simp [bind_apply .univ κ.measurable] + +lemma deterministic_comp_eq_map {f : α → β} (hf : Measurable f) : + Kernel.deterministic f hf ∘ₘ μ = μ.map f := + Measure.bind_dirac_eq_map μ hf + +@[simp] +lemma id_comp : Kernel.id ∘ₘ μ = μ := by rw [Kernel.id, deterministic_comp_eq_map, Measure.map_id] + +lemma swap_comp {μ : Measure (α × β)} : (Kernel.swap α β) ∘ₘ μ = μ.map Prod.swap := + deterministic_comp_eq_map measurable_swap + +@[simp] +lemma const_comp {ν : Measure β} : (Kernel.const α ν) ∘ₘ μ = μ Set.univ • ν := μ.bind_const + +end MeasureTheory.Measure diff --git a/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean index a9ab89e8cf8f4d..3b0cddf917b806 100644 --- a/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean @@ -45,7 +45,7 @@ namespace ProbabilityTheory theorem hasFiniteIntegral_prod_mk_left (a : α) {s : Set (β × γ)} (h2s : (κ ⊗ₖ η) a s ≠ ∞) : HasFiniteIntegral (fun b => (η (a, b) (Prod.mk b ⁻¹' s)).toReal) (κ a) := by let t := toMeasurable ((κ ⊗ₖ η) a) s - simp_rw [hasFiniteIntegral_iff_nnnorm, ennnorm_eq_ofReal toReal_nonneg] + simp_rw [hasFiniteIntegral_iff_enorm, enorm_eq_ofReal toReal_nonneg] calc ∫⁻ b, ENNReal.ofReal (η (a, b) (Prod.mk b ⁻¹' s)).toReal ∂κ a _ ≤ ∫⁻ b, η (a, b) (Prod.mk b ⁻¹' t) ∂κ a := by @@ -82,19 +82,19 @@ theorem hasFiniteIntegral_compProd_iff ⦃f : β × γ → E⦄ (h1f : StronglyM HasFiniteIntegral f ((κ ⊗ₖ η) a) ↔ (∀ᵐ x ∂κ a, HasFiniteIntegral (fun y => f (x, y)) (η (a, x))) ∧ HasFiniteIntegral (fun x => ∫ y, ‖f (x, y)‖ ∂η (a, x)) (κ a) := by - simp only [hasFiniteIntegral_iff_nnnorm] - rw [Kernel.lintegral_compProd _ _ _ h1f.ennnorm] + simp only [hasFiniteIntegral_iff_enorm] + rw [lintegral_compProd _ _ _ h1f.enorm] have : ∀ x, ∀ᵐ y ∂η (a, x), 0 ≤ ‖f (x, y)‖ := fun x => Eventually.of_forall fun y => norm_nonneg _ simp_rw [integral_eq_lintegral_of_nonneg_ae (this _) (h1f.norm.comp_measurable measurable_prod_mk_left).aestronglyMeasurable, - ennnorm_eq_ofReal toReal_nonneg, ofReal_norm_eq_coe_nnnorm] + enorm_eq_ofReal toReal_nonneg, ofReal_norm_eq_enorm] have : ∀ {p q r : Prop} (_ : r → p), (r ↔ p ∧ q) ↔ p → (r ↔ q) := fun {p q r} h1 => by rw [← and_congr_right_iff, and_iff_right_of_imp h1] rw [this] · intro h2f; rw [lintegral_congr_ae] filter_upwards [h2f] with x hx rw [ofReal_toReal]; rw [← lt_top_iff_ne_top]; exact hx - · intro h2f; refine ae_lt_top ?_ h2f.ne; exact h1f.ennnorm.lintegral_kernel_prod_right'' + · intro h2f; refine ae_lt_top ?_ h2f.ne; exact h1f.enorm.lintegral_kernel_prod_right'' theorem hasFiniteIntegral_compProd_iff' ⦃f : β × γ → E⦄ (h1f : AEStronglyMeasurable f ((κ ⊗ₖ η) a)) : @@ -200,21 +200,18 @@ theorem Kernel.continuous_integral_integral : refine tendsto_integral_of_L1 _ (L1.integrable_coeFn g).integral_compProd (Eventually.of_forall fun h => (L1.integrable_coeFn h).integral_compProd) ?_ - simp_rw [← - Kernel.lintegral_fn_integral_sub (fun x => (‖x‖₊ : ℝ≥0∞)) (L1.integrable_coeFn _) - (L1.integrable_coeFn g)] + simp_rw [← lintegral_fn_integral_sub (‖·‖ₑ) (L1.integrable_coeFn _) (L1.integrable_coeFn g)] apply tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds _ (fun i => zero_le _) _ - · exact fun i => ∫⁻ x, ∫⁻ y, ‖i (x, y) - g (x, y)‖₊ ∂η (a, x) ∂κ a - swap; · exact fun i => lintegral_mono fun x => ennnorm_integral_le_lintegral_ennnorm _ + · exact fun i => ∫⁻ x, ∫⁻ y, ‖i (x, y) - g (x, y)‖ₑ ∂η (a, x) ∂κ a + swap; · exact fun i => lintegral_mono fun x => enorm_integral_le_lintegral_enorm _ show Tendsto - (fun i : β × γ →₁[(κ ⊗ₖ η) a] E => ∫⁻ x, ∫⁻ y : γ, ‖i (x, y) - g (x, y)‖₊ ∂η (a, x) ∂κ a) + (fun i : β × γ →₁[(κ ⊗ₖ η) a] E => ∫⁻ x, ∫⁻ y : γ, ‖i (x, y) - g (x, y)‖ₑ ∂η (a, x) ∂κ a) (𝓝 g) (𝓝 0) - have : ∀ i : (MeasureTheory.Lp (α := β × γ) E 1 (((κ ⊗ₖ η) a) : Measure (β × γ))), - Measurable fun z => (‖i z - g z‖₊ : ℝ≥0∞) := fun i => - ((Lp.stronglyMeasurable i).sub (Lp.stronglyMeasurable g)).ennnorm - simp_rw [← Kernel.lintegral_compProd _ _ _ (this _), ← L1.ofReal_norm_sub_eq_lintegral, ← - ofReal_zero] + have this (i : Lp (α := β × γ) E 1 (((κ ⊗ₖ η) a) : Measure (β × γ))) : + Measurable fun z => ‖i z - g z‖ₑ := + ((Lp.stronglyMeasurable i).sub (Lp.stronglyMeasurable g)).enorm + simp_rw [← lintegral_compProd _ _ _ (this _), ← L1.ofReal_norm_sub_eq_lintegral, ← ofReal_zero] refine (continuous_ofReal.tendsto 0).comp ?_ rw [← tendsto_iff_norm_sub_tendsto_zero] exact tendsto_id diff --git a/Mathlib/Probability/Kernel/Composition/Lemmas.lean b/Mathlib/Probability/Kernel/Composition/Lemmas.lean new file mode 100644 index 00000000000000..f4afde59594803 --- /dev/null +++ b/Mathlib/Probability/Kernel/Composition/Lemmas.lean @@ -0,0 +1,103 @@ +/- +Copyright (c) 2025 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne, Lorenzo Luccioli +-/ +import Mathlib.Probability.Kernel.Composition.MeasureComp +import Mathlib.Probability.Kernel.Composition.ParallelComp + +/-! +# Lemmas relating different ways to compose measures and kernels + +This file contains lemmas about the composition of measures and kernels that do not fit in any of +the other files in this directory, because they involve several types of compositions/products. + +## Main statements + +* `parallelComp_comp_parallelComp`: `(η ∥ₖ η') ∘ₖ (κ ∥ₖ κ') = (η ∘ₖ κ) ∥ₖ (η' ∘ₖ κ')` + +-/ + +open MeasureTheory ProbabilityTheory + +variable {α β γ δ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + {mγ : MeasurableSpace γ} {mδ : MeasurableSpace δ} + {μ : Measure α} {ν : Measure β} {κ : Kernel α β} + +namespace MeasureTheory.Measure + +lemma compProd_eq_parallelComp_comp_copy_comp [SFinite μ] [IsSFiniteKernel κ] : + μ ⊗ₘ κ = (Kernel.id ∥ₖ κ) ∘ₘ Kernel.copy α ∘ₘ μ := by + rw [compProd_eq_comp_prod, ← Kernel.parallelComp_comp_copy, Measure.comp_assoc] + +lemma prod_comp_right [SFinite ν] {κ : Kernel β γ} [IsSFiniteKernel κ] : + μ.prod (κ ∘ₘ ν) = (Kernel.id ∥ₖ κ) ∘ₘ (μ.prod ν) := by + ext s hs + rw [Measure.prod_apply hs, Measure.bind_apply hs (Kernel.measurable _)] + simp_rw [Measure.bind_apply (measurable_prod_mk_left hs) (Kernel.measurable _)] + rw [MeasureTheory.lintegral_prod] + swap; · exact (Kernel.measurable_coe _ hs).aemeasurable + congr with a + congr with b + rw [Kernel.parallelComp_apply, Kernel.id_apply, Measure.prod_apply hs, lintegral_dirac'] + exact measurable_measure_prod_mk_left hs + +lemma prod_comp_left [SFinite μ] [SFinite ν] {κ : Kernel α γ} [IsSFiniteKernel κ] : + (κ ∘ₘ μ).prod ν = (κ ∥ₖ Kernel.id) ∘ₘ (μ.prod ν) := by + have h1 : (κ ∘ₘ μ).prod ν = (ν.prod (κ ∘ₘ μ)).map Prod.swap := by rw [Measure.prod_swap] + have h2 : (κ ∥ₖ Kernel.id) ∘ₘ (μ.prod ν) = ((Kernel.id ∥ₖ κ) ∘ₘ (ν.prod μ)).map Prod.swap := by + calc (κ ∥ₖ Kernel.id) ∘ₘ (μ.prod ν) + _ = (κ ∥ₖ Kernel.id) ∘ₘ ((ν.prod μ).map Prod.swap) := by rw [Measure.prod_swap] + _ = (κ ∥ₖ Kernel.id) ∘ₘ ((Kernel.swap _ _) ∘ₘ (ν.prod μ)) := by + rw [Kernel.swap, Measure.deterministic_comp_eq_map] + _ = (Kernel.swap _ _) ∘ₘ ((Kernel.id ∥ₖ κ) ∘ₘ (ν.prod μ)) := by + rw [Measure.comp_assoc, Measure.comp_assoc, Kernel.swap_parallelComp] + _ = ((Kernel.id ∥ₖ κ) ∘ₘ (ν.prod μ)).map Prod.swap := by + rw [Kernel.swap, Measure.deterministic_comp_eq_map] + rw [← Measure.prod_comp_right, ← h1] at h2 + exact h2.symm + +end MeasureTheory.Measure + +namespace ProbabilityTheory.Kernel + +variable {α' β' γ' : Type*} {mα' : MeasurableSpace α'} {mβ' : MeasurableSpace β'} + {mγ' : MeasurableSpace γ'} + [IsSFiniteKernel κ] + +lemma parallelComp_id_left_comp_parallelComp + {η : Kernel α' γ} [IsSFiniteKernel η] {ξ : Kernel γ δ} [IsSFiniteKernel ξ] : + (Kernel.id ∥ₖ ξ) ∘ₖ (κ ∥ₖ η) = κ ∥ₖ (ξ ∘ₖ η) := by + ext a + rw [parallelComp_apply, comp_apply, comp_apply, parallelComp_apply, Measure.prod_comp_right] + +lemma parallelComp_id_right_comp_parallelComp {η : Kernel α' γ} [IsSFiniteKernel η] + {ξ : Kernel γ δ} [IsSFiniteKernel ξ] : + (ξ ∥ₖ Kernel.id) ∘ₖ (η ∥ₖ κ) = (ξ ∘ₖ η) ∥ₖ κ := by + ext a + rw [parallelComp_apply, comp_apply, comp_apply, parallelComp_apply, Measure.prod_comp_left] + +lemma parallelComp_comp_parallelComp {η : Kernel β γ} [IsSFiniteKernel η] + {κ' : Kernel α' β'} [IsSFiniteKernel κ'] {η' : Kernel β' γ'} [IsSFiniteKernel η'] : + (η ∥ₖ η') ∘ₖ (κ ∥ₖ κ') = (η ∘ₖ κ) ∥ₖ (η' ∘ₖ κ') := by + rw [← parallelComp_id_left_comp_parallelComp, ← parallelComp_id_right_comp_parallelComp, + ← comp_assoc, parallelComp_id_left_comp_parallelComp, comp_id] + +lemma parallelComp_comp_prod {η : Kernel β γ} [IsSFiniteKernel η] + {κ' : Kernel α β'} [IsSFiniteKernel κ'] {η' : Kernel β' γ'} [IsSFiniteKernel η'] : + (η ∥ₖ η') ∘ₖ (κ ×ₖ κ') = (η ∘ₖ κ) ×ₖ (η' ∘ₖ κ') := by + rw [← parallelComp_comp_copy, ← comp_assoc, parallelComp_comp_parallelComp, + ← parallelComp_comp_copy] + +lemma parallelComp_comm {η : Kernel γ δ} [IsSFiniteKernel η] : + (Kernel.id ∥ₖ κ) ∘ₖ (η ∥ₖ Kernel.id) = (η ∥ₖ Kernel.id) ∘ₖ (Kernel.id ∥ₖ κ) := by + rw [parallelComp_id_left_comp_parallelComp, parallelComp_id_right_comp_parallelComp, + comp_id, comp_id] + +end ProbabilityTheory.Kernel + +lemma MeasureTheory.Measure.parallelComp_comp_compProd [SFinite μ] + [IsSFiniteKernel κ] {η : Kernel β γ} [IsSFiniteKernel η] : + (Kernel.id ∥ₖ η) ∘ₘ (μ ⊗ₘ κ) = μ ⊗ₘ (η ∘ₖ κ) := by + rw [Measure.compProd_eq_comp_prod, Measure.compProd_eq_comp_prod, Measure.comp_assoc, + Kernel.parallelComp_comp_prod, Kernel.id_comp] diff --git a/Mathlib/Probability/Kernel/Composition/MeasureComp.lean b/Mathlib/Probability/Kernel/Composition/MeasureComp.lean new file mode 100644 index 00000000000000..80b35782496b39 --- /dev/null +++ b/Mathlib/Probability/Kernel/Composition/MeasureComp.lean @@ -0,0 +1,112 @@ +/- +Copyright (c) 2025 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne, Lorenzo Luccioli +-/ +import Mathlib.Probability.Kernel.Composition.CompNotation +import Mathlib.Probability.Kernel.Composition.MeasureCompProd + +/-! +# Lemmas about the composition of a measure and a kernel + +Basic lemmas about the composition `κ ∘ₘ μ` of a kernel `κ` and a measure `μ`. + +-/ + +open scoped ENNReal + +open ProbabilityTheory MeasureTheory + +namespace MeasureTheory.Measure + +variable {α β γ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} + {μ ν : Measure α} {κ η : Kernel α β} + +lemma comp_assoc {η : Kernel β γ} : η ∘ₘ (κ ∘ₘ μ) = (η ∘ₖ κ) ∘ₘ μ := + Measure.bind_bind κ.measurable η.measurable + +@[simp] +lemma snd_compProd (μ : Measure α) [SFinite μ] (κ : Kernel α β) [IsSFiniteKernel κ] : + (μ ⊗ₘ κ).snd = κ ∘ₘ μ := by + ext s hs + rw [bind_apply hs κ.measurable, snd_apply hs, compProd_apply] + · rfl + · exact measurable_snd hs + +instance [SFinite μ] [IsSFiniteKernel κ] : SFinite (κ ∘ₘ μ) := by + rw [← snd_compProd]; infer_instance + +instance [IsFiniteMeasure μ] [IsFiniteKernel κ] : IsFiniteMeasure (κ ∘ₘ μ) := by + rw [← snd_compProd]; infer_instance + +instance [IsProbabilityMeasure μ] [IsMarkovKernel κ] : IsProbabilityMeasure (κ ∘ₘ μ) := by + rw [← snd_compProd]; infer_instance + +instance [IsZeroOrProbabilityMeasure μ] [IsZeroOrMarkovKernel κ] : + IsZeroOrProbabilityMeasure (κ ∘ₘ μ) := by + rw [← snd_compProd]; infer_instance + +lemma map_comp (μ : Measure α) (κ : Kernel α β) {f : β → γ} (hf : Measurable f) : + (κ ∘ₘ μ).map f = (κ.map f) ∘ₘ μ := by + ext s hs + rw [Measure.map_apply hf hs, Measure.bind_apply (hf hs) κ.measurable, + Measure.bind_apply hs (Kernel.measurable _)] + simp_rw [Kernel.map_apply' _ hf _ hs] + +section CompProd + +lemma compProd_eq_comp_prod (μ : Measure α) [SFinite μ] (κ : Kernel α β) [IsSFiniteKernel κ] : + μ ⊗ₘ κ = (Kernel.id ×ₖ κ) ∘ₘ μ := by + rw [compProd, Kernel.compProd_prodMkLeft_eq_comp] + rfl + +lemma compProd_id_eq_copy_comp [SFinite μ] : μ ⊗ₘ Kernel.id = Kernel.copy α ∘ₘ μ := by + rw [compProd_id, Kernel.copy, deterministic_comp_eq_map] + +end CompProd + +section AddSMul + +@[simp] +lemma comp_add [SFinite μ] [SFinite ν] [IsSFiniteKernel κ] : + κ ∘ₘ (μ + ν) = κ ∘ₘ μ + κ ∘ₘ ν := by + simp_rw [← snd_compProd, compProd_add_left] + simp + +lemma add_comp [SFinite μ] [IsSFiniteKernel κ] [IsSFiniteKernel η] : + (κ + η) ∘ₘ μ = κ ∘ₘ μ + η ∘ₘ μ := by + simp_rw [← snd_compProd, compProd_add_right] + simp + +/-- Same as `add_comp` except that it uses `⇑κ + ⇑η` instead of `⇑(κ + η)` in order to have +a simp-normal form on the left of the equality. -/ +@[simp] +lemma add_comp' [SFinite μ] [IsSFiniteKernel κ] [IsSFiniteKernel η] : + (⇑κ + ⇑η) ∘ₘ μ = κ ∘ₘ μ + η ∘ₘ μ := by rw [← Kernel.coe_add, add_comp] + +lemma comp_smul (a : ℝ≥0∞) : κ ∘ₘ (a • μ) = a • (κ ∘ₘ μ) := by + ext s hs + simp only [bind_apply hs κ.measurable, lintegral_smul_measure, smul_apply, smul_eq_mul] + +end AddSMul + +section AbsolutelyContinuous + +variable [SFinite μ] [SFinite ν] [IsSFiniteKernel κ] [IsSFiniteKernel η] + +lemma AbsolutelyContinuous.comp (hμν : μ ≪ ν) (hκη : ∀ᵐ a ∂μ, κ a ≪ η a) : + κ ∘ₘ μ ≪ η ∘ₘ ν := by + simp_rw [← snd_compProd, Measure.snd] + exact (hμν.compProd hκη).map measurable_snd + +lemma AbsolutelyContinuous.comp_right (hμν : μ ≪ ν) (κ : Kernel α γ) [IsSFiniteKernel κ] : + κ ∘ₘ μ ≪ κ ∘ₘ ν := + hμν.comp (ae_of_all μ fun _ _ a ↦ a) + +lemma AbsolutelyContinuous.comp_left (μ : Measure α) [SFinite μ] (hκη : ∀ᵐ a ∂μ, κ a ≪ η a) : + κ ∘ₘ μ ≪ η ∘ₘ μ := + μ.absolutelyContinuous_refl.comp hκη + +end AbsolutelyContinuous + +end MeasureTheory.Measure diff --git a/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean index 64933b8111dc2f..d39c73a222f5cc 100644 --- a/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean @@ -81,6 +81,17 @@ lemma _root_.ProbabilityTheory.Kernel.compProd_apply_eq_compProd_sectR {γ : Typ simp_rw [Kernel.compProd_apply hs, compProd_apply hs, Kernel.sectR_apply] rfl +lemma compProd_id [SFinite μ] : μ ⊗ₘ Kernel.id = μ.map (fun x ↦ (x, x)) := by + ext s hs + rw [compProd_apply hs, map_apply (measurable_id.prod measurable_id) hs] + have h_meas a : MeasurableSet (Prod.mk a ⁻¹' s) := measurable_prod_mk_left hs + simp_rw [Kernel.id_apply, dirac_apply' _ (h_meas _)] + calc ∫⁻ a, (Prod.mk a ⁻¹' s).indicator 1 a ∂μ + _ = ∫⁻ a, ((fun x ↦ (x, x)) ⁻¹' s).indicator 1 a ∂μ := rfl + _ = μ ((fun x ↦ (x, x)) ⁻¹' s) := by + rw [lintegral_indicator_one] + exact (measurable_id.prod measurable_id) hs + lemma compProd_congr [IsSFiniteKernel κ] [IsSFiniteKernel η] (h : κ =ᵐ[μ] η) : μ ⊗ₘ κ = μ ⊗ₘ η := by by_cases hμ : SFinite μ @@ -207,6 +218,17 @@ instance [IsFiniteMeasure μ] [IsFiniteKernel κ] : IsFiniteMeasure (μ ⊗ₘ instance [IsProbabilityMeasure μ] [IsMarkovKernel κ] : IsProbabilityMeasure (μ ⊗ₘ κ) := by rw [compProd]; infer_instance +instance [IsZeroOrProbabilityMeasure μ] [IsZeroOrMarkovKernel κ] : + IsZeroOrProbabilityMeasure (μ ⊗ₘ κ) := by + rw [compProd] + rcases eq_zero_or_isProbabilityMeasure μ with rfl | h + · simp only [Kernel.const_zero, Kernel.compProd_zero_left, Kernel.zero_apply] + infer_instance + rcases eq_zero_or_isMarkovKernel κ with rfl | hκ + · simp only [Kernel.prodMkLeft_zero, Kernel.compProd_zero_right, Kernel.zero_apply] + infer_instance + · infer_instance + section AbsolutelyContinuous lemma AbsolutelyContinuous.compProd_left [SFinite ν] (hμν : μ ≪ ν) (κ : Kernel α β) : diff --git a/Mathlib/Probability/Kernel/CondDistrib.lean b/Mathlib/Probability/Kernel/CondDistrib.lean index 8868754b1be920..a0933088f0f3dd 100644 --- a/Mathlib/Probability/Kernel/CondDistrib.lean +++ b/Mathlib/Probability/Kernel/CondDistrib.lean @@ -321,7 +321,7 @@ theorem _root_.MeasureTheory.Integrable.comp_snd_map_prod_mk by_cases hX : AEMeasurable X μ · have hf := hf_int.1.comp_snd_map_prod_mk X (mΩ := mΩ) (mβ := mβ) refine ⟨hf, ?_⟩ - rw [hasFiniteIntegral_iff_nnnorm, lintegral_map' hf.ennnorm (hX.prod_mk aemeasurable_id)] + rw [hasFiniteIntegral_iff_enorm, lintegral_map' hf.enorm (hX.prod_mk aemeasurable_id)] exact hf_int.2 · rw [Measure.map_of_not_aemeasurable] · simp diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean index 7a63f0c847a347..25bf36467a9843 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean @@ -198,7 +198,7 @@ lemma integrable_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℚ Integrable (fun a ↦ (preCDF ρ x a).toReal) ρ.fst := by refine integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) ?_ fun t _ _ ↦ ?_ · exact measurable_preCDF.ennreal_toReal.aestronglyMeasurable - · simp_rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_of_nonneg ENNReal.toReal_nonneg] + · simp_rw [← ofReal_norm_eq_enorm, Real.norm_of_nonneg ENNReal.toReal_nonneg] rw [← lintegral_one] refine (setLIntegral_le_lintegral _ _).trans (lintegral_mono_ae ?_) filter_upwards [preCDF_le_one ρ] with a ha using ENNReal.ofReal_toReal_le.trans (ha _) diff --git a/Mathlib/Probability/Kernel/MeasurableIntegral.lean b/Mathlib/Probability/Kernel/MeasurableIntegral.lean index d3eb71066d3ede..f226af6adbb258 100644 --- a/Mathlib/Probability/Kernel/MeasurableIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableIntegral.lean @@ -32,7 +32,7 @@ theorem ProbabilityTheory.measurableSet_kernel_integrable ⦃f : α → β → E (hf : StronglyMeasurable (uncurry f)) : MeasurableSet {x | Integrable (f x) (κ x)} := by simp_rw [Integrable, hf.of_uncurry_left.aestronglyMeasurable, true_and] - exact measurableSet_lt (Measurable.lintegral_kernel_prod_right hf.ennnorm) measurable_const + exact measurableSet_lt (Measurable.lintegral_kernel_prod_right hf.enorm) measurable_const open ProbabilityTheory.Kernel diff --git a/Mathlib/Probability/Kernel/Proper.lean b/Mathlib/Probability/Kernel/Proper.lean index b7cb185ae7528f..724d6cff74ea3a 100644 --- a/Mathlib/Probability/Kernel/Proper.lean +++ b/Mathlib/Probability/Kernel/Proper.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Yaël Dillies, Kalle Kytölä, Kin Yau James Wong. All rights Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Kalle Kytölä, Kin Yau James Wong -/ -import Mathlib.Probability.Kernel.Basic +import Mathlib.Probability.Kernel.Composition.CompNotation /-! # Proper kernels @@ -66,14 +66,17 @@ alias ⟨IsProper.restrict_eq_indicator_smul, IsProper.of_restrict_eq_indicator_ alias ⟨IsProper.inter_eq_indicator_mul, IsProper.of_inter_eq_indicator_mul⟩ := isProper_iff_inter_eq_indicator_mul -lemma IsProper.setLIntegral_eq_bind (hπ : IsProper π) (h𝓑𝓧 : 𝓑 ≤ 𝓧) {μ : Measure[𝓧] X} +lemma IsProper.setLIntegral_eq_comp (hπ : IsProper π) (h𝓑𝓧 : 𝓑 ≤ 𝓧) {μ : Measure[𝓧] X} (hA : MeasurableSet[𝓧] A) (hB : MeasurableSet[𝓑] B) : - ∫⁻ a in B, π a A ∂μ = μ.bind π (A ∩ B) := by + ∫⁻ a in B, π a A ∂μ = (π ∘ₘ μ) (A ∩ B) := by rw [Measure.bind_apply (by measurability) (π.measurable.mono h𝓑𝓧 le_rfl)] simp only [hπ.inter_eq_indicator_mul h𝓑𝓧 hA hB, ← indicator_mul_const, Pi.one_apply, one_mul] rw [← lintegral_indicator (h𝓑𝓧 _ hB)] rfl +@[deprecated (since := "2025-01-24")] +alias IsProper.setLIntegral_eq_bind := IsProper.setLIntegral_eq_comp + /-- Auxiliary lemma for `IsProper.lintegral_mul` and `IsProper.setLIntegral_eq_indicator_mul_lintegral`. -/ private lemma IsProper.lintegral_indicator_mul_indicator (hπ : IsProper π) (h𝓑𝓧 : 𝓑 ≤ 𝓧) diff --git a/Mathlib/Probability/Martingale/Convergence.lean b/Mathlib/Probability/Martingale/Convergence.lean index 78e152f66e7705..00ffa32b2454bd 100644 --- a/Mathlib/Probability/Martingale/Convergence.lean +++ b/Mathlib/Probability/Martingale/Convergence.lean @@ -154,7 +154,7 @@ theorem Submartingale.upcrossings_ae_lt_top' [IsFiniteMeasure μ] (hf : Submarti rw [mul_comm, ← ENNReal.le_div_iff_mul_le] at this · refine (lt_of_le_of_lt this (ENNReal.div_lt_top ?_ ?_)).ne · have hR' : ∀ n, ∫⁻ ω, ‖f n ω - a‖₊ ∂μ ≤ R + ‖a‖₊ * μ Set.univ := by - simp_rw [eLpNorm_one_eq_lintegral_nnnorm] at hbdd + simp_rw [eLpNorm_one_eq_lintegral_enorm] at hbdd intro n refine (lintegral_mono ?_ : ∫⁻ ω, ‖f n ω - a‖₊ ∂μ ≤ ∫⁻ ω, ‖f n ω‖₊ + ‖a‖₊ ∂μ).trans ?_ · intro ω diff --git a/Mathlib/Probability/Moments/Basic.lean b/Mathlib/Probability/Moments/Basic.lean index ec9aaa1d554221..edcfe1a187ce9e 100644 --- a/Mathlib/Probability/Moments/Basic.lean +++ b/Mathlib/Probability/Moments/Basic.lean @@ -184,6 +184,12 @@ theorem mgf_pos [IsProbabilityMeasure μ] (h_int_X : Integrable (fun ω => exp ( 0 < mgf X μ t := mgf_pos' (IsProbabilityMeasure.ne_zero μ) h_int_X +lemma mgf_pos_iff [hμ : NeZero μ] : + 0 < mgf X μ t ↔ Integrable (fun ω ↦ exp (t * X ω)) μ := by + refine ⟨fun h ↦ ?_, fun h ↦ mgf_pos' hμ.out h⟩ + contrapose! h with h + simp [mgf_undef h] + lemma exp_cgf_of_neZero [hμ : NeZero μ] (hX : Integrable (fun ω ↦ exp (t * X ω)) μ) : exp (cgf X μ t) = mgf X μ t := by rw [cgf, exp_log (mgf_pos' hμ.out hX)] @@ -196,6 +202,11 @@ lemma mgf_id_map (hX : AEMeasurable X μ) : mgf id (μ.map X) = mgf X μ := by · rfl · exact (measurable_const_mul _).exp.aestronglyMeasurable +lemma mgf_congr_identDistrib {Ω' : Type*} {mΩ' : MeasurableSpace Ω'} {μ' : Measure Ω'} + {Y : Ω' → ℝ} (h : IdentDistrib X Y μ μ') : + mgf X μ = mgf Y μ' := by + rw [← mgf_id_map h.aemeasurable_fst, ← mgf_id_map h.aemeasurable_snd, h.map_eq] + theorem mgf_neg : mgf (-X) μ t = mgf X μ (-t) := by simp_rw [mgf, Pi.neg_apply, mul_neg, neg_mul] theorem cgf_neg : cgf (-X) μ t = cgf X μ (-t) := by simp_rw [cgf, mgf_neg] diff --git a/Mathlib/Probability/Moments/ComplexMGF.lean b/Mathlib/Probability/Moments/ComplexMGF.lean new file mode 100644 index 00000000000000..1849cec292373a --- /dev/null +++ b/Mathlib/Probability/Moments/ComplexMGF.lean @@ -0,0 +1,298 @@ +/- +Copyright (c) 2025 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Analysis.Analytic.IsolatedZeros +import Mathlib.Analysis.Calculus.ParametricIntegral +import Mathlib.Analysis.Complex.CauchyIntegral +import Mathlib.Probability.Moments.Basic +import Mathlib.Probability.Moments.IntegrableExpMul + +/-! +# The complex-valued moment generating function + +The moment generating function (mgf) is `t : ℝ ↦ μ[fun ω ↦ rexp (t * X ω)]`. It can be extended to +a complex function `z : ℂ ↦ μ[fun ω ↦ cexp (z * X ω)]`, which we call `complexMGF X μ`. +That function is holomorphic on the vertical strip with base the interior of the interval +of definition of the mgf. +On the vertical line that goes through 0, `complexMGF X μ` is equal to the characteristic function. +This allows us to link properties of the characteristic function and the mgf (mostly deducing +properties of the mgf from those of the characteristic function). + +## Main definitions + +* `complexMGF X μ`: the function `z : ℂ ↦ μ[fun ω ↦ cexp (z * X ω)]`. + +## Main results + +* `complexMGF_ofReal`: for `x : ℝ`, `complexMGF X μ x = mgf X μ x`. + +* `hasDerivAt_complexMGF`: for all `z : ℂ` such that the real part `z.re` belongs to the interior + of the interval of definition of the mgf, `complexMGF X μ` is differentiable at `z` + with derivative `μ[X * exp (z * X)]`. +* `differentiableOn_complexMGF`: `complexMGF X μ` is holomorphic on the vertical strip + `{z | z.re ∈ interior (integrableExpSet X μ)}`. +* `analyticOn_complexMGF`: `complexMGF X μ` is analytic on the vertical strip + `{z | z.re ∈ interior (integrableExpSet X μ)}`. + +* `eqOn_complexMGF_of_mgf`: if two random variables have the same moment generating function, + then they have the same `complexMGF` on the vertical strip + `{z | z.re ∈ interior (integrableExpSet X μ)}`. + Once we know that equal `mgf` implies equal distributions, we will be able to show that + the `complexMGF` are equal everywhere, not only on the strip. + This lemma will be used in the proof of the equality of distributions. + +## TODO + +Once we have a definition for the characteristic function, we will be able to prove the following. + +* `x : ℝ ↦ complexMGF X μ (I * x)` is equal to the characteristic function of + the random variable `X`. +* As a consequence, if two random variables have same `mgf`, then they have the same + characteristic function and the same distribution. + +-/ + + +open MeasureTheory Filter Finset Real Complex + +open scoped MeasureTheory ProbabilityTheory ENNReal NNReal Topology + +namespace ProbabilityTheory + +variable {Ω ι : Type*} {m : MeasurableSpace Ω} {X : Ω → ℝ} {μ : Measure Ω} {t u v : ℝ} {z ε : ℂ} + +/-- Complex extension of the moment generating function. -/ +noncomputable +def complexMGF (X : Ω → ℝ) (μ : Measure Ω) (z : ℂ) : ℂ := μ[fun ω ↦ cexp (z * X ω)] + +lemma complexMGF_undef (hX : AEMeasurable X μ) (h : ¬ Integrable (fun ω ↦ rexp (z.re * X ω)) μ) : + complexMGF X μ z = 0 := by + rw [complexMGF, integral_undef] + rw [← integrable_norm_iff (AEMeasurable.aestronglyMeasurable <| by fun_prop)] + simpa [Complex.norm_eq_abs, Complex.abs_exp] using h + +lemma complexMGF_id_map (hX : AEMeasurable X μ) : complexMGF id (μ.map X) = complexMGF X μ := by + ext t + rw [complexMGF, integral_map hX] + · rfl + · exact AEMeasurable.aestronglyMeasurable <| by fun_prop + +lemma complexMGF_congr_identDistrib {Ω' : Type*} {mΩ' : MeasurableSpace Ω'} {μ' : Measure Ω'} + {Y : Ω' → ℝ} (h : IdentDistrib X Y μ μ') : + complexMGF X μ = complexMGF Y μ' := by + rw [← complexMGF_id_map h.aemeasurable_fst, ← complexMGF_id_map h.aemeasurable_snd, h.map_eq] + +lemma abs_complexMGF_le_mgf : abs (complexMGF X μ z) ≤ mgf X μ z.re := by + rw [complexMGF, ← re_add_im z] + simp_rw [add_mul, Complex.exp_add, re_add_im, ← Complex.norm_eq_abs] + calc ‖∫ ω, cexp (z.re * X ω) * cexp (z.im * I * X ω) ∂μ‖ + _ ≤ ∫ ω, ‖cexp (z.re * X ω) * cexp (z.im * I * X ω)‖ ∂μ := norm_integral_le_integral_norm _ + _ = ∫ ω, rexp (z.re * X ω) ∂μ := by simp [Complex.abs_exp] + +lemma complexMGF_ofReal (x : ℝ) : complexMGF X μ x = mgf X μ x := by + rw [complexMGF, mgf, ← integral_complex_ofReal] + norm_cast + +lemma re_complexMGF_ofReal (x : ℝ) : (complexMGF X μ x).re = mgf X μ x := by + simp [complexMGF_ofReal] + +lemma re_complexMGF_ofReal' : (fun x : ℝ ↦ (complexMGF X μ x).re) = mgf X μ := by + ext x + exact re_complexMGF_ofReal x + +section Analytic + +/-- For `z : ℂ` with `z.re ∈ interior (integrableExpSet X μ)`, the derivative of the function +`z' ↦ μ[X ^ n * cexp (z' * X)]` at `z` is `μ[X ^ (n + 1) * cexp (z * X)]`. -/ +lemma hasDerivAt_integral_pow_mul_exp (hz : z.re ∈ interior (integrableExpSet X μ)) (n : ℕ) : + HasDerivAt (fun z ↦ μ[fun ω ↦ X ω ^ n * cexp (z * X ω)]) + μ[fun ω ↦ X ω ^ (n + 1) * cexp (z * X ω)] z := by + have hX : AEMeasurable X μ := aemeasurable_of_mem_interior_integrableExpSet hz + have hz' := hz + rw [mem_interior_iff_mem_nhds, mem_nhds_iff_exists_Ioo_subset] at hz' + obtain ⟨l, u, hlu, h_subset⟩ := hz' + let t := ((z.re - l) ⊓ (u - z.re)) / 2 + have h_pos : 0 < (z.re - l) ⊓ (u - z.re) := by simp [hlu.1, hlu.2] + have ht : 0 < t := half_pos h_pos + refine (hasDerivAt_integral_of_dominated_loc_of_deriv_le + (bound := fun ω ↦ |X ω| ^ (n + 1) * rexp (z.re * X ω + t/2 * |X ω|)) + (F := fun z ω ↦ X ω ^ n * cexp (z * X ω)) + (F' := fun z ω ↦ X ω ^ (n + 1) * cexp (z * X ω)) (half_pos ht) ?_ ?_ ?_ ?_ ?_ ?_).2 + · exact .of_forall fun z ↦ AEMeasurable.aestronglyMeasurable (by fun_prop) + · exact integrable_pow_mul_cexp_of_re_mem_interior_integrableExpSet hz n + · exact AEMeasurable.aestronglyMeasurable (by fun_prop) + · refine ae_of_all _ fun ω ε hε ↦ ?_ + simp only [norm_mul, norm_pow, norm_real, Real.norm_eq_abs, Complex.norm_eq_abs] + rw [Complex.abs_ofReal, Complex.abs_exp] + simp only [mul_re, ofReal_re, ofReal_im, mul_zero, sub_zero] + gcongr + have : ε = z + (ε - z) := by simp + rw [this, add_re, add_mul] + gcongr _ + ?_ + refine (le_abs_self _).trans ?_ + rw [abs_mul] + gcongr + refine (abs_re_le_abs _).trans ?_ + simp only [Metric.mem_ball, dist_eq_norm, Complex.norm_eq_abs] at hε + exact hε.le + · refine integrable_pow_abs_mul_exp_add_of_integrable_exp_mul ?_ ?_ ?_ ?_ (t := t) (n + 1) + · exact h_subset (add_half_inf_sub_mem_Ioo hlu) + · exact h_subset (sub_half_inf_sub_mem_Ioo hlu) + · positivity + · exact lt_of_lt_of_le (by simp [ht]) (le_abs_self _) + · refine ae_of_all _ fun ω ε hε ↦ ?_ + simp only + simp_rw [pow_succ, mul_assoc] + refine HasDerivAt.const_mul _ ?_ + simp_rw [← smul_eq_mul, Complex.exp_eq_exp_ℂ] + convert hasDerivAt_exp_smul_const (X ω : ℂ) ε using 1 + rw [smul_eq_mul, mul_comm] + +/-- For all `z : ℂ` with `z.re ∈ interior (integrableExpSet X μ)`, +`complexMGF X μ` is differentiable at `z` with derivative `μ[X * exp (z * X)]`. -/ +theorem hasDerivAt_complexMGF (hz : z.re ∈ interior (integrableExpSet X μ)) : + HasDerivAt (complexMGF X μ) μ[fun ω ↦ X ω * cexp (z * X ω)] z := by + convert hasDerivAt_integral_pow_mul_exp hz 0 + · simp [complexMGF] + · simp + +/-- `complexMGF X μ` is holomorphic on the vertical strip +`{z | z.re ∈ interior (integrableExpSet X μ)}`. -/ +theorem differentiableOn_complexMGF : + DifferentiableOn ℂ (complexMGF X μ) {z | z.re ∈ interior (integrableExpSet X μ)} := by + intro z hz + have h := hasDerivAt_complexMGF hz + rw [hasDerivAt_iff_hasFDerivAt] at h + exact h.hasFDerivWithinAt.differentiableWithinAt + +theorem analyticOnNhd_complexMGF : + AnalyticOnNhd ℂ (complexMGF X μ) {z | z.re ∈ interior (integrableExpSet X μ)} := + differentiableOn_complexMGF.analyticOnNhd (isOpen_interior.preimage Complex.continuous_re) + +/-- `complexMGF X μ` is analytic on the vertical strip + `{z | z.re ∈ interior (integrableExpSet X μ)}`. -/ +theorem analyticOn_complexMGF : + AnalyticOn ℂ (complexMGF X μ) {z | z.re ∈ interior (integrableExpSet X μ)} := + analyticOnNhd_complexMGF.analyticOn + +/-- `complexMGF X μ` is analytic at any point `z` with `z.re ∈ interior (integrableExpSet X μ)`. -/ +lemma analyticAt_complexMGF (hz : z.re ∈ interior (integrableExpSet X μ)) : + AnalyticAt ℂ (complexMGF X μ) z := + analyticOnNhd_complexMGF z hz + +end Analytic + +section Deriv + +/-! ### Iterated derivatives of `complexMGF` -/ + +lemma hasDerivAt_iteratedDeriv_complexMGF (hz : z.re ∈ interior (integrableExpSet X μ)) (n : ℕ) : + HasDerivAt (iteratedDeriv n (complexMGF X μ)) μ[fun ω ↦ X ω ^ (n + 1) * cexp (z * X ω)] z := by + induction n generalizing z with + | zero => simp [hasDerivAt_complexMGF hz] + | succ n hn => + rw [iteratedDeriv_succ] + have : deriv (iteratedDeriv n (complexMGF X μ)) + =ᶠ[𝓝 z] fun z ↦ μ[fun ω ↦ X ω ^ (n + 1) * cexp (z * X ω)] := by + have h_mem : ∀ᶠ y in 𝓝 z, y.re ∈ interior (integrableExpSet X μ) := by + refine IsOpen.eventually_mem ?_ hz + exact isOpen_interior.preimage Complex.continuous_re + filter_upwards [h_mem] with y hy using HasDerivAt.deriv (hn hy) + rw [EventuallyEq.hasDerivAt_iff this] + exact hasDerivAt_integral_pow_mul_exp hz (n + 1) + +/-- For `z : ℂ` with `z.re ∈ interior (integrableExpSet X μ)`, the n-th derivative of the function +`complexMGF X μ` at `z` is `μ[X ^ n * cexp (z * X)]`. -/ +lemma iteratedDeriv_complexMGF (hz : z.re ∈ interior (integrableExpSet X μ)) (n : ℕ) : + iteratedDeriv n (complexMGF X μ) z = μ[fun ω ↦ X ω ^ n * cexp (z * X ω)] := by + induction n generalizing z with + | zero => simp [complexMGF] + | succ n hn => + rw [iteratedDeriv_succ] + exact (hasDerivAt_iteratedDeriv_complexMGF hz n).deriv + +end Deriv + +section EqOfMGF + +/-! We prove that if two random variables have the same `mgf`, then +they also have the same `complexMGF`.-/ + +variable {Ω' : Type*} {mΩ' : MeasurableSpace Ω'} {Y : Ω' → ℝ} {μ' : Measure Ω'} + +/-- If two random variables have the same moment generating function then they have +the same `integrableExpSet`. -/ +lemma integrableExpSet_eq_of_mgf' (hXY : mgf X μ = mgf Y μ') (hμμ' : μ = 0 ↔ μ' = 0) : + integrableExpSet X μ = integrableExpSet Y μ' := by + ext t + simp only [integrableExpSet, Set.mem_setOf_eq] + by_cases hμ : μ = 0 + · simp [hμ, hμμ'.mp hμ] + have : NeZero μ := ⟨hμ⟩ + have : NeZero μ' := ⟨(not_iff_not.mpr hμμ').mp hμ⟩ + rw [← mgf_pos_iff, ← mgf_pos_iff, hXY] + +/-- If two random variables have the same moment generating function then they have +the same `integrableExpSet`. -/ +lemma integrableExpSet_eq_of_mgf [IsProbabilityMeasure μ] + (hXY : mgf X μ = mgf Y μ') : + integrableExpSet X μ = integrableExpSet Y μ' := by + refine integrableExpSet_eq_of_mgf' hXY ?_ + simp only [IsProbabilityMeasure.ne_zero, false_iff] + suffices mgf Y μ' 0 ≠ 0 by + intro h_contra + simp [h_contra] at this + rw [← hXY] + exact (mgf_pos (by simp)).ne' + +/-- If two random variables have the same moment generating function then they have +the same `complexMGF` on the vertical strip `{z | z.re ∈ interior (integrableExpSet X μ)}`. + +TODO: once we know that equal `mgf` implies equal distributions, we will be able to show that +the `complexMGF` are equal everywhere, not only on the strip. +This lemma will be used in the proof of the equality of distributions. -/ +lemma eqOn_complexMGF_of_mgf' (hXY : mgf X μ = mgf Y μ') (hμμ' : μ = 0 ↔ μ' = 0) : + Set.EqOn (complexMGF X μ) (complexMGF Y μ') {z | z.re ∈ interior (integrableExpSet X μ)} := by + by_cases h_empty : interior (integrableExpSet X μ) = ∅ + · simp [h_empty] + rw [← ne_eq, ← Set.nonempty_iff_ne_empty] at h_empty + obtain ⟨t, ht⟩ := h_empty + have hX : AnalyticOnNhd ℂ (complexMGF X μ) {z | z.re ∈ interior (integrableExpSet X μ)} := + analyticOnNhd_complexMGF + have hY : AnalyticOnNhd ℂ (complexMGF Y μ') {z | z.re ∈ interior (integrableExpSet Y μ')} := + analyticOnNhd_complexMGF + rw [integrableExpSet_eq_of_mgf' hXY hμμ'] at hX ht ⊢ + refine AnalyticOnNhd.eqOn_of_preconnected_of_frequently_eq hX hY + (convex_integrableExpSet.interior.linear_preimage reLm).isPreconnected + (z₀ := (t : ℂ)) (by simp [ht]) ?_ + have h_real : ∃ᶠ (x : ℝ) in 𝓝[≠] t, complexMGF X μ x = complexMGF Y μ' x := by + refine .of_forall fun y ↦ ?_ + rw [complexMGF_ofReal, complexMGF_ofReal, hXY] + rw [frequently_iff_seq_forall] at h_real ⊢ + obtain ⟨xs, hx_tendsto, hx_eq⟩ := h_real + refine ⟨fun n ↦ xs n, ?_, fun n ↦ ?_⟩ + · rw [tendsto_nhdsWithin_iff] at hx_tendsto ⊢ + constructor + · rw [tendsto_ofReal_iff] + exact hx_tendsto.1 + · simpa using hx_tendsto.2 + · simp [hx_eq] + +/-- If two random variables have the same moment generating function then they have +the same `complexMGF` on the vertical strip `{z | z.re ∈ interior (integrableExpSet X μ)}`. -/ +lemma eqOn_complexMGF_of_mgf [IsProbabilityMeasure μ] + (hXY : mgf X μ = mgf Y μ') : + Set.EqOn (complexMGF X μ) (complexMGF Y μ') {z | z.re ∈ interior (integrableExpSet X μ)} := by + refine eqOn_complexMGF_of_mgf' hXY ?_ + simp only [IsProbabilityMeasure.ne_zero, false_iff] + suffices mgf Y μ' 0 ≠ 0 by + intro h_contra + simp [h_contra] at this + rw [← hXY] + exact (mgf_pos (by simp)).ne' + +end EqOfMGF + +end ProbabilityTheory diff --git a/Mathlib/Probability/Moments/IntegrableExpMul.lean b/Mathlib/Probability/Moments/IntegrableExpMul.lean index f2135d63a157ed..74d8ee95ca9403 100644 --- a/Mathlib/Probability/Moments/IntegrableExpMul.lean +++ b/Mathlib/Probability/Moments/IntegrableExpMul.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.MeasureTheory.Function.L1Space +import Mathlib.MeasureTheory.Function.L1Space.Integrable import Mathlib.MeasureTheory.Order.Group.Lattice /-! @@ -539,6 +539,63 @@ lemma memℒp_of_mem_interior_integrableExpSet (h : 0 ∈ interior (integrableEx simp only [norm_eq_abs, ENNReal.coe_toReal] exact integrable_rpow_abs_of_mem_interior_integrableExpSet h p.2 +section Complex + +open Complex + +variable {z : ℂ} + +lemma integrable_cexp_mul_of_re_mem_integrableExpSet (hX : AEMeasurable X μ) + (hz : z.re ∈ integrableExpSet X μ) : + Integrable (fun ω ↦ cexp (z * X ω)) μ := by + rw [← integrable_norm_iff] + · simpa [Complex.norm_eq_abs, Complex.abs_exp] using hz + · exact AEMeasurable.aestronglyMeasurable (by fun_prop) + +lemma integrable_cexp_mul_of_re_mem_interior_integrableExpSet + (hz : z.re ∈ interior (integrableExpSet X μ)) : + Integrable (fun ω ↦ cexp (z * X ω)) μ := + integrable_cexp_mul_of_re_mem_integrableExpSet + (aemeasurable_of_mem_interior_integrableExpSet hz) (interior_subset hz) + +lemma integrable_rpow_abs_mul_cexp_of_re_mem_interior_integrableExpSet + (hz : z.re ∈ interior (integrableExpSet X μ)) {p : ℝ} (hp : 0 ≤ p) : + Integrable (fun ω ↦ (|X ω| ^ p : ℝ) * cexp (z * X ω)) μ := by + have hX : AEMeasurable X μ := aemeasurable_of_mem_interior_integrableExpSet hz + rw [← integrable_norm_iff] + swap; · exact AEMeasurable.aestronglyMeasurable (by fun_prop) + simpa [abs_rpow_of_nonneg (abs_nonneg _), Complex.abs_exp] + using integrable_rpow_abs_mul_exp_of_mem_interior_integrableExpSet hz hp + +lemma integrable_pow_abs_mul_cexp_of_re_mem_interior_integrableExpSet + (hz : z.re ∈ interior (integrableExpSet X μ)) (n : ℕ) : + Integrable (fun ω ↦ |X ω| ^ n * cexp (z * X ω)) μ := by + convert integrable_rpow_abs_mul_cexp_of_re_mem_interior_integrableExpSet hz (Nat.cast_nonneg n) + simp + +lemma integrable_rpow_mul_cexp_of_re_mem_interior_integrableExpSet + (hz : z.re ∈ interior (integrableExpSet X μ)) {p : ℝ} (hp : 0 ≤ p) : + Integrable (fun ω ↦ (X ω ^ p : ℝ) * cexp (z * X ω)) μ := by + have hX : AEMeasurable X μ := aemeasurable_of_mem_interior_integrableExpSet hz + rw [← integrable_norm_iff] + swap; · exact AEMeasurable.aestronglyMeasurable (by fun_prop) + simp only [norm_mul, norm_real, Real.norm_eq_abs, Complex.norm_eq_abs, Complex.abs_exp, mul_re, + ofReal_re, ofReal_im, mul_zero, sub_zero, Complex.abs_ofReal] + refine (integrable_rpow_abs_mul_exp_of_mem_interior_integrableExpSet hz hp).mono ?_ ?_ + · exact AEMeasurable.aestronglyMeasurable (by fun_prop) + refine ae_of_all _ fun ω ↦ ?_ + simp only [norm_mul, Real.norm_eq_abs, abs_abs, Real.abs_exp] + gcongr + exact abs_rpow_le_abs_rpow _ _ + +lemma integrable_pow_mul_cexp_of_re_mem_interior_integrableExpSet + (hz : z.re ∈ interior (integrableExpSet X μ)) (n : ℕ) : + Integrable (fun ω ↦ X ω ^ n * cexp (z * X ω)) μ := by + convert integrable_rpow_mul_cexp_of_re_mem_interior_integrableExpSet hz (Nat.cast_nonneg n) + simp + +end Complex + end IntegrableExpSet end FiniteMoments diff --git a/Mathlib/Probability/Moments/MGFAnalytic.lean b/Mathlib/Probability/Moments/MGFAnalytic.lean new file mode 100644 index 00000000000000..57f5d8242d2e64 --- /dev/null +++ b/Mathlib/Probability/Moments/MGFAnalytic.lean @@ -0,0 +1,124 @@ +/- +Copyright (c) 2025 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Probability.Moments.ComplexMGF +import Mathlib.Analysis.SpecialFunctions.Complex.Analytic + +/-! +# The moment generating function is analytic + +The moment generating function `mgf X μ` of a random variable `X` with respect to a measure `μ` +is analytic on the interior of `integrableExpSet X μ`, the interval on which it is defined. + +## Main results + +* `analyticOn_mgf`: the moment generating function is analytic on the interior of the interval + on which it is defined. +* `iteratedDeriv_mgf`: the n-th derivative of the mgf at `t` is `μ[X ^ n * exp (t * X)]`. + +-/ + + +open MeasureTheory Filter Finset Real + +open scoped MeasureTheory ProbabilityTheory ENNReal NNReal Topology Nat + +namespace ProbabilityTheory + +variable {Ω ι : Type*} {m : MeasurableSpace Ω} {X : Ω → ℝ} {μ : Measure Ω} {t u v : ℝ} + +section Deriv + +/-- For `t : ℝ` with `t ∈ interior (integrableExpSet X μ)`, the derivative of the function +`x ↦ μ[X ^ n * rexp (x * X)]` at `t` is `μ[X ^ (n + 1) * rexp (t * X)]`. -/ +lemma hasDerivAt_integral_pow_mul_exp_real (ht : t ∈ interior (integrableExpSet X μ)) (n : ℕ) : + HasDerivAt (fun t ↦ μ[fun ω ↦ X ω ^ n * rexp (t * X ω)]) + μ[fun ω ↦ X ω ^ (n + 1) * rexp (t * X ω)] t := by + have h_re_of_mem n t (ht' : t ∈ interior (integrableExpSet X μ)) : + (∫ ω, X ω ^ n * Complex.exp (t * X ω) ∂μ).re = ∫ ω, X ω ^ n * rexp (t * X ω) ∂μ := by + rw [← RCLike.re_eq_complex_re, ← integral_re] + · norm_cast + · refine integrable_pow_mul_cexp_of_re_mem_interior_integrableExpSet ?_ n + simpa using ht' + have h_re n : ∀ᶠ t' : ℝ in 𝓝 t, (∫ ω, X ω ^ n * Complex.exp (t' * X ω) ∂μ).re + = ∫ ω, X ω ^ n * rexp (t' * X ω) ∂μ := by + filter_upwards [isOpen_interior.eventually_mem ht] with t ht' using h_re_of_mem n t ht' + rw [← EventuallyEq.hasDerivAt_iff (h_re _), ← h_re_of_mem _ t ht] + exact (hasDerivAt_integral_pow_mul_exp (by simp [ht]) n).real_of_complex + +/-- For `t ∈ interior (integrableExpSet X μ)`, the derivative of `mgf X μ` at `t` is +`μ[X * exp (t * X)]`. -/ +lemma hasDerivAt_mgf (h : t ∈ interior (integrableExpSet X μ)) : + HasDerivAt (mgf X μ) (μ[fun ω ↦ X ω * exp (t * X ω)]) t := by + convert hasDerivAt_integral_pow_mul_exp_real h 0 + · simp [mgf] + · simp + +lemma hasDerivAt_iteratedDeriv_mgf (ht : t ∈ interior (integrableExpSet X μ)) (n : ℕ) : + HasDerivAt (iteratedDeriv n (mgf X μ)) μ[fun ω ↦ X ω ^ (n + 1) * exp (t * X ω)] t := by + induction n generalizing t with + | zero => simp [hasDerivAt_mgf ht] + | succ n hn => + rw [iteratedDeriv_succ] + have : deriv (iteratedDeriv n (mgf X μ)) + =ᶠ[𝓝 t] fun t ↦ μ[fun ω ↦ X ω ^ (n + 1) * exp (t * X ω)] := by + have h_mem : ∀ᶠ y in 𝓝 t, y ∈ interior (integrableExpSet X μ) := + isOpen_interior.eventually_mem ht + filter_upwards [h_mem] with y hy using HasDerivAt.deriv (hn hy) + rw [EventuallyEq.hasDerivAt_iff this] + exact hasDerivAt_integral_pow_mul_exp_real ht (n + 1) + +/-- For `t ∈ interior (integrableExpSet X μ)`, the n-th derivative of `mgf X μ` at `t` is +`μ[X ^ n * exp (t * X)]`. -/ +lemma iteratedDeriv_mgf (ht : t ∈ interior (integrableExpSet X μ)) (n : ℕ) : + iteratedDeriv n (mgf X μ) t = μ[fun ω ↦ X ω ^ n * exp (t * X ω)] := by + induction n generalizing t with + | zero => simp [mgf] + | succ n hn => + rw [iteratedDeriv_succ] + exact (hasDerivAt_iteratedDeriv_mgf ht n).deriv + +/-- The derivatives of the moment generating function at zero are the moments. -/ +lemma iteratedDeriv_mgf_zero (h : 0 ∈ interior (integrableExpSet X μ)) (n : ℕ) : + iteratedDeriv n (mgf X μ) 0 = μ[X ^ n] := by + simp [iteratedDeriv_mgf h n] + +/-- For `t ∈ interior (integrableExpSet X μ)`, the derivative of `mgf X μ` at `t` is +`μ[X * exp (t * X)]`. -/ +lemma deriv_mgf (h : t ∈ interior (integrableExpSet X μ)) : + deriv (mgf X μ) t = μ[fun ω ↦ X ω * exp (t * X ω)] := + (hasDerivAt_mgf h).deriv + +lemma deriv_mgf_zero (h : 0 ∈ interior (integrableExpSet X μ)) : deriv (mgf X μ) 0 = μ[X] := by + simp [deriv_mgf h] + +end Deriv + +section Analytic + +/-- The moment generating function is analytic at every `t ∈ interior (integrableExpSet X μ)`. -/ +lemma analyticAt_mgf (ht : t ∈ interior (integrableExpSet X μ)) : + AnalyticAt ℝ (mgf X μ) t := by + rw [← re_complexMGF_ofReal'] + exact (analyticAt_complexMGF (by simp [ht])).re_ofReal + +lemma analyticOnNhd_mgf : AnalyticOnNhd ℝ (mgf X μ) (interior (integrableExpSet X μ)) := + fun _ hx ↦ analyticAt_mgf hx + +/-- The moment generating function is analytic on the interior of the interval on which it is +defined. -/ +lemma analyticOn_mgf : AnalyticOn ℝ (mgf X μ) (interior (integrableExpSet X μ)) := + analyticOnNhd_mgf.analyticOn + +lemma hasFPowerSeriesAt_mgf (hv : v ∈ interior (integrableExpSet X μ)) : + HasFPowerSeriesAt (mgf X μ) + (FormalMultilinearSeries.ofScalars ℝ + (fun n ↦ (μ[fun ω ↦ X ω ^ n * exp (v * X ω)] : ℝ) / n !)) v := by + convert (analyticAt_mgf hv).hasFPowerSeriesAt + rw [iteratedDeriv_mgf hv] + +end Analytic + +end ProbabilityTheory diff --git a/Mathlib/Probability/Process/Filtration.lean b/Mathlib/Probability/Process/Filtration.lean index 82b1ace233b965..1c4556148126c7 100644 --- a/Mathlib/Probability/Process/Filtration.lean +++ b/Mathlib/Probability/Process/Filtration.lean @@ -33,7 +33,7 @@ filtration, stochastic process open Filter Order TopologicalSpace -open scoped Classical MeasureTheory NNReal ENNReal Topology +open scoped MeasureTheory NNReal ENNReal Topology namespace MeasureTheory @@ -133,6 +133,7 @@ theorem sSup_def (s : Set (Filtration ι m)) (i : ι) : sSup s i = sSup ((fun f : Filtration ι m => f i) '' s) := rfl +open scoped Classical in noncomputable instance : InfSet (Filtration ι m) := ⟨fun s => { seq := fun i => if Set.Nonempty s then sInf ((fun f : Filtration ι m => f i) '' s) else m @@ -151,6 +152,7 @@ noncomputable instance : InfSet (Filtration ι m) := obtain ⟨f, hf_mem⟩ := h_nonempty exact le_trans (sInf_le ⟨f, hf_mem, rfl⟩) (f.le i) }⟩ +open scoped Classical in theorem sInf_def (s : Set (Filtration ι m)) (i : ι) : sInf s i = if Set.Nonempty s then sInf ((fun f : Filtration ι m => f i) '' s) else m := rfl @@ -295,6 +297,7 @@ section Limit variable {E : Type*} [Zero E] [TopologicalSpace E] {ℱ : Filtration ι m} {f : ι → Ω → E} {μ : Measure Ω} +open scoped Classical in /-- Given a process `f` and a filtration `ℱ`, if `f` converges to some `g` almost everywhere and `g` is `⨆ n, ℱ n`-measurable, then `limitProcess f ℱ μ` chooses said `g`, else it returns 0. diff --git a/Mathlib/Probability/Process/HittingTime.lean b/Mathlib/Probability/Process/HittingTime.lean index 5ac75975d64832..7610ba2b87de1d 100644 --- a/Mathlib/Probability/Process/HittingTime.lean +++ b/Mathlib/Probability/Process/HittingTime.lean @@ -53,7 +53,6 @@ noncomputable def hitting [Preorder ι] [InfSet ι] (u : ι → Ω → β) fun x => if ∃ j ∈ Set.Icc n m, u j x ∈ s then sInf (Set.Icc n m ∩ {i : ι | u i x ∈ s}) else m -#adaptation_note /-- nightly-2024-03-16: added to replace simp [hitting] -/ open scoped Classical in theorem hitting_def [Preorder ι] [InfSet ι] (u : ι → Ω → β) (s : Set β) (n m : ι) : hitting u s n m = diff --git a/Mathlib/Probability/Variance.lean b/Mathlib/Probability/Variance.lean index 5c3c9f6c42be48..886ceef0e7ebe4 100644 --- a/Mathlib/Probability/Variance.lean +++ b/Mathlib/Probability/Variance.lean @@ -49,7 +49,7 @@ namespace ProbabilityTheory /-- The `ℝ≥0∞`-valued variance of a real-valued random variable defined as the Lebesgue integral of `(X - 𝔼[X])^2`. -/ def evariance {Ω : Type*} {_ : MeasurableSpace Ω} (X : Ω → ℝ) (μ : Measure Ω) : ℝ≥0∞ := - ∫⁻ ω, (‖X ω - μ[X]‖₊ : ℝ≥0∞) ^ 2 ∂μ + ∫⁻ ω, ‖X ω - μ[X]‖ₑ ^ 2 ∂μ /-- The `ℝ`-valued variance of a real-valued random variable defined by applying `ENNReal.toReal` to `evariance`. -/ @@ -82,7 +82,7 @@ variable {Ω : Type*} {m : MeasurableSpace Ω} {X : Ω → ℝ} {μ : Measure Ω theorem _root_.MeasureTheory.Memℒp.evariance_lt_top [IsFiniteMeasure μ] (hX : Memℒp X 2 μ) : evariance X μ < ∞ := by have := ENNReal.pow_lt_top (hX.sub <| memℒp_const <| μ[X]).2 2 - rw [eLpNorm_eq_lintegral_rpow_nnnorm two_ne_zero ENNReal.two_ne_top, ← ENNReal.rpow_two] at this + rw [eLpNorm_eq_lintegral_rpow_enorm two_ne_zero ENNReal.two_ne_top, ← ENNReal.rpow_two] at this simp only [ENNReal.toReal_ofNat, Pi.sub_apply, ENNReal.one_toReal, one_div] at this rw [← ENNReal.rpow_mul, inv_mul_cancel₀ (two_ne_zero : (2 : ℝ) ≠ 0), ENNReal.rpow_one] at this simp_rw [ENNReal.rpow_two] at this @@ -94,7 +94,7 @@ theorem evariance_eq_top [IsFiniteMeasure μ] (hXm : AEStronglyMeasurable X μ) rw [← Ne, ← lt_top_iff_ne_top] at h have : Memℒp (fun ω => X ω - μ[X]) 2 μ := by refine ⟨hXm.sub aestronglyMeasurable_const, ?_⟩ - rw [eLpNorm_eq_lintegral_rpow_nnnorm two_ne_zero ENNReal.two_ne_top] + rw [eLpNorm_eq_lintegral_rpow_enorm two_ne_zero ENNReal.two_ne_top] simp only [ENNReal.toReal_ofNat, ENNReal.one_toReal, ENNReal.rpow_two, Ne] exact ENNReal.rpow_lt_top_of_nonneg (by linarith) h.ne refine hX ?_ @@ -117,11 +117,8 @@ theorem _root_.MeasureTheory.Memℒp.ofReal_variance_eq [IsFiniteMeasure μ] (hX theorem evariance_eq_lintegral_ofReal (X : Ω → ℝ) (μ : Measure Ω) : evariance X μ = ∫⁻ ω, ENNReal.ofReal ((X ω - μ[X]) ^ 2) ∂μ := by rw [evariance] - congr - ext1 ω - rw [pow_two, ← ENNReal.coe_mul, ← nnnorm_mul, ← pow_two] - congr - exact (Real.toNNReal_eq_nnnorm_of_nonneg <| sq_nonneg _).symm + congr with ω + rw [pow_two, ← enorm_mul, ← pow_two, Real.enorm_of_nonneg (sq_nonneg _)] theorem _root_.MeasureTheory.Memℒp.variance_eq_of_integral_eq_zero (hX : Memℒp X 2 μ) (hXint : μ[X] = 0) : variance X μ = μ[X ^ (2 : Nat)] := by @@ -152,25 +149,16 @@ theorem evariance_eq_zero_iff (hX : AEMeasurable X μ) : evariance X μ = 0 ↔ X =ᵐ[μ] fun _ => μ[X] := by rw [evariance, lintegral_eq_zero_iff'] constructor <;> intro hX <;> filter_upwards [hX] with ω hω - · simpa only [Pi.zero_apply, sq_eq_zero_iff, ENNReal.coe_eq_zero, nnnorm_eq_zero, sub_eq_zero] - using hω - · rw [hω] - simp - · exact (hX.sub_const _).ennnorm.pow_const _ -- TODO `measurability` and `fun_prop` fail + · simpa [sub_eq_zero] using hω + · simp [hω] + · exact (hX.sub_const _).enorm.pow_const _ -- TODO `measurability` and `fun_prop` fail theorem evariance_mul (c : ℝ) (X : Ω → ℝ) (μ : Measure Ω) : evariance (fun ω => c * X ω) μ = ENNReal.ofReal (c ^ 2) * evariance X μ := by rw [evariance, evariance, ← lintegral_const_mul' _ _ ENNReal.ofReal_lt_top.ne] - congr - ext1 ω - rw [ENNReal.ofReal, ← ENNReal.coe_pow, ← ENNReal.coe_pow, ← ENNReal.coe_mul] - congr - rw [← sq_abs, ← Real.rpow_two, Real.toNNReal_rpow_of_nonneg (abs_nonneg _), NNReal.rpow_two, - ← mul_pow, Real.toNNReal_mul_nnnorm _ (abs_nonneg _)] - conv_rhs => rw [← nnnorm_norm, norm_mul, norm_abs_eq_norm, ← norm_mul, nnnorm_norm, mul_sub] - congr - rw [mul_comm] - simp_rw [← smul_eq_mul, ← integral_smul_const, smul_eq_mul, mul_comm] + congr with ω + rw [integral_mul_left, ← mul_sub, enorm_mul, mul_pow, ← enorm_pow, + Real.enorm_of_nonneg (sq_nonneg _)] @[simp] theorem variance_zero (μ : Measure Ω) : variance 0 μ = 0 := by @@ -228,20 +216,20 @@ theorem variance_le_expectation_sq [IsProbabilityMeasure μ] {X : Ω → ℝ} · exact (AEMeasurable.pow_const (hm.aemeasurable.sub_const _) _).aestronglyMeasurable theorem evariance_def' [IsProbabilityMeasure μ] {X : Ω → ℝ} (hX : AEStronglyMeasurable X μ) : - evariance X μ = (∫⁻ ω, (‖X ω‖₊ ^ 2 :) ∂μ) - ENNReal.ofReal (μ[X] ^ 2) := by + evariance X μ = (∫⁻ ω, ‖X ω‖ₑ ^ 2 ∂μ) - ENNReal.ofReal (μ[X] ^ 2) := by by_cases hℒ : Memℒp X 2 μ · rw [← hℒ.ofReal_variance_eq, variance_def' hℒ, ENNReal.ofReal_sub _ (sq_nonneg _)] congr + simp_rw [← enorm_pow, enorm] rw [lintegral_coe_eq_integral] - · congr 2 with ω - simp only [Pi.pow_apply, NNReal.coe_pow, coe_nnnorm, Real.norm_eq_abs, Even.pow_abs even_two] - · exact hℒ.abs.integrable_sq + · simp + · simpa using hℒ.abs.integrable_sq · symm rw [evariance_eq_top hX hℒ, ENNReal.sub_eq_top_iff] refine ⟨?_, ENNReal.ofReal_ne_top⟩ rw [Memℒp, not_and] at hℒ specialize hℒ hX - simp only [eLpNorm_eq_lintegral_rpow_nnnorm two_ne_zero ENNReal.two_ne_top, not_lt, top_le_iff, + simp only [eLpNorm_eq_lintegral_rpow_enorm two_ne_zero ENNReal.two_ne_top, not_lt, top_le_iff, ENNReal.toReal_ofNat, one_div, ENNReal.rpow_eq_top_iff, inv_lt_zero, inv_pos, and_true, or_iff_not_imp_left, not_and_or, zero_lt_two] at hℒ exact mod_cast hℒ fun _ => zero_le_two @@ -255,7 +243,7 @@ theorem meas_ge_le_evariance_div_sq {X : Ω → ℝ} (hX : AEStronglyMeasurable · congr simp only [Pi.sub_apply, ENNReal.coe_le_coe, ← Real.norm_eq_abs, ← coe_nnnorm, NNReal.coe_le_coe, ENNReal.ofReal_coe_nnreal] - · rw [eLpNorm_eq_lintegral_rpow_nnnorm two_ne_zero ENNReal.two_ne_top] + · rw [eLpNorm_eq_lintegral_rpow_enorm two_ne_zero ENNReal.two_ne_top] simp only [show ENNReal.ofNNReal (c ^ 2) = (ENNReal.ofNNReal c) ^ 2 by norm_cast, ENNReal.toReal_ofNat, one_div, Pi.sub_apply] rw [div_eq_mul_inv, ENNReal.inv_pow, mul_comm, ENNReal.rpow_two] @@ -263,6 +251,7 @@ theorem meas_ge_le_evariance_div_sq {X : Ω → ℝ} (hX : AEStronglyMeasurable simp_rw [← ENNReal.rpow_mul, inv_mul_cancel₀ (two_ne_zero : (2 : ℝ) ≠ 0), ENNReal.rpow_two, ENNReal.rpow_one, evariance] + /-- **Chebyshev's inequality**: one can control the deviation probability of a real random variable from its expectation in terms of the variance. -/ theorem meas_ge_le_variance_div_sq [IsFiniteMeasure μ] {X : Ω → ℝ} (hX : Memℒp X 2 μ) {c : ℝ} diff --git a/Mathlib/RepresentationTheory/Character.lean b/Mathlib/RepresentationTheory/Character.lean index ada933887ae231..6dc6591865ab82 100644 --- a/Mathlib/RepresentationTheory/Character.lean +++ b/Mathlib/RepresentationTheory/Character.lean @@ -115,7 +115,7 @@ theorem char_orthonormal (V W : FDRep k G) [Simple V] [Simple W] : rw [char_iso (FDRep.dualTensorIsoLinHom W.ρ V)] -- The average over the group of the character of a representation equals the dimension of the -- space of invariants. - rw [average_char_eq_finrank_invariants, ← FDRep.endMulEquiv_comp_ρ (of _), + rw [average_char_eq_finrank_invariants, ← FDRep.endRingEquiv_comp_ρ (of _), FDRep.of_ρ (linHom W.ρ V.ρ)] -- The space of invariants of `Hom(W, V)` is the subspace of `G`-equivariant linear maps, -- `Hom_G(W, V)`. diff --git a/Mathlib/RepresentationTheory/FDRep.lean b/Mathlib/RepresentationTheory/FDRep.lean index dd5870c18de169..dae66e0f2aa478 100644 --- a/Mathlib/RepresentationTheory/FDRep.lean +++ b/Mathlib/RepresentationTheory/FDRep.lean @@ -82,15 +82,16 @@ instance (V W : FDRep k G) : FiniteDimensional k (V ⟶ W) := /-- The monoid homomorphism corresponding to the action of `G` onto `V : FDRep k G`. -/ def ρ (V : FDRep k G) : G →* V →ₗ[k] V := - (ModuleCat.endMulEquiv _).toMonoidHom.comp (Action.ρ V) + (ModuleCat.endRingEquiv _).toMonoidHom.comp (Action.ρ V) @[simp] -lemma endMulEquiv_symm_comp_ρ (V : FDRep k G) : - (MonoidHomClass.toMonoidHom (ModuleCat.endMulEquiv V.V.obj).symm).comp (ρ V) = Action.ρ V := rfl +lemma endRingEquiv_symm_comp_ρ (V : FDRep k G) : + (MonoidHomClass.toMonoidHom (ModuleCat.endRingEquiv V.V.obj).symm).comp (ρ V) = Action.ρ V := + rfl @[simp] -lemma endMulEquiv_comp_ρ (V : FDRep k G) : - (MonoidHomClass.toMonoidHom (ModuleCat.endMulEquiv V.V.obj)).comp (Action.ρ V) = ρ V := rfl +lemma endRingEquiv_comp_ρ (V : FDRep k G) : + (MonoidHomClass.toMonoidHom (ModuleCat.endRingEquiv V.V.obj)).comp (Action.ρ V) = ρ V := rfl @[simp] lemma hom_action_ρ (V : FDRep k G) (g : G) : (Action.ρ V g).hom = ρ V g := rfl @@ -111,7 +112,7 @@ theorem Iso.conj_ρ {V W : FDRep k G} (i : V ≅ W) (g : G) : @[simps ρ] def of {V : Type u} [AddCommGroup V] [Module k V] [FiniteDimensional k V] (ρ : Representation k G V) : FDRep k G := - ⟨FGModuleCat.of k V, ρ ≫ MonCat.ofHom (ModuleCat.endMulEquiv _).symm.toMonoidHom⟩ + ⟨FGModuleCat.of k V, ρ ≫ MonCat.ofHom (ModuleCat.endRingEquiv _).symm.toMonoidHom⟩ instance : HasForget₂ (FDRep k G) (Rep k G) where forget₂ := (forget₂ (FGModuleCat k) (ModuleCat k)).mapAction (MonCat.of G) diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean index 756771dc9da34d..c620c7d2c03c88 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean @@ -146,8 +146,7 @@ and the homogeneous `linearYonedaObjResolution`. -/ -- https://github.com/leanprover-community/mathlib4/issues/5164 change d n A f g = diagonalHomEquiv (n + 1) A ((resolution k G).d (n + 1) n ≫ (diagonalHomEquiv n A).symm f) g - rw [diagonalHomEquiv_apply, Action.comp_hom, ModuleCat.hom_comp, LinearMap.comp_apply, - resolution.d_eq] + rw [diagonalHomEquiv_apply, Action.comp_hom, ConcreteCategory.comp_apply, resolution.d_eq] erw [resolution.d_of (Fin.partialProd g)] simp only [map_sum, ← Finsupp.smul_single_one _ ((-1 : k) ^ _)] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 @@ -188,8 +187,9 @@ noncomputable abbrev inhomogeneousCochains : CochainComplex (ModuleCat k) ℕ := ((linearYonedaObjResolution A).d_comp_d n (n + 1) (n + 2))) simp only [ModuleCat.hom_comp, LinearMap.comp_apply] at this dsimp only - simp only [d_eq, LinearEquiv.toModuleIso_inv_hom, LinearEquiv.toModuleIso_hom_hom, - ModuleCat.hom_comp, LinearMap.comp_apply, LinearEquiv.coe_coe, ModuleCat.hom_zero] + simp only [d_eq, LinearEquiv.toModuleIso_inv, LinearEquiv.toModuleIso_hom, + ModuleCat.hom_comp, LinearMap.comp_apply, LinearEquiv.coe_coe, ModuleCat.hom_zero, + ModuleCat.ofHom_comp] /- Porting note: I can see I need to rewrite `LinearEquiv.coe_coe` twice to at least reduce the need for `symm_apply_apply` to be an `erw`. However, even `erw` refuses to rewrite the second `coe_coe`... -/ @@ -213,11 +213,11 @@ def inhomogeneousCochainsIso : inhomogeneousCochains A ≅ linearYonedaObjResolu ext simp only [ChainComplex.linearYonedaObj_X, linearYoneda_obj_obj_carrier, CochainComplex.of_x, linearYoneda_obj_obj_isAddCommGroup, linearYoneda_obj_obj_isModule, Iso.symm_hom, - ChainComplex.linearYonedaObj_d, ModuleCat.hom_comp, linearYoneda_obj_map_hom, - Quiver.Hom.unop_op, LinearEquiv.toModuleIso_inv_hom, LinearMap.coe_comp, Function.comp_apply, - Linear.leftComp_apply, inhomogeneousCochains.d_def, d_eq, LinearEquiv.toModuleIso_hom_hom, - ModuleCat.ofHom_comp, Category.assoc, LinearEquiv.comp_coe, LinearEquiv.self_trans_symm, - LinearEquiv.refl_toLinearMap, LinearMap.id_comp, LinearEquiv.coe_coe] + LinearEquiv.toModuleIso_inv, ChainComplex.linearYonedaObj_d, ModuleCat.hom_comp, + ModuleCat.hom_ofHom, LinearMap.coe_comp, Function.comp_apply, Linear.leftComp_apply, + inhomogeneousCochains.d_def, d_eq, LinearEquiv.toModuleIso_hom, ModuleCat.ofHom_comp, + Category.assoc, LinearEquiv.comp_coe, LinearEquiv.self_trans_symm, LinearEquiv.refl_toLinearMap, + LinearMap.id_comp, LinearEquiv.coe_coe] rfl /-- The `n`-cocycles `Zⁿ(G, A)` of a `k`-linear `G`-representation `A`, i.e. the kernel of the diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index 1c8844cd09b486..ccb8715cbcaa61 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -205,7 +205,7 @@ theorem dTwo_comp_dOne : dTwo A ∘ₗ dOne A = 0 := by show (ModuleCat.ofHom (dOne A) ≫ ModuleCat.ofHom (dTwo A)).hom = _ have h1 := congr_arg ModuleCat.ofHom (dOne_comp_eq A) have h2 := congr_arg ModuleCat.ofHom (dTwo_comp_eq A) - simp only [ModuleCat.ofHom_comp, ModuleCat.ofHom_comp, ← LinearEquiv.toModuleIso_hom_hom] at h1 h2 + simp only [ModuleCat.ofHom_comp, ModuleCat.ofHom_comp, ← LinearEquiv.toModuleIso_hom] at h1 h2 simp only [(Iso.eq_inv_comp _).2 h2, (Iso.eq_inv_comp _).2 h1, ModuleCat.ofHom_hom, ModuleCat.hom_ofHom, Category.assoc, Iso.hom_inv_id_assoc, HomologicalComplex.d_comp_d_assoc, zero_comp, comp_zero, ModuleCat.hom_zero] diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index b53ab4e4d4058c..ad038c15186b45 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -494,7 +494,7 @@ theorem d_eq (n : ℕ) : ((groupCohomology.resolution k G).d (n + 1) n).hom = /- Porting note: want to rewrite `LinearMap.smul_apply` but simp/simp_rw won't do it; I need erw, so using Finset.sum_congr to get rid of the binder -/ refine Finset.sum_congr rfl fun _ _ => ?_ - simp only [ModuleCat.hom_smul, SimplexCategory.len_mk] + simp only [ModuleCat.hom_smul, SimplexCategory.len_mk, ModuleCat.hom_ofHom] erw [LinearMap.smul_apply] rw [Finsupp.lmapDomain_apply, Finsupp.mapDomain_single, Finsupp.smul_single', mul_one] rfl @@ -560,7 +560,8 @@ theorem forget₂ToModuleCatHomotopyEquiv_f_0_eq : LinearMap.id fun i => rfl, LinearMap.id_comp] rfl - · congr + · rw [ModuleCat.hom_comp] + congr · ext x dsimp (config := { unfoldPartialApp := true }) [HomotopyEquiv.ofIso, Finsupp.LinearEquiv.finsuppUnique] diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 68e82a8670b12a..5e1201f82dfade 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -61,11 +61,11 @@ instance (V : Rep k G) : Module k V := by -/ def ρ (V : Rep k G) : Representation k G V := -- Porting note: was `V.ρ` - (ModuleCat.endMulEquiv V.V).toMonoidHom.comp (Action.ρ V) + (ModuleCat.endRingEquiv V.V).toMonoidHom.comp (Action.ρ V) /-- Lift an unbundled representation to `Rep`. -/ def of {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) : Rep k G := - ⟨ModuleCat.of k V, MonCat.ofHom ((ModuleCat.endMulEquiv _).symm.toMonoidHom.comp ρ) ⟩ + ⟨ModuleCat.of k V, MonCat.ofHom ((ModuleCat.endRingEquiv _).symm.toMonoidHom.comp ρ) ⟩ @[simp] theorem coe_of {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) : @@ -77,7 +77,7 @@ theorem of_ρ {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k rfl theorem Action_ρ_eq_ρ {A : Rep k G} : - Action.ρ A = (ModuleCat.endMulEquiv _).symm.toMonoidHom.comp A.ρ := + Action.ρ A = (ModuleCat.endRingEquiv _).symm.toMonoidHom.comp A.ρ := rfl @[simp] @@ -296,7 +296,7 @@ noncomputable def leftRegularHom (A : Rep k G) (x : A) : Rep.ofMulAction k G G theorem leftRegularHom_apply {A : Rep k G} (x : A) : (leftRegularHom A x).hom (Finsupp.single 1 1) = x := by -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [leftRegularHom_hom_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul, + erw [leftRegularHom_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul, A.ρ.map_one, LinearMap.one_apply] rw [zero_smul] @@ -314,9 +314,10 @@ noncomputable def leftRegularHomEquiv (A : Rep k G) : (Rep.ofMulAction k G G ⟶ f.hom ((ofMulAction k G G).ρ x (Finsupp.single (1 : G) (1 : k))) = A.ρ x (f.hom (Finsupp.single (1 : G) (1 : k))) := LinearMap.ext_iff.1 (ModuleCat.hom_ext_iff.mp (f.comm x)) (Finsupp.single 1 1) - simp only [leftRegularHom_hom_hom, LinearMap.comp_apply, Finsupp.lsingle_apply, + simp only [leftRegularHom_hom, LinearMap.comp_apply, Finsupp.lsingle_apply, Finsupp.lift_apply, ← this, coe_of, of_ρ, Representation.ofMulAction_single x (1 : G) (1 : k), - smul_eq_mul, mul_one, zero_smul, Finsupp.sum_single_index, one_smul] + smul_eq_mul, mul_one, zero_smul, Finsupp.sum_single_index, one_smul, + ConcreteCategory.hom_ofHom] -- Mismatched `Zero k` instances rfl right_inv x := leftRegularHom_apply x @@ -324,7 +325,7 @@ noncomputable def leftRegularHomEquiv (A : Rep k G) : (Rep.ofMulAction k G G ⟶ theorem leftRegularHomEquiv_symm_single {A : Rep k G} (x : A) (g : G) : ((leftRegularHomEquiv A).symm x).hom (Finsupp.single g 1) = A.ρ g x := by -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [leftRegularHomEquiv_symm_apply, leftRegularHom_hom_hom, Finsupp.lift_apply, + erw [leftRegularHomEquiv_symm_apply, leftRegularHom_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul] rw [zero_smul] diff --git a/Mathlib/RingTheory/Adjoin/Tower.lean b/Mathlib/RingTheory/Adjoin/Tower.lean index 5e0b6870afb102..f0bec2c20ba932 100644 --- a/Mathlib/RingTheory/Adjoin/Tower.lean +++ b/Mathlib/RingTheory/Adjoin/Tower.lean @@ -141,7 +141,8 @@ variable [Algebra A B] [Algebra B C] [Algebra A C] [IsScalarTower A B C] A is noetherian, and C is algebra-finite over A, and C is module-finite over B, then B is algebra-finite over A. -References: Atiyah--Macdonald Proposition 7.8; Stacks 00IS; Altman--Kleiman 16.17. -/ +References: Atiyah--Macdonald Proposition 7.8; Altman--Kleiman 16.17. -/ +@[stacks 00IS] theorem fg_of_fg_of_fg [IsNoetherianRing A] (hAC : (⊤ : Subalgebra A C).FG) (hBC : (⊤ : Submodule B C).FG) (hBCi : Function.Injective (algebraMap B C)) : (⊤ : Subalgebra A B).FG := diff --git a/Mathlib/RingTheory/Artinian/Ring.lean b/Mathlib/RingTheory/Artinian/Ring.lean index 4889c2cb83fb23..455b601e23f615 100644 --- a/Mathlib/RingTheory/Artinian/Ring.lean +++ b/Mathlib/RingTheory/Artinian/Ring.lean @@ -91,6 +91,26 @@ theorem isField_of_isReduced_of_isLocalRing [IsReduced R] [IsLocalRing R] : IsFi (IsArtinianRing.equivPi R).trans (RingEquiv.piUnique _) |>.toMulEquiv.isField _ (Ideal.Quotient.field _).toIsField +section IsUnit + +open nonZeroDivisors + +/-- If an element of an artinian ring is not a zero divisor then it is a unit. -/ +theorem isUnit_of_mem_nonZeroDivisors {a : R} (ha : a ∈ R⁰) : IsUnit a := + IsUnit.isUnit_iff_mulLeft_bijective.mpr <| + IsArtinian.bijective_of_injective_endomorphism (LinearMap.mulLeft R a) + fun _ _ ↦ (mul_cancel_left_mem_nonZeroDivisors ha).mp + +/-- In an artinian ring, an element is a unit iff it is a non-zero-divisor. +See also `isUnit_iff_mem_nonZeroDivisors_of_finite`.-/ +theorem isUnit_iff_mem_nonZeroDivisors {a : R} : IsUnit a ↔ a ∈ R⁰ := + ⟨IsUnit.mem_nonZeroDivisors, isUnit_of_mem_nonZeroDivisors⟩ + +theorem isUnit_submonoid_eq : IsUnit.submonoid R = R⁰ := by + ext; simp [IsUnit.mem_submonoid_iff, isUnit_iff_mem_nonZeroDivisors] + +end IsUnit + section Localization variable (S : Submonoid R) (L : Type*) [CommSemiring L] [Algebra R L] [IsLocalization S L] diff --git a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean index dafc8270e1007d..e8efc2bbe44ae3 100644 --- a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean +++ b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean @@ -57,7 +57,7 @@ noncomputable instance TensorProduct.instCoalgebra : Coalgebra R (M ⊗[R] N) := map_comp_comul := by rw [CoalgebraCat.ofComonObjCoalgebraStruct_comul] simp [-Mon_.monMonoidalStruct_tensorObj_X, - ModuleCat.MonoidalCategory.instMonoidalCategoryStruct_tensorHom_hom, + ModuleCat.MonoidalCategory.instMonoidalCategoryStruct_tensorHom, ModuleCat.hom_comp, ModuleCat.of, ModuleCat.ofHom, ModuleCat.MonoidalCategory.tensorμ_eq_tensorTensorTensorComm] } diff --git a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean index 110ec9cd323fbc..d7d5139cc0b3e8 100644 --- a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean +++ b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean @@ -403,6 +403,9 @@ def adicCompletion := instance : Field (v.adicCompletion K) := inferInstanceAs <| Field (@UniformSpace.Completion K v.adicValued.toUniformSpace) +instance : Algebra K (v.adicCompletion K) := + RingHom.toAlgebra (@UniformSpace.Completion.coeRingHom K _ v.adicValued.toUniformSpace _ _) + instance : Inhabited (v.adicCompletion K) := ⟨0⟩ diff --git a/Mathlib/RingTheory/Etale/Basic.lean b/Mathlib/RingTheory/Etale/Basic.lean index 1f11834fdfbaa8..1b85886b995e0f 100644 --- a/Mathlib/RingTheory/Etale/Basic.lean +++ b/Mathlib/RingTheory/Etale/Basic.lean @@ -167,11 +167,9 @@ section variable (R : Type u) [CommRing R] variable (A : Type u) [CommRing A] [Algebra R A] -/-- An `R`-algebra `A` is étale if it is formally étale and of finite presentation. - -Note that the definition in the stacks project is -different, but shows that it is equivalent -to the definition here. -/ +/-- An `R`-algebra `A` is étale if it is formally étale and of finite presentation. -/ +@[stacks 00U1 "Note that this is a different definition from this Stacks entry, but + shows that it is equivalent to the definition here."] class Etale : Prop where formallyEtale : FormallyEtale R A := by infer_instance finitePresentation : FinitePresentation R A := by infer_instance diff --git a/Mathlib/RingTheory/FinitePresentation.lean b/Mathlib/RingTheory/FinitePresentation.lean index 06867a417b8d1c..dbc3a551101b0b 100644 --- a/Mathlib/RingTheory/FinitePresentation.lean +++ b/Mathlib/RingTheory/FinitePresentation.lean @@ -212,8 +212,8 @@ theorem trans [Algebra A B] [IsScalarTower R A B] [FinitePresentation R A] open MvPolynomial --- We follow the proof of https://stacks.math.columbia.edu/tag/0561 -- TODO: extract out helper lemmas and tidy proof. +@[stacks 0561] theorem of_restrict_scalars_finitePresentation [Algebra A B] [IsScalarTower R A B] [FinitePresentation.{w₁, w₃} R B] [FiniteType R A] : FinitePresentation.{w₂, w₃} A B := by diff --git a/Mathlib/RingTheory/Finiteness/Nakayama.lean b/Mathlib/RingTheory/Finiteness/Nakayama.lean index a5ae263188b4da..d17b83e9cc8881 100644 --- a/Mathlib/RingTheory/Finiteness/Nakayama.lean +++ b/Mathlib/RingTheory/Finiteness/Nakayama.lean @@ -26,8 +26,8 @@ variable {R : Type*} {M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] open Set -/-- **Nakayama's Lemma**. Atiyah-Macdonald 2.5, Eisenbud 4.7, Matsumura 2.2, -[Stacks 00DV](https://stacks.math.columbia.edu/tag/00DV) -/ +/-- **Nakayama's Lemma**. Atiyah-Macdonald 2.5, Eisenbud 4.7, Matsumura 2.2. -/ +@[stacks 00DV] theorem exists_sub_one_mem_and_smul_eq_zero_of_fg_of_le_smul {R : Type*} [CommRing R] {M : Type*} [AddCommGroup M] [Module R M] (I : Ideal R) (N : Submodule R M) (hn : N.FG) (hin : N ≤ I • N) : ∃ r : R, r - 1 ∈ I ∧ ∀ n ∈ N, r • n = (0 : M) := by diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index 5db79f945e9dd8..bbf91f8c56ddb2 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -714,7 +714,7 @@ def awayMapₐ : Away 𝒜 f →ₐ[𝒜 0] Away 𝒜 x where /-- This is a convenient constructor for `Away 𝒜 f` when `f` is homogeneous. `Away.mk 𝒜 hf n x hx` is the fraction `x / f ^ n`. -/ protected def Away.mk {d : ι} (hf : f ∈ 𝒜 d) (n : ℕ) (x : A) (hx : x ∈ 𝒜 (n • d)) : Away 𝒜 f := - .mk ⟨n • d, ⟨x, hx⟩, ⟨f ^ n, SetLike.pow_mem_graded n hf⟩, ⟨n, rfl⟩⟩ + HomogeneousLocalization.mk ⟨n • d, ⟨x, hx⟩, ⟨f ^ n, SetLike.pow_mem_graded n hf⟩, ⟨n, rfl⟩⟩ @[simp] lemma Away.val_mk {d : ι} (n : ℕ) (hf : f ∈ 𝒜 d) (x : A) (hx : x ∈ 𝒜 (n • d)) : diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index ee0467668b791b..de11be3c0f851c 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -356,22 +356,18 @@ theorem prod_mem_prod {ι : Type*} {s : Finset ι} {I : ι → Ideal R} {x : ι theorem mul_le_right : I * J ≤ I := Ideal.mul_le.2 fun _ hr _ _ => I.mul_mem_right _ hr -#adaptation_note -/-- -On nightly-2024-11-12, we had to add `nolint simpNF` to the following lemma, +#adaptation_note /-- nightly-2024-11-12 +we had to add `nolint simpNF` to the following lemma, as otherwise we get a deterministic timeout in typeclass inference. -This should be investigated. --/ +This should be investigated. -/ @[simp, nolint simpNF] theorem sup_mul_right_self : I ⊔ I * J = I := sup_eq_left.2 Ideal.mul_le_right -#adaptation_note -/-- -On nightly-2024-11-12, we had to add `nolint simpNF` to the following lemma, +#adaptation_note /-- nightly-2024-11-12 +we had to add `nolint simpNF` to the following lemma, as otherwise we get a deterministic timeout in typeclass inference. -This should be investigated. --/ +This should be investigated. -/ @[simp, nolint simpNF] theorem mul_right_self_sup : I * J ⊔ I = I := sup_eq_right.2 Ideal.mul_le_right @@ -1035,7 +1031,8 @@ theorem subset_union_prime' {R : Type u} [CommRing R] {s : Finset ι} {f : ι exact hs <| Or.inr <| Set.mem_biUnion hjt <| add_sub_cancel_left r s ▸ (f j).sub_mem hj <| hr j hjt -/-- Prime avoidance. Atiyah-Macdonald 1.11, Eisenbud 3.3, Stacks 00DS, Matsumura Ex.1.6. -/ +/-- Prime avoidance. Atiyah-Macdonald 1.11, Eisenbud 3.3, Matsumura Ex.1.6. -/ +@[stacks 00DS] theorem subset_union_prime {R : Type u} [CommRing R] {s : Finset ι} {f : ι → Ideal R} (a b : ι) (hp : ∀ i ∈ s, i ≠ a → i ≠ b → IsPrime (f i)) {I : Ideal R} : ((I : Set R) ⊆ ⋃ i ∈ (↑s : Set ι), f i) ↔ ∃ i ∈ s, I ≤ f i := diff --git a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean index 0f9874f3e5ae84..c8e7297bf65f1e 100644 --- a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean @@ -41,6 +41,11 @@ theorem Algebra.IsIntegral.of_injective (f : A →ₐ[R] B) (hf : Function.Injec [Algebra.IsIntegral R B] : Algebra.IsIntegral R A := ⟨fun _ ↦ (isIntegral_algHom_iff f hf).mp (isIntegral _)⟩ +/-- Homomorphic image of an integral algebra is an integral algebra. -/ +theorem Algebra.IsIntegral.of_surjective [Algebra.IsIntegral R A] + (f : A →ₐ[R] B) (hf : Function.Surjective f) : Algebra.IsIntegral R B := + isIntegral_def.mpr fun b ↦ let ⟨a, ha⟩ := hf b; ha ▸ (isIntegral_def.mp ‹_› a).map f + theorem AlgEquiv.isIntegral_iff (e : A ≃ₐ[R] B) : Algebra.IsIntegral R A ↔ Algebra.IsIntegral R B := ⟨fun h ↦ h.of_injective e.symm e.symm.injective, fun h ↦ h.of_injective e e.injective⟩ @@ -205,8 +210,15 @@ def integralClosure : Subalgebra R A where mul_mem' := IsIntegral.mul algebraMap_mem' _ := isIntegral_algebraMap -end - -theorem mem_integralClosure_iff (R A : Type*) [CommRing R] [CommRing A] [Algebra R A] {a : A} : - a ∈ integralClosure R A ↔ IsIntegral R a := +theorem mem_integralClosure_iff {a : A} : a ∈ integralClosure R A ↔ IsIntegral R a := Iff.rfl + +variable {R} {A B : Type*} [Ring A] [Algebra R A] [Ring B] [Algebra R B] + +/-- Product of two integral algebras is an integral algebra. -/ +instance Algebra.IsIntegral.prod [Algebra.IsIntegral R A] [Algebra.IsIntegral R B] : + Algebra.IsIntegral R (A × B) := + Algebra.isIntegral_def.mpr fun x ↦ + (Algebra.isIntegral_def.mp ‹_› x.1).pair (Algebra.isIntegral_def.mp ‹_› x.2) + +end diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean index 361784bfb1bc20..5e75b8a8f6cd12 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean @@ -201,8 +201,37 @@ theorem fg_adjoin_of_finite {s : Set A} (hfs : s.Finite) (his : ∀ x ∈ s, IsI (his a <| Set.mem_insert a s).fg_adjoin_singleton) his +theorem Algebra.finite_adjoin_of_finite_of_isIntegral {s : Set A} (hf : s.Finite) + (hi : ∀ x ∈ s, IsIntegral R x) : Module.Finite R (adjoin R s) := + Module.Finite.iff_fg.mpr <| fg_adjoin_of_finite hf hi + +theorem Algebra.finite_adjoin_simple_of_isIntegral {x : A} (hi : IsIntegral R x) : + Module.Finite R (adjoin R {x}) := + Module.Finite.iff_fg.mpr hi.fg_adjoin_singleton + theorem isNoetherian_adjoin_finset [IsNoetherianRing R] (s : Finset A) (hs : ∀ x ∈ s, IsIntegral R x) : IsNoetherian R (Algebra.adjoin R (s : Set A)) := isNoetherian_of_fg_of_noetherian _ (fg_adjoin_of_finite s.finite_toSet hs) end + +section Prod + +variable {R A B : Type*} +variable [CommRing R] [Ring A] [Ring B] [Algebra R A] [Algebra R B] + +/-- An element of a product algebra is integral if each component is integral. -/ +theorem IsIntegral.pair {x : A × B} (hx₁ : IsIntegral R x.1) (hx₂ : IsIntegral R x.2) : + IsIntegral R x := by + obtain ⟨p₁, ⟨hp₁Monic, hp₁Eval⟩⟩ := hx₁ + obtain ⟨p₂, ⟨hp₂Monic, hp₂Eval⟩⟩ := hx₂ + refine ⟨p₁ * p₂, ⟨hp₁Monic.mul hp₂Monic, ?_⟩⟩ + rw [← aeval_def] at * + rw [aeval_prod_apply, aeval_mul, hp₁Eval, zero_mul, aeval_mul, hp₂Eval, mul_zero, + Prod.zero_eq_mk] + +/-- An element of a product algebra is integral iff each component is integral. -/ +theorem IsIntegral.pair_iff {x : A × B} : IsIntegral R x ↔ IsIntegral R x.1 ∧ IsIntegral R x.2 := + ⟨fun h ↦ ⟨h.map (AlgHom.fst R A B), h.map (AlgHom.snd R A B)⟩, fun h ↦ h.1.pair h.2⟩ + +end Prod diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean index c62e33c4ab1b2e..ff4bf59739f942 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean @@ -132,7 +132,7 @@ theorem le_integralClosure_iff_isIntegral {S : Subalgebra R A} : Algebra.isIntegral_def.symm theorem Algebra.IsIntegral.adjoin {S : Set A} (hS : ∀ x ∈ S, IsIntegral R x) : - Algebra.IsIntegral R (Algebra.adjoin R S) := + Algebra.IsIntegral R (adjoin R S) := le_integralClosure_iff_isIntegral.mp <| adjoin_le hS theorem integralClosure_eq_top_iff : integralClosure R A = ⊤ ↔ Algebra.IsIntegral R A := by diff --git a/Mathlib/RingTheory/Jacobson/Ring.lean b/Mathlib/RingTheory/Jacobson/Ring.lean index 68129d02d2fcbc..aca96b49168c82 100644 --- a/Mathlib/RingTheory/Jacobson/Ring.lean +++ b/Mathlib/RingTheory/Jacobson/Ring.lean @@ -702,11 +702,8 @@ lemma finite_of_finite_type_of_isJacobsonRing (R S : Type*) [CommRing R] [Field obtain ⟨ι, hι, f, hf⟩ := Algebra.FiniteType.iff_quotient_mvPolynomial'.mp ‹_› have : (algebraMap R S).IsIntegral := by rw [← f.comp_algebraMap] - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/6024 - we needed to write `f.toRingHom` instead of just `f`, to avoid unification issues. - -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 + we needed to write `f.toRingHom` instead of just `f`, to avoid unification issues. -/ exact MvPolynomial.comp_C_integral_of_surjective_of_isJacobsonRing f.toRingHom hf have : Algebra.IsIntegral R S := Algebra.isIntegral_def.mpr this exact Algebra.IsIntegral.finite diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index cf8a5776479281..27e5f145a98c1c 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -14,6 +14,7 @@ import Mathlib.RingTheory.Localization.FractionRing import Mathlib.Topology.UniformSpace.Cauchy import Mathlib.Algebra.Group.Int.TypeTags + /-! # Laurent Series @@ -60,8 +61,8 @@ such that the `X`-adic valuation `v` satisfies `v (f - Q) < γ`. `LaurentSeries.exists_powerSeries_of_memIntegers` show that an element in the completion of `RatFunc K` is in the unit ball if and only if it comes from a power series through the isomorphism `LaurentSeriesRingEquiv`. -* `LaurentSeries.powerSeriesRingEquiv` is the ring isomorphism between `K⟦X⟧` and the unit ball -inside the `X`-adic completion of `RatFunc K`. +* `LaurentSeries.powerSeriesAlgEquiv` is the `K`-algebra isomorphism between `K⟦X⟧` +and the unit ball inside the `X`-adic completion of `RatFunc K`. ## Implementation details @@ -75,18 +76,13 @@ that it is complete and contains `RatFunc K` as a dense subspace. The isomorphis equivalence, expressing the mathematical idea that the completion "is unique". It is `LaurentSeries.comparePkg`. * For applications to `K⟦X⟧` it is actually more handy to use the *inverse* of the above -equivalence: `LaurentSeries.LaurentSeriesRingEquiv` is the *topological, ring equivalence* -`K⸨X⸩ ≃+* RatFuncAdicCompl K`. +equivalence: `LaurentSeries.LaurentSeriesAlgEquiv` is the *topological, algebra equivalence* +`K⸨X⸩ ≃ₐ[K] RatFuncAdicCompl K`. * In order to compare `K⟦X⟧` with the valuation subring in the `X`-adic completion of `RatFunc K` we consider its alias `LaurentSeries.powerSeries_as_subring` as a subring of `K⸨X⸩`, that is itself clearly isomorphic (via the inverse of `LaurentSeries.powerSeriesEquivSubring`) to `K⟦X⟧`. -## To Do -* The `AdicCompletion` construction is currently done for ideals in rings and does not take into -account the relation with algebra structures on the ring, hence it does not yield a `K`-algebra -structure on the `X`-adic completion of `K⸨X⸩`. Once this will be available, we should update -`LaurentSeries.LaurentSeriesRingEquiv` to an algebra equivalence. -/ universe u @@ -1061,6 +1057,11 @@ equivalence: it goes from `K⸨X⸩` to `RatFuncAdicCompl K` -/ abbrev LaurentSeriesRingEquiv : K⸨X⸩ ≃+* RatFuncAdicCompl K := (ratfuncAdicComplRingEquiv K).symm +@[simp] +lemma LaurentSeriesRingEquiv_def (f : K⟦X⟧) : + (LaurentSeriesRingEquiv K) f = (LaurentSeriesPkg K).compare ratfuncAdicComplPkg (f : K⸨X⸩) := + rfl + @[simp] theorem ratfuncAdicComplRingEquiv_apply (x : RatFuncAdicCompl K) : ratfuncAdicComplRingEquiv K x = ratfuncAdicComplPkg.compare (LaurentSeriesPkg K) x := rfl @@ -1071,6 +1072,17 @@ theorem coe_X_compare : rw [PowerSeries.coe_X, ← RatFunc.coe_X, ← LaurentSeries_coe, ← compare_coe] rfl +theorem algebraMap_apply (a : K) : algebraMap K K⸨X⸩ a = HahnSeries.C a := by + simp [RingHom.algebraMap_toAlgebra] + +instance : Algebra K (RatFuncAdicCompl K) := + RingHom.toAlgebra ((LaurentSeriesRingEquiv K).toRingHom.comp HahnSeries.C) + +/-- The algebra equivalence between `K⸨X⸩` and the `X`-adic completion of `RatFunc X` -/ +def LaurentSeriesAlgEquiv : K⸨X⸩ ≃ₐ[K] RatFuncAdicCompl K := + AlgEquiv.ofRingEquiv (f := LaurentSeriesRingEquiv K) + (fun a ↦ by simp [RingHom.algebraMap_toAlgebra]) + open Filter WithZero open scoped WithZeroTopology Topology Multiplicative @@ -1132,22 +1144,30 @@ section PowerSeries /-- In order to compare `K⟦X⟧` with the valuation subring in the `X`-adic completion of `RatFunc K` we consider its alias as a subring of `K⸨X⸩`. -/ abbrev powerSeries_as_subring : Subring K⸨X⸩ := - RingHom.range (HahnSeries.ofPowerSeries ℤ K) +-- RingHom.range (HahnSeries.ofPowerSeries ℤ K) + Subring.map (HahnSeries.ofPowerSeries ℤ K) ⊤ /-- The ring `K⟦X⟧` is isomorphic to the subring `powerSeries_as_subring K` -/ -abbrev powerSeriesEquivSubring : K⟦X⟧ ≃+* powerSeries_as_subring K := by - rw [powerSeries_as_subring, RingHom.range_eq_map] - exact ((Subring.topEquiv).symm).trans (Subring.equivMapOfInjective ⊤ (ofPowerSeries ℤ K) +abbrev powerSeriesEquivSubring : K⟦X⟧ ≃+* powerSeries_as_subring K := + ((Subring.topEquiv).symm).trans (Subring.equivMapOfInjective ⊤ (ofPowerSeries ℤ K) ofPowerSeries_injective) +lemma powerSeriesEquivSubring_apply (f : K⟦X⟧) : + powerSeriesEquivSubring K f = + ⟨HahnSeries.ofPowerSeries ℤ K f, Subring.mem_map.mpr ⟨f, trivial, rfl⟩⟩ := + rfl + +lemma powerSeriesEquivSubring_coe_apply (f : K⟦X⟧) : + (powerSeriesEquivSubring K f : K⸨X⸩) = ofPowerSeries ℤ K f := + rfl + /- Through the isomorphism `LaurentSeriesRingEquiv`, power series land in the unit ball inside the completion of `RatFunc K`. -/ theorem mem_integers_of_powerSeries (F : K⟦X⟧) : (LaurentSeriesRingEquiv K) F ∈ (idealX K).adicCompletionIntegers (RatFunc K) := by - have : (LaurentSeriesRingEquiv K) F = - (LaurentSeriesPkg K).compare ratfuncAdicComplPkg (F : K⸨X⸩) := rfl simp only [Subring.mem_map, exists_prop, ValuationSubring.mem_toSubring, - mem_adicCompletionIntegers, this, valuation_compare, val_le_one_iff_eq_coe] + mem_adicCompletionIntegers, LaurentSeriesRingEquiv_def, + valuation_compare, val_le_one_iff_eq_coe] exact ⟨F, rfl⟩ /- Conversely, all elements in the unit ball inside the completion of `RatFunc K` come from a power @@ -1166,13 +1186,13 @@ theorem powerSeries_ext_subring : Subring.map (LaurentSeriesRingEquiv K).toRingHom (powerSeries_as_subring K) = ((idealX K).adicCompletionIntegers (RatFunc K)).toSubring := by ext x - refine ⟨fun ⟨f, ⟨F, coe_F⟩, hF⟩ ↦ ?_, fun H ↦ ?_⟩ + refine ⟨fun ⟨f, ⟨F, _, coe_F⟩, hF⟩ ↦ ?_, fun H ↦ ?_⟩ · simp only [ValuationSubring.mem_toSubring, ← hF, ← coe_F] apply mem_integers_of_powerSeries · obtain ⟨F, hF⟩ := exists_powerSeries_of_memIntegers K H simp only [Equiv.toFun_as_coe, UniformEquiv.coe_toEquiv, exists_exists_eq_and, UniformEquiv.coe_symm_toEquiv, Subring.mem_map, Equiv.invFun_as_coe] - exact ⟨F, ⟨F, rfl⟩, hF⟩ + exact ⟨F, ⟨F, trivial, rfl⟩, hF⟩ /-- The ring isomorphism between `K⟦X⟧` and the unit ball inside the `X`-adic completion of `RatFunc K`. -/ @@ -1180,6 +1200,43 @@ abbrev powerSeriesRingEquiv : K⟦X⟧ ≃+* (idealX K).adicCompletionIntegers ( ((powerSeriesEquivSubring K).trans (LaurentSeriesRingEquiv K).subringMap).trans <| RingEquiv.subringCongr (powerSeries_ext_subring K) +lemma powerSeriesRingEquiv_coe_apply (f : K⟦X⟧) : + powerSeriesRingEquiv K f = LaurentSeriesRingEquiv K (f : K⸨X⸩) := + rfl + +lemma LaurentSeriesRingEquiv_mem_valuationSubring (f : K⟦X⟧) : + LaurentSeriesRingEquiv K f ∈ Valued.v.valuationSubring := by + simp only [Valuation.mem_valuationSubring_iff] + rw [LaurentSeriesRingEquiv_def, valuation_compare, val_le_one_iff_eq_coe] + use f + +lemma algebraMap_C_mem_adicCompletionIntegers (x : K) : + ((LaurentSeriesRingEquiv K).toRingHom.comp HahnSeries.C) x ∈ + adicCompletionIntegers (RatFunc K) (idealX K) := by + have : HahnSeries.C x = ofPowerSeries ℤ K (PowerSeries.C K x) := by + simp [C_apply, ofPowerSeries_C] + simp only [RingHom.comp_apply, RingEquiv.toRingHom_eq_coe, RingHom.coe_coe, this] + apply LaurentSeriesRingEquiv_mem_valuationSubring + +instance : Algebra K ((idealX K).adicCompletionIntegers (RatFunc K)) := + RingHom.toAlgebra <| + ((LaurentSeriesRingEquiv K).toRingHom.comp HahnSeries.C).codRestrict _ + (algebraMap_C_mem_adicCompletionIntegers K) + +instance : IsScalarTower K ((idealX K).adicCompletionIntegers (RatFunc K)) + ((idealX K).adicCompletion (RatFunc K)) := + IsScalarTower.of_algebraMap_eq (fun _ ↦ by rfl) + +/-- The algebra isomorphism between `K⟦X⟧` and the unit ball inside the `X`-adic completion of +`RatFunc K`. -/ +def powerSeriesAlgEquiv : K⟦X⟧ ≃ₐ[K] (idealX K).adicCompletionIntegers (RatFunc K) := by + apply AlgEquiv.ofRingEquiv (f := powerSeriesRingEquiv K) + intro a + rw [PowerSeries.algebraMap_eq, RingHom.algebraMap_toAlgebra, ← Subtype.coe_inj, + powerSeriesRingEquiv_coe_apply, + RingHom.codRestrict_apply _ _ (algebraMap_C_mem_adicCompletionIntegers K)] + simp + end PowerSeries end Comparison diff --git a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean index 2c5a8113251ff2..e2b6096cd016dd 100644 --- a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean @@ -86,8 +86,8 @@ section variable [CommSemiring R] [IsLocalRing R] [CommSemiring S] [IsLocalRing S] /-- A ring homomorphism between local rings is a local ring hom iff it reflects units, -i.e. any preimage of a unit is still a unit. https://stacks.math.columbia.edu/tag/07BJ --/ +i.e. any preimage of a unit is still a unit. -/ +@[stacks 07BJ] theorem local_hom_TFAE (f : R →+* S) : List.TFAE [IsLocalHom f, f '' (maximalIdeal R).1 ⊆ maximalIdeal S, diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index 22ac6c09fc61ed..8b0b1e1417be31 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -296,6 +296,11 @@ theorem isLocalization_iff_of_isLocalization [IsLocalization M S] [IsLocalizatio ⟨fun _ ↦ isLocalization_of_algEquiv N (algEquiv M S P), fun _ ↦ isLocalization_of_algEquiv M (algEquiv N S P)⟩ +theorem iff_of_le_of_exists_dvd (N : Submonoid R) (h₁ : M ≤ N) (h₂ : ∀ n ∈ N, ∃ m ∈ M, n ∣ m) : + IsLocalization M S ↔ IsLocalization N S := + have : IsLocalization N (Localization M) := of_le_of_exists_dvd _ _ h₁ h₂ + isLocalization_iff_of_isLocalization _ _ (Localization M) + end variable (M) diff --git a/Mathlib/RingTheory/Localization/Pi.lean b/Mathlib/RingTheory/Localization/Pi.lean new file mode 100644 index 00000000000000..28a3911c43cd75 --- /dev/null +++ b/Mathlib/RingTheory/Localization/Pi.lean @@ -0,0 +1,79 @@ +/- +Copyright (c) 2024 Madison Crim. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Madison Crim +-/ +import Mathlib.Algebra.Algebra.Pi +import Mathlib.Algebra.BigOperators.Pi +import Mathlib.Algebra.Divisibility.Prod +import Mathlib.Algebra.Group.Submonoid.BigOperators +import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.RingTheory.Localization.Basic +import Mathlib.Algebra.Group.Pi.Units + +/-! +# Localizing a product of commutative rings + +## Main Result + + * `bijective_lift_piRingHom_algebraMap_comp_piEvalRingHom`: the canonical map from a + localization of a finite product of rings `R i `at a monoid `M` to the direct product of + localizations `R i` at the projection of `M` onto each corresponding factor is bijective. + +## Implementation notes + +See `Mathlib/RingTheory/Localization/Defs.lean` for a design overview. + +## Tags +localization, commutative ring +-/ + +namespace IsLocalization + +variable {ι : Type*} (R S : ι → Type*) + [Π i, CommSemiring (R i)] [Π i, CommSemiring (S i)] [Π i, Algebra (R i) (S i)] + +/-- If `S i` is a localization of `R i` at the submonoid `M i` for each `i`, +then `Π i, S i` is a localization of `Π i, R i` at the product submonoid. -/ +instance (M : Π i, Submonoid (R i)) [∀ i, IsLocalization (M i) (S i)] : + IsLocalization (.pi .univ M) (Π i, S i) where + map_units' m := Pi.isUnit_iff.mpr fun i ↦ map_units _ ⟨m.1 i, m.2 i ⟨⟩⟩ + surj' z := by + choose rm h using fun i ↦ surj (M := M i) (z i) + exact ⟨(fun i ↦ (rm i).1, ⟨_, fun i _ ↦ (rm i).2.2⟩), funext h⟩ + exists_of_eq {x y} eq := by + choose c hc using fun i ↦ exists_of_eq (M := M i) (congr_fun eq i) + exact ⟨⟨_, fun i _ ↦ (c i).2⟩, funext hc⟩ + +variable (S' : Type*) [CommSemiring S'] [Algebra (Π i, R i) S'] (M : Submonoid (Π i, R i)) + +theorem iff_map_piEvalRingHom [Finite ι] : + IsLocalization M S' ↔ IsLocalization (.pi .univ fun i ↦ M.map (Pi.evalRingHom R i)) S' := + iff_of_le_of_exists_dvd M _ (fun m hm i _ ↦ ⟨m, hm, rfl⟩) fun n hn ↦ by + choose m mem eq using hn + have := Fintype.ofFinite ι + refine ⟨∏ i, m i ⟨⟩, prod_mem fun i _ ↦ mem i _, pi_dvd_iff.mpr fun i ↦ ?_⟩ + rw [Fintype.prod_apply] + exact (eq i ⟨⟩).symm.dvd.trans (Finset.dvd_prod_of_mem _ <| Finset.mem_univ _) + +variable [∀ i, IsLocalization (M.map (Pi.evalRingHom R i)) (S i)] + +/-- Let `M` be a submonoid of a direct product of commutative rings `R i`, and let `M' i` denote +the projection of `M` onto each corresponding factor. Given a ring homomorphism from the direct +product `Π i, R i` to the product of the localizations of each `R i` at `M' i`, every `y : M` +maps to a unit under this homomorphism. -/ +lemma isUnit_piRingHom_algebraMap_comp_piEvalRingHom (y : M) : + IsUnit ((Pi.ringHom fun i ↦ (algebraMap (R i) (S i)).comp (Pi.evalRingHom R i)) y) := + Pi.isUnit_iff.mpr fun i ↦ map_units _ (⟨y.1 i, y, y.2, rfl⟩ : M.map (Pi.evalRingHom R i)) + +/-- Let `M` be a submonoid of a direct product of commutative rings `R i`, and let `M' i` denote +the projection of `M` onto each factor. Then the canonical map from the localization of the direct +product `Π i, R i` at `M` to the direct product of the localizations of each `R i` at `M' i` +is bijective. -/ +theorem bijective_lift_piRingHom_algebraMap_comp_piEvalRingHom [IsLocalization M S'] [Finite ι] : + Function.Bijective (lift (S := S') (isUnit_piRingHom_algebraMap_comp_piEvalRingHom R S M)) := + have := (iff_map_piEvalRingHom R (Π i, S i) M).mpr inferInstance + (ringEquivOfRingEquiv (M := M) (T := M) _ _ (.refl _) <| + Submonoid.map_equiv_eq_comap_symm _ _).bijective + +end IsLocalization diff --git a/Mathlib/RingTheory/Nakayama.lean b/Mathlib/RingTheory/Nakayama.lean index ad4d01141b8f1f..55c8610c70b934 100644 --- a/Mathlib/RingTheory/Nakayama.lean +++ b/Mathlib/RingTheory/Nakayama.lean @@ -16,13 +16,13 @@ This file contains some alternative statements of Nakayama's Lemma as found in ## Main statements * `Submodule.eq_smul_of_le_smul_of_le_jacobson` - A version of (2) in - [Stacks: Nakayama's Lemma](https://stacks.math.columbia.edu/tag/00DV)., + [Stacks: Nakayama's Lemma](https://stacks.math.columbia.edu/tag/00DV), generalising to the Jacobson of any ideal. * `Submodule.eq_bot_of_le_smul_of_le_jacobson_bot` - Statement (2) in [Stacks: Nakayama's Lemma](https://stacks.math.columbia.edu/tag/00DV). * `Submodule.sup_smul_eq_sup_smul_of_le_smul_of_le_jacobson` - A version of (4) in - [Stacks: Nakayama's Lemma](https://stacks.math.columbia.edu/tag/00DV)., + [Stacks: Nakayama's Lemma](https://stacks.math.columbia.edu/tag/00DV), generalising to the Jacobson of any ideal. * `Submodule.smul_le_of_le_smul_of_le_jacobson_bot` - Statement (4) in [Stacks: Nakayama's Lemma](https://stacks.math.columbia.edu/tag/00DV). diff --git a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean index 35474da03eabb6..ecf3da187d103a 100644 --- a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean @@ -342,6 +342,17 @@ instance center.instNonUnitalCommRing : NonUnitalCommRing (center R) := { NonUnitalSubsemiring.center.instNonUnitalCommSemiring R, inferInstanceAs <| NonUnitalNonAssocRing (center R) with } +variable {R} + +/-- The center of isomorphic (not necessarily unital or associative) rings are isomorphic. -/ +@[simps!] def centerCongr {S} [NonUnitalNonAssocRing S] (e : R ≃+* S) : center R ≃+* center S := + NonUnitalSubsemiring.centerCongr e + +/-- The center of a (not necessarily uintal or associative) ring +is isomorphic to the center of its opposite. -/ +@[simps!] def centerToMulOpposite : center R ≃+* center Rᵐᵒᵖ := + NonUnitalSubsemiring.centerToMulOpposite + end NonUnitalNonAssocRing section NonUnitalRing diff --git a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean index 61c5c358fc66c0..76b1cd54c68799 100644 --- a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean @@ -10,8 +10,10 @@ import Mathlib.Algebra.Group.Subsemigroup.Operations import Mathlib.Algebra.GroupWithZero.Center import Mathlib.Algebra.Ring.Center import Mathlib.Algebra.Ring.Centralizer +import Mathlib.Algebra.Ring.Opposite import Mathlib.Algebra.Ring.Prod import Mathlib.Data.Set.Finite.Range +import Mathlib.GroupTheory.Submonoid.Center import Mathlib.GroupTheory.Subsemigroup.Centralizer import Mathlib.RingTheory.NonUnitalSubsemiring.Defs @@ -245,6 +247,19 @@ lemma _root_.Set.mem_center_iff_addMonoidHom (a : R) : rw [Set.mem_center_iff, isMulCentral_iff] simp [DFunLike.ext_iff] +variable {R} + +/-- The center of isomorphic (not necessarily unital or associative) semirings are isomorphic. -/ +@[simps!] def centerCongr [NonUnitalNonAssocSemiring S] (e : R ≃+* S) : center R ≃+* center S where + __ := Subsemigroup.centerCongr e + map_add' _ _ := Subtype.ext <| by exact map_add e .. + +/-- The center of a (not necessarily unital or associative) semiring +is isomorphic to the center of its opposite. -/ +@[simps!] def centerToMulOpposite : center R ≃+* center Rᵐᵒᵖ where + __ := Subsemigroup.centerToMulOpposite + map_add' _ _ := rfl + end NonUnitalNonAssocSemiring section NonUnitalSemiring diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 682fe26b642a96..27ce39be124cc5 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -860,7 +860,7 @@ end MvPolynomial end Prime -/-- Hilbert basis theorem: a polynomial ring over a noetherian ring is a noetherian ring. -/ +/-- **Hilbert basis theorem**: a polynomial ring over a Noetherian ring is a Noetherian ring. -/ protected theorem Polynomial.isNoetherianRing [inst : IsNoetherianRing R] : IsNoetherianRing R[X] := isNoetherianRing_iff.2 ⟨fun I : Ideal R[X] => diff --git a/Mathlib/RingTheory/RingHomProperties.lean b/Mathlib/RingTheory/RingHomProperties.lean index c2ac311deb3894..ae523c83039b31 100644 --- a/Mathlib/RingTheory/RingHomProperties.lean +++ b/Mathlib/RingTheory/RingHomProperties.lean @@ -48,14 +48,13 @@ theorem RespectsIso.cancel_left_isIso (hP : RespectsIso @P) {R S T : CommRingCat (g : S ⟶ T) [IsIso f] : P (g.hom.comp f.hom) ↔ P g.hom := ⟨fun H => by convert hP.2 (f ≫ g).hom (asIso f).symm.commRingCatIsoToRingEquiv H - exact (IsIso.inv_hom_id_assoc _ _).symm, hP.2 g.hom (asIso f).commRingCatIsoToRingEquiv⟩ + simp [← CommRingCat.hom_comp], hP.2 g.hom (asIso f).commRingCatIsoToRingEquiv⟩ theorem RespectsIso.cancel_right_isIso (hP : RespectsIso @P) {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) [IsIso g] : P (g.hom.comp f.hom) ↔ P f.hom := ⟨fun H => by convert hP.1 (f ≫ g).hom (asIso g).symm.commRingCatIsoToRingEquiv H - change f = f ≫ g ≫ inv g - simp, hP.1 f.hom (asIso g).commRingCatIsoToRingEquiv⟩ + simp [← CommRingCat.hom_comp], hP.1 f.hom (asIso g).commRingCatIsoToRingEquiv⟩ theorem RespectsIso.is_localization_away_iff (hP : RingHom.RespectsIso @P) {R S : Type u} (R' S' : Type u) [CommRing R] [CommRing S] [CommRing R'] [CommRing S'] [Algebra R R'] diff --git a/Mathlib/RingTheory/Smooth/Basic.lean b/Mathlib/RingTheory/Smooth/Basic.lean index 76b6df39845a51..9ccb330ef6a71d 100644 --- a/Mathlib/RingTheory/Smooth/Basic.lean +++ b/Mathlib/RingTheory/Smooth/Basic.lean @@ -340,12 +340,9 @@ section variable (R : Type u) [CommSemiring R] variable (A : Type u) [Semiring A] [Algebra R A] -/-- An `R` algebra `A` is smooth if it is formally smooth and of finite presentation. - -In the stacks project, the definition of smooth is completely different, and tag - proves that their definition is equivalent -to this. --/ +/-- An `R` algebra `A` is smooth if it is formally smooth and of finite presentation. -/ +@[stacks 00T2 "In the stacks project, the definition of smooth is completely different, and tag + proves that their definition is equivalent to this."] class Smooth [CommSemiring R] (A : Type u) [Semiring A] [Algebra R A] : Prop where formallySmooth : FormallySmooth R A := by infer_instance finitePresentation : FinitePresentation R A := by infer_instance diff --git a/Mathlib/RingTheory/Spectrum/Maximal/Localization.lean b/Mathlib/RingTheory/Spectrum/Maximal/Localization.lean index d0723e00ec9b57..b41a0bebd9a729 100644 --- a/Mathlib/RingTheory/Spectrum/Maximal/Localization.lean +++ b/Mathlib/RingTheory/Spectrum/Maximal/Localization.lean @@ -15,8 +15,6 @@ Localization results. noncomputable section -open scoped Classical - variable (R S P : Type*) [CommSemiring R] [CommSemiring S] [CommSemiring P] namespace MaximalSpectrum @@ -121,6 +119,7 @@ variable {ι} (R : ι → Type*) [∀ i, CommSemiring (R i)] [∀ i, Nontrivial theorem toPiLocalization_not_surjective_of_infinite [Infinite ι] : ¬ Function.Surjective (toPiLocalization (Π i, R i)) := fun surj ↦ by + classical have ⟨J, max, nmem⟩ := PrimeSpectrum.exists_maximal_nmem_range_sigmaToPi_of_infinite R obtain ⟨r, hr⟩ := surj (Function.update 0 ⟨J, max⟩ 1) have : r = 0 := funext fun i ↦ toPiLocalization_injective _ <| funext fun I ↦ by @@ -172,6 +171,7 @@ localizations at maximal ideals. -/ def piLocalizationToMaximal : PiLocalization R →+* MaximalSpectrum.PiLocalization R := Pi.ringHom fun I ↦ Pi.evalRingHom _ I.toPrimeSpectrum +open scoped Classical in theorem piLocalizationToMaximal_surjective : Function.Surjective (piLocalizationToMaximal R) := fun r ↦ ⟨fun I ↦ if h : I.1.IsMaximal then r ⟨_, h⟩ else 0, funext fun _ ↦ dif_pos _⟩ @@ -197,6 +197,7 @@ variable {S} theorem isMaximal_of_toPiLocalization_surjective (surj : Function.Surjective (toPiLocalization R)) (I : PrimeSpectrum R) : I.1.IsMaximal := by + classical have ⟨J, max, le⟩ := I.1.exists_le_maximal I.2.ne_top obtain ⟨r, hr⟩ := surj (Function.update 0 ⟨J, max.isPrime⟩ 1) by_contra h diff --git a/Mathlib/RingTheory/Spectrum/Prime/IsOpenComapC.lean b/Mathlib/RingTheory/Spectrum/Prime/IsOpenComapC.lean index 4a6c542cd56932..2117b0ff267734 100644 --- a/Mathlib/RingTheory/Spectrum/Prime/IsOpenComapC.lean +++ b/Mathlib/RingTheory/Spectrum/Prime/IsOpenComapC.lean @@ -59,11 +59,8 @@ theorem imageOfDf_eq_comap_C_compl_zeroLocus : · rintro ⟨xli, complement, rfl⟩ exact comap_C_mem_imageOfDf complement -/-- The morphism `C⁺ : Spec R[x] → Spec R` is open. -Stacks Project "Lemma 00FB", first part. - -https://stacks.math.columbia.edu/tag/00FB --/ +/-- The morphism `C⁺ : Spec R[x] → Spec R` is open. -/ +@[stacks 00FB "First part"] theorem isOpenMap_comap_C : IsOpenMap (PrimeSpectrum.comap (C : R →+* R[X])) := by rintro U ⟨s, z⟩ rw [← compl_compl U, ← z, ← iUnion_of_singleton_coe s, zeroLocus_iUnion, compl_iInter, diff --git a/Mathlib/RingTheory/Spectrum/Prime/Topology.lean b/Mathlib/RingTheory/Spectrum/Prime/Topology.lean index 52f45bb2373677..c846ca364238f1 100644 --- a/Mathlib/RingTheory/Spectrum/Prime/Topology.lean +++ b/Mathlib/RingTheory/Spectrum/Prime/Topology.lean @@ -976,11 +976,9 @@ section CommSemiring variable [CommSemiring R] open PrimeSpectrum in -/-- -[Stacks: Lemma 00ES (3)](https://stacks.math.columbia.edu/tag/00ES) -Zero loci of minimal prime ideals of `R` are irreducible components in `Spec R` and any -irreducible component is a zero locus of some minimal prime ideal. --/ +/-- Zero loci of minimal prime ideals of `R` are irreducible components in `Spec R` and any +irreducible component is a zero locus of some minimal prime ideal. -/ +@[stacks 00ES] protected def minimalPrimes.equivIrreducibleComponents : minimalPrimes R ≃o (irreducibleComponents <| PrimeSpectrum R)ᵒᵈ := by let e : {p : Ideal R | p.IsPrime ∧ ⊥ ≤ p} ≃o PrimeSpectrum R := diff --git a/Mathlib/RingTheory/TensorProduct/Basic.lean b/Mathlib/RingTheory/TensorProduct/Basic.lean index 23244caec6a64f..d1e5a3d2fd7c12 100644 --- a/Mathlib/RingTheory/TensorProduct/Basic.lean +++ b/Mathlib/RingTheory/TensorProduct/Basic.lean @@ -636,11 +636,9 @@ tensors can be directly applied by the caller (without needing `TensorProduct.on def algHomOfLinearMapTensorProduct (f : A ⊗[R] B →ₗ[S] C) (h_mul : ∀ (a₁ a₂ : A) (b₁ b₂ : B), f ((a₁ * a₂) ⊗ₜ (b₁ * b₂)) = f (a₁ ⊗ₜ b₁) * f (a₂ ⊗ₜ b₂)) (h_one : f (1 ⊗ₜ[R] 1) = 1) : A ⊗[R] B →ₐ[S] C := - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/4119 we either need to specify - the `(R := S) (A := A ⊗[R] B)` arguments, or use `set_option maxSynthPendingDepth 2 in`. - -/ + #adaptation_note /-- https://github.com/leanprover/lean4/pull/4119 + we either need to specify the `(R := S) (A := A ⊗[R] B)` arguments, + or use `set_option maxSynthPendingDepth 2 in`. -/ AlgHom.ofLinearMap f h_one <| (f.map_mul_iff (R := S) (A := A ⊗[R] B)).2 <| by -- these instances are needed by the statement of `ext`, but not by the current definition. letI : Algebra R C := RestrictScalars.algebra R S C diff --git a/Mathlib/RingTheory/Unramified/Basic.lean b/Mathlib/RingTheory/Unramified/Basic.lean index f24642133de7fa..8d6c5320108197 100644 --- a/Mathlib/RingTheory/Unramified/Basic.lean +++ b/Mathlib/RingTheory/Unramified/Basic.lean @@ -282,12 +282,9 @@ section variable (R : Type*) [CommRing R] variable (A : Type*) [CommRing A] [Algebra R A] -/-- An `R`-algebra `A` is unramified if it is formally unramified and of finite type. - -Note that the Stacks project has a different definition of unramified, and tag - shows that their definition is the -same as this one. --/ +/-- An `R`-algebra `A` is unramified if it is formally unramified and of finite type. -/ +@[stacks 00UT "Note that the Stacks project has a different definition of unramified, and tag + shows that their definition is the same as this one."] class Unramified : Prop where formallyUnramified : FormallyUnramified R A := by infer_instance finiteType : FiniteType R A := by infer_instance diff --git a/Mathlib/RingTheory/Valuation/Basic.lean b/Mathlib/RingTheory/Valuation/Basic.lean index 2d1fe63a3a6d8b..6a320028e373a0 100644 --- a/Mathlib/RingTheory/Valuation/Basic.lean +++ b/Mathlib/RingTheory/Valuation/Basic.lean @@ -253,20 +253,6 @@ lemma map_apply (f : Γ₀ →*₀ Γ'₀) (hf : Monotone f) (v : Valuation R Γ def IsEquiv (v₁ : Valuation R Γ₀) (v₂ : Valuation R Γ'₀) : Prop := ∀ r s, v₁ r ≤ v₁ s ↔ v₂ r ≤ v₂ s -/-- An ordered monoid isomorphism `Γ₀ ≃ Γ'₀` induces an equivalence -`Valuation R Γ₀ ≃ Valuation R Γ'₀`. -/ -def congr (f : Γ₀ ≃*o Γ'₀) : Valuation R Γ₀ ≃ Valuation R Γ'₀ where - toFun := map f f.toOrderIso.monotone - invFun := map f.symm f.toOrderIso.symm.monotone - left_inv ν := by ext; simp - right_inv ν := by ext; simp - -end Monoid - -section Group - -variable [LinearOrderedCommGroupWithZero Γ₀] (v : Valuation R Γ₀) {x y : R} - @[simp] theorem map_neg (x : R) : v (-x) = v x := v.toMonoidWithZeroHom.toMonoidHom.map_neg x @@ -274,12 +260,6 @@ theorem map_neg (x : R) : v (-x) = v x := theorem map_sub_swap (x y : R) : v (x - y) = v (y - x) := v.toMonoidWithZeroHom.toMonoidHom.map_sub_swap x y -theorem map_inv {R : Type*} [DivisionRing R] (v : Valuation R Γ₀) : ∀ x, v x⁻¹ = (v x)⁻¹ := - map_inv₀ _ - -theorem map_div {R : Type*} [DivisionRing R] (v : Valuation R Γ₀) : ∀ x y, v (x / y) = v x / v y := - map_div₀ _ - theorem map_sub (x y : R) : v (x - y) ≤ max (v x) (v y) := calc v (x - y) = v (x + -y) := by rw [sub_eq_add_neg] @@ -290,6 +270,8 @@ theorem map_sub_le {x y g} (hx : v x ≤ g) (hy : v y ≤ g) : v (x - y) ≤ g : rw [sub_eq_add_neg] exact v.map_add_le hx (le_trans (le_of_eq (v.map_neg y)) hy) +variable {x y : R} + theorem map_add_of_distinct_val (h : v x ≠ v y) : v (x + y) = max (v x) (v y) := by suffices ¬v (x + y) < max (v x) (v y) from or_iff_not_imp_right.1 (le_iff_eq_or_lt.1 (v.map_add x y)) this @@ -339,6 +321,26 @@ theorem map_one_sub_of_lt (h : v x < 1) : v (1 - x) = 1 := by rw [sub_eq_add_neg 1 x] simpa only [v.map_one, v.map_neg] using v.map_add_eq_of_lt_left h +/-- An ordered monoid isomorphism `Γ₀ ≃ Γ'₀` induces an equivalence +`Valuation R Γ₀ ≃ Valuation R Γ'₀`. -/ +def congr (f : Γ₀ ≃*o Γ'₀) : Valuation R Γ₀ ≃ Valuation R Γ'₀ where + toFun := map f f.toOrderIso.monotone + invFun := map f.symm f.toOrderIso.symm.monotone + left_inv ν := by ext; simp + right_inv ν := by ext; simp + +end Monoid + +section Group + +variable [LinearOrderedCommGroupWithZero Γ₀] (v : Valuation R Γ₀) {x y : R} + +theorem map_inv {R : Type*} [DivisionRing R] (v : Valuation R Γ₀) : ∀ x, v x⁻¹ = (v x)⁻¹ := + map_inv₀ _ + +theorem map_div {R : Type*} [DivisionRing R] (v : Valuation R Γ₀) : ∀ x y, v (x / y) = v x / v y := + map_div₀ _ + theorem one_lt_val_iff (v : Valuation K Γ₀) {x : K} (h : x ≠ 0) : 1 < v x ↔ v x⁻¹ < 1 := by simp [inv_lt_one₀ (v.pos_iff.2 h)] @@ -782,20 +784,6 @@ lemma map_apply (f : Γ₀ →+ Γ'₀) (ht : f ⊤ = ⊤) (hf : Monotone f) (v def IsEquiv (v₁ : AddValuation R Γ₀) (v₂ : AddValuation R Γ'₀) : Prop := Valuation.IsEquiv v₁ v₂ -end Monoid - -section Group - -variable [LinearOrderedAddCommGroupWithTop Γ₀] [Ring R] (v : AddValuation R Γ₀) {x y : R} - -@[simp] -theorem map_inv (v : AddValuation K Γ₀) {x : K} : v x⁻¹ = - (v x) := - map_inv₀ (toValuation v) x - -@[simp] -theorem map_div (v : AddValuation K Γ₀) {x y : K} : v (x / y) = v x - v y := - map_div₀ (toValuation v) x y - @[simp] theorem map_neg (x : R) : v (-x) = v x := Valuation.map_neg v x @@ -809,6 +797,8 @@ theorem map_sub (x y : R) : min (v x) (v y) ≤ v (x - y) := theorem map_le_sub {x y : R} {g : Γ₀} (hx : g ≤ v x) (hy : g ≤ v y) : g ≤ v (x - y) := Valuation.map_sub_le v hx hy +variable {x y : R} + theorem map_add_of_distinct_val (h : v x ≠ v y) : v (x + y) = @Min.min Γ₀ _ (v x) (v y) := Valuation.map_add_of_distinct_val v h @@ -831,6 +821,20 @@ theorem map_sub_eq_of_lt_right {x y : R} (hx : v y < v x) : theorem map_eq_of_lt_sub (h : v x < v (y - x)) : v y = v x := Valuation.map_eq_of_sub_lt v h +end Monoid + +section Group + +variable [LinearOrderedAddCommGroupWithTop Γ₀] [Ring R] (v : AddValuation R Γ₀) {x y : R} + +@[simp] +theorem map_inv (v : AddValuation K Γ₀) {x : K} : v x⁻¹ = - (v x) := + map_inv₀ (toValuation v) x + +@[simp] +theorem map_div (v : AddValuation K Γ₀) {x y : K} : v (x / y) = v x - v y := + map_div₀ (toValuation v) x y + end Group end Basic diff --git a/Mathlib/SetTheory/Descriptive/Tree.lean b/Mathlib/SetTheory/Descriptive/Tree.lean new file mode 100644 index 00000000000000..43a7f2f810ce28 --- /dev/null +++ b/Mathlib/SetTheory/Descriptive/Tree.lean @@ -0,0 +1,46 @@ +/- +Copyright (c) 2024 Sven Manthe. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sven Manthe +-/ +import Mathlib.Order.CompleteLattice.SetLike + +/-! +# Trees in the sense of descriptive set theory + +This file defines trees of depth `ω` in the sense of descriptive set theory as sets of finite +sequences that are stable under taking prefixes. + +## Main declarations + +* `tree A`: a (possibly infinite) tree of depth at most `ω` with nodes in `A` +-/ + +namespace Descriptive + +/-- A tree is a set of finite sequences, implemented as `List A`, that is stable under + taking prefixes. For the definition we use the equivalent property `x ++ [a] ∈ T → x ∈ T`, + which is more convenient to check. We define `tree A` as a complete sublattice of + `Set (List A)`, which coerces to the type of trees on `A`. -/ +def tree (A : Type*) : CompleteSublattice (Set (List A)) := + CompleteSublattice.mk' {T | ∀ ⦃x : List A⦄ ⦃a : A⦄, x ++ [a] ∈ T → x ∈ T} + (by rintro S hS x a ⟨t, ht, hx⟩; use t, ht, hS ht hx) + (by rintro S hS x a h T hT; exact hS hT <| h T hT) + +instance (A : Type*) : SetLike (tree A) (List A) := SetLike.instSubtypeSet + +variable {A : Type*} {T : tree A} + +lemma mem_of_append {x y : List A} (h : x ++ y ∈ T) : x ∈ T := by + induction' y with y ys ih generalizing x + · simpa using h + · exact T.prop (ih (by simpa)) + +lemma mem_of_prefix {x y : List A} (h' : x <+: y) (h : y ∈ T) : x ∈ T := by + obtain ⟨_, rfl⟩ := h'; exact mem_of_append h + +@[simp] lemma tree_eq_bot : T = ⊥ ↔ [] ∉ T where + mp := by rintro rfl; simp + mpr h := by ext x; simpa using fun h' ↦ h <| mem_of_prefix x.nil_prefix h' + +end Descriptive diff --git a/Mathlib/SetTheory/Game/Impartial.lean b/Mathlib/SetTheory/Game/Impartial.lean index e06e5beb85b350..2d6cde0fec9aa1 100644 --- a/Mathlib/SetTheory/Game/Impartial.lean +++ b/Mathlib/SetTheory/Game/Impartial.lean @@ -103,16 +103,11 @@ termination_by G variable (G : PGame) [Impartial G] theorem nonpos : ¬0 < G := by - intro h - have h' := neg_lt_neg_iff.2 h - rw [neg_zero, lt_congr_left (Equiv.symm (neg_equiv_self G))] at h' - exact (h.trans h').false + apply (lt_asymm · ?_) + rwa [← neg_lt_neg_iff, neg_zero, ← lt_congr_right (neg_equiv_self G)] theorem nonneg : ¬G < 0 := by - intro h - have h' := neg_lt_neg_iff.2 h - rw [neg_zero, lt_congr_right (Equiv.symm (neg_equiv_self G))] at h' - exact (h.trans h').false + simpa using nonpos (-G) /-- In an impartial game, either the first player always wins, or the second player always wins. -/ theorem equiv_or_fuzzy_zero : (G ≈ 0) ∨ G ‖ 0 := by diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index 8186fa7fea3ebe..ac6caee7a31665 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -1363,6 +1363,26 @@ theorem moveRight_neg_symm {x : PGame} (i) : theorem moveRight_neg_symm' {x : PGame} (i) : x.moveRight i = -(-x).moveLeft (toLeftMovesNeg i) := by simp +@[simp] +theorem forall_leftMoves_neg {x : PGame} {p : (-x).LeftMoves → Prop} : + (∀ i : (-x).LeftMoves, p i) ↔ (∀ i : x.RightMoves, p (toLeftMovesNeg i)) := + toLeftMovesNeg.forall_congr_right.symm + +@[simp] +theorem exists_leftMoves_neg {x : PGame} {p : (-x).LeftMoves → Prop} : + (∃ i : (-x).LeftMoves, p i) ↔ (∃ i : x.RightMoves, p (toLeftMovesNeg i)) := + toLeftMovesNeg.exists_congr_right.symm + +@[simp] +theorem forall_rightMoves_neg {x : PGame} {p : (-x).RightMoves → Prop} : + (∀ i : (-x).RightMoves, p i) ↔ (∀ i : x.LeftMoves, p (toRightMovesNeg i)) := + toRightMovesNeg.forall_congr_right.symm + +@[simp] +theorem exists_rightMoves_neg {x : PGame} {p : (-x).RightMoves → Prop} : + (∃ i : (-x).RightMoves, p i) ↔ (∃ i : x.LeftMoves, p (toRightMovesNeg i)) := + toRightMovesNeg.exists_congr_right.symm + theorem leftMoves_neg_cases {x : PGame} (k) {P : (-x).LeftMoves → Prop} (h : ∀ i, P <| toLeftMovesNeg i) : P k := by diff --git a/Mathlib/SetTheory/Ordinal/Exponential.lean b/Mathlib/SetTheory/Ordinal/Exponential.lean index 59b0607db2df3b..c58e66d7ac7e64 100644 --- a/Mathlib/SetTheory/Ordinal/Exponential.lean +++ b/Mathlib/SetTheory/Ordinal/Exponential.lean @@ -21,16 +21,24 @@ universe u v w namespace Ordinal -/-- The ordinal exponential, defined by transfinite recursion. -/ -instance pow : Pow Ordinal Ordinal := - ⟨fun a b => if a = 0 then 1 - b else limitRecOn b 1 (fun _ IH => IH * a) fun b _ => bsup.{u, u} b⟩ +/-- The ordinal exponential, defined by transfinite recursion. -theorem opow_def (a b : Ordinal) : - a ^ b = if a = 0 then 1 - b else limitRecOn b 1 (fun _ IH => IH * a) fun b _ => bsup.{u, u} b := - rfl +We call this `opow` in theorems in order to disambiguate from other exponentials. -/ +instance instPow : Pow Ordinal Ordinal := + ⟨fun a b ↦ if a = 0 then 1 - b else + limitRecOn b 1 (fun _ x ↦ x * a) fun o _ f ↦ ⨆ x : Iio o, f x.1 x.2⟩ + +private theorem opow_of_ne_zero {a b : Ordinal} (h : a ≠ 0) : a ^ b = + limitRecOn b 1 (fun _ x ↦ x * a) fun o _ f ↦ ⨆ x : Iio o, f x.1 x.2 := + if_neg h --- Porting note: `if_pos rfl` → `if_true` -theorem zero_opow' (a : Ordinal) : 0 ^ a = 1 - a := by simp only [opow_def, if_true] +/-- `0 ^ a = 1` if `a = 0` and `0 ^ a = 0` otherwise. -/ +theorem zero_opow' (a : Ordinal) : 0 ^ a = 1 - a := + if_pos rfl + +theorem zero_opow_le (a : Ordinal) : (0 : Ordinal) ^ a ≤ 1 := by + rw [zero_opow'] + exact sub_le_self 1 a @[simp] theorem zero_opow {a : Ordinal} (a0 : a ≠ 0) : (0 : Ordinal) ^ a = 0 := by @@ -38,29 +46,34 @@ theorem zero_opow {a : Ordinal} (a0 : a ≠ 0) : (0 : Ordinal) ^ a = 0 := by @[simp] theorem opow_zero (a : Ordinal) : a ^ (0 : Ordinal) = 1 := by - by_cases h : a = 0 - · simp only [opow_def, if_pos h, sub_zero] - · simp only [opow_def, if_neg h, limitRecOn_zero] + obtain rfl | h := eq_or_ne a 0 + · rw [zero_opow', Ordinal.sub_zero] + · rw [opow_of_ne_zero h, limitRecOn_zero] @[simp] -theorem opow_succ (a b : Ordinal) : a ^ succ b = a ^ b * a := - if h : a = 0 then by subst a; simp only [zero_opow (succ_ne_zero _), mul_zero] - else by simp only [opow_def, limitRecOn_succ, if_neg h] +theorem opow_succ (a b : Ordinal) : a ^ succ b = a ^ b * a := by + obtain rfl | h := eq_or_ne a 0 + · rw [zero_opow (succ_ne_zero b), mul_zero] + · rw [opow_of_ne_zero h, opow_of_ne_zero h, limitRecOn_succ] -theorem opow_limit {a b : Ordinal} (a0 : a ≠ 0) (h : IsLimit b) : - a ^ b = bsup.{u, u} b fun c _ => a ^ c := by - simp only [opow_def, if_neg a0]; rw [limitRecOn_limit _ _ _ _ h] +theorem opow_limit {a b : Ordinal} (ha : a ≠ 0) (hb : IsLimit b) : + a ^ b = ⨆ x : Iio b, a ^ x.1 := by + simp_rw [opow_of_ne_zero ha, limitRecOn_limit _ _ _ _ hb] theorem opow_le_of_limit {a b c : Ordinal} (a0 : a ≠ 0) (h : IsLimit b) : - a ^ b ≤ c ↔ ∀ b' < b, a ^ b' ≤ c := by rw [opow_limit a0 h, bsup_le_iff] + a ^ b ≤ c ↔ ∀ b' < b, a ^ b' ≤ c := by + rw [opow_limit a0 h, Ordinal.iSup_le_iff, Subtype.forall] + rfl theorem lt_opow_of_limit {a b c : Ordinal} (b0 : b ≠ 0) (h : IsLimit c) : a < b ^ c ↔ ∃ c' < c, a < b ^ c' := by - rw [← not_iff_not, not_exists]; simp only [not_lt, opow_le_of_limit b0 h, exists_prop, not_and] + rw [← not_iff_not, not_exists] + simp only [not_lt, opow_le_of_limit b0 h, exists_prop, not_and] @[simp] theorem opow_one (a : Ordinal) : a ^ (1 : Ordinal) = a := by - rw [← succ_zero, opow_succ]; simp only [opow_zero, one_mul] + rw [← succ_zero, opow_succ] + simp only [opow_zero, one_mul] @[simp] theorem one_opow (a : Ordinal) : (1 : Ordinal) ^ a = 1 := by @@ -317,8 +330,7 @@ theorem lt_opow_succ_log_self {b : Ordinal} (hb : 1 < b) (x : Ordinal) : theorem opow_log_le_self (b : Ordinal) {x : Ordinal} (hx : x ≠ 0) : b ^ log b x ≤ x := by rcases eq_or_ne b 0 with (rfl | b0) - · rw [zero_opow'] - exact (sub_le_self _ _).trans (one_le_iff_ne_zero.2 hx) + · exact (zero_opow_le _).trans (one_le_iff_ne_zero.2 hx) rcases lt_or_eq_of_le (one_le_iff_ne_zero.2 b0) with (hb | rfl) · refine le_of_not_lt fun h => (lt_succ (log b x)).not_le ?_ have := @csInf_le' _ _ { o | x < b ^ o } _ h diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 9d81bb8ff98acd..acc0ddcba26a4c 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -749,11 +749,11 @@ instance nf_opow (o₁ o₂) [NF o₁] [NF o₂] : NF (o₁ ^ o₂) := by haveI := (nf_repr_split' e₂).1 cases' a with a0 n a' · cases' m with m - · by_cases o₂ = 0 <;> simp only [(· ^ ·), Pow.pow, pow, opow, opowAux2, *] <;> decide + · by_cases o₂ = 0 <;> simp only [(· ^ ·), Pow.pow, opow, opowAux2, *] <;> decide · by_cases m = 0 - · simp only [(· ^ ·), Pow.pow, pow, opow, opowAux2, *, zero_def] + · simp only [(· ^ ·), Pow.pow, opow, opowAux2, *, zero_def] decide - · simp only [(· ^ ·), Pow.pow, pow, opow, opowAux2, mulNat_eq_mul, ofNat, *] + · simp only [(· ^ ·), Pow.pow, opow, opowAux2, mulNat_eq_mul, ofNat, *] infer_instance · simp only [(· ^ ·), Pow.pow, opow, opowAux2, e₁, split_eq_scale_split' e₂, mulNat_eq_mul] have := na.fst diff --git a/Mathlib/Tactic/CC/Addition.lean b/Mathlib/Tactic/CC/Addition.lean index 4ad4dc22d38f02..2fb9ddbc68213f 100644 --- a/Mathlib/Tactic/CC/Addition.lean +++ b/Mathlib/Tactic/CC/Addition.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Miyahara Kō -/ import Mathlib.Data.Option.Defs -import Mathlib.Lean.Expr.Basic import Mathlib.Tactic.CC.MkProof /-! @@ -199,12 +198,11 @@ def pushSubsingletonEq (a b : Expr) : CCM Unit := do let B ← normalize (← inferType b) -- TODO(Leo): check if the following test is a performance bottleneck if ← pureIsDefEq A B then - -- TODO(Leo): to improve performance we can create the following proof lazily - let proof ← mkAppM ``Subsingleton.elim #[a, b] + let proof ← mkAppM ``FastSubsingleton.elim #[a, b] pushEq a b proof else let some AEqB ← getEqProof A B | failure - let proof ← mkAppM ``Subsingleton.helim #[AEqB, a, b] + let proof ← mkAppM ``FastSubsingleton.helim #[AEqB, a, b] pushHEq a b proof /-- Given the equivalent expressions `oldRoot` and `newRoot` the root of `oldRoot` is @@ -940,9 +938,7 @@ partial def applySimpleEqvs (e : Expr) : CCM Unit := do to the todo list or register `e` as the canonical form of itself. -/ partial def processSubsingletonElem (e : Expr) : CCM Unit := do let type ← inferType e - -- TODO: this is likely to become a bottleneck. See e.g. - -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/convert.20is.20often.20slow/near/433830798 - let ss ← synthInstance? (← mkAppM ``Subsingleton #[type]) + let ss ← synthInstance? (← mkAppM ``FastSubsingleton #[type]) if ss.isNone then return -- type is not a subsingleton let type ← normalize type -- Make sure type has been internalized diff --git a/Mathlib/Tactic/CC/Datatypes.lean b/Mathlib/Tactic/CC/Datatypes.lean index 6dc3e10fd6aa5c..11aac8d38934ed 100644 --- a/Mathlib/Tactic/CC/Datatypes.lean +++ b/Mathlib/Tactic/CC/Datatypes.lean @@ -3,9 +3,9 @@ Copyright (c) 2016 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Miyahara Kō -/ -import Lean.Meta.Tactic.Rfl import Batteries.Data.RBMap.Basic import Mathlib.Lean.Meta.Basic +import Mathlib.Lean.Meta.CongrTheorems import Mathlib.Data.Ordering.Basic /-! @@ -428,7 +428,8 @@ Note that this only works for two-argument relations: `ModEq n` and `ModEq m` ar same. -/ abbrev SymmCongruences := Std.HashMap SymmCongruencesKey (List (Expr × Name)) -/-- Stores the root representatives of subsingletons. -/ +/-- Stores the root representatives of subsingletons, this uses `FastSingleton` instead of +`Subsingleton`. -/ abbrev SubsingletonReprs := RBExprMap Expr /-- Stores the root representatives of `.instImplicit` arguments. -/ diff --git a/Mathlib/Tactic/CC/MkProof.lean b/Mathlib/Tactic/CC/MkProof.lean index 5f17a75846cc5a..25140ed57c1d82 100644 --- a/Mathlib/Tactic/CC/MkProof.lean +++ b/Mathlib/Tactic/CC/MkProof.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Miyahara Kō -/ import Batteries.Data.RBMap.Alter -import Mathlib.Logic.Basic import Mathlib.Tactic.CC.Datatypes import Mathlib.Tactic.CC.Lemmas import Mathlib.Tactic.Relation.Rfl diff --git a/Mathlib/Tactic/ComputeDegree.lean b/Mathlib/Tactic/ComputeDegree.lean index f94993bf7de959..f82b8e01897655 100644 --- a/Mathlib/Tactic/ComputeDegree.lean +++ b/Mathlib/Tactic/ComputeDegree.lean @@ -440,14 +440,14 @@ def miscomputedDegree? (deg : Expr) : List Expr → List MessageData * `degree f ≤ d`, * `coeff f d = r`, if `d` is the degree of `f`. -The tactic may leave goals of the form `d' = d` `d' ≤ d`, or `r ≠ 0`, where `d'` in `ℕ` or +The tactic may leave goals of the form `d' = d`, `d' ≤ d`, or `r ≠ 0`, where `d'` in `ℕ` or `WithBot ℕ` is the tactic's guess of the degree, and `r` is the coefficient's guess of the leading coefficient of `f`. `compute_degree` applies `norm_num` to the left-hand side of all side goals, trying to close them. The variant `compute_degree!` first applies `compute_degree`. -Then it uses `norm_num` on all the whole remaining goals and tries `assumption`. +Then it uses `norm_num` on all the remaining goals and tries `assumption`. -/ syntax (name := computeDegree) "compute_degree" "!"? : tactic diff --git a/Mathlib/Tactic/CongrExclamation.lean b/Mathlib/Tactic/CongrExclamation.lean index 3bdb8ef60024a9..dfa8e21eef0fe2 100644 --- a/Mathlib/Tactic/CongrExclamation.lean +++ b/Mathlib/Tactic/CongrExclamation.lean @@ -425,11 +425,6 @@ def Lean.MVarId.congrImplies?' (mvarId : MVarId) : MetaM (Option (List MVarId)) | throwError "unexpected number of goals" return [mvarId₁, mvarId₂] -protected theorem FastSubsingleton.helim {α β : Sort u} [FastSubsingleton α] - (h₂ : α = β) (a : α) (b : β) : HEq a b := by - have : Subsingleton α := FastSubsingleton.inst - exact Subsingleton.helim h₂ a b - /-- Try to apply `Subsingleton.helim` if the goal is a `HEq`. Tries synthesizing a `Subsingleton` instance for both the LHS and the RHS. diff --git a/Mathlib/Tactic/FunProp.lean b/Mathlib/Tactic/FunProp.lean index 1e9f65989489ad..a52e272d5b9561 100644 --- a/Mathlib/Tactic/FunProp.lean +++ b/Mathlib/Tactic/FunProp.lean @@ -37,7 +37,7 @@ example (y : ℝ) (hy : y ≠ 0) : ContinuousAt (fun x : ℝ => 1/x) y := by fun **Basic debugging:** The most common issue is that a function is missing the appropriate theorem. For example: ```lean -import Mathlib.Data.Complex.Exponential +import Mathlib.Data.Complex.Trigonometric example : Continuous (fun x : ℝ => x * Real.sin x) := by fun_prop ``` Fails with the error: diff --git a/Mathlib/Tactic/Linter/TextBased.lean b/Mathlib/Tactic/Linter/TextBased.lean index c0df2b14e30dad..5786f3360de338 100644 --- a/Mathlib/Tactic/Linter/TextBased.lean +++ b/Mathlib/Tactic/Linter/TextBased.lean @@ -51,8 +51,8 @@ deriving BEq /-- Possible errors that text-based linters can report. -/ -- We collect these in one inductive type to centralise error reporting. inductive StyleError where - /-- The bare string "Adaptation note" (or variants thereof): instead, the - #adaptation_note command should be used. -/ + /-- The bare string "Adaptation note" (or variants thereof): + instead, the #adaptation_note command should be used. -/ | adaptationNote /-- A line ends with windows line endings (\r\n) instead of unix ones (\n). -/ | windowsLineEnding diff --git a/Mathlib/Tactic/NormNum/Basic.lean b/Mathlib/Tactic/NormNum/Basic.lean index a1d9e39f9a55bd..91ce6bebb6980a 100644 --- a/Mathlib/Tactic/NormNum/Basic.lean +++ b/Mathlib/Tactic/NormNum/Basic.lean @@ -25,14 +25,6 @@ See other files in this directory for many more plugins. universe u -#adaptation_note -/-- -Since https://github.com/leanprover/lean4/pull/5338, -the unused variable linter can not see usages of variables in -`haveI' : ⋯ =Q ⋯ := ⟨⟩` clauses, so generates many false positives. --/ -set_option linter.unusedVariables false - namespace Mathlib open Lean open Meta diff --git a/Mathlib/Tactic/NormNum/DivMod.lean b/Mathlib/Tactic/NormNum/DivMod.lean index 90f6938e16c6ba..c07a80e7e9788d 100644 --- a/Mathlib/Tactic/NormNum/DivMod.lean +++ b/Mathlib/Tactic/NormNum/DivMod.lean @@ -13,14 +13,6 @@ This file adds support for the `%`, `/`, and `∣` (divisibility) operators on ` to the `norm_num` tactic. -/ -#adaptation_note -/-- -Since https://github.com/leanprover/lean4/pull/5338, -the unused variable linter can not see usages of variables in -`haveI' : ⋯ =Q ⋯ := ⟨⟩` clauses, so generates many false positives. --/ -set_option linter.unusedVariables false - namespace Mathlib open Lean open Meta diff --git a/Mathlib/Tactic/NormNum/Eq.lean b/Mathlib/Tactic/NormNum/Eq.lean index d467ab0571729b..277f0479a4a029 100644 --- a/Mathlib/Tactic/NormNum/Eq.lean +++ b/Mathlib/Tactic/NormNum/Eq.lean @@ -11,14 +11,6 @@ import Mathlib.Tactic.NormNum.Inv variable {α : Type*} -#adaptation_note -/-- -Since https://github.com/leanprover/lean4/pull/5338, -the unused variable linter can not see usages of variables in -`haveI' : ⋯ =Q ⋯ := ⟨⟩` clauses, so generates many false positives. --/ -set_option linter.unusedVariables false - open Lean Meta Qq namespace Mathlib.Meta.NormNum diff --git a/Mathlib/Tactic/NormNum/GCD.lean b/Mathlib/Tactic/NormNum/GCD.lean index c9bcc4f4c2c986..cb0a2aa2d169a4 100644 --- a/Mathlib/Tactic/NormNum/GCD.lean +++ b/Mathlib/Tactic/NormNum/GCD.lean @@ -16,14 +16,6 @@ Note that `Nat.coprime` is reducible and defined in terms of `Nat.gcd`, so the ` also indirectly provides a `Nat.coprime` extension. -/ -#adaptation_note -/-- -Since https://github.com/leanprover/lean4/pull/5338, -the unused variable linter can not see usages of variables in -`haveI' : ⋯ =Q ⋯ := ⟨⟩` clauses, so generates many false positives. --/ -set_option linter.unusedVariables false - namespace Tactic namespace NormNum diff --git a/Mathlib/Tactic/NormNum/Ineq.lean b/Mathlib/Tactic/NormNum/Ineq.lean index 44013a8b76cedb..0c27395a582bef 100644 --- a/Mathlib/Tactic/NormNum/Ineq.lean +++ b/Mathlib/Tactic/NormNum/Ineq.lean @@ -13,14 +13,6 @@ import Mathlib.Algebra.Order.Ring.Cast # `norm_num` extensions for inequalities. -/ -#adaptation_note -/-- -Since https://github.com/leanprover/lean4/pull/5338, -the unused variable linter can not see usages of variables in -`haveI' : ⋯ =Q ⋯ := ⟨⟩` clauses, so generates many false positives. --/ -set_option linter.unusedVariables false - open Lean Meta Qq namespace Mathlib.Meta.NormNum diff --git a/Mathlib/Tactic/NormNum/Inv.lean b/Mathlib/Tactic/NormNum/Inv.lean index c722b8cf4c5870..82c34adb85e5a1 100644 --- a/Mathlib/Tactic/NormNum/Inv.lean +++ b/Mathlib/Tactic/NormNum/Inv.lean @@ -13,14 +13,6 @@ import Mathlib.Algebra.Field.Basic variable {u : Lean.Level} -#adaptation_note -/-- -Since https://github.com/leanprover/lean4/pull/5338, -the unused variable linter can not see usages of variables in -`haveI' : ⋯ =Q ⋯ := ⟨⟩` clauses, so generates many false positives. --/ -set_option linter.unusedVariables false - namespace Mathlib.Meta.NormNum open Lean.Meta Qq diff --git a/Mathlib/Tactic/NormNum/LegendreSymbol.lean b/Mathlib/Tactic/NormNum/LegendreSymbol.lean index 166fc060cfcdca..dfb3def6ba7c4e 100644 --- a/Mathlib/Tactic/NormNum/LegendreSymbol.lean +++ b/Mathlib/Tactic/NormNum/LegendreSymbol.lean @@ -356,14 +356,6 @@ namespace NormNum open Lean Elab Tactic Qq Mathlib.Meta.NormNum -#adaptation_note -/-- -Since https://github.com/leanprover/lean4/pull/5338, -the unused variable linter can not see usages of variables in -`haveI' : ⋯ =Q ⋯ := ⟨⟩` clauses, so generates many false positives. --/ -set_option linter.unusedVariables false - /-- This is the `norm_num` plug-in that evaluates Jacobi symbols. -/ @[norm_num jacobiSym _ _] def evalJacobiSym : NormNumExt where eval {u α} e := do diff --git a/Mathlib/Tactic/NormNum/Pow.lean b/Mathlib/Tactic/NormNum/Pow.lean index 467a97ad68b3f0..72ced934fd53d0 100644 --- a/Mathlib/Tactic/NormNum/Pow.lean +++ b/Mathlib/Tactic/NormNum/Pow.lean @@ -10,14 +10,6 @@ import Mathlib.Tactic.NormNum.Basic ## `norm_num` plugin for `^`. -/ -#adaptation_note -/-- -Since https://github.com/leanprover/lean4/pull/5338, -the unused variable linter can not see usages of variables in -`haveI' : ⋯ =Q ⋯ := ⟨⟩` clauses, so generates many false positives. --/ -set_option linter.unusedVariables false - namespace Mathlib open Lean open Meta @@ -235,9 +227,7 @@ theorem isRat_zpow_neg {α : Type*} [DivisionRing α] {a : α} {b : ℤ} {nb : IsRat (a^b) num den := by rwa [pb.out, Int.cast_negOfNat, zpow_neg, zpow_natCast] -#adaptation_note -/-- -Prior to https://github.com/leanprover/lean4/pull/4096, +#adaptation_note /-- https://github.com/leanprover/lean4/pull/4096 the repeated ``` have h : $e =Q (HPow.hPow (γ := $α) $a $b) := ⟨⟩ diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 28b8320d761a63..649dc2de97a8bb 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -813,10 +813,11 @@ theorem continuous_of_continuousAt_one₂ {H M : Type*} [CommMonoid M] [Topologi @[to_additive] lemma TopologicalGroup.isInducing_iff_nhds_one - {H : Type*} [Group H] [TopologicalSpace H] [TopologicalGroup H] {f : G →* H} : + {H : Type*} [Group H] [TopologicalSpace H] [TopologicalGroup H] {F : Type*} + [FunLike F G H] [MonoidHomClass F G H] {f : F} : Topology.IsInducing f ↔ 𝓝 (1 : G) = (𝓝 (1 : H)).comap f := by rw [Topology.isInducing_iff_nhds] - refine ⟨(f.map_one ▸ · 1), fun hf x ↦ ?_⟩ + refine ⟨(map_one f ▸ · 1), fun hf x ↦ ?_⟩ rw [← nhds_translation_mul_inv, ← nhds_translation_mul_inv (f x), Filter.comap_comap, hf, Filter.comap_comap] congr 1 diff --git a/Mathlib/Topology/Algebra/Indicator.lean b/Mathlib/Topology/Algebra/Indicator.lean index 189c9d6f29038b..87a704835566a7 100644 --- a/Mathlib/Topology/Algebra/Indicator.lean +++ b/Mathlib/Topology/Algebra/Indicator.lean @@ -5,6 +5,7 @@ Authors: PFR contributors -/ import Mathlib.Algebra.Group.Indicator import Mathlib.Topology.ContinuousOn +import Mathlib.Topology.Clopen /-! # Continuity of indicator functions @@ -36,3 +37,8 @@ theorem ContinuousOn.continuousAt_mulIndicator (hf : ContinuousOn f (interior s) ⟨interior s, hs, Set.eqOn_mulIndicator.symm.mono interior_subset⟩ · exact ContinuousAt.congr continuousAt_const <| Filter.eventuallyEq_iff_exists_mem.mpr ⟨sᶜ, mem_interior_iff_mem_nhds.mp h, Set.eqOn_mulIndicator'.symm⟩ + +@[to_additive] +lemma IsClopen.continuous_mulIndicator (hs : IsClopen s) (hf : Continuous f) : + Continuous (s.mulIndicator f) := + hf.mulIndicator (by simp [isClopen_iff_frontier_eq_empty.mp hs]) diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean b/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean index afef581ea8ffb4..70a2cc032d21ed 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean @@ -55,6 +55,32 @@ theorem cauchySeq_prod_of_tendsto_cofinite_one {f : α → G} (hf : Tendsto f co intro i hi simpa using Finset.disjoint_left.mp ht hi +/-- Let `G` be a nonarchimedean abelian group, and let `f : ℕ → G` be a function +such that the quotients `f (n + 1) / f n` tend to one. Then the function is a Cauchy sequence. -/ +@[to_additive "Let `G` be a nonarchimedean additive abelian group, and let `f : ℕ → G` be a +function such that the differences `f (n + 1) - f n` tend to zero. +Then the function is a Cauchy sequence."] +lemma cauchySeq_of_tendsto_div_nhds_one {f : ℕ → G} + (hf : Tendsto (fun n ↦ f (n + 1) / f n) atTop (𝓝 1)) : + CauchySeq f := by + suffices Tendsto (fun p : ℕ × ℕ ↦ f p.2 / f p.1) atTop (𝓝 1) by simpa [CauchySeq, + cauchy_map_iff, prod_atTop_atTop_eq, uniformity_eq_comap_nhds_one G, atTop_neBot] + rw [tendsto_atTop'] + intro s hs + obtain ⟨t, ht⟩ := is_nonarchimedean s hs + obtain ⟨N, hN⟩ : ∃ N : ℕ, ∀ b, N ≤ b → f (b + 1) / f b ∈ t := by + simpa using tendsto_def.mp hf t t.mem_nhds_one + refine ⟨(N, N), ?_⟩ + rintro ⟨M, M'⟩ ⟨(hMN : N ≤ M), (hMN' : N ≤ M')⟩ + apply ht + wlog h : M ≤ M' generalizing M M' + · simpa [inv_div] using t.inv_mem <| this _ _ hMN' hMN (le_of_not_ge h) + obtain ⟨k, rfl⟩ := Nat.exists_eq_add_of_le h + clear h hMN' + induction k with + | zero => simpa using one_mem t + | succ k ih => simpa using t.mul_mem (hN _ (by omega : N ≤ M + k)) ih + /-- Let `G` be a complete nonarchimedean multiplicative abelian group, and let `f : α → G` be a function that tends to one on the filter of cofinite sets. Then `f` is unconditionally multipliable. -/ diff --git a/Mathlib/Topology/Algebra/Module/Equiv.lean b/Mathlib/Topology/Algebra/Module/Equiv.lean index 9df3e155b17c7e..faaa18273a54cc 100644 --- a/Mathlib/Topology/Algebra/Module/Equiv.lean +++ b/Mathlib/Topology/Algebra/Module/Equiv.lean @@ -852,6 +852,28 @@ def piFinTwo (M : Fin 2 → Type*) [∀ i, AddCommMonoid (M i)] [∀ i, Module R def finTwoArrow : (Fin 2 → M) ≃L[R] M × M := { piFinTwo R fun _ => M with toLinearEquiv := LinearEquiv.finTwoArrow R M } +section +variable {n : ℕ} {R : Type*} {M : Fin n.succ → Type*} {N : Type*} +variable [Semiring R] +variable [∀ i, AddCommMonoid (M i)] [∀ i, Module R (M i)] [∀ i, TopologicalSpace (M i)] + +variable (R M) in +/-- `Fin.consEquiv` as a continuous linear equivalence. -/ +@[simps!] +def _root_.Fin.consEquivL : (M 0 × Π i, M (Fin.succ i)) ≃L[R] (Π i, M i) where + __ := Fin.consLinearEquiv R M + continuous_toFun := continuous_id.fst.finCons continuous_id.snd + continuous_invFun := .prod_mk (continuous_apply 0) (by continuity) + +/-- `Fin.cons` in the codomain of continuous linear maps. -/ +abbrev _root_.ContinuousLinearMap.finCons + [AddCommMonoid N] [Module R N] [TopologicalSpace N] + (f : N →L[R] M 0) (fs : N →L[R] Π i, M (Fin.succ i)) : + N →L[R] Π i, M i := + Fin.consEquivL R M ∘L f.prod fs + +end + end end ContinuousLinearEquiv diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index f136b5fd395686..025c3e373cc410 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -9,6 +9,7 @@ import Mathlib.Topology.Algebra.MulAction import Mathlib.Algebra.BigOperators.Pi import Mathlib.Algebra.Group.ULift import Mathlib.Topology.ContinuousMap.Defs +import Mathlib.Algebra.Group.Submonoid.Basic /-! # Theory of topological monoids @@ -567,6 +568,12 @@ theorem exists_open_nhds_one_mul_subset {U : Set M} (hU : U ∈ 𝓝 (1 : M)) : ∃ V : Set M, IsOpen V ∧ (1 : M) ∈ V ∧ V * V ⊆ U := by simpa only [mul_subset_iff] using exists_open_nhds_one_split hU +@[to_additive] +theorem Filter.HasBasis.mul_self {p : ι → Prop} {s : ι → Set M} (h : (𝓝 1).HasBasis p s) : + (𝓝 1).HasBasis p fun i => s i * s i := by + rw [← nhds_mul_nhds_one, ← map₂_mul, ← map_uncurry_prod] + simpa only [← image_mul_prod] using h.prod_self.map _ + end MulOneClass section ContinuousMul diff --git a/Mathlib/Topology/Algebra/Ring/Basic.lean b/Mathlib/Topology/Algebra/Ring/Basic.lean index 613f1181d73b99..84c683b719cfe3 100644 --- a/Mathlib/Topology/Algebra/Ring/Basic.lean +++ b/Mathlib/Topology/Algebra/Ring/Basic.lean @@ -183,7 +183,8 @@ instance [NonUnitalNonAssocRing α] [NonUnitalNonAssocRing β] [TopologicalRing end -#adaptation_note /-- nightly-2024-04-08, needed to help `Pi.instTopologicalSemiring` -/ +#adaptation_note /-- nightly-2024-04-08 +needed to help `Pi.instTopologicalSemiring` -/ instance {β : Type*} {C : β → Type*} [∀ b, TopologicalSpace (C b)] [∀ b, NonUnitalNonAssocSemiring (C b)] [∀ b, TopologicalSemiring (C b)] : ContinuousAdd ((b : β) → C b) := diff --git a/Mathlib/Topology/Category/TopCat/Opens.lean b/Mathlib/Topology/Category/TopCat/Opens.lean index 4e4930958fd0e1..cd24b21e34cb9b 100644 --- a/Mathlib/Topology/Category/TopCat/Opens.lean +++ b/Mathlib/Topology/Category/TopCat/Opens.lean @@ -154,16 +154,13 @@ theorem map_id_obj (U : Opens X) : (map (𝟙 X)).obj U = U := let ⟨_, _⟩ := U rfl -@[simp 1100] +@[simp] theorem map_id_obj' (U) (p) : (map (𝟙 X)).obj ⟨U, p⟩ = ⟨U, p⟩ := rfl -@[simp 1100] -theorem map_id_obj_unop (U : (Opens X)ᵒᵖ) : (map (𝟙 X)).obj (unop U) = unop U := - let ⟨_, _⟩ := U.unop - rfl +theorem map_id_obj_unop (U : (Opens X)ᵒᵖ) : (map (𝟙 X)).obj (unop U) = unop U := by + simp -@[simp 1100] theorem op_map_id_obj (U : (Opens X)ᵒᵖ) : (map (𝟙 X)).op.obj U = U := by simp @[simp] diff --git a/Mathlib/Topology/Compactness/CompactlyGeneratedSpace.lean b/Mathlib/Topology/Compactness/CompactlyGeneratedSpace.lean index 09080f8e326cda..0939e3bde7b5ed 100644 --- a/Mathlib/Topology/Compactness/CompactlyGeneratedSpace.lean +++ b/Mathlib/Topology/Compactness/CompactlyGeneratedSpace.lean @@ -94,12 +94,9 @@ instance (X : Type v) [t : TopologicalSpace X] [DiscreteTopology X] : rw [DiscreteTopology.eq_bot (t := t)] exact bot_le -#adaptation_note -/-- -The new unused variable linter in -https://github.com/leanprover/lean4/pull/5338 -flags `[tY : TopologicalSpace Y]`, but we want to use this as a named argument. --/ +#adaptation_note /-- https://github.com/leanprover/lean4/pull/5338 +The new unused variable linter flags `[tY : TopologicalSpace Y]`, +but we want to use this as a named argument. -/ set_option linter.unusedVariables false in /-- Let `f : X → Y`. Suppose that to prove that `f` is continuous, it suffices to show that for every compact Hausdorff space `K` and every continuous map `g : K → X`, `f ∘ g` is continuous. diff --git a/Mathlib/Topology/EMetricSpace/Defs.lean b/Mathlib/Topology/EMetricSpace/Defs.lean index e3582222d36109..0c9f840cb9455f 100644 --- a/Mathlib/Topology/EMetricSpace/Defs.lean +++ b/Mathlib/Topology/EMetricSpace/Defs.lean @@ -596,7 +596,7 @@ theorem edist_pos {x y : γ} : 0 < edist x y ↔ x ≠ y := by simp [← not_le] /-- Two points coincide if their distance is `< ε` for all positive ε -/ theorem eq_of_forall_edist_le {x y : γ} (h : ∀ ε > 0, edist x y ≤ ε) : x = y := - eq_of_edist_eq_zero (eq_of_le_of_forall_le_of_dense bot_le h) + eq_of_edist_eq_zero (eq_of_le_of_forall_lt_imp_le_of_dense bot_le h) /-- Auxiliary function to replace the uniformity on an emetric space with a uniformity which is equal to the original one, but maybe not defeq. diff --git a/Mathlib/Topology/Instances/Matrix.lean b/Mathlib/Topology/Instances/Matrix.lean index 13decc88c4c0a3..c465b090b65452 100644 --- a/Mathlib/Topology/Instances/Matrix.lean +++ b/Mathlib/Topology/Instances/Matrix.lean @@ -73,16 +73,17 @@ theorem Continuous.matrix_elem {A : X → Matrix m n R} (hA : Continuous A) (i : Continuous fun x => A x i j := (continuous_apply_apply i j).comp hA -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_map [TopologicalSpace S] {A : X → Matrix m n S} {f : S → R} (hA : Continuous A) (hf : Continuous f) : Continuous fun x => (A x).map f := continuous_matrix fun _ _ => hf.comp <| hA.matrix_elem _ _ -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_transpose {A : X → Matrix m n R} (hA : Continuous A) : Continuous fun x => (A x)ᵀ := continuous_matrix fun i j => hA.matrix_elem j i +@[continuity, fun_prop] theorem Continuous.matrix_conjTranspose [Star R] [ContinuousStar R] {A : X → Matrix m n R} (hA : Continuous A) : Continuous fun x => (A x)ᴴ := hA.matrix_transpose.matrix_map continuous_star @@ -90,22 +91,22 @@ theorem Continuous.matrix_conjTranspose [Star R] [ContinuousStar R] {A : X → M instance [Star R] [ContinuousStar R] : ContinuousStar (Matrix m m R) := ⟨continuous_id.matrix_conjTranspose⟩ -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_col {ι : Type*} {A : X → n → R} (hA : Continuous A) : Continuous fun x => col ι (A x) := continuous_matrix fun i _ => (continuous_apply i).comp hA -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_row {ι : Type*} {A : X → n → R} (hA : Continuous A) : Continuous fun x => row ι (A x) := continuous_matrix fun _ _ => (continuous_apply _).comp hA -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_diagonal [Zero R] [DecidableEq n] {A : X → n → R} (hA : Continuous A) : Continuous fun x => diagonal (A x) := continuous_matrix fun i _ => ((continuous_apply i).comp hA).if_const _ continuous_zero -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_dotProduct [Fintype n] [Mul R] [AddCommMonoid R] [ContinuousAdd R] [ContinuousMul R] {A : X → n → R} {B : X → n → R} (hA : Continuous A) (hB : Continuous B) : Continuous fun x => dotProduct (A x) (B x) := @@ -113,7 +114,7 @@ theorem Continuous.matrix_dotProduct [Fintype n] [Mul R] [AddCommMonoid R] [Cont ((continuous_apply i).comp hA).mul ((continuous_apply i).comp hB) /-- For square matrices the usual `continuous_mul` can be used. -/ -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_mul [Fintype n] [Mul R] [AddCommMonoid R] [ContinuousAdd R] [ContinuousMul R] {A : X → Matrix m n R} {B : X → Matrix n p R} (hA : Continuous A) (hB : Continuous B) : Continuous fun x => A x * B x := @@ -130,34 +131,34 @@ instance [Fintype n] [NonUnitalNonAssocSemiring R] [TopologicalSemiring R] : instance Matrix.topologicalRing [Fintype n] [NonUnitalNonAssocRing R] [TopologicalRing R] : TopologicalRing (Matrix n n R) where -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_vecMulVec [Mul R] [ContinuousMul R] {A : X → m → R} {B : X → n → R} (hA : Continuous A) (hB : Continuous B) : Continuous fun x => vecMulVec (A x) (B x) := continuous_matrix fun _ _ => ((continuous_apply _).comp hA).mul ((continuous_apply _).comp hB) -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_mulVec [NonUnitalNonAssocSemiring R] [ContinuousAdd R] [ContinuousMul R] [Fintype n] {A : X → Matrix m n R} {B : X → n → R} (hA : Continuous A) (hB : Continuous B) : Continuous fun x => A x *ᵥ B x := continuous_pi fun i => ((continuous_apply i).comp hA).matrix_dotProduct hB -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_vecMul [NonUnitalNonAssocSemiring R] [ContinuousAdd R] [ContinuousMul R] [Fintype m] {A : X → m → R} {B : X → Matrix m n R} (hA : Continuous A) (hB : Continuous B) : Continuous fun x => A x ᵥ* B x := continuous_pi fun _i => hA.matrix_dotProduct <| continuous_pi fun _j => hB.matrix_elem _ _ -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_submatrix {A : X → Matrix l n R} (hA : Continuous A) (e₁ : m → l) (e₂ : p → n) : Continuous fun x => (A x).submatrix e₁ e₂ := continuous_matrix fun _i _j => hA.matrix_elem _ _ -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_reindex {A : X → Matrix l n R} (hA : Continuous A) (e₁ : l ≃ m) (e₂ : n ≃ p) : Continuous fun x => reindex e₁ e₂ (A x) := hA.matrix_submatrix _ _ -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_diag {A : X → Matrix n n R} (hA : Continuous A) : Continuous fun x => Matrix.diag (A x) := continuous_pi fun _ => hA.matrix_elem _ _ @@ -166,19 +167,19 @@ theorem Continuous.matrix_diag {A : X → Matrix n n R} (hA : Continuous A) : theorem continuous_matrix_diag : Continuous (Matrix.diag : Matrix n n R → n → R) := show Continuous fun x : Matrix n n R => Matrix.diag x from continuous_id.matrix_diag -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_trace [Fintype n] [AddCommMonoid R] [ContinuousAdd R] {A : X → Matrix n n R} (hA : Continuous A) : Continuous fun x => trace (A x) := continuous_finset_sum _ fun _ _ => hA.matrix_elem _ _ -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_det [Fintype n] [DecidableEq n] [CommRing R] [TopologicalRing R] {A : X → Matrix n n R} (hA : Continuous A) : Continuous fun x => (A x).det := by simp_rw [Matrix.det_apply] refine continuous_finset_sum _ fun l _ => Continuous.const_smul ?_ _ exact continuous_finset_prod _ fun l _ => hA.matrix_elem _ _ -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_updateCol [DecidableEq n] (i : n) {A : X → Matrix m n R} {B : X → m → R} (hA : Continuous A) (hB : Continuous B) : Continuous fun x => (A x).updateCol i (B x) := @@ -189,18 +190,18 @@ theorem Continuous.matrix_updateCol [DecidableEq n] (i : n) {A : X → Matrix m @[deprecated (since := "2024-12-11")] alias Continuous.matrix_updateColumn := Continuous.matrix_updateCol -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_updateRow [DecidableEq m] (i : m) {A : X → Matrix m n R} {B : X → n → R} (hA : Continuous A) (hB : Continuous B) : Continuous fun x => (A x).updateRow i (B x) := hA.update i hB -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_cramer [Fintype n] [DecidableEq n] [CommRing R] [TopologicalRing R] {A : X → Matrix n n R} {B : X → n → R} (hA : Continuous A) (hB : Continuous B) : Continuous fun x => cramer (A x) (B x) := continuous_pi fun _ => (hA.matrix_updateCol _ hB).matrix_det -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_adjugate [Fintype n] [DecidableEq n] [CommRing R] [TopologicalRing R] {A : X → Matrix n n R} (hA : Continuous A) : Continuous fun x => (A x).adjugate := continuous_matrix fun _j k => @@ -215,7 +216,7 @@ theorem continuousAt_matrix_inv [Fintype n] [DecidableEq n] [CommRing R] [Topolo -- lemmas about functions in `Data/Matrix/Block.lean` section BlockMatrices -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_fromBlocks {A : X → Matrix n l R} {B : X → Matrix n m R} {C : X → Matrix p l R} {D : X → Matrix p m R} (hA : Continuous A) (hB : Continuous B) (hC : Continuous C) (hD : Continuous D) : @@ -223,18 +224,18 @@ theorem Continuous.matrix_fromBlocks {A : X → Matrix n l R} {B : X → Matrix continuous_matrix <| by rintro (i | i) (j | j) <;> refine Continuous.matrix_elem ?_ i j <;> assumption -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_blockDiagonal [Zero R] [DecidableEq p] {A : X → p → Matrix m n R} (hA : Continuous A) : Continuous fun x => blockDiagonal (A x) := continuous_matrix fun ⟨i₁, i₂⟩ ⟨j₁, _j₂⟩ => (((continuous_apply i₂).comp hA).matrix_elem i₁ j₁).if_const _ continuous_zero -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_blockDiag {A : X → Matrix (m × p) (n × p) R} (hA : Continuous A) : Continuous fun x => blockDiag (A x) := continuous_pi fun _i => continuous_matrix fun _j _k => hA.matrix_elem _ _ -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_blockDiagonal' [Zero R] [DecidableEq l] {A : X → ∀ i, Matrix (m' i) (n' i) R} (hA : Continuous A) : Continuous fun x => blockDiagonal' (A x) := @@ -245,7 +246,7 @@ theorem Continuous.matrix_blockDiagonal' [Zero R] [DecidableEq l] exact ((continuous_apply i₁).comp hA).matrix_elem i₂ j₂ · exact continuous_const -@[continuity] +@[continuity, fun_prop] theorem Continuous.matrix_blockDiag' {A : X → Matrix (Σi, m' i) (Σi, n' i) R} (hA : Continuous A) : Continuous fun x => blockDiag' (A x) := continuous_pi fun _i => continuous_matrix fun _j _k => hA.matrix_elem _ _ diff --git a/Mathlib/Topology/List.lean b/Mathlib/Topology/List.lean index be04e98105dbd3..10a7e8b0991c05 100644 --- a/Mathlib/Topology/List.lean +++ b/Mathlib/Topology/List.lean @@ -56,8 +56,6 @@ theorem nhds_list (as : List α) : 𝓝 as = traverse 𝓝 as := by have : List.Forall₂ (fun a s => IsOpen s ∧ a ∈ s) u v := by refine List.Forall₂.flip ?_ replace hv := hv.flip - #adaptation_note /-- nightly-2024-03-16: simp was - simp only [List.forall₂_and_left, flip] at hv ⊢ -/ simp only [List.forall₂_and_left, Function.flip_def] at hv ⊢ exact ⟨hv.1, hu.flip⟩ refine mem_of_superset ?_ hvs diff --git a/Mathlib/Topology/MetricSpace/Defs.lean b/Mathlib/Topology/MetricSpace/Defs.lean index e678bdcfedb6b9..0ada8be44573bc 100644 --- a/Mathlib/Topology/MetricSpace/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Defs.lean @@ -78,7 +78,7 @@ theorem dist_pos {x y : γ} : 0 < dist x y ↔ x ≠ y := by simpa only [not_le] using not_congr dist_le_zero theorem eq_of_forall_dist_le {x y : γ} (h : ∀ ε > 0, dist x y ≤ ε) : x = y := - eq_of_dist_eq_zero (eq_of_le_of_forall_le_of_dense dist_nonneg h) + eq_of_dist_eq_zero (eq_of_le_of_forall_lt_imp_le_of_dense dist_nonneg h) /-- Deduce the equality of points from the vanishing of the nonnegative distance -/ theorem eq_of_nndist_eq_zero {x y : γ} : nndist x y = 0 → x = y := by diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean index a564dcfd1a9098..caf9feed00c8f0 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean @@ -311,7 +311,7 @@ theorem hausdorffDist_optimal {X : Type u} [MetricSpace X] [CompactSpace X] [Non let Fb := candidatesBOfCandidates F Fgood have : hausdorffDist (range (optimalGHInjl X Y)) (range (optimalGHInjr X Y)) ≤ HD Fb := hausdorffDist_optimal_le_HD _ _ (candidatesBOfCandidates_mem F Fgood) - refine le_trans this (le_of_forall_le_of_dense fun r hr => ?_) + refine le_trans this (le_of_forall_gt_imp_ge_of_dense fun r hr => ?_) have I1 : ∀ x : X, (⨅ y, Fb (inl x, inr y)) ≤ r := by intro x have : f (inl x) ∈ (p : Set _) := Φrange ▸ (mem_range_self _) diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean index f0124bb1d3151b..5eedc8d457f845 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean @@ -480,7 +480,7 @@ the Hausdorff distance in the optimal coupling, although we only prove here the we need. -/ theorem hausdorffDist_optimal_le_HD {f} (h : f ∈ candidatesB X Y) : hausdorffDist (range (optimalGHInjl X Y)) (range (optimalGHInjr X Y)) ≤ HD f := by - refine le_trans (le_of_forall_le_of_dense fun r hr => ?_) (HD_optimalGHDist_le X Y f h) + refine le_trans (le_of_forall_gt_imp_ge_of_dense fun r hr => ?_) (HD_optimalGHDist_le X Y f h) have A : ∀ x ∈ range (optimalGHInjl X Y), ∃ y ∈ range (optimalGHInjr X Y), dist x y ≤ r := by rintro _ ⟨z, rfl⟩ have I1 : (⨆ x, ⨅ y, optimalGHDist X Y (inl x, inr y)) < r := diff --git a/Mathlib/Topology/MetricSpace/HausdorffDimension.lean b/Mathlib/Topology/MetricSpace/HausdorffDimension.lean index ec335783eba45e..0235fc0d813222 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDimension.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDimension.lean @@ -232,7 +232,7 @@ theorem bsupr_limsup_dimH (s : Set X) : ⨆ x ∈ s, limsup dimH (𝓝[s] x).sma refine le_antisymm (iSup₂_le fun x _ => ?_) ?_ · refine limsup_le_of_le isCobounded_le_of_bot ?_ exact eventually_smallSets.2 ⟨s, self_mem_nhdsWithin, fun t => dimH_mono⟩ - · refine le_of_forall_ge_of_dense fun r hr => ?_ + · refine le_of_forall_lt_imp_le_of_dense fun r hr => ?_ rcases exists_mem_nhdsWithin_lt_dimH_of_lt_dimH hr with ⟨x, hxs, hxr⟩ refine le_iSup₂_of_le x hxs ?_; rw [limsup_eq]; refine le_sInf fun b hb => ?_ rcases eventually_smallSets.1 hb with ⟨t, htx, ht⟩ diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean b/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean index 6b6cd86da53a26..9f284fb00b8813 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean @@ -91,7 +91,7 @@ lemma exists_isCompact_closedBall [WeaklyLocallyCompactSpace α] (x : α) : theorem biInter_gt_closedBall (x : α) (r : ℝ) : ⋂ r' > r, closedBall x r' = closedBall x r := by ext - simp [forall_gt_ge_iff] + simp [forall_gt_imp_ge_iff_le_of_dense] theorem biInter_gt_ball (x : α) (r : ℝ) : ⋂ r' > r, ball x r' = closedBall x r := by ext @@ -100,7 +100,7 @@ theorem biInter_gt_ball (x : α) (r : ℝ) : ⋂ r' > r, ball x r' = closedBall theorem biUnion_lt_ball (x : α) (r : ℝ) : ⋃ r' < r, ball x r' = ball x r := by ext rw [← not_iff_not] - simp [forall_lt_le_iff] + simp [forall_lt_imp_le_iff_le_of_dense] theorem biUnion_lt_closedBall (x : α) (r : ℝ) : ⋃ r' < r, closedBall x r' = ball x r := by ext diff --git a/Mathlib/Topology/NoetherianSpace.lean b/Mathlib/Topology/NoetherianSpace.lean index 931e8748cbfcef..e26ec3b5c4b6d7 100644 --- a/Mathlib/Topology/NoetherianSpace.lean +++ b/Mathlib/Topology/NoetherianSpace.lean @@ -70,7 +70,7 @@ protected theorem _root_.Topology.IsInducing.noetherianSpace [NoetherianSpace α @[deprecated (since := "2024-10-28")] alias _root_.Inducing.noetherianSpace := IsInducing.noetherianSpace -/-- [Stacks: Lemma 0052 (1)](https://stacks.math.columbia.edu/tag/0052)-/ +@[stacks 0052 "(1)"] instance NoetherianSpace.set [NoetherianSpace α] (s : Set α) : NoetherianSpace s := IsInducing.subtypeVal.noetherianSpace @@ -191,7 +191,7 @@ theorem NoetherianSpace.exists_finset_irreducible [NoetherianSpace α] (s : Clos simpa [Set.exists_finite_iff_finset, Finset.sup_id_eq_sSup] using NoetherianSpace.exists_finite_set_closeds_irreducible s -/-- [Stacks: Lemma 0052 (2)](https://stacks.math.columbia.edu/tag/0052) -/ +@[stacks 0052 "(2)"] theorem NoetherianSpace.finite_irreducibleComponents [NoetherianSpace α] : (irreducibleComponents α).Finite := by obtain ⟨S : Set (Set α), hSf, hSc, hSi, hSU⟩ := @@ -201,7 +201,7 @@ theorem NoetherianSpace.finite_irreducibleComponents [NoetherianSpace α] : rcases isIrreducible_iff_sUnion_isClosed.1 hs.1 S hSc (hSU ▸ Set.subset_univ _) with ⟨t, htS, ht⟩ rwa [ht.antisymm (hs.2 (hSi _ htS) ht)] -/-- [Stacks: Lemma 0052 (3)](https://stacks.math.columbia.edu/tag/0052) -/ +@[stacks 0052 "(3)"] theorem NoetherianSpace.exists_open_ne_empty_le_irreducibleComponent [NoetherianSpace α] (Z : Set α) (H : Z ∈ irreducibleComponents α) : ∃ o : Set α, IsOpen o ∧ o ≠ ∅ ∧ o ≤ Z := by diff --git a/Mathlib/Topology/PartialHomeomorph.lean b/Mathlib/Topology/PartialHomeomorph.lean index a178c181fd1cdd..c0004b4df3cbab 100644 --- a/Mathlib/Topology/PartialHomeomorph.lean +++ b/Mathlib/Topology/PartialHomeomorph.lean @@ -1326,15 +1326,14 @@ theorem subtypeRestr_symm_eqOn_of_le {U V : Opens X} (hU : Nonempty U) (hV : Non end subtypeRestr variable {X X' Z : Type*} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Z] - [Nonempty X] [Nonempty Z] {f : X → X'} + [Nonempty Z] {f : X → X'} /-- Extend a partial homeomorphism `e : X → Z` to `X' → Z`, using an open embedding `ι : X → X'`. On `ι(X)`, the extension is specified by `e`; its value elsewhere is arbitrary (and uninteresting). -/ -noncomputable def lift_openEmbedding - (e : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) : +noncomputable def lift_openEmbedding (e : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) : PartialHomeomorph X' Z where - toFun := extend f e (Classical.arbitrary (α := X' → Z)) + toFun := extend f e (fun _ ↦ (Classical.arbitrary Z)) invFun := f ∘ e.invFun source := f '' e.source target := e.target @@ -1352,7 +1351,9 @@ noncomputable def lift_openEmbedding open_source := hf.isOpenMap _ e.open_source open_target := e.open_target continuousOn_toFun := by - set F := (extend f (↑e) (Classical.arbitrary (X' → Z))) with F_eq + by_cases Nonempty X; swap + · intro x hx; simp_all + set F := (extend f e (fun _ ↦ (Classical.arbitrary Z))) with F_eq have heq : EqOn F (e ∘ (hf.toPartialHomeomorph).symm) (f '' e.source) := by intro x ⟨x₀, hx₀, hxx₀⟩ rw [← hxx₀, F_eq, hf.injective.extend_apply e, comp_apply, hf.toPartialHomeomorph_left_inv] @@ -1360,15 +1361,15 @@ noncomputable def lift_openEmbedding apply e.continuousOn_toFun.comp; swap · intro x' ⟨x, hx, hx'x⟩ rw [← hx'x, hf.toPartialHomeomorph_left_inv]; exact hx - have : ContinuousOn (hf.toPartialHomeomorph).symm (f '' univ) := by - apply (hf.toPartialHomeomorph).continuousOn_invFun + have : ContinuousOn (hf.toPartialHomeomorph).symm (f '' univ) := + (hf.toPartialHomeomorph).continuousOn_invFun apply this.mono; gcongr; exact fun ⦃a⦄ a ↦ trivial - apply ContinuousOn.congr this heq + exact ContinuousOn.congr this heq continuousOn_invFun := hf.continuous.comp_continuousOn e.continuousOn_invFun @[simp] lemma lift_openEmbedding_toFun (e : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) : - (e.lift_openEmbedding hf) = extend f e (Classical.arbitrary (α := (X' → Z))) := rfl + (e.lift_openEmbedding hf) = extend f e (fun _ ↦ (Classical.arbitrary Z)) := rfl lemma lift_openEmbedding_apply (e : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) {x : X} : (lift_openEmbedding e hf) (f x) = e x := by @@ -1404,8 +1405,8 @@ lemma lift_openEmbedding_trans_apply @[simp] lemma lift_openEmbedding_trans (e e' : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) : (e.lift_openEmbedding hf).symm.trans (e'.lift_openEmbedding hf) = e.symm.trans e' := by - ext x - · exact e.lift_openEmbedding_trans_apply e' hf x + ext z + · exact e.lift_openEmbedding_trans_apply e' hf z · simp [hf.injective.extend_apply e] · simp_rw [PartialHomeomorph.trans_source, e.lift_openEmbedding_symm_source, e.symm_source, e.lift_openEmbedding_symm, e'.lift_openEmbedding_source] diff --git a/Mathlib/Topology/SeparatedMap.lean b/Mathlib/Topology/SeparatedMap.lean index 61b7da07704282..5773590ffc3172 100644 --- a/Mathlib/Topology/SeparatedMap.lean +++ b/Mathlib/Topology/SeparatedMap.lean @@ -183,25 +183,11 @@ section eqLocus variable {f : X → Y} {g₁ g₂ : A → X} (h₁ : Continuous g₁) (h₂ : Continuous g₂) include h₁ h₂ -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/5338, -the unused variable linter flags `g` here, -but it is used in a type ascription to direct `fun_prop`. --/ -set_option linter.unusedVariables false in theorem IsSeparatedMap.isClosed_eqLocus (sep : IsSeparatedMap f) (he : f ∘ g₁ = f ∘ g₂) : IsClosed {a | g₁ a = g₂ a} := let g : A → f.Pullback f := fun a ↦ ⟨⟨g₁ a, g₂ a⟩, congr_fun he a⟩ (isSeparatedMap_iff_isClosed_diagonal.mp sep).preimage (by fun_prop : Continuous g) -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/5338, -the unused variable linter flags `g` here, -but it is used in a type ascription to direct `fun_prop`. --/ -set_option linter.unusedVariables false in theorem IsLocallyInjective.isOpen_eqLocus (inj : IsLocallyInjective f) (he : f ∘ g₁ = f ∘ g₂) : IsOpen {a | g₁ a = g₂ a} := let g : A → f.Pullback f := fun a ↦ ⟨⟨g₁ a, g₂ a⟩, congr_fun he a⟩ diff --git a/Mathlib/Topology/Separation/Regular.lean b/Mathlib/Topology/Separation/Regular.lean index 3be8602edb54b5..ae49d2084217e1 100644 --- a/Mathlib/Topology/Separation/Regular.lean +++ b/Mathlib/Topology/Separation/Regular.lean @@ -3,8 +3,9 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Topology.Separation.Hausdorff +import Mathlib.Tactic.StacksAttribute import Mathlib.Topology.Compactness.Lindelof +import Mathlib.Topology.Separation.Hausdorff /-! # Regular, normal, T₃, T₄ and T₅ spaces @@ -637,8 +638,8 @@ theorem connectedComponent_eq_iInter_isClopen [T2Space X] [CompactSpace X] (x : ⟨s ∩ v, H2, mem_inter H.2.1 h1⟩ /-- `ConnectedComponents X` is Hausdorff when `X` is Hausdorff and compact -/ +@[stacks 0900 "The Stacks entry proves profiniteness."] instance ConnectedComponents.t2 [T2Space X] [CompactSpace X] : T2Space (ConnectedComponents X) := by - -- Proof follows that of: https://stacks.math.columbia.edu/tag/0900 -- Fix 2 distinct connected components, with points a and b refine ⟨ConnectedComponents.surjective_coe.forall₂.2 fun a b ne => ?_⟩ rw [ConnectedComponents.coe_ne_coe] at ne diff --git a/Mathlib/Topology/UniformSpace/Dini.lean b/Mathlib/Topology/UniformSpace/Dini.lean index aff1384ce1d4fd..97b9a70caecb89 100644 --- a/Mathlib/Topology/UniformSpace/Dini.lean +++ b/Mathlib/Topology/UniformSpace/Dini.lean @@ -42,7 +42,7 @@ variable {F : ι → α → G} {f : α → G} namespace Monotone -/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions +/-- **Dini's theorem**: if `F n` is a monotone increasing collection of continuous functions converging pointwise to a continuous function `f`, then `F n` converges locally uniformly to `f`. -/ lemma tendstoLocallyUniformly_of_forall_tendsto (hF_cont : ∀ i, Continuous (F i)) (hF_mono : Monotone F) (hf : Continuous f) @@ -64,7 +64,7 @@ lemma tendstoLocallyUniformly_of_forall_tendsto simp only [abs_of_nonpos (sub_nonpos_of_le (F_le_f _ _)), neg_sub, sub_le_sub_iff_left] exact hF_mono hnm z -/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions on a +/-- **Dini's theorem**: if `F n` is a monotone increasing collection of continuous functions on a set `s` converging pointwise to a continuous function `f`, then `F n` converges locally uniformly to `f`. -/ lemma tendstoLocallyUniformlyOn_of_forall_tendsto {s : Set α} @@ -75,7 +75,7 @@ lemma tendstoLocallyUniformlyOn_of_forall_tendsto {s : Set α} exact tendstoLocallyUniformly_of_forall_tendsto (hF_cont · |>.restrict) (fun _ _ h x ↦ hF_mono _ x.2 h) hf.restrict (fun x ↦ h_tendsto x x.2) -/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions on a +/-- **Dini's theorem**: if `F n` is a monotone increasing collection of continuous functions on a compact space converging pointwise to a continuous function `f`, then `F n` converges uniformly to `f`. -/ lemma tendstoUniformly_of_forall_tendsto [CompactSpace α] (hF_cont : ∀ i, Continuous (F i)) @@ -84,7 +84,7 @@ lemma tendstoUniformly_of_forall_tendsto [CompactSpace α] (hF_cont : ∀ i, Con tendstoLocallyUniformly_iff_tendstoUniformly_of_compactSpace.mp <| tendstoLocallyUniformly_of_forall_tendsto hF_cont hF_mono hf h_tendsto -/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions on a +/-- **Dini's theorem**: if `F n` is a monotone increasing collection of continuous functions on a compact set `s` converging pointwise to a continuous function `f`, then `F n` converges uniformly to `f`. -/ lemma tendstoUniformlyOn_of_forall_tendsto {s : Set α} (hs : IsCompact s) @@ -98,7 +98,7 @@ end Monotone namespace Antitone -/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +/-- **Dini's theorem**: if `F n` is a monotone decreasing collection of continuous functions on a converging pointwise to a continuous function `f`, then `F n` converges locally uniformly to `f`. -/ lemma tendstoLocallyUniformly_of_forall_tendsto (hF_cont : ∀ i, Continuous (F i)) (hF_anti : Antitone F) (hf : Continuous f) @@ -106,7 +106,7 @@ lemma tendstoLocallyUniformly_of_forall_tendsto TendstoLocallyUniformly F f atTop := Monotone.tendstoLocallyUniformly_of_forall_tendsto (G := Gᵒᵈ) hF_cont hF_anti hf h_tendsto -/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +/-- **Dini's theorem**: if `F n` is a monotone decreasing collection of continuous functions on a set `s` converging pointwise to a continuous function `f`, then `F n` converges locally uniformly to `f`. -/ lemma tendstoLocallyUniformlyOn_of_forall_tendsto {s : Set α} @@ -115,7 +115,7 @@ lemma tendstoLocallyUniformlyOn_of_forall_tendsto {s : Set α} TendstoLocallyUniformlyOn F f atTop s := Monotone.tendstoLocallyUniformlyOn_of_forall_tendsto (G := Gᵒᵈ) hF_cont hF_anti hf h_tendsto -/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +/-- **Dini's theorem**: if `F n` is a monotone decreasing collection of continuous functions on a compact space converging pointwise to a continuous function `f`, then `F n` converges uniformly to `f`. -/ lemma tendstoUniformly_of_forall_tendsto [CompactSpace α] (hF_cont : ∀ i, Continuous (F i)) @@ -123,7 +123,7 @@ lemma tendstoUniformly_of_forall_tendsto [CompactSpace α] (hF_cont : ∀ i, Con TendstoUniformly F f atTop := Monotone.tendstoUniformly_of_forall_tendsto (G := Gᵒᵈ) hF_cont hF_anti hf h_tendsto -/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +/-- **Dini's theorem**: if `F n` is a monotone decreasing collection of continuous functions on a compact set `s` converging pointwise to a continuous `f`, then `F n` converges uniformly to `f`. -/ lemma tendstoUniformlyOn_of_forall_tendsto {s : Set α} (hs : IsCompact s) (hF_cont : ∀ i, ContinuousOn (F i) s) (hF_anti : ∀ x ∈ s, Antitone (F · x)) @@ -139,7 +139,7 @@ namespace ContinuousMap variable {F : ι → C(α, G)} {f : C(α, G)} -/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions +/-- **Dini's theorem**: if `F n` is a monotone increasing collection of continuous functions converging pointwise to a continuous function `f`, then `F n` converges to `f` in the compact-open topology. -/ lemma tendsto_of_monotone_of_pointwise (hF_mono : Monotone F) @@ -148,7 +148,7 @@ lemma tendsto_of_monotone_of_pointwise (hF_mono : Monotone F) tendsto_of_tendstoLocallyUniformly <| hF_mono.tendstoLocallyUniformly_of_forall_tendsto (F · |>.continuous) f.continuous h_tendsto -/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions +/-- **Dini's theorem**: if `F n` is a monotone decreasing collection of continuous functions converging pointwise to a continuous function `f`, then `F n` converges to `f` in the compact-open topology. -/ lemma tendsto_of_antitone_of_pointwise (hF_anti : Antitone F) diff --git a/MathlibTest/Header.lean b/MathlibTest/Header.lean index dcfe5788d4f319..6553027d5e3fc4 100644 --- a/MathlibTest/Header.lean +++ b/MathlibTest/Header.lean @@ -9,7 +9,7 @@ import Lake import Mathlib.Tactic.Linter.Header import /- -/ Mathlib.Tactic -- the `TextBased` linter does not flag this `broadImport` import Mathlib.Tactic.Have -import Mathlib.Deprecated.Subfield +import Mathlib.Deprecated.MinMax /-- warning: In the past, importing 'Lake' in mathlib has led to dramatic slow-downs of the linter (see e.g. https://github.com/leanprover-community/mathlib4/pull/13779). Please consider carefully if this import is useful and make sure to benchmark it. If this is fine, feel free to silence this linter. diff --git a/MathlibTest/MoveAdd.lean b/MathlibTest/MoveAdd.lean index 74d7a9fd44d709..e091e08e058a7a 100644 --- a/MathlibTest/MoveAdd.lean +++ b/MathlibTest/MoveAdd.lean @@ -108,13 +108,12 @@ example {a b c d e : Prop} (h : a ∨ b ∨ c ∨ d ∨ e) : a ∨ c ∨ e ∨ b end left_assoc -#adaptation_note /-- nightly-2024-03-11 -This test is now failing with `unknown free variable '_fvar.36787'` -/ --- example (k : ℕ) (h0 : 0 + 2 = 9 + 0) (h9 : k + 2 = k + 9) : k + 2 = 9 + k := by --- induction' k with k _ih --- · exact h0 --- · move_add [9] --- exact h9 +example (k : ℕ) (h0 : 0 + 2 = 9 + 0) (h9 : k + 2 = k + 9) : k + 2 = 9 + k := by + induction k with + | zero => exact h0 + | succ n _ => + move_add [9] + exact h9 -- Testing internals of the tactic `move_add`. section tactic diff --git a/MathlibTest/Perm.lean b/MathlibTest/Perm.lean new file mode 100644 index 00000000000000..72aaf2dee0de92 --- /dev/null +++ b/MathlibTest/Perm.lean @@ -0,0 +1,14 @@ + +import Mathlib.GroupTheory.Perm.Cycle.Concrete +import Mathlib.Tactic.DeriveFintype + +open Equiv + +#guard with_decl_name% ex₁ unsafe (reprStr (1 : Perm (Fin 4))) == "1" +#guard with_decl_name% ex₂ unsafe (reprStr (c[0, 1] : Perm (Fin 4))) == "c[0, 1]" +#guard with_decl_name% ex₃ + unsafe (reprStr (c[0, 1] * c[2, 3] : Perm (Fin 4))) == "c[0, 1] * c[2, 3]" +#guard with_decl_name% ex₄ + unsafe (reprPrec (c[0, 1] * c[2, 3] : Perm (Fin 4)) 70).pretty == "(c[0, 1] * c[2, 3])" +#guard with_decl_name% ex₅ + unsafe (reprPrec (c[0, 1] * c[1, 2] : Perm (Fin 4)) 70).pretty == "c[0, 1, 2]" diff --git a/MathlibTest/Recall.lean b/MathlibTest/Recall.lean index a665edc1a8f344..01092a624dc923 100644 --- a/MathlibTest/Recall.lean +++ b/MathlibTest/Recall.lean @@ -1,7 +1,7 @@ import Mathlib.Tactic.Recall import Mathlib.Analysis.Calculus.Deriv.Basic import Mathlib.Analysis.SpecialFunctions.Trigonometric.Basic -import Mathlib.Data.Complex.Exponential +import Mathlib.Data.Complex.Trigonometric set_option linter.style.setOption false -- Remark: When the test is run by make/CI, this option is not set, so we set it here. diff --git a/MathlibTest/cc.lean b/MathlibTest/cc.lean index 71962960104261..1263e28bba6fa8 100644 --- a/MathlibTest/cc.lean +++ b/MathlibTest/cc.lean @@ -250,7 +250,7 @@ axiom C : (a : A) → B a → Type axiom D : (a : A) → (ba : B a) → C a ba → Type axiom E : (a : A) → (ba : B a) → (cba : C a ba) → D a ba cba → Type axiom F : (a : A) → (ba : B a) → (cba : C a ba) → (dcba : D a ba cba) → E a ba cba dcba → Type -axiom C_ss : ∀ a ba, Subsingleton (C a ba) +axiom C_ss : ∀ a ba, Lean.Meta.FastSubsingleton (C a ba) axiom a1 : A axiom a2 : A axiom a3 : A diff --git a/MathlibTest/positivity.lean b/MathlibTest/positivity.lean index 80b80dab4bb9ed..a9b0c4816cbd62 100644 --- a/MathlibTest/positivity.lean +++ b/MathlibTest/positivity.lean @@ -1,5 +1,5 @@ import Mathlib.Tactic.Positivity -import Mathlib.Data.Complex.Exponential +import Mathlib.Data.Complex.Trigonometric import Mathlib.Data.Real.Sqrt import Mathlib.Analysis.Normed.Group.Basic import Mathlib.Analysis.SpecialFunctions.Pow.Real diff --git a/docs/1000.yaml b/docs/1000.yaml index fc2fe3e686f5ea..ca2e502ee182b5 100644 --- a/docs/1000.yaml +++ b/docs/1000.yaml @@ -9,10 +9,7 @@ # and by `make_site.py` in the leanprover-community.github.io repo. # # Current TODOs/unresolved questions: -# - add infrastructure for updating the information upstream when formalisations are added # - perhaps add a wikidata version of the stacks attribute, and generate this file automatically -# -# TODO: display a less ambiguous title for e.g. "Liouville's theorem" Q11518: title: Pythagorean theorem @@ -272,7 +269,7 @@ Q257387: comment: "mathlib4 pull request at https://github.com/leanprover-community/mathlib4/pull/20722" Q258374: - title: Carathéodory's theorem + title: Carathéodory's theorem (measure theory) decl: MeasureTheory.OuterMeasure.caratheodory comment: Hard to say what exactly is meant. @@ -280,7 +277,7 @@ Q260928: title: Jordan curve theorem Q266291: - title: Bloch's theorem + title: Bloch's theorem (complex variables) Q268031: title: Abel's binomial theorem @@ -383,6 +380,7 @@ Q384142: Q386292: title: Prime number theorem + comment: "Ongoing formalisation of various versions in https://github.com/AlexKontorovich/PrimeNumberTheoremAnd" # also in 100.yaml Q388525: @@ -464,8 +462,8 @@ Q504843: title: Mycielski's theorem Q505798: - title: Lagrange's theorem - # group theorem result; exist additive and multiplicative version + title: Lagrange's theorem (group theory) + # exist additive and multiplicative version decl: Subgroup.card_subgroup_dvd_card Q510197: @@ -526,7 +524,7 @@ Q574902: title: Tarski's indefinability theorem Q576478: - title: Liouville's theorem + title: Liouville's theorem (complex analysis) decl: Differentiable.apply_eq_apply_of_bounded Q578555: @@ -554,6 +552,9 @@ Q594571: Q595466: title: Dini's theorem + decl: Monotone.tendstoLocallyUniformly_of_forall_tendsto + authors: Jireh Loreaux + date: 2024 Q599876: title: Cochran's theorem @@ -620,7 +621,7 @@ Q649977: title: Shirshov–Cohn theorem Q650738: - title: Folk theorem + title: Folk theorem (game theory) Q651593: title: Norton's theorem @@ -634,7 +635,9 @@ Q656198: Q656645: title: Hilbert's basis theorem - # search mathlib or zulip; mathlib is close to this! + decl: Polynomial.isNoetherianRing + authors: Kenny Lau + date: 2019 Q656772: title: Cayley–Hamilton theorem @@ -656,7 +659,7 @@ Q657903: title: Hilbert–Waring theorem Q660799: - title: Darboux's theorem + title: Darboux's theorem (analysis) # "all derivatives have the intermediate value property" # TODO should exist in mathlib, add decl! @@ -818,6 +821,9 @@ Q794269: Q810431: title: Basel problem + decl: hasSum_zeta_two + authors: David Loeffler + date: 2022 Q828284: title: Parallel axis theorem @@ -844,7 +850,7 @@ Q837506: comment: "will enter mathlib in 2024" Q837551: - title: Frobenius theorem + title: Frobenius theorem (differential topology) Q841893: title: Desargues's theorem @@ -911,7 +917,7 @@ Q859122: title: Cauchy–Hadamard theorem Q865665: - title: Birkhoff's theorem + title: Birkhoff's theorem (relativity) Q866116: title: Hahn–Banach theorem @@ -936,7 +942,7 @@ Q890875: title: Bohr–van Leeuwen theorem Q897769: - title: König's theorem + title: Kőnig's theorem (graph theory) Q899002: title: Pascal's theorem @@ -1024,12 +1030,13 @@ Q943246: decl: littleWedderburn Q944297: - title: Open mapping theorem - # TODO: this entry points to Wikipedia's disambiguation page => wikipedia should be fixed - # Banach open mapping theorem/Banach-Schauder theorem in functional analysis is in mathlib (ContinuousLinearMap.isOpenMap) + title: Open mapping theorem (functional analysis) + decl: ContinuousLinearMap.isOpenMap + authors: Sébastien Gouëzel + date: 2019 Q948664: - title: Kneser's addition theorem + title: Kneser's addition theorem (combinatorics) authors: Mantas Bakšys, Yaël Dillies url: https://github.com/YaelDillies/LeanCamCombi/blob/master/LeanCamCombi/Kneser/Kneser.lean date: 2022 @@ -1050,11 +1057,11 @@ Q966837: title: Cramér–Wold theorem Q967972: - title: Open mapping theorem + title: Open mapping theorem (complex analysis) decl: AnalyticOnNhd.is_constant_or_isOpen Q973359: - title: Rado's theorem + title: Radó's theorem (harmonic functions) Q974405: title: Hille–Yosida theorem @@ -1161,7 +1168,7 @@ Q1057919: Q1058662: title: Darboux's theorem - # symplectic geometry result + # symplectic geometry Q1059151: title: FWL theorem @@ -1213,7 +1220,7 @@ Q1076274: title: Choquet–Bishop–de Leeuw theorem Q1077462: - title: König's theorem + title: König's theorem (set theory) Q1077741: title: Hairy ball theorem @@ -1268,7 +1275,7 @@ Q1135706: title: Harnack's theorem Q1136043: - title: Hurwitz's theorem + title: Hurwitz's theorem (complex analysis) Q1136262: title: Optical theorem @@ -1285,12 +1292,11 @@ Q1137206: authors: Moritz Doll Q1139041: - title: Cauchy's theorem - # for finite groups + title: Cauchy's theorem (group theory) decl: exists_prime_orderOf_dvd_card Q1139430: - title: Hurwitz's theorem + title: Hurwitz's theorem (normed division algebras) Q1139524: title: Rationality theorem @@ -1302,7 +1308,7 @@ Q1140200: title: PCP theorem Q1141747: - title: Carnot's theorem + title: Carnot's theorem (inradius, circumradius) Q1143540: title: Heckscher–Ohlin theorem @@ -1421,7 +1427,7 @@ Q1217677: authors: Yury G. Kudryashov (first) and Benjamin Davidson (second) Q1227061: - title: Khinchin's theorem + title: Khinchin's theorem on Diophantine approximations Q1227702: title: Dirichlet's unit theorem @@ -1475,7 +1481,7 @@ Q1317367: title: Japanese theorem for concyclic polygons Q1322892: - title: Dirac's theorems + title: Dirac's theorem on chordal graphs Q1330788: title: Weierstrass factorization theorem @@ -1500,7 +1506,7 @@ Q1357684: title: Riesz representation theorem Q1361031: - title: Frobenius theorem + title: Frobenius theorem (real division algebras) Q1361393: title: Luzin's theorem @@ -1564,6 +1570,9 @@ Q1455794: Q1457052: title: Well-ordering theorem + decl: exists_wellOrder + authors: Mario Carneiro, Chris Hughes, Violeta Hernández + date: 2017 Q1471282: title: Radon's theorem @@ -1646,7 +1655,7 @@ Q1621180: title: Orlicz–Pettis theorem Q1630588: - title: Hurwitz's theorem + title: Hurwitz's theorem (number theory) Q1631749: title: Lester's theorem @@ -1729,6 +1738,12 @@ Q1765521: Q1766814: title: Smn theorem + decls: + - Nat.Partrec.Code.smn + - Nat.Partrec'.part_iff + authors: Mario Carneiro + date: 2018 + comment: "This theorem is trivial when using Mathlib's computability definition, but this definition is equivalent to the usual one, as shown in `Nat.Partrec'.part_iff`." Q1785610: title: Poncelet's closure theorem @@ -1781,6 +1796,9 @@ Q1930577: Q1933521: title: Kleene's recursion theorem + decl: Nat.Partrec.Code.fixed_point₂ + authors: Mario Carneiro + date: 2018 Q1938251: title: Hasse–Minkowski theorem @@ -1811,7 +1829,7 @@ Q2000090: title: Art gallery theorem Q2008549: - title: Hilbert's theorem + title: Hilbert's theorem (differential geometry) Q2013213: title: Reuschle's theorem @@ -1991,7 +2009,7 @@ Q2345282: title: Shannon's theorem Q2347139: - title: Carnot's theorem + title: Carnot's theorem (thermodynamics) Q2360531: title: Jordan–Schur theorem @@ -2028,7 +2046,7 @@ Q2449984X: title: Pivot theorem Q2471737: - title: Carathéodory's theorem + title: Carathéodory's theorem (convex hull) decl: convexHull_eq_union Q2473965: @@ -2054,7 +2072,7 @@ Q2518048: title: Kodaira vanishing theorem Q2524990: - title: Jackson's theorem + title: Jackson's theorem (queueing theory) Q2525646: title: Jordan–Hölder theorem @@ -2109,7 +2127,7 @@ Q2800071: title: Perfect graph theorem Q2862403: - title: Bombieri's theorem + title: Schneider–Lang theorem Q2893695: title: Unmixedness theorem @@ -2175,7 +2193,11 @@ Q3345678: title: Moore–Aronszajn theorem Q3424029: - title: Chasles's theorems + # TODO: refer to the individual entries instead + # The list in wikipedia has already been changed. This will will change automatically + # when the 1000+ theorem project synchronises its data with wikipedia and this file is + # synchronised with the 1000+ theorem's data. + title: Chasles's theorem (disambiguation) Q3443426: title: Bondy's theorem @@ -2223,10 +2245,10 @@ Q3527009: title: Baker's theorem Q3527011: - title: Beck's theorem + title: Beck's theorem (geometry) Q3527015: - title: Bernstein's theorem + title: Bernstein's theorem on monotone functions Q3527017: title: Beurling–Lax theorem @@ -2235,7 +2257,7 @@ Q3527019: title: Birch's theorem Q3527034: - title: Cauchy's theorem + title: Cauchy's theorem (geometry) Q3527040: title: Chomsky–Schützenberger representation theorem @@ -2473,7 +2495,7 @@ Q4378868: date: 2022 Q4378889: - title: Carathéodory's theorem + title: Carathéodory's theorem (conformal mapping) Q4408070: title: De Finetti's theorem @@ -2497,8 +2519,7 @@ Q4454973: title: Lindelöf's theorem Q4454976: - title: Liouville's theorem - # theorem in conformal mappings + title: Liouville's theorem (conformal mappings) Q4454989: title: Meusnier's theorem @@ -2585,7 +2606,7 @@ Q4734844: title: Alperin–Brauer–Gorenstein theorem Q4736422: - title: Denjoy theorem + title: Denjoy–Luzin theorem Q4746948: title: Amitsur–Levitzki theorem @@ -2664,6 +2685,9 @@ Q4877965: Q4878586: title: Beck's monadicity theorem + decl: CategoryTheory.Monad.monadicOfCreatesGSplitCoequalizers + authors: Bhavik Mehta + date: 2020 Q4880958: title: Behnke–Stein theorem @@ -2678,7 +2702,7 @@ Q4891633: title: Berger–Kazdan comparison theorem Q4894565: - title: Bernstein's theorem + title: Bernstein's theorem (approximation theory) Q4914101: title: Bing's recognition theorem @@ -2986,7 +3010,7 @@ Q5494134: title: Fraňková–Helly selection theorem Q5498822: - title: Birkhoff's theorem + title: Birkhoff's ergodic theorem Q5499906: title: Shirshov–Witt theorem @@ -3057,10 +3081,10 @@ Q5609384: title: Grinberg's theorem Q5610188: - title: Gromov's compactness theorem + title: Gromov's compactness theorem (geometry) Q5610190: - title: Gromov's compactness theorem + title: Gromov's compactness theorem (topology) Q5610718: title: Grothendieck's connectedness theorem @@ -3174,8 +3198,7 @@ Q6402928: title: Lagrange reversion theorem Q6403282: - title: Lagrange's theorem - # number theory + title: Lagrange's theorem (number theory) Q6407842: title: Killing–Hopf theorem @@ -3653,8 +3676,7 @@ Q7818977: title: Tomita's theorem Q7820874: - title: Tonelli's theorem - decl: MeasureTheory.lintegral_prod + title: Tonelli's theorem (functional analysis) Q7824894: title: Topkis's theorem @@ -3718,7 +3740,7 @@ Q7966545: title: Walter theorem Q7978735: - title: Weber's theorem + title: Weber's theorem (Algebraic curves) Q7980241: title: Weinberg–Witten theorem @@ -3905,7 +3927,7 @@ Q18205730: title: Byers–Yang theorem Q18206032: - title: Cartan's theorem + title: Closed subgroup theorem Q18206266: title: Euclid–Euler theorem @@ -3946,7 +3968,7 @@ Q25304301: title: Bregman–Minc inequality Q25345219: - title: BBD decomposition theorem + title: Decomposition theorem of Beilinson, Bernstein and Deligne Q25378690: title: Ionescu-Tulcea theorem @@ -3958,7 +3980,7 @@ Q28194853: title: Lovelock's theorem Q28458131: - title: Glivenko's theorem + title: Glivenko's theorem (probability theory) Q31838822: title: Kirchberger's theorem diff --git a/lake-manifest.json b/lake-manifest.json index bc1e91750309a4..f3bfaf39ccd3aa 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "62398edfce9977467f273b8d104f1591aae16843", + "rev": "b198ae51d27902453b27d01fc0ac07eeadcbe2e6", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "fin-find", diff --git a/scripts/lint-style.lean b/scripts/lint-style.lean index b4b7d93b9f2839..7f1b2cdd34467c 100644 --- a/scripts/lint-style.lean +++ b/scripts/lint-style.lean @@ -29,6 +29,12 @@ def findImports (path : System.FilePath) : IO (Array Lean.Name) := do return (← Lean.parseImports' (← IO.FS.readFile path) path.toString) |>.map (fun imp ↦ imp.module) |>.erase `Init +/-- Additional imports generated by `mk_all`. -/ +def explicitImports : Array Lean.Name := #[`Batteries, `Std] + +/-- Remove the additional imports generated by `mk_all` so that only mathlib modules remain. -/ +def eraseExplicitImports (names : Array Lean.Name) : Array Lean.Name := + explicitImports.foldl Array.erase names /-- Check that `Mathlib.Init` is transitively imported in all of Mathlib. @@ -39,7 +45,7 @@ Return `true` iff there was an error. def checkInitImports : IO Bool := do -- Find any file in the Mathlib directory which does not contain any Mathlib import. -- We simply parse `Mathlib.lean`, as CI ensures this file is up to date. - let allModuleNames := (← findImports "Mathlib.lean").erase `Batteries + let allModuleNames := eraseExplicitImports (← findImports "Mathlib.lean") let mut modulesWithoutMathlibImports := #[] let mut importsHeaderLinter := #[] for module in allModuleNames do @@ -107,8 +113,9 @@ def lintStyleCli (args : Cli.Parsed) : IO UInt32 := do let mut allModuleNames := #[] for s in ["Archive.lean", "Counterexamples.lean", "Mathlib.lean"] do allModuleNames := allModuleNames.append (← findImports s) - -- Note: since "Batteries" is added explicitly to "Mathlib.lean", we remove it here manually. - allModuleNames := allModuleNames.erase `Batteries + -- Note: since "Batteries" and "Std" are added explicitly to "Mathlib.lean", we remove them here + -- manually. + allModuleNames := eraseExplicitImports allModuleNames -- Read the `nolints` file, with manual exceptions for the linter. -- NB. We pass these lints to `lintModules` explicitly to prevent cache invalidation bugs: diff --git a/scripts/mk_all.lean b/scripts/mk_all.lean index 00bacbc8eaefb8..ae79948610542f 100644 --- a/scripts/mk_all.lean +++ b/scripts/mk_all.lean @@ -55,9 +55,9 @@ def mkAllCLI (args : Parsed) : IO UInt32 := do for d in libs.reverse do -- reverse to create `Mathlib/Tactic.lean` before `Mathlib.lean` let fileName := addExtension d "lean" let mut allFiles ← getAllModulesSorted git d - -- mathlib exception: manually import Batteries in `Mathlib.lean` + -- mathlib exception: manually import Std and Batteries in `Mathlib.lean` if d == "Mathlib" then - allFiles := #["Batteries"] ++ allFiles + allFiles := #["Std", "Batteries"] ++ allFiles let fileContent := ("\n".intercalate (allFiles.map ("import " ++ ·)).toList).push '\n' if !(← pathExists fileName) then if check then diff --git a/scripts/nolints_prime_decls.txt b/scripts/nolints_prime_decls.txt index be856e67cdf248..fd259826572ac5 100644 --- a/scripts/nolints_prime_decls.txt +++ b/scripts/nolints_prime_decls.txt @@ -1117,7 +1117,7 @@ DoubleCentralizer.nnnorm_def' DoubleCentralizer.norm_def' dvd_antisymm' dvd_geom_sum₂_iff_of_dvd_sub' -edist_eq_coe_nnnorm' +edist_zero_eq_enorm' EllipticCurve.coe_inv_map_Δ' EllipticCurve.coe_inv_variableChange_Δ' EllipticCurve.coe_map_Δ' @@ -2823,7 +2823,7 @@ MeasureTheory.lintegral_mono_fn' MeasureTheory.lintegral_mono_set' MeasureTheory.lintegral_mul_const' MeasureTheory.lintegral_mul_const'' -MeasureTheory.lintegral_rpow_nnnorm_eq_rpow_eLpNorm' +MeasureTheory.lintegral_rpow_enorm_eq_rpow_eLpNorm' MeasureTheory.lintegral_singleton' MeasureTheory.lintegral_sub' MeasureTheory.lintegral_sub_le' @@ -3494,7 +3494,7 @@ Num.of_to_nat' Num.succ_ofInt' odd_add_one_self' odd_add_self_one' -ofReal_norm_eq_coe_nnnorm' +ofReal_norm_eq_enorm' OmegaCompletePartialOrder.const_continuous' OmegaCompletePartialOrder.ContinuousHom.bind_continuous' OmegaCompletePartialOrder.ContinuousHom.forall_forall_merge' diff --git a/scripts/noshake.json b/scripts/noshake.json index 5a6590b022253c..2732e7a608dbe0 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -303,8 +303,10 @@ "Mathlib.Tactic.CategoryTheory.BicategoryCoherence": ["Mathlib.CategoryTheory.Bicategory.Coherence", "Mathlib.Tactic.CategoryTheory.BicategoricalComp"], + "Mathlib.Tactic.CC.Datatypes": + ["Mathlib.Lean.Meta.CongrTheorems", "Mathlib.Lean.Meta.Datatypes"], "Mathlib.Tactic.CC.MkProof": - ["Mathlib.Logic.Basic", "Mathlib.Tactic.CC.Lemmas"], + ["Mathlib.Tactic.CC.Lemmas"], "Mathlib.Tactic.ByContra": ["Batteries.Tactic.Init"], "Mathlib.Tactic.Bound.Init": ["Aesop.Frontend.Command"], "Mathlib.Tactic.Bound.Attribute": @@ -340,6 +342,8 @@ "Mathlib.Probability.Notation": ["Mathlib.MeasureTheory.Decomposition.Lebesgue", "Mathlib.MeasureTheory.Function.ConditionalExpectation.Basic"], + "Mathlib.Probability.Kernel.Proper": + ["Mathlib.Probability.Kernel.Composition.CompNotation"], "Mathlib.Order.RelClasses": ["Batteries.WF"], "Mathlib.Order.Filter.ListTraverse": ["Mathlib.Control.Traversable.Instances"], diff --git a/scripts/technical-debt-metrics.sh b/scripts/technical-debt-metrics.sh index efe28e4c7e85d6..0aeb3052699316 100755 --- a/scripts/technical-debt-metrics.sh +++ b/scripts/technical-debt-metrics.sh @@ -73,11 +73,14 @@ computeDiff () { # The script uses the fact that a line represents a technical debt if and only if the text before # the first `|` is a number. This is then used for comparison and formatting. tdc () { +# We perform word-splitting "by hand" in the "middle" entries. +# See also the comment on the `read` line in the for-loop that follows the definition of this array. titlesPathsAndRegexes=( "porting notes" "*" "Porting note" "backwards compatibility flags" "*" "set_option.*backward" "skipAssignedInstances flags" "*" "set_option tactic.skipAssignedInstances" - "adaptation notes" "*" "adaptation_note" + "adaptation notes" ":^Mathlib/Tactic/AdaptationNote.lean :^Mathlib/Tactic/Linter" + "^[· ]*#adaptation_note" "disabled simpNF lints" "*" "nolint simpNF" "erw" "*" "erw \[" "maxHeartBeats modifications" ":^MathlibTest" "^ *set_option .*maxHeartbeats" @@ -87,13 +90,15 @@ for i in ${!titlesPathsAndRegexes[@]}; do # loop on every 3rd entry and name that entry and the following two if (( i % 3 == 0 )); then title="${titlesPathsAndRegexes[$i]}" - pathspec="${titlesPathsAndRegexes[$(( i + 1 ))]}" + # Here we perform word-splitting: `pathspec` is an array whose entries are the "words" in + # the string `"${titlesPathsAndRegexes[$(( i + 1 ))]}"`. + read -r -a pathspec <<< "${titlesPathsAndRegexes[$(( i + 1 ))]}" regex="${titlesPathsAndRegexes[$(( i + 2 ))]}" if [ "${title}" == "porting notes" ] then fl="-i" # just for porting notes we ignore the case in the regex else fl="--" fi - printf '%s|%s\n' "$(git grep "${fl}" "${regex}" -- "${pathspec}" | wc -l)" "${title}" + printf '%s|%s\n' "$(git grep "${fl}" "${regex}" -- ":^scripts" "${pathspec[@]}" | wc -l)" "${title}" fi done