From a8021cc6ba9f0b5cf3fa251cf16ec4af56db09d6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Micha=C5=82=20Staromiejski?= Date: Sat, 1 Feb 2025 06:30:14 +0000 Subject: [PATCH 1/8] feat(RingTheory/LocalRing): `IsLocalRing` for subrings (#21168) Results with some sufficient conditions for subrings of local rings to be local. --- Mathlib.lean | 1 + Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 3 + Mathlib/Algebra/Ring/Subsemiring/Defs.lean | 2 + .../RingTheory/LocalRing/LocalSubring.lean | 121 ++++++++++++++++ Mathlib/RingTheory/LocalRing/Subring.lean | 137 ++++-------------- .../RingTheory/Valuation/LocalSubring.lean | 2 +- 6 files changed, 158 insertions(+), 108 deletions(-) create mode 100644 Mathlib/RingTheory/LocalRing/LocalSubring.lean diff --git a/Mathlib.lean b/Mathlib.lean index 212c9a1c65d49..822965e01a44a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4634,6 +4634,7 @@ import Mathlib.RingTheory.LocalProperties.Reduced import Mathlib.RingTheory.LocalProperties.Submodule import Mathlib.RingTheory.LocalRing.Basic import Mathlib.RingTheory.LocalRing.Defs +import Mathlib.RingTheory.LocalRing.LocalSubring import Mathlib.RingTheory.LocalRing.MaximalIdeal.Basic import Mathlib.RingTheory.LocalRing.MaximalIdeal.Defs import Mathlib.RingTheory.LocalRing.Module diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index 65b72af50f1f4..c1789b31e7cc8 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -751,6 +751,9 @@ open RingHom def inclusion {S T : Subsemiring R} (h : S ≤ T) : S →+* T := S.subtype.codRestrict _ fun x => h x.2 +theorem inclusion_injective {S T : Subsemiring R} (h : S ≤ T) : + Function.Injective (inclusion h) := Set.inclusion_injective h + @[simp] theorem rangeS_subtype (s : Subsemiring R) : s.subtype.rangeS = s := SetLike.coe_injective <| (coe_rangeS _).trans Subtype.range_coe diff --git a/Mathlib/Algebra/Ring/Subsemiring/Defs.lean b/Mathlib/Algebra/Ring/Subsemiring/Defs.lean index f409e13f8b5bc..d2114e093146e 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Defs.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Defs.lean @@ -293,6 +293,8 @@ def subtype : s →+* R := theorem coe_subtype : ⇑s.subtype = ((↑) : s → R) := rfl +theorem subtype_injective : Function.Injective s.subtype := Subtype.coe_injective + protected theorem nsmul_mem {x : R} (hx : x ∈ s) (n : ℕ) : n • x ∈ s := nsmul_mem hx n diff --git a/Mathlib/RingTheory/LocalRing/LocalSubring.lean b/Mathlib/RingTheory/LocalRing/LocalSubring.lean new file mode 100644 index 0000000000000..90d074bc764e8 --- /dev/null +++ b/Mathlib/RingTheory/LocalRing/LocalSubring.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2024 Andrew Yang, Yaël Dillies, Javier López-Contreras. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang, Yaël Dillies, Javier López-Contreras +-/ +import Mathlib.Tactic.FieldSimp +import Mathlib.RingTheory.LocalRing.RingHom.Basic +import Mathlib.RingTheory.Localization.AtPrime + + +/-! +# Local subrings of fields + +# Main result +- `LocalSubring` : The class of local subrings of a commutative ring. +- `LocalSubring.ofPrime`: The localization of a subring as a `LocalSubring`. +-/ + +open IsLocalRing Set + +variable {R S : Type*} [CommRing R] [CommRing S] +variable {K : Type*} [Field K] + +instance [Nontrivial S] (f : R →+* S) (s : Subring R) [IsLocalRing s] : IsLocalRing (s.map f) := + .of_surjective' (f.restrict s _ (fun _ ↦ Set.mem_image_of_mem f)) + (fun ⟨_, a, ha, e⟩ ↦ ⟨⟨a, ha⟩, Subtype.ext e⟩) + +instance isLocalRing_top [IsLocalRing R] : IsLocalRing (⊤ : Subring R) := + Subring.topEquiv.symm.isLocalRing + +variable (R) in +/-- The class of local subrings of a commutative ring. -/ +@[ext] +structure LocalSubring where + /-- The underlying subring of a local subring. -/ + toSubring : Subring R + [isLocalRing : IsLocalRing toSubring] + +namespace LocalSubring + +attribute [instance] isLocalRing + +lemma toSubring_injective : Function.Injective (toSubring (R := R)) := by + rintro ⟨a, b⟩ ⟨c, d⟩ rfl; rfl + +/-- Copy of a local subring with a new `carrier` equal to the old one. +Useful to fix definitional equalities. -/ +protected def copy (S : LocalSubring R) (s : Set R) (hs : s = ↑S.toSubring) : LocalSubring R := + LocalSubring.mk (S.toSubring.copy s hs) (isLocalRing := hs ▸ S.2) + +/-- The image of a `LocalSubring` as a `LocalSubring`. -/ +@[simps! toSubring] +def map [Nontrivial S] (f : R →+* S) (s : LocalSubring R) : LocalSubring S := + mk (s.1.map f) + +/-- The range of a ringhom from a local ring as a `LocalSubring`. -/ +@[simps! toSubring] +def range [IsLocalRing R] [Nontrivial S] (f : R →+* S) : LocalSubring S := + .copy (map f (mk ⊤)) f.range (by ext x; exact congr(x ∈ $(Set.image_univ.symm))) + +/-- +The domination order on local subrings. +`A` dominates `B` if and only if `B ≤ A` (as subrings) and `m_A ∩ B = m_B`. +-/ +@[stacks 00I9] +instance : PartialOrder (LocalSubring R) where + le A B := ∃ h : A.1 ≤ B.1, IsLocalHom (Subring.inclusion h) + le_refl a := ⟨le_rfl, ⟨fun _ ↦ id⟩⟩ + le_trans A B C h₁ h₂ := ⟨h₁.1.trans h₂.1, @RingHom.isLocalHom_comp _ _ _ _ _ _ _ _ h₂.2 h₁.2⟩ + le_antisymm A B h₁ h₂ := toSubring_injective (le_antisymm h₁.1 h₂.1) + +/-- `A` dominates `B` if and only if `B ≤ A` (as subrings) and `m_A ∩ B = m_B`. -/ +lemma le_def {A B : LocalSubring R} : + A ≤ B ↔ ∃ h : A.toSubring ≤ B.toSubring, IsLocalHom (Subring.inclusion h) := Iff.rfl + +lemma toSubring_mono : Monotone (toSubring (R := R)) := + fun _ _ e ↦ e.1 + +section ofPrime + +variable (A : Subring K) (P : Ideal A) [P.IsPrime] + +/-- The localization of a subring at a prime, as a local subring. +Also see `Localization.subalgebra.ofField` -/ +noncomputable +def ofPrime (A : Subring K) (P : Ideal A) [P.IsPrime] : LocalSubring K := + range (IsLocalization.lift (M := P.primeCompl) (S := Localization.AtPrime P) + (g := A.subtype) (by simp [Ideal.primeCompl, not_imp_not])) + +lemma le_ofPrime : A ≤ (ofPrime A P).toSubring := by + intro x hx + exact ⟨algebraMap A _ ⟨x, hx⟩, by simp⟩ + +noncomputable +instance : Algebra A (ofPrime A P).toSubring := (Subring.inclusion (le_ofPrime A P)).toAlgebra + +instance : IsScalarTower A (ofPrime A P).toSubring K := .of_algebraMap_eq (fun _ ↦ rfl) + +/-- The localization of a subring at a prime is indeed isomorphic to its abstract localization. -/ +noncomputable +def ofPrimeEquiv : Localization.AtPrime P ≃ₐ[A] (ofPrime A P).toSubring := by + refine AlgEquiv.ofInjective (IsLocalization.liftAlgHom (M := P.primeCompl) + (S := Localization.AtPrime P) (f := Algebra.ofId A K) _) ?_ + intro x y e + obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective P.primeCompl x + obtain ⟨y, t, rfl⟩ := IsLocalization.mk'_surjective P.primeCompl y + have H (x : P.primeCompl) : x.1 ≠ 0 := by aesop + have : x.1 = y.1 * t.1.1⁻¹ * s.1.1 := by + simpa [IsLocalization.lift_mk', Algebra.ofId_apply, H, + Algebra.algebraMap_ofSubring_apply, IsUnit.coe_liftRight] using congr($e * s.1.1) + rw [IsLocalization.mk'_eq_iff_eq] + congr 1 + ext + field_simp [H t, this, mul_comm] + +instance : IsLocalization.AtPrime (ofPrime A P).toSubring P := + IsLocalization.isLocalization_of_algEquiv _ (ofPrimeEquiv A P) + +end ofPrime + +end LocalSubring diff --git a/Mathlib/RingTheory/LocalRing/Subring.lean b/Mathlib/RingTheory/LocalRing/Subring.lean index 90d074bc764e8..e8b5746196ff6 100644 --- a/Mathlib/RingTheory/LocalRing/Subring.lean +++ b/Mathlib/RingTheory/LocalRing/Subring.lean @@ -1,121 +1,44 @@ /- -Copyright (c) 2024 Andrew Yang, Yaël Dillies, Javier López-Contreras. All rights reserved. +Copyright (c) 2025 Michal Staromiejski. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Andrew Yang, Yaël Dillies, Javier López-Contreras +Authors: Michal Staromiejski -/ -import Mathlib.Tactic.FieldSimp -import Mathlib.RingTheory.LocalRing.RingHom.Basic -import Mathlib.RingTheory.Localization.AtPrime - +import Mathlib.Algebra.Ring.Subsemiring.Basic +import Mathlib.Algebra.GroupWithZero.NonZeroDivisors +import Mathlib.RingTheory.LocalRing.Defs /-! -# Local subrings of fields +# Subrings of local rings -# Main result -- `LocalSubring` : The class of local subrings of a commutative ring. -- `LocalSubring.ofPrime`: The localization of a subring as a `LocalSubring`. +We prove basic properties of subrings of local rings. -/ -open IsLocalRing Set - -variable {R S : Type*} [CommRing R] [CommRing S] -variable {K : Type*} [Field K] - -instance [Nontrivial S] (f : R →+* S) (s : Subring R) [IsLocalRing s] : IsLocalRing (s.map f) := - .of_surjective' (f.restrict s _ (fun _ ↦ Set.mem_image_of_mem f)) - (fun ⟨_, a, ha, e⟩ ↦ ⟨⟨a, ha⟩, Subtype.ext e⟩) - -instance isLocalRing_top [IsLocalRing R] : IsLocalRing (⊤ : Subring R) := - Subring.topEquiv.symm.isLocalRing - -variable (R) in -/-- The class of local subrings of a commutative ring. -/ -@[ext] -structure LocalSubring where - /-- The underlying subring of a local subring. -/ - toSubring : Subring R - [isLocalRing : IsLocalRing toSubring] - -namespace LocalSubring - -attribute [instance] isLocalRing - -lemma toSubring_injective : Function.Injective (toSubring (R := R)) := by - rintro ⟨a, b⟩ ⟨c, d⟩ rfl; rfl - -/-- Copy of a local subring with a new `carrier` equal to the old one. -Useful to fix definitional equalities. -/ -protected def copy (S : LocalSubring R) (s : Set R) (hs : s = ↑S.toSubring) : LocalSubring R := - LocalSubring.mk (S.toSubring.copy s hs) (isLocalRing := hs ▸ S.2) - -/-- The image of a `LocalSubring` as a `LocalSubring`. -/ -@[simps! toSubring] -def map [Nontrivial S] (f : R →+* S) (s : LocalSubring R) : LocalSubring S := - mk (s.1.map f) - -/-- The range of a ringhom from a local ring as a `LocalSubring`. -/ -@[simps! toSubring] -def range [IsLocalRing R] [Nontrivial S] (f : R →+* S) : LocalSubring S := - .copy (map f (mk ⊤)) f.range (by ext x; exact congr(x ∈ $(Set.image_univ.symm))) - -/-- -The domination order on local subrings. -`A` dominates `B` if and only if `B ≤ A` (as subrings) and `m_A ∩ B = m_B`. --/ -@[stacks 00I9] -instance : PartialOrder (LocalSubring R) where - le A B := ∃ h : A.1 ≤ B.1, IsLocalHom (Subring.inclusion h) - le_refl a := ⟨le_rfl, ⟨fun _ ↦ id⟩⟩ - le_trans A B C h₁ h₂ := ⟨h₁.1.trans h₂.1, @RingHom.isLocalHom_comp _ _ _ _ _ _ _ _ h₂.2 h₁.2⟩ - le_antisymm A B h₁ h₂ := toSubring_injective (le_antisymm h₁.1 h₂.1) - -/-- `A` dominates `B` if and only if `B ≤ A` (as subrings) and `m_A ∩ B = m_B`. -/ -lemma le_def {A B : LocalSubring R} : - A ≤ B ↔ ∃ h : A.toSubring ≤ B.toSubring, IsLocalHom (Subring.inclusion h) := Iff.rfl - -lemma toSubring_mono : Monotone (toSubring (R := R)) := - fun _ _ e ↦ e.1 - -section ofPrime - -variable (A : Subring K) (P : Ideal A) [P.IsPrime] - -/-- The localization of a subring at a prime, as a local subring. -Also see `Localization.subalgebra.ofField` -/ -noncomputable -def ofPrime (A : Subring K) (P : Ideal A) [P.IsPrime] : LocalSubring K := - range (IsLocalization.lift (M := P.primeCompl) (S := Localization.AtPrime P) - (g := A.subtype) (by simp [Ideal.primeCompl, not_imp_not])) - -lemma le_ofPrime : A ≤ (ofPrime A P).toSubring := by - intro x hx - exact ⟨algebraMap A _ ⟨x, hx⟩, by simp⟩ +namespace IsLocalRing -noncomputable -instance : Algebra A (ofPrime A P).toSubring := (Subring.inclusion (le_ofPrime A P)).toAlgebra +variable {R S} [Semiring R] [Semiring S] -instance : IsScalarTower A (ofPrime A P).toSubring K := .of_algebraMap_eq (fun _ ↦ rfl) +open nonZeroDivisors -/-- The localization of a subring at a prime is indeed isomorphic to its abstract localization. -/ -noncomputable -def ofPrimeEquiv : Localization.AtPrime P ≃ₐ[A] (ofPrime A P).toSubring := by - refine AlgEquiv.ofInjective (IsLocalization.liftAlgHom (M := P.primeCompl) - (S := Localization.AtPrime P) (f := Algebra.ofId A K) _) ?_ - intro x y e - obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective P.primeCompl x - obtain ⟨y, t, rfl⟩ := IsLocalization.mk'_surjective P.primeCompl y - have H (x : P.primeCompl) : x.1 ≠ 0 := by aesop - have : x.1 = y.1 * t.1.1⁻¹ * s.1.1 := by - simpa [IsLocalization.lift_mk', Algebra.ofId_apply, H, - Algebra.algebraMap_ofSubring_apply, IsUnit.coe_liftRight] using congr($e * s.1.1) - rw [IsLocalization.mk'_eq_iff_eq] - congr 1 - ext - field_simp [H t, this, mul_comm] +/-- If a (semi)ring `R` in which every element is either invertible or a zero divisor +embeds in a local (semi)ring `S`, then `R` is local. -/ +theorem of_injective [IsLocalRing S] {f : R →+* S} (hf : Function.Injective f) + (h : ∀ a, a ∈ R⁰ → IsUnit a) : IsLocalRing R := by + haveI : Nontrivial R := f.domain_nontrivial + refine .of_is_unit_or_is_unit_of_add_one fun {a b} hab ↦ + (IsLocalRing.isUnit_or_isUnit_of_add_one (map_add f .. ▸ map_one f ▸ congrArg f hab)).imp ?_ ?_ + <;> exact h _ ∘ mem_nonZeroDivisor_of_injective hf ∘ IsUnit.mem_nonZeroDivisors -instance : IsLocalization.AtPrime (ofPrime A P).toSubring P := - IsLocalization.isLocalization_of_algEquiv _ (ofPrimeEquiv A P) +/-- If in a sub(semi)ring `R` of a local (semi)ring `S` every element is either +invertible or a zero divisor, then `R` is local. -/ +theorem of_subring [IsLocalRing S] {R : Subsemiring S} (h : ∀ a, a ∈ R⁰ → IsUnit a) : + IsLocalRing R := + of_injective R.subtype_injective h -end ofPrime +/-- If in a sub(semi)ring `R` of a local (semi)ring `R'` every element is either +invertible or a zero divisor, then `R` is local. +This version is for `R` and `R'` that are both sub(semi)rings of a (semi)ring `S`. -/ +theorem of_subring' {R R' : Subsemiring S} [IsLocalRing R'] (inc : R ≤ R') + (h : ∀ a, a ∈ R⁰ → IsUnit a) : IsLocalRing R := + of_injective (Subsemiring.inclusion_injective inc) h -end LocalSubring +end IsLocalRing diff --git a/Mathlib/RingTheory/Valuation/LocalSubring.lean b/Mathlib/RingTheory/Valuation/LocalSubring.lean index ae5fef7208b40..ac0c4253b0f3e 100644 --- a/Mathlib/RingTheory/Valuation/LocalSubring.lean +++ b/Mathlib/RingTheory/Valuation/LocalSubring.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang, Yaël Dillies, Javier López-Contreras -/ import Mathlib.RingTheory.Ideal.Over -import Mathlib.RingTheory.LocalRing.Subring +import Mathlib.RingTheory.LocalRing.LocalSubring import Mathlib.RingTheory.Polynomial.Ideal import Mathlib.RingTheory.Valuation.ValuationSubring From 05f71fad0df43db93c048f8a30253e7bcdbf1d31 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Sat, 1 Feb 2025 07:23:26 +0000 Subject: [PATCH 2/8] feat(GroupTheory/Perm/Centralizer): study the centralizer of a permutation (#17522) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is the core of the work on the centralizer of a permutation. It is the sequel of several PR which lay out basic useful results. Let `α : Type` with `Fintype α` (and `DecidableEq α`). The main goal of this file is to compute the cardinality of conjugacy classes in `Equiv.Perm α`. Every `g : Equiv.Perm α` has a `cycleType α : Multiset ℕ`. By `Equiv.Perm.isConj_iff_cycleType_eq`, two permutations are conjugate in `Equiv.Perm α` iff their cycle types are equal. To compute the cardinality of the conjugacy classes, we could use a purely combinatorial approach and compute the number of permutations with given cycle type but we resorted to a more algebraic approach. A subsequent PR #17047 treats the case of the alternating group. Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: Johan Commelin --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Subgroup/Finite.lean | 22 + Mathlib/GroupTheory/NoncommCoprod.lean | 38 +- Mathlib/GroupTheory/Perm/Basic.lean | 7 + Mathlib/GroupTheory/Perm/Centralizer.lean | 647 ++++++++++++++++++++ Mathlib/GroupTheory/Perm/Cycle/Basic.lean | 36 +- Mathlib/GroupTheory/Perm/Cycle/Factors.lean | 50 ++ Mathlib/GroupTheory/Perm/Finite.lean | 45 ++ Mathlib/GroupTheory/Perm/Support.lean | 17 +- 9 files changed, 847 insertions(+), 16 deletions(-) create mode 100644 Mathlib/GroupTheory/Perm/Centralizer.lean diff --git a/Mathlib.lean b/Mathlib.lean index 822965e01a44a..0fab6812c75fa 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3362,6 +3362,7 @@ import Mathlib.GroupTheory.OreLocalization.Cardinality import Mathlib.GroupTheory.OreLocalization.OreSet import Mathlib.GroupTheory.PGroup import Mathlib.GroupTheory.Perm.Basic +import Mathlib.GroupTheory.Perm.Centralizer import Mathlib.GroupTheory.Perm.Closure import Mathlib.GroupTheory.Perm.ClosureSwap import Mathlib.GroupTheory.Perm.ConjAct diff --git a/Mathlib/Algebra/Group/Subgroup/Finite.lean b/Mathlib/Algebra/Group/Subgroup/Finite.lean index c5f14ed4bd10a..104d627327f02 100644 --- a/Mathlib/Algebra/Group/Subgroup/Finite.lean +++ b/Mathlib/Algebra/Group/Subgroup/Finite.lean @@ -151,6 +151,18 @@ theorem card_le_card_group [Finite G] : Nat.card H ≤ Nat.card G := theorem card_le_of_le {H K : Subgroup G} [Finite K] (h : H ≤ K) : Nat.card H ≤ Nat.card K := Nat.card_le_card_of_injective _ (Subgroup.inclusion_injective h) +@[to_additive] +theorem card_map_of_injective {H : Type*} [Group H] {K : Subgroup G} {f : G →* H} + (hf : Function.Injective f) : + Nat.card (map f K) = Nat.card K := by + -- simp only [← SetLike.coe_sort_coe] + apply Nat.card_image_of_injective hf + +@[to_additive] +theorem card_subtype (K : Subgroup G) (L : Subgroup K) : + Nat.card (map K.subtype L) = Nat.card L := + card_map_of_injective K.subtype_injective + end Subgroup namespace Subgroup @@ -275,4 +287,14 @@ Note: this instance can form a diamond with `Subtype.fintype` or `Subgroup.finty instance fintypeRange [Fintype G] [DecidableEq N] (f : G →* N) : Fintype (range f) := Set.fintypeRange f +lemma _root_.Fintype.card_coeSort_mrange {M N : Type*} [Monoid M] [Monoid N] [Fintype M] + [DecidableEq N] {f : M →* N} (hf : Function.Injective f) : + Fintype.card (mrange f) = Fintype.card M := + Set.card_range_of_injective hf + +lemma _root_.Fintype.card_coeSort_range [Fintype G] [DecidableEq N] {f : G →* N} + (hf : Function.Injective f) : + Fintype.card (range f) = Fintype.card G := + Set.card_range_of_injective hf + end MonoidHom diff --git a/Mathlib/GroupTheory/NoncommCoprod.lean b/Mathlib/GroupTheory/NoncommCoprod.lean index f6064b9bd922f..8bb21af12efde 100644 --- a/Mathlib/GroupTheory/NoncommCoprod.lean +++ b/Mathlib/GroupTheory/NoncommCoprod.lean @@ -5,6 +5,9 @@ Authors: Antoine Chambert-Loir -/ import Mathlib.Algebra.Group.Commute.Hom import Mathlib.Algebra.Group.Prod +import Mathlib.Algebra.Group.Subgroup.Ker +import Mathlib.Algebra.Group.Subgroup.Lattice +import Mathlib.Order.Disjoint /-! # Canonical homomorphism from a pair of monoids @@ -53,7 +56,6 @@ theorem noncommCoprod_apply' (comm) (mn : M × N) : (f.noncommCoprod g comm) mn = g mn.2 * f mn.1 := by rw [← comm, noncommCoprod_apply] - @[to_additive] theorem comp_noncommCoprod {Q : Type*} [Semigroup Q] (h : P →ₙ* Q) (comm : ∀ m n, Commute (f m) (g n)) : @@ -114,4 +116,38 @@ theorem comp_noncommCoprod {Q : Type*} [Monoid Q] (h : P →* Q) : (h.comp f).noncommCoprod (h.comp g) (fun m n ↦ (comm m n).map h) := ext fun x => by simp +section group + +open Subgroup + +lemma noncommCoprod_injective {M N P : Type*} [Group M] [Group N] [Group P] + (f : M →* P) (g : N →* P) (comm : ∀ (m : M) (n : N), Commute (f m) (g n)) : + Function.Injective (noncommCoprod f g comm) ↔ + (Function.Injective f ∧ Function.Injective g ∧ _root_.Disjoint f.range g.range) := by + simp only [injective_iff_map_eq_one, disjoint_iff_inf_le, + noncommCoprod_apply, Prod.forall, Prod.mk_eq_one] + refine ⟨fun h ↦ ⟨fun x ↦ ?_, fun x ↦ ?_, ?_⟩, ?_⟩ + · simpa using h x 1 + · simpa using h 1 x + · intro x ⟨⟨y, hy⟩, z, hz⟩ + rwa [(h y z⁻¹ (by rw [map_inv, hy, hz, mul_inv_cancel])).1, map_one, eq_comm] at hy + · intro ⟨hf, hg, hp⟩ a b h + have key := hp ⟨⟨a⁻¹, by rwa [map_inv, inv_eq_iff_mul_eq_one]⟩, b, rfl⟩ + exact ⟨hf a (by rwa [key, mul_one] at h), hg b key⟩ + +lemma noncommCoprod_range {M N P : Type*} [Group M] [Group N] [Group P] + (f : M →* P) (g : N →* P) (comm : ∀ (m : M) (n : N), Commute (f m) (g n)) : + (noncommCoprod f g comm).range = f.range ⊔ g.range := by + apply le_antisymm + · rintro - ⟨a, rfl⟩ + exact mul_mem (mem_sup_left ⟨a.1, rfl⟩) (mem_sup_right ⟨a.2, rfl⟩) + · rw [sup_le_iff] + constructor + · rintro - ⟨a, rfl⟩ + exact ⟨(a, 1), by rw [noncommCoprod_apply, map_one, mul_one]⟩ + · rintro - ⟨a, rfl⟩ + exact ⟨(1, a), by rw [noncommCoprod_apply, map_one, one_mul]⟩ + +end group + end MonoidHom diff --git a/Mathlib/GroupTheory/Perm/Basic.lean b/Mathlib/GroupTheory/Perm/Basic.lean index 70e6c423d6dfd..3ab26e2457edb 100644 --- a/Mathlib/GroupTheory/Perm/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Basic.lean @@ -399,6 +399,13 @@ theorem mem_iff_ofSubtype_apply_mem (f : Perm (Subtype p)) (x : α) : simpa only [h, true_iff, MonoidHom.coe_mk, ofSubtype_apply_of_mem f h] using (f ⟨x, h⟩).2 else by simp [h, ofSubtype_apply_of_not_mem f h] +theorem ofSubtype_injective : Function.Injective (ofSubtype : Perm (Subtype p) → Perm α) := by + intro x y h + rw [Perm.ext_iff] at h ⊢ + intro a + specialize h a + rwa [ofSubtype_apply_coe, ofSubtype_apply_coe, SetCoe.ext_iff] at h + @[simp] theorem subtypePerm_ofSubtype (f : Perm (Subtype p)) : subtypePerm (ofSubtype f) (mem_iff_ofSubtype_apply_mem f) = f := diff --git a/Mathlib/GroupTheory/Perm/Centralizer.lean b/Mathlib/GroupTheory/Perm/Centralizer.lean new file mode 100644 index 0000000000000..fb283ed718095 --- /dev/null +++ b/Mathlib/GroupTheory/Perm/Centralizer.lean @@ -0,0 +1,647 @@ +/- +Copyright (c) 2023 Antoine Chambert-Loir. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir +-/ + +import Mathlib.Algebra.Order.BigOperators.GroupWithZero.Multiset +import Mathlib.Algebra.Order.BigOperators.Ring.Finset +import Mathlib.GroupTheory.NoncommCoprod +import Mathlib.GroupTheory.NoncommPiCoprod +import Mathlib.GroupTheory.Perm.ConjAct +import Mathlib.GroupTheory.Perm.Cycle.Factors +import Mathlib.GroupTheory.Perm.Cycle.PossibleTypes +import Mathlib.GroupTheory.Perm.DomMulAct +import Mathlib.GroupTheory.Perm.Finite + +/-! # Centralizer of a permutation and cardinality of conjugacy classes + # in the symmetric groups + +Let `α : Type` with `Fintype α` (and `DecidableEq α`). +The main goal of this file is to compute the cardinality of +conjugacy classes in `Equiv.Perm α`. +Every `g : Equiv.Perm α` has a `cycleType α : Multiset ℕ`. +By `Equiv.Perm.isConj_iff_cycleType_eq`, +two permutations are conjugate in `Equiv.Perm α` iff +their cycle types are equal. +To compute the cardinality of the conjugacy classes, we could use +a purely combinatorial approach and compute the number of permutations +with given cycle type but we resorted to a more algebraic approach +based on the study of the centralizer of a permutation `g`. + +Given `g : Equiv.Perm α`, the conjugacy class of `g` is the orbit +of `g` under the action `ConjAct (Equiv.Perm α)`, and we use the +orbit-stabilizer theorem +(`MulAction.card_orbit_mul_card_stabilizer_eq_card_group`) to reduce +the computation to the computation of the centralizer of `g`, the +subgroup of `Equiv.Perm α` consisting of all permutations which +commute with `g`. It is accessed here as `MulAction.stabilizer +(ConjAct (Equiv.Perm α)) g` and +`Equiv.Perm.centralizer_eq_comap_stabilizer`. + +We compute this subgroup as follows. + +* If `h : Subgroup.centralizer {g}`, then the action of `ConjAct.toConjAct h` + by conjugation on `Equiv.Perm α` stabilizes `g.cycleFactorsFinset`. + That induces an action of `Subgroup.centralizer {g}` on + `g.cycleFactorsFinset` which is defined via + `Equiv.Perm.OnCycleFactors.subMulActionOnCycleFactors` + +* This action defines a group morphism `Equiv.Perm.OnCycleFactors.toPermHom g` + from `Subgroup.centralizer {g}` to `Equiv.Perm g.cycleFactorsFinset`. + +* `Equiv.Perm.OnCycleFactors.range_toPermHom'` is the subgroup of + `Equiv.Perm g.cycleFactorsFinset` consisting of permutations that + preserve the cardinality of the support. + +* `Equiv.Perm.OnCycleFactors.range_toPermHom_eq_range_toPermHom' shows that + the range of `Equiv.Perm.OnCycleFactors.toPermHom g` + is the subgroup `toPermHom_range' g` of `Equiv.Perm g.cycleFactorsFinset`. + +This is shown by constructing a right inverse +`Equiv.Perm.Basis.toCentralizer`, as established by +`Equiv.Perm.Basis.toPermHom_apply_toCentralizer`. + +* `Equiv.Perm.OnCycleFactors.nat_card_range_toPermHom` computes the + cardinality of `(Equiv.Perm.OnCycleFactors.toPermHom g).range` + as a product of factorials. + +* `Equiv.Perm.OnCycleFactors.mem_ker_toPermHom_iff` proves that + `k : Subgroup.centralizer {g}` belongs to the kernel of + `Equiv.Perm.OnCycleFactors.toPermHom g` if and only if it commutes with + each cycle of `g`. This is equivalent to the conjunction of two properties: + * `k` preserves the set of fixed points of `g`; + * on each cycle `c`, `k` acts as a power of that cycle. + +This allows to give a description of the kernel of +`Equiv.Perm.OnCycleFactors.toPermHom g` as the product of a +symmetric group and of a product of cyclic groups. This analysis +starts with the morphism `Equiv.Perm.OnCycleFactors.θ`, its +injectivity `Equiv.Perm.OnCycleFactors.θ_injective`, its range +`Equiv.Perm.OnCycleFactors.θ_range_eq`, and its cardinality +`Equiv.Perm.OnCycleFactors.θ_range_card`. + +* `Equiv.Perm.nat_card_centralizer g` computes the cardinality + of the centralizer of `g`. + +* `Equiv.Perm.card_isConj_mul_eq g`computes the cardinality + of the conjugacy class of `g`. + +* We now can compute the cardinality of the set of permutations with given cycle type. + The condition for this cardinality to be zero is given by + `Equiv.Perm.card_of_cycleType_eq_zero_iff` + which is itself derived from `Equiv.Perm.exists_with_cycleType_iff`. + +* `Equiv.Perm.card_of_cycleType_mul_eq m` and `Equiv.Perm.card_of_cycleType m` + compute this cardinality. + +-/ + +open scoped Pointwise + +namespace Equiv.Perm + +open MulAction Equiv Subgroup + +variable {α : Type*} [DecidableEq α] [Fintype α] {g : Equiv.Perm α} + +namespace OnCycleFactors + +variable (g) + +variable {g} in +lemma Subgroup.Centralizer.toConjAct_smul_mem_cycleFactorsFinset {k c : Perm α} + (k_mem : k ∈ centralizer {g}) (c_mem : c ∈ g.cycleFactorsFinset) : + ConjAct.toConjAct k • c ∈ g.cycleFactorsFinset := by + suffices (g.cycleFactorsFinset : Set (Perm α)) = + (ConjAct.toConjAct k) • g.cycleFactorsFinset by + rw [← Finset.mem_coe, this] + simp only [Set.smul_mem_smul_set_iff, Finset.mem_coe, Finset.coe_mem, c_mem] + have this := cycleFactorsFinset_conj_eq (ConjAct.toConjAct (k : Perm α)) g + rw [ConjAct.toConjAct_smul, mem_centralizer_singleton_iff.mp k_mem, mul_assoc] at this + simp only [mul_inv_cancel, mul_one] at this + conv_lhs => rw [this] + simp only [Finset.coe_smul_finset] + +/-- The action by conjugation of `Subgroup.centraliser {g}` + on the cycles of a given permutation -/ +def Subgroup.Centralizer.cycleFactorsFinset_mulAction : + MulAction (centralizer {g}) g.cycleFactorsFinset where + smul k c := ⟨ConjAct.toConjAct (k : Perm α) • (c : Perm α), + Subgroup.Centralizer.toConjAct_smul_mem_cycleFactorsFinset k.prop c.prop⟩ + one_smul c := by + rw [← Subtype.coe_inj] + change ConjAct.toConjAct (1 : Perm α) • (c : Perm α) = c + simp only [map_one, one_smul] + mul_smul k l c := by + simp only [← Subtype.coe_inj] + change ConjAct.toConjAct (k * l : Perm α) • (c : Perm α) = + ConjAct.toConjAct (k : Perm α) • (ConjAct.toConjAct (l : Perm α)) • (c : Perm α) + simp only [map_mul, mul_smul] + +/-- The conjugation action of `Subgroup.centralizer {g}` on `g.cycleFactorsFinset` -/ +scoped instance : MulAction (centralizer {g}) (g.cycleFactorsFinset) := + (Subgroup.Centralizer.cycleFactorsFinset_mulAction g) + +/-- The canonical morphism from `Subgroup.centralizer {g}` + to the group of permutations of `g.cycleFactorsFinset` -/ +def toPermHom := MulAction.toPermHom (centralizer {g}) g.cycleFactorsFinset + +theorem centralizer_smul_def (k : centralizer {g}) (c : g.cycleFactorsFinset) : + k • c = ⟨k * c * k⁻¹, + Subgroup.Centralizer.toConjAct_smul_mem_cycleFactorsFinset k.prop c.prop⟩ := + rfl + +@[simp] +theorem val_centralizer_smul (k : Subgroup.centralizer {g}) (c : g.cycleFactorsFinset) : + ((k • c :) : Perm α) = k * c * k⁻¹ := + rfl + +theorem toPermHom_apply (k : centralizer {g}) (c : g.cycleFactorsFinset) : + (toPermHom g k c) = k • c := rfl + +theorem coe_toPermHom (k : centralizer {g}) (c : g.cycleFactorsFinset) : + (toPermHom g k c : Perm α) = k * c * (k : Perm α)⁻¹ := rfl + +/-- The range of `Equiv.Perm.OnCycleFactors.toPermHom`. + +The equality is proved by `Equiv.Perm.OnCycleFactors.range_toPermHom_eq_range_toPermHom'`. -/ +def range_toPermHom' : Subgroup (Perm g.cycleFactorsFinset) where + carrier := {τ | ∀ c, (τ c : Perm α).support.card = (c : Perm α).support.card} + one_mem' := by + simp only [Set.mem_setOf_eq, coe_one, id_eq, eq_self_iff_true, imp_true_iff] + mul_mem' hσ hτ := by + simp only [Subtype.forall, Set.mem_setOf_eq, coe_mul, Function.comp_apply] + simp only [Subtype.forall, Set.mem_setOf_eq] at hσ hτ + intro c hc + rw [hσ, hτ] + inv_mem' hσ := by + simp only [Subtype.forall, Set.mem_setOf_eq] at hσ ⊢ + intro c hc + rw [← hσ] + · simp only [Finset.coe_mem, Subtype.coe_eta, apply_inv_self] + · simp only [Finset.coe_mem] + +variable {g} in +theorem mem_range_toPermHom'_iff {τ : Perm g.cycleFactorsFinset} : + τ ∈ range_toPermHom' g ↔ + ∀ c, (τ c : Perm α).support.card = (c : Perm α).support.card := + Iff.rfl + +variable (k : centralizer {g}) + +/-- `k : Subgroup.centralizer {g}` belongs to the kernel of `toPermHom g` + iff it commutes with each cycle of `g` -/ +theorem mem_ker_toPermHom_iff : + k ∈ (toPermHom g).ker ↔ ∀ c ∈ g.cycleFactorsFinset, Commute (k : Perm α) c := by + simp only [toPermHom, MonoidHom.mem_ker, DFunLike.ext_iff, Subtype.forall] + refine forall₂_congr (fun _ _ ↦ ?_) + simp [← Subtype.coe_inj, commute_iff_eq, mul_inv_eq_iff_eq_mul] + +end OnCycleFactors + +open OnCycleFactors Subgroup + +/-- A `Basis` of a permutation is a choice of an element in each of its cycles -/ +structure Basis (g : Equiv.Perm α) where + /-- A choice of elements in each cycle -/ + (toFun : g.cycleFactorsFinset → α) + /-- For each cycle, the chosen element belongs to the cycle -/ + (mem_support_self' : ∀ (c : g.cycleFactorsFinset), toFun c ∈ (c : Perm α).support) + +instance (g : Perm α) : + DFunLike (Basis g) (g.cycleFactorsFinset) (fun _ => α) where + coe a := a.toFun + coe_injective' a a' _ := by cases a; cases a'; congr + +namespace Basis + +theorem nonempty (g : Perm α) : Nonempty (Basis g) := by + have (c : g.cycleFactorsFinset) : (c : Perm α).support.Nonempty := + IsCycle.nonempty_support (mem_cycleFactorsFinset_iff.mp c.prop).1 + exact ⟨fun c ↦ (this c).choose, fun c ↦ (this c).choose_spec⟩ + +variable (a : Basis g) (c : g.cycleFactorsFinset) + +theorem mem_support_self : + a c ∈ (c : Perm α).support := a.mem_support_self' c + +theorem injective : Function.Injective a := by + intro c d h + rw [← Subtype.coe_inj] + apply g.cycleFactorsFinset_pairwise_disjoint.eq c.prop d.prop + simp only [Disjoint, not_forall, not_or] + use a c + conv_rhs => rw [h] + simp only [← Perm.mem_support, a.mem_support_self c, a.mem_support_self d, and_self] + +theorem cycleOf_eq : g.cycleOf (a c) = c := + (cycle_is_cycleOf (a.mem_support_self c) c.prop).symm + +theorem sameCycle {x : α} (hx : g.cycleOf x ∈ g.cycleFactorsFinset) : + g.SameCycle (a ⟨g.cycleOf x, hx⟩) x := + (mem_support_cycleOf_iff.mp (a.mem_support_self ⟨g.cycleOf x, hx⟩)).1.symm + +variable (τ : range_toPermHom' g) + +/-- The function that will provide a right inverse `toCentralizer` to `toPermHom` -/ +def ofPermHomFun (x : α) : α := + if hx : g.cycleOf x ∈ g.cycleFactorsFinset + then + (g ^ (Nat.find (a.sameCycle hx).exists_nat_pow_eq)) + (a ((τ : Perm g.cycleFactorsFinset) ⟨g.cycleOf x, hx⟩)) + else x + +theorem mem_fixedPoints_or_exists_zpow_eq (x : α) : + x ∈ Function.fixedPoints g ∨ + ∃ (c : g.cycleFactorsFinset) (_ : x ∈ (c : Perm α).support) (m : ℤ), (g ^ m) (a c) = x := by + rw [Classical.or_iff_not_imp_left] + intro hx + rw [Function.mem_fixedPoints_iff, ← ne_eq, ← mem_support, + ← cycleOf_mem_cycleFactorsFinset_iff] at hx + refine ⟨⟨g.cycleOf x, hx⟩, ?_, (a.sameCycle hx)⟩ + rw [mem_support_cycleOf_iff, ← cycleOf_mem_cycleFactorsFinset_iff] + simp [SameCycle.rfl, hx, and_self] + +theorem ofPermHomFun_apply_of_cycleOf_mem {x : α} {c : g.cycleFactorsFinset} + (hx : x ∈ (c : Perm α).support) {m : ℤ} (hm : (g ^ m) (a c) = x) : + ofPermHomFun a τ x = (g ^ m) (a ((τ : Perm g.cycleFactorsFinset) c)) := by + have hx' : c = g.cycleOf x := cycle_is_cycleOf hx (Subtype.prop c) + have hx'' : g.cycleOf x ∈ g.cycleFactorsFinset := hx' ▸ c.prop + set n := Nat.find (a.sameCycle hx'').exists_nat_pow_eq + have hn : (g ^ (n : ℤ)) (a c) = x := by + rw [← Nat.find_spec (a.sameCycle hx'').exists_nat_pow_eq, zpow_natCast] + congr + rw [← Subtype.coe_inj, hx'] + suffices ofPermHomFun a τ x = (g ^ (n : ℤ)) (a ((τ : Perm g.cycleFactorsFinset) c)) by + rw [this, IsCycleOn.zpow_apply_eq_zpow_apply + (isCycleOn_support_of_mem_cycleFactorsFinset ((τ : Perm g.cycleFactorsFinset) c).prop) + (mem_support_self a ((τ : Perm g.cycleFactorsFinset) c))] + simp only [τ.prop c] + rw [← IsCycleOn.zpow_apply_eq_zpow_apply + (isCycleOn_support_of_mem_cycleFactorsFinset c.prop) (mem_support_self a c)] + rw [hn, hm] + simp only [ofPermHomFun, dif_pos hx''] + congr + exact hx'.symm + +theorem ofPermHomFun_apply_of_mem_fixedPoints {x : α} (hx : x ∈ Function.fixedPoints g) : + ofPermHomFun a τ x = x := by + rw [ofPermHomFun, dif_neg] + rw [cycleOf_mem_cycleFactorsFinset_iff, not_mem_support] + exact hx + +theorem ofPermHomFun_apply_mem_support_cycle_iff {x : α} {c : g.cycleFactorsFinset} : + ofPermHomFun a τ x ∈ ((τ : Perm g.cycleFactorsFinset) c : Perm α).support ↔ + x ∈ (c : Perm α).support := by + rcases mem_fixedPoints_or_exists_zpow_eq a x with (hx | ⟨d, hd, m, hm⟩) + · simp only [ofPermHomFun_apply_of_mem_fixedPoints a τ hx] + suffices ∀ (d : g.cycleFactorsFinset), x ∉ (d : Perm α).support by + simp only [this] + intro d hx' + rw [Function.mem_fixedPoints_iff, ← not_mem_support] at hx + apply hx + exact mem_cycleFactorsFinset_support_le d.prop hx' + · rw [ofPermHomFun_apply_of_cycleOf_mem a τ hd hm] -- + rw [zpow_apply_mem_support_of_mem_cycleFactorsFinset_iff] + by_cases h : c = d + · simp only [h, hd, iff_true, mem_support_self] + · have H : Disjoint (c : Perm α) (d : Perm α) := + cycleFactorsFinset_pairwise_disjoint g c.prop d.prop (Subtype.coe_ne_coe.mpr h) + have H' : Disjoint ((τ : Perm g.cycleFactorsFinset) c : Perm α) + ((τ : Perm g.cycleFactorsFinset) d : Perm α) := + cycleFactorsFinset_pairwise_disjoint g ((τ : Perm g.cycleFactorsFinset) c).prop + ((τ : Perm g.cycleFactorsFinset) d).prop (by + intro h'; apply h + simpa only [Subtype.coe_inj, EmbeddingLike.apply_eq_iff_eq] using h') + rw [disjoint_iff_disjoint_support, Finset.disjoint_right] at H H' + simp only [H hd, H' (mem_support_self a _)] + +theorem ofPermHomFun_commute_zpow_apply (x : α) (j : ℤ) : + ofPermHomFun a τ ((g ^ j) x) = (g ^ j) (ofPermHomFun a τ x) := by + rcases mem_fixedPoints_or_exists_zpow_eq a x with (hx | hx) + · rw [ofPermHomFun_apply_of_mem_fixedPoints a τ hx, ofPermHomFun_apply_of_mem_fixedPoints] + rw [Function.mem_fixedPoints_iff] + simp only [← mul_apply, ← zpow_one_add, add_comm] + conv_rhs => rw [← hx, ← mul_apply, ← zpow_add_one] + · obtain ⟨c, hc, m, hm⟩ := hx + have hm' : (g ^ (j + m)) (a c) = (g ^ j) x := by rw [zpow_add, mul_apply, hm] + rw [ofPermHomFun_apply_of_cycleOf_mem a τ hc hm, ofPermHomFun_apply_of_cycleOf_mem a τ _ hm', + ← mul_apply, ← zpow_add] + exact zpow_apply_mem_support_of_mem_cycleFactorsFinset_iff.mpr hc + +theorem ofPermHomFun_mul (σ τ : range_toPermHom' g) (x) : + ofPermHomFun a (σ * τ) x = (ofPermHomFun a σ) (ofPermHomFun a τ x) := by + rcases mem_fixedPoints_or_exists_zpow_eq a x with (hx | ⟨c, hc, m, hm⟩) + · simp only [ofPermHomFun_apply_of_mem_fixedPoints a _ hx] + · simp only [ofPermHomFun_apply_of_cycleOf_mem a _ hc hm] + rw [ofPermHomFun_apply_of_cycleOf_mem a _ _ rfl] + · rfl + · rw [zpow_apply_mem_support_of_mem_cycleFactorsFinset_iff] + apply mem_support_self + +theorem ofPermHomFun_one (x : α) : (ofPermHomFun a 1) x = x := by + rcases mem_fixedPoints_or_exists_zpow_eq a x with (hx | ⟨c, hc, m, hm⟩) + · rw [ofPermHomFun_apply_of_mem_fixedPoints a _ hx] + · rw [ofPermHomFun_apply_of_cycleOf_mem a _ hc hm, OneMemClass.coe_one, coe_one, id_eq, hm] + +/-- Given `a : g.Basis` and a permutation of `g.cycleFactorsFinset` that + preserve the lengths of the cycles, a permutation of `α` that + moves the `Basis` and commutes with `g` -/ +noncomputable def ofPermHom : range_toPermHom' g →* Perm α where + toFun τ := { + toFun := ofPermHomFun a τ + invFun := ofPermHomFun a τ⁻¹ + left_inv := fun x ↦ by rw [← ofPermHomFun_mul, inv_mul_cancel, ofPermHomFun_one] + right_inv := fun x ↦ by rw [← ofPermHomFun_mul, mul_inv_cancel, ofPermHomFun_one] } + map_one' := ext fun x ↦ ofPermHomFun_one a x + map_mul' := fun σ τ ↦ ext fun x ↦ by simp [mul_apply, ofPermHomFun_mul a σ τ x] + +theorem ofPermHom_mem_centralizer : + a.ofPermHom τ ∈ centralizer {g} := by + rw [mem_centralizer_singleton_iff] + ext x + simp only [mul_apply] + exact ofPermHomFun_commute_zpow_apply a τ x 1 + +/-- Given `a : Equiv.Perm.Basis g`, +we define a right inverse of `Equiv.Perm.OnCycleFactors.toPermHom`, +on `range_toPermHom' g` -/ +noncomputable def toCentralizer : + range_toPermHom' g →* centralizer {g} where + toFun τ := ⟨ofPermHom a τ, ofPermHom_mem_centralizer a τ⟩ + map_one' := by simp only [map_one, mk_eq_one] + map_mul' σ τ := by simp only [map_mul, MulMemClass.mk_mul_mk] + +theorem toCentralizer_apply (x) : (toCentralizer a τ : Perm α) x = ofPermHomFun a τ x := rfl + +theorem toCentralizer_equivariant : + (toCentralizer a τ) • c = (τ : Perm g.cycleFactorsFinset) c := by + simp only [← Subtype.coe_inj, val_centralizer_smul, InvMemClass.coe_inv, mul_inv_eq_iff_eq_mul] + ext x + simp only [mul_apply, toCentralizer_apply] + by_cases hx : x ∈ (c : Perm α).support + · rw [(mem_cycleFactorsFinset_iff.mp c.prop).2 x hx] + have := ofPermHomFun_commute_zpow_apply a τ x 1 + simp only [zpow_one] at this + rw [this, ← (mem_cycleFactorsFinset_iff.mp ((τ : Perm g.cycleFactorsFinset) c).prop).2] + rw [ofPermHomFun_apply_mem_support_cycle_iff] + exact hx + · rw [not_mem_support.mp hx, eq_comm, ← not_mem_support, + ofPermHomFun_apply_mem_support_cycle_iff] + exact hx + +theorem toPermHom_apply_toCentralizer : + (toPermHom g) (toCentralizer a τ) = (τ : Perm g.cycleFactorsFinset) := by + apply ext + intro c + rw [OnCycleFactors.toPermHom_apply, toCentralizer_equivariant] + +end Basis + +namespace OnCycleFactors + +open Basis BigOperators Nat Equiv.Perm Equiv Subgroup + +theorem mem_range_toPermHom_iff {τ} : τ ∈ (toPermHom g).range ↔ + ∀ c, (τ c : Perm α).support.card = (c : Perm α).support.card := by + constructor + · rintro ⟨k, rfl⟩ c + rw [coe_toPermHom, Equiv.Perm.support_conj] + apply Finset.card_map + · obtain ⟨a⟩ := Basis.nonempty g + exact fun hτ ↦ ⟨toCentralizer a ⟨τ, hτ⟩, toPermHom_apply_toCentralizer a ⟨τ, hτ⟩⟩ + +/-- Unapplied variant of `Equiv.Perm.mem_range_toPermHom_iff` -/ +theorem mem_range_toPermHom_iff' {τ} : τ ∈ (toPermHom g).range ↔ + (fun (c : g.cycleFactorsFinset) ↦ (c : Perm α).support.card) ∘ τ = + fun (c : g.cycleFactorsFinset) ↦ (c : Perm α).support.card := by + rw [mem_range_toPermHom_iff, funext_iff] + simp only [Finset.coe_sort_coe, Subtype.forall, Function.comp_apply] + +/-- Computes the range of `Equiv.Perm.toPermHom g` -/ +theorem range_toPermHom_eq_range_toPermHom' : + (toPermHom g).range = range_toPermHom' g := by + ext τ + rw [mem_range_toPermHom_iff, mem_range_toPermHom'_iff] + +theorem nat_card_range_toPermHom : + Nat.card (toPermHom g).range = + ∏ n ∈ g.cycleType.toFinset, (g.cycleType.count n)! := by + classical + set sc : (c : g.cycleFactorsFinset) → ℕ := fun c ↦ (c : Perm α).support.card with hsc + suffices Fintype.card (toPermHom g).range = + Fintype.card { k : Perm g.cycleFactorsFinset | sc ∘ k = sc } by + simp only [Nat.card_eq_fintype_card, this, Set.coe_setOf, DomMulAct.stabilizer_card', hsc, + Finset.univ_eq_attach] + simp_rw [← CycleType.count_def] + apply Finset.prod_congr _ (fun _ _ => rfl) + ext n + simp only [Finset.univ_eq_attach, Finset.mem_image, Finset.mem_attach, + hsc, true_and, Subtype.exists, exists_prop, Multiset.mem_toFinset] + simp only [cycleType_def, Function.comp_apply, Multiset.mem_map, Finset.mem_val] + simp only [← SetLike.coe_sort_coe, Fintype.card_eq_nat_card] + congr + ext + rw [SetLike.mem_coe, mem_range_toPermHom_iff', Set.mem_setOf_eq, ← beq_eq_beq] + +section Kernel +/- Here, we describe the kernel of `g.OnCycleFactors.toPermHom` -/ + +open BigOperators Nat OnCycleFactors Subgroup + +variable (g) in +/-- The parametrization of the kernel of `toPermHom` -/ +def kerParam : (Perm (Function.fixedPoints g)) × + ((c : g.cycleFactorsFinset) → Subgroup.zpowers (c : Perm α)) →* Perm α := + MonoidHom.noncommCoprod ofSubtype (Subgroup.noncommPiCoprod g.pairwise_commute_of_mem_zpowers) + g.commute_ofSubtype_noncommPiCoprod + + +theorem kerParam_apply {u : Perm (Function.fixedPoints g)} + {v : (c : g.cycleFactorsFinset) → Subgroup.zpowers (c : Perm α)} {x : α} : + kerParam g (u,v) x = + if hx : g.cycleOf x ∈ g.cycleFactorsFinset + then (v ⟨g.cycleOf x, hx⟩ : Perm α) x + else ofSubtype u x := by + split_ifs with hx + · have hx' := hx + rw [cycleOf_mem_cycleFactorsFinset_iff, mem_support, Ne, ← Function.mem_fixedPoints_iff] at hx' + rw [kerParam, MonoidHom.noncommCoprod_apply', mul_apply, ofSubtype_apply_of_not_mem u hx', + noncommPiCoprod_apply, ← Finset.noncommProd_erase_mul _ (Finset.mem_univ ⟨g.cycleOf x, hx⟩), + mul_apply, ← not_mem_support] + contrapose! hx' + obtain ⟨a, ha1, ha2⟩ := mem_support_of_mem_noncommProd_support hx' + simp only [Finset.mem_erase, Finset.mem_univ, and_true, Ne, Subtype.ext_iff] at ha1 + have key := cycleFactorsFinset_pairwise_disjoint g a.2 hx ha1 + rw [disjoint_iff_disjoint_support, Finset.disjoint_left] at key + obtain ⟨k, hk⟩ := mem_zpowers_iff.mp (v a).2 + replace ha2 := key (support_zpow_le a.1 k (hk ▸ ha2)) + obtain ⟨k, hk⟩ := mem_zpowers_iff.mp (v ⟨g.cycleOf x, hx⟩).2 + rwa [← hk, zpow_apply_mem_support, not_mem_support, cycleOf_apply_self] at ha2 + · rw [cycleOf_mem_cycleFactorsFinset_iff] at hx + rw [kerParam, MonoidHom.noncommCoprod_apply, mul_apply, Equiv.apply_eq_iff_eq, + ← not_mem_support] + contrapose! hx + obtain ⟨a, -, ha⟩ := mem_support_of_mem_noncommProd_support + (comm := fun a ha b hb h ↦ g.pairwise_commute_of_mem_zpowers h (v a) (v b) (v a).2 (v b).2) hx + exact support_zpowers_of_mem_cycleFactorsFinset_le (v a) ha + +theorem kerParam_injective (g : Perm α) : Function.Injective (kerParam g) := by + rw [kerParam, MonoidHom.noncommCoprod_injective] + refine ⟨ofSubtype_injective, ?_, ?_⟩ + · apply MonoidHom.injective_noncommPiCoprod_of_iSupIndep + · intro a + simp only [range_subtype, ne_eq] + simp only [zpowers_eq_closure, ← closure_iUnion] + apply disjoint_closure_of_disjoint_support + rintro - ⟨-⟩ - ⟨-, ⟨b, rfl⟩, -, ⟨h, rfl⟩, ⟨-⟩⟩ + rw [← disjoint_iff_disjoint_support] + apply cycleFactorsFinset_pairwise_disjoint g a.2 b.2 + simp only [ne_eq, ← Subtype.ext_iff] + exact ne_comm.mp h + · exact fun i ↦ subtype_injective _ + · rw [noncommPiCoprod_range, ← ofSubtype.range.closure_eq] + simp only [zpowers_eq_closure, ← closure_iUnion] + apply disjoint_closure_of_disjoint_support + rintro - ⟨a, rfl⟩ - ⟨-, ⟨b, rfl⟩, ⟨-⟩⟩ + exact (ofSubtype_support_disjoint a).mono_right (mem_cycleFactorsFinset_support_le b.2) + +theorem kerParam_range_eq : + (kerParam g).range = (toPermHom g).ker.map (Subgroup.subtype _) := by + apply le_antisymm + · rw [kerParam, MonoidHom.noncommCoprod_range, sup_le_iff, noncommPiCoprod_range, iSup_le_iff] + simp only [zpowers_le] + constructor + · rintro - ⟨a, rfl⟩ + refine ⟨⟨ofSubtype a, ?_⟩, ?_, rfl⟩ + · rw [mem_centralizer_singleton_iff] + exact Disjoint.commute (disjoint_iff_disjoint_support.mpr (ofSubtype_support_disjoint a)) + · exact Perm.ext fun x ↦ Subtype.ext (disjoint_iff_disjoint_support.mpr + ((ofSubtype_support_disjoint a).mono_right + (mem_cycleFactorsFinset_support_le x.2))).commute.mul_inv_cancel + · intro i + refine ⟨⟨i, mem_centralizer_singleton_iff.mpr (self_mem_cycle_factors_commute i.2)⟩, ?_, rfl⟩ + exact Perm.ext fun x ↦ Subtype.ext (cycleFactorsFinset_mem_commute' g i.2 x.2).mul_inv_cancel + · rintro - ⟨p, hp, rfl⟩ + simp only [coeSubtype] + set u : Perm (Function.fixedPoints g) := + subtypePerm p (fun x ↦ mem_fixedPoints_iff_apply_mem_of_mem_centralizer p.2) + simp only [SetLike.mem_coe, mem_ker_toPermHom_iff, IsCycle.forall_commute_iff] at hp + set v : (c : g.cycleFactorsFinset) → (Subgroup.zpowers (c : Perm α)) := + fun c => ⟨ofSubtype + (p.1.subtypePerm (Classical.choose (hp c.val c.prop))), + Classical.choose_spec (hp c.val c.prop)⟩ + use (u, v) + ext x + rw [kerParam_apply] + split_ifs with hx + · rw [cycleOf_mem_cycleFactorsFinset_iff, mem_support] at hx + rw [ofSubtype_apply_of_mem, subtypePerm_apply] + rwa [mem_support, cycleOf_apply_self, ne_eq] + · rw [cycleOf_mem_cycleFactorsFinset_iff, not_mem_support] at hx + rwa [ofSubtype_apply_of_mem, subtypePerm_apply] + +theorem kerParam_range_card (g : Equiv.Perm α) : + Fintype.card (kerParam g).range = (Fintype.card α - g.cycleType.sum)! * g.cycleType.prod := by + rw [Fintype.card_coeSort_range (kerParam_injective g)] + rw [Fintype.card_prod, Fintype.card_perm, Fintype.card_pi, card_fixedPoints] + apply congr_arg + rw [Finset.univ_eq_attach, g.cycleFactorsFinset.prod_attach (fun i ↦ Fintype.card (zpowers i)), + cycleType, Finset.prod_map_val] + refine Finset.prod_congr rfl (fun x hx ↦ ?_) + rw [Fintype.card_zpowers, (mem_cycleFactorsFinset_iff.mp hx).1.orderOf, Function.comp_apply] + +end Kernel + +end OnCycleFactors + +open Nat + +variable (g : Perm α) + +-- Should one parenthesize the product ? +/-- Cardinality of the centralizer in `Equiv.Perm α` of a permutation given `cycleType` -/ +theorem nat_card_centralizer : + Nat.card (centralizer {g}) = + (Fintype.card α - g.cycleType.sum)! * g.cycleType.prod * + (∏ n ∈ g.cycleType.toFinset, (g.cycleType.count n)!) := by + rw [← (toPermHom g).ker.card_mul_index, index_ker, nat_card_range_toPermHom, + ← kerParam_range_card, ← Nat.card_eq_fintype_card, kerParam_range_eq, card_subtype] + +theorem card_isConj_mul_eq : + Nat.card {h : Perm α | IsConj g h} * + ((Fintype.card α - g.cycleType.sum)! * + g.cycleType.prod * + (∏ n ∈ g.cycleType.toFinset, (g.cycleType.count n)!)) = + (Fintype.card α)! := by + classical + rw [Nat.card_eq_fintype_card, ← nat_card_centralizer g] + rw [Subgroup.nat_card_centralizer_nat_card_stabilizer, Nat.card_eq_fintype_card] + convert MulAction.card_orbit_mul_card_stabilizer_eq_card_group (ConjAct (Perm α)) g + · ext h + simp only [Set.mem_setOf_eq, ConjAct.mem_orbit_conjAct, isConj_comm] + · rw [ConjAct.card, Fintype.card_perm] + +/-- Cardinality of a conjugacy class in `Equiv.Perm α` of a given `cycleType` -/ +theorem card_isConj_eq : + Nat.card {h : Perm α | IsConj g h} = + (Fintype.card α)! / + ((Fintype.card α - g.cycleType.sum)! * + g.cycleType.prod * + (∏ n ∈ g.cycleType.toFinset, (g.cycleType.count n)!)) := by + rw [← card_isConj_mul_eq g, Nat.div_eq_of_eq_mul_left _] + · rfl + -- This is the cardinal of the centralizer + · rw [← nat_card_centralizer g] + apply Nat.card_pos + +variable (α) + +theorem card_of_cycleType_eq_zero_iff {m : Multiset ℕ} : + ({g | g.cycleType = m} : Finset (Perm α)).card = 0 + ↔ ¬ ((m.sum ≤ Fintype.card α ∧ ∀ a ∈ m, 2 ≤ a)) := by + rw [Finset.card_eq_zero, Finset.filter_eq_empty_iff, + ← exists_with_cycleType_iff, not_exists] + aesop + +theorem card_of_cycleType_mul_eq (m : Multiset ℕ) : + ({g | g.cycleType = m} : Finset (Perm α)).card * + ((Fintype.card α - m.sum)! * m.prod * + (∏ n ∈ m.toFinset, (m.count n)!)) = + if (m.sum ≤ Fintype.card α ∧ ∀ a ∈ m, 2 ≤ a) then (Fintype.card α)! else 0 := by + split_ifs with hm + · -- nonempty case + classical + obtain ⟨g, rfl⟩ := (exists_with_cycleType_iff α).mpr hm + convert card_isConj_mul_eq g + simp_rw [Set.coe_setOf, Nat.card_eq_fintype_card, ← Fintype.card_coe, Finset.mem_filter, + Finset.mem_univ, true_and, ← isConj_iff_cycleType_eq, isConj_comm (g := g)] + · -- empty case + convert MulZeroClass.zero_mul _ + exact (card_of_cycleType_eq_zero_iff α).mpr hm + +/-- Cardinality of the `Finset` of `Equiv.Perm α` of given `cycleType` -/ +theorem card_of_cycleType (m : Multiset ℕ) : + ({g | g.cycleType = m} : Finset (Perm α)).card = + if m.sum ≤ Fintype.card α ∧ ∀ a ∈ m, 2 ≤ a then + (Fintype.card α)! / + ((Fintype.card α - m.sum)! * m.prod * (∏ n ∈ m.toFinset, (m.count n)!)) + else 0 := by + split_ifs with hm + · -- nonempty case + apply symm + apply Nat.div_eq_of_eq_mul_left + · apply Nat.mul_pos + · apply Nat.mul_pos + · apply Nat.factorial_pos + · apply Multiset.prod_pos + exact fun a ha ↦ lt_of_lt_of_le (zero_lt_two) (hm.2 a ha) + · exact Finset.prod_pos (fun _ _ ↦ Nat.factorial_pos _) + rw [card_of_cycleType_mul_eq, if_pos hm] + · -- empty case + exact (card_of_cycleType_eq_zero_iff α).mpr hm + +end Equiv.Perm + + diff --git a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean index dd80ddc889ffc..05e5a64c67a95 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean @@ -174,23 +174,31 @@ theorem sameCycle_extendDomain {p : β → Prop} [DecidablePred p] {f : α ≃ S alias ⟨_, SameCycle.extendDomain⟩ := sameCycle_extendDomain theorem SameCycle.exists_pow_eq' [Finite α] : SameCycle f x y → ∃ i < orderOf f, (f ^ i) x = y := by - classical - rintro ⟨k, rfl⟩ - use (k % orderOf f).natAbs - have h₀ := Int.natCast_pos.mpr (orderOf_pos f) - have h₁ := Int.emod_nonneg k h₀.ne' - rw [← zpow_natCast, Int.natAbs_of_nonneg h₁, zpow_mod_orderOf] - refine ⟨?_, by rfl⟩ - rw [← Int.ofNat_lt, Int.natAbs_of_nonneg h₁] - exact Int.emod_lt_of_pos _ h₀ + rintro ⟨k, rfl⟩ + use (k % orderOf f).natAbs + have h₀ := Int.natCast_pos.mpr (orderOf_pos f) + have h₁ := Int.emod_nonneg k h₀.ne' + rw [← zpow_natCast, Int.natAbs_of_nonneg h₁, zpow_mod_orderOf] + refine ⟨?_, by rfl⟩ + rw [← Int.ofNat_lt, Int.natAbs_of_nonneg h₁] + exact Int.emod_lt_of_pos _ h₀ theorem SameCycle.exists_pow_eq'' [Finite α] (h : SameCycle f x y) : ∃ i : ℕ, 0 < i ∧ i ≤ orderOf f ∧ (f ^ i) x = y := by - classical - obtain ⟨_ | i, hi, rfl⟩ := h.exists_pow_eq' - · refine ⟨orderOf f, orderOf_pos f, le_rfl, ?_⟩ - rw [pow_orderOf_eq_one, pow_zero] - · exact ⟨i.succ, i.zero_lt_succ, hi.le, by rfl⟩ + obtain ⟨_ | i, hi, rfl⟩ := h.exists_pow_eq' + · refine ⟨orderOf f, orderOf_pos f, le_rfl, ?_⟩ + rw [pow_orderOf_eq_one, pow_zero] + · exact ⟨i.succ, i.zero_lt_succ, hi.le, by rfl⟩ + +theorem SameCycle.exists_fin_pow_eq [Finite α] (h : SameCycle f x y) : + ∃ i : Fin (orderOf f), (f ^ (i : ℕ)) x = y := by + obtain ⟨i, hi, hx⟩ := SameCycle.exists_pow_eq' h + exact ⟨⟨i, hi⟩, hx⟩ + +theorem SameCycle.exists_nat_pow_eq [Finite α] (h : SameCycle f x y) : + ∃ i : ℕ, (f ^ i) x = y := by + obtain ⟨i, _, hi⟩ := h.exists_pow_eq' + exact ⟨i, hi⟩ instance (f : Perm α) [DecidableRel (SameCycle f)] : DecidableRel (SameCycle f⁻¹) := fun x y => diff --git a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean index 516cc85380d9d..e5f7bc27cc91e 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean @@ -6,6 +6,7 @@ Authors: Chris Hughes, Yaël Dillies import Mathlib.Data.List.Iterate import Mathlib.GroupTheory.Perm.Cycle.Basic +import Mathlib.GroupTheory.NoncommPiCoprod /-! # Cycle factors of a permutation @@ -523,9 +524,18 @@ theorem cycleFactorsFinset_pairwise_disjoint : (cycleFactorsFinset f : Set (Perm α)).Pairwise Disjoint := (cycleFactorsFinset_eq_finset.mp rfl).2.choose +/-- Two cycles of a permutation commute. -/ theorem cycleFactorsFinset_mem_commute : (cycleFactorsFinset f : Set (Perm α)).Pairwise Commute := (cycleFactorsFinset_pairwise_disjoint _).mono' fun _ _ => Disjoint.commute +/-- Two cycles of a permutation commute. -/ +theorem cycleFactorsFinset_mem_commute' {g1 g2 : Perm α} + (h1 : g1 ∈ f.cycleFactorsFinset) (h2 : g2 ∈ f.cycleFactorsFinset) : + Commute g1 g2 := by + rcases eq_or_ne g1 g2 with rfl | h + · apply Commute.refl + · exact Equiv.Perm.cycleFactorsFinset_mem_commute f h1 h2 h + /-- The product of cycle factors is equal to the original `f : perm α`. -/ theorem cycleFactorsFinset_noncommProd (comm : (cycleFactorsFinset f : Set (Perm α)).Pairwise Commute := @@ -576,6 +586,40 @@ lemma support_zpowers_of_mem_cycleFactorsFinset_le {g : Perm α} simp only [← hm] exact le_trans (support_zpow_le _ _) (mem_cycleFactorsFinset_support_le c.prop) +theorem pairwise_disjoint_of_mem_zpowers : + Pairwise fun (i j : f.cycleFactorsFinset) ↦ + ∀ (x y : Perm α), x ∈ Subgroup.zpowers ↑i → y ∈ Subgroup.zpowers ↑j → Disjoint x y := + fun c d hcd ↦ fun x y hx hy ↦ by + obtain ⟨m, hm⟩ := hx; obtain ⟨n, hn⟩ := hy + simp only [← hm, ← hn] + apply Disjoint.zpow_disjoint_zpow + exact f.cycleFactorsFinset_pairwise_disjoint c.prop d.prop (Subtype.coe_ne_coe.mpr hcd) + +lemma pairwise_commute_of_mem_zpowers : + Pairwise fun (i j : f.cycleFactorsFinset) ↦ + ∀ (x y : Perm α), x ∈ Subgroup.zpowers ↑i → y ∈ Subgroup.zpowers ↑j → Commute x y := + f.pairwise_disjoint_of_mem_zpowers.mono + (fun _ _ ↦ forall₂_imp (fun _ _ h hx hy ↦ (h hx hy).commute)) + +lemma disjoint_ofSubtype_noncommPiCoprod (u : Perm (Function.fixedPoints f)) + (v : (c : { x // x ∈ f.cycleFactorsFinset }) → (Subgroup.zpowers (c : Perm α))) : + Disjoint (ofSubtype u) ((Subgroup.noncommPiCoprod f.pairwise_commute_of_mem_zpowers) v) := by + apply Finset.noncommProd_induction + · intro a _ b _ h + apply f.pairwise_commute_of_mem_zpowers h <;> simp only [Subgroup.coeSubtype, SetLike.coe_mem] + · intro x y + exact Disjoint.mul_right + · exact disjoint_one_right _ + · intro c _ + simp only [Subgroup.coeSubtype] + exact Disjoint.mono (disjoint_ofSubtype_of_memFixedPoints_self u) + le_rfl (support_zpowers_of_mem_cycleFactorsFinset_le (v c)) + +lemma commute_ofSubtype_noncommPiCoprod (u : Perm (Function.fixedPoints f)) + (v : (c : { x // x ∈ f.cycleFactorsFinset }) → (Subgroup.zpowers (c : Perm α))) : + Commute (ofSubtype u) ((Subgroup.noncommPiCoprod f.pairwise_commute_of_mem_zpowers) v) := + Disjoint.commute (f.disjoint_ofSubtype_noncommPiCoprod u v) + theorem mem_support_iff_mem_support_of_mem_cycleFactorsFinset {g : Equiv.Perm α} {x : α} : x ∈ g.support ↔ ∃ c ∈ g.cycleFactorsFinset, x ∈ c.support := by constructor @@ -681,6 +725,12 @@ theorem eq_cycleOf_of_mem_cycleFactorsFinset_iff rw [mem_support, cycleOf_apply_self, ne_eq, ← cycleOf_eq_one_iff] exact (mem_cycleFactorsFinset_iff.mp hc).left.ne_one +theorem zpow_apply_mem_support_of_mem_cycleFactorsFinset_iff {g : Perm α} + {x :α} {m : ℤ} {c : g.cycleFactorsFinset} : + (g ^ m) x ∈ (c : Perm α).support ↔ x ∈ (c : Perm α).support := by + rw [← g.eq_cycleOf_of_mem_cycleFactorsFinset_iff _ c.prop, cycleOf_self_apply_zpow, + eq_cycleOf_of_mem_cycleFactorsFinset_iff _ _ c.prop] + /-- A permutation `c` is a cycle of `g` iff `k * c * k⁻¹` is a cycle of `k * g * k⁻¹` -/ theorem mem_cycleFactorsFinset_conj (g k c : Perm α) : k * c * k⁻¹ ∈ (k * g * k⁻¹).cycleFactorsFinset ↔ c ∈ g.cycleFactorsFinset := by diff --git a/Mathlib/GroupTheory/Perm/Finite.lean b/Mathlib/GroupTheory/Perm/Finite.lean index ce64b56eddb1d..11496da19bea0 100644 --- a/Mathlib/GroupTheory/Perm/Finite.lean +++ b/Mathlib/GroupTheory/Perm/Finite.lean @@ -248,6 +248,51 @@ theorem support_pow_coprime {σ : Perm α} {n : ℕ} (h : Nat.Coprime n (orderOf le_antisymm (support_pow_le σ n) (le_trans (ge_of_eq (congr_arg support hm)) (support_pow_le (σ ^ n) m)) +lemma ofSubtype_support_disjoint {σ : Perm α} (x : Perm (Function.fixedPoints σ)) : + _root_.Disjoint x.ofSubtype.support σ.support := by + rw [Finset.disjoint_iff_ne] + rintro a ha b hb rfl + rw [mem_support] at ha hb + exact ha (ofSubtype_apply_of_not_mem x (mt Function.mem_fixedPoints_iff.mp hb)) + +open Subgroup + +lemma disjoint_of_disjoint_support {H K : Subgroup (Perm α)} + (h : ∀ a ∈ H, ∀ b ∈ K, _root_.Disjoint a.support b.support) : + _root_.Disjoint H K := by + rw [disjoint_iff_inf_le] + intro x ⟨hx1, hx2⟩ + specialize h x hx1 x hx2 + rwa [disjoint_self, Finset.bot_eq_empty, support_eq_empty_iff] at h + +lemma support_closure_subset_union (S : Set (Perm α)) : + ∀ a ∈ closure S, (a.support : Set α) ⊆ ⋃ b ∈ S, b.support := by + apply closure_induction + · exact fun x hx ↦ Set.subset_iUnion₂_of_subset x hx subset_rfl + · simp only [support_one, Finset.coe_empty, Set.empty_subset] + · intro a b ha hb hc hd + refine (Finset.coe_subset.mpr (support_mul_le a b)).trans ?_ + rw [Finset.sup_eq_union, Finset.coe_union, Set.union_subset_iff] + exact ⟨hc, hd⟩ + · simp only [support_inv, imp_self, implies_true] + +lemma disjoint_support_closure_of_disjoint_support {S T : Set (Perm α)} + (h : ∀ a ∈ S, ∀ b ∈ T, _root_.Disjoint a.support b.support) : + ∀ a ∈ closure S, ∀ b ∈ closure T, _root_.Disjoint a.support b.support := by + intro a ha b hb + have key1 := support_closure_subset_union S a ha + have key2 := support_closure_subset_union T b hb + have key := Set.disjoint_of_subset key1 key2 + simp_rw [Set.disjoint_iUnion_left, Set.disjoint_iUnion_right, Finset.disjoint_coe] at key + exact key h + +lemma disjoint_closure_of_disjoint_support {S T : Set (Perm α)} + (h : ∀ a ∈ S, ∀ b ∈ T, _root_.Disjoint a.support b.support) : + _root_.Disjoint (closure S) (closure T) := by + apply disjoint_of_disjoint_support + apply disjoint_support_closure_of_disjoint_support + exact h + end Fintype end Equiv.Perm diff --git a/Mathlib/GroupTheory/Perm/Support.lean b/Mathlib/GroupTheory/Perm/Support.lean index e151ea67cc71e..148582b1e77eb 100644 --- a/Mathlib/GroupTheory/Perm/Support.lean +++ b/Mathlib/GroupTheory/Perm/Support.lean @@ -29,7 +29,7 @@ Assume `α` is a Fintype: -/ -open Equiv Finset +open Equiv Finset Function namespace Equiv.Perm @@ -377,6 +377,21 @@ theorem support_ofSubtype {p : α → Prop} [DecidablePred p] (u : Perm (Subtype · simp only [forall_prop_of_true hx, ofSubtype_apply_of_mem u hx, ← Subtype.coe_inj] · simp only [forall_prop_of_false hx, true_iff, ofSubtype_apply_of_not_mem u hx] +theorem mem_support_of_mem_noncommProd_support {α β : Type*} [DecidableEq β] [Fintype β] + {s : Finset α} {f : α → Perm β} + {comm : (s : Set α).Pairwise (Commute on f)} {x : β} (hx : x ∈ (s.noncommProd f comm).support) : + ∃ a ∈ s, x ∈ (f a).support := by + contrapose! hx + classical + revert hx comm s + apply Finset.induction + · simp + · intro a s ha ih comm hs + rw [Finset.noncommProd_insert_of_not_mem s a f comm ha] + apply mt (Finset.mem_of_subset (support_mul_le _ _)) + rw [Finset.sup_eq_union, Finset.not_mem_union] + exact ⟨hs a (s.mem_insert_self a), ih (fun a ha ↦ hs a (Finset.mem_insert_of_mem ha))⟩ + theorem pow_apply_mem_support {n : ℕ} {x : α} : (f ^ n) x ∈ f.support ↔ x ∈ f.support := by simp only [mem_support, ne_eq, apply_pow_apply_eq_iff] From 74bf56fd21d80405dbf30378aea2a84da88787ef Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 1 Feb 2025 07:23:28 +0000 Subject: [PATCH 3/8] chore(Data/ZMod/Basic): split `ZMod.valMinAbs` off (#21308) --- .../Wiedijk100Theorems/FriendshipGraphs.lean | 3 - .../CanonicallyOrderedCommSemiringTwoMul.lean | 3 +- .../CliffordAlgebraNotInjective.lean | 5 - Counterexamples/HomogeneousPrimeNotPrime.lean | 3 +- Counterexamples/QuadraticForm.lean | 2 - Mathlib.lean | 1 + Mathlib/Algebra/Module/ZMod.lean | 2 - Mathlib/Data/Nat/Defs.lean | 3 + Mathlib/Data/Nat/Totient.lean | 2 - Mathlib/Data/ZMod/Basic.lean | 177 +----------------- Mathlib/Data/ZMod/Units.lean | 1 - Mathlib/Data/ZMod/ValMinAbs.lean | 176 +++++++++++++++++ Mathlib/FieldTheory/Finite/Basic.lean | 4 +- Mathlib/GroupTheory/Coxeter/Length.lean | 1 + .../GroupTheory/SpecificGroups/Cyclic.lean | 4 - .../SpecificGroups/Quaternion.lean | 3 - .../CliffordAlgebra/Grading.lean | 1 - Mathlib/NumberTheory/LucasLehmer.lean | 2 + .../ModularForms/CongruenceSubgroups.lean | 3 - Mathlib/NumberTheory/Pell.lean | 2 +- Mathlib/NumberTheory/PythagoreanTriples.lean | 7 +- Mathlib/NumberTheory/SumFourSquares.lean | 4 - Mathlib/RingTheory/Fintype.lean | 1 - 23 files changed, 191 insertions(+), 219 deletions(-) create mode 100644 Mathlib/Data/ZMod/ValMinAbs.lean diff --git a/Archive/Wiedijk100Theorems/FriendshipGraphs.lean b/Archive/Wiedijk100Theorems/FriendshipGraphs.lean index c5d0604033257..f3e377c74d1dd 100644 --- a/Archive/Wiedijk100Theorems/FriendshipGraphs.lean +++ b/Archive/Wiedijk100Theorems/FriendshipGraphs.lean @@ -5,9 +5,6 @@ Authors: Aaron Anderson, Jalex Stark, Kyle Miller -/ import Mathlib.Combinatorics.SimpleGraph.AdjMatrix import Mathlib.LinearAlgebra.Matrix.Charpoly.FiniteField -import Mathlib.Data.Int.ModEq -import Mathlib.Data.ZMod.Basic -import Mathlib.Tactic.IntervalCases /-! # The Friendship Theorem diff --git a/Counterexamples/CanonicallyOrderedCommSemiringTwoMul.lean b/Counterexamples/CanonicallyOrderedCommSemiringTwoMul.lean index 59a7f0d9f5e26..a7bd5fad38057 100644 --- a/Counterexamples/CanonicallyOrderedCommSemiringTwoMul.lean +++ b/Counterexamples/CanonicallyOrderedCommSemiringTwoMul.lean @@ -3,9 +3,8 @@ Copyright (c) 2021 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ -import Mathlib.Data.ZMod.Basic -import Mathlib.Algebra.Order.Monoid.Basic import Mathlib.Algebra.Ring.Subsemiring.Order +import Mathlib.Data.ZMod.Basic /-! diff --git a/Counterexamples/CliffordAlgebraNotInjective.lean b/Counterexamples/CliffordAlgebraNotInjective.lean index cb8014abd1bf9..98b40e868ce29 100644 --- a/Counterexamples/CliffordAlgebraNotInjective.lean +++ b/Counterexamples/CliffordAlgebraNotInjective.lean @@ -5,12 +5,7 @@ Authors: Eric Wieser -/ import Mathlib.Algebra.CharP.Pi import Mathlib.Algebra.CharP.Quotient -import Mathlib.Algebra.CharP.Two -import Mathlib.Algebra.MvPolynomial.CommRing -import Mathlib.Data.ZMod.Basic -import Mathlib.LinearAlgebra.CliffordAlgebra.Basic import Mathlib.LinearAlgebra.CliffordAlgebra.Contraction -import Mathlib.LinearAlgebra.Finsupp.SumProd import Mathlib.RingTheory.MvPolynomial.Basic import Mathlib.RingTheory.MvPolynomial.Ideal diff --git a/Counterexamples/HomogeneousPrimeNotPrime.lean b/Counterexamples/HomogeneousPrimeNotPrime.lean index 61c133a80e292..f53eceac454ba 100644 --- a/Counterexamples/HomogeneousPrimeNotPrime.lean +++ b/Counterexamples/HomogeneousPrimeNotPrime.lean @@ -4,9 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Eric Wieser, Jujian Zhang -/ import Mathlib.Algebra.Divisibility.Prod +import Mathlib.Data.Fintype.Units import Mathlib.RingTheory.GradedAlgebra.HomogeneousIdeal -import Mathlib.Data.ZMod.Basic -import Mathlib.Tactic.DeriveFintype /-! # A homogeneous ideal that is homogeneously prime but not prime diff --git a/Counterexamples/QuadraticForm.lean b/Counterexamples/QuadraticForm.lean index c691445111fa7..d073178a21c1d 100644 --- a/Counterexamples/QuadraticForm.lean +++ b/Counterexamples/QuadraticForm.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.LinearAlgebra.QuadraticForm.Basic -import Mathlib.Algebra.CharP.Two -import Mathlib.Data.ZMod.Basic /-! # `QuadraticForm R M` and `Subtype LinearMap.IsSymm` are distinct notions in characteristic 2 diff --git a/Mathlib.lean b/Mathlib.lean index 0fab6812c75fa..9412ac3b75fdf 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3095,6 +3095,7 @@ import Mathlib.Data.ZMod.IntUnitsPower import Mathlib.Data.ZMod.QuotientGroup import Mathlib.Data.ZMod.QuotientRing import Mathlib.Data.ZMod.Units +import Mathlib.Data.ZMod.ValMinAbs import Mathlib.Deprecated.AlgebraClasses import Mathlib.Deprecated.Aliases import Mathlib.Deprecated.ByteArray diff --git a/Mathlib/Algebra/Module/ZMod.lean b/Mathlib/Algebra/Module/ZMod.lean index 8b3f756361977..ff66bd7736466 100644 --- a/Mathlib/Algebra/Module/ZMod.lean +++ b/Mathlib/Algebra/Module/ZMod.lean @@ -3,8 +3,6 @@ Copyright (c) 2023 Lawrence Wu. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Lawrence Wu -/ -import Mathlib.Algebra.Module.Submodule.Lattice -import Mathlib.Data.ZMod.Basic import Mathlib.GroupTheory.Sylow /-! diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 89a1dbbbbec52..2bf65bae0c070 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -394,6 +394,9 @@ protected lemma add_le_mul {a : ℕ} (ha : 2 ≤ a) : ∀ {b : ℕ} (_ : 2 ≤ b /-! ### `div` -/ +lemma le_div_two_iff_mul_two_le {n m : ℕ} : m ≤ n / 2 ↔ (m : ℤ) * 2 ≤ n := by + rw [Nat.le_div_iff_mul_le Nat.zero_lt_two, ← Int.ofNat_le, Int.ofNat_mul]; rfl + attribute [simp] Nat.div_self lemma div_le_iff_le_mul_add_pred (hb : 0 < b) : a / b ≤ c ↔ a ≤ b * c + (b - 1) := by diff --git a/Mathlib/Data/Nat/Totient.lean b/Mathlib/Data/Nat/Totient.lean index 57626c563de64..41552a50cc4d4 100644 --- a/Mathlib/Data/Nat/Totient.lean +++ b/Mathlib/Data/Nat/Totient.lean @@ -7,8 +7,6 @@ import Mathlib.Algebra.CharP.Two import Mathlib.Data.Nat.Factorization.Basic import Mathlib.Data.Nat.Factorization.Induction import Mathlib.Data.Nat.Periodic -import Mathlib.Data.ZMod.Basic -import Mathlib.NumberTheory.Divisors /-! # Euler's totient function diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index ef6d34d2aedac..4bc9564fba5c8 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -6,11 +6,10 @@ Authors: Chris Hughes import Mathlib.Algebra.CharP.Basic import Mathlib.Algebra.Module.End import Mathlib.Algebra.Ring.Prod +import Mathlib.Data.Fintype.Units import Mathlib.GroupTheory.GroupAction.SubMulAction import Mathlib.GroupTheory.OrderOfElement import Mathlib.Tactic.FinCases -import Mathlib.Tactic.Linarith -import Mathlib.Data.Fintype.Units /-! # Integers mod `n` @@ -26,8 +25,6 @@ Definition of the integers mod n, and the field structure on the integers mod p. - for `a : ZMod 0` it is the absolute value of `a` - for `a : ZMod n` with `0 < n` it is the least natural number in the equivalence class -* `valMinAbs` returns the integer closest to zero in the equivalence class. - * A coercion `cast` is defined from `ZMod n` into any ring. This is a ring hom if the ring has characteristic dividing `n` @@ -1136,168 +1133,6 @@ theorem val_pow_le {m n : ℕ} [Fact (1 < n)] {a : ZMod n} : (a ^ m).val ≤ a.v apply le_trans (ZMod.val_mul_le _ _) apply Nat.mul_le_mul_right _ ih -/-- `valMinAbs x` returns the integer in the same equivalence class as `x` that is closest to `0`, - The result will be in the interval `(-n/2, n/2]`. -/ -def valMinAbs : ∀ {n : ℕ}, ZMod n → ℤ - | 0, x => x - | n@(_ + 1), x => if x.val ≤ n / 2 then x.val else (x.val : ℤ) - n - -@[simp] -theorem valMinAbs_def_zero (x : ZMod 0) : valMinAbs x = x := - rfl - -theorem valMinAbs_def_pos {n : ℕ} [NeZero n] (x : ZMod n) : - valMinAbs x = if x.val ≤ n / 2 then (x.val : ℤ) else x.val - n := by - cases n - · cases NeZero.ne 0 rfl - · rfl - -@[simp, norm_cast] -theorem coe_valMinAbs : ∀ {n : ℕ} (x : ZMod n), (x.valMinAbs : ZMod n) = x - | 0, _ => Int.cast_id - | k@(n + 1), x => by - rw [valMinAbs_def_pos] - split_ifs - · rw [Int.cast_natCast, natCast_zmod_val] - · rw [Int.cast_sub, Int.cast_natCast, natCast_zmod_val, Int.cast_natCast, natCast_self, - sub_zero] - -theorem injective_valMinAbs {n : ℕ} : (valMinAbs : ZMod n → ℤ).Injective := - Function.injective_iff_hasLeftInverse.2 ⟨_, coe_valMinAbs⟩ - -theorem _root_.Nat.le_div_two_iff_mul_two_le {n m : ℕ} : m ≤ n / 2 ↔ (m : ℤ) * 2 ≤ n := by - rw [Nat.le_div_iff_mul_le zero_lt_two, ← Int.ofNat_le, Int.ofNat_mul, Nat.cast_two] - -theorem valMinAbs_nonneg_iff {n : ℕ} [NeZero n] (x : ZMod n) : 0 ≤ x.valMinAbs ↔ x.val ≤ n / 2 := by - rw [valMinAbs_def_pos]; split_ifs with h - · exact iff_of_true (Nat.cast_nonneg _) h - · exact iff_of_false (sub_lt_zero.2 <| Int.ofNat_lt.2 x.val_lt).not_le h - -theorem valMinAbs_mul_two_eq_iff {n : ℕ} (a : ZMod n) : a.valMinAbs * 2 = n ↔ 2 * a.val = n := by - cases' n with n - · simp - by_cases h : a.val ≤ n.succ / 2 - · dsimp [valMinAbs] - rw [if_pos h, ← Int.natCast_inj, Nat.cast_mul, Nat.cast_two, mul_comm] - apply iff_of_false _ (mt _ h) - · intro he - rw [← a.valMinAbs_nonneg_iff, ← mul_nonneg_iff_left_nonneg_of_pos, he] at h - exacts [h (Nat.cast_nonneg _), zero_lt_two] - · rw [mul_comm] - exact fun h => (Nat.le_div_iff_mul_le zero_lt_two).2 h.le - -theorem valMinAbs_mem_Ioc {n : ℕ} [NeZero n] (x : ZMod n) : - x.valMinAbs * 2 ∈ Set.Ioc (-n : ℤ) n := by - simp_rw [valMinAbs_def_pos, Nat.le_div_two_iff_mul_two_le]; split_ifs with h - · refine ⟨(neg_lt_zero.2 <| mod_cast NeZero.pos n).trans_le (mul_nonneg ?_ ?_), h⟩ - exacts [Nat.cast_nonneg _, zero_le_two] - · refine ⟨?_, le_trans (mul_nonpos_of_nonpos_of_nonneg ?_ zero_le_two) <| Nat.cast_nonneg _⟩ - · linarith only [h] - · rw [sub_nonpos, Int.ofNat_le] - exact x.val_lt.le - -theorem valMinAbs_spec {n : ℕ} [NeZero n] (x : ZMod n) (y : ℤ) : - x.valMinAbs = y ↔ x = y ∧ y * 2 ∈ Set.Ioc (-n : ℤ) n := - ⟨by - rintro rfl - exact ⟨x.coe_valMinAbs.symm, x.valMinAbs_mem_Ioc⟩, fun h => - by - rw [← sub_eq_zero] - apply @Int.eq_zero_of_abs_lt_dvd n - · rw [← intCast_zmod_eq_zero_iff_dvd, Int.cast_sub, coe_valMinAbs, h.1, sub_self] - rw [← mul_lt_mul_right (@zero_lt_two ℤ _ _ _ _ _)] - nth_rw 1 [← abs_eq_self.2 (@zero_le_two ℤ _ _ _ _)] - rw [← abs_mul, sub_mul, abs_lt] - constructor <;> linarith only [x.valMinAbs_mem_Ioc.1, x.valMinAbs_mem_Ioc.2, h.2.1, h.2.2]⟩ - -theorem natAbs_valMinAbs_le {n : ℕ} [NeZero n] (x : ZMod n) : x.valMinAbs.natAbs ≤ n / 2 := by - rw [Nat.le_div_two_iff_mul_two_le] - cases' x.valMinAbs.natAbs_eq with h h - · rw [← h] - exact x.valMinAbs_mem_Ioc.2 - · rw [← neg_le_neg_iff, ← neg_mul, ← h] - exact x.valMinAbs_mem_Ioc.1.le - -@[simp] -theorem valMinAbs_zero : ∀ n, (0 : ZMod n).valMinAbs = 0 - | 0 => by simp only [valMinAbs_def_zero] - | n + 1 => by simp only [valMinAbs_def_pos, if_true, Int.ofNat_zero, zero_le, val_zero] - -@[simp] -theorem valMinAbs_eq_zero {n : ℕ} (x : ZMod n) : x.valMinAbs = 0 ↔ x = 0 := by - cases' n with n - · simp - rw [← valMinAbs_zero n.succ] - apply injective_valMinAbs.eq_iff - -theorem natCast_natAbs_valMinAbs {n : ℕ} [NeZero n] (a : ZMod n) : - (a.valMinAbs.natAbs : ZMod n) = if a.val ≤ (n : ℕ) / 2 then a else -a := by - have : (a.val : ℤ) - n ≤ 0 := by - rw [sub_nonpos, Int.ofNat_le] - exact a.val_le - rw [valMinAbs_def_pos] - split_ifs - · rw [Int.natAbs_ofNat, natCast_zmod_val] - · rw [← Int.cast_natCast, Int.ofNat_natAbs_of_nonpos this, Int.cast_neg, Int.cast_sub, - Int.cast_natCast, Int.cast_natCast, natCast_self, sub_zero, natCast_zmod_val] - -@[deprecated (since := "2024-04-17")] -alias nat_cast_natAbs_valMinAbs := natCast_natAbs_valMinAbs - -theorem valMinAbs_neg_of_ne_half {n : ℕ} {a : ZMod n} (ha : 2 * a.val ≠ n) : - (-a).valMinAbs = -a.valMinAbs := by - cases' eq_zero_or_neZero n with h h - · subst h - rfl - refine (valMinAbs_spec _ _).2 ⟨?_, ?_, ?_⟩ - · rw [Int.cast_neg, coe_valMinAbs] - · rw [neg_mul, neg_lt_neg_iff] - exact a.valMinAbs_mem_Ioc.2.lt_of_ne (mt a.valMinAbs_mul_two_eq_iff.1 ha) - · linarith only [a.valMinAbs_mem_Ioc.1] - -@[simp] -theorem natAbs_valMinAbs_neg {n : ℕ} (a : ZMod n) : (-a).valMinAbs.natAbs = a.valMinAbs.natAbs := by - by_cases h2a : 2 * a.val = n - · rw [a.neg_eq_self_iff.2 (Or.inr h2a)] - · rw [valMinAbs_neg_of_ne_half h2a, Int.natAbs_neg] - -theorem val_eq_ite_valMinAbs {n : ℕ} [NeZero n] (a : ZMod n) : - (a.val : ℤ) = a.valMinAbs + if a.val ≤ n / 2 then 0 else n := by - rw [valMinAbs_def_pos] - split_ifs <;> simp [add_zero, sub_add_cancel] - -theorem prime_ne_zero (p q : ℕ) [hp : Fact p.Prime] [hq : Fact q.Prime] (hpq : p ≠ q) : - (q : ZMod p) ≠ 0 := by - rwa [← Nat.cast_zero, Ne, eq_iff_modEq_nat, Nat.modEq_zero_iff_dvd, ← - hp.1.coprime_iff_not_dvd, Nat.coprime_primes hp.1 hq.1] - -variable {n a : ℕ} - -theorem valMinAbs_natAbs_eq_min {n : ℕ} [hpos : NeZero n] (a : ZMod n) : - a.valMinAbs.natAbs = min a.val (n - a.val) := by - rw [valMinAbs_def_pos] - have := a.val_lt - omega - -theorem valMinAbs_natCast_of_le_half (ha : a ≤ n / 2) : (a : ZMod n).valMinAbs = a := by - cases n - · simp - · simp [valMinAbs_def_pos, val_natCast, Nat.mod_eq_of_lt (ha.trans_lt <| Nat.div_lt_self' _ 0), - ha] - -theorem valMinAbs_natCast_of_half_lt (ha : n / 2 < a) (ha' : a < n) : - (a : ZMod n).valMinAbs = a - n := by - cases n - · cases not_lt_bot ha' - · simp [valMinAbs_def_pos, val_natCast, Nat.mod_eq_of_lt ha', ha.not_le] - --- Porting note: There was an extraneous `nat_` in the mathlib3 name -@[simp] -theorem valMinAbs_natCast_eq_self [NeZero n] : (a : ZMod n).valMinAbs = a ↔ a ≤ n / 2 := by - refine ⟨fun ha => ?_, valMinAbs_natCast_of_le_half⟩ - rw [← Int.natAbs_ofNat a, ← ha] - exact natAbs_valMinAbs_le a - theorem natAbs_min_of_le_div_two (n : ℕ) (x y : ℤ) (he : (x : ZMod n) = y) (hl : x.natAbs ≤ n / 2) : x.natAbs ≤ y.natAbs := by rw [intCast_eq_intCast_iff_dvd_sub] at he @@ -1313,14 +1148,6 @@ theorem natAbs_min_of_le_div_two (n : ℕ) (x y : ℤ) (he : (x : ZMod n) = y) ( refine le_trans ?_ (Nat.le_mul_of_pos_right _ <| Int.natAbs_pos.2 hm) rw [← mul_two]; apply Nat.div_mul_le_self -theorem natAbs_valMinAbs_add_le {n : ℕ} (a b : ZMod n) : - (a + b).valMinAbs.natAbs ≤ (a.valMinAbs + b.valMinAbs).natAbs := by - cases' n with n - · rfl - apply natAbs_min_of_le_div_two n.succ - · simp_rw [Int.cast_add, coe_valMinAbs] - · apply natAbs_valMinAbs_le - variable (p : ℕ) [Fact p.Prime] private theorem mul_inv_cancel_aux (a : ZMod p) (h : a ≠ 0) : a * a⁻¹ = 1 := by @@ -1572,5 +1399,3 @@ def Nat.residueClassesEquiv (N : ℕ) [NeZero N] : ℕ ≃ ZMod N × ℕ where cast_id', id_eq, zero_add] · simp only [add_comm p.1.val, mul_add_div (NeZero.pos _), (Nat.div_eq_zero_iff).2 <| .inr p.1.val_lt, add_zero] - -set_option linter.style.longFile 1700 diff --git a/Mathlib/Data/ZMod/Units.lean b/Mathlib/Data/ZMod/Units.lean index ff194571eb8a8..649c738a4f821 100644 --- a/Mathlib/Data/ZMod/Units.lean +++ b/Mathlib/Data/ZMod/Units.lean @@ -5,7 +5,6 @@ Authors: Moritz Firsching, Ashvni Narayanan, Michael Stoll -/ import Mathlib.Algebra.BigOperators.Associated import Mathlib.Data.ZMod.Basic -import Mathlib.Data.Nat.PrimeFin import Mathlib.RingTheory.Coprime.Lemmas /-! diff --git a/Mathlib/Data/ZMod/ValMinAbs.lean b/Mathlib/Data/ZMod/ValMinAbs.lean new file mode 100644 index 0000000000000..cfbf77ffae802 --- /dev/null +++ b/Mathlib/Data/ZMod/ValMinAbs.lean @@ -0,0 +1,176 @@ +/- +Copyright (c) 2019 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import Mathlib.Data.ZMod.Basic +import Mathlib.Tactic.Linarith + +/-! +# Absolute value in `ZMod n` +-/ + +namespace ZMod +variable {n : ℕ} {a : ZMod n} + +/-- Returns the integer in the same equivalence class as `x` that is closest to `0`. + +The result will be in the interval `(-n/2, n/2]`. -/ +def valMinAbs : ∀ {n : ℕ}, ZMod n → ℤ + | 0, x => x + | n@(_ + 1), x => if x.val ≤ n / 2 then x.val else (x.val : ℤ) - n + +@[simp] lemma valMinAbs_def_zero (x : ZMod 0) : valMinAbs x = x := rfl + +lemma valMinAbs_def_pos : ∀ {n : ℕ} [NeZero n] (x : ZMod n), + valMinAbs x = if x.val ≤ n / 2 then (x.val : ℤ) else x.val - n + | 0, _, x => by cases NeZero.ne 0 rfl + | n + 1, _, x => rfl + +@[simp, norm_cast] +lemma coe_valMinAbs : ∀ {n : ℕ} (x : ZMod n), (x.valMinAbs : ZMod n) = x + | 0, _ => Int.cast_id + | k@(n + 1), x => by + rw [valMinAbs_def_pos] + split_ifs + · rw [Int.cast_natCast, natCast_zmod_val] + · rw [Int.cast_sub, Int.cast_natCast, natCast_zmod_val, Int.cast_natCast, natCast_self, + sub_zero] + +lemma injective_valMinAbs : (valMinAbs : ZMod n → ℤ).Injective := + Function.injective_iff_hasLeftInverse.2 ⟨_, coe_valMinAbs⟩ + +lemma valMinAbs_nonneg_iff [NeZero n] (x : ZMod n) : 0 ≤ x.valMinAbs ↔ x.val ≤ n / 2 := by + rw [valMinAbs_def_pos]; split_ifs with h + · exact iff_of_true (Nat.cast_nonneg _) h + · exact iff_of_false (sub_lt_zero.2 <| Int.ofNat_lt.2 x.val_lt).not_le h + +lemma valMinAbs_mul_two_eq_iff (a : ZMod n) : a.valMinAbs * 2 = n ↔ 2 * a.val = n := by + cases' n with n + · simp + by_cases h : a.val ≤ n.succ / 2 + · dsimp [valMinAbs] + rw [if_pos h, ← Int.natCast_inj, Nat.cast_mul, Nat.cast_two, mul_comm] + apply iff_of_false _ (mt _ h) + · intro he + rw [← a.valMinAbs_nonneg_iff, ← mul_nonneg_iff_left_nonneg_of_pos, he] at h + exacts [h (Nat.cast_nonneg _), zero_lt_two] + · rw [mul_comm] + exact fun h => (Nat.le_div_iff_mul_le zero_lt_two).2 h.le + +lemma valMinAbs_mem_Ioc [NeZero n] (x : ZMod n) : x.valMinAbs * 2 ∈ Set.Ioc (-n : ℤ) n := by + simp_rw [valMinAbs_def_pos, Nat.le_div_two_iff_mul_two_le]; split_ifs with h + · refine ⟨(neg_lt_zero.2 <| mod_cast NeZero.pos n).trans_le (mul_nonneg ?_ ?_), h⟩ + exacts [Nat.cast_nonneg _, zero_le_two] + · refine ⟨?_, le_trans (mul_nonpos_of_nonpos_of_nonneg ?_ zero_le_two) <| Nat.cast_nonneg _⟩ + · linarith only [h] + · rw [sub_nonpos, Int.ofNat_le] + exact x.val_lt.le + +lemma valMinAbs_spec [NeZero n] (x : ZMod n) (y : ℤ) : + x.valMinAbs = y ↔ x = y ∧ y * 2 ∈ Set.Ioc (-n : ℤ) n where + mp := by rintro rfl; exact ⟨x.coe_valMinAbs.symm, x.valMinAbs_mem_Ioc⟩ + mpr h := by + rw [← sub_eq_zero] + apply @Int.eq_zero_of_abs_lt_dvd n + · rw [← intCast_zmod_eq_zero_iff_dvd, Int.cast_sub, coe_valMinAbs, h.1, sub_self] + rw [← mul_lt_mul_right (@zero_lt_two ℤ _ _ _ _ _)] + nth_rw 1 [← abs_eq_self.2 (@zero_le_two ℤ _ _ _ _)] + rw [← abs_mul, sub_mul, abs_lt] + constructor <;> linarith only [x.valMinAbs_mem_Ioc.1, x.valMinAbs_mem_Ioc.2, h.2.1, h.2.2] + +lemma natAbs_valMinAbs_le [NeZero n] (x : ZMod n) : x.valMinAbs.natAbs ≤ n / 2 := by + rw [Nat.le_div_two_iff_mul_two_le] + cases' x.valMinAbs.natAbs_eq with h h + · rw [← h] + exact x.valMinAbs_mem_Ioc.2 + · rw [← neg_le_neg_iff, ← neg_mul, ← h] + exact x.valMinAbs_mem_Ioc.1.le + +@[simp] +lemma valMinAbs_zero : ∀ n, (0 : ZMod n).valMinAbs = 0 + | 0 => by simp only [valMinAbs_def_zero] + | n + 1 => by simp only [valMinAbs_def_pos, if_true, Int.ofNat_zero, zero_le, val_zero] + +@[simp] +lemma valMinAbs_eq_zero (x : ZMod n) : x.valMinAbs = 0 ↔ x = 0 := by + cases' n with n + · simp + rw [← valMinAbs_zero n.succ] + apply injective_valMinAbs.eq_iff + +lemma natCast_natAbs_valMinAbs [NeZero n] (a : ZMod n) : + (a.valMinAbs.natAbs : ZMod n) = if a.val ≤ (n : ℕ) / 2 then a else -a := by + have : (a.val : ℤ) - n ≤ 0 := by + rw [sub_nonpos, Int.ofNat_le] + exact a.val_le + rw [valMinAbs_def_pos] + split_ifs + · rw [Int.natAbs_ofNat, natCast_zmod_val] + · rw [← Int.cast_natCast, Int.ofNat_natAbs_of_nonpos this, Int.cast_neg, Int.cast_sub, + Int.cast_natCast, Int.cast_natCast, natCast_self, sub_zero, natCast_zmod_val] + +@[deprecated (since := "2024-04-17")] +alias nat_cast_natAbs_valMinAbs := natCast_natAbs_valMinAbs + +lemma valMinAbs_neg_of_ne_half (ha : 2 * a.val ≠ n) : (-a).valMinAbs = -a.valMinAbs := by + cases' eq_zero_or_neZero n with h h + · subst h + rfl + refine (valMinAbs_spec _ _).2 ⟨?_, ?_, ?_⟩ + · rw [Int.cast_neg, coe_valMinAbs] + · rw [neg_mul, neg_lt_neg_iff] + exact a.valMinAbs_mem_Ioc.2.lt_of_ne (mt a.valMinAbs_mul_two_eq_iff.1 ha) + · linarith only [a.valMinAbs_mem_Ioc.1] + +@[simp] +lemma natAbs_valMinAbs_neg (a : ZMod n) : (-a).valMinAbs.natAbs = a.valMinAbs.natAbs := by + by_cases h2a : 2 * a.val = n + · rw [a.neg_eq_self_iff.2 (Or.inr h2a)] + · rw [valMinAbs_neg_of_ne_half h2a, Int.natAbs_neg] + +lemma val_eq_ite_valMinAbs [NeZero n] (a : ZMod n) : + (a.val : ℤ) = a.valMinAbs + if a.val ≤ n / 2 then 0 else n := by + rw [valMinAbs_def_pos] + split_ifs <;> simp [add_zero, sub_add_cancel] + +lemma prime_ne_zero (p q : ℕ) [hp : Fact p.Prime] [hq : Fact q.Prime] (hpq : p ≠ q) : + (q : ZMod p) ≠ 0 := by + rwa [← Nat.cast_zero, Ne, eq_iff_modEq_nat, Nat.modEq_zero_iff_dvd, ← + hp.1.coprime_iff_not_dvd, Nat.coprime_primes hp.1 hq.1] + +variable {n a : ℕ} + +lemma valMinAbs_natAbs_eq_min [hpos : NeZero n] (a : ZMod n) : + a.valMinAbs.natAbs = min a.val (n - a.val) := by + rw [valMinAbs_def_pos] + have := a.val_lt + omega + +lemma valMinAbs_natCast_of_le_half (ha : a ≤ n / 2) : (a : ZMod n).valMinAbs = a := by + cases n + · simp + · simp [valMinAbs_def_pos, val_natCast, Nat.mod_eq_of_lt (ha.trans_lt <| Nat.div_lt_self' _ 0), + ha] + +lemma valMinAbs_natCast_of_half_lt (ha : n / 2 < a) (ha' : a < n) : + (a : ZMod n).valMinAbs = a - n := by + cases n + · cases not_lt_bot ha' + · simp [valMinAbs_def_pos, val_natCast, Nat.mod_eq_of_lt ha', ha.not_le] + +@[simp] +lemma valMinAbs_natCast_eq_self [NeZero n] : (a : ZMod n).valMinAbs = a ↔ a ≤ n / 2 := by + refine ⟨fun ha => ?_, valMinAbs_natCast_of_le_half⟩ + rw [← Int.natAbs_ofNat a, ← ha] + exact natAbs_valMinAbs_le (n := n) a + +lemma natAbs_valMinAbs_add_le (a b : ZMod n) : + (a + b).valMinAbs.natAbs ≤ (a.valMinAbs + b.valMinAbs).natAbs := by + cases' n with n + · rfl + apply natAbs_min_of_le_div_two n.succ + · simp_rw [Int.cast_add, coe_valMinAbs] + · apply natAbs_valMinAbs_le + +end ZMod diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index 9f98c44c3ceac..caa9cfdf653ab 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -3,10 +3,10 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Joey van Langen, Casper Putz -/ +import Mathlib.Algebra.CharP.Reduced +import Mathlib.Data.ZMod.ValMinAbs import Mathlib.FieldTheory.Separable import Mathlib.RingTheory.IntegralDomain -import Mathlib.Algebra.CharP.Reduced -import Mathlib.Tactic.ApplyFun /-! # Finite fields diff --git a/Mathlib/GroupTheory/Coxeter/Length.lean b/Mathlib/GroupTheory/Coxeter/Length.lean index c987abb55aa89..bbbf38b3ed503 100644 --- a/Mathlib/GroupTheory/Coxeter/Length.lean +++ b/Mathlib/GroupTheory/Coxeter/Length.lean @@ -5,6 +5,7 @@ Authors: Mitchell Lee -/ import Mathlib.Data.ZMod.Basic import Mathlib.GroupTheory.Coxeter.Basic +import Mathlib.Tactic.Linarith import Mathlib.Tactic.Zify /-! diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index d000af6ee0f64..028fb46fc8ecb 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -3,15 +3,11 @@ 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 -/ -import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Data.Nat.Totient import Mathlib.Data.ZMod.Aut import Mathlib.Data.ZMod.QuotientGroup -import Mathlib.GroupTheory.OrderOfElement import Mathlib.GroupTheory.SpecificGroups.Dihedral import Mathlib.GroupTheory.Subgroup.Simple -import Mathlib.Tactic.Group -import Mathlib.GroupTheory.Exponent /-! # Cyclic groups diff --git a/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean b/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean index 318e8a450260d..8ba51e337fe2a 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean @@ -3,9 +3,6 @@ Copyright (c) 2021 Julian Kuelshammer. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Julian Kuelshammer -/ -import Mathlib.Data.ZMod.Basic -import Mathlib.Tactic.IntervalCases -import Mathlib.GroupTheory.SpecificGroups.Dihedral import Mathlib.GroupTheory.SpecificGroups.Cyclic /-! diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean index 65ef7c84f0abe..d81762c66cd8f 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.LinearAlgebra.CliffordAlgebra.Basic -import Mathlib.Data.ZMod.Basic import Mathlib.RingTheory.GradedAlgebra.Basic /-! diff --git a/Mathlib/NumberTheory/LucasLehmer.lean b/Mathlib/NumberTheory/LucasLehmer.lean index 8a59a2f806ddd..e457c2e2a106c 100644 --- a/Mathlib/NumberTheory/LucasLehmer.lean +++ b/Mathlib/NumberTheory/LucasLehmer.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kim Morrison, Ainsley Pahljina -/ import Mathlib.RingTheory.Fintype +import Mathlib.Tactic.NormNum +import Mathlib.Tactic.Zify /-! # The Lucas-Lehmer test for Mersenne primes. diff --git a/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean b/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean index cf81fe88e9ba2..db56d6f23d174 100644 --- a/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean +++ b/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean @@ -3,9 +3,6 @@ Copyright (c) 2022 Chris Birkbeck. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ -import Mathlib.Algebra.Group.Subgroup.Pointwise -import Mathlib.Data.ZMod.Basic -import Mathlib.GroupTheory.GroupAction.ConjAct import Mathlib.LinearAlgebra.Matrix.SpecialLinearGroup /-! diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index b69923c045a69..f6a1948f7d8c0 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -3,10 +3,10 @@ Copyright (c) 2023 Michael Stoll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Geißer, Michael Stoll -/ -import Mathlib.Tactic.Qify import Mathlib.Data.ZMod.Basic import Mathlib.NumberTheory.DiophantineApproximation.Basic import Mathlib.NumberTheory.Zsqrtd.Basic +import Mathlib.Tactic.Qify /-! # Pell's Equation diff --git a/Mathlib/NumberTheory/PythagoreanTriples.lean b/Mathlib/NumberTheory/PythagoreanTriples.lean index 7286c49b7dae7..b1155f6b4d739 100644 --- a/Mathlib/NumberTheory/PythagoreanTriples.lean +++ b/Mathlib/NumberTheory/PythagoreanTriples.lean @@ -3,13 +3,10 @@ Copyright (c) 2020 Paul van Wamelen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Paul van Wamelen -/ -import Mathlib.Algebra.Field.Basic -import Mathlib.Algebra.Order.Ring.Basic -import Mathlib.RingTheory.Int.Basic -import Mathlib.Tactic.Ring -import Mathlib.Tactic.FieldSimp import Mathlib.Data.Int.NatPrime import Mathlib.Data.ZMod.Basic +import Mathlib.RingTheory.Int.Basic +import Mathlib.Tactic.FieldSimp /-! # Pythagorean Triples diff --git a/Mathlib/NumberTheory/SumFourSquares.lean b/Mathlib/NumberTheory/SumFourSquares.lean index dccd30ae02ba5..2c2192d76a391 100644 --- a/Mathlib/NumberTheory/SumFourSquares.lean +++ b/Mathlib/NumberTheory/SumFourSquares.lean @@ -3,10 +3,6 @@ Copyright (c) 2019 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Ring.Int.Parity -import Mathlib.Algebra.Ring.Int.Units -import Mathlib.Data.Fintype.BigOperators -import Mathlib.Data.ZMod.Basic import Mathlib.FieldTheory.Finite.Basic /-! diff --git a/Mathlib/RingTheory/Fintype.lean b/Mathlib/RingTheory/Fintype.lean index e13486bc2013a..6d2ce8286520f 100644 --- a/Mathlib/RingTheory/Fintype.lean +++ b/Mathlib/RingTheory/Fintype.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.Data.Fintype.Units import Mathlib.Data.ZMod.Basic /-! From e736b5ace8446e2d7393256721044c464692b39c Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Sat, 1 Feb 2025 08:46:20 +0000 Subject: [PATCH 4/8] feat(Analysis/Normed/Module/Dual): polar in a normed space as a submodule (#20084) Definition of the polar of a subspace closed under scalar multiplication as a submodule. - Follows on from #14458 - Motivation #14369 Co-authored-by: Christopher Hoskin --- Mathlib/Analysis/Normed/Module/Dual.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/Analysis/Normed/Module/Dual.lean b/Mathlib/Analysis/Normed/Module/Dual.lean index 279e098d12e2c..ce73b8866528d 100644 --- a/Mathlib/Analysis/Normed/Module/Dual.lean +++ b/Mathlib/Analysis/Normed/Module/Dual.lean @@ -152,12 +152,32 @@ def polar (𝕜 : Type*) [NontriviallyNormedField 𝕜] {E : Type*} [SeminormedA [NormedSpace 𝕜 E] : Set E → Set (Dual 𝕜 E) := (dualPairing 𝕜 E).flip.polar +/-- Given a subset `s` in a normed space `E` (over a field `𝕜`) closed under scalar multiplication, + the polar `polarSubmodule 𝕜 s` is the submodule of `Dual 𝕜 E` consisting of those functionals which +evaluate to zero at all points `z ∈ s`. -/ +def polarSubmodule (𝕜 : Type*) [NontriviallyNormedField 𝕜] {E : Type*} [SeminormedAddCommGroup E] + [NormedSpace 𝕜 E] {S : Type*} [SetLike S E] [SMulMemClass S 𝕜 E] (m : S) : + Submodule 𝕜 (Dual 𝕜 E) := (dualPairing 𝕜 E).flip.polarSubmodule m + variable (𝕜 : Type*) [NontriviallyNormedField 𝕜] variable {E : Type*} [SeminormedAddCommGroup E] [NormedSpace 𝕜 E] +lemma polarSubmodule_eq_polar (m : SubMulAction 𝕜 E) : + (polarSubmodule 𝕜 m : Set (Dual 𝕜 E)) = polar 𝕜 m := rfl + theorem mem_polar_iff {x' : Dual 𝕜 E} (s : Set E) : x' ∈ polar 𝕜 s ↔ ∀ z ∈ s, ‖x' z‖ ≤ 1 := Iff.rfl +lemma polarSubmodule_eq_setOf {S : Type*} [SetLike S E] [SMulMemClass S 𝕜 E] (m : S) : + polarSubmodule 𝕜 m = { y : Dual 𝕜 E | ∀ x ∈ m, y x = 0 } := + (dualPairing 𝕜 E).flip.polar_subMulAction _ + +lemma mem_polarSubmodule {S : Type*} [SetLike S E] [SMulMemClass S 𝕜 E] (m : S) (y : Dual 𝕜 E) : + y ∈ polarSubmodule 𝕜 m ↔ ∀ x ∈ m, y x = 0 := by + have := polarSubmodule_eq_setOf 𝕜 m + apply_fun (y ∈ ·) at this + rwa [propext_iff] at this + @[simp] theorem zero_mem_polar (s : Set E) : (0 : Dual 𝕜 E) ∈ polar 𝕜 s := LinearMap.zero_mem_polar _ s From fb51fb6989d720b67072c59191b075498a20cc6a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 1 Feb 2025 09:03:36 +0000 Subject: [PATCH 5/8] chore: delete declarations deprecated between 2024-01 and 2024-07 (#21271) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The first commit was done by replacing all matches of `@\[deprecated.*\(since := "2024-0[1-7]-.."\)\][\s\n]?(protected )?alias [_a-zA-Zα-ωΑ-Ωϊ-ϻἀ-῾℀-⅏𝒜-𝖟'.₀₁₂₃₄₅₆₇₈₉]* :=[\s\n]+[_a-zA-Zα-ωΑ-Ωϊ-ϻἀ-῾℀-⅏𝒜-𝖟'.₀₁₂₃₄₅₆₇₈₉]*\n\n?` with nothing. The second commit was done manually. --- Mathlib.lean | 5 - Mathlib/Algebra/Algebra/Equiv.lean | 15 - Mathlib/Algebra/Algebra/Hom.lean | 49 --- Mathlib/Algebra/Algebra/Quasispectrum.lean | 3 +- Mathlib/Algebra/Algebra/Rat.lean | 3 - Mathlib/Algebra/Algebra/Spectrum.lean | 8 +- Mathlib/Algebra/Algebra/Subalgebra/Basic.lean | 3 - .../BigOperators/Group/Finset/Basic.lean | 3 - .../BigOperators/Group/Finset/Defs.lean | 3 - .../BigOperators/Group/List/Basic.lean | 2 - .../BigOperators/Group/List/Lemmas.lean | 10 - Mathlib/Algebra/BigOperators/Ring.lean | 2 - Mathlib/Algebra/CharP/ExpChar.lean | 3 - Mathlib/Algebra/CharZero/Lemmas.lean | 10 - Mathlib/Algebra/DirectSum/Internal.lean | 6 - Mathlib/Algebra/Field/Defs.lean | 4 - Mathlib/Algebra/Field/Subfield/Defs.lean | 6 - Mathlib/Algebra/Group/AddChar.lean | 18 - Mathlib/Algebra/Group/Basic.lean | 3 - Mathlib/Algebra/Group/Even.lean | 4 - Mathlib/Algebra/Group/Hom/Instances.lean | 3 - .../Group/Subgroup/ZPowers/Lemmas.lean | 6 - Mathlib/Algebra/Group/Units/Basic.lean | 9 - Mathlib/Algebra/GroupWithZero/Associated.lean | 3 - Mathlib/Algebra/GroupWithZero/Center.lean | 8 - .../Algebra/GroupWithZero/Units/Basic.lean | 3 - .../Algebra/Homology/ShortComplex/Exact.lean | 9 - Mathlib/Algebra/Lie/Weights/Linear.lean | 3 - Mathlib/Algebra/Module/Basic.lean | 21 -- Mathlib/Algebra/Module/End.lean | 6 - Mathlib/Algebra/Module/NatInt.lean | 7 - Mathlib/Algebra/Module/Rat.lean | 6 - Mathlib/Algebra/MonoidAlgebra/Defs.lean | 12 - Mathlib/Algebra/Order/AddTorsor.lean | 6 - Mathlib/Algebra/Order/Floor.lean | 2 - .../Algebra/Order/Group/Unbundled/Int.lean | 2 - .../Order/GroupWithZero/Unbundled.lean | 15 - Mathlib/Algebra/Order/Interval/Set/Group.lean | 18 - .../Algebra/Order/Monoid/Canonical/Defs.lean | 3 - .../Algebra/Order/Monoid/Unbundled/Basic.lean | 3 - .../Order/Monoid/Unbundled/WithTop.lean | 8 - Mathlib/Algebra/Order/Nonneg/Basic.lean | 6 - Mathlib/Algebra/Order/Rearrangement.lean | 18 - Mathlib/Algebra/Order/ToIntervalMod.lean | 18 - Mathlib/Algebra/Polynomial/AlgebraMap.lean | 9 - Mathlib/Algebra/Polynomial/Basic.lean | 12 - Mathlib/Algebra/Polynomial/Coeff.lean | 12 - .../Polynomial/Degree/Definitions.lean | 12 - .../Polynomial/Degree/TrailingDegree.lean | 9 - Mathlib/Algebra/Polynomial/Derivative.lean | 21 -- Mathlib/Algebra/Polynomial/Div.lean | 3 - Mathlib/Algebra/Polynomial/Eval/Coeff.lean | 6 - Mathlib/Algebra/Polynomial/Eval/Defs.lean | 38 -- Mathlib/Algebra/Polynomial/Smeval.lean | 6 - Mathlib/Algebra/Quaternion.lean | 93 ----- Mathlib/Algebra/Ring/CentroidHom.lean | 18 - Mathlib/Algebra/Ring/Parity.lean | 2 - Mathlib/Algebra/Ring/Rat.lean | 2 - Mathlib/Algebra/Ring/Subring/Defs.lean | 2 - Mathlib/Algebra/Ring/Subsemiring/Defs.lean | 2 - Mathlib/Algebra/Squarefree/Basic.lean | 2 - Mathlib/Algebra/Star/Module.lean | 15 - Mathlib/Algebra/TrivSqZeroExt.lean | 18 - Mathlib/Algebra/Vertex/HVertexOperator.lean | 11 - Mathlib/AlgebraicGeometry/AffineScheme.lean | 34 -- Mathlib/AlgebraicGeometry/Cover/Open.lean | 1 - .../EllipticCurve/Affine.lean | 30 -- .../GammaSpecAdjunction.lean | 13 - .../AlgebraicGeometry/Morphisms/Basic.lean | 5 - .../Morphisms/Constructors.lean | 14 - .../ProjectiveSpectrum/StructureSheaf.lean | 4 - Mathlib/AlgebraicGeometry/Restrict.lean | 3 - Mathlib/AlgebraicGeometry/Scheme.lean | 3 - Mathlib/AlgebraicGeometry/StructureSheaf.lean | 4 - Mathlib/Analysis/Asymptotics/Lemmas.lean | 6 - Mathlib/Analysis/CStarAlgebra/Multiplier.lean | 18 - Mathlib/Analysis/Calculus/FDeriv/Add.lean | 30 -- Mathlib/Analysis/Calculus/FDeriv/Extend.lean | 9 - .../Analysis/Calculus/LocalExtr/Basic.lean | 3 - Mathlib/Analysis/Complex/Basic.lean | 3 - Mathlib/Analysis/Complex/Circle.lean | 31 -- Mathlib/Analysis/Convex/Integral.lean | 4 - .../FunctionalSpaces/SobolevInequality.lean | 15 - Mathlib/Analysis/Normed/Group/AddCircle.lean | 3 - Mathlib/Analysis/Normed/Group/Basic.lean | 3 - Mathlib/Analysis/Normed/Group/Int.lean | 2 - Mathlib/Analysis/Normed/Lp/lpSpace.lean | 12 - Mathlib/Analysis/NormedSpace/Int.lean | 2 - Mathlib/Analysis/PSeries.lean | 6 - .../SpecialFunctions/Complex/Circle.lean | 27 -- Mathlib/Analysis/SpecialFunctions/Exp.lean | 8 - .../Analysis/SpecialFunctions/Log/Base.lean | 6 - .../Analysis/SpecialFunctions/Log/Basic.lean | 9 - .../Analysis/SpecialFunctions/PolarCoord.lean | 2 - .../SpecialFunctions/Pow/Complex.lean | 6 - .../Analysis/SpecialFunctions/Pow/NNReal.lean | 43 --- .../Analysis/SpecialFunctions/Pow/Real.lean | 6 - .../SpecialFunctions/Trigonometric/Angle.lean | 3 - Mathlib/Analysis/SpecificLimits/Normed.lean | 6 - Mathlib/CategoryTheory/Abelian/Exact.lean | 8 - .../Adjunction/FullyFaithful.lean | 3 - Mathlib/CategoryTheory/Adjunction/Mates.lean | 4 - Mathlib/CategoryTheory/Adjunction/Over.lean | 13 - Mathlib/CategoryTheory/DiscreteCategory.lean | 6 - Mathlib/CategoryTheory/Equivalence.lean | 15 - Mathlib/CategoryTheory/EssentialImage.lean | 3 - Mathlib/CategoryTheory/Functor/Category.lean | 5 - .../CategoryTheory/Functor/FullyFaithful.lean | 30 -- .../CategoryTheory/Functor/ReflectsIso.lean | 2 - .../CategoryTheory/Idempotents/Karoubi.lean | 3 - Mathlib/CategoryTheory/Iso.lean | 3 - .../MorphismProperty/Basic.lean | 8 - .../Sites/DenseSubsite/InducedTopology.lean | 3 - .../Sites/DenseSubsite/SheafEquiv.lean | 6 - Mathlib/CategoryTheory/Sites/Sieves.lean | 2 - .../Enumerative/Composition.lean | 21 -- Mathlib/Control/LawfulFix.lean | 47 +-- Mathlib/Data/Bool/Basic.lean | 25 -- Mathlib/Data/Complex/Basic.lean | 30 -- Mathlib/Data/ENNReal/Basic.lean | 7 - Mathlib/Data/ENNReal/Inv.lean | 3 - Mathlib/Data/ENNReal/Operations.lean | 3 - Mathlib/Data/ENNReal/Real.lean | 20 -- Mathlib/Data/Fin/Basic.lean | 17 - Mathlib/Data/Fin/Tuple/Basic.lean | 7 - Mathlib/Data/Finset/Card.lean | 16 - Mathlib/Data/Finset/NoncommProd.lean | 10 - Mathlib/Data/Finset/SDiff.lean | 3 - Mathlib/Data/Finsupp/Weight.lean | 15 - Mathlib/Data/Int/AbsoluteValue.lean | 3 - Mathlib/Data/Int/Cast/Lemmas.lean | 20 -- Mathlib/Data/Int/Cast/Pi.lean | 3 - Mathlib/Data/Int/CharZero.lean | 6 - Mathlib/Data/Int/Defs.lean | 23 -- Mathlib/Data/Int/GCD.lean | 2 - Mathlib/Data/Int/Lemmas.lean | 9 - Mathlib/Data/Int/ModEq.lean | 2 - Mathlib/Data/Int/Order/Basic.lean | 3 - Mathlib/Data/Int/SuccPred.lean | 3 - Mathlib/Data/List/Basic.lean | 75 ---- Mathlib/Data/List/Defs.lean | 4 - Mathlib/Data/List/Enum.lean | 4 - Mathlib/Data/List/Flatten.lean | 12 - Mathlib/Data/List/GetD.lean | 6 - Mathlib/Data/List/InsertIdx.lean | 3 - Mathlib/Data/List/Lex.lean | 11 +- Mathlib/Data/List/MinMax.lean | 6 - Mathlib/Data/List/OfFn.lean | 4 - Mathlib/Data/List/Permutation.lean | 6 - Mathlib/Data/List/ToFinsupp.lean | 11 - Mathlib/Data/List/Zip.lean | 20 -- Mathlib/Data/Multiset/Basic.lean | 2 - Mathlib/Data/NNRat/Defs.lean | 2 - Mathlib/Data/NNReal/Defs.lean | 26 -- Mathlib/Data/Nat/Cast/Basic.lean | 3 - Mathlib/Data/Nat/Cast/Commute.lean | 8 - Mathlib/Data/Nat/Choose/Multinomial.lean | 2 - Mathlib/Data/Nat/Defs.lean | 22 -- Mathlib/Data/Nat/Factorization/Defs.lean | 4 - Mathlib/Data/Nat/Factors.lean | 2 - Mathlib/Data/Nat/PrimeFin.lean | 3 - Mathlib/Data/Nat/Squarefree.lean | 6 - Mathlib/Data/Num/Lemmas.lean | 9 - Mathlib/Data/Part.lean | 2 - Mathlib/Data/Prod/Basic.lean | 11 - Mathlib/Data/Rat/Cast/CharZero.lean | 2 - Mathlib/Data/Rat/Defs.lean | 10 - Mathlib/Data/Rat/Floor.lean | 2 - Mathlib/Data/Rat/Lemmas.lean | 15 - Mathlib/Data/Real/Irrational.lean | 6 - Mathlib/Data/Real/Sign.lean | 3 - Mathlib/Data/Seq/Seq.lean | 6 - Mathlib/Data/Set/Card.lean | 16 - Mathlib/Data/Set/Image.lean | 4 - Mathlib/Data/Set/Lattice.lean | 3 - Mathlib/Data/Set/List.lean | 1 - Mathlib/Data/Sigma/Basic.lean | 7 - Mathlib/Data/String/Basic.lean | 6 - Mathlib/Data/Subtype.lean | 2 - Mathlib/Data/UInt.lean | 6 - Mathlib/Data/Vector/Basic.lean | 8 - Mathlib/Data/Vector/Defs.lean | 2 - Mathlib/Data/ZMod/Basic.lean | 98 ------ Mathlib/Data/ZMod/ValMinAbs.lean | 3 - Mathlib/Deprecated/AlgebraClasses.lean | 174 +-------- Mathlib/Deprecated/Aliases.lean | 6 - Mathlib/Deprecated/Cardinal/Finite.lean | 2 - Mathlib/Deprecated/Combinator.lean | 25 -- Mathlib/Deprecated/HashMap.lean | 58 --- Mathlib/Deprecated/LazyList.lean | 112 ------ Mathlib/Deprecated/Logic.lean | 2 - Mathlib/Deprecated/NatLemmas.lean | 27 -- Mathlib/Deprecated/RelClasses.lean | 40 --- Mathlib/Dynamics/Ergodic/Ergodic.lean | 2 - .../IntermediateField/Adjoin/Defs.lean | 3 - .../FieldTheory/IntermediateField/Basic.lean | 3 - Mathlib/FieldTheory/IsAlgClosed/Basic.lean | 10 - Mathlib/FieldTheory/PerfectClosure.lean | 9 - Mathlib/Geometry/Manifold/Instances/Real.lean | 2 - .../Geometry/Manifold/Instances/Sphere.lean | 2 - Mathlib/Geometry/Manifold/MFDeriv/Defs.lean | 6 - Mathlib/Geometry/RingedSpace/Stalks.lean | 3 - Mathlib/GroupTheory/Exponent.lean | 13 - Mathlib/GroupTheory/Index.lean | 10 - Mathlib/GroupTheory/Nilpotent.lean | 2 - Mathlib/GroupTheory/OrderOfElement.lean | 4 - Mathlib/GroupTheory/Perm/Cycle/Factors.lean | 5 - Mathlib/Lean/Expr/Basic.lean | 2 - Mathlib/LinearAlgebra/BilinearForm/Basic.lean | 20 -- Mathlib/LinearAlgebra/BilinearForm/Hom.lean | 67 ---- .../BilinearForm/Orthogonal.lean | 10 - .../LinearAlgebra/FiniteDimensional/Defs.lean | 5 - Mathlib/LinearAlgebra/Matrix/ZPow.lean | 2 - Mathlib/LinearAlgebra/Prod.lean | 128 ------- Mathlib/Logic/Basic.lean | 32 -- Mathlib/Logic/Equiv/Defs.lean | 10 - Mathlib/Logic/Equiv/Fin.lean | 20 -- .../Decomposition/RadonNikodym.lean | 27 -- .../Function/AEEqOfIntegral.lean | 82 ----- .../Function/AEEqOfLIntegral.lean | 19 - .../ConditionalExpectation/Basic.lean | 8 - .../ConditionalExpectation/CondexpL1.lean | 22 -- .../ConditionalExpectation/CondexpL2.lean | 29 -- .../Function/ConditionalExpectation/Real.lean | 11 - .../ConditionalExpectation/Unique.lean | 15 - .../Function/ContinuousMapDense.lean | 9 - .../Function/ConvergenceInMeasure.lean | 13 - Mathlib/MeasureTheory/Function/L2Space.lean | 17 - .../Function/LpSeminorm/Basic.lean | 332 ------------------ .../Function/LpSeminorm/ChebyshevMarkov.lean | 12 - .../Function/LpSeminorm/CompareExp.lean | 49 --- .../LpSeminorm/TriangleInequality.lean | 30 -- .../Function/LpSeminorm/Trim.lean | 12 - Mathlib/MeasureTheory/Function/LpSpace.lean | 77 +--- .../Function/SimpleFuncDenseLp.lean | 13 - .../Function/UniformIntegrable.lean | 24 -- .../Group/FundamentalDomain.lean | 18 - Mathlib/MeasureTheory/Integral/Bochner.lean | 29 -- .../Integral/DominatedConvergence.lean | 3 - .../MeasureTheory/Integral/IntegrableOn.lean | 3 - .../Integral/IntegralEqImproper.lean | 10 - .../Integral/IntervalIntegral.lean | 3 - Mathlib/MeasureTheory/Integral/Lebesgue.lean | 87 ----- .../Integral/MeanInequalities.lean | 7 - Mathlib/MeasureTheory/Integral/Prod.lean | 4 - .../MeasureTheory/Integral/SetIntegral.lean | 187 ---------- .../Measure/Haar/NormedSpace.lean | 6 - .../Measure/Lebesgue/VolumeOfBalls.lean | 5 - .../MeasureTheory/Measure/MeasureSpace.lean | 2 - Mathlib/MeasureTheory/Measure/Tilted.lean | 12 - .../MeasureTheory/Measure/Typeclasses.lean | 3 - .../MeasureTheory/Measure/WithDensity.lean | 33 -- Mathlib/MeasureTheory/OuterMeasure/Basic.lean | 73 ---- .../OuterMeasure/Caratheodory.lean | 3 - Mathlib/NumberTheory/ArithmeticFunction.lean | 10 - .../Cyclotomic/PrimitiveRoots.lean | 28 -- .../DirichletCharacter/Basic.lean | 4 - Mathlib/NumberTheory/Divisors.lean | 6 - Mathlib/NumberTheory/LSeries/RiemannZeta.lean | 15 - .../LegendreSymbol/AddCharacter.lean | 2 - .../LegendreSymbol/QuadraticChar/Basic.lean | 5 - Mathlib/NumberTheory/LucasLehmer.lean | 9 - Mathlib/NumberTheory/MulChar/Basic.lean | 27 -- .../NumberTheory/Padics/PadicIntegers.lean | 11 - .../NumberTheory/Padics/PadicVal/Defs.lean | 8 - Mathlib/NumberTheory/Padics/RingHoms.lean | 6 - Mathlib/NumberTheory/Zsqrtd/Basic.lean | 22 -- Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean | 11 - Mathlib/Order/Atoms.lean | 3 - Mathlib/Order/CompleteLattice.lean | 3 - .../ConditionallyCompleteLattice/Basic.lean | 2 - Mathlib/Order/Cover.lean | 3 - Mathlib/Order/Defs/PartialOrder.lean | 4 - .../Order/Filter/AtTopBot/Archimedean.lean | 42 --- Mathlib/Order/Filter/AtTopBot/Field.lean | 12 - Mathlib/Order/Filter/Germ/Basic.lean | 3 - Mathlib/Order/Fin/Basic.lean | 8 - Mathlib/Order/Fin/Tuple.lean | 8 - Mathlib/Order/Lattice.lean | 8 - Mathlib/Order/LiminfLimsup.lean | 18 - Mathlib/Order/OmegaCompletePartialOrder.lean | 190 ---------- Mathlib/Order/Part.lean | 2 - Mathlib/Order/RelClasses.lean | 17 - Mathlib/Order/WithBot.lean | 10 + Mathlib/Probability/Density.lean | 6 - Mathlib/Probability/IdentDistrib.lean | 3 - Mathlib/Probability/Kernel/Basic.lean | 15 - .../Probability/Kernel/Composition/Basic.lean | 9 - .../Kernel/Composition/IntegralCompProd.lean | 9 - .../Kernel/Composition/MeasureCompProd.lean | 6 - Mathlib/Probability/Kernel/CondDistrib.lean | 6 - Mathlib/Probability/Kernel/Defs.lean | 2 - .../Kernel/Disintegration/CDFToKernel.lean | 28 -- .../Kernel/Disintegration/CondCDF.lean | 12 - .../Kernel/Disintegration/Density.lean | 35 -- .../Kernel/Disintegration/Integral.lean | 42 --- .../Kernel/Disintegration/StandardBorel.lean | 44 --- Mathlib/Probability/Kernel/Integral.lean | 15 - .../Kernel/MeasurableLIntegral.lean | 11 - Mathlib/Probability/Kernel/RadonNikodym.lean | 3 - Mathlib/Probability/Martingale/Basic.lean | 21 -- .../Probability/Martingale/BorelCantelli.lean | 6 - .../Probability/Martingale/Convergence.lean | 20 -- Mathlib/Probability/Process/Filtration.lean | 3 - Mathlib/RepresentationTheory/FDRep.lean | 3 - Mathlib/RingTheory/Binomial.lean | 3 - Mathlib/RingTheory/Congruence/Defs.lean | 6 - .../DedekindDomain/AdicValuation.lean | 36 -- Mathlib/RingTheory/Derivation/Basic.lean | 3 - Mathlib/RingTheory/FiniteType.lean | 17 - Mathlib/RingTheory/FractionalIdeal/Basic.lean | 3 - Mathlib/RingTheory/Ideal/Operations.lean | 3 - Mathlib/RingTheory/Int/Basic.lean | 8 - Mathlib/RingTheory/IntegralDomain.lean | 3 - .../RingTheory/Localization/Away/Basic.lean | 4 - Mathlib/RingTheory/Localization/Basic.lean | 6 - Mathlib/RingTheory/Localization/NumDen.lean | 4 - Mathlib/RingTheory/Multiplicity.lean | 2 - Mathlib/RingTheory/Noetherian/Basic.lean | 4 +- Mathlib/RingTheory/Polynomial/Basic.lean | 12 - Mathlib/RingTheory/Spectrum/Prime/Defs.lean | 2 - .../UniqueFactorizationDomain/FactorSet.lean | 2 - Mathlib/RingTheory/WittVector/Basic.lean | 12 - Mathlib/RingTheory/WittVector/Truncated.lean | 6 - Mathlib/SetTheory/Cardinal/UnivLE.lean | 2 +- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 44 --- Mathlib/SetTheory/Ordinal/Basic.lean | 3 - Mathlib/Tactic/ComputeDegree.lean | 9 - Mathlib/Tactic/Linarith/Lemmas.lean | 3 - Mathlib/Tactic/NormNum/Basic.lean | 9 - Mathlib/Tactic/Qify.lean | 3 - Mathlib/Tactic/Rify.lean | 3 - Mathlib/Tactic/Zify.lean | 3 - .../Algebra/Module/Multilinear/Topology.lean | 1 - Mathlib/Topology/Algebra/Support.lean | 3 - Mathlib/Topology/Compactness/Compact.lean | 5 - Mathlib/Topology/ContinuousMap/Algebra.lean | 12 - Mathlib/Topology/Defs/Sequences.lean | 2 - Mathlib/Topology/EMetricSpace/Lipschitz.lean | 2 - Mathlib/Topology/ExtremallyDisconnected.lean | 4 - Mathlib/Topology/List.lean | 8 - Mathlib/Topology/LocalAtTarget.lean | 8 - Mathlib/Topology/Maps/Proper/Basic.lean | 2 - Mathlib/Topology/Order/Bounded.lean | 40 --- Mathlib/Topology/Sheaves/Stalks.lean | 2 - Mathlib/Topology/UniformSpace/Basic.lean | 2 - Mathlib/Topology/UniformSpace/Cauchy.lean | 3 - .../UniformConvergenceTopology.lean | 10 - 348 files changed, 30 insertions(+), 5088 deletions(-) delete mode 100644 Mathlib/Deprecated/Combinator.lean delete mode 100644 Mathlib/Deprecated/HashMap.lean delete mode 100644 Mathlib/Deprecated/LazyList.lean delete mode 100644 Mathlib/Deprecated/RelClasses.lean delete mode 100644 Mathlib/Topology/Order/Bounded.lean diff --git a/Mathlib.lean b/Mathlib.lean index 9412ac3b75fdf..aa7d8e6665449 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3102,15 +3102,11 @@ import Mathlib.Deprecated.ByteArray import Mathlib.Deprecated.Cardinal.Continuum import Mathlib.Deprecated.Cardinal.Finite import Mathlib.Deprecated.Cardinal.PartENat -import Mathlib.Deprecated.Combinator import Mathlib.Deprecated.Equiv -import Mathlib.Deprecated.HashMap -import Mathlib.Deprecated.LazyList import Mathlib.Deprecated.Logic import Mathlib.Deprecated.MinMax import Mathlib.Deprecated.NatLemmas import Mathlib.Deprecated.Order -import Mathlib.Deprecated.RelClasses import Mathlib.Dynamics.BirkhoffSum.Average import Mathlib.Dynamics.BirkhoffSum.Basic import Mathlib.Dynamics.BirkhoffSum.NormedSpace @@ -5562,7 +5558,6 @@ import Mathlib.Topology.OmegaCompletePartialOrder import Mathlib.Topology.Order import Mathlib.Topology.Order.Basic import Mathlib.Topology.Order.Bornology -import Mathlib.Topology.Order.Bounded import Mathlib.Topology.Order.Category.AlexDisc import Mathlib.Topology.Order.Category.FrameAdjunction import Mathlib.Topology.Order.Compact diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index 9a0fa07bd87e7..d50b2aa6bba88 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -729,21 +729,6 @@ instance _root_.Finite.algEquiv [Finite (A₁ →ₐ[R] A₂)] : Finite (A₁ end Semiring -section Ring - -variable [CommSemiring R] [Ring A₁] [Ring A₂] -variable [Algebra R A₁] [Algebra R A₂] (e : A₁ ≃ₐ[R] A₂) - -@[deprecated map_neg (since := "2024-06-20")] -protected theorem map_neg (x) : e (-x) = -e x := - map_neg e x - -@[deprecated map_sub (since := "2024-06-20")] -protected theorem map_sub (x y) : e (x - y) = e x - e y := - map_sub e x y - -end Ring - end AlgEquiv namespace MulSemiringAction diff --git a/Mathlib/Algebra/Algebra/Hom.lean b/Mathlib/Algebra/Algebra/Hom.lean index 355643e02267e..0fb346aae4195 100644 --- a/Mathlib/Algebra/Algebra/Hom.lean +++ b/Mathlib/Algebra/Algebra/Hom.lean @@ -194,35 +194,6 @@ theorem commutes (r : R) : φ (algebraMap R A r) = algebraMap R B r := theorem comp_algebraMap : (φ : A →+* B).comp (algebraMap R A) = algebraMap R B := RingHom.ext <| φ.commutes -@[deprecated map_add (since := "2024-06-26")] -protected theorem map_add (r s : A) : φ (r + s) = φ r + φ s := - map_add _ _ _ - -@[deprecated map_zero (since := "2024-06-26")] -protected theorem map_zero : φ 0 = 0 := - map_zero _ - -@[deprecated map_mul (since := "2024-06-26")] -protected theorem map_mul (x y) : φ (x * y) = φ x * φ y := - map_mul _ _ _ - -@[deprecated map_one (since := "2024-06-26")] -protected theorem map_one : φ 1 = 1 := - map_one _ - -@[deprecated map_pow (since := "2024-06-26")] -protected theorem map_pow (x : A) (n : ℕ) : φ (x ^ n) = φ x ^ n := - map_pow _ _ _ - -@[deprecated map_smul (since := "2024-06-26")] -protected theorem map_smul (r : R) (x : A) : φ (r • x) = r • φ x := - map_smul _ _ _ - -@[deprecated map_sum (since := "2024-06-26")] -protected theorem map_sum {ι : Type*} (f : ι → A) (s : Finset ι) : - φ (∑ x ∈ s, f x) = ∑ x ∈ s, φ (f x) := - map_sum _ _ _ - /-- If a `RingHom` is `R`-linear, then it is an `AlgHom`. -/ def mk' (f : A →+* B) (h : ∀ (c : R) (x), f (c • x) = c • f x) : A →ₐ[R] B := { f with @@ -334,10 +305,6 @@ theorem map_smul_of_tower {R'} [SMul R' A] [SMul R' B] [LinearMap.CompatibleSMul (x : A) : φ (r • x) = r • φ x := φ.toLinearMap.map_smul_of_tower r x -@[deprecated map_list_prod (since := "2024-06-26")] -protected theorem map_list_prod (s : List A) : φ s.prod = (s.map φ).prod := - map_list_prod φ s - @[simps (config := .lemmasOnly) toSemigroup_toMul_mul toOne_one] instance End : Monoid (A →ₐ[R] A) where mul := comp @@ -359,22 +326,6 @@ theorem algebraMap_eq_apply (f : A →ₐ[R] B) {y : R} {x : A} (h : algebraMap h ▸ (f.commutes _).symm end Semiring - -section Ring - -variable [CommSemiring R] [Ring A] [Ring B] -variable [Algebra R A] [Algebra R B] (φ : A →ₐ[R] B) - -@[deprecated map_neg (since := "2024-06-26")] -protected theorem map_neg (x) : φ (-x) = -φ x := - map_neg _ _ - -@[deprecated map_sub (since := "2024-06-26")] -protected theorem map_sub (x y) : φ (x - y) = φ x - φ y := - map_sub _ _ _ - -end Ring - end AlgHom namespace RingHom diff --git a/Mathlib/Algebra/Algebra/Quasispectrum.lean b/Mathlib/Algebra/Algebra/Quasispectrum.lean index fa1469c46f1a2..41b77543522ae 100644 --- a/Mathlib/Algebra/Algebra/Quasispectrum.lean +++ b/Mathlib/Algebra/Algebra/Quasispectrum.lean @@ -3,8 +3,9 @@ Copyright (c) 2024 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Algebra.Algebra.Unitization import Mathlib.Algebra.Algebra.Spectrum +import Mathlib.Algebra.Algebra.Tower +import Mathlib.Algebra.Algebra.Unitization /-! # Quasiregularity and quasispectrum diff --git a/Mathlib/Algebra/Algebra/Rat.lean b/Mathlib/Algebra/Algebra/Rat.lean index 8cdd3179627ce..d8b6989c91ed8 100644 --- a/Mathlib/Algebra/Algebra/Rat.lean +++ b/Mathlib/Algebra/Algebra/Rat.lean @@ -106,9 +106,6 @@ instance instSMulCommClass' [SMulCommClass S R S] : SMulCommClass R ℚ S := end DivisionRing -@[deprecated Algebra.id.map_eq_id (since := "2024-07-30")] -lemma _root_.algebraMap_rat_rat : algebraMap ℚ ℚ = RingHom.id ℚ := rfl - instance algebra_rat_subsingleton {R} [Semiring R] : Subsingleton (Algebra ℚ R) := ⟨fun x y => Algebra.algebra_ext x y <| RingHom.congr_fun <| Subsingleton.elim _ _⟩ diff --git a/Mathlib/Algebra/Algebra/Spectrum.lean b/Mathlib/Algebra/Algebra/Spectrum.lean index 25fc6d68d6f85..6ce7386ea0a1f 100644 --- a/Mathlib/Algebra/Algebra/Spectrum.lean +++ b/Mathlib/Algebra/Algebra/Spectrum.lean @@ -3,7 +3,8 @@ Copyright (c) 2021 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Algebra.Star.Subalgebra +import Mathlib.Algebra.Algebra.Subalgebra.Basic +import Mathlib.Algebra.Star.Pointwise import Mathlib.RingTheory.Ideal.Maps import Mathlib.RingTheory.Ideal.Nonunits import Mathlib.Tactic.NoncommRing @@ -273,11 +274,6 @@ theorem subset_subalgebra {S R A : Type*} [CommSemiring R] [Ring A] [Algebra R A spectrum R (a : A) ⊆ spectrum R a := Set.compl_subset_compl.mpr fun _ ↦ IsUnit.map (SubalgebraClass.val s) -@[deprecated subset_subalgebra (since := "2024-07-19")] -theorem subset_starSubalgebra [StarRing R] [StarRing A] [StarModule R A] {S : StarSubalgebra R A} - (a : S) : spectrum R (a : A) ⊆ spectrum R a := - subset_subalgebra a - theorem singleton_add_eq (a : A) (r : R) : {r} + σ a = σ (↑ₐ r + a) := ext fun x => by rw [singleton_add, image_add_left, mem_preimage, add_comm, add_mem_iff, map_neg, neg_neg] diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index 64279f792eac2..c26e4462891f9 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -175,9 +175,6 @@ protected theorem intCast_mem {R : Type u} {A : Type v} [CommRing R] [Ring A] [A (S : Subalgebra R A) (n : ℤ) : (n : A) ∈ S := intCast_mem S n -@[deprecated natCast_mem (since := "2024-04-05")] alias coe_nat_mem := Subalgebra.natCast_mem -@[deprecated intCast_mem (since := "2024-04-05")] alias coe_int_mem := Subalgebra.intCast_mem - /-- The projection from a subalgebra of `A` to an additive submonoid of `A`. -/ @[simps coe] def toAddSubmonoid {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] diff --git a/Mathlib/Algebra/BigOperators/Group/Finset/Basic.lean b/Mathlib/Algebra/BigOperators/Group/Finset/Basic.lean index c18a296e64f53..37c6b4dec36b7 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset/Basic.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset/Basic.lean @@ -1612,9 +1612,6 @@ theorem toFinset_sum_count_eq (s : Multiset α) : ∑ a ∈ s.toFinset, s.count congr with a simpa using hms a -@[deprecated sum_count_eq_card (since := "2024-07-21")] -theorem sum_count_eq [Fintype α] (s : Multiset α) : ∑ a, s.count a = Multiset.card s := by simp - @[simp] theorem toFinset_sum_count_nsmul_eq (s : Multiset α) : ∑ a ∈ s.toFinset, s.count a • {a} = s := by diff --git a/Mathlib/Algebra/BigOperators/Group/Finset/Defs.lean b/Mathlib/Algebra/BigOperators/Group/Finset/Defs.lean index 4a1c814db33f9..6dd51d4dcc7a9 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset/Defs.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset/Defs.lean @@ -312,9 +312,6 @@ theorem prod_empty : ∏ x ∈ ∅, f x = 1 := theorem prod_of_isEmpty [IsEmpty α] (s : Finset α) : ∏ i ∈ s, f i = 1 := by rw [eq_empty_of_isEmpty s, prod_empty] -@[deprecated (since := "2024-06-11")] alias prod_of_empty := prod_of_isEmpty -@[deprecated (since := "2024-06-11")] alias sum_of_empty := sum_of_isEmpty - @[to_additive (attr := simp)] theorem prod_const_one : (∏ _x ∈ s, (1 : β)) = 1 := by simp only [Finset.prod, Multiset.map_const', Multiset.prod_replicate, one_pow] diff --git a/Mathlib/Algebra/BigOperators/Group/List/Basic.lean b/Mathlib/Algebra/BigOperators/Group/List/Basic.lean index 17499ca3064d6..3aeeb77c4a50b 100644 --- a/Mathlib/Algebra/BigOperators/Group/List/Basic.lean +++ b/Mathlib/Algebra/BigOperators/Group/List/Basic.lean @@ -473,8 +473,6 @@ namespace MonoidHom protected theorem map_list_prod (f : M →* N) (l : List M) : f l.prod = (l.map f).prod := map_list_prod f l -attribute [deprecated map_list_sum (since := "2024-05-02")] AddMonoidHom.map_list_sum - end MonoidHom end MonoidHom diff --git a/Mathlib/Algebra/BigOperators/Group/List/Lemmas.lean b/Mathlib/Algebra/BigOperators/Group/List/Lemmas.lean index 0b1fac94d6687..4fbfe6be74017 100644 --- a/Mathlib/Algebra/BigOperators/Group/List/Lemmas.lean +++ b/Mathlib/Algebra/BigOperators/Group/List/Lemmas.lean @@ -147,16 +147,6 @@ lemma drop_take_succ_flatten_eq_getElem (L : List (List α)) (i : Nat) (h : i < simp only [this, length_map, take_sum_flatten, drop_sum_flatten, drop_take_succ_eq_cons_getElem, h, flatten, append_nil] -@[deprecated (since := "2024-06-11")] -alias drop_take_succ_join_eq_getElem := drop_take_succ_flatten_eq_getElem - -@[deprecated drop_take_succ_flatten_eq_getElem (since := "2024-06-11")] -lemma drop_take_succ_join_eq_get (L : List (List α)) (i : Fin L.length) : - (L.flatten.take ((L.map length).take (i + 1)).sum).drop - ((L.map length).take i).sum = get L i := by - rw [drop_take_succ_flatten_eq_getElem _ _ i.2] - simp - end List diff --git a/Mathlib/Algebra/BigOperators/Ring.lean b/Mathlib/Algebra/BigOperators/Ring.lean index a548b362058b9..12d5c406ca4ca 100644 --- a/Mathlib/Algebra/BigOperators/Ring.lean +++ b/Mathlib/Algebra/BigOperators/Ring.lean @@ -266,8 +266,6 @@ theorem prod_range_natCast_sub (n k : ℕ) : · rw [← mem_range] at hnk rw [prod_eq_zero hnk, prod_eq_zero hnk] <;> simp -@[deprecated (since := "2024-05-27")] alias prod_range_cast_nat_sub := prod_range_natCast_sub - end CommRing diff --git a/Mathlib/Algebra/CharP/ExpChar.lean b/Mathlib/Algebra/CharP/ExpChar.lean index 9053124e1a464..f211e34841757 100644 --- a/Mathlib/Algebra/CharP/ExpChar.lean +++ b/Mathlib/Algebra/CharP/ExpChar.lean @@ -68,8 +68,5 @@ theorem frobenius_add : frobenius R p (x + y) = frobenius R p x + frobenius R p theorem frobenius_natCast (n : ℕ) : frobenius R p n = n := map_natCast (frobenius R p) n -@[deprecated (since := "2024-04-17")] -alias frobenius_nat_cast := frobenius_natCast - end CommSemiring end frobenius diff --git a/Mathlib/Algebra/CharZero/Lemmas.lean b/Mathlib/Algebra/CharZero/Lemmas.lean index 5c3ac417809dd..5c3ee47cef6fc 100644 --- a/Mathlib/Algebra/CharZero/Lemmas.lean +++ b/Mathlib/Algebra/CharZero/Lemmas.lean @@ -61,15 +61,9 @@ namespace Function lemma support_natCast (hn : n ≠ 0) : support (n : α → M) = univ := support_const <| Nat.cast_ne_zero.2 hn -@[deprecated (since := "2024-04-17")] -alias support_nat_cast := support_natCast - lemma mulSupport_natCast (hn : n ≠ 1) : mulSupport (n : α → M) = univ := mulSupport_const <| Nat.cast_ne_one.2 hn -@[deprecated (since := "2024-04-17")] -alias mulSupport_nat_cast := mulSupport_natCast - end Function end AddMonoidWithOne @@ -108,13 +102,9 @@ variable {R : Type*} [DivisionSemiring R] [NeZero (2 : R)] @[simp] lemma add_self_div_two (a : R) : (a + a) / 2 = a := by rw [← mul_two, mul_div_cancel_right₀ a two_ne_zero] -@[deprecated (since := "2024-07-16")] alias half_add_self := add_self_div_two - @[simp] theorem add_halves (a : R) : a / 2 + a / 2 = a := by rw [← add_div, add_self_div_two] -@[deprecated (since := "2024-07-16")] alias add_halves' := add_halves - end section diff --git a/Mathlib/Algebra/DirectSum/Internal.lean b/Mathlib/Algebra/DirectSum/Internal.lean index 35229ff31057a..9a252c382591b 100644 --- a/Mathlib/Algebra/DirectSum/Internal.lean +++ b/Mathlib/Algebra/DirectSum/Internal.lean @@ -68,9 +68,6 @@ theorem SetLike.natCast_mem_graded [Zero ι] [AddMonoidWithOne R] [SetLike σ R] rw [Nat.cast_succ] exact add_mem n_ih (SetLike.one_mem_graded _) -@[deprecated (since := "2024-04-17")] -alias SetLike.nat_cast_mem_graded := SetLike.natCast_mem_graded - theorem SetLike.intCast_mem_graded [Zero ι] [AddGroupWithOne R] [SetLike σ R] [AddSubgroupClass σ R] (A : ι → σ) [SetLike.GradedOne A] (z : ℤ) : (z : R) ∈ A 0 := by induction z @@ -79,9 +76,6 @@ theorem SetLike.intCast_mem_graded [Zero ι] [AddGroupWithOne R] [SetLike σ R] · rw [Int.cast_negSucc] exact neg_mem (SetLike.natCast_mem_graded _ _) -@[deprecated (since := "2024-04-17")] -alias SetLike.int_cast_mem_graded := SetLike.intCast_mem_graded - section DirectSum variable [DecidableEq ι] diff --git a/Mathlib/Algebra/Field/Defs.lean b/Mathlib/Algebra/Field/Defs.lean index d3b64ac26a5f3..9e1786cf0fabc 100644 --- a/Mathlib/Algebra/Field/Defs.lean +++ b/Mathlib/Algebra/Field/Defs.lean @@ -190,8 +190,6 @@ variable (K) @[simp] lemma smul_one_eq_cast (q : ℚ≥0) : q • (1 : K) = q := by rw [NNRat.smul_def, mul_one] -@[deprecated (since := "2024-05-03")] alias smul_one_eq_coe := smul_one_eq_cast - end NNRat namespace Rat @@ -210,8 +208,6 @@ theorem smul_def (a : ℚ) (x : K) : a • x = ↑a * x := DivisionRing.qsmul_de theorem smul_one_eq_cast (A : Type*) [DivisionRing A] (m : ℚ) : m • (1 : A) = ↑m := by rw [Rat.smul_def, mul_one] -@[deprecated (since := "2024-05-03")] alias smul_one_eq_coe := smul_one_eq_cast - end Rat /-- `OfScientific.ofScientific` is the simp-normal form. -/ diff --git a/Mathlib/Algebra/Field/Subfield/Defs.lean b/Mathlib/Algebra/Field/Subfield/Defs.lean index f7955cd9e0f5e..4703a93ef4269 100644 --- a/Mathlib/Algebra/Field/Subfield/Defs.lean +++ b/Mathlib/Algebra/Field/Subfield/Defs.lean @@ -89,10 +89,6 @@ lemma nnqsmul_mem (s : S) (q : ℚ≥0) (hx : x ∈ s) : q • x ∈ s := by lemma qsmul_mem (s : S) (q : ℚ) (hx : x ∈ s) : q • x ∈ s := by simpa only [Rat.smul_def] using mul_mem (ratCast_mem _ _) hx -@[deprecated (since := "2024-04-05")] alias coe_rat_cast := coe_ratCast -@[deprecated (since := "2024-04-05")] alias coe_rat_mem := ratCast_mem -@[deprecated (since := "2024-04-05")] alias rat_smul_mem := qsmul_mem - @[aesop safe apply (rule_sets := [SetLike])] lemma ofScientific_mem (s : S) {b : Bool} {n m : ℕ} : (OfScientific.ofScientific n b m : K) ∈ s := @@ -256,8 +252,6 @@ protected theorem zsmul_mem {x : K} (hx : x ∈ s) (n : ℤ) : n • x ∈ s := protected theorem intCast_mem (n : ℤ) : (n : K) ∈ s := intCast_mem s n -@[deprecated (since := "2024-04-05")] alias coe_int_mem := intCast_mem - theorem zpow_mem {x : K} (hx : x ∈ s) (n : ℤ) : x ^ n ∈ s := by cases n · simpa using s.pow_mem hx _ diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index 91b84bb8823ee..09523de0a4085 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -105,9 +105,6 @@ instance instFunLike : FunLike (AddChar A M) A M where /-- An additive character maps sums to products. -/ lemma map_add_eq_mul (ψ : AddChar A M) (x y : A) : ψ (x + y) = ψ x * ψ y := ψ.map_add_eq_mul' x y -@[deprecated (since := "2024-06-06")] alias map_zero_one := map_zero_eq_one -@[deprecated (since := "2024-06-06")] alias map_add_mul := map_add_eq_mul - /-- Interpret an additive character as a monoid homomorphism. -/ def toMonoidHom (φ : AddChar A M) : Multiplicative A →* M where toFun := φ.toFun @@ -124,8 +121,6 @@ def toMonoidHom (φ : AddChar A M) : Multiplicative A →* M where lemma map_nsmul_eq_pow (ψ : AddChar A M) (n : ℕ) (x : A) : ψ (n • x) = ψ x ^ n := ψ.toMonoidHom.map_pow x n -@[deprecated (since := "2024-06-06")] alias map_nsmul_pow := map_nsmul_eq_pow - /-- Additive characters `A → M` are the same thing as monoid homomorphisms from `Multiplicative A` to `M`. -/ def toMonoidHomEquiv : AddChar A M ≃ (Multiplicative A →* M) where @@ -254,16 +249,6 @@ lemma eq_zero_iff : ψ = 0 ↔ ∀ x, ψ x = 1 := DFunLike.ext_iff lemma ne_one_iff : ψ ≠ 1 ↔ ∃ x, ψ x ≠ 1 := DFunLike.ne_iff lemma ne_zero_iff : ψ ≠ 0 ↔ ∃ x, ψ x ≠ 1 := DFunLike.ne_iff -/-- An additive character is *nontrivial* if it takes a value `≠ 1`. -/ -@[deprecated "No deprecation message was provided." (since := "2024-06-06")] -def IsNontrivial (ψ : AddChar A M) : Prop := ∃ a : A, ψ a ≠ 1 - -set_option linter.deprecated false in -/-- An additive character is nontrivial iff it is not the trivial character. -/ -@[deprecated ne_one_iff (since := "2024-06-06")] -lemma isNontrivial_iff_ne_trivial (ψ : AddChar A M) : IsNontrivial ψ ↔ ψ ≠ 1 := - not_forall.symm.trans (DFunLike.ext_iff (f := ψ) (g := 1)).symm.not - noncomputable instance : DecidableEq (AddChar A M) := Classical.decEq _ end Basic @@ -401,9 +386,6 @@ lemma map_neg_eq_inv (ψ : AddChar A M) (a : A) : ψ (-a) = (ψ a)⁻¹ := by lemma map_zsmul_eq_zpow (ψ : AddChar A M) (n : ℤ) (a : A) : ψ (n • a) = (ψ a) ^ n := ψ.toMonoidHom.map_zpow a n -@[deprecated (since := "2024-06-06")] alias map_neg_inv := map_neg_eq_inv -@[deprecated (since := "2024-06-06")] alias map_zsmul_zpow := map_zsmul_eq_zpow - end fromAddGrouptoDivisionMonoid section fromAddCommGrouptoDivisionCommMonoid diff --git a/Mathlib/Algebra/Group/Basic.lean b/Mathlib/Algebra/Group/Basic.lean index 342a421455323..f95267bddc8a0 100644 --- a/Mathlib/Algebra/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Basic.lean @@ -1031,6 +1031,3 @@ theorem multiplicative_of_isTotal (p : α → Prop) (hswap : ∀ {a b}, p a → exacts [⟨pa, pb⟩, ⟨pb, pc⟩, ⟨pa, pc⟩] end multiplicative - --- The name `add_sub_cancel` was reused --- @[deprecated (since := "2024-03-20")] alias add_sub_cancel := add_sub_cancel_right diff --git a/Mathlib/Algebra/Group/Even.lean b/Mathlib/Algebra/Group/Even.lean index 6e598f103f8a3..af9fde64df633 100644 --- a/Mathlib/Algebra/Group/Even.lean +++ b/Mathlib/Algebra/Group/Even.lean @@ -152,8 +152,6 @@ variable [DivisionMonoid α] {a : α} @[to_additive Even.zsmul_right] lemma IsSquare.zpow (n : ℤ) : IsSquare a → IsSquare (a ^ n) := by rintro ⟨r, rfl⟩; exact ⟨r ^ n, (Commute.refl _).mul_zpow _⟩ -@[deprecated (since := "2024-01-19")] alias Even.zsmul := Even.zsmul_right - end DivisionMonoid @[to_additive] @@ -163,5 +161,3 @@ lemma IsSquare.div [DivisionCommMonoid α] {a b : α} (ha : IsSquare a) (hb : Is @[to_additive (attr := simp) Even.zsmul_left] lemma Even.isSquare_zpow [Group α] {n : ℤ} : Even n → ∀ a : α, IsSquare (a ^ n) := by rintro ⟨m, rfl⟩ a; exact ⟨a ^ m, zpow_add _ _ _⟩ - -@[deprecated (since := "2024-01-07")] alias Even.zsmul' := Even.zsmul_left diff --git a/Mathlib/Algebra/Group/Hom/Instances.lean b/Mathlib/Algebra/Group/Hom/Instances.lean index eb5b76b157ba5..4e84eda29a94f 100644 --- a/Mathlib/Algebra/Group/Hom/Instances.lean +++ b/Mathlib/Algebra/Group/Hom/Instances.lean @@ -93,9 +93,6 @@ theorem AddMonoid.End.intCast_apply [AddCommGroup M] (z : ℤ) (m : M) : (↑z : AddMonoid.End M) m = z • m := rfl -@[deprecated (since := "2024-04-17")] -alias AddMonoid.End.int_cast_apply := AddMonoid.End.intCast_apply - @[to_additive (attr := simp)] lemma MonoidHom.pow_apply {M N : Type*} [MulOneClass M] [CommMonoid N] (f : M →* N) (n : ℕ) (x : M) : (f ^ n) x = (f x) ^ n := diff --git a/Mathlib/Algebra/Group/Subgroup/ZPowers/Lemmas.lean b/Mathlib/Algebra/Group/Subgroup/ZPowers/Lemmas.lean index 60b1bc0f5fff1..6c16701b39a5d 100644 --- a/Mathlib/Algebra/Group/Subgroup/ZPowers/Lemmas.lean +++ b/Mathlib/Algebra/Group/Subgroup/ZPowers/Lemmas.lean @@ -45,16 +45,10 @@ variable {R : Type*} [Ring R] (r : R) (k : ℤ) theorem intCast_mul_mem_zmultiples : ↑(k : ℤ) * r ∈ zmultiples r := by simpa only [← zsmul_eq_mul] using zsmul_mem_zmultiples r k -@[deprecated (since := "2024-04-17")] -alias int_cast_mul_mem_zmultiples := intCast_mul_mem_zmultiples - @[simp] theorem intCast_mem_zmultiples_one : ↑(k : ℤ) ∈ zmultiples (1 : R) := mem_zmultiples_iff.mp ⟨k, by simp⟩ -@[deprecated (since := "2024-04-17")] -alias int_cast_mem_zmultiples_one := intCast_mem_zmultiples_one - end Ring end AddSubgroup diff --git a/Mathlib/Algebra/Group/Units/Basic.lean b/Mathlib/Algebra/Group/Units/Basic.lean index 20b162ed67d8c..e1e58c3e22f7e 100644 --- a/Mathlib/Algebra/Group/Units/Basic.lean +++ b/Mathlib/Algebra/Group/Units/Basic.lean @@ -463,12 +463,3 @@ end IsUnit -- namespace end IsUnit - -attribute [deprecated div_mul_cancel_right (since := "2024-03-20")] IsUnit.div_mul_left -attribute [deprecated sub_add_cancel_right (since := "2024-03-20")] IsAddUnit.sub_add_left -attribute [deprecated div_mul_cancel_left (since := "2024-03-20")] IsUnit.div_mul_right -attribute [deprecated sub_add_cancel_left (since := "2024-03-20")] IsAddUnit.sub_add_right --- The names `IsUnit.mul_div_cancel` and `IsAddUnit.add_sub_cancel` have been reused --- @[deprecated (since := "2024-03-20")] alias IsUnit.mul_div_cancel := IsUnit.mul_div_cancel_right --- @[deprecated (since := "2024-03-20")] --- alias IsAddUnit.add_sub_cancel := IsAddUnit.add_sub_cancel_right diff --git a/Mathlib/Algebra/GroupWithZero/Associated.lean b/Mathlib/Algebra/GroupWithZero/Associated.lean index f26266a8ec15e..e82388bd07f57 100644 --- a/Mathlib/Algebra/GroupWithZero/Associated.lean +++ b/Mathlib/Algebra/GroupWithZero/Associated.lean @@ -504,9 +504,6 @@ instance uniqueUnits : Unique (Associates M)ˣ where exact Quotient.inductionOn₂ a b <| fun a b hab hba ↦ Units.ext <| Quotient.sound <| associated_one_of_associated_mul_one <| Quotient.exact hab -@[deprecated (since := "2024-07-22")] alias mul_eq_one_iff := mul_eq_one -@[deprecated (since := "2024-07-22")] protected alias units_eq_one := Subsingleton.elim - @[simp] theorem coe_unit_eq_one (u : (Associates M)ˣ) : (u : Associates M) = 1 := by simp [eq_iff_true_of_subsingleton] diff --git a/Mathlib/Algebra/GroupWithZero/Center.lean b/Mathlib/Algebra/GroupWithZero/Center.lean index f38ccd5587ab2..78f4a2b058124 100644 --- a/Mathlib/Algebra/GroupWithZero/Center.lean +++ b/Mathlib/Algebra/GroupWithZero/Center.lean @@ -53,13 +53,5 @@ lemma center_units_eq : center G₀ˣ = ((↑) : G₀ˣ → G₀) ⁻¹' center a / b ∈ centralizer s := by simpa only [div_eq_mul_inv] using mul_mem_centralizer ha (inv_mem_centralizer₀ hb) -@[deprecated inv_mem_center (since := "2024-06-17")] -theorem inv_mem_center₀ (ha : a ∈ Set.center G₀) : a⁻¹ ∈ Set.center G₀ := - inv_mem_center ha - -@[deprecated div_mem_center (since := "2024-06-17")] -theorem div_mem_center₀ (ha : a ∈ Set.center G₀) (hb : b ∈ Set.center G₀) : a / b ∈ Set.center G₀ := - div_mem_center ha hb - end GroupWithZero end Set diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index c4108f058926d..fe21cad70966f 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -365,9 +365,6 @@ lemma zpow_ne_zero {a : G₀} : ∀ n : ℤ, a ≠ 0 → a ^ n ≠ 0 lemma eq_zero_of_zpow_eq_zero {n : ℤ} : a ^ n = 0 → a = 0 := not_imp_not.1 (zpow_ne_zero _) -@[deprecated (since := "2024-05-07")] alias zpow_ne_zero_of_ne_zero := zpow_ne_zero -@[deprecated (since := "2024-05-07")] alias zpow_eq_zero := eq_zero_of_zpow_eq_zero - lemma zpow_eq_zero_iff {n : ℤ} (hn : n ≠ 0) : a ^ n = 0 ↔ a = 0 := ⟨eq_zero_of_zpow_eq_zero, fun ha => ha.symm ▸ zero_zpow _ hn⟩ diff --git a/Mathlib/Algebra/Homology/ShortComplex/Exact.lean b/Mathlib/Algebra/Homology/ShortComplex/Exact.lean index 4e407ab86ecdf..7b7ba4b2c4b1b 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Exact.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Exact.lean @@ -909,15 +909,6 @@ lemma Exact.liftFromProjective_comp rw [← toCycles_i, Projective.factorThru_comp_assoc, liftCycles_i] -@[deprecated (since := "2024-07-09")] alias _root_.CategoryTheory.Exact.lift := - Exact.liftFromProjective -@[deprecated (since := "2024-07-09")] alias _root_.CategoryTheory.Exact.lift_comp := - Exact.liftFromProjective_comp -@[deprecated (since := "2024-07-09")] alias _root_.CategoryTheory.Injective.Exact.desc := - Exact.descToInjective -@[deprecated (since := "2024-07-09")] alias _root_.CategoryTheory.Injective.Exact.comp_desc := - Exact.comp_descToInjective - end Abelian end ShortComplex diff --git a/Mathlib/Algebra/Lie/Weights/Linear.lean b/Mathlib/Algebra/Lie/Weights/Linear.lean index f52cb53419a1b..4fac19ae47958 100644 --- a/Mathlib/Algebra/Lie/Weights/Linear.lean +++ b/Mathlib/Algebra/Lie/Weights/Linear.lean @@ -128,9 +128,6 @@ lemma trace_comp_toEnd_genWeightSpace_eq (χ : L → R) : rw [LinearMap.comp_apply, LieHom.coe_toLinearMap, h₁, map_add, h₂] simp [mul_comm (χ x)] -@[deprecated (since := "2024-04-06")] -alias trace_comp_toEnd_weight_space_eq := trace_comp_toEnd_genWeightSpace_eq - variable {R L M} in lemma zero_lt_finrank_genWeightSpace {χ : L → R} (hχ : genWeightSpace M χ ≠ ⊥) : 0 < finrank R (genWeightSpace M χ) := by diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index 1c97089ef1615..196c467a47072 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -35,9 +35,6 @@ theorem invOf_two_smul_add_invOf_two_smul (R) [Semiring R] [AddCommMonoid M] [Mo (⅟ 2 : R) • x + (⅟ 2 : R) • x = x := Convex.combo_self invOf_two_add_invOf_two _ -@[deprecated (since := "2024-04-17")] -alias map_nat_cast_smul := map_natCast_smul - theorem map_inv_natCast_smul [AddCommMonoid M] [AddCommMonoid M₂] {F : Type*} [FunLike F M M₂] [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [DivisionSemiring R] [DivisionSemiring S] [Module R M] @@ -55,9 +52,6 @@ theorem map_inv_natCast_smul [AddCommMonoid M] [AddCommMonoid M₂] {F : Type*} rw [← smul_inv_smul₀ hR x, map_natCast_smul f R S, hS, zero_smul] · rw [← inv_smul_smul₀ hS (f _), ← map_natCast_smul f R S, smul_inv_smul₀ hR] -@[deprecated (since := "2024-04-17")] -alias map_inv_nat_cast_smul := map_inv_natCast_smul - theorem map_inv_intCast_smul [AddCommGroup M] [AddCommGroup M₂] {F : Type*} [FunLike F M M₂] [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [DivisionRing R] [DivisionRing S] [Module R M] [Module S M₂] (z : ℤ) (x : M) : f ((z⁻¹ : R) • x) = (z⁻¹ : S) • f x := by @@ -66,9 +60,6 @@ theorem map_inv_intCast_smul [AddCommGroup M] [AddCommGroup M₂] {F : Type*} [F · simp_rw [Int.cast_neg, Int.cast_natCast, inv_neg, neg_smul, map_neg, map_inv_natCast_smul _ R S] -@[deprecated (since := "2024-04-17")] -alias map_inv_int_cast_smul := map_inv_intCast_smul - /-- If `E` is a vector space over two division semirings `R` and `S`, then scalar multiplications agree on inverses of natural numbers in `R` and `S`. -/ theorem inv_natCast_smul_eq {E : Type*} (R S : Type*) [AddCommMonoid E] [DivisionSemiring R] @@ -76,18 +67,12 @@ theorem inv_natCast_smul_eq {E : Type*} (R S : Type*) [AddCommMonoid E] [Divisio (n⁻¹ : R) • x = (n⁻¹ : S) • x := map_inv_natCast_smul (AddMonoidHom.id E) R S n x -@[deprecated (since := "2024-04-17")] -alias inv_nat_cast_smul_eq := inv_natCast_smul_eq - /-- If `E` is a vector space over two division rings `R` and `S`, then scalar multiplications agree on inverses of integer numbers in `R` and `S`. -/ theorem inv_intCast_smul_eq {E : Type*} (R S : Type*) [AddCommGroup E] [DivisionRing R] [DivisionRing S] [Module R E] [Module S E] (n : ℤ) (x : E) : (n⁻¹ : R) • x = (n⁻¹ : S) • x := map_inv_intCast_smul (AddMonoidHom.id E) R S n x -@[deprecated (since := "2024-04-17")] -alias inv_int_cast_smul_eq := inv_intCast_smul_eq - /-- If `E` is a vector space over a division semiring `R` and has a monoid action by `α`, then that action commutes by scalar multiplication of inverses of natural numbers in `R`. -/ theorem inv_natCast_smul_comm {α E : Type*} (R : Type*) [AddCommMonoid E] [DivisionSemiring R] @@ -95,9 +80,6 @@ theorem inv_natCast_smul_comm {α E : Type*} (R : Type*) [AddCommMonoid E] [Divi (n⁻¹ : R) • s • x = s • (n⁻¹ : R) • x := (map_inv_natCast_smul (DistribMulAction.toAddMonoidHom E s) R R n x).symm -@[deprecated (since := "2024-04-17")] -alias inv_nat_cast_smul_comm := inv_natCast_smul_comm - /-- If `E` is a vector space over a division ring `R` and has a monoid action by `α`, then that action commutes by scalar multiplication of inverses of integers in `R` -/ theorem inv_intCast_smul_comm {α E : Type*} (R : Type*) [AddCommGroup E] [DivisionRing R] @@ -105,9 +87,6 @@ theorem inv_intCast_smul_comm {α E : Type*} (R : Type*) [AddCommGroup E] [Divis (n⁻¹ : R) • s • x = s • (n⁻¹ : R) • x := (map_inv_intCast_smul (DistribMulAction.toAddMonoidHom E s) R R n x).symm -@[deprecated (since := "2024-04-17")] -alias inv_int_cast_smul_comm := inv_intCast_smul_comm - namespace Function lemma support_smul_subset_left [Zero R] [Zero M] [SMulWithZero R M] (f : α → R) (g : α → M) : diff --git a/Mathlib/Algebra/Module/End.lean b/Mathlib/Algebra/Module/End.lean index fcc7b16410795..a0a9057609f45 100644 --- a/Mathlib/Algebra/Module/End.lean +++ b/Mathlib/Algebra/Module/End.lean @@ -96,12 +96,6 @@ lemma Int.cast_smul_eq_zsmul (n : ℤ) (b : M) : (n : R) • b = n • b := simp DFunLike.congr_fun this n -@[deprecated (since := "2024-07-23")] alias intCast_smul := Int.cast_smul_eq_zsmul - -/-- `zsmul` is equal to any other module structure via a cast. -/ -@[deprecated Int.cast_smul_eq_zsmul (since := "2024-07-23")] -theorem zsmul_eq_smul_cast (n : ℤ) (b : M) : n • b = (n : R) • b := (Int.cast_smul_eq_zsmul ..).symm - end /-- Convert back any exotic `ℤ`-smul to the canonical instance. This should not be needed since in diff --git a/Mathlib/Algebra/Module/NatInt.lean b/Mathlib/Algebra/Module/NatInt.lean index 12ce588590dbe..4797de84d5a83 100644 --- a/Mathlib/Algebra/Module/NatInt.lean +++ b/Mathlib/Algebra/Module/NatInt.lean @@ -101,10 +101,6 @@ lemma Nat.cast_smul_eq_nsmul (n : ℕ) (b : M) : (n : R) • b = n • b := by lemma ofNat_smul_eq_nsmul (n : ℕ) [n.AtLeastTwo] (b : M) : (ofNat(n) : R) • b = ofNat(n) • b := Nat.cast_smul_eq_nsmul .. -/-- `nsmul` is equal to any other module structure via a cast. -/ -@[deprecated Nat.cast_smul_eq_nsmul (since := "2024-07-23")] -lemma nsmul_eq_smul_cast (n : ℕ) (b : M) : n • b = (n : R) • b := (Nat.cast_smul_eq_nsmul ..).symm - end /-- Convert back any exotic `ℕ`-smul to the canonical instance. This should not be needed since in @@ -137,6 +133,3 @@ theorem Nat.smul_one_eq_cast {R : Type*} [NonAssocSemiring R] (m : ℕ) : m • theorem Int.smul_one_eq_cast {R : Type*} [NonAssocRing R] (m : ℤ) : m • (1 : R) = ↑m := by rw [zsmul_eq_mul, mul_one] - -@[deprecated (since := "2024-05-03")] alias Nat.smul_one_eq_coe := Nat.smul_one_eq_cast -@[deprecated (since := "2024-05-03")] alias Int.smul_one_eq_coe := Int.smul_one_eq_cast diff --git a/Mathlib/Algebra/Module/Rat.lean b/Mathlib/Algebra/Module/Rat.lean index 3ce5e5ac011a8..82bef1eddeacb 100644 --- a/Mathlib/Algebra/Module/Rat.lean +++ b/Mathlib/Algebra/Module/Rat.lean @@ -30,9 +30,6 @@ theorem map_ratCast_smul [AddCommGroup M] [AddCommGroup M₂] {F : Type*} [FunLi rw [Rat.cast_def, Rat.cast_def, div_eq_mul_inv, div_eq_mul_inv, mul_smul, mul_smul, map_intCast_smul f R S, map_inv_natCast_smul f R S] -@[deprecated (since := "2024-04-17")] -alias map_rat_cast_smul := map_ratCast_smul - theorem map_nnrat_smul [AddCommMonoid M] [AddCommMonoid M₂] [_instM : Module ℚ≥0 M] [_instM₂ : Module ℚ≥0 M₂] {F : Type*} [FunLike F M M₂] [AddMonoidHomClass F M M₂] @@ -67,9 +64,6 @@ theorem ratCast_smul_eq {E : Type*} (R S : Type*) [AddCommGroup E] [DivisionRing [DivisionRing S] [Module R E] [Module S E] (r : ℚ) (x : E) : (r : R) • x = (r : S) • x := map_ratCast_smul (AddMonoidHom.id E) R S r x -@[deprecated (since := "2024-04-17")] -alias rat_cast_smul_eq := ratCast_smul_eq - instance IsScalarTower.nnrat {R : Type u} {M : Type v} [Semiring R] [AddCommMonoid M] [Module R M] [Module ℚ≥0 R] [Module ℚ≥0 M] : IsScalarTower ℚ≥0 R M where smul_assoc r x y := map_nnrat_smul ((smulAddHom R M).flip y) r x diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index d6da6c7a4ee38..b8b39be30757e 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -256,9 +256,6 @@ instance nonAssocSemiring : NonAssocSemiring (MonoidAlgebra k G) := theorem natCast_def (n : ℕ) : (n : MonoidAlgebra k G) = single (1 : G) (n : k) := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_def := natCast_def - end MulOneClass /-! #### Semiring structure -/ @@ -326,9 +323,6 @@ theorem intCast_def [Ring k] [MulOneClass G] (z : ℤ) : (z : MonoidAlgebra k G) = single (1 : G) (z : k) := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_def := intCast_def - instance ring [Ring k] [Monoid G] : Ring (MonoidAlgebra k G) := { MonoidAlgebra.nonAssocRing, MonoidAlgebra.semiring with } @@ -1064,9 +1058,6 @@ instance nonAssocSemiring : NonAssocSemiring k[G] := theorem natCast_def (n : ℕ) : (n : k[G]) = single (0 : G) (n : k) := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_def := natCast_def - end MulOneClass /-! #### Semiring structure -/ @@ -1138,9 +1129,6 @@ theorem intCast_def [Ring k] [AddZeroClass G] (z : ℤ) : (z : k[G]) = single (0 : G) (z : k) := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_def := intCast_def - instance ring [Ring k] [AddMonoid G] : Ring k[G] := { AddMonoidAlgebra.nonAssocRing, AddMonoidAlgebra.semiring with } diff --git a/Mathlib/Algebra/Order/AddTorsor.lean b/Mathlib/Algebra/Order/AddTorsor.lean index dbe6e284362a8..7a6a85dbb68a5 100644 --- a/Mathlib/Algebra/Order/AddTorsor.lean +++ b/Mathlib/Algebra/Order/AddTorsor.lean @@ -56,8 +56,6 @@ class IsOrderedVAdd (G P : Type*) [LE G] [LE P] [VAdd G P] : Prop where protected vadd_le_vadd_left : ∀ a b : P, a ≤ b → ∀ c : G, c +ᵥ a ≤ c +ᵥ b protected vadd_le_vadd_right : ∀ c d : G, c ≤ d → ∀ a : P, c +ᵥ a ≤ d +ᵥ a -@[deprecated (since := "2024-07-15")] alias OrderedVAdd := IsOrderedVAdd - /-- An ordered scalar multiplication is a bi-monotone scalar multiplication. Note that this is different from `OrderedSMul`, which uses strict inequality, requires `G` to be a semiring, and the defining conditions are restricted to positive elements of `G`. -/ @@ -92,8 +90,6 @@ class IsCancelVAdd (G P : Type*) [VAdd G P] : Prop where protected left_cancel : ∀ (a : G) (b c : P), a +ᵥ b = a +ᵥ c → b = c protected right_cancel : ∀ (a b : G) (c : P), a +ᵥ c = b +ᵥ c → a = b -@[deprecated (since := "2024-07-15")] alias CancelVAdd := IsCancelVAdd - /-- A scalar multiplication is cancellative if it is pointwise injective on the left and right. -/ @[to_additive] class IsCancelSMul (G P : Type*) [SMul G P] : Prop where @@ -106,8 +102,6 @@ class IsOrderedCancelVAdd (G P : Type*) [LE G] [LE P] [VAdd G P] extends protected le_of_vadd_le_vadd_left : ∀ (a : G) (b c : P), a +ᵥ b ≤ a +ᵥ c → b ≤ c protected le_of_vadd_le_vadd_right : ∀ (a b : G) (c : P), a +ᵥ c ≤ b +ᵥ c → a ≤ b -@[deprecated (since := "2024-07-15")] alias OrderedCancelVAdd := IsOrderedCancelVAdd - /-- An ordered cancellative scalar multiplication is an ordered scalar multiplication that is cancellative. -/ @[to_additive] diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 94a5729037a50..e8c16c967f796 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -160,8 +160,6 @@ theorem floor_natCast (n : ℕ) : ⌊(n : α)⌋₊ = n := rw [le_floor_iff, Nat.cast_le] exact n.cast_nonneg -@[deprecated (since := "2024-06-08")] alias floor_coe := floor_natCast - @[simp] theorem floor_zero : ⌊(0 : α)⌋₊ = 0 := by rw [← Nat.cast_zero, floor_natCast] diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Int.lean b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean index af82f1eae06a5..599fbbf23e791 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Int.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean @@ -30,8 +30,6 @@ namespace Int theorem natCast_strictMono : StrictMono (· : ℕ → ℤ) := fun _ _ ↦ Int.ofNat_lt.2 -@[deprecated (since := "2024-05-25")] alias coe_nat_strictMono := natCast_strictMono - /-! ### Miscellaneous lemmas -/ theorem abs_eq_natAbs : ∀ a : ℤ, |a| = natAbs a diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index a3293d637c8a9..f538946d0fc94 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -275,9 +275,6 @@ theorem mul_le_mul_of_nonneg [PosMulMono α] [MulPosMono α] (h₁ : a ≤ b) (h₂ : c ≤ d) (a0 : 0 ≤ a) (d0 : 0 ≤ d) : a * c ≤ b * d := (mul_le_mul_of_nonneg_left h₂ a0).trans (mul_le_mul_of_nonneg_right h₁ d0) -@[deprecated (since := "2024-07-13")] -alias mul_le_mul_of_le_of_le := mul_le_mul_of_nonneg - theorem mul_le_mul_of_nonneg' [PosMulMono α] [MulPosMono α] (h₁ : a ≤ b) (h₂ : c ≤ d) (c0 : 0 ≤ c) (b0 : 0 ≤ b) : a * c ≤ b * d := (mul_le_mul_of_nonneg_right h₁ c0).trans (mul_le_mul_of_nonneg_left h₂ b0) @@ -294,9 +291,6 @@ theorem mul_lt_mul_of_le_of_lt_of_nonneg_of_pos [PosMulStrictMono α] [MulPosMon alias mul_lt_mul_of_nonneg_of_pos' := mul_lt_mul_of_le_of_lt_of_nonneg_of_pos -@[deprecated (since := "2024-07-13")] -alias mul_lt_mul_of_le_of_le' := mul_lt_mul_of_le_of_lt_of_nonneg_of_pos - theorem mul_lt_mul_of_lt_of_le_of_nonneg_of_pos [PosMulMono α] [MulPosStrictMono α] (h₁ : a < b) (h₂ : c ≤ d) (a0 : 0 ≤ a) (d0 : 0 < d) : a * c < b * d := (mul_le_mul_of_nonneg_left h₂ a0).trans_lt (mul_lt_mul_of_pos_right h₁ d0) @@ -309,23 +303,14 @@ theorem mul_lt_mul_of_lt_of_le_of_pos_of_nonneg [PosMulMono α] [MulPosStrictMon alias mul_lt_mul_of_pos_of_nonneg' := mul_lt_mul_of_lt_of_le_of_pos_of_nonneg -@[deprecated (since := "2024-07-13")] -alias mul_lt_mul_of_le_of_lt' := mul_lt_mul_of_lt_of_le_of_pos_of_nonneg - theorem mul_lt_mul_of_pos [PosMulStrictMono α] [MulPosStrictMono α] (h₁ : a < b) (h₂ : c < d) (a0 : 0 < a) (d0 : 0 < d) : a * c < b * d := (mul_lt_mul_of_pos_left h₂ a0).trans (mul_lt_mul_of_pos_right h₁ d0) -@[deprecated (since := "2024-07-13")] -alias mul_lt_mul_of_pos_of_pos := mul_lt_mul_of_pos - theorem mul_lt_mul_of_pos' [PosMulStrictMono α] [MulPosStrictMono α] (h₁ : a < b) (h₂ : c < d) (c0 : 0 < c) (b0 : 0 < b) : a * c < b * d := (mul_lt_mul_of_pos_right h₁ c0).trans (mul_lt_mul_of_pos_left h₂ b0) -@[deprecated (since := "2024-07-13")] -alias mul_lt_mul_of_lt_of_lt' := mul_lt_mul_of_pos' - alias mul_le_mul := mul_le_mul_of_nonneg' attribute [gcongr] mul_le_mul diff --git a/Mathlib/Algebra/Order/Interval/Set/Group.lean b/Mathlib/Algebra/Order/Interval/Set/Group.lean index 74957b1640241..01c3d7cf1ec1f 100644 --- a/Mathlib/Algebra/Order/Interval/Set/Group.lean +++ b/Mathlib/Algebra/Order/Interval/Set/Group.lean @@ -215,46 +215,28 @@ theorem pairwise_disjoint_Ioc_add_intCast : simpa only [zsmul_one, Int.cast_add, Int.cast_one, ← add_assoc] using pairwise_disjoint_Ioc_add_zsmul a (1 : α) -@[deprecated (since := "2024-04-17")] -alias pairwise_disjoint_Ioc_add_int_cast := pairwise_disjoint_Ioc_add_intCast - theorem pairwise_disjoint_Ico_add_intCast : Pairwise (Disjoint on fun n : ℤ => Ico (a + n) (a + n + 1)) := by simpa only [zsmul_one, Int.cast_add, Int.cast_one, ← add_assoc] using pairwise_disjoint_Ico_add_zsmul a (1 : α) -@[deprecated (since := "2024-04-17")] -alias pairwise_disjoint_Ico_add_int_cast := pairwise_disjoint_Ico_add_intCast - theorem pairwise_disjoint_Ioo_add_intCast : Pairwise (Disjoint on fun n : ℤ => Ioo (a + n) (a + n + 1)) := by simpa only [zsmul_one, Int.cast_add, Int.cast_one, ← add_assoc] using pairwise_disjoint_Ioo_add_zsmul a (1 : α) -@[deprecated (since := "2024-04-17")] -alias pairwise_disjoint_Ioo_add_int_cast := pairwise_disjoint_Ioo_add_intCast - variable (α) theorem pairwise_disjoint_Ico_intCast : Pairwise (Disjoint on fun n : ℤ => Ico (n : α) (n + 1)) := by simpa only [zero_add] using pairwise_disjoint_Ico_add_intCast (0 : α) -@[deprecated (since := "2024-04-17")] -alias pairwise_disjoint_Ico_int_cast := pairwise_disjoint_Ico_intCast - theorem pairwise_disjoint_Ioo_intCast : Pairwise (Disjoint on fun n : ℤ => Ioo (n : α) (n + 1)) := by simpa only [zero_add] using pairwise_disjoint_Ioo_add_intCast (0 : α) -@[deprecated (since := "2024-04-17")] -alias pairwise_disjoint_Ioo_int_cast := pairwise_disjoint_Ioo_intCast - theorem pairwise_disjoint_Ioc_intCast : Pairwise (Disjoint on fun n : ℤ => Ioc (n : α) (n + 1)) := by simpa only [zero_add] using pairwise_disjoint_Ioc_add_intCast (0 : α) -@[deprecated (since := "2024-04-17")] -alias pairwise_disjoint_Ioc_int_cast := pairwise_disjoint_Ioc_intCast - end OrderedRing end PairwiseDisjoint diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean index fac09c291a486..c3dac155b9453 100644 --- a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean @@ -275,9 +275,6 @@ end PartialOrder end CommMonoid -@[deprecated (since := "2024-07-24")] alias mul_eq_one_iff := mul_eq_one -@[deprecated (since := "2024-07-24")] alias add_eq_zero_iff := add_eq_zero - namespace NeZero theorem pos {M} [AddZeroClass M] [PartialOrder M] [CanonicallyOrderedAdd M] diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean index aa2b6f7455877..8af4d0dbd4795 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean @@ -1025,9 +1025,6 @@ theorem mul_eq_one_iff_of_one_le [MulLeftMono α] And.intro ‹a = 1› ‹b = 1›) (by rintro ⟨rfl, rfl⟩; rw [mul_one]) -@[deprecated (since := "2024-07-24")] alias mul_eq_one_iff' := mul_eq_one_iff_of_one_le -@[deprecated (since := "2024-07-24")] alias add_eq_zero_iff' := add_eq_zero_iff_of_nonneg - section Left variable [MulLeftMono α] {a b : α} diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean index b9bc99747bd14..6d9fc46645bc1 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean @@ -320,10 +320,6 @@ instance addMonoidWithOne : AddMonoidWithOne (WithTop α) := @[simp] lemma natCast_ne_top (n : ℕ) : (n : WithTop α) ≠ ⊤ := coe_ne_top @[simp] lemma natCast_lt_top [LT α] (n : ℕ) : (n : WithTop α) < ⊤ := coe_lt_top _ -@[deprecated (since := "2024-04-05")] alias coe_nat := coe_natCast -@[deprecated (since := "2024-04-05")] alias nat_ne_top := natCast_ne_top -@[deprecated (since := "2024-04-05")] alias top_ne_nat := top_ne_natCast - @[simp] lemma coe_ofNat (n : ℕ) [n.AtLeastTwo] : ((ofNat(n) : α) : WithTop α) = ofNat(n) := rfl @[simp] lemma coe_eq_ofNat (n : ℕ) [n.AtLeastTwo] (m : α) : @@ -528,10 +524,6 @@ instance addMonoidWithOne : AddMonoidWithOne (WithBot α) := WithTop.addMonoidWi @[simp] lemma bot_ne_natCast (n : ℕ) : (⊥ : WithBot α) ≠ n := bot_ne_coe -@[deprecated (since := "2024-04-05")] alias coe_nat := coe_natCast -@[deprecated (since := "2024-04-05")] alias nat_ne_bot := natCast_ne_bot -@[deprecated (since := "2024-04-05")] alias bot_ne_nat := bot_ne_natCast - @[simp] lemma coe_ofNat (n : ℕ) [n.AtLeastTwo] : ((ofNat(n) : α) : WithBot α) = ofNat(n) := rfl @[simp] lemma coe_eq_ofNat (n : ℕ) [n.AtLeastTwo] (m : α) : diff --git a/Mathlib/Algebra/Order/Nonneg/Basic.lean b/Mathlib/Algebra/Order/Nonneg/Basic.lean index b802e57fd0526..254c54915e020 100644 --- a/Mathlib/Algebra/Order/Nonneg/Basic.lean +++ b/Mathlib/Algebra/Order/Nonneg/Basic.lean @@ -160,16 +160,10 @@ instance natCast : NatCast { x : α // 0 ≤ x } := protected theorem coe_natCast (n : ℕ) : ((↑n : { x : α // 0 ≤ x }) : α) = n := rfl -@[deprecated (since := "2024-04-17")] -alias coe_nat_cast := Nonneg.coe_natCast - @[simp] theorem mk_natCast (n : ℕ) : (⟨n, n.cast_nonneg'⟩ : { x : α // 0 ≤ x }) = n := rfl -@[deprecated (since := "2024-04-17")] -alias mk_nat_cast := mk_natCast - instance addMonoidWithOne : AddMonoidWithOne { x : α // 0 ≤ x } := { Nonneg.one (α := α) with toNatCast := Nonneg.natCast diff --git a/Mathlib/Algebra/Order/Rearrangement.lean b/Mathlib/Algebra/Order/Rearrangement.lean index d9ab756af9db1..082ae948b73c1 100644 --- a/Mathlib/Algebra/Order/Rearrangement.lean +++ b/Mathlib/Algebra/Order/Rearrangement.lean @@ -230,9 +230,6 @@ theorem AntivaryOn.sum_comp_perm_smul_eq_sum_smul_iff (hfg : AntivaryOn f g s) ∑ i ∈ s, f (σ i) • g i = ∑ i ∈ s, f i • g i ↔ AntivaryOn (f ∘ σ) g s := (hfg.dual_right.sum_comp_perm_smul_eq_sum_smul_iff hσ).trans monovaryOn_toDual_right -@[deprecated (since := "2024-06-25")] -alias AntivaryOn.sum_smul_eq_sum_comp_perm_smul_iff := AntivaryOn.sum_comp_perm_smul_eq_sum_smul_iff - variable [Fintype ι] /-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and @@ -256,9 +253,6 @@ theorem Antivary.sum_smul_comp_perm_eq_sum_smul_iff (hfg : Antivary f g) : ∑ i, f i • g (σ i) = ∑ i, f i • g i ↔ Antivary f (g ∘ σ) := by simp [(hfg.antivaryOn _).sum_smul_comp_perm_eq_sum_smul_iff fun _ _ ↦ mem_univ _] -@[deprecated (since := "2024-06-25")] -alias Antivary.sum_smul_eq_sum_smul_comp_perm_iff := Antivary.sum_smul_comp_perm_eq_sum_smul_iff - /-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which antivary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` antivary together. Stated by permuting the entries of `f`. -/ @@ -266,9 +260,6 @@ theorem Antivary.sum_comp_perm_smul_eq_sum_smul_iff (hfg : Antivary f g) : ∑ i, f (σ i) • g i = ∑ i, f i • g i ↔ Antivary (f ∘ σ) g := by simp [(hfg.antivaryOn _).sum_comp_perm_smul_eq_sum_smul_iff fun _ _ ↦ mem_univ _] -@[deprecated (since := "2024-06-25")] -alias Antivary.sum_smul_eq_sum_comp_perm_smul_iff := Antivary.sum_comp_perm_smul_eq_sum_smul_iff - end equality_case /-! #### Strict rearrangement inequality -/ @@ -303,9 +294,6 @@ theorem MonovaryOn.sum_comp_perm_smul_lt_sum_smul_iff (hfg : MonovaryOn f g s) simp [← hfg.sum_comp_perm_smul_eq_sum_smul_iff hσ, lt_iff_le_and_ne, hfg.sum_comp_perm_smul_le_sum_smul hσ] -@[deprecated (since := "2024-06-25")] -alias AntivaryOn.sum_smul_eq_sum_smul_comp_perm_iff := AntivaryOn.sum_smul_comp_perm_eq_sum_smul_iff - /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which antivary together on `s`, is strictly decreased by a permutation if and only if `f ∘ σ` and `g` do not antivary together on `s`. Stated by permuting the entries of `f`. -/ @@ -437,9 +425,6 @@ theorem AntivaryOn.sum_comp_perm_mul_eq_sum_mul_iff (hfg : AntivaryOn f g s) ∑ i ∈ s, f (σ i) * g i = ∑ i ∈ s, f i * g i ↔ AntivaryOn (f ∘ σ) g s := hfg.sum_comp_perm_smul_eq_sum_smul_iff hσ -@[deprecated (since := "2024-06-25")] -alias AntivaryOn.sum_mul_eq_sum_comp_perm_mul_iff := AntivaryOn.sum_comp_perm_mul_eq_sum_mul_iff - /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise multiplication of `f` and `g`, which antivary together on `s`, is strictly decreased by a permutation if and only if `f ∘ σ` and `g` do not antivary together on `s`. Stated by permuting the entries of `f`. -/ @@ -523,9 +508,6 @@ theorem Antivary.sum_comp_perm_mul_eq_sum_mul_iff (hfg : Antivary f g) : ∑ i, f (σ i) * g i = ∑ i, f i * g i ↔ Antivary (f ∘ σ) g := hfg.sum_comp_perm_smul_eq_sum_smul_iff -@[deprecated (since := "2024-06-25")] -alias Antivary.sum_mul_eq_sum_comp_perm_mul_iff := Antivary.sum_comp_perm_mul_eq_sum_mul_iff - /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise multiplication of `f` and `g`, which antivary together, is strictly decreased by a permutation if and only if `f ∘ σ` and `g` do not antivary together. Stated by permuting the entries of `f`. -/ diff --git a/Mathlib/Algebra/Order/ToIntervalMod.lean b/Mathlib/Algebra/Order/ToIntervalMod.lean index 4d64c73ff7da2..a3b326dc85ff6 100644 --- a/Mathlib/Algebra/Order/ToIntervalMod.lean +++ b/Mathlib/Algebra/Order/ToIntervalMod.lean @@ -946,43 +946,25 @@ theorem iUnion_Ioc_add_intCast : ⋃ n : ℤ, Ioc (a + n) (a + n + 1) = Set.univ simpa only [zsmul_one, Int.cast_add, Int.cast_one, ← add_assoc] using iUnion_Ioc_add_zsmul zero_lt_one a -@[deprecated (since := "2024-04-17")] -alias iUnion_Ioc_add_int_cast := iUnion_Ioc_add_intCast - theorem iUnion_Ico_add_intCast : ⋃ n : ℤ, Ico (a + n) (a + n + 1) = Set.univ := by simpa only [zsmul_one, Int.cast_add, Int.cast_one, ← add_assoc] using iUnion_Ico_add_zsmul zero_lt_one a -@[deprecated (since := "2024-04-17")] -alias iUnion_Ico_add_int_cast := iUnion_Ico_add_intCast - theorem iUnion_Icc_add_intCast : ⋃ n : ℤ, Icc (a + n) (a + n + 1) = Set.univ := by simpa only [zsmul_one, Int.cast_add, Int.cast_one, ← add_assoc] using iUnion_Icc_add_zsmul zero_lt_one a -@[deprecated (since := "2024-04-17")] -alias iUnion_Icc_add_int_cast := iUnion_Icc_add_intCast - variable (α) theorem iUnion_Ioc_intCast : ⋃ n : ℤ, Ioc (n : α) (n + 1) = Set.univ := by simpa only [zero_add] using iUnion_Ioc_add_intCast (0 : α) -@[deprecated (since := "2024-04-17")] -alias iUnion_Ioc_int_cast := iUnion_Ioc_intCast - theorem iUnion_Ico_intCast : ⋃ n : ℤ, Ico (n : α) (n + 1) = Set.univ := by simpa only [zero_add] using iUnion_Ico_add_intCast (0 : α) -@[deprecated (since := "2024-04-17")] -alias iUnion_Ico_int_cast := iUnion_Ico_intCast - theorem iUnion_Icc_intCast : ⋃ n : ℤ, Icc (n : α) (n + 1) = Set.univ := by simpa only [zero_add] using iUnion_Icc_add_intCast (0 : α) -@[deprecated (since := "2024-04-17")] -alias iUnion_Icc_int_cast := iUnion_Icc_intCast - end LinearOrderedRing end Union diff --git a/Mathlib/Algebra/Polynomial/AlgebraMap.lean b/Mathlib/Algebra/Polynomial/AlgebraMap.lean index 2129497e1006c..eb69993de2f62 100644 --- a/Mathlib/Algebra/Polynomial/AlgebraMap.lean +++ b/Mathlib/Algebra/Polynomial/AlgebraMap.lean @@ -136,17 +136,11 @@ theorem ringHom_eval₂_intCastRingHom {R S : Type*} [Ring R] [Ring S] (p : ℤ[ (r : R) : f (eval₂ (Int.castRingHom R) r p) = eval₂ (Int.castRingHom S) (f r) p := algHom_eval₂_algebraMap p f.toIntAlgHom r -@[deprecated (since := "2024-05-27")] -alias ringHom_eval₂_cast_int_ringHom := ringHom_eval₂_intCastRingHom - @[simp] theorem eval₂_intCastRingHom_X {R : Type*} [Ring R] (p : ℤ[X]) (f : ℤ[X] →+* R) : eval₂ (Int.castRingHom R) (f X) p = f p := eval₂_algebraMap_X p f.toIntAlgHom -@[deprecated (since := "2024-04-17")] -alias eval₂_int_castRingHom_X := eval₂_intCastRingHom_X - /-- `Polynomial.eval₂` as an `AlgHom` for noncommutative algebras. This is `Polynomial.eval₂RingHom'` for `AlgHom`s. -/ @@ -268,9 +262,6 @@ theorem aeval_one : aeval x (1 : R[X]) = 1 := theorem aeval_natCast (n : ℕ) : aeval x (n : R[X]) = n := map_natCast _ _ -@[deprecated (since := "2024-04-17")] -alias aeval_nat_cast := aeval_natCast - theorem aeval_mul : aeval x (p * q) = aeval x p * aeval x q := map_mul _ _ _ diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index 880597eb13c1f..6ed5fbfe49788 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -475,9 +475,6 @@ theorem C_pow : C (a ^ n) = C a ^ n := theorem C_eq_natCast (n : ℕ) : C (n : R) = (n : R[X]) := map_natCast C n -@[deprecated (since := "2024-04-17")] -alias C_eq_nat_cast := C_eq_natCast - @[simp] theorem C_mul_monomial : C a * monomial n b = monomial n (a * b) := by simp only [← monomial_zero_left, monomial_mul_monomial, zero_add] @@ -655,9 +652,6 @@ lemma coeff_C_succ {r : R} {n : ℕ} : coeff (C r) (n + 1) = 0 := by simp [coeff theorem coeff_natCast_ite : (Nat.cast m : R[X]).coeff n = ite (n = 0) m 0 := by simp only [← C_eq_natCast, coeff_C, Nat.cast_ite, Nat.cast_zero] -@[deprecated (since := "2024-04-17")] -alias coeff_nat_cast_ite := coeff_natCast_ite - @[simp] theorem coeff_ofNat_zero (a : ℕ) [a.AtLeastTwo] : coeff (ofNat(a) : R[X]) 0 = ofNat(a) := @@ -837,9 +831,6 @@ theorem binomial_eq_binomial {k l m n : ℕ} {u v : R} (hu : u ≠ 0) (hv : v theorem natCast_mul (n : ℕ) (p : R[X]) : (n : R[X]) * p = n • p := (nsmul_eq_mul _ _).symm -@[deprecated (since := "2024-04-17")] -alias nat_cast_mul := natCast_mul - /-- Summing the values of a function applied to the coefficients of a polynomial -/ def sum {S : Type*} [AddCommMonoid S] (p : R[X]) (f : ℕ → R → S) : S := ∑ n ∈ p.support, f n (p.coeff n) @@ -1097,9 +1088,6 @@ theorem support_neg {p : R[X]} : (-p).support = p.support := by theorem C_eq_intCast (n : ℤ) : C (n : R) = n := by simp -@[deprecated (since := "2024-04-17")] -alias C_eq_int_cast := C_eq_intCast - theorem C_neg : C (-a) = -C a := RingHom.map_neg C a diff --git a/Mathlib/Algebra/Polynomial/Coeff.lean b/Mathlib/Algebra/Polynomial/Coeff.lean index de59da07ebfc2..4d5288226c18d 100644 --- a/Mathlib/Algebra/Polynomial/Coeff.lean +++ b/Mathlib/Algebra/Polynomial/Coeff.lean @@ -342,9 +342,6 @@ section cast theorem natCast_coeff_zero {n : ℕ} {R : Type*} [Semiring R] : (n : R[X]).coeff 0 = n := by simp only [coeff_natCast_ite, ite_true] -@[deprecated (since := "2024-04-17")] -alias nat_cast_coeff_zero := natCast_coeff_zero - @[norm_cast] theorem natCast_inj {m n : ℕ} {R : Type*} [Semiring R] [CharZero R] : (↑m : R[X]) = ↑n ↔ m = n := by @@ -355,16 +352,10 @@ theorem natCast_inj {m n : ℕ} {R : Type*} [Semiring R] [CharZero R] : · rintro rfl rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_inj := natCast_inj - @[simp] theorem intCast_coeff_zero {i : ℤ} {R : Type*} [Ring R] : (i : R[X]).coeff 0 = i := by cases i <;> simp -@[deprecated (since := "2024-04-17")] -alias int_cast_coeff_zero := intCast_coeff_zero - @[norm_cast] theorem intCast_inj {m n : ℤ} {R : Type*} [Ring R] [CharZero R] : (↑m : R[X]) = ↑n ↔ m = n := by constructor @@ -374,9 +365,6 @@ theorem intCast_inj {m n : ℤ} {R : Type*} [Ring R] [CharZero R] : (↑m : R[X] · rintro rfl rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_inj := intCast_inj - end cast instance charZero [CharZero R] : CharZero R[X] where cast_injective _x _y := natCast_inj.mp diff --git a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean index a7f16ee7d9aa6..e9124609802e8 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean @@ -174,9 +174,6 @@ theorem natDegree_one : natDegree (1 : R[X]) = 0 := theorem natDegree_natCast (n : ℕ) : natDegree (n : R[X]) = 0 := by simp only [← C_eq_natCast, natDegree_C] -@[deprecated (since := "2024-04-17")] -alias natDegree_nat_cast := natDegree_natCast - @[simp] theorem natDegree_ofNat (n : ℕ) [Nat.AtLeastTwo n] : natDegree (ofNat(n) : R[X]) = 0 := @@ -184,9 +181,6 @@ theorem natDegree_ofNat (n : ℕ) [Nat.AtLeastTwo n] : theorem degree_natCast_le (n : ℕ) : degree (n : R[X]) ≤ 0 := degree_le_of_natDegree_le (by simp) -@[deprecated (since := "2024-04-17")] -alias degree_nat_cast_le := degree_natCast_le - @[simp] theorem degree_monomial (n : ℕ) (ha : a ≠ 0) : degree (monomial n a) = n := by rw [degree, support_monomial n ha, max_singleton, Nat.cast_withBot] @@ -287,14 +281,8 @@ theorem natDegree_neg_le_of_le {p : R[X]} (hp : natDegree p ≤ m) : natDegree ( theorem natDegree_intCast (n : ℤ) : natDegree (n : R[X]) = 0 := by rw [← C_eq_intCast, natDegree_C] -@[deprecated (since := "2024-04-17")] -alias natDegree_int_cast := natDegree_intCast - theorem degree_intCast_le (n : ℤ) : degree (n : R[X]) ≤ 0 := degree_le_of_natDegree_le (by simp) -@[deprecated (since := "2024-04-17")] -alias degree_int_cast_le := degree_intCast_le - @[simp] theorem leadingCoeff_neg (p : R[X]) : (-p).leadingCoeff = -p.leadingCoeff := by rw [leadingCoeff, leadingCoeff, natDegree_neg, coeff_neg] diff --git a/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean b/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean index 062fb2042f27c..571099425f307 100644 --- a/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean +++ b/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean @@ -199,9 +199,6 @@ theorem natTrailingDegree_one : natTrailingDegree (1 : R[X]) = 0 := theorem natTrailingDegree_natCast (n : ℕ) : natTrailingDegree (n : R[X]) = 0 := by simp only [← C_eq_natCast, natTrailingDegree_C] -@[deprecated (since := "2024-04-17")] -alias natTrailingDegree_nat_cast := natTrailingDegree_natCast - @[simp] theorem trailingDegree_C_mul_X_pow (n : ℕ) (ha : a ≠ 0) : trailingDegree (C a * X ^ n) = n := by rw [C_mul_X_pow_eq_monomial, trailingDegree_monomial ha] @@ -395,9 +392,6 @@ theorem natTrailingDegree_neg (p : R[X]) : natTrailingDegree (-p) = natTrailingD theorem natTrailingDegree_intCast (n : ℤ) : natTrailingDegree (n : R[X]) = 0 := by simp only [← C_eq_intCast, natTrailingDegree_C] -@[deprecated (since := "2024-04-17")] -alias natTrailingDegree_int_cast := natTrailingDegree_intCast - end Ring section Semiring @@ -456,9 +450,6 @@ lemma eq_X_pow_iff_natTrailingDegree_eq_natDegree (h₁ : p.Monic) : p = X ^ p.natDegree ↔ p.natTrailingDegree = p.natDegree := h₁.eq_X_pow_iff_natDegree_le_natTrailingDegree.trans (natTrailingDegree_le_natDegree p).ge_iff_eq -@[deprecated (since := "2024-04-26")] -alias ⟨_, eq_X_pow_of_natTrailingDegree_eq_natDegree⟩ := eq_X_pow_iff_natTrailingDegree_eq_natDegree - end Monic end Semiring diff --git a/Mathlib/Algebra/Polynomial/Derivative.lean b/Mathlib/Algebra/Polynomial/Derivative.lean index e0154c45d7370..c5d69bf7dd0b7 100644 --- a/Mathlib/Algebra/Polynomial/Derivative.lean +++ b/Mathlib/Algebra/Polynomial/Derivative.lean @@ -190,9 +190,6 @@ theorem derivative_natCast {n : ℕ} : derivative (n : R[X]) = 0 := by rw [← map_natCast C n] exact derivative_C -@[deprecated (since := "2024-04-17")] -alias derivative_nat_cast := derivative_natCast - @[simp] theorem derivative_ofNat (n : ℕ) [n.AtLeastTwo] : derivative (ofNat(n) : R[X]) = 0 := @@ -286,17 +283,11 @@ theorem derivative_natCast_mul {n : ℕ} {f : R[X]} : derivative ((n : R[X]) * f) = n * derivative f := by simp -@[deprecated (since := "2024-04-17")] -alias derivative_nat_cast_mul := derivative_natCast_mul - @[simp] theorem iterate_derivative_natCast_mul {n k : ℕ} {f : R[X]} : derivative^[k] ((n : R[X]) * f) = n * derivative^[k] f := by induction' k with k ih generalizing f <;> simp [*] -@[deprecated (since := "2024-04-17")] -alias iterate_derivative_nat_cast_mul := iterate_derivative_natCast_mul - theorem mem_support_derivative [NoZeroSMulDivisors ℕ R] (p : R[X]) (n : ℕ) : n ∈ (derivative p).support ↔ n + 1 ∈ p.support := by suffices ¬p.coeff (n + 1) * (n + 1 : ℕ) = 0 ↔ coeff p (n + 1) ≠ 0 by @@ -496,9 +487,6 @@ theorem iterate_derivative_X_pow_eq_natCast_mul (n k : ℕ) : Nat.descFactorial_succ, Nat.sub_sub, Nat.cast_mul] simp [mul_comm, mul_assoc, mul_left_comm] -@[deprecated (since := "2024-04-17")] -alias iterate_derivative_X_pow_eq_nat_cast_mul := iterate_derivative_X_pow_eq_natCast_mul - theorem iterate_derivative_X_pow_eq_C_mul (n k : ℕ) : derivative^[k] (X ^ n : R[X]) = C (Nat.descFactorial n k : R) * X ^ (n - k) := by rw [iterate_derivative_X_pow_eq_natCast_mul n k, C_eq_natCast] @@ -587,24 +575,15 @@ theorem derivative_intCast {n : ℤ} : derivative (n : R[X]) = 0 := by rw [← C_eq_intCast n] exact derivative_C -@[deprecated (since := "2024-04-17")] -alias derivative_int_cast := derivative_intCast - theorem derivative_intCast_mul {n : ℤ} {f : R[X]} : derivative ((n : R[X]) * f) = n * derivative f := by simp -@[deprecated (since := "2024-04-17")] -alias derivative_int_cast_mul := derivative_intCast_mul - @[simp] theorem iterate_derivative_intCast_mul {n : ℤ} {k : ℕ} {f : R[X]} : derivative^[k] ((n : R[X]) * f) = n * derivative^[k] f := by induction' k with k ih generalizing f <;> simp [*] -@[deprecated (since := "2024-04-17")] -alias iterate_derivative_int_cast_mul := iterate_derivative_intCast_mul - end Ring section CommRing diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 6d1a1c3d96070..1047af9fab54a 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -434,9 +434,6 @@ theorem mul_divByMonic_cancel_left (p : R[X]) {q : R[X]} (hmo : q.Monic) : rw [degree_zero] exact Ne.bot_lt fun h => hmo.ne_zero (degree_eq_bot.1 h) -@[deprecated (since := "2024-06-30")] -alias mul_div_mod_by_monic_cancel_left := mul_divByMonic_cancel_left - lemma coeff_divByMonic_X_sub_C_rec (p : R[X]) (a : R) (n : ℕ) : (p /ₘ (X - C a)).coeff n = coeff p (n + 1) + a * (p /ₘ (X - C a)).coeff (n + 1) := by nontriviality R diff --git a/Mathlib/Algebra/Polynomial/Eval/Coeff.lean b/Mathlib/Algebra/Polynomial/Eval/Coeff.lean index cd0d1add2ccd8..9e3de4a5c384d 100644 --- a/Mathlib/Algebra/Polynomial/Eval/Coeff.lean +++ b/Mathlib/Algebra/Polynomial/Eval/Coeff.lean @@ -157,9 +157,6 @@ theorem eval_natCast_map (f : R →+* S) (p : R[X]) (n : ℕ) : | h_monomial n r => simp only [map_natCast f, eval_monomial, map_monomial, f.map_pow, f.map_mul] -@[deprecated (since := "2024-04-17")] -alias eval_nat_cast_map := eval_natCast_map - @[simp] theorem eval_intCast_map {R S : Type*} [Ring R] [Ring S] (f : R →+* S) (p : R[X]) (i : ℤ) : (p.map f).eval (i : S) = f (p.eval i) := by @@ -169,9 +166,6 @@ theorem eval_intCast_map {R S : Type*} [Ring R] [Ring S] (f : R →+* S) (p : R[ | h_monomial n r => simp only [map_intCast, eval_monomial, map_monomial, map_pow, map_mul] -@[deprecated (since := "2024-04-17")] -alias eval_int_cast_map := eval_intCast_map - end Map section HomEval₂ diff --git a/Mathlib/Algebra/Polynomial/Eval/Defs.lean b/Mathlib/Algebra/Polynomial/Eval/Defs.lean index 1cb42ee480336..4a88165f66aed 100644 --- a/Mathlib/Algebra/Polynomial/Eval/Defs.lean +++ b/Mathlib/Algebra/Polynomial/Eval/Defs.lean @@ -100,9 +100,6 @@ theorem eval₂_natCast (n : ℕ) : (n : R[X]).eval₂ f x = n := by | zero => simp only [eval₂_zero, Nat.cast_zero] | succ n ih => rw [n.cast_succ, eval₂_add, ih, eval₂_one, n.cast_succ] -@[deprecated (since := "2024-04-17")] -alias eval₂_nat_cast := eval₂_natCast - @[simp] lemma eval₂_ofNat {S : Type*} [Semiring S] (n : ℕ) [n.AtLeastTwo] (f : R →+* S) (a : S) : (ofNat(n) : R[X]).eval₂ f a = ofNat(n) := by @@ -260,9 +257,6 @@ theorem eval₂_at_natCast {S : Type*} [Semiring S] (f : R →+* S) (n : ℕ) : convert eval₂_at_apply (p := p) f n simp -@[deprecated (since := "2024-04-17")] -alias eval₂_at_nat_cast := eval₂_at_natCast - @[simp] theorem eval₂_at_ofNat {S : Type*} [Semiring S] (f : R →+* S) (n : ℕ) [n.AtLeastTwo] : p.eval₂ f ofNat(n) = f (p.eval (ofNat(n))) := by @@ -275,9 +269,6 @@ theorem eval_C : (C a).eval x = a := @[simp] theorem eval_natCast {n : ℕ} : (n : R[X]).eval x = n := by simp only [← C_eq_natCast, eval_C] -@[deprecated (since := "2024-04-17")] -alias eval_nat_cast := eval_natCast - @[simp] lemma eval_ofNat (n : ℕ) [n.AtLeastTwo] (a : R) : (ofNat(n) : R[X]).eval a = ofNat(n) := by @@ -315,9 +306,6 @@ theorem eval_C_mul : (C a * p).eval x = a * p.eval x := by theorem eval_natCast_mul {n : ℕ} : ((n : R[X]) * p).eval x = n * p.eval x := by rw [← C_eq_natCast, eval_C_mul] -@[deprecated (since := "2024-04-17")] -alias eval_nat_cast_mul := eval_natCast_mul - @[simp] theorem eval_mul_X : (p * X).eval x = p.eval x * x := by induction p using Polynomial.induction_on' with @@ -392,9 +380,6 @@ theorem C_comp : (C a).comp p = C a := @[simp] theorem natCast_comp {n : ℕ} : (n : R[X]).comp p = n := by rw [← C_eq_natCast, C_comp] -@[deprecated (since := "2024-04-17")] -alias nat_cast_comp := natCast_comp - @[simp] theorem ofNat_comp (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : R[X]).comp p = n := natCast_comp @@ -451,16 +436,10 @@ theorem C_mul_comp : (C a * p).comp r = C a * p.comp r := by theorem natCast_mul_comp {n : ℕ} : ((n : R[X]) * p).comp r = n * p.comp r := by rw [← C_eq_natCast, C_mul_comp] -@[deprecated (since := "2024-04-17")] -alias nat_cast_mul_comp := natCast_mul_comp - theorem mul_X_add_natCast_comp {n : ℕ} : (p * (X + (n : R[X]))).comp q = p.comp q * (q + n) := by rw [mul_add, add_comp, mul_X_comp, ← Nat.cast_comm, natCast_mul_comp, Nat.cast_comm, mul_add] -@[deprecated (since := "2024-04-17")] -alias mul_X_add_nat_cast_comp := mul_X_add_natCast_comp - @[simp] theorem mul_comp {R : Type*} [CommSemiring R] (p q r : R[X]) : (p * q).comp r = p.comp r * q.comp r := @@ -545,9 +524,6 @@ theorem coe_mapRingHom (f : R →+* S) : ⇑(mapRingHom f) = map f := protected theorem map_natCast (n : ℕ) : (n : R[X]).map f = n := map_natCast (mapRingHom f) n -@[deprecated (since := "2024-04-17")] -alias map_nat_cast := map_natCast - @[simp] protected theorem map_ofNat (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : R[X]).map f = ofNat(n) := @@ -738,16 +714,10 @@ protected theorem map_neg {S} [Ring S] (f : R →+* S) : (-p).map f = -p.map f : @[simp] protected lemma map_intCast {S} [Ring S] (f : R →+* S) (n : ℤ) : map f ↑n = ↑n := map_intCast (mapRingHom f) n -@[deprecated (since := "2024-04-17")] -alias map_int_cast := map_intCast - @[simp] theorem eval_intCast {n : ℤ} {x : R} : (n : R[X]).eval x = n := by simp only [← C_eq_intCast, eval_C] -@[deprecated (since := "2024-04-17")] -alias eval_int_cast := eval_intCast - @[simp] theorem eval₂_neg {S} [Ring S] (f : R →+* S) {x : S} : (-p).eval₂ f x = -p.eval₂ f x := by rw [eq_neg_iff_add_eq_zero, ← eval₂_add, neg_add_cancel, eval₂_zero] @@ -779,24 +749,16 @@ theorem sub_comp : (p - q).comp r = p.comp r - q.comp r := @[simp] theorem intCast_comp (i : ℤ) : comp (i : R[X]) p = i := by cases i <;> simp -@[deprecated (since := "2024-05-27")] alias cast_int_comp := intCast_comp - @[simp] theorem eval₂_at_intCast {S : Type*} [Ring S] (f : R →+* S) (n : ℤ) : p.eval₂ f n = f (p.eval n) := by convert eval₂_at_apply (p := p) f n simp -@[deprecated (since := "2024-04-17")] -alias eval₂_at_int_cast := eval₂_at_intCast - theorem mul_X_sub_intCast_comp {n : ℕ} : (p * (X - (n : R[X]))).comp q = p.comp q * (q - n) := by rw [mul_sub, sub_comp, mul_X_comp, ← Nat.cast_comm, natCast_mul_comp, Nat.cast_comm, mul_sub] -@[deprecated (since := "2024-04-17")] -alias mul_X_sub_int_cast_comp := mul_X_sub_intCast_comp - end Ring end Polynomial diff --git a/Mathlib/Algebra/Polynomial/Smeval.lean b/Mathlib/Algebra/Polynomial/Smeval.lean index ad2b3ce21ccc3..058efa23b2789 100644 --- a/Mathlib/Algebra/Polynomial/Smeval.lean +++ b/Mathlib/Algebra/Polynomial/Smeval.lean @@ -113,9 +113,6 @@ theorem smeval_natCast (n : ℕ) : (n : R[X]).smeval x = n • x ^ 0 := by | zero => simp only [smeval_zero, Nat.cast_zero, zero_smul] | succ n ih => rw [n.cast_succ, smeval_add, ih, smeval_one, ← add_nsmul] -@[deprecated (since := "2024-04-17")] -alias smeval_nat_cast := smeval_natCast - @[simp] theorem smeval_smul (r : R) : (r • p).smeval x = r • p.smeval x := by induction p using Polynomial.induction_on' with @@ -203,9 +200,6 @@ theorem smeval_at_natCast (q : ℕ[X]) : ∀(n : ℕ), q.smeval (n : S) = q.smev intro n rw [smeval_monomial, smeval_monomial, nsmul_eq_mul, smul_eq_mul, Nat.cast_mul, Nat.cast_npow] -@[deprecated (since := "2024-04-17")] -alias smeval_at_nat_cast := smeval_at_natCast - theorem smeval_at_zero : p.smeval (0 : S) = (p.coeff 0) • (1 : S) := by induction p using Polynomial.induction_on' with | h_add p q ph qh => diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index bfc6a75772dcc..6a4475fd7cc95 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -401,44 +401,26 @@ instance : AddCommGroupWithOne ℍ[R,c₁,c₂,c₃] where theorem natCast_re (n : ℕ) : (n : ℍ[R,c₁,c₂,c₃]).re = n := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_re := natCast_re - @[simp, norm_cast] theorem natCast_imI (n : ℕ) : (n : ℍ[R,c₁,c₂,c₃]).imI = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_imI := natCast_imI - @[simp, norm_cast] theorem natCast_imJ (n : ℕ) : (n : ℍ[R,c₁,c₂,c₃]).imJ = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_imJ := natCast_imJ - @[simp, norm_cast] theorem natCast_imK (n : ℕ) : (n : ℍ[R,c₁,c₂,c₃]).imK = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_imK := natCast_imK - @[simp, norm_cast] theorem natCast_im (n : ℕ) : (n : ℍ[R,c₁,c₂,c₃]).im = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_im := natCast_im - @[norm_cast] theorem coe_natCast (n : ℕ) : ↑(n : R) = (n : ℍ[R,c₁,c₂,c₃]) := rfl -@[deprecated (since := "2024-04-17")] -alias coe_nat_cast := coe_natCast - @[simp, norm_cast] theorem intCast_re (z : ℤ) : (z : ℍ[R,c₁,c₂,c₃]).re = z := rfl @@ -458,44 +440,26 @@ theorem ofNat_imK (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℍ[R,c₁,c₂,c₃]). @[simp] theorem ofNat_im (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℍ[R,c₁,c₂,c₃]).im = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_re := intCast_re - @[simp, norm_cast] theorem intCast_imI (z : ℤ) : (z : ℍ[R,c₁,c₂,c₃]).imI = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_imI := intCast_imI - @[simp, norm_cast] theorem intCast_imJ (z : ℤ) : (z : ℍ[R,c₁,c₂,c₃]).imJ = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_imJ := intCast_imJ - @[simp, norm_cast] theorem intCast_imK (z : ℤ) : (z : ℍ[R,c₁,c₂,c₃]).imK = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_imK := intCast_imK - @[simp, norm_cast] theorem intCast_im (z : ℤ) : (z : ℍ[R,c₁,c₂,c₃]).im = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_im := intCast_im - @[norm_cast] theorem coe_intCast (z : ℤ) : ↑(z : R) = (z : ℍ[R,c₁,c₂,c₃]) := rfl -@[deprecated (since := "2024-04-17")] -alias coe_int_cast := coe_intCast - end AddCommGroupWithOne -- For the remainder of the file we assume `CommRing R`. @@ -952,75 +916,39 @@ theorem coe_pow (n : ℕ) : (↑(x ^ n) : ℍ[R]) = (x : ℍ[R]) ^ n := @[simp, norm_cast] theorem natCast_re (n : ℕ) : (n : ℍ[R]).re = n := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_re := natCast_re - @[simp, norm_cast] theorem natCast_imI (n : ℕ) : (n : ℍ[R]).imI = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_imI := natCast_imI - @[simp, norm_cast] theorem natCast_imJ (n : ℕ) : (n : ℍ[R]).imJ = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_imJ := natCast_imJ - @[simp, norm_cast] theorem natCast_imK (n : ℕ) : (n : ℍ[R]).imK = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_imK := natCast_imK - @[simp, norm_cast] theorem natCast_im (n : ℕ) : (n : ℍ[R]).im = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_im := natCast_im - @[norm_cast] theorem coe_natCast (n : ℕ) : ↑(n : R) = (n : ℍ[R]) := rfl -@[deprecated (since := "2024-04-17")] -alias coe_nat_cast := coe_natCast - @[simp, norm_cast] theorem intCast_re (z : ℤ) : (z : ℍ[R]).re = z := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_re := intCast_re - @[simp, norm_cast] theorem intCast_imI (z : ℤ) : (z : ℍ[R]).imI = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_imI := intCast_imI - @[simp, norm_cast] theorem intCast_imJ (z : ℤ) : (z : ℍ[R]).imJ = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_imJ := intCast_imJ - @[simp, norm_cast] theorem intCast_imK (z : ℤ) : (z : ℍ[R]).imK = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_imK := intCast_imK - @[simp, norm_cast] theorem intCast_im (z : ℤ) : (z : ℍ[R]).im = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_im := intCast_im - @[norm_cast] theorem coe_intCast (z : ℤ) : ↑(z : R) = (z : ℍ[R]) := rfl -@[deprecated (since := "2024-04-17")] -alias coe_int_cast := coe_intCast - theorem coe_injective : Function.Injective (coe : R → ℍ[R]) := QuaternionAlgebra.coe_injective @@ -1175,16 +1103,10 @@ theorem normSq_star : normSq (star a) = normSq a := by simp [normSq_def'] theorem normSq_natCast (n : ℕ) : normSq (n : ℍ[R]) = (n : R) ^ 2 := by rw [← coe_natCast, normSq_coe] -@[deprecated (since := "2024-04-17")] -alias normSq_nat_cast := normSq_natCast - @[norm_cast] theorem normSq_intCast (z : ℤ) : normSq (z : ℍ[R]) = (z : R) ^ 2 := by rw [← coe_intCast, normSq_coe] -@[deprecated (since := "2024-04-17")] -alias normSq_int_cast := normSq_intCast - @[simp] theorem normSq_neg : normSq (-a) = normSq a := by simp only [normSq_def, star_neg, neg_mul_neg] @@ -1303,22 +1225,10 @@ instance instRatCast : RatCast ℍ[R] where ratCast q := (q : R) @[simp, norm_cast] lemma ratCast_imJ (q : ℚ) : (q : ℍ[R]).imJ = 0 := rfl @[simp, norm_cast] lemma ratCast_imK (q : ℚ) : (q : ℍ[R]).imK = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias rat_cast_imI := ratCast_imI - -@[deprecated (since := "2024-04-17")] -alias rat_cast_imJ := ratCast_imJ - -@[deprecated (since := "2024-04-17")] -alias rat_cast_imK := ratCast_imK - @[norm_cast] lemma coe_nnratCast (q : ℚ≥0) : ↑(q : R) = (q : ℍ[R]) := rfl @[norm_cast] lemma coe_ratCast (q : ℚ) : ↑(q : R) = (q : ℍ[R]) := rfl -@[deprecated (since := "2024-04-17")] -alias coe_rat_cast := coe_ratCast - instance instDivisionRing : DivisionRing ℍ[R] where __ := Quaternion.instGroupWithZero __ := Quaternion.instRing @@ -1342,9 +1252,6 @@ theorem normSq_zpow (z : ℤ) : normSq (a ^ z) = normSq a ^ z := theorem normSq_ratCast (q : ℚ) : normSq (q : ℍ[R]) = (q : ℍ[R]) ^ 2 := by rw [← coe_ratCast, normSq_coe, coe_pow] -@[deprecated (since := "2024-04-17")] -alias normSq_rat_cast := normSq_ratCast - end Field end Quaternion diff --git a/Mathlib/Algebra/Ring/CentroidHom.lean b/Mathlib/Algebra/Ring/CentroidHom.lean index b1b90bb34b9d1..fce16a8613db8 100644 --- a/Mathlib/Algebra/Ring/CentroidHom.lean +++ b/Mathlib/Algebra/Ring/CentroidHom.lean @@ -337,15 +337,9 @@ instance : NatCast (CentroidHom α) where natCast n := n • (1 : CentroidHom α theorem coe_natCast (n : ℕ) : ⇑(n : CentroidHom α) = n • (CentroidHom.id α) := rfl -@[deprecated (since := "2024-04-17")] -alias coe_nat_cast := coe_natCast - theorem natCast_apply (n : ℕ) (m : α) : (n : CentroidHom α) m = n • m := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_apply := natCast_apply - @[simp] theorem toEnd_one : (1 : CentroidHom α).toEnd = 1 := rfl @@ -362,9 +356,6 @@ theorem toEnd_pow (x : CentroidHom α) (n : ℕ) : (x ^ n).toEnd = x.toEnd ^ n : theorem toEnd_natCast (n : ℕ) : (n : CentroidHom α).toEnd = ↑n := rfl -@[deprecated (since := "2024-04-17")] -alias toEnd_nat_cast := toEnd_natCast - -- cf `add_monoid.End.semiring` instance : Semiring (CentroidHom α) := toEnd_injective.semiring _ toEnd_zero toEnd_one toEnd_add toEnd_mul toEnd_smul toEnd_pow @@ -574,15 +565,9 @@ instance : IntCast (CentroidHom α) where intCast z := z • (1 : CentroidHom α theorem coe_intCast (z : ℤ) : ⇑(z : CentroidHom α) = z • (CentroidHom.id α) := rfl -@[deprecated (since := "2024-04-17")] -alias coe_int_cast := coe_intCast - theorem intCast_apply (z : ℤ) (m : α) : (z : CentroidHom α) m = z • m := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_apply := intCast_apply - @[simp] theorem toEnd_neg (x : CentroidHom α) : (-x).toEnd = -x.toEnd := rfl @@ -615,9 +600,6 @@ theorem sub_apply (f g : CentroidHom α) (a : α) : (f - g) a = f a - g a := theorem toEnd_intCast (z : ℤ) : (z : CentroidHom α).toEnd = ↑z := rfl -@[deprecated (since := "2024-04-17")] -alias toEnd_int_cast := toEnd_intCast - instance instRing : Ring (CentroidHom α) := toEnd_injective.ring _ toEnd_zero toEnd_one toEnd_add toEnd_mul toEnd_neg toEnd_sub toEnd_smul toEnd_smul toEnd_pow toEnd_natCast toEnd_intCast diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index ed97a0457e825..f5b3128adee62 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -56,8 +56,6 @@ end DivisionMonoid @[simp] lemma IsSquare.zero [MulZeroClass α] : IsSquare (0 : α) := ⟨0, (mul_zero _).symm⟩ -@[deprecated (since := "2024-01-07")] alias isSquare_zero := IsSquare.zero - section Semiring variable [Semiring α] [Semiring β] {a b : α} {m n : ℕ} diff --git a/Mathlib/Algebra/Ring/Rat.lean b/Mathlib/Algebra/Ring/Rat.lean index 26cd561224df0..3435bf0effdfb 100644 --- a/Mathlib/Algebra/Ring/Rat.lean +++ b/Mathlib/Algebra/Ring/Rat.lean @@ -94,6 +94,4 @@ lemma natCast_eq_divInt (n : ℕ) : ↑n = n /. 1 := by rw [← Int.cast_natCast @[simp] lemma den_mul_eq_num (q : ℚ) : q.den * q = q.num := by rw [mul_comm, mul_den_eq_num] -@[deprecated (since := "2024-04-07")] alias coe_nat_eq_divInt := natCast_eq_divInt - end Rat diff --git a/Mathlib/Algebra/Ring/Subring/Defs.lean b/Mathlib/Algebra/Ring/Subring/Defs.lean index 69b8b9aab8fb8..c99262c124f95 100644 --- a/Mathlib/Algebra/Ring/Subring/Defs.lean +++ b/Mathlib/Algebra/Ring/Subring/Defs.lean @@ -88,8 +88,6 @@ variable [SetLike S R] [hSR : SubringClass S R] (s : S) @[aesop safe apply (rule_sets := [SetLike])] theorem intCast_mem (n : ℤ) : (n : R) ∈ s := by simp only [← zsmul_one, zsmul_mem, one_mem] -@[deprecated _root_.intCast_mem (since := "2024-04-05")] alias coe_int_mem := intCast_mem - namespace SubringClass instance (priority := 75) toHasIntCast : IntCast s := diff --git a/Mathlib/Algebra/Ring/Subsemiring/Defs.lean b/Mathlib/Algebra/Ring/Subsemiring/Defs.lean index d2114e093146e..3a6eaaf2abc0b 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Defs.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Defs.lean @@ -28,8 +28,6 @@ variable {S R : Type*} [AddMonoidWithOne R] [SetLike S R] (s : S) theorem natCast_mem [AddSubmonoidWithOneClass S R] (n : ℕ) : (n : R) ∈ s := by induction n <;> simp [zero_mem, add_mem, one_mem, *] -@[deprecated (since := "2024-04-05")] alias coe_nat_mem := natCast_mem - @[aesop safe apply (rule_sets := [SetLike])] lemma ofNat_mem [AddSubmonoidWithOneClass S R] (s : S) (n : ℕ) [n.AtLeastTwo] : ofNat(n) ∈ s := by diff --git a/Mathlib/Algebra/Squarefree/Basic.lean b/Mathlib/Algebra/Squarefree/Basic.lean index 6096792244d94..d058b461dda33 100644 --- a/Mathlib/Algebra/Squarefree/Basic.lean +++ b/Mathlib/Algebra/Squarefree/Basic.lean @@ -281,6 +281,4 @@ theorem squarefree_natAbs {n : ℤ} : Squarefree n.natAbs ↔ Squarefree n := by theorem squarefree_natCast {n : ℕ} : Squarefree (n : ℤ) ↔ Squarefree n := by rw [← squarefree_natAbs, natAbs_ofNat] -@[deprecated (since := "2024-04-05")] alias squarefree_coe_nat := squarefree_natCast - end Int diff --git a/Mathlib/Algebra/Star/Module.lean b/Mathlib/Algebra/Star/Module.lean index adbbbc109e362..dd4fc7cc5f809 100644 --- a/Mathlib/Algebra/Star/Module.lean +++ b/Mathlib/Algebra/Star/Module.lean @@ -40,41 +40,26 @@ theorem star_natCast_smul [Semiring R] [AddCommMonoid M] [Module R M] [StarAddMo (x : M) : star ((n : R) • x) = (n : R) • star x := map_natCast_smul (starAddEquiv : M ≃+ M) R R n x -@[deprecated (since := "2024-04-17")] -alias star_nat_cast_smul := star_natCast_smul - @[simp] theorem star_intCast_smul [Ring R] [AddCommGroup M] [Module R M] [StarAddMonoid M] (n : ℤ) (x : M) : star ((n : R) • x) = (n : R) • star x := map_intCast_smul (starAddEquiv : M ≃+ M) R R n x -@[deprecated (since := "2024-04-17")] -alias star_int_cast_smul := star_intCast_smul - @[simp] theorem star_inv_natCast_smul [DivisionSemiring R] [AddCommMonoid M] [Module R M] [StarAddMonoid M] (n : ℕ) (x : M) : star ((n⁻¹ : R) • x) = (n⁻¹ : R) • star x := map_inv_natCast_smul (starAddEquiv : M ≃+ M) R R n x -@[deprecated (since := "2024-04-17")] -alias star_inv_nat_cast_smul := star_inv_natCast_smul - @[simp] theorem star_inv_intCast_smul [DivisionRing R] [AddCommGroup M] [Module R M] [StarAddMonoid M] (n : ℤ) (x : M) : star ((n⁻¹ : R) • x) = (n⁻¹ : R) • star x := map_inv_intCast_smul (starAddEquiv : M ≃+ M) R R n x -@[deprecated (since := "2024-04-17")] -alias star_inv_int_cast_smul := star_inv_intCast_smul - @[simp] theorem star_ratCast_smul [DivisionRing R] [AddCommGroup M] [Module R M] [StarAddMonoid M] (n : ℚ) (x : M) : star ((n : R) • x) = (n : R) • star x := map_ratCast_smul (starAddEquiv : M ≃+ M) _ _ _ x -@[deprecated (since := "2024-04-17")] -alias star_rat_cast_smul := star_ratCast_smul - /-! Per the naming convention, these two lemmas call `(q • ·)` `nnrat_smul` and `rat_smul` respectively, rather than `nnqsmul` and `qsmul` because the latter are reserved to the actions coming from diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index 027817e6fef82..310e881988587 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -491,23 +491,14 @@ instance addMonoidWithOne [AddMonoidWithOne R] [AddMonoid M] : AddMonoidWithOne theorem fst_natCast [AddMonoidWithOne R] [AddMonoid M] (n : ℕ) : (n : tsze R M).fst = n := rfl -@[deprecated (since := "2024-04-17")] -alias fst_nat_cast := fst_natCast - @[simp] theorem snd_natCast [AddMonoidWithOne R] [AddMonoid M] (n : ℕ) : (n : tsze R M).snd = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias snd_nat_cast := snd_natCast - @[simp] theorem inl_natCast [AddMonoidWithOne R] [AddMonoid M] (n : ℕ) : (inl n : tsze R M) = n := rfl -@[deprecated (since := "2024-04-17")] -alias inl_nat_cast := inl_natCast - instance addGroupWithOne [AddGroupWithOne R] [AddGroup M] : AddGroupWithOne (tsze R M) := { TrivSqZeroExt.addGroup, TrivSqZeroExt.addMonoidWithOne with intCast := fun z => inl z @@ -518,23 +509,14 @@ instance addGroupWithOne [AddGroupWithOne R] [AddGroup M] : AddGroupWithOne (tsz theorem fst_intCast [AddGroupWithOne R] [AddGroup M] (z : ℤ) : (z : tsze R M).fst = z := rfl -@[deprecated (since := "2024-04-17")] -alias fst_int_cast := fst_intCast - @[simp] theorem snd_intCast [AddGroupWithOne R] [AddGroup M] (z : ℤ) : (z : tsze R M).snd = 0 := rfl -@[deprecated (since := "2024-04-17")] -alias snd_int_cast := snd_intCast - @[simp] theorem inl_intCast [AddGroupWithOne R] [AddGroup M] (z : ℤ) : (inl z : tsze R M) = z := rfl -@[deprecated (since := "2024-04-17")] -alias inl_int_cast := inl_intCast - instance nonAssocSemiring [Semiring R] [AddCommMonoid M] [Module R M] [Module Rᵐᵒᵖ M] : NonAssocSemiring (tsze R M) := { TrivSqZeroExt.addMonoidWithOne, TrivSqZeroExt.mulOneClass, TrivSqZeroExt.addCommMonoid with diff --git a/Mathlib/Algebra/Vertex/HVertexOperator.lean b/Mathlib/Algebra/Vertex/HVertexOperator.lean index c9461bf6789ec..421058df2ff18 100644 --- a/Mathlib/Algebra/Vertex/HVertexOperator.lean +++ b/Mathlib/Algebra/Vertex/HVertexOperator.lean @@ -53,8 +53,6 @@ open HahnModule theorem ext (A B : HVertexOperator Γ R V W) (h : ∀ v : V, A v = B v) : A = B := LinearMap.ext h -@[deprecated (since := "2024-06-18")] alias _root_.VertexAlg.HetVertexOperator.ext := ext - /-- The coefficient of a heterogeneous vertex operator, viewed as a formal power series with coefficients in linear maps. -/ @[simps] @@ -64,23 +62,16 @@ def coeff (A : HVertexOperator Γ R V W) (n : Γ) : V →ₗ[R] W where map_smul' _ _ := by simp only [map_smul, RingHom.id_apply, of_symm_smul, HahnSeries.coeff_smul] -@[deprecated (since := "2024-06-18")] alias _root_.VertexAlg.coeff := coeff - theorem coeff_isPWOsupport (A : HVertexOperator Γ R V W) (v : V) : ((of R).symm (A v)).coeff.support.IsPWO := ((of R).symm (A v)).isPWO_support' -@[deprecated (since := "2024-06-18")] -alias _root_.VertexAlg.coeff_isPWOsupport := coeff_isPWOsupport - @[ext] theorem coeff_inj : Function.Injective (coeff : HVertexOperator Γ R V W → Γ → (V →ₗ[R] W)) := by intro _ _ h ext v n exact congrFun (congrArg DFunLike.coe (congrFun h n)) v -@[deprecated (since := "2024-06-18")] alias _root_.VertexAlg.coeff_inj := coeff_inj - /-- Given a coefficient function valued in linear maps satisfying a partially well-ordered support condition, we produce a heterogeneous vertex operator. -/ @[simps] @@ -90,8 +81,6 @@ def of_coeff (f : Γ → V →ₗ[R] W) map_add' _ _ := by ext; simp map_smul' _ _ := by ext; simp -@[deprecated (since := "2024-06-18")] alias _root_.VertexAlg.HetVertexOperator.of_coeff := of_coeff - @[simp] theorem coeff_add (A B : HVertexOperator Γ R V W) : (A + B).coeff = A.coeff + B.coeff := by ext diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 4772bc08cc2a9..cc8afa1e32204 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -1178,38 +1178,4 @@ def Scheme.arrowStalkMapSpecIso {R S : CommRingCat.{u}} (f : R ⟶ S) (p : Prime simp end Stalks - -@[deprecated (since := "2024-06-21"), nolint defLemma] -alias isAffineAffineScheme := isAffine_affineScheme -@[deprecated (since := "2024-06-21"), nolint defLemma] -alias SpecIsAffine := isAffine_Spec -@[deprecated (since := "2024-06-21")] -alias isAffineOfIso := isAffine_of_isIso -@[deprecated (since := "2024-06-21")] -alias rangeIsAffineOpenOfOpenImmersion := isAffineOpen_opensRange -@[deprecated (since := "2024-06-21")] -alias topIsAffineOpen := isAffineOpen_top -@[deprecated (since := "2024-06-21"), nolint defLemma] -alias Scheme.affineCoverIsAffine := Scheme.isAffine_affineCover -@[deprecated (since := "2024-06-21"), nolint defLemma] -alias Scheme.affineBasisCoverIsAffine := Scheme.isAffine_affineBasisCover -@[deprecated (since := "2024-06-21")] -alias IsAffineOpen.fromSpec_range := IsAffineOpen.range_fromSpec -@[deprecated (since := "2024-06-21")] -alias IsAffineOpen.imageIsOpenImmersion := IsAffineOpen.image_of_isOpenImmersion -@[deprecated (since := "2024-06-21"), nolint defLemma] -alias Scheme.quasi_compact_of_affine := Scheme.compactSpace_of_isAffine -@[deprecated (since := "2024-06-21")] -alias IsAffineOpen.fromSpec_base_preimage := IsAffineOpen.fromSpec_preimage_self -@[deprecated (since := "2024-06-21")] -alias IsAffineOpen.fromSpec_map_basicOpen' := IsAffineOpen.fromSpec_preimage_basicOpen' -@[deprecated (since := "2024-06-21")] -alias IsAffineOpen.fromSpec_map_basicOpen := IsAffineOpen.fromSpec_preimage_basicOpen -@[deprecated (since := "2024-06-21")] -alias IsAffineOpen.opensFunctor_map_basicOpen := IsAffineOpen.fromSpec_image_basicOpen -@[deprecated (since := "2024-06-21")] -alias IsAffineOpen.basicOpenIsAffine := IsAffineOpen.basicOpen -@[deprecated (since := "2024-06-21")] -alias IsAffineOpen.mapRestrictBasicOpen := IsAffineOpen.ι_basicOpen_preimage - end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Cover/Open.lean b/Mathlib/AlgebraicGeometry/Cover/Open.lean index 49feb36d291ee..a6ec6c2082cb2 100644 --- a/Mathlib/AlgebraicGeometry/Cover/Open.lean +++ b/Mathlib/AlgebraicGeometry/Cover/Open.lean @@ -32,7 +32,6 @@ namespace Scheme /-- An open cover of a scheme `X` is a cover where all component maps are open immersions. -/ abbrev OpenCover (X : Scheme.{u}) : Type _ := Cover.{v} @IsOpenImmersion X -@[deprecated (since := "2024-06-23")] alias OpenCover.Covers := Cover.covers @[deprecated (since := "2024-11-06")] alias OpenCover.IsOpen := Cover.map_prop variable {X Y Z : Scheme.{u}} (𝒰 : OpenCover X) (f : X ⟶ Z) (g : Y ⟶ Z) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index 97e4ff84b8c19..159080b8f8eae 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -234,13 +234,6 @@ lemma evalEval_polynomialY (x y : R) : lemma evalEval_polynomialY_zero : W.polynomialY.evalEval 0 0 = W.a₃ := by simp only [evalEval_polynomialY, zero_add, mul_zero] -@[deprecated (since := "2024-06-19")] alias eval_polynomial := evalEval_polynomial -@[deprecated (since := "2024-06-19")] alias eval_polynomial_zero := evalEval_polynomial_zero -@[deprecated (since := "2024-06-19")] alias eval_polynomialX := evalEval_polynomialX -@[deprecated (since := "2024-06-19")] alias eval_polynomialX_zero := evalEval_polynomialX_zero -@[deprecated (since := "2024-06-19")] alias eval_polynomialY := evalEval_polynomialY -@[deprecated (since := "2024-06-19")] alias eval_polynomialY_zero := evalEval_polynomialY_zero - /-- The proposition that an affine point $(x, y)$ in `W` is nonsingular. In other words, either $W_X(x, y) \ne 0$ or $W_Y(x, y) \ne 0$. @@ -699,15 +692,6 @@ lemma add_of_X_ne' {x₁ x₂ y₁ y₂ : F} {h₁ : W.Nonsingular x₁ y₁} {h (hx : x₁ ≠ x₂) : some h₁ + some h₂ = -some (nonsingular_negAdd h₁ h₂ fun h => (hx h).elim) := add_of_X_ne hx -@[deprecated (since := "2024-06-03")] alias some_add_some_of_Yeq := add_of_Y_eq -@[deprecated (since := "2024-06-03")] alias some_add_self_of_Yeq := add_self_of_Y_eq -@[deprecated (since := "2024-06-03")] alias some_add_some_of_Yne := add_of_Y_ne -@[deprecated (since := "2024-06-03")] alias some_add_some_of_Yne' := add_of_Y_ne' -@[deprecated (since := "2024-06-03")] alias some_add_self_of_Yne := add_self_of_Y_ne -@[deprecated (since := "2024-06-03")] alias some_add_self_of_Yne' := add_self_of_Y_ne' -@[deprecated (since := "2024-06-03")] alias some_add_some_of_Xne := add_of_X_ne -@[deprecated (since := "2024-06-03")] alias some_add_some_of_Xne' := add_of_X_ne' - end Point end Group @@ -926,20 +910,6 @@ end Point end BaseChange -@[deprecated (since := "2024-06-03")] alias addY' := negAddY -@[deprecated (since := "2024-06-03")] alias - nonsingular_add_of_eval_derivative_ne_zero := nonsingular_negAdd_of_eval_derivative_ne_zero -@[deprecated (since := "2024-06-03")] alias slope_of_Yeq := slope_of_Y_eq -@[deprecated (since := "2024-06-03")] alias slope_of_Yne := slope_of_Y_ne -@[deprecated (since := "2024-06-03")] alias slope_of_Xne := slope_of_X_ne -@[deprecated (since := "2024-06-03")] alias slope_of_Yne_eq_eval := slope_of_Y_ne_eq_eval -@[deprecated (since := "2024-06-03")] alias Yeq_of_Xeq := Y_eq_of_X_eq -@[deprecated (since := "2024-06-03")] alias Yeq_of_Yne := Y_eq_of_Y_ne -@[deprecated (since := "2024-06-03")] alias equation_add' := equation_negAdd -@[deprecated (since := "2024-06-03")] alias nonsingular_add' := nonsingular_negAdd -@[deprecated (since := "2024-06-03")] alias baseChange_addY' := baseChange_negAddY -@[deprecated (since := "2024-06-03")] alias map_addY' := map_negAddY - /-! ## Elliptic curves -/ section EllipticCurve diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 30a706d61de6e..b897c89abc30f 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -477,15 +477,6 @@ theorem ΓSpecIso_obj_hom {X : Scheme.{u}} (U : X.Opens) : (Scheme.ΓSpecIso Γ(X, U)).hom = (Spec.map U.topIso.inv).appTop ≫ U.toScheme.toSpecΓ.appTop ≫ U.topIso.hom := by simp -@[deprecated (since := "2024-07-24")] -alias ΓSpec.adjunction_unit_naturality := Scheme.toSpecΓ_naturality -@[deprecated (since := "2024-07-24")] -alias ΓSpec.adjunction_unit_naturality_assoc := Scheme.toSpecΓ_naturality_assoc -@[deprecated (since := "2024-07-24")] -alias ΓSpec.adjunction_unit_app_app_top := Scheme.toSpecΓ_appTop -@[deprecated (since := "2024-07-24")] -alias ΓSpec.adjunction_unit_map_basicOpen := Scheme.toSpecΓ_preimage_basicOpen - /-! Immediate consequences of the adjunction. -/ @@ -579,8 +570,4 @@ instance : Reflective Spec.toLocallyRingedSpace where instance Spec.reflective : Reflective Scheme.Spec where adj := ΓSpec.adjunction -@[deprecated (since := "2024-07-02")] -alias LocallyRingedSpace.toΓSpec_preim_basicOpen_eq := - LocallyRingedSpace.toΓSpec_preimage_basicOpen_eq - end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean index 070c60d31cc77..5b5e0f2bcc6d8 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean @@ -338,11 +338,6 @@ theorem cancel_right_of_respectsIso P (f ≫ g) ↔ P f := by rw [← P.toProperty_apply, ← P.toProperty_apply, P.toProperty.cancel_right_of_respectsIso] -@[deprecated (since := "2024-07-02")] alias affine_cancel_left_isIso := - AffineTargetMorphismProperty.cancel_left_of_respectsIso -@[deprecated (since := "2024-07-02")] alias affine_cancel_right_isIso := - AffineTargetMorphismProperty.cancel_right_of_respectsIso - theorem arrow_mk_iso_iff (P : AffineTargetMorphismProperty) [P.toProperty.RespectsIso] {X Y X' Y' : Scheme} {f : X ⟶ Y} {f' : X' ⟶ Y'} diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean index 7a34d9b050ddd..3cd9d1cbcdddb 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean @@ -330,18 +330,4 @@ lemma isStableUnderBaseChange_of_isStableUnderBaseChangeOnAffine_of_isLocalAtTar end AffineTargetMorphismProperty -@[deprecated (since := "2024-06-22")] -alias diagonalTargetAffineLocallyOfOpenCover := HasAffineProperty.diagonal_of_openCover - -@[deprecated (since := "2024-06-22")] -alias AffineTargetMorphismProperty.diagonalOfTargetAffineLocally := - HasAffineProperty.diagonal_of_diagonal_of_isPullback - -@[deprecated (since := "2024-06-22")] -alias universallyIsLocalAtTarget := universally_isLocalAtTarget - -@[deprecated (since := "2024-06-22")] -alias universallyIsLocalAtTargetOfMorphismRestrict := - universally_isLocalAtTarget - end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean index 7fc0bc4d9218a..a2e93bb4b4bc8 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean @@ -269,8 +269,6 @@ theorem stalkToFiberRingHom_germ (U : Opens (ProjectiveSpectrum.top 𝒜)) stalkToFiberRingHom 𝒜 x ((Proj.structureSheaf 𝒜).presheaf.germ _ x hx s) = s.1 ⟨x, hx⟩ := RingHom.ext_iff.1 (CommRingCat.hom_ext_iff.mp (germ_comp_stalkToFiberRingHom 𝒜 U x hx)) s -@[deprecated (since := "2024-07-30")] alias stalkToFiberRingHom_germ' := stalkToFiberRingHom_germ - theorem mem_basicOpen_den (x : ProjectiveSpectrum.top 𝒜) (f : HomogeneousLocalization.NumDenSameDeg 𝒜 x.asHomogeneousIdeal.toIdeal.primeCompl) : x ∈ ProjectiveSpectrum.basicOpen 𝒜 f.den := by @@ -358,8 +356,6 @@ theorem Proj.stalkIso'_germ (U : Opens (ProjectiveSpectrum.top 𝒜)) Proj.stalkIso' 𝒜 x ((Proj.structureSheaf 𝒜).presheaf.germ _ x hx s) = s.1 ⟨x, hx⟩ := stalkToFiberRingHom_germ 𝒜 U x hx s -@[deprecated (since := "2024-07-30")] alias Proj.stalkIso'_germ' := Proj.stalkIso'_germ - @[simp] theorem Proj.stalkIso'_symm_mk (x) (f) : (Proj.stalkIso' 𝒜 x).symm (.mk f) = (Proj.structureSheaf 𝒜).presheaf.germ _ diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index 1ce2fe0fcee26..d82bc0979eb00 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -176,9 +176,6 @@ def Scheme.openCoverOfISupEqTop {s : Type*} (X : Scheme.{u}) (U : s → X.Opens) have : x ∈ ⨆ i, U i := hU.symm ▸ show x ∈ (⊤ : X.Opens) by trivial exact (Opens.mem_iSup.mp this).choose_spec -@[deprecated (since := "2024-07-24")] -noncomputable alias Scheme.openCoverOfSuprEqTop := Scheme.openCoverOfISupEqTop - /-- The open sets of an open subscheme corresponds to the open sets containing in the subset. -/ @[simps!] def opensRestrict : diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index fdd3809d8e81f..ef49d929d8d24 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -290,9 +290,6 @@ theorem comp_appTop {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : (f ≫ g).appTop = g.appTop ≫ f.appTop := rfl -@[deprecated (since := "2024-06-23")] alias comp_val_c_app := comp_app -@[deprecated (since := "2024-06-23")] alias comp_val_c_app_assoc := comp_app_assoc - theorem appLE_comp_appLE {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) (U V W e₁ e₂) : g.appLE U V e₁ ≫ f.appLE V W e₂ = (f ≫ g).appLE U W (e₂.trans ((Opens.map f.base).map (homOfLE e₁)).le) := by diff --git a/Mathlib/AlgebraicGeometry/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/StructureSheaf.lean index 9e79bc6c7f8b9..9a5af37277991 100644 --- a/Mathlib/AlgebraicGeometry/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/StructureSheaf.lean @@ -418,8 +418,6 @@ theorem toOpen_Γgerm_apply (x : PrimeSpectrum.Top R) (f : R) : (structureSheaf R).presheaf.Γgerm x (toOpen R ⊤ f) = toStalk R x f := rfl -@[deprecated (since := "2024-07-30")] alias germ_to_top := toOpen_Γgerm_apply - theorem isUnit_to_basicOpen_self (f : R) : IsUnit (toOpen R (PrimeSpectrum.basicOpen f) f) := isUnit_of_mul_eq_one _ (const R 1 f (PrimeSpectrum.basicOpen f) fun _ => id) <| by rw [toOpen_eq_const, const_mul_rev] @@ -499,8 +497,6 @@ theorem stalkToFiberRingHom_germ (U : Opens (PrimeSpectrum.Top R)) stalkToFiberRingHom R x ((structureSheaf R).presheaf.germ U x hx s) = s.1 ⟨x, hx⟩ := RingHom.ext_iff.mp (CommRingCat.hom_ext_iff.mp (germ_comp_stalkToFiberRingHom R U x hx)) s -@[deprecated (since := "2024-07-30")] alias stalkToFiberRingHom_germ' := stalkToFiberRingHom_germ - @[simp] theorem toStalk_comp_stalkToFiberRingHom (x : PrimeSpectrum.Top R) : toStalk R x ≫ stalkToFiberRingHom R x = CommRingCat.ofHom (algebraMap _ _) := by diff --git a/Mathlib/Analysis/Asymptotics/Lemmas.lean b/Mathlib/Analysis/Asymptotics/Lemmas.lean index a42c822104f4a..b70736a7ea6c5 100644 --- a/Mathlib/Analysis/Asymptotics/Lemmas.lean +++ b/Mathlib/Analysis/Asymptotics/Lemmas.lean @@ -626,17 +626,11 @@ theorem IsBigO.natCast_atTop {R : Type*} [StrictOrderedSemiring R] [Archimedean (fun (n : ℕ) => f n) =O[atTop] (fun n => g n) := IsBigO.comp_tendsto h tendsto_natCast_atTop_atTop -@[deprecated (since := "2024-04-17")] -alias IsBigO.nat_cast_atTop := IsBigO.natCast_atTop - theorem IsLittleO.natCast_atTop {R : Type*} [StrictOrderedSemiring R] [Archimedean R] {f : R → E} {g : R → F} (h : f =o[atTop] g) : (fun (n : ℕ) => f n) =o[atTop] (fun n => g n) := IsLittleO.comp_tendsto h tendsto_natCast_atTop_atTop -@[deprecated (since := "2024-04-17")] -alias IsLittleO.nat_cast_atTop := IsLittleO.natCast_atTop - theorem isBigO_atTop_iff_eventually_exists {α : Type*} [SemilatticeSup α] [Nonempty α] {f : α → E} {g : α → F} : f =O[atTop] g ↔ ∀ᶠ n₀ in atTop, ∃ c, ∀ n ≥ n₀, ‖f n‖ ≤ c * ‖g n‖ := by rw [isBigO_iff, exists_eventually_atTop] diff --git a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean index 2463324f04481..3af2409bab158 100644 --- a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean +++ b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean @@ -225,16 +225,10 @@ theorem one_toProd : (1 : 𝓜(𝕜, A)).toProd = 1 := theorem natCast_toProd (n : ℕ) : (n : 𝓜(𝕜, A)).toProd = n := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_toProd := natCast_toProd - @[simp] theorem intCast_toProd (n : ℤ) : (n : 𝓜(𝕜, A)).toProd = n := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_toProd := intCast_toProd - @[simp] theorem pow_toProd (n : ℕ) (a : 𝓜(𝕜, A)) : (a ^ n).toProd = a.toProd ^ n := rfl @@ -280,27 +274,15 @@ theorem mul_snd (a b : 𝓜(𝕜, A)) : (a * b).snd = b.snd * a.snd := theorem natCast_fst (n : ℕ) : (n : 𝓜(𝕜, A)).fst = n := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_fst := natCast_fst - theorem natCast_snd (n : ℕ) : (n : 𝓜(𝕜, A)).snd = n := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_snd := natCast_snd - theorem intCast_fst (n : ℤ) : (n : 𝓜(𝕜, A)).fst = n := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_fst := intCast_fst - theorem intCast_snd (n : ℤ) : (n : 𝓜(𝕜, A)).snd = n := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_snd := intCast_snd - theorem pow_fst (n : ℕ) (a : 𝓜(𝕜, A)) : (a ^ n).fst = a.fst ^ n := rfl diff --git a/Mathlib/Analysis/Calculus/FDeriv/Add.lean b/Mathlib/Analysis/Calculus/FDeriv/Add.lean index 2677296a22c2e..81b1e2cd9bb26 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Add.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Add.lean @@ -615,29 +615,14 @@ theorem DifferentiableAt.sub_const (hf : DifferentiableAt 𝕜 f x) (c : F) : DifferentiableAt 𝕜 (fun y => f y - c) x := (hf.hasFDerivAt.sub_const c).differentiableAt -@[deprecated DifferentiableAt.sub_iff_left (since := "2024-07-11")] -theorem differentiableAt_sub_const_iff (c : F) : - DifferentiableAt 𝕜 (fun y => f y - c) x ↔ DifferentiableAt 𝕜 f x := - (differentiableAt_const _).sub_iff_left - @[fun_prop] theorem DifferentiableOn.sub_const (hf : DifferentiableOn 𝕜 f s) (c : F) : DifferentiableOn 𝕜 (fun y => f y - c) s := fun x hx => (hf x hx).sub_const c -@[deprecated DifferentiableOn.sub_iff_left (since := "2024-07-11")] -theorem differentiableOn_sub_const_iff (c : F) : - DifferentiableOn 𝕜 (fun y => f y - c) s ↔ DifferentiableOn 𝕜 f s := - (differentiableOn_const _).sub_iff_left - @[fun_prop] theorem Differentiable.sub_const (hf : Differentiable 𝕜 f) (c : F) : Differentiable 𝕜 fun y => f y - c := fun x => (hf x).sub_const c -@[deprecated Differentiable.sub_iff_left (since := "2024-07-11")] -theorem differentiable_sub_const_iff (c : F) : - (Differentiable 𝕜 fun y => f y - c) ↔ Differentiable 𝕜 f := - (differentiable_const _).sub_iff_left - theorem fderivWithin_sub_const (hxs : UniqueDiffWithinAt 𝕜 s x) (c : F) : fderivWithin 𝕜 (fun y => f y - c) s x = fderivWithin 𝕜 f s x := by simp only [sub_eq_add_neg, fderivWithin_add_const hxs] @@ -679,29 +664,14 @@ theorem DifferentiableAt.const_sub (hf : DifferentiableAt 𝕜 f x) (c : F) : DifferentiableAt 𝕜 (fun y => c - f y) x := (hf.hasFDerivAt.const_sub c).differentiableAt -@[deprecated DifferentiableAt.sub_iff_right (since := "2024-07-11")] -theorem differentiableAt_const_sub_iff (c : F) : - DifferentiableAt 𝕜 (fun y => c - f y) x ↔ DifferentiableAt 𝕜 f x := - (differentiableAt_const _).sub_iff_right - @[fun_prop] theorem DifferentiableOn.const_sub (hf : DifferentiableOn 𝕜 f s) (c : F) : DifferentiableOn 𝕜 (fun y => c - f y) s := fun x hx => (hf x hx).const_sub c -@[deprecated DifferentiableOn.sub_iff_right (since := "2024-07-11")] -theorem differentiableOn_const_sub_iff (c : F) : - DifferentiableOn 𝕜 (fun y => c - f y) s ↔ DifferentiableOn 𝕜 f s := - (differentiableOn_const _).sub_iff_right - @[fun_prop] theorem Differentiable.const_sub (hf : Differentiable 𝕜 f) (c : F) : Differentiable 𝕜 fun y => c - f y := fun x => (hf x).const_sub c -@[deprecated Differentiable.sub_iff_right (since := "2024-07-11")] -theorem differentiable_const_sub_iff (c : F) : - (Differentiable 𝕜 fun y => c - f y) ↔ Differentiable 𝕜 f := - (differentiable_const _).sub_iff_right - theorem fderivWithin_const_sub (hxs : UniqueDiffWithinAt 𝕜 s x) (c : F) : fderivWithin 𝕜 (fun y => c - f y) s x = -fderivWithin 𝕜 f s x := by simp only [sub_eq_add_neg, fderivWithin_const_add, fderivWithin_neg, hxs] diff --git a/Mathlib/Analysis/Calculus/FDeriv/Extend.lean b/Mathlib/Analysis/Calculus/FDeriv/Extend.lean index c83f208cc9e1a..60fdacfba0cad 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Extend.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Extend.lean @@ -100,9 +100,6 @@ theorem hasFDerivWithinAt_closure_of_tendsto_fderiv {f : E → F} {s : Set E} {x tendsto_const_nhds.mul (Tendsto.comp continuous_norm.continuousAt <| tendsto_snd.sub tendsto_fst) -@[deprecated (since := "2024-07-10")] alias has_fderiv_at_boundary_of_tendsto_fderiv := - hasFDerivWithinAt_closure_of_tendsto_fderiv - /-- If a function is differentiable on the right of a point `a : ℝ`, continuous at `a`, and its derivative also converges at `a`, then `f` is differentiable on the right at `a`. -/ theorem hasDerivWithinAt_Ici_of_tendsto_deriv {s : Set ℝ} {e : E} {a : ℝ} {f : ℝ → E} @@ -137,9 +134,6 @@ theorem hasDerivWithinAt_Ici_of_tendsto_deriv {s : Set ℝ} {e : E} {a : ℝ} {f exact hasFDerivWithinAt_closure_of_tendsto_fderiv t_diff t_conv t_open t_cont t_diff' exact this.mono_of_mem_nhdsWithin (Icc_mem_nhdsGE ab) -@[deprecated (since := "2024-07-10")] alias has_deriv_at_interval_left_endpoint_of_tendsto_deriv := - hasDerivWithinAt_Ici_of_tendsto_deriv - /-- If a function is differentiable on the left of a point `a : ℝ`, continuous at `a`, and its derivative also converges at `a`, then `f` is differentiable on the left at `a`. -/ theorem hasDerivWithinAt_Iic_of_tendsto_deriv {s : Set ℝ} {e : E} {a : ℝ} @@ -175,9 +169,6 @@ theorem hasDerivWithinAt_Iic_of_tendsto_deriv {s : Set ℝ} {e : E} {a : ℝ} exact hasFDerivWithinAt_closure_of_tendsto_fderiv t_diff t_conv t_open t_cont t_diff' exact this.mono_of_mem_nhdsWithin (Icc_mem_nhdsLE ba) -@[deprecated (since := "2024-07-10")] alias has_deriv_at_interval_right_endpoint_of_tendsto_deriv := - hasDerivWithinAt_Iic_of_tendsto_deriv - /-- If a real function `f` has a derivative `g` everywhere but at a point, and `f` and `g` are continuous at this point, then `g` is also the derivative of `f` at this point. -/ theorem hasDerivAt_of_hasDerivAt_of_ne {f g : ℝ → E} {x : ℝ} diff --git a/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean b/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean index 4538d75d4267a..ba440227664f5 100644 --- a/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean +++ b/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean @@ -102,9 +102,6 @@ theorem mem_posTangentConeAt_of_segment_subset (h : [x -[ℝ] x + y] ⊆ s) : rw [segment_eq_image', add_sub_cancel_left] exact mem_image_of_mem _ ⟨le_of_lt ht₀, ht₁⟩ -@[deprecated (since := "2024-07-13")] -- cleanup docstrings when we drop this alias -alias mem_posTangentConeAt_of_segment_subset' := mem_posTangentConeAt_of_segment_subset - theorem sub_mem_posTangentConeAt_of_segment_subset (h : segment ℝ x y ⊆ s) : y - x ∈ posTangentConeAt s x := mem_posTangentConeAt_of_segment_subset <| by rwa [add_sub_cancel] diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index 3d4217f58338e..382e0690c6bdf 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -689,9 +689,6 @@ lemma zero_not_mem_slitPlane : 0 ∉ slitPlane := mt ofReal_mem_slitPlane.1 (lt_ lemma natCast_mem_slitPlane {n : ℕ} : ↑n ∈ slitPlane ↔ n ≠ 0 := by simpa [pos_iff_ne_zero] using @ofReal_mem_slitPlane n -@[deprecated (since := "2024-04-17")] -alias nat_cast_mem_slitPlane := natCast_mem_slitPlane - @[simp] lemma ofNat_mem_slitPlane (n : ℕ) [n.AtLeastTwo] : ofNat(n) ∈ slitPlane := natCast_mem_slitPlane.2 (NeZero.ne n) diff --git a/Mathlib/Analysis/Complex/Circle.lean b/Mathlib/Analysis/Complex/Circle.lean index e4a9a2e283278..a1c4c1c117e8c 100644 --- a/Mathlib/Analysis/Complex/Circle.lean +++ b/Mathlib/Analysis/Complex/Circle.lean @@ -38,23 +38,6 @@ open Complex Function Metric open ComplexConjugate -/-- The unit circle in `ℂ`, here given the structure of a submonoid of `ℂ`. - -Please use `Circle` when referring to the circle as a type. -/ -@[deprecated "No deprecation message was provided." (since := "2024-07-24")] -def circle : Submonoid ℂ := - Submonoid.unitSphere ℂ - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-07-24")] -theorem mem_circle_iff_abs {z : ℂ} : z ∈ circle ↔ abs z = 1 := - mem_sphere_zero_iff_norm - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-07-24")] -theorem mem_circle_iff_normSq {z : ℂ} : z ∈ circle ↔ normSq z = 1 := by - simp [Complex.abs, mem_circle_iff_abs] - /-- The unit circle in `ℂ`. -/ def Circle : Type := Submonoid.unitSphere ℂ deriving TopologicalSpace @@ -94,13 +77,6 @@ def coeHom : Circle →* ℂ where map_one' := coe_one map_mul' := coe_mul -@[deprecated (since := "2024-07-24")] alias _root_.abs_coe_circle := abs_coe -@[deprecated (since := "2024-07-24")] alias _root_.normSq_eq_of_mem_circle := normSq_coe -@[deprecated (since := "2024-07-24")] alias _root_.ne_zero_of_mem_circle := coe_ne_zero -@[deprecated (since := "2024-07-24")] alias _root_.coe_inv_circle := coe_inv -@[deprecated (since := "2024-07-24")] alias _root_.coe_inv_circle_eq_conj := coe_inv_eq_conj -@[deprecated (since := "2024-07-24")] alias _root_.coe_div_circle := coe_div - /-- The elements of the circle embed into the units. -/ def toUnits : Circle →* Units ℂ := unitSphereToUnits ℂ @@ -179,11 +155,4 @@ protected lemma norm_smul {E : Type*} [SeminormedAddCommGroup E] [NormedSpace ‖u • v‖ = ‖v‖ := by rw [Submonoid.smul_def, norm_smul, norm_eq_of_mem_sphere, one_mul] -@[deprecated (since := "2024-07-24")] noncomputable alias _root_.expMapCircle := exp -@[deprecated (since := "2024-07-24")] noncomputable alias _root_.expMapCircle_apply := coe_exp -@[deprecated (since := "2024-07-24")] noncomputable alias _root_.expMapCircle_zero := exp_zero -@[deprecated (since := "2024-07-24")] noncomputable alias _root_.expMapCircle_sub := exp_sub -@[deprecated (since := "2024-07-24")] noncomputable alias _root_.norm_circle_smul := - Circle.norm_smul - end Circle diff --git a/Mathlib/Analysis/Convex/Integral.lean b/Mathlib/Analysis/Convex/Integral.lean index eb9be5b719ffe..829db0eb9ae61 100644 --- a/Mathlib/Analysis/Convex/Integral.lean +++ b/Mathlib/Analysis/Convex/Integral.lean @@ -337,7 +337,3 @@ theorem ae_eq_const_or_norm_setIntegral_lt_of_norm_le_const [StrictConvexSpace haveI := Fact.mk ht.lt_top rw [← restrict_apply_univ] exact ae_eq_const_or_norm_integral_lt_of_norm_le_const h_le - -@[deprecated (since := "2024-04-17")] -alias ae_eq_const_or_norm_set_integral_lt_of_norm_le_const := - ae_eq_const_or_norm_setIntegral_lt_of_norm_le_const diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 1a638bc818025..b521bb8650ec2 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -452,9 +452,6 @@ theorem eLpNorm_le_eLpNorm_fderiv_one {u : E → F} (hu : ContDiff ℝ 1 u) (h2 inv_mul_cancel₀ h0p.ne', NNReal.rpow_one] exact lintegral_pow_le_pow_lintegral_fderiv μ hu h2u hp.coe -@[deprecated (since := "2024-07-27")] -alias snorm_le_snorm_fderiv_one := eLpNorm_le_eLpNorm_fderiv_one - /-- The constant factor occurring in the conclusion of `eLpNorm_le_eLpNorm_fderiv_of_eq_inner`. It only depends on `E`, `μ` and `p`. -/ def eLpNormLESNormFDerivOfEqInnerConst (p : ℝ) : ℝ≥0 := @@ -590,9 +587,6 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq_inner {u : E → F'} left positivity -@[deprecated (since := "2024-07-27")] -alias snorm_le_snorm_fderiv_of_eq_inner := eLpNorm_le_eLpNorm_fderiv_of_eq_inner - variable (F) in /-- The constant factor occurring in the conclusion of `eLpNorm_le_eLpNorm_fderiv_of_eq`. It only depends on `E`, `F`, `μ` and `p`. -/ @@ -647,9 +641,6 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq [FiniteDimensional ℝ F] push_cast simp_rw [mul_assoc] -@[deprecated (since := "2024-07-27")] -alias snorm_le_snorm_fderiv_of_eq := eLpNorm_le_eLpNorm_fderiv_of_eq - variable (F) in /-- The constant factor occurring in the conclusion of `eLpNorm_le_eLpNorm_fderiv_of_le`. @@ -718,9 +709,6 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_le [FiniteDimensional ℝ F] _ = eLpNormLESNormFDerivOfLeConst F μ s p q * eLpNorm (fderiv ℝ u) p μ := by simp_rw [eLpNormLESNormFDerivOfLeConst, ENNReal.coe_mul]; ring -@[deprecated (since := "2024-07-27")] -alias snorm_le_snorm_fderiv_of_le := eLpNorm_le_eLpNorm_fderiv_of_le - /-- The **Gagliardo-Nirenberg-Sobolev inequality**. Let `u` be a continuously differentiable function `u` supported in a bounded set `s` in a normed space `E` of finite dimension `n`, equipped with Haar measure, and let `1 < p < n`. @@ -738,7 +726,4 @@ theorem eLpNorm_le_eLpNorm_fderiv [FiniteDimensional ℝ F] simp only [tsub_le_iff_right, le_add_iff_nonneg_right] positivity -@[deprecated (since := "2024-07-27")] -alias snorm_le_snorm_fderiv := eLpNorm_le_eLpNorm_fderiv - end MeasureTheory diff --git a/Mathlib/Analysis/Normed/Group/AddCircle.lean b/Mathlib/Analysis/Normed/Group/AddCircle.lean index 8cd44a44c9c69..ec64ee09eb697 100644 --- a/Mathlib/Analysis/Normed/Group/AddCircle.lean +++ b/Mathlib/Analysis/Normed/Group/AddCircle.lean @@ -229,9 +229,6 @@ theorem norm_div_natCast {m n : ℕ} : have : p⁻¹ * (↑m / ↑n * p) = ↑m / ↑n := by rw [mul_comm _ p, inv_mul_cancel_left₀ hp.out.ne.symm] rw [norm_eq' p hp.out, this, abs_sub_round_div_natCast_eq] -@[deprecated (since := "2024-04-17")] -alias norm_div_nat_cast := norm_div_natCast - theorem exists_norm_eq_of_isOfFinAddOrder {u : AddCircle p} (hu : IsOfFinAddOrder u) : ∃ k : ℕ, ‖u‖ = p * (k / addOrderOf u) := by let n := addOrderOf u diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index d773e1f72c4f5..d2f4094dcf547 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -1298,9 +1298,6 @@ theorem le_norm_self (r : ℝ) : r ≤ ‖r‖ := @[simp 1100] lemma nnnorm_natCast (n : ℕ) : ‖(n : ℝ)‖₊ = n := NNReal.eq <| norm_natCast _ @[simp 1100] lemma enorm_natCast (n : ℕ) : ‖(n : ℝ)‖ₑ = n := by simp [enorm] -@[deprecated (since := "2024-04-05")] alias norm_coe_nat := norm_natCast -@[deprecated (since := "2024-04-05")] alias nnnorm_coe_nat := nnnorm_natCast - @[simp 1100] lemma norm_ofNat (n : ℕ) [n.AtLeastTwo] : ‖(ofNat(n) : ℝ)‖ = ofNat(n) := norm_natCast n diff --git a/Mathlib/Analysis/Normed/Group/Int.lean b/Mathlib/Analysis/Normed/Group/Int.lean index 61239bc61172c..da0a26bda53f3 100644 --- a/Mathlib/Analysis/Normed/Group/Int.lean +++ b/Mathlib/Analysis/Normed/Group/Int.lean @@ -28,8 +28,6 @@ theorem norm_eq_abs (n : ℤ) : ‖n‖ = |(n : ℝ)| := @[simp] theorem norm_natCast (n : ℕ) : ‖(n : ℤ)‖ = n := by simp [Int.norm_eq_abs] -@[deprecated (since := "2024-04-05")] alias norm_coe_nat := norm_natCast - theorem _root_.NNReal.natCast_natAbs (n : ℤ) : (n.natAbs : ℝ≥0) = ‖n‖₊ := NNReal.eq <| calc diff --git a/Mathlib/Analysis/Normed/Lp/lpSpace.lean b/Mathlib/Analysis/Normed/Lp/lpSpace.lean index 356ca2b7ac0dc..4d0398250852a 100644 --- a/Mathlib/Analysis/Normed/Lp/lpSpace.lean +++ b/Mathlib/Analysis/Normed/Lp/lpSpace.lean @@ -802,15 +802,9 @@ theorem _root_.Memℓp.infty_pow {f : ∀ i, B i} (hf : Memℓp f ∞) (n : ℕ) theorem _root_.natCast_memℓp_infty (n : ℕ) : Memℓp (n : ∀ i, B i) ∞ := natCast_mem (lpInftySubring B) n -@[deprecated (since := "2024-04-17")] -alias _root_.nat_cast_memℓp_infty := _root_.natCast_memℓp_infty - theorem _root_.intCast_memℓp_infty (z : ℤ) : Memℓp (z : ∀ i, B i) ∞ := intCast_mem (lpInftySubring B) z -@[deprecated (since := "2024-04-17")] -alias _root_.int_cast_memℓp_infty := _root_.intCast_memℓp_infty - @[simp] theorem infty_coeFn_one : ⇑(1 : lp B ∞) = 1 := rfl @@ -823,16 +817,10 @@ theorem infty_coeFn_pow (f : lp B ∞) (n : ℕ) : ⇑(f ^ n) = (⇑f) ^ n := theorem infty_coeFn_natCast (n : ℕ) : ⇑(n : lp B ∞) = n := rfl -@[deprecated (since := "2024-04-17")] -alias infty_coeFn_nat_cast := infty_coeFn_natCast - @[simp] theorem infty_coeFn_intCast (z : ℤ) : ⇑(z : lp B ∞) = z := rfl -@[deprecated (since := "2024-04-17")] -alias infty_coeFn_int_cast := infty_coeFn_intCast - instance [Nonempty I] : NormOneClass (lp B ∞) where norm_one := by simp_rw [lp.norm_eq_ciSup, infty_coeFn_one, Pi.one_apply, norm_one, ciSup_const] diff --git a/Mathlib/Analysis/NormedSpace/Int.lean b/Mathlib/Analysis/NormedSpace/Int.lean index d8a954c1320a1..85022940d76de 100644 --- a/Mathlib/Analysis/NormedSpace/Int.lean +++ b/Mathlib/Analysis/NormedSpace/Int.lean @@ -32,8 +32,6 @@ theorem nnnorm_natCast (n : ℕ) : ‖(n : ℤ)‖₊ = n := @[simp] lemma enorm_natCast (n : ℕ) : ‖(n : ℤ)‖ₑ = n := Real.enorm_natCast _ -@[deprecated (since := "2024-04-05")] alias nnnorm_coe_nat := nnnorm_natCast - @[simp] theorem toNat_add_toNat_neg_eq_nnnorm (n : ℤ) : ↑n.toNat + ↑(-n).toNat = ‖n‖₊ := by rw [← Nat.cast_add, toNat_add_toNat_neg_eq_natAbs, NNReal.natCast_natAbs] diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 38fe34d857d76..fc4539be91cb9 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -332,16 +332,10 @@ theorem not_summable_natCast_inv : ¬Summable (fun n => n⁻¹ : ℕ → ℝ) := mt (summable_nat_pow_inv (p := 1)).1 (lt_irrefl 1) simpa -@[deprecated (since := "2024-04-17")] -alias not_summable_nat_cast_inv := not_summable_natCast_inv - /-- Harmonic series is not unconditionally summable. -/ theorem not_summable_one_div_natCast : ¬Summable (fun n => 1 / n : ℕ → ℝ) := by simpa only [inv_eq_one_div] using not_summable_natCast_inv -@[deprecated (since := "2024-04-17")] -alias not_summable_one_div_nat_cast := not_summable_one_div_natCast - /-- **Divergence of the Harmonic Series** -/ theorem tendsto_sum_range_one_div_nat_succ_atTop : Tendsto (fun n => ∑ i ∈ Finset.range n, (1 / (i + 1) : ℝ)) atTop atTop := by diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean index fb2afeb724d28..6b1c4f5b1691f 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean @@ -35,9 +35,6 @@ theorem arg_exp {x : ℝ} (h₁ : -π < x) (h₂ : x ≤ π) : arg (exp x) = x : theorem exp_arg (z : Circle) : exp (arg z) = z := injective_arg <| arg_exp (neg_pi_lt_arg _) (arg_le_pi _) -@[deprecated (since := "2024-07-25")] alias _root_.arg_expMapCircle := arg_exp -@[deprecated (since := "2024-07-25")] alias _root_.expMapCircle_arg := exp_arg - /-- `Complex.arg ∘ (↑)` and `expMapCircle` define a partial equivalence between `circle` and `ℝ` with `source = Set.univ` and `target = Set.Ioc (-π) π`. -/ @[simps (config := .asFn)] @@ -89,19 +86,6 @@ lemma exp_inj {r s : ℝ} : exp r = exp s ↔ r ≡ s [PMOD (2 * π)] := by lemma exp_sub_two_pi (x : ℝ) : exp (x - 2 * π) = exp x := periodic_exp.sub_eq x lemma exp_add_two_pi (x : ℝ) : exp (x + 2 * π) = exp x := periodic_exp x -@[deprecated (since := "2024-07-25")] -alias _root_.leftInverse_expMapCircle_arg := leftInverse_exp_arg - -@[deprecated (since := "2024-07-25")] -alias _root_.surjOn_expMapCircle_neg_pi_pi := surjOn_exp_neg_pi_pi - -@[deprecated (since := "2024-07-25")] alias _root_.invOn_arg_expMapCircle := invOn_arg_exp -@[deprecated (since := "2024-07-25")] alias _root_.expMapCircle_eq_expMapCircle := exp_eq_exp -@[deprecated (since := "2024-07-25")] alias _root_.periodic_expMapCircle := periodic_exp -@[deprecated (since := "2024-07-25")] alias _root_.expMapCircle_two_pi := exp_two_pi -@[deprecated (since := "2024-07-25")] alias _root_.expMapCircle_sub_two_pi := exp_sub_two_pi -@[deprecated (since := "2024-07-25")] alias _root_.expMapCircle_add_two_pi := exp_add_two_pi - end Circle namespace Real.Angle @@ -131,14 +115,6 @@ lemma coe_toCircle (θ : Angle) : (θ.toCircle : ℂ) = θ.cos + θ.sin * I := b rw [toCircle_coe, Circle.coe_exp, exp_mul_I, ← ofReal_cos, ← ofReal_sin, ← Real.Angle.cos_coe, ← Real.Angle.sin_coe, arg_cos_add_sin_mul_I_coe_angle] -@[deprecated (since := "2024-07-25")] alias expMapCircle := toCircle -@[deprecated (since := "2024-07-25")] alias expMapCircle_coe := toCircle_coe -@[deprecated (since := "2024-07-25")] alias coe_expMapCircle := coe_toCircle -@[deprecated (since := "2024-07-25")] alias expMapCircle_zero := toCircle_zero -@[deprecated (since := "2024-07-25")] alias expMapCircle_neg := toCircle_neg -@[deprecated (since := "2024-07-25")] alias expMapCircle_add := toCircle_add -@[deprecated (since := "2024-07-25")] alias arg_expMapCircle := arg_toCircle - end Real.Angle namespace AddCircle @@ -218,6 +194,3 @@ open AddCircle lemma isLocalHomeomorph_circleExp : IsLocalHomeomorph Circle.exp := by have : Fact (0 < 2 * π) := ⟨by positivity⟩ exact homeomorphCircle'.isLocalHomeomorph.comp (isLocalHomeomorph_coe (2 * π)) - -@[deprecated (since := "2024-07-25")] -alias isLocalHomeomorph_expMapCircle := isLocalHomeomorph_circleExp diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index 615698798410f..e138613487e90 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -187,25 +187,17 @@ nonrec theorem ContinuousWithinAt.rexp (h : ContinuousWithinAt f s x) : ContinuousWithinAt (fun y ↦ exp (f y)) s x := h.rexp -@[deprecated (since := "2024-05-09")] alias ContinuousWithinAt.exp := ContinuousWithinAt.rexp - @[fun_prop] nonrec theorem ContinuousAt.rexp (h : ContinuousAt f x) : ContinuousAt (fun y ↦ exp (f y)) x := h.rexp -@[deprecated (since := "2024-05-09")] alias ContinuousAt.exp := ContinuousAt.rexp - @[fun_prop] theorem ContinuousOn.rexp (h : ContinuousOn f s) : ContinuousOn (fun y ↦ exp (f y)) s := fun x hx ↦ (h x hx).rexp -@[deprecated (since := "2024-05-09")] alias ContinuousOn.exp := ContinuousOn.rexp - @[fun_prop] theorem Continuous.rexp (h : Continuous f) : Continuous fun y ↦ exp (f y) := continuous_iff_continuousAt.2 fun _ ↦ h.continuousAt.rexp -@[deprecated (since := "2024-05-09")] alias Continuous.exp := Continuous.rexp - end RealContinuousExpComp namespace Real diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean index 5e688563f5ae7..eeaafecef2008 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean @@ -367,9 +367,6 @@ theorem floor_logb_natCast {b : ℕ} {r : ℝ} (hr : 0 ≤ r) : · simp_all only [CharP.cast_eq_zero, logb_zero_left, Int.floor_zero, Int.log_zero_left] · simp_all only [Nat.cast_one, logb_one_left, Int.floor_zero, Int.log_one_left] -@[deprecated (since := "2024-04-17")] -alias floor_logb_nat_cast := floor_logb_natCast - theorem ceil_logb_natCast {b : ℕ} {r : ℝ} (hr : 0 ≤ r) : ⌈logb b r⌉ = Int.clog b r := by obtain rfl | hr := hr.eq_or_lt @@ -387,9 +384,6 @@ theorem ceil_logb_natCast {b : ℕ} {r : ℝ} (hr : 0 ≤ r) : · simp_all only [CharP.cast_eq_zero, logb_zero_left, Int.ceil_zero, Int.clog_zero_left] · simp_all only [Nat.cast_one, logb_one_left, Int.ceil_zero, Int.clog_one_left] -@[deprecated (since := "2024-04-17")] -alias ceil_logb_nat_cast := ceil_logb_natCast - lemma natLog_le_logb (a b : ℕ) : Nat.log b a ≤ Real.logb b a := by apply le_trans _ (Int.floor_le ((b : ℝ).logb a)) rw [Real.floor_logb_natCast (Nat.cast_nonneg a), Int.log_natCast, Int.cast_natCast] diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index 3944edf91c60c..40dcee0cf0a06 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -197,16 +197,10 @@ theorem log_natCast_nonneg (n : ℕ) : 0 ≤ log n := by have : (1 : ℝ) ≤ n := mod_cast Nat.one_le_of_lt <| Nat.pos_of_ne_zero hn exact log_nonneg this -@[deprecated (since := "2024-04-17")] -alias log_nat_cast_nonneg := log_natCast_nonneg - theorem log_neg_natCast_nonneg (n : ℕ) : 0 ≤ log (-n) := by rw [← log_neg_eq_log, neg_neg] exact log_natCast_nonneg _ -@[deprecated (since := "2024-04-17")] -alias log_neg_nat_cast_nonneg := log_neg_natCast_nonneg - theorem log_intCast_nonneg (n : ℤ) : 0 ≤ log n := by cases lt_trichotomy 0 n with | inl hn => @@ -220,9 +214,6 @@ theorem log_intCast_nonneg (n : ℤ) : 0 ≤ log n := by rw [← log_neg_eq_log] exact log_nonneg this -@[deprecated (since := "2024-04-17")] -alias log_int_cast_nonneg := log_intCast_nonneg - theorem strictMonoOn_log : StrictMonoOn log (Set.Ioi 0) := fun _ hx _ _ hxy => log_lt_log hx hxy theorem strictAntiOn_log : StrictAntiOn log (Set.Iio 0) := by diff --git a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean index 86233fe6cde8c..fd6883ec03808 100644 --- a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean +++ b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean @@ -198,8 +198,6 @@ theorem measurableEquivRealProd_symm_polarCoord_symm_apply (p : ℝ × ℝ) : theorem polarCoord_symm_abs (p : ℝ × ℝ) : Complex.abs (Complex.polarCoord.symm p) = |p.1| := by simp -@[deprecated (since := "2024-07-15")] alias polardCoord_symm_abs := polarCoord_symm_abs - protected theorem integral_comp_polarCoord_symm {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] (f : ℂ → E) : (∫ p in polarCoord.target, p.1 • f (Complex.polarCoord.symm p)) = ∫ p, f p := by diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean index 30cb5251b049e..1fa2d130b2f55 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean @@ -126,9 +126,6 @@ lemma cpow_mul_ofNat (x y : ℂ) (n : ℕ) [n.AtLeastTwo] : @[simp, norm_cast] theorem cpow_natCast (x : ℂ) (n : ℕ) : x ^ (n : ℂ) = x ^ n := by simpa using cpow_nat_mul x n 1 -@[deprecated (since := "2024-04-17")] -alias cpow_nat_cast := cpow_natCast - @[simp] lemma cpow_ofNat (x : ℂ) (n : ℕ) [n.AtLeastTwo] : x ^ (ofNat(n) : ℂ) = x ^ ofNat(n) := @@ -139,9 +136,6 @@ theorem cpow_two (x : ℂ) : x ^ (2 : ℂ) = x ^ (2 : ℕ) := cpow_ofNat x 2 @[simp, norm_cast] theorem cpow_intCast (x : ℂ) (n : ℤ) : x ^ (n : ℂ) = x ^ n := by simpa using cpow_int_mul x n 1 -@[deprecated (since := "2024-04-17")] -alias cpow_int_cast := cpow_intCast - @[simp] theorem cpow_nat_inv_pow (x : ℂ) {n : ℕ} (hn : n ≠ 0) : (x ^ (n⁻¹ : ℂ)) ^ n = x := by rw [← cpow_nat_mul, mul_inv_cancel₀, cpow_one] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index 3356da74c63dd..18488e7c4d8c2 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -171,9 +171,6 @@ theorem sqrt_eq_rpow (x : ℝ≥0) : sqrt x = x ^ (1 / (2 : ℝ)) := by push_cast exact Real.sqrt_eq_rpow x.1 -@[deprecated (since := "2024-04-17")] -alias rpow_nat_cast := rpow_natCast - @[simp] lemma rpow_ofNat (x : ℝ≥0) (n : ℕ) [n.AtLeastTwo] : x ^ (ofNat(n) : ℝ) = x ^ (OfNat.ofNat n : ℕ) := @@ -260,17 +257,9 @@ theorem rpow_le_rpow_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ^ z ≤ y ^ theorem le_rpow_inv_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ≤ y ^ z⁻¹ ↔ x ^ z ≤ y := by rw [← rpow_le_rpow_iff hz, ← one_div, rpow_self_rpow_inv hz.ne'] -@[deprecated le_rpow_inv_iff (since := "2024-07-10")] -theorem le_rpow_one_div_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ≤ y ^ (1 / z) ↔ x ^ z ≤ y := by - rw [← rpow_le_rpow_iff hz, rpow_self_rpow_inv hz.ne'] - theorem rpow_inv_le_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ^ z⁻¹ ≤ y ↔ x ≤ y ^ z := by rw [← rpow_le_rpow_iff hz, ← one_div, rpow_self_rpow_inv hz.ne'] -@[deprecated rpow_inv_le_iff (since := "2024-07-10")] -theorem rpow_one_div_le_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ^ (1 / z) ≤ y ↔ x ≤ y ^ z := by - rw [← rpow_le_rpow_iff hz, rpow_self_rpow_inv hz.ne'] - theorem lt_rpow_inv_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x < y ^ z⁻¹ ↔ x ^z < y := by simp only [← not_le, rpow_inv_le_iff hz] @@ -393,17 +382,9 @@ theorem rpow_left_bijective {x : ℝ} (hx : x ≠ 0) : Function.Bijective fun y theorem eq_rpow_inv_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x = y ^ z⁻¹ ↔ x ^ z = y := by rw [← rpow_eq_rpow_iff hz, ← one_div, rpow_self_rpow_inv hz] -@[deprecated eq_rpow_inv_iff (since := "2024-07-10")] -theorem eq_rpow_one_div_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x = y ^ (1 / z) ↔ x ^ z = y := by - rw [← rpow_eq_rpow_iff hz, rpow_self_rpow_inv hz] - theorem rpow_inv_eq_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x ^ z⁻¹ = y ↔ x = y ^ z := by rw [← rpow_eq_rpow_iff hz, ← one_div, rpow_self_rpow_inv hz] -@[deprecated rpow_inv_eq_iff (since := "2024-07-10")] -theorem rpow_one_div_eq_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x ^ (1 / z) = y ↔ x = y ^ z := by - rw [← rpow_eq_rpow_iff hz, rpow_self_rpow_inv hz] - @[simp] lemma rpow_rpow_inv {y : ℝ} (hy : y ≠ 0) (x : ℝ≥0) : (x ^ y) ^ y⁻¹ = x := by rw [← rpow_mul, mul_inv_cancel₀ hy, rpow_one] @@ -630,9 +611,6 @@ theorem rpow_natCast (x : ℝ≥0∞) (n : ℕ) : x ^ (n : ℝ) = x ^ n := by · cases n <;> simp [top_rpow_of_pos (Nat.cast_add_one_pos _), top_pow (Nat.succ_pos _)] · simp [← coe_rpow_of_nonneg _ (Nat.cast_nonneg n)] -@[deprecated (since := "2024-04-17")] -alias rpow_nat_cast := rpow_natCast - @[simp] lemma rpow_ofNat (x : ℝ≥0∞) (n : ℕ) [n.AtLeastTwo] : x ^ (ofNat(n) : ℝ) = x ^ (OfNat.ofNat n) := @@ -643,9 +621,6 @@ lemma rpow_intCast (x : ℝ≥0∞) (n : ℤ) : x ^ (n : ℝ) = x ^ n := by cases n <;> simp only [Int.ofNat_eq_coe, Int.cast_natCast, rpow_natCast, zpow_natCast, Int.cast_negSucc, rpow_neg, zpow_negSucc] -@[deprecated (since := "2024-04-17")] -alias rpow_int_cast := rpow_intCast - theorem rpow_two (x : ℝ≥0∞) : x ^ (2 : ℝ) = x ^ 2 := rpow_ofNat x 2 theorem mul_rpow_eq_ite (x y : ℝ≥0∞) (z : ℝ) : @@ -760,12 +735,6 @@ theorem le_rpow_inv_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ≤ y ^ z nth_rw 1 [← @mul_inv_cancel₀ _ _ z hz.ne'] rw [rpow_mul, @rpow_le_rpow_iff _ _ z⁻¹ (by simp [hz])] -@[deprecated le_rpow_inv_iff (since := "2024-07-10")] -theorem le_rpow_one_div_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ≤ y ^ (1 / z) ↔ x ^ z ≤ y := by - nth_rw 1 [← rpow_one x] - nth_rw 1 [← @mul_inv_cancel₀ _ _ z hz.ne'] - rw [rpow_mul, ← one_div, @rpow_le_rpow_iff _ _ (1 / z) (by simp [hz])] - theorem rpow_inv_lt_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z⁻¹ < y ↔ x < y ^ z := by simp only [← not_le, le_rpow_inv_iff hz] @@ -774,23 +743,11 @@ theorem lt_rpow_inv_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x < y ^ z⁻ nth_rw 1 [← @mul_inv_cancel₀ _ _ z (ne_of_lt hz).symm] rw [rpow_mul, @rpow_lt_rpow_iff _ _ z⁻¹ (by simp [hz])] -@[deprecated lt_rpow_inv_iff (since := "2024-07-10")] -theorem lt_rpow_one_div_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x < y ^ (1 / z) ↔ x ^ z < y := by - nth_rw 1 [← rpow_one x] - nth_rw 1 [← @mul_inv_cancel₀ _ _ z (ne_of_lt hz).symm] - rw [rpow_mul, ← one_div, @rpow_lt_rpow_iff _ _ (1 / z) (by simp [hz])] - theorem rpow_inv_le_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z⁻¹ ≤ y ↔ x ≤ y ^ z := by nth_rw 1 [← ENNReal.rpow_one y] nth_rw 1 [← @mul_inv_cancel₀ _ _ z hz.ne.symm] rw [ENNReal.rpow_mul, ENNReal.rpow_le_rpow_iff (inv_pos.2 hz)] -@[deprecated rpow_inv_le_iff (since := "2024-07-10")] -theorem rpow_one_div_le_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ (1 / z) ≤ y ↔ x ≤ y ^ z := by - nth_rw 1 [← ENNReal.rpow_one y] - nth_rw 2 [← @mul_inv_cancel₀ _ _ z hz.ne.symm] - rw [ENNReal.rpow_mul, ← one_div, ENNReal.rpow_le_rpow_iff (one_div_pos.2 hz)] - theorem rpow_lt_rpow_of_exponent_lt {x : ℝ≥0∞} {y z : ℝ} (hx : 1 < x) (hx' : x ≠ ⊤) (hyz : y < z) : x ^ y < x ^ z := by lift x to ℝ≥0 using hx' diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index 0a8b78fdc039e..a7e211c380d7f 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -53,15 +53,9 @@ theorem rpow_intCast (x : ℝ) (n : ℤ) : x ^ (n : ℝ) = x ^ n := by simp only [rpow_def, ← Complex.ofReal_zpow, Complex.cpow_intCast, Complex.ofReal_intCast, Complex.ofReal_re] -@[deprecated (since := "2024-04-17")] -alias rpow_int_cast := rpow_intCast - @[simp, norm_cast] theorem rpow_natCast (x : ℝ) (n : ℕ) : x ^ (n : ℝ) = x ^ n := by simpa using rpow_intCast x n -@[deprecated (since := "2024-04-17")] -alias rpow_nat_cast := rpow_natCast - @[simp] theorem exp_one_rpow (x : ℝ) : exp 1 ^ x = exp x := by rw [← exp_mul, one_mul] diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean index 486b4177f1eb6..24100361acca4 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean @@ -98,9 +98,6 @@ theorem natCast_mul_eq_nsmul (x : ℝ) (n : ℕ) : ↑((n : ℝ) * x) = n • ( theorem intCast_mul_eq_zsmul (x : ℝ) (n : ℤ) : ↑((n : ℝ) * x : ℝ) = n • (↑x : Angle) := by simpa only [zsmul_eq_mul] using coeHom.map_zsmul x n -@[deprecated (since := "2024-05-25")] alias coe_nat_mul_eq_nsmul := natCast_mul_eq_nsmul -@[deprecated (since := "2024-05-25")] alias coe_int_mul_eq_zsmul := intCast_mul_eq_zsmul - theorem angle_eq_iff_two_pi_dvd_sub {ψ θ : ℝ} : (θ : Angle) = ψ ↔ ∃ k : ℤ, θ - ψ = 2 * π * k := by simp only [QuotientAddGroup.eq, AddSubgroup.zmultiples_eq_closure, AddSubgroup.mem_closure_singleton, zsmul_eq_mul', (sub_eq_neg_add _ _).symm, eq_comm] diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index 10783716eb430..7080ebbb39699 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -242,9 +242,6 @@ theorem tsum_geometric_le_of_norm_lt_one (x : R) (h : ‖x‖ < 1) : variable [HasSummableGeomSeries R] -@[deprecated (since := "2024-07-27")] -alias NormedRing.tsum_geometric_of_norm_lt_one := tsum_geometric_le_of_norm_lt_one - theorem geom_series_mul_neg (x : R) (h : ‖x‖ < 1) : (∑' i : ℕ, x ^ i) * (1 - x) = 1 := by have := (summable_geometric_of_norm_lt_one h).hasSum.mul_right (1 - x) refine tendsto_nhds_unique this.tendsto_sum_nat ?_ @@ -298,9 +295,6 @@ lemma isUnit_one_sub_of_norm_lt_one {x : R} (h : ‖x‖ < 1) : IsUnit (1 - x) : end HasSummableGeometricSeries -@[deprecated (since := "2024-07-27")] -alias NormedRing.summable_geometric_of_norm_lt_one := summable_geometric_of_norm_lt_one - section Geometric variable {K : Type*} [NormedDivisionRing K] {ξ : K} diff --git a/Mathlib/CategoryTheory/Abelian/Exact.lean b/Mathlib/CategoryTheory/Abelian/Exact.lean index 30b6b7aab74df..dc62313a74207 100644 --- a/Mathlib/CategoryTheory/Abelian/Exact.lean +++ b/Mathlib/CategoryTheory/Abelian/Exact.lean @@ -212,9 +212,6 @@ end Functor namespace Functor -@[deprecated (since := "2024-07-09")] alias CategoryTheory.Functor.map_exact := - ShortComplex.Exact.map - open Limits Abelian variable {A : Type u₁} {B : Type u₂} [Category.{v₁} A] [Category.{v₂} B] @@ -264,11 +261,6 @@ lemma preservesHomology_of_map_exact : L.PreservesHomology where exact (hL (ShortComplex.mk _ _ (kernel.condition f)) (ShortComplex.exact_of_f_is_kernel _ (kernelIsKernel f))).fIsKernel -@[deprecated (since := "2024-07-09")] alias preservesKernelsOfMapExact := - PreservesHomology.preservesKernels -@[deprecated (since := "2024-07-09")] alias preservesCokernelsOfMapExact := - PreservesHomology.preservesCokernels - end section diff --git a/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean b/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean index b646309bd0974..589815ad1ab03 100644 --- a/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean @@ -225,9 +225,6 @@ theorem mem_essImage_of_unit_isIso (A : C) [IsIso (h.unit.app A)] : A ∈ R.essImage := ⟨L.obj A, ⟨(asIso (h.unit.app A)).symm⟩⟩ -@[deprecated (since := "2024-06-19")] alias _root_.CategoryTheory.mem_essImage_of_unit_isIso := - mem_essImage_of_unit_isIso - lemma isIso_unit_app_of_iso [R.Faithful] [R.Full] {X : D} {Y : C} (e : Y ≅ R.obj X) : IsIso (h.unit.app Y) := (isIso_unit_app_iff_mem_essImage h).mpr ⟨X, ⟨e.symm⟩⟩ diff --git a/Mathlib/CategoryTheory/Adjunction/Mates.lean b/Mathlib/CategoryTheory/Adjunction/Mates.lean index 322bbb825e516..59a44bbd78cc5 100644 --- a/Mathlib/CategoryTheory/Adjunction/Mates.lean +++ b/Mathlib/CategoryTheory/Adjunction/Mates.lean @@ -96,8 +96,6 @@ def mateEquiv : (G ⋙ L₂ ⟶ L₁ ⋙ H) ≃ (R₁ ⋙ G ⟶ H ⋙ R₂) wher rw [← assoc, ← Functor.comp_map, assoc, ← β.naturality, ← assoc, Functor.comp_map, ← G.map_comp, right_triangle_components, map_id, id_comp] -@[deprecated (since := "2024-07-09")] alias transferNatTrans := mateEquiv - /-- A component of a transposed version of the mates correspondence. -/ theorem mateEquiv_counit (α : G ⋙ L₂ ⟶ L₁ ⋙ H) (d : D) : L₂.map ((mateEquiv adj₁ adj₂ α).app _) ≫ adj₂.counit.app _ = @@ -335,8 +333,6 @@ def conjugateEquiv : (L₂ ⟶ L₁) ≃ (R₁ ⟶ R₂) := _ ≃ _ := mateEquiv adj₁ adj₂ _ ≃ (R₁ ⟶ R₂) := R₁.rightUnitor.homCongr R₂.leftUnitor -@[deprecated (since := "2024-07-09")] alias transferNatTransSelf := conjugateEquiv - /-- A component of a transposed form of the conjugation definition. -/ theorem conjugateEquiv_counit (α : L₂ ⟶ L₁) (d : D) : L₂.map ((conjugateEquiv adj₁ adj₂ α).app _) ≫ adj₂.counit.app d = diff --git a/Mathlib/CategoryTheory/Adjunction/Over.lean b/Mathlib/CategoryTheory/Adjunction/Over.lean index 949e581673aea..3409887a63bda 100644 --- a/Mathlib/CategoryTheory/Adjunction/Over.lean +++ b/Mathlib/CategoryTheory/Adjunction/Over.lean @@ -55,12 +55,6 @@ def pullback {X Y : C} (f : X ⟶ Y) : Over Y ⥤ Over X where Over.homMk (pullback.lift (pullback.fst _ _ ≫ k.left) (pullback.snd _ _) (by simp [pullback.condition])) -@[deprecated (since := "2024-05-15")] -noncomputable alias Limits.baseChange := Over.pullback - -@[deprecated (since := "2024-07-08")] -noncomputable alias baseChange := pullback - /-- `Over.map f` is left adjoint to `Over.pullback f`. -/ @[simps! unit_app counit_app] def mapPullbackAdj {X Y : C} (f : X ⟶ Y) : Over.map f ⊣ pullback f := @@ -78,9 +72,6 @@ def mapPullbackAdj {X Y : C} (f : X ⟶ Y) : Over.map f ⊣ pullback f := · simp · simpa using (Over.w v).symm } } -@[deprecated (since := "2024-07-08")] -noncomputable alias mapAdjunction := mapPullbackAdj - /-- pullback (𝟙 X) : Over X ⥤ Over X is the identity functor. -/ def pullbackId {X : C} : pullback (𝟙 X) ≅ 𝟭 _ := conjugateIsoEquiv (mapPullbackAdj (𝟙 _)) (Adjunction.id (C := Over _)) (Over.mapId _).symm @@ -117,10 +108,6 @@ instance [HasBinaryProducts C] : (forget X).IsLeftAdjoint := end Over -@[deprecated (since := "2024-05-18")] noncomputable alias star := Over.star - -@[deprecated (since := "2024-05-18")] noncomputable alias forgetAdjStar := Over.forgetAdjStar - namespace Under variable [HasPushouts C] diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index f97365afa177a..d30dc736344e6 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -170,16 +170,10 @@ theorem functor_obj {I : Type u₁} (F : I → C) (i : I) : @[simp] theorem functor_map {I : Type u₁} (F : I → C) {i : Discrete I} (f : i ⟶ i) : (Discrete.functor F).map f = 𝟙 (F i.as) := by aesop_cat -@[deprecated (since := "2024-07-16")] -alias CategoryTheory.FreeMonoidalCategory.discrete_functor_map_eq_id := functor_map - @[simp] theorem functor_obj_eq_as {I : Type u₁} (F : I → C) (X : Discrete I) : (Discrete.functor F).obj X = F X.as := rfl -@[deprecated (since := "2024-07-16")] -alias CategoryTheory.FreeMonoidalCategory.discrete_functor_obj_eq_as := functor_obj_eq_as - /-- The discrete functor induced by a composition of maps can be written as a composition of two discrete functors. -/ diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index 073a622465f51..de003e519b9d2 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -627,19 +627,4 @@ def isoInverseComp {G : C ≌ D} (i : G.functor ⋙ H ≅ F) : H ≅ G.inverse end Iso -@[deprecated (since := "2024-04-06")] alias IsEquivalence := Functor.IsEquivalence -@[deprecated (since := "2024-04-06")] alias IsEquivalence.fun_inv_map := Functor.fun_inv_map -@[deprecated (since := "2024-04-06")] alias IsEquivalence.inv_fun_map := Functor.inv_fun_map -@[deprecated (since := "2024-04-06")] alias IsEquivalence.ofIso := Equivalence.changeFunctor -@[deprecated (since := "2024-04-06")] -alias IsEquivalence.ofIso_trans := Equivalence.changeFunctor_trans -@[deprecated (since := "2024-04-06")] -alias IsEquivalence.ofIso_refl := Equivalence.changeFunctor_refl -@[deprecated (since := "2024-04-06")] -alias IsEquivalence.equivOfIso := Functor.isEquivalence_iff_of_iso -@[deprecated (since := "2024-04-06")] -alias IsEquivalence.cancelCompRight := Functor.isEquivalence_of_comp_right -@[deprecated (since := "2024-04-06")] -alias IsEquivalence.cancelCompLeft := Functor.isEquivalence_of_comp_left - end CategoryTheory diff --git a/Mathlib/CategoryTheory/EssentialImage.lean b/Mathlib/CategoryTheory/EssentialImage.lean index f7fd17a62ab59..f809eabd0a256 100644 --- a/Mathlib/CategoryTheory/EssentialImage.lean +++ b/Mathlib/CategoryTheory/EssentialImage.lean @@ -162,7 +162,4 @@ lemma essSurj_of_comp_fully_faithful (F : C ⥤ D) (G : D ⥤ E) [(F ⋙ G).EssS end Functor -@[deprecated (since := "2024-04-06")] alias EssSurj := Functor.EssSurj -@[deprecated (since := "2024-04-06")] alias Iso.map_essSurj := Functor.essSurj_of_iso - end CategoryTheory diff --git a/Mathlib/CategoryTheory/Functor/Category.lean b/Mathlib/CategoryTheory/Functor/Category.lean index 64d1a3f601533..c8c33aded9b15 100644 --- a/Mathlib/CategoryTheory/Functor/Category.lean +++ b/Mathlib/CategoryTheory/Functor/Category.lean @@ -168,9 +168,4 @@ theorem map_inv_hom_id_app {X Y : C} (e : X ≅ Y) (F : C ⥤ D ⥤ E) (Z : D) : end Iso -@[deprecated (since := "2024-06-09")] alias map_hom_inv_app := Iso.map_hom_inv_id_app -@[deprecated (since := "2024-06-09")] alias map_inv_hom_app := Iso.map_inv_hom_id_app -@[deprecated (since := "2024-06-09")] alias map_hom_inv_app_assoc := Iso.map_hom_inv_id_app_assoc -@[deprecated (since := "2024-06-09")] alias map_inv_hom_app_assoc := Iso.map_inv_hom_id_app_assoc - end CategoryTheory diff --git a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean index 98607af7c86c4..21b9228ed4deb 100644 --- a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean @@ -363,34 +363,4 @@ theorem fullyFaithfulCancelRight_inv_app {F G : C ⥤ D} {H : D ⥤ E} [Full H] rfl end Functor - -@[deprecated (since := "2024-04-06")] alias Full := Functor.Full -@[deprecated (since := "2024-04-06")] alias Faithful := Functor.Faithful -@[deprecated (since := "2024-04-06")] alias preimage_id := Functor.preimage_id -@[deprecated (since := "2024-04-06")] alias preimage_comp := Functor.preimage_comp -@[deprecated (since := "2024-04-06")] alias preimage_map := Functor.preimage_map -@[deprecated (since := "2024-04-06")] alias Faithful.of_comp := Functor.Faithful.of_comp -@[deprecated (since := "2024-04-06")] alias Full.ofIso := Functor.Full.of_iso -@[deprecated (since := "2024-04-06")] alias Faithful.of_iso := Functor.Faithful.of_iso -@[deprecated (since := "2024-04-06")] alias Faithful.of_comp_iso := Functor.Faithful.of_comp_iso -@[deprecated (since := "2024-04-06")] alias Faithful.of_comp_eq := Functor.Faithful.of_comp_eq -@[deprecated (since := "2024-04-06")] alias Faithful.div := Functor.Faithful.div -@[deprecated (since := "2024-04-06")] alias Faithful.div_comp := Functor.Faithful.div_comp -@[deprecated (since := "2024-04-06")] alias Faithful.div_faithful := Functor.Faithful.div_faithful -@[deprecated (since := "2024-04-06")] alias Full.ofCompFaithful := Functor.Full.of_comp_faithful - -@[deprecated (since := "2024-04-06")] -alias Full.ofCompFaithfulIso := Functor.Full.of_comp_faithful_iso - -@[deprecated (since := "2024-04-06")] -alias fullyFaithfulCancelRight := Functor.fullyFaithfulCancelRight - -@[deprecated (since := "2024-04-06")] -alias fullyFaithfulCancelRight_hom_app := Functor.fullyFaithfulCancelRight_hom_app - -@[deprecated (since := "2024-04-06")] -alias fullyFaithfulCancelRight_inv_app := Functor.fullyFaithfulCancelRight_inv_app - -@[deprecated (since := "2024-04-26")] alias Functor.image_preimage := Functor.map_preimage - end CategoryTheory diff --git a/Mathlib/CategoryTheory/Functor/ReflectsIso.lean b/Mathlib/CategoryTheory/Functor/ReflectsIso.lean index af27e1c18d46a..ec4bd387fd710 100644 --- a/Mathlib/CategoryTheory/Functor/ReflectsIso.lean +++ b/Mathlib/CategoryTheory/Functor/ReflectsIso.lean @@ -39,8 +39,6 @@ class Functor.ReflectsIsomorphisms (F : C ⥤ D) : Prop where /-- For any `f`, if `F.map f` is an iso, then so was `f`-/ reflects : ∀ {A B : C} (f : A ⟶ B) [IsIso (F.map f)], IsIso f -@[deprecated (since := "2024-04-06")] alias ReflectsIsomorphisms := Functor.ReflectsIsomorphisms - /-- If `F` reflects isos and `F.map f` is an iso, then `f` is an iso. -/ theorem isIso_of_reflects_iso {A B : C} (f : A ⟶ B) (F : C ⥤ D) [IsIso (F.map f)] [F.ReflectsIsomorphisms] : IsIso f := diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index f022f13162297..57fd2c7d41e03 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -110,9 +110,6 @@ theorem comp_f {P Q R : Karoubi C} (f : P ⟶ Q) (g : Q ⟶ R) : (f ≫ g).f = f @[simp] theorem id_f {P : Karoubi C} : Hom.f (𝟙 P) = P.p := rfl -@[deprecated "No deprecation message was provided." (since := "2024-07-15")] -theorem id_eq {P : Karoubi C} : 𝟙 P = ⟨P.p, by repeat' rw [P.idem]⟩ := rfl - /-- It is possible to coerce an object of `C` into an object of `Karoubi C`. See also the functor `toKaroubi`. -/ instance coe : CoeTC C (Karoubi C) := diff --git a/Mathlib/CategoryTheory/Iso.lean b/Mathlib/CategoryTheory/Iso.lean index 781198c903a31..cee28cf05442e 100644 --- a/Mathlib/CategoryTheory/Iso.lean +++ b/Mathlib/CategoryTheory/Iso.lean @@ -325,9 +325,6 @@ theorem eq_inv_of_inv_hom_id {f : X ⟶ Y} [IsIso f] {g : Y ⟶ X} (inv_hom_id : instance id (X : C) : IsIso (𝟙 X) := ⟨⟨𝟙 X, by simp⟩⟩ -@[deprecated (since := "2024-05-15")] alias of_iso := CategoryTheory.Iso.isIso_hom -@[deprecated (since := "2024-05-15")] alias of_iso_inv := CategoryTheory.Iso.isIso_inv - variable {f : X ⟶ Y} {h : Y ⟶ Z} instance inv_isIso [IsIso f] : IsIso (inv f) := diff --git a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean index 6d548e45cfabe..a54f6901a6672 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean @@ -415,14 +415,6 @@ instance RespectsIso.isomorphisms : RespectsIso (isomorphisms C) := by intro exact IsIso.comp_isIso -@[deprecated (since := "2024-07-02")] alias RespectsIso.cancel_left_isIso := - cancel_left_of_respectsIso -@[deprecated (since := "2024-07-02")] alias RespectsIso.cancel_right_isIso := - cancel_right_of_respectsIso -@[deprecated (since := "2024-07-02")] alias RespectsIso.arrow_iso_iff := arrow_iso_iff -@[deprecated (since := "2024-07-02")] alias RespectsIso.arrow_mk_iso_iff := arrow_mk_iso_iff -@[deprecated (since := "2024-07-02")] alias RespectsIso.isoClosure_eq := isoClosure_eq_self - /-- If `W₁` and `W₂` are morphism properties on two categories `C₁` and `C₂`, this is the induced morphism property on `C₁ × C₂`. -/ def prod {C₁ C₂ : Type*} [Category C₁] [Category C₂] diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite/InducedTopology.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/InducedTopology.lean index 4710ae5d4bf61..28534b95a6405 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite/InducedTopology.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/InducedTopology.lean @@ -119,9 +119,6 @@ instance (priority := 900) locallyCoverDense_of_isCoverDense [G.IsCoverDense K] instance (priority := 900) [G.IsCoverDense K] : G.IsDenseSubsite (G.inducedTopology K) K where functorPushforward_mem_iff := Iff.rfl -@[deprecated (since := "2024-07-23")] -alias inducedTopologyOfIsCoverDense := inducedTopology - variable (J) instance over_forget_locallyCoverDense (X : C) : (Over.forget X).LocallyCoverDense J where diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean index ecb6633e6f906..cf245bb9301a1 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean @@ -134,10 +134,4 @@ abbrev sheafEquivSheafificationCompatibility : end IsDenseSubsite -@[deprecated (since := "2024-07-23")] -alias IsCoverDense.sheafEquivOfCoverPreservingCoverLifting := IsDenseSubsite.sheafEquiv -@[deprecated (since := "2024-07-23")] -alias IsCoverDense.sheafEquivOfCoverPreservingCoverLiftingSheafificationCompatibility := - IsDenseSubsite.sheafEquivSheafificationCompatibility - end CategoryTheory.Functor diff --git a/Mathlib/CategoryTheory/Sites/Sieves.lean b/Mathlib/CategoryTheory/Sites/Sieves.lean index 23b36b70f4d19..f4ebdb71b3f5d 100644 --- a/Mathlib/CategoryTheory/Sites/Sieves.lean +++ b/Mathlib/CategoryTheory/Sites/Sieves.lean @@ -409,8 +409,6 @@ theorem generate_le_iff (R : Presieve X) (S : Sieve X) : generate R ≤ S ↔ R rintro ⟨Z, f, g, hg, rfl⟩ exact S.downward_closed (ss Z hg) f⟩ -@[deprecated (since := "2024-07-13")] alias sets_iff_generate := generate_le_iff - /-- Show that there is a galois insertion (generate, set_over). -/ def giGenerate : GaloisInsertion (generate : Presieve X → Sieve X) arrows where gc := generate_le_iff diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index 46dc3c0b92821..fcf38ce535510 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -647,27 +647,6 @@ theorem getElem_splitWrtComposition (l : List α) (c : Composition n) (l.splitWrtComposition c)[i] = (l.take (c.sizeUpTo (i + 1))).drop (c.sizeUpTo i) := getElem_splitWrtComposition' _ _ h -@[deprecated getElem_splitWrtCompositionAux (since := "2024-06-12")] -theorem get_splitWrtCompositionAux (l : List α) (ns : List ℕ) {i : ℕ} (hi) : - (l.splitWrtCompositionAux ns).get ⟨i, hi⟩ = - (l.take (ns.take (i + 1)).sum).drop (ns.take i).sum := by - simp [getElem_splitWrtCompositionAux] - -/-- The `i`-th sublist in the splitting of a list `l` along a composition `c`, is the slice of `l` -between the indices `c.sizeUpTo i` and `c.sizeUpTo (i+1)`, i.e., the indices in the `i`-th -block of the composition. -/ -@[deprecated getElem_splitWrtComposition' (since := "2024-06-12")] -theorem get_splitWrtComposition' (l : List α) (c : Composition n) {i : ℕ} - (hi : i < (l.splitWrtComposition c).length) : - (l.splitWrtComposition c).get ⟨i, hi⟩ = (l.take (c.sizeUpTo (i + 1))).drop (c.sizeUpTo i) := by - simp [getElem_splitWrtComposition'] - -@[deprecated getElem_splitWrtComposition (since := "2024-06-12")] -theorem get_splitWrtComposition (l : List α) (c : Composition n) - (i : Fin (l.splitWrtComposition c).length) : - get (l.splitWrtComposition c) i = (l.take (c.sizeUpTo (i + 1))).drop (c.sizeUpTo i) := by - simp [getElem_splitWrtComposition] - theorem flatten_splitWrtCompositionAux {ns : List ℕ} : ∀ {l : List α}, ns.sum = l.length → (l.splitWrtCompositionAux ns).flatten = l := by induction' ns with n ns IH <;> intro l h <;> simp at h diff --git a/Mathlib/Control/LawfulFix.lean b/Mathlib/Control/LawfulFix.lean index 9c9a30c5b82a3..5a974ad8b9016 100644 --- a/Mathlib/Control/LawfulFix.lean +++ b/Mathlib/Control/LawfulFix.lean @@ -179,20 +179,7 @@ theorem fix_eq_of_ωScottContinuous (hc : ωScottContinuous g) : variable {f} -set_option linter.deprecated false in -@[deprecated fix_eq_of_ωScottContinuous (since := "2024-08-26")] -theorem fix_eq (hc : Continuous f) : Part.fix f = f (Part.fix f) := by - rw [fix_eq_ωSup f, hc] - apply le_antisymm - · apply ωSup_le_ωSup_of_le _ - intro i - exists i - intro x - -- intros x y hx, - apply le_f_of_mem_approx _ ⟨i, rfl⟩ - · apply ωSup_le_ωSup_of_le _ - intro i - exists i.succ +@[deprecated (since := "2024-08-26")] alias fix_eq := fix_eq_of_ωScottContinuous end Part @@ -210,13 +197,7 @@ theorem ωScottContinuous_toUnitMono (f : Part α → Part α) (hc : ωScottCont dsimp [OmegaCompletePartialOrder.ωSup] erw [hc.map_ωSup, Chain.map_comp]; rfl -set_option linter.deprecated false in -@[deprecated ωScottContinuous_toUnitMono (since := "2024-08-26")] -theorem to_unit_cont (f : Part α →o Part α) (hc : Continuous f) : Continuous (toUnitMono f) - | _ => by - ext ⟨⟩ : 1 - dsimp [OmegaCompletePartialOrder.ωSup] - erw [hc, Chain.map_comp]; rfl +@[deprecated (since := "2024-08-26")] alias to_unit_cont := ωScottContinuous_toUnitMono instance lawfulFix : LawfulFix (Part α) := ⟨fun {f : Part α → Part α} hc ↦ show Part.fix (toUnitMono ⟨f,hc.monotone⟩) () = _ by @@ -263,13 +244,7 @@ theorem ωScottContinuous_curry : rw [map_comp, map_comp] rfl -set_option linter.deprecated false in -@[deprecated ωScottContinuous_curry (since := "2024-08-26")] -theorem continuous_curry : Continuous <| monotoneCurry α β γ := fun c ↦ by - ext x y - dsimp [curry, ωSup] - rw [map_comp, map_comp] - rfl +@[deprecated (since := "2024-08-26")] alias continuous_curry := ωScottContinuous_curry theorem ωScottContinuous_uncurry : ωScottContinuous (monotoneUncurry α β γ) := @@ -279,13 +254,7 @@ theorem ωScottContinuous_uncurry : rw [map_comp, map_comp] rfl -set_option linter.deprecated false in -@[deprecated ωScottContinuous_uncurry (since := "2024-08-26")] -theorem continuous_uncurry : Continuous <| monotoneUncurry α β γ := fun c ↦ by - ext ⟨x, y⟩ - dsimp [uncurry, ωSup] - rw [map_comp, map_comp] - rfl +@[deprecated (since := "2024-08-26")] alias continuous_uncurry := ωScottContinuous_uncurry end Monotone @@ -306,12 +275,8 @@ theorem uncurry_curry_ωScottContinuous (hc : ωScottContinuous f) : monotoneCurry α β γ := (ωScottContinuous_uncurry _ _ _).comp (hc.comp (ωScottContinuous_curry _ _ _)) -set_option linter.deprecated false in -@[deprecated uncurry_curry_ωScottContinuous (since := "2024-08-26")] -theorem uncurry_curry_continuous {f : ((x : _) → (y : β x) → γ x y) →o (x : _) → (y : β x) → γ x y} - (hc : Continuous f) : - Continuous <| (monotoneUncurry α β γ).comp <| f.comp <| monotoneCurry α β γ := - continuous_comp _ _ (continuous_comp _ _ (continuous_curry _ _ _) hc) (continuous_uncurry _ _ _) +@[deprecated (since := "2024-08-26")] +alias uncurry_curry_continuous := uncurry_curry_ωScottContinuous end Curry diff --git a/Mathlib/Data/Bool/Basic.lean b/Mathlib/Data/Bool/Basic.lean index 0857104a5b71e..e717ff8783b52 100644 --- a/Mathlib/Data/Bool/Basic.lean +++ b/Mathlib/Data/Bool/Basic.lean @@ -94,38 +94,15 @@ theorem of_decide_false {p : Prop} [Decidable p] : decide p = false → ¬p := theorem decide_congr {p q : Prop} [Decidable p] [Decidable q] (h : p ↔ q) : decide p = decide q := decide_eq_decide.mpr h -@[deprecated (since := "2024-06-07")] alias coe_or_iff := or_eq_true_iff - -@[deprecated (since := "2024-06-07")] alias coe_and_iff := and_eq_true_iff - theorem coe_xor_iff (a b : Bool) : xor a b ↔ Xor' (a = true) (b = true) := by cases a <;> cases b <;> decide end -@[deprecated (since := "2024-06-07")] alias decide_True := decide_true - -@[deprecated (since := "2024-06-07")] alias decide_False := decide_false - -@[deprecated (since := "2024-06-07")] alias coe_decide := decide_eq_true_iff - -@[deprecated decide_eq_true_iff (since := "2024-06-07")] -alias of_decide_iff := decide_eq_true_iff - -@[deprecated (since := "2024-06-07")] alias decide_not := decide_not - -@[deprecated (since := "2024-06-07")] alias not_false' := false_ne_true - -@[deprecated (since := "2024-06-07")] alias eq_iff_eq_true_iff := eq_iff_iff - theorem dichotomy (b : Bool) : b = false ∨ b = true := by cases b <;> simp theorem not_ne_id : not ≠ id := fun h ↦ false_ne_true <| congrFun h true -@[deprecated (since := "2024-06-07")] alias eq_true_of_ne_false := eq_true_of_ne_false - -@[deprecated (since := "2024-06-07")] alias eq_false_of_ne_true := eq_false_of_ne_true - theorem or_inl {a b : Bool} (H : a) : a || b := by simp [H] theorem or_inr {a b : Bool} (H : b) : a || b := by cases a <;> simp [H] @@ -143,8 +120,6 @@ lemma not_eq_iff : ∀ {a b : Bool}, !a = b ↔ a ≠ b := by decide theorem ne_not {a b : Bool} : a ≠ !b ↔ a = b := not_eq_not -@[deprecated (since := "2024-06-07")] alias not_ne := not_not_eq - lemma not_ne_self : ∀ b : Bool, (!b) ≠ b := by decide lemma self_ne_not : ∀ b : Bool, b ≠ !b := by decide diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 2ff347eef5e78..cea4fbe732b1b 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -428,9 +428,6 @@ noncomputable instance instRatCast : RatCast ℂ where ratCast q := ofReal q @[simp, norm_cast] lemma ofReal_nnratCast (q : ℚ≥0) : ofReal q = q := rfl @[simp, norm_cast] lemma ofReal_ratCast (q : ℚ) : ofReal q = q := rfl -@[deprecated (since := "2024-04-17")] -alias ofReal_rat_cast := ofReal_ratCast - @[simp] lemma re_ofNat (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℂ).re = ofNat(n) := rfl @[simp] lemma im_ofNat (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℂ).im = 0 := rfl @@ -452,9 +449,6 @@ lemma im_zsmul (n : ℤ) (z : ℂ) : (n • z).im = n • z.im := smul_im .. @[simp] lemma re_qsmul (q : ℚ) (z : ℂ) : (q • z).re = q • z.re := smul_re .. @[simp] lemma im_qsmul (q : ℚ) (z : ℂ) : (q • z).im = q • z.im := smul_im .. -@[deprecated (since := "2024-04-17")] -alias rat_cast_im := ratCast_im - @[norm_cast] lemma ofReal_nsmul (n : ℕ) (r : ℝ) : ↑(n • r) = n • (r : ℂ) := by simp @[norm_cast] lemma ofReal_zsmul (n : ℤ) (r : ℝ) : ↑(n • r) = n • (r : ℂ) := by simp @@ -488,9 +482,6 @@ theorem conj_I : conj I = -I := theorem conj_natCast (n : ℕ) : conj (n : ℂ) = n := map_natCast _ _ -@[deprecated (since := "2024-04-17")] -alias conj_nat_cast := conj_natCast - theorem conj_ofNat (n : ℕ) [n.AtLeastTwo] : conj (ofNat(n) : ℂ) = ofNat(n) := map_ofNat _ _ @@ -542,21 +533,12 @@ theorem normSq_ofReal (r : ℝ) : normSq r = r * r := by @[simp] theorem normSq_natCast (n : ℕ) : normSq n = n * n := normSq_ofReal _ -@[deprecated (since := "2024-04-17")] -alias normSq_nat_cast := normSq_natCast - @[simp] theorem normSq_intCast (z : ℤ) : normSq z = z * z := normSq_ofReal _ -@[deprecated (since := "2024-04-17")] -alias normSq_int_cast := normSq_intCast - @[simp] theorem normSq_ratCast (q : ℚ) : normSq q = q * q := normSq_ofReal _ -@[deprecated (since := "2024-04-17")] -alias normSq_rat_cast := normSq_ratCast - @[simp] theorem normSq_ofNat (n : ℕ) [n.AtLeastTwo] : normSq (ofNat(n) : ℂ) = ofNat(n) * ofNat(n) := @@ -774,21 +756,12 @@ lemma div_ofReal (z : ℂ) (x : ℝ) : z / x = ⟨z.re / x, z.im / x⟩ := by lemma div_natCast (z : ℂ) (n : ℕ) : z / n = ⟨z.re / n, z.im / n⟩ := mod_cast div_ofReal z n -@[deprecated (since := "2024-04-17")] -alias div_nat_cast := div_natCast - lemma div_intCast (z : ℂ) (n : ℤ) : z / n = ⟨z.re / n, z.im / n⟩ := mod_cast div_ofReal z n -@[deprecated (since := "2024-04-17")] -alias div_int_cast := div_intCast - lemma div_ratCast (z : ℂ) (x : ℚ) : z / x = ⟨z.re / x, z.im / x⟩ := mod_cast div_ofReal z x -@[deprecated (since := "2024-04-17")] -alias div_rat_cast := div_ratCast - lemma div_ofNat (z : ℂ) (n : ℕ) [n.AtLeastTwo] : z / ofNat(n) = ⟨z.re / ofNat(n), z.im / ofNat(n)⟩ := div_natCast z n @@ -802,9 +775,6 @@ lemma div_ofNat (z : ℂ) (n : ℕ) [n.AtLeastTwo] : @[simp] lemma div_ratCast_re (z : ℂ) (x : ℚ) : (z / x).re = z.re / x := by rw [div_ratCast] @[simp] lemma div_ratCast_im (z : ℂ) (x : ℚ) : (z / x).im = z.im / x := by rw [div_ratCast] -@[deprecated (since := "2024-04-17")] -alias div_rat_cast_im := div_ratCast_im - @[simp] lemma div_ofNat_re (z : ℂ) (n : ℕ) [n.AtLeastTwo] : (z / ofNat(n)).re = z.re / ofNat(n) := div_natCast_re z n diff --git a/Mathlib/Data/ENNReal/Basic.lean b/Mathlib/Data/ENNReal/Basic.lean index 485b5a9e784f0..ec49d5c42cb5d 100644 --- a/Mathlib/Data/ENNReal/Basic.lean +++ b/Mathlib/Data/ENNReal/Basic.lean @@ -565,13 +565,6 @@ theorem natCast_lt_coe {n : ℕ} : n < (r : ℝ≥0∞) ↔ n < r := ENNReal.coe theorem coe_lt_natCast {n : ℕ} : (r : ℝ≥0∞) < n ↔ r < n := ENNReal.coe_natCast n ▸ coe_lt_coe -@[deprecated (since := "2024-04-05")] alias coe_nat := coe_natCast -@[deprecated (since := "2024-04-05")] alias ofReal_coe_nat := ofReal_natCast -@[deprecated (since := "2024-04-05")] alias nat_ne_top := natCast_ne_top -@[deprecated (since := "2024-04-05")] alias top_ne_nat := top_ne_natCast -@[deprecated (since := "2024-04-05")] alias coe_nat_lt_coe := natCast_lt_coe -@[deprecated (since := "2024-04-05")] alias coe_lt_coe_nat := coe_lt_natCast - protected theorem exists_nat_gt {r : ℝ≥0∞} (h : r ≠ ∞) : ∃ n : ℕ, r < n := by lift r to ℝ≥0 using h rcases exists_nat_gt r with ⟨n, hn⟩ diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index db3e33069465b..db6b70a9538a9 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -914,9 +914,6 @@ lemma finsetSum_iSup_of_monotone {α ι : Type*} [Preorder ι] [IsDirected ι ( {f : α → ι → ℝ≥0∞} (hf : ∀ a, Monotone (f a)) : (∑ a ∈ s, iSup (f a)) = ⨆ n, ∑ a ∈ s, f a n := finsetSum_iSup fun i j ↦ (exists_ge_ge i j).imp fun _k ⟨hi, hj⟩ a ↦ ⟨hf a hi, hf a hj⟩ -@[deprecated (since := "2024-07-14")] -alias finset_sum_iSup_nat := finsetSum_iSup_of_monotone - lemma le_iInf_mul_iInf {g : κ → ℝ≥0∞} (hf : ∃ i, f i ≠ ∞) (hg : ∃ j, g j ≠ ∞) (ha : ∀ i j, a ≤ f i * g j) : a ≤ (⨅ i, f i) * ⨅ j, g j := by rw [← iInf_ne_top_subtype] diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index df6744aa82316..835dd0a62d9f3 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -325,9 +325,6 @@ theorem sub_ne_top (ha : a ≠ ∞) : a - b ≠ ∞ := mt sub_eq_top_iff.mp <| m theorem natCast_sub (m n : ℕ) : ↑(m - n) = (m - n : ℝ≥0∞) := by rw [← coe_natCast, Nat.cast_tsub, coe_sub, coe_natCast, coe_natCast] -@[deprecated (since := "2024-04-17")] -alias nat_cast_sub := natCast_sub - /-- See `ENNReal.sub_eq_of_eq_add'` for a version assuming that `a = c + b` itself is finite rather than `b`. -/ protected theorem sub_eq_of_eq_add (hb : b ≠ ∞) : a = c + b → a - b = c := diff --git a/Mathlib/Data/ENNReal/Real.lean b/Mathlib/Data/ENNReal/Real.lean index bf44809f4e036..fcb6094a0a231 100644 --- a/Mathlib/Data/ENNReal/Real.lean +++ b/Mathlib/Data/ENNReal/Real.lean @@ -207,9 +207,6 @@ alias ⟨_, ofReal_of_nonpos⟩ := ofReal_eq_zero lemma ofReal_lt_natCast {p : ℝ} {n : ℕ} (hn : n ≠ 0) : ENNReal.ofReal p < n ↔ p < n := by exact mod_cast ofReal_lt_ofReal_iff (Nat.cast_pos.2 hn.bot_lt) -@[deprecated (since := "2024-04-17")] -alias ofReal_lt_nat_cast := ofReal_lt_natCast - @[simp] lemma ofReal_lt_one {p : ℝ} : ENNReal.ofReal p < 1 ↔ p < 1 := by exact mod_cast ofReal_lt_natCast one_ne_zero @@ -223,9 +220,6 @@ lemma ofReal_lt_ofNat {p : ℝ} {n : ℕ} [n.AtLeastTwo] : lemma natCast_le_ofReal {n : ℕ} {p : ℝ} (hn : n ≠ 0) : n ≤ ENNReal.ofReal p ↔ n ≤ p := by simp only [← not_lt, ofReal_lt_natCast hn] -@[deprecated (since := "2024-04-17")] -alias nat_cast_le_ofReal := natCast_le_ofReal - @[simp] lemma one_le_ofReal {p : ℝ} : 1 ≤ ENNReal.ofReal p ↔ 1 ≤ p := by exact mod_cast natCast_le_ofReal one_ne_zero @@ -239,9 +233,6 @@ lemma ofNat_le_ofReal {n : ℕ} [n.AtLeastTwo] {p : ℝ} : lemma ofReal_le_natCast {r : ℝ} {n : ℕ} : ENNReal.ofReal r ≤ n ↔ r ≤ n := coe_le_coe.trans Real.toNNReal_le_natCast -@[deprecated (since := "2024-04-17")] -alias ofReal_le_nat_cast := ofReal_le_natCast - @[simp] lemma ofReal_le_one {r : ℝ} : ENNReal.ofReal r ≤ 1 ↔ r ≤ 1 := coe_le_coe.trans Real.toNNReal_le_one @@ -255,9 +246,6 @@ lemma ofReal_le_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : lemma natCast_lt_ofReal {n : ℕ} {r : ℝ} : n < ENNReal.ofReal r ↔ n < r := coe_lt_coe.trans Real.natCast_lt_toNNReal -@[deprecated (since := "2024-04-17")] -alias nat_cast_lt_ofReal := natCast_lt_ofReal - @[simp] lemma one_lt_ofReal {r : ℝ} : 1 < ENNReal.ofReal r ↔ 1 < r := coe_lt_coe.trans Real.one_lt_toNNReal @@ -270,9 +258,6 @@ lemma ofNat_lt_ofReal {n : ℕ} [n.AtLeastTwo] {r : ℝ} : lemma ofReal_eq_natCast {r : ℝ} {n : ℕ} (h : n ≠ 0) : ENNReal.ofReal r = n ↔ r = n := ENNReal.coe_inj.trans <| Real.toNNReal_eq_natCast h -@[deprecated (since := "2024-04-17")] -alias ofReal_eq_nat_cast := ofReal_eq_natCast - @[simp] lemma ofReal_eq_one {r : ℝ} : ENNReal.ofReal r = 1 ↔ r = 1 := ENNReal.coe_inj.trans Real.toNNReal_eq_one @@ -560,14 +545,9 @@ theorem iInf_sum {α : Type*} {f : ι → α → ℝ≥0∞} {s : Finset α} [No end iInf -section iSup theorem sup_eq_zero {a b : ℝ≥0∞} : a ⊔ b = 0 ↔ a = 0 ∧ b = 0 := sup_eq_bot_iff -@[deprecated (since := "2024-04-05")] alias iSup_coe_nat := iSup_natCast - -end iSup - end ENNReal namespace Mathlib.Meta.Positivity diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 331cbae4643f7..b06760d8dfe35 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -390,9 +390,6 @@ theorem ofNat'_eq_cast (n : ℕ) [NeZero n] (a : ℕ) : Fin.ofNat' n a = a := @[simp] lemma val_natCast (a n : ℕ) [NeZero n] : (a : Fin n).val = a % n := rfl -@[deprecated (since := "2024-04-17")] -alias val_nat_cast := val_natCast - -- Porting note: is this the right name for things involving `Nat.cast`? /-- Converting an in-range number to `Fin (n + 1)` produces a result whose value is the original number. -/ @@ -411,20 +408,12 @@ in the same value. -/ -- This is a special case of `CharP.cast_eq_zero` that doesn't require typeclass search @[simp high] lemma natCast_self (n : ℕ) [NeZero n] : (n : Fin n) = 0 := by ext; simp -@[deprecated (since := "2024-04-17")] -alias nat_cast_self := natCast_self - @[simp] lemma natCast_eq_zero {a n : ℕ} [NeZero n] : (a : Fin n) = 0 ↔ n ∣ a := by simp [Fin.ext_iff, Nat.dvd_iff_mod_eq_zero] -@[deprecated (since := "2024-04-17")] -alias nat_cast_eq_zero := natCast_eq_zero - @[simp] theorem natCast_eq_last (n) : (n : Fin (n + 1)) = Fin.last n := by ext; simp -@[deprecated (since := "2024-05-04")] alias cast_nat_eq_last := natCast_eq_last - theorem le_val_last (i : Fin (n + 1)) : i ≤ n := by rw [Fin.natCast_eq_last] exact Fin.le_last i @@ -646,8 +635,6 @@ theorem le_of_castSucc_lt_of_succ_lt {a b : Fin (n + 1)} {i : Fin n} theorem castSucc_lt_or_lt_succ (p : Fin (n + 1)) (i : Fin n) : castSucc i < p ∨ p < i.succ := by simp [Fin.lt_def, -val_fin_lt]; omega -@[deprecated (since := "2024-05-30")] alias succAbove_lt_gt := castSucc_lt_or_lt_succ - theorem succ_le_or_le_castSucc (p : Fin (n + 1)) (i : Fin n) : succ i ≤ p ∨ p ≤ i.castSucc := by rw [le_castSucc_iff, ← castSucc_lt_iff_succ_le] exact p.castSucc_lt_or_lt_succ i @@ -1047,10 +1034,6 @@ lemma succAbove_ne_last {a : Fin (n + 2)} {b : Fin (n + 1)} (ha : a ≠ last _) lemma succAbove_last_apply (i : Fin n) : succAbove (last n) i = castSucc i := by rw [succAbove_last] -@[deprecated "No deprecation message was provided." (since := "2024-05-30")] -lemma succAbove_lt_ge (p : Fin (n + 1)) (i : Fin n) : - castSucc i < p ∨ p ≤ castSucc i := Nat.lt_or_ge (castSucc i) p - /-- Embedding `i : Fin n` into `Fin (n + 1)` using a pivot `p` that is greater results in a value that is less than `p`. -/ lemma succAbove_lt_iff_castSucc_lt (p : Fin (n + 1)) (i : Fin n) : diff --git a/Mathlib/Data/Fin/Tuple/Basic.lean b/Mathlib/Data/Fin/Tuple/Basic.lean index 49034b164fdbd..41bfa4341ff87 100644 --- a/Mathlib/Data/Fin/Tuple/Basic.lean +++ b/Mathlib/Data/Fin/Tuple/Basic.lean @@ -1013,13 +1013,6 @@ not a definitional equality. -/ @[simp] lemma insertNthEquiv_last (n : ℕ) (α : Type*) : insertNthEquiv (fun _ ↦ α) (last n) = snocEquiv (fun _ ↦ α) := by ext; simp -/-- Separates an `n+1`-tuple, returning a selected index and then the rest of the tuple. -Functional form of `Equiv.piFinSuccAbove`. -/ -@[deprecated removeNth (since := "2024-06-19")] -def extractNth {α : Fin (n + 1) → Type*} (i : Fin (n + 1)) (f : (∀ j, α j)) : - α i × ∀ j, α (i.succAbove j) := - (f i, removeNth i f) - end InsertNth section Find diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 2648f2ebda7ac..6b3fe50a32275 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -350,8 +350,6 @@ lemma card_bij (i : ∀ a ∈ s, β) (hi : ∀ a ha, i a ha ∈ t) · obtain ⟨_, _, rfl⟩ := mem_image.1 h; apply hi · obtain ⟨a, ha, rfl⟩ := i_surj b h; exact mem_image.2 ⟨⟨a, ha⟩, by simp⟩ -@[deprecated (since := "2024-05-04")] alias card_congr := card_bij - /-- Reorder a finset. The difference with `Finset.card_bij` is that the bijection is specified with an inverse, rather @@ -409,8 +407,6 @@ lemma card_le_card_of_injOn (f : α → β) (hf : ∀ a ∈ s, f a ∈ t) (f_inj calc #s = #(s.image f) := (card_image_of_injOn f_inj).symm _ ≤ #t := card_le_card <| image_subset_iff.2 hf -@[deprecated (since := "2024-06-01")] alias card_le_card_of_inj_on := card_le_card_of_injOn - lemma card_le_card_of_injective {f : s → t} (hf : f.Injective) : #s ≤ #t := by rcases s.eq_empty_or_nonempty with rfl | ⟨a₀, ha₀⟩ · simp @@ -589,18 +585,6 @@ lemma exists_subsuperset_card_eq (hst : s ⊆ t) (hsn : #s ≤ n) (hnt : n ≤ # lemma exists_subset_card_eq (hns : n ≤ #s) : ∃ t ⊆ s, #t = n := by simpa using exists_subsuperset_card_eq s.empty_subset (by simp) hns -/-- Given a set `A` and a set `B` inside it, we can shrink `A` to any appropriate size, and keep `B` -inside it. -/ -@[deprecated exists_subsuperset_card_eq (since := "2024-06-23")] -theorem exists_intermediate_set {A B : Finset α} (i : ℕ) (h₁ : i + #B ≤ #A) (h₂ : B ⊆ A) : - ∃ C : Finset α, B ⊆ C ∧ C ⊆ A ∧ #C = i + #B := - exists_subsuperset_card_eq h₂ (Nat.le_add_left ..) h₁ - -/-- We can shrink `A` to any smaller size. -/ -@[deprecated exists_subset_card_eq (since := "2024-06-23")] -theorem exists_smaller_set (A : Finset α) (i : ℕ) (h₁ : i ≤ #A) : - ∃ B : Finset α, B ⊆ A ∧ #B = i := exists_subset_card_eq h₁ - theorem le_card_iff_exists_subset_card : n ≤ #s ↔ ∃ t ⊆ s, #t = n := by refine ⟨fun h => ?_, fun ⟨t, hst, ht⟩ => ht ▸ card_le_card hst⟩ exact exists_subset_card_eq h diff --git a/Mathlib/Data/Finset/NoncommProd.lean b/Mathlib/Data/Finset/NoncommProd.lean index c103380bb1a12..39fd28452094c 100644 --- a/Mathlib/Data/Finset/NoncommProd.lean +++ b/Mathlib/Data/Finset/NoncommProd.lean @@ -174,13 +174,6 @@ theorem map_noncommProd [MonoidHomClass F α β] (s : Multiset α) (comm) (f : F induction s using Quotient.inductionOn simpa using map_list_prod f _ -@[deprecated (since := "2024-07-23")] alias noncommProd_map := map_noncommProd -@[deprecated (since := "2024-07-23")] alias noncommSum_map := map_noncommSum -@[deprecated (since := "2024-07-23")] -protected alias noncommProd_map_aux := Multiset.map_noncommProd_aux -@[deprecated (since := "2024-07-23")] -protected alias noncommSum_map_aux := Multiset.map_noncommSum_aux - @[to_additive noncommSum_eq_card_nsmul] theorem noncommProd_eq_pow_card (s : Multiset α) (comm) (m : α) (h : ∀ x ∈ s, x = m) : s.noncommProd comm = m ^ Multiset.card s := by @@ -321,9 +314,6 @@ theorem map_noncommProd [MonoidHomClass F β γ] (s : Finset α) (f : α → β) s.noncommProd (fun i => g (f i)) fun _ hx _ hy _ => (comm.of_refl hx hy).map g := by simp [noncommProd, Multiset.map_noncommProd] -@[deprecated (since := "2024-07-23")] alias noncommProd_map := map_noncommProd -@[deprecated (since := "2024-07-23")] alias noncommSum_map := map_noncommSum - @[to_additive noncommSum_eq_card_nsmul] theorem noncommProd_eq_pow_card (s : Finset α) (f : α → β) (comm) (m : β) (h : ∀ x ∈ s, f x = m) : s.noncommProd f comm = m ^ s.card := by diff --git a/Mathlib/Data/Finset/SDiff.lean b/Mathlib/Data/Finset/SDiff.lean index 8e0c211942a4a..dc90b5d429031 100644 --- a/Mathlib/Data/Finset/SDiff.lean +++ b/Mathlib/Data/Finset/SDiff.lean @@ -85,9 +85,6 @@ lemma sdiff_inter_right_comm (s t u : Finset α) : s \ t ∩ u = (s ∩ u) \ t : lemma inter_sdiff_left_comm (s t u : Finset α) : s ∩ (t \ u) = t ∩ (s \ u) := inf_sdiff_left_comm .. -@[deprecated inter_sdiff_assoc (since := "2024-05-01")] -theorem inter_sdiff (s t u : Finset α) : s ∩ (t \ u) = (s ∩ t) \ u := (inter_sdiff_assoc _ _ _).symm - @[simp] theorem sdiff_inter_self (s₁ s₂ : Finset α) : s₂ \ s₁ ∩ s₁ = ∅ := inf_sdiff_self_left diff --git a/Mathlib/Data/Finsupp/Weight.lean b/Mathlib/Data/Finsupp/Weight.lean index 8209db86bda0f..c66c9b61ae184 100644 --- a/Mathlib/Data/Finsupp/Weight.lean +++ b/Mathlib/Data/Finsupp/Weight.lean @@ -75,15 +75,9 @@ with respect to `w : σ → M` is the sum `∑ i, f i • w i`. -/ noncomputable def weight : (σ →₀ R) →+ M := (Finsupp.linearCombination R w).toAddMonoidHom -@[deprecated weight (since := "2024-07-20")] -alias _root_.MvPolynomial.weightedDegree := weight - theorem weight_apply (f : σ →₀ R) : weight w f = Finsupp.sum f (fun i c => c • w i) := rfl -@[deprecated weight_apply (since := "2024-07-20")] -alias _root_.MvPolynomial.weightedDegree_apply := weight_apply - theorem weight_single_index [DecidableEq σ] (s : σ) (c : M) (f : σ →₀ R) : weight (Pi.single s c) f = f s • c := linearCombination_single_index σ M R c s f @@ -213,9 +207,6 @@ variable {R : Type*} [AddCommMonoid R] /-- The degree of a finsupp function. -/ def degree (d : σ →₀ R) : R := ∑ i ∈ d.support, d i -@[deprecated degree (since := "2024-07-20")] -alias _root_.MvPolynomial.degree := degree - @[simp] theorem degree_add (a b : σ →₀ R) : (a + b).degree = a.degree + b.degree := sum_add_index' (h := fun _ ↦ id) (congrFun rfl) fun _ _ ↦ congrFun rfl @@ -233,9 +224,6 @@ lemma degree_eq_zero_iff {R : Type*} [OrderedAddCommMonoid R] [CanonicallyOrdere simp only [degree, Finset.sum_eq_zero_iff, mem_support_iff, ne_eq, _root_.not_imp_self, DFunLike.ext_iff, coe_zero, Pi.zero_apply] -@[deprecated degree_eq_zero_iff (since := "2024-07-20")] -alias _root_.MvPolynomial.degree_eq_zero_iff := degree_eq_zero_iff - theorem le_degree {R : Type*} [OrderedAddCommMonoid R] [CanonicallyOrderedAdd R] (s : σ) (f : σ →₀ R) : f s ≤ degree f := by @@ -249,9 +237,6 @@ theorem degree_eq_weight_one {R : Type*} [Semiring R] : ext d simp only [degree, weight_apply, Pi.one_apply, smul_eq_mul, mul_one, Finsupp.sum] -@[deprecated degree_eq_weight_one (since := "2024-07-20")] -alias _root_.MvPolynomial.weightedDegree_one := degree_eq_weight_one - theorem finite_of_degree_le [Finite σ] (n : ℕ) : {f : σ →₀ ℕ | degree f ≤ n}.Finite := by simp_rw [degree_eq_weight_one] diff --git a/Mathlib/Data/Int/AbsoluteValue.lean b/Mathlib/Data/Int/AbsoluteValue.lean index ac42e74e70507..c6407d72caaac 100644 --- a/Mathlib/Data/Int/AbsoluteValue.lean +++ b/Mathlib/Data/Int/AbsoluteValue.lean @@ -31,9 +31,6 @@ theorem AbsoluteValue.map_units_int (abv : AbsoluteValue ℤ S) (x : ℤˣ) : ab theorem AbsoluteValue.map_units_intCast [Nontrivial R] (abv : AbsoluteValue R S) (x : ℤˣ) : abv ((x : ℤ) : R) = 1 := by rcases Int.units_eq_one_or x with (rfl | rfl) <;> simp -@[deprecated (since := "2024-04-17")] -alias AbsoluteValue.map_units_int_cast := AbsoluteValue.map_units_intCast - @[simp] theorem AbsoluteValue.map_units_int_smul (abv : AbsoluteValue R S) (x : ℤˣ) (y : R) : abv (x • y) = abv y := by rcases Int.units_eq_one_or x with (rfl | rfl) <;> simp diff --git a/Mathlib/Data/Int/Cast/Lemmas.lean b/Mathlib/Data/Int/Cast/Lemmas.lean index 628317d302d5e..93be770cd9c1b 100644 --- a/Mathlib/Data/Int/Cast/Lemmas.lean +++ b/Mathlib/Data/Int/Cast/Lemmas.lean @@ -121,8 +121,6 @@ end Ring theorem cast_dvd_cast [CommRing α] (m n : ℤ) (h : m ∣ n) : (m : α) ∣ (n : α) := RingHom.map_dvd (Int.castRingHom α) h -@[deprecated (since := "2024-05-25")] alias coe_int_dvd := cast_dvd_cast - end cast end Int @@ -141,10 +139,6 @@ variable [Ring α] {a x y : α} @[simp] lemma intCast_mul_intCast_mul (h : SemiconjBy a x y) (m n : ℤ) : SemiconjBy (m * a) (n * x) (n * y) := (h.intCast_mul_left m).intCast_mul_right n -@[deprecated (since := "2024-05-27")] alias cast_int_mul_right := intCast_mul_right -@[deprecated (since := "2024-05-27")] alias cast_int_mul_left := intCast_mul_left -@[deprecated (since := "2024-05-27")] alias cast_int_mul_cast_int_mul := intCast_mul_intCast_mul - end SemiconjBy namespace Commute @@ -155,9 +149,6 @@ variable [NonAssocRing α] {a : α} {n : ℤ} @[simp] lemma intCast_right : Commute a n := Int.commute_cast _ _ -@[deprecated (since := "2024-05-27")] alias cast_int_right := intCast_right -@[deprecated (since := "2024-05-27")] alias cast_int_left := intCast_left - end NonAssocRing section Ring @@ -181,14 +172,6 @@ lemma intCast_mul_self : Commute ((n : α) * a) a := (Commute.refl a).intCast_mu lemma self_intCast_mul_intCast_mul : Commute (m * a : α) (n * a : α) := (Commute.refl a).intCast_mul_intCast_mul m n -@[deprecated (since := "2024-05-27")] alias cast_int_mul_right := intCast_mul_right -@[deprecated (since := "2024-05-27")] alias cast_int_mul_left := intCast_mul_left -@[deprecated (since := "2024-05-27")] alias cast_int_mul_cast_int_mul := intCast_mul_intCast_mul -@[deprecated (since := "2024-05-27")] alias self_cast_int_mul := self_intCast_mul -@[deprecated (since := "2024-05-27")] alias cast_int_mul_self := intCast_mul_self -@[deprecated (since := "2024-05-27")] -alias self_cast_int_mul_cast_int_mul := self_intCast_mul_intCast_mul - end Ring end Commute @@ -211,9 +194,6 @@ variable [AddGroupWithOne A] theorem eq_intCastAddHom (f : ℤ →+ A) (h1 : f 1 = 1) : f = Int.castAddHom A := ext_int <| by simp [h1] -@[deprecated (since := "2024-04-17")] -alias eq_int_castAddHom := eq_intCastAddHom - end AddMonoidHom theorem eq_intCast' [AddGroupWithOne α] [FunLike F ℤ α] [AddMonoidHomClass F ℤ α] diff --git a/Mathlib/Data/Int/Cast/Pi.lean b/Mathlib/Data/Int/Cast/Pi.lean index 654d33c17c3f9..06d0ce120185e 100644 --- a/Mathlib/Data/Int/Cast/Pi.lean +++ b/Mathlib/Data/Int/Cast/Pi.lean @@ -33,9 +33,6 @@ theorem intCast_apply (n : ℤ) (i : ι) : (n : ∀ i, π i) i = n := theorem intCast_def (n : ℤ) : (n : ∀ i, π i) = fun _ => ↑n := rfl -@[deprecated (since := "2024-04-05")] alias int_apply := intCast_apply -@[deprecated (since := "2024-04-05")] alias coe_int := intCast_def - end Pi theorem Sum.elim_intCast_intCast {α β γ : Type*} [IntCast γ] (n : ℤ) : diff --git a/Mathlib/Data/Int/CharZero.lean b/Mathlib/Data/Int/CharZero.lean index 49a66742d5cd6..2580be563498e 100644 --- a/Mathlib/Data/Int/CharZero.lean +++ b/Mathlib/Data/Int/CharZero.lean @@ -44,13 +44,7 @@ variable [AddGroupWithOne β] [CharZero β] {n : ℤ} lemma support_intCast (hn : n ≠ 0) : support (n : α → β) = univ := support_const <| Int.cast_ne_zero.2 hn -@[deprecated (since := "2024-04-17")] -alias support_int_cast := support_intCast - lemma mulSupport_intCast (hn : n ≠ 1) : mulSupport (n : α → β) = univ := mulSupport_const <| Int.cast_ne_one.2 hn -@[deprecated (since := "2024-04-17")] -alias mulSupport_int_cast := mulSupport_intCast - end Function diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 74a86d6296cf7..5be28e0bf2570 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -499,10 +499,6 @@ lemma ofNat_add_negSucc_of_ge {m n : ℕ} (h : n.succ ≤ m) : lemma natAbs_le_of_dvd_ne_zero (hmn : m ∣ n) (hn : n ≠ 0) : natAbs m ≤ natAbs n := not_lt.mp (mt (eq_zero_of_dvd_of_natAbs_lt_natAbs hmn) hn) -@[deprecated (since := "2024-04-02")] alias coe_nat_dvd := natCast_dvd_natCast -@[deprecated (since := "2024-04-02")] alias coe_nat_dvd_right := dvd_natCast -@[deprecated (since := "2024-04-02")] alias coe_nat_dvd_left := natCast_dvd - /-! #### `/` and ordering -/ lemma natAbs_eq_of_dvd_dvd (hmn : m ∣ n) (hnm : n ∣ m) : natAbs m = natAbs n := @@ -562,9 +558,6 @@ lemma lt_of_toNat_lt {a b : ℤ} (h : toNat a < toNat b) : a < b := theorem toNat_sub_of_le {a b : ℤ} (h : b ≤ a) : (toNat (a - b) : ℤ) = a - b := Int.toNat_of_nonneg (Int.sub_nonneg_of_le h) -@[deprecated (since := "2024-04-05")] alias coe_nat_pos := natCast_pos -@[deprecated (since := "2024-04-05")] alias coe_nat_succ_pos := natCast_succ_pos - lemma toNat_lt' {n : ℕ} (hn : n ≠ 0) : m.toNat < n ↔ m < n := by rw [← toNat_lt_toNat, toNat_natCast]; omega @@ -576,23 +569,7 @@ lemma natMod_lt {n : ℕ} (hn : n ≠ 0) : m.natMod n < n := attribute [simp] natCast_pow -@[deprecated (since := "2024-05-25")] alias coe_nat_pow := natCast_pow - -- Porting note: this was added in an ad hoc port for use in `Tactic/NormNum/Basic` @[simp] lemma pow_eq (m : ℤ) (n : ℕ) : m.pow n = m ^ n := rfl -@[deprecated (since := "2024-04-02")] alias ofNat_eq_cast := ofNat_eq_natCast -@[deprecated (since := "2024-04-02")] alias cast_eq_cast_iff_Nat := natCast_inj -@[deprecated (since := "2024-04-02")] alias coe_nat_sub := Int.natCast_sub -@[deprecated (since := "2024-04-02")] alias coe_nat_nonneg := natCast_nonneg -@[deprecated (since := "2024-04-02")] alias sign_coe_add_one := sign_natCast_add_one -@[deprecated (since := "2024-04-02")] alias nat_succ_eq_int_succ := natCast_succ -@[deprecated (since := "2024-04-02")] alias succ_neg_nat_succ := succ_neg_natCast_succ -@[deprecated (since := "2024-04-02")] alias coe_pred_of_pos := natCast_pred_of_pos -@[deprecated (since := "2024-04-02")] alias coe_nat_div := natCast_div -@[deprecated (since := "2024-04-02")] alias coe_nat_ediv := natCast_ediv -@[deprecated (since := "2024-04-02")] alias sign_coe_nat_of_nonzero := sign_natCast_of_ne_zero -@[deprecated (since := "2024-04-02")] alias toNat_coe_nat := toNat_natCast -@[deprecated (since := "2024-04-02")] alias toNat_coe_nat_add_one := toNat_natCast_add_one - end Int diff --git a/Mathlib/Data/Int/GCD.lean b/Mathlib/Data/Int/GCD.lean index 2216a7674165e..22963db04965e 100644 --- a/Mathlib/Data/Int/GCD.lean +++ b/Mathlib/Data/Int/GCD.lean @@ -153,8 +153,6 @@ theorem gcd_def (i j : ℤ) : gcd i j = Nat.gcd i.natAbs j.natAbs := rfl @[simp, norm_cast] protected lemma gcd_natCast_natCast (m n : ℕ) : gcd ↑m ↑n = m.gcd n := rfl -@[deprecated (since := "2024-05-25")] alias coe_nat_gcd := Int.gcd_natCast_natCast - /-- The extended GCD `a` value in the equation `gcd x y = x * a + y * b`. -/ def gcdA : ℤ → ℤ → ℤ | ofNat m, n => m.gcdA n.natAbs diff --git a/Mathlib/Data/Int/Lemmas.lean b/Mathlib/Data/Int/Lemmas.lean index c51bbeb9c0106..4671f26ec4eb0 100644 --- a/Mathlib/Data/Int/Lemmas.lean +++ b/Mathlib/Data/Int/Lemmas.lean @@ -121,15 +121,6 @@ theorem div2_bit (b n) : div2 (bit b n) = n := by rw [Nat.div_eq_of_lt] <;> simp · decide -@[deprecated (since := "2024-04-02")] alias le_coe_nat_sub := le_natCast_sub -@[deprecated (since := "2024-04-02")] alias succ_coe_nat_pos := succ_natCast_pos -@[deprecated (since := "2024-04-02")] alias coe_natAbs := natCast_natAbs -@[deprecated (since := "2024-04-02")] alias coe_nat_eq_zero := natCast_eq_zero -@[deprecated (since := "2024-04-02")] alias coe_nat_ne_zero := natCast_ne_zero -@[deprecated (since := "2024-04-02")] alias coe_nat_ne_zero_iff_pos := natCast_ne_zero_iff_pos -@[deprecated (since := "2024-04-02")] alias abs_coe_nat := abs_natCast -@[deprecated (since := "2024-04-02")] alias coe_nat_nonpos_iff := natCast_nonpos_iff - /-- Like `Int.ediv_emod_unique`, but permitting negative `b`. -/ theorem ediv_emod_unique' {a b r q : Int} (h : b ≠ 0) : a / b = q ∧ a % b = r ↔ r + b * q = a ∧ 0 ≤ r ∧ r < |b| := by diff --git a/Mathlib/Data/Int/ModEq.lean b/Mathlib/Data/Int/ModEq.lean index 8374afec3980a..2ee262f6a9d0b 100644 --- a/Mathlib/Data/Int/ModEq.lean +++ b/Mathlib/Data/Int/ModEq.lean @@ -266,6 +266,4 @@ theorem mod_mul_right_mod (a b c : ℤ) : a % (b * c) % b = a % b := theorem mod_mul_left_mod (a b c : ℤ) : a % (b * c) % c = a % c := (mod_modEq _ _).of_mul_left _ -@[deprecated (since := "2024-04-02")] alias coe_nat_modEq_iff := natCast_modEq_iff - end Int diff --git a/Mathlib/Data/Int/Order/Basic.lean b/Mathlib/Data/Int/Order/Basic.lean index 04758c04bf6f5..8d8e534deec64 100644 --- a/Mathlib/Data/Int/Order/Basic.lean +++ b/Mathlib/Data/Int/Order/Basic.lean @@ -40,9 +40,6 @@ instance instLinearOrder : LinearOrder ℤ where decidableLE := by infer_instance decidableLT := by infer_instance -@[deprecated (since := "2024-07-27")] alias mul_neg_eq_neg_mul_symm := Int.mul_neg -@[deprecated (since := "2024-07-27")] alias neg_mul_eq_neg_mul_symm := Int.neg_mul - protected theorem eq_zero_or_eq_zero_of_mul_eq_zero {a b : ℤ} (h : a * b = 0) : a = 0 ∨ b = 0 := Int.mul_eq_zero.mp h diff --git a/Mathlib/Data/Int/SuccPred.lean b/Mathlib/Data/Int/SuccPred.lean index 51f82b5b2556a..ce0979c26e99d 100644 --- a/Mathlib/Data/Int/SuccPred.lean +++ b/Mathlib/Data/Int/SuccPred.lean @@ -83,6 +83,3 @@ theorem natCast_covBy {a b : ℕ} : (a : ℤ) ⋖ b ↔ a ⋖ b := by end Int alias ⟨_, CovBy.intCast⟩ := Int.natCast_covBy - -@[deprecated (since := "2024-05-27")] alias Nat.cast_int_covBy_iff := Int.natCast_covBy -@[deprecated (since := "2024-05-27")] alias CovBy.cast_int := CovBy.intCast diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 187b3b0cb21bc..569361c5a5b42 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -37,10 +37,6 @@ universe u v w variable {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {l₁ l₂ : List α} -/-- `≤` implies not `>` for lists. -/ -@[deprecated "No deprecation message was provided." (since := "2024-07-27")] -theorem le_eq_not_gt [LT α] : ∀ l₁ l₂ : List α, (l₁ ≤ l₂) = ¬l₂ < l₁ := fun _ _ => rfl - -- Porting note: Delete this attribute -- attribute [inline] List.head! @@ -347,8 +343,6 @@ lemma getLast_filter' {p : α → Bool} : @[deprecated (since := "2024-09-06")] alias getLast?_eq_none := getLast?_eq_none_iff -@[deprecated (since := "2024-06-20")] alias getLast?_isNone := getLast?_eq_none - theorem mem_getLast?_eq_getLast : ∀ {l : List α} {x : α}, x ∈ l.getLast? → ∃ h, x = getLast l h | [], x, hx => False.elim <| by simp at hx | [a], x, hx => @@ -624,11 +618,6 @@ lemma cons_sublist_cons' {a b : α} : a :: l₁ <+ b :: l₂ ↔ a :: l₁ <+ l theorem sublist_cons_of_sublist (a : α) (h : l₁ <+ l₂) : l₁ <+ a :: l₂ := h.cons _ -@[deprecated "No deprecation message was provided." (since := "2024-04-07")] -theorem sublist_of_cons_sublist_cons {a} (h : a :: l₁ <+ a :: l₂) : l₁ <+ l₂ := h.of_cons_cons - -@[deprecated (since := "2024-04-07")] alias cons_sublist_cons_iff := cons_sublist_cons - -- Porting note: this lemma seems to have been renamed on the occasion of its move to Batteries alias sublist_nil_iff_eq_nil := sublist_nil @@ -743,20 +732,10 @@ section deprecated @[simp] theorem getElem?_length (l : List α) : l[l.length]? = none := getElem?_eq_none le_rfl -@[deprecated getElem?_length (since := "2024-06-12")] -theorem get?_length (l : List α) : l.get? l.length = none := get?_eq_none le_rfl - -@[deprecated (since := "2024-05-03")] alias get?_injective := get?_inj - /-- A version of `getElem_map` that can be used for rewriting. -/ theorem getElem_map_rev (f : α → β) {l} {n : Nat} {h : n < l.length} : f l[n] = (map f l)[n]'((l.length_map f).symm ▸ h) := Eq.symm (getElem_map _) -/-- A version of `get_map` that can be used for rewriting. -/ -@[deprecated getElem_map_rev (since := "2024-06-12")] -theorem get_map_rev (f : α → β) {l n} : - f (get l n) = get (map f l) ⟨n.1, (l.length_map f).symm ▸ n.2⟩ := Eq.symm (getElem_map _) - theorem get_length_sub_one {l : List α} (h : l.length - 1 < l.length) : l.get ⟨l.length - 1, h⟩ = l.getLast (by rintro rfl; exact Nat.lt_irrefl 0 h) := (getLast_eq_getElem l _).symm @@ -826,14 +805,6 @@ theorem indexOf_inj [DecidableEq α] {l : List α} {x y : α} (hx : x ∈ l) (hy simp only [h] simp only [indexOf_get] at x_eq_y; exact x_eq_y, fun h => by subst h; rfl⟩ -@[deprecated getElem_reverse (since := "2024-06-12")] -theorem get_reverse (l : List α) (i : Nat) (h1 h2) : - get (reverse l) ⟨length l - 1 - i, h1⟩ = get l ⟨i, h2⟩ := by - rw [get_eq_getElem, get_eq_getElem, getElem_reverse] - congr - dsimp - omega - theorem get_reverse' (l : List α) (n) (hn') : l.reverse.get n = l.get ⟨l.length - 1 - n, hn'⟩ := by simp @@ -851,8 +822,6 @@ alias modifyNthTail_modifyNthTail_le := modifyTailIdx_modifyTailIdx_le @[deprecated (since := "2024-10-21")] alias modifyNthTail_modifyNthTail_same := modifyTailIdx_modifyTailIdx_self -@[deprecated (since := "2024-05-04")] alias removeNth_eq_nthTail := eraseIdx_eq_modifyTailIdx - @[deprecated (since := "2024-10-21")] alias modifyNth_eq_set := modify_eq_set @[simp] @@ -862,12 +831,6 @@ theorem getElem_set_of_ne {l : List α} {i j : ℕ} (h : i ≠ j) (a : α) rw [← Option.some_inj, ← List.getElem?_eq_getElem, List.getElem?_set_ne h, List.getElem?_eq_getElem] -@[deprecated getElem_set_of_ne (since := "2024-06-12")] -theorem get_set_of_ne {l : List α} {i j : ℕ} (h : i ≠ j) (a : α) - (hj : j < (l.set i a).length) : - (l.set i a).get ⟨j, hj⟩ = l.get ⟨j, by simpa using hj⟩ := by - simp [getElem_set_of_ne, h] - /-! ### map -/ -- `List.map_const` (the version with `Function.const` instead of a lambda) is already tagged @@ -875,8 +838,6 @@ theorem get_set_of_ne {l : List α} {i j : ℕ} (h : i ≠ j) (a : α) -- TODO: Upstream the tagging to Core? attribute [simp] map_const' -@[deprecated (since := "2024-06-21")] alias map_congr := map_congr_left - theorem flatMap_pure_eq_map (f : α → β) (l : List α) : l.flatMap (pure ∘ f) = map f l := .symm <| map_eq_flatMap .. @@ -1184,20 +1145,12 @@ theorem getElem?_scanl_zero : (scanl f b l)[0]? = some b := by · simp [scanl_nil] · simp [scanl_cons, singleton_append] -@[deprecated getElem?_scanl_zero (since := "2024-06-12")] -theorem get?_zero_scanl : (scanl f b l).get? 0 = some b := by - simp [getElem?_scanl_zero] - @[simp] theorem getElem_scanl_zero {h : 0 < (scanl f b l).length} : (scanl f b l)[0] = b := by cases l · simp [scanl_nil] · simp [scanl_cons, singleton_append] -@[deprecated getElem_scanl_zero (since := "2024-06-12")] -theorem get_zero_scanl {h : 0 < (scanl f b l).length} : (scanl f b l).get ⟨0, h⟩ = b := by - simp [getElem_scanl_zero] - theorem get?_succ_scanl {i : ℕ} : (scanl f b l).get? (i + 1) = ((scanl f b l).get? i).bind fun x => (l.get? i).map fun y => f x y := by induction' l with hd tl hl generalizing b i @@ -1538,17 +1491,6 @@ theorem sizeOf_lt_sizeOf_of_mem [SizeOf α] {x : α} {l : List α} (hx : x ∈ l · specialize ih ‹_› omega -@[deprecated attach_map_coe (since := "2024-07-29")] alias attach_map_coe' := attach_map_coe -@[deprecated attach_map_val (since := "2024-07-29")] alias attach_map_val' := attach_map_val - -/-! ### find -/ - -section find? - -@[deprecated (since := "2024-05-05")] alias find?_mem := mem_of_find?_eq_some - -end find? - /-! ### lookmap -/ section Lookmap @@ -2012,23 +1954,6 @@ theorem length_dropSlice_lt (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs (List.dropSlice i j xs).length < xs.length := by simp; omega -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-07-25")] -theorem sizeOf_dropSlice_lt [SizeOf α] (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs.length) : - SizeOf.sizeOf (List.dropSlice i j xs) < SizeOf.sizeOf xs := by - induction xs generalizing i j hj with - | nil => cases hi - | cons x xs xs_ih => - cases i <;> simp only [List.dropSlice] - · cases j with - | zero => contradiction - | succ n => - dsimp only [drop]; apply lt_of_le_of_lt (drop_sizeOf_le xs n) - simp only [cons.sizeOf_spec]; omega - · simp only [cons.sizeOf_spec, Nat.add_lt_add_iff_left] - apply xs_ih _ j hj - apply lt_of_succ_lt_succ hi - section Disjoint /-- The images of disjoint lists under a partially defined map are disjoint -/ diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index a0351827a88e0..ef8b393fe73e1 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -50,10 +50,6 @@ def getLastI [Inhabited α] : List α → α | [_, b] => b | _ :: _ :: l => getLastI l -/-- List with a single given element. -/ -@[inline, deprecated List.pure (since := "2024-03-24")] -protected def ret {α : Type u} (a : α) : List α := [a] - /-- "Inhabited" `take` function: Take `n` elements from a list `l`. If `l` has less than `n` elements, append `n - length l` elements `default`. -/ def takeI [Inhabited α] (n : Nat) (l : List α) : List α := diff --git a/Mathlib/Data/List/Enum.lean b/Mathlib/Data/List/Enum.lean index a574ea756a233..38f58ee6b8223 100644 --- a/Mathlib/Data/List/Enum.lean +++ b/Mathlib/Data/List/Enum.lean @@ -26,14 +26,10 @@ theorem get?_enumFrom (n) (l : List α) (m) : get? (enumFrom n l) m = (get? l m).map fun a => (n + m, a) := by simp -@[deprecated (since := "2024-04-06")] alias enumFrom_get? := get?_enumFrom - @[deprecated getElem?_enum (since := "2024-08-15")] theorem get?_enum (l : List α) (n) : get? (enum l) n = (get? l n).map fun a => (n, a) := by simp -@[deprecated (since := "2024-04-06")] alias enum_get? := get?_enum - @[deprecated getElem_enumFrom (since := "2024-08-15")] theorem get_enumFrom (l : List α) (n) (i : Fin (l.enumFrom n).length) : (l.enumFrom n).get i = (n + i, l.get (i.cast enumFrom_length)) := by diff --git a/Mathlib/Data/List/Flatten.lean b/Mathlib/Data/List/Flatten.lean index 797ba9dc5c297..6f8beb4c8bad3 100644 --- a/Mathlib/Data/List/Flatten.lean +++ b/Mathlib/Data/List/Flatten.lean @@ -109,11 +109,6 @@ theorem drop_take_succ_eq_cons_getElem (L : List α) (i : Nat) (h : i < L.length · simp · simpa using ih _ (by simpa using h) -@[deprecated drop_take_succ_eq_cons_getElem (since := "2024-06-11")] -theorem drop_take_succ_eq_cons_get (L : List α) (i : Fin L.length) : - (L.take (i + 1)).drop i = [get L i] := by - simp [drop_take_succ_eq_cons_getElem] - set_option linter.deprecated false in /-- In a flatten of sublists, taking the slice between the indices `A` and `B - 1` gives back the original sublist of index `i` if `A` is the sum of the lengths of sublists of index `< i`, and @@ -132,13 +127,6 @@ theorem drop_take_succ_flatten_eq_getElem' (L : List (List α)) (i : Nat) (h : i @[deprecated (since := "2024-10-15")] alias drop_take_succ_join_eq_getElem' := drop_take_succ_flatten_eq_getElem' -set_option linter.deprecated false in -@[deprecated drop_take_succ_flatten_eq_getElem' (since := "2024-06-11")] -theorem drop_take_succ_join_eq_get' (L : List (List α)) (i : Fin L.length) : - (L.flatten.take (Nat.sum ((L.map length).take (i + 1)))).drop - (Nat.sum ((L.map length).take i)) = get L i := by - simp [drop_take_succ_flatten_eq_getElem'] - theorem flatten_drop_length_sub_one {L : List (List α)} (h : L ≠ []) : (L.drop (L.length - 1)).flatten = L.getLast h := by induction L using List.reverseRecOn diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index 9e93d425fe171..6a47041fdcfea 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -62,18 +62,12 @@ def decidableGetDNilNe (a : α) : DecidablePred fun i : ℕ => getD ([] : List @[simp] theorem getElem?_getD_singleton_default_eq (n : ℕ) : [d][n]?.getD d = d := by cases n <;> simp -@[deprecated (since := "2024-06-12")] -alias getD_singleton_default_eq := getElem?_getD_singleton_default_eq - @[simp] theorem getElem?_getD_replicate_default_eq (r n : ℕ) : (replicate r d)[n]?.getD d = d := by induction r generalizing n with | zero => simp | succ n ih => simp at ih; cases n <;> simp [ih, replicate_succ] -@[deprecated (since := "2024-06-12")] -alias getD_replicate_default_eq := getElem?_getD_replicate_default_eq - theorem getD_replicate {y i n} (h : i < n) : getD (replicate n x) i y = x := by rw [getD_eq_getElem, getElem_replicate] diff --git a/Mathlib/Data/List/InsertIdx.lean b/Mathlib/Data/List/InsertIdx.lean index 431b9703de66b..97beb602d1bac 100644 --- a/Mathlib/Data/List/InsertIdx.lean +++ b/Mathlib/Data/List/InsertIdx.lean @@ -63,11 +63,8 @@ theorem insertIdx_injective (n : ℕ) (x : α) : Function.Injective (insertIdx n @[deprecated (since := "2024-10-21")] alias insertNth_succ_cons := insertIdx_succ_cons @[deprecated (since := "2024-10-21")] alias length_insertNth := length_insertIdx @[deprecated (since := "2024-10-21")] alias removeNth_insertIdx := eraseIdx_insertIdx -@[deprecated (since := "2024-05-04")] alias removeNth_insertNth := eraseIdx_insertIdx @[deprecated (since := "2024-10-21")] alias insertNth_eraseIdx_of_ge := insertIdx_eraseIdx_of_ge -@[deprecated (since := "2024-05-04")] alias insertNth_removeNth_of_ge := insertIdx_eraseIdx_of_ge @[deprecated (since := "2024-10-21")] alias insertNth_eraseIdx_of_le := insertIdx_eraseIdx_of_le -@[deprecated (since := "2024-05-04")] alias insertIdx_removeNth_of_le := insertIdx_eraseIdx_of_le @[deprecated (since := "2024-10-21")] alias insertNth_comm := insertIdx_comm @[deprecated (since := "2024-10-21")] alias mem_insertNth := mem_insertIdx @[deprecated (since := "2024-10-21")] alias insertNth_of_length_lt := insertIdx_of_length_lt diff --git a/Mathlib/Data/List/Lex.lean b/Mathlib/Data/List/Lex.lean index 4c166bc4b024a..bb22beeacd721 100644 --- a/Mathlib/Data/List/Lex.lean +++ b/Mathlib/Data/List/Lex.lean @@ -88,11 +88,6 @@ instance isAsymm (r : α → α → Prop) [IsAsymm α r] : IsAsymm (List α) (Le | _, _, Lex.cons _, Lex.rel h₂ => asymm h₂ h₂ | _, _, Lex.cons h₁, Lex.cons h₂ => aux _ _ h₁ h₂ -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -instance isStrictTotalOrder (r : α → α → Prop) [IsStrictTotalOrder α r] : - IsStrictTotalOrder (List α) (Lex r) := - { isStrictWeakOrder_of_isOrderConnected with } - instance decidableRel [DecidableEq α] (r : α → α → Prop) [DecidableRel r] : DecidableRel (Lex r) | l₁, [] => isFalse fun h => by cases h | [], _ :: _ => isTrue Lex.nil @@ -147,6 +142,12 @@ end Lex instance LT' [LT α] : LT (List α) := ⟨Lex (· < ·)⟩ +-- TODO: This deprecated instance is still used (by the instance just below) +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] +instance isStrictTotalOrder (r : α → α → Prop) [IsStrictTotalOrder α r] : + IsStrictTotalOrder (List α) (Lex r) := + { isStrictWeakOrder_of_isOrderConnected with } + instance [LinearOrder α] : LinearOrder (List α) := linearOrderOfSTO (Lex (· < ·)) diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index 241a87b598d88..1adc6b35b85fc 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -281,16 +281,10 @@ theorem minimum_mem {l : List α} {m : α} : (minimum l : WithBot α) = m → m theorem maximum_eq_bot {l : List α} : l.maximum = ⊥ ↔ l = [] := argmax_eq_none -@[simp, deprecated maximum_eq_bot "Don't mix Option and WithBot" (since := "2024-05-27")] -theorem maximum_eq_none {l : List α} : l.maximum = none ↔ l = [] := maximum_eq_bot - @[simp] theorem minimum_eq_top {l : List α} : l.minimum = ⊤ ↔ l = [] := argmin_eq_none -@[simp, deprecated minimum_eq_top "Don't mix Option and WithTop" (since := "2024-05-27")] -theorem minimum_eq_none {l : List α} : l.minimum = none ↔ l = [] := minimum_eq_top - theorem not_maximum_lt_of_mem : a ∈ l → (maximum l : WithBot α) = m → ¬m < a := not_lt_of_mem_argmax diff --git a/Mathlib/Data/List/OfFn.lean b/Mathlib/Data/List/OfFn.lean index ac14755a5fd6d..0e77e65fdc740 100644 --- a/Mathlib/Data/List/OfFn.lean +++ b/Mathlib/Data/List/OfFn.lean @@ -123,10 +123,6 @@ theorem ofFn_getElem_eq_map {β : Type*} (l : List α) (f : α → β) : ofFn (fun i : Fin l.length => f <| l[(i : Nat)]) = l.map f := by rw [← Function.comp_def, ← map_ofFn, ofFn_getElem] -@[deprecated ofFn_getElem_eq_map (since := "2024-06-12")] -theorem ofFn_get_eq_map {β : Type*} (l : List α) (f : α → β) : ofFn (f <| l.get ·) = l.map f := by - simp - -- not registered as a simp lemma, as otherwise it fires before `forall_mem_ofFn_iff` which -- is much more useful theorem mem_ofFn {n} (f : Fin n → α) (a : α) : a ∈ ofFn f ↔ a ∈ Set.range f := by diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index d598a0ba42d67..8e661b3806c9c 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -417,12 +417,6 @@ theorem length_permutations'Aux (s : List α) (x : α) : · simp · simpa using IH -@[deprecated "No deprecation message was provided." (since := "2024-06-12")] -theorem permutations'Aux_get_zero (s : List α) (x : α) - (hn : 0 < length (permutations'Aux x s) := (by simp)) : - (permutations'Aux x s).get ⟨0, hn⟩ = x :: s := - get_permutations'Aux _ _ _ _ - theorem injective_permutations'Aux (x : α) : Function.Injective (permutations'Aux x) := by intro s t h apply insertIdx_injective s.length x diff --git a/Mathlib/Data/List/ToFinsupp.lean b/Mathlib/Data/List/ToFinsupp.lean index 2e45c00adad82..3af2f3ea04c13 100644 --- a/Mathlib/Data/List/ToFinsupp.lean +++ b/Mathlib/Data/List/ToFinsupp.lean @@ -80,17 +80,6 @@ theorem toFinsupp_singleton (x : M) [DecidablePred (getD [x] · 0 ≠ 0)] : toFinsupp [x] = Finsupp.single 0 x := by ext ⟨_ | i⟩ <;> simp [Finsupp.single_apply, (Nat.zero_lt_succ _).ne] -@[deprecated "This lemma is unused, and can be proved by `simp`." (since := "2024-06-12")] -theorem toFinsupp_cons_apply_zero (x : M) (xs : List M) - [DecidablePred (getD (x::xs) · 0 ≠ 0)] : (x::xs).toFinsupp 0 = x := - rfl - -@[deprecated "This lemma is unused, and can be proved by `simp`." (since := "2024-06-12")] -theorem toFinsupp_cons_apply_succ (x : M) (xs : List M) (n : ℕ) - [DecidablePred (getD (x::xs) · 0 ≠ 0)] [DecidablePred (getD xs · 0 ≠ 0)] : - (x::xs).toFinsupp n.succ = xs.toFinsupp n := - rfl - theorem toFinsupp_append {R : Type*} [AddZeroClass R] (l₁ l₂ : List R) [DecidablePred (getD (l₁ ++ l₂) · 0 ≠ 0)] [DecidablePred (getD l₁ · 0 ≠ 0)] [DecidablePred (getD l₂ · 0 ≠ 0)] : diff --git a/Mathlib/Data/List/Zip.lean b/Mathlib/Data/List/Zip.lean index b7f0009d75359..e810b4c1315ab 100644 --- a/Mathlib/Data/List/Zip.lean +++ b/Mathlib/Data/List/Zip.lean @@ -117,39 +117,19 @@ theorem reverse_revzip (l : List α) : reverse l.revzip = revzip l.reverse := by theorem revzip_swap (l : List α) : (revzip l).map Prod.swap = revzip l.reverse := by simp [revzip] -@[deprecated (since := "2024-07-29")] alias getElem?_zip_with := getElem?_zipWith' - theorem get?_zipWith' (f : α → β → γ) (l₁ : List α) (l₂ : List β) (i : ℕ) : (zipWith f l₁ l₂).get? i = ((l₁.get? i).map f).bind fun g => (l₂.get? i).map g := by simp [getElem?_zipWith'] -@[deprecated (since := "2024-07-29")] alias get?_zip_with := get?_zipWith' -@[deprecated (since := "2024-07-29")] alias getElem?_zip_with_eq_some := getElem?_zipWith_eq_some - theorem get?_zipWith_eq_some (f : α → β → γ) (l₁ : List α) (l₂ : List β) (z : γ) (i : ℕ) : (zipWith f l₁ l₂).get? i = some z ↔ ∃ x y, l₁.get? i = some x ∧ l₂.get? i = some y ∧ f x y = z := by simp [getElem?_zipWith_eq_some] -@[deprecated (since := "2024-07-29")] alias get?_zip_with_eq_some := get?_zipWith_eq_some - theorem get?_zip_eq_some (l₁ : List α) (l₂ : List β) (z : α × β) (i : ℕ) : (zip l₁ l₂).get? i = some z ↔ l₁.get? i = some z.1 ∧ l₂.get? i = some z.2 := by simp [getElem?_zip_eq_some] -@[deprecated getElem_zipWith (since := "2024-06-12")] -theorem get_zipWith {f : α → β → γ} {l : List α} {l' : List β} {i : Fin (zipWith f l l').length} : - (zipWith f l l').get i = - f (l.get ⟨i, lt_length_left_of_zipWith i.isLt⟩) - (l'.get ⟨i, lt_length_right_of_zipWith i.isLt⟩) := by - simp - -@[deprecated getElem_zip (since := "2024-06-12")] -theorem get_zip {l : List α} {l' : List β} {i : Fin (zip l l').length} : - (zip l l').get i = - (l.get ⟨i, lt_length_left_of_zip i.isLt⟩, l'.get ⟨i, lt_length_right_of_zip i.isLt⟩) := by - simp - theorem mem_zip_inits_tails {l : List α} {init tail : List α} : (init, tail) ∈ zip l.inits l.tails ↔ init ++ tail = l := by induction' l with hd tl ih generalizing init tail <;> simp_rw [tails, inits, zip_cons_cons] diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index bb263451f18a0..d3c3bf82924e0 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1529,8 +1529,6 @@ theorem filter_add_not (s : Multiset α) : filter p s + filter (fun a => ¬p a) theorem filter_map (f : β → α) (s : Multiset β) : filter p (map f s) = map f (filter (p ∘ f) s) := Quot.inductionOn s fun l => by simp [List.filter_map]; rfl -@[deprecated (since := "2024-06-16")] alias map_filter := filter_map - -- TODO: rename to `map_filter` when the deprecated alias above is removed. lemma map_filter' {f : α → β} (hf : Injective f) (s : Multiset α) [DecidablePred fun b => ∃ a, p a ∧ f a = b] : diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index 053a63b88459a..b7bd33f4fb44f 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -194,8 +194,6 @@ def coeHom : ℚ≥0 →+* ℚ where theorem mk_natCast (n : ℕ) : @Eq ℚ≥0 (⟨(n : ℚ), Nat.cast_nonneg' n⟩ : ℚ≥0) n := rfl -@[deprecated (since := "2024-04-05")] alias mk_coe_nat := mk_natCast - @[simp] theorem coe_coeHom : ⇑coeHom = ((↑) : ℚ≥0 → ℚ) := rfl diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index be4890f48ea14..b443b6bbf71cb 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -267,9 +267,6 @@ variable {ι : Type*} {f : ι → ℝ} protected theorem coe_natCast (n : ℕ) : (↑(↑n : ℝ≥0) : ℝ) = n := map_natCast toRealHom n -@[deprecated (since := "2024-04-17")] -alias coe_nat_cast := NNReal.coe_natCast - @[simp, norm_cast] protected theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : ((ofNat(n) : ℝ≥0) : ℝ) = ofNat(n) := rfl @@ -315,8 +312,6 @@ theorem _root_.Real.toNNReal_coe {r : ℝ≥0} : Real.toNNReal r = r := theorem mk_natCast (n : ℕ) : @Eq ℝ≥0 (⟨(n : ℝ), n.cast_nonneg⟩ : ℝ≥0) n := NNReal.eq (NNReal.coe_natCast n).symm -@[deprecated (since := "2024-04-05")] alias mk_coe_nat := mk_natCast - -- Porting note: place this in the `Real` namespace @[simp] theorem toNNReal_coe_nat (n : ℕ) : Real.toNNReal n = n := @@ -524,9 +519,6 @@ lemma toNNReal_eq_one {r : ℝ} : r.toNNReal = 1 ↔ r = 1 := toNNReal_eq_iff_eq lemma toNNReal_eq_natCast {r : ℝ} {n : ℕ} (hn : n ≠ 0) : r.toNNReal = n ↔ r = n := mod_cast toNNReal_eq_iff_eq_coe <| Nat.cast_ne_zero.2 hn -@[deprecated (since := "2024-04-17")] -alias toNNReal_eq_nat_cast := toNNReal_eq_natCast - @[simp] lemma toNNReal_eq_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : r.toNNReal = ofNat(n) ↔ r = OfNat.ofNat n := @@ -548,16 +540,10 @@ lemma one_lt_toNNReal {r : ℝ} : 1 < r.toNNReal ↔ 1 < r := by lemma toNNReal_le_natCast {r : ℝ} {n : ℕ} : r.toNNReal ≤ n ↔ r ≤ n := by simpa using toNNReal_le_toNNReal_iff n.cast_nonneg -@[deprecated (since := "2024-04-17")] -alias toNNReal_le_nat_cast := toNNReal_le_natCast - @[simp] lemma natCast_lt_toNNReal {r : ℝ} {n : ℕ} : n < r.toNNReal ↔ n < r := by simpa only [not_le] using toNNReal_le_natCast.not -@[deprecated (since := "2024-04-17")] -alias nat_cast_lt_toNNReal := natCast_lt_toNNReal - @[simp] lemma toNNReal_le_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : r.toNNReal ≤ ofNat(n) ↔ r ≤ n := @@ -604,26 +590,14 @@ lemma toNNReal_lt_one {r : ℝ} : r.toNNReal < 1 ↔ r < 1 := by simp only [← lemma natCastle_toNNReal' {n : ℕ} {r : ℝ} : ↑n ≤ r.toNNReal ↔ n ≤ r ∨ n = 0 := by simpa [n.cast_nonneg.le_iff_eq] using toNNReal_le_toNNReal_iff' (r := n) -@[deprecated (since := "2024-04-17")] -alias nat_cast_le_toNNReal' := natCastle_toNNReal' - @[simp] lemma toNNReal_lt_natCast' {n : ℕ} {r : ℝ} : r.toNNReal < n ↔ r < n ∧ n ≠ 0 := by simpa [pos_iff_ne_zero] using toNNReal_lt_toNNReal_iff' (r := r) (p := n) -@[deprecated (since := "2024-04-17")] -alias toNNReal_lt_nat_cast' := toNNReal_lt_natCast' - lemma natCast_le_toNNReal {n : ℕ} {r : ℝ} (hn : n ≠ 0) : ↑n ≤ r.toNNReal ↔ n ≤ r := by simp [hn] -@[deprecated (since := "2024-04-17")] -alias nat_cast_le_toNNReal := natCast_le_toNNReal - lemma toNNReal_lt_natCast {r : ℝ} {n : ℕ} (hn : n ≠ 0) : r.toNNReal < n ↔ r < n := by simp [hn] -@[deprecated (since := "2024-04-17")] -alias toNNReal_lt_nat_cast := toNNReal_lt_natCast - @[simp] lemma toNNReal_lt_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] : r.toNNReal < ofNat(n) ↔ r < OfNat.ofNat n := diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index 617f4537a4023..84a0af847aad0 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -194,9 +194,6 @@ theorem natCast_apply (n : ℕ) (a : α) : (n : ∀ a, π a) a = n := theorem natCast_def (n : ℕ) : (n : ∀ a, π a) = fun _ ↦ ↑n := rfl -@[deprecated (since := "2024-04-05")] alias nat_apply := natCast_apply -@[deprecated (since := "2024-04-05")] alias coe_nat := natCast_def - @[simp] theorem ofNat_apply (n : ℕ) [n.AtLeastTwo] (a : α) : (OfNat.ofNat n : ∀ a, π a) a = n := rfl diff --git a/Mathlib/Data/Nat/Cast/Commute.lean b/Mathlib/Data/Nat/Cast/Commute.lean index a2d80c1d27b64..a7bba36ad0578 100644 --- a/Mathlib/Data/Nat/Cast/Commute.lean +++ b/Mathlib/Data/Nat/Cast/Commute.lean @@ -78,12 +78,4 @@ lemma natCast_mul_self : Commute (n * a) a := (Commute.refl a).natCast_mul_left lemma self_natCast_mul_natCast_mul : Commute (m * a) (n * a) := (Commute.refl a).natCast_mul_natCast_mul m n -@[deprecated (since := "2024-05-27")] alias cast_nat_mul_right := natCast_mul_right -@[deprecated (since := "2024-05-27")] alias cast_nat_mul_left := natCast_mul_left -@[deprecated (since := "2024-05-27")] alias cast_nat_mul_cast_nat_mul := natCast_mul_natCast_mul -@[deprecated (since := "2024-05-27")] alias self_cast_nat_mul := self_natCast_mul -@[deprecated (since := "2024-05-27")] alias cast_nat_mul_self := natCast_mul_self -@[deprecated (since := "2024-05-27")] -alias self_cast_nat_mul_cast_nat_mul := self_natCast_mul_natCast_mul - end Commute diff --git a/Mathlib/Data/Nat/Choose/Multinomial.lean b/Mathlib/Data/Nat/Choose/Multinomial.lean index 98037ce2808a7..1500c67f6a583 100644 --- a/Mathlib/Data/Nat/Choose/Multinomial.lean +++ b/Mathlib/Data/Nat/Choose/Multinomial.lean @@ -51,8 +51,6 @@ theorem multinomial_spec : (∏ i ∈ s, (f i)!) * multinomial s f = (∑ i ∈ @[simp] lemma multinomial_empty : multinomial ∅ f = 1 := by simp [multinomial] -@[deprecated (since := "2024-06-01")] alias multinomial_nil := multinomial_empty - variable {s f} lemma multinomial_cons (ha : a ∉ s) (f : α → ℕ) : diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 2bf65bae0c070..1ba1022898043 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -222,12 +222,6 @@ attribute [simp] le_add_left le_add_right Nat.lt_add_left_iff_pos Nat.lt_add_rig -- We want to use these two lemmas earlier than the lemmas simp can prove them with @[simp, nolint simpNF] protected alias add_left_inj := Nat.add_right_cancel_iff @[simp, nolint simpNF] protected alias add_right_inj := Nat.add_left_cancel_iff - --- Sometimes a bare `Nat.add` or similar appears as a consequence of unfolding during pattern --- matching. These lemmas package them back up as typeclass mediated operations. -@[deprecated (since := "2024-04-05")] alias add_def := add_eq - --- We want to use these two lemmas earlier than the lemmas simp can prove them with @[simp, nolint simpNF] protected lemma add_eq_left : a + b = a ↔ b = 0 := by omega @[simp, nolint simpNF] protected lemma add_eq_right : a + b = b ↔ a = 0 := by omega @@ -478,8 +472,6 @@ lemma div_mul_div_comm : b ∣ a → d ∣ c → (a / b) * (c / d) = (a * c) / ( protected lemma mul_div_mul_comm (hba : b ∣ a) (hdc : d ∣ c) : a * c / (b * d) = a / b * (c / d) := (div_mul_div_comm hba hdc).symm -@[deprecated (since := "2024-05-29")] alias mul_div_mul_comm_of_dvd_dvd := Nat.mul_div_mul_comm - lemma eq_zero_of_le_div (hn : 2 ≤ n) (h : m ≤ m / n) : m = 0 := eq_zero_of_mul_le hn <| by rw [Nat.mul_comm]; exact (Nat.le_div_iff_mul_le (Nat.lt_of_lt_of_le (by decide) hn)).1 h @@ -745,14 +737,6 @@ lemma leRec_succ_left {motive : (m : ℕ) → n ≤ m → Sort*} leRec (motive := motive) refl le_succ_of_le h1 := by rw [leRec_trans _ _ (le_succ n) h2, leRec_succ'] -/-- Recursion starting at a non-zero number: given a map `C k → C (k+1)` for each `k ≥ n`, -there is a map from `C n` to each `C m`, `n ≤ m`. - -Prefer `Nat.leRec`, which can be used as `induction h using Nat.leRec`. -/ -@[elab_as_elim, deprecated Nat.leRec (since := "2024-07-05")] -def leRecOn' {C : ℕ → Sort*} : ∀ {m}, n ≤ m → (∀ ⦃k⦄, n ≤ k → C k → C (k + 1)) → C n → C m := - fun h of_succ self => Nat.leRec self of_succ h - /-- Recursion starting at a non-zero number: given a map `C k → C (k + 1)` for each `k`, there is a map from `C n` to each `C m`, `n ≤ m`. For a version where the assumption is only made when `k ≥ n`, see `Nat.leRec`. -/ @@ -970,10 +954,6 @@ attribute [simp] Nat.dvd_zero lemma mod_two_ne_one : n % 2 ≠ 1 ↔ n % 2 = 0 := mod_two_not_eq_one lemma mod_two_ne_zero : n % 2 ≠ 0 ↔ n % 2 = 1 := mod_two_not_eq_zero -@[deprecated mod_mul_right_div_self (since := "2024-05-29")] -lemma div_mod_eq_mod_mul_div (a b c : ℕ) : a / b % c = a % (b * c) / b := - (mod_mul_right_div_self a b c).symm - /-- Variant of `Nat.lt_div_iff_mul_lt` that assumes `d ∣ n`. -/ protected lemma lt_div_iff_mul_lt' (hdn : d ∣ n) (a : ℕ) : a < n / d ↔ d * a < n := by obtain rfl | hd := d.eq_zero_or_pos @@ -1227,8 +1207,6 @@ lemma dvd_div_of_mul_dvd (h : a * b ∣ c) : b ∣ c / a := @[simp] lemma dvd_div_iff_mul_dvd (hbc : c ∣ b) : a ∣ b / c ↔ c * a ∣ b := ⟨fun h => mul_dvd_of_dvd_div hbc h, fun h => dvd_div_of_mul_dvd h⟩ -@[deprecated (since := "2024-06-18")] alias dvd_div_iff := dvd_div_iff_mul_dvd - lemma dvd_mul_of_div_dvd (h : b ∣ a) (hdiv : a / b ∣ c) : a ∣ b * c := by obtain ⟨e, rfl⟩ := hdiv rw [← Nat.mul_assoc, Nat.mul_comm _ (a / b), Nat.div_mul_cancel h] diff --git a/Mathlib/Data/Nat/Factorization/Defs.lean b/Mathlib/Data/Nat/Factorization/Defs.lean index 19201465a196a..2aeb32e45164e 100644 --- a/Mathlib/Data/Nat/Factorization/Defs.lean +++ b/Mathlib/Data/Nat/Factorization/Defs.lean @@ -77,10 +77,6 @@ theorem factorization_eq_primeFactorsList_multiset (n : ℕ) : ext p simp -@[deprecated (since := "2024-07-16")] alias factors_count_eq := primeFactorsList_count_eq -@[deprecated (since := "2024-07-16")] -alias factorization_eq_factors_multiset := factorization_eq_primeFactorsList_multiset - theorem Prime.factorization_pos_of_dvd {n p : ℕ} (hp : p.Prime) (hn : n ≠ 0) (h : p ∣ n) : 0 < n.factorization p := by rwa [← primeFactorsList_count_eq, count_pos_iff, mem_primeFactorsList_iff_dvd hn hp] diff --git a/Mathlib/Data/Nat/Factors.lean b/Mathlib/Data/Nat/Factors.lean index 67bf2864fa197..fc045dc74c821 100644 --- a/Mathlib/Data/Nat/Factors.lean +++ b/Mathlib/Data/Nat/Factors.lean @@ -39,8 +39,6 @@ def primeFactorsList : ℕ → List ℕ m :: primeFactorsList ((k + 2) / m) decreasing_by exact factors_lemma -@[deprecated (since := "2024-06-14")] alias factors := primeFactorsList - @[simp] theorem primeFactorsList_zero : primeFactorsList 0 = [] := by rw [primeFactorsList] diff --git a/Mathlib/Data/Nat/PrimeFin.lean b/Mathlib/Data/Nat/PrimeFin.lean index a9054c2d80174..f2b24d6d23b4c 100644 --- a/Mathlib/Data/Nat/PrimeFin.lean +++ b/Mathlib/Data/Nat/PrimeFin.lean @@ -45,9 +45,6 @@ lemma primeFactors_mono (hmn : m ∣ n) (hn : n ≠ 0) : primeFactors m ⊆ prim lemma mem_primeFactors_iff_mem_primeFactorsList : p ∈ n.primeFactors ↔ p ∈ n.primeFactorsList := by simp only [primeFactors, List.mem_toFinset] -@[deprecated (since := "2024-07-16")] -alias mem_primeFactors_iff_mem_factors := mem_primeFactors_iff_mem_primeFactorsList - lemma prime_of_mem_primeFactors (hp : p ∈ n.primeFactors) : p.Prime := (mem_primeFactors.1 hp).1 lemma dvd_of_mem_primeFactors (hp : p ∈ n.primeFactors) : p ∣ n := (mem_primeFactors.1 hp).2.1 diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index c81e7815857a8..892e4531a63c4 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -29,17 +29,11 @@ theorem squarefree_iff_nodup_primeFactorsList {n : ℕ} (h0 : n ≠ 0) : rw [UniqueFactorizationMonoid.squarefree_iff_nodup_normalizedFactors h0, Nat.factors_eq] simp -@[deprecated (since := "2024-07-17")] -alias squarefree_iff_nodup_factors := squarefree_iff_nodup_primeFactorsList - end Nat theorem Squarefree.nodup_primeFactorsList {n : ℕ} (hn : Squarefree n) : n.primeFactorsList.Nodup := (Nat.squarefree_iff_nodup_primeFactorsList hn.ne_zero).mp hn -@[deprecated (since := "2024-07-17")] -alias Squarefree.nodup_factors := Squarefree.nodup_primeFactorsList - namespace Nat variable {s : Finset ℕ} {m n p : ℕ} diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index 777ee93d18575..99f36fb91d9d2 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -418,9 +418,6 @@ theorem to_of_nat : ∀ n : ℕ, ((n : Num) : ℕ) = n theorem of_natCast {α} [AddMonoidWithOne α] (n : ℕ) : ((n : Num) : α) = n := by rw [← cast_to_nat, to_of_nat] -@[deprecated (since := "2024-04-17")] -alias of_nat_cast := of_natCast - @[norm_cast] theorem of_nat_inj {m n : ℕ} : (m : Num) = n ↔ m = n := ⟨fun h => Function.LeftInverse.injective to_of_nat h, congr_arg _⟩ @@ -1355,16 +1352,10 @@ theorem of_nat_toZNumNeg (n : ℕ) : Num.toZNumNeg n = -n := by rw [← of_nat_t theorem of_intCast [AddGroupWithOne α] (n : ℤ) : ((n : ZNum) : α) = n := by rw [← cast_to_int, to_of_int] -@[deprecated (since := "2024-04-17")] -alias of_int_cast := of_intCast - @[simp, norm_cast] theorem of_natCast [AddGroupWithOne α] (n : ℕ) : ((n : ZNum) : α) = n := by rw [← Int.cast_natCast, of_intCast, Int.cast_natCast] -@[deprecated (since := "2024-04-17")] -alias of_nat_cast := of_natCast - @[simp, norm_cast] theorem dvd_to_int (m n : ZNum) : (m : ℤ) ∣ n ↔ m ∣ n := ⟨fun ⟨k, e⟩ => ⟨k, by rw [← of_to_int n, e]; simp⟩, fun ⟨k, e⟩ => ⟨k, by simp [e]⟩⟩ diff --git a/Mathlib/Data/Part.lean b/Mathlib/Data/Part.lean index 5a508e4709c66..328d4bcd6571a 100644 --- a/Mathlib/Data/Part.lean +++ b/Mathlib/Data/Part.lean @@ -67,8 +67,6 @@ def toOption (o : Part α) [Decidable o.Dom] : Option α := @[simp] lemma toOption_eq_none (o : Part α) [Decidable o.Dom] : o.toOption = none ↔ ¬o.Dom := by by_cases h : o.Dom <;> simp [h, toOption] -@[deprecated (since := "2024-06-20")] alias toOption_isNone := toOption_eq_none - /-- `Part` extensionality -/ theorem ext' : ∀ {o p : Part α}, (o.Dom ↔ p.Dom) → (∀ h₁ h₂, o.get h₁ = p.get h₂) → o = p | ⟨od, o⟩, ⟨pd, p⟩, H1, H2 => by diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index d68eab7f60f81..fdb50a6c831d6 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -17,8 +17,6 @@ It also defines better delaborators for product projections. variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} -@[deprecated (since := "2024-05-08")] alias Prod_map := Prod.map_apply - namespace Prod lemma swap_eq_iff_eq_swap {x : α × β} {y : β × α} : x.swap = y ↔ x = y.swap := by aesop @@ -88,8 +86,6 @@ theorem id_prod : (fun p : α × β ↦ (p.1, p.2)) = id := theorem map_iterate (f : α → α) (g : β → β) (n : ℕ) : (Prod.map f g)^[n] = Prod.map f^[n] g^[n] := by induction n <;> simp [*, Prod.map_comp_map] -@[deprecated (since := "2024-07-03")] alias iterate_prod_map := Prod.map_iterate - theorem fst_surjective [h : Nonempty β] : Function.Surjective (@fst α β) := fun x ↦ h.elim fun y ↦ ⟨⟨x, y⟩, rfl⟩ @@ -231,13 +227,6 @@ theorem Involutive.prodMap {f : α → α} {g : β → β} : Involutive f → Involutive g → Involutive (map f g) := LeftInverse.prodMap -@[deprecated (since := "2024-05-08")] alias Injective.Prod_map := Injective.prodMap -@[deprecated (since := "2024-05-08")] alias Surjective.Prod_map := Surjective.prodMap -@[deprecated (since := "2024-05-08")] alias Bijective.Prod_map := Bijective.prodMap -@[deprecated (since := "2024-05-08")] alias LeftInverse.Prod_map := LeftInverse.prodMap -@[deprecated (since := "2024-05-08")] alias RightInverse.Prod_map := RightInverse.prodMap -@[deprecated (since := "2024-05-08")] alias Involutive.Prod_map := Involutive.prodMap - end Function namespace Prod diff --git a/Mathlib/Data/Rat/Cast/CharZero.lean b/Mathlib/Data/Rat/Cast/CharZero.lean index fac876b3aa557..c6c2cd6b73207 100644 --- a/Mathlib/Data/Rat/Cast/CharZero.lean +++ b/Mathlib/Data/Rat/Cast/CharZero.lean @@ -54,8 +54,6 @@ def castHom : ℚ →+* α where @[simp] lemma coe_castHom : ⇑(castHom α) = ((↑) : ℚ → α) := rfl -@[deprecated (since := "2024-07-22")] alias coe_cast_hom := coe_castHom - @[simp, norm_cast] lemma cast_inv (p : ℚ) : ↑(p⁻¹) = (p⁻¹ : α) := map_inv₀ (castHom α) _ @[simp, norm_cast] lemma cast_div (p q : ℚ) : ↑(p / q) = (p / q : α) := map_div₀ (castHom α) .. diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index 59d83bde6c20c..da68f98d28c08 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -60,9 +60,6 @@ theorem ofInt_eq_cast (n : ℤ) : ofInt n = Int.cast n := @[simp, norm_cast] lemma den_intCast (n : ℤ) : (n : ℚ).den = 1 := rfl -@[deprecated (since := "2024-04-29")] alias coe_int_num := num_intCast -@[deprecated (since := "2024-04-29")] alias coe_int_den := den_intCast - lemma intCast_injective : Injective (Int.cast : ℤ → ℚ) := fun _ _ ↦ congr_arg num lemma natCast_injective : Injective (Nat.cast : ℕ → ℚ) := intCast_injective.comp fun _ _ ↦ Int.natCast_inj.1 @@ -194,8 +191,6 @@ lemma mul_eq_mkRat (q r : ℚ) : q * r = mkRat (q.num * r.num) (q.den * r.den) : -- TODO: Rename `divInt_eq_iff` in Batteries to `divInt_eq_divInt` alias divInt_eq_divInt := divInt_eq_iff -@[deprecated (since := "2024-04-29")] alias mul_num_den := mul_eq_mkRat - instance instPowNat : Pow ℚ ℕ where pow q n := ⟨q.num ^ n, q.den ^ n, by simp [Nat.pow_eq_zero], by rw [Int.natAbs_pow]; exact q.reduced.pow _ _⟩ @@ -235,8 +230,6 @@ lemma inv_def' (q : ℚ) : q⁻¹ = q.den /. q.num := by rw [← inv_divInt', nu lemma div_def' (q r : ℚ) : q / r = (q.num * r.den) /. (q.den * r.num) := by rw [← divInt_div_divInt, num_divInt_den, num_divInt_den] -@[deprecated (since := "2024-04-15")] alias div_num_den := div_def' - variable (a b c : ℚ) protected lemma add_zero : a + 0 = a := by simp [add_def, normalize_eq_mkRat] @@ -438,9 +431,6 @@ theorem den_eq_one_iff (r : ℚ) : r.den = 1 ↔ ↑r.num = r := instance canLift : CanLift ℚ ℤ (↑) fun q => q.den = 1 := ⟨fun q hq => ⟨q.num, coe_int_num_of_den_eq_one hq⟩⟩ -@[deprecated (since := "2024-04-05")] alias coe_int_eq_divInt := intCast_eq_divInt -@[deprecated (since := "2024-04-05")] alias coe_int_div_eq_divInt := intCast_div_eq_divInt - -- Will be subsumed by `Int.coe_inj` after we have defined -- `LinearOrderedField ℚ` (which implies characteristic zero). theorem coe_int_inj (m n : ℤ) : (m : ℚ) = n ↔ m = n := diff --git a/Mathlib/Data/Rat/Floor.lean b/Mathlib/Data/Rat/Floor.lean index 42464ab2589e4..7d92662ba0466 100644 --- a/Mathlib/Data/Rat/Floor.lean +++ b/Mathlib/Data/Rat/Floor.lean @@ -87,8 +87,6 @@ theorem natFloor_natCast_div_natCast (n d : ℕ) : ⌊(↑n / ↑d : ℚ)⌋₊ push_cast exact floor_intCast_div_natCast n d -@[deprecated (since := "2024-07-23")] alias floor_int_div_nat_eq_div := floor_intCast_div_natCast - @[simp, norm_cast] theorem floor_cast (x : ℚ) : ⌊(x : α)⌋ = ⌊x⌋ := floor_eq_iff.2 (mod_cast floor_eq_iff.1 (Eq.refl ⌊x⌋)) diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index e0ac6e8bf3883..1c317a9896de2 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -229,8 +229,6 @@ theorem den_div_intCast_eq_one_iff (m n : ℤ) (hn : n ≠ 0) : ((m : ℚ) / n). theorem den_div_natCast_eq_one_iff (m n : ℕ) (hn : n ≠ 0) : ((m : ℚ) / n).den = 1 ↔ n ∣ m := (den_div_intCast_eq_one_iff m n (Int.ofNat_ne_zero.mpr hn)).trans Int.ofNat_dvd -@[deprecated (since := "2024-05-11")] alias den_div_cast_eq_one_iff := den_div_intCast_eq_one_iff - theorem inv_intCast_num_of_pos {a : ℤ} (ha0 : 0 < a) : (a : ℚ)⁻¹.num = 1 := by rw [← ofInt_eq_cast, ofInt, mk_eq_divInt, Rat.inv_divInt', divInt_eq_div, Nat.cast_one] apply num_div_eq_of_coprime ha0 @@ -285,19 +283,6 @@ theorem inv_intCast_den (a : ℤ) : (a : ℚ)⁻¹.den = if a = 0 then 1 else a. theorem inv_natCast_den (a : ℕ) : (a : ℚ)⁻¹.den = if a = 0 then 1 else a := by simpa [-inv_intCast_den, ofInt_eq_cast] using inv_intCast_den a -@[deprecated (since := "2024-04-05")] alias coe_int_div_self := intCast_div_self -@[deprecated (since := "2024-04-05")] alias coe_nat_div_self := natCast_div_self -@[deprecated (since := "2024-04-05")] alias coe_int_div := intCast_div -@[deprecated (since := "2024-04-05")] alias coe_nat_div := natCast_div -@[deprecated (since := "2024-04-05")] alias inv_coe_int_num_of_pos := inv_intCast_num_of_pos -@[deprecated (since := "2024-04-05")] alias inv_coe_nat_num_of_pos := inv_natCast_num_of_pos -@[deprecated (since := "2024-04-05")] alias inv_coe_int_den_of_pos := inv_intCast_den_of_pos -@[deprecated (since := "2024-04-05")] alias inv_coe_nat_den_of_pos := inv_natCast_den_of_pos -@[deprecated (since := "2024-04-05")] alias inv_coe_int_num := inv_intCast_num -@[deprecated (since := "2024-04-05")] alias inv_coe_nat_num := inv_natCast_num -@[deprecated (since := "2024-04-05")] alias inv_coe_int_den := inv_intCast_den -@[deprecated (since := "2024-04-05")] alias inv_coe_nat_den := inv_natCast_den - @[simp] theorem inv_ofNat_den (a : ℕ) [a.AtLeastTwo] : (ofNat(a) : ℚ)⁻¹.den = OfNat.ofNat a := diff --git a/Mathlib/Data/Real/Irrational.lean b/Mathlib/Data/Real/Irrational.lean index 70f269455316c..53b89aed5b1d2 100644 --- a/Mathlib/Data/Real/Irrational.lean +++ b/Mathlib/Data/Real/Irrational.lean @@ -133,12 +133,6 @@ theorem Nat.Prime.irrational_sqrt {p : ℕ} (hp : Nat.Prime p) : Irrational (√ theorem irrational_sqrt_two : Irrational (√2) := by simpa using Nat.prime_two.irrational_sqrt -@[deprecated irrational_sqrt_ratCast_iff (since := "2024-06-16")] -theorem irrational_sqrt_rat_iff (q : ℚ) : - Irrational (√q) ↔ Rat.sqrt q * Rat.sqrt q ≠ q ∧ 0 ≤ q := by - rw [irrational_sqrt_ratCast_iff, ne_eq, ← Rat.exists_mul_self] - simp only [eq_comm, IsSquare] - /-- This can be used as ```lean diff --git a/Mathlib/Data/Real/Sign.lean b/Mathlib/Data/Real/Sign.lean index d4b23cdfb07e2..6937f915c1d8a 100644 --- a/Mathlib/Data/Real/Sign.lean +++ b/Mathlib/Data/Real/Sign.lean @@ -69,9 +69,6 @@ theorem sign_intCast (z : ℤ) : sign (z : ℝ) = ↑(Int.sign z) := by · rw [Int.cast_zero, sign_zero, Int.sign_zero, Int.cast_zero] · rw [sign_of_pos (Int.cast_pos.mpr hp), Int.sign_eq_one_of_pos hp, Int.cast_one] -@[deprecated (since := "2024-04-17")] -alias sign_int_cast := sign_intCast - theorem sign_neg {r : ℝ} : sign (-r) = -sign r := by obtain hn | rfl | hp := lt_trichotomy r (0 : ℝ) · rw [sign_of_neg hn, sign_of_pos (neg_pos.mpr hn), neg_neg] diff --git a/Mathlib/Data/Seq/Seq.lean b/Mathlib/Data/Seq/Seq.lean index 172c4fc696654..08d384759f0b2 100644 --- a/Mathlib/Data/Seq/Seq.lean +++ b/Mathlib/Data/Seq/Seq.lean @@ -418,13 +418,9 @@ def ofMLList : MLList Id α → Seq α := | .none => none | .some (a, l') => some (a, l') -@[deprecated (since := "2024-07-26")] alias ofLazyList := ofMLList - instance coeMLList : Coe (MLList Id α) (Seq α) := ⟨ofMLList⟩ -@[deprecated (since := "2024-07-26")] alias coeLazyList := coeMLList - /-- Translate a sequence into a `MLList`. -/ unsafe def toMLList : Seq α → MLList Id α | s => @@ -432,8 +428,6 @@ unsafe def toMLList : Seq α → MLList Id α | none => .nil | some (a, s') => .cons a (toMLList s') -@[deprecated (since := "2024-07-26")] alias toLazyList := toMLList - end MLList /-- Translate a sequence to a list. This function will run forever if diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index 3355f09f3689d..fbc51e1fc27e5 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -906,22 +906,6 @@ lemma exists_subsuperset_card_eq {n : ℕ} (hst : s ⊆ t) (hsn : s.ncard ≤ n) lemma exists_subset_card_eq {n : ℕ} (hns : n ≤ s.ncard) : ∃ t ⊆ s, t.ncard = n := by simpa using exists_subsuperset_card_eq s.empty_subset (by simp) hns -/-- Given a set `t` and a set `s` inside it, we can shrink `t` to any appropriate size, and keep `s` - inside it. -/ -@[deprecated exists_subsuperset_card_eq (since := "2024-06-24")] -theorem exists_intermediate_Set (i : ℕ) (h₁ : i + s.ncard ≤ t.ncard) (h₂ : s ⊆ t) : - ∃ r : Set α, s ⊆ r ∧ r ⊆ t ∧ r.ncard = i + s.ncard := - exists_subsuperset_card_eq h₂ (Nat.le_add_left ..) h₁ - -@[deprecated exists_subsuperset_card_eq (since := "2024-06-24")] -theorem exists_intermediate_set' {m : ℕ} (hs : s.ncard ≤ m) (ht : m ≤ t.ncard) (h : s ⊆ t) : - ∃ r : Set α, s ⊆ r ∧ r ⊆ t ∧ r.ncard = m := exists_subsuperset_card_eq h hs ht - -/-- We can shrink `s` to any smaller size. -/ -@[deprecated exists_subset_card_eq (since := "2024-06-23")] -theorem exists_smaller_set (s : Set α) (i : ℕ) (h₁ : i ≤ s.ncard) : - ∃ t : Set α, t ⊆ s ∧ t.ncard = i := exists_subset_card_eq h₁ - theorem Infinite.exists_subset_ncard_eq {s : Set α} (hs : s.Infinite) (k : ℕ) : ∃ t, t ⊆ s ∧ t.Finite ∧ t.ncard = k := by have := hs.to_subtype diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index ac85ff8a86b20..3a93bcdb1c839 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -686,10 +686,6 @@ theorem exists_subset_range_and_iff {f : α → β} {p : Set β → Prop} : (∃ s, s ⊆ range f ∧ p s) ↔ ∃ s, p (f '' s) := by rw [← exists_range_iff, range_image]; rfl -@[deprecated exists_subset_range_and_iff (since := "2024-06-06")] -theorem exists_subset_range_iff {f : α → β} {p : Set β → Prop} : - (∃ (s : _) (_ : s ⊆ range f), p s) ↔ ∃ s, p (f '' s) := by simp - @[simp] theorem forall_subset_range_iff {f : α → β} {p : Set β → Prop} : (∀ s, s ⊆ range f → p s) ↔ ∀ s, p (f '' s) := by diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index 232017be6260c..e92c7a3eb1bc6 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -1807,9 +1807,6 @@ theorem directedOn_iUnion {r} {f : ι → Set α} (hd : Directed (· ⊆ ·) f) let ⟨x, xf, xa₁, xa₂⟩ := h z a₁ (zb₁ fb₁) a₂ (zb₂ fb₂) ⟨x, ⟨z, xf⟩, xa₁, xa₂⟩ -@[deprecated (since := "2024-05-05")] -alias directed_on_iUnion := directedOn_iUnion - theorem directedOn_sUnion {r} {S : Set (Set α)} (hd : DirectedOn (· ⊆ ·) S) (h : ∀ x ∈ S, DirectedOn r x) : DirectedOn r (⋃₀ S) := by rw [sUnion_eq_iUnion] diff --git a/Mathlib/Data/Set/List.lean b/Mathlib/Data/Set/List.lean index 0d3cf309c7743..08512b126e266 100644 --- a/Mathlib/Data/Set/List.lean +++ b/Mathlib/Data/Set/List.lean @@ -36,7 +36,6 @@ theorem range_list_map_coe (s : Set α) : range (map ((↑) : s → α)) = { l | theorem range_list_get : range l.get = { x | x ∈ l } := by ext x rw [mem_setOf_eq, mem_iff_get, mem_range] -@[deprecated (since := "2024-04-22")] alias range_list_nthLe := range_list_get theorem range_list_get? : range l.get? = insert none (some '' { x | x ∈ l }) := by rw [← range_list_get, ← range_comp] diff --git a/Mathlib/Data/Sigma/Basic.lean b/Mathlib/Data/Sigma/Basic.lean index e9e037d20b88d..85bb264783fc7 100644 --- a/Mathlib/Data/Sigma/Basic.lean +++ b/Mathlib/Data/Sigma/Basic.lean @@ -213,8 +213,6 @@ def elim {γ} (f : ∀ a, β a → γ) (a : PSigma β) : γ := theorem elim_val {γ} (f : ∀ a, β a → γ) (a b) : PSigma.elim f ⟨a, b⟩ = f a b := rfl -@[deprecated (since := "2024-07-27")] alias ex_of_psig := ex_of_PSigma - instance [Inhabited α] [Inhabited (β default)] : Inhabited (PSigma β) := ⟨⟨default, default⟩⟩ @@ -238,11 +236,6 @@ theorem mk.inj_iff {a₁ a₂ : α} {b₁ : β a₁} {b₂ : β a₂} : match a₁, a₂, b₁, b₂, h₁, h₂ with | _, _, _, _, Eq.refl _, HEq.refl _ => rfl -@[deprecated PSigma.ext_iff (since := "2024-07-27")] -protected theorem eq {α : Sort*} {β : α → Sort*} : ∀ {p₁ p₂ : Σ' a, β a} (h₁ : p₁.1 = p₂.1), - (Eq.recOn h₁ p₁.2 : β p₂.1) = p₂.2 → p₁ = p₂ - | ⟨_, _⟩, _, rfl, rfl => rfl - -- This should not be a simp lemma, since its discrimination tree key would just be `→`. theorem «forall» {p : (Σ'a, β a) → Prop} : (∀ x, p x) ↔ ∀ a b, p ⟨a, b⟩ := ⟨fun h a b ↦ h ⟨a, b⟩, fun h ⟨a, b⟩ ↦ h a b⟩ diff --git a/Mathlib/Data/String/Basic.lean b/Mathlib/Data/String/Basic.lean index 7bd4035fff1aa..b0612b0034c39 100644 --- a/Mathlib/Data/String/Basic.lean +++ b/Mathlib/Data/String/Basic.lean @@ -116,8 +116,6 @@ theorem toList_inj {s₁ s₂ : String} : s₁.toList = s₂.toList ↔ s₁ = s theorem asString_nil : [].asString = "" := rfl -@[deprecated (since := "2024-06-04")] alias nil_asString_eq_empty := asString_nil - @[simp] theorem toList_empty : "".toList = [] := rfl @@ -125,8 +123,6 @@ theorem toList_empty : "".toList = [] := theorem asString_toList (s : String) : s.toList.asString = s := rfl -@[deprecated (since := "2024-06-04")] alias asString_inv_toList := asString_toList - theorem toList_nonempty : ∀ {s : String}, s ≠ "" → s.toList = s.head :: (s.drop 1).toList | ⟨s⟩, h => by cases s with @@ -168,8 +164,6 @@ namespace List theorem toList_asString (l : List Char) : l.asString.toList = l := rfl -@[deprecated (since := "2024-06-04")] alias toList_inv_asString := toList_asString - @[simp] theorem length_asString (l : List Char) : l.asString.length = l.length := rfl diff --git a/Mathlib/Data/Subtype.lean b/Mathlib/Data/Subtype.lean index ebde31a801bf5..5d5e44eef3eb1 100644 --- a/Mathlib/Data/Subtype.lean +++ b/Mathlib/Data/Subtype.lean @@ -99,8 +99,6 @@ theorem val_inj {a b : Subtype p} : a.val = b.val ↔ a = b := lemma coe_ne_coe {a b : Subtype p} : (a : α) ≠ b ↔ a ≠ b := coe_injective.ne_iff -@[deprecated (since := "2024-04-04")] alias ⟨ne_of_val_ne, _⟩ := coe_ne_coe - @[simp] theorem _root_.exists_eq_subtype_mk_iff {a : Subtype p} {b : α} : (∃ h : p b, a = Subtype.mk b h) ↔ ↑a = b := diff --git a/Mathlib/Data/UInt.lean b/Mathlib/Data/UInt.lean index 001eb16bc105a..49f26afdb371a 100644 --- a/Mathlib/Data/UInt.lean +++ b/Mathlib/Data/UInt.lean @@ -117,12 +117,6 @@ def isASCIIDigit (c : UInt8) : Bool := def isASCIIAlphanum (c : UInt8) : Bool := c.isASCIIAlpha || c.isASCIIDigit -@[deprecated (since := "2024-06-06")] alias isUpper := isASCIIUpper -@[deprecated (since := "2024-06-06")] alias isLower := isASCIILower -@[deprecated (since := "2024-06-06")] alias isAlpha := isASCIIAlpha -@[deprecated (since := "2024-06-06")] alias isDigit := isASCIIDigit -@[deprecated (since := "2024-06-06")] alias isAlphanum := isASCIIAlphanum - /-- The numbers from 0 to 256 are all valid UTF-8 characters, so we can embed one in the other. -/ def toChar (n : UInt8) : Char := ⟨n.toUInt32, .inl (Nat.lt_trans n.toBitVec.isLt (by decide))⟩ diff --git a/Mathlib/Data/Vector/Basic.lean b/Mathlib/Data/Vector/Basic.lean index 40f05c5f9d267..17854c944ceed 100644 --- a/Mathlib/Data/Vector/Basic.lean +++ b/Mathlib/Data/Vector/Basic.lean @@ -541,15 +541,10 @@ theorem eraseIdx_val {i : Fin n} : ∀ {v : Vector α n}, (eraseIdx i v).val = v | _ => rfl @[deprecated (since := "2024-10-21")] alias eraseNth_val := eraseIdx_val -@[deprecated (since := "2024-05-04")] alias removeNth_val := eraseIdx_val - theorem eraseIdx_insertIdx {v : Vector α n} {i : Fin (n + 1)} : eraseIdx i (insertIdx a i v) = v := Subtype.eq <| List.eraseIdx_insertIdx i.1 v.1 -@[deprecated (since := "2024-05-04")] alias eraseIdx_insertNth := eraseIdx_insertIdx -@[deprecated (since := "2024-05-04")] alias removeNth_insertNth := eraseIdx_insertIdx - /-- Erasing an element after inserting an element, at different indices. -/ theorem eraseIdx_insertIdx' {v : Vector α (n + 1)} : ∀ {i : Fin (n + 1)} {j : Fin (n + 2)}, @@ -571,9 +566,6 @@ theorem eraseIdx_insertIdx' {v : Vector α (n + 1)} : · simpa · simpa [not_lt] using hij -@[deprecated (since := "2024-05-04")] alias eraseIdx_insertNth' := eraseIdx_insertIdx' -@[deprecated (since := "2024-05-04")] alias removeNth_insertNth' := eraseIdx_insertIdx' - theorem insertIdx_comm (a b : α) (i j : Fin (n + 1)) (h : i ≤ j) : ∀ v : Vector α n, (v.insertIdx a i).insertIdx b j.succ = (v.insertIdx b j).insertIdx a (Fin.castSucc i) diff --git a/Mathlib/Data/Vector/Defs.lean b/Mathlib/Data/Vector/Defs.lean index 342677efda76a..e41fc6aa29297 100644 --- a/Mathlib/Data/Vector/Defs.lean +++ b/Mathlib/Data/Vector/Defs.lean @@ -142,8 +142,6 @@ def take (i : ℕ) : Vector α n → Vector α (min i n) def eraseIdx (i : Fin n) : Vector α n → Vector α (n - 1) | ⟨l, p⟩ => ⟨List.eraseIdx l i.1, by rw [l.length_eraseIdx_of_lt] <;> rw [p]; exact i.2⟩ -@[deprecated (since := "2024-05-04")] alias removeNth := eraseIdx - /-- Vector of length `n` from a function on `Fin n`. -/ def ofFn : ∀ {n}, (Fin n → α) → Vector α n | 0, _ => nil diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 4bc9564fba5c8..65e7040c6ce47 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -85,9 +85,6 @@ theorem val_natCast {n : ℕ} (a : ℕ) : (a : ZMod n).val = a % n := by exact Int.natAbs_ofNat a · apply Fin.val_natCast -@[deprecated (since := "2024-04-17")] -alias val_nat_cast := val_natCast - theorem val_unit' {n : ZMod 0} : IsUnit n ↔ n.val = 1 := by simp only [val] rw [Int.isUnit_iff, Int.natAbs_eq_iff, Nat.cast_one] @@ -98,9 +95,6 @@ lemma eq_one_of_isUnit_natCast {n : ℕ} (h : IsUnit (n : ZMod 0)) : n = 1 := by theorem val_natCast_of_lt {n a : ℕ} (h : a < n) : (a : ZMod n).val = a := by rwa [val_natCast, Nat.mod_eq_of_lt] -@[deprecated (since := "2024-04-17")] -alias val_nat_cast_of_lt := val_natCast_of_lt - instance charP (n : ℕ) : CharP (ZMod n) n where cast_eq_zero_iff' := by intro k @@ -135,16 +129,10 @@ theorem ringChar_zmod_n (n : ℕ) : ringChar (ZMod n) = n := by theorem natCast_self (n : ℕ) : (n : ZMod n) = 0 := CharP.cast_eq_zero (ZMod n) n -@[deprecated (since := "2024-04-17")] -alias nat_cast_self := natCast_self - @[simp] theorem natCast_self' (n : ℕ) : (n + 1 : ZMod (n + 1)) = 0 := by rw [← Nat.cast_add_one, natCast_self (n + 1)] -@[deprecated (since := "2024-04-17")] -alias nat_cast_self' := natCast_self' - section UniversalProperty variable {n : ℕ} {R : Type*} @@ -196,21 +184,12 @@ theorem natCast_zmod_val {n : ℕ} [NeZero n] (a : ZMod n) : (a.val : ZMod n) = · cases NeZero.ne 0 rfl · apply Fin.cast_val_eq_self -@[deprecated (since := "2024-04-17")] -alias nat_cast_zmod_val := natCast_zmod_val - theorem natCast_rightInverse [NeZero n] : Function.RightInverse val ((↑) : ℕ → ZMod n) := natCast_zmod_val -@[deprecated (since := "2024-04-17")] -alias nat_cast_rightInverse := natCast_rightInverse - theorem natCast_zmod_surjective [NeZero n] : Function.Surjective ((↑) : ℕ → ZMod n) := natCast_rightInverse.surjective -@[deprecated (since := "2024-04-17")] -alias nat_cast_zmod_surjective := natCast_zmod_surjective - /-- So-named because the outer coercion is `Int.cast` into `ZMod`. For `Int.cast` into an arbitrary ring, see `ZMod.intCast_cast`. -/ @[norm_cast] @@ -220,21 +199,12 @@ theorem intCast_zmod_cast (a : ZMod n) : ((cast a : ℤ) : ZMod n) = a := by · dsimp [ZMod.cast, ZMod] erw [Int.cast_natCast, Fin.cast_val_eq_self] -@[deprecated (since := "2024-04-17")] -alias int_cast_zmod_cast := intCast_zmod_cast - theorem intCast_rightInverse : Function.RightInverse (cast : ZMod n → ℤ) ((↑) : ℤ → ZMod n) := intCast_zmod_cast -@[deprecated (since := "2024-04-17")] -alias int_cast_rightInverse := intCast_rightInverse - theorem intCast_surjective : Function.Surjective ((↑) : ℤ → ZMod n) := intCast_rightInverse.surjective -@[deprecated (since := "2024-04-17")] -alias int_cast_surjective := intCast_surjective - lemma «forall» {P : ZMod n → Prop} : (∀ x, P x) ↔ ∀ x : ℤ, P x := intCast_surjective.forall lemma «exists» {P : ZMod n → Prop} : (∃ x, P x) ↔ ∃ x : ℤ, P x := intCast_surjective.exists @@ -255,9 +225,6 @@ theorem natCast_comp_val [NeZero n] : ((↑) : ℕ → R) ∘ (val : ZMod n → · cases NeZero.ne 0 rfl rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_comp_val := natCast_comp_val - /-- The coercions are respectively `Int.cast`, `ZMod.cast`, and `ZMod.cast`. -/ @[simp] theorem intCast_comp_cast : ((↑) : ℤ → R) ∘ (cast : ZMod n → ℤ) = cast := by @@ -266,25 +233,16 @@ theorem intCast_comp_cast : ((↑) : ℤ → R) ∘ (cast : ZMod n → ℤ) = ca · ext simp [ZMod, ZMod.cast] -@[deprecated (since := "2024-04-17")] -alias int_cast_comp_cast := intCast_comp_cast - variable {R} @[simp] theorem natCast_val [NeZero n] (i : ZMod n) : (i.val : R) = cast i := congr_fun (natCast_comp_val R) i -@[deprecated (since := "2024-04-17")] -alias nat_cast_val := natCast_val - @[simp] theorem intCast_cast (i : ZMod n) : ((cast i : ℤ) : R) = cast i := congr_fun (intCast_comp_cast R) i -@[deprecated (since := "2024-04-17")] -alias int_cast_cast := intCast_cast - theorem cast_add_eq_ite {n : ℕ} (a b : ZMod n) : (cast (a + b) : ℤ) = if (n : ℤ) ≤ cast a + cast b then (cast a + cast b - n : ℤ) else cast a + cast b := by @@ -368,16 +326,10 @@ theorem cast_pow (h : m ∣ n) (a : ZMod n) (k : ℕ) : (cast (a ^ k : ZMod n) : theorem cast_natCast (h : m ∣ n) (k : ℕ) : (cast (k : ZMod n) : R) = k := map_natCast (castHom h R) k -@[deprecated (since := "2024-04-17")] -alias cast_nat_cast := cast_natCast - @[simp, norm_cast] theorem cast_intCast (h : m ∣ n) (k : ℤ) : (cast (k : ZMod n) : R) = k := map_intCast (castHom h R) k -@[deprecated (since := "2024-04-17")] -alias cast_int_cast := cast_intCast - end CharDvd section CharEq @@ -411,16 +363,10 @@ theorem cast_pow' (a : ZMod n) (k : ℕ) : (cast (a ^ k : ZMod n) : R) = (cast a theorem cast_natCast' (k : ℕ) : (cast (k : ZMod n) : R) = k := cast_natCast dvd_rfl k -@[deprecated (since := "2024-04-17")] -alias cast_nat_cast' := cast_natCast' - @[simp, norm_cast] theorem cast_intCast' (k : ℤ) : (cast (k : ZMod n) : R) = k := cast_intCast dvd_rfl k -@[deprecated (since := "2024-04-17")] -alias cast_int_cast' := cast_intCast' - variable (R) theorem castHom_injective : Function.Injective (ZMod.castHom (dvd_refl n) R) := by @@ -511,8 +457,6 @@ lemma ringEquivCongr_intCast {a b : ℕ} (h : a = b) (z : ℤ) : subst h cases a <;> rfl -@[deprecated (since := "2024-05-25")] alias int_coe_ringEquivCongr := ringEquivCongr_intCast - end CharEq end UniversalProperty @@ -529,62 +473,35 @@ theorem val_eq_zero : ∀ {n : ℕ} (a : ZMod n), a.val = 0 ↔ a = 0 theorem intCast_eq_intCast_iff (a b : ℤ) (c : ℕ) : (a : ZMod c) = (b : ZMod c) ↔ a ≡ b [ZMOD c] := CharP.intCast_eq_intCast (ZMod c) c -@[deprecated (since := "2024-04-17")] -alias int_cast_eq_int_cast_iff := intCast_eq_intCast_iff - theorem intCast_eq_intCast_iff' (a b : ℤ) (c : ℕ) : (a : ZMod c) = (b : ZMod c) ↔ a % c = b % c := ZMod.intCast_eq_intCast_iff a b c -@[deprecated (since := "2024-04-17")] -alias int_cast_eq_int_cast_iff' := intCast_eq_intCast_iff' - theorem val_intCast {n : ℕ} (a : ℤ) [NeZero n] : ↑(a : ZMod n).val = a % n := by have hle : (0 : ℤ) ≤ ↑(a : ZMod n).val := Int.natCast_nonneg _ have hlt : ↑(a : ZMod n).val < (n : ℤ) := Int.ofNat_lt.mpr (ZMod.val_lt a) refine (Int.emod_eq_of_lt hle hlt).symm.trans ?_ rw [← ZMod.intCast_eq_intCast_iff', Int.cast_natCast, ZMod.natCast_val, ZMod.cast_id] -@[deprecated (since := "2024-04-17")] -alias val_int_cast := val_intCast - theorem natCast_eq_natCast_iff (a b c : ℕ) : (a : ZMod c) = (b : ZMod c) ↔ a ≡ b [MOD c] := by simpa [Int.natCast_modEq_iff] using ZMod.intCast_eq_intCast_iff a b c -@[deprecated (since := "2024-04-17")] -alias nat_cast_eq_nat_cast_iff := natCast_eq_natCast_iff - theorem natCast_eq_natCast_iff' (a b c : ℕ) : (a : ZMod c) = (b : ZMod c) ↔ a % c = b % c := ZMod.natCast_eq_natCast_iff a b c -@[deprecated (since := "2024-04-17")] -alias nat_cast_eq_nat_cast_iff' := natCast_eq_natCast_iff' - theorem intCast_zmod_eq_zero_iff_dvd (a : ℤ) (b : ℕ) : (a : ZMod b) = 0 ↔ (b : ℤ) ∣ a := by rw [← Int.cast_zero, ZMod.intCast_eq_intCast_iff, Int.modEq_zero_iff_dvd] -@[deprecated (since := "2024-04-17")] -alias int_cast_zmod_eq_zero_iff_dvd := intCast_zmod_eq_zero_iff_dvd - theorem intCast_eq_intCast_iff_dvd_sub (a b : ℤ) (c : ℕ) : (a : ZMod c) = ↑b ↔ ↑c ∣ b - a := by rw [ZMod.intCast_eq_intCast_iff, Int.modEq_iff_dvd] -@[deprecated (since := "2024-04-17")] -alias int_cast_eq_int_cast_iff_dvd_sub := intCast_eq_intCast_iff_dvd_sub - theorem natCast_zmod_eq_zero_iff_dvd (a b : ℕ) : (a : ZMod b) = 0 ↔ b ∣ a := by rw [← Nat.cast_zero, ZMod.natCast_eq_natCast_iff, Nat.modEq_zero_iff_dvd] -@[deprecated (since := "2024-04-17")] -alias nat_cast_zmod_eq_zero_iff_dvd := natCast_zmod_eq_zero_iff_dvd - theorem coe_intCast (a : ℤ) : cast (a : ZMod n) = a % n := by cases n · rw [Int.ofNat_zero, Int.emod_zero, Int.cast_id]; rfl · rw [← val_intCast, val]; rfl -@[deprecated (since := "2024-04-17")] -alias coe_int_cast := coe_intCast - lemma intCast_cast_add (x y : ZMod n) : (cast (x + y) : ℤ) = (cast x + cast y) % n := by rw [← ZMod.coe_intCast, Int.cast_add, ZMod.intCast_zmod_cast, ZMod.intCast_zmod_cast] @@ -644,26 +561,17 @@ theorem intCast_eq_iff (p : ℕ) (n : ℤ) (z : ZMod p) [NeZero p] : rw [Int.cast_add, Int.cast_mul, Int.cast_natCast, Int.cast_natCast, natCast_val, ZMod.natCast_self, zero_mul, add_zero, cast_id] -@[deprecated (since := "2024-05-25")] alias nat_coe_zmod_eq_iff := natCast_eq_iff -@[deprecated (since := "2024-05-25")] alias int_coe_zmod_eq_iff := intCast_eq_iff - @[push_cast, simp] theorem intCast_mod (a : ℤ) (b : ℕ) : ((a % b : ℤ) : ZMod b) = (a : ZMod b) := by rw [ZMod.intCast_eq_intCast_iff] apply Int.mod_modEq -@[deprecated (since := "2024-04-17")] -alias int_cast_mod := intCast_mod - theorem ker_intCastAddHom (n : ℕ) : (Int.castAddHom (ZMod n)).ker = AddSubgroup.zmultiples (n : ℤ) := by ext rw [Int.mem_zmultiples_iff, AddMonoidHom.mem_ker, Int.coe_castAddHom, intCast_zmod_eq_zero_iff_dvd] -@[deprecated (since := "2024-04-17")] -alias ker_int_castAddHom := ker_intCastAddHom - theorem cast_injective_of_le {m n : ℕ} [nzm : NeZero m] (h : m ≤ n) : Function.Injective (@cast (ZMod n) _ m) := by cases m with @@ -688,9 +596,6 @@ theorem natCast_toNat (p : ℕ) : ∀ {z : ℤ} (_h : 0 ≤ z), (z.toNat : ZMod | (n : ℕ), _h => by simp only [Int.cast_natCast, Int.toNat_natCast] | Int.negSucc n, h => by simp at h -@[deprecated (since := "2024-04-17")] -alias nat_cast_toNat := natCast_toNat - theorem val_injective (n : ℕ) [NeZero n] : Function.Injective (val : ZMod n → ℕ) := by cases n · cases NeZero.ne 0 rfl @@ -821,9 +726,6 @@ theorem natCast_mod (a : ℕ) (n : ℕ) : ((a % n : ℕ) : ZMod n) = a := by rw [← Nat.mod_add_div a n] simp -@[deprecated (since := "2024-04-17")] -alias nat_cast_mod := natCast_mod - theorem eq_iff_modEq_nat (n : ℕ) {a b : ℕ} : (a : ZMod n) = b ↔ a ≡ b [MOD n] := by cases n · simp [Nat.ModEq, Int.natCast_inj, Nat.mod_zero] diff --git a/Mathlib/Data/ZMod/ValMinAbs.lean b/Mathlib/Data/ZMod/ValMinAbs.lean index cfbf77ffae802..26912ab823a9e 100644 --- a/Mathlib/Data/ZMod/ValMinAbs.lean +++ b/Mathlib/Data/ZMod/ValMinAbs.lean @@ -110,9 +110,6 @@ lemma natCast_natAbs_valMinAbs [NeZero n] (a : ZMod n) : · rw [← Int.cast_natCast, Int.ofNat_natAbs_of_nonpos this, Int.cast_neg, Int.cast_sub, Int.cast_natCast, Int.cast_natCast, natCast_self, sub_zero, natCast_zmod_val] -@[deprecated (since := "2024-04-17")] -alias nat_cast_natAbs_valMinAbs := natCast_natAbs_valMinAbs - lemma valMinAbs_neg_of_ne_half (ha : 2 * a.val ≠ n) : (-a).valMinAbs = -a.valMinAbs := by cases' eq_zero_or_neZero n with h h · subst h diff --git a/Mathlib/Deprecated/AlgebraClasses.lean b/Mathlib/Deprecated/AlgebraClasses.lean index e9e879b6c6967..5322e9f336146 100644 --- a/Mathlib/Deprecated/AlgebraClasses.lean +++ b/Mathlib/Deprecated/AlgebraClasses.lean @@ -3,8 +3,7 @@ Copyright (c) 2017 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ -import Mathlib.Order.Defs.Unbundled -import Mathlib.Order.Defs.LinearOrder +import Mathlib.Init /-! # Note about deprecated files @@ -33,174 +32,3 @@ class IsLeftCancel (α : Sort u) (op : α → α → α) : Prop where @[deprecated "No deprecation message was provided." (since := "2024-09-11")] class IsRightCancel (α : Sort u) (op : α → α → α) : Prop where right_cancel : ∀ a b c, op a b = op c b → a = c - -/-- `IsTotalPreorder X r` means that the binary relation `r` on `X` is total and a preorder. -/ -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -class IsTotalPreorder (α : Sort u) (r : α → α → Prop) extends IsTrans α r, IsTotal α r : Prop - -/-- Every total pre-order is a pre-order. -/ -instance (priority := 100) isTotalPreorder_isPreorder (α : Sort u) (r : α → α → Prop) - [s : IsTotalPreorder α r] : IsPreorder α r where - trans := s.trans - refl a := Or.elim (@IsTotal.total _ r _ a a) id id - -/-- `IsIncompTrans X lt` means that for `lt` a binary relation on `X`, the incomparable relation -`fun a b => ¬ lt a b ∧ ¬ lt b a` is transitive. -/ -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -class IsIncompTrans (α : Sort u) (lt : α → α → Prop) : Prop where - incomp_trans : ∀ a b c, ¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -instance (priority := 100) (α : Sort u) (lt : α → α → Prop) [IsStrictWeakOrder α lt] : - IsIncompTrans α lt := { ‹IsStrictWeakOrder α lt› with } - -section - -variable {r : α → α → Prop} - -local infixl:50 " ≺ " => r - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem incomp_trans [IsIncompTrans α r] {a b c : α} : - ¬a ≺ b ∧ ¬b ≺ a → ¬b ≺ c ∧ ¬c ≺ b → ¬a ≺ c ∧ ¬c ≺ a := - IsIncompTrans.incomp_trans _ _ _ - -section ExplicitRelationVariants - -variable (r) - -@[elab_without_expected_type, - deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem incomp_trans_of [IsIncompTrans α r] {a b c : α} : - ¬a ≺ b ∧ ¬b ≺ a → ¬b ≺ c ∧ ¬c ≺ b → ¬a ≺ c ∧ ¬c ≺ a := - incomp_trans - -end ExplicitRelationVariants - -end - -namespace StrictWeakOrder - -section - -variable {r : α → α → Prop} - -local infixl:50 " ≺ " => r - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -def Equiv (a b : α) : Prop := - ¬a ≺ b ∧ ¬b ≺ a - -local infixl:50 " ≈ " => @Equiv _ r - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem esymm {a b : α} : a ≈ b → b ≈ a := fun ⟨h₁, h₂⟩ => ⟨h₂, h₁⟩ - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem not_lt_of_equiv {a b : α} : a ≈ b → ¬a ≺ b := fun h => h.1 - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem not_lt_of_equiv' {a b : α} : a ≈ b → ¬b ≺ a := fun h => h.2 - -variable [IsStrictWeakOrder α r] - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem erefl (a : α) : a ≈ a := - ⟨irrefl a, irrefl a⟩ - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem etrans {a b c : α} : a ≈ b → b ≈ c → a ≈ c := - incomp_trans - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -instance isEquiv : IsEquiv α (@Equiv _ r) where - refl := erefl - trans _ _ _ := etrans - symm _ _ := esymm - -end - -/-- The equivalence relation induced by `lt` -/ -notation:50 a " ≈[" lt "]" b:50 => @Equiv _ lt a b - -end StrictWeakOrder - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem isStrictWeakOrder_of_isTotalPreorder {α : Sort u} {le : α → α → Prop} {lt : α → α → Prop} - [DecidableRel le] [IsTotalPreorder α le] (h : ∀ a b, lt a b ↔ ¬le b a) : - IsStrictWeakOrder α lt := - { trans := fun a b c hab hbc => - have nba : ¬le b a := Iff.mp (h _ _) hab - have ncb : ¬le c b := Iff.mp (h _ _) hbc - have hab : le a b := Or.resolve_left (total_of le b a) nba - have nca : ¬le c a := fun hca : le c a => - have hcb : le c b := trans_of le hca hab - absurd hcb ncb - Iff.mpr (h _ _) nca - irrefl := fun a hlt => absurd (refl_of le a) (Iff.mp (h _ _) hlt) - incomp_trans := fun a b c ⟨nab, nba⟩ ⟨nbc, ncb⟩ => - have hba : le b a := Decidable.of_not_not (Iff.mp (not_congr (h _ _)) nab) - have hab : le a b := Decidable.of_not_not (Iff.mp (not_congr (h _ _)) nba) - have hcb : le c b := Decidable.of_not_not (Iff.mp (not_congr (h _ _)) nbc) - have hbc : le b c := Decidable.of_not_not (Iff.mp (not_congr (h _ _)) ncb) - have hac : le a c := trans_of le hab hbc - have hca : le c a := trans_of le hcb hba - And.intro (fun n => absurd hca (Iff.mp (h _ _) n)) fun n => absurd hac (Iff.mp (h _ _) n) } - -section LinearOrder -variable {α : Type*} [LinearOrder α] - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -instance : IsTotalPreorder α (· ≤ ·) where - trans := @le_trans _ _ - total := le_total - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -instance isStrictWeakOrder_of_linearOrder : IsStrictWeakOrder α (· < ·) := - have : IsTotalPreorder α (· ≤ ·) := by infer_instance -- Porting note: added - isStrictWeakOrder_of_isTotalPreorder lt_iff_not_ge - -end LinearOrder - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem lt_of_lt_of_incomp {α : Sort u} {lt : α → α → Prop} [IsStrictWeakOrder α lt] - [DecidableRel lt] : ∀ {a b c}, lt a b → ¬lt b c ∧ ¬lt c b → lt a c := - @fun a b c hab ⟨nbc, ncb⟩ => - have nca : ¬lt c a := fun hca => absurd (trans_of lt hca hab) ncb - Decidable.byContradiction fun nac : ¬lt a c => - have : ¬lt a b ∧ ¬lt b a := incomp_trans_of lt ⟨nac, nca⟩ ⟨ncb, nbc⟩ - absurd hab this.1 - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem lt_of_incomp_of_lt {α : Sort u} {lt : α → α → Prop} [IsStrictWeakOrder α lt] - [DecidableRel lt] : ∀ {a b c}, ¬lt a b ∧ ¬lt b a → lt b c → lt a c := - @fun a b c ⟨nab, nba⟩ hbc => - have nca : ¬lt c a := fun hca => absurd (trans_of lt hbc hca) nba - Decidable.byContradiction fun nac : ¬lt a c => - have : ¬lt b c ∧ ¬lt c b := incomp_trans_of lt ⟨nba, nab⟩ ⟨nac, nca⟩ - absurd hbc this.1 - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem eq_of_incomp {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] {a b} : - ¬lt a b ∧ ¬lt b a → a = b := fun ⟨nab, nba⟩ => - match trichotomous_of lt a b with - | Or.inl hab => absurd hab nab - | Or.inr (Or.inl hab) => hab - | Or.inr (Or.inr hba) => absurd hba nba - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem eq_of_eqv_lt {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] {a b} : - a ≈[lt]b → a = b := - eq_of_incomp - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem incomp_iff_eq {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] [IsIrrefl α lt] (a b) : - ¬lt a b ∧ ¬lt b a ↔ a = b := - Iff.intro eq_of_incomp fun hab => hab ▸ And.intro (irrefl_of lt a) (irrefl_of lt a) - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem eqv_lt_iff_eq {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] [IsIrrefl α lt] (a b) : - a ≈[lt]b ↔ a = b := - incomp_iff_eq a b diff --git a/Mathlib/Deprecated/Aliases.lean b/Mathlib/Deprecated/Aliases.lean index 6266e7c3f17d2..458f6a85b1f4a 100644 --- a/Mathlib/Deprecated/Aliases.lean +++ b/Mathlib/Deprecated/Aliases.lean @@ -4,14 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.Init -import Batteries.Tactic.Alias /-! Deprecated aliases can be dumped here if they are no longer used in Mathlib, to avoid needing their imports if they are otherwise unnecessary. -/ -namespace String - -@[deprecated (since := "2024-06-04")] alias getRest := dropPrefix? - -end String diff --git a/Mathlib/Deprecated/Cardinal/Finite.lean b/Mathlib/Deprecated/Cardinal/Finite.lean index f28497050cfda..aa79f7f68de01 100644 --- a/Mathlib/Deprecated/Cardinal/Finite.lean +++ b/Mathlib/Deprecated/Cardinal/Finite.lean @@ -116,6 +116,4 @@ theorem card_eq_coe_natCard (α : Type*) [Finite α] : card α = Nat.card α := exact Finite.cast_card_eq_mk -@[deprecated (since := "2024-05-25")] alias card_eq_coe_nat_card := card_eq_coe_natCard - end PartENat diff --git a/Mathlib/Deprecated/Combinator.lean b/Mathlib/Deprecated/Combinator.lean deleted file mode 100644 index 366fd5997f982..0000000000000 --- a/Mathlib/Deprecated/Combinator.lean +++ /dev/null @@ -1,25 +0,0 @@ -/- -Copyright (c) 2014 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura, Mario Carneiro --/ - -import Mathlib.Init - -/-! -# Deprecated combinators, ported from Lean 3 core. --/ - -namespace Combinator - -universe u v w -variable {α : Sort u} {β : Sort v} {γ : Sort w} - -@[deprecated "No deprecation message was provided." (since := "2024-07-27")] -def I (a : α) := a -@[deprecated "No deprecation message was provided." (since := "2024-07-27")] -def K (a : α) (_b : β) := a -@[deprecated "No deprecation message was provided." (since := "2024-07-27")] -def S (x : α → β → γ) (y : α → β) (z : α) := x z (y z) - -end Combinator diff --git a/Mathlib/Deprecated/HashMap.lean b/Mathlib/Deprecated/HashMap.lean deleted file mode 100644 index b5c815bd6e5a0..0000000000000 --- a/Mathlib/Deprecated/HashMap.lean +++ /dev/null @@ -1,58 +0,0 @@ -/- -Copyright (c) 2022 Kim Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kim Morrison --/ - -import Mathlib.Init -import Mathlib.Tactic.TypeStar -import Batteries.Data.HashMap.Basic -import Batteries.Data.RBMap.Basic - -/-! -# Additional API for `HashMap` and `RBSet`. - -As `HashMap` has been completely reimplemented in `Batteries`, -nothing from the mathlib3 file `data.hash_map` is reflected here. -The porting header is just here to mark that no further work on `data.hash_map` is desired. --/ - -variable {α β : Type*} - -namespace Batteries.HashMap - --- not an exact match, the Lean3 version was dependently-typed - -variable [BEq α] [Hashable α] - -/-- The list of keys in a `HashMap`. -/ -@[deprecated "This declaration is unused in Mathlib: if you need it, \ - please file an issue in the Batteries repository." (since := "2024-06-12")] -def keys (m : HashMap α β) : List α := - m.fold (fun ks k _ => k :: ks) [] - -/-- The list of values in a `HashMap`. -/ -@[deprecated "This declaration is unused in Mathlib: if you need it, \ - please file an issue in the Batteries repository." (since := "2024-06-12")] -def values (m : HashMap α β) : List β := - m.fold (fun vs _ v => v :: vs) [] - -/-- Add a value to a `HashMap α (List β)` viewed as a multimap. -/ -@[deprecated "This declaration is unused in Mathlib: if you need it, \ - please file an issue in the Batteries repository." (since := "2024-06-12")] -def consVal (self : HashMap α (List β)) (a : α) (b : β) : HashMap α (List β) := - match self.find? a with - | none => self.insert a [b] - | some L => self.insert a (b::L) - -end Batteries.HashMap - -namespace Batteries.RBSet - -/-- Insert all elements of a list into an `RBSet`. -/ -@[deprecated "This declaration is unused in Mathlib: if you need it, \ - please file an issue in the Batteries repository." (since := "2024-06-12")] -def insertList {cmp} (m : RBSet α cmp) (L : List α) : RBSet α cmp := - L.foldl (fun m a => m.insert a) m - -end Batteries.RBSet diff --git a/Mathlib/Deprecated/LazyList.lean b/Mathlib/Deprecated/LazyList.lean deleted file mode 100644 index 80dbb7f596a7a..0000000000000 --- a/Mathlib/Deprecated/LazyList.lean +++ /dev/null @@ -1,112 +0,0 @@ -/- -Copyright (c) 2018 Simon Hudon. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Simon Hudon --/ -import Mathlib.Control.Traversable.Equiv -import Mathlib.Control.Traversable.Instances -import Batteries.Data.LazyList -import Mathlib.Lean.Thunk - -/-! -## Definitions on lazy lists - -This file is entirely deprecated, and contains various definitions and proofs on lazy lists. --/ - --- The whole file is full of deprecations about LazyList -set_option linter.deprecated false - -universe u - -namespace LazyList - -open Function - -/-- Isomorphism between strict and lazy lists. -/ -@[deprecated "No deprecation message was provided." (since := "2024-07-22")] -def listEquivLazyList (α : Type*) : List α ≃ LazyList α where - toFun := LazyList.ofList - invFun := LazyList.toList - right_inv := by - intro xs - induction xs using toList.induct - · simp [toList, ofList] - · simp [toList, ofList, *]; rfl - left_inv := by - intro xs - induction xs - · simp [toList, ofList] - · simpa [ofList, toList] - -@[deprecated "No deprecation message was provided." (since := "2024-07-22")] -instance : Traversable LazyList where - map := @LazyList.traverse Id _ - traverse := @LazyList.traverse - -@[deprecated "No deprecation message was provided." (since := "2024-07-22")] -instance : LawfulTraversable LazyList := by - apply Equiv.isLawfulTraversable' listEquivLazyList <;> intros <;> ext <;> rename_i f xs - · induction xs using LazyList.rec with - | nil => - simp only [Functor.map, LazyList.traverse, pure, Equiv.map, listEquivLazyList, - Equiv.coe_fn_symm_mk, toList, List.map_nil, Equiv.coe_fn_mk, ofList] - | cons => - simpa only [Functor.map, LazyList.traverse, Seq.seq, Equiv.map, listEquivLazyList, - Equiv.coe_fn_symm_mk, toList, List.map_cons, Equiv.coe_fn_mk, ofList, cons.injEq, true_and] - | mk _ ih => ext; apply ih - · simp only [Functor.mapConst, comp, Equiv.map, listEquivLazyList, Equiv.coe_fn_symm_mk, - List.map_eq_map, List.map_const, Equiv.coe_fn_mk] - induction xs using LazyList.rec with - | nil => simp [LazyList.traverse, pure, Functor.map, toList, ofList] - | cons => - simpa [toList, ofList, LazyList.traverse, Seq.seq, Functor.map, cons.injEq, true_and] - | mk _ ih => congr; apply ih - · simp only [traverse, Equiv.traverse, listEquivLazyList, Equiv.coe_fn_mk, Equiv.coe_fn_symm_mk] - induction xs using LazyList.rec with - | nil => simp only [LazyList.traverse, toList, List.traverse, map_pure, ofList] - | cons _ tl ih => - replace ih : tl.get.traverse f = ofList <$> tl.get.toList.traverse f := ih - simp [traverse.eq_2, ih, Functor.map_map, seq_map_assoc, toList, List.traverse, map_seq, - Function.comp_def, Thunk.pure, ofList] - | mk _ ih => apply ih - -@[deprecated "No deprecation message was provided." (since := "2024-07-22"), simp] -theorem bind_singleton {α} (x : LazyList α) : x.bind singleton = x := by - induction x using LazyList.rec (motive_2 := fun xs => xs.get.bind singleton = xs.get) with - | nil => simp [LazyList.bind] - | cons h t ih => - simp only [LazyList.bind, singleton, append, Thunk.get_pure, Thunk.get_mk, cons.injEq, true_and] - ext - simp [ih] - | mk f ih => simp_all - -@[deprecated "No deprecation message was provided." (since := "2024-07-22")] -instance : LawfulMonad LazyList := LawfulMonad.mk' - (id_map := by - intro α xs - induction xs using LazyList.rec (motive_2 := fun xs => id <$> xs.get = xs) with - | nil => simp only [Functor.map, comp_id, LazyList.bind] - | cons h t _ => simp only [Functor.map, comp_id, bind_singleton] - | mk f _ => ext; simp_all) - (pure_bind := by - intros - simp only [bind, pure, singleton, LazyList.bind, append, Thunk.pure, Thunk.get] - apply append_nil) - (bind_assoc := by - intro _ _ _ xs _ _ - induction xs using LazyList.rec with - | nil => simp only [bind, LazyList.bind] - | cons => simp only [bind, LazyList.bind, append_bind]; congr - | mk _ ih => congr; funext; apply ih) - (bind_pure_comp := by - intro _ _ f xs - simp only [bind, Functor.map, pure, singleton] - induction xs using LazyList.traverse.induct with - | case1 => - simp only [Thunk.pure, LazyList.bind, LazyList.traverse, Id.pure_eq] - | case2 _ _ ih => - simp only [Thunk.pure, LazyList.bind, append, Thunk.get_mk, comp_apply, ← ih] - simp only [Thunk.get, append, singleton, Thunk.pure]) - -end LazyList diff --git a/Mathlib/Deprecated/Logic.lean b/Mathlib/Deprecated/Logic.lean index 52d03620dc96a..8c1f62cfd3fbb 100644 --- a/Mathlib/Deprecated/Logic.lean +++ b/Mathlib/Deprecated/Logic.lean @@ -134,8 +134,6 @@ def recOn_false [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} @[deprecated (since := "2024-09-03")] alias by_cases := byCases @[deprecated (since := "2024-09-03")] alias by_contradiction := byContradiction -@[deprecated (since := "2024-07-27")] alias not_not_iff := not_not - end Decidable @[deprecated (since := "2024-09-03")] alias Or.decidable := instDecidableOr diff --git a/Mathlib/Deprecated/NatLemmas.lean b/Mathlib/Deprecated/NatLemmas.lean index 81b124b4476cd..c014be7eed503 100644 --- a/Mathlib/Deprecated/NatLemmas.lean +++ b/Mathlib/Deprecated/NatLemmas.lean @@ -34,31 +34,4 @@ def discriminate {B : Sort u} {n : ℕ} (H1 : n = 0 → B) (H2 : ∀ m, n = succ | zero => exact H1 rfl | succ n _ => apply H2 _ rfl --- Unused in Mathlib; --- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated "No deprecation message was provided." (since := "2024-07-27")] -theorem one_eq_succ_zero : 1 = succ 0 := - rfl - -/-! induction principles -/ - --- Unused in Mathlib; --- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated "No deprecation message was provided." (since := "2024-07-27")] -def subInduction {P : ℕ → ℕ → Sort u} (H1 : ∀ m, P 0 m) (H2 : ∀ n, P (succ n) 0) - (H3 : ∀ n m, P n m → P (succ n) (succ m)) : ∀ n m : ℕ, P n m - | 0, _m => H1 _ - | succ _n, 0 => H2 _ - | succ n, succ m => H3 _ _ (subInduction H1 H2 H3 n m) - -/-! mod -/ - --- Unused in Mathlib; --- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated "No deprecation message was provided." (since := "2024-07-27")] -theorem cond_decide_mod_two (x : ℕ) [d : Decidable (x % 2 = 1)] : - cond (@decide (x % 2 = 1) d) 1 0 = x % 2 := by - simp only [cond_eq_if, decide_eq_true_eq] - split <;> omega - end Nat diff --git a/Mathlib/Deprecated/RelClasses.lean b/Mathlib/Deprecated/RelClasses.lean deleted file mode 100644 index 8d5f647030d5d..0000000000000 --- a/Mathlib/Deprecated/RelClasses.lean +++ /dev/null @@ -1,40 +0,0 @@ -/- -Copyright (c) 2020 Jeremy Avigad. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad, Mario Carneiro, Yury Kudryashov --/ -import Mathlib.Deprecated.AlgebraClasses -import Mathlib.Order.RelClasses - -/-! -# Note about deprecated files - -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. - -# Unbundled relation classes - -In this file we prove some properties of `Is*` classes defined in -`Mathlib/Deprecated/AlgebraClasses.lean`. - --/ - -set_option linter.deprecated false - -universe u v - -variable {α : Type u} - -open Function - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem IsTotalPreorder.swap (r) [IsTotalPreorder α r] : IsTotalPreorder α (swap r) := - { @IsPreorder.swap α r _, @IsTotal.swap α r _ with } - -@[deprecated "No deprecation message was provided." (since := "2024-08-22")] -instance [LinearOrder α] : IsTotalPreorder α (· ≤ ·) where -@[deprecated "No deprecation message was provided." (since := "2024-08-22")] -instance [LinearOrder α] : IsTotalPreorder α (· ≥ ·) where - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -instance [LinearOrder α] : IsIncompTrans α (· < ·) := by infer_instance diff --git a/Mathlib/Dynamics/Ergodic/Ergodic.lean b/Mathlib/Dynamics/Ergodic/Ergodic.lean index 29c03db776487..e5199e430406b 100644 --- a/Mathlib/Dynamics/Ergodic/Ergodic.lean +++ b/Mathlib/Dynamics/Ergodic/Ergodic.lean @@ -125,8 +125,6 @@ theorem ae_empty_or_univ₀ (hf : QuasiErgodic f μ) (hsm : NullMeasurableSet s s =ᵐ[μ] (∅ : Set α) ∨ s =ᵐ[μ] univ := eventuallyConst_set'.mp <| hf.aeconst_set₀ hsm hs -@[deprecated (since := "2024-07-21")] alias ae_empty_or_univ' := ae_empty_or_univ₀ - /-- For a quasi ergodic map, sets that are almost invariant (rather than strictly invariant) are still either almost empty or full. -/ theorem ae_mem_or_ae_nmem₀ (hf : QuasiErgodic f μ) (hsm : NullMeasurableSet s μ) diff --git a/Mathlib/FieldTheory/IntermediateField/Adjoin/Defs.lean b/Mathlib/FieldTheory/IntermediateField/Adjoin/Defs.lean index c0b0a9bb9f508..b848f5175a8a1 100644 --- a/Mathlib/FieldTheory/IntermediateField/Adjoin/Defs.lean +++ b/Mathlib/FieldTheory/IntermediateField/Adjoin/Defs.lean @@ -625,9 +625,6 @@ theorem adjoin_intCast (n : ℤ) : F⟮(n : E)⟯ = ⊥ := by theorem adjoin_natCast (n : ℕ) : F⟮(n : E)⟯ = ⊥ := adjoin_simple_eq_bot_iff.mpr (natCast_mem ⊥ n) -@[deprecated (since := "2024-04-05")] alias adjoin_int := adjoin_intCast -@[deprecated (since := "2024-04-05")] alias adjoin_nat := adjoin_natCast - end AdjoinIntermediateFieldLattice section Induction diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index 7f59bd51abc1d..1130f9fda6c77 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -230,9 +230,6 @@ end InheritedLemmas theorem natCast_mem (n : ℕ) : (n : L) ∈ S := by simpa using intCast_mem S n -@[deprecated _root_.natCast_mem (since := "2024-04-05")] alias coe_nat_mem := natCast_mem -@[deprecated _root_.intCast_mem (since := "2024-04-05")] alias coe_int_mem := intCast_mem - end IntermediateField /-- Turn a subalgebra closed under inverses into an intermediate field -/ diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index 6c1c121b87fb7..f64c0cb3f0efd 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -184,16 +184,6 @@ theorem algebraMap_surjective_of_isIntegral' {k K : Type*} [Field k] [CommRing K have : Algebra.IsIntegral k K := ⟨hf⟩ algebraMap_surjective_of_isIntegral -/-- -Deprecated: `algebraMap_surjective_of_isIntegral` is identical apart from the `IsIntegral` argument, -which can be found by instance synthesis --/ -@[deprecated algebraMap_surjective_of_isIntegral (since := "2024-05-08")] -theorem algebraMap_surjective_of_isAlgebraic {k K : Type*} [Field k] [Ring K] [IsDomain K] - [IsAlgClosed k] [Algebra k K] [Algebra.IsAlgebraic k K] : - Function.Surjective (algebraMap k K) := - algebraMap_surjective_of_isIntegral - end IsAlgClosed /-- If `k` is algebraically closed, `K / k` is a field extension, `L / k` is an intermediate field diff --git a/Mathlib/FieldTheory/PerfectClosure.lean b/Mathlib/FieldTheory/PerfectClosure.lean index 7f553815a9734..7d5622aa37b72 100644 --- a/Mathlib/FieldTheory/PerfectClosure.lean +++ b/Mathlib/FieldTheory/PerfectClosure.lean @@ -360,16 +360,10 @@ theorem natCast (n x : ℕ) : (x : PerfectClosure K p) = mk K p (n, x) := by rwa [frobenius_natCast K p x] at this apply R.intro -@[deprecated (since := "2024-04-17")] -alias nat_cast := natCast - theorem intCast (x : ℤ) : (x : PerfectClosure K p) = mk K p (0, x) := by induction x <;> simp only [Int.ofNat_eq_coe, Int.cast_natCast, Int.cast_negSucc, natCast K p 0] rfl -@[deprecated (since := "2024-04-17")] -alias int_cast := intCast - theorem natCast_eq_iff (x y : ℕ) : (x : PerfectClosure K p) = y ↔ (x : K) = y := by constructor <;> intro H · rw [natCast K p 0, natCast K p 0, mk_eq_iff] at H @@ -377,9 +371,6 @@ theorem natCast_eq_iff (x y : ℕ) : (x : PerfectClosure K p) = y ↔ (x : K) = simpa only [zero_add, iterate_fixed (frobenius_natCast K p _)] using H rw [natCast K p 0, natCast K p 0, H] -@[deprecated (since := "2024-04-17")] -alias nat_cast_eq_iff := natCast_eq_iff - instance instCharP : CharP (PerfectClosure K p) p := by constructor; intro x; rw [← CharP.cast_eq_zero_iff K] rw [← Nat.cast_zero, natCast_eq_iff, Nat.cast_zero] diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index f18b90045c1e6..c93692ddc6070 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -117,7 +117,6 @@ instance : LocPathConnectedSpace (EuclideanQuadrant n) := theorem range_euclideanHalfSpace (n : ℕ) [NeZero n] : (range fun x : EuclideanHalfSpace n => x.val) = { y | 0 ≤ y 0 } := Subtype.range_val -@[deprecated (since := "2024-04-05")] alias range_half_space := range_euclideanHalfSpace open ENNReal in @[simp] @@ -164,7 +163,6 @@ theorem frontier_halfSpace {n : ℕ} (p : ℝ≥0∞) (a : ℝ) (i : Fin n) : theorem range_euclideanQuadrant (n : ℕ) : (range fun x : EuclideanQuadrant n => x.val) = { y | ∀ i : Fin n, 0 ≤ y i } := Subtype.range_val -@[deprecated (since := "2024-04-05")] alias range_quadrant := range_euclideanQuadrant end diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index d525541360b7d..d5f85d4c10547 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -575,6 +575,4 @@ instance : LieGroup (𝓡 1) ω Circle where theorem contMDiff_circleExp {m : WithTop ℕ∞} : ContMDiff 𝓘(ℝ, ℝ) (𝓡 1) m Circle.exp := (contDiff_exp.comp (contDiff_id.smul contDiff_const)).contMDiff.codRestrict_sphere _ -@[deprecated (since := "2024-07-25")] alias contMDiff_expMapCircle := contMDiff_circleExp - end Circle diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean index 8fae4d69ee6bd..311efe82069d7 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean @@ -224,9 +224,6 @@ theorem mdifferentiableWithinAt_iff' (f : M → M') (s : Set M) (x : M) : ((extChartAt I x).symm ⁻¹' s ∩ range I) ((extChartAt I x) x) := by rw [MDifferentiableWithinAt, liftPropWithinAt_iff']; rfl -@[deprecated (since := "2024-04-30")] -alias mdifferentiableWithinAt_iff_liftPropWithinAt := mdifferentiableWithinAt_iff' - theorem MDifferentiableWithinAt.continuousWithinAt {f : M → M'} {s : Set M} {x : M} (hf : MDifferentiableWithinAt I I' f s x) : ContinuousWithinAt f s x := @@ -259,9 +256,6 @@ theorem mdifferentiableAt_iff (f : M → M') (x : M) : -- Porting note: `rfl` wasn't needed rfl -@[deprecated (since := "2024-04-30")] -alias mdifferentiableAt_iff_liftPropAt := mdifferentiableAt_iff - theorem MDifferentiableAt.continuousAt {f : M → M'} {x : M} (hf : MDifferentiableAt I I' f x) : ContinuousAt f x := mdifferentiableAt_iff .. |>.1 hf |>.1 diff --git a/Mathlib/Geometry/RingedSpace/Stalks.lean b/Mathlib/Geometry/RingedSpace/Stalks.lean index 96927f0280b51..9574828c1ed89 100644 --- a/Mathlib/Geometry/RingedSpace/Stalks.lean +++ b/Mathlib/Geometry/RingedSpace/Stalks.lean @@ -48,9 +48,6 @@ theorem stalkMap_germ {X Y : PresheafedSpace.{_, _, v} C} (α : X ⟶ Y) (U : Op X.presheaf.germ ((Opens.map α.base).obj U) x hx := by rw [Hom.stalkMap, stalkFunctor_map_germ_assoc, stalkPushforward_germ] -@[deprecated (since := "2024-07-30")] alias stalkMap_germ' := stalkMap_germ -@[deprecated (since := "2024-07-30")] alias stalkMap_germ'_assoc := stalkMap_germ - section Restrict /-- For an open embedding `f : U ⟶ X` and a point `x : U`, we get an isomorphism between the stalk diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index 0aa68664c8997..1ec207e2b1a3b 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -498,11 +498,6 @@ section Group variable [Group G] {n m : ℤ} -@[to_additive (attr := deprecated Monoid.one_lt_exponent (since := "2024-02-17")) - AddGroup.one_lt_exponent] -lemma Group.one_lt_exponent [Finite G] [Nontrivial G] : 1 < Monoid.exponent G := - Monoid.one_lt_exponent - @[to_additive] theorem Group.exponent_dvd_card [Fintype G] : Monoid.exponent G ∣ Fintype.card G := Monoid.exponent_dvd.mpr <| fun _ => orderOf_dvd_card @@ -645,14 +640,6 @@ lemma inv_eq_self_of_orderOf_eq_two {x : G} (hx : orderOf x = 2) : x⁻¹ = x := inv_eq_of_mul_eq_one_left <| pow_two (a := x) ▸ hx ▸ pow_orderOf_eq_one x --- TODO: delete -/-- Any group of exponent two is abelian. -/ -@[to_additive (attr := reducible, - deprecated "No deprecation message was provided." (since := "2024-02-17")) - "Any additive group of exponent two is abelian."] -def instCommGroupOfExponentTwo (hG : Monoid.exponent G = 2) : CommGroup G where - mul_comm := mul_comm_of_exponent_two hG - @[to_additive] lemma mul_not_mem_of_orderOf_eq_two {x y : G} (hx : orderOf x = 2) (hy : orderOf y = 2) (hxy : x ≠ y) : x * y ∉ ({x, y, 1} : Set G) := by diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index f8e36f80728c3..f538b163a58df 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -185,8 +185,6 @@ theorem index_top : (⊤ : Subgroup G).index = 1 := theorem index_bot : (⊥ : Subgroup G).index = Nat.card G := Cardinal.toNat_congr QuotientGroup.quotientBot.toEquiv -@[deprecated (since := "2024-06-15")] alias index_bot_eq_card := index_bot - @[to_additive (attr := simp)] theorem relindex_top_left : (⊤ : Subgroup G).relindex H = 1 := index_top @@ -199,8 +197,6 @@ theorem relindex_top_right : H.relindex ⊤ = H.index := by theorem relindex_bot_left : (⊥ : Subgroup G).relindex H = Nat.card H := by rw [relindex, bot_subgroupOf, index_bot] -@[deprecated (since := "2024-06-15")] alias relindex_bot_left_eq_card := relindex_bot_left - @[to_additive (attr := simp)] theorem relindex_bot_right : H.relindex ⊥ = 1 := by rw [relindex, subgroupOf_bot_eq_top, index_top] @@ -220,18 +216,12 @@ theorem card_mul_index : Nat.card H * H.index = Nat.card G := by rw [← relindex_bot_left, ← index_bot] exact relindex_mul_index bot_le -@[deprecated (since := "2024-06-15")] alias nat_card_dvd_of_injective := card_dvd_of_injective - -@[deprecated (since := "2024-06-15")] alias nat_card_dvd_of_le := card_dvd_of_le - @[to_additive] theorem card_dvd_of_surjective (f : G →* G') (hf : Function.Surjective f) : Nat.card G' ∣ Nat.card G := by rw [← Nat.card_congr (QuotientGroup.quotientKerEquivOfSurjective f hf).toEquiv] exact Dvd.intro_left (Nat.card f.ker) f.ker.card_mul_index -@[deprecated (since := "2024-06-15")] alias nat_card_dvd_of_surjective := card_dvd_of_surjective - @[to_additive] theorem card_range_dvd (f : G →* G') : Nat.card f.range ∣ Nat.card G := card_dvd_of_surjective f.rangeRestrict f.rangeRestrict_surjective diff --git a/Mathlib/GroupTheory/Nilpotent.lean b/Mathlib/GroupTheory/Nilpotent.lean index 8c188f330ea88..74feb4afd4251 100644 --- a/Mathlib/GroupTheory/Nilpotent.lean +++ b/Mathlib/GroupTheory/Nilpotent.lean @@ -862,6 +862,4 @@ theorem isNilpotent_of_finite_tfae : | ⟨e⟩ => isNilpotent_of_product_of_sylow_group e tfae_finish -@[deprecated (since := "2024-06-05")] alias isNilpotent_of_finite_tFAE := isNilpotent_of_finite_tfae - end WithFiniteGroup diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 4a207ce12d1ef..a609b7d46308d 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -214,10 +214,6 @@ theorem IsOfFinOrder.mono [Monoid β] {y : β} (hx : IsOfFinOrder x) (h : orderO @[to_additive] theorem pow_ne_one_of_lt_orderOf (n0 : n ≠ 0) (h : n < orderOf x) : x ^ n ≠ 1 := fun j => not_isPeriodicPt_of_pos_of_lt_minimalPeriod n0 h ((isPeriodicPt_mul_iff_pow_eq_one x).mpr j) -@[deprecated (since := "2024-07-20")] alias pow_ne_one_of_lt_orderOf' := pow_ne_one_of_lt_orderOf -@[deprecated (since := "2024-07-20")] alias - nsmul_ne_zero_of_lt_addOrderOf' := nsmul_ne_zero_of_lt_addOrderOf - @[to_additive] theorem orderOf_le_of_pow_eq_one (hn : 0 < n) (h : x ^ n = 1) : orderOf x ≤ n := IsPeriodicPt.minimalPeriod_le hn (by rwa [isPeriodicPt_mul_iff_pow_eq_one]) diff --git a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean index e5f7bc27cc91e..e20ede054c539 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean @@ -263,11 +263,6 @@ theorem two_le_card_support_cycleOf_iff [DecidableEq α] [Fintype α] : rw [← two_le_card_support_cycleOf_iff, ← card_pos, ← Nat.succ_le_iff] exact ⟨fun h => Or.resolve_left h.eq_or_lt (card_support_ne_one _).symm, zero_lt_two.trans_le⟩ -@[deprecated support_cycleOf_nonempty (since := "2024-06-16")] -theorem card_support_cycleOf_pos_iff [DecidableEq α] [Fintype α] : - 0 < #(cycleOf f x).support ↔ f x ≠ x := by - rw [card_pos, support_cycleOf_nonempty] - theorem mem_support_cycleOf_iff [DecidableEq α] [Fintype α] : y ∈ support (f.cycleOf x) ↔ SameCycle f x y ∧ x ∈ support f := mem_support_cycleOf_iff_aux diff --git a/Mathlib/Lean/Expr/Basic.lean b/Mathlib/Lean/Expr/Basic.lean index bfb1353ed2c58..647e543bd9a7b 100644 --- a/Mathlib/Lean/Expr/Basic.lean +++ b/Mathlib/Lean/Expr/Basic.lean @@ -66,8 +66,6 @@ def lastComponentAsString : Name → String | .num _ n => toString n | .anonymous => "" -@[deprecated (since := "2024-05-14")] alias getString := lastComponentAsString - /-- `nm.splitAt n` splits a name `nm` in two parts, such that the *second* part has depth `n`, i.e. `(nm.splitAt n).2.getNumParts = n` (assuming `nm.getNumParts ≥ n`). Example: ``splitAt `foo.bar.baz.back.bat 1 = (`foo.bar.baz.back, `bat)``. -/ diff --git a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean index da171b8d6bfaa..30669576c7a9d 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean @@ -58,10 +58,6 @@ namespace LinearMap namespace BilinForm -@[deprecated "No deprecation message was provided." (since := "2024-04-14")] -theorem coeFn_congr : ∀ {x x' y y' : M}, x = x' → y = y' → B x y = B x' y' - | _, _, _, _, rfl, rfl => rfl - theorem add_left (x y z : M) : B (x + y) z = B x z + B y z := map_add₂ _ _ _ _ theorem smul_left (a : R) (x y : M) : B (a • x) y = a * B x y := map_smul₂ _ _ _ _ @@ -101,36 +97,20 @@ theorem ext (H : ∀ x y : M, B x y = D x y) : B = D := ext₂ H theorem congr_fun (h : B = D) (x y : M) : B x y = D x y := congr_fun₂ h _ _ -@[deprecated "No deprecation message was provided." (since := "2024-04-14")] -theorem coe_zero : ⇑(0 : BilinForm R M) = 0 := - rfl - @[simp] theorem zero_apply (x y : M) : (0 : BilinForm R M) x y = 0 := rfl variable (B D B₁ D₁) -@[deprecated "No deprecation message was provided." (since := "2024-04-14")] -theorem coe_add : ⇑(B + D) = B + D := - rfl - @[simp] theorem add_apply (x y : M) : (B + D) x y = B x y + D x y := rfl -@[deprecated "No deprecation message was provided." (since := "2024-04-14")] -theorem coe_neg : ⇑(-B₁) = -B₁ := - rfl - @[simp] theorem neg_apply (x y : M₁) : (-B₁) x y = -B₁ x y := rfl -@[deprecated "No deprecation message was provided." (since := "2024-04-14")] -theorem coe_sub : ⇑(B₁ - D₁) = B₁ - D₁ := - rfl - @[simp] theorem sub_apply (x y : M₁) : (B₁ - D₁) x y = B₁ x y - D₁ x y := rfl diff --git a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean index c1b171f04d13e..64431f2c30303 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean @@ -57,20 +57,6 @@ section ToLin' /-- Auxiliary definition to define `toLinHom`; see below. -/ def toLinHomAux₁ (A : BilinForm R M) (x : M) : M →ₗ[R] R := A x -/-- Auxiliary definition to define `toLinHom`; see below. -/ -@[deprecated "No deprecation message was provided." (since := "2024-04-26")] -def toLinHomAux₂ (A : BilinForm R M) : M →ₗ[R] M →ₗ[R] R := A - -/-- The linear map obtained from a `BilinForm` by fixing the left co-ordinate and evaluating in -the right. -/ -@[deprecated "No deprecation message was provided." (since := "2024-04-26")] -def toLinHom : BilinForm R M →ₗ[R] M →ₗ[R] M →ₗ[R] R := LinearMap.id - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-04-26")] -theorem toLin'_apply (A : BilinForm R M) (x : M) : toLinHom (M := M) A x = A x := - rfl - variable (B) theorem sum_left {α} (t : Finset α) (g : α → M) (w : M) : @@ -102,59 +88,6 @@ end BilinForm end LinearMap -section EquivLin - -/-- A map with two arguments that is linear in both is a bilinear form. - -This is an auxiliary definition for the full linear equivalence `LinearMap.toBilin`. --/ -def LinearMap.toBilinAux (f : M →ₗ[R] M →ₗ[R] R) : BilinForm R M := f - -set_option linter.deprecated false in -/-- Bilinear forms are linearly equivalent to maps with two arguments that are linear in both. -/ -@[deprecated "No deprecation message was provided." (since := "2024-04-26")] -def LinearMap.BilinForm.toLin : BilinForm R M ≃ₗ[R] M →ₗ[R] M →ₗ[R] R := - { BilinForm.toLinHom with - invFun := LinearMap.toBilinAux - left_inv := fun _ => rfl - right_inv := fun _ => rfl } - -set_option linter.deprecated false in -/-- A map with two arguments that is linear in both is linearly equivalent to bilinear form. -/ -@[deprecated "No deprecation message was provided." (since := "2024-04-26")] -def LinearMap.toBilin : (M →ₗ[R] M →ₗ[R] R) ≃ₗ[R] BilinForm R M := - BilinForm.toLin.symm - -@[deprecated "No deprecation message was provided." (since := "2024-04-26")] -theorem LinearMap.toBilinAux_eq (f : M →ₗ[R] M →ₗ[R] R) : - LinearMap.toBilinAux f = f := - rfl - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-04-26")] -theorem LinearMap.toBilin_symm : - (LinearMap.toBilin.symm : BilinForm R M ≃ₗ[R] _) = BilinForm.toLin := - rfl - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-04-26")] -theorem BilinForm.toLin_symm : - (BilinForm.toLin.symm : _ ≃ₗ[R] BilinForm R M) = LinearMap.toBilin := - LinearMap.toBilin.symm_symm - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-04-26")] -theorem LinearMap.toBilin_apply (f : M →ₗ[R] M →ₗ[R] R) (x y : M) : - toBilin f x y = f x y := - rfl - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-04-26")] -theorem BilinForm.toLin_apply (x : M) : BilinForm.toLin B x = B x := - rfl - -end EquivLin - namespace LinearMap variable {R' : Type*} [CommSemiring R'] [Algebra R' R] [Module R' M] [IsScalarTower R' R M] diff --git a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean index 6cebe6c4de152..b3ce1e567f3b3 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean @@ -210,9 +210,6 @@ theorem nondegenerate_restrict_of_disjoint_orthogonal (B : BilinForm R₁ M₁) simp only [restrict_apply, domRestrict_apply] at b₁ exact isOrtho_def.mpr (b x y b₁) -@[deprecated (since := "2024-05-30")] -alias nondegenerateRestrictOfDisjointOrthogonal := nondegenerate_restrict_of_disjoint_orthogonal - /-- An orthogonal basis with respect to a nondegenerate bilinear form has no self-orthogonal elements. -/ theorem iIsOrtho.not_isOrtho_basis_self_of_nondegenerate {n : Type w} [Nontrivial R] @@ -364,9 +361,6 @@ theorem isCompl_orthogonal_of_restrict_nondegenerate finrank_add_finrank_orthogonal b₁] exact le_self_add -@[deprecated (since := "2024-05-24")] -alias restrict_nondegenerate_of_isCompl_orthogonal := isCompl_orthogonal_of_restrict_nondegenerate - /-- A subspace is complement to its orthogonal complement with respect to some reflexive bilinear form if and only if that bilinear form restricted on to the subspace is nondegenerate. -/ theorem restrict_nondegenerate_iff_isCompl_orthogonal @@ -414,10 +408,6 @@ theorem restrict_nondegenerate_orthogonal_spanSingleton (B : BilinForm K V) (b rw [restrict] at hm erw [add_right, show B m.1 y = 0 by rw [b₂]; exact m.2 y hy, hm, add_zero] -@[deprecated (since := "2024-05-30")] -alias restrictNondegenerateOrthogonalSpanSingleton := - restrict_nondegenerate_orthogonal_spanSingleton - end BilinForm end LinearMap diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index f52cccf477bc3..9d9f144fa4213 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -5,7 +5,6 @@ Authors: Chris Hughes -/ import Mathlib.Algebra.Module.Projective import Mathlib.FieldTheory.Finiteness -import Mathlib.RingTheory.Finiteness.Subalgebra /-! # Finite dimensional vector spaces @@ -760,10 +759,6 @@ instance FiniteDimensional.finiteDimensional_subalgebra [FiniteDimensional F E] (S : Subalgebra F E) : FiniteDimensional F S := FiniteDimensional.of_subalgebra_toSubmodule inferInstance -@[deprecated Subalgebra.finite_bot (since := "2024-04-11")] -theorem Subalgebra.finiteDimensional_bot : FiniteDimensional F (⊥ : Subalgebra F E) := - Subalgebra.finite_bot - end SubalgebraRank namespace Module diff --git a/Mathlib/LinearAlgebra/Matrix/ZPow.lean b/Mathlib/LinearAlgebra/Matrix/ZPow.lean index 3535d8a612953..330e5bd528012 100644 --- a/Mathlib/LinearAlgebra/Matrix/ZPow.lean +++ b/Mathlib/LinearAlgebra/Matrix/ZPow.lean @@ -101,8 +101,6 @@ theorem zpow_neg_natCast (A : M) (n : ℕ) : A ^ (-n : ℤ) = (A ^ n)⁻¹ := by · simp · exact DivInvMonoid.zpow_neg' _ _ -@[deprecated (since := "2024-04-05")] alias zpow_neg_coe_nat := zpow_neg_natCast - theorem _root_.IsUnit.det_zpow {A : M} (h : IsUnit A.det) (n : ℤ) : IsUnit (A ^ n).det := by cases' n with n n · simpa using h.pow n diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 2f6e9cd67d6fb..a1317983523cb 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -5,8 +5,6 @@ Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov, Eric W -/ import Mathlib.Algebra.Algebra.Prod import Mathlib.Algebra.Group.Graph -import Mathlib.Algebra.Order.PartialSups -import Mathlib.Data.Nat.SuccPred import Mathlib.LinearAlgebra.Span.Basic /-! ### Products of modules @@ -825,132 +823,6 @@ later, when we assume `M` is noetherian, this implies that `N` must be trivial, and establishes the strong rank condition for any left-noetherian ring. -/ - -noncomputable section Tunnel - --- (This doesn't work over a semiring: we need to use that `Submodule R M` is a modular lattice, --- which requires cancellation.) -variable [Ring R] -variable {N : Type*} [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] - -open Function - -set_option linter.deprecated false in -/-- An auxiliary construction for `tunnel`. -The composition of `f`, followed by the isomorphism back to `K`, -followed by the inclusion of this submodule back into `M`. -/ -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -def tunnelAux (f : M × N →ₗ[R] M) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : M × N →ₗ[R] M := - (Kφ.1.subtype.comp Kφ.2.symm.toLinearMap).comp f - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -theorem tunnelAux_injective (f : M × N →ₗ[R] M) (i : Injective f) - (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : Injective (tunnelAux f Kφ) := - (Subtype.val_injective.comp Kφ.2.symm.injective).comp i - -set_option linter.deprecated false in -/-- Auxiliary definition for `tunnel`. -/ -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule R M, K ≃ₗ[R] M - | 0 => ⟨⊤, LinearEquiv.ofTop ⊤ rfl⟩ - | n + 1 => - ⟨(Submodule.fst R M N).map (tunnelAux f (tunnel' f i n)), - ((Submodule.fst R M N).equivMapOfInjective _ - (tunnelAux_injective f i (tunnel' f i n))).symm.trans (Submodule.fstEquiv R M N)⟩ - -set_option linter.deprecated false in -/-- Give an injective map `f : M × N →ₗ[R] M` we can find a nested sequence of submodules -all isomorphic to `M`. --/ -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -def tunnel (f : M × N →ₗ[R] M) (i : Injective f) : ℕ →o (Submodule R M)ᵒᵈ := - -- Note: the hint `(α := _)` had to be added in https://github.com/leanprover-community/mathlib4/pull/8386 - ⟨fun n => OrderDual.toDual (α := Submodule R M) (tunnel' f i n).1, - monotone_nat_of_le_succ fun n => by - dsimp [tunnel', tunnelAux] - rw [Submodule.map_comp, Submodule.map_comp] - apply Submodule.map_subtype_le⟩ - -set_option linter.deprecated false in -/-- Give an injective map `f : M × N →ₗ[R] M` we can find a sequence of submodules -all isomorphic to `N`. --/ -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -def tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Submodule R M := - (Submodule.snd R M N).map (tunnelAux f (tunnel' f i n)) - -set_option linter.deprecated false in -/-- Each `tailing f i n` is a copy of `N`. -/ -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -def tailingLinearEquiv (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≃ₗ[R] N := - ((Submodule.snd R M N).equivMapOfInjective _ (tunnelAux_injective f i (tunnel' f i n))).symm.trans - (Submodule.sndEquiv R M N) - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -theorem tailing_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : - tailing f i n ≤ OrderDual.ofDual (α := Submodule R M) (tunnel f i n) := by - dsimp [tailing, tunnelAux] - rw [Submodule.map_comp, Submodule.map_comp] - apply Submodule.map_subtype_le - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : - Disjoint (tailing f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by - rw [disjoint_iff] - dsimp [tailing, tunnel, tunnel'] - rw [Submodule.map_inf_eq_map_inf_comap, - Submodule.comap_map_eq_of_injective (tunnelAux_injective _ i _), inf_comm, - Submodule.fst_inf_snd, Submodule.map_bot] - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : - tailing f i n ⊔ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) ≤ - (OrderDual.ofDual (α := Submodule R M) <| tunnel f i n) := by - dsimp [tailing, tunnel, tunnel', tunnelAux] - rw [← Submodule.map_sup, sup_comm, Submodule.fst_sup_snd, Submodule.map_comp, Submodule.map_comp] - apply Submodule.map_subtype_le - -set_option linter.deprecated false in -/-- The supremum of all the copies of `N` found inside the tunnel. -/ -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -def tailings (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → Submodule R M := - partialSups (tailing f i) - -set_option linter.deprecated false in -@[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] -theorem tailings_zero (f : M × N →ₗ[R] M) (i : Injective f) : tailings f i 0 = tailing f i 0 := by - simp [tailings] - -set_option linter.deprecated false in -@[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] -theorem tailings_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : - tailings f i (n + 1) = tailings f i n ⊔ tailing f i (n + 1) := by simp [tailings] - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : - Disjoint (tailings f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by - induction' n with n ih - · simp only [tailings_zero] - apply tailing_disjoint_tunnel_succ - · simp only [tailings_succ] - refine Disjoint.disjoint_sup_left_of_disjoint_sup_right ?_ ?_ - · apply tailing_disjoint_tunnel_succ - · apply Disjoint.mono_right _ ih - apply tailing_sup_tunnel_succ_le_tunnel - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-06-05")] -theorem tailings_disjoint_tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : - Disjoint (tailings f i n) (tailing f i (n + 1)) := - Disjoint.mono_right (tailing_le_tunnel f i _) (tailings_disjoint_tunnel f i _) - -end Tunnel - section Graph variable [Semiring R] [AddCommMonoid M] [AddCommMonoid M₂] [AddCommGroup M₃] [AddCommGroup M₄] diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index a5d92199fe79a..7d45608bb2ad6 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -720,14 +720,6 @@ noncomputable def existsCases {α C : Sort*} {p : α → Prop} (H0 : C) (H : ∀ theorem some_spec₂ {α : Sort*} {p : α → Prop} {h : ∃ a, p a} (q : α → Prop) (hpq : ∀ a, p a → q a) : q (choose h) := hpq _ <| choose_spec _ -/-- A version of `Classical.indefiniteDescription` which is definitionally equal to a pair. - -In Lean 4, this definition is defeq to `Classical.indefiniteDescription`, -so it is deprecated. -/ -@[deprecated Classical.indefiniteDescription (since := "2024-07-04")] -noncomputable def subtype_of_exists {α : Type*} {P : α → Prop} (h : ∃ x, P x) : { x // P x } := - ⟨Classical.choose h, Classical.choose_spec h⟩ - /-- A version of `byContradiction` that uses types instead of propositions. -/ protected noncomputable def byContradiction' {α : Sort*} (H : ¬(α → False)) : α := Classical.choice <| (peirce _ False) fun h ↦ (H fun a ↦ h ⟨a⟩).elim @@ -752,23 +744,6 @@ alias by_contradiction := byContradiction -- TODO: remove? rename in core? alias prop_complete := propComplete -- TODO: remove? rename in core? -@[elab_as_elim, deprecated "No deprecation message was provided." (since := "2024-07-27")] -theorem cases_true_false (p : Prop → Prop) - (h1 : p True) (h2 : p False) (a : Prop) : p a := - Or.elim (prop_complete a) (fun ht : a = True ↦ ht.symm ▸ h1) fun hf : a = False ↦ hf.symm ▸ h2 - -@[deprecated "No deprecation message was provided." (since := "2024-07-27")] -theorem eq_false_or_eq_true (a : Prop) : a = False ∨ a = True := (prop_complete a).symm - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-07-27")] -theorem cases_on (a : Prop) {p : Prop → Prop} (h1 : p True) (h2 : p False) : p a := - @cases_true_false p h1 h2 a - -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-07-27")] -theorem cases {p : Prop → Prop} (h1 : p True) (h2 : p False) (a) : p a := cases_on a h1 h2 - end Classical /-- This function has the same type as `Exists.recOn`, and can be used to case on an equality, @@ -792,13 +767,6 @@ theorem BEx.elim {b : Prop} : (∃ x h, P x h) → (∀ a h, P a h → b) → b theorem BEx.intro (a : α) (h₁ : p a) (h₂ : P a h₁) : ∃ (x : _) (h : p x), P x h := ⟨a, h₁, h₂⟩ -@[deprecated exists_eq_left (since := "2024-04-06")] -theorem bex_eq_left {a : α} : (∃ (x : _) (_ : x = a), p x) ↔ p a := by - simp only [exists_prop, exists_eq_left] - -@[deprecated (since := "2024-04-06")] alias ball_congr := forall₂_congr -@[deprecated (since := "2024-04-06")] alias bex_congr := exists₂_congr - theorem BAll.imp_right (H : ∀ x h, P x h → Q x h) (h₁ : ∀ x h, P x h) (x h) : Q x h := H _ _ <| h₁ _ _ diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index c1382615c0f6e..b3ca33b7c9ca3 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -721,8 +721,6 @@ protected lemma forall_congr_right : (∀ a, q (e a)) ↔ ∀ b, q b := protected lemma forall_congr_left : (∀ a, p a) ↔ ∀ b, p (e.symm b) := e.symm.forall_congr_right.symm -@[deprecated (since := "2024-06-11")] alias forall_congr_left' := Equiv.forall_congr_left - protected lemma forall_congr (h : ∀ a, p a ↔ q (e a)) : (∀ a, p a) ↔ ∀ b, q b := e.forall_congr_left.trans (by simp [h]) @@ -747,17 +745,9 @@ protected lemma existsUnique_congr_right : (∃! a, q (e a)) ↔ ∃! b, q b := protected lemma existsUnique_congr_left : (∃! a, p a) ↔ ∃! b, p (e.symm b) := e.symm.existsUnique_congr_right.symm -@[deprecated (since := "2024-06-11")] -alias exists_unique_congr_left := Equiv.existsUnique_congr_right - -@[deprecated (since := "2024-06-11")] -alias exists_unique_congr_left' := Equiv.existsUnique_congr_left - protected lemma existsUnique_congr (h : ∀ a, p a ↔ q (e a)) : (∃! a, p a) ↔ ∃! b, q b := e.existsUnique_congr_left.trans <| by simp [h] -@[deprecated (since := "2024-06-11")] alias exists_unique_congr := Equiv.existsUnique_congr - protected lemma existsUnique_congr' (h : ∀ b, p (e.symm b) ↔ q b) : (∃! a, p a) ↔ ∃! b, q b := e.existsUnique_congr_left.trans <| by simp [h] diff --git a/Mathlib/Logic/Equiv/Fin.lean b/Mathlib/Logic/Equiv/Fin.lean index a8dc54a1bbf67..c3622dc3f66bf 100644 --- a/Mathlib/Logic/Equiv/Fin.lean +++ b/Mathlib/Logic/Equiv/Fin.lean @@ -179,20 +179,6 @@ theorem finSuccEquivLast_symm_some (i : Fin n) : @[simp] theorem finSuccEquivLast_symm_none : finSuccEquivLast.symm none = Fin.last n := finSuccEquiv'_symm_none _ -/-- Equivalence between `Π j : Fin (n + 1), α j` and `α i × Π j : Fin n, α (Fin.succAbove i j)`. -/ -@[simps (config := .asFn), deprecated Fin.insertNthEquiv (since := "2024-07-12")] -def Equiv.piFinSuccAbove (α : Fin (n + 1) → Type u) (i : Fin (n + 1)) : - (∀ j, α j) ≃ α i × ∀ j, α (i.succAbove j) where - toFun f := (f i, i.removeNth f) - invFun f := i.insertNth f.1 f.2 - left_inv f := by simp - right_inv f := by simp - -/-- Equivalence between `Fin (n + 1) → β` and `β × (Fin n → β)`. -/ -@[simps! (config := .asFn), deprecated Fin.consEquiv (since := "2024-07-12")] -def Equiv.piFinSucc (n : ℕ) (β : Type u) : (Fin (n + 1) → β) ≃ β × (Fin n → β) := - (Fin.insertNthEquiv (fun _ => β) 0).symm - /-- An embedding `e : Fin (n+1) ↪ ι` corresponds to an embedding `f : Fin n ↪ ι` (corresponding the last `n` coordinates of `e`) together with a value not taken by `f` (corresponding to `e 0`). -/ def Equiv.embeddingFinSucc (n : ℕ) (ι : Type*) : @@ -212,12 +198,6 @@ def Equiv.embeddingFinSucc (n : ℕ) (ι : Type*) : ext i exact Fin.cases rfl (fun j ↦ rfl) i -/-- Equivalence between `Fin (n + 1) → β` and `β × (Fin n → β)` which separates out the last -element of the tuple. -/ -@[simps! (config := .asFn), deprecated Fin.snocEquiv (since := "2024-07-12")] -def Equiv.piFinCastSucc (n : ℕ) (β : Type u) : (Fin (n + 1) → β) ≃ β × (Fin n → β) := - (Fin.insertNthEquiv (fun _ => β) (.last _)).symm - /-- Equivalence between `Fin m ⊕ Fin n` and `Fin (m + n)` -/ def finSumFinEquiv : Fin m ⊕ Fin n ≃ Fin (m + n) where toFun := Sum.elim (Fin.castAdd n) (Fin.natAdd m) diff --git a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean index f8f7b82e72071..b793c87ee48e0 100644 --- a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean +++ b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean @@ -302,9 +302,6 @@ lemma setLIntegral_rnDeriv_le (s : Set α) : ∫⁻ x in s, μ.rnDeriv ν x ∂ν ≤ μ s := (withDensity_apply_le _ _).trans (Measure.le_iff'.1 (withDensity_rnDeriv_le μ ν) s) -@[deprecated (since := "2024-06-29")] -alias set_lintegral_rnDeriv_le := setLIntegral_rnDeriv_le - lemma lintegral_rnDeriv_le : ∫⁻ x, μ.rnDeriv ν x ∂ν ≤ μ Set.univ := (setLIntegral_univ _).symm ▸ Measure.setLIntegral_rnDeriv_le Set.univ @@ -313,17 +310,11 @@ lemma setLIntegral_rnDeriv' [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ ν ∫⁻ x in s, μ.rnDeriv ν x ∂ν = μ s := by rw [← withDensity_apply _ hs, Measure.withDensity_rnDeriv_eq _ _ hμν] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_rnDeriv' := setLIntegral_rnDeriv' - lemma setLIntegral_rnDeriv [HaveLebesgueDecomposition μ ν] [SFinite ν] (hμν : μ ≪ ν) (s : Set α) : ∫⁻ x in s, μ.rnDeriv ν x ∂ν = μ s := by rw [← withDensity_apply' _ s, Measure.withDensity_rnDeriv_eq _ _ hμν] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_rnDeriv := setLIntegral_rnDeriv - lemma lintegral_rnDeriv [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ ν) : ∫⁻ x, μ.rnDeriv ν x ∂ν = μ Set.univ := by rw [← setLIntegral_univ, setLIntegral_rnDeriv' hμν MeasurableSet.univ] @@ -341,9 +332,6 @@ lemma setIntegral_toReal_rnDeriv_eq_withDensity' [SigmaFinite μ] simp · exact ae_restrict_of_ae (Measure.rnDeriv_lt_top _ _) -@[deprecated (since := "2024-04-17")] -alias set_integral_toReal_rnDeriv_eq_withDensity' := setIntegral_toReal_rnDeriv_eq_withDensity' - lemma setIntegral_toReal_rnDeriv_eq_withDensity [SigmaFinite μ] [SFinite ν] (s : Set α) : ∫ x in s, (μ.rnDeriv ν x).toReal ∂ν = (ν.withDensity (μ.rnDeriv ν) s).toReal := by rw [integral_toReal (Measure.measurable_rnDeriv _ _).aemeasurable] @@ -351,9 +339,6 @@ lemma setIntegral_toReal_rnDeriv_eq_withDensity [SigmaFinite μ] [SFinite ν] (s simp · exact ae_restrict_of_ae (Measure.rnDeriv_lt_top _ _) -@[deprecated (since := "2024-04-17")] -alias set_integral_toReal_rnDeriv_eq_withDensity := setIntegral_toReal_rnDeriv_eq_withDensity - lemma setIntegral_toReal_rnDeriv_le [SigmaFinite μ] {s : Set α} (hμs : μ s ≠ ∞) : ∫ x in s, (μ.rnDeriv ν x).toReal ∂ν ≤ (μ s).toReal := by set t := toMeasurable μ s with ht @@ -371,24 +356,15 @@ lemma setIntegral_toReal_rnDeriv_le [SigmaFinite μ] {s : Set α} (hμs : μ s · apply withDensity_rnDeriv_le _ = (μ s).toReal := by rw [measure_toMeasurable s] -@[deprecated (since := "2024-04-17")] -alias set_integral_toReal_rnDeriv_le := setIntegral_toReal_rnDeriv_le - lemma setIntegral_toReal_rnDeriv' [SigmaFinite μ] [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ ν) {s : Set α} (hs : MeasurableSet s) : ∫ x in s, (μ.rnDeriv ν x).toReal ∂ν = (μ s).toReal := by rw [setIntegral_toReal_rnDeriv_eq_withDensity' hs, Measure.withDensity_rnDeriv_eq _ _ hμν] -@[deprecated (since := "2024-04-17")] -alias set_integral_toReal_rnDeriv' := setIntegral_toReal_rnDeriv' - lemma setIntegral_toReal_rnDeriv [SigmaFinite μ] [SigmaFinite ν] (hμν : μ ≪ ν) (s : Set α) : ∫ x in s, (μ.rnDeriv ν x).toReal ∂ν = (μ s).toReal := by rw [setIntegral_toReal_rnDeriv_eq_withDensity s, Measure.withDensity_rnDeriv_eq _ _ hμν] -@[deprecated (since := "2024-04-17")] -alias set_integral_toReal_rnDeriv := setIntegral_toReal_rnDeriv - lemma integral_toReal_rnDeriv [SigmaFinite μ] [SigmaFinite ν] (hμν : μ ≪ ν) : ∫ x, (μ.rnDeriv ν x).toReal ∂ν = (μ Set.univ).toReal := by rw [← setIntegral_univ, setIntegral_toReal_rnDeriv hμν Set.univ] @@ -544,9 +520,6 @@ lemma setLIntegral_rnDeriv_mul [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ rw [setLIntegral_withDensity_eq_lintegral_mul₀ (measurable_rnDeriv μ ν).aemeasurable hf hs] simp only [Pi.mul_apply] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_rnDeriv_mul := setLIntegral_rnDeriv_mul - variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [HaveLebesgueDecomposition μ ν] [SigmaFinite μ] {f : α → E} diff --git a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean index 99e6748be5ef4..102b8db19c4b6 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean @@ -143,18 +143,6 @@ theorem ae_nonneg_of_forall_setIntegral_nonneg (hf : Integrable f μ) rw [← μ.restrict_toMeasurable mus.ne] exact hf_zero _ (measurableSet_toMeasurable ..) (by rwa [measure_toMeasurable]) -@[deprecated (since := "2024-04-17")] -alias ae_nonneg_of_forall_set_integral_nonneg_of_stronglyMeasurable := - ae_nonneg_of_forall_setIntegral_nonneg - -@[deprecated (since := "2024-07-12")] -alias ae_nonneg_of_forall_setIntegral_nonneg_of_stronglyMeasurable := - ae_nonneg_of_forall_setIntegral_nonneg - -@[deprecated (since := "2024-04-17")] -alias ae_nonneg_of_forall_set_integral_nonneg := - ae_nonneg_of_forall_setIntegral_nonneg - theorem ae_le_of_forall_setIntegral_le {f g : α → ℝ} (hf : Integrable f μ) (hg : Integrable g μ) (hf_le : ∀ s, MeasurableSet s → μ s < ∞ → (∫ x in s, f x ∂μ) ≤ ∫ x in s, g x ∂μ) : f ≤ᵐ[μ] g := by @@ -163,9 +151,6 @@ theorem ae_le_of_forall_setIntegral_le {f g : α → ℝ} (hf : Integrable f μ) rw [integral_sub' hg.integrableOn hf.integrableOn, sub_nonneg] exact hf_le s hs -@[deprecated (since := "2024-04-17")] -alias ae_le_of_forall_set_integral_le := ae_le_of_forall_setIntegral_le - theorem ae_nonneg_restrict_of_forall_setIntegral_nonneg_inter {f : α → ℝ} {t : Set α} (hf : IntegrableOn f t μ) (hf_zero : ∀ s, MeasurableSet s → μ (s ∩ t) < ∞ → 0 ≤ ∫ x in s ∩ t, f x ∂μ) : @@ -175,10 +160,6 @@ theorem ae_nonneg_restrict_of_forall_setIntegral_nonneg_inter {f : α → ℝ} { apply hf_zero s hs rwa [Measure.restrict_apply hs] at h's -@[deprecated (since := "2024-04-17")] -alias ae_nonneg_restrict_of_forall_set_integral_nonneg_inter := - ae_nonneg_restrict_of_forall_setIntegral_nonneg_inter - theorem ae_nonneg_of_forall_setIntegral_nonneg_of_sigmaFinite [SigmaFinite μ] {f : α → ℝ} (hf_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn f s μ) (hf_zero : ∀ s, MeasurableSet s → μ s < ∞ → 0 ≤ ∫ x in s, f x ∂μ) : 0 ≤ᵐ[μ] f := by @@ -190,10 +171,6 @@ theorem ae_nonneg_of_forall_setIntegral_nonneg_of_sigmaFinite [SigmaFinite μ] { hf_zero _ (s_meas.inter t_meas) (lt_of_le_of_lt (measure_mono (Set.inter_subset_right)) t_lt_top) -@[deprecated (since := "2024-04-17")] -alias ae_nonneg_of_forall_set_integral_nonneg_of_sigmaFinite := - ae_nonneg_of_forall_setIntegral_nonneg_of_sigmaFinite - theorem AEFinStronglyMeasurable.ae_nonneg_of_forall_setIntegral_nonneg {f : α → ℝ} (hf : AEFinStronglyMeasurable f μ) (hf_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn f s μ) @@ -211,10 +188,6 @@ theorem AEFinStronglyMeasurable.ae_nonneg_of_forall_setIntegral_nonneg {f : α rw [Measure.restrict_apply hs] at hμts exact hf_zero (s ∩ t) (hs.inter hf.measurableSet) hμts -@[deprecated (since := "2024-04-17")] -alias AEFinStronglyMeasurable.ae_nonneg_of_forall_set_integral_nonneg := - AEFinStronglyMeasurable.ae_nonneg_of_forall_setIntegral_nonneg - theorem ae_nonneg_restrict_of_forall_setIntegral_nonneg {f : α → ℝ} (hf_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn f s μ) (hf_zero : ∀ s, MeasurableSet s → μ s < ∞ → 0 ≤ ∫ x in s, f x ∂μ) {t : Set α} @@ -225,10 +198,6 @@ theorem ae_nonneg_restrict_of_forall_setIntegral_nonneg {f : α → ℝ} refine hf_zero (s ∩ t) (hs.inter ht) ?_ exact (measure_mono Set.inter_subset_right).trans_lt (lt_top_iff_ne_top.mpr hμt) -@[deprecated (since := "2024-04-17")] -alias ae_nonneg_restrict_of_forall_set_integral_nonneg := - ae_nonneg_restrict_of_forall_setIntegral_nonneg - theorem ae_eq_zero_restrict_of_forall_setIntegral_eq_zero_real {f : α → ℝ} (hf_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn f s μ) (hf_zero : ∀ s, MeasurableSet s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) {t : Set α} @@ -250,10 +219,6 @@ theorem ae_eq_zero_restrict_of_forall_setIntegral_eq_zero_real {f : α → ℝ} rw [integral_neg, neg_nonneg] exact (hf_zero s hs hμs).le -@[deprecated (since := "2024-04-17")] -alias ae_eq_zero_restrict_of_forall_set_integral_eq_zero_real := - ae_eq_zero_restrict_of_forall_setIntegral_eq_zero_real - end Real theorem ae_eq_zero_restrict_of_forall_setIntegral_eq_zero {f : α → E} @@ -270,10 +235,6 @@ theorem ae_eq_zero_restrict_of_forall_setIntegral_eq_zero {f : α → E} rw [ContinuousLinearMap.integral_comp_comm c (hf_int_finite s hs hμs), hf_zero s hs hμs] exact ContinuousLinearMap.map_zero _ -@[deprecated (since := "2024-04-17")] -alias ae_eq_zero_restrict_of_forall_set_integral_eq_zero := - ae_eq_zero_restrict_of_forall_setIntegral_eq_zero - theorem ae_eq_restrict_of_forall_setIntegral_eq {f g : α → E} (hf_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn f s μ) (hg_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn g s μ) @@ -288,10 +249,6 @@ theorem ae_eq_restrict_of_forall_setIntegral_eq {f g : α → E} (hf_int_finite s hs hμs).sub (hg_int_finite s hs hμs) exact ae_eq_zero_restrict_of_forall_setIntegral_eq_zero hfg_int hfg' ht hμt -@[deprecated (since := "2024-04-17")] -alias ae_eq_restrict_of_forall_set_integral_eq := - ae_eq_restrict_of_forall_setIntegral_eq - theorem ae_eq_zero_of_forall_setIntegral_eq_of_sigmaFinite [SigmaFinite μ] {f : α → E} (hf_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn f s μ) (hf_zero : ∀ s : Set α, MeasurableSet s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) : f =ᵐ[μ] 0 := by @@ -305,10 +262,6 @@ theorem ae_eq_zero_of_forall_setIntegral_eq_of_sigmaFinite [SigmaFinite μ] {f : rw [← Measure.restrict_apply' h_meas_n] exact ae_eq_zero_restrict_of_forall_setIntegral_eq_zero hf_int_finite hf_zero h_meas_n hμn.ne -@[deprecated (since := "2024-04-17")] -alias ae_eq_zero_of_forall_set_integral_eq_of_sigmaFinite := - ae_eq_zero_of_forall_setIntegral_eq_of_sigmaFinite - theorem ae_eq_of_forall_setIntegral_eq_of_sigmaFinite [SigmaFinite μ] {f g : α → E} (hf_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn f s μ) (hg_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn g s μ) @@ -323,10 +276,6 @@ theorem ae_eq_of_forall_setIntegral_eq_of_sigmaFinite [SigmaFinite μ] {f g : α (hf_int_finite s hs hμs).sub (hg_int_finite s hs hμs) exact ae_eq_zero_of_forall_setIntegral_eq_of_sigmaFinite hfg_int hfg -@[deprecated (since := "2024-04-17")] -alias ae_eq_of_forall_set_integral_eq_of_sigmaFinite := - ae_eq_of_forall_setIntegral_eq_of_sigmaFinite - theorem AEFinStronglyMeasurable.ae_eq_zero_of_forall_setIntegral_eq_zero {f : α → E} (hf_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn f s μ) (hf_zero : ∀ s : Set α, MeasurableSet s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) @@ -345,10 +294,6 @@ theorem AEFinStronglyMeasurable.ae_eq_zero_of_forall_setIntegral_eq_zero {f : α rw [Measure.restrict_apply hs] at hμs exact hf_zero _ (hs.inter hf.measurableSet) hμs -@[deprecated (since := "2024-04-17")] -alias AEFinStronglyMeasurable.ae_eq_zero_of_forall_set_integral_eq_zero := - AEFinStronglyMeasurable.ae_eq_zero_of_forall_setIntegral_eq_zero - theorem AEFinStronglyMeasurable.ae_eq_of_forall_setIntegral_eq {f g : α → E} (hf_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn f s μ) (hg_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn g s μ) @@ -363,20 +308,12 @@ theorem AEFinStronglyMeasurable.ae_eq_of_forall_setIntegral_eq {f g : α → E} (hf_int_finite s hs hμs).sub (hg_int_finite s hs hμs) exact (hf.sub hg).ae_eq_zero_of_forall_setIntegral_eq_zero hfg_int hfg -@[deprecated (since := "2024-04-17")] -alias AEFinStronglyMeasurable.ae_eq_of_forall_set_integral_eq := - AEFinStronglyMeasurable.ae_eq_of_forall_setIntegral_eq - theorem Lp.ae_eq_zero_of_forall_setIntegral_eq_zero (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn f s μ) (hf_zero : ∀ s : Set α, MeasurableSet s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) : f =ᵐ[μ] 0 := AEFinStronglyMeasurable.ae_eq_zero_of_forall_setIntegral_eq_zero hf_int_finite hf_zero (Lp.finStronglyMeasurable _ hp_ne_zero hp_ne_top).aefinStronglyMeasurable -@[deprecated (since := "2024-04-17")] -alias Lp.ae_eq_zero_of_forall_set_integral_eq_zero := - Lp.ae_eq_zero_of_forall_setIntegral_eq_zero - theorem Lp.ae_eq_of_forall_setIntegral_eq (f g : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn f s μ) (hg_int_finite : ∀ s, MeasurableSet s → μ s < ∞ → IntegrableOn g s μ) @@ -386,9 +323,6 @@ theorem Lp.ae_eq_of_forall_setIntegral_eq (f g : Lp E p μ) (hp_ne_zero : p ≠ (Lp.finStronglyMeasurable _ hp_ne_zero hp_ne_top).aefinStronglyMeasurable (Lp.finStronglyMeasurable _ hp_ne_zero hp_ne_top).aefinStronglyMeasurable -@[deprecated (since := "2024-04-17")] -alias Lp.ae_eq_of_forall_set_integral_eq := Lp.ae_eq_of_forall_setIntegral_eq - theorem ae_eq_zero_of_forall_setIntegral_eq_of_finStronglyMeasurable_trim (hm : m ≤ m0) {f : α → E} (hf_int_finite : ∀ s, MeasurableSet[m] s → μ s < ∞ → IntegrableOn f s μ) (hf_zero : ∀ s : Set α, MeasurableSet[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) @@ -417,10 +351,6 @@ theorem ae_eq_zero_of_forall_setIntegral_eq_of_finStronglyMeasurable_trim (hm : rw [← integral_trim hm hf_meas_m] exact hf_zero _ (hs.inter ht_meas) hμs -@[deprecated (since := "2024-04-17")] -alias ae_eq_zero_of_forall_set_integral_eq_of_finStronglyMeasurable_trim := - ae_eq_zero_of_forall_setIntegral_eq_of_finStronglyMeasurable_trim - theorem Integrable.ae_eq_zero_of_forall_setIntegral_eq_zero {f : α → E} (hf : Integrable f μ) (hf_zero : ∀ s, MeasurableSet s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) : f =ᵐ[μ] 0 := by have hf_Lp : Memℒp f 1 μ := memℒp_one_iff_integrable.mpr hf @@ -433,10 +363,6 @@ theorem Integrable.ae_eq_zero_of_forall_setIntegral_eq_zero {f : α → E} (hf : rw [integral_congr_ae (ae_restrict_of_ae hf_f_Lp.symm)] exact hf_zero s hs hμs -@[deprecated (since := "2024-04-17")] -alias Integrable.ae_eq_zero_of_forall_set_integral_eq_zero := - Integrable.ae_eq_zero_of_forall_setIntegral_eq_zero - theorem Integrable.ae_eq_of_forall_setIntegral_eq (f g : α → E) (hf : Integrable f μ) (hg : Integrable g μ) (hfg : ∀ s : Set α, MeasurableSet s → μ s < ∞ → ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ) : @@ -448,10 +374,6 @@ theorem Integrable.ae_eq_of_forall_setIntegral_eq (f g : α → E) (hf : Integra exact sub_eq_zero.mpr (hfg s hs hμs) exact Integrable.ae_eq_zero_of_forall_setIntegral_eq_zero (hf.sub hg) hfg' -@[deprecated (since := "2024-04-17")] -alias Integrable.ae_eq_of_forall_set_integral_eq := - Integrable.ae_eq_of_forall_setIntegral_eq - variable {β : Type*} [TopologicalSpace β] [MeasurableSpace β] [BorelSpace β] /-- If an integrable function has zero integral on all closed sets, then it is zero @@ -472,10 +394,6 @@ lemma ae_eq_zero_of_forall_setIntegral_isClosed_eq_zero {μ : Measure β} {f : | compl s hs ihs => exact A s hs ihs | iUnion g g_disj g_meas hg => simp [integral_iUnion g_meas g_disj hf.integrableOn, hg] -@[deprecated (since := "2024-04-17")] -alias ae_eq_zero_of_forall_set_integral_isClosed_eq_zero := - ae_eq_zero_of_forall_setIntegral_isClosed_eq_zero - /-- If an integrable function has zero integral on all compact sets in a sigma-compact space, then it is zero almost everywhere. -/ lemma ae_eq_zero_of_forall_setIntegral_isCompact_eq_zero diff --git a/Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean index 41ac33f231c08..e3efebeaaeaf0 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean @@ -125,19 +125,11 @@ theorem ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ [SigmaFinite μ] _ ≤ ∑' n, μ (s n) := measure_iUnion_le _ _ = 0 := by simp only [μs, tsum_zero] -@[deprecated (since := "2024-06-29")] -alias ae_le_of_forall_set_lintegral_le_of_sigmaFinite₀ := - ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ - theorem ae_le_of_forall_setLIntegral_le_of_sigmaFinite [SigmaFinite μ] {f g : α → ℝ≥0∞} (hf : Measurable f) (h : ∀ s, MeasurableSet s → μ s < ∞ → (∫⁻ x in s, f x ∂μ) ≤ ∫⁻ x in s, g x ∂μ) : f ≤ᵐ[μ] g := ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ hf.aemeasurable h -@[deprecated (since := "2024-06-29")] -alias ae_le_of_forall_set_lintegral_le_of_sigmaFinite := - ae_le_of_forall_setLIntegral_le_of_sigmaFinite - theorem ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ [SigmaFinite μ] {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) (hg : AEMeasurable g μ) (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : f =ᵐ[μ] g := by @@ -147,19 +139,11 @@ theorem ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ [SigmaFinite μ] ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ hg fun s hs h's => ge_of_eq (h s hs h's) filter_upwards [A, B] with x using le_antisymm -@[deprecated (since := "2024-06-29")] -alias ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite₀ := - ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ - theorem ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite [SigmaFinite μ] {f g : α → ℝ≥0∞} (hf : Measurable f) (hg : Measurable g) (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : f =ᵐ[μ] g := ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ hf.aemeasurable hg.aemeasurable h -@[deprecated (since := "2024-06-29")] -alias ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite := - ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite - theorem AEMeasurable.ae_eq_of_forall_setLIntegral_eq {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) (hg : AEMeasurable g μ) (hfi : ∫⁻ x, f x ∂μ ≠ ∞) (hgi : ∫⁻ x, g x ∂μ ≠ ∞) (hfg : ∀ ⦃s⦄, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : @@ -188,9 +172,6 @@ theorem AEMeasurable.ae_eq_of_forall_setLIntegral_eq {f g : α → ℝ≥0∞} ( rw [Measure.restrict_apply hu] at huμ exact hfg (hu.inter (hf'.measurableSet.union hg'.measurableSet)) huμ -@[deprecated (since := "2024-06-29")] -alias AEMeasurable.ae_eq_of_forall_set_lintegral_eq := AEMeasurable.ae_eq_of_forall_setLIntegral_eq - section WithDensity variable {m : MeasurableSpace α} {μ : Measure α} diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean index 2c2f035853d21..c41713785c0eb 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean @@ -231,8 +231,6 @@ theorem setIntegral_condExp (hm : m ≤ m₀) [SigmaFinite (μ.trim hm)] (hf : I @[deprecated (since := "2025-01-21")] alias setIntegral_condexp := setIntegral_condExp -@[deprecated (since := "2024-04-17")] alias set_integral_condexp := setIntegral_condExp - theorem integral_condExp (hm : m ≤ m₀) [hμm : SigmaFinite (μ.trim hm)] : ∫ x, (μ[f|m]) x ∂μ = ∫ x, f x ∂μ := by by_cases hf : Integrable f μ @@ -267,12 +265,6 @@ theorem ae_eq_condExp_of_forall_setIntegral_eq (hm : m ≤ m₀) [SigmaFinite ( @[deprecated (since := "2025-01-21")] alias ae_eq_condexp_of_forall_setIntegral_eq := ae_eq_condExp_of_forall_setIntegral_eq -@[deprecated (since := "2024-04-17")] -alias ae_eq_condExp_of_forall_set_integral_eq := ae_eq_condExp_of_forall_setIntegral_eq - -@[deprecated (since := "2025-01-21")] -alias ae_eq_condexp_of_forall_set_integral_eq := ae_eq_condExp_of_forall_set_integral_eq - theorem condExp_bot' [hμ : NeZero μ] (f : α → E) : μ[f|⊥] = fun _ => (μ Set.univ).toReal⁻¹ • ∫ x, f x ∂μ := by by_cases hμ_finite : IsFiniteMeasure μ diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean index 5bb82053bb176..1052ad5648143 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean @@ -375,11 +375,6 @@ theorem setIntegral_condExpInd (hs : MeasurableSet[m] s) (ht : MeasurableSet t) @[deprecated (since := "2025-01-21")] alias setIntegral_condexpInd := setIntegral_condExpInd -@[deprecated (since := "2024-04-17")] -alias set_integral_condExpInd := setIntegral_condExpInd - -@[deprecated (since := "2025-01-21")] alias set_integral_condexpInd := set_integral_condExpInd - theorem condExpInd_of_measurable (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) (c : G) : condExpInd G hm μ s c = indicatorConstLp 1 (hm s hs) hμs c := by ext1 @@ -470,13 +465,6 @@ theorem setIntegral_condExpL1CLM_of_measure_ne_top (f : α →₁[μ] F') (hs : @[deprecated (since := "2025-01-21")] alias setIntegral_condexpL1CLM_of_measure_ne_top := setIntegral_condExpL1CLM_of_measure_ne_top -@[deprecated (since := "2024-04-17")] -alias set_integral_condexpL1CLM_of_measure_ne_top := - setIntegral_condExpL1CLM_of_measure_ne_top - -@[deprecated (since := "2025-01-21")] -alias setIntegral_condexpL1CLM := set_integral_condexpL1CLM_of_measure_ne_top - /-- The integral of the conditional expectation `condExpL1CLM` over an `m`-measurable set is equal to the integral of `f` on that set. See also `setIntegral_condExp`, the similar statement for `condExp`. -/ @@ -514,11 +502,6 @@ theorem setIntegral_condExpL1CLM (f : α →₁[μ] F') (hs : MeasurableSet[m] s rw [h_eq_forall] at h_left exact tendsto_nhds_unique h_left h_right -@[deprecated (since := "2024-04-17")] -alias set_integral_condExpL1CLM := setIntegral_condExpL1CLM - -@[deprecated (since := "2025-01-21")] alias set_integral_condexpL1CLM := set_integral_condExpL1CLM - theorem aestronglyMeasurable_condExpL1CLM (f : α →₁[μ] F') : AEStronglyMeasurable[m] (condExpL1CLM F' hm μ f) μ := by refine @Lp.induction _ _ _ _ _ _ _ ENNReal.one_ne_top @@ -644,11 +627,6 @@ theorem setIntegral_condExpL1 (hf : Integrable f μ) (hs : MeasurableSet[m] s) : @[deprecated (since := "2025-01-21")] alias setIntegral_condexpL1 := setIntegral_condExpL1 -@[deprecated (since := "2024-04-17")] -alias set_integral_condExpL1 := setIntegral_condExpL1 - -@[deprecated (since := "2025-01-21")] alias set_integral_condexpL1 := set_integral_condExpL1 - theorem condExpL1_add (hf : Integrable f μ) (hg : Integrable g μ) : condExpL1 hm μ (f + g) = condExpL1 hm μ f + condExpL1 hm μ g := setToFun_add _ hf hg diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean index c51e17e623243..c1a9c3790e423 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean @@ -120,11 +120,6 @@ theorem eLpNorm_condExpL2_le (hm : m ≤ m0) (f : α →₂[μ] E) : @[deprecated (since := "2025-01-21")] alias eLpNorm_condexpL2_le := eLpNorm_condExpL2_le -@[deprecated (since := "2024-07-27")] -alias snorm_condExpL2_le := eLpNorm_condExpL2_le - -@[deprecated (since := "2025-01-21")] alias snorm_condexpL2_le := snorm_condExpL2_le - theorem norm_condExpL2_coe_le (hm : m ≤ m0) (f : α →₂[μ] E) : ‖(condExpL2 E 𝕜 hm f : α →₂[μ] E)‖ ≤ ‖f‖ := by rw [Lp.norm_def, Lp.norm_def, ← lpMeas_coe] @@ -390,12 +385,6 @@ theorem setLIntegral_nnnorm_condExpL2_indicator_le (hm : m ≤ m0) (hs : Measura @[deprecated (since := "2025-01-21")] alias setLIntegral_nnnorm_condexpL2_indicator_le := setLIntegral_nnnorm_condExpL2_indicator_le -@[deprecated (since := "2024-06-29")] -alias set_lintegral_nnnorm_condExpL2_indicator_le := setLIntegral_nnnorm_condExpL2_indicator_le - -@[deprecated (since := "2025-01-21")] -alias set_lintegral_nnnorm_condexpL2_indicator_le := set_lintegral_nnnorm_condExpL2_indicator_le - theorem lintegral_nnnorm_condExpL2_indicator_le (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : E') [SigmaFinite (μ.trim hm)] : ∫⁻ a, ‖(condExpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) : α → E') a‖₊ ∂μ ≤ μ s * ‖x‖₊ := by @@ -495,12 +484,6 @@ theorem setLIntegral_nnnorm_condExpIndSMul_le (hm : m ≤ m0) (hs : MeasurableSe @[deprecated (since := "2025-01-21")] alias setLIntegral_nnnorm_condexpIndSMul_le := setLIntegral_nnnorm_condExpIndSMul_le -@[deprecated (since := "2024-06-29")] -alias set_lintegral_nnnorm_condExpIndSMul_le := setLIntegral_nnnorm_condExpIndSMul_le - -@[deprecated (since := "2025-01-21")] -alias set_lintegral_nnnorm_condexpIndSMul_le := set_lintegral_nnnorm_condExpIndSMul_le - theorem lintegral_nnnorm_condExpIndSMul_le (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) [SigmaFinite (μ.trim hm)] : ∫⁻ a, ‖condExpIndSMul hm hs hμs x a‖₊ ∂μ ≤ μ s * ‖x‖₊ := by refine lintegral_le_of_forall_fin_meas_trim_le hm (μ s * ‖x‖₊) fun t ht hμt => ?_ @@ -544,12 +527,6 @@ theorem setIntegral_condExpL2_indicator (hs : MeasurableSet[m] s) (ht : Measurab @[deprecated (since := "2025-01-21")] alias setIntegral_condexpL2_indicator := setIntegral_condExpL2_indicator -@[deprecated (since := "2024-04-17")] -alias set_integral_condExpL2_indicator := setIntegral_condExpL2_indicator - -@[deprecated (since := "2025-01-21")] -alias set_integral_condexpL2_indicator := set_integral_condExpL2_indicator - theorem setIntegral_condExpIndSMul (hs : MeasurableSet[m] s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (x : G') : ∫ a in s, (condExpIndSMul hm ht hμt x) a ∂μ = (μ (t ∩ s)).toReal • x := @@ -564,12 +541,6 @@ theorem setIntegral_condExpIndSMul (hs : MeasurableSet[m] s) (ht : MeasurableSet @[deprecated (since := "2025-01-21")] alias setIntegral_condexpIndSMul := setIntegral_condExpIndSMul -@[deprecated (since := "2024-04-17")] -alias set_integral_condExpIndSMul := setIntegral_condExpIndSMul - -@[deprecated (since := "2025-01-21")] -alias set_integral_condexpIndSMul := set_integral_condExpIndSMul - theorem condExpL2_indicator_nonneg (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) [SigmaFinite (μ.trim hm)] : (0 : α → ℝ) ≤ᵐ[μ] condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) := by diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean index 3f4e72c29f95a..038f169b51ba4 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean @@ -90,11 +90,6 @@ theorem eLpNorm_one_condExp_le_eLpNorm (f : α → ℝ) : eLpNorm (μ[f|m]) 1 μ @[deprecated (since := "2025-01-21")] alias eLpNorm_one_condexp_le_eLpNorm := eLpNorm_one_condExp_le_eLpNorm -@[deprecated (since := "2024-07-27")] -alias snorm_one_condExp_le_snorm := eLpNorm_one_condExp_le_eLpNorm - -@[deprecated (since := "2025-01-21")] alias snorm_one_condexp_le_snorm := snorm_one_condExp_le_snorm - theorem integral_abs_condExp_le (f : α → ℝ) : ∫ x, |(μ[f|m]) x| ∂μ ≤ ∫ x, |f x| ∂μ := by by_cases hm : m ≤ m0 swap @@ -145,12 +140,6 @@ theorem setIntegral_abs_condExp_le {s : Set α} (hs : MeasurableSet[m] s) (f : @[deprecated (since := "2025-01-21")] alias setIntegral_abs_condexp_le := setIntegral_abs_condExp_le -@[deprecated (since := "2024-04-17")] -alias set_integral_abs_condExp_le := setIntegral_abs_condExp_le - -@[deprecated (since := "2025-01-21")] -alias set_integral_abs_condexp_le := set_integral_abs_condExp_le - /-- If the real valued function `f` is bounded almost everywhere by `R`, then so is its conditional expectation. -/ theorem ae_bdd_condExp_of_ae_bdd {R : ℝ≥0} {f : α → ℝ} (hbdd : ∀ᵐ x ∂μ, |f x| ≤ R) : diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean index 490cbe24cc817..03ad4be2de84f 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean @@ -63,10 +63,6 @@ theorem lpMeas.ae_eq_zero_of_forall_setIntegral_eq_zero (hm : m ≤ m0) (f : lpM rw [integral_congr_ae hfg_restrict.symm] exact hf_zero s hs hμs -@[deprecated (since := "2024-04-17")] -alias lpMeas.ae_eq_zero_of_forall_set_integral_eq_zero := - lpMeas.ae_eq_zero_of_forall_setIntegral_eq_zero - variable (𝕜) include 𝕜 in @@ -89,10 +85,6 @@ theorem Lp.ae_eq_zero_of_forall_setIntegral_eq_zero' (hm : m ≤ m0) (f : Lp E' rw [integral_congr_ae hfg_restrict.symm] exact hf_zero s hs hμs -@[deprecated (since := "2024-04-17")] -alias Lp.ae_eq_zero_of_forall_set_integral_eq_zero' := - Lp.ae_eq_zero_of_forall_setIntegral_eq_zero' - include 𝕜 in /-- **Uniqueness of the conditional expectation** -/ theorem Lp.ae_eq_of_forall_setIntegral_eq' (hm : m ≤ m0) (f g : Lp E' p μ) (hp_ne_zero : p ≠ 0) @@ -115,9 +107,6 @@ theorem Lp.ae_eq_of_forall_setIntegral_eq' (hm : m ≤ m0) (f g : Lp E' p μ) (h exact Lp.ae_eq_zero_of_forall_setIntegral_eq_zero' 𝕜 hm (f - g) hp_ne_zero hp_ne_top hfg_int hfg' <| (hf_meas.sub hg_meas).congr (Lp.coeFn_sub f g).symm -@[deprecated (since := "2024-04-17")] -alias Lp.ae_eq_of_forall_set_integral_eq' := Lp.ae_eq_of_forall_setIntegral_eq' - variable {𝕜} theorem ae_eq_of_forall_setIntegral_eq_of_sigmaFinite' (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] @@ -161,10 +150,6 @@ theorem ae_eq_of_forall_setIntegral_eq_of_sigmaFinite' (hm : m ≤ m0) [SigmaFin exact hfg_eq s hs hμs exact ae_eq_of_forall_setIntegral_eq_of_sigmaFinite hf_mk_int_finite hg_mk_int_finite hfg_mk_eq -@[deprecated (since := "2024-04-17")] -alias ae_eq_of_forall_set_integral_eq_of_sigmaFinite' := - ae_eq_of_forall_setIntegral_eq_of_sigmaFinite' - end UniquenessOfConditionalExpectation section IntegralNormLE diff --git a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean index 8110b019c55bb..36d4c24c918e9 100644 --- a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean +++ b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean @@ -129,9 +129,6 @@ theorem exists_continuous_eLpNorm_sub_le_of_closed [μ.OuterRegular] (hp : p ≠ gc_support.trans inter_subset_left, gc_mem⟩ exact hη _ ((measure_mono (diff_subset_diff inter_subset_right Subset.rfl)).trans hV.le) -@[deprecated (since := "2024-07-27")] -alias exists_continuous_snorm_sub_le_of_closed := exists_continuous_eLpNorm_sub_le_of_closed - /-- In a locally compact space, any function in `ℒp` can be approximated by compactly supported continuous functions when `p < ∞`, version in terms of `eLpNorm`. -/ theorem Memℒp.exists_hasCompactSupport_eLpNorm_sub_le @@ -186,9 +183,6 @@ theorem Memℒp.exists_hasCompactSupport_eLpNorm_sub_le contrapose! hx exact interior_subset (f_support hx) -@[deprecated (since := "2024-07-27")] -alias Memℒp.exists_hasCompactSupport_snorm_sub_le := Memℒp.exists_hasCompactSupport_eLpNorm_sub_le - /-- In a locally compact space, any function in `ℒp` can be approximated by compactly supported continuous functions when `0 < p < ∞`, version in terms of `∫`. -/ theorem Memℒp.exists_hasCompactSupport_integral_rpow_sub_le @@ -282,9 +276,6 @@ theorem Memℒp.exists_boundedContinuous_eLpNorm_sub_le [μ.WeaklyRegular] (hp : refine ⟨f, I3, f_cont, f_mem, ?_⟩ exact (BoundedContinuousFunction.ofNormedAddCommGroup f f_cont _ f_bound).isBounded_range -@[deprecated (since := "2024-07-27")] -alias Memℒp.exists_boundedContinuous_snorm_sub_le := Memℒp.exists_boundedContinuous_eLpNorm_sub_le - /-- Any function in `ℒp` can be approximated by bounded continuous functions when `0 < p < ∞`, version in terms of `∫`. -/ theorem Memℒp.exists_boundedContinuous_integral_rpow_sub_le [μ.WeaklyRegular] {p : ℝ} (hp : 0 < p) diff --git a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean index 3e55e291670f1..1130352ecbb23 100644 --- a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean +++ b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean @@ -275,10 +275,6 @@ theorem tendstoInMeasure_of_tendsto_eLpNorm_of_stronglyMeasurable (hp_ne_zero : · rw [Ne, ENNReal.ofReal_eq_zero, not_le] exact Or.inl (Real.rpow_pos_of_pos hε _) -@[deprecated (since := "2024-07-27")] -alias tendstoInMeasure_of_tendsto_snorm_of_stronglyMeasurable := - tendstoInMeasure_of_tendsto_eLpNorm_of_stronglyMeasurable - /-- This lemma is superseded by `MeasureTheory.tendstoInMeasure_of_tendsto_eLpNorm` where we allow `p = ∞`. -/ theorem tendstoInMeasure_of_tendsto_eLpNorm_of_ne_top (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) @@ -292,9 +288,6 @@ theorem tendstoInMeasure_of_tendsto_eLpNorm_of_ne_top (hp_ne_zero : p ≠ 0) (hp rw [this] exact hfg -@[deprecated (since := "2024-07-27")] -alias tendstoInMeasure_of_tendsto_snorm_of_ne_top := tendstoInMeasure_of_tendsto_eLpNorm_of_ne_top - /-- See also `MeasureTheory.tendstoInMeasure_of_tendsto_eLpNorm` which work for general Lp-convergence for all `p ≠ 0`. -/ theorem tendstoInMeasure_of_tendsto_eLpNorm_top {E} [NormedAddCommGroup E] {f : ι → α → E} @@ -318,9 +311,6 @@ theorem tendstoInMeasure_of_tendsto_eLpNorm_top {E} [NormedAddCommGroup E] {f : coe_nnnorm, Set.mem_setOf_eq, Set.mem_compl_iff] rw [← dist_eq_norm (f n x) (g x)] -@[deprecated (since := "2024-07-27")] -alias tendstoInMeasure_of_tendsto_snorm_top := tendstoInMeasure_of_tendsto_eLpNorm_top - /-- Convergence in Lp implies convergence in measure. -/ theorem tendstoInMeasure_of_tendsto_eLpNorm {l : Filter ι} (hp_ne_zero : p ≠ 0) (hf : ∀ n, AEStronglyMeasurable (f n) μ) (hg : AEStronglyMeasurable g μ) @@ -330,9 +320,6 @@ theorem tendstoInMeasure_of_tendsto_eLpNorm {l : Filter ι} (hp_ne_zero : p ≠ exact tendstoInMeasure_of_tendsto_eLpNorm_top hfg · exact tendstoInMeasure_of_tendsto_eLpNorm_of_ne_top hp_ne_zero hp_ne_top hf hg hfg -@[deprecated (since := "2024-07-27")] -alias tendstoInMeasure_of_tendsto_snorm := tendstoInMeasure_of_tendsto_eLpNorm - /-- Convergence in Lp implies convergence in measure. -/ theorem tendstoInMeasure_of_tendsto_Lp [hp : Fact (1 ≤ p)] {f : ι → Lp E p μ} {g : Lp E p μ} {l : Filter ι} (hfg : Tendsto f l (𝓝 g)) : TendstoInMeasure μ (fun n => f n) l g := diff --git a/Mathlib/MeasureTheory/Function/L2Space.lean b/Mathlib/MeasureTheory/Function/L2Space.lean index 44dc28cd5f2ba..4fd62f0ddf78d 100644 --- a/Mathlib/MeasureTheory/Function/L2Space.lean +++ b/Mathlib/MeasureTheory/Function/L2Space.lean @@ -111,9 +111,6 @@ theorem eLpNorm_rpow_two_norm_lt_top (f : Lp F 2 μ) : rw [eLpNorm_norm_rpow f zero_lt_two, one_mul, h_two] exact ENNReal.rpow_lt_top_of_nonneg zero_le_two (Lp.eLpNorm_ne_top f) -@[deprecated (since := "2024-07-27")] -alias snorm_rpow_two_norm_lt_top := eLpNorm_rpow_two_norm_lt_top - theorem eLpNorm_inner_lt_top (f g : α →₂[μ] E) : eLpNorm (fun x : α => ⟪f x, g x⟫) 1 μ < ∞ := by have h : ∀ x, ‖⟪f x, g x⟫‖ ≤ ‖‖f x‖ ^ (2 : ℝ) + ‖g x‖ ^ (2 : ℝ)‖ := by intro x @@ -131,9 +128,6 @@ theorem eLpNorm_inner_lt_top (f g : α →₂[μ] E) : eLpNorm (fun x : α => rw [ENNReal.add_lt_top] exact ⟨eLpNorm_rpow_two_norm_lt_top f, eLpNorm_rpow_two_norm_lt_top g⟩ -@[deprecated (since := "2024-07-27")] -alias snorm_inner_lt_top := eLpNorm_inner_lt_top - section InnerProductSpace open scoped ComplexConjugate @@ -159,9 +153,6 @@ theorem integral_inner_eq_sq_eLpNorm (f : α →₂[μ] E) : ENNReal.ofReal_rpow_of_nonneg (norm_nonneg _) zero_le_two, ofReal_norm_eq_enorm] norm_cast -@[deprecated (since := "2024-07-27")] -alias integral_inner_eq_sq_snorm := integral_inner_eq_sq_eLpNorm - private theorem norm_sq_eq_inner' (f : α →₂[μ] E) : ‖f‖ ^ 2 = RCLike.re ⟪f, f⟫ := by have h_two : (2 : ℝ≥0∞).toReal = 2 := by simp rw [inner_def, integral_inner_eq_sq_eLpNorm, norm_def, ← ENNReal.toReal_pow, RCLike.ofReal_re, @@ -245,10 +236,6 @@ theorem inner_indicatorConstLp_eq_setIntegral_inner (f : Lp E 2 μ) (hs : Measur exact inner_zero_left _ rw [h_left, h_right, add_zero] -@[deprecated (since := "2024-04-17")] -alias inner_indicatorConstLp_eq_set_integral_inner := - inner_indicatorConstLp_eq_setIntegral_inner - /-- The inner product in `L2` of the indicator of a set `indicatorConstLp 2 hs hμs c` and `f` is equal to the inner product of the constant `c` and the integral of `f` over `s`. -/ theorem inner_indicatorConstLp_eq_inner_setIntegral [CompleteSpace E] [NormedSpace ℝ E] @@ -257,10 +244,6 @@ theorem inner_indicatorConstLp_eq_inner_setIntegral [CompleteSpace E] [NormedSpa rw [← integral_inner (integrableOn_Lp_of_measure_ne_top f fact_one_le_two_ennreal.elim hμs), L2.inner_indicatorConstLp_eq_setIntegral_inner] -@[deprecated (since := "2024-04-17")] -alias inner_indicatorConstLp_eq_inner_set_integral := - inner_indicatorConstLp_eq_inner_setIntegral - variable {𝕜} /-- The inner product in `L2` of the indicator of a set `indicatorConstLp 2 hs hμs (1 : 𝕜)` and diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 34188e7455aaa..dcb4374b447f5 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -93,26 +93,17 @@ def eLpNorm {_ : MeasurableSpace α} (f : α → ε) (p : ℝ≥0∞) (μ : Measure α := by volume_tac) : ℝ≥0∞ := if p = 0 then 0 else if p = ∞ then eLpNormEssSup f μ else eLpNorm' f (ENNReal.toReal p) μ -@[deprecated (since := "2024-07-26")] noncomputable alias snorm := eLpNorm - theorem eLpNorm_eq_eLpNorm' (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {f : α → F} : eLpNorm f p μ = eLpNorm' f (ENNReal.toReal p) μ := by simp [eLpNorm, hp_ne_zero, hp_ne_top] -@[deprecated (since := "2024-07-27")] alias snorm_eq_snorm' := eLpNorm_eq_eLpNorm' - lemma eLpNorm_nnreal_eq_eLpNorm' {f : α → F} {p : ℝ≥0} (hp : p ≠ 0) : eLpNorm f p μ = eLpNorm' f p μ := eLpNorm_eq_eLpNorm' (by exact_mod_cast hp) ENNReal.coe_ne_top -@[deprecated (since := "2024-07-27")] alias snorm_nnreal_eq_snorm' := eLpNorm_nnreal_eq_eLpNorm' - 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_enorm - @[deprecated (since := "2025-01-17")] alias eLpNorm_eq_lintegral_rpow_nnnorm := eLpNorm_eq_lintegral_rpow_enorm @@ -120,8 +111,6 @@ lemma eLpNorm_nnreal_eq_lintegral {f : α → F} {p : ℝ≥0} (hp : p ≠ 0) : 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_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] @@ -129,15 +118,9 @@ theorem eLpNorm_one_eq_lintegral_enorm {f : α → F} : eLpNorm f 1 μ = ∫⁻ @[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_enorm - @[simp] theorem eLpNorm_exponent_top {f : α → F} : eLpNorm f ∞ μ = eLpNormEssSup f μ := by simp [eLpNorm] -@[deprecated (since := "2024-07-27")] -alias snorm_exponent_top := eLpNorm_exponent_top - /-- The property that `f:α→E` is ae strongly measurable and `(∫ ‖f a‖^p ∂μ)^(1/p)` is finite if `p < ∞`, or `essSup f < ∞` if `p = ∞`. -/ def Memℒp {α} {_ : MeasurableSpace α} [TopologicalSpace ε] (f : α → ε) (p : ℝ≥0∞) @@ -156,17 +139,11 @@ theorem lintegral_rpow_enorm_eq_rpow_eLpNorm' {f : α → F} (hq0_lt : 0 < q) : @[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_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 simp [eLpNorm_eq_eLpNorm' (by exact_mod_cast hp) ENNReal.coe_ne_top, 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 - end ℒpSpaceDefinition section Top @@ -174,24 +151,14 @@ section Top theorem Memℒp.eLpNorm_lt_top {f : α → E} (hfp : Memℒp f p μ) : eLpNorm f p μ < ∞ := hfp.2 -@[deprecated (since := "2024-07-27")] -alias Memℒp.snorm_lt_top := Memℒp.eLpNorm_lt_top - theorem Memℒp.eLpNorm_ne_top {f : α → E} (hfp : Memℒp f p μ) : eLpNorm f p μ ≠ ∞ := ne_of_lt hfp.2 -@[deprecated (since := "2024-07-27")] -alias Memℒp.snorm_ne_top := Memℒp.eLpNorm_ne_top - 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_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 @@ -202,9 +169,6 @@ theorem lintegral_rpow_enorm_lt_top_of_eLpNorm_lt_top {f : α → F} (hp_ne_zero · 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_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 @@ -218,10 +182,6 @@ theorem eLpNorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top {f : α → F} (hp_ne_ze 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")] -alias snorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top := - eLpNorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top - end Top section Zero @@ -230,15 +190,9 @@ section Zero theorem eLpNorm'_exponent_zero {f : α → F} : eLpNorm' f 0 μ = 1 := by rw [eLpNorm', div_zero, ENNReal.rpow_zero] -@[deprecated (since := "2024-07-27")] -alias snorm'_exponent_zero := eLpNorm'_exponent_zero - @[simp] theorem eLpNorm_exponent_zero {f : α → F} : eLpNorm f 0 μ = 0 := by simp [eLpNorm] -@[deprecated (since := "2024-07-27")] -alias snorm_exponent_zero := eLpNorm_exponent_zero - @[simp] theorem memℒp_zero_iff_aestronglyMeasurable {f : α → E} : Memℒp f 0 μ ↔ AEStronglyMeasurable f μ := by simp [Memℒp, eLpNorm_exponent_zero] @@ -247,25 +201,16 @@ theorem memℒp_zero_iff_aestronglyMeasurable {f : α → E} : theorem eLpNorm'_zero (hp0_lt : 0 < q) : eLpNorm' (0 : α → F) q μ = 0 := by simp [eLpNorm'_eq_lintegral_enorm, hp0_lt] -@[deprecated (since := "2024-07-27")] -alias snorm'_zero := eLpNorm'_zero - @[simp] 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_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 [eLpNormEssSup, ← bot_eq_zero', essSup_const_bot] -@[deprecated (since := "2024-07-27")] -alias snormEssSup_zero := eLpNormEssSup_zero - @[simp] theorem eLpNorm_zero : eLpNorm (0 : α → F) p μ = 0 := by by_cases h0 : p = 0 @@ -275,15 +220,9 @@ theorem eLpNorm_zero : eLpNorm (0 : α → F) p μ = 0 := by rw [← Ne] at h0 simp [eLpNorm_eq_eLpNorm' h0 h_top, ENNReal.toReal_pos h0 h_top] -@[deprecated (since := "2024-07-27")] -alias snorm_zero := eLpNorm_zero - @[simp] theorem eLpNorm_zero' : eLpNorm (fun _ : α => (0 : F)) p μ = 0 := by convert eLpNorm_zero (F := F) -@[deprecated (since := "2024-07-27")] -alias snorm_zero' := eLpNorm_zero' - @[simp] lemma Memℒp.zero : Memℒp (0 : α → E) p μ := ⟨aestronglyMeasurable_zero, by rw [eLpNorm_zero]; exact ENNReal.coe_lt_top⟩ @@ -297,28 +236,16 @@ variable [MeasurableSpace α] theorem eLpNorm'_measure_zero_of_pos {f : α → F} (hq_pos : 0 < q) : eLpNorm' f q (0 : Measure α) = 0 := by simp [eLpNorm', hq_pos] -@[deprecated (since := "2024-07-27")] -alias snorm'_measure_zero_of_pos := eLpNorm'_measure_zero_of_pos - theorem eLpNorm'_measure_zero_of_exponent_zero {f : α → F} : eLpNorm' f 0 (0 : Measure α) = 1 := by simp [eLpNorm'] -@[deprecated (since := "2024-07-27")] -alias snorm'_measure_zero_of_exponent_zero := eLpNorm'_measure_zero_of_exponent_zero - theorem eLpNorm'_measure_zero_of_neg {f : α → F} (hq_neg : q < 0) : eLpNorm' f q (0 : Measure α) = ∞ := by simp [eLpNorm', hq_neg] -@[deprecated (since := "2024-07-27")] -alias snorm'_measure_zero_of_neg := eLpNorm'_measure_zero_of_neg - @[simp] theorem eLpNormEssSup_measure_zero {f : α → F} : eLpNormEssSup f (0 : Measure α) = 0 := by simp [eLpNormEssSup] -@[deprecated (since := "2024-07-27")] -alias snormEssSup_measure_zero := eLpNormEssSup_measure_zero - @[simp] theorem eLpNorm_measure_zero {f : α → F} : eLpNorm f p (0 : Measure α) = 0 := by by_cases h0 : p = 0 @@ -328,9 +255,6 @@ theorem eLpNorm_measure_zero {f : α → F} : eLpNorm f p (0 : Measure α) = 0 : rw [← Ne] at h0 simp [eLpNorm_eq_eLpNorm' h0 h_top, eLpNorm', ENNReal.toReal_pos h0 h_top] -@[deprecated (since := "2024-07-27")] -alias snorm_measure_zero := eLpNorm_measure_zero - @[simp] lemma memℒp_measure_zero {f : α → F} : Memℒp f p (0 : Measure α) := by simp [Memℒp] end Zero @@ -341,9 +265,6 @@ section Neg theorem eLpNorm'_neg (f : α → F) (q : ℝ) (μ : Measure α) : eLpNorm' (-f) q μ = eLpNorm' f q μ := by simp [eLpNorm'_eq_lintegral_enorm] -@[deprecated (since := "2024-07-27")] -alias snorm'_neg := eLpNorm'_neg - @[simp] theorem eLpNorm_neg (f : α → F) (p : ℝ≥0∞) (μ : Measure α) : eLpNorm (-f) p μ = eLpNorm f p μ := by by_cases h0 : p = 0 @@ -355,9 +276,6 @@ theorem eLpNorm_neg (f : α → F) (p : ℝ≥0∞) (μ : Measure α) : eLpNorm lemma eLpNorm_sub_comm (f g : α → E) (p : ℝ≥0∞) (μ : Measure α) : eLpNorm (f - g) p μ = eLpNorm (g - f) p μ := by simp [← eLpNorm_neg (f := f - g)] -@[deprecated (since := "2024-07-27")] -alias snorm_neg := eLpNorm_neg - theorem Memℒp.neg {f : α → E} (hf : Memℒp f p μ) : Memℒp (-f) p μ := ⟨AEStronglyMeasurable.neg hf.1, by simp [hf.right]⟩ @@ -377,9 +295,6 @@ theorem eLpNorm'_const (c : F) (hq_pos : 0 < q) : suffices hq_cancel : q * (1 / q) = 1 by rw [hq_cancel, ENNReal.rpow_one] rw [one_div, mul_inv_cancel₀ (ne_of_lt hq_pos).symm] -@[deprecated (since := "2024-07-27")] -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‖ₑ * μ Set.univ ^ (1 / q) := by rw [eLpNorm'_eq_lintegral_enorm, lintegral_const, @@ -391,37 +306,22 @@ theorem eLpNorm'_const' [IsFiniteMeasure μ] (c : F) (hc_ne_zero : c ≠ 0) (hq_ · rw [Ne, ENNReal.rpow_eq_top_iff, not_or, not_and_or, not_and_or] 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‖ₑ := 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‖ₑ := 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‖ₑ * μ 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] -@[deprecated (since := "2024-07-27")] -alias snorm_const := eLpNorm_const - theorem eLpNorm_const' (c : F) (h0 : p ≠ 0) (h_top : p ≠ ∞) : 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")] -alias snorm_const' := eLpNorm_const' - theorem eLpNorm_const_lt_top_iff {p : ℝ≥0∞} {c : F} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : eLpNorm (fun _ : α => c) p μ < ∞ ↔ c = 0 ∨ μ Set.univ < ∞ := by have hp : 0 < p.toReal := ENNReal.toReal_pos hp_ne_zero hp_ne_top @@ -437,9 +337,6 @@ theorem eLpNorm_const_lt_top_iff {p : ℝ≥0∞} {c : F} (hp_ne_zero : p ≠ 0) 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 - theorem memℒp_const (c : E) [IsFiniteMeasure μ] : Memℒp (fun _ : α => c) p μ := by refine ⟨aestronglyMeasurable_const, ?_⟩ by_cases h0 : p = 0 @@ -471,51 +368,30 @@ lemma eLpNorm'_mono_nnnorm_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : dsimp [enorm] gcongr -@[deprecated (since := "2024-07-27")] -alias snorm'_mono_nnnorm_ae := eLpNorm'_mono_nnnorm_ae - theorem eLpNorm'_mono_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : eLpNorm' f q μ ≤ eLpNorm' g q μ := eLpNorm'_mono_nnnorm_ae hq h -@[deprecated (since := "2024-07-27")] -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 : (‖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 - theorem eLpNorm'_congr_norm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖ = ‖g x‖) : eLpNorm' f q μ = eLpNorm' g q μ := eLpNorm'_congr_nnnorm_ae <| hfg.mono fun _x hx => NNReal.eq hx -@[deprecated (since := "2024-07-27")] -alias snorm'_congr_norm_ae := eLpNorm'_congr_norm_ae - theorem eLpNorm'_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : eLpNorm' f q μ = eLpNorm' g q μ := eLpNorm'_congr_nnnorm_ae (hfg.fun_comp _) -@[deprecated (since := "2024-07-27")] -alias snorm'_congr_ae := eLpNorm'_congr_ae - theorem eLpNormEssSup_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : eLpNormEssSup f μ = eLpNormEssSup g μ := essSup_congr_ae (hfg.fun_comp (((↑) : ℝ≥0 → ℝ≥0∞) ∘ nnnorm)) -@[deprecated (since := "2024-07-27")] -alias snormEssSup_congr_ae := eLpNormEssSup_congr_ae - theorem eLpNormEssSup_mono_nnnorm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : eLpNormEssSup f μ ≤ eLpNormEssSup g μ := essSup_mono_ae <| hfg.mono fun _x hx => ENNReal.coe_le_coe.mpr hx -@[deprecated (since := "2024-07-27")] -alias snormEssSup_mono_nnnorm_ae := eLpNormEssSup_mono_nnnorm_ae - theorem eLpNorm_mono_nnnorm_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : eLpNorm f p μ ≤ eLpNorm g p μ := by simp only [eLpNorm] @@ -524,73 +400,43 @@ theorem eLpNorm_mono_nnnorm_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ · exact essSup_mono_ae (h.mono fun x hx => ENNReal.coe_le_coe.mpr hx) · exact eLpNorm'_mono_nnnorm_ae ENNReal.toReal_nonneg h -@[deprecated (since := "2024-07-27")] -alias snorm_mono_nnnorm_ae := eLpNorm_mono_nnnorm_ae - theorem eLpNorm_mono_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : eLpNorm f p μ ≤ eLpNorm g p μ := eLpNorm_mono_nnnorm_ae h -@[deprecated (since := "2024-07-27")] -alias snorm_mono_ae := eLpNorm_mono_ae - theorem eLpNorm_mono_ae_real {f : α → F} {g : α → ℝ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ g x) : eLpNorm f p μ ≤ eLpNorm g p μ := eLpNorm_mono_ae <| h.mono fun _x hx => hx.trans ((le_abs_self _).trans (Real.norm_eq_abs _).symm.le) -@[deprecated (since := "2024-07-27")] -alias snorm_mono_ae_real := eLpNorm_mono_ae_real - theorem eLpNorm_mono_nnnorm {f : α → F} {g : α → G} (h : ∀ x, ‖f x‖₊ ≤ ‖g x‖₊) : eLpNorm f p μ ≤ eLpNorm g p μ := eLpNorm_mono_nnnorm_ae (Eventually.of_forall fun x => h x) -@[deprecated (since := "2024-07-27")] -alias snorm_mono_nnnorm := eLpNorm_mono_nnnorm - theorem eLpNorm_mono {f : α → F} {g : α → G} (h : ∀ x, ‖f x‖ ≤ ‖g x‖) : eLpNorm f p μ ≤ eLpNorm g p μ := eLpNorm_mono_ae (Eventually.of_forall fun x => h x) -@[deprecated (since := "2024-07-27")] -alias snorm_mono := eLpNorm_mono - theorem eLpNorm_mono_real {f : α → F} {g : α → ℝ} (h : ∀ x, ‖f x‖ ≤ g x) : eLpNorm f p μ ≤ eLpNorm g p μ := eLpNorm_mono_ae_real (Eventually.of_forall fun x => h x) -@[deprecated (since := "2024-07-27")] -alias snorm_mono_real := eLpNorm_mono_real - theorem eLpNormEssSup_le_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : eLpNormEssSup f μ ≤ C := essSup_le_of_ae_le (C : ℝ≥0∞) <| hfC.mono fun _x hx => ENNReal.coe_le_coe.mpr hx -@[deprecated (since := "2024-07-27")] -alias snormEssSup_le_of_ae_nnnorm_bound := eLpNormEssSup_le_of_ae_nnnorm_bound - theorem eLpNormEssSup_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : eLpNormEssSup f μ ≤ ENNReal.ofReal C := eLpNormEssSup_le_of_ae_nnnorm_bound <| hfC.mono fun _x hx => hx.trans C.le_coe_toNNReal -@[deprecated (since := "2024-07-27")] -alias snormEssSup_le_of_ae_bound := eLpNormEssSup_le_of_ae_bound - theorem eLpNormEssSup_lt_top_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : eLpNormEssSup f μ < ∞ := (eLpNormEssSup_le_of_ae_nnnorm_bound hfC).trans_lt ENNReal.coe_lt_top -@[deprecated (since := "2024-07-27")] -alias snormEssSup_lt_top_of_ae_nnnorm_bound := eLpNormEssSup_lt_top_of_ae_nnnorm_bound - theorem eLpNormEssSup_lt_top_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : eLpNormEssSup f μ < ∞ := (eLpNormEssSup_le_of_ae_bound hfC).trans_lt ENNReal.ofReal_lt_top -@[deprecated (since := "2024-07-27")] -alias snormEssSup_lt_top_of_ae_bound := eLpNormEssSup_lt_top_of_ae_bound - theorem eLpNorm_le_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : eLpNorm f p μ ≤ C • μ Set.univ ^ p.toReal⁻¹ := by rcases eq_zero_or_neZero μ with rfl | hμ @@ -601,55 +447,34 @@ theorem eLpNorm_le_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ refine (eLpNorm_mono_ae this).trans_eq ?_ 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 - theorem eLpNorm_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : eLpNorm f p μ ≤ μ Set.univ ^ p.toReal⁻¹ * ENNReal.ofReal C := by rw [← mul_comm] exact eLpNorm_le_of_ae_nnnorm_bound (hfC.mono fun x hx => hx.trans C.le_coe_toNNReal) -@[deprecated (since := "2024-07-27")] -alias snorm_le_of_ae_bound := eLpNorm_le_of_ae_bound - theorem eLpNorm_congr_nnnorm_ae {f : α → F} {g : α → G} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ = ‖g x‖₊) : eLpNorm f p μ = eLpNorm g p μ := le_antisymm (eLpNorm_mono_nnnorm_ae <| EventuallyEq.le hfg) (eLpNorm_mono_nnnorm_ae <| (EventuallyEq.symm hfg).le) -@[deprecated (since := "2024-07-27")] -alias snorm_congr_nnnorm_ae := eLpNorm_congr_nnnorm_ae - theorem eLpNorm_congr_norm_ae {f : α → F} {g : α → G} (hfg : ∀ᵐ x ∂μ, ‖f x‖ = ‖g x‖) : eLpNorm f p μ = eLpNorm g p μ := eLpNorm_congr_nnnorm_ae <| hfg.mono fun _x hx => NNReal.eq hx -@[deprecated (since := "2024-07-27")] -alias snorm_congr_norm_ae := eLpNorm_congr_norm_ae - open scoped symmDiff in theorem eLpNorm_indicator_sub_indicator (s t : Set α) (f : α → E) : eLpNorm (s.indicator f - t.indicator f) p μ = eLpNorm ((s ∆ t).indicator f) p μ := eLpNorm_congr_norm_ae <| ae_of_all _ fun x ↦ by simp only [Pi.sub_apply, Set.apply_indicator_symmDiff norm_neg] -@[deprecated (since := "2024-07-27")] -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_enorm] -@[deprecated (since := "2024-07-27")] -alias snorm'_norm := eLpNorm'_norm - @[simp] theorem eLpNorm_norm (f : α → F) : eLpNorm (fun x => ‖f x‖) p μ = eLpNorm f p μ := eLpNorm_congr_norm_ae <| Eventually.of_forall fun _ => norm_norm _ -@[deprecated (since := "2024-07-27")] -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', ← ENNReal.rpow_mul, ← one_div_mul_one_div, one_div, @@ -657,9 +482,6 @@ theorem eLpNorm'_norm_rpow (f : α → F) (p q : ℝ) (hq_pos : 0 < q) : 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] -@[deprecated (since := "2024-07-27")] -alias snorm'_norm_rpow := eLpNorm'_norm_rpow - theorem eLpNorm_norm_rpow (f : α → F) (hq_pos : 0 < q) : eLpNorm (fun x => ‖f x‖ ^ q) p μ = eLpNorm f (p * ENNReal.ofReal q) μ ^ q := by by_cases h0 : p = 0 @@ -685,15 +507,9 @@ theorem eLpNorm_norm_rpow (f : α → F) (hq_pos : 0 < q) : rw [ENNReal.toReal_mul, ENNReal.toReal_ofReal hq_pos.le] exact eLpNorm'_norm_rpow f p.toReal q hq_pos -@[deprecated (since := "2024-07-27")] -alias snorm_norm_rpow := eLpNorm_norm_rpow - theorem eLpNorm_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : eLpNorm f p μ = eLpNorm g p μ := eLpNorm_congr_norm_ae <| hfg.mono fun _x hx => hx ▸ rfl -@[deprecated (since := "2024-07-27")] -alias snorm_congr_ae := eLpNorm_congr_ae - theorem memℒp_congr_ae {f g : α → E} (hfg : f =ᵐ[μ] g) : Memℒp f p μ ↔ Memℒp g p μ := by simp only [Memℒp, eLpNorm_congr_ae hfg, aestronglyMeasurable_congr hfg] @@ -742,18 +558,12 @@ theorem eLpNorm'_mono_measure (f : α → F) (hμν : ν ≤ μ) (hq : 0 ≤ q) gcongr exact lintegral_mono' hμν le_rfl -@[deprecated (since := "2024-07-27")] -alias snorm'_mono_measure := eLpNorm'_mono_measure - @[gcongr, mono] theorem eLpNormEssSup_mono_measure (f : α → F) (hμν : ν ≪ μ) : eLpNormEssSup f ν ≤ eLpNormEssSup f μ := by simp_rw [eLpNormEssSup] exact essSup_mono_measure hμν -@[deprecated (since := "2024-07-27")] -alias snormEssSup_mono_measure := eLpNormEssSup_mono_measure - @[gcongr, mono] theorem eLpNorm_mono_measure (f : α → F) (hμν : ν ≤ μ) : eLpNorm f p ν ≤ eLpNorm f p μ := by by_cases hp0 : p = 0 @@ -763,9 +573,6 @@ theorem eLpNorm_mono_measure (f : α → F) (hμν : ν ≤ μ) : eLpNorm f p ν simp_rw [eLpNorm_eq_eLpNorm' hp0 hp_top] exact eLpNorm'_mono_measure f hμν ENNReal.toReal_nonneg -@[deprecated (since := "2024-07-27")] -alias snorm_mono_measure := eLpNorm_mono_measure - theorem Memℒp.mono_measure {f : α → E} (hμν : ν ≤ μ) (hf : Memℒp f p μ) : Memℒp f p ν := ⟨hf.1.mono_measure hμν, (eLpNorm_mono_measure f hμν).trans_lt hf.2⟩ @@ -791,9 +598,6 @@ lemma eLpNorm_indicator_eq_eLpNorm_restrict (hs : MeasurableSet s) : -- `∘` well. exact (Set.indicator_comp_of_zero (g := fun x : ℝ≥0∞ => x ^ p.toReal) h_zero).symm -@[deprecated (since := "2024-07-27")] -alias snorm_indicator_eq_restrict := eLpNorm_indicator_eq_eLpNorm_restrict - @[deprecated (since := "2025-01-07")] alias eLpNorm_indicator_eq_restrict := eLpNorm_indicator_eq_eLpNorm_restrict @@ -805,9 +609,6 @@ lemma eLpNorm_restrict_le (f : α → F) (p : ℝ≥0∞) (μ : Measure α) (s : eLpNorm f p (μ.restrict s) ≤ eLpNorm f p μ := eLpNorm_mono_measure f Measure.restrict_le_self -@[deprecated (since := "2024-07-27")] -alias snorm_restrict_le := eLpNorm_restrict_le - lemma eLpNorm_indicator_le (f : α → E) : eLpNorm (s.indicator f) p μ ≤ eLpNorm f p μ := by refine eLpNorm_mono_ae <| .of_forall fun x ↦ ?_ suffices ‖s.indicator f x‖₊ ≤ ‖f x‖₊ by exact NNReal.coe_mono this @@ -946,9 +747,6 @@ theorem eLpNorm_restrict_eq_of_support_subset {s : Set α} {f : α → F} (hsf : have : ¬(p.toReal ≤ 0) := by simpa only [not_le] using ENNReal.toReal_pos hp0 hp_top simpa [this] using hsf -@[deprecated (since := "2024-07-27")] -alias snorm_restrict_eq_of_support_subset := eLpNorm_restrict_eq_of_support_subset - theorem Memℒp.restrict (s : Set α) {f : α → E} (hf : Memℒp f p μ) : Memℒp f p (μ.restrict s) := hf.mono_measure Measure.restrict_le_self @@ -957,9 +755,6 @@ theorem eLpNorm'_smul_measure {p : ℝ} (hp : 0 ≤ p) {f : α → F} (c : ℝ rw [eLpNorm', lintegral_smul_measure, ENNReal.mul_rpow_of_nonneg, eLpNorm'] simp [hp] -@[deprecated (since := "2024-07-27")] -alias snorm'_smul_measure := eLpNorm'_smul_measure - section SMul variable {R : Type*} [Zero R] [SMulWithZero R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] [NoZeroSMulDivisors R ℝ≥0∞] {c : R} @@ -971,9 +766,6 @@ variable {R : Type*} [Zero R] [SMulWithZero R ℝ≥0∞] [IsScalarTower R ℝ end SMul -@[deprecated (since := "2024-07-27")] -alias snormEssSup_smul_measure := eLpNormEssSup_smul_measure - /-- Use `eLpNorm_smul_measure_of_ne_top` instead. -/ private theorem eLpNorm_smul_measure_of_ne_zero_of_ne_top {p : ℝ≥0∞} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {f : α → F} (c : ℝ≥0∞) : @@ -984,9 +776,6 @@ private theorem eLpNorm_smul_measure_of_ne_zero_of_ne_top {p : ℝ≥0∞} (hp_n simp_rw [one_div] rw [ENNReal.toReal_inv] -@[deprecated (since := "2024-07-27")] -alias snorm_smul_measure_of_ne_zero_of_ne_top := eLpNorm_smul_measure_of_ne_zero_of_ne_top - /-- See `eLpNorm_smul_measure_of_ne_zero'` for a version with scalar multiplication by `ℝ≥0`. -/ theorem eLpNorm_smul_measure_of_ne_zero {c : ℝ≥0∞} (hc : c ≠ 0) (f : α → F) (p : ℝ≥0∞) (μ : Measure α) : eLpNorm f p (c • μ) = c ^ (1 / p).toReal • eLpNorm f p μ := by @@ -996,9 +785,6 @@ theorem eLpNorm_smul_measure_of_ne_zero {c : ℝ≥0∞} (hc : c ≠ 0) (f : α · simp [hp_top, eLpNormEssSup_smul_measure hc] exact eLpNorm_smul_measure_of_ne_zero_of_ne_top hp0 hp_top c -@[deprecated (since := "2024-07-27")] -alias snorm_smul_measure_of_ne_zero := eLpNorm_smul_measure_of_ne_zero - /-- See `eLpNorm_smul_measure_of_ne_zero` for a version with scalar multiplication by `ℝ≥0∞`. -/ lemma eLpNorm_smul_measure_of_ne_zero' {c : ℝ≥0} (hc : c ≠ 0) (f : α → F) (p : ℝ≥0∞) (μ : Measure α) : eLpNorm f p (c • μ) = c ^ p.toReal⁻¹ • eLpNorm f p μ := @@ -1017,17 +803,11 @@ lemma eLpNorm_smul_measure_of_ne_top' (hp : p ≠ ∞) (c : ℝ≥0) (f : α → have : 0 ≤ p.toReal⁻¹ := by positivity refine (eLpNorm_smul_measure_of_ne_top hp ..).trans ?_ simp [ENNReal.smul_def, ENNReal.coe_rpow_of_nonneg, this] -@[deprecated (since := "2024-07-27")] -alias snorm_smul_measure_of_ne_top := eLpNorm_smul_measure_of_ne_top - theorem eLpNorm_one_smul_measure {f : α → F} (c : ℝ≥0∞) : eLpNorm f 1 (c • μ) = c * eLpNorm f 1 μ := by rw [@eLpNorm_smul_measure_of_ne_top _ _ _ μ _ 1 (@ENNReal.coe_ne_top 1) f c] simp -@[deprecated (since := "2024-07-27")] -alias snorm_one_smul_measure := eLpNorm_one_smul_measure - theorem Memℒp.of_measure_le_smul {μ' : Measure α} (c : ℝ≥0∞) (hc : c ≠ ∞) (hμ'_le : μ' ≤ c • μ) {f : α → E} (hf : Memℒp f p μ) : Memℒp f p μ' := by refine ⟨hf.1.mono_ac (Measure.absolutelyContinuous_of_le_smul hμ'_le), ?_⟩ @@ -1047,23 +827,14 @@ theorem eLpNorm_one_add_measure (f : α → F) (μ ν : Measure α) : simp_rw [eLpNorm_one_eq_lintegral_enorm] rw [lintegral_add_measure _ μ ν] -@[deprecated (since := "2024-07-27")] -alias snorm_one_add_measure := eLpNorm_one_add_measure - theorem eLpNorm_le_add_measure_right (f : α → F) (μ ν : Measure α) {p : ℝ≥0∞} : eLpNorm f p μ ≤ eLpNorm f p (μ + ν) := eLpNorm_mono_measure f <| Measure.le_add_right <| le_refl _ -@[deprecated (since := "2024-07-27")] -alias snorm_le_add_measure_right := eLpNorm_le_add_measure_right - theorem eLpNorm_le_add_measure_left (f : α → F) (μ ν : Measure α) {p : ℝ≥0∞} : eLpNorm f p ν ≤ eLpNorm f p (μ + ν) := eLpNorm_mono_measure f <| Measure.le_add_left <| le_refl _ -@[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‖ₑ := essSup_eq_iSup hμ _ @@ -1086,16 +857,10 @@ theorem memℒp_norm_iff {f : α → E} (hf : AEStronglyMeasurable f μ) : theorem eLpNorm'_eq_zero_of_ae_zero {f : α → F} (hq0_lt : 0 < q) (hf_zero : f =ᵐ[μ] 0) : eLpNorm' f q μ = 0 := by rw [eLpNorm'_congr_ae hf_zero, eLpNorm'_zero hq0_lt] -@[deprecated (since := "2024-07-27")] -alias snorm'_eq_zero_of_ae_zero := eLpNorm'_eq_zero_of_ae_zero - theorem eLpNorm'_eq_zero_of_ae_zero' (hq0_ne : q ≠ 0) (hμ : μ ≠ 0) {f : α → F} (hf_zero : f =ᵐ[μ] 0) : eLpNorm' f q μ = 0 := by rw [eLpNorm'_congr_ae hf_zero, eLpNorm'_zero' hq0_ne hμ] -@[deprecated (since := "2024-07-27")] -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 simp only [eLpNorm'_eq_lintegral_enorm, lintegral_eq_zero_iff' (hf.enorm.pow_const q), one_div, @@ -1105,30 +870,18 @@ theorem ae_eq_zero_of_eLpNorm'_eq_zero {f : α → E} (hq0 : 0 ≤ q) (hf : AESt 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 - theorem eLpNorm'_eq_zero_iff (hq0_lt : 0 < q) {f : α → E} (hf : AEStronglyMeasurable f μ) : eLpNorm' f q μ = 0 ↔ f =ᵐ[μ] 0 := ⟨ae_eq_zero_of_eLpNorm'_eq_zero (le_of_lt hq0_lt) hf, eLpNorm'_eq_zero_of_ae_zero hq0_lt⟩ -@[deprecated (since := "2024-07-27")] -alias snorm'_eq_zero_iff := eLpNorm'_eq_zero_iff - theorem coe_nnnorm_ae_le_eLpNormEssSup {_ : MeasurableSpace α} (f : α → F) (μ : Measure α) : ∀ᵐ 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_enorm] -@[deprecated (since := "2024-07-27")] -alias snormEssSup_eq_zero_iff := eLpNormEssSup_eq_zero_iff - theorem eLpNorm_eq_zero_iff {f : α → E} (hf : AEStronglyMeasurable f μ) (h0 : p ≠ 0) : eLpNorm f p μ = 0 ↔ f =ᵐ[μ] 0 := by by_cases h_top : p = ∞ @@ -1136,9 +889,6 @@ theorem eLpNorm_eq_zero_iff {f : α → E} (hf : AEStronglyMeasurable f μ) (h0 rw [eLpNorm_eq_eLpNorm' h0 h_top] exact eLpNorm'_eq_zero_iff (ENNReal.toReal_pos h0 h_top) hf -@[deprecated (since := "2024-07-27")] -alias snorm_eq_zero_iff := eLpNorm_eq_zero_iff - theorem eLpNorm_eq_zero_of_ae_zero {f : α → E} (hf : f =ᵐ[μ] 0) : eLpNorm f p μ = 0 := by rw [← eLpNorm_zero (p := p) (μ := μ) (α := α) (F := E)] exact eLpNorm_congr_ae hf @@ -1146,9 +896,6 @@ theorem eLpNorm_eq_zero_of_ae_zero {f : α → E} (hf : f =ᵐ[μ] 0) : eLpNorm theorem ae_le_eLpNormEssSup {f : α → F} : ∀ᵐ y ∂μ, ‖f y‖ₑ ≤ eLpNormEssSup f μ := ae_le_essSup -@[deprecated (since := "2024-07-27")] -alias ae_le_snormEssSup := ae_le_eLpNormEssSup - lemma eLpNormEssSup_lt_top_iff_isBoundedUnder : eLpNormEssSup f μ < ⊤ ↔ IsBoundedUnder (· ≤ ·) (ae μ) fun x ↦ ‖f x‖₊ where mp h := ⟨(eLpNormEssSup f μ).toNNReal, by @@ -1158,9 +905,6 @@ lemma eLpNormEssSup_lt_top_iff_isBoundedUnder : theorem meas_eLpNormEssSup_lt {f : α → F} : μ { y | eLpNormEssSup f μ < ‖f y‖ₑ } = 0 := meas_essSup_lt -@[deprecated (since := "2024-07-27")] -alias meas_snormEssSup_lt := meas_eLpNormEssSup_lt - lemma eLpNorm_lt_top_of_finite [Finite α] [IsFiniteMeasure μ] : eLpNorm f p μ < ∞ := by obtain rfl | hp₀ := eq_or_ne p 0 · simp @@ -1180,12 +924,6 @@ lemma eLpNorm_lt_top_of_finite [Finite α] [IsFiniteMeasure μ] : eLpNorm f p μ @[simp] lemma eLpNorm_of_isEmpty [IsEmpty α] (f : α → E) (p : ℝ≥0∞) : eLpNorm f p μ = 0 := by simp [Subsingleton.elim f 0] -@[deprecated (since := "2024-07-27")] -alias snormEssSup_piecewise := eLpNormEssSup_piecewise - -@[deprecated (since := "2024-07-27")] -alias snorm_top_piecewise := eLpNorm_top_piecewise - section MapMeasure variable {β : Type*} {mβ : MeasurableSpace β} {f : α → β} {g : β → E} @@ -1194,9 +932,6 @@ 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.enorm hf -@[deprecated (since := "2024-07-27")] -alias snormEssSup_map_measure := eLpNormEssSup_map_measure - theorem eLpNorm_map_measure (hg : AEStronglyMeasurable g (Measure.map f μ)) (hf : AEMeasurable f μ) : eLpNorm g p (Measure.map f μ) = eLpNorm (g ∘ f) p μ := by by_cases hp_zero : p = 0 @@ -1208,9 +943,6 @@ theorem eLpNorm_map_measure (hg : AEStronglyMeasurable g (Measure.map f μ)) rw [lintegral_map' (hg.enorm.pow_const p.toReal) hf] rfl -@[deprecated (since := "2024-07-27")] -alias snorm_map_measure := eLpNorm_map_measure - theorem memℒp_map_measure_iff (hg : AEStronglyMeasurable g (Measure.map f μ)) (hf : AEMeasurable f μ) : Memℒp g p (Measure.map f μ) ↔ Memℒp (g ∘ f) p μ := by simp [Memℒp, eLpNorm_map_measure hg hf, hg.comp_aemeasurable hf, hg] @@ -1223,18 +955,12 @@ theorem eLpNorm_comp_measurePreserving {ν : MeasureTheory.Measure β} (hg : AES (hf : MeasurePreserving f μ ν) : eLpNorm (g ∘ f) p μ = eLpNorm g p ν := Eq.symm <| hf.map_eq ▸ eLpNorm_map_measure (hf.map_eq ▸ hg) hf.aemeasurable -@[deprecated (since := "2024-07-27")] -alias snorm_comp_measurePreserving := eLpNorm_comp_measurePreserving - theorem AEEqFun.eLpNorm_compMeasurePreserving {ν : MeasureTheory.Measure β} (g : β →ₘ[ν] E) (hf : MeasurePreserving f μ ν) : eLpNorm (g.compMeasurePreserving f hf) p μ = eLpNorm g p ν := by rw [eLpNorm_congr_ae (g.coeFn_compMeasurePreserving _)] exact eLpNorm_comp_measurePreserving g.aestronglyMeasurable hf -@[deprecated (since := "2024-07-27")] -alias AEEqFun.snorm_compMeasurePreserving := AEEqFun.eLpNorm_compMeasurePreserving - theorem Memℒp.comp_measurePreserving {ν : MeasureTheory.Measure β} (hg : Memℒp g p ν) (hf : MeasurePreserving f μ ν) : Memℒp (g ∘ f) p μ := .comp_of_map (hf.map_eq.symm ▸ hg) hf.aemeasurable @@ -1243,10 +969,6 @@ theorem _root_.MeasurableEmbedding.eLpNormEssSup_map_measure {g : β → F} (hf : MeasurableEmbedding f) : eLpNormEssSup g (Measure.map f μ) = eLpNormEssSup (g ∘ f) μ := hf.essSup_map_measure -@[deprecated (since := "2024-07-27")] -alias _root_.MeasurableEmbedding.snormEssSup_map_measure := - _root_.MeasurableEmbedding.eLpNormEssSup_map_measure - theorem _root_.MeasurableEmbedding.eLpNorm_map_measure {g : β → F} (hf : MeasurableEmbedding f) : eLpNorm g p (Measure.map f μ) = eLpNorm (g ∘ f) p μ := by by_cases hp_zero : p = 0 @@ -1258,9 +980,6 @@ theorem _root_.MeasurableEmbedding.eLpNorm_map_measure {g : β → F} (hf : Meas rw [hf.lintegral_map] rfl -@[deprecated (since := "2024-07-27")] -alias _root_.MeasurableEmbedding.snorm_map_measure := _root_.MeasurableEmbedding.eLpNorm_map_measure - theorem _root_.MeasurableEmbedding.memℒp_map_measure_iff {g : β → F} (hf : MeasurableEmbedding f) : Memℒp g p (Measure.map f μ) ↔ Memℒp (g ∘ f) p μ := by simp_rw [Memℒp, hf.aestronglyMeasurable_map_iff, hf.eLpNorm_map_measure] @@ -1286,9 +1005,6 @@ theorem eLpNorm'_le_nnreal_smul_eLpNorm'_of_ae_le_mul {f : α → F} {g : α → simp_rw [ENNReal.coe_le_coe, ← NNReal.mul_rpow, NNReal.rpow_le_rpow_iff hp] exact h -@[deprecated (since := "2024-07-27")] -alias snorm'_le_nnreal_smul_snorm'_of_ae_le_mul := eLpNorm'_le_nnreal_smul_eLpNorm'_of_ae_le_mul - 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 @@ -1297,10 +1013,6 @@ theorem eLpNormEssSup_le_nnreal_smul_eLpNormEssSup_of_ae_le_mul {f : α → F} { _ = 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 := - eLpNormEssSup_le_nnreal_smul_eLpNormEssSup_of_ae_le_mul - theorem eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) (p : ℝ≥0∞) : eLpNorm f p μ ≤ c • eLpNorm g p μ := by by_cases h0 : p = 0 @@ -1311,9 +1023,6 @@ theorem eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul {f : α → F} {g : α → G simp_rw [eLpNorm_eq_eLpNorm' h0 h_top] exact eLpNorm'_le_nnreal_smul_eLpNorm'_of_ae_le_mul h (ENNReal.toReal_pos h0 h_top) -@[deprecated (since := "2024-07-27")] -alias snorm_le_nnreal_smul_snorm_of_ae_le_mul := eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul - -- TODO: add the whole family of lemmas? private theorem le_mul_iff_eq_zero_of_nonneg_of_neg_of_nonneg {α} [LinearOrderedSemiring α] {a b c : α} (ha : 0 ≤ a) (hb : b < 0) (hc : 0 ≤ c) : a ≤ b * c ↔ a = 0 ∧ c = 0 := by @@ -1335,18 +1044,12 @@ theorem eLpNorm_eq_zero_and_zero_of_ae_le_mul_neg {f : α → F} {g : α → G} change f =ᵐ[μ] 0 ∧ g =ᵐ[μ] 0 at h simp [eLpNorm_congr_ae h.1, eLpNorm_congr_ae h.2] -@[deprecated (since := "2024-07-27")] -alias snorm_eq_zero_and_zero_of_ae_le_mul_neg := eLpNorm_eq_zero_and_zero_of_ae_le_mul_neg - theorem eLpNorm_le_mul_eLpNorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (p : ℝ≥0∞) : eLpNorm f p μ ≤ ENNReal.ofReal c * eLpNorm g p μ := eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul (h.mono fun _x hx => hx.trans <| mul_le_mul_of_nonneg_right c.le_coe_toNNReal (norm_nonneg _)) _ -@[deprecated (since := "2024-07-27")] -alias snorm_le_mul_snorm_of_ae_le_mul := eLpNorm_le_mul_eLpNorm_of_ae_le_mul - theorem Memℒp.of_nnnorm_le_mul {f : α → E} {g : α → F} {c : ℝ≥0} (hg : Memℒp g p μ) (hf : AEStronglyMeasurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : Memℒp f p μ := ⟨hf, @@ -1374,23 +1077,14 @@ variable [BoundedSMul 𝕜 E] [BoundedSMul 𝕜 F] {c : 𝕜} {f : α → F} 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 μ := 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 μ := eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul (Eventually.of_forall fun _ => by simp [nnnorm_smul_le]) _ -@[deprecated (since := "2024-07-27")] -alias snorm_const_smul_le := eLpNorm_const_smul_le - theorem Memℒp.const_smul (hf : Memℒp f p μ) (c : 𝕜) : Memℒp (c • f) p μ := ⟨AEStronglyMeasurable.const_smul hf.1 c, eLpNorm_const_smul_le.trans_lt (ENNReal.mul_lt_top ENNReal.coe_lt_top hf.2)⟩ @@ -1419,17 +1113,11 @@ theorem eLpNorm'_const_smul {f : α → F} (c : 𝕜) (hq_pos : 0 < q) : 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‖ₑ * 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‖ₑ * eLpNorm f p μ := by obtain rfl | hc := eq_or_ne c 0 @@ -1438,9 +1126,6 @@ theorem eLpNorm_const_smul (c : 𝕜) (f : α → F) (p : ℝ≥0∞) (μ : Meas 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 - lemma eLpNorm_nsmul [NormedSpace ℝ F] (n : ℕ) (f : α → F) : eLpNorm (n • f) p μ = n * eLpNorm f p μ := by simpa [Nat.cast_smul_eq_nsmul] using eLpNorm_const_smul (n : ℝ) f .. @@ -1463,12 +1148,6 @@ theorem le_eLpNorm_of_bddBelow (hp : p ≠ 0) (hp' : p ≠ ∞) {f : α → F} ( rwa [coe_le_enorm] · simp [Set.indicator_of_not_mem hxs] -@[deprecated (since := "2024-07-27")] -alias le_snorm_of_bddBelow := le_eLpNorm_of_bddBelow - -@[deprecated (since := "2024-06-26")] -alias snorm_indicator_ge_of_bdd_below := le_snorm_of_bddBelow - section RCLike variable {𝕜 : Type*} [RCLike 𝕜] {f : α → 𝕜} @@ -1521,9 +1200,6 @@ theorem ae_bdd_liminf_atTop_rpow_of_eLpNorm_bdd {p : ℝ≥0∞} {f : ℕ → α sSup_le fun b ⟨a, ha⟩ => (ha a le_rfl).trans ((ENNReal.rpow_inv_le_iff (ENNReal.toReal_pos hp hp')).1 (hbdd _)) -@[deprecated (since := "2024-07-27")] -alias ae_bdd_liminf_atTop_rpow_of_snorm_bdd := ae_bdd_liminf_atTop_rpow_of_eLpNorm_bdd - 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‖ₑ)) atTop < ∞ := by @@ -1552,9 +1228,6 @@ theorem ae_bdd_liminf_atTop_of_eLpNorm_bdd {p : ℝ≥0∞} (hp : p ≠ 0) {f : ENNReal.rpow_mul] exact ENNReal.rpow_lt_top_of_nonneg (inv_nonneg.2 hppos.le) hx.ne -@[deprecated (since := "2024-07-27")] -alias ae_bdd_liminf_atTop_of_snorm_bdd := ae_bdd_liminf_atTop_of_eLpNorm_bdd - end Liminf /-- A continuous function with compact support belongs to `L^∞`. @@ -1585,12 +1258,7 @@ theorem Memℒp.exists_eLpNorm_indicator_compl_lt {β : Type*} [NormedAddCommGro eLpNorm_eq_lintegral_rpow_enorm hp₀ hp_top, one_div, ENNReal.rpow_inv_lt_iff] simp [ENNReal.toReal_pos, *] -@[deprecated (since := "2024-07-27")] -alias Memℒp.exists_snorm_indicator_compl_lt := Memℒp.exists_eLpNorm_indicator_compl_lt - end UnifTight end ℒp end MeasureTheory - -set_option linter.style.longFile 1700 diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean index 968d77d53489e..73181c77d8878 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean @@ -25,9 +25,6 @@ theorem pow_mul_meas_ge_le_eLpNorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞ gcongr 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 - theorem mul_meas_ge_le_pow_eLpNorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : AEStronglyMeasurable f μ) (ε : ℝ≥0∞) : ε * μ { x | ε ≤ (‖f x‖₊ : ℝ≥0∞) ^ p.toReal } ≤ eLpNorm f p μ ^ p.toReal := by @@ -39,9 +36,6 @@ theorem mul_meas_ge_le_pow_eLpNorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞ gcongr exact pow_mul_meas_ge_le_eLpNorm μ hp_ne_zero hp_ne_top hf ε -@[deprecated (since := "2024-07-27")] -alias mul_meas_ge_le_pow_snorm := mul_meas_ge_le_pow_eLpNorm - /-- A version of Chebyshev-Markov's inequality using Lp-norms. -/ theorem mul_meas_ge_le_pow_eLpNorm' (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : AEStronglyMeasurable f μ) (ε : ℝ≥0∞) : @@ -50,9 +44,6 @@ theorem mul_meas_ge_le_pow_eLpNorm' (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ext x rw [ENNReal.rpow_le_rpow_iff (ENNReal.toReal_pos hp_ne_zero hp_ne_top)] -@[deprecated (since := "2024-07-27")] -alias mul_meas_ge_le_pow_snorm' := mul_meas_ge_le_pow_eLpNorm' - theorem meas_ge_le_mul_pow_eLpNorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : AEStronglyMeasurable f μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : μ { x | ε ≤ ‖f x‖₊ } ≤ ε⁻¹ ^ p.toReal * eLpNorm f p μ ^ p.toReal := by @@ -64,9 +55,6 @@ theorem meas_ge_le_mul_pow_eLpNorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞ ENNReal.mul_inv_cancel hεpow hεpow', one_mul] exact mul_meas_ge_le_pow_eLpNorm' μ hp_ne_zero hp_ne_top hf ε -@[deprecated (since := "2024-07-27")] -alias meas_ge_le_mul_pow_snorm := meas_ge_le_mul_pow_eLpNorm - theorem Memℒp.meas_ge_lt_top' {μ : Measure α} (hℒp : Memℒp f p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {ε : ℝ≥0∞} (hε : ε ≠ 0) : μ { x | ε ≤ ‖f x‖₊ } < ∞ := by diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean index f4c629a7fb32b..8ffc8136c61b8 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean @@ -43,9 +43,6 @@ theorem eLpNorm'_le_eLpNorm'_mul_rpow_measure_univ {p q : ℝ} (hp0_lt : 0 < p) _ = (∫⁻ a : α, ↑‖f a‖₊ ^ q ∂μ) ^ (1 / q) * μ Set.univ ^ (1 / p - 1 / q) := by rw [hpqr]; simp [r, g] -@[deprecated (since := "2024-07-27")] -alias snorm'_le_snorm'_mul_rpow_measure_univ := eLpNorm'_le_eLpNorm'_mul_rpow_measure_univ - 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‖ₑ ^ q ∂μ) ≤ ∫⁻ _ : α, eLpNormEssSup f μ ^ q ∂μ := by @@ -58,9 +55,6 @@ theorem eLpNorm'_le_eLpNormEssSup_mul_rpow_measure_univ {q : ℝ} (hq_pos : 0 < gcongr rwa [lintegral_const] at h_le -@[deprecated (since := "2024-07-27")] -alias snorm'_le_snormEssSup_mul_rpow_measure_univ := eLpNorm'_le_eLpNormEssSup_mul_rpow_measure_univ - theorem eLpNorm_le_eLpNorm_mul_rpow_measure_univ {p q : ℝ≥0∞} (hpq : p ≤ q) (hf : AEStronglyMeasurable f μ) : eLpNorm f p μ ≤ eLpNorm f q μ * μ Set.univ ^ (1 / p.toReal - 1 / q.toReal) := by @@ -87,32 +81,20 @@ theorem eLpNorm_le_eLpNorm_mul_rpow_measure_univ {p q : ℝ≥0∞} (hpq : p ≤ have hpq_real : p.toReal ≤ q.toReal := ENNReal.toReal_mono hq_top hpq exact eLpNorm'_le_eLpNorm'_mul_rpow_measure_univ hp_pos hpq_real hf -@[deprecated (since := "2024-07-27")] -alias snorm_le_snorm_mul_rpow_measure_univ := eLpNorm_le_eLpNorm_mul_rpow_measure_univ - theorem eLpNorm'_le_eLpNorm'_of_exponent_le {p q : ℝ} (hp0_lt : 0 < p) (hpq : p ≤ q) (μ : Measure α) [IsProbabilityMeasure μ] (hf : AEStronglyMeasurable f μ) : eLpNorm' f p μ ≤ eLpNorm' f q μ := by have h_le_μ := eLpNorm'_le_eLpNorm'_mul_rpow_measure_univ hp0_lt hpq hf rwa [measure_univ, ENNReal.one_rpow, mul_one] at h_le_μ -@[deprecated (since := "2024-07-27")] -alias snorm'_le_snorm'_of_exponent_le := eLpNorm'_le_eLpNorm'_of_exponent_le - theorem eLpNorm'_le_eLpNormEssSup {q : ℝ} (hq_pos : 0 < q) [IsProbabilityMeasure μ] : eLpNorm' f q μ ≤ eLpNormEssSup f μ := (eLpNorm'_le_eLpNormEssSup_mul_rpow_measure_univ hq_pos).trans_eq (by simp [measure_univ]) -@[deprecated (since := "2024-07-27")] -alias snorm'_le_snormEssSup := eLpNorm'_le_eLpNormEssSup - theorem eLpNorm_le_eLpNorm_of_exponent_le {p q : ℝ≥0∞} (hpq : p ≤ q) [IsProbabilityMeasure μ] (hf : AEStronglyMeasurable f μ) : eLpNorm f p μ ≤ eLpNorm f q μ := (eLpNorm_le_eLpNorm_mul_rpow_measure_univ hpq hf).trans (le_of_eq (by simp [measure_univ])) -@[deprecated (since := "2024-07-27")] -alias snorm_le_snorm_of_exponent_le := eLpNorm_le_eLpNorm_of_exponent_le - theorem eLpNorm'_lt_top_of_eLpNorm'_lt_top_of_exponent_le {p q : ℝ} [IsFiniteMeasure μ] (hf : AEStronglyMeasurable f μ) (hfq_lt_top : eLpNorm' f q μ < ∞) (hp_nonneg : 0 ≤ p) (hpq : p ≤ q) : eLpNorm' f p μ < ∞ := by @@ -128,10 +110,6 @@ theorem eLpNorm'_lt_top_of_eLpNorm'_lt_top_of_exponent_le {p q : ℝ} [IsFiniteM refine Or.inl ⟨hfq_lt_top, ENNReal.rpow_lt_top_of_nonneg ?_ (measure_ne_top μ Set.univ)⟩ rwa [le_sub_comm, sub_zero, one_div, one_div, inv_le_inv₀ hq_pos hp_pos] -@[deprecated (since := "2024-07-27")] -alias snorm'_lt_top_of_snorm'_lt_top_of_exponent_le := - eLpNorm'_lt_top_of_eLpNorm'_lt_top_of_exponent_le - theorem Memℒp.mono_exponent {p q : ℝ≥0∞} [IsFiniteMeasure μ] {f : α → E} (hfq : Memℒp f q μ) (hpq : p ≤ q) : Memℒp f p μ := by cases' hfq with hfq_m hfq_lt_top @@ -227,9 +205,6 @@ theorem eLpNorm_le_eLpNorm_top_mul_eLpNorm (p : ℝ≥0∞) (f : α → E) {g : rw [Ne, ENNReal.toReal_eq_zero_iff, not_or] exact ⟨hp_zero, hp_top⟩ -@[deprecated (since := "2024-07-27")] -alias snorm_le_snorm_top_mul_snorm := eLpNorm_le_eLpNorm_top_mul_eLpNorm - theorem eLpNorm_le_eLpNorm_mul_eLpNorm_top (p : ℝ≥0∞) {f : α → E} (hf : AEStronglyMeasurable f μ) (g : α → F) (b : E → F → G) (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) : eLpNorm (fun x => b (f x) (g x)) p μ ≤ eLpNorm f p μ * eLpNorm g ∞ μ := @@ -238,9 +213,6 @@ theorem eLpNorm_le_eLpNorm_mul_eLpNorm_top (p : ℝ≥0∞) {f : α → E} (hf : eLpNorm_le_eLpNorm_top_mul_eLpNorm p g hf (flip b) <| by simpa only [mul_comm] using h _ = eLpNorm f p μ * eLpNorm g ∞ μ := mul_comm _ _ -@[deprecated (since := "2024-07-27")] -alias snorm_le_snorm_mul_snorm_top := eLpNorm_le_eLpNorm_mul_eLpNorm_top - theorem eLpNorm'_le_eLpNorm'_mul_eLpNorm' {p q r : ℝ} (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasurable g μ) (b : E → F → G) (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) (hp0_lt : 0 < p) (hpq : p < q) @@ -257,9 +229,6 @@ theorem eLpNorm'_le_eLpNorm'_mul_eLpNorm' {p q r : ℝ} (hf : AEStronglyMeasurab simp_rw [eLpNorm', ENNReal.coe_mul] 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' - /-- Hölder's inequality, as an inequality on the `ℒp` seminorm of an elementwise operation `fun x => b (f x) (g x)`. -/ theorem eLpNorm_le_eLpNorm_mul_eLpNorm_of_nnnorm {p q r : ℝ≥0∞} @@ -300,9 +269,6 @@ theorem eLpNorm_le_eLpNorm_mul_eLpNorm_of_nnnorm {p q r : ℝ≥0∞} · simp only [hq_ne_zero, one_div, Ne, ENNReal.inv_eq_top, not_false_iff] · simp only [hr_ne_zero, one_div, Ne, ENNReal.inv_eq_top, not_false_iff] -@[deprecated (since := "2024-07-27")] -alias snorm_le_snorm_mul_snorm_of_nnnorm := eLpNorm_le_eLpNorm_mul_eLpNorm_of_nnnorm - /-- Hölder's inequality, as an inequality on the `ℒp` seminorm of an elementwise operation `fun x => b (f x) (g x)`. -/ theorem eLpNorm_le_eLpNorm_mul_eLpNorm'_of_norm {p q r : ℝ≥0∞} (hf : AEStronglyMeasurable f μ) @@ -311,9 +277,6 @@ theorem eLpNorm_le_eLpNorm_mul_eLpNorm'_of_norm {p q r : ℝ≥0∞} (hf : AEStr eLpNorm (fun x => b (f x) (g x)) p μ ≤ eLpNorm f q μ * eLpNorm g r μ := eLpNorm_le_eLpNorm_mul_eLpNorm_of_nnnorm hf hg b h hpqr -@[deprecated (since := "2024-07-27")] -alias snorm_le_snorm_mul_snorm'_of_norm := eLpNorm_le_eLpNorm_mul_eLpNorm'_of_norm - end Bilinear section BoundedSMul @@ -327,26 +290,17 @@ theorem eLpNorm_smul_le_eLpNorm_top_mul_eLpNorm (p : ℝ≥0∞) (hf : AEStrongl (eLpNorm_le_eLpNorm_top_mul_eLpNorm p φ hf (· • ·) (Eventually.of_forall fun _ => nnnorm_smul_le _ _) :) -@[deprecated (since := "2024-07-27")] -alias snorm_smul_le_snorm_top_mul_snorm := eLpNorm_smul_le_eLpNorm_top_mul_eLpNorm - theorem eLpNorm_smul_le_eLpNorm_mul_eLpNorm_top (p : ℝ≥0∞) (f : α → E) {φ : α → 𝕜} (hφ : AEStronglyMeasurable φ μ) : eLpNorm (φ • f) p μ ≤ eLpNorm φ p μ * eLpNorm f ∞ μ := (eLpNorm_le_eLpNorm_mul_eLpNorm_top p hφ f (· • ·) (Eventually.of_forall fun _ => nnnorm_smul_le _ _) :) -@[deprecated (since := "2024-07-27")] -alias snorm_smul_le_snorm_mul_snorm_top := eLpNorm_smul_le_eLpNorm_mul_eLpNorm_top - theorem eLpNorm'_smul_le_mul_eLpNorm' {p q r : ℝ} {f : α → E} (hf : AEStronglyMeasurable f μ) {φ : α → 𝕜} (hφ : AEStronglyMeasurable φ μ) (hp0_lt : 0 < p) (hpq : p < q) (hpqr : 1 / p = 1 / q + 1 / r) : eLpNorm' (φ • f) p μ ≤ eLpNorm' φ q μ * eLpNorm' f r μ := eLpNorm'_le_eLpNorm'_mul_eLpNorm' hφ hf (· • ·) (Eventually.of_forall fun _ => nnnorm_smul_le _ _) hp0_lt hpq hpqr -@[deprecated (since := "2024-07-27")] -alias snorm'_smul_le_mul_snorm' := eLpNorm'_smul_le_mul_eLpNorm' - /-- Hölder's inequality, as an inequality on the `ℒp` seminorm of a scalar product `φ • f`. -/ theorem eLpNorm_smul_le_mul_eLpNorm {p q r : ℝ≥0∞} {f : α → E} (hf : AEStronglyMeasurable f μ) {φ : α → 𝕜} (hφ : AEStronglyMeasurable φ μ) (hpqr : 1 / p = 1 / q + 1 / r) : @@ -355,9 +309,6 @@ theorem eLpNorm_smul_le_mul_eLpNorm {p q r : ℝ≥0∞} {f : α → E} (hf : AE (Eventually.of_forall fun _ => nnnorm_smul_le _ _) hpqr : _) -@[deprecated (since := "2024-07-27")] -alias snorm_smul_le_mul_snorm := eLpNorm_smul_le_mul_eLpNorm - theorem Memℒp.smul {p q r : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} (hf : Memℒp f r μ) (hφ : Memℒp φ q μ) (hpqr : 1 / p = 1 / q + 1 / r) : Memℒp (φ • f) p μ := ⟨hφ.1.smul hf.1, diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean index b1ebbc6cbffaa..1663c46dd836c 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean @@ -30,9 +30,6 @@ theorem eLpNorm'_add_le (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasura 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.enorm hg.enorm hq1 -@[deprecated (since := "2024-07-27")] -alias snorm'_add_le := eLpNorm'_add_le - theorem eLpNorm'_add_le_of_le_one (hf : AEStronglyMeasurable f μ) (hq0 : 0 ≤ q) (hq1 : q ≤ 1) : eLpNorm' (f + g) q μ ≤ (2 : ℝ≥0∞) ^ (1 / q - 1) * (eLpNorm' f q μ + eLpNorm' g q μ) := calc @@ -43,18 +40,12 @@ theorem eLpNorm'_add_le_of_le_one (hf : AEStronglyMeasurable f μ) (hq0 : 0 ≤ _ ≤ (2 : ℝ≥0∞) ^ (1 / q - 1) * (eLpNorm' f q μ + eLpNorm' g q μ) := 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 - theorem eLpNormEssSup_add_le {f g : α → E} : eLpNormEssSup (f + g) μ ≤ eLpNormEssSup f μ + eLpNormEssSup g μ := by refine le_trans (essSup_mono_ae (Eventually.of_forall fun x => ?_)) (ENNReal.essSup_add_le _ _) simp_rw [Pi.add_apply, enorm_eq_nnnorm, ← ENNReal.coe_add, ENNReal.coe_le_coe] exact nnnorm_add_le _ _ -@[deprecated (since := "2024-07-27")] -alias snormEssSup_add_le := eLpNormEssSup_add_le - theorem eLpNorm_add_le {f g : α → E} (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasurable g μ) (hp1 : 1 ≤ p) : eLpNorm (f + g) p μ ≤ eLpNorm f p μ + eLpNorm g p μ := by by_cases hp0 : p = 0 @@ -66,9 +57,6 @@ theorem eLpNorm_add_le {f g : α → E} (hf : AEStronglyMeasurable f μ) (hg : A repeat rw [eLpNorm_eq_eLpNorm' hp0 hp_top] exact eLpNorm'_add_le hf hg hp1_real -@[deprecated (since := "2024-07-27")] -alias snorm_add_le := eLpNorm_add_le - /-- A constant for the inequality `‖f + g‖_{L^p} ≤ C * (‖f‖_{L^p} + ‖g‖_{L^p})`. It is equal to `1` for `p ≥ 1` or `p = 0`, and `2^(1/p-1)` in the more tricky interval `(0, 1)`. -/ noncomputable def LpAddConst (p : ℝ≥0∞) : ℝ≥0∞ := @@ -104,9 +92,6 @@ theorem eLpNorm_add_le' (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasura · simpa using ENNReal.toReal_mono ENNReal.one_ne_top h'p.le · simpa [LpAddConst_of_one_le h'p] using eLpNorm_add_le hf hg h'p -@[deprecated (since := "2024-07-27")] -alias snorm_add_le' := eLpNorm_add_le' - variable (μ E) /-- Technical lemma to control the addition of functions in `L^p` even for `p < 1`: Given `δ > 0`, @@ -137,16 +122,10 @@ theorem eLpNorm_sub_le' (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasura (p : ℝ≥0∞) : eLpNorm (f - g) p μ ≤ LpAddConst p * (eLpNorm f p μ + eLpNorm g p μ) := by simpa only [sub_eq_add_neg, eLpNorm_neg] using eLpNorm_add_le' hf hg.neg p -@[deprecated (since := "2024-07-27")] -alias snorm_sub_le' := eLpNorm_sub_le' - theorem eLpNorm_sub_le {f g : α → E} (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasurable g μ) (hp : 1 ≤ p) : eLpNorm (f - g) p μ ≤ eLpNorm f p μ + eLpNorm g p μ := by simpa [LpAddConst_of_one_le hp] using eLpNorm_sub_le' hf hg p -@[deprecated (since := "2024-07-27")] -alias snorm_sub_le := eLpNorm_sub_le - theorem eLpNorm_add_lt_top {f g : α → E} (hf : Memℒp f p μ) (hg : Memℒp g p μ) : eLpNorm (f + g) p μ < ∞ := calc @@ -156,9 +135,6 @@ theorem eLpNorm_add_lt_top {f g : α → E} (hf : Memℒp f p μ) (hg : Memℒp apply ENNReal.mul_lt_top (LpAddConst_lt_top p) exact ENNReal.add_lt_top.2 ⟨hf.2, hg.2⟩ -@[deprecated (since := "2024-07-27")] -alias snorm_add_lt_top := eLpNorm_add_lt_top - theorem eLpNorm'_sum_le {ι} {f : ι → α → E} {s : Finset ι} (hfs : ∀ i, i ∈ s → AEStronglyMeasurable (f i) μ) (hq1 : 1 ≤ q) : eLpNorm' (∑ i ∈ s, f i) q μ ≤ ∑ i ∈ s, eLpNorm' (f i) q μ := @@ -166,9 +142,6 @@ theorem eLpNorm'_sum_le {ι} {f : ι → α → E} {s : Finset ι} (fun f => AEStronglyMeasurable f μ) (eLpNorm'_zero (zero_lt_one.trans_le hq1)) (fun _f _g hf hg => eLpNorm'_add_le hf hg hq1) (fun _f _g hf hg => hf.add hg) _ hfs -@[deprecated (since := "2024-07-27")] -alias snorm'_sum_le := eLpNorm'_sum_le - theorem eLpNorm_sum_le {ι} {f : ι → α → E} {s : Finset ι} (hfs : ∀ i, i ∈ s → AEStronglyMeasurable (f i) μ) (hp1 : 1 ≤ p) : eLpNorm (∑ i ∈ s, f i) p μ ≤ ∑ i ∈ s, eLpNorm (f i) p μ := @@ -176,9 +149,6 @@ theorem eLpNorm_sum_le {ι} {f : ι → α → E} {s : Finset ι} (fun f => AEStronglyMeasurable f μ) eLpNorm_zero (fun _f _g hf hg => eLpNorm_add_le hf hg hp1) (fun _f _g hf hg => hf.add hg) _ hfs -@[deprecated (since := "2024-07-27")] -alias snorm_sum_le := eLpNorm_sum_le - theorem Memℒp.add {f g : α → E} (hf : Memℒp f p μ) (hg : Memℒp g p μ) : Memℒp (f + g) p μ := ⟨AEStronglyMeasurable.add hf.1 hg.1, eLpNorm_add_lt_top hf hg⟩ diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean index 0cb5cd60fe951..e6daffcdb1f2d 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean @@ -29,9 +29,6 @@ theorem eLpNorm'_trim (hm : m ≤ m0) {f : α → E} (hf : StronglyMeasurable[m] apply @StronglyMeasurable.measurable exact @StronglyMeasurable.nnnorm α m _ _ _ hf -@[deprecated (since := "2024-07-27")] -alias snorm'_trim := eLpNorm'_trim - theorem limsup_trim (hm : m ≤ m0) {f : α → ℝ≥0∞} (hf : Measurable[m] f) : limsup f (ae (μ.trim hm)) = limsup f (ae μ) := by simp_rw [limsup_eq] @@ -53,9 +50,6 @@ theorem eLpNormEssSup_trim (hm : m ≤ m0) {f : α → E} (hf : StronglyMeasurab eLpNormEssSup f (μ.trim hm) = eLpNormEssSup f μ := essSup_trim _ (@StronglyMeasurable.enorm _ m _ _ _ hf) -@[deprecated (since := "2024-07-27")] -alias snormEssSup_trim := eLpNormEssSup_trim - theorem eLpNorm_trim (hm : m ≤ m0) {f : α → E} (hf : StronglyMeasurable[m] f) : eLpNorm f p (μ.trim hm) = eLpNorm f p μ := by by_cases h0 : p = 0 @@ -64,17 +58,11 @@ theorem eLpNorm_trim (hm : m ≤ m0) {f : α → E} (hf : StronglyMeasurable[m] · simpa only [h_top, eLpNorm_exponent_top] using eLpNormEssSup_trim hm hf simpa only [eLpNorm_eq_eLpNorm' h0 h_top] using eLpNorm'_trim hm hf -@[deprecated (since := "2024-07-27")] -alias snorm_trim := eLpNorm_trim - theorem eLpNorm_trim_ae (hm : m ≤ m0) {f : α → E} (hf : AEStronglyMeasurable[m] f (μ.trim hm)) : eLpNorm f p (μ.trim hm) = eLpNorm f p μ := by rw [eLpNorm_congr_ae hf.ae_eq_mk, eLpNorm_congr_ae (ae_eq_of_ae_eq_trim hf.ae_eq_mk)] exact eLpNorm_trim hm hf.stronglyMeasurable_mk -@[deprecated (since := "2024-07-27")] -alias snorm_trim_ae := eLpNorm_trim_ae - theorem memℒp_of_memℒp_trim (hm : m ≤ m0) {f : α → E} (hf : Memℒp f p (μ.trim hm)) : Memℒp f p μ := ⟨aestronglyMeasurable_of_aestronglyMeasurable_trim hm hf.1, (le_of_eq (eLpNorm_trim_ae hm hf.1).symm).trans_lt hf.2⟩ diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index c7af179f8bdd8..d48dd9ba2bb6f 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -88,16 +88,10 @@ theorem eLpNorm_aeeqFun {α E : Type*} [MeasurableSpace α] {μ : Measure α} [N eLpNorm (AEEqFun.mk f hf) p μ = eLpNorm f p μ := eLpNorm_congr_ae (AEEqFun.coeFn_mk _ _) -@[deprecated (since := "2024-07-27")] -alias snorm_aeeqFun := eLpNorm_aeeqFun - theorem Memℒp.eLpNorm_mk_lt_top {α E : Type*} [MeasurableSpace α] {μ : Measure α} [NormedAddCommGroup E] {p : ℝ≥0∞} {f : α → E} (hfp : Memℒp f p μ) : eLpNorm (AEEqFun.mk f hfp.1) p μ < ∞ := by simp [hfp.2] -@[deprecated (since := "2024-07-27")] -alias Memℒp.snorm_mk_lt_top := Memℒp.eLpNorm_mk_lt_top - /-- Lp space -/ def Lp {α} (E : Type*) {m : MeasurableSpace α} [NormedAddCommGroup E] (p : ℝ≥0∞) (μ : Measure α := by volume_tac) : AddSubgroup (α →ₘ[μ] E) where @@ -161,9 +155,6 @@ theorem ext {f g : Lp E p μ} (h : f =ᵐ[μ] g) : f = g := by theorem mem_Lp_iff_eLpNorm_lt_top {f : α →ₘ[μ] E} : f ∈ Lp E p μ ↔ eLpNorm f p μ < ∞ := Iff.rfl -@[deprecated (since := "2024-07-27")] -alias mem_Lp_iff_snorm_lt_top := mem_Lp_iff_eLpNorm_lt_top - theorem mem_Lp_iff_memℒp {f : α →ₘ[μ] E} : f ∈ Lp E p μ ↔ Memℒp f p μ := by simp [mem_Lp_iff_eLpNorm_lt_top, Memℒp, f.stronglyMeasurable.aestronglyMeasurable] @@ -186,15 +177,9 @@ theorem toLp_coeFn (f : Lp E p μ) (hf : Memℒp f p μ) : hf.toLp f = f := by theorem eLpNorm_lt_top (f : Lp E p μ) : eLpNorm f p μ < ∞ := f.prop -@[deprecated (since := "2024-07-27")] -alias snorm_lt_top := eLpNorm_lt_top - theorem eLpNorm_ne_top (f : Lp E p μ) : eLpNorm f p μ ≠ ∞ := (eLpNorm_lt_top f).ne -@[deprecated (since := "2024-07-27")] -alias snorm_ne_top := eLpNorm_ne_top - @[measurability] protected theorem stronglyMeasurable (f : Lp E p μ) : StronglyMeasurable f := f.val.stronglyMeasurable @@ -500,37 +485,6 @@ For a set `s` with `(hs : MeasurableSet s)` and `(hμs : μ s < ∞)`, we build `indicatorConstLp p hs hμs c`, the element of `Lp` corresponding to `s.indicator (fun _ => c)`. -/ -@[deprecated (since := "2024-07-27")] -alias snormEssSup_indicator_le := eLpNormEssSup_indicator_le - -@[deprecated (since := "2024-07-27")] -alias snormEssSup_indicator_const_le := eLpNormEssSup_indicator_const_le - -@[deprecated (since := "2024-07-27")] -alias snormEssSup_indicator_const_eq := eLpNormEssSup_indicator_const_eq - -@[deprecated (since := "2024-07-27")] -alias snorm_indicator_le := eLpNorm_indicator_le - -@[deprecated (since := "2024-07-27")] -alias snorm_indicator_const₀ := eLpNorm_indicator_const₀ - -@[deprecated (since := "2024-07-27")] -alias snorm_indicator_const := eLpNorm_indicator_const - -@[deprecated (since := "2024-07-27")] -alias snorm_indicator_const' := eLpNorm_indicator_const' - -@[deprecated (since := "2024-07-27")] -alias snorm_indicator_const_le := eLpNorm_indicator_const_le - -@[deprecated (since := "2024-07-27")] -alias snormEssSup_indicator_eq_snormEssSup_restrict := - eLpNormEssSup_indicator_eq_eLpNormEssSup_restrict - -@[deprecated (since := "2024-07-27")] -alias snorm_indicator_eq_snorm_restrict := eLpNorm_indicator_eq_eLpNorm_restrict - /-- The `ℒ^p` norm of the indicator of a set is uniformly small if the set itself has small measure, for any `p < ∞`. Given here as an existential `∀ ε > 0, ∃ η > 0, ...` to avoid later management of `ℝ≥0∞`-arithmetic. -/ @@ -557,9 +511,6 @@ theorem exists_eLpNorm_indicator_le (hp : p ≠ ∞) (c : E) {ε : ℝ≥0∞} ( refine (eLpNorm_indicator_const_le _ _).trans (le_trans ?_ hη_le) exact mul_le_mul_left' (ENNReal.rpow_le_rpow hs hp₀') _ -@[deprecated (since := "2024-07-27")] -alias exists_snorm_indicator_le := exists_eLpNorm_indicator_le - section Topology variable {X : Type*} [TopologicalSpace X] [MeasurableSpace X] {μ : Measure X} [IsFiniteMeasureOnCompacts μ] @@ -1172,9 +1123,6 @@ theorem eLpNorm'_lim_eq_lintegral_liminf {ι} [Nonempty ι] [LinearOrder ι] {f 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 - theorem eLpNorm'_lim_le_liminf_eLpNorm' {E} [NormedAddCommGroup E] {f : ℕ → α → E} {p : ℝ} (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))) : @@ -1193,9 +1141,6 @@ theorem eLpNorm'_lim_le_liminf_eLpNorm' {E} [NormedAddCommGroup E] {f : ℕ → 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")] -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‖ₑ) μ := by @@ -1205,9 +1150,6 @@ theorem eLpNorm_exponent_top_lim_eq_essSup_liminf {ι} [Nonempty ι] [LinearOrde apply (Tendsto.liminf_eq ..).symm 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 - theorem eLpNorm_exponent_top_lim_le_liminf_eLpNorm_exponent_top {ι} [Nonempty ι] [Countable ι] [LinearOrder ι] {f : ι → α → F} {f_lim : α → F} (h_lim : ∀ᵐ x : α ∂μ, Tendsto (fun n => f n x) atTop (𝓝 (f_lim x))) : @@ -1216,10 +1158,6 @@ theorem eLpNorm_exponent_top_lim_le_liminf_eLpNorm_exponent_top {ι} [Nonempty simp_rw [eLpNorm_exponent_top, eLpNormEssSup] exact ENNReal.essSup_liminf_le _ -@[deprecated (since := "2024-07-27")] -alias snorm_exponent_top_lim_le_liminf_snorm_exponent_top := - eLpNorm_exponent_top_lim_le_liminf_eLpNorm_exponent_top - theorem eLpNorm_lim_le_liminf_eLpNorm {E} [NormedAddCommGroup E] {f : ℕ → α → E} (hf : ∀ n, AEStronglyMeasurable (f n) μ) (f_lim : α → E) (h_lim : ∀ᵐ x : α ∂μ, Tendsto (fun n => f n x) atTop (𝓝 (f_lim x))) : @@ -1233,9 +1171,6 @@ theorem eLpNorm_lim_le_liminf_eLpNorm {E} [NormedAddCommGroup E] {f : ℕ → α have hp_pos : 0 < p.toReal := ENNReal.toReal_pos hp0 hp_top exact eLpNorm'_lim_le_liminf_eLpNorm' hp_pos hf h_lim -@[deprecated (since := "2024-07-27")] -alias snorm_lim_le_liminf_snorm := eLpNorm_lim_le_liminf_eLpNorm - /-! ### `Lp` is complete iff Cauchy sequences of `ℒp` have limits in `ℒp` -/ @@ -1339,10 +1274,6 @@ private theorem eLpNorm'_sum_norm_sub_le_tsum_of_cauchy_eLpNorm' {f : ℕ → α refine (Finset.sum_le_sum ?_).trans (sum_le_tsum _ (fun m _ => zero_le _) ENNReal.summable) exact fun m _ => (h_cau m (m + 1) m (Nat.le_succ m) (le_refl m)).le -@[deprecated (since := "2024-07-27")] -alias snorm'_sum_norm_sub_le_tsum_of_cauchy_snorm' := - eLpNorm'_sum_norm_sub_le_tsum_of_cauchy_eLpNorm' - 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) : @@ -1440,9 +1371,6 @@ theorem ae_tendsto_of_cauchy_eLpNorm' [CompleteSpace E] {f : ℕ → α → E} { rw [hf_rw] exact ⟨l + f 0 x, Tendsto.add_const _ hx⟩ -@[deprecated (since := "2024-07-27")] -alias ae_tendsto_of_cauchy_snorm' := ae_tendsto_of_cauchy_eLpNorm' - theorem ae_tendsto_of_cauchy_eLpNorm [CompleteSpace E] {f : ℕ → α → E} (hf : ∀ n, AEStronglyMeasurable (f n) μ) (hp : 1 ≤ p) {B : ℕ → ℝ≥0∞} (hB : ∑' i, B i ≠ ∞) (h_cau : ∀ N n m : ℕ, N ≤ n → N ≤ m → eLpNorm (f n - f m) p μ < B N) : @@ -1474,9 +1402,6 @@ theorem ae_tendsto_of_cauchy_eLpNorm [CompleteSpace E] {f : ℕ → α → E} rwa [eLpNorm_eq_eLpNorm' (zero_lt_one.trans_le hp).ne.symm hp_top] at h_cau exact ae_tendsto_of_cauchy_eLpNorm' hf hp1 hB h_cau' -@[deprecated (since := "2024-07-27")] -alias ae_tendsto_of_cauchy_snorm := ae_tendsto_of_cauchy_eLpNorm - theorem cauchy_tendsto_of_tendsto {f : ℕ → α → E} (hf : ∀ n, AEStronglyMeasurable (f n) μ) (f_lim : α → E) {B : ℕ → ℝ≥0∞} (hB : ∑' i, B i ≠ ∞) (h_cau : ∀ N n m : ℕ, N ≤ n → N ≤ m → eLpNorm (f n - f m) p μ < B N) @@ -1774,4 +1699,4 @@ end Lp end MeasureTheory -set_option linter.style.longFile 1900 +set_option linter.style.longFile 1800 diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index a7d019b6e5ebe..5053cdbe56a20 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -119,9 +119,6 @@ theorem tendsto_approxOn_Lp_eLpNorm [OpensMeasurableSpace E] {f : β → E} (hf -- Then we apply the Dominated Convergence Theorem simpa using tendsto_lintegral_of_dominated_convergence _ hF_meas h_bound h_fin h_lim -@[deprecated (since := "2024-07-27")] -alias tendsto_approxOn_Lp_snorm := tendsto_approxOn_Lp_eLpNorm - theorem memℒp_approxOn [BorelSpace E] {f : β → E} {μ : Measure β} (fmeas : Measurable f) (hf : Memℒp f p μ) {s : Set E} {y₀ : E} (h₀ : y₀ ∈ s) [SeparableSpace s] (hi₀ : Memℒp (fun _ => y₀) p μ) (n : ℕ) : Memℒp (approxOn f fmeas s y₀ h₀ n) p μ := by @@ -160,9 +157,6 @@ theorem tendsto_approxOn_range_Lp_eLpNorm [BorelSpace E] {f : β → E} (hp_ne_t · filter_upwards with x using subset_closure (by simp) · simpa using hf -@[deprecated (since := "2024-07-27")] -alias tendsto_approxOn_range_Lp_snorm := tendsto_approxOn_range_Lp_eLpNorm - theorem memℒp_approxOn_range [BorelSpace E] {f : β → E} {μ : Measure β} (fmeas : Measurable f) [SeparableSpace (range f ∪ {0} : Set E)] (hf : Memℒp f p μ) (n : ℕ) : Memℒp (approxOn f fmeas (range f ∪ {0}) 0 (by simp) n) p μ := @@ -199,10 +193,6 @@ theorem _root_.MeasureTheory.Memℒp.exists_simpleFunc_eLpNorm_sub_lt {E : Type* rw [← eLpNorm_neg, neg_sub] at hn exact ⟨_, hn, memℒp_approxOn_range f'meas hf' _⟩ -@[deprecated (since := "2024-07-27")] -alias _root_.MeasureTheory.Memℒp.exists_simpleFunc_snorm_sub_lt := - _root_.MeasureTheory.Memℒp.exists_simpleFunc_eLpNorm_sub_lt - end Lp /-! ### L1 approximation by simple functions -/ @@ -278,9 +268,6 @@ protected theorem eLpNorm'_eq {p : ℝ} (f : α →ₛ F) (μ : Measure α) : 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 - theorem measure_preimage_lt_top_of_memℒp (hp_pos : p ≠ 0) (hp_ne_top : p ≠ ∞) (f : α →ₛ E) (hf : Memℒp f p μ) (y : E) (hy_ne : y ≠ 0) : μ (f ⁻¹' {y}) < ∞ := by have hp_pos_real : 0 < p.toReal := ENNReal.toReal_pos hp_pos hp_ne_top diff --git a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean index 8ba93267f9ee7..ab96dccc8f1b8 100644 --- a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean @@ -264,9 +264,6 @@ theorem Memℒp.eLpNormEssSup_indicator_norm_ge_eq_zero (hf : Memℒp f ∞ μ) rw [this, eLpNormEssSup_measure_zero] exact measurableSet_le measurable_const hmeas.nnnorm.measurable.subtype_coe -@[deprecated (since := "2024-07-27")] -alias Memℒp.snormEssSup_indicator_norm_ge_eq_zero := Memℒp.eLpNormEssSup_indicator_norm_ge_eq_zero - /- This lemma is slightly weaker than `MeasureTheory.Memℒp.eLpNorm_indicator_norm_ge_pos_le` as the latter provides `0 < M`. -/ theorem Memℒp.eLpNorm_indicator_norm_ge_le (hf : Memℒp f p μ) (hmeas : StronglyMeasurable f) {ε : ℝ} @@ -304,9 +301,6 @@ theorem Memℒp.eLpNorm_indicator_norm_ge_le (hf : Memℒp f p μ) (hmeas : Stro · rw [Set.mem_setOf_eq] rwa [← hiff] -@[deprecated (since := "2024-07-27")] -alias Memℒp.snorm_indicator_norm_ge_le := Memℒp.eLpNorm_indicator_norm_ge_le - /-- This lemma implies that a single function is uniformly integrable (in the probability sense). -/ theorem Memℒp.eLpNorm_indicator_norm_ge_pos_le (hf : Memℒp f p μ) (hmeas : StronglyMeasurable f) {ε : ℝ} (hε : 0 < ε) : @@ -319,9 +313,6 @@ theorem Memℒp.eLpNorm_indicator_norm_ge_pos_le (hf : Memℒp f p μ) (hmeas : rw [Set.mem_setOf_eq] at hx -- removing the `rw` breaks the proof! exact (max_le_iff.1 hx).1 -@[deprecated (since := "2024-07-27")] -alias Memℒp.snorm_indicator_norm_ge_pos_le := Memℒp.eLpNorm_indicator_norm_ge_pos_le - end theorem eLpNorm_indicator_le_of_bound {f : α → β} (hp_top : p ≠ ∞) {ε : ℝ} (hε : 0 < ε) {M : ℝ} @@ -352,9 +343,6 @@ theorem eLpNorm_indicator_le_of_bound {f : α → β} (hp_top : p ≠ ∞) {ε : ENNReal.rpow_le_rpow_iff (ENNReal.toReal_pos hp hp_top), ENNReal.ofReal_div_of_pos hM] · simpa only [ENNReal.ofReal_eq_zero, not_le, Ne] -@[deprecated (since := "2024-07-27")] -alias snorm_indicator_le_of_bound := eLpNorm_indicator_le_of_bound - section variable {f : α → β} @@ -389,9 +377,6 @@ theorem Memℒp.eLpNorm_indicator_le' (hp_one : 1 ≤ p) (hp_top : p ≠ ∞) (h · rw [Pi.add_apply, Set.indicator_of_not_mem, Set.indicator_of_mem, zero_add] <;> simpa using hx -@[deprecated (since := "2024-07-27")] -alias Memℒp.snorm_indicator_le' := Memℒp.eLpNorm_indicator_le' - /-- This lemma is superseded by `MeasureTheory.Memℒp.eLpNorm_indicator_le` which does not require measurability on `f`. -/ theorem Memℒp.eLpNorm_indicator_le_of_meas (hp_one : 1 ≤ p) (hp_top : p ≠ ∞) (hf : Memℒp f p μ) @@ -404,9 +389,6 @@ theorem Memℒp.eLpNorm_indicator_le_of_meas (hp_one : 1 ≤ p) (hp_top : p ≠ ENNReal.mul_div_cancel] <;> norm_num -@[deprecated (since := "2024-07-27")] -alias Memℒp.snorm_indicator_le_of_meas := Memℒp.eLpNorm_indicator_le_of_meas - theorem Memℒp.eLpNorm_indicator_le (hp_one : 1 ≤ p) (hp_top : p ≠ ∞) (hf : Memℒp f p μ) {ε : ℝ} (hε : 0 < ε) : ∃ (δ : ℝ) (_ : 0 < δ), ∀ s, MeasurableSet s → μ s ≤ ENNReal.ofReal δ → @@ -419,9 +401,6 @@ theorem Memℒp.eLpNorm_indicator_le (hp_one : 1 ≤ p) (hp_top : p ≠ ∞) (hf rw [eLpNorm_indicator_eq_eLpNorm_restrict hs, eLpNorm_indicator_eq_eLpNorm_restrict hs] exact eLpNorm_congr_ae heq.restrict -@[deprecated (since := "2024-07-27")] -alias Memℒp.snorm_indicator_le := Memℒp.eLpNorm_indicator_le - /-- A constant function is uniformly integrable. -/ theorem unifIntegrable_const {g : α → β} (hp : 1 ≤ p) (hp_ne_top : p ≠ ∞) (hg : Memℒp g p μ) : UnifIntegrable (fun _ : ι => g) p μ := by @@ -500,9 +479,6 @@ theorem eLpNorm_sub_le_of_dist_bdd (μ : Measure α) refine mul_le_mul_right' (le_of_eq ?_) _ 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 - /-- A sequence of uniformly integrable functions which converges μ-a.e. converges in Lp. -/ theorem tendsto_Lp_finite_of_tendsto_ae_of_meas [IsFiniteMeasure μ] (hp : 1 ≤ p) (hp' : p ≠ ∞) {f : ℕ → α → β} {g : α → β} (hf : ∀ n, StronglyMeasurable (f n)) (hg : StronglyMeasurable g) diff --git a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean index 8844dfbdfd7a6..87c4fc1eac620 100644 --- a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean +++ b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean @@ -245,9 +245,6 @@ theorem setLIntegral_eq_tsum (h : IsFundamentalDomain G s μ) (f : α → ℝ≥ h.lintegral_eq_tsum_of_ac restrict_le_self.absolutelyContinuous _ _ = ∑' g : G, ∫⁻ x in t ∩ g • s, f x ∂μ := by simp only [h.restrict_restrict, inter_comm] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_eq_tsum := setLIntegral_eq_tsum - @[to_additive] theorem setLIntegral_eq_tsum' (h : IsFundamentalDomain G s μ) (f : α → ℝ≥0∞) (t : Set α) : ∫⁻ x in t, f x ∂μ = ∑' g : G, ∫⁻ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := @@ -258,9 +255,6 @@ theorem setLIntegral_eq_tsum' (h : IsFundamentalDomain G s μ) (f : α → ℝ _ = ∑' g : G, ∫⁻ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := tsum_congr fun g => Eq.symm <| (measurePreserving_smul g⁻¹ μ).setLIntegral_comp_emb (measurableEmbedding_const_smul _) _ _ -@[deprecated (since := "2024-06-29")] -alias set_lintegral_eq_tsum' := setLIntegral_eq_tsum' - @[to_additive] theorem measure_eq_tsum_of_ac (h : IsFundamentalDomain G s μ) (hν : ν ≪ μ) (t : Set α) : ν t = ∑' g : G, ν (t ∩ g • s) := by @@ -307,9 +301,6 @@ protected theorem setLIntegral_eq (hs : IsFundamentalDomain G s μ) (ht : IsFund _ = ∑' g : G, ∫⁻ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := by simp only [hf, inter_comm] _ = ∫⁻ x in t, f x ∂μ := (hs.setLIntegral_eq_tsum' _ _).symm -@[deprecated (since := "2024-06-29")] -alias set_lintegral_eq := MeasureTheory.IsFundamentalDomain.setLIntegral_eq - @[to_additive] theorem measure_set_eq (hs : IsFundamentalDomain G s μ) (ht : IsFundamentalDomain G t μ) {A : Set α} (hA₀ : MeasurableSet A) (hA : ∀ g : G, (fun x => g • x) ⁻¹' A = A) : μ (A ∩ s) = μ (A ∩ t) := by @@ -400,9 +391,6 @@ theorem setIntegral_eq_tsum (h : IsFundamentalDomain G s μ) {f : α → E} {t : _ = ∑' g : G, ∫ x in t ∩ g • s, f x ∂μ := by simp only [h.restrict_restrict, measure_smul, inter_comm] -@[deprecated (since := "2024-04-17")] -alias set_integral_eq_tsum := setIntegral_eq_tsum - @[to_additive] theorem setIntegral_eq_tsum' (h : IsFundamentalDomain G s μ) {f : α → E} {t : Set α} (hf : IntegrableOn f t μ) : ∫ x in t, f x ∂μ = ∑' g : G, ∫ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := @@ -414,9 +402,6 @@ theorem setIntegral_eq_tsum' (h : IsFundamentalDomain G s μ) {f : α → E} {t tsum_congr fun g => (measurePreserving_smul g⁻¹ μ).setIntegral_image_emb (measurableEmbedding_const_smul _) _ _ -@[deprecated (since := "2024-04-17")] -alias set_integral_eq_tsum' := setIntegral_eq_tsum' - @[to_additive] protected theorem setIntegral_eq (hs : IsFundamentalDomain G s μ) (ht : IsFundamentalDomain G t μ) {f : α → E} (hf : ∀ (g : G) (x), f (g • x) = f x) : ∫ x in s, f x ∂μ = ∫ x in t, f x ∂μ := by @@ -429,9 +414,6 @@ protected theorem setIntegral_eq (hs : IsFundamentalDomain G s μ) (ht : IsFunda · rw [integral_undef hfs, integral_undef] rwa [hs.integrableOn_iff ht hf] at hfs -@[deprecated (since := "2024-04-17")] -alias set_integral_eq := MeasureTheory.IsFundamentalDomain.setIntegral_eq - /-- If the action of a countable group `G` admits an invariant measure `μ` with a fundamental domain `s`, then every null-measurable set `t` such that the sets `g • t ∩ s` are pairwise a.e.-disjoint has measure at most `μ s`. -/ diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 2b9801d37f84f..ffe1348391420 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -891,10 +891,6 @@ theorem HasFiniteIntegral.tendsto_setIntegral_nhds_zero {ι} {f : α → G} (tendsto_setLIntegral_zero (ne_of_lt hf) hs) (fun i => zero_le _) fun i => enorm_integral_le_lintegral_enorm _ -@[deprecated (since := "2024-04-17")] -alias HasFiniteIntegral.tendsto_set_integral_nhds_zero := - HasFiniteIntegral.tendsto_setIntegral_nhds_zero - /-- If `f` is integrable, then `∫ x in s, f x ∂μ` is absolutely continuous in `s`: it tends to zero as `μ s` tends to zero. -/ theorem Integrable.tendsto_setIntegral_nhds_zero {ι} {f : α → G} (hf : Integrable f μ) @@ -902,10 +898,6 @@ theorem Integrable.tendsto_setIntegral_nhds_zero {ι} {f : α → G} (hf : Integ Tendsto (fun i => ∫ x in s i, f x ∂μ) l (𝓝 0) := hf.2.tendsto_setIntegral_nhds_zero hs -@[deprecated (since := "2024-04-17")] -alias Integrable.tendsto_set_integral_nhds_zero := - Integrable.tendsto_setIntegral_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) μ) @@ -936,9 +928,6 @@ lemma tendsto_setIntegral_of_L1 {ι} (f : α → G) (hfi : Integrable f μ) {F : 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) -@[deprecated (since := "2024-04-17")] -alias tendsto_set_integral_of_L1 := tendsto_setIntegral_of_L1 - /-- 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 ι} @@ -949,9 +938,6 @@ lemma tendsto_setIntegral_of_L1' {ι} (f : α → G) (hfi : Integrable f μ) {F simp_rw [eLpNorm_one_eq_lintegral_enorm, Pi.sub_apply] at hF exact hF -@[deprecated (since := "2024-04-17")] -alias tendsto_set_integral_of_L1' := tendsto_setIntegral_of_L1' - variable {X : Type*} [TopologicalSpace X] [FirstCountableTopology X] theorem continuousWithinAt_of_dominated {F : X → α → G} {x₀ : X} {bound : α → ℝ} {s : Set X} @@ -1331,9 +1317,6 @@ theorem Memℒp.eLpNorm_eq_integral_rpow_norm {f : α → H} {p : ℝ≥0∞} (h rw [A, ← ofReal_rpow_of_nonneg toReal_nonneg (inv_nonneg.2 toReal_nonneg), ofReal_toReal] 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 - end NormedAddCommGroup theorem integral_mono_ae {f g : α → ℝ} (hf : Integrable f μ) (hg : Integrable g μ) (h : f ≤ᵐ[μ] g) : @@ -1638,9 +1621,6 @@ theorem setIntegral_dirac' {mα : MeasurableSpace α} {f : α → E} (hf : Stron · exact integral_dirac' _ _ hf · exact integral_zero_measure _ -@[deprecated (since := "2024-04-17")] -alias set_integral_dirac' := setIntegral_dirac' - theorem setIntegral_dirac [MeasurableSpace α] [MeasurableSingletonClass α] (f : α → E) (a : α) (s : Set α) [Decidable (a ∈ s)] : ∫ x in s, f x ∂Measure.dirac a = if a ∈ s then f a else 0 := by @@ -1649,9 +1629,6 @@ theorem setIntegral_dirac [MeasurableSpace α] [MeasurableSingletonClass α] (f · exact integral_dirac _ _ · exact integral_zero_measure _ -@[deprecated (since := "2024-04-17")] -alias set_integral_dirac := setIntegral_dirac - /-- **Markov's inequality** also known as **Chebyshev's first inequality**. -/ theorem mul_meas_ge_le_integral_of_nonneg {f : α → ℝ} (hf_nonneg : 0 ≤ᵐ[μ] f) (hf_int : Integrable f μ) (ε : ℝ) : ε * (μ { x | ε ≤ f x }).toReal ≤ ∫ x, f x ∂μ := by @@ -1912,18 +1889,12 @@ theorem eLpNorm_one_le_of_le {r : ℝ≥0} (hfint : Integrable f μ) (hfint' : 0 rwa [← two_mul, mul_assoc, mul_le_mul_left (two_pos : (0 : ℝ) < 2)] · exact hfint.neg.sup (integrable_zero _ _ μ) -@[deprecated (since := "2024-07-27")] -alias snorm_one_le_of_le := eLpNorm_one_le_of_le - theorem eLpNorm_one_le_of_le' {r : ℝ} (hfint : Integrable f μ) (hfint' : 0 ≤ ∫ x, f x ∂μ) (hf : ∀ᵐ ω ∂μ, f ω ≤ r) : eLpNorm f 1 μ ≤ 2 * μ Set.univ * ENNReal.ofReal r := by refine eLpNorm_one_le_of_le hfint hfint' ?_ simp only [Real.coe_toNNReal', le_max_iff] filter_upwards [hf] with ω hω using Or.inl hω -@[deprecated (since := "2024-07-27")] -alias snorm_one_le_of_le' := eLpNorm_one_le_of_le' - end SnormBound end MeasureTheory diff --git a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean index 063908dfd80a5..7beda9a175fda 100644 --- a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean +++ b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean @@ -180,9 +180,6 @@ theorem _root_.Antitone.tendsto_setIntegral (hsm : ∀ i, MeasurableSet (s i)) ( exact indicator_le_indicator_of_subset (h_anti (zero_le n)) (fun a => norm_nonneg _) _ · filter_upwards [] with a using le_trans (h_anti.tendsto_indicator _ _ _) (pure_le_nhds _) -@[deprecated (since := "2024-04-17")] -alias _root_.Antitone.tendsto_set_integral := _root_.Antitone.tendsto_setIntegral - end TendstoMono /-! diff --git a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean index d5c7a91ba8e6a..cac07461b168d 100644 --- a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean +++ b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean @@ -342,9 +342,6 @@ theorem IntegrableOn.setLIntegral_lt_top {f : α → ℝ} {s : Set α} (hf : Int (∫⁻ x in s, ENNReal.ofReal (f x) ∂μ) < ∞ := Integrable.lintegral_lt_top hf -@[deprecated (since := "2024-06-29")] -alias IntegrableOn.set_lintegral_lt_top := IntegrableOn.setLIntegral_lt_top - /-- We say that a function `f` is *integrable at filter* `l` if it is integrable on some set `s ∈ l`. Equivalently, it is eventually integrable on `s` in `l.smallSets`. -/ def IntegrableAtFilter (f : α → ε) (l : Filter α) (μ : Measure α := by volume_tac) := diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index e0363619f99bb..a71d150fc2f11 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -599,16 +599,6 @@ theorem integrableOn_Ioc_of_intervalIntegral_norm_bounded_right {I a b₀ : ℝ} (h : ∀ᶠ i in l, (∫ x in Ioc a (b i), ‖f x‖) ≤ I) : IntegrableOn f (Ioc a b₀) := integrableOn_Ioc_of_intervalIntegral_norm_bounded hfi tendsto_const_nhds hb h -@[deprecated (since := "2024-04-06")] -alias integrableOn_Ioc_of_interval_integral_norm_bounded := - integrableOn_Ioc_of_intervalIntegral_norm_bounded -@[deprecated (since := "2024-04-06")] -alias integrableOn_Ioc_of_interval_integral_norm_bounded_left := - integrableOn_Ioc_of_intervalIntegral_norm_bounded_left -@[deprecated (since := "2024-04-06")] -alias integrableOn_Ioc_of_interval_integral_norm_bounded_right := - integrableOn_Ioc_of_intervalIntegral_norm_bounded_right - end IntegrableOfIntervalIntegral section IntegralOfIntervalIntegral diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 338687b657efa..67539055e3558 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -657,9 +657,6 @@ nonrec theorem _root_.RCLike.intervalIntegral_ofReal {𝕜 : Type*} [RCLike 𝕜 {μ : Measure ℝ} {f : ℝ → ℝ} : (∫ x in a..b, (f x : 𝕜) ∂μ) = ↑(∫ x in a..b, f x ∂μ) := by simp only [intervalIntegral, integral_ofReal, RCLike.ofReal_sub] -@[deprecated (since := "2024-04-06")] -alias RCLike.interval_integral_ofReal := RCLike.intervalIntegral_ofReal - nonrec theorem integral_ofReal {a b : ℝ} {μ : Measure ℝ} {f : ℝ → ℝ} : (∫ x in a..b, (f x : ℂ) ∂μ) = ↑(∫ x in a..b, f x ∂μ) := RCLike.intervalIntegral_ofReal diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index 171fca1f4b81a..10a81f8855d86 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -134,22 +134,13 @@ theorem lintegral_one : ∫⁻ _, (1 : ℝ≥0∞) ∂μ = μ univ := by rw [lin theorem setLIntegral_const (s : Set α) (c : ℝ≥0∞) : ∫⁻ _ in s, c ∂μ = c * μ s := by rw [lintegral_const, Measure.restrict_apply_univ] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_const := setLIntegral_const - theorem setLIntegral_one (s) : ∫⁻ _ in s, 1 ∂μ = μ s := by rw [setLIntegral_const, one_mul] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_one := setLIntegral_one - theorem setLIntegral_const_lt_top [IsFiniteMeasure μ] (s : Set α) {c : ℝ≥0∞} (hc : c ≠ ∞) : ∫⁻ _ in s, c ∂μ < ∞ := by rw [lintegral_const] exact ENNReal.mul_lt_top hc.lt_top (measure_lt_top (μ.restrict s) univ) -@[deprecated (since := "2024-06-29")] -alias set_lintegral_const_lt_top := setLIntegral_const_lt_top - theorem lintegral_const_lt_top [IsFiniteMeasure μ] {c : ℝ≥0∞} (hc : c ≠ ∞) : ∫⁻ _, c ∂μ < ∞ := by simpa only [Measure.restrict_univ] using setLIntegral_const_lt_top (univ : Set α) hc @@ -271,37 +262,22 @@ theorem setLIntegral_mono_ae {s : Set α} {f g : α → ℝ≥0∞} (hg : AEMeas · exact hfg.mono fun x hx hxs ↦ (hle x).trans (hx hxs) · exact nullMeasurableSet_le hf'm.aemeasurable hg -@[deprecated (since := "2024-06-29")] -alias set_lintegral_mono_ae := setLIntegral_mono_ae - theorem setLIntegral_mono {s : Set α} {f g : α → ℝ≥0∞} (hg : Measurable g) (hfg : ∀ x ∈ s, f x ≤ g x) : ∫⁻ x in s, f x ∂μ ≤ ∫⁻ x in s, g x ∂μ := setLIntegral_mono_ae hg.aemeasurable (ae_of_all _ hfg) -@[deprecated (since := "2024-06-29")] -alias set_lintegral_mono := setLIntegral_mono - theorem setLIntegral_mono_ae' {s : Set α} {f g : α → ℝ≥0∞} (hs : MeasurableSet s) (hfg : ∀ᵐ x ∂μ, x ∈ s → f x ≤ g x) : ∫⁻ x in s, f x ∂μ ≤ ∫⁻ x in s, g x ∂μ := lintegral_mono_ae <| (ae_restrict_iff' hs).2 hfg -@[deprecated (since := "2024-06-29")] -alias set_lintegral_mono_ae' := setLIntegral_mono_ae' - theorem setLIntegral_mono' {s : Set α} {f g : α → ℝ≥0∞} (hs : MeasurableSet s) (hfg : ∀ x ∈ s, f x ≤ g x) : ∫⁻ x in s, f x ∂μ ≤ ∫⁻ x in s, g x ∂μ := setLIntegral_mono_ae' hs (ae_of_all _ hfg) -@[deprecated (since := "2024-06-29")] -alias set_lintegral_mono' := setLIntegral_mono' - theorem setLIntegral_le_lintegral (s : Set α) (f : α → ℝ≥0∞) : ∫⁻ x in s, f x ∂μ ≤ ∫⁻ x, f x ∂μ := lintegral_mono' Measure.restrict_le_self le_rfl -@[deprecated (since := "2024-06-29")] -alias set_lintegral_le_lintegral := setLIntegral_le_lintegral - theorem lintegral_congr_ae {f g : α → ℝ≥0∞} (h : f =ᵐ[μ] g) : ∫⁻ a, f a ∂μ = ∫⁻ a, g a ∂μ := le_antisymm (lintegral_mono_ae <| h.le) (lintegral_mono_ae <| h.symm.le) @@ -311,18 +287,12 @@ theorem lintegral_congr {f g : α → ℝ≥0∞} (h : ∀ a, f a = g a) : ∫ theorem setLIntegral_congr {f : α → ℝ≥0∞} {s t : Set α} (h : s =ᵐ[μ] t) : ∫⁻ x in s, f x ∂μ = ∫⁻ x in t, f x ∂μ := by rw [Measure.restrict_congr_set h] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_congr := setLIntegral_congr - theorem setLIntegral_congr_fun {f g : α → ℝ≥0∞} {s : Set α} (hs : MeasurableSet s) (hfg : ∀ᵐ x ∂μ, x ∈ s → f x = g x) : ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ := by rw [lintegral_congr_ae] rw [EventuallyEq] rwa [ae_restrict_iff' hs] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_congr_fun := setLIntegral_congr_fun - theorem lintegral_ofReal_le_lintegral_enorm (f : α → ℝ) : ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ ∫⁻ x, ‖f x‖ₑ ∂μ := by simp_rw [← ofReal_norm_eq_enorm] @@ -502,9 +472,6 @@ theorem exists_pos_setLIntegral_lt_of_measure_lt {f : α → ℝ≥0∞} (h : _ ≤ ε₂ - ε₁ + ε₁ := by gcongr; apply mul_div_le _ = ε₂ := tsub_add_cancel_of_le hε₁₂.le -@[deprecated (since := "2024-06-29")] -alias exists_pos_set_lintegral_lt_of_measure_lt := exists_pos_setLIntegral_lt_of_measure_lt - /-- If `f` has finite integral, then `∫⁻ x in s, f x ∂μ` is absolutely continuous in `s`: it tends to zero as `μ s` tends to zero. -/ theorem tendsto_setLIntegral_zero {ι} {f : α → ℝ≥0∞} (h : ∫⁻ x, f x ∂μ ≠ ∞) {l : Filter ι} @@ -516,9 +483,6 @@ theorem tendsto_setLIntegral_zero {ι} {f : α → ℝ≥0∞} (h : ∫⁻ x, f rcases exists_pos_setLIntegral_lt_of_measure_lt h ε0.ne' with ⟨δ, δ0, hδ⟩ exact (hl δ δ0).mono fun i => hδ _ -@[deprecated (since := "2024-06-29")] -alias tendsto_set_lintegral_zero := tendsto_setLIntegral_zero - /-- The sum of the lower Lebesgue integrals of two functions is less than or equal to the integral of their sum. The other inequality needs one of these functions to be (a.e.-)measurable. -/ theorem le_lintegral_add (f g : α → ℝ≥0∞) : @@ -600,9 +564,6 @@ lemma setLIntegral_smul_measure (c : ℝ≥0∞) (f : α → ℝ≥0∞) (s : Se ∫⁻ a in s, f a ∂(c • μ) = c * ∫⁻ a in s, f a ∂μ := by rw [Measure.restrict_smul, lintegral_smul_measure] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_smul_measure := setLIntegral_smul_measure - @[simp] theorem lintegral_zero_measure {m : MeasurableSpace α} (f : α → ℝ≥0∞) : ∫⁻ a, f a ∂(0 : Measure α) = 0 := by @@ -647,23 +608,14 @@ theorem lintegral_of_isEmpty {α} [MeasurableSpace α] [IsEmpty α] (μ : Measur theorem setLIntegral_empty (f : α → ℝ≥0∞) : ∫⁻ x in ∅, f x ∂μ = 0 := by rw [Measure.restrict_empty, lintegral_zero_measure] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_empty := setLIntegral_empty - theorem setLIntegral_univ (f : α → ℝ≥0∞) : ∫⁻ x in univ, f x ∂μ = ∫⁻ x, f x ∂μ := by rw [Measure.restrict_univ] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_univ := setLIntegral_univ - theorem setLIntegral_measure_zero (s : Set α) (f : α → ℝ≥0∞) (hs' : μ s = 0) : ∫⁻ x in s, f x ∂μ = 0 := by convert lintegral_zero_measure _ exact Measure.restrict_eq_zero.2 hs' -@[deprecated (since := "2024-06-29")] -alias set_lintegral_measure_zero := setLIntegral_measure_zero - theorem lintegral_finset_sum' (s : Finset β) {f : β → α → ℝ≥0∞} (hf : ∀ b ∈ s, AEMeasurable (f b) μ) : ∫⁻ a, ∑ b ∈ s, f b a ∂μ = ∑ b ∈ s, ∫⁻ a, f b a ∂μ := by @@ -828,9 +780,6 @@ theorem setLIntegral_eq_const {f : α → ℝ≥0∞} (hf : Measurable f) (r : · rw [lintegral_const, Measure.restrict_apply MeasurableSet.univ, Set.univ_inter] · exact hf (measurableSet_singleton r) -@[deprecated (since := "2024-06-29")] -alias set_lintegral_eq_const := setLIntegral_eq_const - theorem lintegral_indicator_one_le (s : Set α) : ∫⁻ a, s.indicator 1 a ∂μ ≤ μ s := (lintegral_indicator_const_le _ _).trans <| (one_mul _).le @@ -1039,9 +988,6 @@ theorem setLIntegral_strict_mono {f g : α → ℝ≥0∞} {s : Set α} (hsm : M (h : ∀ᵐ x ∂μ, x ∈ s → f x < g x) : ∫⁻ x in s, f x ∂μ < ∫⁻ x in s, g x ∂μ := lintegral_strict_mono (by simp [hs]) hg.aemeasurable hfi ((ae_restrict_iff' hsm).mpr h) -@[deprecated (since := "2024-06-29")] -alias set_lintegral_strict_mono := setLIntegral_strict_mono - /-- Monotone convergence theorem for nonincreasing sequences of functions -/ theorem lintegral_iInf_ae {f : ℕ → α → ℝ≥0∞} (h_meas : ∀ n, Measurable (f n)) (h_mono : ∀ n : ℕ, f n.succ ≤ᵐ[μ] f n) (h_fin : ∫⁻ a, f 0 a ∂μ ≠ ∞) : @@ -1395,9 +1341,6 @@ theorem setLIntegral_max {f g : α → ℝ≥0∞} (hf : Measurable f) (hg : Mea rw [lintegral_max hf hg, restrict_restrict, restrict_restrict, inter_comm s, inter_comm s] exacts [measurableSet_lt hg hf, measurableSet_le hf hg] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_max := setLIntegral_max - theorem lintegral_map {mβ : MeasurableSpace β} {f : β → ℝ≥0∞} {g : α → β} (hf : Measurable f) (hg : Measurable g) : ∫⁻ a, f a ∂map g μ = ∫⁻ a, f (g a) ∂μ := by erw [lintegral_eq_iSup_eapprox_lintegral hf, lintegral_eq_iSup_eapprox_lintegral (hf.comp hg)] @@ -1434,9 +1377,6 @@ theorem setLIntegral_map [MeasurableSpace β] {f : β → ℝ≥0∞} {g : α ∫⁻ y in s, f y ∂map g μ = ∫⁻ x in g ⁻¹' s, f (g x) ∂μ := by rw [restrict_map hg hs, lintegral_map hf hg] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_map := setLIntegral_map - theorem lintegral_indicator_const_comp {mβ : MeasurableSpace β} {f : α → β} {s : Set β} (hf : Measurable f) (hs : MeasurableSet s) (c : ℝ≥0∞) : ∫⁻ a, s.indicator (fun _ => c) (f a) ∂μ = c * μ (f ⁻¹' s) := by @@ -1483,26 +1423,16 @@ theorem MeasurePreserving.setLIntegral_comp_preimage {mb : MeasurableSpace β} { (hf : Measurable f) : ∫⁻ a in g ⁻¹' s, f (g a) ∂μ = ∫⁻ b in s, f b ∂ν := by rw [← hg.map_eq, setLIntegral_map hs hf hg.measurable] -@[deprecated (since := "2024-06-29")] -alias MeasurePreserving.set_lintegral_comp_preimage := MeasurePreserving.setLIntegral_comp_preimage - theorem MeasurePreserving.setLIntegral_comp_preimage_emb {mb : MeasurableSpace β} {ν : Measure β} {g : α → β} (hg : MeasurePreserving g μ ν) (hge : MeasurableEmbedding g) (f : β → ℝ≥0∞) (s : Set β) : ∫⁻ a in g ⁻¹' s, f (g a) ∂μ = ∫⁻ b in s, f b ∂ν := by rw [← hg.map_eq, hge.restrict_map, hge.lintegral_map] -@[deprecated (since := "2024-06-29")] -alias MeasurePreserving.set_lintegral_comp_preimage_emb := - MeasurePreserving.setLIntegral_comp_preimage_emb - theorem MeasurePreserving.setLIntegral_comp_emb {mb : MeasurableSpace β} {ν : Measure β} {g : α → β} (hg : MeasurePreserving g μ ν) (hge : MeasurableEmbedding g) (f : β → ℝ≥0∞) (s : Set α) : ∫⁻ a in s, f (g a) ∂μ = ∫⁻ b in g '' s, f b ∂ν := by rw [← hg.setLIntegral_comp_preimage_emb hge, preimage_image_eq _ hge.injective] -@[deprecated (since := "2024-06-29")] -alias MeasurePreserving.set_lintegral_comp_emb := MeasurePreserving.setLIntegral_comp_emb - theorem lintegral_subtype_comap {s : Set α} (hs : MeasurableSet s) (f : α → ℝ≥0∞) : ∫⁻ x : s, f x ∂(μ.comap (↑)) = ∫⁻ x in s, f x ∂μ := by rw [← (MeasurableEmbedding.subtype_coe hs).lintegral_map, map_comap_subtype_coe hs] @@ -1558,9 +1488,6 @@ theorem exists_measurable_le_setLintegral_eq_of_integrable {f : α → ℝ≥0 end UnifTight -@[deprecated (since := "2024-06-29")] -alias set_lintegral_subtype := setLIntegral_subtype - section DiracAndCount variable [MeasurableSpace α] @@ -1578,9 +1505,6 @@ theorem setLIntegral_dirac' {a : α} {f : α → ℝ≥0∞} (hf : Measurable f) · exact lintegral_dirac' _ hf · exact lintegral_zero_measure _ -@[deprecated (since := "2024-06-29")] -alias set_lintegral_dirac' := setLIntegral_dirac' - theorem setLIntegral_dirac {a : α} (f : α → ℝ≥0∞) (s : Set α) [MeasurableSingletonClass α] [Decidable (a ∈ s)] : ∫⁻ x in s, f x ∂Measure.dirac a = if a ∈ s then f a else 0 := by rw [restrict_dirac] @@ -1588,9 +1512,6 @@ theorem setLIntegral_dirac {a : α} (f : α → ℝ≥0∞) (s : Set α) [Measur · exact lintegral_dirac _ _ · exact lintegral_zero_measure _ -@[deprecated (since := "2024-06-29")] -alias set_lintegral_dirac := setLIntegral_dirac - theorem lintegral_count' {f : α → ℝ≥0∞} (hf : Measurable f) : ∫⁻ a, f a ∂count = ∑' a, f a := by rw [count, lintegral_sum_measure] congr @@ -1703,17 +1624,11 @@ theorem setLIntegral_lt_top_of_bddAbove {s : Set α} (hs : μ s ≠ ∞) {f : α setLIntegral_lt_top_of_le_nnreal hs <| hbdd.imp fun _M hM _x hx ↦ ENNReal.coe_le_coe.2 <| hM (mem_image_of_mem f hx) -@[deprecated (since := "2024-06-29")] -alias set_lintegral_lt_top_of_bddAbove := setLIntegral_lt_top_of_bddAbove - theorem setLIntegral_lt_top_of_isCompact [TopologicalSpace α] {s : Set α} (hs : μ s ≠ ∞) (hsc : IsCompact s) {f : α → ℝ≥0} (hf : Continuous f) : ∫⁻ x in s, f x ∂μ < ∞ := setLIntegral_lt_top_of_bddAbove hs (hsc.image hf).bddAbove -@[deprecated (since := "2024-06-29")] -alias set_lintegral_lt_top_of_isCompact := setLIntegral_lt_top_of_isCompact - theorem _root_.IsFiniteMeasure.lintegral_lt_top_of_bounded_to_ennreal {α : Type*} [MeasurableSpace α] (μ : Measure α) [μ_fin : IsFiniteMeasure μ] {f : α → ℝ≥0∞} (f_bdd : ∃ c : ℝ≥0, ∀ x, f x ≤ c) : ∫⁻ x, f x ∂μ < ∞ := by @@ -1978,8 +1893,6 @@ theorem lintegral_le_of_forall_fin_meas_trim_le {μ : Measure α} (hm : m ≤ m0 rw [setLIntegral_iUnion_of_directed] exact directed_of_isDirected_le hS_mono -@[deprecated lintegral_le_of_forall_fin_meas_trim_le (since := "2024-07-14")] -alias lintegral_le_of_forall_fin_meas_le' := lintegral_le_of_forall_fin_meas_trim_le alias lintegral_le_of_forall_fin_meas_le_of_measurable := lintegral_le_of_forall_fin_meas_trim_le /-- If the Lebesgue integral of a function is bounded by some constant on all sets with finite diff --git a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean index b06ba408c13fc..08b1c97eb5fe6 100644 --- a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean +++ b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean @@ -83,9 +83,6 @@ theorem fun_eq_funMulInvSnorm_mul_eLpNorm {p : ℝ} (f : α → ℝ≥0∞) f a = funMulInvSnorm f p μ a * (∫⁻ c, f c ^ p ∂μ) ^ (1 / p) := by simp [funMulInvSnorm, mul_assoc, ENNReal.inv_mul_cancel, hf_nonzero, hf_top] -@[deprecated (since := "2024-07-27")] -alias fun_eq_funMulInvSnorm_mul_snorm := fun_eq_funMulInvSnorm_mul_eLpNorm - theorem funMulInvSnorm_rpow {p : ℝ} (hp0 : 0 < p) {f : α → ℝ≥0∞} {a : α} : funMulInvSnorm f p μ a ^ p = f a ^ p * (∫⁻ c, f c ^ p ∂μ)⁻¹ := by rw [funMulInvSnorm, mul_rpow_of_nonneg _ _ (le_of_lt hp0)] @@ -385,10 +382,6 @@ theorem lintegral_rpow_add_le_add_eLpNorm_mul_lintegral_rpow_add {p q : ℝ} · exact lintegral_mul_rpow_le_lintegral_rpow_mul_lintegral_rpow hpq hf (hf.add hg) hf_top · exact lintegral_mul_rpow_le_lintegral_rpow_mul_lintegral_rpow hpq hg (hf.add hg) hg_top -@[deprecated (since := "2024-07-27")] -alias lintegral_rpow_add_le_add_snorm_mul_lintegral_rpow_add := - lintegral_rpow_add_le_add_eLpNorm_mul_lintegral_rpow_add - private theorem lintegral_Lp_add_le_aux {p q : ℝ} (hpq : p.IsConjExponent q) {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) (hf_top : (∫⁻ a, f a ^ p ∂μ) ≠ ⊤) (hg : AEMeasurable g μ) (hg_top : (∫⁻ a, g a ^ p ∂μ) ≠ ⊤) (h_add_zero : (∫⁻ a, (f + g) a ^ p ∂μ) ≠ 0) diff --git a/Mathlib/MeasureTheory/Integral/Prod.lean b/Mathlib/MeasureTheory/Integral/Prod.lean index 193e5df6d8fb3..6ac3a42bc03f0 100644 --- a/Mathlib/MeasureTheory/Integral/Prod.lean +++ b/Mathlib/MeasureTheory/Integral/Prod.lean @@ -459,8 +459,6 @@ theorem setIntegral_prod (f : α × β → E) {s : Set α} {t : Set β} simp only [← Measure.prod_restrict s t, IntegrableOn] at hf ⊢ exact integral_prod f hf -@[deprecated (since := "2024-04-17")] alias set_integral_prod := setIntegral_prod - theorem integral_prod_smul {𝕜 : Type*} [RCLike 𝕜] [NormedSpace 𝕜 E] (f : α → 𝕜) (g : β → E) : ∫ z, f z.1 • g z.2 ∂μ.prod ν = (∫ x, f x ∂μ) • ∫ y, g y ∂ν := by by_cases hE : CompleteSpace E; swap; · simp [integral, hE] @@ -483,8 +481,6 @@ theorem setIntegral_prod_mul {L : Type*} [RCLike L] (f : α → L) (g : β → L rw [← Measure.prod_restrict s t] apply integral_prod_mul -@[deprecated (since := "2024-04-17")] alias set_integral_prod_mul := setIntegral_prod_mul - theorem integral_fun_snd (f : β → E) : ∫ z, f z.2 ∂μ.prod ν = (μ univ).toReal • ∫ y, f y ∂ν := by simpa using integral_prod_smul (1 : α → ℝ) f diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 05f685c4df4b5..fe4b888207acb 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -71,16 +71,10 @@ theorem setIntegral_congr_ae₀ (hs : NullMeasurableSet s μ) (h : ∀ᵐ x ∂ ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := integral_congr_ae ((ae_restrict_iff'₀ hs).2 h) -@[deprecated (since := "2024-04-17")] -alias set_integral_congr_ae₀ := setIntegral_congr_ae₀ - theorem setIntegral_congr_ae (hs : MeasurableSet s) (h : ∀ᵐ x ∂μ, x ∈ s → f x = g x) : ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := integral_congr_ae ((ae_restrict_iff' hs).2 h) -@[deprecated (since := "2024-04-17")] -alias set_integral_congr_ae := setIntegral_congr_ae - theorem setIntegral_congr_fun₀ (hs : NullMeasurableSet s μ) (h : EqOn f g s) : ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := setIntegral_congr_ae₀ hs <| Eventually.of_forall h @@ -88,9 +82,6 @@ theorem setIntegral_congr_fun₀ (hs : NullMeasurableSet s μ) (h : EqOn f g s) @[deprecated (since := "2024-10-12")] alias setIntegral_congr₀ := setIntegral_congr_fun₀ -@[deprecated (since := "2024-04-17")] -alias set_integral_congr₀ := setIntegral_congr_fun₀ - theorem setIntegral_congr_fun (hs : MeasurableSet s) (h : EqOn f g s) : ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := setIntegral_congr_ae hs <| Eventually.of_forall h @@ -98,18 +89,12 @@ theorem setIntegral_congr_fun (hs : MeasurableSet s) (h : EqOn f g s) : @[deprecated (since := "2024-10-12")] alias setIntegral_congr := setIntegral_congr_fun -@[deprecated (since := "2024-04-17")] -alias set_integral_congr := setIntegral_congr_fun - theorem setIntegral_congr_set (hst : s =ᵐ[μ] t) : ∫ x in s, f x ∂μ = ∫ x in t, f x ∂μ := by rw [Measure.restrict_congr_set hst] @[deprecated (since := "2024-10-12")] alias setIntegral_congr_set_ae := setIntegral_congr_set -@[deprecated (since := "2024-04-17")] -alias set_integral_congr_set_ae := setIntegral_congr_set - theorem integral_union_ae (hst : AEDisjoint μ s t) (ht : NullMeasurableSet t μ) (hfs : IntegrableOn f s μ) (hft : IntegrableOn f t μ) : ∫ x in s ∪ t, f x ∂μ = ∫ x in s, f x ∂μ + ∫ x in t, f x ∂μ := by @@ -203,9 +188,6 @@ theorem setIntegral_indicator (ht : MeasurableSet t) : ∫ x in s, t.indicator f x ∂μ = ∫ x in s ∩ t, f x ∂μ := by rw [integral_indicator ht, Measure.restrict_restrict ht, Set.inter_comm] -@[deprecated (since := "2024-04-17")] -alias set_integral_indicator := setIntegral_indicator - theorem ofReal_setIntegral_one_of_measure_ne_top {X : Type*} {m : MeasurableSpace X} {μ : Measure X} {s : Set X} (hs : μ s ≠ ∞) : ENNReal.ofReal (∫ _ in s, (1 : ℝ) ∂μ) = μ s := calc @@ -215,16 +197,10 @@ theorem ofReal_setIntegral_one_of_measure_ne_top {X : Type*} {m : MeasurableSpac simpa [ofReal_integral_norm_eq_lintegral_enorm (integrableOn_const.2 (.inr hs.lt_top))] _ = μ s := setLIntegral_one _ -@[deprecated (since := "2024-04-17")] -alias ofReal_set_integral_one_of_measure_ne_top := ofReal_setIntegral_one_of_measure_ne_top - theorem ofReal_setIntegral_one {X : Type*} {_ : MeasurableSpace X} (μ : Measure X) [IsFiniteMeasure μ] (s : Set X) : ENNReal.ofReal (∫ _ in s, (1 : ℝ) ∂μ) = μ s := ofReal_setIntegral_one_of_measure_ne_top (measure_ne_top μ s) -@[deprecated (since := "2024-04-17")] -alias ofReal_set_integral_one := ofReal_setIntegral_one - theorem integral_piecewise [DecidablePred (· ∈ s)] (hs : MeasurableSet s) (hf : IntegrableOn f s μ) (hg : IntegrableOn g sᶜ μ) : ∫ x, s.piecewise f g x ∂μ = ∫ x in s, f x ∂μ + ∫ x in sᶜ, g x ∂μ := by @@ -257,9 +233,6 @@ theorem tendsto_setIntegral_of_monotone exacts [tsub_le_iff_tsub_le.mp hi.1, (hi.2.trans_lt <| ENNReal.add_lt_top.2 ⟨hfi', ENNReal.coe_lt_top⟩).ne] -@[deprecated (since := "2024-04-17")] -alias tendsto_set_integral_of_monotone := tendsto_setIntegral_of_monotone - theorem tendsto_setIntegral_of_antitone {ι : Type*} [Preorder ι] [(atTop : Filter ι).IsCountablyGenerated] {s : ι → Set X} (hsm : ∀ i, MeasurableSet (s i)) (h_anti : Antitone s) @@ -280,9 +253,6 @@ theorem tendsto_setIntegral_of_antitone · rw [← diff_iInter] exact hi₀.mono_set diff_subset -@[deprecated (since := "2024-04-17")] -alias tendsto_set_integral_of_antitone := tendsto_setIntegral_of_antitone - theorem hasSum_integral_iUnion_ae {ι : Type*} [Countable ι] {s : ι → Set X} (hm : ∀ i, NullMeasurableSet (s i) μ) (hd : Pairwise (AEDisjoint μ on s)) (hfi : IntegrableOn f (⋃ i, s i) μ) : @@ -323,16 +293,10 @@ theorem setIntegral_eq_zero_of_ae_eq_zero (ht_eq : ∀ᵐ x ∂μ, x ∈ t → f rw [← this] exact integral_congr_ae hf.ae_eq_mk -@[deprecated (since := "2024-04-17")] -alias set_integral_eq_zero_of_ae_eq_zero := setIntegral_eq_zero_of_ae_eq_zero - theorem setIntegral_eq_zero_of_forall_eq_zero (ht_eq : ∀ x ∈ t, f x = 0) : ∫ x in t, f x ∂μ = 0 := setIntegral_eq_zero_of_ae_eq_zero (Eventually.of_forall ht_eq) -@[deprecated (since := "2024-04-17")] -alias set_integral_eq_zero_of_forall_eq_zero := setIntegral_eq_zero_of_forall_eq_zero - theorem integral_union_eq_left_of_ae_aux (ht_eq : ∀ᵐ x ∂μ.restrict t, f x = 0) (haux : StronglyMeasurable f) (H : IntegrableOn f (s ∪ t) μ) : ∫ x in s ∪ t, f x ∂μ = ∫ x in s, f x ∂μ := by @@ -401,10 +365,6 @@ theorem setIntegral_eq_of_subset_of_ae_diff_eq_zero_aux (hts : s ⊆ t) rw [setIntegral_eq_zero_of_forall_eq_zero this, zero_add] _ = ∫ x in s, f x ∂μ := by rw [integral_inter_add_diff hk (h'aux.mono hts le_rfl)] -@[deprecated (since := "2024-04-17")] -alias set_integral_eq_of_subset_of_ae_diff_eq_zero_aux := - setIntegral_eq_of_subset_of_ae_diff_eq_zero_aux - /-- If a function vanishes almost everywhere on `t \ s` with `s ⊆ t`, then its integrals on `s` and `t` coincide if `t` is null-measurable. -/ theorem setIntegral_eq_of_subset_of_ae_diff_eq_zero (ht : NullMeasurableSet t μ) (hts : s ⊆ t) @@ -426,9 +386,6 @@ theorem setIntegral_eq_of_subset_of_ae_diff_eq_zero (ht : NullMeasurableSet t μ apply ae_restrict_of_ae_restrict_of_subset hts exact h.1.ae_eq_mk.symm -@[deprecated (since := "2024-04-17")] -alias set_integral_eq_of_subset_of_ae_diff_eq_zero := setIntegral_eq_of_subset_of_ae_diff_eq_zero - /-- If a function vanishes on `t \ s` with `s ⊆ t`, then its integrals on `s` and `t` coincide if `t` is measurable. -/ theorem setIntegral_eq_of_subset_of_forall_diff_eq_zero (ht : MeasurableSet t) (hts : s ⊆ t) @@ -436,10 +393,6 @@ theorem setIntegral_eq_of_subset_of_forall_diff_eq_zero (ht : MeasurableSet t) ( setIntegral_eq_of_subset_of_ae_diff_eq_zero ht.nullMeasurableSet hts (Eventually.of_forall fun x hx => h't x hx) -@[deprecated (since := "2024-04-17")] -alias set_integral_eq_of_subset_of_forall_diff_eq_zero := - setIntegral_eq_of_subset_of_forall_diff_eq_zero - /-- If a function vanishes almost everywhere on `sᶜ`, then its integral on `s` coincides with its integral on the whole space. -/ theorem setIntegral_eq_integral_of_ae_compl_eq_zero (h : ∀ᵐ x ∂μ, x ∉ s → f x = 0) : @@ -449,19 +402,12 @@ theorem setIntegral_eq_integral_of_ae_compl_eq_zero (h : ∀ᵐ x ∂μ, x ∉ s apply setIntegral_eq_of_subset_of_ae_diff_eq_zero nullMeasurableSet_univ (subset_univ _) filter_upwards [h] with x hx h'x using hx h'x.2 -@[deprecated (since := "2024-04-17")] -alias set_integral_eq_integral_of_ae_compl_eq_zero := setIntegral_eq_integral_of_ae_compl_eq_zero - /-- If a function vanishes on `sᶜ`, then its integral on `s` coincides with its integral on the whole space. -/ theorem setIntegral_eq_integral_of_forall_compl_eq_zero (h : ∀ x, x ∉ s → f x = 0) : ∫ x in s, f x ∂μ = ∫ x, f x ∂μ := setIntegral_eq_integral_of_ae_compl_eq_zero (Eventually.of_forall h) -@[deprecated (since := "2024-04-17")] -alias set_integral_eq_integral_of_forall_compl_eq_zero := - setIntegral_eq_integral_of_forall_compl_eq_zero - theorem setIntegral_neg_eq_setIntegral_nonpos [LinearOrder E] {f : X → E} (hf : AEStronglyMeasurable f μ) : ∫ x in {x | f x < 0}, f x ∂μ = ∫ x in {x | f x ≤ 0}, f x ∂μ := by @@ -474,9 +420,6 @@ theorem setIntegral_neg_eq_setIntegral_nonpos [LinearOrder E] {f : X → E} refine integral_union_eq_left_of_ae ?_ filter_upwards [ae_restrict_mem₀ B] with x hx using hx -@[deprecated (since := "2024-04-17")] -alias set_integral_neg_eq_set_integral_nonpos := setIntegral_neg_eq_setIntegral_nonpos - theorem integral_norm_eq_pos_sub_neg {f : X → ℝ} (hfi : Integrable f μ) : ∫ x, ‖f x‖ ∂μ = ∫ x in {x | 0 ≤ f x}, f x ∂μ - ∫ x in {x | f x ≤ 0}, f x ∂μ := have h_meas : NullMeasurableSet {x | 0 ≤ f x} μ := @@ -504,9 +447,6 @@ theorem integral_norm_eq_pos_sub_neg {f : X → ℝ} (hfi : Integrable f μ) : theorem setIntegral_const [CompleteSpace E] (c : E) : ∫ _ in s, c ∂μ = (μ s).toReal • c := by rw [integral_const, Measure.restrict_apply_univ] -@[deprecated (since := "2024-04-17")] -alias set_integral_const := setIntegral_const - @[simp] theorem integral_indicator_const [CompleteSpace E] (e : E) ⦃s : Set X⦄ (s_meas : MeasurableSet s) : ∫ x : X, s.indicator (fun _ : X => e) x ∂μ = (μ s).toReal • e := by @@ -525,9 +465,6 @@ theorem setIntegral_indicatorConstLp [CompleteSpace E] rw [setIntegral_congr_ae hs (indicatorConstLp_coeFn.mono fun x hx _ => hx)] _ = (μ (t ∩ s)).toReal • e := by rw [integral_indicator_const _ ht, Measure.restrict_apply ht] -@[deprecated (since := "2024-04-17")] -alias set_integral_indicatorConstLp := setIntegral_indicatorConstLp - theorem integral_indicatorConstLp [CompleteSpace E] {p : ℝ≥0∞} (ht : MeasurableSet t) (hμt : μ t ≠ ∞) (e : E) : ∫ x, indicatorConstLp p ht hμt e x ∂μ = (μ t).toReal • e := @@ -544,17 +481,11 @@ theorem setIntegral_map {Y} [MeasurableSpace Y] {g : X → Y} {f : Y → E} {s : integral_map (hg.mono_measure Measure.restrict_le_self) (hf.mono_measure _)] exact Measure.map_mono_of_aemeasurable Measure.restrict_le_self hg -@[deprecated (since := "2024-04-17")] -alias set_integral_map := setIntegral_map - theorem _root_.MeasurableEmbedding.setIntegral_map {Y} {_ : MeasurableSpace Y} {f : X → Y} (hf : MeasurableEmbedding f) (g : Y → E) (s : Set Y) : ∫ y in s, g y ∂Measure.map f μ = ∫ x in f ⁻¹' s, g (f x) ∂μ := by rw [hf.restrict_map, hf.integral_map] -@[deprecated (since := "2024-04-17")] -alias _root_.MeasurableEmbedding.set_integral_map := _root_.MeasurableEmbedding.setIntegral_map - theorem _root_.Topology.IsClosedEmbedding.setIntegral_map [TopologicalSpace X] [BorelSpace X] {Y} [MeasurableSpace Y] [TopologicalSpace Y] [BorelSpace Y] {g : X → Y} {f : Y → E} (s : Set Y) (hg : IsClosedEmbedding g) : ∫ y in s, f y ∂Measure.map g μ = ∫ x in g ⁻¹' s, f (g x) ∂μ := @@ -563,45 +494,26 @@ theorem _root_.Topology.IsClosedEmbedding.setIntegral_map [TopologicalSpace X] [ @[deprecated (since := "2024-10-20")] alias _root_.ClosedEmbedding.setIntegral_map := IsClosedEmbedding.setIntegral_map -@[deprecated (since := "2024-04-17")] -alias _root_.IsClosedEmbedding.set_integral_map := - IsClosedEmbedding.setIntegral_map - -@[deprecated (since := "2024-10-20")] -alias _root_.ClosedEmbedding.set_integral_map := IsClosedEmbedding.set_integral_map - theorem MeasurePreserving.setIntegral_preimage_emb {Y} {_ : MeasurableSpace Y} {f : X → Y} {ν} (h₁ : MeasurePreserving f μ ν) (h₂ : MeasurableEmbedding f) (g : Y → E) (s : Set Y) : ∫ x in f ⁻¹' s, g (f x) ∂μ = ∫ y in s, g y ∂ν := (h₁.restrict_preimage_emb h₂ s).integral_comp h₂ _ -@[deprecated (since := "2024-04-17")] -alias MeasurePreserving.set_integral_preimage_emb := MeasurePreserving.setIntegral_preimage_emb - theorem MeasurePreserving.setIntegral_image_emb {Y} {_ : MeasurableSpace Y} {f : X → Y} {ν} (h₁ : MeasurePreserving f μ ν) (h₂ : MeasurableEmbedding f) (g : Y → E) (s : Set X) : ∫ y in f '' s, g y ∂ν = ∫ x in s, g (f x) ∂μ := Eq.symm <| (h₁.restrict_image_emb h₂ s).integral_comp h₂ _ -@[deprecated (since := "2024-04-17")] -alias MeasurePreserving.set_integral_image_emb := MeasurePreserving.setIntegral_image_emb - theorem setIntegral_map_equiv {Y} [MeasurableSpace Y] (e : X ≃ᵐ Y) (f : Y → E) (s : Set Y) : ∫ y in s, f y ∂Measure.map e μ = ∫ x in e ⁻¹' s, f (e x) ∂μ := e.measurableEmbedding.setIntegral_map f s -@[deprecated (since := "2024-04-17")] -alias set_integral_map_equiv := setIntegral_map_equiv - theorem norm_setIntegral_le_of_norm_le_const_ae {C : ℝ} (hs : μ s < ∞) (hC : ∀ᵐ x ∂μ.restrict s, ‖f x‖ ≤ C) : ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).toReal := by rw [← Measure.restrict_apply_univ] at * haveI : IsFiniteMeasure (μ.restrict s) := ⟨hs⟩ exact norm_integral_le_of_norm_le_const hC -@[deprecated (since := "2024-04-17")] -alias norm_set_integral_le_of_norm_le_const_ae := norm_setIntegral_le_of_norm_le_const_ae - theorem norm_setIntegral_le_of_norm_le_const_ae' {C : ℝ} (hs : μ s < ∞) (hC : ∀ᵐ x ∂μ, x ∈ s → ‖f x‖ ≤ C) (hfm : AEStronglyMeasurable f (μ.restrict s)) : ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).toReal := by @@ -615,47 +527,29 @@ theorem norm_setIntegral_le_of_norm_le_const_ae' {C : ℝ} (hs : μ s < ∞) filter_upwards [hfm.ae_eq_mk, (ae_restrict_iff B).2 A] with _ h1 _ rwa [h1] -@[deprecated (since := "2024-04-17")] -alias norm_set_integral_le_of_norm_le_const_ae' := norm_setIntegral_le_of_norm_le_const_ae' - theorem norm_setIntegral_le_of_norm_le_const_ae'' {C : ℝ} (hs : μ s < ∞) (hsm : MeasurableSet s) (hC : ∀ᵐ x ∂μ, x ∈ s → ‖f x‖ ≤ C) : ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).toReal := norm_setIntegral_le_of_norm_le_const_ae hs <| by rwa [ae_restrict_eq hsm, eventually_inf_principal] -@[deprecated (since := "2024-04-17")] -alias norm_set_integral_le_of_norm_le_const_ae'' := norm_setIntegral_le_of_norm_le_const_ae'' - theorem norm_setIntegral_le_of_norm_le_const {C : ℝ} (hs : μ s < ∞) (hC : ∀ x ∈ s, ‖f x‖ ≤ C) (hfm : AEStronglyMeasurable f (μ.restrict s)) : ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).toReal := norm_setIntegral_le_of_norm_le_const_ae' hs (Eventually.of_forall hC) hfm -@[deprecated (since := "2024-04-17")] -alias norm_set_integral_le_of_norm_le_const := norm_setIntegral_le_of_norm_le_const - theorem norm_setIntegral_le_of_norm_le_const' {C : ℝ} (hs : μ s < ∞) (hsm : MeasurableSet s) (hC : ∀ x ∈ s, ‖f x‖ ≤ C) : ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).toReal := norm_setIntegral_le_of_norm_le_const_ae'' hs hsm <| Eventually.of_forall hC -@[deprecated (since := "2024-04-17")] -alias norm_set_integral_le_of_norm_le_const' := norm_setIntegral_le_of_norm_le_const' - theorem setIntegral_eq_zero_iff_of_nonneg_ae {f : X → ℝ} (hf : 0 ≤ᵐ[μ.restrict s] f) (hfi : IntegrableOn f s μ) : ∫ x in s, f x ∂μ = 0 ↔ f =ᵐ[μ.restrict s] 0 := integral_eq_zero_iff_of_nonneg_ae hf hfi -@[deprecated (since := "2024-04-17")] -alias set_integral_eq_zero_iff_of_nonneg_ae := setIntegral_eq_zero_iff_of_nonneg_ae - theorem setIntegral_pos_iff_support_of_nonneg_ae {f : X → ℝ} (hf : 0 ≤ᵐ[μ.restrict s] f) (hfi : IntegrableOn f s μ) : (0 < ∫ x in s, f x ∂μ) ↔ 0 < μ (support f ∩ s) := by rw [integral_pos_iff_support_of_nonneg_ae hf hfi, Measure.restrict_apply₀] rw [support_eq_preimage] exact hfi.aestronglyMeasurable.aemeasurable.nullMeasurable (measurableSet_singleton 0).compl -@[deprecated (since := "2024-04-17")] -alias set_integral_pos_iff_support_of_nonneg_ae := setIntegral_pos_iff_support_of_nonneg_ae - theorem setIntegral_gt_gt {R : ℝ} {f : X → ℝ} (hR : 0 ≤ R) (hfint : IntegrableOn f {x | ↑R < f x} μ) (hμ : μ {x | ↑R < f x} ≠ 0) : (μ {x | ↑R < f x}).toReal * R < ∫ x in {x | ↑R < f x}, f x ∂μ := by @@ -675,17 +569,11 @@ theorem setIntegral_gt_gt {R : ℝ} {f : X → ℝ} (hR : 0 ≤ R) · exact nullMeasurableSet_le aemeasurable_zero (hfint.1.aemeasurable.sub aemeasurable_const) · exact Integrable.sub hfint this -@[deprecated (since := "2024-04-17")] -alias set_integral_gt_gt := setIntegral_gt_gt - theorem setIntegral_trim {X} {m m0 : MeasurableSpace X} {μ : Measure X} (hm : m ≤ m0) {f : X → E} (hf_meas : StronglyMeasurable[m] f) {s : Set X} (hs : MeasurableSet[m] s) : ∫ x in s, f x ∂μ = ∫ x in s, f x ∂μ.trim hm := by rwa [integral_trim hm hf_meas, restrict_trim hm μ] -@[deprecated (since := "2024-04-17")] -alias set_integral_trim := setIntegral_trim - /-! ### Lemmas about adding and removing interval boundaries The primed lemmas take explicit arguments about the endpoint having zero measure, while the @@ -763,61 +651,37 @@ theorem setIntegral_mono_ae_restrict (h : f ≤ᵐ[μ.restrict s] g) : ∫ x in s, f x ∂μ ≤ ∫ x in s, g x ∂μ := integral_mono_ae hf hg h -@[deprecated (since := "2024-04-17")] -alias set_integral_mono_ae_restrict := setIntegral_mono_ae_restrict - theorem setIntegral_mono_ae (h : f ≤ᵐ[μ] g) : ∫ x in s, f x ∂μ ≤ ∫ x in s, g x ∂μ := setIntegral_mono_ae_restrict hf hg (ae_restrict_of_ae h) -@[deprecated (since := "2024-04-17")] -alias set_integral_mono_ae := setIntegral_mono_ae - theorem setIntegral_mono_on (hs : MeasurableSet s) (h : ∀ x ∈ s, f x ≤ g x) : ∫ x in s, f x ∂μ ≤ ∫ x in s, g x ∂μ := setIntegral_mono_ae_restrict hf hg (by simp [hs, EventuallyLE, eventually_inf_principal, ae_of_all _ h]) -@[deprecated (since := "2024-04-17")] -alias set_integral_mono_on := setIntegral_mono_on - theorem setIntegral_mono_on_ae (hs : MeasurableSet s) (h : ∀ᵐ x ∂μ, x ∈ s → f x ≤ g x) : ∫ x in s, f x ∂μ ≤ ∫ x in s, g x ∂μ := by refine setIntegral_mono_ae_restrict hf hg ?_; rwa [EventuallyLE, ae_restrict_iff' hs] -@[deprecated (since := "2024-04-17")] -alias set_integral_mono_on_ae := setIntegral_mono_on_ae - theorem setIntegral_mono (h : f ≤ g) : ∫ x in s, f x ∂μ ≤ ∫ x in s, g x ∂μ := integral_mono hf hg h -@[deprecated (since := "2024-04-17")] -alias set_integral_mono := setIntegral_mono - end theorem setIntegral_mono_set (hfi : IntegrableOn f t μ) (hf : 0 ≤ᵐ[μ.restrict t] f) (hst : s ≤ᵐ[μ] t) : ∫ x in s, f x ∂μ ≤ ∫ x in t, f x ∂μ := integral_mono_measure (Measure.restrict_mono_ae hst) hf hfi -@[deprecated (since := "2024-04-17")] -alias set_integral_mono_set := setIntegral_mono_set - theorem setIntegral_le_integral (hfi : Integrable f μ) (hf : 0 ≤ᵐ[μ] f) : ∫ x in s, f x ∂μ ≤ ∫ x, f x ∂μ := integral_mono_measure (Measure.restrict_le_self) hf hfi -@[deprecated (since := "2024-04-17")] -alias set_integral_le_integral := setIntegral_le_integral - theorem setIntegral_ge_of_const_le {c : ℝ} (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (hf : ∀ x ∈ s, c ≤ f x) (hfint : IntegrableOn (fun x : X => f x) s μ) : c * (μ s).toReal ≤ ∫ x in s, f x ∂μ := by rw [mul_comm, ← smul_eq_mul, ← setIntegral_const c] exact setIntegral_mono_on (integrableOn_const.2 (Or.inr hμs.lt_top)) hfint hs hf -@[deprecated (since := "2024-04-17")] -alias set_integral_ge_of_const_le := setIntegral_ge_of_const_le - end Mono section Nonneg @@ -827,29 +691,17 @@ variable {μ : Measure X} {f : X → ℝ} {s : Set X} theorem setIntegral_nonneg_of_ae_restrict (hf : 0 ≤ᵐ[μ.restrict s] f) : 0 ≤ ∫ x in s, f x ∂μ := integral_nonneg_of_ae hf -@[deprecated (since := "2024-04-17")] -alias set_integral_nonneg_of_ae_restrict := setIntegral_nonneg_of_ae_restrict - theorem setIntegral_nonneg_of_ae (hf : 0 ≤ᵐ[μ] f) : 0 ≤ ∫ x in s, f x ∂μ := setIntegral_nonneg_of_ae_restrict (ae_restrict_of_ae hf) -@[deprecated (since := "2024-04-17")] -alias set_integral_nonneg_of_ae := setIntegral_nonneg_of_ae - theorem setIntegral_nonneg (hs : MeasurableSet s) (hf : ∀ x, x ∈ s → 0 ≤ f x) : 0 ≤ ∫ x in s, f x ∂μ := setIntegral_nonneg_of_ae_restrict ((ae_restrict_iff' hs).mpr (ae_of_all μ hf)) -@[deprecated (since := "2024-04-17")] -alias set_integral_nonneg := setIntegral_nonneg - theorem setIntegral_nonneg_ae (hs : MeasurableSet s) (hf : ∀ᵐ x ∂μ, x ∈ s → 0 ≤ f x) : 0 ≤ ∫ x in s, f x ∂μ := setIntegral_nonneg_of_ae_restrict <| by rwa [EventuallyLE, ae_restrict_iff' hs] -@[deprecated (since := "2024-04-17")] -alias set_integral_nonneg_ae := setIntegral_nonneg_ae - theorem setIntegral_le_nonneg {s : Set X} (hs : MeasurableSet s) (hf : StronglyMeasurable f) (hfi : Integrable f μ) : ∫ x in s, f x ∂μ ≤ ∫ x in {y | 0 ≤ f y}, f x ∂μ := by rw [← integral_indicator hs, ← @@ -859,35 +711,20 @@ theorem setIntegral_le_nonneg {s : Set X} (hs : MeasurableSet s) (hf : StronglyM (hfi.indicator (stronglyMeasurable_const.measurableSet_le hf)) (indicator_le_indicator_nonneg s f) -@[deprecated (since := "2024-04-17")] -alias set_integral_le_nonneg := setIntegral_le_nonneg - theorem setIntegral_nonpos_of_ae_restrict (hf : f ≤ᵐ[μ.restrict s] 0) : ∫ x in s, f x ∂μ ≤ 0 := integral_nonpos_of_ae hf -@[deprecated (since := "2024-04-17")] -alias set_integral_nonpos_of_ae_restrict := setIntegral_nonpos_of_ae_restrict - theorem setIntegral_nonpos_of_ae (hf : f ≤ᵐ[μ] 0) : ∫ x in s, f x ∂μ ≤ 0 := setIntegral_nonpos_of_ae_restrict (ae_restrict_of_ae hf) -@[deprecated (since := "2024-04-17")] -alias set_integral_nonpos_of_ae := setIntegral_nonpos_of_ae - theorem setIntegral_nonpos_ae (hs : MeasurableSet s) (hf : ∀ᵐ x ∂μ, x ∈ s → f x ≤ 0) : ∫ x in s, f x ∂μ ≤ 0 := setIntegral_nonpos_of_ae_restrict <| by rwa [EventuallyLE, ae_restrict_iff' hs] -@[deprecated (since := "2024-04-17")] -alias set_integral_nonpos_ae := setIntegral_nonpos_ae - theorem setIntegral_nonpos (hs : MeasurableSet s) (hf : ∀ x, x ∈ s → f x ≤ 0) : ∫ x in s, f x ∂μ ≤ 0 := setIntegral_nonpos_ae hs <| ae_of_all μ hf -@[deprecated (since := "2024-04-17")] -alias set_integral_nonpos := setIntegral_nonpos - theorem setIntegral_nonpos_le {s : Set X} (hs : MeasurableSet s) (hf : StronglyMeasurable f) (hfi : Integrable f μ) : ∫ x in {y | f y ≤ 0}, f x ∂μ ≤ ∫ x in s, f x ∂μ := by rw [← integral_indicator hs, ← @@ -896,9 +733,6 @@ theorem setIntegral_nonpos_le {s : Set X} (hs : MeasurableSet s) (hf : StronglyM integral_mono (hfi.indicator (hf.measurableSet_le stronglyMeasurable_const)) (hfi.indicator hs) (indicator_nonpos_le_indicator s f) -@[deprecated (since := "2024-04-17")] -alias set_integral_nonpos_le := setIntegral_nonpos_le - lemma Integrable.measure_le_integral {f : X → ℝ} (f_int : Integrable f μ) (f_nonneg : 0 ≤ᵐ[μ] f) {s : Set X} (hs : ∀ x ∈ s, 1 ≤ f x) : μ s ≤ ENNReal.ofReal (∫ x, f x ∂μ) := by @@ -1050,9 +884,6 @@ theorem continuous_setIntegral [NormedSpace ℝ E] (s : Set X) : rw [h_comp] exact continuous_integral.comp (LpToLpRestrictCLM X E ℝ μ 1 s).continuous -@[deprecated (since := "2024-04-17")] -alias continuous_set_integral := continuous_setIntegral - end ContinuousSetIntegral end MeasureTheory @@ -1216,9 +1047,6 @@ theorem setIntegral_compLp (L : E →L[𝕜] F) (φ : Lp E p μ) {s : Set X} (hs ∫ x in s, (L.compLp φ) x ∂μ = ∫ x in s, L (φ x) ∂μ := setIntegral_congr_ae hs ((L.coeFn_compLp φ).mono fun _x hx _ => hx) -@[deprecated (since := "2024-04-17")] -alias set_integral_compLp := setIntegral_compLp - theorem continuous_integral_comp_L1 (L : E →L[𝕜] F) : Continuous fun φ : X →₁[μ] E => ∫ x : X, L (φ x) ∂μ := by rw [← funext L.integral_compLp]; exact continuous_integral.comp (L.compLpL 1 μ).continuous @@ -1365,9 +1193,6 @@ theorem setIntegral_re_add_im {f : X → 𝕜} {i : Set X} (hf : IntegrableOn f ∫ x in i, f x ∂μ := integral_re_add_im hf -@[deprecated (since := "2024-04-17")] -alias set_integral_re_add_im := setIntegral_re_add_im - variable [NormedSpace ℝ E] [NormedSpace ℝ F] lemma swap_integral (f : X → E × F) : (∫ x, f x ∂μ).swap = ∫ x, (f x).swap ∂μ := @@ -1466,26 +1291,16 @@ theorem setIntegral_withDensity_eq_setIntegral_smul {f : X → ℝ≥0} (f_meas ∫ x in s, g x ∂μ.withDensity (fun x => f x) = ∫ x in s, f x • g x ∂μ := by rw [restrict_withDensity hs, integral_withDensity_eq_integral_smul f_meas] -@[deprecated (since := "2024-04-17")] -alias set_integral_withDensity_eq_set_integral_smul := setIntegral_withDensity_eq_setIntegral_smul - theorem setIntegral_withDensity_eq_setIntegral_smul₀ {f : X → ℝ≥0} {s : Set X} (hf : AEMeasurable f (μ.restrict s)) (g : X → E) (hs : MeasurableSet s) : ∫ x in s, g x ∂μ.withDensity (fun x => f x) = ∫ x in s, f x • g x ∂μ := by rw [restrict_withDensity hs, integral_withDensity_eq_integral_smul₀ hf] -@[deprecated (since := "2024-04-17")] -alias set_integral_withDensity_eq_set_integral_smul₀ := setIntegral_withDensity_eq_setIntegral_smul₀ - theorem setIntegral_withDensity_eq_setIntegral_smul₀' [SFinite μ] {f : X → ℝ≥0} (s : Set X) (hf : AEMeasurable f (μ.restrict s)) (g : X → E) : ∫ x in s, g x ∂μ.withDensity (fun x => f x) = ∫ x in s, f x • g x ∂μ := by rw [restrict_withDensity' s, integral_withDensity_eq_integral_smul₀ hf] -@[deprecated (since := "2024-04-17")] -alias set_integral_withDensity_eq_set_integral_smul₀' := - setIntegral_withDensity_eq_setIntegral_smul₀' - end section thickenedIndicator @@ -1654,5 +1469,3 @@ lemma continuousOn_integral_of_compact_support hk hf hfs (integrableOn_const.2 (Or.inr hk.measure_lt_top)) (μ := μ) (g := fun _ ↦ 1) end ParametricIntegral - -set_option linter.style.longFile 1700 diff --git a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean index 1e44f801378f7..7566f6ade9f92 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean @@ -141,16 +141,10 @@ theorem setIntegral_comp_smul (f : E → F) {R : ℝ} (s : Set E) (hR : R ≠ 0) rw [mem_smul_set_iff_inv_smul_mem₀ hR] rfl -@[deprecated (since := "2024-04-17")] -alias set_integral_comp_smul := setIntegral_comp_smul - theorem setIntegral_comp_smul_of_pos (f : E → F) {R : ℝ} (s : Set E) (hR : 0 < R) : ∫ x in s, f (R • x) ∂μ = (R ^ finrank ℝ E)⁻¹ • ∫ x in R • s, f x ∂μ := by rw [setIntegral_comp_smul μ f s hR.ne', abs_of_nonneg (inv_nonneg.2 (pow_nonneg hR.le _))] -@[deprecated (since := "2024-04-17")] -alias set_integral_comp_smul_of_pos := setIntegral_comp_smul_of_pos - theorem integral_comp_mul_left (g : ℝ → F) (a : ℝ) : (∫ x : ℝ, g (a * x)) = |a⁻¹| • ∫ y : ℝ, g y := by simp_rw [← smul_eq_mul, Measure.integral_comp_smul, Module.finrank_self, pow_one] diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean index fce565dd079d5..5625c857e3833 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean @@ -342,11 +342,6 @@ theorem volume_closedBall (x : EuclideanSpace ℝ ι) (r : ℝ) : .ofReal (sqrt π ^ card ι / Gamma (card ι / 2 + 1)) := by rw [addHaar_closedBall_eq_addHaar_ball, EuclideanSpace.volume_ball] -@[deprecated (since := "2024-04-06")] -alias Euclidean_space.volume_ball := EuclideanSpace.volume_ball -@[deprecated (since := "2024-04-06")] -alias Euclidean_space.volume_closedBall := EuclideanSpace.volume_closedBall - end EuclideanSpace namespace InnerProductSpace diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 54c1d8c77d045..a48e192e94400 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -1242,8 +1242,6 @@ theorem sum_cond (μ ν : Measure α) : (sum fun b => cond b μ ν) = μ + ν := theorem sum_of_isEmpty [IsEmpty ι] (μ : ι → Measure α) : sum μ = 0 := by rw [← measure_univ_eq_zero, sum_apply _ MeasurableSet.univ, tsum_empty] -@[deprecated (since := "2024-06-11")] alias sum_of_empty := sum_of_isEmpty - theorem sum_add_sum_compl (s : Set ι) (μ : ι → Measure α) : ((sum fun i : s => μ i) + sum fun i : ↥sᶜ => μ i) = sum μ := by ext1 t ht diff --git a/Mathlib/MeasureTheory/Measure/Tilted.lean b/Mathlib/MeasureTheory/Measure/Tilted.lean index a599a9994646b..2dbb796cfa154 100644 --- a/Mathlib/MeasureTheory/Measure/Tilted.lean +++ b/Mathlib/MeasureTheory/Measure/Tilted.lean @@ -161,9 +161,6 @@ lemma setLIntegral_tilted' (f : α → ℝ) (g : α → ℝ≥0∞) {s : Set α} rw [integral_undef hf'] simp -@[deprecated (since := "2024-06-29")] -alias set_lintegral_tilted' := setLIntegral_tilted' - lemma setLIntegral_tilted [SFinite μ] (f : α → ℝ) (g : α → ℝ≥0∞) (s : Set α) : ∫⁻ x in s, g x ∂(μ.tilted f) = ∫⁻ x in s, ENNReal.ofReal (exp (f x) / ∫ x, exp (f x) ∂μ) * g x ∂μ := by @@ -181,9 +178,6 @@ lemma setLIntegral_tilted [SFinite μ] (f : α → ℝ) (g : α → ℝ≥0∞) rw [integral_undef hf'] simp -@[deprecated (since := "2024-06-29")] -alias set_lintegral_tilted := setLIntegral_tilted - lemma lintegral_tilted (f : α → ℝ) (g : α → ℝ≥0∞) : ∫⁻ x, g x ∂(μ.tilted f) = ∫⁻ x, ENNReal.ofReal (exp (f x) / ∫ x, exp (f x) ∂μ) * (g x) ∂μ := by @@ -212,9 +206,6 @@ lemma setIntegral_tilted' (f : α → ℝ) (g : α → E) {s : Set α} (hs : Mea rw [integral_undef hf'] simp -@[deprecated (since := "2024-04-17")] -alias set_integral_tilted' := setIntegral_tilted' - lemma setIntegral_tilted [SFinite μ] (f : α → ℝ) (g : α → E) (s : Set α) : ∫ x in s, g x ∂(μ.tilted f) = ∫ x in s, (exp (f x) / ∫ x, exp (f x) ∂μ) • (g x) ∂μ := by by_cases hf : AEMeasurable f μ @@ -232,9 +223,6 @@ lemma setIntegral_tilted [SFinite μ] (f : α → ℝ) (g : α → E) (s : Set rw [integral_undef hf'] simp -@[deprecated (since := "2024-04-17")] -alias set_integral_tilted := setIntegral_tilted - lemma integral_tilted (f : α → ℝ) (g : α → E) : ∫ x, g x ∂(μ.tilted f) = ∫ x, (exp (f x) / ∫ x, exp (f x) ∂μ) • (g x) ∂μ := by rw [← setIntegral_univ, setIntegral_tilted' f g MeasurableSet.univ, setIntegral_univ] diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index cb4284f411de1..acb64f67292dd 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -1303,9 +1303,6 @@ theorem exists_pos_ball [PseudoMetricSpace α] (x : α) (hμ : μ ≠ 0) : /-- If a set has zero measure in a neighborhood of each of its points, then it has zero measure in a second-countable space. -/ -@[deprecated (since := "2024-05-14")] -alias null_of_locally_null := measure_null_of_locally_null - theorem exists_ne_forall_mem_nhds_pos_measure_preimage {β} [TopologicalSpace β] [T1Space β] [SecondCountableTopology β] [Nonempty β] {f : α → β} (h : ∀ b, ∃ᵐ x ∂μ, f x ≠ b) : ∃ a b : β, a ≠ b ∧ (∀ s ∈ 𝓝 a, 0 < μ (f ⁻¹' s)) ∧ ∀ t ∈ 𝓝 b, 0 < μ (f ⁻¹' t) := by diff --git a/Mathlib/MeasureTheory/Measure/WithDensity.lean b/Mathlib/MeasureTheory/Measure/WithDensity.lean index f2b8b222f5c26..87ec89bf6b0dd 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensity.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensity.lean @@ -360,9 +360,6 @@ theorem setLIntegral_withDensity_eq_setLIntegral_mul (μ : Measure α) {f g : α ∫⁻ x in s, g x ∂μ.withDensity f = ∫⁻ x in s, (f * g) x ∂μ := by rw [restrict_withDensity hs, lintegral_withDensity_eq_lintegral_mul _ hf hg] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_withDensity_eq_set_lintegral_mul := setLIntegral_withDensity_eq_setLIntegral_mul - /-- The Lebesgue integral of `g` with respect to the measure `μ.withDensity f` coincides with the integral of `f * g`. This version assumes that `g` is almost everywhere measurable. For a version without conditions on `g` but requiring that `f` is almost everywhere finite, see @@ -406,9 +403,6 @@ lemma setLIntegral_withDensity_eq_lintegral_mul₀' {μ : Measure α} {f : α rw [← restrict_withDensity hs] exact hg.restrict -@[deprecated (since := "2024-06-29")] -alias set_lintegral_withDensity_eq_lintegral_mul₀' := setLIntegral_withDensity_eq_lintegral_mul₀' - theorem lintegral_withDensity_eq_lintegral_mul₀ {μ : Measure α} {f : α → ℝ≥0∞} (hf : AEMeasurable f μ) {g : α → ℝ≥0∞} (hg : AEMeasurable g μ) : ∫⁻ a, g a ∂μ.withDensity f = ∫⁻ a, (f * g) a ∂μ := @@ -421,9 +415,6 @@ lemma setLIntegral_withDensity_eq_lintegral_mul₀ {μ : Measure α} {f : α → setLIntegral_withDensity_eq_lintegral_mul₀' hf (hg.mono' (MeasureTheory.withDensity_absolutelyContinuous μ f)) hs -@[deprecated (since := "2024-06-29")] -alias set_lintegral_withDensity_eq_lintegral_mul₀ := setLIntegral_withDensity_eq_lintegral_mul₀ - theorem lintegral_withDensity_le_lintegral_mul (μ : Measure α) {f : α → ℝ≥0∞} (f_meas : Measurable f) (g : α → ℝ≥0∞) : (∫⁻ a, g a ∂μ.withDensity f) ≤ ∫⁻ a, (f * g) a ∂μ := by rw [← iSup_lintegral_measurable_le_eq_lintegral, ← iSup_lintegral_measurable_le_eq_lintegral] @@ -463,10 +454,6 @@ theorem setLIntegral_withDensity_eq_setLIntegral_mul_non_measurable (μ : Measur ∫⁻ a in s, g a ∂μ.withDensity f = ∫⁻ a in s, (f * g) a ∂μ := by rw [restrict_withDensity hs, lintegral_withDensity_eq_lintegral_mul_non_measurable _ f_meas hf] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_withDensity_eq_set_lintegral_mul_non_measurable := - setLIntegral_withDensity_eq_setLIntegral_mul_non_measurable - theorem lintegral_withDensity_eq_lintegral_mul_non_measurable₀ (μ : Measure α) {f : α → ℝ≥0∞} (hf : AEMeasurable f μ) (h'f : ∀ᵐ x ∂μ, f x < ∞) (g : α → ℝ≥0∞) : ∫⁻ a, g a ∂μ.withDensity f = ∫⁻ a, (f * g) a ∂μ := by @@ -491,20 +478,12 @@ theorem setLIntegral_withDensity_eq_setLIntegral_mul_non_measurable₀ (μ : Mea ∫⁻ a in s, g a ∂μ.withDensity f = ∫⁻ a in s, (f * g) a ∂μ := by rw [restrict_withDensity hs, lintegral_withDensity_eq_lintegral_mul_non_measurable₀ _ hf h'f] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_withDensity_eq_set_lintegral_mul_non_measurable₀ := - setLIntegral_withDensity_eq_setLIntegral_mul_non_measurable₀ - theorem setLIntegral_withDensity_eq_setLIntegral_mul_non_measurable₀' (μ : Measure α) [SFinite μ] {f : α → ℝ≥0∞} (s : Set α) (hf : AEMeasurable f (μ.restrict s)) (g : α → ℝ≥0∞) (h'f : ∀ᵐ x ∂μ.restrict s, f x < ∞) : ∫⁻ a in s, g a ∂μ.withDensity f = ∫⁻ a in s, (f * g) a ∂μ := by rw [restrict_withDensity' s, lintegral_withDensity_eq_lintegral_mul_non_measurable₀ _ hf h'f] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_withDensity_eq_set_lintegral_mul_non_measurable₀' := - setLIntegral_withDensity_eq_setLIntegral_mul_non_measurable₀' - theorem withDensity_mul₀ {μ : Measure α} {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) (hg : AEMeasurable g μ) : μ.withDensity (f * g) = (μ.withDensity f).withDensity g := by @@ -638,18 +617,6 @@ instance Measure.withDensity.instSFinite [SFinite μ] {f : α → ℝ≥0∞} : rw [key] infer_instance -@[deprecated Measure.withDensity.instSFinite (since := "2024-07-14"), nolint unusedArguments] -lemma sFinite_withDensity_of_sigmaFinite_of_measurable (μ : Measure α) [SigmaFinite μ] - {f : α → ℝ≥0∞} (_hf : Measurable f) : - SFinite (μ.withDensity f) := - inferInstance - -@[deprecated Measure.withDensity.instSFinite (since := "2024-07-14"), nolint unusedArguments] -lemma sFinite_withDensity_of_measurable (μ : Measure α) [SFinite μ] - {f : α → ℝ≥0∞} (_hf : Measurable f) : - SFinite (μ.withDensity f) := - inferInstance - instance [SFinite μ] (c : ℝ≥0∞) : SFinite (c • μ) := by rw [← withDensity_const] infer_instance diff --git a/Mathlib/MeasureTheory/OuterMeasure/Basic.lean b/Mathlib/MeasureTheory/OuterMeasure/Basic.lean index 9ca65df0c95b1..faa25a5ad5c75 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Basic.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Basic.lean @@ -158,65 +158,6 @@ namespace OuterMeasure variable {α β : Type*} {m : OuterMeasure α} -@[deprecated measure_empty (since := "2024-05-14")] -theorem empty' (m : OuterMeasure α) : m ∅ = 0 := measure_empty - -@[deprecated measure_mono (since := "2024-05-14")] -theorem mono' (m : OuterMeasure α) {s₁ s₂} (h : s₁ ⊆ s₂) : m s₁ ≤ m s₂ := by gcongr - -@[deprecated measure_mono_null (since := "2024-05-14")] -theorem mono_null (m : OuterMeasure α) {s t} (h : s ⊆ t) (ht : m t = 0) : m s = 0 := - measure_mono_null h ht - -@[deprecated measure_pos_of_superset (since := "2024-05-14")] -theorem pos_of_subset_ne_zero (m : OuterMeasure α) {a b : Set α} (hs : a ⊆ b) (hnz : m a ≠ 0) : - 0 < m b := - measure_pos_of_superset hs hnz - -@[deprecated measure_iUnion_le (since := "2024-05-14")] -protected theorem iUnion (m : OuterMeasure α) {β} [Countable β] (s : β → Set α) : - m (⋃ i, s i) ≤ ∑' i, m (s i) := - measure_iUnion_le s - -@[deprecated measure_biUnion_null_iff (since := "2024-05-14")] -theorem biUnion_null_iff (m : OuterMeasure α) {s : Set β} (hs : s.Countable) {t : β → Set α} : - m (⋃ i ∈ s, t i) = 0 ↔ ∀ i ∈ s, m (t i) = 0 := - measure_biUnion_null_iff hs - -@[deprecated measure_sUnion_null_iff (since := "2024-05-14")] -theorem sUnion_null_iff (m : OuterMeasure α) {S : Set (Set α)} (hS : S.Countable) : - m (⋃₀ S) = 0 ↔ ∀ s ∈ S, m s = 0 := measure_sUnion_null_iff hS - -@[deprecated measure_iUnion_null_iff (since := "2024-05-14")] -theorem iUnion_null_iff {ι : Sort*} [Countable ι] (m : OuterMeasure α) {s : ι → Set α} : - m (⋃ i, s i) = 0 ↔ ∀ i, m (s i) = 0 := - measure_iUnion_null_iff - -@[deprecated measure_iUnion_null (since := "2024-05-14")] -alias ⟨_, iUnion_null⟩ := iUnion_null_iff - -@[deprecated measure_biUnion_finset_le (since := "2024-05-14")] -protected theorem iUnion_finset (m : OuterMeasure α) (s : β → Set α) (t : Finset β) : - m (⋃ i ∈ t, s i) ≤ ∑ i ∈ t, m (s i) := - measure_biUnion_finset_le t s - -@[deprecated measure_union_le (since := "2024-05-14")] -protected theorem union (m : OuterMeasure α) (s₁ s₂ : Set α) : m (s₁ ∪ s₂) ≤ m s₁ + m s₂ := - measure_union_le s₁ s₂ - -/-- If a set has zero measure in a neighborhood of each of its points, then it has zero measure -in a second-countable space. -/ -@[deprecated measure_null_of_locally_null (since := "2024-05-14")] -theorem null_of_locally_null [TopologicalSpace α] [SecondCountableTopology α] (m : OuterMeasure α) - (s : Set α) (hs : ∀ x ∈ s, ∃ u ∈ 𝓝[s] x, m u = 0) : m s = 0 := - measure_null_of_locally_null s hs - -/-- If `m s ≠ 0`, then for some point `x ∈ s` and any `t ∈ 𝓝[s] x` we have `0 < m t`. -/ -@[deprecated exists_mem_forall_mem_nhdsWithin_pos_measure (since := "2024-05-14")] -theorem exists_mem_forall_mem_nhds_within_pos [TopologicalSpace α] [SecondCountableTopology α] - (m : OuterMeasure α) {s : Set α} (hs : m s ≠ 0) : ∃ x ∈ s, ∀ t ∈ 𝓝[s] x, 0 < m t := - exists_mem_forall_mem_nhdsWithin_pos_measure hs - /-- If `s : ι → Set α` is a sequence of sets, `S = ⋃ n, s n`, and `m (S \ s n)` tends to zero along some nontrivial filter (usually `atTop` on `ι = ℕ`), then `m S = ⨆ n, m (s n)`. -/ theorem iUnion_of_tendsto_zero {ι} (m : OuterMeasure α) {s : ι → Set α} (l : Filter ι) [NeBot l] @@ -246,20 +187,6 @@ theorem iUnion_nat_of_monotone_of_tsum_ne_top (m : OuterMeasure α) {s : ℕ → · rwa [this] · rw [← Nat.succ_le_iff, Nat.succ_eq_add_one, this] -@[deprecated measure_le_inter_add_diff (since := "2024-05-14")] -theorem le_inter_add_diff {m : OuterMeasure α} {t : Set α} (s : Set α) : - m t ≤ m (t ∩ s) + m (t \ s) := - measure_le_inter_add_diff m t s - -@[deprecated measure_diff_null (since := "2024-05-14")] -theorem diff_null (m : OuterMeasure α) (s : Set α) {t : Set α} (ht : m t = 0) : m (s \ t) = m s := - measure_diff_null ht - -@[deprecated measure_union_null (since := "2024-05-14")] -theorem union_null (m : OuterMeasure α) {s₁ s₂ : Set α} (h₁ : m s₁ = 0) (h₂ : m s₂ = 0) : - m (s₁ ∪ s₂) = 0 := - measure_union_null h₁ h₂ - theorem coe_fn_injective : Injective fun (μ : OuterMeasure α) (s : Set α) => μ s := DFunLike.coe_injective diff --git a/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean b/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean index 36c1a59159274..122729c759673 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean @@ -142,9 +142,6 @@ theorem isCaratheodory_iUnion_of_disjoint {s : ℕ → Set α} (h : ∀ i, IsCar refine m.mono (diff_subset_diff_right ?_) exact iUnion₂_subset fun i _ => subset_iUnion _ i -@[deprecated (since := "2024-07-29")] -alias isCaratheodory_iUnion_nat := isCaratheodory_iUnion_of_disjoint - lemma isCaratheodory_iUnion {s : ℕ → Set α} (h : ∀ i, m.IsCaratheodory (s i)) : m.IsCaratheodory (⋃ i, s i) := by rw [← iUnion_disjointed] diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 42ee51c93ceca..8bb2c8b2f8ab4 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -603,18 +603,12 @@ theorem natCast {f : ArithmeticFunction ℕ} [Semiring R] (h : f.IsMultiplicativ -- Porting note: was `by simp [cop, h]` ⟨by simp [h], fun {m n} cop => by simp [h.2 cop]⟩ -@[deprecated (since := "2024-04-17")] -alias nat_cast := natCast - @[arith_mult] theorem intCast {f : ArithmeticFunction ℤ} [Ring R] (h : f.IsMultiplicative) : IsMultiplicative (f : ArithmeticFunction R) := -- Porting note: was `by simp [cop, h]` ⟨by simp [h], fun {m n} cop => by simp [h.2 cop]⟩ -@[deprecated (since := "2024-04-17")] -alias int_cast := intCast - @[arith_mult] theorem mul [CommSemiring R] {f g : ArithmeticFunction R} (hf : f.IsMultiplicative) (hg : g.IsMultiplicative) : IsMultiplicative (f * g) := by @@ -1316,10 +1310,6 @@ theorem _root_.Nat.card_divisors {n : ℕ} (hn : n ≠ 0) : exact Finset.prod_congr n.support_factorization fun _ h => sigma_zero_apply_prime_pow <| Nat.prime_of_mem_primeFactors h -@[deprecated "No deprecation message was provided." (since := "2024-06-09")] -theorem card_divisors (n : ℕ) (hn : n ≠ 0) : - #n.divisors = n.primeFactors.prod (n.factorization · + 1) := Nat.card_divisors hn - theorem _root_.Nat.sum_divisors {n : ℕ} (hn : n ≠ 0) : ∑ d ∈ n.divisors, d = ∏ p ∈ n.primeFactors, ∑ k ∈ .range (n.factorization p + 1), p ^ k := by rw [← sigma_one_apply, isMultiplicative_sigma.multiplicative_factorization _ hn] diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index 3ae3f6217b5ea..2767a4973edb3 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -594,32 +594,4 @@ theorem norm_zeta_pow_sub_one_two {k : ℕ} (hk : 2 ≤ k) norm_sub_one_two (zeta_spec ((2 : ℕ+) ^ k) K L) hk hirr end IsCyclotomicExtension - -@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.pow_sub_one_norm_prime_pow_ne_two := - IsPrimitiveRoot.norm_pow_sub_one_of_prime_pow_ne_two -@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.pow_sub_one_norm_prime_ne_two := - IsPrimitiveRoot.norm_pow_sub_one_of_prime_ne_two -@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.sub_one_norm_prime_ne_two := - IsPrimitiveRoot.norm_sub_one_of_prime_ne_two -@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.sub_one_norm_prime := - IsPrimitiveRoot.norm_sub_one_of_prime_ne_two' -@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.pow_sub_one_norm_two := - IsPrimitiveRoot.norm_pow_sub_one_two -@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.sub_one_norm_two := - IsPrimitiveRoot.norm_sub_one_two -@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.pow_sub_one_norm_prime_pow_of_ne_zero := - IsPrimitiveRoot.norm_pow_sub_one_eq_prime_pow_of_ne_zero -@[deprecated (since := "2024-04-02")] alias IsCyclotomicExtension.isPrimePow_norm_zeta_sub_one := - IsCyclotomicExtension.norm_zeta_sub_one_of_isPrimePow -@[deprecated (since := "2024-04-02")] - alias IsCyclotomicExtension.prime_ne_two_pow_norm_zeta_pow_sub_one := - IsCyclotomicExtension.norm_zeta_pow_sub_one_of_prime_pow_ne_two -@[deprecated (since := "2024-04-02")] - alias IsCyclotomicExtension.prime_ne_two_pow_norm_zeta_sub_one := - IsCyclotomicExtension.norm_zeta_pow_sub_one_of_prime_ne_two -@[deprecated (since := "2024-04-02")] alias IsCyclotomicExtension.prime_ne_two_norm_zeta_sub_one := - IsCyclotomicExtension.norm_zeta_sub_one_of_prime_ne_two -@[deprecated (since := "2024-04-02")] alias IsCyclotomicExtension.two_pow_norm_zeta_sub_one := - IsCyclotomicExtension.norm_zeta_pow_sub_one_two - end Norm diff --git a/Mathlib/NumberTheory/DirichletCharacter/Basic.lean b/Mathlib/NumberTheory/DirichletCharacter/Basic.lean index a811a9135a7d5..6026d444ad49c 100644 --- a/Mathlib/NumberTheory/DirichletCharacter/Basic.lean +++ b/Mathlib/NumberTheory/DirichletCharacter/Basic.lean @@ -233,8 +233,6 @@ variable (χ) /-- A character is primitive if its level is equal to its conductor. -/ def IsPrimitive : Prop := conductor χ = n -@[deprecated (since := "2024-06-16")] alias isPrimitive := IsPrimitive - lemma isPrimitive_def : IsPrimitive χ ↔ conductor χ = n := Iff.rfl lemma isPrimitive_one_level_one : IsPrimitive (1 : DirichletCharacter R 1) := @@ -284,8 +282,6 @@ lemma primitive_mul_isPrimitive {m : ℕ} (ψ : DirichletCharacter R m) : IsPrimitive (primitive_mul χ ψ) := primitiveCharacter_isPrimitive _ -@[deprecated (since := "2024-06-16")] alias isPrimitive.primitive_mul := primitive_mul_isPrimitive - /- ### Even and odd characters -/ diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 4ef71ae103a18..258bafd7a6d60 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -475,17 +475,11 @@ theorem primeFactors_eq_to_filter_divisors_prime (n : ℕ) : · ext q simpa [hn, hn.ne', mem_primeFactorsList] using and_comm -@[deprecated (since := "2024-07-17")] -alias prime_divisors_eq_to_filter_divisors_prime := primeFactors_eq_to_filter_divisors_prime - lemma primeFactors_filter_dvd_of_dvd {m n : ℕ} (hn : n ≠ 0) (hmn : m ∣ n) : {p ∈ n.primeFactors | p ∣ m} = m.primeFactors := by simp_rw [primeFactors_eq_to_filter_divisors_prime, filter_comm, divisors_filter_dvd_of_dvd hn hmn] -@[deprecated (since := "2024-07-17")] -alias prime_divisors_filter_dvd_of_dvd := primeFactors_filter_dvd_of_dvd - @[simp] theorem image_div_divisors_eq_divisors (n : ℕ) : image (fun x : ℕ => n / x) n.divisors = n.divisors := by diff --git a/Mathlib/NumberTheory/LSeries/RiemannZeta.lean b/Mathlib/NumberTheory/LSeries/RiemannZeta.lean index 8642146695e22..bb054b6ad6e3e 100644 --- a/Mathlib/NumberTheory/LSeries/RiemannZeta.lean +++ b/Mathlib/NumberTheory/LSeries/RiemannZeta.lean @@ -225,19 +225,4 @@ theorem tendsto_sub_mul_tsum_nat_rpow : aliases for the old names -/ section aliases -@[deprecated (since := "2024-05-27")] -noncomputable alias riemannCompletedZeta₀ := completedRiemannZeta₀ - -@[deprecated (since := "2024-05-27")] -noncomputable alias riemannCompletedZeta := completedRiemannZeta - -@[deprecated (since := "2024-05-27")] -alias riemannCompletedZeta₀_one_sub := completedRiemannZeta₀_one_sub - -@[deprecated (since := "2024-05-27")] -alias riemannCompletedZeta_one_sub := completedRiemannZeta_one_sub - -@[deprecated (since := "2024-05-27")] -alias riemannCompletedZeta_residue_one := completedRiemannZeta_residue_one - end aliases diff --git a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean index acf857f2bf3e1..8d33816b369e1 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean @@ -223,8 +223,6 @@ noncomputable def FiniteField.primitiveChar (F F' : Type*) [Field F] [Finite F] exact ne_one_iff.2 ⟨a, fun hf => ha <| (ψ.prim.zmod_char_eq_one_iff pp <| Algebra.trace (ZMod p) F a).mp hf⟩ exact ⟨ψ.n, ψ', IsPrimitive.of_ne_one hψ'⟩ -@[deprecated (since := "2024-05-30")] alias primitiveCharFiniteField := FiniteField.primitiveChar - /-! ### The sum of all character values -/ diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index 8368544665f34..a02c87a48f3d5 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -214,11 +214,6 @@ theorem quadraticChar_ne_one (hF : ringChar F ≠ 2) : quadraticChar F ≠ 1 := intro hχ simp only [hχ, one_apply a.isUnit, one_ne_zero, reduceCtorEq] at ha -set_option linter.deprecated false in -@[deprecated quadraticChar_ne_one (since := "2024-06-16")] -theorem quadraticChar_isNontrivial (hF : ringChar F ≠ 2) : (quadraticChar F).IsNontrivial := - (isNontrivial_iff _).mpr <| quadraticChar_ne_one hF - open Finset in /-- The number of solutions to `x^2 = a` is determined by the quadratic character. -/ theorem quadraticChar_card_sqrts (hF : ringChar F ≠ 2) (a : F) : diff --git a/Mathlib/NumberTheory/LucasLehmer.lean b/Mathlib/NumberTheory/LucasLehmer.lean index e457c2e2a106c..4613db8d73fe3 100644 --- a/Mathlib/NumberTheory/LucasLehmer.lean +++ b/Mathlib/NumberTheory/LucasLehmer.lean @@ -152,8 +152,6 @@ theorem Int.natCast_pow_pred (b p : ℕ) (w : 0 < b) : ((b ^ p - 1 : ℕ) : ℤ) have : 1 ≤ b ^ p := Nat.one_le_pow p b w norm_cast -@[deprecated (since := "2024-05-25")] alias Int.coe_nat_pow_pred := Int.natCast_pow_pred - theorem Int.coe_nat_two_pow_pred (p : ℕ) : ((2 ^ p - 1 : ℕ) : ℤ) = (2 ^ p - 1 : ℤ) := Int.natCast_pow_pred 2 p (by decide) @@ -319,19 +317,12 @@ theorem fst_intCast (n : ℤ) : (n : X q).fst = (n : ZMod q) := theorem snd_intCast (n : ℤ) : (n : X q).snd = (0 : ZMod q) := rfl -@[deprecated (since := "2024-05-25")] alias nat_coe_fst := fst_natCast -@[deprecated (since := "2024-05-25")] alias nat_coe_snd := snd_natCast -@[deprecated (since := "2024-05-25")] alias int_coe_fst := fst_intCast -@[deprecated (since := "2024-05-25")] alias int_coe_snd := snd_intCast - @[norm_cast] theorem coe_mul (n m : ℤ) : ((n * m : ℤ) : X q) = (n : X q) * (m : X q) := by ext <;> simp @[norm_cast] theorem coe_natCast (n : ℕ) : ((n : ℤ) : X q) = (n : X q) := by ext <;> simp -@[deprecated (since := "2024-04-05")] alias coe_nat := coe_natCast - /-- The cardinality of `X` is `q^2`. -/ theorem card_eq : Fintype.card (X q) = q ^ 2 := by dsimp [X] diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index 7ddd95b920b4e..69fd911a6513f 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -390,17 +390,6 @@ lemma eq_one_iff {χ : MulChar R R'} : χ = 1 ↔ ∀ a : Rˣ, χ a = 1 := by lemma ne_one_iff {χ : MulChar R R'} : χ ≠ 1 ↔ ∃ a : Rˣ, χ a ≠ 1 := by simp only [Ne, eq_one_iff, not_forall] -/-- A multiplicative character is *nontrivial* if it takes a value `≠ 1` on a unit. -/ -@[deprecated "No deprecation message was provided." (since := "2024-06-16")] -def IsNontrivial (χ : MulChar R R') : Prop := - ∃ a : Rˣ, χ a ≠ 1 - -set_option linter.deprecated false in -/-- A multiplicative character is nontrivial iff it is not the trivial character. -/ -@[deprecated "No deprecation message was provided." (since := "2024-06-16")] -theorem isNontrivial_iff (χ : MulChar R R') : χ.IsNontrivial ↔ χ ≠ 1 := by - simp only [IsNontrivial, Ne, MulChar.ext_iff, not_forall, one_apply_coe] - end nontrivial section quadratic_and_comp @@ -461,16 +450,6 @@ lemma ringHomComp_ne_one_iff {f : R' →+* R''} (hf : Function.Injective f) {χ χ.ringHomComp f ≠ 1 ↔ χ ≠ 1 := (ringHomComp_eq_one_iff hf).not -set_option linter.deprecated false in -/-- Composition with an injective ring homomorphism preserves nontriviality. -/ -@[deprecated ringHomComp_ne_one_iff (since := "2024-06-16")] -theorem IsNontrivial.comp {χ : MulChar R R'} (hχ : χ.IsNontrivial) {f : R' →+* R''} - (hf : Function.Injective f) : (χ.ringHomComp f).IsNontrivial := by - obtain ⟨a, ha⟩ := hχ - use a - simp_rw [ringHomComp_apply, ← RingHom.map_one f] - exact fun h => ha (hf h) - /-- Composition with a ring homomorphism preserves the property of being a quadratic character. -/ theorem IsQuadratic.comp {χ : MulChar R R'} (hχ : χ.IsQuadratic) (f : R' →+* R'') : (χ.ringHomComp f).IsQuadratic := by @@ -572,12 +551,6 @@ theorem sum_eq_zero_of_ne_one [IsDomain R'] {χ : MulChar R R'} (hχ : χ ≠ 1) refine eq_zero_of_mul_eq_self_left hb ?_ simpa only [Finset.mul_sum, ← map_mul] using b.mulLeft_bijective.sum_comp _ -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-06-16")] -lemma IsNontrivial.sum_eq_zero [IsDomain R'] {χ : MulChar R R'} (hχ : χ.IsNontrivial) : - ∑ a, χ a = 0 := - sum_eq_zero_of_ne_one ((isNontrivial_iff _).mp hχ) - /-- The sum over all values of the trivial multiplicative character on a finite ring is the cardinality of its unit group. -/ theorem sum_one_eq_card_units [DecidableEq R] : diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 326ef6043194c..209f7b0e60424 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -137,15 +137,9 @@ instance instCommRing : CommRing ℤ_[p] := (by infer_instance : CommRing (subri @[simp, norm_cast] theorem coe_natCast (n : ℕ) : ((n : ℤ_[p]) : ℚ_[p]) = n := rfl -@[deprecated (since := "2024-04-17")] -alias coe_nat_cast := coe_natCast - @[simp, norm_cast] theorem coe_intCast (z : ℤ) : ((z : ℤ_[p]) : ℚ_[p]) = z := rfl -@[deprecated (since := "2024-04-17")] -alias coe_int_cast := coe_intCast - /-- The coercion from `ℤ_[p]` to `ℚ_[p]` as a ring homomorphism. -/ def Coe.ringHom : ℤ_[p] →+* ℚ_[p] := (subring p).subtype @@ -166,8 +160,6 @@ instance : CharZero ℤ_[p] where @[norm_cast] theorem intCast_eq (z1 z2 : ℤ) : (z1 : ℤ_[p]) = z2 ↔ z1 = z2 := by simp -@[deprecated (since := "2024-04-05")] alias coe_int_eq := intCast_eq - /-- A sequence of integers that is Cauchy with respect to the `p`-adic norm converges to a `p`-adic integer. -/ def ofIntSeq (seq : ℕ → ℤ) (h : IsCauSeq (padicNorm p) fun n => seq n) : ℤ_[p] := @@ -255,9 +247,6 @@ theorem padic_norm_e_of_padicInt (z : ℤ_[p]) : ‖(z : ℚ_[p])‖ = ‖z‖ : theorem norm_intCast_eq_padic_norm (z : ℤ) : ‖(z : ℤ_[p])‖ = ‖(z : ℚ_[p])‖ := by simp [norm_def] -@[deprecated (since := "2024-04-17")] -alias norm_int_cast_eq_padic_norm := norm_intCast_eq_padic_norm - @[simp] theorem norm_eq_padic_norm {q : ℚ_[p]} (hq : ‖q‖ ≤ 1) : @norm ℤ_[p] _ ⟨q, hq⟩ = ‖q‖ := rfl diff --git a/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean b/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean index 46c363d9d0744..b74442c047642 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean @@ -76,10 +76,6 @@ theorem le_emultiplicity_iff_replicate_subperm_primeFactorsList {a b : ℕ} {n : (replicate_subperm_primeFactorsList_iff ha hb).trans pow_dvd_iff_le_emultiplicity |>.symm -@[deprecated (since := "2024-07-17")] -alias le_multiplicity_iff_replicate_subperm_factors := - le_emultiplicity_iff_replicate_subperm_primeFactorsList - theorem le_padicValNat_iff_replicate_subperm_primeFactorsList {a b : ℕ} {n : ℕ} (ha : a.Prime) (hb : b ≠ 0) : n ≤ padicValNat a b ↔ replicate n a <+~ b.primeFactorsList := by @@ -87,7 +83,3 @@ theorem le_padicValNat_iff_replicate_subperm_primeFactorsList {a b : ℕ} {n : Nat.finiteMultiplicity_iff.2 ⟨ha.ne_one, Nat.pos_of_ne_zero hb⟩ |>.emultiplicity_eq_multiplicity, ← padicValNat_def' ha.ne_one (Nat.pos_of_ne_zero hb), Nat.cast_le] - -@[deprecated (since := "2024-07-17")] -alias le_padicValNat_iff_replicate_subperm_factors := - le_padicValNat_iff_replicate_subperm_primeFactorsList diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index f6b9067b3e3c9..7482f2835cb49 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -445,9 +445,6 @@ theorem denseRange_natCast : DenseRange (Nat.cast : ℕ → ℤ_[p]) := by rw [norm_le_pow_iff_mem_span_pow] apply appr_spec -@[deprecated (since := "2024-04-17")] -alias denseRange_nat_cast := denseRange_natCast - theorem denseRange_intCast : DenseRange (Int.cast : ℤ → ℤ_[p]) := by intro x refine DenseRange.induction_on denseRange_natCast x ?_ ?_ @@ -456,9 +453,6 @@ theorem denseRange_intCast : DenseRange (Int.cast : ℤ → ℤ_[p]) := by apply subset_closure exact Set.mem_range_self _ -@[deprecated (since := "2024-04-17")] -alias denseRange_int_cast := denseRange_intCast - end RingHoms section lift diff --git a/Mathlib/NumberTheory/Zsqrtd/Basic.lean b/Mathlib/NumberTheory/Zsqrtd/Basic.lean index 00bc9cbbccdaa..add497247da15 100644 --- a/Mathlib/NumberTheory/Zsqrtd/Basic.lean +++ b/Mathlib/NumberTheory/Zsqrtd/Basic.lean @@ -257,14 +257,6 @@ instance : CharZero (ℤ√d) where cast_injective m n := by simp [Zsqrtd.ext_if @[simp] theorem ofInt_eq_intCast (n : ℤ) : (ofInt n : ℤ√d) = n := by ext <;> simp [ofInt_re, ofInt_im] -@[deprecated (since := "2024-04-05")] alias coe_nat_re := natCast_re -@[deprecated (since := "2024-04-05")] alias coe_nat_im := natCast_im -@[deprecated (since := "2024-04-05")] alias coe_nat_val := natCast_val -@[deprecated (since := "2024-04-05")] alias coe_int_re := intCast_re -@[deprecated (since := "2024-04-05")] alias coe_int_im := intCast_im -@[deprecated (since := "2024-04-05")] alias coe_int_val := intCast_val -@[deprecated (since := "2024-04-05")] alias ofInt_eq_coe := ofInt_eq_intCast - @[simp] theorem smul_val (n x y : ℤ) : (n : ℤ√d) * ⟨x, y⟩ = ⟨n * x, n * y⟩ := by ext <;> simp @@ -286,11 +278,6 @@ theorem decompose {x y : ℤ} : (⟨x, y⟩ : ℤ√d) = x + sqrtd (d := d) * y theorem mul_star {x y : ℤ} : (⟨x, y⟩ * star ⟨x, y⟩ : ℤ√d) = x * x - d * y * y := by ext <;> simp [sub_eq_add_neg, mul_comm] -@[deprecated (since := "2024-05-25")] alias coe_int_add := Int.cast_add -@[deprecated (since := "2024-05-25")] alias coe_int_sub := Int.cast_sub -@[deprecated (since := "2024-05-25")] alias coe_int_mul := Int.cast_mul -@[deprecated (since := "2024-05-25")] alias coe_int_inj := Int.cast_inj - theorem intCast_dvd (z : ℤ) (a : ℤ√d) : ↑z ∣ a ↔ z ∣ a.re ∧ z ∣ a.im := by constructor · rintro ⟨x, rfl⟩ @@ -310,9 +297,6 @@ theorem intCast_dvd_intCast (a b : ℤ) : (a : ℤ√d) ∣ b ↔ a ∣ b := by · rw [intCast_re, intCast_im] exact fun hc => ⟨hc, dvd_zero a⟩ -@[deprecated (since := "2024-05-25")] alias coe_int_dvd_iff := intCast_dvd -@[deprecated (since := "2024-05-25")] alias coe_int_dvd_coe_int := intCast_dvd_intCast - protected theorem eq_of_smul_eq_smul_left {a : ℤ} {b c : ℤ√d} (ha : a ≠ 0) (h : ↑a * b = a * c) : b = c := by rw [Zsqrtd.ext_iff] at h ⊢ @@ -450,16 +434,10 @@ theorem norm_one : norm (1 : ℤ√d) = 1 := by simp [norm] @[simp] theorem norm_intCast (n : ℤ) : norm (n : ℤ√d) = n * n := by simp [norm] -@[deprecated (since := "2024-04-17")] -alias norm_int_cast := norm_intCast - @[simp] theorem norm_natCast (n : ℕ) : norm (n : ℤ√d) = n * n := norm_intCast n -@[deprecated (since := "2024-04-17")] -alias norm_nat_cast := norm_natCast - @[simp] theorem norm_mul (n m : ℤ√d) : norm (n * m) = norm n * norm m := by simp only [norm, mul_im, mul_re] diff --git a/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean b/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean index b3ef1877a2913..7e4d7941bda62 100644 --- a/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean +++ b/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean @@ -127,16 +127,10 @@ theorem toComplex_eq_zero {x : ℤ[i]} : (x : ℂ) = 0 ↔ x = 0 := by theorem intCast_real_norm (x : ℤ[i]) : (x.norm : ℝ) = Complex.normSq (x : ℂ) := by rw [Zsqrtd.norm, normSq]; simp -@[deprecated (since := "2024-04-17")] -alias int_cast_real_norm := intCast_real_norm - @[simp] theorem intCast_complex_norm (x : ℤ[i]) : (x.norm : ℂ) = Complex.normSq (x : ℂ) := by cases x; rw [Zsqrtd.norm, normSq]; simp -@[deprecated (since := "2024-04-17")] -alias int_cast_complex_norm := intCast_complex_norm - theorem norm_nonneg (x : ℤ[i]) : 0 ≤ norm x := Zsqrtd.norm_nonneg (by norm_num) _ @@ -149,15 +143,10 @@ theorem norm_pos {x : ℤ[i]} : 0 < norm x ↔ x ≠ 0 := by theorem abs_natCast_norm (x : ℤ[i]) : (x.norm.natAbs : ℤ) = x.norm := Int.natAbs_of_nonneg (norm_nonneg _) -@[deprecated (since := "2024-04-05")] alias abs_coe_nat_norm := abs_natCast_norm - @[simp] theorem natCast_natAbs_norm {α : Type*} [Ring α] (x : ℤ[i]) : (x.norm.natAbs : α) = x.norm := by rw [← Int.cast_natCast, abs_natCast_norm] -@[deprecated (since := "2024-04-17")] -alias nat_cast_natAbs_norm := natCast_natAbs_norm - theorem natAbs_norm_eq (x : ℤ[i]) : x.norm.natAbs = x.re.natAbs * x.re.natAbs + x.im.natAbs * x.im.natAbs := Int.ofNat.inj <| by simp; simp [Zsqrtd.norm] diff --git a/Mathlib/Order/Atoms.lean b/Mathlib/Order/Atoms.lean index 936ee64e16e19..07d610b43d88c 100644 --- a/Mathlib/Order/Atoms.lean +++ b/Mathlib/Order/Atoms.lean @@ -687,9 +687,6 @@ theorem eq_top_of_lt : b = ⊤ := alias _root_.LT.lt.eq_bot := eq_bot_of_lt alias _root_.LT.lt.eq_top := eq_top_of_lt -@[deprecated (since := "2024-05-29")] alias LT.lt.eq_bot := _root_.LT.lt.eq_bot -@[deprecated (since := "2024-05-29")] alias LT.lt.eq_top := _root_.LT.lt.eq_top - end Preorder section BoundedOrder diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index 5e469fe3ff642..16f0ce9de454e 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -1612,9 +1612,6 @@ protected lemma Antitone.sSup (hs : ∀ f ∈ s, Antitone f) : Antitone (sSup s) protected lemma Antitone.sInf (hs : ∀ f ∈ s, Antitone f) : Antitone (sInf s) := fun _ _ h ↦ iInf_mono fun f ↦ hs f f.2 h -@[deprecated (since := "2024-05-29")] alias monotone_sSup_of_monotone := Monotone.sSup -@[deprecated (since := "2024-05-29")] alias monotone_sInf_of_monotone := Monotone.sInf - protected lemma Monotone.iSup (hf : ∀ i, Monotone (f i)) : Monotone (⨆ i, f i) := Monotone.sSup (by simpa) protected lemma Monotone.iInf (hf : ∀ i, Monotone (f i)) : Monotone (⨅ i, f i) := diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index 84fa0aae1c71a..3b8dab942e6ce 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -102,8 +102,6 @@ theorem WithTop.coe_sSup' [SupSet α] {s : Set α} (hs : BddAbove s) : theorem WithBot.sSup_empty [SupSet α] : sSup (∅ : Set (WithBot α)) = ⊥ := WithTop.sInf_empty (α := αᵒᵈ) -@[deprecated (since := "2024-06-10")] alias WithBot.csSup_empty := WithBot.sSup_empty - @[norm_cast] theorem WithBot.coe_sSup' [SupSet α] {s : Set α} (hs : s.Nonempty) (h's : BddAbove s) : ↑(sSup s) = (sSup ((fun (a : α) ↦ ↑a) '' s) : WithBot α) := diff --git a/Mathlib/Order/Cover.lean b/Mathlib/Order/Cover.lean index 9c99fd65d0036..74e92607af79b 100644 --- a/Mathlib/Order/Cover.lean +++ b/Mathlib/Order/Cover.lean @@ -223,9 +223,6 @@ theorem denselyOrdered_iff_forall_not_covBy : DenselyOrdered α ↔ ∀ a b : α ⟨fun h _ _ => @not_covBy _ _ _ _ h, fun h => ⟨fun _ _ hab => exists_lt_lt_of_not_covBy hab <| h _ _⟩⟩ -@[deprecated (since := "2024-04-04")] -alias densely_ordered_iff_forall_not_covBy := denselyOrdered_iff_forall_not_covBy - @[simp] theorem toDual_covBy_toDual_iff : toDual b ⋖ toDual a ↔ a ⋖ b := and_congr_right' <| forall_congr' fun _ => forall_swap diff --git a/Mathlib/Order/Defs/PartialOrder.lean b/Mathlib/Order/Defs/PartialOrder.lean index d5bac42a0abd8..ac3744d8140c9 100644 --- a/Mathlib/Order/Defs/PartialOrder.lean +++ b/Mathlib/Order/Defs/PartialOrder.lean @@ -46,10 +46,6 @@ lemma lt_iff_le_not_le : a < b ↔ a ≤ b ∧ ¬b ≤ a := Preorder.lt_iff_le_n lemma lt_of_le_not_le (hab : a ≤ b) (hba : ¬ b ≤ a) : a < b := lt_iff_le_not_le.2 ⟨hab, hba⟩ -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem le_not_le_of_lt : ∀ {a b : α}, a < b → a ≤ b ∧ ¬b ≤ a - | _a, _b, hab => lt_iff_le_not_le.mp hab - lemma le_of_eq (hab : a = b) : a ≤ b := by rw [hab] lemma le_of_lt (hab : a < b) : a ≤ b := (lt_iff_le_not_le.1 hab).1 lemma not_le_of_lt (hab : a < b) : ¬ b ≤ a := (lt_iff_le_not_le.1 hab).2 diff --git a/Mathlib/Order/Filter/AtTopBot/Archimedean.lean b/Mathlib/Order/Filter/AtTopBot/Archimedean.lean index 167f93687a47d..68dd2de56d8a8 100644 --- a/Mathlib/Order/Filter/AtTopBot/Archimedean.lean +++ b/Mathlib/Order/Filter/AtTopBot/Archimedean.lean @@ -31,23 +31,14 @@ theorem tendsto_natCast_atTop_iff [StrictOrderedSemiring R] [Archimedean R] {f : {l : Filter α} : Tendsto (fun n => (f n : R)) l atTop ↔ Tendsto f l atTop := tendsto_atTop_embedding (fun _ _ => Nat.cast_le) exists_nat_ge -@[deprecated (since := "2024-04-17")] -alias tendsto_nat_cast_atTop_iff := tendsto_natCast_atTop_iff - theorem tendsto_natCast_atTop_atTop [OrderedSemiring R] [Archimedean R] : Tendsto ((↑) : ℕ → R) atTop atTop := Nat.mono_cast.tendsto_atTop_atTop exists_nat_ge -@[deprecated (since := "2024-04-17")] -alias tendsto_nat_cast_atTop_atTop := tendsto_natCast_atTop_atTop - theorem Filter.Eventually.natCast_atTop [OrderedSemiring R] [Archimedean R] {p : R → Prop} (h : ∀ᶠ (x : R) in atTop, p x) : ∀ᶠ (n : ℕ) in atTop, p n := tendsto_natCast_atTop_atTop.eventually h -@[deprecated (since := "2024-04-17")] -alias Filter.Eventually.nat_cast_atTop := Filter.Eventually.natCast_atTop - @[simp] theorem Int.comap_cast_atTop [StrictOrderedRing R] [Archimedean R] : comap ((↑) : ℤ → R) atTop = atTop := comap_embedding_atTop (fun _ _ => Int.cast_le) fun r => @@ -64,37 +55,22 @@ theorem tendsto_intCast_atTop_iff [StrictOrderedRing R] [Archimedean R] {f : α {l : Filter α} : Tendsto (fun n => (f n : R)) l atTop ↔ Tendsto f l atTop := by rw [← @Int.comap_cast_atTop R, tendsto_comap_iff]; rfl -@[deprecated (since := "2024-04-17")] -alias tendsto_int_cast_atTop_iff := tendsto_intCast_atTop_iff - theorem tendsto_intCast_atBot_iff [StrictOrderedRing R] [Archimedean R] {f : α → ℤ} {l : Filter α} : Tendsto (fun n => (f n : R)) l atBot ↔ Tendsto f l atBot := by rw [← @Int.comap_cast_atBot R, tendsto_comap_iff]; rfl -@[deprecated (since := "2024-04-17")] -alias tendsto_int_cast_atBot_iff := tendsto_intCast_atBot_iff - theorem tendsto_intCast_atTop_atTop [StrictOrderedRing R] [Archimedean R] : Tendsto ((↑) : ℤ → R) atTop atTop := tendsto_intCast_atTop_iff.2 tendsto_id -@[deprecated (since := "2024-04-17")] -alias tendsto_int_cast_atTop_atTop := tendsto_intCast_atTop_atTop - theorem Filter.Eventually.intCast_atTop [StrictOrderedRing R] [Archimedean R] {p : R → Prop} (h : ∀ᶠ (x : R) in atTop, p x) : ∀ᶠ (n : ℤ) in atTop, p n := by rw [← Int.comap_cast_atTop (R := R)]; exact h.comap _ -@[deprecated (since := "2024-04-17")] -alias Filter.Eventually.int_cast_atTop := Filter.Eventually.intCast_atTop - theorem Filter.Eventually.intCast_atBot [StrictOrderedRing R] [Archimedean R] {p : R → Prop} (h : ∀ᶠ (x : R) in atBot, p x) : ∀ᶠ (n : ℤ) in atBot, p n := by rw [← Int.comap_cast_atBot (R := R)]; exact h.comap _ -@[deprecated (since := "2024-04-17")] -alias Filter.Eventually.int_cast_atBot := Filter.Eventually.intCast_atBot - @[simp] theorem Rat.comap_cast_atTop [LinearOrderedField R] [Archimedean R] : comap ((↑) : ℚ → R) atTop = atTop := @@ -111,30 +87,18 @@ theorem tendsto_ratCast_atTop_iff [LinearOrderedField R] [Archimedean R] {f : α {l : Filter α} : Tendsto (fun n => (f n : R)) l atTop ↔ Tendsto f l atTop := by rw [← @Rat.comap_cast_atTop R, tendsto_comap_iff]; rfl -@[deprecated (since := "2024-04-17")] -alias tendsto_rat_cast_atTop_iff := tendsto_ratCast_atTop_iff - theorem tendsto_ratCast_atBot_iff [LinearOrderedField R] [Archimedean R] {f : α → ℚ} {l : Filter α} : Tendsto (fun n => (f n : R)) l atBot ↔ Tendsto f l atBot := by rw [← @Rat.comap_cast_atBot R, tendsto_comap_iff]; rfl -@[deprecated (since := "2024-04-17")] -alias tendsto_rat_cast_atBot_iff := tendsto_ratCast_atBot_iff - theorem Filter.Eventually.ratCast_atTop [LinearOrderedField R] [Archimedean R] {p : R → Prop} (h : ∀ᶠ (x : R) in atTop, p x) : ∀ᶠ (n : ℚ) in atTop, p n := by rw [← Rat.comap_cast_atTop (R := R)]; exact h.comap _ -@[deprecated (since := "2024-04-17")] -alias Filter.Eventually.rat_cast_atTop := Filter.Eventually.ratCast_atTop - theorem Filter.Eventually.ratCast_atBot [LinearOrderedField R] [Archimedean R] {p : R → Prop} (h : ∀ᶠ (x : R) in atBot, p x) : ∀ᶠ (n : ℚ) in atBot, p n := by rw [← Rat.comap_cast_atBot (R := R)]; exact h.comap _ -@[deprecated (since := "2024-04-17")] -alias Filter.Eventually.rat_cast_atBot := Filter.Eventually.ratCast_atBot - theorem atTop_hasAntitoneBasis_of_archimedean [OrderedSemiring R] [Archimedean R] : (atTop : Filter R).HasAntitoneBasis fun n : ℕ => Ici n := hasAntitoneBasis_atTop.comp_mono Nat.mono_cast tendsto_natCast_atTop_atTop @@ -228,12 +192,6 @@ theorem Tendsto.atBot_mul_const_of_neg' (hr : r < 0) (hf : Tendsto f l atBot) : Tendsto (fun x => f x * r) l atTop := by simpa only [mul_neg, tendsto_neg_atBot_iff] using hf.atBot_mul_const' (neg_pos.2 hr) -@[deprecated (since := "2024-05-06")] -alias Tendsto.atTop_mul_neg_const' := Tendsto.atTop_mul_const_of_neg' - -@[deprecated (since := "2024-05-06")] -alias Tendsto.atBot_mul_neg_const' := Tendsto.atBot_mul_const_of_neg' - end LinearOrderedRing section LinearOrderedCancelAddCommMonoid diff --git a/Mathlib/Order/Filter/AtTopBot/Field.lean b/Mathlib/Order/Filter/AtTopBot/Field.lean index 1112b9f9d9e92..c82f42e8df916 100644 --- a/Mathlib/Order/Filter/AtTopBot/Field.lean +++ b/Mathlib/Order/Filter/AtTopBot/Field.lean @@ -304,17 +304,5 @@ theorem tendsto_const_mul_pow_atBot_iff {c : α} {n : ℕ} : Tendsto (fun x => c * x ^ n) atTop atBot ↔ n ≠ 0 ∧ c < 0 := by simp only [← tendsto_neg_atTop_iff, ← neg_mul, tendsto_const_mul_pow_atTop_iff, neg_pos] -@[deprecated (since := "2024-05-06")] -alias Tendsto.neg_const_mul_atTop := Tendsto.const_mul_atTop_of_neg - -@[deprecated (since := "2024-05-06")] -alias Tendsto.atTop_mul_neg_const := Tendsto.atTop_mul_const_of_neg - -@[deprecated (since := "2024-05-06")] -alias Tendsto.neg_const_mul_atBot := Tendsto.const_mul_atBot_of_neg - -@[deprecated (since := "2024-05-06")] -alias Tendsto.atBot_mul_neg_const := Tendsto.atBot_mul_const_of_neg - end LinearOrderedField end Filter diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index f02186b5a0c53..9d6edfa9fcb8b 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -442,9 +442,6 @@ instance instIntCast [IntCast M] : IntCast (Germ l M) where intCast n := (n : α @[simp] theorem intCast_def [IntCast M] (n : ℤ) : ((fun _ ↦ n : α → M) : Germ l M) = n := rfl -@[deprecated (since := "2024-04-05")] alias coe_nat := natCast_def -@[deprecated (since := "2024-04-05")] alias coe_int := intCast_def - instance instAddMonoidWithOne [AddMonoidWithOne M] : AddMonoidWithOne (Germ l M) where natCast_zero := congrArg ofFun <| by simp; rfl natCast_succ _ := congrArg ofFun <| by simp [Function.comp]; rfl diff --git a/Mathlib/Order/Fin/Basic.lean b/Mathlib/Order/Fin/Basic.lean index 8f534eba16e5a..8255757175836 100644 --- a/Mathlib/Order/Fin/Basic.lean +++ b/Mathlib/Order/Fin/Basic.lean @@ -215,25 +215,17 @@ def castOrderIso (eq : n = m) : Fin n ≃o Fin m where toEquiv := ⟨Fin.cast eq, Fin.cast eq.symm, leftInverse_cast eq, rightInverse_cast eq⟩ map_rel_iff' := cast_le_cast eq -@[deprecated (since := "2024-05-23")] alias castIso := castOrderIso - @[simp] lemma symm_castOrderIso (h : n = m) : (castOrderIso h).symm = castOrderIso h.symm := by subst h; rfl -@[deprecated (since := "2024-05-23")] alias symm_castIso := symm_castOrderIso - @[simp] lemma castOrderIso_refl (h : n = n := rfl) : castOrderIso h = OrderIso.refl (Fin n) := by ext; simp -@[deprecated (since := "2024-05-23")] alias castIso_refl := castOrderIso_refl - /-- While in many cases `Fin.castOrderIso` is better than `Equiv.cast`/`cast`, sometimes we want to apply a generic lemma about `cast`. -/ lemma castOrderIso_toEquiv (h : n = m) : (castOrderIso h).toEquiv = Equiv.cast (h ▸ rfl) := by subst h; rfl -@[deprecated (since := "2024-05-23")] alias castIso_to_equiv := castOrderIso_toEquiv - /-- `Fin.rev n` as an order-reversing isomorphism. -/ @[simps! apply toEquiv] def revOrderIso : (Fin n)ᵒᵈ ≃o Fin n := ⟨OrderDual.ofDual.trans revPerm, rev_le_rev⟩ diff --git a/Mathlib/Order/Fin/Tuple.lean b/Mathlib/Order/Fin/Tuple.lean index 26c333e2ed7fa..6a3de6af21eaf 100644 --- a/Mathlib/Order/Fin/Tuple.lean +++ b/Mathlib/Order/Fin/Tuple.lean @@ -154,14 +154,6 @@ not a definitional equality. -/ end Fin -/-- Order isomorphism between `Π j : Fin (n + 1), α j` and -`α i × Π j : Fin n, α (Fin.succAbove i j)`. -/ -@[deprecated Fin.insertNthOrderIso (since := "2024-07-12")] -def OrderIso.piFinSuccAboveIso (α : Fin (n + 1) → Type*) [∀ i, LE (α i)] - (i : Fin (n + 1)) : (∀ j, α j) ≃o α i × ∀ j, α (i.succAbove j) where - toEquiv := (Fin.insertNthEquiv α i).symm - map_rel_iff' := Iff.symm i.forall_iff_succAbove - /-- `Fin.succAbove` as an order isomorphism between `Fin n` and `{x : Fin (n + 1) // x ≠ p}`. -/ def finSuccAboveOrderIso (p : Fin (n + 1)) : Fin n ≃o { x : Fin (n + 1) // x ≠ p } where __ := finSuccAboveEquiv p diff --git a/Mathlib/Order/Lattice.lean b/Mathlib/Order/Lattice.lean index 5cdba7b8a8b2a..4b7df71d39ff3 100644 --- a/Mathlib/Order/Lattice.lean +++ b/Mathlib/Order/Lattice.lean @@ -103,14 +103,10 @@ variable [SemilatticeSup α] {a b c d : α} theorem le_sup_left : a ≤ a ⊔ b := SemilatticeSup.le_sup_left a b -@[deprecated (since := "2024-06-04")] alias le_sup_left' := le_sup_left - @[simp] theorem le_sup_right : b ≤ a ⊔ b := SemilatticeSup.le_sup_right a b -@[deprecated (since := "2024-06-04")] alias le_sup_right' := le_sup_right - theorem le_sup_of_le_left (h : c ≤ a) : c ≤ a ⊔ b := le_trans h le_sup_left @@ -307,14 +303,10 @@ variable [SemilatticeInf α] {a b c d : α} theorem inf_le_left : a ⊓ b ≤ a := SemilatticeInf.inf_le_left a b -@[deprecated (since := "2024-06-04")] alias inf_le_left' := inf_le_left - @[simp] theorem inf_le_right : a ⊓ b ≤ b := SemilatticeInf.inf_le_right a b -@[deprecated (since := "2024-06-04")] alias inf_le_right' := inf_le_right - theorem le_inf : a ≤ b → a ≤ c → a ≤ b ⊓ c := SemilatticeInf.le_inf a b c diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 5956cf5e2f2cf..d55c43748ab79 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -1002,18 +1002,12 @@ theorem _root_.CompleteLatticeHom.apply_limsup_iterate (f : CompleteLatticeHom simp only [zero_add, Function.comp_apply, iSup_le_iff] exact fun i => le_iSup (fun i => f^[i] a) (i + 1) -@[deprecated (since := "2024-07-21")] -alias CompleteLatticeHom.apply_limsup_iterate := CompleteLatticeHom.apply_limsup_iterate - /-- If `f : α → α` is a morphism of complete lattices, then the liminf of its iterates of any `a : α` is a fixed point. -/ theorem _root_.CompleteLatticeHom.apply_liminf_iterate (f : CompleteLatticeHom α α) (a : α) : f (liminf (fun n => f^[n] a) atTop) = liminf (fun n => f^[n] a) atTop := (CompleteLatticeHom.dual f).apply_limsup_iterate _ -@[deprecated (since := "2024-07-21")] -alias CompleteLatticeHom.apply_liminf_iterate := CompleteLatticeHom.apply_liminf_iterate - variable {f g : Filter β} {p q : β → Prop} {u v : β → α} theorem blimsup_mono (h : ∀ x, p x → q x) : blimsup u f p ≤ blimsup u f q := @@ -1097,32 +1091,20 @@ theorem _root_.OrderIso.apply_blimsup [CompleteLattice γ] (e : α ≃o γ) : simp only [blimsup_eq, map_sInf, Function.comp_apply, e.image_eq_preimage, Set.preimage_setOf_eq, e.le_symm_apply] -@[deprecated (since := "2024-07-21")] -alias OrderIso.apply_blimsup := OrderIso.apply_blimsup - theorem _root_.OrderIso.apply_bliminf [CompleteLattice γ] (e : α ≃o γ) : e (bliminf u f p) = bliminf (e ∘ u) f p := e.dual.apply_blimsup -@[deprecated (since := "2024-07-21")] -alias OrderIso.apply_bliminf := OrderIso.apply_bliminf - theorem _root_.sSupHom.apply_blimsup_le [CompleteLattice γ] (g : sSupHom α γ) : g (blimsup u f p) ≤ blimsup (g ∘ u) f p := by simp only [blimsup_eq_iInf_biSup, Function.comp] refine ((OrderHomClass.mono g).map_iInf₂_le _).trans ?_ simp only [_root_.map_iSup, le_refl] -@[deprecated (since := "2024-07-21")] -alias SupHom.apply_blimsup_le := sSupHom.apply_blimsup_le - theorem _root_.sInfHom.le_apply_bliminf [CompleteLattice γ] (g : sInfHom α γ) : bliminf (g ∘ u) f p ≤ g (bliminf u f p) := (sInfHom.dual g).apply_blimsup_le -@[deprecated (since := "2024-07-21")] -alias InfHom.le_apply_bliminf := sInfHom.le_apply_bliminf - end CompleteLattice section CompleteDistribLattice diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index c1fe4b32ab082..b36a75d4c65f4 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -6,7 +6,6 @@ Authors: Simon Hudon, Ira Fesefeldt import Mathlib.Control.Monad.Basic import Mathlib.Dynamics.FixedPoints.Basic import Mathlib.Order.Chain -import Mathlib.Order.Hom.Order import Mathlib.Order.Iterate import Mathlib.Order.Part import Mathlib.Order.ScottContinuity @@ -296,89 +295,6 @@ lemma ωScottContinuous.comp (hg : ωScottContinuous g) (hf : ωScottContinuous lemma ωScottContinuous.const {x : β} : ωScottContinuous (Function.const α x) := by simp [ωScottContinuous, ScottContinuousOn, Set.range_nonempty] -set_option linter.deprecated false in -/-- A monotone function `f : α →o β` is continuous if it distributes over ωSup. - -In order to distinguish it from the (more commonly used) continuity from topology -(see `Mathlib/Topology/Basic.lean`), the present definition is often referred to as -"Scott-continuity" (referring to Dana Scott). It corresponds to continuity -in Scott topological spaces (not defined here). -/ -@[deprecated ωScottContinuous (since := "2024-05-29")] -def Continuous (f : α →o β) : Prop := - ∀ c : Chain α, f (ωSup c) = ωSup (c.map f) - -set_option linter.deprecated false in -/-- `Continuous' f` asserts that `f` is both monotone and continuous. -/ -@[deprecated ωScottContinuous (since := "2024-05-29")] -def Continuous' (f : α → β) : Prop := - ∃ hf : Monotone f, Continuous ⟨f, hf⟩ - -@[deprecated ωScottContinuous.isLUB (since := "2024-05-29")] -lemma isLUB_of_scottContinuous {c : Chain α} {f : α → β} (hf : ScottContinuous f) : - IsLUB (Set.range (Chain.map c ⟨f, (ScottContinuous.monotone hf)⟩)) (f (ωSup c)) := - ωScottContinuous.isLUB hf.scottContinuousOn - -set_option linter.deprecated false in -@[deprecated ScottContinuous.ωScottContinuous (since := "2024-05-29")] -lemma ScottContinuous.continuous' {f : α → β} (hf : ScottContinuous f) : Continuous' f := by - constructor - · intro c - rw [← (ωSup_eq_of_isLUB (isLUB_of_scottContinuous hf))] - simp only [OrderHom.coe_mk] - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.monotone (since := "2024-05-29")] -theorem Continuous'.to_monotone {f : α → β} (hf : Continuous' f) : Monotone f := - hf.fst - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.of_monotone_map_ωSup (since := "2024-05-29")] -theorem Continuous.of_bundled (f : α → β) (hf : Monotone f) (hf' : Continuous ⟨f, hf⟩) : - Continuous' f := - ⟨hf, hf'⟩ - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.of_monotone_map_ωSup (since := "2024-05-29")] -theorem Continuous.of_bundled' (f : α →o β) (hf' : Continuous f) : Continuous' f := - ⟨f.mono, hf'⟩ - -set_option linter.deprecated false in -@[deprecated ωScottContinuous_iff_monotone_map_ωSup (since := "2024-05-29")] -theorem Continuous'.to_bundled (f : α → β) (hf : Continuous' f) : Continuous ⟨f, hf.to_monotone⟩ := - hf.snd - -set_option linter.deprecated false in -@[simp, norm_cast, deprecated ωScottContinuous_iff_monotone_map_ωSup (since := "2024-05-29")] -theorem continuous'_coe : ∀ {f : α →o β}, Continuous' f ↔ Continuous f - | ⟨_, hf⟩ => ⟨fun ⟨_, hc⟩ => hc, fun hc => ⟨hf, hc⟩⟩ - -variable (f : α →o β) (g : β →o γ) - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.id (since := "2024-05-29")] -theorem continuous_id : Continuous (@OrderHom.id α _) := by intro c; rw [c.map_id]; rfl - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.comp (since := "2024-05-29")] -theorem continuous_comp (hfc : Continuous f) (hgc : Continuous g) : Continuous (g.comp f) := by - dsimp [Continuous] at *; intro - rw [hfc, hgc, Chain.map_comp] - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.id (since := "2024-05-29")] -theorem id_continuous' : Continuous' (@id α) := - continuous_id.of_bundled' _ - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.const (since := "2024-05-29")] -theorem continuous_const (x : β) : Continuous (OrderHom.const α x) := fun c => - eq_of_forall_ge_iff fun z => by rw [ωSup_le_iff, Chain.map_coe, OrderHom.const_coe_coe]; simp - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.const (since := "2024-05-29")] -theorem const_continuous' (x : β) : Continuous' (Function.const α x) := - Continuous.of_bundled' (OrderHom.const α x) (continuous_const x) - end Continuity end OmegaCompletePartialOrder @@ -495,19 +411,6 @@ lemma ωScottContinuous.of_apply₂ (hf : ∀ a, ωScottContinuous (f · a)) : lemma ωScottContinuous_iff_apply₂ : ωScottContinuous f ↔ ∀ a, ωScottContinuous (f · a) := ⟨ωScottContinuous.apply₂, ωScottContinuous.of_apply₂⟩ -set_option linter.deprecated false in -@[deprecated ωScottContinuous.apply₂ (since := "2024-05-29")] -theorem flip₁_continuous' (f : ∀ x : α, γ → β x) (a : α) (hf : Continuous' fun x y => f y x) : - Continuous' (f a) := - Continuous.of_bundled _ (fun _ _ h => hf.to_monotone h a) fun c => congr_fun (hf.to_bundled _ c) a - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.of_apply₂ (since := "2024-05-29")] -theorem flip₂_continuous' (f : γ → ∀ x, β x) (hf : ∀ x, Continuous' fun g => f g x) : - Continuous' f := - Continuous.of_bundled _ (fun _ _ h a => (hf a).to_monotone h) - (by intro c; ext a; apply (hf a).to_bundled _ c) - end OmegaCompletePartialOrder end Pi @@ -582,52 +485,6 @@ lemma ωScottContinuous.top : ωScottContinuous (⊤ : α → β) := lemma ωScottContinuous.bot : ωScottContinuous (⊥ : α → β) := by rw [← sSup_empty]; exact ωScottContinuous.sSup (by simp) -set_option linter.deprecated false in -@[deprecated ωScottContinuous.sSup (since := "2024-05-29")] -theorem sSup_continuous (s : Set <| α →o β) (hs : ∀ f ∈ s, Continuous f) : Continuous (sSup s) := by - intro c - apply eq_of_forall_ge_iff - intro z - suffices (∀ f ∈ s, ∀ n, f (c n) ≤ z) ↔ ∀ n, ∀ f ∈ s, f (c n) ≤ z by - simpa (config := { contextual := true }) [ωSup_le_iff, hs _ _ _] using this - exact ⟨fun H n f hf => H f hf n, fun H f hf n => H n f hf⟩ - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.iSup (since := "2024-05-29")] -theorem iSup_continuous {ι : Sort*} {f : ι → α →o β} (h : ∀ i, Continuous (f i)) : - Continuous (⨆ i, f i) := - sSup_continuous _ <| Set.forall_mem_range.2 h - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.sSup (since := "2024-05-29")] -theorem sSup_continuous' (s : Set (α → β)) (hc : ∀ f ∈ s, Continuous' f) : - Continuous' (sSup s) := by - lift s to Set (α →o β) using fun f hf => (hc f hf).to_monotone - simp only [Set.forall_mem_image, continuous'_coe] at hc - rw [sSup_image] - norm_cast - exact iSup_continuous fun f ↦ iSup_continuous fun hf ↦ hc hf - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.sup (since := "2024-05-29")] -theorem sup_continuous {f g : α →o β} (hf : Continuous f) (hg : Continuous g) : - Continuous (f ⊔ g) := by - rw [← sSup_pair]; apply sSup_continuous - rintro f (rfl | rfl | _) <;> assumption - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.top (since := "2024-05-29")] -theorem top_continuous : Continuous (⊤ : α →o β) := by - intro c; apply eq_of_forall_ge_iff; intro z - simp only [OrderHom.instTopOrderHom_top, OrderHom.const_coe_coe, Function.const, top_le_iff, - ωSup_le_iff, Chain.map_coe, Function.comp, forall_const] - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.bot (since := "2024-05-29")] -theorem bot_continuous : Continuous (⊥ : α →o β) := by - rw [← sSup_empty] - exact sSup_continuous _ fun f hf => hf.elim - end CompleteLattice namespace CompleteLattice @@ -650,23 +507,6 @@ lemma ωScottContinuous.inf (hf : ωScottContinuous f) (hg : ωScottContinuous g (h (max j i)).imp (le_trans <| hf.monotone <| c.mono <| le_max_left _ _) (le_trans <| hg.monotone <| c.mono <| le_max_right _ _)⟩ -set_option linter.deprecated false in -@[deprecated ωScottContinuous.inf (since := "2024-05-29")] -theorem inf_continuous (f g : α →o β) (hf : Continuous f) (hg : Continuous g) : - Continuous (f ⊓ g) := by - refine fun c => eq_of_forall_ge_iff fun z => ?_ - simp only [inf_le_iff, hf c, hg c, ωSup_le_iff, ← forall_or_left, ← forall_or_right, - Chain.map_coe, OrderHom.coe_inf, Pi.inf_apply, Function.comp] - exact ⟨fun h _ ↦ h _ _, fun h i j ↦ - (h (max j i)).imp (le_trans <| f.mono <| c.mono <| le_max_left _ _) - (le_trans <| g.mono <| c.mono <| le_max_right _ _)⟩ - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.inf (since := "2024-05-29")] -theorem inf_continuous' {f g : α → β} (hf : Continuous' f) (hg : Continuous' g) : - Continuous' (f ⊓ g) := - ⟨_, inf_continuous _ _ hf.snd hg.snd⟩ - end CompleteLattice namespace OmegaCompletePartialOrder @@ -745,12 +585,6 @@ protected theorem monotone (f : α →𝒄 β) : Monotone f := theorem apply_mono {f g : α →𝒄 β} {x y : α} (h₁ : f ≤ g) (h₂ : x ≤ y) : f x ≤ g y := OrderHom.apply_mono (show (f : α →o β) ≤ g from h₁) h₂ -set_option linter.deprecated false in -@[deprecated "No deprecation message was provided." (since := "2024-07-27")] -theorem ite_continuous' {p : Prop} [hp : Decidable p] (f g : α → β) (hf : Continuous' f) - (hg : Continuous' g) : Continuous' fun x => if p then f x else g x := by - split_ifs <;> simp [*] - theorem ωSup_bind {β γ : Type v} (c : Chain α) (f : α →o Part β) (g : α →o β → Part γ) : ωSup (c.map (f.partBind g)) = ωSup (c.map f) >>= ωSup (c.map g) := by apply eq_of_forall_ge_iff; intro x @@ -794,30 +628,6 @@ lemma ωScottContinuous.seq {β γ} {f : α → Part (β → γ)} {g : α → Pa simp only [seq_eq_bind_map] exact ωScottContinuous.bind hf <| ωScottContinuous.of_apply₂ fun _ ↦ ωScottContinuous.map hg -set_option linter.deprecated false in -@[deprecated ωScottContinuous.bind (since := "2024-05-29")] -theorem bind_continuous' {β γ : Type v} (f : α → Part β) (g : α → β → Part γ) : - Continuous' f → Continuous' g → Continuous' fun x => f x >>= g x - | ⟨hf, hf'⟩, ⟨hg, hg'⟩ => - Continuous.of_bundled' (OrderHom.partBind ⟨f, hf⟩ ⟨g, hg⟩) - (by intro c; rw [ωSup_bind, ← hf', ← hg']; rfl) - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.map (since := "2024-05-29")] -theorem map_continuous' {β γ : Type v} (f : β → γ) (g : α → Part β) (hg : Continuous' g) : - Continuous' fun x => f <$> g x := by - simp only [map_eq_bind_pure_comp]; apply bind_continuous' _ _ hg; apply const_continuous' - -set_option linter.deprecated false in -@[deprecated ωScottContinuous.seq (since := "2024-05-29")] -theorem seq_continuous' {β γ : Type v} (f : α → Part (β → γ)) (g : α → Part β) (hf : Continuous' f) - (hg : Continuous' g) : Continuous' fun x => f x <*> g x := by - simp only [seq_eq_bind_map] - apply bind_continuous' _ _ hf - apply OmegaCompletePartialOrder.flip₂_continuous' - intro - apply map_continuous' _ _ hg - theorem continuous (F : α →𝒄 β) (C : Chain α) : F (ωSup C) = ωSup (C.map F) := F.ωScottContinuous.map_ωSup _ diff --git a/Mathlib/Order/Part.lean b/Mathlib/Order/Part.lean index 095e4622ab550..2dd2b9174d884 100644 --- a/Mathlib/Order/Part.lean +++ b/Mathlib/Order/Part.lean @@ -62,6 +62,4 @@ def partBind (f : α →o Part β) (g : α →o β → Part γ) : α →o Part toFun x := (f x).bind (g x) monotone' := f.2.partBind g.2 -@[deprecated (since := "2024-07-04")] alias bind := partBind - end OrderHom diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index 3ca9b0f2679b2..2b231879970b4 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -86,10 +86,6 @@ theorem IsStrictOrder.swap (r) [IsStrictOrder α r] : IsStrictOrder α (swap r) theorem IsPartialOrder.swap (r) [IsPartialOrder α r] : IsPartialOrder α (swap r) := { @IsPreorder.swap α r _, @IsAntisymm.swap α r _ with } -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem IsLinearOrder.swap (r) [IsLinearOrder α r] : IsLinearOrder α (swap r) := - { @IsPartialOrder.swap α r _, @IsTotal.swap α r _ with } - protected theorem IsAsymm.isAntisymm (r) [IsAsymm α r] : IsAntisymm α r := ⟨fun _ _ h₁ h₂ => (_root_.asymm h₁ h₂).elim⟩ @@ -221,12 +217,6 @@ instance (priority := 100) isStrictOrderConnected_of_isStrictTotalOrder [IsStric ⟨fun _ _ _ h ↦ (trichotomous _ _).imp_right fun o ↦ o.elim (fun e ↦ e ▸ h) fun h' ↦ _root_.trans h' h⟩ --- see Note [lower instance priority] -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -instance (priority := 100) isStrictTotalOrder_of_isStrictTotalOrder [IsStrictTotalOrder α r] : - IsStrictWeakOrder α r := - { isStrictWeakOrder_of_isOrderConnected with } - /-! ### Well-order -/ @@ -272,10 +262,6 @@ theorem WellFounded.psigma_skipLeft (α : Type u) {β : Type v} {s : β → β (hb : WellFounded s) : WellFounded (SkipLeft α s) := psigma_revLex emptyWf.wf hb -@[deprecated (since := "2024-07-24")] alias PSigma.lex_wf := WellFounded.psigma_lex -@[deprecated (since := "2024-07-24")] alias PSigma.revLex_wf := WellFounded.psigma_revLex -@[deprecated (since := "2024-07-24")] alias PSigma.skipLeft_wf := WellFounded.psigma_skipLeft - end PSigma namespace IsWellFounded @@ -800,9 +786,6 @@ instance [LinearOrder α] : IsStrictTotalOrder α (· < ·) where instance [LinearOrder α] : IsOrderConnected α (· < ·) := by infer_instance -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -instance [LinearOrder α] : IsStrictWeakOrder α (· < ·) := by infer_instance - theorem transitive_le [Preorder α] : Transitive (@LE.le α _) := transitive_of_trans _ diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index 8cbe3710d0429..6aa3ac9b37fcd 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -187,10 +187,12 @@ theorem coe_le_coe : (a : WithBot α) ≤ b ↔ a ≤ b := by instance orderBot : OrderBot (WithBot α) where bot_le _ := fun _ h => Option.noConfusion h +-- TODO: This deprecated lemma is still used (through simp) @[simp, deprecated coe_le_coe "Don't mix Option and WithBot" (since := "2024-05-27")] theorem some_le_some : @LE.le (WithBot α) _ (Option.some a) (Option.some b) ↔ a ≤ b := coe_le_coe +-- TODO: This deprecated lemma is still used (through simp) @[simp, deprecated bot_le "Don't mix Option and WithBot" (since := "2024-05-27")] theorem none_le {a : WithBot α} : @LE.le (WithBot α) _ none a := bot_le @@ -262,13 +264,16 @@ theorem bot_lt_coe (a : α) : ⊥ < (a : WithBot α) := protected theorem not_lt_bot (a : WithBot α) : ¬a < ⊥ := fun ⟨_, h, _⟩ => Option.not_mem_none _ h +-- TODO: This deprecated lemma is still used (through simp) @[simp, deprecated coe_lt_coe "Don't mix Option and WithBot" (since := "2024-05-27")] theorem some_lt_some : @LT.lt (WithBot α) _ (Option.some a) (Option.some b) ↔ a < b := coe_lt_coe +-- TODO: This deprecated lemma is still used (through simp) @[simp, deprecated bot_lt_coe "Don't mix Option and WithBot" (since := "2024-05-27")] theorem none_lt_some (a : α) : @LT.lt (WithBot α) _ none (some a) := bot_lt_coe _ +-- TODO: This deprecated lemma is still used (through simp) @[simp, deprecated not_lt_bot "Don't mix Option and WithBot" (since := "2024-05-27")] theorem not_lt_none (a : WithBot α) : ¬@LT.lt (WithBot α) _ a none := WithBot.not_lt_bot _ @@ -793,6 +798,7 @@ theorem ofDual_le_ofDual_iff {a b : WithTop αᵒᵈ} : WithTop.ofDual a ≤ Wit theorem coe_le_coe : (a : WithTop α) ≤ b ↔ a ≤ b := by simp only [← toDual_le_toDual_iff, toDual_apply_coe, WithBot.coe_le_coe, toDual_le_toDual] +-- TODO: This deprecated lemma is still used (through simp) @[simp, deprecated coe_le_coe "Don't mix Option and WithTop" (since := "2024-05-27")] theorem some_le_some : @LE.le (WithTop α) _ (Option.some a) (Option.some b) ↔ a ≤ b := coe_le_coe @@ -800,6 +806,7 @@ theorem some_le_some : @LE.le (WithTop α) _ (Option.some a) (Option.some b) ↔ instance orderTop : OrderTop (WithTop α) where le_top := fun _ => toDual_le_toDual_iff.mp bot_le +-- TODO: This deprecated lemma is still used (through simp) @[simp, deprecated le_top "Don't mix Option and WithTop" (since := "2024-05-27")] theorem le_none {a : WithTop α} : @LE.le (WithTop α) _ a none := le_top @@ -1051,12 +1058,15 @@ protected theorem not_top_lt (a : WithTop α) : ¬⊤ < a := by rw [← toDual_lt_toDual_iff] exact WithBot.not_lt_bot _ +-- TODO: This deprecated lemma is still used (through simp) @[simp, deprecated coe_lt_coe "Don't mix Option and WithTop" (since := "2024-05-27")] theorem some_lt_some : @LT.lt (WithTop α) _ (Option.some a) (Option.some b) ↔ a < b := coe_lt_coe +-- TODO: This deprecated lemma is still used (through simp) @[simp, deprecated coe_lt_top "Don't mix Option and WithTop" (since := "2024-05-27")] theorem some_lt_none (a : α) : @LT.lt (WithTop α) _ (Option.some a) none := coe_lt_top a +-- TODO: This deprecated lemma is still used (through simp) @[simp, deprecated not_top_lt "Don't mix Option and WithTop" (since := "2024-05-27")] theorem not_none_lt (a : WithTop α) : ¬@LT.lt (WithTop α) _ none a := WithTop.not_top_lt _ diff --git a/Mathlib/Probability/Density.lean b/Mathlib/Probability/Density.lean index 3a90248fdc473..ec410d68a63d8 100644 --- a/Mathlib/Probability/Density.lean +++ b/Mathlib/Probability/Density.lean @@ -163,9 +163,6 @@ theorem setLIntegral_pdf_le_map {m : MeasurableSpace Ω} (X : Ω → E) (ℙ : M apply (withDensity_apply_le _ s).trans exact withDensity_pdf_le_map _ _ _ s -@[deprecated (since := "2024-06-29")] -alias set_lintegral_pdf_le_map := setLIntegral_pdf_le_map - theorem map_eq_withDensity_pdf {m : MeasurableSpace Ω} (X : Ω → E) (ℙ : Measure Ω) (μ : Measure E := by volume_tac) [hX : HasPDF X ℙ μ] : map X ℙ = μ.withDensity (pdf X ℙ μ) := by @@ -176,9 +173,6 @@ theorem map_eq_setLIntegral_pdf {m : MeasurableSpace Ω} (X : Ω → E) (ℙ : M (hs : MeasurableSet s) : map X ℙ s = ∫⁻ x in s, pdf X ℙ μ x ∂μ := by rw [← withDensity_apply _ hs, map_eq_withDensity_pdf X ℙ μ] -@[deprecated (since := "2024-06-29")] -alias map_eq_set_lintegral_pdf := map_eq_setLIntegral_pdf - namespace pdf variable {m : MeasurableSpace Ω} {ℙ : Measure Ω} {μ : Measure E} diff --git a/Mathlib/Probability/IdentDistrib.lean b/Mathlib/Probability/IdentDistrib.lean index 9ce855cb48da8..9f22fb58ba75f 100644 --- a/Mathlib/Probability/IdentDistrib.lean +++ b/Mathlib/Probability/IdentDistrib.lean @@ -202,9 +202,6 @@ theorem eLpNorm_eq [NormedAddCommGroup γ] [OpensMeasurableSpace γ] (h : IdentD exact h.comp (Measurable.pow_const (measurable_coe_nnreal_ennreal.comp measurable_nnnorm) p.toReal) -@[deprecated (since := "2024-07-27")] -alias snorm_eq := eLpNorm_eq - theorem memℒp_snd [NormedAddCommGroup γ] [BorelSpace γ] {p : ℝ≥0∞} (h : IdentDistrib f g μ ν) (hf : Memℒp f p μ) : Memℒp g p ν := by refine ⟨h.aestronglyMeasurable_snd hf.aestronglyMeasurable, ?_⟩ diff --git a/Mathlib/Probability/Kernel/Basic.lean b/Mathlib/Probability/Kernel/Basic.lean index 6317d5f6bd5e0..58cc5877dabca 100644 --- a/Mathlib/Probability/Kernel/Basic.lean +++ b/Mathlib/Probability/Kernel/Basic.lean @@ -86,18 +86,12 @@ theorem setLIntegral_deterministic' {f : β → ℝ≥0∞} {g : α → β} {a : ∫⁻ x in s, f x ∂deterministic g hg a = if g a ∈ s then f (g a) else 0 := by rw [deterministic_apply, setLIntegral_dirac' hf hs] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_deterministic' := setLIntegral_deterministic' - @[simp] theorem setLIntegral_deterministic {f : β → ℝ≥0∞} {g : α → β} {a : α} (hg : Measurable g) [MeasurableSingletonClass β] (s : Set β) [Decidable (g a ∈ s)] : ∫⁻ x in s, f x ∂deterministic g hg a = if g a ∈ s then f (g a) else 0 := by rw [deterministic_apply, setLIntegral_dirac f s] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_deterministic := setLIntegral_deterministic - end Deterministic section Id @@ -216,9 +210,6 @@ theorem lintegral_const {f : β → ℝ≥0∞} {μ : Measure β} {a : α} : theorem setLIntegral_const {f : β → ℝ≥0∞} {μ : Measure β} {a : α} {s : Set β} : ∫⁻ x in s, f x ∂const α μ a = ∫⁻ x in s, f x ∂μ := by rw [const_apply] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_const := setLIntegral_const - end Const /-- In a countable space with measurable singletons, every function `α → MeasureTheory.Measure β` @@ -262,9 +253,6 @@ theorem setLIntegral_restrict (κ : Kernel α β) (hs : MeasurableSet s) (a : α (t : Set β) : ∫⁻ b in t, f b ∂κ.restrict hs a = ∫⁻ b in t ∩ s, f b ∂κ a := by rw [restrict_apply, Measure.restrict_restrict' hs] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_restrict := setLIntegral_restrict - instance IsFiniteKernel.restrict (κ : Kernel α β) [IsFiniteKernel κ] (hs : MeasurableSet s) : IsFiniteKernel (κ.restrict hs) := by @@ -386,9 +374,6 @@ theorem setLIntegral_piecewise (a : α) (g : β → ℝ≥0∞) (t : Set β) : if a ∈ s then ∫⁻ b in t, g b ∂κ a else ∫⁻ b in t, g b ∂η a := by simp_rw [piecewise_apply]; split_ifs <;> rfl -@[deprecated (since := "2024-06-29")] -alias set_lintegral_piecewise := setLIntegral_piecewise - end Piecewise lemma exists_ae_eq_isMarkovKernel {μ : Measure α} diff --git a/Mathlib/Probability/Kernel/Composition/Basic.lean b/Mathlib/Probability/Kernel/Composition/Basic.lean index 3a6ccc33841e1..9749d7ab19190 100644 --- a/Mathlib/Probability/Kernel/Composition/Basic.lean +++ b/Mathlib/Probability/Kernel/Composition/Basic.lean @@ -464,27 +464,18 @@ theorem setLIntegral_compProd (κ : Kernel α β) [IsSFiniteKernel κ] (η : Ker simp_rw [← Kernel.restrict_apply (κ ⊗ₖ η) (hs.prod ht), ← compProd_restrict hs ht, lintegral_compProd _ _ _ hf, Kernel.restrict_apply] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_compProd := setLIntegral_compProd - theorem setLIntegral_compProd_univ_right (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel (α × β) γ) [IsSFiniteKernel η] (a : α) {f : β × γ → ℝ≥0∞} (hf : Measurable f) {s : Set β} (hs : MeasurableSet s) : ∫⁻ z in s ×ˢ Set.univ, f z ∂(κ ⊗ₖ η) a = ∫⁻ x in s, ∫⁻ y, f (x, y) ∂η (a, x) ∂κ a := by simp_rw [setLIntegral_compProd κ η a hf hs MeasurableSet.univ, Measure.restrict_univ] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_compProd_univ_right := setLIntegral_compProd_univ_right - theorem setLIntegral_compProd_univ_left (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel (α × β) γ) [IsSFiniteKernel η] (a : α) {f : β × γ → ℝ≥0∞} (hf : Measurable f) {t : Set γ} (ht : MeasurableSet t) : ∫⁻ z in Set.univ ×ˢ t, f z ∂(κ ⊗ₖ η) a = ∫⁻ x, ∫⁻ y in t, f (x, y) ∂η (a, x) ∂κ a := by simp_rw [setLIntegral_compProd κ η a hf MeasurableSet.univ ht, Measure.restrict_univ] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_compProd_univ_left := setLIntegral_compProd_univ_left - end Lintegral theorem compProd_eq_tsum_compProd (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel (α × β) γ) diff --git a/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean index 3b0cddf917b80..502cec66aa178 100644 --- a/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean @@ -251,23 +251,14 @@ theorem setIntegral_compProd {f : β × γ → E} {s : Set β} {t : Set γ} (hs · simp_rw [Kernel.restrict_apply] · rw [compProd_restrict, Kernel.restrict_apply]; exact hf -@[deprecated (since := "2024-04-17")] -alias set_integral_compProd := setIntegral_compProd - theorem setIntegral_compProd_univ_right (f : β × γ → E) {s : Set β} (hs : MeasurableSet s) (hf : IntegrableOn f (s ×ˢ univ) ((κ ⊗ₖ η) a)) : ∫ z in s ×ˢ univ, f z ∂(κ ⊗ₖ η) a = ∫ x in s, ∫ y, f (x, y) ∂η (a, x) ∂κ a := by simp_rw [setIntegral_compProd hs MeasurableSet.univ hf, Measure.restrict_univ] -@[deprecated (since := "2024-04-17")] -alias set_integral_compProd_univ_right := setIntegral_compProd_univ_right - theorem setIntegral_compProd_univ_left (f : β × γ → E) {t : Set γ} (ht : MeasurableSet t) (hf : IntegrableOn f (univ ×ˢ t) ((κ ⊗ₖ η) a)) : ∫ z in univ ×ˢ t, f z ∂(κ ⊗ₖ η) a = ∫ x, ∫ y in t, f (x, y) ∂η (a, x) ∂κ a := by simp_rw [setIntegral_compProd MeasurableSet.univ ht hf, Measure.restrict_univ] -@[deprecated (since := "2024-04-17")] -alias set_integral_compProd_univ_left := setIntegral_compProd_univ_left - end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean index d39c73a222f5c..a5ef4c3f8bdad 100644 --- a/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean @@ -163,9 +163,6 @@ lemma setLIntegral_compProd [SFinite μ] [IsSFiniteKernel κ] rw [compProd, Kernel.setLIntegral_compProd _ _ _ hf hs ht] simp -@[deprecated (since := "2024-06-29")] -alias set_lintegral_compProd := setLIntegral_compProd - lemma integrable_compProd_iff [SFinite μ] [IsSFiniteKernel κ] {E : Type*} [NormedAddCommGroup E] {f : α × β → E} (hf : AEStronglyMeasurable f (μ ⊗ₘ κ)) : Integrable f (μ ⊗ₘ κ) ↔ @@ -189,9 +186,6 @@ lemma setIntegral_compProd [SFinite μ] [IsSFiniteKernel κ] {E : Type*} rw [compProd, ProbabilityTheory.setIntegral_compProd hs ht hf] simp -@[deprecated (since := "2024-04-17")] -alias set_integral_compProd := setIntegral_compProd - end Integral lemma dirac_compProd_apply [MeasurableSingletonClass α] {a : α} [IsSFiniteKernel κ] diff --git a/Mathlib/Probability/Kernel/CondDistrib.lean b/Mathlib/Probability/Kernel/CondDistrib.lean index a0933088f0f3d..2b7cb94c1d9ee 100644 --- a/Mathlib/Probability/Kernel/CondDistrib.lean +++ b/Mathlib/Probability/Kernel/CondDistrib.lean @@ -191,18 +191,12 @@ theorem setLIntegral_preimage_condDistrib (hX : Measurable X) (hY : AEMeasurable Measure.fst_map_prod_mk₀ hY, Measure.setLIntegral_condKernel_eq_measure_prod ht hs, Measure.map_apply_of_aemeasurable (hX.aemeasurable.prod_mk hY) (ht.prod hs), mk_preimage_prod] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_preimage_condDistrib := setLIntegral_preimage_condDistrib - theorem setLIntegral_condDistrib_of_measurableSet (hX : Measurable X) (hY : AEMeasurable Y μ) (hs : MeasurableSet s) {t : Set α} (ht : MeasurableSet[mβ.comap X] t) : ∫⁻ a in t, condDistrib Y X μ (X a) s ∂μ = μ (t ∩ Y ⁻¹' s) := by obtain ⟨t', ht', rfl⟩ := ht rw [setLIntegral_preimage_condDistrib hX hY hs ht'] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_condDistrib_of_measurableSet := setLIntegral_condDistrib_of_measurableSet - /-- For almost every `a : α`, the `condDistrib Y X μ` kernel applied to `X a` and a measurable set `s` is equal to the conditional expectation of the indicator of `Y ⁻¹' s`. -/ theorem condDistrib_ae_eq_condExp (hX : Measurable X) (hY : Measurable Y) (hs : MeasurableSet s) : diff --git a/Mathlib/Probability/Kernel/Defs.lean b/Mathlib/Probability/Kernel/Defs.lean index ba2c845a6a796..6939623a21ea4 100644 --- a/Mathlib/Probability/Kernel/Defs.lean +++ b/Mathlib/Probability/Kernel/Defs.lean @@ -59,8 +59,6 @@ structure Kernel (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] where Do not use this lemma directly. Use `Kernel.measurable` instead. -/ measurable' : Measurable toFun -@[deprecated (since := "2024-07-22")] alias kernel := Kernel - /-- Notation for `Kernel` with respect to a non-standard σ-algebra in the domain. -/ scoped notation "Kernel[" mα "]" α:arg β:arg => @Kernel α β mα _ diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index 18477b6312708..a8b49524b71b9 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -99,9 +99,6 @@ lemma setIntegral_stieltjesOfMeasurableRat_rat (hf : IsRatCondKernelCDF f κ ν) rw [setIntegral_congr_ae hs (g := fun b ↦ f (a, b) q) ?_, hf.setIntegral a hs] filter_upwards [stieltjesOfMeasurableRat_ae_eq hf a q] with b hb using fun _ ↦ hb -@[deprecated (since := "2024-04-17")] -alias set_integral_stieltjesOfMeasurableRat_rat := setIntegral_stieltjesOfMeasurableRat_rat - lemma setLIntegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel κ] (hf : IsRatCondKernelCDF f κ ν) (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : ∫⁻ b in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) q) ∂(ν a) @@ -114,9 +111,6 @@ lemma setLIntegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel κ] (hf : IsRatC exact hf.integrable a q · exact ae_of_all _ (fun x ↦ stieltjesOfMeasurableRat_nonneg _ _ _) -@[deprecated (since := "2024-06-29")] -alias set_lintegral_stieltjesOfMeasurableRat_rat := setLIntegral_stieltjesOfMeasurableRat_rat - lemma setLIntegral_stieltjesOfMeasurableRat [IsFiniteKernel κ] (hf : IsRatCondKernelCDF f κ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : ∫⁻ b in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) x) ∂(ν a) @@ -176,9 +170,6 @@ lemma setLIntegral_stieltjesOfMeasurableRat [IsFiniteKernel κ] (hf : IsRatCondK · exact fun i ↦ (hs.prod measurableSet_Iic).nullMeasurableSet · exact ⟨h_nonempty.some, measure_ne_top _ _⟩ -@[deprecated (since := "2024-06-29")] -alias set_lintegral_stieltjesOfMeasurableRat := setLIntegral_stieltjesOfMeasurableRat - lemma lintegral_stieltjesOfMeasurableRat [IsFiniteKernel κ] (hf : IsRatCondKernelCDF f κ ν) (a : α) (x : ℝ) : ∫⁻ b, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) x) ∂(ν a) @@ -213,9 +204,6 @@ lemma setIntegral_stieltjesOfMeasurableRat [IsFiniteKernel κ] (hf : IsRatCondKe · exact (integrable_stieltjesOfMeasurableRat hf _ _).restrict · exact ae_of_all _ (fun _ ↦ stieltjesOfMeasurableRat_nonneg _ _ _) -@[deprecated (since := "2024-04-17")] -alias set_integral_stieltjesOfMeasurableRat := setIntegral_stieltjesOfMeasurableRat - lemma integral_stieltjesOfMeasurableRat [IsFiniteKernel κ] (hf : IsRatCondKernelCDF f κ ν) (a : α) (x : ℝ) : ∫ b, stieltjesOfMeasurableRat f hf.measurable (a, b) x ∂(ν a) @@ -380,10 +368,6 @@ lemma IsRatCondKernelCDFAux.setIntegral_iInf_rat_gt (hf : IsRatCondKernelCDFAux · exact (hf.integrable_iInf_rat_gt _ _).integrableOn · filter_upwards [hf.mono a] with c h_mono using le_ciInf (fun r ↦ h_mono (le_of_lt r.prop)) -@[deprecated (since := "2024-04-17")] -alias IsRatCondKernelCDFAux.set_integral_iInf_rat_gt := - IsRatCondKernelCDFAux.setIntegral_iInf_rat_gt - lemma IsRatCondKernelCDFAux.iInf_rat_gt_eq (hf : IsRatCondKernelCDFAux f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : ∀ᵐ t ∂(ν a), ∀ q : ℚ, ⨅ r : Ioi q, f (a, t) r = f (a, t) q := by @@ -447,9 +431,6 @@ lemma IsCondKernelCDF.setLIntegral [IsFiniteKernel κ] (ae_of_all _ (fun _ ↦ hf.nonneg _ _)), hf.setIntegral a hs x, ENNReal.ofReal_toReal] exact measure_ne_top _ _ -@[deprecated (since := "2024-06-29")] -alias IsCondKernelCDF.set_lintegral := IsCondKernelCDF.setLIntegral - lemma IsCondKernelCDF.lintegral [IsFiniteKernel κ] {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f κ ν) (a : α) (x : ℝ) : ∫⁻ b, ENNReal.ofReal (f (a, b) x) ∂(ν a) = κ a (univ ×ˢ Iic x) := by @@ -504,9 +485,6 @@ lemma setLIntegral_toKernel_Iic [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ simp_rw [IsCondKernelCDF.toKernel_Iic] exact hf.setLIntegral _ hs _ -@[deprecated (since := "2024-06-29")] -alias set_lintegral_toKernel_Iic := setLIntegral_toKernel_Iic - lemma setLIntegral_toKernel_univ [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) (a : α) {s : Set β} (hs : MeasurableSet s) : ∫⁻ b in s, hf.toKernel f (a, b) univ ∂(ν a) = κ a (s ×ˢ univ) := by @@ -526,9 +504,6 @@ lemma setLIntegral_toKernel_univ [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ · refine Monotone.directed_le fun i j hij t ↦ measure_mono (Iic_subset_Iic.mpr ?_) exact mod_cast hij -@[deprecated (since := "2024-06-29")] -alias set_lintegral_toKernel_univ := setLIntegral_toKernel_univ - lemma lintegral_toKernel_univ [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) (a : α) : ∫⁻ b, hf.toKernel f (a, b) univ ∂(ν a) = κ a univ := by rw [← setLIntegral_univ, setLIntegral_toKernel_univ hf a MeasurableSet.univ, univ_prod_univ] @@ -573,9 +548,6 @@ lemma setLIntegral_toKernel_prod [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ · exact fun i ↦ ((Kernel.measurable_coe _ (hf_meas i)).comp measurable_prod_mk_left).aemeasurable.restrict -@[deprecated (since := "2024-06-29")] -alias set_lintegral_toKernel_prod := setLIntegral_toKernel_prod - open scoped Function in -- required for scoped `on` notation lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) (a : α) {s : Set (β × ℝ)} (hs : MeasurableSet s) : diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean index 25bf36467a984..44f48c01bbfef 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean @@ -154,9 +154,6 @@ theorem setLIntegral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) {s : Set α · rw [(_ : (1 : α → ℝ≥0∞) = fun _ ↦ 1)] exacts [measurable_const, rfl] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_preCDF_fst := setLIntegral_preCDF_fst - lemma lintegral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) [IsFiniteMeasure ρ] : ∫⁻ x, preCDF ρ r x ∂ρ.fst = ρ.IicSnd r univ := by rw [← setLIntegral_univ, setLIntegral_preCDF_fst ρ r MeasurableSet.univ] @@ -187,9 +184,6 @@ lemma setIntegral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) {s : Set α} ( filter_upwards [preCDF_le_one ρ] with a ha exact (ha r).trans_lt ENNReal.one_lt_top -@[deprecated (since := "2024-04-17")] -alias set_integral_preCDF_fst := setIntegral_preCDF_fst - lemma integral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) [IsFiniteMeasure ρ] : ∫ x, (preCDF ρ r x).toReal ∂ρ.fst = (ρ.IicSnd r univ).toReal := by rw [← setIntegral_univ, setIntegral_preCDF_fst ρ _ MeasurableSet.univ] @@ -296,9 +290,6 @@ theorem setLIntegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x ∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x) := (isCondKernelCDF_condCDF ρ).setLIntegral () hs x -@[deprecated (since := "2024-06-29")] -alias set_lintegral_condCDF := setLIntegral_condCDF - theorem lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : ∫⁻ a, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (univ ×ˢ Iic x) := (isCondKernelCDF_condCDF ρ).lintegral () x @@ -311,9 +302,6 @@ theorem setIntegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : (hs : MeasurableSet s) : ∫ a in s, condCDF ρ a x ∂ρ.fst = (ρ (s ×ˢ Iic x)).toReal := (isCondKernelCDF_condCDF ρ).setIntegral () hs x -@[deprecated (since := "2024-04-17")] -alias set_integral_condCDF := setIntegral_condCDF - theorem integral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : ∫ a, condCDF ρ a x ∂ρ.fst = (ρ (univ ×ˢ Iic x)).toReal := (isCondKernelCDF_condCDF ρ).integral () x diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 1ac6c0d62bdfe..a1aa1648a693e 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -186,9 +186,6 @@ lemma eLpNorm_densityProcess_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (s : densityProcess_le_one hκν n a x s] · simp -@[deprecated (since := "2024-07-27")] -alias snorm_densityProcess_le := eLpNorm_densityProcess_le - lemma integrable_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : Integrable (fun x ↦ densityProcess κ ν n a x s) (ν a) := by @@ -238,9 +235,6 @@ lemma setIntegral_densityProcess_of_mem (hκν : fst κ ≤ ν) [hν : IsFiniteK rw [div_eq_mul_inv, mul_assoc, ENNReal.inv_mul_cancel h0, mul_one] exact measure_ne_top _ _ -@[deprecated (since := "2024-04-17")] -alias set_integral_densityProcess_of_mem := setIntegral_densityProcess_of_mem - open scoped Function in -- required for scoped `on` notation lemma setIntegral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} @@ -271,9 +265,6 @@ lemma setIntegral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] · exact h_disj · exact (integrable_densityProcess hκν _ _ hs).integrableOn -@[deprecated (since := "2024-04-17")] -alias set_integral_densityProcess := setIntegral_densityProcess - lemma integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : ∫ x, densityProcess κ ν n a x s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by @@ -285,9 +276,6 @@ lemma setIntegral_densityProcess_of_le (hκν : fst κ ≤ ν) ∫ x in A, densityProcess κ ν m a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal := setIntegral_densityProcess hκν m a hs ((countableFiltration γ).mono hnm A hA) -@[deprecated (since := "2024-04-17")] -alias set_integral_densityProcess_of_le := setIntegral_densityProcess_of_le - lemma condExp_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] {i j : ℕ} (hij : i ≤ j) (a : α) {s : Set β} (hs : MeasurableSet s) : (ν a)[fun x ↦ densityProcess κ ν j a x s | countableFiltration γ i] @@ -420,10 +408,6 @@ lemma tendsto_eLpNorm_one_densityProcess_limitProcess (hκν : fst κ ≤ ν) [I exact mod_cast (densityProcess_le_one hκν _ _ _ _) · simp -@[deprecated (since := "2024-07-27")] -alias tendsto_snorm_one_densityProcess_limitProcess := - tendsto_eLpNorm_one_densityProcess_limitProcess - lemma tendsto_eLpNorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel ν] (hκν : fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : Tendsto (fun n ↦ eLpNorm ((fun x ↦ densityProcess κ ν n a x s) @@ -433,10 +417,6 @@ lemma tendsto_eLpNorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel (tendsto_eLpNorm_one_densityProcess_limitProcess hκν a hs) (fun _ ↦ zero_le') (fun _ ↦ eLpNorm_restrict_le _ _ _ _) -@[deprecated (since := "2024-07-27")] -alias tendsto_snorm_one_restrict_densityProcess_limitProcess := - tendsto_eLpNorm_one_restrict_densityProcess_limitProcess - end DensityProcess section Density @@ -507,9 +487,6 @@ lemma eLpNorm_density_le (hκν : fst κ ≤ ν) (a : α) (s : Set β) : density_le_one hκν a t s] · simp -@[deprecated (since := "2024-07-27")] -alias snorm_density_le := eLpNorm_density_le - lemma integrable_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : Integrable (fun x ↦ density κ ν a x s) (ν a) := by @@ -529,9 +506,6 @@ lemma tendsto_setIntegral_densityProcess (hκν : fst κ ≤ ν) refine eLpNorm_congr_ae ?_ exact EventuallyEq.rfl.sub (density_ae_eq_limitProcess hκν a hs).symm -@[deprecated (since := "2024-04-17")] -alias tendsto_set_integral_densityProcess := tendsto_setIntegral_densityProcess - /-- Auxiliary lemma for `setIntegral_density`. -/ lemma setIntegral_density_of_measurableSet (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} @@ -551,9 +525,6 @@ lemma setIntegral_density_of_measurableSet (hκν : fst κ ≤ ν) have h := tendsto_setIntegral_densityProcess hκν a hs A rw [h.limsup_eq] -@[deprecated (since := "2024-04-17")] -alias set_integral_density_of_measurableSet := setIntegral_density_of_measurableSet - lemma integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : ∫ x, density κ ν a x s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by @@ -597,9 +568,6 @@ lemma setIntegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] · exact hf_disj.mono fun _ _ h ↦ h.set_prod_left _ _ · exact fun i ↦ (hf i).prod hs -@[deprecated (since := "2024-04-17")] -alias set_integral_density := setIntegral_density - lemma setLIntegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : ∫⁻ x in A, ENNReal.ofReal (density κ ν a x s) ∂(ν a) = κ a (A ×ˢ s) := by @@ -610,9 +578,6 @@ lemma setLIntegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] · exact (integrable_density hκν a hs).restrict · exact ae_of_all _ (fun _ ↦ density_nonneg hκν _ _ _) -@[deprecated (since := "2024-06-29")] -alias set_lintegral_density := setLIntegral_density - lemma lintegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : ∫⁻ x, ENNReal.ofReal (density κ ν a x s) ∂(ν a) = κ a (univ ×ˢ s) := by diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index b3b312597fb26..73fe8baaff3c5 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -55,9 +55,6 @@ lemma setLIntegral_condKernel_eq_measure_prod (a : α) {s : Set β} (hs : Measur simp_rw [this] rw [lintegral_indicator hs] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_condKernel_eq_measure_prod := setLIntegral_condKernel_eq_measure_prod - lemma lintegral_condKernel (hf : Measurable f) (a : α) : ∫⁻ b, ∫⁻ ω, f (b, ω) ∂(Kernel.condKernel κ (a, b)) ∂(Kernel.fst κ a) = ∫⁻ x, f x ∂(κ a) := by conv_rhs => rw [← κ.disintegrate κ.condKernel] @@ -70,27 +67,18 @@ lemma setLIntegral_condKernel (hf : Measurable f) (a : α) {s : Set β} conv_rhs => rw [← κ.disintegrate κ.condKernel] rw [Kernel.setLIntegral_compProd _ _ _ hf hs ht] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_condKernel := setLIntegral_condKernel - lemma setLIntegral_condKernel_univ_right (hf : Measurable f) (a : α) {s : Set β} (hs : MeasurableSet s) : ∫⁻ b in s, ∫⁻ ω, f (b, ω) ∂(Kernel.condKernel κ (a, b)) ∂(Kernel.fst κ a) = ∫⁻ x in s ×ˢ Set.univ, f x ∂(κ a) := by rw [← setLIntegral_condKernel hf a hs MeasurableSet.univ]; simp_rw [Measure.restrict_univ] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_condKernel_univ_right := setLIntegral_condKernel_univ_right - lemma setLIntegral_condKernel_univ_left (hf : Measurable f) (a : α) {t : Set Ω} (ht : MeasurableSet t) : ∫⁻ b, ∫⁻ ω in t, f (b, ω) ∂(Kernel.condKernel κ (a, b)) ∂(Kernel.fst κ a) = ∫⁻ x in Set.univ ×ˢ t, f x ∂(κ a) := by rw [← setLIntegral_condKernel hf a MeasurableSet.univ ht]; simp_rw [Measure.restrict_univ] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_condKernel_univ_left := setLIntegral_condKernel_univ_left - end Lintegral section Integral @@ -119,27 +107,18 @@ lemma setIntegral_condKernel (a : α) {s : Set β} (hs : MeasurableSet s) rw [← κ.disintegrate κ.condKernel] at hf rw [setIntegral_compProd hs ht hf] -@[deprecated (since := "2024-04-17")] -alias set_integral_condKernel := setIntegral_condKernel - lemma setIntegral_condKernel_univ_right (a : α) {s : Set β} (hs : MeasurableSet s) (hf : IntegrableOn f (s ×ˢ Set.univ) (κ a)) : ∫ b in s, ∫ ω, f (b, ω) ∂(Kernel.condKernel κ (a, b)) ∂(Kernel.fst κ a) = ∫ x in s ×ˢ Set.univ, f x ∂(κ a) := by rw [← setIntegral_condKernel a hs MeasurableSet.univ hf]; simp_rw [Measure.restrict_univ] -@[deprecated (since := "2024-04-17")] -alias set_integral_condKernel_univ_right := setIntegral_condKernel_univ_right - lemma setIntegral_condKernel_univ_left (a : α) {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (Set.univ ×ˢ t) (κ a)) : ∫ b, ∫ ω in t, f (b, ω) ∂(Kernel.condKernel κ (a, b)) ∂(Kernel.fst κ a) = ∫ x in Set.univ ×ˢ t, f x ∂(κ a) := by rw [← setIntegral_condKernel a MeasurableSet.univ ht hf]; simp_rw [Measure.restrict_univ] -@[deprecated (since := "2024-04-17")] -alias set_integral_condKernel_univ_left := setIntegral_condKernel_univ_left - end Integral end ProbabilityTheory @@ -174,9 +153,6 @@ lemma setLIntegral_condKernel_eq_measure_prod {s : Set β} (hs : MeasurableSet s simp_rw [this] rw [lintegral_indicator hs] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_condKernel_eq_measure_prod := setLIntegral_condKernel_eq_measure_prod - lemma lintegral_condKernel (hf : Measurable f) : ∫⁻ b, ∫⁻ ω, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫⁻ x, f x ∂ρ := by conv_rhs => rw [← ρ.disintegrate ρ.condKernel] @@ -189,27 +165,18 @@ lemma setLIntegral_condKernel (hf : Measurable f) {s : Set β} conv_rhs => rw [← ρ.disintegrate ρ.condKernel] rw [setLIntegral_compProd hf hs ht] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_condKernel := setLIntegral_condKernel - lemma setLIntegral_condKernel_univ_right (hf : Measurable f) {s : Set β} (hs : MeasurableSet s) : ∫⁻ b in s, ∫⁻ ω, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫⁻ x in s ×ˢ Set.univ, f x ∂ρ := by rw [← setLIntegral_condKernel hf hs MeasurableSet.univ]; simp_rw [Measure.restrict_univ] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_condKernel_univ_right := setLIntegral_condKernel_univ_right - lemma setLIntegral_condKernel_univ_left (hf : Measurable f) {t : Set Ω} (ht : MeasurableSet t) : ∫⁻ b, ∫⁻ ω in t, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫⁻ x in Set.univ ×ˢ t, f x ∂ρ := by rw [← setLIntegral_condKernel hf MeasurableSet.univ ht]; simp_rw [Measure.restrict_univ] -@[deprecated (since := "2024-06-29")] -alias set_lintegral_condKernel_univ_left := setLIntegral_condKernel_univ_left - end Lintegral section Integral @@ -236,25 +203,16 @@ lemma setIntegral_condKernel {s : Set β} (hs : MeasurableSet s) rw [← ρ.disintegrate ρ.condKernel] at hf rw [setIntegral_compProd hs ht hf] -@[deprecated (since := "2024-04-17")] -alias set_integral_condKernel := setIntegral_condKernel - lemma setIntegral_condKernel_univ_right {s : Set β} (hs : MeasurableSet s) (hf : IntegrableOn f (s ×ˢ Set.univ) ρ) : ∫ b in s, ∫ ω, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫ x in s ×ˢ Set.univ, f x ∂ρ := by rw [← setIntegral_condKernel hs MeasurableSet.univ hf]; simp_rw [Measure.restrict_univ] -@[deprecated (since := "2024-04-17")] -alias set_integral_condKernel_univ_right := setIntegral_condKernel_univ_right - lemma setIntegral_condKernel_univ_left {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (Set.univ ×ˢ t) ρ) : ∫ b, ∫ ω in t, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫ x in Set.univ ×ˢ t, f x ∂ρ := by rw [← setIntegral_condKernel MeasurableSet.univ ht hf]; simp_rw [Measure.restrict_univ] -@[deprecated (since := "2024-04-17")] -alias set_integral_condKernel_univ_left := setIntegral_condKernel_univ_left - end Integral end MeasureTheory.Measure diff --git a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean index e0368a6da934d..04a7544f86693 100644 --- a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean @@ -159,10 +159,6 @@ instance condKernelUnitReal.instIsCondKernel (κ : Kernel Unit (α × ℝ)) [IsF κ.IsCondKernel κ.condKernelUnitReal where disintegrate := by rw [condKernelUnitReal, compProd_toKernel]; ext; simp -@[deprecated disintegrate (since := "2024-07-26")] -lemma compProd_fst_condKernelUnitReal (κ : Kernel Unit (α × ℝ)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernelUnitReal κ = κ := disintegrate _ _ - end Real section BorelSnd @@ -327,10 +323,6 @@ instance condKernelBorel.instIsCondKernel (κ : Kernel α (γ × Ω)) [IsFiniteK disintegrate := by rw [condKernelBorel, compProd_fst_borelMarkovFromReal _ _ (compProd_fst_condKernelReal _)] -@[deprecated disintegrate (since := "2024-07-26")] -lemma compProd_fst_condKernelBorel (κ : Kernel α (γ × Ω)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernelBorel κ = κ := disintegrate _ _ - end CountablyGenerated section Unit @@ -352,10 +344,6 @@ instance condKernelUnitBorel.instIsCondKernel : κ.IsCondKernel κ.condKernelUni disintegrate := by rw [condKernelUnitBorel, compProd_fst_borelMarkovFromReal _ _ (disintegrate _ _)] -@[deprecated disintegrate (since := "2024-07-26")] -lemma compProd_fst_condKernelUnitBorel (κ : Kernel Unit (α × Ω)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernelUnitBorel κ = κ := disintegrate _ _ - end Unit section Measure @@ -391,22 +379,6 @@ instance _root_.MeasureTheory.Measure.instIsMarkovKernelCondKernel rw [Measure.condKernel] infer_instance -/-- **Disintegration** of finite product measures on `α × Ω`, where `Ω` is standard Borel. Such a -measure can be written as the composition-product of `ρ.fst` (marginal measure over `α`) and -a Markov kernel from `α` to `Ω`. We call that Markov kernel `ρ.condKernel`. -/ -@[deprecated Measure.disintegrate (since := "2024-07-24")] -lemma _root_.MeasureTheory.Measure.compProd_fst_condKernel - (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : - ρ.fst ⊗ₘ ρ.condKernel = ρ := ρ.disintegrate ρ.condKernel - -set_option linter.unusedVariables false in -/-- Auxiliary lemma for `condKernel_apply_of_ne_zero`. -/ -@[deprecated Measure.IsCondKernel.apply_of_ne_zero (since := "2024-07-24"), nolint unusedArguments] -lemma _root_.MeasureTheory.Measure.condKernel_apply_of_ne_zero_of_measurableSet - [MeasurableSingletonClass α] {x : α} (hx : ρ.fst {x} ≠ 0) {s : Set Ω} (hs : MeasurableSet s) : - ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := - Measure.IsCondKernel.apply_of_ne_zero _ _ hx _ - /-- If the singleton `{x}` has non-zero mass for `ρ.fst`, then for all `s : Set Ω`, `ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s)` . -/ lemma _root_.MeasureTheory.Measure.condKernel_apply_of_ne_zero [MeasurableSingletonClass α] @@ -416,17 +388,6 @@ lemma _root_.MeasureTheory.Measure.condKernel_apply_of_ne_zero [MeasurableSingle end Measure -section Countable - -variable [Countable α] - -@[deprecated disintegrate (since := "2024-07-24")] -lemma compProd_fst_condKernelCountable (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernelCountable (fun a ↦ (κ a).condKernel) - (fun x y h ↦ by simp [apply_congr_of_mem_measurableAtom _ h]) = κ := disintegrate _ _ - -end Countable - section CountableOrCountablyGenerated variable [h : CountableOrCountablyGenerated α β] (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] @@ -451,11 +412,6 @@ instance instIsMarkovKernelCondKernel : IsMarkovKernel (condKernel κ) := by instance condKernel.instIsCondKernel : κ.IsCondKernel κ.condKernel where disintegrate := by rw [condKernel_def]; split_ifs with hα <;> exact disintegrate _ _ -/-- **Disintegration** of finite kernels. -The composition-product of `fst κ` and `condKernel κ` is equal to `κ`. -/ -@[deprecated Kernel.disintegrate (since := "2024-07-26")] -lemma compProd_fst_condKernel : fst κ ⊗ₖ condKernel κ = κ := κ.disintegrate κ.condKernel - end CountableOrCountablyGenerated end ProbabilityTheory.Kernel diff --git a/Mathlib/Probability/Kernel/Integral.lean b/Mathlib/Probability/Kernel/Integral.lean index d586ae6a66158..fa375f0e70761 100644 --- a/Mathlib/Probability/Kernel/Integral.lean +++ b/Mathlib/Probability/Kernel/Integral.lean @@ -59,18 +59,12 @@ theorem setIntegral_deterministic' (hg : Measurable g) ∫ x in s, f x ∂deterministic g hg a = if g a ∈ s then f (g a) else 0 := by rw [deterministic_apply, setIntegral_dirac' hf _ hs] -@[deprecated (since := "2024-04-17")] -alias set_integral_deterministic' := setIntegral_deterministic' - @[simp] theorem setIntegral_deterministic [MeasurableSingletonClass β] (hg : Measurable g) (s : Set β) [Decidable (g a ∈ s)] : ∫ x in s, f x ∂deterministic g hg a = if g a ∈ s then f (g a) else 0 := by rw [deterministic_apply, setIntegral_dirac f _ s] -@[deprecated (since := "2024-04-17")] -alias set_integral_deterministic := setIntegral_deterministic - end Deterministic section Const @@ -83,9 +77,6 @@ theorem integral_const {μ : Measure β} : ∫ x, f x ∂const α μ a = ∫ x, theorem setIntegral_const {μ : Measure β} {s : Set β} : ∫ x in s, f x ∂const α μ a = ∫ x in s, f x ∂μ := by rw [const_apply] -@[deprecated (since := "2024-04-17")] -alias set_integral_const := setIntegral_const - end Const section Restrict @@ -102,9 +93,6 @@ theorem setIntegral_restrict (hs : MeasurableSet s) (t : Set β) : ∫ x in t, f x ∂κ.restrict hs a = ∫ x in t ∩ s, f x ∂κ a := by rw [restrict_apply, Measure.restrict_restrict' hs] -@[deprecated (since := "2024-04-17")] -alias set_integral_restrict := setIntegral_restrict - end Restrict section Piecewise @@ -120,9 +108,6 @@ theorem setIntegral_piecewise (a : α) (g : β → E) (t : Set β) : if a ∈ s then ∫ b in t, g b ∂κ a else ∫ b in t, g b ∂η a := by simp_rw [piecewise_apply]; split_ifs <;> rfl -@[deprecated (since := "2024-04-17")] -alias set_integral_piecewise := setIntegral_piecewise - end Piecewise end Kernel diff --git a/Mathlib/Probability/Kernel/MeasurableLIntegral.lean b/Mathlib/Probability/Kernel/MeasurableLIntegral.lean index a5c878aab3b61..7bf0464954ee3 100644 --- a/Mathlib/Probability/Kernel/MeasurableLIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableLIntegral.lean @@ -173,10 +173,6 @@ theorem _root_.Measurable.setLIntegral_kernel_prod_right {f : α → β → ℝ Measurable fun a => ∫⁻ b in s, f a b ∂κ a := by simp_rw [← lintegral_restrict κ hs]; exact hf.lintegral_kernel_prod_right -@[deprecated (since := "2024-06-29")] -alias _root_.Measurable.set_lintegral_kernel_prod_right := - _root_.Measurable.setLIntegral_kernel_prod_right - theorem _root_.Measurable.lintegral_kernel_prod_left' {f : β × α → ℝ≥0∞} (hf : Measurable f) : Measurable fun y => ∫⁻ x, f (x, y) ∂κ y := (measurable_swap_iff.mpr hf).lintegral_kernel_prod_right' @@ -190,10 +186,6 @@ theorem _root_.Measurable.setLIntegral_kernel_prod_left {f : β → α → ℝ Measurable fun b => ∫⁻ a in s, f a b ∂κ b := by simp_rw [← lintegral_restrict κ hs]; exact hf.lintegral_kernel_prod_left -@[deprecated (since := "2024-06-29")] -alias _root_.Measurable.set_lintegral_kernel_prod_left := - _root_.Measurable.setLIntegral_kernel_prod_left - theorem _root_.Measurable.lintegral_kernel {f : β → ℝ≥0∞} (hf : Measurable f) : Measurable fun a => ∫⁻ b, f b ∂κ a := Measurable.lintegral_kernel_prod_right (hf.comp measurable_snd) @@ -204,9 +196,6 @@ theorem _root_.Measurable.setLIntegral_kernel {f : β → ℝ≥0∞} (hf : Meas refine Measurable.setLIntegral_kernel_prod_right ?_ hs convert hf.comp measurable_snd -@[deprecated (since := "2024-06-29")] -alias _root_.Measurable.set_lintegral_kernel := _root_.Measurable.setLIntegral_kernel - end Lintegral end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index 3ba66e8d9aa5b..3e4d2a566b0d3 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -147,9 +147,6 @@ lemma setLIntegral_rnDerivAux (κ η : Kernel α γ) [IsFiniteKernel κ] [IsFini congr with x simp -@[deprecated (since := "2024-06-29")] -alias set_lintegral_rnDerivAux := setLIntegral_rnDerivAux - lemma withDensity_rnDerivAux (κ η : Kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel η] : withDensity (κ + η) (fun a x ↦ Real.toNNReal (rnDerivAux κ (κ + η) a x)) = κ := by ext a s hs diff --git a/Mathlib/Probability/Martingale/Basic.lean b/Mathlib/Probability/Martingale/Basic.lean index 9b7789b528d40..99f0a32ea083c 100644 --- a/Mathlib/Probability/Martingale/Basic.lean +++ b/Mathlib/Probability/Martingale/Basic.lean @@ -102,9 +102,6 @@ theorem setIntegral_eq [SigmaFiniteFiltration μ ℱ] (hf : Martingale f ℱ μ) refine setIntegral_congr_ae (ℱ.le i s hs) ?_ filter_upwards [hf.2 i j hij] with _ heq _ using heq.symm -@[deprecated (since := "2024-04-17")] -alias set_integral_eq := setIntegral_eq - theorem add (hf : Martingale f ℱ μ) (hg : Martingale g ℱ μ) : Martingale (f + g) ℱ μ := by refine ⟨hf.adapted.add hg.adapted, fun i j hij => ?_⟩ exact (condExp_add (hf.integrable j) (hg.integrable j) _).trans @@ -165,9 +162,6 @@ theorem setIntegral_le [SigmaFiniteFiltration μ ℱ] {f : ι → Ω → ℝ} (h refine setIntegral_mono_ae integrable_condExp.integrableOn (hf.integrable i).integrableOn ?_ filter_upwards [hf.2.1 i j hij] with _ heq using heq -@[deprecated (since := "2024-04-17")] -alias set_integral_le := setIntegral_le - theorem add [Preorder E] [AddLeftMono E] (hf : Supermartingale f ℱ μ) (hg : Supermartingale g ℱ μ) : Supermartingale (f + g) ℱ μ := by refine ⟨hf.1.add hg.1, fun i j hij => ?_, fun i => (hf.2.2 i).add (hg.2.2 i)⟩ @@ -232,9 +226,6 @@ theorem setIntegral_le [SigmaFiniteFiltration μ ℱ] {f : ι → Ω → ℝ} (h rw [← neg_le_neg_iff, ← integral_neg, ← integral_neg] exact Supermartingale.setIntegral_le hf.neg hij hs -@[deprecated (since := "2024-04-17")] -alias set_integral_le := setIntegral_le - theorem sub_supermartingale [Preorder E] [AddLeftMono E] (hf : Submartingale f ℱ μ) (hg : Supermartingale g ℱ μ) : Submartingale (f - g) ℱ μ := by rw [sub_eq_add_neg]; exact hf.add hg.neg @@ -279,9 +270,6 @@ theorem submartingale_of_setIntegral_le [IsFiniteMeasure μ] {f : ι → Ω → integral_sub' integrable_condExp.integrableOn (hint i).integrableOn, sub_nonneg, setIntegral_condExp (ℱ.le i) (hint j) hs] -@[deprecated (since := "2024-04-17")] -alias submartingale_of_set_integral_le := submartingale_of_setIntegral_le - theorem submartingale_of_condExp_sub_nonneg [IsFiniteMeasure μ] {f : ι → Ω → ℝ} (hadp : Adapted ℱ f) (hint : ∀ i, Integrable (f i) μ) (hf : ∀ i j, i ≤ j → 0 ≤ᵐ[μ] μ[f j - f i|ℱ i]) : Submartingale f ℱ μ := by @@ -379,9 +367,6 @@ theorem submartingale_of_setIntegral_le_succ [IsFiniteMeasure μ] {f : ℕ → · exact le_rfl · exact le_trans hk₂ (hf k s (𝒢.mono hk₁ _ hs)) -@[deprecated (since := "2024-04-17")] -alias submartingale_of_set_integral_le_succ := submartingale_of_setIntegral_le_succ - theorem supermartingale_of_setIntegral_succ_le [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} (hadp : Adapted 𝒢 f) (hint : ∀ i, Integrable (f i) μ) (hf : ∀ i, ∀ s : Set Ω, MeasurableSet[𝒢 i] s → ∫ ω in s, f (i + 1) ω ∂μ ≤ ∫ ω in s, f i ω ∂μ) : @@ -390,9 +375,6 @@ theorem supermartingale_of_setIntegral_succ_le [IsFiniteMeasure μ] {f : ℕ → refine (submartingale_of_setIntegral_le_succ hadp.neg (fun i => (hint i).neg) ?_).neg simpa only [integral_neg, Pi.neg_apply, neg_le_neg_iff] -@[deprecated (since := "2024-04-17")] -alias supermartingale_of_set_integral_succ_le := supermartingale_of_setIntegral_succ_le - theorem martingale_of_setIntegral_eq_succ [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} (hadp : Adapted 𝒢 f) (hint : ∀ i, Integrable (f i) μ) (hf : ∀ i, ∀ s : Set Ω, MeasurableSet[𝒢 i] s → ∫ ω in s, f i ω ∂μ = ∫ ω in s, f (i + 1) ω ∂μ) : @@ -400,9 +382,6 @@ theorem martingale_of_setIntegral_eq_succ [IsFiniteMeasure μ] {f : ℕ → Ω martingale_iff.2 ⟨supermartingale_of_setIntegral_succ_le hadp hint fun i s hs => (hf i s hs).ge, submartingale_of_setIntegral_le_succ hadp hint fun i s hs => (hf i s hs).le⟩ -@[deprecated (since := "2024-04-17")] -alias martingale_of_set_integral_eq_succ := martingale_of_setIntegral_eq_succ - theorem submartingale_nat [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} (hadp : Adapted 𝒢 f) (hint : ∀ i, Integrable (f i) μ) (hf : ∀ i, f i ≤ᵐ[μ] μ[f (i + 1)|𝒢 i]) : Submartingale f 𝒢 μ := by diff --git a/Mathlib/Probability/Martingale/BorelCantelli.lean b/Mathlib/Probability/Martingale/BorelCantelli.lean index 7edfbb36da6b5..2c5bfaa03eea6 100644 --- a/Mathlib/Probability/Martingale/BorelCantelli.lean +++ b/Mathlib/Probability/Martingale/BorelCantelli.lean @@ -131,9 +131,6 @@ theorem Submartingale.stoppedValue_leastGE_eLpNorm_le [IsFiniteMeasure μ] (hf : refine le_trans ?_ ((hf.stoppedValue_leastGE r).setIntegral_le (zero_le _) MeasurableSet.univ) simp_rw [stoppedValue, leastGE, hitting_of_le le_rfl, hf0, integral_zero', le_rfl] -@[deprecated (since := "2024-07-27")] -alias Submartingale.stoppedValue_leastGE_snorm_le := Submartingale.stoppedValue_leastGE_eLpNorm_le - theorem Submartingale.stoppedValue_leastGE_eLpNorm_le' [IsFiniteMeasure μ] (hf : Submartingale f ℱ μ) (hr : 0 ≤ r) (hf0 : f 0 = 0) (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) (i : ℕ) : @@ -142,9 +139,6 @@ theorem Submartingale.stoppedValue_leastGE_eLpNorm_le' [IsFiniteMeasure μ] refine (hf.stoppedValue_leastGE_eLpNorm_le hr hf0 hbdd i).trans ?_ simp [ENNReal.coe_toNNReal (measure_ne_top μ _), ENNReal.coe_toNNReal] -@[deprecated (since := "2024-07-27")] -alias Submartingale.stoppedValue_leastGE_snorm_le' := Submartingale.stoppedValue_leastGE_eLpNorm_le' - /-- This lemma is superseded by `Submartingale.bddAbove_iff_exists_tendsto`. -/ theorem Submartingale.exists_tendsto_of_abs_bddAbove_aux [IsFiniteMeasure μ] (hf : Submartingale f ℱ μ) (hf0 : f 0 = 0) (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : diff --git a/Mathlib/Probability/Martingale/Convergence.lean b/Mathlib/Probability/Martingale/Convergence.lean index 00ffa32b2454b..6f4c9d1007a3e 100644 --- a/Mathlib/Probability/Martingale/Convergence.lean +++ b/Mathlib/Probability/Martingale/Convergence.lean @@ -314,9 +314,6 @@ theorem Submartingale.tendsto_eLpNorm_one_limitProcess (hf : Submartingale f ℱ (memℒp_limitProcess_of_eLpNorm_bdd hmeas hR) hunif.2.1 (tendstoInMeasure_of_tendsto_ae hmeas <| hf.ae_tendsto_limitProcess hR) -@[deprecated (since := "2024-07-27")] -alias Submartingale.tendsto_snorm_one_limitProcess := Submartingale.tendsto_eLpNorm_one_limitProcess - theorem Submartingale.ae_tendsto_limitProcess_of_uniformIntegrable (hf : Submartingale f ℱ μ) (hunif : UniformIntegrable f 1 μ) : ∀ᵐ ω ∂μ, Tendsto (fun n => f n ω) atTop (𝓝 (ℱ.limitProcess f μ ω)) := @@ -343,12 +340,6 @@ theorem Martingale.eq_condExp_of_tendsto_eLpNorm {μ : Measure Ω} (hf : Marting @[deprecated (since := "2025-01-21")] alias Martingale.eq_condexp_of_tendsto_eLpNorm := Martingale.eq_condExp_of_tendsto_eLpNorm -@[deprecated (since := "2024-07-27")] -alias Martingale.eq_condExp_of_tendsto_snorm := Martingale.eq_condExp_of_tendsto_eLpNorm - -@[deprecated (since := "2025-01-21")] -alias Martingale.eq_condexp_of_tendsto_snorm := Martingale.eq_condExp_of_tendsto_snorm - /-- Part b of the **L¹ martingale convergence theorem**: if `f` is a uniformly integrable martingale adapted to the filtration `ℱ`, then for all `n`, `f n` is almost everywhere equal to the conditional expectation of its limiting process wrt. `ℱ n`. -/ @@ -437,12 +428,6 @@ theorem Integrable.tendsto_eLpNorm_condExp (hg : Integrable g μ) @[deprecated (since := "2025-01-21")] alias Integrable.tendsto_eLpNorm_condexp := Integrable.tendsto_eLpNorm_condExp -@[deprecated (since := "2024-07-27")] -alias Integrable.tendsto_snorm_condExp := Integrable.tendsto_eLpNorm_condExp - -@[deprecated (since := "2025-01-21")] -alias Integrable.tendsto_snorm_condexp := Integrable.tendsto_snorm_condExp - /-- **Lévy's upward theorem**, almost everywhere version: given a function `g` and a filtration `ℱ`, the sequence defined by `𝔼[g | ℱ n]` converges almost everywhere to `𝔼[g | ⨆ n, ℱ n]`. -/ theorem tendsto_ae_condExp (g : Ω → ℝ) : @@ -471,11 +456,6 @@ theorem tendsto_eLpNorm_condExp (g : Ω → ℝ) : @[deprecated (since := "2025-01-21")] alias tendsto_eLpNorm_condexp := tendsto_eLpNorm_condExp -@[deprecated (since := "2024-07-27")] -alias tendsto_snorm_condExp := tendsto_eLpNorm_condExp - -@[deprecated (since := "2025-01-21")] alias tendsto_snorm_condexp := tendsto_snorm_condExp - end L1Convergence end MeasureTheory diff --git a/Mathlib/Probability/Process/Filtration.lean b/Mathlib/Probability/Process/Filtration.lean index 1c4556148126c..75a03f87e3c8a 100644 --- a/Mathlib/Probability/Process/Filtration.lean +++ b/Mathlib/Probability/Process/Filtration.lean @@ -331,9 +331,6 @@ theorem memℒp_limitProcess_of_eLpNorm_bdd {R : ℝ≥0} {p : ℝ≥0∞} {F : exact sSup_le fun b ⟨a, ha⟩ => (ha a le_rfl).trans (hbdd _) · exact Memℒp.zero -@[deprecated (since := "2024-07-27")] -alias memℒp_limitProcess_of_snorm_bdd := memℒp_limitProcess_of_eLpNorm_bdd - end Limit variable {α : Type*} diff --git a/Mathlib/RepresentationTheory/FDRep.lean b/Mathlib/RepresentationTheory/FDRep.lean index f0a43daf7ef44..70ee74d5a86d4 100644 --- a/Mathlib/RepresentationTheory/FDRep.lean +++ b/Mathlib/RepresentationTheory/FDRep.lean @@ -48,9 +48,6 @@ open CategoryTheory.Limits abbrev FDRep (k G : Type u) [Field k] [Monoid G] := Action (FGModuleCat.{u} k) (MonCat.of G) -@[deprecated (since := "2024-07-05")] -alias FdRep := FDRep - namespace FDRep variable {k G : Type u} [Field k] [Monoid G] diff --git a/Mathlib/RingTheory/Binomial.lean b/Mathlib/RingTheory/Binomial.lean index 5fa7bcb787490..ed46beb83286a 100644 --- a/Mathlib/RingTheory/Binomial.lean +++ b/Mathlib/RingTheory/Binomial.lean @@ -379,9 +379,6 @@ theorem choose_natCast [NatPowAssoc R] (n k : ℕ) : choose (n : R) k = Nat.choo ← descPochhammer_eq_factorial_smul_choose, nsmul_eq_mul, ← Nat.cast_mul, ← Nat.descFactorial_eq_factorial_mul_choose, ← descPochhammer_smeval_eq_descFactorial] -@[deprecated (since := "2024-04-17")] -alias choose_nat_cast := choose_natCast - @[simp] theorem choose_zero_right' (r : R) : choose r 0 = (r + 1) ^ 0 := by dsimp only [choose] diff --git a/Mathlib/RingTheory/Congruence/Defs.lean b/Mathlib/RingTheory/Congruence/Defs.lean index 79d50bcfd7bb2..ac99fedc86a19 100644 --- a/Mathlib/RingTheory/Congruence/Defs.lean +++ b/Mathlib/RingTheory/Congruence/Defs.lean @@ -287,9 +287,6 @@ instance : NatCast c.Quotient := theorem coe_natCast (n : ℕ) : (↑(n : R) : c.Quotient) = n := rfl -@[deprecated (since := "2024-04-17")] -alias coe_nat_cast := coe_natCast - end NatCast section IntCast @@ -303,9 +300,6 @@ instance : IntCast c.Quotient := theorem coe_intCast (n : ℕ) : (↑(n : R) : c.Quotient) = n := rfl -@[deprecated (since := "2024-04-17")] -alias coe_int_cast := coe_intCast - end IntCast instance [Inhabited R] [Add R] [Mul R] (c : RingCon R) : Inhabited c.Quotient := diff --git a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean index d7d5139cc0b3e..6da94a280d02f 100644 --- a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean +++ b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean @@ -99,24 +99,15 @@ theorem intValuation_ne_zero (x : R) (hx : x ≠ 0) : v.intValuationDef x ≠ 0 rw [intValuationDef, if_neg hx] exact WithZero.coe_ne_zero -@[deprecated (since := "2024-07-09")] -alias int_valuation_ne_zero := intValuation_ne_zero - /-- Nonzero divisors have nonzero valuation. -/ theorem intValuation_ne_zero' (x : nonZeroDivisors R) : v.intValuationDef x ≠ 0 := v.intValuation_ne_zero x (nonZeroDivisors.coe_ne_zero x) -@[deprecated (since := "2024-07-09")] -alias int_valuation_ne_zero' := intValuation_ne_zero' - /-- Nonzero divisors have valuation greater than zero. -/ theorem intValuation_zero_le (x : nonZeroDivisors R) : 0 < v.intValuationDef x := by rw [v.intValuationDef_if_neg (nonZeroDivisors.coe_ne_zero x)] exact WithZero.zero_lt_coe _ -@[deprecated (since := "2024-07-09")] -alias int_valuation_zero_le := intValuation_zero_le - /-- The `v`-adic valuation on `R` is bounded above by 1. -/ theorem intValuation_le_one (x : R) : v.intValuationDef x ≤ 1 := by rw [intValuationDef] @@ -126,9 +117,6 @@ theorem intValuation_le_one (x : R) : v.intValuationDef x ≤ 1 := by Right.neg_nonpos_iff] exact Int.natCast_nonneg _ -@[deprecated (since := "2024-07-09")] -alias int_valuation_le_one := intValuation_le_one - /-- The `v`-adic valuation of `r ∈ R` is less than 1 if and only if `v` divides the ideal `(r)`. -/ theorem intValuation_lt_one_iff_dvd (r : R) : v.intValuationDef r < 1 ↔ v.asIdeal ∣ Ideal.span {r} := by @@ -143,9 +131,6 @@ theorem intValuation_lt_one_iff_dvd (r : R) : exact hr apply Associates.count_ne_zero_iff_dvd h (by apply v.irreducible) -@[deprecated (since := "2024-07-09")] -alias int_valuation_lt_one_iff_dvd := intValuation_lt_one_iff_dvd - /-- The `v`-adic valuation of `r ∈ R` is less than `Multiplicative.ofAdd (-n)` if and only if `vⁿ` divides the ideal `(r)`. -/ theorem intValuation_le_pow_iff_dvd (r : R) (n : ℕ) : @@ -159,16 +144,10 @@ theorem intValuation_le_pow_iff_dvd (r : R) (n : ℕ) : Associates.prime_pow_dvd_iff_le (Associates.mk_ne_zero'.mpr hr) (by apply v.associates_irreducible)] -@[deprecated (since := "2024-07-09")] -alias int_valuation_le_pow_iff_dvd := intValuation_le_pow_iff_dvd - /-- The `v`-adic valuation of `0 : R` equals 0. -/ theorem intValuation.map_zero' : v.intValuationDef 0 = 0 := v.intValuationDef_if_pos (Eq.refl 0) -@[deprecated (since := "2024-07-09")] -alias IntValuation.map_zero' := intValuation.map_zero' - /-- The `v`-adic valuation of `1 : R` equals 1. -/ theorem intValuation.map_one' : v.intValuationDef 1 = 1 := by classical @@ -177,9 +156,6 @@ theorem intValuation.map_one' : v.intValuationDef 1 = 1 := by Associates.count_zero (by apply v.associates_irreducible), Int.ofNat_zero, neg_zero, ofAdd_zero, WithZero.coe_one] -@[deprecated (since := "2024-07-09")] -alias IntValuation.map_one' := intValuation.map_one' - /-- The `v`-adic valuation of a product equals the product of the valuations. -/ theorem intValuation.map_mul' (x y : R) : v.intValuationDef (x * y) = v.intValuationDef x * v.intValuationDef y := by @@ -195,9 +171,6 @@ theorem intValuation.map_mul' (x y : R) : (by apply Associates.mk_ne_zero'.mpr hy) (by apply v.associates_irreducible)] rfl -@[deprecated (since := "2024-07-09")] -alias IntValuation.map_mul' := intValuation.map_mul' - theorem intValuation.le_max_iff_min_le {a b c : ℕ} : Multiplicative.ofAdd (-c : ℤ) ≤ max (Multiplicative.ofAdd (-a : ℤ)) (Multiplicative.ofAdd (-b : ℤ)) ↔ @@ -205,9 +178,6 @@ theorem intValuation.le_max_iff_min_le {a b c : ℕ} : rw [le_max_iff, ofAdd_le, ofAdd_le, neg_le_neg_iff, neg_le_neg_iff, Int.ofNat_le, Int.ofNat_le, ← min_le_iff] -@[deprecated (since := "2024-07-09")] -alias IntValuation.le_max_iff_min_le := intValuation.le_max_iff_min_le - /-- The `v`-adic valuation of a sum is bounded above by the maximum of the valuations. -/ theorem intValuation.map_add_le_max' (x y : R) : v.intValuationDef (x + y) ≤ max (v.intValuationDef x) (v.intValuationDef y) := by @@ -244,9 +214,6 @@ theorem intValuation.map_add_le_max' (x y : R) : · exact h_dvd_xy apply v.associates_irreducible -@[deprecated (since := "2024-07-09")] -alias IntValuation.map_add_le_max' := intValuation.map_add_le_max' - /-- The `v`-adic valuation on `R`. -/ @[simps] def intValuation : Valuation R ℤₘ₀ where @@ -285,9 +252,6 @@ theorem intValuation_exists_uniformizer : rw [Associates.mk_pow, Associates.prime_pow_dvd_iff_le hπ hv, not_le] at nmem exact Nat.eq_of_le_of_lt_succ mem nmem -@[deprecated (since := "2024-07-09")] -alias int_valuation_exists_uniformizer := intValuation_exists_uniformizer - /-- The `I`-adic valuation of a generator of `I` equals `(-1 : ℤₘ₀)` -/ theorem intValuation_singleton {r : R} (hr : r ≠ 0) (hv : v.asIdeal = Ideal.span {r}) : v.intValuation r = Multiplicative.ofAdd (-1 : ℤ) := by diff --git a/Mathlib/RingTheory/Derivation/Basic.lean b/Mathlib/RingTheory/Derivation/Basic.lean index f530128d4efbe..fb8af0111f43f 100644 --- a/Mathlib/RingTheory/Derivation/Basic.lean +++ b/Mathlib/RingTheory/Derivation/Basic.lean @@ -439,9 +439,6 @@ protected theorem map_sub : D (a - b) = D a - D b := theorem map_intCast (n : ℤ) : D (n : A) = 0 := by rw [← zsmul_one, D.map_smul_of_tower n, map_one_eq_zero, smul_zero] -@[deprecated (since := "2024-04-05")] alias map_coe_nat := map_natCast -@[deprecated (since := "2024-04-05")] alias map_coe_int := map_intCast - theorem leibniz_of_mul_eq_one {a b : A} (h : a * b = 1) : D a = -a ^ 2 • D b := by rw [neg_smul] refine eq_neg_of_add_eq_zero_left ?_ diff --git a/Mathlib/RingTheory/FiniteType.lean b/Mathlib/RingTheory/FiniteType.lean index c8b20cf8bab7c..21fa6f29ac59e 100644 --- a/Mathlib/RingTheory/FiniteType.lean +++ b/Mathlib/RingTheory/FiniteType.lean @@ -732,20 +732,3 @@ instance (priority := 100) CommRing.orzechProperty | smul a x _ hx => exact smul_mem _ a hx end Orzech - -section Vasconcelos - -/-- A theorem by Vasconcelos, given a finite module `M` over a commutative ring, any -surjective endomorphism of `M` is also injective. -It is a consequence of the fact `CommRing.orzechProperty` -that any commutative ring `R` satisfies the `OrzechProperty`; -please use `OrzechProperty.injective_of_surjective_endomorphism` instead. -This is similar to `IsNoetherian.injective_of_surjective_endomorphism` but only applies in the -commutative case, but does not use a Noetherian hypothesis. -/ -@[deprecated OrzechProperty.injective_of_surjective_endomorphism (since := "2024-05-30")] -theorem Module.Finite.injective_of_surjective_endomorphism {R : Type*} [CommRing R] {M : Type*} - [AddCommGroup M] [Module R M] [Module.Finite R M] (f : M →ₗ[R] M) - (f_surj : Function.Surjective f) : Function.Injective f := - OrzechProperty.injective_of_surjective_endomorphism f f_surj - -end Vasconcelos diff --git a/Mathlib/RingTheory/FractionalIdeal/Basic.lean b/Mathlib/RingTheory/FractionalIdeal/Basic.lean index e151dbde1d83b..17d1057a60b9b 100644 --- a/Mathlib/RingTheory/FractionalIdeal/Basic.lean +++ b/Mathlib/RingTheory/FractionalIdeal/Basic.lean @@ -577,9 +577,6 @@ theorem coe_natCast (n : ℕ) : ((n : FractionalIdeal S P) : Submodule R P) = n show ((n.unaryCast : FractionalIdeal S P) : Submodule R P) = n by induction n <;> simp [*, Nat.unaryCast] -@[deprecated (since := "2024-04-17")] -alias coe_nat_cast := coe_natCast - instance commSemiring : CommSemiring (FractionalIdeal S P) := Function.Injective.commSemiring _ Subtype.coe_injective coe_zero coe_one coe_add coe_mul (fun _ _ => coe_nsmul _ _) coe_pow coe_natCast diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 4d41090176a74..24405d05da768 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -643,9 +643,6 @@ theorem _root_.IsCoprime.sup_eq (h : IsCoprime I J) : I ⊔ J = ⊤ := isCoprime theorem inf_eq_mul_of_isCoprime (coprime : IsCoprime I J) : I ⊓ J = I * J := (Ideal.mul_eq_inf_of_coprime coprime.sup_eq).symm -@[deprecated (since := "2024-05-28")] -alias inf_eq_mul_of_coprime := inf_eq_mul_of_isCoprime - theorem isCoprime_span_singleton_iff (x y : R) : IsCoprime (span <| singleton x) (span <| singleton y) ↔ IsCoprime x y := by simp_rw [isCoprime_iff_codisjoint, codisjoint_iff, eq_top_iff_one, mem_span_singleton_sup, diff --git a/Mathlib/RingTheory/Int/Basic.lean b/Mathlib/RingTheory/Int/Basic.lean index 6222efbb76ead..313dae1b37c2b 100644 --- a/Mathlib/RingTheory/Int/Basic.lean +++ b/Mathlib/RingTheory/Int/Basic.lean @@ -127,20 +127,12 @@ theorem eq_pow_of_mul_eq_pow_odd_left {a b c : ℤ} (hab : IsCoprime a b) {k : rw [associated_iff_natAbs, natAbs_eq_natAbs_iff, ← hk.neg_pow] at hd obtain rfl | rfl := hd <;> exact ⟨_, rfl⟩ -@[deprecated (since := "2024-07-12")] -alias eq_pow_of_mul_eq_pow_bit1_left := eq_pow_of_mul_eq_pow_odd_left - theorem eq_pow_of_mul_eq_pow_odd_right {a b c : ℤ} (hab : IsCoprime a b) {k : ℕ} (hk : Odd k) (h : a * b = c ^ k) : ∃ d, b = d ^ k := eq_pow_of_mul_eq_pow_odd_left (c := c) hab.symm hk (by rwa [mul_comm] at h) -@[deprecated (since := "2024-07-12")] -alias eq_pow_of_mul_eq_pow_bit1_right := eq_pow_of_mul_eq_pow_odd_right - theorem eq_pow_of_mul_eq_pow_odd {a b c : ℤ} (hab : IsCoprime a b) {k : ℕ} (hk : Odd k) (h : a * b = c ^ k) : (∃ d, a = d ^ k) ∧ ∃ e, b = e ^ k := ⟨eq_pow_of_mul_eq_pow_odd_left hab hk h, eq_pow_of_mul_eq_pow_odd_right hab hk h⟩ -@[deprecated (since := "2024-07-12")] alias eq_pow_of_mul_eq_pow_bit1 := eq_pow_of_mul_eq_pow_odd - end Int diff --git a/Mathlib/RingTheory/IntegralDomain.lean b/Mathlib/RingTheory/IntegralDomain.lean index a44906bcb806f..d88c930f3e8ee 100644 --- a/Mathlib/RingTheory/IntegralDomain.lean +++ b/Mathlib/RingTheory/IntegralDomain.lean @@ -170,9 +170,6 @@ end EuclideanDivision variable [Fintype G] -@[deprecated (since := "2024-06-10")] -alias card_fiber_eq_of_mem_range := MonoidHom.card_fiber_eq_of_mem_range - /-- In an integral domain, a sum indexed by a nontrivial homomorphism from a finite group is zero. -/ theorem sum_hom_units_eq_zero (f : G →* R) (hf : f ≠ 1) : ∑ g : G, f g = 0 := by diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index d6ca20ff422c7..22264fcf09a34 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -502,10 +502,6 @@ theorem selfZPow_sub_natCast {n m : ℕ} : IsLocalization.mk'_eq_iff_eq] simp [Submonoid.pow_apply, ← pow_add, Nat.sub_add_cancel (le_of_not_le h)] -@[deprecated (since := "2024-04-05")] alias selfZPow_coe_nat := selfZPow_natCast -@[deprecated (since := "2024-04-05")] alias selfZPow_neg_coe_nat := selfZPow_neg_natCast -@[deprecated (since := "2024-04-05")] alias selfZPow_sub_cast_nat := selfZPow_sub_natCast - @[simp] theorem selfZPow_add {n m : ℤ} : selfZPow x B (n + m) = selfZPow x B n * selfZPow x B m := by rcases le_or_lt 0 n with hn | hn <;> rcases le_or_lt 0 m with hm | hm diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index 8b0b1e1417be3..841319542fe18 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -356,9 +356,6 @@ open IsLocalization theorem mk_natCast (m : ℕ) : (mk m 1 : Localization M) = m := by simpa using mk_algebraMap (R := R) (A := ℕ) _ -@[deprecated (since := "2024-04-17")] -alias mk_nat_cast := mk_natCast - variable [IsLocalization M S] section @@ -412,9 +409,6 @@ namespace Localization theorem mk_intCast (m : ℤ) : (mk m 1 : Localization M) = m := by simpa using mk_algebraMap (R := R) (A := ℤ) _ -@[deprecated (since := "2024-04-17")] -alias mk_int_cast := mk_intCast - end Localization open IsLocalization diff --git a/Mathlib/RingTheory/Localization/NumDen.lean b/Mathlib/RingTheory/Localization/NumDen.lean index 1df3ce301fd17..4b94d345d32e9 100644 --- a/Mathlib/RingTheory/Localization/NumDen.lean +++ b/Mathlib/RingTheory/Localization/NumDen.lean @@ -122,10 +122,6 @@ theorem isUnit_den_iff (x : K) : IsUnit (den A x : A) ↔ IsLocalization.IsInteg theorem isUnit_den_zero : IsUnit (den A (0 : K) : A) := by simp [isUnit_den_iff, IsLocalization.isInteger_zero] -@[deprecated isUnit_den_zero (since := "2024-07-11")] -theorem isUnit_den_of_num_eq_zero {x : K} (h : num A x = 0) : IsUnit (den A x : A) := - eq_zero_of_num_eq_zero h ▸ isUnit_den_zero - lemma associated_den_num_inv (x : K) (hx : x ≠ 0) : Associated (den A x : A) (num A x⁻¹) := associated_of_dvd_dvd (IsRelPrime.dvd_of_dvd_mul_right (IsFractionRing.num_den_reduced A x).symm <| diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index 56b13008a6c47..679b325f13247 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -220,8 +220,6 @@ theorem Int.natCast_emultiplicity (a b : ℕ) : theorem Int.natCast_multiplicity (a b : ℕ) : multiplicity (a : ℤ) (b : ℤ) = multiplicity a b := multiplicity_eq_of_emultiplicity_eq (natCast_emultiplicity a b) -@[deprecated (since := "2024-04-05")] alias Int.coe_nat_multiplicity := Int.natCast_multiplicity - theorem FiniteMultiplicity.not_iff_forall : ¬FiniteMultiplicity a b ↔ ∀ n : ℕ, a ^ n ∣ b := ⟨fun h n => Nat.casesOn n diff --git a/Mathlib/RingTheory/Noetherian/Basic.lean b/Mathlib/RingTheory/Noetherian/Basic.lean index be7045b5493a5..02aaf2a20128e 100644 --- a/Mathlib/RingTheory/Noetherian/Basic.lean +++ b/Mathlib/RingTheory/Noetherian/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Mario Carneiro, Kevin Buzzard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kevin Buzzard -/ +import Mathlib.Algebra.Order.PartialSups import Mathlib.LinearAlgebra.Quotient.Basic import Mathlib.RingTheory.Noetherian.Defs import Mathlib.RingTheory.Finiteness.Cardinality @@ -82,9 +83,6 @@ instance isNoetherian_quotient {A M : Type*} [Ring A] [AddCommGroup M] [SMul R A isNoetherian_of_surjective M ((Submodule.mkQ N).restrictScalars R) <| LinearMap.range_eq_top.mpr N.mkQ_surjective -@[deprecated (since := "2024-04-27"), nolint defLemma] -alias Submodule.Quotient.isNoetherian := isNoetherian_quotient - theorem isNoetherian_of_linearEquiv (f : M ≃ₗ[R] P) [IsNoetherian R M] : IsNoetherian R P := isNoetherian_of_surjective _ f.toLinearMap f.range diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 27ce39be124cc..7d01da51326bf 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -259,33 +259,23 @@ def coeffs (p : R[X]) : Finset R := letI := Classical.decEq R Finset.image (fun n => p.coeff n) p.support -@[deprecated (since := "2024-05-17")] noncomputable alias frange := coeffs - @[simp] theorem coeffs_zero : coeffs (0 : R[X]) = ∅ := rfl -@[deprecated (since := "2024-05-17")] alias frange_zero := coeffs_zero - theorem mem_coeffs_iff {p : R[X]} {c : R} : c ∈ p.coeffs ↔ ∃ n ∈ p.support, c = p.coeff n := by simp [coeffs, eq_comm, (Finset.mem_image)] -@[deprecated (since := "2024-05-17")] alias mem_frange_iff := mem_coeffs_iff - theorem coeffs_one : coeffs (1 : R[X]) ⊆ {1} := by classical simp_rw [coeffs, Finset.image_subset_iff] simp_all [coeff_one] -@[deprecated (since := "2024-05-17")] alias frange_one := coeffs_one - theorem coeff_mem_coeffs (p : R[X]) (n : ℕ) (h : p.coeff n ≠ 0) : p.coeff n ∈ p.coeffs := by classical simp only [coeffs, exists_prop, mem_support_iff, Finset.mem_image, Ne] exact ⟨n, h, rfl⟩ -@[deprecated (since := "2024-05-17")] alias coeff_mem_frange := coeff_mem_coeffs - theorem coeffs_monomial (n : ℕ) {c : R} (hc : c ≠ 0) : (monomial n c).coeffs = {c} := by rw [coeffs, support_monomial n hc] simp @@ -491,8 +481,6 @@ theorem coeffs_ofSubring {p : T[X]} : (↑(p.ofSubring T).coeffs : Set R) ⊆ T rw [← h'n, coeff_ofSubring] exact Subtype.mem (coeff p n : T) -@[deprecated (since := "2024-05-17")] alias frange_ofSubring := coeffs_ofSubring - end Ring end Polynomial diff --git a/Mathlib/RingTheory/Spectrum/Prime/Defs.lean b/Mathlib/RingTheory/Spectrum/Prime/Defs.lean index 93e0da1ca6ac5..d2f24c116cc32 100644 --- a/Mathlib/RingTheory/Spectrum/Prime/Defs.lean +++ b/Mathlib/RingTheory/Spectrum/Prime/Defs.lean @@ -31,6 +31,4 @@ structure PrimeSpectrum (R : Type*) [CommSemiring R] where asIdeal : Ideal R isPrime : asIdeal.IsPrime -@[deprecated (since := "2024-06-22")] alias PrimeSpectrum.IsPrime := PrimeSpectrum.isPrime - attribute [instance] PrimeSpectrum.isPrime diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean index d2a4e0e814f99..9611b9809525a 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean @@ -247,8 +247,6 @@ theorem factors_eq_top_iff_zero {a : Associates α} : a.factors = ⊤ ↔ a = 0 nontriviality α exact ⟨fun h ↦ by rwa [← factors_prod a, FactorSet.prod_eq_zero_iff], fun h ↦ h ▸ factors_zero⟩ -@[deprecated (since := "2024-04-16")] alias factors_eq_none_iff_zero := factors_eq_top_iff_zero - theorem factors_eq_some_iff_ne_zero {a : Associates α} : (∃ s : Multiset { p : Associates α // Irreducible p }, a.factors = s) ↔ a ≠ 0 := by simp_rw [@eq_comm _ a.factors, ← WithTop.ne_top_iff_exists] diff --git a/Mathlib/RingTheory/WittVector/Basic.lean b/Mathlib/RingTheory/WittVector/Basic.lean index 9258b5955fdb9..ea0018b2a39dc 100644 --- a/Mathlib/RingTheory/WittVector/Basic.lean +++ b/Mathlib/RingTheory/WittVector/Basic.lean @@ -117,16 +117,10 @@ theorem natCast (n : ℕ) : mapFun f (n : 𝕎 R) = n := show mapFun f n.unaryCast = (n : WittVector p S) by induction n <;> simp [*, Nat.unaryCast, add, one, zero] <;> rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast := natCast - theorem intCast (n : ℤ) : mapFun f (n : 𝕎 R) = n := show mapFun f n.castDef = (n : WittVector p S) by cases n <;> simp [*, Int.castDef, add, one, neg, zero, natCast] <;> rfl -@[deprecated (since := "2024-04-17")] -alias int_cast := intCast - end mapFun end WittVector @@ -183,9 +177,6 @@ private theorem ghostFun_natCast (i : ℕ) : ghostFun (i : 𝕎 R) = i := induction i <;> simp [*, Nat.unaryCast, ghostFun_zero, ghostFun_one, ghostFun_add, -Pi.natCast_def] -@[deprecated (since := "2024-04-17")] -alias ghostFun_nat_cast := ghostFun_natCast - private theorem ghostFun_sub : ghostFun (x - y) = ghostFun x - ghostFun y := by ghost_fun_tac X 0 - X 1, ![x.coeff, y.coeff] @@ -199,9 +190,6 @@ private theorem ghostFun_intCast (i : ℤ) : ghostFun (i : 𝕎 R) = i := cases i <;> simp [*, Int.castDef, ghostFun_natCast, ghostFun_neg, -Pi.natCast_def, -Pi.intCast_def] -@[deprecated (since := "2024-04-17")] -alias ghostFun_int_cast := ghostFun_intCast - private lemma ghostFun_nsmul (m : ℕ) (x : WittVector p R) : ghostFun (m • x) = m • ghostFun x := by ghost_fun_tac m • (X 0), ![x.coeff] diff --git a/Mathlib/RingTheory/WittVector/Truncated.lean b/Mathlib/RingTheory/WittVector/Truncated.lean index cb39757374dc6..f7766d24fda63 100644 --- a/Mathlib/RingTheory/WittVector/Truncated.lean +++ b/Mathlib/RingTheory/WittVector/Truncated.lean @@ -254,14 +254,8 @@ theorem truncateFun_pow (x : 𝕎 R) (m : ℕ) : truncateFun n (x ^ m) = truncat theorem truncateFun_natCast (m : ℕ) : truncateFun n (m : 𝕎 R) = m := rfl -@[deprecated (since := "2024-04-17")] -alias truncateFun_nat_cast := truncateFun_natCast - theorem truncateFun_intCast (m : ℤ) : truncateFun n (m : 𝕎 R) = m := rfl -@[deprecated (since := "2024-04-17")] -alias truncateFun_int_cast := truncateFun_intCast - end WittVector namespace TruncatedWittVector diff --git a/Mathlib/SetTheory/Cardinal/UnivLE.lean b/Mathlib/SetTheory/Cardinal/UnivLE.lean index dbc592b08c9b4..fbb30020977b4 100644 --- a/Mathlib/SetTheory/Cardinal/UnivLE.lean +++ b/Mathlib/SetTheory/Cardinal/UnivLE.lean @@ -34,6 +34,6 @@ theorem Ordinal.univLE_of_injective {f : Ordinal.{u} → Ordinal.{v}} (h : f.Inj UnivLE.{u, v} := univLE_iff_exists_embedding.2 ⟨f, h⟩ -/-- Together with transitivity, this shows UnivLE "IsTotalPreorder". -/ +/-- Together with transitivity, this shows `UnivLE` is a total preorder. -/ theorem univLE_total : UnivLE.{u, v} ∨ UnivLE.{v, u} := by simp_rw [univLE_iff_cardinal_le]; apply le_total diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index c2e154cef7e9e..01ef9f18e6ec8 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -2244,9 +2244,6 @@ theorem one_add_natCast (m : ℕ) : 1 + (m : Ordinal) = succ m := by rw [← Nat.cast_one, ← Nat.cast_add, add_comm] rfl -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias one_add_nat_cast := one_add_natCast - @[simp] theorem one_add_ofNat (m : ℕ) [m.AtLeastTwo] : 1 + (ofNat(m) : Ordinal) = Order.succ (OfNat.ofNat m : Ordinal) := @@ -2257,45 +2254,24 @@ theorem natCast_mul (m : ℕ) : ∀ n : ℕ, ((m * n : ℕ) : Ordinal) = m * n | 0 => by simp | n + 1 => by rw [Nat.mul_succ, Nat.cast_add, natCast_mul m n, Nat.cast_succ, mul_add_one] -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias nat_cast_mul := natCast_mul - @[deprecated Nat.cast_le (since := "2024-10-17")] theorem natCast_le {m n : ℕ} : (m : Ordinal) ≤ n ↔ m ≤ n := Nat.cast_le -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias nat_cast_le := natCast_le - @[deprecated Nat.cast_inj (since := "2024-10-17")] theorem natCast_inj {m n : ℕ} : (m : Ordinal) = n ↔ m = n := Nat.cast_inj -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias nat_cast_inj := natCast_inj - @[deprecated Nat.cast_lt (since := "2024-10-17")] theorem natCast_lt {m n : ℕ} : (m : Ordinal) < n ↔ m < n := Nat.cast_lt -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias nat_cast_lt := natCast_lt - @[deprecated Nat.cast_eq_zero (since := "2024-10-17")] theorem natCast_eq_zero {n : ℕ} : (n : Ordinal) = 0 ↔ n = 0 := Nat.cast_eq_zero -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias nat_cast_eq_zero := natCast_eq_zero - @[deprecated Nat.cast_ne_zero (since := "2024-10-17")] theorem natCast_ne_zero {n : ℕ} : (n : Ordinal) ≠ 0 ↔ n ≠ 0 := Nat.cast_ne_zero -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias nat_cast_ne_zero := natCast_ne_zero - @[deprecated Nat.cast_pos' (since := "2024-10-17")] theorem natCast_pos {n : ℕ} : (0 : Ordinal) < n ↔ 0 < n := Nat.cast_pos' -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias nat_cast_pos := natCast_pos - @[simp, norm_cast] theorem natCast_sub (m n : ℕ) : ((m - n : ℕ) : Ordinal) = m - n := by rcases le_total m n with h | h @@ -2303,9 +2279,6 @@ theorem natCast_sub (m n : ℕ) : ((m - n : ℕ) : Ordinal) = m - n := by · rw [← add_left_cancel_iff (a := ↑n), ← Nat.cast_add, add_tsub_cancel_of_le h, Ordinal.add_sub_cancel_of_le (Nat.cast_le.2 h)] -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias nat_cast_sub := natCast_sub - @[simp, norm_cast] theorem natCast_div (m n : ℕ) : ((m / n : ℕ) : Ordinal) = m / n := by rcases eq_or_ne n 0 with (rfl | hn) @@ -2318,25 +2291,16 @@ theorem natCast_div (m n : ℕ) : ((m / n : ℕ) : Ordinal) = m / n := by ← Nat.div_lt_iff_lt_mul (Nat.pos_of_ne_zero hn)] apply Nat.lt_succ_self -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias nat_cast_div := natCast_div - @[simp, norm_cast] theorem natCast_mod (m n : ℕ) : ((m % n : ℕ) : Ordinal) = m % n := by rw [← add_left_cancel_iff, div_add_mod, ← natCast_div, ← natCast_mul, ← Nat.cast_add, Nat.div_add_mod] -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias nat_cast_mod := natCast_mod - @[simp] theorem lift_natCast : ∀ n : ℕ, lift.{u, v} n = n | 0 => by simp | n + 1 => by simp [lift_natCast n] -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias lift_nat_cast := lift_natCast - @[simp] theorem lift_ofNat (n : ℕ) [n.AtLeastTwo] : lift.{u, v} ofNat(n) = OfNat.ofNat n := @@ -2405,14 +2369,6 @@ alias omega_le := omega0_le theorem iSup_natCast : iSup Nat.cast = ω := (Ordinal.iSup_le fun n => (nat_lt_omega0 n).le).antisymm <| omega0_le.2 <| Ordinal.le_iSup _ -set_option linter.deprecated false in -@[deprecated iSup_natCast (since := "2024-04-17")] -theorem sup_natCast : sup Nat.cast = ω := - iSup_natCast - -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias sup_nat_cast := sup_natCast - theorem nat_lt_limit {o} (h : IsLimit o) : ∀ n : ℕ, ↑n < o | 0 => h.pos | n + 1 => h.succ_lt (nat_lt_limit h n) diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 710ae041e6110..d49bc20e3c033 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -966,9 +966,6 @@ theorem card_succ (o : Ordinal) : card (succ o) = card o + 1 := by theorem natCast_succ (n : ℕ) : ↑n.succ = succ (n : Ordinal) := rfl -@[deprecated "No deprecation message was provided." (since := "2024-04-17")] -alias nat_cast_succ := natCast_succ - instance uniqueIioOne : Unique (Iio (1 : Ordinal)) where default := ⟨0, zero_lt_one' Ordinal⟩ uniq a := Subtype.ext <| lt_one_iff_zero.1 a.2 diff --git a/Mathlib/Tactic/ComputeDegree.lean b/Mathlib/Tactic/ComputeDegree.lean index f82b8e0189765..fef8f75ae065b 100644 --- a/Mathlib/Tactic/ComputeDegree.lean +++ b/Mathlib/Tactic/ComputeDegree.lean @@ -95,9 +95,6 @@ theorem natDegree_natCast_le (n : ℕ) : natDegree (n : R[X]) ≤ 0 := (natDegre theorem natDegree_zero_le : natDegree (0 : R[X]) ≤ 0 := natDegree_zero.le theorem natDegree_one_le : natDegree (1 : R[X]) ≤ 0 := natDegree_one.le -@[deprecated (since := "2024-04-17")] -alias natDegree_nat_cast_le := natDegree_natCast_le - theorem coeff_add_of_eq {n : ℕ} {a b : R} {f g : R[X]} (h_add_left : f.coeff n = a) (h_add_right : g.coeff n = b) : (f + g).coeff n = a + b := by subst ‹_› ‹_›; apply coeff_add @@ -181,18 +178,12 @@ variable [Ring R] theorem natDegree_intCast_le (n : ℤ) : natDegree (n : R[X]) ≤ 0 := (natDegree_intCast _).le -@[deprecated (since := "2024-04-17")] -alias natDegree_int_cast_le := natDegree_intCast_le - theorem coeff_sub_of_eq {n : ℕ} {a b : R} {f g : R[X]} (hf : f.coeff n = a) (hg : g.coeff n = b) : (f - g).coeff n = a - b := by subst hf hg; apply coeff_sub theorem coeff_intCast_ite {n : ℕ} {a : ℤ} : (Int.cast a : R[X]).coeff n = ite (n = 0) a 0 := by simp only [← C_eq_intCast, coeff_C, Int.cast_ite, Int.cast_zero] -@[deprecated (since := "2024-04-17")] -alias coeff_int_cast_ite := coeff_intCast_ite - end ring end recursion_lemmas diff --git a/Mathlib/Tactic/Linarith/Lemmas.lean b/Mathlib/Tactic/Linarith/Lemmas.lean index 5802a9ec5a4da..5876ef2848c1d 100644 --- a/Mathlib/Tactic/Linarith/Lemmas.lean +++ b/Mathlib/Tactic/Linarith/Lemmas.lean @@ -100,9 +100,6 @@ lemma zero_mul_eq {α} {R : α → α → Prop} [Semiring α] {a b : α} (h : a lemma natCast_nonneg (α : Type u) [OrderedSemiring α] (n : ℕ) : (0 : α) ≤ n := Nat.cast_nonneg n -@[deprecated (since := "2024-04-17")] -alias nat_cast_nonneg := natCast_nonneg - end Linarith section diff --git a/Mathlib/Tactic/NormNum/Basic.lean b/Mathlib/Tactic/NormNum/Basic.lean index 91ce6bebb6980..bcdcf5f0c519c 100644 --- a/Mathlib/Tactic/NormNum/Basic.lean +++ b/Mathlib/Tactic/NormNum/Basic.lean @@ -104,9 +104,6 @@ theorem isNat_natAbs_neg : {n : ℤ} → {a : ℕ} → IsInt n (.negOfNat a) → theorem isNat_natCast {R} [AddMonoidWithOne R] (n m : ℕ) : IsNat n m → IsNat (n : R) m := by rintro ⟨⟨⟩⟩; exact ⟨rfl⟩ -@[deprecated (since := "2024-04-17")] -alias isNat_cast := isNat_natCast - /-- The `norm_num` extension which identifies an expression `Nat.cast n`, returning `n`. -/ @[norm_num Nat.cast _, NatCast.natCast _] def evalNatCast : NormNumExt where eval {u α} e := do let sα ← inferAddMonoidWithOne α @@ -119,15 +116,9 @@ alias isNat_cast := isNat_natCast theorem isNat_intCast {R} [Ring R] (n : ℤ) (m : ℕ) : IsNat n m → IsNat (n : R) m := by rintro ⟨⟨⟩⟩; exact ⟨by simp⟩ -@[deprecated (since := "2024-04-17")] -alias isNat_int_cast := isNat_intCast - theorem isintCast {R} [Ring R] (n m : ℤ) : IsInt n m → IsInt (n : R) m := by rintro ⟨⟨⟩⟩; exact ⟨rfl⟩ -@[deprecated (since := "2024-04-17")] -alias isInt_cast := isintCast - /-- The `norm_num` extension which identifies an expression `Int.cast n`, returning `n`. -/ @[norm_num Int.cast _, IntCast.intCast _] def evalIntCast : NormNumExt where eval {u α} e := do let rα ← inferRing α diff --git a/Mathlib/Tactic/Qify.lean b/Mathlib/Tactic/Qify.lean index ccba4938adc9a..886f1b4cc03af 100644 --- a/Mathlib/Tactic/Qify.lean +++ b/Mathlib/Tactic/Qify.lean @@ -70,9 +70,6 @@ macro_rules @[qify_simps] lemma intCast_ne (a b : ℤ) : a ≠ b ↔ (a : ℚ) ≠ (b : ℚ) := by simp only [ne_eq, Int.cast_inj] -@[deprecated (since := "2024-04-17")] -alias int_cast_ne := intCast_ne - end Qify end Mathlib.Tactic diff --git a/Mathlib/Tactic/Rify.lean b/Mathlib/Tactic/Rify.lean index 77801b38118d5..dda6b2772a13a 100644 --- a/Mathlib/Tactic/Rify.lean +++ b/Mathlib/Tactic/Rify.lean @@ -79,9 +79,6 @@ macro_rules @[rify_simps] lemma ratCast_lt (a b : ℚ) : a < b ↔ (a : ℝ) < (b : ℝ) := by simp @[rify_simps] lemma ratCast_ne (a b : ℚ) : a ≠ b ↔ (a : ℝ) ≠ (b : ℝ) := by simp -@[deprecated (since := "2024-04-17")] -alias rat_cast_ne := ratCast_ne - @[rify_simps] lemma ofNat_rat_real (a : ℕ) [a.AtLeastTwo] : ((ofNat(a) : ℚ) : ℝ) = (ofNat(a) : ℝ) := rfl diff --git a/Mathlib/Tactic/Zify.lean b/Mathlib/Tactic/Zify.lean index f270be4613f51..22a81b43bdfa3 100644 --- a/Mathlib/Tactic/Zify.lean +++ b/Mathlib/Tactic/Zify.lean @@ -102,9 +102,6 @@ def zifyProof (simpArgs : Option (Syntax.TSepArray `Lean.Parser.Tactic.simpStar -- TODO: is it worth adding lemmas for Prime and Coprime as well? -- Doing so in this file would require adding imports. -@[deprecated (since := "2024-04-17")] -alias nat_cast_dvd := natCast_dvd - -- `Nat.cast_sub` is already tagged as `norm_cast` but it does allow to use assumptions like -- `m < n` or more generally `m + k ≤ n`. We add two lemmas to increase the probability that diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean index 138c67bdb9188..3cf9b6e0a42dd 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean @@ -225,7 +225,6 @@ instance : ContinuousEvalConst (ContinuousMultilinearMap 𝕜 E F) (Π i, E i) F (uniformContinuous_eval_const x).continuous @[deprecated (since := "2024-10-05")] protected alias continuous_eval_const := continuous_eval_const -@[deprecated (since := "2024-04-10")] alias continuous_eval_left := continuous_eval_const @[deprecated (since := "2024-10-05")] protected alias continuous_coe_fun := continuous_coeFun instance instT2Space [T2Space F] : T2Space (ContinuousMultilinearMap 𝕜 E F) := diff --git a/Mathlib/Topology/Algebra/Support.lean b/Mathlib/Topology/Algebra/Support.lean index 0b2ba5734b60b..e255b35efc416 100644 --- a/Mathlib/Topology/Algebra/Support.lean +++ b/Mathlib/Topology/Algebra/Support.lean @@ -340,9 +340,6 @@ theorem HasCompactSupport.smul_right (hf : HasCompactSupport f) : HasCompactSupp rw [hasCompactSupport_iff_eventuallyEq] at hf ⊢ exact hf.mono fun x hx => by simp_rw [Pi.smul_apply', hx, Pi.zero_apply, zero_smul] -@[deprecated (since := "2024-06-05")] -alias HasCompactSupport.smul_left' := HasCompactSupport.smul_left - end SMulWithZero section MulZeroClass diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 3482d2379b813..d0db78c7c2d29 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -584,11 +584,6 @@ theorem disjoint_cocompact_right (f : Filter X) : simp_rw [hasBasis_cocompact.disjoint_iff_right, compl_compl] tauto -@[deprecated "see `cocompact_eq_atTop` with `import Mathlib.Topology.Instances.Nat`" - (since := "2024-02-07")] -theorem _root_.Nat.cocompact_eq : cocompact ℕ = atTop := - (cocompact_eq_cofinite ℕ).trans Nat.cofinite_eq_atTop - theorem Tendsto.isCompact_insert_range_of_cocompact {f : X → Y} {y} (hf : Tendsto f (cocompact X) (𝓝 y)) (hfc : Continuous f) : IsCompact (insert y (range f)) := by intro l hne hle diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index 7d5a552739ded..370e2f20a2b73 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -92,16 +92,10 @@ instance [NatCast β] : NatCast C(α, β) := theorem coe_natCast [NatCast β] (n : ℕ) : ((n : C(α, β)) : α → β) = n := rfl -@[deprecated (since := "2024-04-17")] -alias coe_nat_cast := coe_natCast - @[simp] theorem natCast_apply [NatCast β] (n : ℕ) (x : α) : (n : C(α, β)) x = n := rfl -@[deprecated (since := "2024-04-17")] -alias nat_cast_apply := natCast_apply - /-! ### `Int.cast` -/ instance [IntCast β] : IntCast C(α, β) := @@ -111,16 +105,10 @@ instance [IntCast β] : IntCast C(α, β) := theorem coe_intCast [IntCast β] (n : ℤ) : ((n : C(α, β)) : α → β) = n := rfl -@[deprecated (since := "2024-04-17")] -alias coe_int_cast := coe_intCast - @[simp] theorem intCast_apply [IntCast β] (n : ℤ) (x : α) : (n : C(α, β)) x = n := rfl -@[deprecated (since := "2024-04-17")] -alias int_cast_apply := intCast_apply - /-! ### `nsmul` and `pow` -/ instance instNSMul [AddMonoid β] [ContinuousAdd β] : SMul ℕ C(α, β) := diff --git a/Mathlib/Topology/Defs/Sequences.lean b/Mathlib/Topology/Defs/Sequences.lean index cac021631b4a5..99489bd589a8f 100644 --- a/Mathlib/Topology/Defs/Sequences.lean +++ b/Mathlib/Topology/Defs/Sequences.lean @@ -77,8 +77,6 @@ class SeqCompactSpace : Prop where export SeqCompactSpace (isSeqCompact_univ) -@[deprecated (since := "2024-07-25")] alias seq_compact_univ := isSeqCompact_univ - /-- A topological space is called a *Fréchet-Urysohn space*, if the sequential closure of any set is equal to its closure. Since one of the inclusions is trivial, we require only the non-trivial one in the definition. -/ diff --git a/Mathlib/Topology/EMetricSpace/Lipschitz.lean b/Mathlib/Topology/EMetricSpace/Lipschitz.lean index 2a72f625ee41f..18d504702f609 100644 --- a/Mathlib/Topology/EMetricSpace/Lipschitz.lean +++ b/Mathlib/Topology/EMetricSpace/Lipschitz.lean @@ -84,8 +84,6 @@ lemma LocallyLipschitzOn.mono (hf : LocallyLipschitzOn t f) (h : s ⊆ t) : Loca @[simp] lemma lipschitzOnWith_univ : LipschitzOnWith K f univ ↔ LipschitzWith K f := by simp [LipschitzOnWith, LipschitzWith] -@[deprecated (since := "2024-07-17")] alias lipschitzOn_univ := lipschitzOnWith_univ - @[simp] lemma locallyLipschitzOn_univ : LocallyLipschitzOn univ f ↔ LocallyLipschitz f := by simp [LocallyLipschitzOn, LocallyLipschitz] diff --git a/Mathlib/Topology/ExtremallyDisconnected.lean b/Mathlib/Topology/ExtremallyDisconnected.lean index 78aa72828e83e..09146b4cd39bf 100644 --- a/Mathlib/Topology/ExtremallyDisconnected.lean +++ b/Mathlib/Topology/ExtremallyDisconnected.lean @@ -279,10 +279,6 @@ protected theorem CompactT2.projective_iff_extremallyDisconnected [CompactSpace Projective A ↔ ExtremallyDisconnected A := ⟨Projective.extremallyDisconnected, fun _ => ExtremallyDisconnected.projective⟩ -@[deprecated (since := "2024-05-26")] -alias CompactT2.projective_iff_extremallyDisconnnected := - CompactT2.projective_iff_extremallyDisconnected - end -- Note: It might be possible to use Gleason for this instead diff --git a/Mathlib/Topology/List.lean b/Mathlib/Topology/List.lean index 10a7e8b0991c0..b2b581ce1f833 100644 --- a/Mathlib/Topology/List.lean +++ b/Mathlib/Topology/List.lean @@ -153,13 +153,9 @@ theorem tendsto_eraseIdx : dsimp [eraseIdx] exact tendsto_fst.cons ((@tendsto_eraseIdx n l).comp tendsto_snd) -@[deprecated (since := "2024-05-04")] alias tendsto_removeNth := tendsto_eraseIdx - theorem continuous_eraseIdx {n : ℕ} : Continuous fun l : List α => eraseIdx l n := continuous_iff_continuousAt.mpr fun _a => tendsto_eraseIdx -@[deprecated (since := "2024-05-04")] alias continuous_removeNth := continuous_eraseIdx - @[to_additive] theorem tendsto_prod [Monoid α] [ContinuousMul α] {l : List α} : Tendsto List.prod (𝓝 l) (𝓝 l.prod) := by @@ -221,12 +217,8 @@ theorem continuousAt_eraseIdx {n : ℕ} {i : Fin (n + 1)} : simp only [Vector.eraseIdx_val] exact Tendsto.comp List.tendsto_eraseIdx continuousAt_subtype_val -@[deprecated (since := "2024-05-04")] alias continuousAt_removeNth := continuousAt_eraseIdx - theorem continuous_eraseIdx {n : ℕ} {i : Fin (n + 1)} : Continuous (List.Vector.eraseIdx i : List.Vector α (n + 1) → List.Vector α n) := continuous_iff_continuousAt.mpr fun ⟨_a, _l⟩ => continuousAt_eraseIdx -@[deprecated (since := "2024-05-04")] alias continuous_removeNth := continuous_eraseIdx - end Vector diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 3557289a1439f..d1dcfdbdec44d 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -82,10 +82,6 @@ theorem IsClosedMap.restrictPreimage (H : IsClosedMap f) (s : Set β) : simpa [isClosed_induced_iff] exact fun u hu e => ⟨f '' u, H u hu, by simp [← e, image_restrictPreimage]⟩ -@[deprecated "No deprecation message was provided." (since := "2024-04-02")] -theorem Set.restrictPreimage_isClosedMap (s : Set β) (H : IsClosedMap f) : - IsClosedMap (s.restrictPreimage f) := H.restrictPreimage s - theorem IsOpenMap.restrictPreimage (H : IsOpenMap f) (s : Set β) : IsOpenMap (s.restrictPreimage f) := by intro t @@ -94,10 +90,6 @@ theorem IsOpenMap.restrictPreimage (H : IsOpenMap f) (s : Set β) : simpa [isOpen_induced_iff] exact fun u hu e => ⟨f '' u, H u hu, by simp [← e, image_restrictPreimage]⟩ -@[deprecated "No deprecation message was provided." (since := "2024-04-02")] -theorem Set.restrictPreimage_isOpenMap (s : Set β) (H : IsOpenMap f) : - IsOpenMap (s.restrictPreimage f) := H.restrictPreimage s - variable (hU : iSup U = ⊤) include hU diff --git a/Mathlib/Topology/Maps/Proper/Basic.lean b/Mathlib/Topology/Maps/Proper/Basic.lean index f7cb0d696daac..6cc72fdf53225 100644 --- a/Mathlib/Topology/Maps/Proper/Basic.lean +++ b/Mathlib/Topology/Maps/Proper/Basic.lean @@ -297,8 +297,6 @@ alias isProperMap_restr_of_proper_of_closed := IsProperMap.restrict lemma IsProperMap.isClosed_range (hf : IsProperMap f) : IsClosed (range f) := hf.isClosedMap.isClosed_range -@[deprecated (since := "2024-05-08")] alias IsProperMap.closed_range := IsProperMap.isClosed_range - /-- Version of `isProperMap_iff_isClosedMap_and_compact_fibers` in terms of `cofinite` and `cocompact`. Only works when the codomain is `T1`. -/ lemma isProperMap_iff_isClosedMap_and_tendsto_cofinite [T1Space Y] : diff --git a/Mathlib/Topology/Order/Bounded.lean b/Mathlib/Topology/Order/Bounded.lean deleted file mode 100644 index 03593ab41f4ef..0000000000000 --- a/Mathlib/Topology/Order/Bounded.lean +++ /dev/null @@ -1,40 +0,0 @@ -/- -Copyright (c) 2023 Kalle Kytölä. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kalle Kytölä --/ -import Mathlib.Topology.Bornology.Basic -import Mathlib.Order.LiminfLimsup -import Mathlib.Topology.Instances.Real.Defs - -/-! -# Relating order and metric boundedness - -In spaces equipped with both an order and a metric, there are separate notions of boundedness -associated with each of the two structures. In specific cases such as ℝ, there are results which -relate the two notions. - -## Tags - -bounded, bornology, order, metric --/ - -open Set Filter - -section Real - -@[deprecated isBoundedUnder_of (since := "2024-06-07")] -lemma Filter.isBounded_le_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} - (h : Bornology.IsBounded (Set.range f)) : - (F.map f).IsBounded (· ≤ ·) := by - obtain ⟨c, hc⟩ := h.bddAbove - exact isBoundedUnder_of ⟨c, by simpa [mem_upperBounds] using hc⟩ - -@[deprecated isBoundedUnder_of (since := "2024-06-07")] -lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} - (h : Bornology.IsBounded (Set.range f)) : - (F.map f).IsBounded (· ≥ ·) := by - obtain ⟨c, hc⟩ := h.bddBelow - apply isBoundedUnder_of ⟨c, by simpa [mem_lowerBounds] using hc⟩ - -end Real diff --git a/Mathlib/Topology/Sheaves/Stalks.lean b/Mathlib/Topology/Sheaves/Stalks.lean index cfc4885a55d29..c977d29a84f28 100644 --- a/Mathlib/Topology/Sheaves/Stalks.lean +++ b/Mathlib/Topology/Sheaves/Stalks.lean @@ -351,8 +351,6 @@ theorem germ_stalkSpecializes (F : X.Presheaf C) F.germ U y hy ≫ F.stalkSpecializes h = F.germ U x (h.mem_open U.isOpen hy) := colimit.ι_desc _ _ -@[deprecated (since := "2024-07-30")] alias germ_stalkSpecializes' := germ_stalkSpecializes - @[simp] theorem stalkSpecializes_refl {C : Type*} [Category C] [Limits.HasColimits C] {X : TopCat} (F : X.Presheaf C) (x : X) : F.stalkSpecializes (specializes_refl x) = 𝟙 _ := by diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index f5abe4496c48a..ff1b5e147a906 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -1575,8 +1575,6 @@ instance Sum.instUniformSpace : UniformSpace (α ⊕ β) where cases x <;> simp [mem_comap', -mem_comap, nhds_inl, nhds_inr, nhds_eq_comap_uniformity, Prod.ext_iff] -@[reducible, deprecated (since := "2024-02-15")] alias Sum.uniformSpace := Sum.instUniformSpace - /-- The union of an entourage of the diagonal in each set of a disjoint union is again an entourage of the diagonal. -/ theorem union_mem_uniformity_sum {a : Set (α × α)} (ha : a ∈ 𝓤 α) {b : Set (β × β)} (hb : b ∈ 𝓤 β) : diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index 18f55b5f95253..0547bd463636b 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -480,9 +480,6 @@ theorem TotallyBounded.subset {s₁ s₂ : Set α} (hs : s₁ ⊆ s₂) (h : Tot let ⟨t, ht₁, ht₂⟩ := h d hd ⟨t, ht₁, Subset.trans hs ht₂⟩ -@[deprecated (since := "2024-06-01")] -alias totallyBounded_subset := TotallyBounded.subset - /-- The closure of a totally bounded set is totally bounded. -/ theorem TotallyBounded.closure {s : Set α} (h : TotallyBounded s) : TotallyBounded (closure s) := uniformity_hasBasis_closed.totallyBounded_iff.2 fun V hV => diff --git a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean index a4f69392a8c0c..b427d2dd13581 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean @@ -1066,16 +1066,6 @@ theorem isClosed_setOf_continuous [TopologicalSpace α] (h : RestrictGenTopology rw [← tendsto_id', UniformOnFun.tendsto_iff_tendstoUniformlyOn] at huf exact (huf s hs).continuousOn <| hu fun _ ↦ Continuous.continuousOn -/-- Suppose that the topology on `α` is defined by its restrictions to the sets of `𝔖`. - -Then the set of continuous functions is closed -in the topology of uniform convergence on the sets of `𝔖`. -/ -@[deprecated isClosed_setOf_continuous (since := "2024-06-29")] -theorem isClosed_setOf_continuous_of_le [t : TopologicalSpace α] - (h : t ≤ ⨆ s ∈ 𝔖, .coinduced (Subtype.val : s → α) inferInstance) : - IsClosed {f : α →ᵤ[𝔖] β | Continuous (toFun 𝔖 f)} := - isClosed_setOf_continuous ⟨fun u hu ↦ h _ <| by simpa only [isOpen_iSup_iff, isOpen_coinduced]⟩ - variable (𝔖) in theorem uniformSpace_eq_inf_precomp_of_cover {δ₁ δ₂ : Type*} (φ₁ : δ₁ → α) (φ₂ : δ₂ → α) (𝔗₁ : Set (Set δ₁)) (𝔗₂ : Set (Set δ₂)) From 19755f8f141cfea2483af7ddcd26d514611beb27 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Sat, 1 Feb 2025 11:21:25 +0000 Subject: [PATCH 6/8] feat(CategoryTheory): transfer being iso along an iso in the arrow category (#21310) This is probably the minmal set of necessary lemmas --- Mathlib/CategoryTheory/Comma/Arrow.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/Mathlib/CategoryTheory/Comma/Arrow.lean b/Mathlib/CategoryTheory/Comma/Arrow.lean index ea02bd9201175..1779eaa423ca1 100644 --- a/Mathlib/CategoryTheory/Comma/Arrow.lean +++ b/Mathlib/CategoryTheory/Comma/Arrow.lean @@ -176,6 +176,23 @@ instance isIso_right [IsIso sq] : IsIso sq.right where eq_self_iff_true, and_self_iff] simp +lemma isIso_of_isIso {X Y : T} {f : X ⟶ Y} {g : Arrow T} (sq : mk f ⟶ g) [IsIso sq] [IsIso f] : + IsIso g.hom := by + rw [iso_w' (asIso sq)] + infer_instance + +lemma isIso_hom_iff_isIso_hom_of_isIso {f g : Arrow T} (sq : f ⟶ g) [IsIso sq] : + IsIso f.hom ↔ IsIso g.hom := + ⟨fun _ => isIso_of_isIso sq, fun _ => isIso_of_isIso (inv sq)⟩ + +lemma isIso_iff_isIso_of_isIso {W X Y Z : T} {f : W ⟶ X} {g : Y ⟶ Z} (sq : mk f ⟶ mk g) [IsIso sq] : + IsIso f ↔ IsIso g := + isIso_hom_iff_isIso_hom_of_isIso sq + +lemma isIso_hom_iff_isIso_of_isIso {Y Z : T} {f : Arrow T} {g : Y ⟶ Z} (sq : f ⟶ mk g) [IsIso sq] : + IsIso f.hom ↔ IsIso g := + isIso_hom_iff_isIso_hom_of_isIso sq + @[simp] theorem inv_left [IsIso sq] : (inv sq).left = inv sq.left := IsIso.eq_inv_of_hom_inv_id <| by rw [← Comma.comp_left, IsIso.hom_inv_id, id_left] From 81a72b96bdfd703ae0ac6032fcfb159b823739b0 Mon Sep 17 00:00:00 2001 From: Stefan Kebekus Date: Sat, 1 Feb 2025 12:30:25 +0000 Subject: [PATCH 7/8] doc: fix several typos (#21315) Fix several typos in the documentation that I myself introduced yesterday :-( I apologize for causing this extra work. --- Mathlib/Analysis/SpecialFunctions/Log/PosLog.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/Log/PosLog.lean b/Mathlib/Analysis/SpecialFunctions/Log/PosLog.lean index 7257efd994537..b38cbdfd2ac5d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/PosLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/PosLog.lean @@ -8,8 +8,9 @@ import Mathlib.Analysis.SpecialFunctions.Log.Basic /-! # The Positive Part of the Logarithm -This file defines the function `Real.posLog = r ↦ max 0 (log r)`, introduces the notation `log⁺, -For a finite length-`n` sequence `f i` of reals, it establishes the following standard estimates. +This file defines the function `Real.posLog = r ↦ max 0 (log r)` and introduces the notation +`log⁺`. For a finite length-`n` sequence `f i` of reals, it establishes the following standard +estimates. - `theorem posLog_prod : log⁺ (∏ i, f i) ≤ ∑ i, log⁺ (f i)` @@ -25,7 +26,7 @@ namespace Real /-- Definition: the positive part of the logarithm. -/ noncomputable def posLog : ℝ → ℝ := fun r ↦ max 0 (log r) -/-- Notation `log⁺` for the real part of the logarithm. -/ +/-- Notation `log⁺` for the positive part of the logarithm. -/ scoped notation "log⁺" => posLog /-- Definition of the positive part of the logarithm, formulated as a theorem. -/ @@ -63,7 +64,7 @@ theorem posLog_eq_zero_iff (x : ℝ) : log⁺ x = 0 ↔ |x| ≤ 1 := by rw [← posLog_abs, ← log_nonpos_iff (abs_nonneg x)] simp [posLog] -/-- The function `log⁺` equals `log` outside of in the interval (-1,1). -/ +/-- The function `log⁺` equals `log` outside of the interval (-1,1). -/ theorem posLog_eq_log {x : ℝ} (hx : 1 ≤ |x|) : log⁺ x = log x := by simp only [posLog, sup_eq_right] rw [← log_abs] @@ -75,7 +76,7 @@ theorem log_of_nat_eq_posLog {n : ℕ} : log⁺ n = log n := by · simp [hn, posLog] · simp [posLog_eq_log, Nat.one_le_iff_ne_zero.2 hn] -/-- The function `log⁺` is monotone in the positive axis. -/ +/-- The function `log⁺` is monotone on the positive axis. -/ theorem monotoneOn_posLog : MonotoneOn log⁺ (Set.Ici 0) := by intro x hx y hy hxy simp only [posLog, le_sup_iff, sup_le_iff, le_refl, true_and] From 213a36c22c8034e3b9ea96f90bcc04d44c20ee7c Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 1 Feb 2025 13:19:29 +0000 Subject: [PATCH 8/8] chore: cleanup imports in Archive/IfNormalization (#21318) --- Archive/Examples/IfNormalization/Result.lean | 2 -- Archive/Examples/IfNormalization/WithoutAesop.lean | 4 +--- 2 files changed, 1 insertion(+), 5 deletions(-) diff --git a/Archive/Examples/IfNormalization/Result.lean b/Archive/Examples/IfNormalization/Result.lean index 30ac2f21636cd..93aad7675d064 100644 --- a/Archive/Examples/IfNormalization/Result.lean +++ b/Archive/Examples/IfNormalization/Result.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import Archive.Examples.IfNormalization.Statement -import Mathlib.Algebra.Order.Monoid.Canonical.Defs -import Mathlib.Algebra.Order.Monoid.Unbundled.MinMax import Mathlib.Data.List.AList import Mathlib.Tactic.Recall diff --git a/Archive/Examples/IfNormalization/WithoutAesop.lean b/Archive/Examples/IfNormalization/WithoutAesop.lean index ca384da0f2ffe..3f26989f133f5 100644 --- a/Archive/Examples/IfNormalization/WithoutAesop.lean +++ b/Archive/Examples/IfNormalization/WithoutAesop.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Kim Morrison -/ import Archive.Examples.IfNormalization.Statement -import Mathlib.Algebra.Order.Monoid.Canonical.Defs -import Mathlib.Algebra.Order.Monoid.Unbundled.MinMax import Mathlib.Data.List.AList /-! @@ -20,7 +18,7 @@ we put primes on the declarations in the file.) namespace IfExpr attribute [local simp] eval normalized hasNestedIf hasConstantIf hasRedundantIf disjoint vars - List.disjoint max_add_add_right max_mul_mul_left Nat.lt_add_one_iff le_add_of_le_right + List.disjoint theorem eval_ite_ite' {a b c d e : IfExpr} {f : ℕ → Bool} : (ite (ite a b c) d e).eval f = (ite a (ite b d e) (ite c d e)).eval f := by