diff --git a/.github/build.in.yml b/.github/build.in.yml index c3b0ebc6aad3d..3facc8b60a92b 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -133,8 +133,8 @@ jobs: run: | rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init - lake exe cache get Mathlib.Init - lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" + lake exe cache get #Mathlib.Init + #lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean id: mk_all diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index e8f0399a83c22..207aee8482616 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -143,8 +143,8 @@ jobs: run: | rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init - lake exe cache get Mathlib.Init - lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" + lake exe cache get #Mathlib.Init + #lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean id: mk_all diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 72978c1079fe8..c47cd306ed4da 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -150,8 +150,8 @@ jobs: run: | rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init - lake exe cache get Mathlib.Init - lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" + lake exe cache get #Mathlib.Init + #lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean id: mk_all diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 3bfe7e18fee50..c431ad5e7d2f2 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -147,8 +147,8 @@ jobs: run: | rm -rf .lake/build/lib/Mathlib/ # Fail quickly if the cache is completely cold, by checking for Mathlib.Init - lake exe cache get Mathlib.Init - lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" + lake exe cache get #Mathlib.Init + #lake build --no-build Mathlib.Init && lake exe cache get || echo "No cache for 'Mathlib.Init' available" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean id: mk_all diff --git a/.github/workflows/lean4checker.yml b/.github/workflows/lean4checker.yml index 8405cde825320..8976cc19badb6 100644 --- a/.github/workflows/lean4checker.yml +++ b/.github/workflows/lean4checker.yml @@ -70,7 +70,7 @@ jobs: run: | git clone https://github.com/leanprover/lean4checker cd lean4checker - git checkout v4.12.0-rc1 + git checkout v4.13.0-rc3 # Now that the git hash is embedded in each olean, # we need to compile lean4checker on the same toolchain cp ../lean-toolchain . diff --git a/Archive/Imo/Imo1962Q1.lean b/Archive/Imo/Imo1962Q1.lean index dd0ef63d96b8b..d151c6df95fc3 100644 --- a/Archive/Imo/Imo1962Q1.lean +++ b/Archive/Imo/Imo1962Q1.lean @@ -107,7 +107,7 @@ lemma case_more_digits {c n : ℕ} (hc : (digits 10 c).length ≥ 6) (hpp : Prob calc n ≥ 10 * c := le.intro hpp.left.symm _ ≥ 10 ^ (digits 10 c).length := base_pow_length_digits_le 10 c (by decide) hnz - _ ≥ 10 ^ 6 := pow_le_pow_right (by decide) hc + _ ≥ 10 ^ 6 := pow_right_mono₀ (by decide) hc _ ≥ 153846 := by norm_num /-! diff --git a/Archive/Imo/Imo2024Q1.lean b/Archive/Imo/Imo2024Q1.lean index 0e94469d1e0d9..00575c6abc94a 100644 --- a/Archive/Imo/Imo2024Q1.lean +++ b/Archive/Imo/Imo2024Q1.lean @@ -71,7 +71,7 @@ lemma mem_Ico_one_of_mem_Ioo (h : α ∈ Set.Ioo 0 2) : α ∈ Set.Ico 1 2 := by by_contra! hn have hr : 1 < ⌈α⁻¹⌉₊ := by rw [Nat.lt_ceil] - exact_mod_cast one_lt_inv h0 hn + exact_mod_cast (one_lt_inv₀ h0).2 hn apply hr.ne' suffices ⌈α⁻¹⌉₊ = (1 : ℤ) from mod_cast this apply Int.eq_one_of_dvd_one (Int.zero_le_ofNat _) @@ -153,7 +153,7 @@ lemma not_condition_of_mem_Ioo {α : ℝ} (h : α ∈ Set.Ioo 0 2) : ¬Condition convert hna using 1 field_simp rw [sub_eq_add_neg, ← le_sub_iff_add_le', neg_le, neg_sub] at hna' - rw [le_inv (by linarith) (mod_cast hn), ← not_lt] at hna' + rw [le_inv_comm₀ (by linarith) (mod_cast hn), ← not_lt] at hna' apply hna' exact_mod_cast Nat.lt_floor_add_one (_ : ℝ) diff --git a/Archive/Wiedijk100Theorems/AbelRuffini.lean b/Archive/Wiedijk100Theorems/AbelRuffini.lean index df179eb859924..36026dc2ddd76 100644 --- a/Archive/Wiedijk100Theorems/AbelRuffini.lean +++ b/Archive/Wiedijk100Theorems/AbelRuffini.lean @@ -115,7 +115,7 @@ theorem real_roots_Phi_ge_aux (hab : b < a) : · have hf1 : f 1 < 0 := by simp [hf, hb] have hfa : 0 ≤ f a := by simp_rw [hf, ← sq] - refine add_nonneg (sub_nonneg.mpr (pow_le_pow_right ha ?_)) ?_ <;> norm_num + refine add_nonneg (sub_nonneg.mpr (pow_right_mono₀ ha ?_)) ?_ <;> norm_num obtain ⟨x, ⟨-, hx1⟩, hx2⟩ := intermediate_value_Ico' hle (hc _) (Set.mem_Ioc.mpr ⟨hf1, hf0⟩) obtain ⟨y, ⟨hy1, -⟩, hy2⟩ := intermediate_value_Ioc ha (hc _) (Set.mem_Ioc.mpr ⟨hf1, hfa⟩) exact ⟨x, y, (hx1.trans hy1).ne, hx2, hy2⟩ @@ -126,7 +126,7 @@ theorem real_roots_Phi_ge_aux (hab : b < a) : f (-a) = (a : ℝ) ^ 2 - (a : ℝ) ^ 5 + b := by norm_num [hf, ← sq, sub_eq_add_neg, add_comm, Odd.neg_pow (by decide : Odd 5)] _ ≤ (a : ℝ) ^ 2 - (a : ℝ) ^ 3 + (a - 1) := by - refine add_le_add (sub_le_sub_left (pow_le_pow_right ha ?_) _) ?_ <;> linarith + refine add_le_add (sub_le_sub_left (pow_right_mono₀ ha ?_) _) ?_ <;> linarith _ = -((a : ℝ) - 1) ^ 2 * (a + 1) := by ring _ ≤ 0 := by nlinarith have ha' := neg_nonpos.mpr (hle.trans ha) diff --git a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean index 74017293ce14b..e30881c10059f 100644 --- a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean +++ b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean @@ -154,7 +154,7 @@ theorem card_le_two_pow {x k : ℕ} : card M₁ ≤ card (image f K) := card_le_card h _ ≤ card K := card_image_le _ ≤ 2 ^ card (image Nat.succ (range k)) := by simp only [K, card_powerset]; rfl - _ ≤ 2 ^ card (range k) := pow_le_pow_right one_le_two card_image_le + _ ≤ 2 ^ card (range k) := pow_right_mono₀ one_le_two card_image_le _ = 2 ^ k := by rw [card_range k] /-- diff --git a/Counterexamples/SorgenfreyLine.lean b/Counterexamples/SorgenfreyLine.lean index 746700b0769de..bdebf6085e105 100644 --- a/Counterexamples/SorgenfreyLine.lean +++ b/Counterexamples/SorgenfreyLine.lean @@ -110,7 +110,7 @@ theorem nhds_countable_basis_Ico_inv_pnat (a : ℝₗ) : theorem nhds_antitone_basis_Ico_inv_pnat (a : ℝₗ) : (𝓝 a).HasAntitoneBasis fun n : ℕ+ => Ico a (a + (n : ℝₗ)⁻¹) := ⟨nhds_basis_Ico_inv_pnat a, monotone_const.Ico <| Antitone.const_add - (fun k _l hkl => inv_le_inv_of_le (Nat.cast_pos.2 k.2) + (fun k _l hkl => inv_anti₀ (Nat.cast_pos.2 k.2) (Nat.mono_cast <| Subtype.coe_le_coe.2 hkl)) _⟩ theorem isOpen_iff {s : Set ℝₗ} : IsOpen s ↔ ∀ x ∈ s, ∃ y > x, Ico x y ⊆ s := @@ -141,7 +141,7 @@ theorem continuous_toReal : Continuous toReal := exact inf_le_left instance : OrderClosedTopology ℝₗ := - ⟨isClosed_le_prod.preimage (continuous_toReal.prod_map continuous_toReal)⟩ + ⟨isClosed_le_prod.preimage (continuous_toReal.prodMap continuous_toReal)⟩ instance : ContinuousAdd ℝₗ := by refine ⟨continuous_iff_continuousAt.2 ?_⟩ diff --git a/Mathlib.lean b/Mathlib.lean index 0c0cd4d93a5eb..7ba04a7c6a7f7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1023,6 +1023,7 @@ import Mathlib.Analysis.Calculus.FDeriv.Extend import Mathlib.Analysis.Calculus.FDeriv.Linear import Mathlib.Analysis.Calculus.FDeriv.Measurable import Mathlib.Analysis.Calculus.FDeriv.Mul +import Mathlib.Analysis.Calculus.FDeriv.Norm import Mathlib.Analysis.Calculus.FDeriv.Pi import Mathlib.Analysis.Calculus.FDeriv.Prod import Mathlib.Analysis.Calculus.FDeriv.RestrictScalars @@ -1096,6 +1097,7 @@ import Mathlib.Analysis.ConstantSpeed import Mathlib.Analysis.Convex.AmpleSet import Mathlib.Analysis.Convex.Basic import Mathlib.Analysis.Convex.Between +import Mathlib.Analysis.Convex.Birkhoff import Mathlib.Analysis.Convex.Body import Mathlib.Analysis.Convex.Caratheodory import Mathlib.Analysis.Convex.Combination @@ -2354,6 +2356,7 @@ import Mathlib.Data.Matrix.CharP import Mathlib.Data.Matrix.ColumnRowPartitioned import Mathlib.Data.Matrix.Composition import Mathlib.Data.Matrix.DMatrix +import Mathlib.Data.Matrix.DoublyStochastic import Mathlib.Data.Matrix.DualNumber import Mathlib.Data.Matrix.Hadamard import Mathlib.Data.Matrix.Invertible @@ -3613,6 +3616,7 @@ import Mathlib.Order.Filter.AtTopBot import Mathlib.Order.Filter.AtTopBot.Archimedean import Mathlib.Order.Filter.AtTopBot.BigOperators import Mathlib.Order.Filter.AtTopBot.Field +import Mathlib.Order.Filter.AtTopBot.Floor import Mathlib.Order.Filter.AtTopBot.Group import Mathlib.Order.Filter.AtTopBot.ModEq import Mathlib.Order.Filter.AtTopBot.Monoid @@ -3625,6 +3629,7 @@ import Mathlib.Order.Filter.Cofinite import Mathlib.Order.Filter.CountableInter import Mathlib.Order.Filter.CountableSeparatingOn import Mathlib.Order.Filter.Curry +import Mathlib.Order.Filter.Defs import Mathlib.Order.Filter.ENNReal import Mathlib.Order.Filter.EventuallyConst import Mathlib.Order.Filter.Extr @@ -3880,11 +3885,13 @@ import Mathlib.RingTheory.Derivation.ToSquareZero import Mathlib.RingTheory.DiscreteValuationRing.Basic import Mathlib.RingTheory.DiscreteValuationRing.TFAE import Mathlib.RingTheory.Discriminant +import Mathlib.RingTheory.DualNumber import Mathlib.RingTheory.EisensteinCriterion import Mathlib.RingTheory.EssentialFiniteness import Mathlib.RingTheory.Etale.Basic import Mathlib.RingTheory.EuclideanDomain import Mathlib.RingTheory.Filtration +import Mathlib.RingTheory.FiniteLength import Mathlib.RingTheory.FinitePresentation import Mathlib.RingTheory.FiniteStability import Mathlib.RingTheory.FiniteType @@ -3965,7 +3972,6 @@ import Mathlib.RingTheory.LocalRing.Module import Mathlib.RingTheory.LocalRing.ResidueField.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Defs import Mathlib.RingTheory.LocalRing.RingHom.Basic -import Mathlib.RingTheory.LocalRing.RingHom.Defs import Mathlib.RingTheory.Localization.Algebra import Mathlib.RingTheory.Localization.AsSubring import Mathlib.RingTheory.Localization.AtPrime @@ -4160,6 +4166,7 @@ import Mathlib.SetTheory.Lists import Mathlib.SetTheory.Ordinal.Arithmetic import Mathlib.SetTheory.Ordinal.Basic import Mathlib.SetTheory.Ordinal.CantorNormalForm +import Mathlib.SetTheory.Ordinal.Enum import Mathlib.SetTheory.Ordinal.Exponential import Mathlib.SetTheory.Ordinal.FixedPoint import Mathlib.SetTheory.Ordinal.FixedPointApproximants @@ -4321,6 +4328,7 @@ import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Linter.Lint import Mathlib.Tactic.Linter.MinImports import Mathlib.Tactic.Linter.OldObtain +import Mathlib.Tactic.Linter.PPRoundtrip import Mathlib.Tactic.Linter.RefineLinter import Mathlib.Tactic.Linter.Style import Mathlib.Tactic.Linter.TextBased @@ -4386,6 +4394,7 @@ import Mathlib.Tactic.RewriteSearch import Mathlib.Tactic.Rify import Mathlib.Tactic.Ring import Mathlib.Tactic.Ring.Basic +import Mathlib.Tactic.Ring.Compare import Mathlib.Tactic.Ring.PNat import Mathlib.Tactic.Ring.RingNF import Mathlib.Tactic.Sat.FromLRAT @@ -4612,6 +4621,7 @@ import Mathlib.Topology.ContinuousMap.CocompactMap import Mathlib.Topology.ContinuousMap.Compact import Mathlib.Topology.ContinuousMap.CompactlySupported import Mathlib.Topology.ContinuousMap.ContinuousMapZero +import Mathlib.Topology.ContinuousMap.Defs import Mathlib.Topology.ContinuousMap.Ideals import Mathlib.Topology.ContinuousMap.LocallyConstant import Mathlib.Topology.ContinuousMap.Ordered @@ -4741,7 +4751,7 @@ import Mathlib.Topology.MetricSpace.ThickenedIndicator import Mathlib.Topology.MetricSpace.Thickening import Mathlib.Topology.MetricSpace.Ultra.Basic import Mathlib.Topology.MetricSpace.Ultra.ContinuousMaps -import Mathlib.Topology.MetricSpace.Ultra.TotallyDisconnected +import Mathlib.Topology.MetricSpace.Ultra.TotallySeparated import Mathlib.Topology.Metrizable.Basic import Mathlib.Topology.Metrizable.ContinuousMap import Mathlib.Topology.Metrizable.Uniformity diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index a7a16de103ee9..f4a24f5977006 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -23,12 +23,12 @@ universe u v w u₁ v₁ namespace Algebra -variable {R : Type u} {S : Type v} {A : Type w} {B : Type*} +variable {R : Type u} {A : Type w} section Semiring -variable [CommSemiring R] [CommSemiring S] -variable [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] +variable [CommSemiring R] +variable [Semiring A] [Algebra R A] section PUnit @@ -319,7 +319,6 @@ section IsScalarTower variable {R : Type*} [CommSemiring R] variable (A : Type*) [Semiring A] [Algebra R A] variable {M : Type*} [AddCommMonoid M] [Module A M] [Module R M] [IsScalarTower R A M] -variable {N : Type*} [AddCommMonoid N] [Module A N] [Module R N] [IsScalarTower R A N] theorem algebra_compatible_smul (r : R) (m : M) : r • m = (algebraMap R A) r • m := by rw [← one_smul A m, ← smul_assoc, Algebra.smul_def, mul_one, one_smul] diff --git a/Mathlib/Algebra/Algebra/Spectrum.lean b/Mathlib/Algebra/Algebra/Spectrum.lean index 2b7e28f8e2c29..65c69f706fb00 100644 --- a/Mathlib/Algebra/Algebra/Spectrum.lean +++ b/Mathlib/Algebra/Algebra/Spectrum.lean @@ -400,7 +400,7 @@ end CommSemiring section CommRing -variable {F R A B : Type*} [CommRing R] [Ring A] [Algebra R A] [Ring B] [Algebra R B] +variable {F R A : Type*} [CommRing R] [Ring A] [Algebra R A] variable [FunLike F A R] [AlgHomClass F R A R] local notation "σ" => spectrum R diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index 6eb45fd0b5183..47308730de6b6 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -248,6 +248,12 @@ theorem Irreducible.dvd_comm [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irr p ∣ q ↔ q ∣ p := ⟨hp.dvd_symm hq, hq.dvd_symm hp⟩ +theorem Irreducible.of_map {F : Type*} [Monoid M] [Monoid N] [FunLike F M N] [MonoidHomClass F M N] + {f : F} [IsLocalRingHom f] {x} (hfx : Irreducible (f x)) : Irreducible x := + ⟨fun hu ↦ hfx.not_unit <| hu.map f, + by rintro p q rfl + exact (hfx.isUnit_or_isUnit <| map_mul f p q).imp (.of_map f _) (.of_map f _)⟩ + section variable [Monoid M] diff --git a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean index e0919abefa13d..36546a9af0623 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean @@ -77,7 +77,7 @@ instance hasForgetToRing : HasForget₂ (AlgebraCat.{v} R) RingCat.{v} where instance hasForgetToModule : HasForget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R) where forget₂ := { obj := fun M => ModuleCat.of R M - map := fun f => ModuleCat.ofHom f.toLinearMap } + map := fun f => ModuleCat.asHom f.toLinearMap } @[simp] lemma forget₂_module_obj (X : AlgebraCat.{v} R) : @@ -86,7 +86,7 @@ lemma forget₂_module_obj (X : AlgebraCat.{v} R) : @[simp] lemma forget₂_module_map {X Y : AlgebraCat.{v} R} (f : X ⟶ Y) : - (forget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R)).map f = ModuleCat.ofHom f.toLinearMap := + (forget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R)).map f = ModuleCat.asHom f.toLinearMap := rfl /-- The object in the category of R-algebras associated to a type equipped with the appropriate diff --git a/Mathlib/Algebra/Category/BialgebraCat/Basic.lean b/Mathlib/Algebra/Category/BialgebraCat/Basic.lean index 3cf9cd5c6fcbe..0d3b3d8260dad 100644 --- a/Mathlib/Algebra/Category/BialgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/BialgebraCat/Basic.lean @@ -59,7 +59,7 @@ lemma of_counit {X : Type v} [Ring X] [Bialgebra R X] : /-- A type alias for `BialgHom` to avoid confusion between the categorical and algebraic spellings of composition. -/ @[ext] -structure Hom (V W : BialgebraCat.{v} R) := +structure Hom (V W : BialgebraCat.{v} R) where /-- The underlying `BialgHom` -/ toBialgHom : V →ₐc[R] W diff --git a/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean b/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean index 2e3227b072f57..cbf7f0ebab743 100644 --- a/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean @@ -62,7 +62,7 @@ lemma of_counit {X : Type v} [AddCommGroup X] [Module R X] [Coalgebra R X] : /-- A type alias for `CoalgHom` to avoid confusion between the categorical and algebraic spellings of composition. -/ @[ext] -structure Hom (V W : CoalgebraCat.{v} R) := +structure Hom (V W : CoalgebraCat.{v} R) where /-- The underlying `CoalgHom` -/ toCoalgHom : V →ₗc[R] W diff --git a/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean b/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean index e6957905eb7f9..0a138d78803f1 100644 --- a/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean +++ b/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean @@ -38,8 +38,8 @@ variable {R : Type u} [CommRing R] /-- An `R`-coalgebra is a comonoid object in the category of `R`-modules. -/ @[simps] def toComonObj (X : CoalgebraCat R) : Comon_ (ModuleCat R) where X := ModuleCat.of R X - counit := ModuleCat.ofHom Coalgebra.counit - comul := ModuleCat.ofHom Coalgebra.comul + counit := ModuleCat.asHom Coalgebra.counit + comul := ModuleCat.asHom Coalgebra.comul counit_comul := by simpa only [ModuleCat.of_coe] using Coalgebra.rTensor_counit_comp_comul comul_counit := by simpa only [ModuleCat.of_coe] using Coalgebra.lTensor_counit_comp_comul comul_assoc := by simp_rw [ModuleCat.of_coe]; exact Coalgebra.coassoc.symm @@ -50,7 +50,7 @@ variable (R) in def toComon : CoalgebraCat R ⥤ Comon_ (ModuleCat R) where obj X := toComonObj X map f := - { hom := ModuleCat.ofHom f.1 + { hom := ModuleCat.asHom f.1 hom_counit := f.1.counit_comp hom_comul := f.1.map_comp_comul.symm } diff --git a/Mathlib/Algebra/Category/FGModuleCat/Basic.lean b/Mathlib/Algebra/Category/FGModuleCat/Basic.lean index 01ae8aab6bb4d..50f9e8d2c035a 100644 --- a/Mathlib/Algebra/Category/FGModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/FGModuleCat/Basic.lean @@ -278,8 +278,8 @@ end FGModuleCat @[simp] theorem LinearMap.comp_id_fgModuleCat {R} [Ring R] {G : FGModuleCat.{u} R} {H : Type u} [AddCommGroup H] [Module R H] (f : G →ₗ[R] H) : f.comp (𝟙 G) = f := - Category.id_comp (ModuleCat.ofHom f) + Category.id_comp (ModuleCat.asHom f) @[simp] theorem LinearMap.id_fgModuleCat_comp {R} [Ring R] {G : Type u} [AddCommGroup G] [Module R G] {H : FGModuleCat.{u} R} (f : G →ₗ[R] H) : LinearMap.comp (𝟙 H) f = f := - Category.comp_id (ModuleCat.ofHom f) + Category.comp_id (ModuleCat.asHom f) diff --git a/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean b/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean index b690a7f5f0bb5..413f5f5b36aba 100644 --- a/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean @@ -58,7 +58,7 @@ lemma of_counit {X : Type v} [Ring X] [HopfAlgebra R X] : /-- A type alias for `BialgHom` to avoid confusion between the categorical and algebraic spellings of composition. -/ @[ext] -structure Hom (V W : HopfAlgebraCat.{v} R) := +structure Hom (V W : HopfAlgebraCat.{v} R) where /-- The underlying `BialgHom`. -/ toBialgHom : V →ₐc[R] W diff --git a/Mathlib/Algebra/Category/ModuleCat/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Basic.lean index 48187e0d55415..ec69a10b56fc3 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Basic.lean @@ -151,18 +151,6 @@ theorem forget₂_map (X Y : ModuleCat R) (f : X ⟶ Y) : (forget₂ (ModuleCat R) AddCommGrp).map f = LinearMap.toAddMonoidHom f := rfl --- Porting note (#11215): TODO: `ofHom` and `asHom` are duplicates! - -/-- Typecheck a `LinearMap` as a morphism in `Module R`. -/ -def ofHom {R : Type u} [Ring R] {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] - [Module R Y] (f : X →ₗ[R] Y) : of R X ⟶ of R Y := - f - -@[simp 1100] -theorem ofHom_apply {R : Type u} [Ring R] {X Y : Type v} [AddCommGroup X] [Module R X] - [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) (x : X) : ofHom f x = f x := - rfl - instance : Inhabited (ModuleCat R) := ⟨of R PUnit⟩ @@ -218,14 +206,25 @@ end ModuleCat variable {R} variable {X₁ X₂ : Type v} +open ModuleCat + /-- Reinterpreting a linear map in the category of `R`-modules. -/ def ModuleCat.asHom [AddCommGroup X₁] [Module R X₁] [AddCommGroup X₂] [Module R X₂] : (X₁ →ₗ[R] X₂) → (ModuleCat.of R X₁ ⟶ ModuleCat.of R X₂) := id +@[deprecated (since := "2024-10-06")] alias ModuleCat.ofHom := ModuleCat.asHom + /-- Reinterpreting a linear map in the category of `R`-modules -/ scoped[ModuleCat] notation "↟" f:1024 => ModuleCat.asHom f +@[simp 1100] +theorem ModuleCat.asHom_apply {R : Type u} [Ring R] {X Y : Type v} [AddCommGroup X] [Module R X] + [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) (x : X) : (↟ f) x = f x := + rfl + +@[deprecated (since := "2024-10-06")] alias ModuleCat.ofHom_apply := ModuleCat.asHom_apply + /-- Reinterpreting a linear map in the category of `R`-modules. -/ def ModuleCat.asHomRight [AddCommGroup X₁] [Module R X₁] {X₂ : ModuleCat.{v} R} : (X₁ →ₗ[R] X₂) → (ModuleCat.of R X₁ ⟶ X₂) := @@ -441,8 +440,9 @@ end ModuleCat @[simp] theorem LinearMap.comp_id_moduleCat {R} [Ring R] {G : ModuleCat.{u} R} {H : Type u} [AddCommGroup H] [Module R H] (f : G →ₗ[R] H) : f.comp (𝟙 G) = f := - Category.id_comp (ModuleCat.ofHom f) + Category.id_comp (ModuleCat.asHom f) @[simp] theorem LinearMap.id_moduleCat_comp {R} [Ring R] {G : Type u} [AddCommGroup G] [Module R G] {H : ModuleCat.{u} R} (f : G →ₗ[R] H) : LinearMap.comp (𝟙 H) f = f := - Category.comp_id (ModuleCat.ofHom f) + Category.comp_id (ModuleCat.asHom f) + diff --git a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean index 445b26e61ac33..4bb37bb584fd6 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean @@ -152,8 +152,8 @@ of modules. -/ noncomputable def lequivProdOfRightSplitExact {f : B →ₗ[R] M} (hj : Function.Injective j) (exac : LinearMap.range j = LinearMap.ker g) (h : g.comp f = LinearMap.id) : (A × B) ≃ₗ[R] M := ((ShortComplex.Splitting.ofExactOfSection _ - (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.ofHom j) - (ModuleCat.ofHom g) exac) (asHom f) h + (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.asHom j) + (ModuleCat.asHom g) exac) (asHom f) h (by simpa only [ModuleCat.mono_iff_injective])).isoBinaryBiproduct ≪≫ biprodIsoProd _ _ ).symm.toLinearEquiv @@ -162,8 +162,8 @@ of modules. -/ noncomputable def lequivProdOfLeftSplitExact {f : M →ₗ[R] A} (hg : Function.Surjective g) (exac : LinearMap.range j = LinearMap.ker g) (h : f.comp j = LinearMap.id) : (A × B) ≃ₗ[R] M := ((ShortComplex.Splitting.ofExactOfRetraction _ - (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.ofHom j) - (ModuleCat.ofHom g) exac) (ModuleCat.ofHom f) h + (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.asHom j) + (ModuleCat.asHom g) exac) (ModuleCat.asHom f) h (by simpa only [ModuleCat.epi_iff_surjective] using hg)).isoBinaryBiproduct ≪≫ biprodIsoProd _ _).symm.toLinearEquiv diff --git a/Mathlib/Algebra/Category/ModuleCat/Images.lean b/Mathlib/Algebra/Category/ModuleCat/Images.lean index f08f291b4ebfa..cf1d2081b789e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Images.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Images.lean @@ -97,12 +97,12 @@ noncomputable def imageIsoRange {G H : ModuleCat.{v} R} (f : G ⟶ H) : @[simp, reassoc, elementwise] theorem imageIsoRange_inv_image_ι {G H : ModuleCat.{v} R} (f : G ⟶ H) : - (imageIsoRange f).inv ≫ Limits.image.ι f = ModuleCat.ofHom f.range.subtype := + (imageIsoRange f).inv ≫ Limits.image.ι f = ModuleCat.asHom f.range.subtype := IsImage.isoExt_inv_m _ _ @[simp, reassoc, elementwise] theorem imageIsoRange_hom_subtype {G H : ModuleCat.{v} R} (f : G ⟶ H) : - (imageIsoRange f).hom ≫ ModuleCat.ofHom f.range.subtype = Limits.image.ι f := by + (imageIsoRange f).hom ≫ ModuleCat.asHom f.range.subtype = Limits.image.ι f := by erw [← imageIsoRange_inv_image_ι f, Iso.hom_inv_id_assoc] end ModuleCat diff --git a/Mathlib/Algebra/Category/Ring/Instances.lean b/Mathlib/Algebra/Category/Ring/Instances.lean index c1e005f022104..fc588d27c8248 100644 --- a/Mathlib/Algebra/Category/Ring/Instances.lean +++ b/Mathlib/Algebra/Category/Ring/Instances.lean @@ -38,7 +38,7 @@ instance Localization.epi' {R : CommRingCat} (M : Submonoid R) : instance CommRingCat.isLocalRingHom_comp {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) [IsLocalRingHom g] [IsLocalRingHom f] : IsLocalRingHom (f ≫ g) := - _root_.isLocalRingHom_comp _ _ + RingHom.isLocalRingHom_comp _ _ theorem isLocalRingHom_of_iso {R S : CommRingCat} (f : R ≅ S) : IsLocalRingHom f.hom := { map_nonunit := fun a ha => by diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean b/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean index be6b7195e1964..e2e525d76dcdc 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean @@ -94,12 +94,9 @@ theorem one_le_succ_nth_stream_b {ifp_succ_n : IntFractPair K} ∃ ifp_n, IntFractPair.stream v n = some ifp_n ∧ ifp_n.fr ≠ 0 ∧ IntFractPair.of ifp_n.fr⁻¹ = ifp_succ_n := succ_nth_stream_eq_some_iff.1 succ_nth_stream_eq - suffices 1 ≤ ifp_n.fr⁻¹ by rwa [IntFractPair.of, le_floor, cast_one] - suffices ifp_n.fr ≤ 1 by - have h : 0 < ifp_n.fr := - lt_of_le_of_ne (nth_stream_fr_nonneg nth_stream_eq) stream_nth_fr_ne_zero.symm - apply one_le_inv h this - simp only [le_of_lt (nth_stream_fr_lt_one nth_stream_eq)] + rw [IntFractPair.of, le_floor, cast_one, one_le_inv₀ + ((nth_stream_fr_nonneg nth_stream_eq).lt_of_ne' stream_nth_fr_ne_zero)] + exact (nth_stream_fr_lt_one nth_stream_eq).le /-- Shows that the `n + 1`th integer part `bₙ₊₁` of the stream is smaller or equal than the inverse of diff --git a/Mathlib/Algebra/DirectSum/LinearMap.lean b/Mathlib/Algebra/DirectSum/LinearMap.lean index c3b5c570dd966..25ffc82193de4 100644 --- a/Mathlib/Algebra/DirectSum/LinearMap.lean +++ b/Mathlib/Algebra/DirectSum/LinearMap.lean @@ -78,12 +78,12 @@ lemma trace_eq_sum_trace_restrict' (h : IsInternal N) (hN : {i | N i ≠ ⊥}.Fi rw [← Finset.sum_coe_sort, trace_eq_sum_trace_restrict (isInternal_ne_bot_iff.mpr h) _] exact Fintype.sum_equiv hN.subtypeEquivToFinset _ _ (fun i ↦ rfl) -lemma trace_eq_zero_of_mapsTo_ne (h : IsInternal N) [hn : IsNoetherian R M] +lemma trace_eq_zero_of_mapsTo_ne (h : IsInternal N) [IsNoetherian R M] (σ : ι → ι) (hσ : ∀ i, σ i ≠ i) {f : Module.End R M} (hf : ∀ i, MapsTo f (N i) (N <| σ i)) : trace R M f = 0 := by - have hN : {i | N i ≠ ⊥}.Finite := CompleteLattice.WellFounded.finite_ne_bot_of_independent - hn.wf h.submodule_independent + have hN : {i | N i ≠ ⊥}.Finite := CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent + h.submodule_independent let s := hN.toFinset let κ := fun i ↦ Module.Free.ChooseBasisIndex R (N i) let b : (i : s) → Basis (κ i) R (N i) := fun i ↦ Module.Free.chooseBasis R (N i) @@ -100,27 +100,26 @@ lemma trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero [IsDomain R] [IsPrincipalIdealRing R] [Module.Free R M] [Module.Finite R M] {f g : Module.End R M} (h_comm : Commute f g) - (hf : ⨆ μ, ⨆ k, f.genEigenspace μ k = ⊤) - (hg : ∀ μ, trace R _ (g.restrict (f.mapsTo_iSup_genEigenspace_of_comm h_comm μ)) = 0) : + (hf : ⨆ μ, f.maxGenEigenspace μ = ⊤) + (hg : ∀ μ, trace R _ (g.restrict (f.mapsTo_maxGenEigenspace_of_comm h_comm μ)) = 0) : trace R _ (g ∘ₗ f) = 0 := by have hfg : ∀ μ, - MapsTo (g ∘ₗ f) ↑(⨆ k, f.genEigenspace μ k) ↑(⨆ k, f.genEigenspace μ k) := - fun μ ↦ (f.mapsTo_iSup_genEigenspace_of_comm h_comm μ).comp - (f.mapsTo_iSup_genEigenspace_of_comm rfl μ) + MapsTo (g ∘ₗ f) ↑(f.maxGenEigenspace μ) ↑(f.maxGenEigenspace μ) := + fun μ ↦ (f.mapsTo_maxGenEigenspace_of_comm h_comm μ).comp + (f.mapsTo_maxGenEigenspace_of_comm rfl μ) suffices ∀ μ, trace R _ ((g ∘ₗ f).restrict (hfg μ)) = 0 by classical have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - f.independent_genEigenspace hf - have h_fin : {μ | ⨆ k, f.genEigenspace μ k ≠ ⊥}.Finite := - CompleteLattice.WellFounded.finite_ne_bot_of_independent IsWellFounded.wf - f.independent_genEigenspace + f.independent_maxGenEigenspace hf + have h_fin : {μ | f.maxGenEigenspace μ ≠ ⊥}.Finite := + CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent f.independent_maxGenEigenspace simp [trace_eq_sum_trace_restrict' hds h_fin hfg, this] intro μ - replace h_comm : Commute (g.restrict (f.mapsTo_iSup_genEigenspace_of_comm h_comm μ)) - (f.restrict (f.mapsTo_iSup_genEigenspace_of_comm rfl μ)) := + replace h_comm : Commute (g.restrict (f.mapsTo_maxGenEigenspace_of_comm h_comm μ)) + (f.restrict (f.mapsTo_maxGenEigenspace_of_comm rfl μ)) := restrict_commute h_comm.symm _ _ rw [restrict_comp, trace_comp_eq_mul_of_commute_of_isNilpotent μ h_comm - (f.isNilpotent_restrict_iSup_sub_algebraMap μ), hg, mul_zero] + (f.isNilpotent_restrict_maxGenEigenspace_sub_algebraMap μ), hg, mul_zero] lemma mapsTo_biSup_of_mapsTo {ι : Type*} {N : ι → Submodule R M} (s : Set ι) {f : Module.End R M} (hf : ∀ i, MapsTo f (N i) (N i)) : diff --git a/Mathlib/Algebra/Field/Defs.lean b/Mathlib/Algebra/Field/Defs.lean index 46e889528e3e6..4bb07df74d908 100644 --- a/Mathlib/Algebra/Field/Defs.lean +++ b/Mathlib/Algebra/Field/Defs.lean @@ -194,7 +194,7 @@ variable (K) end NNRat namespace Rat -variable [DivisionRing K] {a b : K} +variable [DivisionRing K] lemma cast_def (q : ℚ) : (q : K) = q.num / q.den := DivisionRing.ratCast_def _ diff --git a/Mathlib/Algebra/Group/Action/Defs.lean b/Mathlib/Algebra/Group/Action/Defs.lean index 91a530d0f239f..733f6bd366afc 100644 --- a/Mathlib/Algebra/Group/Action/Defs.lean +++ b/Mathlib/Algebra/Group/Action/Defs.lean @@ -48,7 +48,7 @@ assert_not_exists MonoidWithZero open Function (Injective Surjective) -variable {M N G H A B α β γ δ : Type*} +variable {M N G H α β γ δ : Type*} /-! ### Faithful actions -/ @@ -487,7 +487,7 @@ lemma smul_inv_smul (g : G) (a : α) : g • g⁻¹ • a = a := by rw [smul_smu ⟨fun h ↦ by rw [h, smul_inv_smul], fun h ↦ by rw [← h, inv_smul_smul]⟩ section Mul -variable [Mul H] [MulAction G H] [SMulCommClass G H H] [IsScalarTower G H H] {g : G} {a b : H} +variable [Mul H] [MulAction G H] [SMulCommClass G H H] [IsScalarTower G H H] {a b : H} @[simp] lemma Commute.smul_right_iff : Commute a (g • b) ↔ Commute a b := ⟨fun h ↦ inv_smul_smul g b ▸ h.smul_right g⁻¹, fun h ↦ h.smul_right g⟩ diff --git a/Mathlib/Algebra/Group/Action/Opposite.lean b/Mathlib/Algebra/Group/Action/Opposite.lean index 8092a68ff91c7..1217487ac1eda 100644 --- a/Mathlib/Algebra/Group/Action/Opposite.lean +++ b/Mathlib/Algebra/Group/Action/Opposite.lean @@ -27,7 +27,7 @@ With `open scoped RightActions`, this provides: assert_not_exists MonoidWithZero -variable {R M N α : Type*} +variable {M N α β : Type*} /-! ### Actions _on_ the opposite type @@ -97,7 +97,7 @@ In lemma names this is still called `op_vadd`. -/ scoped notation3:73 m:73 " <+ᵥ " r:74 => AddOpposite.op r +ᵥ m section examples -variable {α β : Type*} [SMul α β] [SMul αᵐᵒᵖ β] [VAdd α β] [VAdd αᵃᵒᵖ β] {a a₁ a₂ a₃ a₄ : α} {b : β} +variable [SMul α β] [SMul αᵐᵒᵖ β] [VAdd α β] [VAdd αᵃᵒᵖ β] {a a₁ a₂ a₃ a₄ : α} {b : β} -- Left and right actions are just notation around the general `•` and `+ᵥ` notations example : a •> b = a • b := rfl @@ -124,7 +124,7 @@ end examples end RightActions section -variable {α β : Type*} [Monoid α] [MulAction αᵐᵒᵖ β] +variable [Monoid α] [MulAction αᵐᵒᵖ β] open scoped RightActions diff --git a/Mathlib/Algebra/Group/Action/Sum.lean b/Mathlib/Algebra/Group/Action/Sum.lean index 5bc1602680cfc..0286b8e19b6e9 100644 --- a/Mathlib/Algebra/Group/Action/Sum.lean +++ b/Mathlib/Algebra/Group/Action/Sum.lean @@ -20,7 +20,7 @@ This file defines instances for additive and multiplicative actions on the binar assert_not_exists MonoidWithZero -variable {M N P α β γ : Type*} +variable {M N α β : Type*} namespace Sum diff --git a/Mathlib/Algebra/Group/Basic.lean b/Mathlib/Algebra/Group/Basic.lean index 1fd34ff37fdbc..244c63da588ad 100644 --- a/Mathlib/Algebra/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Basic.lean @@ -172,7 +172,7 @@ end CommSemigroup attribute [local simp] mul_assoc sub_eq_add_neg section Monoid -variable [Monoid M] {a b c : M} {m n : ℕ} +variable [Monoid M] {a b : M} {m n : ℕ} @[to_additive boole_nsmul] lemma pow_boole (P : Prop) [Decidable P] (a : M) : @@ -316,7 +316,7 @@ end InvolutiveInv section DivInvMonoid -variable [DivInvMonoid G] {a b c : G} +variable [DivInvMonoid G] @[to_additive, field_simps] -- The attributes are out of order on purpose theorem inv_eq_one_div (x : G) : x⁻¹ = 1 / x := by rw [div_eq_mul_inv, one_mul] diff --git a/Mathlib/Algebra/Group/Center.lean b/Mathlib/Algebra/Group/Center.lean index de1591f22ebe2..786c03ef1f1db 100644 --- a/Mathlib/Algebra/Group/Center.lean +++ b/Mathlib/Algebra/Group/Center.lean @@ -70,7 +70,7 @@ attribute [to_additive existing] isMulCentral_iff namespace IsMulCentral -variable {a b c : M} [Mul M] +variable {a c : M} [Mul M] -- cf. `Commute.left_comm` @[to_additive] diff --git a/Mathlib/Algebra/Group/Commute/Defs.lean b/Mathlib/Algebra/Group/Commute/Defs.lean index b7bcb09ff56a4..324484366c26a 100644 --- a/Mathlib/Algebra/Group/Commute/Defs.lean +++ b/Mathlib/Algebra/Group/Commute/Defs.lean @@ -181,7 +181,7 @@ end Monoid section DivisionMonoid -variable [DivisionMonoid G] {a b c d : G} +variable [DivisionMonoid G] {a b : G} @[to_additive] protected theorem mul_inv (hab : Commute a b) : (a * b)⁻¹ = a⁻¹ * b⁻¹ := by rw [hab.eq, mul_inv_rev] diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index cd436c8a8fdfc..7f7d7ece643cd 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -35,7 +35,7 @@ actions and register the following instances: - `SMul ℕ M` for additive monoids `M`, and `SMul ℤ G` for additive groups `G`. `SMul` is typically, but not exclusively, used for scalar multiplication-like operators. -See the module `Algebra.AddTorsor` for a motivating example for the name `VAdd` (vector addition)`. +See the module `Algebra.AddTorsor` for a motivating example for the name `VAdd` (vector addition). ## Notation @@ -807,7 +807,7 @@ class DivInvMonoid (G : Type u) extends Monoid G, Inv G, Div G where /-- `a ^ 0 = 1` -/ protected zpow_zero' : ∀ a : G, zpow 0 a = 1 := by intros; rfl /-- `a ^ (n + 1) = a ^ n * a` -/ - protected zpow_succ' (n : ℕ) (a : G) : zpow (Int.ofNat n.succ) a = zpow (Int.ofNat n) a * a := by + protected zpow_succ' (n : ℕ) (a : G) : zpow n.succ a = zpow n a * a := by intros; rfl /-- `a ^ -(n + 1) = (a ^ (n + 1))⁻¹` -/ protected zpow_neg' (n : ℕ) (a : G) : zpow (Int.negSucc n) a = (zpow n.succ a)⁻¹ := by intros; rfl @@ -848,7 +848,7 @@ class SubNegMonoid (G : Type u) extends AddMonoid G, Neg G, Sub G where protected zsmul : ℤ → G → G protected zsmul_zero' : ∀ a : G, zsmul 0 a = 0 := by intros; rfl protected zsmul_succ' (n : ℕ) (a : G) : - zsmul (Int.ofNat n.succ) a = zsmul (Int.ofNat n) a + a := by + zsmul n.succ a = zsmul n a + a := by intros; rfl protected zsmul_neg' (n : ℕ) (a : G) : zsmul (Int.negSucc n) a = -zsmul n.succ a := by intros; rfl @@ -1050,7 +1050,7 @@ attribute [to_additive] Group section Group -variable [Group G] {a b c : G} +variable [Group G] {a b : G} @[to_additive (attr := simp)] theorem inv_mul_cancel (a : G) : a⁻¹ * a = 1 := diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index c95faf764777a..cb72fdfe7db40 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -28,7 +28,7 @@ Equiv, MulEquiv, AddEquiv open Function -variable {F α β A B M N P Q G H : Type*} +variable {F α β M N P G H : Type*} /-- Makes a `OneHom` inverse from the bijective inverse of a `OneHom` -/ @[to_additive (attr := simps) @@ -168,7 +168,7 @@ theorem MulEquivClass.toMulEquiv_injective [Mul α] [Mul β] [MulEquivClass F α namespace MulEquiv section Mul -variable [Mul M] [Mul N] [Mul P] [Mul Q] +variable [Mul M] [Mul N] [Mul P] section coe diff --git a/Mathlib/Algebra/Group/Even.lean b/Mathlib/Algebra/Group/Even.lean index 428970edc8bf0..5bd24647c2f65 100644 --- a/Mathlib/Algebra/Group/Even.lean +++ b/Mathlib/Algebra/Group/Even.lean @@ -34,7 +34,7 @@ assert_not_exists DenselyOrdered open MulOpposite -variable {F α β R : Type*} +variable {F α β : Type*} section Mul variable [Mul α] diff --git a/Mathlib/Algebra/Group/Fin/Basic.lean b/Mathlib/Algebra/Group/Fin/Basic.lean index d9710b7aedc23..9385f63b70721 100644 --- a/Mathlib/Algebra/Group/Fin/Basic.lean +++ b/Mathlib/Algebra/Group/Fin/Basic.lean @@ -23,7 +23,7 @@ assert_not_exists MonoidWithZero open Nat namespace Fin -variable {m n : ℕ} +variable {n : ℕ} /-! ### Instances -/ diff --git a/Mathlib/Algebra/Group/Hom/Basic.lean b/Mathlib/Algebra/Group/Hom/Basic.lean index 094fd70d45cc5..9eb317e8d7b71 100644 --- a/Mathlib/Algebra/Group/Hom/Basic.lean +++ b/Mathlib/Algebra/Group/Hom/Basic.lean @@ -16,7 +16,7 @@ import Mathlib.Algebra.Group.Hom.Defs -- `Algebra.Group` folder. assert_not_imported Mathlib.Algebra.NeZero -variable {α β M N P : Type*} +variable {α M N P : Type*} -- monoids variable {G : Type*} {H : Type*} @@ -98,7 +98,7 @@ end MulHom namespace MonoidHom section Group -variable [Group G] [CommGroup H] +variable [Group G] /-- A homomorphism from a group to a monoid is injective iff its kernel is trivial. For the iff statement on the triviality of the kernel, see `injective_iff_map_eq_one'`. -/ @@ -125,8 +125,6 @@ theorem _root_.injective_iff_map_eq_one' {G H} [Group G] [MulOneClass H] (injective_iff_map_eq_one f).trans <| forall_congr' fun _ => ⟨fun h => ⟨h, fun H => H.symm ▸ map_one f⟩, Iff.mp⟩ -variable [MulOneClass M] - /-- Makes a group homomorphism from a proof that the map preserves right division `fun x y => x * y⁻¹`. See also `MonoidHom.of_map_div` for a version using `fun x y => x / y`. -/ diff --git a/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean b/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean index 5f9f1cea90e7a..0bdd4b963dd4a 100644 --- a/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean +++ b/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean @@ -39,7 +39,7 @@ section MonoidHomCompTriple namespace MonoidHom /-- Class of composing triples -/ -class CompTriple {M N P : Type*} [Monoid M] [Monoid N] [Monoid P] +class CompTriple {M N P : Type*} [Monoid M] [Monoid N] [Monoid P] (φ : M →* N) (ψ : N →* P) (χ : outParam (M →* P)) : Prop where /-- The maps form a commuting triangle -/ comp_eq : ψ.comp φ = χ @@ -48,7 +48,6 @@ attribute [simp] CompTriple.comp_eq namespace CompTriple -variable {M' : Type*} [Monoid M'] variable {M N P : Type*} [Monoid M] [Monoid N] [Monoid P] /-- Class of Id maps -/ diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index 2279253985c32..1a3bf932f98fd 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -959,8 +959,6 @@ instance [MulOneClass M] [MulOneClass N] : Inhabited (M →* N) := ⟨1⟩ namespace MonoidHom -variable [Group G] [CommGroup H] - @[to_additive (attr := simp)] theorem one_comp [MulOneClass M] [MulOneClass N] [MulOneClass P] (f : M →* N) : (1 : N →* P).comp f = 1 := rfl diff --git a/Mathlib/Algebra/Group/Hom/End.lean b/Mathlib/Algebra/Group/Hom/End.lean index 00405f5550e0e..caa7de5333f3d 100644 --- a/Mathlib/Algebra/Group/Hom/End.lean +++ b/Mathlib/Algebra/Group/Hom/End.lean @@ -19,9 +19,9 @@ They are separate, and if someone would like to split this file in two that may -/ -universe uM uN uP uQ +universe uM -variable {M : Type uM} {N : Type uN} {P : Type uP} {Q : Type uQ} +variable {M : Type uM} namespace AddMonoid.End @@ -121,7 +121,7 @@ end Semiring section CommSemiring -variable {R S : Type*} [NonUnitalNonAssocCommSemiring R] +variable {R : Type*} [NonUnitalNonAssocCommSemiring R] namespace AddMonoid.End diff --git a/Mathlib/Algebra/Group/InjSurj.lean b/Mathlib/Algebra/Group/InjSurj.lean index a47bbde22c75a..c25210c59545a 100644 --- a/Mathlib/Algebra/Group/InjSurj.lean +++ b/Mathlib/Algebra/Group/InjSurj.lean @@ -55,7 +55,7 @@ a semigroup. See note [reducible non-instances]. -/ injective map that preserves `+` to an additive semigroup."] protected abbrev semigroup [Semigroup M₂] (f : M₁ → M₂) (hf : Injective f) (mul : ∀ x y, f (x * y) = f x * f y) : Semigroup M₁ := - { ‹Mul M₁› with mul_assoc := fun x y z => hf <| by erw [mul, mul, mul, mul, mul_assoc] } + { ‹Mul M₁› with mul_assoc := fun x y z => hf <| by rw [mul, mul, mul, mul, mul_assoc] } /-- A type endowed with `*` is a commutative magma, if it admits a surjective map that preserves `*` from a commutative magma. -/ @@ -83,7 +83,7 @@ semigroup, if it admits an injective map that preserves `+` to an additive left protected abbrev leftCancelSemigroup [LeftCancelSemigroup M₂] (f : M₁ → M₂) (hf : Injective f) (mul : ∀ x y, f (x * y) = f x * f y) : LeftCancelSemigroup M₁ := { hf.semigroup f mul with - mul_left_cancel := fun x y z H => hf <| (mul_right_inj (f x)).1 <| by erw [← mul, ← mul, H] } + mul_left_cancel := fun x y z H => hf <| (mul_right_inj (f x)).1 <| by rw [← mul, ← mul, H] } /-- A type endowed with `*` is a right cancel semigroup, if it admits an injective map that preserves `*` to a right cancel semigroup. See note [reducible non-instances]. -/ @@ -93,7 +93,7 @@ semigroup."] protected abbrev rightCancelSemigroup [RightCancelSemigroup M₂] (f : M₁ → M₂) (hf : Injective f) (mul : ∀ x y, f (x * y) = f x * f y) : RightCancelSemigroup M₁ := { hf.semigroup f mul with - mul_right_cancel := fun x y z H => hf <| (mul_left_inj (f y)).1 <| by erw [← mul, ← mul, H] } + mul_right_cancel := fun x y z H => hf <| (mul_left_inj (f y)).1 <| by rw [← mul, ← mul, H] } variable [One M₁] @@ -105,8 +105,8 @@ injective map that preserves `0` and `+` to an `AddZeroClass`."] protected abbrev mulOneClass [MulOneClass M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) : MulOneClass M₁ := { ‹One M₁›, ‹Mul M₁› with - one_mul := fun x => hf <| by erw [mul, one, one_mul], - mul_one := fun x => hf <| by erw [mul, one, mul_one] } + one_mul := fun x => hf <| by rw [mul, one, one_mul], + mul_one := fun x => hf <| by rw [mul, one, mul_one] } variable [Pow M₁ ℕ] @@ -120,8 +120,8 @@ protected abbrev monoid [Monoid M₂] (f : M₁ → M₂) (hf : Injective f) (on (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : Monoid M₁ := { hf.mulOneClass f one mul, hf.semigroup f mul with npow := fun n x => x ^ n, - npow_zero := fun x => hf <| by erw [npow, one, pow_zero], - npow_succ := fun n x => hf <| by erw [npow, pow_succ, mul, npow] } + npow_zero := fun x => hf <| by rw [npow, one, pow_zero], + npow_succ := fun n x => hf <| by rw [npow, pow_succ, mul, npow] } /-- A type endowed with `0`, `1` and `+` is an additive monoid with one, if it admits an injective map that preserves `0`, `1` and `+` to an additive monoid with one. @@ -132,8 +132,8 @@ protected abbrev addMonoidWithOne {M₁} [Zero M₁] [One M₁] [Add M₁] [SMul (natCast : ∀ n : ℕ, f n = n) : AddMonoidWithOne M₁ := { hf.addMonoid f zero add (swap nsmul) with natCast := Nat.cast, - natCast_zero := hf (by erw [natCast, Nat.cast_zero, zero]), - natCast_succ := fun n => hf (by erw [natCast, Nat.cast_succ, add, one, natCast]), one := 1 } + natCast_zero := hf (by rw [natCast, Nat.cast_zero, zero]), + natCast_succ := fun n => hf (by rw [natCast, Nat.cast_succ, add, one, natCast]), one := 1 } /-- A type endowed with `1` and `*` is a left cancel monoid, if it admits an injective map that preserves `1` and `*` to a left cancel monoid. See note [reducible non-instances]. -/ @@ -215,7 +215,7 @@ injective map that preserves `0` and unary `-` to an `NegZeroClass`."] protected abbrev invOneClass [InvOneClass M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (inv : ∀ x, f (x⁻¹) = (f x)⁻¹) : InvOneClass M₁ := { ‹One M₁›, ‹Inv M₁› with - inv_one := hf <| by erw [inv, one, inv_one] } + inv_one := hf <| by rw [inv, one, inv_one] } variable [Div M₁] [Pow M₁ ℤ] @@ -232,10 +232,10 @@ protected abbrev divInvMonoid [DivInvMonoid M₂] (f : M₁ → M₂) (hf : Inje (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : DivInvMonoid M₁ := { hf.monoid f one mul npow, ‹Inv M₁›, ‹Div M₁› with zpow := fun n x => x ^ n, - zpow_zero' := fun x => hf <| by erw [zpow, zpow_zero, one], - zpow_succ' := fun n x => hf <| by erw [zpow, mul, zpow_natCast, pow_succ, zpow, zpow_natCast], - zpow_neg' := fun n x => hf <| by erw [zpow, zpow_negSucc, inv, zpow, zpow_natCast], - div_eq_mul_inv := fun x y => hf <| by erw [div, mul, inv, div_eq_mul_inv] } + zpow_zero' := fun x => hf <| by rw [zpow, zpow_zero, one], + zpow_succ' := fun n x => hf <| by rw [zpow, mul, zpow_natCast, pow_succ, zpow, zpow_natCast], + zpow_neg' := fun n x => hf <| by rw [zpow, zpow_negSucc, inv, zpow, zpow_natCast], + div_eq_mul_inv := fun x y => hf <| by rw [div, mul, inv, div_eq_mul_inv] } /-- A type endowed with `1`, `*`, `⁻¹`, and `/` is a `DivInvOneMonoid` if it admits an injective map that preserves `1`, `*`, `⁻¹`, and `/` to a `DivInvOneMonoid`. See note @@ -263,9 +263,9 @@ protected abbrev divisionMonoid [DivisionMonoid M₂] (f : M₁ → M₂) (hf : (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : DivisionMonoid M₁ := { hf.divInvMonoid f one mul inv div npow zpow, hf.involutiveInv f inv with - mul_inv_rev := fun x y => hf <| by erw [inv, mul, mul_inv_rev, mul, inv, inv], + mul_inv_rev := fun x y => hf <| by rw [inv, mul, mul_inv_rev, mul, inv, inv], inv_eq_of_mul := fun x y h => hf <| by - erw [inv, inv_eq_of_mul_eq_one_right (by erw [← mul, h, one])] } + rw [inv, inv_eq_of_mul_eq_one_right (by rw [← mul, h, one])] } /-- A type endowed with `1`, `*`, `⁻¹`, and `/` is a `DivisionCommMonoid` if it admits an injective map that preserves `1`, `*`, `⁻¹`, and `/` to a `DivisionCommMonoid`. @@ -291,7 +291,7 @@ protected abbrev group [Group M₂] (f : M₁ → M₂) (hf : Injective f) (one (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : Group M₁ := { hf.divInvMonoid f one mul inv div npow zpow with - inv_mul_cancel := fun x => hf <| by erw [mul, inv, inv_mul_cancel, one] } + inv_mul_cancel := fun x => hf <| by rw [mul, inv, inv_mul_cancel, one] } /-- A type endowed with `0`, `1` and `+` is an additive group with one, if it admits an injective map that preserves `0`, `1` and `+` to an additive group with one. See note @@ -306,7 +306,7 @@ protected abbrev addGroupWithOne {M₁} [Zero M₁] [One M₁] [Add M₁] [SMul hf.addMonoidWithOne f zero one add nsmul natCast with intCast := Int.cast, intCast_ofNat := fun n => hf (by rw [natCast, ← Int.cast, intCast, Int.cast_natCast]), - intCast_negSucc := fun n => hf (by erw [intCast, neg, natCast, Int.cast_negSucc] ) } + intCast_negSucc := fun n => hf (by rw [intCast, neg, natCast, Int.cast_negSucc] ) } /-- A type endowed with `1`, `*` and `⁻¹` is a commutative group, if it admits an injective map that preserves `1`, `*` and `⁻¹` to a commutative group. See note [reducible non-instances]. -/ @@ -358,7 +358,7 @@ protected abbrev semigroup [Semigroup M₁] (f : M₁ → M₂) (hf : Surjective a surjective map that preserves `+` from an additive commutative semigroup."] protected abbrev commMagma [CommMagma M₁] (f : M₁ → M₂) (hf : Surjective f) (mul : ∀ x y, f (x * y) = f x * f y) : CommMagma M₂ where - mul_comm := hf.forall₂.2 fun x y => by erw [← mul, ← mul, mul_comm] + mul_comm := hf.forall₂.2 fun x y => by rw [← mul, ← mul, mul_comm] /-- A type endowed with `*` is a commutative semigroup, if it admits a surjective map that preserves `*` from a commutative semigroup. See note [reducible non-instances]. -/ @@ -380,8 +380,8 @@ surjective map that preserves `0` and `+` to an `AddZeroClass`."] protected abbrev mulOneClass [MulOneClass M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) : MulOneClass M₂ := { ‹One M₂›, ‹Mul M₂› with - one_mul := hf.forall.2 fun x => by erw [← one, ← mul, one_mul], - mul_one := hf.forall.2 fun x => by erw [← one, ← mul, mul_one] } + one_mul := hf.forall.2 fun x => by rw [← one, ← mul, one_mul], + mul_one := hf.forall.2 fun x => by rw [← one, ← mul, mul_one] } variable [Pow M₂ ℕ] @@ -395,10 +395,10 @@ protected abbrev monoid [Monoid M₁] (f : M₁ → M₂) (hf : Surjective f) (o (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : Monoid M₂ := { hf.semigroup f mul, hf.mulOneClass f one mul with npow := fun n x => x ^ n, - npow_zero := hf.forall.2 fun x => by dsimp only; erw [← npow, pow_zero, ← one], + npow_zero := hf.forall.2 fun x => by dsimp only; rw [← npow, pow_zero, ← one], npow_succ := fun n => hf.forall.2 fun x => by dsimp only - erw [← npow, pow_succ, ← npow, ← mul] } + rw [← npow, pow_succ, ← npow, ← mul] } /-- A type endowed with `0`, `1` and `+` is an additive monoid with one, if it admits a surjective map that preserves `0`, `1` and `*` from an additive monoid with one. See note @@ -441,7 +441,7 @@ preserves `-` to a type which has an involutive negation."] protected abbrev involutiveInv {M₂ : Type*} [Inv M₂] [InvolutiveInv M₁] (f : M₁ → M₂) (hf : Surjective f) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) : InvolutiveInv M₂ where inv := Inv.inv - inv_inv := hf.forall.2 fun x => by erw [← inv, ← inv, inv_inv] + inv_inv := hf.forall.2 fun x => by rw [← inv, ← inv, inv_inv] variable [Inv M₂] [Div M₂] [Pow M₂ ℤ] @@ -457,14 +457,14 @@ protected abbrev divInvMonoid [DivInvMonoid M₁] (f : M₁ → M₂) (hf : Surj (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : DivInvMonoid M₂ := { hf.monoid f one mul npow, ‹Div M₂›, ‹Inv M₂› with zpow := fun n x => x ^ n, - zpow_zero' := hf.forall.2 fun x => by dsimp only; erw [← zpow, zpow_zero, ← one], + zpow_zero' := hf.forall.2 fun x => by dsimp only; rw [← zpow, zpow_zero, ← one], zpow_succ' := fun n => hf.forall.2 fun x => by dsimp only - erw [← zpow, ← zpow, zpow_natCast, zpow_natCast, pow_succ, ← mul], + rw [← zpow, ← zpow, zpow_natCast, zpow_natCast, pow_succ, ← mul], zpow_neg' := fun n => hf.forall.2 fun x => by dsimp only - erw [← zpow, ← zpow, zpow_negSucc, zpow_natCast, inv], - div_eq_mul_inv := hf.forall₂.2 fun x y => by erw [← inv, ← mul, ← div, div_eq_mul_inv] } + rw [← zpow, ← zpow, zpow_negSucc, zpow_natCast, inv], + div_eq_mul_inv := hf.forall₂.2 fun x y => by rw [← inv, ← mul, ← div, div_eq_mul_inv] } /-- A type endowed with `1`, `*` and `⁻¹` is a group, if it admits a surjective map that preserves `1`, `*` and `⁻¹` to a group. See note [reducible non-instances]. -/ @@ -476,7 +476,7 @@ protected abbrev group [Group M₁] (f : M₁ → M₂) (hf : Surjective f) (one (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : Group M₂ := { hf.divInvMonoid f one mul inv div npow zpow with - inv_mul_cancel := hf.forall.2 fun x => by erw [← inv, ← mul, inv_mul_cancel, one] } + inv_mul_cancel := hf.forall.2 fun x => by rw [← inv, ← mul, inv_mul_cancel, one] } /-- A type endowed with `0`, `1`, `+` is an additive group with one, if it admits a surjective map that preserves `0`, `1`, and `+` to an additive group with one. diff --git a/Mathlib/Algebra/Group/Int.lean b/Mathlib/Algebra/Group/Int.lean index 720440a2282fb..e30cfc2cfe4b0 100644 --- a/Mathlib/Algebra/Group/Int.lean +++ b/Mathlib/Algebra/Group/Int.lean @@ -47,7 +47,7 @@ instance instAddCommGroup : AddCommGroup ℤ where zsmul := (·*·) zsmul_zero' := Int.zero_mul zsmul_succ' m n := by - simp only [ofNat_eq_coe, ofNat_succ, Int.add_mul, Int.add_comm, Int.one_mul] + simp only [ofNat_succ, Int.add_mul, Int.add_comm, Int.one_mul] zsmul_neg' m n := by simp only [negSucc_coe, ofNat_succ, Int.neg_mul] sub_eq_add_neg _ _ := Int.sub_eq_add_neg diff --git a/Mathlib/Algebra/Group/Opposite.lean b/Mathlib/Algebra/Group/Opposite.lean index d7ea15a86adfe..a66f4838ff8fd 100644 --- a/Mathlib/Algebra/Group/Opposite.lean +++ b/Mathlib/Algebra/Group/Opposite.lean @@ -161,7 +161,6 @@ instance instDivInvMonoid [DivInvMonoid α] : DivInvMonoid αᵐᵒᵖ where zpow n a := op <| a.unop ^ n zpow_zero' _ := unop_injective <| zpow_zero _ zpow_succ' _ _ := unop_injective <| by - simp only [Int.ofNat_eq_coe] rw [unop_op, zpow_natCast, pow_succ', unop_mul, unop_op, zpow_natCast] zpow_neg' _ _ := unop_injective <| DivInvMonoid.zpow_neg' _ _ diff --git a/Mathlib/Algebra/Group/Pi/Basic.lean b/Mathlib/Algebra/Group/Pi/Basic.lean index 480759b083256..7a88ae8220a12 100644 --- a/Mathlib/Algebra/Group/Pi/Basic.lean +++ b/Mathlib/Algebra/Group/Pi/Basic.lean @@ -400,18 +400,6 @@ theorem extend_div [Div γ] (f : α → β) (g₁ g₂ : α → γ) (e₁ e₂ : end Extend -theorem surjective_pi_map {F : ∀ i, f i → g i} (hF : ∀ i, Surjective (F i)) : - Surjective fun x : ∀ i, f i => fun i => F i (x i) := fun y => - ⟨fun i => (hF i (y i)).choose, funext fun i => (hF i (y i)).choose_spec⟩ - -theorem injective_pi_map {F : ∀ i, f i → g i} (hF : ∀ i, Injective (F i)) : - Injective fun x : ∀ i, f i => fun i => F i (x i) := - fun _ _ h => funext fun i => hF i <| (congr_fun h i : _) - -theorem bijective_pi_map {F : ∀ i, f i → g i} (hF : ∀ i, Bijective (F i)) : - Bijective fun x : ∀ i, f i => fun i => F i (x i) := - ⟨injective_pi_map fun i => (hF i).injective, surjective_pi_map fun i => (hF i).surjective⟩ - lemma comp_eq_const_iff (b : β) (f : α → β) {g : β → γ} (hg : Injective g) : g ∘ f = Function.const _ (g b) ↔ f = Function.const _ b := hg.comp_left.eq_iff' rfl diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 521a3b832afdc..5e5ebeefca3c2 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -420,7 +420,7 @@ theorem union_mul_inter_subset_union : (s₁ ∪ s₂) * (t₁ ∩ t₂) ⊆ s `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' + t'`."] theorem subset_mul {s t : Set α} : ↑u ⊆ s * t → ∃ s' t' : Finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' * t' := - subset_image₂ + subset_set_image₂ @[to_additive] theorem image_mul [DecidableEq β] : (s * t).image (f : α → β) = s.image f * t.image f := @@ -606,7 +606,7 @@ theorem union_div_inter_subset_union : (s₁ ∪ s₂) / (t₁ ∩ t₂) ⊆ s `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' - t'`."] theorem subset_div {s t : Set α} : ↑u ⊆ s / t → ∃ s' t' : Finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' / t' := - subset_image₂ + subset_set_image₂ @[to_additive (attr := simp (default + 1))] lemma sup_div_le [SemilatticeSup β] [OrderBot β] {s t : Finset α} {f : α → β} {a : β} : @@ -1200,7 +1200,7 @@ theorem union_smul_inter_subset_union [DecidableEq α] : (s₁ ∪ s₂) • (t finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' +ᵥ t'`."] theorem subset_smul {s : Set α} {t : Set β} : ↑u ⊆ s • t → ∃ (s' : Finset α) (t' : Finset β), ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' • t' := - subset_image₂ + subset_set_image₂ end SMul @@ -1310,7 +1310,7 @@ end finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' -ᵥ t'`. -/ theorem subset_vsub {s t : Set β} : ↑u ⊆ s -ᵥ t → ∃ s' t' : Finset β, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' -ᵥ t' := - subset_image₂ + subset_set_image₂ end VSub diff --git a/Mathlib/Algebra/Group/Prod.lean b/Mathlib/Algebra/Group/Prod.lean index 5870914629f9a..d7cb9b0f80a16 100644 --- a/Mathlib/Algebra/Group/Prod.lean +++ b/Mathlib/Algebra/Group/Prod.lean @@ -34,7 +34,7 @@ assert_not_exists MonoidWithZero -- assert_not_exists AddMonoidWithOne assert_not_exists DenselyOrdered -variable {A : Type*} {B : Type*} {G : Type*} {H : Type*} {M : Type*} {N : Type*} {P : Type*} +variable {G : Type*} {H : Type*} {M : Type*} {N : Type*} {P : Type*} namespace Prod diff --git a/Mathlib/Algebra/Group/Semiconj/Defs.lean b/Mathlib/Algebra/Group/Semiconj/Defs.lean index 88ad2b00a8933..cfeabeb0d12db 100644 --- a/Mathlib/Algebra/Group/Semiconj/Defs.lean +++ b/Mathlib/Algebra/Group/Semiconj/Defs.lean @@ -115,7 +115,7 @@ end Monoid section Group -variable [Group G] {a x y : G} +variable [Group G] /-- `a` semiconjugates `x` to `a * x * a⁻¹`. -/ @[to_additive "`a` semiconjugates `x` to `a + x + -a`."] diff --git a/Mathlib/Algebra/Group/Subgroup/Order.lean b/Mathlib/Algebra/Group/Subgroup/Order.lean index 0578e6adbde06..7e9cd2eefee49 100644 --- a/Mathlib/Algebra/Group/Subgroup/Order.lean +++ b/Mathlib/Algebra/Group/Subgroup/Order.lean @@ -21,7 +21,7 @@ theorem mabs_mem_iff {S G} [Group G] [LinearOrder G] {_ : SetLike S G} section ModularLattice -variable {C : Type*} [CommGroup C] {s t : Subgroup C} {x : C} +variable {C : Type*} [CommGroup C] @[to_additive] instance : IsModularLattice (Subgroup C) := diff --git a/Mathlib/Algebra/Group/Submonoid/Basic.lean b/Mathlib/Algebra/Group/Submonoid/Basic.lean index ddb7f8577ed46..623c641fbcfa6 100644 --- a/Mathlib/Algebra/Group/Submonoid/Basic.lean +++ b/Mathlib/Algebra/Group/Submonoid/Basic.lean @@ -65,14 +65,14 @@ variable [MulOneClass M] {s : Set M} variable [AddZeroClass A] {t : Set A} /-- `OneMemClass S M` says `S` is a type of subsets `s ≤ M`, such that `1 ∈ s` for all `s`. -/ -class OneMemClass (S : Type*) (M : Type*) [One M] [SetLike S M] : Prop where +class OneMemClass (S : Type*) (M : outParam Type*) [One M] [SetLike S M] : Prop where /-- By definition, if we have `OneMemClass S M`, we have `1 ∈ s` for all `s : S`. -/ one_mem : ∀ s : S, (1 : M) ∈ s export OneMemClass (one_mem) /-- `ZeroMemClass S M` says `S` is a type of subsets `s ≤ M`, such that `0 ∈ s` for all `s`. -/ -class ZeroMemClass (S : Type*) (M : Type*) [Zero M] [SetLike S M] : Prop where +class ZeroMemClass (S : Type*) (M : outParam Type*) [Zero M] [SetLike S M] : Prop where /-- By definition, if we have `ZeroMemClass S M`, we have `0 ∈ s` for all `s : S`. -/ zero_mem : ∀ s : S, (0 : M) ∈ s @@ -96,7 +96,7 @@ add_decl_doc Submonoid.toSubsemigroup /-- `SubmonoidClass S M` says `S` is a type of subsets `s ≤ M` that contain `1` and are closed under `(*)` -/ -class SubmonoidClass (S : Type*) (M : Type*) [MulOneClass M] [SetLike S M] extends +class SubmonoidClass (S : Type*) (M : outParam Type*) [MulOneClass M] [SetLike S M] extends MulMemClass S M, OneMemClass S M : Prop section @@ -115,7 +115,7 @@ add_decl_doc AddSubmonoid.toAddSubsemigroup /-- `AddSubmonoidClass S M` says `S` is a type of subsets `s ≤ M` that contain `0` and are closed under `(+)` -/ -class AddSubmonoidClass (S : Type*) (M : Type*) [AddZeroClass M] [SetLike S M] extends +class AddSubmonoidClass (S : Type*) (M : outParam Type*) [AddZeroClass M] [SetLike S M] extends AddMemClass S M, ZeroMemClass S M : Prop attribute [to_additive] Submonoid SubmonoidClass diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index e1ec1fb022a46..3f5261db6d422 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -440,7 +440,7 @@ abbrev groupPowers {x : M} {n : ℕ} (hpos : 0 < n) (hx : x ^ n = 1) : Group (po ← pow_eq_pow_mod _ hx, pow_mul, pow_mul] zpow_succ' m x := Subtype.ext <| by obtain ⟨_, k, rfl⟩ := x - simp only [← pow_mul, Int.natMod, Int.ofNat_eq_coe, SubmonoidClass.coe_pow, coe_mul] + simp only [← pow_mul, Int.natMod, SubmonoidClass.coe_pow, coe_mul] norm_cast iterate 2 rw [Int.toNat_natCast, mul_comm, pow_mul, ← pow_eq_pow_mod _ hx] rw [← pow_mul _ m, mul_comm, pow_mul, ← pow_succ, ← pow_mul, mul_comm, pow_mul] diff --git a/Mathlib/Algebra/Group/Subsemigroup/Basic.lean b/Mathlib/Algebra/Group/Subsemigroup/Basic.lean index 88cda9b9d5583..c5801a38f6bdb 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Basic.lean @@ -57,14 +57,14 @@ variable [Mul M] {s : Set M} variable [Add A] {t : Set A} /-- `MulMemClass S M` says `S` is a type of sets `s : Set M` that are closed under `(*)` -/ -class MulMemClass (S : Type*) (M : Type*) [Mul M] [SetLike S M] : Prop where +class MulMemClass (S : Type*) (M : outParam Type*) [Mul M] [SetLike S M] : Prop where /-- A substructure satisfying `MulMemClass` is closed under multiplication. -/ mul_mem : ∀ {s : S} {a b : M}, a ∈ s → b ∈ s → a * b ∈ s export MulMemClass (mul_mem) /-- `AddMemClass S M` says `S` is a type of sets `s : Set M` that are closed under `(+)` -/ -class AddMemClass (S : Type*) (M : Type*) [Add M] [SetLike S M] : Prop where +class AddMemClass (S : Type*) (M : outParam Type*) [Add M] [SetLike S M] : Prop where /-- A substructure satisfying `AddMemClass` is closed under addition. -/ add_mem : ∀ {s : S} {a b : M}, a ∈ s → b ∈ s → a + b ∈ s diff --git a/Mathlib/Algebra/Group/Units.lean b/Mathlib/Algebra/Group/Units.lean index da042a1542c09..1c93daf4e3123 100644 --- a/Mathlib/Algebra/Group/Units.lean +++ b/Mathlib/Algebra/Group/Units.lean @@ -379,7 +379,7 @@ theorem Units.val_mkOfMulEqOne [CommMonoid α] {a b : α} (h : a * b = 1) : section Monoid -variable [Monoid α] {a b c : α} +variable [Monoid α] {a : α} /-- Partial division. It is defined when the second argument is invertible, and unlike the division operator diff --git a/Mathlib/Algebra/Group/Units/Equiv.lean b/Mathlib/Algebra/Group/Units/Equiv.lean index a3ae5f7904381..12d6fb5b2285c 100644 --- a/Mathlib/Algebra/Group/Units/Equiv.lean +++ b/Mathlib/Algebra/Group/Units/Equiv.lean @@ -13,7 +13,7 @@ import Mathlib.Algebra.Group.Units.Hom assert_not_exists MonoidWithZero assert_not_exists DenselyOrdered -variable {F α β A B M N P Q G H : Type*} +variable {F α M N G : Type*} /-- A group is isomorphic to its group of units. -/ @[to_additive (attr := simps apply_val symm_apply) @@ -31,7 +31,7 @@ lemma toUnits_val_apply {G : Type*} [Group G] (x : Gˣ) : toUnits (x : G) = x := namespace Units -variable [Monoid M] [Monoid N] [Monoid P] +variable [Monoid M] [Monoid N] /-- A multiplicative equivalence of monoids defines a multiplicative equivalence of their groups of units. -/ @@ -192,11 +192,10 @@ def MulEquiv.inv (G : Type*) [DivisionCommMonoid G] : G ≃* G := theorem MulEquiv.inv_symm (G : Type*) [DivisionCommMonoid G] : (MulEquiv.inv G).symm = MulEquiv.inv G := rfl --- Porting note: no `add_equiv.neg_symm` in `mathlib3` -@[to_additive] -protected -theorem MulEquiv.map_isUnit_iff {M N} [Monoid M] [Monoid N] [EquivLike F M N] [MonoidHomClass F M N] - (f : F) {m : M} : IsUnit (f m) ↔ IsUnit m := - isUnit_map_of_leftInverse (MonoidHom.inverse (f : M →* N) (EquivLike.inv f) - (EquivLike.left_inv f) <| EquivLike.right_inv f) (EquivLike.left_inv f) +instance isLocalRingHom_equiv [Monoid M] [Monoid N] [EquivLike F M N] + [MulEquivClass F M N] (f : F) : IsLocalRingHom f where + map_nonunit a ha := by + convert ha.map (f : M ≃* N).symm + rw [MulEquiv.eq_symm_apply] + rfl -- note to reviewers: ugly `rfl` diff --git a/Mathlib/Algebra/Group/Units/Hom.lean b/Mathlib/Algebra/Group/Units/Hom.lean index 1a84f0441934c..09f626ae87904 100644 --- a/Mathlib/Algebra/Group/Units/Hom.lean +++ b/Mathlib/Algebra/Group/Units/Hom.lean @@ -17,6 +17,18 @@ also contains unrelated results about `Units` that depend on `MonoidHom`. * `Units.map`: Turn a homomorphism from `α` to `β` monoids into a homomorphism from `αˣ` to `βˣ`. * `MonoidHom.toHomUnits`: Turn a homomorphism from a group `α` to `β` into a homomorphism from `α` to `βˣ`. +* `IsLocalRingHom`: A predicate on monoid maps, requiring that it maps nonunits + to nonunits. For local rings, this means that the image of the unique maximal ideal is again + contained in the unique maximal ideal. This is developed earlier, and in the generality of + monoids, as it allows its use in non-local-ring related contexts, but it does have the + strange consequence that it does not require local rings, or even rings. + +## TODO + +The results that don't mention homomorphisms should be proved (earlier?) in a different file and be +used to golf the basic `Group` lemmas. + +Add a `@[to_additive]` version of `IsLocalRingHom`. -/ assert_not_exists MonoidWithZero @@ -152,7 +164,7 @@ end MonoidHom namespace IsUnit -variable {F G α M N : Type*} [FunLike F M N] [FunLike G N M] +variable {F G M N : Type*} [FunLike F M N] [FunLike G N M] section Monoid @@ -167,6 +179,7 @@ theorem of_leftInverse [MonoidHomClass G N M] {f : F} {x : M} (g : G) (hfg : Function.LeftInverse g f) (h : IsUnit (f x)) : IsUnit x := by simpa only [hfg x] using h.map g +/-- Prefer `IsLocalRingHom.of_leftInverse`, but we can't get rid of this because of `ToAdditive`. -/ @[to_additive] theorem _root_.isUnit_map_of_leftInverse [MonoidHomClass F M N] [MonoidHomClass G N M] {f : F} {x : M} (g : G) (hfg : Function.LeftInverse g f) : @@ -194,3 +207,49 @@ theorem liftRight_inv_mul (f : M →* N) (h : ∀ x, IsUnit (f x)) (x) : end Monoid end IsUnit + +section IsLocalRingHom + +variable {G R S T F : Type*} + +variable [Monoid R] [Monoid S] [Monoid T] [FunLike F R S] + +/-- A local ring homomorphism is a map `f` between monoids such that `a` in the domain + is a unit if `f a` is a unit for any `a`. See `LocalRing.local_hom_TFAE` for other equivalent + definitions in the local ring case - from where this concept originates, but it is useful in + other contexts, so we allow this generalisation in mathlib. -/ +class IsLocalRingHom (f : F) : Prop where + /-- A local ring homomorphism `f : R ⟶ S` will send nonunits of `R` to nonunits of `S`. -/ + map_nonunit : ∀ a, IsUnit (f a) → IsUnit a + +@[simp] +theorem IsUnit.of_map (f : F) [IsLocalRingHom f] (a : R) (h : IsUnit (f a)) : IsUnit a := + IsLocalRingHom.map_nonunit a h + +-- TODO : remove alias, change the parenthesis of `f` and `a` +alias isUnit_of_map_unit := IsUnit.of_map + +variable [MonoidHomClass F R S] + +@[simp] +theorem isUnit_map_iff (f : F) [IsLocalRingHom f] (a : R) : IsUnit (f a) ↔ IsUnit a := + ⟨IsLocalRingHom.map_nonunit a, IsUnit.map f⟩ + +theorem isLocalRingHom_of_leftInverse [FunLike G S R] [MonoidHomClass G S R] + {f : F} (g : G) (hfg : Function.LeftInverse g f) : IsLocalRingHom f where + map_nonunit a ha := by rwa [isUnit_map_of_leftInverse g hfg] at ha + +instance MonoidHom.isLocalRingHom_comp (g : S →* T) (f : R →* S) [IsLocalRingHom g] + [IsLocalRingHom f] : IsLocalRingHom (g.comp f) where + map_nonunit a := IsLocalRingHom.map_nonunit a ∘ IsLocalRingHom.map_nonunit (f := g) (f a) + +-- see note [lower instance priority] +instance (priority := 100) isLocalRingHom_toMonoidHom (f : F) [IsLocalRingHom f] : + IsLocalRingHom (f : R →* S) := + ⟨IsLocalRingHom.map_nonunit (f := f)⟩ + +theorem MonoidHom.isLocalRingHom_of_comp (f : R →* S) (g : S →* T) [IsLocalRingHom (g.comp f)] : + IsLocalRingHom f := + ⟨fun _ ha => (isUnit_map_iff (g.comp f) _).mp (ha.map g)⟩ + +end IsLocalRingHom diff --git a/Mathlib/Algebra/Group/WithOne/Defs.lean b/Mathlib/Algebra/Group/WithOne/Defs.lean index cb45380e57fa6..39ed722267c5d 100644 --- a/Mathlib/Algebra/Group/WithOne/Defs.lean +++ b/Mathlib/Algebra/Group/WithOne/Defs.lean @@ -43,7 +43,7 @@ assert_not_exists DenselyOrdered universe u v w -variable {α : Type u} {β : Type v} {γ : Type w} +variable {α : Type u} /-- Add an extra element `1` to a type -/ @[to_additive "Add an extra element `0` to a type"] diff --git a/Mathlib/Algebra/GroupWithZero/Action/Defs.lean b/Mathlib/Algebra/GroupWithZero/Action/Defs.lean index 28a7f1306d857..d30f29f25d6f4 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Defs.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Defs.lean @@ -49,7 +49,7 @@ assert_not_exists Ring open Function -variable {R R' M M' N G A B α β : Type*} +variable {M N A B α β : Type*} /-- `Monoid.toMulAction` is faithful on nontrivial cancellative monoids with zero. -/ instance CancelMonoidWithZero.faithfulSMul [CancelMonoidWithZero α] [Nontrivial α] : diff --git a/Mathlib/Algebra/GroupWithZero/Action/Opposite.lean b/Mathlib/Algebra/GroupWithZero/Action/Opposite.lean index e78af0c456d38..27bab44a1aaed 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Opposite.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Opposite.lean @@ -27,7 +27,7 @@ With `open scoped RightActions`, this provides: * `p <+ᵥ v` as an alias for `AddOpposite.op v +ᵥ p` -/ -variable {R M N α : Type*} +variable {M α : Type*} /-! ### Actions _on_ the opposite type diff --git a/Mathlib/Algebra/GroupWithZero/Action/Prod.lean b/Mathlib/Algebra/GroupWithZero/Action/Prod.lean index b3ca53b145019..b825703247594 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Prod.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Prod.lean @@ -26,7 +26,7 @@ namespace Prod section -variable [SMul M α] [SMul M β] [SMul N α] [SMul N β] (a : M) (x : α × β) +variable [SMul M α] [SMul M β] theorem smul_zero_mk {α : Type*} [Monoid M] [AddMonoid α] [DistribMulAction M α] (a : M) (c : β) : a • ((0 : α), c) = (0, a • c) := by rw [Prod.smul_mk, smul_zero] diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index 26f0af2649201..b6f9678b599f4 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -37,7 +37,7 @@ assert_not_exists DenselyOrdered open Function -variable {α M₀ G₀ M₀' G₀' F F' : Type*} +variable {M₀ G₀ : Type*} section @@ -136,7 +136,7 @@ theorem right_ne_zero_of_mul_eq_one (h : a * b = 1) : b ≠ 0 := end section MonoidWithZero -variable [MonoidWithZero M₀] {a : M₀} {m n : ℕ} +variable [MonoidWithZero M₀] {a : M₀} {n : ℕ} @[simp] lemma zero_pow : ∀ {n : ℕ}, n ≠ 0 → (0 : M₀) ^ n = 0 | n + 1, _ => by rw [pow_succ, mul_zero] @@ -234,7 +234,7 @@ end CancelMonoidWithZero section GroupWithZero -variable [GroupWithZero G₀] {a b c g h x : G₀} +variable [GroupWithZero G₀] {a b x : G₀} theorem GroupWithZero.mul_left_injective (h : x ≠ 0) : Function.Injective fun y => x * y := fun y y' w => by diff --git a/Mathlib/Algebra/GroupWithZero/Defs.lean b/Mathlib/Algebra/GroupWithZero/Defs.lean index b3d73d7ba0a35..14ace092e132d 100644 --- a/Mathlib/Algebra/GroupWithZero/Defs.lean +++ b/Mathlib/Algebra/GroupWithZero/Defs.lean @@ -26,7 +26,7 @@ universe u -- We have to fix the universe of `G₀` here, since the default argument to -- `GroupWithZero.div'` cannot contain a universe metavariable. -variable {G₀ : Type u} {M₀ M₀' G₀' : Type*} +variable {G₀ : Type u} {M₀ : Type*} /-- Typeclass for expressing that a type `M₀` with multiplication and a zero satisfies `0 * a = 0` and `a * 0 = 0` for all `a : M₀`. -/ @@ -159,17 +159,17 @@ class MulDivCancelClass (M₀ : Type*) [MonoidWithZero M₀] [Div M₀] : Prop w protected mul_div_cancel (a b : M₀) : b ≠ 0 → a * b / b = a section MulDivCancelClass -variable [MonoidWithZero M₀] [Div M₀] [MulDivCancelClass M₀] {a b : M₀} +variable [MonoidWithZero M₀] [Div M₀] [MulDivCancelClass M₀] -@[simp] lemma mul_div_cancel_right₀ (a : M₀) (hb : b ≠ 0) : a * b / b = a := +@[simp] lemma mul_div_cancel_right₀ (a : M₀) {b : M₀} (hb : b ≠ 0) : a * b / b = a := MulDivCancelClass.mul_div_cancel _ _ hb end MulDivCancelClass section MulDivCancelClass -variable [CommMonoidWithZero M₀] [Div M₀] [MulDivCancelClass M₀] {a b : M₀} +variable [CommMonoidWithZero M₀] [Div M₀] [MulDivCancelClass M₀] -@[simp] lemma mul_div_cancel_left₀ (b : M₀) (ha : a ≠ 0) : a * b / a = b := by +@[simp] lemma mul_div_cancel_left₀ (b : M₀) {a : M₀} (ha : a ≠ 0) : a * b / a = b := by rw [mul_comm, mul_div_cancel_right₀ _ ha] end MulDivCancelClass @@ -216,7 +216,7 @@ end section GroupWithZero -variable [GroupWithZero G₀] {a b c g h x : G₀} +variable [GroupWithZero G₀] {a b : G₀} @[simp] theorem mul_inv_cancel_right₀ (h : b ≠ 0) (a : G₀) : a * b * b⁻¹ = a := diff --git a/Mathlib/Algebra/GroupWithZero/Indicator.lean b/Mathlib/Algebra/GroupWithZero/Indicator.lean index 2b8dcaded4211..52d9d31a8d50e 100644 --- a/Mathlib/Algebra/GroupWithZero/Indicator.lean +++ b/Mathlib/Algebra/GroupWithZero/Indicator.lean @@ -16,7 +16,7 @@ variable {ι κ G₀ M₀ R : Type*} namespace Set section MulZeroClass -variable [MulZeroClass M₀] {s t : Set ι} {f g : ι → M₀} {i : ι} +variable [MulZeroClass M₀] {s t : Set ι} {i : ι} lemma indicator_mul (s : Set ι) (f g : ι → M₀) : indicator s (fun i ↦ f i * g i) = fun i ↦ indicator s f i * indicator s g i := by @@ -40,6 +40,12 @@ lemma indicator_mul_right (s : Set ι) (f g : ι → M₀) : · rfl · rw [mul_zero] +lemma indicator_mul_const (s : Set ι) (f : ι → M₀) (a : M₀) (i : ι) : + s.indicator (f · * a) i = s.indicator f i * a := by rw [indicator_mul_left] + +lemma indicator_const_mul (s : Set ι) (f : ι → M₀) (a : M₀) (i : ι) : + s.indicator (a * f ·) i = a * s.indicator f i := by rw [indicator_mul_right] + lemma inter_indicator_mul (f g : ι → M₀) (i : ι) : (s ∩ t).indicator (fun j ↦ f j * g j) i = s.indicator f i * t.indicator g i := by rw [← Set.indicator_indicator] diff --git a/Mathlib/Algebra/GroupWithZero/InjSurj.lean b/Mathlib/Algebra/GroupWithZero/InjSurj.lean index 26278384167c1..5372c67cb3f27 100644 --- a/Mathlib/Algebra/GroupWithZero/InjSurj.lean +++ b/Mathlib/Algebra/GroupWithZero/InjSurj.lean @@ -149,7 +149,7 @@ end MonoidWithZero section CancelMonoidWithZero -variable [CancelMonoidWithZero M₀] {a b c : M₀} +variable [CancelMonoidWithZero M₀] /-- Pull back a `CancelMonoidWithZero` along an injective function. See note [reducible non-instances]. -/ @@ -167,7 +167,7 @@ end CancelMonoidWithZero section CancelCommMonoidWithZero -variable [CancelCommMonoidWithZero M₀] {a b c : M₀} +variable [CancelCommMonoidWithZero M₀] /-- Pull back a `CancelCommMonoidWithZero` along an injective function. See note [reducible non-instances]. -/ @@ -181,7 +181,7 @@ end CancelCommMonoidWithZero section GroupWithZero -variable [GroupWithZero G₀] {a b c g h x : G₀} +variable [GroupWithZero G₀] /-- Pull back a `GroupWithZero` along an injective function. See note [reducible non-instances]. -/ @@ -215,7 +215,7 @@ end GroupWithZero section CommGroupWithZero -variable [CommGroupWithZero G₀] {a b c d : G₀} +variable [CommGroupWithZero G₀] /-- Pull back a `CommGroupWithZero` along an injective function. See note [reducible non-instances]. -/ diff --git a/Mathlib/Algebra/GroupWithZero/Pi.lean b/Mathlib/Algebra/GroupWithZero/Pi.lean index 547f6a5013092..f11c9daf124b3 100644 --- a/Mathlib/Algebra/GroupWithZero/Pi.lean +++ b/Mathlib/Algebra/GroupWithZero/Pi.lean @@ -22,7 +22,7 @@ variable {ι : Type*} {α : ι → Type*} namespace Pi section MulZeroClass -variable [∀ i, MulZeroClass (α i)] [DecidableEq ι] {i j : ι} {f : ∀ i, α i} +variable [∀ i, MulZeroClass (α i)] [DecidableEq ι] {i : ι} {f : ∀ i, α i} instance mulZeroClass : MulZeroClass (∀ i, α i) where zero_mul := by intros; ext; exact zero_mul _ diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index a32ad81fb22e8..4daea455977cf 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -23,7 +23,7 @@ We also define `Ring.inverse`, a globally defined function on any ring assert_not_exists Multiplicative assert_not_exists DenselyOrdered -variable {α M₀ G₀ M₀' G₀' F F' : Type*} +variable {α M₀ G₀ : Type*} variable [MonoidWithZero M₀] namespace Units @@ -87,6 +87,12 @@ noncomputable def inverse : M₀ → M₀ := fun x => if h : IsUnit x then ((h.u theorem inverse_unit (u : M₀ˣ) : inverse (u : M₀) = (u⁻¹ : M₀ˣ) := by rw [inverse, dif_pos u.isUnit, IsUnit.unit_of_val_units] +theorem IsUnit.ringInverse {x : M₀} (h : IsUnit x) : IsUnit (inverse x) := + match h with + | ⟨u, hu⟩ => hu ▸ ⟨u⁻¹, (inverse_unit u).symm⟩ + +theorem inverse_of_isUnit {x : M₀} (h : IsUnit x) : inverse x = ((h.unit⁻¹ : M₀ˣ) : M₀) := dif_pos h + /-- By definition, if `x` is not invertible then `inverse x = 0`. -/ @[simp] theorem inverse_non_unit (x : M₀) (h : ¬IsUnit x) : inverse x = 0 := @@ -152,7 +158,6 @@ theorem isUnit_ring_inverse {a : M₀} : IsUnit (Ring.inverse a) ↔ IsUnit a := namespace Units variable [GroupWithZero G₀] -variable {a b : G₀} /-- Embed a non-zero element of a `GroupWithZero` into the unit group. By combining this function with the operations on units, @@ -208,7 +213,7 @@ theorem _root_.GroupWithZero.eq_zero_or_unit (a : G₀) : a = 0 ∨ ∃ u : G₀ end Units section GroupWithZero -variable [GroupWithZero G₀] {a b c d : G₀} {m n : ℕ} +variable [GroupWithZero G₀] {a b c : G₀} {m n : ℕ} theorem IsUnit.mk0 (x : G₀) (hx : x ≠ 0) : IsUnit x := (Units.mk0 x hx).isUnit diff --git a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean index bbf5c0a8dc88f..407a0fb1b6c75 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean @@ -15,10 +15,34 @@ import Mathlib.Algebra.GroupWithZero.Hom assert_not_exists DenselyOrdered -variable {α M₀ G₀ M₀' G₀' F F' : Type*} +variable {M M₀ G₀ M₀' G₀' F F' : Type*} variable [MonoidWithZero M₀] +section Monoid + +variable [Monoid M] [GroupWithZero G₀] + +lemma isLocalRingHom_of_exists_map_ne_one [FunLike F G₀ M] [MonoidHomClass F G₀ M] {f : F} + (hf : ∃ x : G₀, f x ≠ 1) : IsLocalRingHom f where + map_nonunit a h := by + rcases eq_or_ne a 0 with (rfl | h) + · obtain ⟨t, ht⟩ := hf + refine (ht ?_).elim + have := map_mul f t 0 + rw [← one_mul (f (t * 0)), mul_zero] at this + exact (h.mul_right_cancel this).symm + · exact ⟨⟨a, a⁻¹, mul_inv_cancel₀ h, inv_mul_cancel₀ h⟩, rfl⟩ + +instance [GroupWithZero G₀] [FunLike F G₀ M₀] [MonoidWithZeroHomClass F G₀ M₀] [Nontrivial M₀] + (f : F) : IsLocalRingHom f := + isLocalRingHom_of_exists_map_ne_one ⟨0, by simp⟩ + +end Monoid + +section GroupWithZero + namespace Commute + variable [GroupWithZero G₀] {a b c d : G₀} /-- The `MonoidWithZero` version of `div_eq_div_iff_mul_eq_mul`. -/ @@ -107,3 +131,5 @@ end Units theorem map_zpow₀ {F G₀ G₀' : Type*} [GroupWithZero G₀] [GroupWithZero G₀'] [FunLike F G₀ G₀'] [MonoidWithZeroHomClass F G₀ G₀'] (f : F) (x : G₀) (n : ℤ) : f (x ^ n) = f x ^ n := map_zpow' f (map_inv₀ f) x n + +end GroupWithZero diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean index b49421f3f25b8..85cb5a4752ba7 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean @@ -61,7 +61,7 @@ noncomputable def hom : (descCochain _ 0 (Cochain.ofHom (inr (f ≫ g))) (neg_add_cancel 1)) (by ext p _ rfl simp [mappingConeCompTriangle, map, ext_from_iff _ _ _ rfl, - inl_v_d_assoc _ (p+1) p (p+2) (by linarith) (by linarith)]) + inl_v_d_assoc _ (p+1) p (p+2) (by omega) (by omega)]) /-- Given two composable morphisms `f` and `g` in the category of cochain complexes, this is the canonical morphism (which is an homotopy equivalence) from the mapping cone of @@ -72,7 +72,7 @@ noncomputable def inv : mappingCone (mappingConeCompTriangle f g).mor₁ ⟶ map ext p rw [ext_from_iff _ (p + 1) _ rfl, ext_to_iff _ _ (p + 1) rfl] simp [map, δ_zero_cochain_comp, - Cochain.comp_v _ _ (add_neg_cancel 1) p (p+1) p (by linarith) (by linarith)]) + Cochain.comp_v _ _ (add_neg_cancel 1) p (p+1) p (by omega) (by omega)]) @[reassoc (attr := simp)] lemma hom_inv_id : hom f g ≫ inv f g = 𝟙 _ := by @@ -86,44 +86,25 @@ this is the `homotopyInvHomId` field of the homotopy equivalence the morphism `mappingCone f ⟶ mappingCone (f ≫ g)`. -/ noncomputable def homotopyInvHomId : Homotopy (inv f g ≫ hom f g) (𝟙 _) := (Cochain.equivHomotopy _ _).symm ⟨-((snd _).comp ((fst (f ≫ g)).1.comp - ((inl f).comp (inl _) (by linarith)) (show 1 + (-2) = -1 by linarith)) (zero_add (-1))), by + ((inl f).comp (inl _) (by omega)) (show 1 + (-2) = -1 by omega)) (zero_add (-1))), by rw [δ_neg, δ_zero_cochain_comp _ _ _ (neg_add_cancel 1), Int.negOnePow_neg, Int.negOnePow_one, Units.neg_smul, one_smul, - δ_comp _ _ (show 1 + (-2) = -1 by linarith) 2 (-1) 0 (by linarith) - (by linarith) (by linarith), - δ_comp _ _ (show (-1) + (-1) = -2 by linarith) 0 0 (-1) (by linarith) - (by linarith) (by linarith), Int.negOnePow_neg, Int.negOnePow_neg, - Int.negOnePow_even 2 ⟨1, by linarith⟩, Int.negOnePow_one, Units.neg_smul, + δ_comp _ _ (show 1 + (-2) = -1 by omega) 2 (-1) 0 (by omega) + (by omega) (by omega), + δ_comp _ _ (show (-1) + (-1) = -2 by omega) 0 0 (-1) (by omega) + (by omega) (by omega), Int.negOnePow_neg, Int.negOnePow_neg, + Int.negOnePow_even 2 ⟨1, by omega⟩, Int.negOnePow_one, Units.neg_smul, one_smul, one_smul, δ_inl, δ_inl, δ_snd, Cocycle.δ_eq_zero, Cochain.zero_comp, add_zero, Cochain.neg_comp, neg_neg] ext n rw [ext_from_iff _ (n + 1) n rfl, ext_from_iff _ (n + 1) n rfl, - ext_from_iff _ (n + 2) (n + 1) (by linarith)] - simp? [hom, inv, ext_to_iff _ n (n + 1) rfl, map, Cochain.comp_v _ _ - (add_neg_cancel 1) n (n + 1) n (by linarith) (by linarith), - Cochain.comp_v _ _ (show 1 + -2 = -1 by linarith) (n + 1) (n + 2) n - (by linarith) (by linarith), - Cochain.comp_v _ _ (show (-1) + -1 = -2 by linarith) (n + 2) (n + 1) n - (by linarith) (by linarith)] says - simp only [mappingConeCompTriangle_obj₁, mappingConeCompTriangle_obj₂, - mappingConeCompTriangle_mor₁, map, Int.reduceNeg, inv, hom, Cochain.ofHom_comp, - ofHom_desc, ofHom_lift, descCocycle_coe, AddSubmonoid.coe_zero, - Cochain.comp_zero_cochain_v, inl_v_descCochain_v_assoc, Cochain.zero_cochain_comp_v, - assoc, inl_v_snd_v_assoc, zero_comp, Cochain.id_comp, - Cochain.comp_assoc_of_first_is_zero_cochain, Cochain.comp_add, Cochain.comp_neg, - Cochain.comp_assoc_of_second_is_zero_cochain, neg_add_rev, neg_neg, Cochain.add_v, - Cochain.neg_v, - Cochain.comp_v _ _ (add_neg_cancel 1) n (n + 1) n (by linarith) (by linarith), - Cochain.comp_v _ _ (show 1 + -2 = -1 by linarith) (n + 1) (n + 2) n (by linarith) - (by linarith), - Cochain.comp_v _ _ (show (-1) + -1 = -2 by linarith) (n + 2) (n + 1) n (by linarith) - (by linarith), - Cochain.ofHom_v, HomologicalComplex.id_f, Preadditive.comp_add, Preadditive.comp_neg, - inl_v_fst_v_assoc, neg_zero, add_zero, comp_id, neg_add_cancel, inr_f_snd_v_assoc, - inr_f_descCochain_v_assoc, inr_f_fst_v_assoc, comp_zero, zero_add, - ext_to_iff _ n (n + 1) rfl, liftCochain_v_fst_v, inl_v_descCochain_v, inl_v_fst_v, - liftCochain_v_snd_v, Cochain.zero_v, inl_v_snd_v, and_self, neg_add_cancel_right, - inr_f_descCochain_v, inr_f_fst_v, inr_f_snd_v]⟩ + ext_from_iff _ (n + 2) (n + 1) (by omega)] + simp [hom, inv, ext_to_iff _ n (n + 1) rfl, map, Cochain.comp_v _ _ + (add_neg_cancel 1) n (n + 1) n (by omega) (by omega), + Cochain.comp_v _ _ (show 1 + -2 = -1 by omega) (n + 1) (n + 2) n + (by omega) (by omega), + Cochain.comp_v _ _ (show (-1) + -1 = -2 by omega) (n + 2) (n + 1) n + (by omega) (by omega)]⟩ end MappingConeCompHomotopyEquiv diff --git a/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean b/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean index 7523eadd214dd..59034f288455b 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean @@ -33,7 +33,7 @@ linear maps `f` and `g` and the vanishing of their composition. -/ def moduleCatMk {X₁ X₂ X₃ : Type v} [AddCommGroup X₁] [AddCommGroup X₂] [AddCommGroup X₃] [Module R X₁] [Module R X₂] [Module R X₃] (f : X₁ →ₗ[R] X₂) (g : X₂ →ₗ[R] X₃) (hfg : g.comp f = 0) : ShortComplex (ModuleCat.{v} R) := - ShortComplex.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) hfg + ShortComplex.mk (ModuleCat.asHom f) (ModuleCat.asHom g) hfg variable (S : ShortComplex (ModuleCat.{v} R)) @@ -138,7 +138,7 @@ def moduleCatLeftHomologyData : S.LeftHomologyData where erw [Submodule.Quotient.mk_eq_zero] rw [LinearMap.mem_range] apply exists_apply_eq_apply - hπ := ModuleCat.cokernelIsColimit (ModuleCat.ofHom S.moduleCatToCycles) + hπ := ModuleCat.cokernelIsColimit (ModuleCat.asHom S.moduleCatToCycles) @[simp] lemma moduleCatLeftHomologyData_f' : diff --git a/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean b/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean index 77769d32f3b80..8fe3dbac06932 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean @@ -381,7 +381,7 @@ variable (S₁ S₂ S₃ : SnakeInput C) /-- A morphism of snake inputs involve four morphisms of short complexes which make the obvious diagram commute. -/ @[ext] -structure Hom := +structure Hom where /-- a morphism between the zeroth lines -/ f₀ : S₁.L₀ ⟶ S₂.L₀ /-- a morphism between the first lines -/ diff --git a/Mathlib/Algebra/Lie/Basic.lean b/Mathlib/Algebra/Lie/Basic.lean index 40207bffb861a..cb37eac2841ca 100644 --- a/Mathlib/Algebra/Lie/Basic.lean +++ b/Mathlib/Algebra/Lie/Basic.lean @@ -269,10 +269,10 @@ attribute [coe] LieHom.toLinearMap instance : Coe (L₁ →ₗ⁅R⁆ L₂) (L₁ →ₗ[R] L₂) := ⟨LieHom.toLinearMap⟩ -instance : FunLike (L₁ →ₗ⁅R⁆ L₂) L₁ L₂ := - { coe := fun f => f.toFun, - coe_injective' := fun x y h => - by cases x; cases y; simp at h; simp [h] } +instance : FunLike (L₁ →ₗ⁅R⁆ L₂) L₁ L₂ where + coe f := f.toFun + coe_injective' x y h := by + cases x; cases y; simp at h; simp [h] initialize_simps_projections LieHom (toFun → apply) @@ -471,13 +471,12 @@ instance hasCoeToLieHom : Coe (L₁ ≃ₗ⁅R⁆ L₂) (L₁ →ₗ⁅R⁆ L₂ instance hasCoeToLinearEquiv : Coe (L₁ ≃ₗ⁅R⁆ L₂) (L₁ ≃ₗ[R] L₂) := ⟨toLinearEquiv⟩ -instance : EquivLike (L₁ ≃ₗ⁅R⁆ L₂) L₁ L₂ := - { coe := fun f => f.toFun, - inv := fun f => f.invFun, - left_inv := fun f => f.left_inv, - right_inv := fun f => f.right_inv, - coe_injective' := fun f g h₁ h₂ => - by cases f; cases g; simp at h₁ h₂; simp [*] } +instance : EquivLike (L₁ ≃ₗ⁅R⁆ L₂) L₁ L₂ where + coe f := f.toFun + inv f := f.invFun + left_inv f := f.left_inv + right_inv f := f.right_inv + coe_injective' f g h₁ h₂ := by cases f; cases g; simp at h₁ h₂; simp [*] theorem coe_to_lieHom (e : L₁ ≃ₗ⁅R⁆ L₂) : ⇑(e : L₁ →ₗ⁅R⁆ L₂) = e := rfl @@ -622,10 +621,9 @@ attribute [coe] LieModuleHom.toLinearMap instance : CoeOut (M →ₗ⁅R,L⁆ N) (M →ₗ[R] N) := ⟨LieModuleHom.toLinearMap⟩ -instance : FunLike (M →ₗ⁅R, L⁆ N) M N := - { coe := fun f => f.toFun, - coe_injective' := fun x y h => - by cases x; cases y; simp at h; simp [h] } +instance : FunLike (M →ₗ⁅R, L⁆ N) M N where + coe f := f.toFun + coe_injective' x y h := by cases x; cases y; simp at h; simp [h] initialize_simps_projections LieModuleHom (toFun → apply) @@ -855,13 +853,12 @@ instance hasCoeToLieModuleHom : Coe (M ≃ₗ⁅R,L⁆ N) (M →ₗ⁅R,L⁆ N) instance hasCoeToLinearEquiv : CoeOut (M ≃ₗ⁅R,L⁆ N) (M ≃ₗ[R] N) := ⟨toLinearEquiv⟩ -instance : EquivLike (M ≃ₗ⁅R,L⁆ N) M N := - { coe := fun f => f.toFun, - inv := fun f => f.invFun, - left_inv := fun f => f.left_inv, - right_inv := fun f => f.right_inv, - coe_injective' := fun f g h₁ h₂ => - by cases f; cases g; simp at h₁ h₂; simp [*] } +instance : EquivLike (M ≃ₗ⁅R,L⁆ N) M N where + coe f := f.toFun + inv f := f.invFun + left_inv f := f.left_inv + right_inv f := f.right_inv + coe_injective' f g h₁ h₂ := by cases f; cases g; simp at h₁ h₂; simp [*] @[simp] lemma coe_coe (e : M ≃ₗ⁅R,L⁆ N) : ⇑(e : M →ₗ⁅R,L⁆ N) = e := rfl diff --git a/Mathlib/Algebra/Lie/Killing.lean b/Mathlib/Algebra/Lie/Killing.lean index ba952dab65ff0..4efe977e0587c 100644 --- a/Mathlib/Algebra/Lie/Killing.lean +++ b/Mathlib/Algebra/Lie/Killing.lean @@ -46,7 +46,7 @@ namespace LieAlgebra NB: This is not standard terminology (the literature does not seem to name Lie algebras with this property). -/ -class IsKilling : Prop := +class IsKilling : Prop where /-- We say a Lie algebra is Killing if its Killing form is non-singular. -/ killingCompl_top_eq_bot : LieIdeal.killingCompl R L ⊤ = ⊥ diff --git a/Mathlib/Algebra/Lie/Solvable.lean b/Mathlib/Algebra/Lie/Solvable.lean index f51e86e77306c..1ce5fd8fdcf58 100644 --- a/Mathlib/Algebra/Lie/Solvable.lean +++ b/Mathlib/Algebra/Lie/Solvable.lean @@ -266,8 +266,8 @@ def radical := /-- The radical of a Noetherian Lie algebra is solvable. -/ instance radicalIsSolvable [IsNoetherian R L] : IsSolvable R (radical R L) := by - have hwf := (LieSubmodule.wellFoundedGT_of_noetherian R L L).wf - rw [← CompleteLattice.isSupClosedCompact_iff_wellFounded] at hwf + have hwf := LieSubmodule.wellFoundedGT_of_noetherian R L L + rw [← CompleteLattice.isSupClosedCompact_iff_wellFoundedGT] at hwf refine hwf { I : LieIdeal R L | IsSolvable R I } ⟨⊥, ?_⟩ fun I hI J hJ => ?_ · exact LieAlgebra.isSolvableBot R L · rw [Set.mem_setOf_eq] at hI hJ ⊢ diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index 6344226a618c7..bddebc215744c 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -226,9 +226,9 @@ lemma traceForm_eq_sum_genWeightSpaceOf convert finite_genWeightSpaceOf_ne_bot R L M z exact LieSubmodule.coeSubmodule_eq_bot_iff (genWeightSpaceOf M _ _) classical - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_genWeightSpaceOf R L M z) - (IsTriangularizable.iSup_eq_top z) + have h := LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_genWeightSpaceOf R L M z + have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top h <| by + simp [← LieSubmodule.iSup_coe_toSubmodule] simp only [LinearMap.coeFn_sum, Finset.sum_apply, traceForm_apply_apply, LinearMap.trace_eq_sum_trace_restrict' hds hfin hxy] exact Finset.sum_congr (by simp) (fun χ _ ↦ rfl) @@ -276,9 +276,9 @@ lemma lowerCentralSeries_one_inf_center_le_ker_traceForm [Module.Free R M] [Modu intro y exact y.induction_on rfl (fun a u ↦ by simp [hzc u]) (fun u v hu hv ↦ by simp [hu, hv]) apply LinearMap.trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero - · exact IsTriangularizable.iSup_eq_top (1 ⊗ₜ[R] x) + · simpa only [Module.End.maxGenEigenspace_def] using IsTriangularizable.iSup_eq_top (1 ⊗ₜ[R] x) · exact fun μ ↦ trace_toEnd_eq_zero_of_mem_lcs A (A ⊗[R] L) - (genWeightSpaceOf (A ⊗[R] M) μ (1 ⊗ₜ x)) (le_refl 1) hz + (genWeightSpaceOf (A ⊗[R] M) μ ((1:A) ⊗ₜ[R] x)) (le_refl 1) hz · exact commute_toEnd_of_mem_center_right (A ⊗[R] M) hzc (1 ⊗ₜ x) /-- A nilpotent Lie algebra with a representation whose trace form is non-singular is Abelian. -/ diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index d991a7a0c5937..e0f030da4cacc 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -707,13 +707,11 @@ lemma independent_genWeightSpaceOf [NoZeroSMulDivisors R M] (x : L) : lemma finite_genWeightSpaceOf_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] (x : L) : {χ : R | genWeightSpaceOf M χ x ≠ ⊥}.Finite := - CompleteLattice.WellFounded.finite_ne_bot_of_independent - IsWellFounded.wf (independent_genWeightSpaceOf R L M x) + CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent (independent_genWeightSpaceOf R L M x) lemma finite_genWeightSpace_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] : {χ : L → R | genWeightSpace M χ ≠ ⊥}.Finite := - CompleteLattice.WellFounded.finite_ne_bot_of_independent - IsWellFounded.wf (independent_genWeightSpace R L M) + CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent (independent_genWeightSpace R L M) instance Weight.instFinite [NoZeroSMulDivisors R M] [IsNoetherian R M] : Finite (Weight R L M) := by @@ -726,7 +724,7 @@ noncomputable instance Weight.instFintype [NoZeroSMulDivisors R M] [IsNoetherian /-- A Lie module `M` of a Lie algebra `L` is triangularizable if the endomorhpism of `M` defined by any `x : L` is triangularizable. -/ -class IsTriangularizable : Prop := +class IsTriangularizable : Prop where iSup_eq_top : ∀ x, ⨆ φ, ⨆ k, (toEnd R L M x).genEigenspace φ k = ⊤ instance (L' : LieSubalgebra R L) [IsTriangularizable R L M] : IsTriangularizable R L' M where @@ -782,8 +780,10 @@ lemma iSup_genWeightSpace_eq_top [IsTriangularizable K L M] : ⨆ χ : L → K, genWeightSpace M χ = ⊤ := by simp only [← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.iSup_coe_toSubmodule, LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.top_coeSubmodule, genWeightSpace] - exact Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo (toEnd K L M) - (fun x y φ z ↦ (genWeightSpaceOf M φ y).lie_mem) (IsTriangularizable.iSup_eq_top) + refine Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo (toEnd K L M) + (fun x y φ z ↦ (genWeightSpaceOf M φ y).lie_mem) ?_ + simp_rw [Module.End.maxGenEigenspace_def] + apply IsTriangularizable.iSup_eq_top lemma iSup_genWeightSpace_eq_top' [IsTriangularizable K L M] : ⨆ χ : Weight K L M, genWeightSpace M χ = ⊤ := by diff --git a/Mathlib/Algebra/Lie/Weights/Linear.lean b/Mathlib/Algebra/Lie/Weights/Linear.lean index 9798876d617f9..741b68fdabb92 100644 --- a/Mathlib/Algebra/Lie/Weights/Linear.lean +++ b/Mathlib/Algebra/Lie/Weights/Linear.lean @@ -48,7 +48,7 @@ namespace LieModule /-- A typeclass encoding the fact that a given Lie module has linear weights, vanishing on the derived ideal. -/ -class LinearWeights [LieAlgebra.IsNilpotent R L] : Prop := +class LinearWeights [LieAlgebra.IsNilpotent R L] : Prop where map_add : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ x y, χ (x + y) = χ x + χ y map_smul : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ (t : R) x, χ (t • x) = t • χ x map_lie : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ x y : L, χ ⁅x, y⁆ = 0 diff --git a/Mathlib/Algebra/Module/Injective.lean b/Mathlib/Algebra/Module/Injective.lean index 9462f04c45e08..a2e077efec5f7 100644 --- a/Mathlib/Algebra/Module/Injective.lean +++ b/Mathlib/Algebra/Module/Injective.lean @@ -67,8 +67,8 @@ theorem Module.injective_module_of_injective_object [inj : CategoryTheory.Injective <| ModuleCat.of R Q] : Module.Injective R Q where out X Y _ _ _ _ f hf g := by - have : CategoryTheory.Mono (ModuleCat.ofHom f) := (ModuleCat.mono_iff_injective _).mpr hf - obtain ⟨l, rfl⟩ := inj.factors (ModuleCat.ofHom g) (ModuleCat.ofHom f) + have : CategoryTheory.Mono (ModuleCat.asHom f) := (ModuleCat.mono_iff_injective _).mpr hf + obtain ⟨l, rfl⟩ := inj.factors (ModuleCat.asHom g) (ModuleCat.asHom f) exact ⟨l, fun _ ↦ rfl⟩ theorem Module.injective_iff_injective_object : diff --git a/Mathlib/Algebra/Module/LinearMap/Basic.lean b/Mathlib/Algebra/Module/LinearMap/Basic.lean index f50429fa64814..1a68ecbd89506 100644 --- a/Mathlib/Algebra/Module/LinearMap/Basic.lean +++ b/Mathlib/Algebra/Module/LinearMap/Basic.lean @@ -21,41 +21,41 @@ open Function universe u u' v w x y z -variable {R R₁ R₂ R₃ k S S₃ T M M₁ M₂ M₃ N₁ N₂ N₃ ι : Type*} +variable {R R' S M M' : Type*} namespace LinearMap section SMul -variable [Semiring R] [Semiring R₂] -variable [AddCommMonoid M] [AddCommMonoid M₂] -variable [Module R M] [Module R₂ M₂] -variable {σ₁₂ : R →+* R₂} +variable [Semiring R] [Semiring R'] +variable [AddCommMonoid M] [AddCommMonoid M'] +variable [Module R M] [Module R' M'] +variable {σ₁₂ : R →+* R'} variable {S' T' : Type*} variable [Monoid S'] [DistribMulAction S' M] [SMulCommClass R S' M] variable [Monoid T'] [DistribMulAction T' M] [SMulCommClass R T' M] -instance : SMul S'ᵈᵐᵃ (M →ₛₗ[σ₁₂] M₂) where +instance : SMul S'ᵈᵐᵃ (M →ₛₗ[σ₁₂] M') where smul a f := - { toFun := a • (f : M → M₂) + { toFun := a • (f : M → M') map_add' := fun x y ↦ by simp only [DomMulAct.smul_apply, f.map_add, smul_add] map_smul' := fun c x ↦ by simp_rw [DomMulAct.smul_apply, ← smul_comm, f.map_smulₛₗ] } -theorem _root_.DomMulAct.smul_linearMap_apply (a : S'ᵈᵐᵃ) (f : M →ₛₗ[σ₁₂] M₂) (x : M) : +theorem _root_.DomMulAct.smul_linearMap_apply (a : S'ᵈᵐᵃ) (f : M →ₛₗ[σ₁₂] M') (x : M) : (a • f) x = f (DomMulAct.mk.symm a • x) := rfl @[simp] -theorem _root_.DomMulAct.mk_smul_linearMap_apply (a : S') (f : M →ₛₗ[σ₁₂] M₂) (x : M) : +theorem _root_.DomMulAct.mk_smul_linearMap_apply (a : S') (f : M →ₛₗ[σ₁₂] M') (x : M) : (DomMulAct.mk a • f) x = f (a • x) := rfl -theorem _root_.DomMulAct.coe_smul_linearMap (a : S'ᵈᵐᵃ) (f : M →ₛₗ[σ₁₂] M₂) : - (a • f : M →ₛₗ[σ₁₂] M₂) = a • (f : M → M₂) := +theorem _root_.DomMulAct.coe_smul_linearMap (a : S'ᵈᵐᵃ) (f : M →ₛₗ[σ₁₂] M') : + (a • f : M →ₛₗ[σ₁₂] M') = a • (f : M → M') := rfl -instance [SMulCommClass S' T' M] : SMulCommClass S'ᵈᵐᵃ T'ᵈᵐᵃ (M →ₛₗ[σ₁₂] M₂) := +instance [SMulCommClass S' T' M] : SMulCommClass S'ᵈᵐᵃ T'ᵈᵐᵃ (M →ₛₗ[σ₁₂] M') := ⟨fun s t f ↦ ext fun m ↦ by simp_rw [DomMulAct.smul_linearMap_apply, smul_comm]⟩ end SMul @@ -63,19 +63,15 @@ end SMul section Actions -variable [Semiring R] [Semiring R₂] [Semiring R₃] -variable [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] -variable [Module R M] [Module R₂ M₂] [Module R₃ M₃] -variable {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] +variable [Semiring R] [Semiring R'] +variable [AddCommMonoid M] [AddCommMonoid M'] +variable [Module R M] [Module R' M'] +variable {σ₁₂ : R →+* R'} section SMul -variable [Monoid S] [DistribMulAction S M₂] [SMulCommClass R₂ S M₂] -variable [Monoid S₃] [DistribMulAction S₃ M₃] [SMulCommClass R₃ S₃ M₃] -variable [Monoid T] [DistribMulAction T M₂] [SMulCommClass R₂ T M₂] - instance {S'} [Monoid S'] [DistribMulAction S' M] [SMulCommClass R S' M] : - DistribMulAction S'ᵈᵐᵃ (M →ₛₗ[σ₁₂] M₂) where + DistribMulAction S'ᵈᵐᵃ (M →ₛₗ[σ₁₂] M') where one_smul _ := ext fun _ ↦ congr_arg _ (one_smul _ _) mul_smul _ _ _ := ext fun _ ↦ congr_arg _ (mul_smul _ _ _) smul_add _ _ _ := ext fun _ ↦ rfl @@ -85,12 +81,12 @@ end SMul section Module -variable [Semiring S] [Module S M] [Module S M₂] [SMulCommClass R₂ S M₂] +variable [Semiring S] [Module S M] [Module S M'] [SMulCommClass R' S M'] -instance [NoZeroSMulDivisors S M₂] : NoZeroSMulDivisors S (M →ₛₗ[σ₁₂] M₂) := +instance [NoZeroSMulDivisors S M'] : NoZeroSMulDivisors S (M →ₛₗ[σ₁₂] M') := coe_injective.noZeroSMulDivisors _ rfl coe_smul -instance [SMulCommClass R S M] : Module Sᵈᵐᵃ (M →ₛₗ[σ₁₂] M₂) where +instance [SMulCommClass R S M] : Module Sᵈᵐᵃ (M →ₛₗ[σ₁₂] M') where add_smul _ _ _ := ext fun _ ↦ by simp_rw [add_apply, DomMulAct.smul_linearMap_apply, ← map_add, ← add_smul]; rfl zero_smul _ := ext fun _ ↦ by erw [DomMulAct.smul_linearMap_apply, zero_smul, map_zero]; rfl diff --git a/Mathlib/Algebra/MvPolynomial/PDeriv.lean b/Mathlib/Algebra/MvPolynomial/PDeriv.lean index b7aa67a2aad48..2f36da703dccd 100644 --- a/Mathlib/Algebra/MvPolynomial/PDeriv.lean +++ b/Mathlib/Algebra/MvPolynomial/PDeriv.lean @@ -65,12 +65,12 @@ theorem pderiv_def [DecidableEq σ] (i : σ) : pderiv i = mkDerivation R (Pi.sin theorem pderiv_monomial {i : σ} : pderiv i (monomial s a) = monomial (s - single i 1) (a * s i) := by classical - simp only [pderiv_def, mkDerivation_monomial, Finsupp.smul_sum, smul_eq_mul, ← smul_mul_assoc, - ← (monomial _).map_smul] - refine (Finset.sum_eq_single i (fun j _ hne => ?_) fun hi => ?_).trans ?_ - · simp [Pi.single_eq_of_ne hne] - · rw [Finsupp.not_mem_support_iff] at hi; simp [hi] - · simp + simp only [pderiv_def, mkDerivation_monomial, Finsupp.smul_sum, smul_eq_mul, ← smul_mul_assoc, + ← (monomial _).map_smul] + refine (Finset.sum_eq_single i (fun j _ hne => ?_) fun hi => ?_).trans ?_ + · simp [Pi.single_eq_of_ne hne] + · rw [Finsupp.not_mem_support_iff] at hi; simp [hi] + · simp theorem pderiv_C {i : σ} : pderiv i (C a) = 0 := derivation_C _ _ @@ -115,9 +115,10 @@ theorem pderiv_map {S} [CommSemiring S] {φ : R →+* S} {f : MvPolynomial σ R} · simp [eq] · simp [eq, h] -lemma pderiv_rename {τ : Type*} [DecidableEq τ] [DecidableEq σ] {f : σ → τ} - (hf : Function.Injective f) (x : σ) (p : MvPolynomial σ R) : +lemma pderiv_rename {τ : Type*} {f : σ → τ} (hf : Function.Injective f) + (x : σ) (p : MvPolynomial σ R) : pderiv (f x) (rename f p) = rename f (pderiv x p) := by + classical induction' p using MvPolynomial.induction_on with a p q hp hq p a h · simp · simp [hp, hq] diff --git a/Mathlib/Algebra/Order/Archimedean/Basic.lean b/Mathlib/Algebra/Order/Archimedean/Basic.lean index d316eb2536aab..5ba6683625dc7 100644 --- a/Mathlib/Algebra/Order/Archimedean/Basic.lean +++ b/Mathlib/Algebra/Order/Archimedean/Basic.lean @@ -241,7 +241,7 @@ theorem exists_mem_Ico_zpow (hx : 0 < x) (hy : 1 < y) : ∃ n : ℤ, x ∈ Ico ( le_of_lt (by rw [zpow_neg y ↑N, zpow_natCast] - exact (inv_lt hx (lt_trans (inv_pos.2 hx) hN)).1 hN)⟩ + exact (inv_lt_comm₀ hx (lt_trans (inv_pos.2 hx) hN)).1 hN)⟩ let ⟨M, hM⟩ := pow_unbounded_of_one_lt x hy have hb : ∃ b : ℤ, ∀ m, y ^ m ≤ x → m ≤ b := ⟨M, fun m hm => @@ -257,8 +257,8 @@ but with ≤ and < the other way around. -/ theorem exists_mem_Ioc_zpow (hx : 0 < x) (hy : 1 < y) : ∃ n : ℤ, x ∈ Ioc (y ^ n) (y ^ (n + 1)) := let ⟨m, hle, hlt⟩ := exists_mem_Ico_zpow (inv_pos.2 hx) hy have hyp : 0 < y := lt_trans zero_lt_one hy - ⟨-(m + 1), by rwa [zpow_neg, inv_lt (zpow_pos_of_pos hyp _) hx], by - rwa [neg_add, neg_add_cancel_right, zpow_neg, le_inv hx (zpow_pos_of_pos hyp _)]⟩ + ⟨-(m + 1), by rwa [zpow_neg, inv_lt_comm₀ (zpow_pos_of_pos hyp _) hx], by + rwa [neg_add, neg_add_cancel_right, zpow_neg, le_inv_comm₀ hx (zpow_pos_of_pos hyp _)]⟩ /-- For any `y < 1` and any positive `x`, there exists `n : ℕ` with `y ^ n < x`. -/ theorem exists_pow_lt_of_lt_one (hx : 0 < x) (hy : y < 1) : ∃ n : ℕ, y ^ n < x := by @@ -267,18 +267,18 @@ theorem exists_pow_lt_of_lt_one (hx : 0 < x) (hy : y < 1) : ∃ n : ℕ, y ^ n < simp only [pow_one] exact y_pos.trans_lt hx rw [not_le] at y_pos - rcases pow_unbounded_of_one_lt x⁻¹ (one_lt_inv y_pos hy) with ⟨q, hq⟩ - exact ⟨q, by rwa [inv_pow, inv_lt_inv hx (pow_pos y_pos _)] at hq⟩ + rcases pow_unbounded_of_one_lt x⁻¹ ((one_lt_inv₀ y_pos).2 hy) with ⟨q, hq⟩ + exact ⟨q, by rwa [inv_pow, inv_lt_inv₀ hx (pow_pos y_pos _)] at hq⟩ /-- Given `x` and `y` between `0` and `1`, `x` is between two successive powers of `y`. This is the same as `exists_nat_pow_near`, but for elements between `0` and `1` -/ theorem exists_nat_pow_near_of_lt_one (xpos : 0 < x) (hx : x ≤ 1) (ypos : 0 < y) (hy : y < 1) : ∃ n : ℕ, y ^ (n + 1) < x ∧ x ≤ y ^ n := by - rcases exists_nat_pow_near (one_le_inv_iff.2 ⟨xpos, hx⟩) (one_lt_inv_iff.2 ⟨ypos, hy⟩) with + rcases exists_nat_pow_near (one_le_inv_iff₀.2 ⟨xpos, hx⟩) (one_lt_inv_iff₀.2 ⟨ypos, hy⟩) with ⟨n, hn, h'n⟩ refine ⟨n, ?_, ?_⟩ - · rwa [inv_pow, inv_lt_inv xpos (pow_pos ypos _)] at h'n - · rwa [inv_pow, inv_le_inv (pow_pos ypos _) xpos] at hn + · rwa [inv_pow, inv_lt_inv₀ xpos (pow_pos ypos _)] at h'n + · rwa [inv_pow, inv_le_inv₀ (pow_pos ypos _) xpos] at hn end LinearOrderedSemifield diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index c7a0670ed2b05..d5cdc6b8f86e7 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -210,6 +210,16 @@ theorem prod_le_prod_fiberwise_of_prod_fiber_le_one' {t : Finset ι'} {g : ι end OrderedCommMonoid +@[to_additive] +lemma max_prod_le [LinearOrderedCommMonoid M] {f g : ι → M} {s : Finset ι} : + max (s.prod f) (s.prod g) ≤ s.prod (fun i ↦ max (f i) (g i)) := + Multiset.max_prod_le + +@[to_additive] +lemma prod_min_le [LinearOrderedCommMonoid M] {f g : ι → M} {s : Finset ι} : + s.prod (fun i ↦ min (f i) (g i)) ≤ min (s.prod f) (s.prod g) := + Multiset.prod_min_le + theorem abs_sum_le_sum_abs {G : Type*} [LinearOrderedAddCommGroup G] (f : ι → G) (s : Finset ι) : |∑ i ∈ s, f i| ≤ ∑ i ∈ s, |f i| := le_sum_of_subadditive _ abs_zero abs_add s f diff --git a/Mathlib/Algebra/Order/BigOperators/Group/List.lean b/Mathlib/Algebra/Order/BigOperators/Group/List.lean index b440e3dea8596..db7ba413249d4 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/List.lean @@ -119,6 +119,24 @@ lemma one_le_prod_of_one_le [Preorder M] [MulLeftMono M] {l : List M} rw [prod_cons] exact one_le_mul (hl₁ hd (mem_cons_self hd tl)) (ih fun x h => hl₁ x (mem_cons_of_mem hd h)) +@[to_additive] +lemma max_prod_le (l : List α) (f g : α → M) [LinearOrder M] + [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] : + max (l.map f).prod (l.map g).prod ≤ (l.map fun i ↦ max (f i) (g i)).prod := by + rw [max_le_iff] + constructor <;> apply List.prod_le_prod' <;> intros + · apply le_max_left + · apply le_max_right + +@[to_additive] +lemma prod_min_le [LinearOrder M] [CovariantClass M M (· * ·) (· ≤ ·)] + [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] (l : List α) (f g : α → M) : + (l.map fun i ↦ min (f i) (g i)).prod ≤ min (l.map f).prod (l.map g).prod := by + rw [le_min_iff] + constructor <;> apply List.prod_le_prod' <;> intros + · apply min_le_left + · apply min_le_right + end Monoid -- TODO: develop theory of tropical rings diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Multiset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Multiset.lean index 52ae9a3e331ce..ad405fe75fb40 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Multiset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Multiset.lean @@ -156,6 +156,20 @@ lemma max_le_of_forall_le {α : Type*} [LinearOrder α] [OrderBot α] (l : Multi induction l using Quotient.inductionOn simpa using List.max_le_of_forall_le _ _ h +@[to_additive] +lemma max_prod_le [LinearOrderedCommMonoid α] {s : Multiset ι} {f g : ι → α} : + max (s.map f).prod (s.map g).prod ≤ (s.map fun i ↦ max (f i) (g i)).prod := by + obtain ⟨l⟩ := s + simp_rw [Multiset.quot_mk_to_coe'', Multiset.map_coe, Multiset.prod_coe] + apply List.max_prod_le + +@[to_additive] +lemma prod_min_le [LinearOrderedCommMonoid α] {s : Multiset ι} {f g : ι → α} : + (s.map fun i ↦ min (f i) (g i)).prod ≤ min (s.map f).prod (s.map g).prod := by + obtain ⟨l⟩ := s + simp_rw [Multiset.quot_mk_to_coe'', Multiset.map_coe, Multiset.prod_coe] + apply List.prod_min_le + lemma abs_sum_le_sum_abs [LinearOrderedAddCommGroup α] {s : Multiset α} : |s.sum| ≤ (s.map abs).sum := le_sum_of_subadditive _ abs_zero abs_add s diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 7121bb47a6885..4e76076e79d17 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -8,7 +8,6 @@ import Mathlib.Algebra.Order.Field.Defs import Mathlib.Algebra.Order.GroupWithZero.Unbundled.Lemmas import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Order.Bounds.OrderIso -import Mathlib.Tactic.Bound.Attribute import Mathlib.Tactic.Positivity.Core /-! @@ -86,10 +85,6 @@ theorem div_le_of_nonneg_of_le_mul (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ c * lemma mul_le_of_nonneg_of_le_div (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ b / c) : a * c ≤ b := mul_le_of_le_div₀ hb hc h -attribute [bound] div_le_one_of_le₀ -attribute [bound] mul_inv_le_one_of_le₀ -attribute [bound] inv_mul_le_one_of_le₀ - @[deprecated div_le_one_of_le₀ (since := "2024-10-03")] theorem div_le_one_of_le (h : a ≤ b) (hb : 0 ≤ b) : a / b ≤ 1 := div_le_one_of_le₀ h hb @@ -103,77 +98,68 @@ lemma inv_mul_le_one_of_le (h : a ≤ b) (hb : 0 ≤ b) : b⁻¹ * a ≤ 1 := in ### Bi-implications of inequalities using inversions -/ -@[gcongr, bound] -theorem inv_le_inv_of_le (ha : 0 < a) (h : a ≤ b) : b⁻¹ ≤ a⁻¹ := by - rwa [← one_div a, le_div_iff₀' ha, ← div_eq_mul_inv, div_le_iff₀ (ha.trans_le h), one_mul] +@[deprecated inv_anti₀ (since := "2024-10-05")] +theorem inv_le_inv_of_le (ha : 0 < a) (h : a ≤ b) : b⁻¹ ≤ a⁻¹ := inv_anti₀ ha h /-- See `inv_le_inv_of_le` for the implication from right-to-left with one fewer assumption. -/ -theorem inv_le_inv (ha : 0 < a) (hb : 0 < b) : a⁻¹ ≤ b⁻¹ ↔ b ≤ a := by - rw [← one_div, div_le_iff₀ ha, ← div_eq_inv_mul, le_div_iff₀ hb, one_mul] +@[deprecated inv_le_inv₀ (since := "2024-10-05")] +theorem inv_le_inv (ha : 0 < a) (hb : 0 < b) : a⁻¹ ≤ b⁻¹ ↔ b ≤ a := inv_le_inv₀ ha hb /-- In a linear ordered field, for positive `a` and `b` we have `a⁻¹ ≤ b ↔ b⁻¹ ≤ a`. See also `inv_le_of_inv_le` for a one-sided implication with one fewer assumption. -/ -theorem inv_le (ha : 0 < a) (hb : 0 < b) : a⁻¹ ≤ b ↔ b⁻¹ ≤ a := by - rw [← inv_le_inv hb (inv_pos.2 ha), inv_inv] +@[deprecated inv_le_comm₀ (since := "2024-10-05")] +theorem inv_le (ha : 0 < a) (hb : 0 < b) : a⁻¹ ≤ b ↔ b⁻¹ ≤ a := inv_le_comm₀ ha hb -theorem inv_le_of_inv_le (ha : 0 < a) (h : a⁻¹ ≤ b) : b⁻¹ ≤ a := - (inv_le ha ((inv_pos.2 ha).trans_le h)).1 h +@[deprecated inv_le_of_inv_le₀ (since := "2024-10-05")] +theorem inv_le_of_inv_le (ha : 0 < a) (h : a⁻¹ ≤ b) : b⁻¹ ≤ a := inv_le_of_inv_le₀ ha h -theorem le_inv (ha : 0 < a) (hb : 0 < b) : a ≤ b⁻¹ ↔ b ≤ a⁻¹ := by - rw [← inv_le_inv (inv_pos.2 hb) ha, inv_inv] +@[deprecated le_inv_comm₀ (since := "2024-10-05")] +theorem le_inv (ha : 0 < a) (hb : 0 < b) : a ≤ b⁻¹ ↔ b ≤ a⁻¹ := le_inv_comm₀ ha hb /-- See `inv_lt_inv_of_lt` for the implication from right-to-left with one fewer assumption. -/ -theorem inv_lt_inv (ha : 0 < a) (hb : 0 < b) : a⁻¹ < b⁻¹ ↔ b < a := - lt_iff_lt_of_le_iff_le (inv_le_inv hb ha) +@[deprecated inv_lt_inv₀ (since := "2024-10-05")] +theorem inv_lt_inv (ha : 0 < a) (hb : 0 < b) : a⁻¹ < b⁻¹ ↔ b < a := inv_lt_inv₀ ha hb -@[gcongr] -theorem inv_lt_inv_of_lt (hb : 0 < b) (h : b < a) : a⁻¹ < b⁻¹ := - (inv_lt_inv (hb.trans h) hb).2 h +@[deprecated inv_strictAnti₀ (since := "2024-10-05")] +theorem inv_lt_inv_of_lt (hb : 0 < b) (h : b < a) : a⁻¹ < b⁻¹ := inv_strictAnti₀ hb h /-- In a linear ordered field, for positive `a` and `b` we have `a⁻¹ < b ↔ b⁻¹ < a`. See also `inv_lt_of_inv_lt` for a one-sided implication with one fewer assumption. -/ -theorem inv_lt (ha : 0 < a) (hb : 0 < b) : a⁻¹ < b ↔ b⁻¹ < a := - lt_iff_lt_of_le_iff_le (le_inv hb ha) +@[deprecated inv_lt_comm₀ (since := "2024-10-05")] +theorem inv_lt (ha : 0 < a) (hb : 0 < b) : a⁻¹ < b ↔ b⁻¹ < a := inv_lt_comm₀ ha hb -theorem inv_lt_of_inv_lt (ha : 0 < a) (h : a⁻¹ < b) : b⁻¹ < a := - (inv_lt ha ((inv_pos.2 ha).trans h)).1 h +@[deprecated inv_lt_of_inv_lt₀ (since := "2024-10-05")] +theorem inv_lt_of_inv_lt (ha : 0 < a) (h : a⁻¹ < b) : b⁻¹ < a := inv_lt_of_inv_lt₀ ha h -theorem lt_inv (ha : 0 < a) (hb : 0 < b) : a < b⁻¹ ↔ b < a⁻¹ := - lt_iff_lt_of_le_iff_le (inv_le hb ha) +@[deprecated lt_inv_comm₀ (since := "2024-10-05")] +theorem lt_inv (ha : 0 < a) (hb : 0 < b) : a < b⁻¹ ↔ b < a⁻¹ := lt_inv_comm₀ ha hb -theorem inv_lt_one (ha : 1 < a) : a⁻¹ < 1 := by - rwa [inv_lt (zero_lt_one.trans ha) zero_lt_one, inv_one] +@[deprecated inv_lt_one_of_one_lt₀ (since := "2024-10-05")] +theorem inv_lt_one (ha : 1 < a) : a⁻¹ < 1 := inv_lt_one_of_one_lt₀ ha -theorem one_lt_inv (h₁ : 0 < a) (h₂ : a < 1) : 1 < a⁻¹ := by - rwa [lt_inv (@zero_lt_one α _ _ _ _ _) h₁, inv_one] +@[deprecated one_lt_inv₀ (since := "2024-10-05")] +theorem one_lt_inv (h₁ : 0 < a) (h₂ : a < 1) : 1 < a⁻¹ := (one_lt_inv₀ h₁).2 h₂ -@[bound] -theorem inv_le_one (ha : 1 ≤ a) : a⁻¹ ≤ 1 := by - rwa [inv_le (zero_lt_one.trans_le ha) zero_lt_one, inv_one] +@[deprecated inv_le_one_of_one_le₀ (since := "2024-10-05")] +theorem inv_le_one (ha : 1 ≤ a) : a⁻¹ ≤ 1 := inv_le_one_of_one_le₀ ha -theorem one_le_inv (h₁ : 0 < a) (h₂ : a ≤ 1) : 1 ≤ a⁻¹ := by - rwa [le_inv (@zero_lt_one α _ _ _ _ _) h₁, inv_one] +@[deprecated one_le_inv₀ (since := "2024-10-05")] +theorem one_le_inv (h₁ : 0 < a) (h₂ : a ≤ 1) : 1 ≤ a⁻¹ := (one_le_inv₀ h₁).2 h₂ -theorem inv_lt_one_iff_of_pos (h₀ : 0 < a) : a⁻¹ < 1 ↔ 1 < a := - ⟨fun h₁ => inv_inv a ▸ one_lt_inv (inv_pos.2 h₀) h₁, inv_lt_one⟩ +@[deprecated inv_lt_one₀ (since := "2024-10-05")] +theorem inv_lt_one_iff_of_pos (h₀ : 0 < a) : a⁻¹ < 1 ↔ 1 < a := inv_lt_one₀ h₀ -theorem inv_lt_one_iff : a⁻¹ < 1 ↔ a ≤ 0 ∨ 1 < a := by - rcases le_or_lt a 0 with ha | ha - · simp [ha, (inv_nonpos.2 ha).trans_lt zero_lt_one] - · simp only [ha.not_le, false_or, inv_lt_one_iff_of_pos ha] +@[deprecated inv_lt_one_iff₀ (since := "2024-10-05")] +theorem inv_lt_one_iff : a⁻¹ < 1 ↔ a ≤ 0 ∨ 1 < a := inv_lt_one_iff₀ -theorem one_lt_inv_iff : 1 < a⁻¹ ↔ 0 < a ∧ a < 1 := - ⟨fun h => ⟨inv_pos.1 (zero_lt_one.trans h), - inv_inv a ▸ inv_lt_one h⟩, and_imp.2 one_lt_inv⟩ +@[deprecated one_lt_inv_iff₀ (since := "2024-10-05")] +theorem one_lt_inv_iff : 1 < a⁻¹ ↔ 0 < a ∧ a < 1 := one_lt_inv_iff₀ -theorem inv_le_one_iff : a⁻¹ ≤ 1 ↔ a ≤ 0 ∨ 1 ≤ a := by - rcases em (a = 1) with (rfl | ha) - · simp [le_rfl] - · simp only [Ne.le_iff_lt (Ne.symm ha), Ne.le_iff_lt (mt inv_eq_one.1 ha), inv_lt_one_iff] +@[deprecated inv_le_one_iff₀ (since := "2024-10-05")] +theorem inv_le_one_iff : a⁻¹ ≤ 1 ↔ a ≤ 0 ∨ 1 ≤ a := inv_le_one_iff₀ -theorem one_le_inv_iff : 1 ≤ a⁻¹ ↔ 0 < a ∧ a ≤ 1 := - ⟨fun h => ⟨inv_pos.1 (zero_lt_one.trans_le h), - inv_inv a ▸ inv_le_one h⟩, and_imp.2 one_le_inv⟩ +@[deprecated one_le_inv_iff₀ (since := "2024-10-05")] +theorem one_le_inv_iff : 1 ≤ a⁻¹ ↔ 0 < a ∧ a ≤ 1 := one_le_inv_iff₀ /-! ### Relating two divisions. @@ -194,11 +180,11 @@ lemma div_lt_div_of_pos_right (h : a < b) (hc : 0 < c) : a / c < b / c := by @[gcongr] lemma div_le_div_of_nonneg_left (ha : 0 ≤ a) (hc : 0 < c) (h : c ≤ b) : a / b ≤ a / c := by rw [div_eq_mul_inv, div_eq_mul_inv] - exact mul_le_mul_of_nonneg_left ((inv_le_inv (hc.trans_le h) hc).mpr h) ha + exact mul_le_mul_of_nonneg_left ((inv_le_inv₀ (hc.trans_le h) hc).mpr h) ha @[gcongr, bound] lemma div_lt_div_of_pos_left (ha : 0 < a) (hc : 0 < c) (h : c < b) : a / b < a / c := by - simpa only [div_eq_mul_inv, mul_lt_mul_left ha, inv_lt_inv (hc.trans h) hc] + simpa only [div_eq_mul_inv, mul_lt_mul_left ha, inv_lt_inv₀ (hc.trans h) hc] @[deprecated (since := "2024-02-16")] alias div_le_div_of_le_of_nonneg := div_le_div_of_nonneg_right @[deprecated (since := "2024-02-16")] alias div_lt_div_of_lt := div_lt_div_of_pos_right @@ -217,7 +203,7 @@ theorem div_lt_div_right (hc : 0 < c) : a / c < b / c ↔ a < b := lt_iff_lt_of_le_iff_le <| div_le_div_right hc theorem div_lt_div_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a / b < a / c ↔ c < b := by - simp only [div_eq_mul_inv, mul_lt_mul_left ha, inv_lt_inv hb hc] + simp only [div_eq_mul_inv, mul_lt_mul_left ha, inv_lt_inv₀ hb hc] theorem div_le_div_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a / b ≤ a / c ↔ c ≤ b := le_iff_le_iff_lt_iff_lt.2 (div_lt_div_left ha hc hb) @@ -265,13 +251,17 @@ theorem one_lt_div (hb : 0 < b) : 1 < a / b ↔ b < a := by rw [lt_div_iff₀ hb theorem div_lt_one (hb : 0 < b) : a / b < 1 ↔ a < b := by rw [div_lt_iff₀ hb, one_mul] -theorem one_div_le (ha : 0 < a) (hb : 0 < b) : 1 / a ≤ b ↔ 1 / b ≤ a := by simpa using inv_le ha hb +theorem one_div_le (ha : 0 < a) (hb : 0 < b) : 1 / a ≤ b ↔ 1 / b ≤ a := by + simpa using inv_le_comm₀ ha hb -theorem one_div_lt (ha : 0 < a) (hb : 0 < b) : 1 / a < b ↔ 1 / b < a := by simpa using inv_lt ha hb +theorem one_div_lt (ha : 0 < a) (hb : 0 < b) : 1 / a < b ↔ 1 / b < a := by + simpa using inv_lt_comm₀ ha hb -theorem le_one_div (ha : 0 < a) (hb : 0 < b) : a ≤ 1 / b ↔ b ≤ 1 / a := by simpa using le_inv ha hb +theorem le_one_div (ha : 0 < a) (hb : 0 < b) : a ≤ 1 / b ↔ b ≤ 1 / a := by + simpa using le_inv_comm₀ ha hb -theorem lt_one_div (ha : 0 < a) (hb : 0 < b) : a < 1 / b ↔ b < 1 / a := by simpa using lt_inv ha hb +theorem lt_one_div (ha : 0 < a) (hb : 0 < b) : a < 1 / b ↔ b < 1 / a := by + simpa using lt_inv_comm₀ ha hb @[bound] lemma Bound.one_lt_div_of_pos_of_lt (b0 : 0 < b) : b < a → 1 < a / b := (one_lt_div b0).mpr @@ -283,7 +273,7 @@ theorem lt_one_div (ha : 0 < a) (hb : 0 < b) : a < 1 / b ↔ b < 1 / a := by sim theorem one_div_le_one_div_of_le (ha : 0 < a) (h : a ≤ b) : 1 / b ≤ 1 / a := by - simpa using inv_le_inv_of_le ha h + simpa using inv_anti₀ ha h theorem one_div_lt_one_div_of_lt (ha : 0 < a) (h : a < b) : 1 / b < 1 / a := by rwa [lt_div_iff₀' ha, ← div_eq_mul_one_div, div_lt_one (ha.trans h)] @@ -415,7 +405,7 @@ theorem one_div_strictAntiOn : StrictAntiOn (fun x : α => 1 / x) (Set.Ioi 0) := theorem one_div_pow_le_one_div_pow_of_le (a1 : 1 ≤ a) {m n : ℕ} (mn : m ≤ n) : 1 / a ^ n ≤ 1 / a ^ m := by - refine (one_div_le_one_div ?_ ?_).mpr (pow_le_pow_right a1 mn) <;> + refine (one_div_le_one_div ?_ ?_).mpr (pow_right_mono₀ a1 mn) <;> exact pow_pos (zero_lt_one.trans_le a1) _ theorem one_div_pow_lt_one_div_pow_of_lt (a1 : 1 < a) {m n : ℕ} (mn : m < n) : @@ -430,7 +420,7 @@ theorem one_div_pow_strictAnti (a1 : 1 < a) : StrictAnti fun n : ℕ => 1 / a ^ one_div_pow_lt_one_div_pow_of_lt a1 theorem inv_strictAntiOn : StrictAntiOn (fun x : α => x⁻¹) (Set.Ioi 0) := fun _ hx _ hy xy => - (inv_lt_inv hy hx).2 xy + (inv_lt_inv₀ hy hx).2 xy theorem inv_pow_le_inv_pow_of_le (a1 : 1 ≤ a) {m n : ℕ} (mn : m ≤ n) : (a ^ n)⁻¹ ≤ (a ^ m)⁻¹ := by convert one_div_pow_le_one_div_pow_of_le a1 mn using 1 <;> simp @@ -555,7 +545,7 @@ theorem lt_inv_of_neg (ha : a < 0) (hb : b < 0) : a < b⁻¹ ↔ b < a⁻¹ := theorem sub_inv_antitoneOn_Ioi : AntitoneOn (fun x ↦ (x-c)⁻¹) (Set.Ioi c) := antitoneOn_iff_forall_lt.mpr fun _ ha _ hb hab ↦ - inv_le_inv (sub_pos.mpr hb) (sub_pos.mpr ha) |>.mpr <| sub_le_sub (le_of_lt hab) le_rfl + inv_le_inv₀ (sub_pos.mpr hb) (sub_pos.mpr ha) |>.mpr <| sub_le_sub (le_of_lt hab) le_rfl theorem sub_inv_antitoneOn_Iio : AntitoneOn (fun x ↦ (x-c)⁻¹) (Set.Iio c) := @@ -709,11 +699,11 @@ theorem add_sub_div_two_lt (h : a < b) : a + (b - a) / 2 < b := by /-- An inequality involving `2`. -/ theorem sub_one_div_inv_le_two (a2 : 2 ≤ a) : (1 - 1 / a)⁻¹ ≤ 2 := by -- Take inverses on both sides to obtain `2⁻¹ ≤ 1 - 1 / a` - refine (inv_le_inv_of_le (inv_pos.2 <| zero_lt_two' α) ?_).trans_eq (inv_inv (2 : α)) + refine (inv_anti₀ (inv_pos.2 <| zero_lt_two' α) ?_).trans_eq (inv_inv (2 : α)) -- move `1 / a` to the left and `2⁻¹` to the right. rw [le_sub_iff_add_le, add_comm, ← le_sub_iff_add_le] -- take inverses on both sides and use the assumption `2 ≤ a`. - convert (one_div a).le.trans (inv_le_inv_of_le zero_lt_two a2) using 1 + convert (one_div a).le.trans (inv_anti₀ zero_lt_two a2) using 1 -- show `1 - 1 / 2 = 1 / 2`. rw [sub_eq_iff_eq_add, ← two_mul, mul_inv_cancel₀ two_ne_zero] diff --git a/Mathlib/Algebra/Order/Field/Defs.lean b/Mathlib/Algebra/Order/Field/Defs.lean index ad7ae9d36bb7a..86284a304dde6 100644 --- a/Mathlib/Algebra/Order/Field/Defs.lean +++ b/Mathlib/Algebra/Order/Field/Defs.lean @@ -36,7 +36,7 @@ instance (priority := 100) LinearOrderedField.toLinearOrderedSemifield [LinearOr LinearOrderedSemifield α := { LinearOrderedRing.toLinearOrderedSemiring, ‹LinearOrderedField α› with } -variable [LinearOrderedSemifield α] {a b : α} +variable [LinearOrderedSemifield α] {a b c : α} /-- Equality holds when `a ≠ 0`. See `mul_inv_cancel`. -/ lemma mul_inv_le_one : a * a⁻¹ ≤ 1 := by obtain rfl | ha := eq_or_ne a 0 <;> simp [*] @@ -75,3 +75,27 @@ lemma inv_mul_right_le (ha : 0 ≤ a) : a * b⁻¹ * b ≤ a := by /-- Equality holds when `b ≠ 0`. See `inv_mul_cancel_right`. -/ lemma le_inv_mul_right (ha : a ≤ 0) : a ≤ a * b⁻¹ * b := by obtain rfl | hb := eq_or_ne b 0 <;> simp [*] + +/-- Equality holds when `c ≠ 0`. See `mul_div_mul_left`. -/ +lemma mul_div_mul_left_le (h : 0 ≤ a / b) : c * a / (c * b) ≤ a / b := by + obtain rfl | hc := eq_or_ne c 0 + · simpa + · rw [mul_div_mul_left _ _ hc] + +/-- Equality holds when `c ≠ 0`. See `mul_div_mul_left`. -/ +lemma le_mul_div_mul_left (h : a / b ≤ 0) : a / b ≤ c * a / (c * b) := by + obtain rfl | hc := eq_or_ne c 0 + · simpa + · rw [mul_div_mul_left _ _ hc] + +/-- Equality holds when `c ≠ 0`. See `mul_div_mul_right`. -/ +lemma mul_div_mul_right_le (h : 0 ≤ a / b) : a * c / (b * c) ≤ a / b := by + obtain rfl | hc := eq_or_ne c 0 + · simpa + · rw [mul_div_mul_right _ _ hc] + +/-- Equality holds when `c ≠ 0`. See `mul_div_mul_right`. -/ +lemma le_mul_div_mul_right (h : a / b ≤ 0) : a / b ≤ a * c / (b * c) := by + obtain rfl | hc := eq_or_ne c 0 + · simpa + · rw [mul_div_mul_right _ _ hc] diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 6bd51a8b46a96..09bdb68063e38 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -222,21 +222,13 @@ theorem floor_eq_zero : ⌊a⌋₊ = 0 ↔ a < 1 := by rw [← lt_one_iff, ← @cast_one α] exact floor_lt' Nat.one_ne_zero -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem floor_eq_iff (ha : 0 ≤ a) : ⌊a⌋₊ = n ↔ ↑n ≤ a ∧ a < ↑n + 1 := by rw [← le_floor_iff ha, ← Nat.cast_one, ← Nat.cast_add, ← floor_lt ha, Nat.lt_add_one_iff, - le_antisymm_iff, _root_.and_comm] + le_antisymm_iff, and_comm] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem floor_eq_iff' (hn : n ≠ 0) : ⌊a⌋₊ = n ↔ ↑n ≤ a ∧ a < ↑n + 1 := by rw [← le_floor_iff' hn, ← Nat.cast_one, ← Nat.cast_add, ← floor_lt' (Nat.add_one_ne_zero n), - Nat.lt_add_one_iff, le_antisymm_iff, _root_.and_comm] + Nat.lt_add_one_iff, le_antisymm_iff, and_comm] theorem floor_eq_on_Ico (n : ℕ) : ∀ a ∈ (Set.Ico n (n + 1) : Set α), ⌊a⌋₊ = n := fun _ ⟨h₀, h₁⟩ => (floor_eq_iff <| n.cast_nonneg.trans h₀).mpr ⟨h₀, h₁⟩ @@ -328,14 +320,10 @@ theorem floor_lt_ceil_of_lt_of_pos {a b : α} (h : a < b) (h' : 0 < b) : ⌊a⌋ exact h.trans_le (le_ceil _) · rwa [floor_of_nonpos ha.le, lt_ceil, Nat.cast_zero] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem ceil_eq_iff (hn : n ≠ 0) : ⌈a⌉₊ = n ↔ ↑(n - 1) < a ∧ a ≤ n := by rw [← ceil_le, ← not_le, ← ceil_le, not_le, tsub_lt_iff_right (Nat.add_one_le_iff.2 (pos_iff_ne_zero.2 hn)), Nat.lt_add_one_iff, - le_antisymm_iff, _root_.and_comm] + le_antisymm_iff, and_comm] @[simp] theorem preimage_ceil_zero : (Nat.ceil : α → ℕ) ⁻¹' {0} = Iic 0 := @@ -1273,7 +1261,7 @@ lemma ceil_div_ceil_inv_sub_one (ha : 1 ≤ a) : ⌈⌈(a - 1)⁻¹⌉ / a⌉ = have : 0 < ⌈(a - 1)⁻¹⌉ := ceil_pos.2 <| by positivity refine le_antisymm (ceil_le.2 <| div_le_self (by positivity) ha.le) <| ?_ rw [le_ceil_iff, sub_lt_comm, div_eq_mul_inv, ← mul_one_sub, - ← lt_div_iff₀ (sub_pos.2 <| inv_lt_one ha)] + ← lt_div_iff₀ (sub_pos.2 <| inv_lt_one_of_one_lt₀ ha)] convert ceil_lt_add_one _ using 1 field_simp diff --git a/Mathlib/Algebra/Order/Floor/Div.lean b/Mathlib/Algebra/Order/Floor/Div.lean index 35d76dbf57197..20b853e6e5f94 100644 --- a/Mathlib/Algebra/Order/Floor/Div.lean +++ b/Mathlib/Algebra/Order/Floor/Div.lean @@ -119,7 +119,7 @@ end OrderedAddCommMonoid section LinearOrderedAddCommMonoid variable [LinearOrderedAddCommMonoid α] [OrderedAddCommMonoid β] [SMulZeroClass α β] - [PosSMulReflectLE α β] [FloorDiv α β] [CeilDiv α β] {a : α} {b c : β} + [PosSMulReflectLE α β] [FloorDiv α β] [CeilDiv α β] {a : α} {b : β} lemma floorDiv_le_ceilDiv : b ⌊/⌋ a ≤ b ⌈/⌉ a := by obtain ha | ha := le_or_lt a 0 diff --git a/Mathlib/Algebra/Order/Floor/Prime.lean b/Mathlib/Algebra/Order/Floor/Prime.lean index 6c418d3d3d34b..f3545648ae511 100644 --- a/Mathlib/Algebra/Order/Floor/Prime.lean +++ b/Mathlib/Algebra/Order/Floor/Prime.lean @@ -3,40 +3,44 @@ Copyright (c) 2022 Yuyang Zhao. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuyang Zhao -/ - -import Mathlib.Algebra.Order.Floor import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Topology.Algebra.Order.Floor /-! # Existence of a sufficiently large prime for which `a * c ^ p / (p - 1)! < 1` This is a technical result used in the proof of the Lindemann-Weierstrass theorem. --/ -namespace FloorRing +TODO: delete this file, as all its lemmas have been deprecated. +-/ open scoped Nat +@[deprecated eventually_mul_pow_lt_factorial_sub (since := "2024-09-25")] +theorem Nat.exists_prime_mul_pow_lt_factorial (n a c : ℕ) : + ∃ p > n, p.Prime ∧ a * c ^ p < (p - 1)! := + ((Filter.frequently_atTop.mpr Nat.exists_infinite_primes).and_eventually + (eventually_mul_pow_lt_factorial_sub a c 1)).forall_exists_of_atTop (n + 1) + +namespace FloorRing + variable {K : Type*} +@[deprecated FloorSemiring.eventually_mul_pow_lt_factorial_sub (since := "2024-09-25")] theorem exists_prime_mul_pow_lt_factorial [LinearOrderedRing K] [FloorRing K] (n : ℕ) (a c : K) : - ∃ p > n, p.Prime ∧ a * c ^ p < (p - 1)! := by - obtain ⟨p, pn, pp, h⟩ := n.exists_prime_mul_pow_lt_factorial ⌈|a|⌉.natAbs ⌈|c|⌉.natAbs - use p, pn, pp - calc a * c ^ p - _ ≤ |a * c ^ p| := le_abs_self _ - _ ≤ ⌈|a|⌉ * (⌈|c|⌉ : K) ^ p := ?_ - _ = ↑(Int.natAbs ⌈|a|⌉ * Int.natAbs ⌈|c|⌉ ^ p) := ?_ - _ < ↑(p - 1)! := Nat.cast_lt.mpr h - · rw [abs_mul, abs_pow] - gcongr <;> try first | positivity | apply Int.le_ceil - · simp_rw [Nat.cast_mul, Nat.cast_pow, Int.cast_natAbs, - abs_eq_self.mpr (Int.ceil_nonneg (abs_nonneg (_ : K)))] + ∃ p > n, p.Prime ∧ a * c ^ p < (p - 1)! := + ((Filter.frequently_atTop.mpr Nat.exists_infinite_primes).and_eventually + (FloorSemiring.eventually_mul_pow_lt_factorial_sub a c 1)).forall_exists_of_atTop (n + 1) +@[deprecated FloorSemiring.tendsto_mul_pow_div_factorial_sub_atTop (since := "2024-09-25")] theorem exists_prime_mul_pow_div_factorial_lt_one [LinearOrderedField K] [FloorRing K] (n : ℕ) (a c : K) : - ∃ p > n, p.Prime ∧ a * c ^ p / (p - 1)! < 1 := by - simp_rw [div_lt_one (α := K) (Nat.cast_pos.mpr (Nat.factorial_pos _))] - exact exists_prime_mul_pow_lt_factorial .. + ∃ p > n, p.Prime ∧ a * c ^ p / (p - 1)! < 1 := + letI := Preorder.topology K + haveI : OrderTopology K := ⟨rfl⟩ + ((Filter.frequently_atTop.mpr Nat.exists_infinite_primes).and_eventually + (eventually_lt_of_tendsto_lt zero_lt_one + (FloorSemiring.tendsto_mul_pow_div_factorial_sub_atTop a c 1))).forall_exists_of_atTop + (n + 1) end FloorRing diff --git a/Mathlib/Algebra/Order/Group/Cone.lean b/Mathlib/Algebra/Order/Group/Cone.lean index 59f273599f45a..6d794594f08df 100644 --- a/Mathlib/Algebra/Order/Group/Cone.lean +++ b/Mathlib/Algebra/Order/Group/Cone.lean @@ -24,7 +24,7 @@ class AddGroupConeClass (S : Type*) (G : outParam Type*) [AddCommGroup G] [SetLi /-- `GroupConeClass S G` says that `S` is a type of cones in `G`. -/ @[to_additive] -class GroupConeClass (S G : Type*) [CommGroup G] [SetLike S G] extends +class GroupConeClass (S : Type*) (G : outParam Type*) [CommGroup G] [SetLike S G] extends SubmonoidClass S G : Prop where eq_one_of_mem_of_inv_mem {C : S} {a : G} : a ∈ C → a⁻¹ ∈ C → a = 1 diff --git a/Mathlib/Algebra/Order/Group/DenselyOrdered.lean b/Mathlib/Algebra/Order/Group/DenselyOrdered.lean index 3236b10869608..70408fea662ab 100644 --- a/Mathlib/Algebra/Order/Group/DenselyOrdered.lean +++ b/Mathlib/Algebra/Order/Group/DenselyOrdered.lean @@ -18,7 +18,7 @@ section DenselyOrdered variable [Group α] [LinearOrder α] variable [MulLeftMono α] -variable [DenselyOrdered α] {a b c : α} +variable [DenselyOrdered α] {a b : α} @[to_additive] theorem le_of_forall_lt_one_mul_le (h : ∀ ε < 1, a * ε ≤ b) : a ≤ b := diff --git a/Mathlib/Algebra/Order/Group/Indicator.lean b/Mathlib/Algebra/Order/Group/Indicator.lean index 6ae4c61cca65e..a8e229c800374 100644 --- a/Mathlib/Algebra/Order/Group/Indicator.lean +++ b/Mathlib/Algebra/Order/Group/Indicator.lean @@ -106,18 +106,25 @@ lemma mulIndicator_le_mulIndicator' (h : a ∈ s → f a ≤ g a) : mulIndicator s f a ≤ mulIndicator s g a := mulIndicator_rel_mulIndicator le_rfl h -@[to_additive] +@[to_additive (attr := mono, gcongr)] lemma mulIndicator_le_mulIndicator (h : f a ≤ g a) : mulIndicator s f a ≤ mulIndicator s g a := mulIndicator_rel_mulIndicator le_rfl fun _ ↦ h -attribute [mono] mulIndicator_le_mulIndicator indicator_le_indicator +@[to_additive (attr := gcongr)] +lemma mulIndicator_mono (h : f ≤ g) : s.mulIndicator f ≤ s.mulIndicator g := + fun _ ↦ mulIndicator_le_mulIndicator (h _) @[to_additive] -lemma mulIndicator_le_mulIndicator_of_subset (h : s ⊆ t) (hf : ∀ a, 1 ≤ f a) (a : α) : +lemma mulIndicator_le_mulIndicator_apply_of_subset (h : s ⊆ t) (hf : 1 ≤ f a) : mulIndicator s f a ≤ mulIndicator t f a := mulIndicator_apply_le' (fun ha ↦ le_mulIndicator_apply (fun _ ↦ le_rfl) fun hat ↦ (hat <| h ha).elim) fun _ ↦ - one_le_mulIndicator_apply fun _ ↦ hf _ + one_le_mulIndicator_apply fun _ ↦ hf + +@[to_additive] +lemma mulIndicator_le_mulIndicator_of_subset (h : s ⊆ t) (hf : 1 ≤ f) : + mulIndicator s f ≤ mulIndicator t f := + fun _ ↦ mulIndicator_le_mulIndicator_apply_of_subset h (hf _) @[to_additive] lemma mulIndicator_le_self' (hf : ∀ x ∉ s, 1 ≤ f x) : mulIndicator s f ≤ f := @@ -174,6 +181,23 @@ lemma mulIndicator_iInter_apply (h1 : (⊥ : M) = 1) (s : ι → Set α) (f : α refine le_antisymm (by simp only [← h1, le_iInf_iff, bot_le, forall_const]) ?_ simpa [mulIndicator_of_not_mem hj] using (iInf_le (fun i ↦ (s i).mulIndicator f) j) x +@[to_additive] +lemma iSup_mulIndicator {ι : Type*} [Preorder ι] [IsDirected ι (· ≤ ·)] {f : ι → α → M} + {s : ι → Set α} (h1 : (⊥ : M) = 1) (hf : Monotone f) (hs : Monotone s) : + ⨆ i, (s i).mulIndicator (f i) = (⋃ i, s i).mulIndicator (⨆ i, f i) := by + simp only [le_antisymm_iff, iSup_le_iff] + refine ⟨fun i ↦ (mulIndicator_mono (le_iSup _ _)).trans (mulIndicator_le_mulIndicator_of_subset + (subset_iUnion _ _) (fun _ ↦ by simp [← h1])), fun a ↦ ?_⟩ + by_cases ha : a ∈ ⋃ i, s i + · obtain ⟨i, hi⟩ : ∃ i, a ∈ s i := by simpa using ha + rw [mulIndicator_of_mem ha, iSup_apply, iSup_apply] + refine iSup_le fun j ↦ ?_ + obtain ⟨k, hik, hjk⟩ := exists_ge_ge i j + refine le_iSup_of_le k <| (hf hjk _).trans_eq ?_ + rw [mulIndicator_of_mem (hs hik hi)] + · rw [mulIndicator_of_not_mem ha, ← h1] + exact bot_le + end CompleteLattice section CanonicallyOrderedCommMonoid diff --git a/Mathlib/Algebra/Order/Group/MinMax.lean b/Mathlib/Algebra/Order/Group/MinMax.lean index aa3dd0282e1d0..06f9007ec5e6f 100644 --- a/Mathlib/Algebra/Order/Group/MinMax.lean +++ b/Mathlib/Algebra/Order/Group/MinMax.lean @@ -30,7 +30,7 @@ end section LinearOrderedCommGroup -variable {α : Type*} [LinearOrderedCommGroup α] {a b c : α} +variable {α : Type*} [LinearOrderedCommGroup α] @[to_additive min_neg_neg] theorem min_inv_inv' (a b : α) : min a⁻¹ b⁻¹ = (max a b)⁻¹ := @@ -62,7 +62,7 @@ end LinearOrderedCommGroup section LinearOrderedAddCommGroup -variable {α : Type*} [LinearOrderedAddCommGroup α] {a b c : α} +variable {α : Type*} [LinearOrderedAddCommGroup α] theorem max_sub_max_le_max (a b c d : α) : max a b - max c d ≤ max (a - c) (b - d) := by simp only [sub_le_iff_le_add, max_le_iff]; constructor diff --git a/Mathlib/Algebra/Order/Group/PosPart.lean b/Mathlib/Algebra/Order/Group/PosPart.lean index 464120cf3a627..66e862110c728 100644 --- a/Mathlib/Algebra/Order/Group/PosPart.lean +++ b/Mathlib/Algebra/Order/Group/PosPart.lean @@ -242,12 +242,12 @@ end covariantmul end LinearOrder namespace Pi -variable {ι : Type*} {α : ι → Type*} [∀ i, Lattice (α i)] [∀ i, AddCommGroup (α i)] +variable {ι : Type*} {α : ι → Type*} [∀ i, Lattice (α i)] [∀ i, Group (α i)] -@[to_additive (attr := simp)] lemma oneLePart_apply (f : ∀ i, α i) (i : ι) : f⁺ i = (f i)⁺ := rfl -@[to_additive (attr := simp)] lemma leOnePart_apply (f : ∀ i, α i) (i : ι) : f⁻ i = (f i)⁻ := rfl +@[to_additive (attr := simp)] lemma oneLePart_apply (f : ∀ i, α i) (i : ι) : f⁺ᵐ i = (f i)⁺ᵐ := rfl +@[to_additive (attr := simp)] lemma leOnePart_apply (f : ∀ i, α i) (i : ι) : f⁻ᵐ i = (f i)⁻ᵐ := rfl -@[to_additive] lemma oneLePart_def (f : ∀ i, α i) : f⁺ = fun i ↦ (f i)⁺ := rfl -@[to_additive] lemma leOnePart_def (f : ∀ i, α i) : f⁻ = fun i ↦ (f i)⁻ := rfl +@[to_additive] lemma oneLePart_def (f : ∀ i, α i) : f⁺ᵐ = fun i ↦ (f i)⁺ᵐ := rfl +@[to_additive] lemma leOnePart_def (f : ∀ i, α i) : f⁻ᵐ = fun i ↦ (f i)⁻ᵐ := rfl end Pi diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 1e831d8b42d2b..7007a0ea5cd76 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -33,7 +33,7 @@ variable [Group α] section TypeclassesLeftLE -variable [LE α] [MulLeftMono α] {a b c d : α} +variable [LE α] [MulLeftMono α] {a b c : α} /-- Uses `left` co(ntra)variant. -/ @[to_additive (attr := simp) "Uses `left` co(ntra)variant."] @@ -471,7 +471,7 @@ variable [Group α] [LE α] section Right -variable [MulRightMono α] {a b c d : α} +variable [MulRightMono α] {a b c : α} @[to_additive] theorem div_le_div_iff_right (c : α) : a / c ≤ b / c ↔ a ≤ b := by @@ -594,7 +594,7 @@ variable [Group α] [LT α] section Right -variable [MulRightStrictMono α] {a b c d : α} +variable [MulRightStrictMono α] {a b c : α} @[to_additive (attr := simp)] theorem div_lt_div_iff_right (c : α) : a / c < b / c ↔ a < b := by @@ -722,7 +722,7 @@ variable [MulLeftMono α] section VariableNames -variable {a b c : α} +variable {a b : α} @[to_additive] theorem le_of_forall_one_lt_lt_mul (h : ∀ ε : α, 1 < ε → a < b * ε) : a ≤ b := diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index c0ac4cc1ccec8..6a2aa7896bd6e 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -130,11 +130,6 @@ instance (priority := 100) LinearOrderedCommGroupWithZero.toMulPosStrictMono : MulPosStrictMono α where elim a b c hbc := by by_contra! h; exact hbc.not_le <| (mul_le_mul_right a.2).1 h -/-- Alias of `mul_le_one'` for unification. -/ -@[deprecated mul_le_one' (since := "2024-08-21")] -theorem mul_le_one₀ (ha : a ≤ 1) (hb : b ≤ 1) : a * b ≤ 1 := - mul_le_one' ha hb - /-- Alias of `one_le_mul'` for unification. -/ @[deprecated one_le_mul (since := "2024-08-21")] theorem one_le_mul₀ (ha : 1 ≤ a) (hb : 1 ≤ b) : 1 ≤ a * b := @@ -185,17 +180,9 @@ theorem inv_mul_lt_of_lt_mul₀ (h : a < b * c) : b⁻¹ * a < c := by theorem mul_lt_right₀ (c : α) (h : a < b) (hc : c ≠ 0) : a * c < b * c := mul_lt_mul_of_pos_right h (zero_lt_iff.2 hc) -theorem inv_lt_inv₀ (ha : a ≠ 0) (hb : b ≠ 0) : a⁻¹ < b⁻¹ ↔ b < a := - show (Units.mk0 a ha)⁻¹ < (Units.mk0 b hb)⁻¹ ↔ Units.mk0 b hb < Units.mk0 a ha from - inv_lt_inv_iff - -theorem inv_le_inv₀ (ha : a ≠ 0) (hb : b ≠ 0) : a⁻¹ ≤ b⁻¹ ↔ b ≤ a := - show (Units.mk0 a ha)⁻¹ ≤ (Units.mk0 b hb)⁻¹ ↔ Units.mk0 b hb ≤ Units.mk0 a ha from - inv_le_inv_iff - theorem lt_of_mul_lt_mul_of_le₀ (h : a * b < c * d) (hc : 0 < c) (hh : c ≤ a) : b < d := by have ha : a ≠ 0 := ne_of_gt (lt_of_lt_of_le hc hh) - simp_rw [← inv_le_inv₀ ha (ne_of_gt hc)] at hh + rw [← inv_le_inv₀ (zero_lt_iff.2 ha) hc] at hh have := mul_lt_mul_of_lt_of_le₀ hh (inv_ne_zero (ne_of_gt hc)) h simpa [inv_mul_cancel_left₀ ha, inv_mul_cancel_left₀ (ne_of_gt hc)] using this @@ -211,7 +198,8 @@ theorem div_le_div_right₀ (hc : c ≠ 0) : a / c ≤ b / c ↔ a ≤ b := by rw [div_eq_mul_inv, div_eq_mul_inv, mul_le_mul_right (zero_lt_iff.2 (inv_ne_zero hc))] theorem div_le_div_left₀ (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) : a / b ≤ a / c ↔ c ≤ b := by - simp only [div_eq_mul_inv, mul_le_mul_left (zero_lt_iff.2 ha), inv_le_inv₀ hb hc] + simp only [div_eq_mul_inv, mul_le_mul_left (zero_lt_iff.2 ha), + inv_le_inv₀ (zero_lt_iff.2 hb) (zero_lt_iff.2 hc)] /-- `Equiv.mulLeft₀` as an `OrderIso` on a `LinearOrderedCommGroupWithZero.`. diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index a83a71f9e5b2e..d89b2b8a4fc2b 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Group.Pi.Basic import Mathlib.Algebra.GroupWithZero.Units.Basic import Mathlib.Algebra.Order.Monoid.Unbundled.Defs import Mathlib.Algebra.Order.ZeroLEOne +import Mathlib.Tactic.Bound.Attribute import Mathlib.Tactic.GCongr.CoreAttrs import Mathlib.Tactic.Nontriviality @@ -935,12 +936,16 @@ section MonoidWithZero variable [MonoidWithZero M₀] section Preorder -variable [Preorder M₀] {a b c d : M₀} {n : ℕ} +variable [Preorder M₀] {a b : M₀} {m n : ℕ} @[simp] lemma pow_nonneg [ZeroLEOneClass M₀] [PosMulMono M₀] (ha : 0 ≤ a) : ∀ n, 0 ≤ a ^ n | 0 => pow_zero a ▸ zero_le_one | n + 1 => pow_succ a n ▸ mul_nonneg (pow_nonneg ha _) ha +lemma zero_pow_le_one [ZeroLEOneClass M₀] : ∀ n : ℕ, (0 : M₀) ^ n ≤ 1 + | 0 => (pow_zero _).le + | n + 1 => by rw [zero_pow n.succ_ne_zero]; exact zero_le_one + lemma pow_le_pow_of_le_one [ZeroLEOneClass M₀] [PosMulMono M₀] [MulPosMono M₀] (ha₀ : 0 ≤ a) (ha₁ : a ≤ 1) : ∀ {m n : ℕ}, m ≤ n → a ^ n ≤ a ^ m | _, _, Nat.le.refl => le_rfl @@ -958,9 +963,6 @@ lemma sq_le [ZeroLEOneClass M₀] [PosMulMono M₀] [MulPosMono M₀] (h₀ : 0 lemma one_le_mul_of_one_le_of_one_le [ZeroLEOneClass M₀] [PosMulMono M₀] (ha : 1 ≤ a) (hb : 1 ≤ b) : (1 : M₀) ≤ a * b := Left.one_le_mul_of_le_of_le ha hb <| zero_le_one.trans ha -lemma mul_le_one [ZeroLEOneClass M₀] [PosMulMono M₀] [MulPosMono M₀] (ha : a ≤ 1) (hb₀ : 0 ≤ b) - (hb : b ≤ 1) : a * b ≤ 1 := one_mul (1 : M₀) ▸ mul_le_mul ha hb hb₀ zero_le_one - lemma one_lt_mul_of_le_of_lt [ZeroLEOneClass M₀] [MulPosMono M₀] (ha : 1 ≤ a) (hb : 1 < b) : 1 < a * b := hb.trans_le <| le_mul_of_one_le_left (zero_le_one.trans hb.le) ha @@ -975,6 +977,43 @@ lemma mul_lt_one_of_nonneg_of_lt_one_left [PosMulMono M₀] (ha₀ : 0 ≤ a) (h lemma mul_lt_one_of_nonneg_of_lt_one_right [MulPosMono M₀] (ha : a ≤ 1) (hb₀ : 0 ≤ b) (hb : b < 1) : a * b < 1 := (mul_le_of_le_one_left hb₀ ha).trans_lt hb +section +variable [ZeroLEOneClass M₀] [PosMulMono M₀] [MulPosMono M₀] + +lemma mul_le_one₀ (ha : a ≤ 1) (hb₀ : 0 ≤ b) (hb : b ≤ 1) : a * b ≤ 1 := + one_mul (1 : M₀) ▸ mul_le_mul ha hb hb₀ zero_le_one + +lemma pow_le_one₀ : ∀ {n : ℕ}, 0 ≤ a → a ≤ 1 → a ^ n ≤ 1 + | 0, _, _ => (pow_zero a).le + | n + 1, h₀, h₁ => (pow_succ a n).le.trans (mul_le_one₀ (pow_le_one₀ h₀ h₁) h₀ h₁) + +lemma pow_lt_one₀ (h₀ : 0 ≤ a) (h₁ : a < 1) : ∀ {n : ℕ}, n ≠ 0 → a ^ n < 1 + | 0, h => (h rfl).elim + | n + 1, _ => by + rw [pow_succ']; exact mul_lt_one_of_nonneg_of_lt_one_left h₀ h₁ (pow_le_one₀ h₀ h₁.le) + +lemma one_le_pow₀ (ha : 1 ≤ a) : ∀ {n : ℕ}, 1 ≤ a ^ n + | 0 => by rw [pow_zero] + | n + 1 => by + simpa only [pow_succ', mul_one] + using mul_le_mul ha (one_le_pow₀ ha) zero_le_one (zero_le_one.trans ha) + +lemma one_lt_pow₀ (ha : 1 < a) : ∀ {n : ℕ}, n ≠ 0 → 1 < a ^ n + | 0, h => (h rfl).elim + | n + 1, _ => by rw [pow_succ']; exact one_lt_mul_of_lt_of_le ha (one_le_pow₀ ha.le) + +lemma pow_right_mono₀ (h : 1 ≤ a) : Monotone (a ^ ·) := + monotone_nat_of_le_succ fun n => by + rw [pow_succ']; exact le_mul_of_one_le_left (pow_nonneg (zero_le_one.trans h) _) h + +@[gcongr] +lemma pow_le_pow_right₀ (ha : 1 ≤ a) (hmn : m ≤ n) : a ^ m ≤ a ^ n := pow_right_mono₀ ha hmn + +lemma le_self_pow₀ (ha : 1 ≤ a) (hn : n ≠ 0) : a ≤ a ^ n := by + simpa only [pow_one] using pow_le_pow_right₀ ha <| Nat.pos_iff_ne_zero.2 hn + +end + variable [Preorder α] {f g : α → M₀} lemma monotone_mul_left_of_nonneg [PosMulMono M₀] (ha : 0 ≤ a) : Monotone fun x ↦ a * x := @@ -1122,7 +1161,7 @@ lemma div_self_le_one (a : G₀) : a / a ≤ 1 := by obtain rfl | ha := eq_or_ne end Preorder section PartialOrder -variable [PartialOrder G₀] [ZeroLEOneClass G₀] [PosMulReflectLT G₀] {a b c d : G₀} +variable [PartialOrder G₀] [ZeroLEOneClass G₀] [PosMulReflectLT G₀] {a b c : G₀} @[simp] lemma inv_pos : 0 < a⁻¹ ↔ 0 < a := suffices ∀ a : G₀, 0 < a → 0 < a⁻¹ from ⟨fun h ↦ inv_inv a ▸ this _ h, this a⟩ @@ -1170,13 +1209,21 @@ lemma inv_mul_le_iff₀ (hc : 0 < c) : c⁻¹ * b ≤ a ↔ b ≤ c * a where lemma one_le_inv_mul₀ (ha : 0 < a) : 1 ≤ a⁻¹ * b ↔ a ≤ b := by rw [le_inv_mul_iff₀ ha, mul_one] lemma inv_mul_le_one₀ (ha : 0 < a) : a⁻¹ * b ≤ 1 ↔ b ≤ a := by rw [inv_mul_le_iff₀ ha, mul_one] -lemma one_le_inv₀ (ha : 0 < a) : 1 ≤ a⁻¹ ↔ a ≤ 1 := by simpa using one_le_inv_mul₀ ha (b := 1) -lemma inv_le_one₀ (ha : 0 < a) : a⁻¹ ≤ 1 ↔ 1 ≤ a := by simpa using inv_mul_le_one₀ ha (b := 1) - /-- See `inv_le_iff_one_le_mul₀` for a version with multiplication on the other side. -/ lemma inv_le_iff_one_le_mul₀' (ha : 0 < a) : a⁻¹ ≤ b ↔ 1 ≤ a * b := by rw [← inv_mul_le_iff₀ ha, mul_one] +lemma one_le_inv₀ (ha : 0 < a) : 1 ≤ a⁻¹ ↔ a ≤ 1 := by simpa using one_le_inv_mul₀ ha (b := 1) +lemma inv_le_one₀ (ha : 0 < a) : a⁻¹ ≤ 1 ↔ 1 ≤ a := by simpa using inv_mul_le_one₀ ha (b := 1) + +@[bound] +lemma inv_le_one_of_one_le₀ (ha : 1 ≤ a) : a⁻¹ ≤ 1 := (inv_le_one₀ <| zero_lt_one.trans_le ha).2 ha + +lemma one_le_inv_iff₀ : 1 ≤ a⁻¹ ↔ 0 < a ∧ a ≤ 1 where + mp h := ⟨inv_pos.1 (zero_lt_one.trans_le h), + inv_inv a ▸ (inv_le_one₀ <| zero_lt_one.trans_le h).2 h⟩ + mpr h := (one_le_inv₀ h.1).2 h.2 + /-- One direction of `le_inv_mul_iff₀` where `c` is allowed to be `0` (but `b` must be nonnegative). -/ lemma mul_le_of_le_inv_mul₀ (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ c⁻¹ * b) : c * a ≤ b := by @@ -1191,6 +1238,7 @@ lemma inv_mul_le_of_le_mul₀ (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ b * c) : · simp [hc] · rwa [inv_mul_le_iff₀ hb] +@[bound] lemma inv_mul_le_one_of_le₀ (h : a ≤ b) (hb : 0 ≤ b) : b⁻¹ * a ≤ 1 := inv_mul_le_of_le_mul₀ hb zero_le_one <| by rwa [mul_one] @@ -1217,13 +1265,13 @@ lemma le_div_iff₀ (hc : 0 < c) : a ≤ b / c ↔ a * c ≤ b := by lemma div_le_iff₀ (hc : 0 < c) : b / c ≤ a ↔ b ≤ a * c := by rw [div_eq_mul_inv, mul_inv_le_iff₀ hc] -lemma one_le_div₀ (hb : 0 < b) : 1 ≤ a / b ↔ b ≤ a := by rw [le_div_iff₀ hb, one_mul] -lemma div_le_one₀ (hb : 0 < b) : a / b ≤ 1 ↔ a ≤ b := by rw [div_le_iff₀ hb, one_mul] - /-- See `inv_le_iff_one_le_mul₀'` for a version with multiplication on the other side. -/ lemma inv_le_iff_one_le_mul₀ (ha : 0 < a) : a⁻¹ ≤ b ↔ 1 ≤ b * a := by rw [← mul_inv_le_iff₀ ha, one_mul] +lemma one_le_div₀ (hb : 0 < b) : 1 ≤ a / b ↔ b ≤ a := by rw [le_div_iff₀ hb, one_mul] +lemma div_le_one₀ (hb : 0 < b) : a / b ≤ 1 ↔ a ≤ b := by rw [div_le_iff₀ hb, one_mul] + /-- One direction of `le_mul_inv_iff₀` where `c` is allowed to be `0` (but `b` must be nonnegative). -/ lemma mul_le_of_le_mul_inv₀ (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ b * c⁻¹) : a * c ≤ b := by @@ -1246,15 +1294,40 @@ lemma mul_le_of_le_div₀ (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ b / c) : a * lemma div_le_of_le_mul₀ (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ c * b) : a / b ≤ c := div_eq_mul_inv a _ ▸ mul_inv_le_of_le_mul₀ hb hc h +@[bound] lemma mul_inv_le_one_of_le₀ (h : a ≤ b) (hb : 0 ≤ b) : a * b⁻¹ ≤ 1 := mul_inv_le_of_le_mul₀ hb zero_le_one <| by rwa [one_mul] +@[bound] lemma div_le_one_of_le₀ (h : a ≤ b) (hb : 0 ≤ b) : a / b ≤ 1 := div_le_of_le_mul₀ hb zero_le_one <| by rwa [one_mul] @[deprecated (since := "2024-08-21")] alias le_div_iff := le_div_iff₀ @[deprecated (since := "2024-08-21")] alias div_le_iff := div_le_iff₀ +variable [PosMulMono G₀] + +/-- See `inv_anti₀` for the implication from right-to-left with one fewer assumption. -/ +lemma inv_le_inv₀ (ha : 0 < a) (hb : 0 < b) : a⁻¹ ≤ b⁻¹ ↔ b ≤ a := by + rw [inv_le_iff_one_le_mul₀' ha, le_mul_inv_iff₀ hb, one_mul] + +@[gcongr, bound] +lemma inv_anti₀ (hb : 0 < b) (hba : b ≤ a) : a⁻¹ ≤ b⁻¹ := (inv_le_inv₀ (hb.trans_le hba) hb).2 hba + +/-- See also `inv_le_of_inv_le₀` for a one-sided implication with one fewer assumption. -/ +lemma inv_le_comm₀ (ha : 0 < a) (hb : 0 < b) : a⁻¹ ≤ b ↔ b⁻¹ ≤ a := by + rw [← inv_le_inv₀ hb (inv_pos.2 ha), inv_inv] + +lemma inv_le_of_inv_le₀ (ha : 0 < a) (h : a⁻¹ ≤ b) : b⁻¹ ≤ a := + (inv_le_comm₀ ha <| (inv_pos.2 ha).trans_le h).1 h + +/-- See also `le_inv_of_le_inv₀` for a one-sided implication with one fewer assumption. -/ +lemma le_inv_comm₀ (ha : 0 < a) (hb : 0 < b) : a ≤ b⁻¹ ↔ b ≤ a⁻¹ := by + rw [← inv_le_inv₀ (inv_pos.2 hb) ha, inv_inv] + +lemma le_inv_of_le_inv₀ (ha : 0 < a) (h : a ≤ b⁻¹) : b ≤ a⁻¹ := + (le_inv_comm₀ ha <| inv_pos.1 <| ha.trans_le h).1 h + end MulPosMono section PosMulStrictMono @@ -1270,15 +1343,22 @@ lemma inv_mul_lt_iff₀ (hc : 0 < c) : c⁻¹ * b < a ↔ b < c * a where mp h := by simpa [hc.ne'] using mul_lt_mul_of_pos_left h hc mpr h := by simpa [hc.ne'] using mul_lt_mul_of_pos_left h (inv_pos.2 hc) +/-- See `inv_lt_iff_one_lt_mul₀` for a version with multiplication on the other side. -/ +lemma inv_lt_iff_one_lt_mul₀' (ha : 0 < a) : a⁻¹ < b ↔ 1 < a * b := by + rw [← inv_mul_lt_iff₀ ha, mul_one] + lemma one_lt_inv_mul₀ (ha : 0 < a) : 1 < a⁻¹ * b ↔ a < b := by rw [lt_inv_mul_iff₀ ha, mul_one] lemma inv_mul_lt_one₀ (ha : 0 < a) : a⁻¹ * b < 1 ↔ b < a := by rw [inv_mul_lt_iff₀ ha, mul_one] lemma one_lt_inv₀ (ha : 0 < a) : 1 < a⁻¹ ↔ a < 1 := by simpa using one_lt_inv_mul₀ ha (b := 1) lemma inv_lt_one₀ (ha : 0 < a) : a⁻¹ < 1 ↔ 1 < a := by simpa using inv_mul_lt_one₀ ha (b := 1) -/-- See `inv_lt_iff_one_lt_mul₀` for a version with multiplication on the other side. -/ -lemma inv_lt_iff_one_lt_mul₀' (ha : 0 < a) : a⁻¹ < b ↔ 1 < a * b := by - rw [← inv_mul_lt_iff₀ ha, mul_one] +@[bound] +lemma inv_lt_one_of_one_lt₀ (ha : 1 < a) : a⁻¹ < 1 := (inv_lt_one₀ <| zero_lt_one.trans ha).2 ha + +lemma one_lt_inv_iff₀ : 1 < a⁻¹ ↔ 0 < a ∧ a < 1 where + mp h := ⟨inv_pos.1 (zero_lt_one.trans h), inv_inv a ▸ (inv_lt_one₀ <| zero_lt_one.trans h).2 h⟩ + mpr h := (one_lt_inv₀ h.1).2 h.2 end PosMulStrictMono @@ -1307,11 +1387,35 @@ lemma div_lt_iff₀ (hc : 0 < c) : b / c < a ↔ b < a * c := by lemma inv_lt_iff_one_lt_mul₀ (ha : 0 < a) : a⁻¹ < b ↔ 1 < b * a := by rw [← mul_inv_lt_iff₀ ha, one_mul] +variable [PosMulStrictMono G₀] + +/-- See `inv_strictAnti₀` for the implication from right-to-left with one fewer assumption. -/ +lemma inv_lt_inv₀ (ha : 0 < a) (hb : 0 < b) : a⁻¹ < b⁻¹ ↔ b < a := by + rw [inv_lt_iff_one_lt_mul₀' ha, lt_mul_inv_iff₀ hb, one_mul] + +@[gcongr, bound] +lemma inv_strictAnti₀ (hb : 0 < b) (hba : b < a) : a⁻¹ < b⁻¹ := + (inv_lt_inv₀ (hb.trans hba) hb).2 hba + +/-- See also `inv_lt_of_inv_lt₀` for a one-sided implication with one fewer assumption. -/ +lemma inv_lt_comm₀ (ha : 0 < a) (hb : 0 < b) : a⁻¹ < b ↔ b⁻¹ < a := by + rw [← inv_lt_inv₀ hb (inv_pos.2 ha), inv_inv] + +lemma inv_lt_of_inv_lt₀ (ha : 0 < a) (h : a⁻¹ < b) : b⁻¹ < a := + (inv_lt_comm₀ ha <| (inv_pos.2 ha).trans h).1 h + +/-- See also `lt_inv_of_lt_inv₀` for a one-sided implication with one fewer assumption. -/ +lemma lt_inv_comm₀ (ha : 0 < a) (hb : 0 < b) : a < b⁻¹ ↔ b < a⁻¹ := by + rw [← inv_lt_inv₀ (inv_pos.2 hb) ha, inv_inv] + +lemma lt_inv_of_lt_inv₀ (ha : 0 < a) (h : a < b⁻¹) : b < a⁻¹ := + (lt_inv_comm₀ ha <| inv_pos.1 <| ha.trans h).1 h + end MulPosStrictMono end PartialOrder section LinearOrder -variable [LinearOrder G₀] [ZeroLEOneClass G₀] [PosMulReflectLT G₀] {a b c : G₀} +variable [LinearOrder G₀] [ZeroLEOneClass G₀] [PosMulReflectLT G₀] {a b : G₀} @[simp] lemma inv_neg'' : a⁻¹ < 0 ↔ a < 0 := by simp only [← not_le, inv_nonneg] @[simp] lemma inv_nonpos : a⁻¹ ≤ 0 ↔ a ≤ 0 := by simp only [← not_lt, inv_pos] @@ -1324,6 +1428,12 @@ lemma one_div_nonpos : 1 / a ≤ 0 ↔ a ≤ 0 := one_div a ▸ inv_nonpos lemma div_nonpos_of_nonneg_of_nonpos [PosMulMono G₀] (ha : 0 ≤ a) (hb : b ≤ 0) : a / b ≤ 0 := by rw [div_eq_mul_inv]; exact mul_nonpos_of_nonneg_of_nonpos ha (inv_nonpos.2 hb) +lemma inv_lt_one_iff₀ [PosMulStrictMono G₀] : a⁻¹ < 1 ↔ a ≤ 0 ∨ 1 < a := by + simp_rw [← not_le, one_le_inv_iff₀, not_and_or, not_lt] + +lemma inv_le_one_iff₀ [PosMulStrictMono G₀] : a⁻¹ ≤ 1 ↔ a ≤ 0 ∨ 1 ≤ a := by + simp only [← not_lt, one_lt_inv_iff₀, not_and_or] + end GroupWithZero.LinearOrder section CommSemigroupHasZero @@ -1438,3 +1548,5 @@ lemma div_lt_comm₀ (hb : 0 < b) (hc : 0 < c) : a / b < c ↔ a / c < b := by end PosMulStrictMono end CommGroupWithZero + +set_option linter.style.longFile 1700 diff --git a/Mathlib/Algebra/Order/Interval/Set/Instances.lean b/Mathlib/Algebra/Order/Interval/Set/Instances.lean index da5a81838c6e1..4624971afef4b 100644 --- a/Mathlib/Algebra/Order/Interval/Set/Instances.lean +++ b/Mathlib/Algebra/Order/Interval/Set/Instances.lean @@ -3,7 +3,8 @@ Copyright (c) 2022 Stuart Presnell. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Stuart Presnell, Eric Wieser, Yaël Dillies, Patrick Massot, Kim Morrison -/ -import Mathlib.Algebra.Order.Ring.Basic +import Mathlib.Algebra.GroupWithZero.InjSurj +import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Algebra.Ring.Regular import Mathlib.Order.Interval.Set.Basic @@ -98,7 +99,7 @@ theorem le_one {t : Icc (0 : α) 1} : t ≤ 1 := t.2.2 instance mul : Mul (Icc (0 : α) 1) where - mul p q := ⟨p * q, ⟨mul_nonneg p.2.1 q.2.1, mul_le_one p.2.2 q.2.1 q.2.2⟩⟩ + mul p q := ⟨p * q, ⟨mul_nonneg p.2.1 q.2.1, mul_le_one₀ p.2.2 q.2.1 q.2.2⟩⟩ instance pow : Pow (Icc (0 : α) 1) ℕ where pow p n := ⟨p.1 ^ n, ⟨pow_nonneg p.2.1 n, pow_le_one₀ p.2.1 p.2.2⟩⟩ @@ -236,7 +237,7 @@ theorem le_one {t : Ioc (0 : α) 1} : t ≤ 1 := t.2.2 instance mul : Mul (Ioc (0 : α) 1) where - mul p q := ⟨p.1 * q.1, ⟨mul_pos p.2.1 q.2.1, mul_le_one p.2.2 (le_of_lt q.2.1) q.2.2⟩⟩ + mul p q := ⟨p.1 * q.1, ⟨mul_pos p.2.1 q.2.1, mul_le_one₀ p.2.2 (le_of_lt q.2.1) q.2.2⟩⟩ instance pow : Pow (Ioc (0 : α) 1) ℕ where pow p n := ⟨p.1 ^ n, ⟨pow_pos p.2.1 n, pow_le_one₀ (le_of_lt p.2.1) p.2.2⟩⟩ diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean index 5968a01cd301c..4425eebe275f4 100644 --- a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean @@ -58,7 +58,7 @@ instance (priority := 100) CanonicallyOrderedCommMonoid.existsMulOfLE (α : Type section CanonicallyOrderedCommMonoid -variable [CanonicallyOrderedCommMonoid α] {a b c d : α} +variable [CanonicallyOrderedCommMonoid α] {a b c : α} @[to_additive] theorem le_self_mul : a ≤ a * c := diff --git a/Mathlib/Algebra/Order/Monoid/Defs.lean b/Mathlib/Algebra/Order/Monoid/Defs.lean index d04efccbdc0e8..32322b7807649 100644 --- a/Mathlib/Algebra/Order/Monoid/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Defs.lean @@ -15,7 +15,7 @@ This file provides the definitions of ordered monoids. open Function -variable {α β : Type*} +variable {α : Type*} /-- An ordered (additive) commutative monoid is a commutative monoid with a partial order such that addition is monotone. -/ diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Defs.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Defs.lean index 410847d5fa72a..911093e32c699 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Defs.lean @@ -328,7 +328,7 @@ theorem Group.mulRightReflectLT_of_mulRightStrictMono [Group N] [LT N] section Trans -variable [IsTrans N r] (m n : M) {a b c d : N} +variable [IsTrans N r] (m : M) {a b c : N} -- Lemmas with 3 elements. theorem act_rel_of_rel_of_act_rel (ab : r a b) (rl : r (μ m b) c) : r (μ m a) c := @@ -361,7 +361,7 @@ theorem rel_of_act_rel_act (m : M) {a b : N} (ab : r (μ m a) (μ m b)) : r a b section Trans -variable [IsTrans N r] (m n : M) {a b c d : N} +variable [IsTrans N r] (m : M) {a b c : N} -- Lemmas with 3 elements. theorem act_rel_of_act_rel_of_rel_act_rel (ab : r (μ m a) b) (rl : r (μ m b) (μ m c)) : diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean index a2140c0410d85..1765144a330e3 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean @@ -26,7 +26,7 @@ variable [Preorder M] namespace Left -variable [CovariantClass M M (· * ·) (· ≤ ·)] {a b : M} +variable [MulLeftMono M] {a : M} @[to_additive Left.nsmul_nonneg] theorem one_le_pow_of_le (ha : 1 ≤ a) : ∀ n : ℕ, 1 ≤ a ^ n @@ -58,7 +58,7 @@ end Left section Left -variable [MulLeftMono M] {x : M} +variable [MulLeftMono M] @[to_additive nsmul_left_monotone] theorem pow_right_monotone {a : M} (ha : 1 ≤ a) : Monotone fun n : ℕ ↦ a ^ n := diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean index b926fcd7773f6..e0b7747198a3b 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean @@ -80,7 +80,7 @@ end One section Add -variable [Add α] {a b c d : WithTop α} {x y : α} +variable [Add α] {a b c d : WithTop α} {x : α} instance add : Add (WithTop α) := ⟨Option.map₂ (· + ·)⟩ diff --git a/Mathlib/Algebra/Order/Monovary.lean b/Mathlib/Algebra/Order/Monovary.lean index f742a97bcf60e..e085b565b3140 100644 --- a/Mathlib/Algebra/Order/Monovary.lean +++ b/Mathlib/Algebra/Order/Monovary.lean @@ -238,19 +238,19 @@ variable [LinearOrderedSemifield α] [LinearOrderedSemifield β] {s : Set ι} {f @[simp] lemma monovaryOn_inv_left₀ (hf : ∀ i ∈ s, 0 < f i) : MonovaryOn f⁻¹ g s ↔ AntivaryOn f g s := - forall₅_congr fun _i hi _j hj _ ↦ inv_le_inv (hf _ hi) (hf _ hj) + forall₅_congr fun _i hi _j hj _ ↦ inv_le_inv₀ (hf _ hi) (hf _ hj) @[simp] lemma antivaryOn_inv_left₀ (hf : ∀ i ∈ s, 0 < f i) : AntivaryOn f⁻¹ g s ↔ MonovaryOn f g s := - forall₅_congr fun _i hi _j hj _ ↦ inv_le_inv (hf _ hj) (hf _ hi) + forall₅_congr fun _i hi _j hj _ ↦ inv_le_inv₀ (hf _ hj) (hf _ hi) @[simp] lemma monovaryOn_inv_right₀ (hg : ∀ i ∈ s, 0 < g i) : MonovaryOn f g⁻¹ s ↔ AntivaryOn f g s := - forall₂_swap.trans <| forall₄_congr fun i hi j hj ↦ by erw [inv_lt_inv (hg _ hj) (hg _ hi)] + forall₂_swap.trans <| forall₄_congr fun i hi j hj ↦ by erw [inv_lt_inv₀ (hg _ hj) (hg _ hi)] @[simp] lemma antivaryOn_inv_right₀ (hg : ∀ i ∈ s, 0 < g i) : AntivaryOn f g⁻¹ s ↔ MonovaryOn f g s := - forall₂_swap.trans <| forall₄_congr fun i hi j hj ↦ by erw [inv_lt_inv (hg _ hj) (hg _ hi)] + forall₂_swap.trans <| forall₄_congr fun i hi j hj ↦ by erw [inv_lt_inv₀ (hg _ hj) (hg _ hi)] lemma monovaryOn_inv₀ (hf : ∀ i ∈ s, 0 < f i) (hg : ∀ i ∈ s, 0 < g i) : MonovaryOn f⁻¹ g⁻¹ s ↔ MonovaryOn f g s := by @@ -260,16 +260,16 @@ lemma antivaryOn_inv₀ (hf : ∀ i ∈ s, 0 < f i) (hg : ∀ i ∈ s, 0 < g i) rw [antivaryOn_inv_left₀ hf, monovaryOn_inv_right₀ hg] @[simp] lemma monovary_inv_left₀ (hf : StrongLT 0 f) : Monovary f⁻¹ g ↔ Antivary f g := - forall₃_congr fun _i _j _ ↦ inv_le_inv (hf _) (hf _) + forall₃_congr fun _i _j _ ↦ inv_le_inv₀ (hf _) (hf _) @[simp] lemma antivary_inv_left₀ (hf : StrongLT 0 f) : Antivary f⁻¹ g ↔ Monovary f g := - forall₃_congr fun _i _j _ ↦ inv_le_inv (hf _) (hf _) + forall₃_congr fun _i _j _ ↦ inv_le_inv₀ (hf _) (hf _) @[simp] lemma monovary_inv_right₀ (hg : StrongLT 0 g) : Monovary f g⁻¹ ↔ Antivary f g := - forall_swap.trans <| forall₂_congr fun i j ↦ by erw [inv_lt_inv (hg _) (hg _)] + forall_swap.trans <| forall₂_congr fun i j ↦ by erw [inv_lt_inv₀ (hg _) (hg _)] @[simp] lemma antivary_inv_right₀ (hg : StrongLT 0 g) : Antivary f g⁻¹ ↔ Monovary f g := - forall_swap.trans <| forall₂_congr fun i j ↦ by erw [inv_lt_inv (hg _) (hg _)] + forall_swap.trans <| forall₂_congr fun i j ↦ by erw [inv_lt_inv₀ (hg _) (hg _)] lemma monovary_inv₀ (hf : StrongLT 0 f) (hg : StrongLT 0 g) : Monovary f⁻¹ g⁻¹ ↔ Monovary f g := by rw [monovary_inv_left₀ hf, antivary_inv_right₀ hg] diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index 9f307d96d0851..81771b41efaf0 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -37,10 +37,6 @@ section OrderedSemiring variable [OrderedSemiring R] {a b x y : R} {n m : ℕ} -theorem zero_pow_le_one : ∀ n : ℕ, (0 : R) ^ n ≤ 1 - | 0 => (pow_zero _).le - | n + 1 => by rw [zero_pow n.succ_ne_zero]; exact zero_le_one - theorem pow_add_pow_le (hx : 0 ≤ x) (hy : 0 ≤ y) (hn : n ≠ 0) : x ^ n + y ^ n ≤ (x + y) ^ n := by rcases Nat.exists_eq_add_one_of_ne_zero hn with ⟨k, rfl⟩ induction k with @@ -60,48 +56,21 @@ theorem pow_add_pow_le (hx : 0 ≤ x) (hy : 0 ≤ y) (hn : n ≠ 0) : x ^ n + y rw [pow_succ' _ n] exact mul_le_mul_of_nonneg_left (ih (Nat.succ_ne_zero k)) h2 -@[bound] -theorem pow_le_one₀ : ∀ {n : ℕ}, 0 ≤ a → a ≤ 1 → a ^ n ≤ 1 - | 0, _, _ => (pow_zero a).le - | n + 1, h₀, h₁ => (pow_succ a n).le.trans (mul_le_one (pow_le_one₀ h₀ h₁) h₀ h₁) - -theorem pow_lt_one₀ (h₀ : 0 ≤ a) (h₁ : a < 1) : ∀ {n : ℕ}, n ≠ 0 → a ^ n < 1 - | 0, h => (h rfl).elim - | n + 1, _ => by - rw [pow_succ'] - exact mul_lt_one_of_nonneg_of_lt_one_left h₀ h₁ (pow_le_one₀ h₀ h₁.le) - -@[bound] -theorem one_le_pow₀ (H : 1 ≤ a) : ∀ {n : ℕ}, 1 ≤ a ^ n - | 0 => by rw [pow_zero] - | n + 1 => by - simpa only [pow_succ', mul_one] - using mul_le_mul H (one_le_pow₀ H) zero_le_one (zero_le_one.trans H) - -lemma one_lt_pow₀ (ha : 1 < a) : ∀ {n : ℕ}, n ≠ 0 → 1 < a ^ n - | 0, h => (h rfl).elim - | n + 1, _ => by rw [pow_succ']; exact one_lt_mul_of_lt_of_le ha (one_le_pow₀ ha.le) +attribute [bound] pow_le_one₀ one_le_pow₀ +@[deprecated (since := "2024-09-28")] alias mul_le_one := mul_le_one₀ @[deprecated (since := "2024-09-28")] alias pow_le_one := pow_le_one₀ @[deprecated (since := "2024-09-28")] alias pow_lt_one := pow_lt_one₀ @[deprecated (since := "2024-09-28")] alias one_le_pow_of_one_le := one_le_pow₀ @[deprecated (since := "2024-09-28")] alias one_lt_pow := one_lt_pow₀ - -theorem pow_right_mono (h : 1 ≤ a) : Monotone (a ^ ·) := - monotone_nat_of_le_succ fun n => by - rw [pow_succ'] - exact le_mul_of_one_le_left (pow_nonneg (zero_le_one.trans h) _) h - -@[gcongr] -theorem pow_le_pow_right (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := pow_right_mono ha h - -theorem le_self_pow (ha : 1 ≤ a) (h : m ≠ 0) : a ≤ a ^ m := by - simpa only [pow_one] using pow_le_pow_right ha <| Nat.pos_iff_ne_zero.2 h +@[deprecated (since := "2024-10-04")] alias pow_right_mono := pow_right_mono₀ +@[deprecated (since := "2024-10-04")] alias pow_le_pow_right := pow_le_pow_right₀ +@[deprecated (since := "2024-10-04")] alias le_self_pow := le_self_pow₀ /-- The `bound` tactic can't handle `m ≠ 0` goals yet, so we express as `0 < m` -/ @[bound] lemma Bound.le_self_pow_of_pos {m : ℕ} (ha : 1 ≤ a) (h : 0 < m) : a ≤ a ^ m := - le_self_pow ha h.ne' + le_self_pow₀ ha h.ne' @[mono, gcongr, bound] theorem pow_le_pow_left {a b : R} (ha : 0 ≤ a) (hab : a ≤ b) : ∀ n, a ^ n ≤ b ^ n @@ -119,7 +88,7 @@ lemma pow_add_pow_le' (ha : 0 ≤ a) (hb : 0 ≤ b) : a ^ n + b ^ n ≤ 2 * (a + lemma Bound.pow_le_pow_right_of_le_one_or_one_le (h : 1 ≤ a ∧ n ≤ m ∨ 0 ≤ a ∧ a ≤ 1 ∧ m ≤ n) : a ^ n ≤ a ^ m := by rcases h with ⟨a1, nm⟩ | ⟨a0, a1, mn⟩ - · exact pow_le_pow_right a1 nm + · exact pow_right_mono₀ a1 nm · exact pow_le_pow_of_le_one a0 a1 mn end OrderedSemiring diff --git a/Mathlib/Algebra/Order/Sub/Defs.lean b/Mathlib/Algebra/Order/Sub/Defs.lean index 9dfdf76181e73..5cb6b7acd468c 100644 --- a/Mathlib/Algebra/Order/Sub/Defs.lean +++ b/Mathlib/Algebra/Order/Sub/Defs.lean @@ -41,7 +41,7 @@ TODO: generalize `Nat.le_of_le_of_sub_le_sub_right`, `Nat.sub_le_sub_right_iff`, -/ -variable {α β : Type*} +variable {α : Type*} /-- `OrderedSub α` means that `α` has a subtraction characterized by `a - b ≤ c ↔ a ≤ c + b`. In other words, `a - b` is the least `c` such that `a ≤ b + c`. @@ -60,7 +60,7 @@ theorem tsub_le_iff_right [LE α] [Add α] [Sub α] [OrderedSub α] {a b c : α} a - b ≤ c ↔ a ≤ c + b := OrderedSub.tsub_le_iff_right a b c -variable [Preorder α] [Add α] [Sub α] [OrderedSub α] {a b c d : α} +variable [Preorder α] [Add α] [Sub α] [OrderedSub α] {a b : α} /-- See `add_tsub_cancel_right` for the equality if `AddLeftReflectLE α`. -/ theorem add_tsub_le_right : a + b - b ≤ a := @@ -210,7 +210,7 @@ end Contra end AddCommSemigroup -variable [AddCommMonoid α] [Sub α] [OrderedSub α] {a b c d : α} +variable [AddCommMonoid α] [Sub α] [OrderedSub α] {a b : α} theorem tsub_nonpos : a - b ≤ 0 ↔ a ≤ b := by rw [tsub_le_iff_left, add_zero] @@ -375,7 +375,7 @@ end OrderedAddCommSemigroup section LinearOrder -variable {a b c d : α} [LinearOrder α] [AddCommSemigroup α] [Sub α] [OrderedSub α] +variable {a b c : α} [LinearOrder α] [AddCommSemigroup α] [Sub α] [OrderedSub α] /-- See `lt_of_tsub_lt_tsub_right_of_le` for a weaker statement in a partial order. -/ theorem lt_of_tsub_lt_tsub_right (h : a - c < b - c) : a < b := diff --git a/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean b/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean index 02b66c70ad0b8..6d9cedb09fb7c 100644 --- a/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean +++ b/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean @@ -16,7 +16,7 @@ variable {α β : Type*} section Add -variable [Preorder α] [Add α] [Sub α] [OrderedSub α] {a b c d : α} +variable [Preorder α] [Add α] [Sub α] [OrderedSub α] theorem AddHom.le_map_tsub [Preorder β] [Add β] [Sub β] [OrderedSub β] (f : AddHom α β) (hf : Monotone f) (a b : α) : f a - f b ≤ f (a - b) := by @@ -49,7 +49,7 @@ theorem OrderIso.map_tsub {M N : Type*} [Preorder M] [Add M] [Sub M] [OrderedSub section Preorder variable [Preorder α] -variable [AddCommMonoid α] [Sub α] [OrderedSub α] {a b c d : α} +variable [AddCommMonoid α] [Sub α] [OrderedSub α] theorem AddMonoidHom.le_map_tsub [Preorder β] [AddCommMonoid β] [Sub β] [OrderedSub β] (f : α →+ β) (hf : Monotone f) (a b : α) : f a - f b ≤ f (a - b) := diff --git a/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean b/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean index 01b1c8ec0561e..8b5296df327a3 100644 --- a/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean +++ b/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean @@ -61,7 +61,7 @@ noncomputable def cardPowDegree : AbsoluteValue Fq[X] ℤ := · simp only [hpq, hp, hq, eq_self_iff_true, if_true, if_false] exact add_nonneg (pow_pos _).le (pow_pos _).le simp only [hpq, hp, hq, if_false] - refine le_trans (pow_le_pow_right (by omega) (Polynomial.natDegree_add_le _ _)) ?_ + refine le_trans (pow_right_mono₀ (by omega) (Polynomial.natDegree_add_le _ _)) ?_ refine le_trans (le_max_iff.mpr ?_) (max_le_add_of_nonneg (pow_nonneg (by omega) _) (pow_nonneg (by omega) _)) diff --git a/Mathlib/Algebra/Polynomial/Expand.lean b/Mathlib/Algebra/Polynomial/Expand.lean index e1db94a64141b..ae212d4a08496 100644 --- a/Mathlib/Algebra/Polynomial/Expand.lean +++ b/Mathlib/Algebra/Polynomial/Expand.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ import Mathlib.RingTheory.Polynomial.Basic -import Mathlib.RingTheory.LocalRing.RingHom.Basic /-! # Expand a polynomial by a factor of p, so `∑ aₙ xⁿ` becomes `∑ aₙ xⁿᵖ`. @@ -269,9 +268,8 @@ section IsDomain variable (R : Type u) [CommRing R] [IsDomain R] -theorem isLocalRingHom_expand {p : ℕ} (hp : 0 < p) : - IsLocalRingHom (↑(expand R p) : R[X] →+* R[X]) := by - refine ⟨fun f hf1 => ?_⟩; norm_cast at hf1 +theorem isLocalRingHom_expand {p : ℕ} (hp : 0 < p) : IsLocalRingHom (expand R p) := by + refine ⟨fun f hf1 => ?_⟩ have hf2 := eq_C_of_degree_eq_zero (degree_eq_zero_of_isUnit hf1) rw [coeff_expand hp, if_pos (dvd_zero _), p.zero_div] at hf2 rw [hf2, isUnit_C] at hf1; rw [expand_eq_C hp] at hf2; rwa [hf2, isUnit_C] @@ -281,7 +279,7 @@ variable {R} theorem of_irreducible_expand {p : ℕ} (hp : p ≠ 0) {f : R[X]} (hf : Irreducible (expand R p f)) : Irreducible f := let _ := isLocalRingHom_expand R hp.bot_lt - of_irreducible_map (↑(expand R p)) hf + hf.of_map theorem of_irreducible_expand_pow {p : ℕ} (hp : p ≠ 0) {f : R[X]} {n : ℕ} : Irreducible (expand R (p ^ n) f) → Irreducible f := diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index ab1ca1c76d94d..87a240f259415 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -548,6 +548,26 @@ theorem rootMultiplicity_mul' {p q : R[X]} {x : R} theorem Monic.comp_X_sub_C {p : R[X]} (hp : p.Monic) (r : R) : (p.comp (X - C r)).Monic := by simpa using hp.comp_X_add_C (-r) +@[simp] +theorem comp_neg_X_leadingCoeff_eq (p : R[X]) : + (p.comp (-X)).leadingCoeff = (-1) ^ p.natDegree * p.leadingCoeff := by + nontriviality R + by_cases h : p = 0 + · simp [h] + rw [Polynomial.leadingCoeff, natDegree_comp_eq_of_mul_ne_zero, coeff_comp_degree_mul_degree] <;> + simp [mul_comm, h] + +theorem Monic.neg_one_pow_natDegree_mul_comp_neg_X {p : R[X]} (hp : p.Monic) : + ((-1) ^ p.natDegree * p.comp (-X)).Monic := by + simp only [Monic] + calc + ((-1) ^ p.natDegree * p.comp (-X)).leadingCoeff = + (p.comp (-X) * C ((-1) ^ p.natDegree)).leadingCoeff := by + simp [mul_comm] + _ = 1 := by + apply monic_mul_C_of_leadingCoeff_mul_eq_one + simp [← pow_add, hp] + variable [IsDomain R] {p q : R[X]} @[simp] diff --git a/Mathlib/Algebra/Polynomial/Smeval.lean b/Mathlib/Algebra/Polynomial/Smeval.lean index 2eb2eecb587a6..d24a7089bfc93 100644 --- a/Mathlib/Algebra/Polynomial/Smeval.lean +++ b/Mathlib/Algebra/Polynomial/Smeval.lean @@ -180,7 +180,7 @@ the defining structures independently. For non-associative power-associative al octonions), we replace the `[Semiring S]` with `[NonAssocSemiring S] [Pow S ℕ] [NatPowAssoc S]`. -/ -variable (R : Type*) [Semiring R] {p : R[X]} (r : R) (p q : R[X]) {S : Type*} +variable (R : Type*) [Semiring R] (r : R) (p q : R[X]) {S : Type*} [NonAssocSemiring S] [Module R S] [Pow S ℕ] (x : S) theorem smeval_C_mul : (C r * p).smeval x = r • p.smeval x := by diff --git a/Mathlib/Algebra/Quandle.lean b/Mathlib/Algebra/Quandle.lean index a3921e4072e72..f037884dabe10 100644 --- a/Mathlib/Algebra/Quandle.lean +++ b/Mathlib/Algebra/Quandle.lean @@ -100,9 +100,9 @@ class Shelf (α : Type u) where A *unital shelf* is a shelf equipped with an element `1` such that, for all elements `x`, we have both `x ◃ 1` and `1 ◃ x` equal `x`. -/ -class UnitalShelf (α : Type u) extends Shelf α, One α := -(one_act : ∀ a : α, act 1 a = a) -(act_one : ∀ a : α, act a 1 = a) +class UnitalShelf (α : Type u) extends Shelf α, One α where + one_act : ∀ a : α, act 1 a = a + act_one : ∀ a : α, act a 1 = a /-- The type of homomorphisms between shelves. This is also the notion of rack and quandle homomorphisms. diff --git a/Mathlib/Algebra/Ring/Action/Basic.lean b/Mathlib/Algebra/Ring/Action/Basic.lean index 25558ac78f579..d085e72efe2c3 100644 --- a/Mathlib/Algebra/Ring/Action/Basic.lean +++ b/Mathlib/Algebra/Ring/Action/Basic.lean @@ -42,8 +42,8 @@ class MulSemiringAction (M : Type u) (R : Type v) [Monoid M] [Semiring R] extend section Semiring -variable (M N G : Type*) [Monoid M] [Monoid N] [Group G] -variable (A R S F : Type v) [AddMonoid A] [Semiring R] [CommSemiring S] +variable (M N : Type*) [Monoid M] [Monoid N] +variable (R : Type v) [Semiring R] -- note we could not use `extends` since these typeclasses are made with `old_structure_cmd` instance (priority := 100) MulSemiringAction.toMulDistribMulAction [h : MulSemiringAction M R] : @@ -92,8 +92,6 @@ end section SimpLemmas -variable {M G A R F} - attribute [simp] smul_one smul_mul' smul_zero smul_add end SimpLemmas diff --git a/Mathlib/Algebra/Ring/Commute.lean b/Mathlib/Algebra/Ring/Commute.lean index 54cf0da485262..27887a07dc5bb 100644 --- a/Mathlib/Algebra/Ring/Commute.lean +++ b/Mathlib/Algebra/Ring/Commute.lean @@ -21,9 +21,9 @@ For the definitions of semirings and rings see `Mathlib.Algebra.Ring.Defs`. -/ -universe u v w x +universe u -variable {α : Type u} {β : Type v} {γ : Type w} {R : Type x} +variable {R : Type u} open Function diff --git a/Mathlib/Algebra/Ring/Defs.lean b/Mathlib/Algebra/Ring/Defs.lean index 171d45485e20d..8ff6a13b2f8f9 100644 --- a/Mathlib/Algebra/Ring/Defs.lean +++ b/Mathlib/Algebra/Ring/Defs.lean @@ -45,9 +45,9 @@ assert_not_exists DivisionMonoid.toDivInvOneMonoid assert_not_exists mul_rotate -universe u v w x +universe u v -variable {α : Type u} {β : Type v} {γ : Type w} {R : Type x} +variable {α : Type u} {R : Type v} open Function @@ -250,7 +250,7 @@ instance (priority := 100) CommSemiring.toCommMonoidWithZero [CommSemiring α] : section CommSemiring -variable [CommSemiring α] {a b c : α} +variable [CommSemiring α] theorem add_mul_self_eq (a b : α) : (a + b) * (a + b) = a * a + 2 * a * b + b * b := by simp only [two_mul, add_mul, mul_add, add_assoc, mul_comm b] diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index eea3f38baa20c..51c833ecb1d78 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -5,7 +5,6 @@ Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Opposite -import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Algebra.GroupWithZero.InjSurj import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Logic.Equiv.Set @@ -409,7 +408,7 @@ end Opposite section NonUnitalSemiring -variable [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] (f : R ≃+* S) (x y : R) +variable [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] (f : R ≃+* S) (x : R) /-- A ring isomorphism sends zero to zero. -/ protected theorem map_zero : f 0 = 0 := @@ -533,7 +532,7 @@ end NonUnitalSemiring section Semiring -variable [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) (x y : R) +variable [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) (x : R) /-- A ring isomorphism sends one to one. -/ protected theorem map_one : f 1 = 1 := @@ -604,7 +603,7 @@ end NonUnitalRing section Ring -variable [NonAssocRing R] [NonAssocRing S] (f : R ≃+* S) (x y : R) +variable [NonAssocRing R] [NonAssocRing S] (f : R ≃+* S) -- Porting note (#10618): `simp` can now prove that, so we remove the `@[simp]` tag theorem map_neg_one : f (-1) = -1 := @@ -805,9 +804,6 @@ protected theorem map_pow (f : R ≃+* S) (a) : ∀ n : ℕ, f (a ^ n) = f a ^ n end GroupPower -protected theorem isUnit_iff (f : R ≃+* S) {a} : IsUnit (f a) ↔ IsUnit a := - MulEquiv.map_isUnit_iff f - end RingEquiv namespace MulEquiv diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index d257486efbda8..9205b9aeff65a 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -233,14 +233,10 @@ lemma even_or_odd (n : ℕ) : Even n ∨ Odd n := (even_xor_odd n).or lemma even_or_odd' (n : ℕ) : ∃ k, n = 2 * k ∨ n = 2 * k + 1 := by simpa only [← two_mul, exists_or, Odd, Even] using even_or_odd n -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ lemma even_xor_odd' (n : ℕ) : ∃ k, Xor' (n = 2 * k) (n = 2 * k + 1) := by obtain ⟨k, rfl⟩ | ⟨k, rfl⟩ := even_or_odd n <;> use k · simpa only [← two_mul, eq_self_iff_true, xor_true] using (succ_ne_self (2 * k)).symm - · simpa only [xor_true, _root_.xor_comm] using (succ_ne_self _) + · simpa only [xor_true, xor_comm] using (succ_ne_self _) lemma mod_two_add_add_odd_mod_two (m : ℕ) {n : ℕ} (hn : Odd n) : m % 2 + (m + n) % 2 = 1 := ((even_or_odd m).elim fun hm ↦ by rw [even_iff.1 hm, odd_iff.1 (hm.add_odd hn)]) fun hm ↦ by diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index f775d6a04fdda..3fd3c93c9cef4 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -27,7 +27,7 @@ section AddSubmonoidWithOneClass /-- `AddSubmonoidWithOneClass S R` says `S` is a type of subsets `s ≤ R` that contain `0`, `1`, and are closed under `(+)` -/ -class AddSubmonoidWithOneClass (S R : Type*) [AddMonoidWithOne R] +class AddSubmonoidWithOneClass (S : Type*) (R : outParam Type*) [AddMonoidWithOne R] [SetLike S R] extends AddSubmonoidClass S R, OneMemClass S R : Prop variable {S R : Type*} [AddMonoidWithOne R] [SetLike S R] (s : S) diff --git a/Mathlib/Algebra/Ring/Units.lean b/Mathlib/Algebra/Ring/Units.lean index 3e177e160d871..619def14cb01a 100644 --- a/Mathlib/Algebra/Ring/Units.lean +++ b/Mathlib/Algebra/Ring/Units.lean @@ -15,7 +15,7 @@ import Mathlib.Algebra.Ring.Hom.Defs universe u v w x -variable {α : Type u} {β : Type v} {γ : Type w} {R : Type x} +variable {α : Type u} {β : Type v} {R : Type x} open Function @@ -23,7 +23,7 @@ namespace Units section HasDistribNeg -variable [Monoid α] [HasDistribNeg α] {a b : α} +variable [Monoid α] [HasDistribNeg α] /-- Each element of the group of units of a ring has an additive inverse. -/ instance : Neg αˣ := diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index 4eb552d919fc2..c40ab975c18d7 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -147,6 +147,12 @@ theorem snd_comp_inr [Zero R] : snd ∘ (inr : M → tsze R M) = id := end +theorem fst_surjective [Nonempty M] : Function.Surjective (fst : tsze R M → R) := + Prod.fst_surjective + +theorem snd_surjective [Nonempty R] : Function.Surjective (snd : tsze R M → M) := + Prod.snd_surjective + theorem inl_injective [Zero M] : Function.Injective (inl : R → tsze R M) := Function.LeftInverse.injective <| fst_inl _ @@ -226,6 +232,16 @@ instance module [Semiring S] [AddCommMonoid R] [AddCommMonoid M] [Module S R] [M Module S (tsze R M) := Prod.instModule +/-- The trivial square-zero extension is nontrivial if it is over a nontrivial ring. -/ +instance instNontrivial_of_left {R M : Type*} [Nontrivial R] [Nonempty M] : + Nontrivial (TrivSqZeroExt R M) := + fst_surjective.nontrivial + +/-- The trivial square-zero extension is nontrivial if it is over a nontrivial module. -/ +instance instNontrivial_of_right {R M : Type*} [Nonempty R] [Nontrivial M] : + Nontrivial (TrivSqZeroExt R M) := + snd_surjective.nontrivial + @[simp] theorem fst_zero [Zero R] [Zero M] : (0 : tsze R M).fst = 0 := rfl diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean index 4cd462e78f355..565f29a05b7f1 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean @@ -680,16 +680,14 @@ theorem toCharTwoJNeZeroNF_spec (ha₁ : W.a₁ ≠ 0) : · field_simp [toCharTwoJNeZeroNF] linear_combination (W.a₁ ^ 4 * W.a₃ ^ 2 + W.a₁ ^ 5 * W.a₃ * W.a₂) * CharP.cast_eq_zero F 2 -variable [DecidableEq F] - /-- For a `WeierstrassCurve` defined over a field of characteristic = 2, there is an explicit change of variables of it to `WeierstrassCurve.IsCharTwoNF`, that is, $Y^2 + XY = X^3 + a_2X^2 + a_6$ (`WeierstrassCurve.IsCharTwoJNeZeroNF`) or $Y^2 + a_3Y = X^3 + a_4X + a_6$ (`WeierstrassCurve.IsCharTwoJEqZeroNF`). -/ -def toCharTwoNF : VariableChange F := +def toCharTwoNF [DecidableEq F] : VariableChange F := if ha₁ : W.a₁ = 0 then W.toCharTwoJEqZeroNF else W.toCharTwoJNeZeroNF ha₁ -instance toCharTwoNF_spec : (W.variableChange W.toCharTwoNF).IsCharTwoNF := by +instance toCharTwoNF_spec [DecidableEq F] : (W.variableChange W.toCharTwoNF).IsCharTwoNF := by by_cases ha₁ : W.a₁ = 0 · rw [toCharTwoNF, dif_pos ha₁] haveI := W.toCharTwoJEqZeroNF_spec ha₁ @@ -699,8 +697,9 @@ instance toCharTwoNF_spec : (W.variableChange W.toCharTwoNF).IsCharTwoNF := by infer_instance theorem exists_variableChange_isCharTwoNF : - ∃ C : VariableChange F, (W.variableChange C).IsCharTwoNF := - ⟨_, W.toCharTwoNF_spec⟩ + ∃ C : VariableChange F, (W.variableChange C).IsCharTwoNF := by + classical + exact ⟨_, W.toCharTwoNF_spec⟩ end VariableChange diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 6e19f5f554874..1fe4f13ebd574 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -282,6 +282,8 @@ def identityToΓSpec : 𝟭 LocallyRingedSpace.{u} ⟶ Γ.rightOp ⋙ Spec.toLoc = toΓSpecFun X x := by rw [ContinuousMap.coe_mk] erw [this] dsimp [toΓSpecFun] + -- TODO: this instance was found automatically before #6045 + have := @AlgebraicGeometry.LocallyRingedSpace.isLocalRingHomStalkMap X Y -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [← LocalRing.comap_closedPoint (f.stalkMap x), ← PrimeSpectrum.comap_comp_apply, ← PrimeSpectrum.comap_comp_apply] diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index c12db5f61af3a..f3fa7a4fbdf4a 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -10,6 +10,7 @@ import Mathlib.RingTheory.Ideal.MinimalPrime import Mathlib.RingTheory.Ideal.Over import Mathlib.RingTheory.Localization.Away.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Defs +import Mathlib.RingTheory.LocalRing.RingHom.Basic /-! # The Zariski topology on the prime spectrum of a commutative (semi)ring diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Noetherian.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Noetherian.lean index 1e92a49a98670..7522a16f45dbe 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Noetherian.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Noetherian.lean @@ -19,9 +19,8 @@ open TopologicalSpace variable (R : Type u) [CommRing R] [IsNoetherianRing R] -instance : NoetherianSpace (PrimeSpectrum R) := by - apply ((noetherianSpace_TFAE <| PrimeSpectrum R).out 0 1).mpr - exact (closedsEmbedding R).dual.wellFounded IsWellFounded.wf +instance : NoetherianSpace (PrimeSpectrum R) := + ((noetherianSpace_TFAE <| PrimeSpectrum R).out 0 1).mpr (closedsEmbedding R).dual.wellFoundedLT lemma _root_.minimalPrimes.finite_of_isNoetherianRing : (minimalPrimes R).Finite := minimalPrimes.equivIrreducibleComponents R diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 938d2b0845645..133dcd36da1a9 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -294,12 +294,7 @@ theorem mem_carrier_iff_of_mem (hm : 0 < m) (q : Spec.T A⁰_ f) (a : A) {n} (hn trans (HomogeneousLocalization.mk ⟨m * n, ⟨proj 𝒜 n a ^ m, by rw [← smul_eq_mul]; mem_tac⟩, ⟨f ^ n, by rw [mul_comm]; mem_tac⟩, ⟨_, rfl⟩⟩ : A⁰_ f) ∈ q.asIdeal · refine ⟨fun h ↦ h n, fun h i ↦ if hi : i = n then hi ▸ h else ?_⟩ - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/5376 - we need to specify the implicit arguments by `inferInstance`. - -/ - convert @zero_mem _ _ inferInstance inferInstance _ q.asIdeal + convert zero_mem q.asIdeal apply HomogeneousLocalization.val_injective simp only [proj_apply, decompose_of_mem_ne _ hn (Ne.symm hi), zero_pow hm.ne', HomogeneousLocalization.val_mk, Localization.mk_zero, HomogeneousLocalization.val_zero] diff --git a/Mathlib/AlgebraicGeometry/ResidueField.lean b/Mathlib/AlgebraicGeometry/ResidueField.lean index f192506ed485b..e127f7a5bccab 100644 --- a/Mathlib/AlgebraicGeometry/ResidueField.lean +++ b/Mathlib/AlgebraicGeometry/ResidueField.lean @@ -79,7 +79,14 @@ lemma evaluation_ne_zero_iff_mem_basicOpen (x : X) (hx : x ∈ U) (f : Γ(X, U)) variable {X Y : Scheme.{u}} (f : X ⟶ Y) -instance (x) : IsLocalRingHom (f.stalkMap x) := inferInstanceAs (IsLocalRingHom (f.val.stalkMap x)) + +-- TODO: This instance is found before #6045. +-- We need this strange instance for `residueFieldMap`, the type of `F` must be fixed +-- like this. The instance `IsLocalRingHom (f.stalkMap x)` already exists, but does not work for +-- `residueFieldMap`. +instance (x): IsLocalRingHom (F := Y.presheaf.stalk (f.val.base x) →+* X.presheaf.stalk x) + (f.stalkMap x) := + f.2 x /-- If `X ⟶ Y` is a morphism of locally ringed spaces and `x` a point of `X`, we obtain a morphism of residue fields in the other direction. -/ diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index 3a10a69127f49..78b6f5fd8c9f9 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -236,6 +236,8 @@ def Spec.locallyRingedSpaceMap {R S : CommRingCat.{u}} (f : R ⟶ S) : #adaptation_note /-- nightly-2024-04-01 It's this `erw` that is blowing up. The implicit arguments differ significantly. -/ erw [← localRingHom_comp_stalkIso_apply] at ha + -- TODO: this instance was found automatically before #6045 + haveI : IsLocalRingHom (stalkIso (↑S) p).inv := isLocalRingHom_of_isIso _ replace ha := (isUnit_map_iff (stalkIso S p).inv _).mp ha -- Porting note: `f` had to be made explicit replace ha := IsLocalRingHom.map_nonunit diff --git a/Mathlib/AlgebraicGeometry/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/StructureSheaf.lean index 6d18e0fcb5c9f..8cfb094cb6c53 100644 --- a/Mathlib/AlgebraicGeometry/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/StructureSheaf.lean @@ -543,9 +543,15 @@ def stalkIso (x : PrimeSpectrum.Top R) : instance (x : PrimeSpectrum R) : IsIso (stalkToFiberRingHom R x) := (stalkIso R x).isIso_hom +instance (x : PrimeSpectrum R) : IsLocalRingHom (stalkToFiberRingHom R x) := + isLocalRingHom_of_isIso _ + instance (x : PrimeSpectrum R) : IsIso (localizationToStalk R x) := (stalkIso R x).isIso_inv +instance (x : PrimeSpectrum R) : IsLocalRingHom (localizationToStalk R x) := + isLocalRingHom_of_isIso _ + @[simp, reassoc] theorem stalkToFiberRingHom_localizationToStalk (x : PrimeSpectrum.Top R) : stalkToFiberRingHom R x ≫ localizationToStalk R x = 𝟙 _ := diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean index 84980eaca3178..2a8bb0c418f7e 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean @@ -59,7 +59,7 @@ theorem reflTransSymmAux_mem_I (x : I × I) : reflTransSymmAux x ∈ I := by · norm_num · unit_interval · rw [mul_assoc] - apply mul_le_one + apply mul_le_one₀ · unit_interval · apply mul_nonneg · norm_num @@ -69,7 +69,7 @@ theorem reflTransSymmAux_mem_I (x : I × I) : reflTransSymmAux x ∈ I := by · apply mul_nonneg · unit_interval linarith [unitInterval.nonneg x.2, unitInterval.le_one x.2] - · apply mul_le_one + · apply mul_le_one₀ · unit_interval · linarith [unitInterval.nonneg x.2, unitInterval.le_one x.2] · linarith [unitInterval.nonneg x.2, unitInterval.le_one x.2] diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index e72f2e8cd7bec..42a745d3429b8 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -165,6 +165,43 @@ theorem const_comp (x : SimplexCategory) {y z : SimplexCategory} const x y i ≫ f = const x z (f.toOrderHom i) := rfl +theorem const_fac_thru_zero (n m : SimplexCategory) (i : Fin (m.len + 1)) : + const n m i = const n [0] 0 ≫ SimplexCategory.const [0] m i := by + rw [const_comp]; rfl + +theorem eq_const_of_zero {n : SimplexCategory} (f : ([0] : SimplexCategory) ⟶ n) : + f = const _ n (f.toOrderHom 0) := by + ext x; match x with | 0 => rfl + +theorem exists_eq_const_of_zero {n : SimplexCategory} (f : ([0] : SimplexCategory) ⟶ n) : + ∃ a, f = const _ n a := ⟨_, eq_const_of_zero _⟩ + +theorem eq_const_to_zero {n : SimplexCategory} (f : n ⟶ [0]) : + f = const n _ 0 := by + ext : 3 + apply @Subsingleton.elim (Fin 1) + +theorem eq_of_one_to_one (f : ([1] : SimplexCategory) ⟶ [1]) : + (∃ a, f = const [1] _ a) ∨ f = 𝟙 _ := by + match e0 : f.toOrderHom 0, e1 : f.toOrderHom 1 with + | 0, 0 | 1, 1 => + refine .inl ⟨f.toOrderHom 0, ?_⟩ + ext i : 3 + match i with + | 0 => rfl + | 1 => exact e1.trans e0.symm + | 0, 1 => + right + ext i : 3 + match i with + | 0 => exact e0 + | 1 => exact e1 + | 1, 0 => + have := f.toOrderHom.monotone (by decide : (0 : Fin 2) ≤ 1) + rw [e0, e1] at this + exact Not.elim (by decide) this + + /-- Make a morphism `[n] ⟶ [m]` from a monotone map between fin's. This is useful for constructing morphisms between `[n]` directly without identifying `n` with `[n].len`. @@ -173,6 +210,36 @@ without identifying `n` with `[n].len`. def mkHom {n m : ℕ} (f : Fin (n + 1) →o Fin (m + 1)) : ([n] : SimplexCategory) ⟶ [m] := SimplexCategory.Hom.mk f +/-- The morphism `[1] ⟶ [n]` that picks out a specified `h : i ≤ j` in `Fin (n+1)`.-/ +def mkOfLe {n} (i j : Fin (n+1)) (h : i ≤ j) : ([1] : SimplexCategory) ⟶ [n] := + SimplexCategory.mkHom { + toFun := fun | 0 => i | 1 => j + monotone' := fun + | 0, 0, _ | 1, 1, _ => le_rfl + | 0, 1, _ => h + } + +/-- The morphism `[1] ⟶ [n]` that picks out the arrow `i ⟶ i+1` in `Fin (n+1)`.-/ +def mkOfSucc {n} (i : Fin n) : ([1] : SimplexCategory) ⟶ [n] := + SimplexCategory.mkHom { + toFun := fun | 0 => i.castSucc | 1 => i.succ + monotone' := fun + | 0, 0, _ | 1, 1, _ => le_rfl + | 0, 1, _ => Fin.castSucc_le_succ i + } + +/-- The morphism `[2] ⟶ [n]` that picks out a specified composite of morphisms in `Fin (n+1)`.-/ +def mkOfLeComp {n} (i j k : Fin (n + 1)) (h₁ : i ≤ j) (h₂ : j ≤ k) : + ([2] : SimplexCategory) ⟶ [n] := + SimplexCategory.mkHom { + toFun := fun | 0 => i | 1 => j | 2 => k + monotone' := fun + | 0, 0, _ | 1, 1, _ | 2, 2, _ => le_rfl + | 0, 1, _ => h₁ + | 1, 2, _ => h₂ + | 0, 2, _ => Fin.le_trans h₁ h₂ + } + instance (Δ : SimplexCategory) : Subsingleton (Δ ⟶ [0]) where allEq f g := by ext : 3; apply Subsingleton.elim (α := Fin 1) @@ -400,6 +467,40 @@ lemma factor_δ_spec {m n : ℕ} (f : ([m] : SimplexCategory) ⟶ [n+1]) (j : Fi · rwa [succ_le_castSucc_iff, lt_pred_iff] rw [succ_pred] + +theorem eq_of_one_to_two (f : ([1] : SimplexCategory) ⟶ [2]) : + f = (δ (n := 1) 0) ∨ f = (δ (n := 1) 1) ∨ f = (δ (n := 1) 2) ∨ + ∃ a, f = SimplexCategory.const _ _ a := by + have : f.toOrderHom 0 ≤ f.toOrderHom 1 := f.toOrderHom.monotone (by decide : (0 : Fin 2) ≤ 1) + match e0 : f.toOrderHom 0, e1 : f.toOrderHom 1 with + | 1, 2 => + left + ext i : 3 + match i with + | 0 => exact e0 + | 1 => exact e1 + | 0, 2 => + right; left + ext i : 3 + match i with + | 0 => exact e0 + | 1 => exact e1 + | 0, 1 => + right; right; left + ext i : 3 + match i with + | 0 => exact e0 + | 1 => exact e1 + | 0, 0 | 1, 1 | 2, 2 => + right; right; right; use f.toOrderHom 0 + ext i : 3 + match i with + | 0 => rfl + | 1 => exact e1.trans e0.symm + | 1, 0 | 2, 0 | 2, 1 => + rw [e0, e1] at this + exact Not.elim (by decide) this + end Generators section Skeleton @@ -490,6 +591,14 @@ def inclusion {n : ℕ} : SimplexCategory.Truncated n ⥤ SimplexCategory := instance (n : ℕ) : (inclusion : Truncated n ⥤ _).Full := FullSubcategory.full _ instance (n : ℕ) : (inclusion : Truncated n ⥤ _).Faithful := FullSubcategory.faithful _ +/-- A proof that the full subcategory inclusion is fully faithful.-/ +noncomputable def inclusion.fullyFaithful (n : ℕ) : + (inclusion : Truncated n ⥤ _).op.FullyFaithful := Functor.FullyFaithful.ofFullyFaithful _ + +@[ext] +theorem Hom.ext {n} {a b : Truncated n} (f g : a ⟶ b) : + f.toOrderHom = g.toOrderHom → f = g := SimplexCategory.Hom.ext _ _ + end Truncated section Concrete diff --git a/Mathlib/AlgebraicTopology/SimplicialObject.lean b/Mathlib/AlgebraicTopology/SimplicialObject.lean index ea02f63a11c95..c15da7cf1c35d 100644 --- a/Mathlib/AlgebraicTopology/SimplicialObject.lean +++ b/Mathlib/AlgebraicTopology/SimplicialObject.lean @@ -4,7 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Kim Morrison, Adam Topaz -/ import Mathlib.AlgebraicTopology.SimplexCategory +import Mathlib.CategoryTheory.Adjunction.Reflective import Mathlib.CategoryTheory.Comma.Arrow +import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction import Mathlib.CategoryTheory.Limits.FunctorCategory.Basic import Mathlib.CategoryTheory.Opposites @@ -23,7 +25,7 @@ open Opposite open CategoryTheory -open CategoryTheory.Limits +open CategoryTheory.Limits CategoryTheory.Functor universe v u v' u' @@ -241,13 +243,103 @@ variable {C} end Truncated -section Skeleton +section Truncation -/-- The skeleton functor from simplicial objects to truncated simplicial objects. -/ -def sk (n : ℕ) : SimplicialObject C ⥤ SimplicialObject.Truncated C n := +/-- The truncation functor from simplicial objects to truncated simplicial objects. -/ +def truncation (n : ℕ) : SimplicialObject C ⥤ SimplicialObject.Truncated C n := (whiskeringLeft _ _ _).obj SimplexCategory.Truncated.inclusion.op -end Skeleton +end Truncation + + +noncomputable section + +/-- The n-skeleton as a functor `SimplicialObject.Truncated C n ⥤ SimplicialObject C`. -/ +protected abbrev Truncated.sk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), + SimplexCategory.Truncated.inclusion.op.HasLeftKanExtension F] : + SimplicialObject.Truncated C n ⥤ SimplicialObject C := + lan (SimplexCategory.Truncated.inclusion.op) + +/-- The n-coskeleton as a functor `SimplicialObject.Truncated C n ⥤ SimplicialObject C`. -/ +protected abbrev Truncated.cosk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), + SimplexCategory.Truncated.inclusion.op.HasRightKanExtension F] : + SimplicialObject.Truncated C n ⥤ SimplicialObject C := + ran (SimplexCategory.Truncated.inclusion.op) + +/-- The n-skeleton as an endofunctor on `SimplicialObject C`. -/ +abbrev sk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), + SimplexCategory.Truncated.inclusion.op.HasLeftKanExtension F] : + SimplicialObject C ⥤ SimplicialObject C := truncation n ⋙ Truncated.sk n + +/-- The n-coskeleton as an endofunctor on `SimplicialObject C`. -/ +abbrev cosk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), + SimplexCategory.Truncated.inclusion.op.HasRightKanExtension F] : + SimplicialObject C ⥤ SimplicialObject C := truncation n ⋙ Truncated.cosk n + +end + +section adjunctions +/- When the left and right Kan extensions exist, `Truncated.sk n` and `Truncated.cosk n` +respectively define left and right adjoints to `truncation n`.-/ + + +variable (n : ℕ) +variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), + SimplexCategory.Truncated.inclusion.op.HasRightKanExtension F] +variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), + SimplexCategory.Truncated.inclusion.op.HasLeftKanExtension F] + +/-- The adjunction between the n-skeleton and n-truncation.-/ +noncomputable def skAdj : Truncated.sk (C := C) n ⊣ truncation n := + lanAdjunction _ _ + +/-- The adjunction between n-truncation and the n-coskeleton.-/ +noncomputable def coskAdj : truncation (C := C) n ⊣ Truncated.cosk n := + ranAdjunction _ _ + +namespace Truncated +/- When the left and right Kan extensions exist and are pointwise Kan extensions, +`skAdj n` and `coskAdj n` are respectively coreflective and reflective.-/ + +variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), + SimplexCategory.Truncated.inclusion.op.HasPointwiseRightKanExtension F] +variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), + SimplexCategory.Truncated.inclusion.op.HasPointwiseLeftKanExtension F] + +instance cosk_reflective : IsIso (coskAdj (C := C) n).counit := + reflective' SimplexCategory.Truncated.inclusion.op + +instance sk_coreflective : IsIso (skAdj (C := C) n).unit := + coreflective' SimplexCategory.Truncated.inclusion.op + +/-- Since `Truncated.inclusion` is fully faithful, so is right Kan extension along it.-/ +noncomputable def cosk.fullyFaithful : + (Truncated.cosk (C := C) n).FullyFaithful := by + apply Adjunction.fullyFaithfulROfIsIsoCounit (coskAdj n) + +instance cosk.full : (Truncated.cosk (C := C) n).Full := FullyFaithful.full (cosk.fullyFaithful _) + +instance cosk.faithful : (Truncated.cosk (C := C) n).Faithful := + FullyFaithful.faithful (cosk.fullyFaithful _) + +noncomputable instance coskAdj.reflective : Reflective (Truncated.cosk (C := C) n) := + Reflective.mk (truncation _) (coskAdj _) + +/-- Since `Truncated.inclusion` is fully faithful, so is left Kan extension along it.-/ +noncomputable def sk.fullyFaithful : (Truncated.sk (C := C) n).FullyFaithful := + Adjunction.fullyFaithfulLOfIsIsoUnit (skAdj n) + +instance sk.full : (Truncated.sk (C := C) n).Full := FullyFaithful.full (sk.fullyFaithful _) + +instance sk.faithful : (Truncated.sk (C := C) n).Faithful := + FullyFaithful.faithful (sk.fullyFaithful _) + +noncomputable instance skAdj.coreflective : Coreflective (Truncated.sk (C := C) n) := + Coreflective.mk (truncation _) (skAdj _) + +end Truncated + +end adjunctions variable (C) @@ -576,13 +668,13 @@ variable {C} end Truncated -section Skeleton +section Truncation -/-- The skeleton functor from cosimplicial objects to truncated cosimplicial objects. -/ -def sk (n : ℕ) : CosimplicialObject C ⥤ CosimplicialObject.Truncated C n := +/-- The truncation functor from cosimplicial objects to truncated cosimplicial objects. -/ +def truncation (n : ℕ) : CosimplicialObject C ⥤ CosimplicialObject.Truncated C n := (whiskeringLeft _ _ _).obj SimplexCategory.Truncated.inclusion -end Skeleton +end Truncation variable (C) diff --git a/Mathlib/AlgebraicTopology/SimplicialSet.lean b/Mathlib/AlgebraicTopology/SimplicialSet.lean index a8198f848a6c7..93e29ca1a49c3 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet.lean @@ -33,7 +33,7 @@ a morphism `Δ[n] ⟶ ∂Δ[n]`. universe v u -open CategoryTheory CategoryTheory.Limits +open CategoryTheory CategoryTheory.Limits CategoryTheory.Functor open Simplicial @@ -335,18 +335,91 @@ instance Truncated.hasColimits {n : ℕ} : HasColimits (Truncated n) := by dsimp only [Truncated] infer_instance +/-- The ulift functor `SSet.Truncated.{u} ⥤ SSet.Truncated.{max u v}` on truncated +simplicial sets. -/ +def Truncated.uliftFunctor (k : ℕ) : SSet.Truncated.{u} k ⥤ SSet.Truncated.{max u v} k := + (whiskeringRight _ _ _).obj CategoryTheory.uliftFunctor.{v, u} + -- Porting note (#5229): added an `ext` lemma. @[ext] lemma Truncated.hom_ext {n : ℕ} {X Y : Truncated n} {f g : X ⟶ Y} (w : ∀ n, f.app n = g.app n) : f = g := NatTrans.ext (funext w) -/-- The skeleton functor on simplicial sets. -/ -def sk (n : ℕ) : SSet ⥤ SSet.Truncated n := - SimplicialObject.sk n +/-- The truncation functor on simplicial sets. -/ +abbrev truncation (n : ℕ) : SSet ⥤ SSet.Truncated n := SimplicialObject.truncation n instance {n} : Inhabited (SSet.Truncated n) := - ⟨(sk n).obj <| Δ[0]⟩ + ⟨(truncation n).obj <| Δ[0]⟩ + + +open SimplexCategory + +noncomputable section + +/-- The n-skeleton as a functor `SSet.Truncated n ⥤ SSet`. -/ +protected abbrev Truncated.sk (n : ℕ) : SSet.Truncated n ⥤ SSet.{u} := + SimplicialObject.Truncated.sk n + +/-- The n-coskeleton as a functor `SSet.Truncated n ⥤ SSet`. -/ +protected abbrev Truncated.cosk (n : ℕ) : SSet.Truncated n ⥤ SSet.{u} := + SimplicialObject.Truncated.cosk n + +/-- The n-skeleton as an endofunctor on `SSet`. -/ +abbrev sk (n : ℕ) : SSet ⥤ SSet := SimplicialObject.sk n + +/-- The n-coskeleton as an endofunctor on `SSet`. -/ +abbrev cosk (n : ℕ) : SSet ⥤ SSet := SimplicialObject.cosk n + +end + +section adjunctions + +/-- The adjunction between the n-skeleton and n-truncation.-/ +noncomputable def skAdj (n : ℕ) : Truncated.sk n ⊣ truncation.{u} n := + SimplicialObject.skAdj n + +/-- The adjunction between n-truncation and the n-coskeleton.-/ +noncomputable def coskAdj (n : ℕ) : truncation.{u} n ⊣ Truncated.cosk n := + SimplicialObject.coskAdj n + +namespace Truncated + +instance cosk_reflective (n) : IsIso (coskAdj n).counit := + SimplicialObject.Truncated.cosk_reflective n + +instance sk_coreflective (n) : IsIso (skAdj n).unit := + SimplicialObject.Truncated.sk_coreflective n + +/-- Since `Truncated.inclusion` is fully faithful, so is right Kan extension along it.-/ +noncomputable def cosk.fullyFaithful (n) : + (Truncated.cosk n).FullyFaithful := + SimplicialObject.Truncated.cosk.fullyFaithful n + +instance cosk.full (n) : (Truncated.cosk n).Full := + SimplicialObject.Truncated.cosk.full n + +instance cosk.faithful (n) : (Truncated.cosk n).Faithful := + SimplicialObject.Truncated.cosk.faithful n + +noncomputable instance coskAdj.reflective (n) : Reflective (Truncated.cosk n) := + SimplicialObject.Truncated.coskAdj.reflective n + +/-- Since `Truncated.inclusion` is fully faithful, so is left Kan extension along it.-/ +noncomputable def sk.fullyFaithful (n) : + (Truncated.sk n).FullyFaithful := SimplicialObject.Truncated.sk.fullyFaithful n + +instance sk.full (n) : (Truncated.sk n).Full := SimplicialObject.Truncated.sk.full n + +instance sk.faithful (n) : (Truncated.sk n).Faithful := + SimplicialObject.Truncated.sk.faithful n + +noncomputable instance skAdj.coreflective (n) : Coreflective (Truncated.sk n) := + SimplicialObject.Truncated.skAdj.coreflective n + +end Truncated + +end adjunctions /-- The category of augmented simplicial sets, as a particular case of augmented simplicial objects. -/ diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 49ef36b0ddbd9..c73176c9a5ec5 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -473,7 +473,7 @@ theorem comp_summable_nnreal (q : FormalMultilinearSeries 𝕜 F G) (p : FormalM simp only [Finset.prod_mul_distrib, Finset.prod_pow_eq_pow_sum, c.sum_blocksFun] _ ≤ ∏ _i : Fin c.length, Cp := Finset.prod_le_prod' fun i _ => hCp _ _ = Cp ^ c.length := by simp - _ ≤ Cp ^ n := pow_le_pow_right hCp1 c.length_le + _ ≤ Cp ^ n := pow_right_mono₀ hCp1 c.length_le calc ‖q.compAlongComposition p c‖₊ * r ^ n ≤ (‖q c.length‖₊ * ∏ i, ‖p (c.blocksFun i)‖₊) * r ^ n := diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 6c6b7be8b5ffe..15ae960d1d41e 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -251,7 +251,7 @@ theorem isLittleO_iff_nat_mul_le_aux (h₀ : (∀ x, 0 ≤ ‖f x‖) ∨ ∀ x, rcases exists_nat_gt ε⁻¹ with ⟨n, hn⟩ have hn₀ : (0 : ℝ) < n := (inv_pos.2 ε0).trans hn refine ((isBigOWith_inv hn₀).2 (H n)).bound.mono fun x hfg => ?_ - refine hfg.trans (mul_le_mul_of_nonneg_right (inv_le_of_inv_le ε0 hn.le) ?_) + refine hfg.trans (mul_le_mul_of_nonneg_right (inv_le_of_inv_le₀ ε0 hn.le) ?_) refine h₀.elim (fun hf => nonneg_of_mul_nonneg_right ((hf x).trans hfg) ?_) fun h => h x exact inv_pos.2 hn₀ @@ -1251,6 +1251,9 @@ theorem IsLittleO.trans_tendsto (hfg : f'' =o[l] g'') (hg : Tendsto g'' l (𝓝 Tendsto f'' l (𝓝 0) := hfg.isBigO.trans_tendsto hg +lemma isLittleO_id_one [One F''] [NeZero (1 : F'')] : (fun x : E'' => x) =o[𝓝 0] (1 : E'' → F'') := + isLittleO_id_const one_ne_zero + /-! ### Multiplication by a constant -/ @@ -1447,7 +1450,7 @@ theorem IsBigOWith.inv_rev {f : α → 𝕜} {g : α → 𝕜'} (h : IsBigOWith rcases eq_or_ne (f x) 0 with hx | hx · simp only [hx, h₀ hx, inv_zero, norm_zero, mul_zero, le_rfl] · have hc : 0 < c := pos_of_mul_pos_left ((norm_pos_iff.2 hx).trans_le hle) (norm_nonneg _) - replace hle := inv_le_inv_of_le (norm_pos_iff.2 hx) hle + replace hle := inv_anti₀ (norm_pos_iff.2 hx) hle simpa only [norm_inv, mul_inv, ← div_eq_inv_mul, div_le_iff₀ hc] using hle theorem IsBigO.inv_rev {f : α → 𝕜} {g : α → 𝕜'} (h : f =O[l] g) @@ -1915,6 +1918,13 @@ theorem isBigO_atTop_iff_eventually_exists_pos {α : Type*} f =O[atTop] g ↔ ∀ᶠ n₀ in atTop, ∃ c > 0, ∀ n ≥ n₀, c * ‖f n‖ ≤ ‖g n‖ := by simp_rw [isBigO_iff'', ← exists_prop, Subtype.exists', exists_eventually_atTop] +lemma isBigO_mul_iff_isBigO_div {f g h : α → 𝕜} (hf : ∀ᶠ x in l, f x ≠ 0) : + (fun x ↦ f x * g x) =O[l] h ↔ g =O[l] (fun x ↦ h x / f x) := by + rw [isBigO_iff', isBigO_iff'] + refine ⟨fun ⟨c, hc, H⟩ ↦ ⟨c, hc, ?_⟩, fun ⟨c, hc, H⟩ ↦ ⟨c, hc, ?_⟩⟩ <;> + · refine H.congr <| Eventually.mp hf <| Eventually.of_forall fun x hx ↦ ?_ + rw [norm_mul, norm_div, ← mul_div_assoc, le_div_iff₀' (norm_pos_iff.mpr hx)] + end Asymptotics open Asymptotics diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index a298ba04ac176..361cdc8aefebc 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -692,7 +692,7 @@ lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : let f : C(spectrum R a, σₙ R a) := ⟨_, continuous_inclusion <| spectrum_subset_quasispectrum R a⟩ refine (cfcHom_closedEmbedding ha).comp <| - (UniformInducing.isUniformEmbedding ⟨?_⟩).toClosedEmbedding + (IsUniformInducing.isUniformEmbedding ⟨?_⟩).toClosedEmbedding have := uniformSpace_eq_inf_precomp_of_cover (β := R) f (0 : C(Unit, σₙ R a)) (map_continuous f).isProperMap (map_continuous 0).isProperMap <| by simp only [← Subtype.val_injective.image_injective.eq_iff, f, ContinuousMap.coe_mk, diff --git a/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean b/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean index afee7eea2b49e..82554bb415dee 100644 --- a/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean +++ b/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean @@ -167,7 +167,8 @@ instance [u : UniformSpace E] : UniformSpace (C⋆ᵐᵒᵈ E) := u.comap <| equ instance [Bornology E] : Bornology (C⋆ᵐᵒᵈ E) := Bornology.induced <| equiv E /-- `WithCStarModule.equiv` as a uniform equivalence between `C⋆ᵐᵒᵈ E` and `E`. -/ -def uniformEquiv [UniformSpace E] : C⋆ᵐᵒᵈ E ≃ᵤ E := equiv E |>.toUniformEquivOfUniformInducing ⟨rfl⟩ +def uniformEquiv [UniformSpace E] : C⋆ᵐᵒᵈ E ≃ᵤ E := + equiv E |>.toUniformEquivOfIsUniformInducing ⟨rfl⟩ instance [UniformSpace E] [CompleteSpace E] : CompleteSpace (C⋆ᵐᵒᵈ E) := uniformEquiv.completeSpace_iff.mpr inferInstance diff --git a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean index d9c622e26c2c4..3c84852d4f405 100644 --- a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean +++ b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean @@ -542,7 +542,7 @@ theorem isUniformEmbedding_toProdMulOpposite : alias uniformEmbedding_toProdMulOpposite := isUniformEmbedding_toProdMulOpposite instance [CompleteSpace A] : CompleteSpace 𝓜(𝕜, A) := by - rw [completeSpace_iff_isComplete_range isUniformEmbedding_toProdMulOpposite.toUniformInducing] + rw [completeSpace_iff_isComplete_range isUniformEmbedding_toProdMulOpposite.isUniformInducing] apply IsClosed.isComplete simp only [range_toProdMulOpposite, Set.setOf_forall] refine isClosed_iInter fun x => isClosed_iInter fun y => isClosed_eq ?_ ?_ diff --git a/Mathlib/Analysis/CStarAlgebra/Spectrum.lean b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean index 337a423b3d120..6d91e465acfeb 100644 --- a/Mathlib/Analysis/CStarAlgebra/Spectrum.lean +++ b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean @@ -78,7 +78,7 @@ theorem unitary.spectrum_subset_circle (u : unitary E) : rw [← inv_inv (unitary.toUnits u), ← spectrum.map_inv, Set.mem_inv] at hk have : ‖k‖⁻¹ ≤ ‖(↑(unitary.toUnits u)⁻¹ : E)‖ := by simpa only [norm_inv] using norm_le_norm_of_mem hk - simpa using inv_le_of_inv_le (norm_pos_iff.mpr hnk) this + simpa using inv_le_of_inv_le₀ (norm_pos_iff.mpr hnk) this theorem spectrum.subset_circle_of_unitary {u : E} (h : u ∈ unitary E) : spectrum 𝕜 u ⊆ Metric.sphere 0 1 := diff --git a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean index 75f2e43860603..cc97d67c10ce6 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean @@ -449,7 +449,7 @@ theorem norm_iteratedFDerivWithin_comp_le_aux {Fu Gu : Type u} [NormedAddCommGro exact Nat.add_sub_of_le (Finset.mem_range_succ_iff.1 hi) _ ≤ ∑ i ∈ Finset.range (n + 1), (n ! : ℝ) * 1 * C * D ^ (n + 1) * 1 := by gcongr with i - apply inv_le_one + apply inv_le_one_of_one_le₀ simpa only [Nat.one_le_cast] using (n - i).factorial_pos _ = (n + 1)! * C * D ^ (n + 1) := by simp only [mul_assoc, mul_one, Finset.sum_const, Finset.card_range, nsmul_eq_mul, diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index ff7ff85e3d49e..3f934a6deec9b 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -347,7 +347,7 @@ protected theorem ContDiffWithinAt.eventually {n : ℕ} (h : ContDiffWithinAt ∀ᶠ y in 𝓝[insert x s] x, ContDiffWithinAt 𝕜 n f s y := by rcases h.contDiffOn le_rfl with ⟨u, hu, _, hd⟩ have : ∀ᶠ y : E in 𝓝[insert x s] x, u ∈ 𝓝[insert x s] y ∧ y ∈ u := - (eventually_nhdsWithin_nhdsWithin.2 hu).and hu + (eventually_eventually_nhdsWithin.2 hu).and hu refine this.mono fun y hy => (hd y hy.2).mono_of_mem ?_ exact nhdsWithin_mono y (subset_insert _ _) hy.1 diff --git a/Mathlib/Analysis/Calculus/Deriv/Abs.lean b/Mathlib/Analysis/Calculus/Deriv/Abs.lean index 55365ae78213a..77194eadb8363 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Abs.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Abs.lean @@ -21,7 +21,7 @@ absolute value, derivative open Filter Real Set variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] -variable {n : ℕ∞} {f g : E → ℝ} {f' : E →L[ℝ] ℝ} {s : Set E} {x : E} +variable {n : ℕ∞} {f : E → ℝ} {f' : E →L[ℝ] ℝ} {s : Set E} {x : E} theorem contDiffAt_abs {x : ℝ} (hx : x ≠ 0) : ContDiffAt ℝ n (|·|) x := contDiffAt_norm ℝ hx diff --git a/Mathlib/Analysis/Calculus/Deriv/Mul.lean b/Mathlib/Analysis/Calculus/Deriv/Mul.lean index 64628ceb448d0..c89e9ca6d7c04 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Mul.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Mul.lean @@ -324,22 +324,26 @@ end HasDeriv variable {ι : Type*} {𝔸' : Type*} [NormedCommRing 𝔸'] [NormedAlgebra 𝕜 𝔸'] {u : Finset ι} {f : ι → 𝕜 → 𝔸'} {f' : ι → 𝔸'} +@[fun_prop] theorem DifferentiableAt.finset_prod (hd : ∀ i ∈ u, DifferentiableAt 𝕜 (f i) x) : DifferentiableAt 𝕜 (∏ i ∈ u, f i ·) x := by classical exact (HasDerivAt.finset_prod (fun i hi ↦ DifferentiableAt.hasDerivAt (hd i hi))).differentiableAt +@[fun_prop] theorem DifferentiableWithinAt.finset_prod (hd : ∀ i ∈ u, DifferentiableWithinAt 𝕜 (f i) s x) : DifferentiableWithinAt 𝕜 (∏ i ∈ u, f i ·) s x := by classical exact (HasDerivWithinAt.finset_prod (fun i hi ↦ DifferentiableWithinAt.hasDerivWithinAt (hd i hi))).differentiableWithinAt +@[fun_prop] theorem DifferentiableOn.finset_prod (hd : ∀ i ∈ u, DifferentiableOn 𝕜 (f i) s) : DifferentiableOn 𝕜 (∏ i ∈ u, f i ·) s := fun x hx ↦ .finset_prod (fun i hi ↦ hd i hi x hx) +@[fun_prop] theorem Differentiable.finset_prod (hd : ∀ i ∈ u, Differentiable 𝕜 (f i)) : Differentiable 𝕜 (∏ i ∈ u, f i ·) := fun x ↦ .finset_prod (fun i hi ↦ hd i hi x) @@ -362,19 +366,21 @@ theorem HasStrictDerivAt.div_const (hc : HasStrictDerivAt c c' x) (d : 𝕜') : HasStrictDerivAt (fun x => c x / d) (c' / d) x := by simpa only [div_eq_mul_inv] using hc.mul_const d⁻¹ +@[fun_prop] theorem DifferentiableWithinAt.div_const (hc : DifferentiableWithinAt 𝕜 c s x) (d : 𝕜') : DifferentiableWithinAt 𝕜 (fun x => c x / d) s x := (hc.hasDerivWithinAt.div_const _).differentiableWithinAt -@[simp] +@[simp, fun_prop] theorem DifferentiableAt.div_const (hc : DifferentiableAt 𝕜 c x) (d : 𝕜') : DifferentiableAt 𝕜 (fun x => c x / d) x := (hc.hasDerivAt.div_const _).differentiableAt +@[fun_prop] theorem DifferentiableOn.div_const (hc : DifferentiableOn 𝕜 c s) (d : 𝕜') : DifferentiableOn 𝕜 (fun x => c x / d) s := fun x hx => (hc x hx).div_const d -@[simp] +@[simp, fun_prop] theorem Differentiable.div_const (hc : Differentiable 𝕜 c) (d : 𝕜') : Differentiable 𝕜 fun x => c x / d := fun x => (hc x).div_const d diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 0e90e3c0a221f..80735fc0538cf 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -361,9 +361,9 @@ protected theorem AnalyticOn.fderivWithin (h : AnalyticOn 𝕜 f s) (hu : Unique refine ⟨p.derivSeries, r, ?_⟩ refine ⟨hr.r_le.trans p.radius_le_radius_derivSeries, hr.r_pos, fun {y} hy h'y ↦ ?_⟩ apply hr.hasSum_derivSeries_of_hasFDerivWithinAt (by simpa [edist_eq_coe_nnnorm] using h'y) hy - rw [insert_eq_of_mem hx] at hy ⊢ - apply DifferentiableWithinAt.hasFDerivWithinAt - · exact h.differentiableOn _ hy + · rw [insert_eq_of_mem hx] at hy ⊢ + apply DifferentiableWithinAt.hasFDerivWithinAt + exact h.differentiableOn _ hy · rwa [insert_eq_of_mem hx] /-- If a function is analytic on a set `s`, so are its successive Fréchet derivative within this diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index f156437148642..88e55627bf87e 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -895,7 +895,7 @@ theorem Filter.EventuallyEq.fderivWithin_eq (hs : f₁ =ᶠ[𝓝[s] x] f) (hx : theorem Filter.EventuallyEq.fderivWithin' (hs : f₁ =ᶠ[𝓝[s] x] f) (ht : t ⊆ s) : fderivWithin 𝕜 f₁ t =ᶠ[𝓝[s] x] fderivWithin 𝕜 f t := - (eventually_nhdsWithin_nhdsWithin.2 hs).mp <| + (eventually_eventually_nhdsWithin.2 hs).mp <| eventually_mem_nhdsWithin.mono fun _y hys hs => EventuallyEq.fderivWithin_eq (hs.filter_mono <| nhdsWithin_mono _ ht) (hs.self_of_nhdsWithin hys) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean index beaf6fad6b9fe..410d6d1cf3e69 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean @@ -172,7 +172,7 @@ theorem Differentiable.comp_differentiableOn {g : F → G} (hg : Differentiable protected theorem HasStrictFDerivAt.comp {g : F → G} {g' : F →L[𝕜] G} (hg : HasStrictFDerivAt g g' (f x)) (hf : HasStrictFDerivAt f f' x) : HasStrictFDerivAt (fun x => g (f x)) (g'.comp f') x := - ((hg.comp_tendsto (hf.continuousAt.prod_map' hf.continuousAt)).trans_isBigO + ((hg.comp_tendsto (hf.continuousAt.prodMap' hf.continuousAt)).trans_isBigO hf.isBigO_sub).triangle <| by simpa only [g'.map_sub, f'.coe_comp'] using (g'.isBigO_comp _ _).trans_isLittleO hf diff --git a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean index 1c6738f520fcc..ac47ac267e108 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean @@ -339,7 +339,7 @@ inverse function. -/ theorem HasStrictFDerivAt.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] F} {g : F → E} {a : F} (hg : ContinuousAt g a) (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) (g a)) (hfg : ∀ᶠ y in 𝓝 a, f (g y) = y) : HasStrictFDerivAt g (f'.symm : F →L[𝕜] E) a := by - replace hg := hg.prod_map' hg + replace hg := hg.prodMap' hg replace hfg := hfg.prod_mk_nhds hfg have : (fun p : F × F => g p.1 - g p.2 - f'.symm (p.1 - p.2)) =O[𝓝 (a, a)] fun p : F × F => @@ -355,7 +355,7 @@ theorem HasStrictFDerivAt.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] · refine (hf.isBigO_sub_rev.comp_tendsto hg).congr' (Eventually.of_forall fun _ => rfl) (hfg.mono ?_) rintro p ⟨hp1, hp2⟩ - simp only [(· ∘ ·), hp1, hp2] + simp only [(· ∘ ·), hp1, hp2, Prod.map] /-- If `f (g y) = y` for `y` in some neighborhood of `a`, `g` is continuous at `a`, and `f` has an invertible derivative `f'` at `g a`, then `g` has the derivative `f'⁻¹` at `a`. diff --git a/Mathlib/Analysis/Calculus/FDeriv/Norm.lean b/Mathlib/Analysis/Calculus/FDeriv/Norm.lean new file mode 100644 index 0000000000000..b43a36ea81d3a --- /dev/null +++ b/Mathlib/Analysis/Calculus/FDeriv/Norm.lean @@ -0,0 +1,200 @@ +/- +Copyright (c) 2024 Etienne Marion. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Etienne Marion +-/ +import Mathlib.Analysis.Calculus.Deriv.Abs +import Mathlib.Analysis.Calculus.LineDeriv.Basic + +/-! +# Differentiability of the norm in a real normed vector space + +This file provides basic results about the differentiability of the norm in a real vector space. +Most are of the following kind: if the norm has some differentiability property +(`DifferentiableAt`, `ContDiffAt`, `HasStrictFDerivAt`, `HasFDerivAt`) at `x`, then so it has +at `t • x` when `t ≠ 0`. + +## Main statements + +* `ContDiffAt.contDiffAt_norm_smul`: If the norm is continuously differentiable up to order `n` + at `x`, then so it is at `t • x` when `t ≠ 0`. +* `differentiableAt_norm_smul`: If `t ≠ 0`, the norm is differentiable at `x` if and only if + it is at `t • x`. +* `HasFDerivAt.hasFDerivAt_norm_smul`: If the norm has a Fréchet derivative `f` at `x` and `t ≠ 0`, + then it has `(SignType t) • f` as a Fréchet derivative at `t · x`. +* `fderiv_norm_smul` : `fderiv ℝ (‖·‖) (t • x) = (SignType.sign t : ℝ) • (fderiv ℝ (‖·‖) x)`, + this holds without any differentiability assumptions. +* `DifferentiableAt.fderiv_norm_self`: if the norm is differentiable at `x`, + then `fderiv ℝ (‖·‖) x x = ‖x‖`. +* `norm_fderiv_norm`: if the norm is differentiable at `x` then the operator norm of its derivative + is `1` (on a non trivial space). + +## Tags + +differentiability, norm + +-/ + +open ContinuousLinearMap Filter NNReal Real Set + +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] +variable {n : ℕ∞} {f : E →L[ℝ] ℝ} {x : E} {t : ℝ} + +variable (E) in +theorem not_differentiableAt_norm_zero [Nontrivial E] : + ¬DifferentiableAt ℝ (‖·‖) (0 : E) := by + obtain ⟨x, hx⟩ := NormedSpace.exists_lt_norm ℝ E 0 + intro h + have : DifferentiableAt ℝ (fun t : ℝ ↦ ‖t • x‖) 0 := DifferentiableAt.comp _ (by simpa) (by simp) + have : DifferentiableAt ℝ (|·|) (0 : ℝ) := by + simp_rw [norm_smul, norm_eq_abs] at this + have aux : abs = fun t ↦ (1 / ‖x‖) * (|t| * ‖x‖) := by field_simp + rw [aux] + exact this.const_mul _ + exact not_differentiableAt_abs_zero this + +theorem ContDiffAt.contDiffAt_norm_smul (ht : t ≠ 0) (h : ContDiffAt ℝ n (‖·‖) x) : + ContDiffAt ℝ n (‖·‖) (t • x) := by + have h1 : ContDiffAt ℝ n (fun y ↦ t⁻¹ • y) (t • x) := (contDiff_const_smul t⁻¹).contDiffAt + have h2 : ContDiffAt ℝ n (fun y ↦ |t| * ‖y‖) x := h.const_smul |t| + conv at h2 => enter [4]; rw [← one_smul ℝ x, ← inv_mul_cancel₀ ht, mul_smul] + convert h2.comp (t • x) h1 using 1 + ext y + simp only [Function.comp_apply] + rw [norm_smul, ← mul_assoc, norm_eq_abs, ← abs_mul, mul_inv_cancel₀ ht, abs_one, one_mul] + +theorem contDiffAt_norm_smul_iff (ht : t ≠ 0) : + ContDiffAt ℝ n (‖·‖) x ↔ ContDiffAt ℝ n (‖·‖) (t • x) where + mp h := h.contDiffAt_norm_smul ht + mpr hd := by + convert hd.contDiffAt_norm_smul (inv_ne_zero ht) + rw [smul_smul, inv_mul_cancel₀ ht, one_smul] + +theorem ContDiffAt.contDiffAt_norm_of_smul (h : ContDiffAt ℝ n (‖·‖) (t • x)) : + ContDiffAt ℝ n (‖·‖) x := by + obtain rfl | hn : n = 0 ∨ 1 ≤ n := by + rw [← ENat.lt_one_iff_eq_zero] + exact lt_or_le .. + · rw [contDiffAt_zero] + exact ⟨univ, univ_mem, continuous_norm.continuousOn⟩ + obtain rfl | ht := eq_or_ne t 0 + · by_cases hE : Nontrivial E + · rw [zero_smul] at h + exact (mt (ContDiffAt.differentiableAt · hn)) (not_differentiableAt_norm_zero E) h |>.elim + · rw [not_nontrivial_iff_subsingleton] at hE + rw [eq_const_of_subsingleton (‖·‖) 0] + exact contDiffAt_const + · exact contDiffAt_norm_smul_iff ht |>.2 h + +theorem HasStrictFDerivAt.hasStrictFDerivAt_norm_smul + (ht : t ≠ 0) (h : HasStrictFDerivAt (‖·‖) f x) : + HasStrictFDerivAt (‖·‖) ((SignType.sign t : ℝ) • f) (t • x) := by + have h1 : HasStrictFDerivAt (fun y ↦ t⁻¹ • y) (t⁻¹ • ContinuousLinearMap.id ℝ E) (t • x) := + hasStrictFDerivAt_id (t • x) |>.const_smul t⁻¹ + have h2 : HasStrictFDerivAt (fun y ↦ |t| * ‖y‖) (|t| • f) x := h.const_smul |t| + conv at h2 => enter [3]; rw [← one_smul ℝ x, ← inv_mul_cancel₀ ht, mul_smul] + convert h2.comp (t • x) h1 with y + · rw [norm_smul, ← mul_assoc, norm_eq_abs, ← abs_mul, mul_inv_cancel₀ ht, abs_one, one_mul] + ext y + simp only [coe_smul', Pi.smul_apply, smul_eq_mul, comp_smulₛₗ, map_inv₀, RingHom.id_apply, + comp_id] + rw [eq_inv_mul_iff_mul_eq₀ ht, ← mul_assoc, self_mul_sign] + +theorem HasStrictFDerivAt.hasStrictDerivAt_norm_smul_neg + (ht : t < 0) (h : HasStrictFDerivAt (‖·‖) f x) : + HasStrictFDerivAt (‖·‖) (-f) (t • x) := by + simpa [ht] using h.hasStrictFDerivAt_norm_smul ht.ne + +theorem HasStrictFDerivAt.hasStrictDerivAt_norm_smul_pos + (ht : 0 < t) (h : HasStrictFDerivAt (‖·‖) f x) : + HasStrictFDerivAt (‖·‖) f (t • x) := by + simpa [ht] using h.hasStrictFDerivAt_norm_smul ht.ne' + +theorem HasFDerivAt.hasFDerivAt_norm_smul + (ht : t ≠ 0) (h : HasFDerivAt (‖·‖) f x) : + HasFDerivAt (‖·‖) ((SignType.sign t : ℝ) • f) (t • x) := by + have h1 : HasFDerivAt (fun y ↦ t⁻¹ • y) (t⁻¹ • ContinuousLinearMap.id ℝ E) (t • x) := + hasFDerivAt_id (t • x) |>.const_smul t⁻¹ + have h2 : HasFDerivAt (fun y ↦ |t| * ‖y‖) (|t| • f) x := h.const_smul |t| + conv at h2 => enter [3]; rw [← one_smul ℝ x, ← inv_mul_cancel₀ ht, mul_smul] + convert h2.comp (t • x) h1 using 2 with y + · simp only [Function.comp_apply] + rw [norm_smul, ← mul_assoc, norm_eq_abs, ← abs_mul, mul_inv_cancel₀ ht, abs_one, one_mul] + · ext y + simp only [coe_smul', Pi.smul_apply, smul_eq_mul, comp_smulₛₗ, map_inv₀, RingHom.id_apply, + comp_id] + rw [eq_inv_mul_iff_mul_eq₀ ht, ← mul_assoc, self_mul_sign] + +theorem HasFDerivAt.hasFDerivAt_norm_smul_neg + (ht : t < 0) (h : HasFDerivAt (‖·‖) f x) : + HasFDerivAt (‖·‖) (-f) (t • x) := by + simpa [ht] using h.hasFDerivAt_norm_smul ht.ne + +theorem HasFDerivAt.hasFDerivAt_norm_smul_pos + (ht : 0 < t) (h : HasFDerivAt (‖·‖) f x) : + HasFDerivAt (‖·‖) f (t • x) := by + simpa [ht] using h.hasFDerivAt_norm_smul ht.ne' + +theorem differentiableAt_norm_smul (ht : t ≠ 0) : + DifferentiableAt ℝ (‖·‖) x ↔ DifferentiableAt ℝ (‖·‖) (t • x) where + mp hd := (hd.hasFDerivAt.hasFDerivAt_norm_smul ht).differentiableAt + mpr hd := by + convert (hd.hasFDerivAt.hasFDerivAt_norm_smul (inv_ne_zero ht)).differentiableAt + rw [smul_smul, inv_mul_cancel₀ ht, one_smul] + +theorem DifferentiableAt.differentiableAt_norm_of_smul (h : DifferentiableAt ℝ (‖·‖) (t • x)) : + DifferentiableAt ℝ (‖·‖) x := by + obtain rfl | ht := eq_or_ne t 0 + · by_cases hE : Nontrivial E + · rw [zero_smul] at h + exact not_differentiableAt_norm_zero E h |>.elim + · rw [not_nontrivial_iff_subsingleton] at hE + exact (hasFDerivAt_of_subsingleton _ _).differentiableAt + · exact differentiableAt_norm_smul ht |>.2 h + +theorem DifferentiableAt.fderiv_norm_self {x : E} (h : DifferentiableAt ℝ (‖·‖) x) : + fderiv ℝ (‖·‖) x x = ‖x‖ := by + rw [← h.lineDeriv_eq_fderiv, lineDeriv] + have this (t : ℝ) : ‖x + t • x‖ = |1 + t| * ‖x‖ := by + rw [← norm_eq_abs, ← norm_smul, add_smul, one_smul] + simp_rw [this] + rw [deriv_mul_const] + · conv_lhs => enter [1, 1]; change _root_.abs ∘ (fun t ↦ 1 + t) + rw [deriv.comp, deriv_abs, deriv_const_add] + · simp + · exact differentiableAt_abs (by norm_num) + · exact differentiableAt_id.const_add _ + · exact (differentiableAt_abs (by norm_num)).comp _ (differentiableAt_id.const_add _) + +variable (x t) in +theorem fderiv_norm_smul : + fderiv ℝ (‖·‖) (t • x) = (SignType.sign t : ℝ) • (fderiv ℝ (‖·‖) x) := by + by_cases hE : Nontrivial E + · by_cases hd : DifferentiableAt ℝ (‖·‖) x + · obtain rfl | ht := eq_or_ne t 0 + · simp only [zero_smul, _root_.sign_zero, SignType.coe_zero] + exact fderiv_zero_of_not_differentiableAt <| not_differentiableAt_norm_zero E + · rw [(hd.hasFDerivAt.hasFDerivAt_norm_smul ht).fderiv] + · rw [fderiv_zero_of_not_differentiableAt hd, fderiv_zero_of_not_differentiableAt] + · simp + · exact mt DifferentiableAt.differentiableAt_norm_of_smul hd + · rw [not_nontrivial_iff_subsingleton] at hE + simp_rw [(hasFDerivAt_of_subsingleton _ _).fderiv, smul_zero] + +theorem fderiv_norm_smul_pos (ht : 0 < t) : + fderiv ℝ (‖·‖) (t • x) = fderiv ℝ (‖·‖) x := by + simp [fderiv_norm_smul, ht] + +theorem fderiv_norm_smul_neg (ht : t < 0) : + fderiv ℝ (‖·‖) (t • x) = -fderiv ℝ (‖·‖) x := by + simp [fderiv_norm_smul, ht] + +theorem norm_fderiv_norm [Nontrivial E] (h : DifferentiableAt ℝ (‖·‖) x) : + ‖fderiv ℝ (‖·‖) x‖ = 1 := by + have : x ≠ 0 := fun hx ↦ not_differentiableAt_norm_zero E (hx ▸ h) + refine le_antisymm (NNReal.coe_one ▸ norm_fderiv_le_of_lipschitz ℝ lipschitzWith_one_norm) ?_ + apply le_of_mul_le_mul_right _ (norm_pos_iff.2 this) + calc + 1 * ‖x‖ = fderiv ℝ (‖·‖) x x := by rw [one_mul, h.fderiv_norm_self] + _ ≤ ‖fderiv ℝ (‖·‖) x x‖ := le_norm_self _ + _ ≤ ‖fderiv ℝ (‖·‖) x‖ * ‖x‖ := le_opNorm _ _ diff --git a/Mathlib/Analysis/Calculus/LHopital.lean b/Mathlib/Analysis/Calculus/LHopital.lean index 653749f47e528..610346d0119ac 100644 --- a/Mathlib/Analysis/Calculus/LHopital.lean +++ b/Mathlib/Analysis/Calculus/LHopital.lean @@ -143,7 +143,7 @@ theorem lhopital_zero_atTop_on_Ioi (hff' : ∀ x ∈ Ioi a, HasDerivAt f (f' x) ⟨lt_of_le_of_lt (le_max_left a 0) (lt_one_add _), lt_of_le_of_lt (le_max_right a 0) (lt_one_add _)⟩⟩ have fact1 : ∀ x : ℝ, x ∈ Ioo 0 a'⁻¹ → x ≠ 0 := fun _ hx => (ne_of_lt hx.1).symm - have fact2 : ∀ x ∈ Ioo 0 a'⁻¹, a < x⁻¹ := fun _ hx => lt_trans haa' ((lt_inv ha' hx.1).mpr hx.2) + have fact2 (x) (hx : x ∈ Ioo 0 a'⁻¹) : a < x⁻¹ := lt_trans haa' ((lt_inv_comm₀ ha' hx.1).mpr hx.2) have hdnf : ∀ x ∈ Ioo 0 a'⁻¹, HasDerivAt (f ∘ Inv.inv) (f' x⁻¹ * -(x ^ 2)⁻¹) x := fun x hx => comp x (hff' x⁻¹ <| fact2 x hx) (hasDerivAt_inv <| fact1 x hx) have hdng : ∀ x ∈ Ioo 0 a'⁻¹, HasDerivAt (g ∘ Inv.inv) (g' x⁻¹ * -(x ^ 2)⁻¹) x := fun x hx => diff --git a/Mathlib/Analysis/Calculus/Rademacher.lean b/Mathlib/Analysis/Calculus/Rademacher.lean index 260a55686a38c..67771abec0dce 100644 --- a/Mathlib/Analysis/Calculus/Rademacher.lean +++ b/Mathlib/Analysis/Calculus/Rademacher.lean @@ -386,3 +386,18 @@ theorem LipschitzWith.ae_differentiableAt {f : E → F} (h : LipschitzWith C f) ∀ᵐ x ∂μ, DifferentiableAt ℝ f x := by rw [← lipschitzOnWith_univ] at h simpa [differentiableWithinAt_univ] using h.ae_differentiableWithinAt_of_mem + +/-- In a real finite-dimensional normed vector space, + the norm is almost everywhere differentiable. -/ +theorem ae_differentiableAt_norm : + ∀ᵐ x ∂μ, DifferentiableAt ℝ (‖·‖) x := lipschitzWith_one_norm.ae_differentiableAt + +omit [MeasurableSpace E] in +/-- In a real finite-dimensional normed vector space, + the set of points where the norm is differentiable at is dense. -/ +theorem dense_differentiableAt_norm : + Dense {x : E | DifferentiableAt ℝ (‖·‖) x} := + let _ : MeasurableSpace E := borel E + have _ : BorelSpace E := ⟨rfl⟩ + let w := Basis.ofVectorSpace ℝ E + MeasureTheory.Measure.dense_of_ae (ae_differentiableAt_norm (μ := w.addHaar)) diff --git a/Mathlib/Analysis/Calculus/TangentCone.lean b/Mathlib/Analysis/Calculus/TangentCone.lean index 53b15bc072b28..f3ac765b70a30 100644 --- a/Mathlib/Analysis/Calculus/TangentCone.lean +++ b/Mathlib/Analysis/Calculus/TangentCone.lean @@ -81,7 +81,7 @@ theorem mem_tangentConeAt_of_pow_smul {r : 𝕜} (hr₀ : r ≠ 0) (hr : ‖r‖ (hs : ∀ᶠ n : ℕ in atTop, x + r ^ n • y ∈ s) : y ∈ tangentConeAt 𝕜 s x := by refine ⟨fun n ↦ (r ^ n)⁻¹, fun n ↦ r ^ n • y, hs, ?_, ?_⟩ · simp only [norm_inv, norm_pow, ← inv_pow] - exact tendsto_pow_atTop_atTop_of_one_lt <| one_lt_inv (norm_pos_iff.2 hr₀) hr + exact tendsto_pow_atTop_atTop_of_one_lt <| (one_lt_inv₀ (norm_pos_iff.2 hr₀)).2 hr · simp only [inv_smul_smul₀ (pow_ne_zero _ hr₀), tendsto_const_nhds] theorem tangentCone_univ : tangentConeAt 𝕜 univ x = univ := diff --git a/Mathlib/Analysis/Complex/RemovableSingularity.lean b/Mathlib/Analysis/Complex/RemovableSingularity.lean index 9702422a17c49..b6d73c9132378 100644 --- a/Mathlib/Analysis/Complex/RemovableSingularity.lean +++ b/Mathlib/Analysis/Complex/RemovableSingularity.lean @@ -48,7 +48,7 @@ theorem differentiableOn_compl_singleton_and_continuousAt_iff {f : ℂ → E} {s rcases eq_or_ne x c with (rfl | hne) · refine (analyticAt_of_differentiable_on_punctured_nhds_of_continuousAt ?_ hc).differentiableAt.differentiableWithinAt - refine eventually_nhdsWithin_iff.2 ((eventually_mem_nhds.2 hs).mono fun z hz hzx => ?_) + refine eventually_nhdsWithin_iff.2 ((eventually_mem_nhds_iff.2 hs).mono fun z hz hzx => ?_) exact hd.differentiableAt (inter_mem hz (isOpen_ne.mem_nhds hzx)) · simpa only [DifferentiableWithinAt, HasFDerivWithinAt, hne.nhdsWithin_diff_singleton] using hd x ⟨hx, hne⟩ diff --git a/Mathlib/Analysis/Convex/Basic.lean b/Mathlib/Analysis/Convex/Basic.lean index 1eb8cdb030baa..a06894150ee98 100644 --- a/Mathlib/Analysis/Convex/Basic.lean +++ b/Mathlib/Analysis/Convex/Basic.lean @@ -505,7 +505,8 @@ theorem convex_iff_div : theorem Convex.mem_smul_of_zero_mem (h : Convex 𝕜 s) {x : E} (zero_mem : (0 : E) ∈ s) (hx : x ∈ s) {t : 𝕜} (ht : 1 ≤ t) : x ∈ t • s := by rw [mem_smul_set_iff_inv_smul_mem₀ (zero_lt_one.trans_le ht).ne'] - exact h.smul_mem_of_zero_mem zero_mem hx ⟨inv_nonneg.2 (zero_le_one.trans ht), inv_le_one ht⟩ + exact h.smul_mem_of_zero_mem zero_mem hx + ⟨inv_nonneg.2 (zero_le_one.trans ht), inv_le_one_of_one_le₀ ht⟩ theorem Convex.exists_mem_add_smul_eq (h : Convex 𝕜 s) {x y : E} {p q : 𝕜} (hx : x ∈ s) (hy : y ∈ s) (hp : 0 ≤ p) (hq : 0 ≤ q) : ∃ z ∈ s, (p + q) • z = p • x + q • y := by diff --git a/Mathlib/Analysis/Convex/Between.lean b/Mathlib/Analysis/Convex/Between.lean index f9d6c966bcb5f..d1be44a83c178 100644 --- a/Mathlib/Analysis/Convex/Between.lean +++ b/Mathlib/Analysis/Convex/Between.lean @@ -5,6 +5,7 @@ Authors: Joseph Myers -/ import Mathlib.Algebra.CharP.Invertible import Mathlib.Algebra.Order.Interval.Set.Group +import Mathlib.Analysis.Convex.Basic import Mathlib.Analysis.Convex.Segment import Mathlib.LinearAlgebra.AffineSpace.FiniteDimensional import Mathlib.Tactic.FieldSimp @@ -128,6 +129,14 @@ variable {R} lemma mem_segment_iff_wbtw {x y z : V} : y ∈ segment R x z ↔ Wbtw R x y z := by rw [Wbtw, affineSegment_eq_segment] +alias ⟨_, Wbtw.mem_segment⟩ := mem_segment_iff_wbtw + +lemma Convex.mem_of_wbtw {p₀ p₁ p₂ : V} {s : Set V} (hs : Convex R s) (h₀₁₂ : Wbtw R p₀ p₁ p₂) + (h₀ : p₀ ∈ s) (h₂ : p₂ ∈ s) : p₁ ∈ s := hs.segment_subset h₀ h₂ h₀₁₂.mem_segment + +lemma AffineSubspace.mem_of_wbtw {s : AffineSubspace R P} {x y z : P} (hxyz : Wbtw R x y z) + (hx : x ∈ s) (hz : z ∈ s) : y ∈ s := by obtain ⟨ε, -, rfl⟩ := hxyz; exact lineMap_mem _ hx hz + theorem Wbtw.map {x y z : P} (h : Wbtw R x y z) (f : P →ᵃ[R] P') : Wbtw R (f x) (f y) (f z) := by rw [Wbtw, ← affineSegment_image] exact Set.mem_image_of_mem _ h @@ -392,7 +401,7 @@ theorem sbtw_one_zero_iff {x : R} : Sbtw R 1 x 0 ↔ x ∈ Set.Ioo (0 : R) 1 := theorem Wbtw.trans_left {w x y z : P} (h₁ : Wbtw R w y z) (h₂ : Wbtw R w x y) : Wbtw R w x z := by rcases h₁ with ⟨t₁, ht₁, rfl⟩ rcases h₂ with ⟨t₂, ht₂, rfl⟩ - refine ⟨t₂ * t₁, ⟨mul_nonneg ht₂.1 ht₁.1, mul_le_one ht₂.2 ht₁.1 ht₁.2⟩, ?_⟩ + refine ⟨t₂ * t₁, ⟨mul_nonneg ht₂.1 ht₁.1, mul_le_one₀ ht₂.2 ht₁.1 ht₁.2⟩, ?_⟩ rw [lineMap_apply, lineMap_apply, lineMap_vsub_left, smul_smul] theorem Wbtw.trans_right {w x y z : P} (h₁ : Wbtw R w x z) (h₂ : Wbtw R x y z) : Wbtw R w y z := by @@ -574,7 +583,7 @@ end LinearOrderedRing section LinearOrderedField -variable [LinearOrderedField R] [AddCommGroup V] [Module R V] [AddTorsor V P] +variable [LinearOrderedField R] [AddCommGroup V] [Module R V] [AddTorsor V P] {x y z : P} variable {R} theorem wbtw_iff_left_eq_or_right_mem_image_Ici {x y z : P} : @@ -583,14 +592,14 @@ theorem wbtw_iff_left_eq_or_right_mem_image_Ici {x y z : P} : · rcases h with ⟨r, ⟨hr0, hr1⟩, rfl⟩ rcases hr0.lt_or_eq with (hr0' | rfl) · rw [Set.mem_image] - refine Or.inr ⟨r⁻¹, one_le_inv hr0' hr1, ?_⟩ + refine .inr ⟨r⁻¹, (one_le_inv₀ hr0').2 hr1, ?_⟩ simp only [lineMap_apply, smul_smul, vadd_vsub] rw [inv_mul_cancel₀ hr0'.ne', one_smul, vsub_vadd] · simp · rcases h with (rfl | ⟨r, ⟨hr, rfl⟩⟩) · exact wbtw_self_left _ _ _ · rw [Set.mem_Ici] at hr - refine ⟨r⁻¹, ⟨inv_nonneg.2 (zero_le_one.trans hr), inv_le_one hr⟩, ?_⟩ + refine ⟨r⁻¹, ⟨inv_nonneg.2 (zero_le_one.trans hr), inv_le_one_of_one_le₀ hr⟩, ?_⟩ simp only [lineMap_apply, smul_smul, vadd_vsub] rw [inv_mul_cancel₀ (one_pos.trans_le hr).ne', one_smul, vsub_vadd] @@ -653,6 +662,12 @@ theorem Sbtw.left_mem_image_Ioi {x y z : P} (h : Sbtw R x y z) : theorem Sbtw.left_mem_affineSpan {x y z : P} (h : Sbtw R x y z) : x ∈ line[R, z, y] := h.symm.right_mem_affineSpan +lemma AffineSubspace.right_mem_of_wbtw {s : AffineSubspace R P} (hxyz : Wbtw R x y z) (hx : x ∈ s) + (hy : y ∈ s) (hxy : x ≠ y) : z ∈ s := by + obtain ⟨ε, -, rfl⟩ := hxyz + have hε : ε ≠ 0 := by rintro rfl; simp at hxy + simpa [hε] using lineMap_mem ε⁻¹ hx hy + theorem wbtw_smul_vadd_smul_vadd_of_nonneg_of_le (x : P) (v : V) {r₁ r₂ : R} (hr₁ : 0 ≤ r₁) (hr₂ : r₁ ≤ r₂) : Wbtw R x (r₁ • v +ᵥ x) (r₂ • v +ᵥ x) := by refine ⟨r₁ / r₂, ⟨div_nonneg hr₁ (hr₁.trans hr₂), div_le_one_of_le₀ hr₂ (hr₁.trans hr₂)⟩, ?_⟩ @@ -695,8 +710,8 @@ theorem Wbtw.trans_left_right {w x y z : P} (h₁ : Wbtw R w y z) (h₂ : Wbtw R refine ⟨(t₁ - t₂ * t₁) / (1 - t₂ * t₁), ⟨div_nonneg (sub_nonneg.2 (mul_le_of_le_one_left ht₁.1 ht₂.2)) - (sub_nonneg.2 (mul_le_one ht₂.2 ht₁.1 ht₁.2)), - div_le_one_of_le₀ (sub_le_sub_right ht₁.2 _) (sub_nonneg.2 (mul_le_one ht₂.2 ht₁.1 ht₁.2))⟩, + (sub_nonneg.2 (mul_le_one₀ ht₂.2 ht₁.1 ht₁.2)), div_le_one_of_le₀ + (sub_le_sub_right ht₁.2 _) (sub_nonneg.2 (mul_le_one₀ ht₂.2 ht₁.1 ht₁.2))⟩, ?_⟩ simp only [lineMap_apply, smul_smul, ← add_vadd, vsub_vadd_eq_vsub_sub, smul_sub, ← sub_smul, ← add_smul, vadd_vsub, vadd_right_cancel_iff, div_mul_eq_mul_div, div_sub_div_same] diff --git a/Mathlib/Analysis/Convex/Birkhoff.lean b/Mathlib/Analysis/Convex/Birkhoff.lean new file mode 100644 index 0000000000000..a754f5db135c9 --- /dev/null +++ b/Mathlib/Analysis/Convex/Birkhoff.lean @@ -0,0 +1,172 @@ +/- +Copyright (c) 2024 Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Bhavik Mehta +-/ + +import Mathlib.Analysis.Convex.Combination +import Mathlib.Combinatorics.Hall.Basic +import Mathlib.Data.Matrix.DoublyStochastic +import Mathlib.Tactic.Linarith + +/-! +# Birkhoff's theorem + +## Main statements + +* `doublyStochastic_eq_sum_perm`: If `M` is a doubly stochastic matrix, then it is a convex + combination of permutation matrices. +* `doublyStochastic_eq_convexHull_perm`: The set of doubly stochastic matrices is the convex hull + of the permutation matrices. + +## TODO + +* Show that the extreme points of doubly stochastic matrices are the permutation matrices. +* Show that for `x y : n → R`, `x` is majorized by `y` if and only if there is a doubly stochastic + matrix `M` such that `M *ᵥ y = x`. + +## Tags + +Doubly stochastic, Birkhoff's theorem, Birkhoff-von Neumann theorem +-/ + +open Finset Function Matrix + +variable {R n : Type*} [Fintype n] [DecidableEq n] + +section LinearOrderedSemifield + +variable [LinearOrderedSemifield R] {M : Matrix n n R} + +/-- +If M is a positive scalar multiple of a doubly stochastic matrix, then there is a permutation matrix +whose support is contained in the support of M. +-/ +private lemma exists_perm_eq_zero_implies_eq_zero [Nonempty n] {s : R} (hs : 0 < s) + (hM : ∃ M' ∈ doublyStochastic R n, M = s • M') : + ∃ σ : Equiv.Perm n, ∀ i j, M i j = 0 → σ.permMatrix R i j = 0 := by + rw [exists_mem_doublyStochastic_eq_smul_iff hs.le] at hM + let f (i : n) : Finset n := univ.filter (M i · ≠ 0) + have hf (A : Finset n) : A.card ≤ (A.biUnion f).card := by + have (i) : ∑ j ∈ f i, M i j = s := by simp [sum_subset (filter_subset _ _), hM.2.1] + have h₁ : ∑ i ∈ A, ∑ j ∈ f i, M i j = A.card * s := by simp [this] + have h₂ : ∑ i, ∑ j ∈ A.biUnion f, M i j = (A.biUnion f).card * s := by + simp [sum_comm (t := A.biUnion f), hM.2.2, mul_comm s] + suffices A.card * s ≤ (A.biUnion f).card * s by exact_mod_cast le_of_mul_le_mul_right this hs + rw [← h₁, ← h₂] + trans ∑ i ∈ A, ∑ j ∈ A.biUnion f, M i j + · refine sum_le_sum fun i hi => ?_ + exact sum_le_sum_of_subset_of_nonneg (subset_biUnion_of_mem f hi) (by simp [*]) + · exact sum_le_sum_of_subset_of_nonneg (by simp) fun _ _ _ => sum_nonneg fun j _ => hM.1 _ _ + obtain ⟨g, hg, hg'⟩ := (all_card_le_biUnion_card_iff_exists_injective f).1 hf + rw [Finite.injective_iff_bijective] at hg + refine ⟨Equiv.ofBijective g hg, fun i j hij => ?_⟩ + simp only [PEquiv.toMatrix_apply, Option.mem_def, ite_eq_right_iff, one_ne_zero, imp_false, + Equiv.toPEquiv_apply, Equiv.ofBijective_apply, Option.some.injEq] + rintro rfl + simpa [f, hij] using hg' i + +end LinearOrderedSemifield + +section LinearOrderedField + +variable [LinearOrderedField R] {M : Matrix n n R} + +/-- +If M is a scalar multiple of a doubly stochastic matrix, then it is a conical combination of +permutation matrices. This is most useful when M is a doubly stochastic matrix, in which case +the combination is convex. + +This particular formulation is chosen to make the inductive step easier: we no longer need to +rescale each time a permutation matrix is subtracted. +-/ +private lemma doublyStochastic_sum_perm_aux (M : Matrix n n R) + (s : R) (hs : 0 ≤ s) + (hM : ∃ M' ∈ doublyStochastic R n, M = s • M') : + ∃ w : Equiv.Perm n → R, (∀ σ, 0 ≤ w σ) ∧ ∑ σ, w σ • σ.permMatrix R = M := by + rcases isEmpty_or_nonempty n + case inl => exact ⟨1, by simp, Subsingleton.elim _ _⟩ + set d : ℕ := (Finset.univ.filter fun i : n × n => M i.1 i.2 ≠ 0).card with ← hd + clear_value d + induction d using Nat.strongRecOn generalizing M s + case ind d ih => + rcases eq_or_lt_of_le hs with rfl | hs' + case inl => + use 0 + simp only [zero_smul, exists_and_right] at hM + simp [hM] + obtain ⟨σ, hσ⟩ := exists_perm_eq_zero_implies_eq_zero hs' hM + obtain ⟨i, hi, hi'⟩ := exists_min_image _ (fun i => M i (σ i)) univ_nonempty + rw [exists_mem_doublyStochastic_eq_smul_iff hs] at hM + let N : Matrix n n R := M - M i (σ i) • σ.permMatrix R + have hMi' : 0 < M i (σ i) := (hM.1 _ _).lt_of_ne' fun h => by + simpa [Equiv.toPEquiv_apply] using hσ _ _ h + let s' : R := s - M i (σ i) + have hs' : 0 ≤ s' := by + simp only [s', sub_nonneg, ← hM.2.1 i] + exact single_le_sum (fun j _ => hM.1 i j) (by simp) + have : ∃ M' ∈ doublyStochastic R n, N = s' • M' := by + rw [exists_mem_doublyStochastic_eq_smul_iff hs'] + simp only [sub_apply, smul_apply, PEquiv.toMatrix_apply, Equiv.toPEquiv_apply, Option.mem_def, + Option.some.injEq, smul_eq_mul, mul_ite, mul_one, mul_zero, sub_nonneg, + sum_sub_distrib, sum_ite_eq, mem_univ, ↓reduceIte, N] + refine ⟨fun i' j => ?_, by simp [hM.2.1], by simp [← σ.eq_symm_apply, hM]⟩ + split + case isTrue h => exact (hi' i' (by simp)).trans_eq (by rw [h]) + case isFalse h => exact hM.1 _ _ + have hd' : (univ.filter fun i : n × n => N i.1 i.2 ≠ 0).card < d := by + rw [← hd] + refine card_lt_card ?_ + rw [ssubset_iff_of_subset (monotone_filter_right _ _)] + · simp only [ne_eq, mem_filter, mem_univ, true_and, Decidable.not_not, Prod.exists] + refine ⟨i, σ i, hMi'.ne', ?_⟩ + simp [N, Equiv.toPEquiv_apply] + · rintro ⟨i', j'⟩ hN' hM' + dsimp at hN' hM' + simp only [sub_apply, hM', smul_apply, PEquiv.toMatrix_apply, Equiv.toPEquiv_apply, + Option.mem_def, Option.some.injEq, smul_eq_mul, mul_ite, mul_one, mul_zero, zero_sub, + neg_eq_zero, ite_eq_right_iff, Classical.not_imp, N] at hN' + obtain ⟨rfl, _⟩ := hN' + linarith [hi' i' (by simp)] + obtain ⟨w, hw, hw'⟩ := ih _ hd' _ s' hs' this rfl + refine ⟨w + fun σ' => if σ' = σ then M i (σ i) else 0, ?_⟩ + simp only [Pi.add_apply, add_smul, sum_add_distrib, hw', ite_smul, zero_smul, + sum_ite_eq', mem_univ, ↓reduceIte, N, sub_add_cancel, and_true] + intro σ' + split <;> simp [add_nonneg, hw, hM.1] + +/-- +If M is a doubly stochastic matrix, then it is an convex combination of permutation matrices. Note +`doublyStochastic_eq_convexHull_permMatrix` shows `doublyStochastic n` is exactly the convex hull of +the permutation matrices, and this lemma is instead most useful for accessing the coefficients of +each permutation matrices directly. +-/ +lemma exists_eq_sum_perm_of_mem_doublyStochastic (hM : M ∈ doublyStochastic R n) : + ∃ w : Equiv.Perm n → R, (∀ σ, 0 ≤ w σ) ∧ ∑ σ, w σ = 1 ∧ ∑ σ, w σ • σ.permMatrix R = M := by + rcases isEmpty_or_nonempty n + case inl => exact ⟨fun _ => 1, by simp, by simp, Subsingleton.elim _ _⟩ + obtain ⟨w, hw1, hw3⟩ := doublyStochastic_sum_perm_aux M 1 (by simp) ⟨M, hM, by simp⟩ + refine ⟨w, hw1, ?_, hw3⟩ + inhabit n + have : ∑ j, ∑ σ : Equiv.Perm n, w σ • σ.permMatrix R default j = 1 := by + simp only [← smul_apply (m := n), ← Finset.sum_apply, hw3] + rw [sum_row_of_mem_doublyStochastic hM] + simpa [sum_comm (γ := n), Equiv.toPEquiv_apply] using this + +/-- +**Birkhoff's theorem** +The set of doubly stochastic matrices is the convex hull of the permutation matrices. Note +`exists_eq_sum_perm_of_mem_doublyStochastic` gives a convex weighting of each permutation matrix +directly. To show `doublyStochastic n` is convex, use `convex_doublyStochastic`. +-/ +theorem doublyStochastic_eq_convexHull_permMatrix : + doublyStochastic R n = convexHull R {σ.permMatrix R | σ : Equiv.Perm n} := by + refine (convexHull_min ?g1 convex_doublyStochastic).antisymm' fun M hM => ?g2 + case g1 => + rintro x ⟨h, rfl⟩ + exact permMatrix_mem_doublyStochastic + case g2 => + obtain ⟨w, hw1, hw2, hw3⟩ := exists_eq_sum_perm_of_mem_doublyStochastic hM + exact mem_convexHull_of_exists_fintype w (·.permMatrix R) hw1 hw2 (by simp) hw3 + +end LinearOrderedField diff --git a/Mathlib/Analysis/Convex/Gauge.lean b/Mathlib/Analysis/Convex/Gauge.lean index 7338a929f1f82..df43239a123de 100644 --- a/Mathlib/Analysis/Convex/Gauge.lean +++ b/Mathlib/Analysis/Convex/Gauge.lean @@ -316,7 +316,7 @@ theorem comap_gauge_nhds_zero_le (ha : Absorbent ℝ s) (hb : Bornology.IsVonNBo rcases (hb hu).exists_pos with ⟨r, hr₀, hr⟩ filter_upwards [preimage_mem_comap (gt_mem_nhds (inv_pos.2 hr₀))] with x (hx : gauge s x < r⁻¹) rcases exists_lt_of_gauge_lt ha hx with ⟨c, hc₀, hcr, y, hy, rfl⟩ - have hrc := (lt_inv hr₀ hc₀).2 hcr + have hrc := (lt_inv_comm₀ hr₀ hc₀).2 hcr rcases hr c⁻¹ (hrc.le.trans (le_abs_self _)) hy with ⟨z, hz, rfl⟩ simpa only [smul_inv_smul₀ hc₀.ne'] diff --git a/Mathlib/Analysis/Convex/Star.lean b/Mathlib/Analysis/Convex/Star.lean index 35d31a5d32be1..ad0ee8acf1ca5 100644 --- a/Mathlib/Analysis/Convex/Star.lean +++ b/Mathlib/Analysis/Convex/Star.lean @@ -386,7 +386,7 @@ theorem starConvex_iff_div : StarConvex 𝕜 x s ↔ ∀ ⦃y⦄, y ∈ s → theorem StarConvex.mem_smul (hs : StarConvex 𝕜 0 s) (hx : x ∈ s) {t : 𝕜} (ht : 1 ≤ t) : x ∈ t • s := by rw [mem_smul_set_iff_inv_smul_mem₀ (zero_lt_one.trans_le ht).ne'] - exact hs.smul_mem hx (by positivity) (inv_le_one ht) + exact hs.smul_mem hx (by positivity) (inv_le_one_of_one_le₀ ht) end AddCommGroup diff --git a/Mathlib/Analysis/Convex/Strict.lean b/Mathlib/Analysis/Convex/Strict.lean index 43a8732b0e6d3..1eabcd86ec380 100644 --- a/Mathlib/Analysis/Convex/Strict.lean +++ b/Mathlib/Analysis/Convex/Strict.lean @@ -365,7 +365,7 @@ theorem strictConvex_iff_div : theorem StrictConvex.mem_smul_of_zero_mem (hs : StrictConvex 𝕜 s) (zero_mem : (0 : E) ∈ s) (hx : x ∈ s) (hx₀ : x ≠ 0) {t : 𝕜} (ht : 1 < t) : x ∈ t • interior s := by rw [mem_smul_set_iff_inv_smul_mem₀ (by positivity)] - exact hs.smul_mem_of_zero_mem zero_mem hx hx₀ (by positivity) (inv_lt_one ht) + exact hs.smul_mem_of_zero_mem zero_mem hx hx₀ (by positivity) (inv_lt_one_of_one_lt₀ ht) end AddCommGroup diff --git a/Mathlib/Analysis/Convex/Uniform.lean b/Mathlib/Analysis/Convex/Uniform.lean index c57da9a0c876f..9e22c41ca4d16 100644 --- a/Mathlib/Analysis/Convex/Uniform.lean +++ b/Mathlib/Analysis/Convex/Uniform.lean @@ -72,7 +72,8 @@ theorem exists_forall_closed_ball_dist_add_le_two_sub (hε : 0 < ε) : have h₂ : ∀ z : E, ‖z‖ ≤ 1 → 1 - δ' ≤ ‖z‖ → ‖‖z‖⁻¹ • z - z‖ ≤ δ' := by rintro z hz hδz nth_rw 3 [← one_smul ℝ z] - rwa [← sub_smul, norm_smul_of_nonneg (sub_nonneg_of_le <| one_le_inv (hδ'.trans_le hδz) hz), + rwa [← sub_smul, + norm_smul_of_nonneg (sub_nonneg_of_le <| (one_le_inv₀ (hδ'.trans_le hδz)).2 hz), sub_mul, inv_mul_cancel₀ (hδ'.trans_le hδz).ne', one_mul, sub_le_comm] set x' := ‖x‖⁻¹ • x set y' := ‖y‖⁻¹ • y diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index e7a99ac85a5d0..63f73727e30da 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -579,7 +579,7 @@ open MeasureTheory Module /-- A measure `μ` has temperate growth if there is an `n : ℕ` such that `(1 + ‖x‖) ^ (- n)` is `μ`-integrable. -/ -class _root_.MeasureTheory.Measure.HasTemperateGrowth (μ : Measure D) : Prop := +class _root_.MeasureTheory.Measure.HasTemperateGrowth (μ : Measure D) : Prop where exists_integrable : ∃ (n : ℕ), Integrable (fun x ↦ (1 + ‖x‖) ^ (- (n : ℝ))) μ open Classical in @@ -870,8 +870,8 @@ def compCLM {g : D → E} (hg : g.HasTemperateGrowth) rw [mul_pow] have hN₁' := (lt_of_lt_of_le zero_lt_one hN₁).ne' gcongr - · exact le_trans (by simp [hC]) (le_self_pow (by simp [hC]) hN₁') - · refine le_self_pow (one_le_pow₀ ?_) hN₁' + · exact le_trans (by simp [hC]) (le_self_pow₀ (by simp [hC]) hN₁') + · refine le_self_pow₀ (one_le_pow₀ ?_) hN₁' simp only [le_add_iff_nonneg_right, norm_nonneg] have := norm_iteratedFDeriv_comp_le f.smooth' hg.1 le_top x hbound hgrowth' have hxk : ‖x‖ ^ k ≤ (1 + ‖x‖) ^ k := diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 16957cece15d7..0fbe303edbe46 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -413,7 +413,7 @@ lemma norm_iteratedFDeriv_fourierPowSMulRight · rw [← Nat.cast_pow, Nat.cast_le] calc n.descFactorial i ≤ n ^ i := Nat.descFactorial_le_pow _ _ _ ≤ (n + 1) ^ i := pow_le_pow_left (by omega) (by omega) i - _ ≤ (n + 1) ^ k := pow_le_pow_right (by omega) (Finset.mem_range_succ_iff.mp hi) + _ ≤ (n + 1) ^ k := pow_right_mono₀ (by omega) (Finset.mem_range_succ_iff.mp hi) · exact hv _ (by omega) _ (by omega) _ = (2 * n + 2) ^ k * (‖L‖^n * C) := by simp only [← Finset.sum_mul, ← Nat.cast_sum, Nat.sum_range_choose, mul_one, ← mul_assoc, diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 5e1e46a1e699c..8164bbd406fee 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -489,7 +489,7 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq_inner {u : E → F'} have : 0 < p⁻¹ - (n : ℝ)⁻¹ := NNReal.coe_lt_coe.mpr (pos_iff_ne_zero.mpr (inv_ne_zero hp'0)) |>.trans_eq hp' rwa [NNReal.coe_inv, sub_pos, - inv_lt_inv _ (zero_lt_one.trans_le (NNReal.coe_le_coe.mpr hp))] at this + inv_lt_inv₀ _ (zero_lt_one.trans_le (NNReal.coe_le_coe.mpr hp))] at this exact_mod_cast hn have h0n : 2 ≤ n := Nat.succ_le_of_lt <| Nat.one_lt_cast.mp <| hp.trans_lt h2p have hn : NNReal.IsConjExponent n n' := .conjExponent (by norm_cast) @@ -514,7 +514,7 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq_inner {u : E → F'} have h0p' : p' ≠ 0 := by suffices 0 < (p' : ℝ) from (show 0 < p' from this) |>.ne' rw [← inv_pos, hp', sub_pos] - exact inv_lt_inv_of_lt hq.pos h2p + exact inv_strictAnti₀ hq.pos h2p have h2q : 1 / n' - 1 / q = 1 / p' := by simp_rw (config := {zeta := false}) [one_div, hp'] rw [← hq.one_sub_inv, ← hn.coe.one_sub_inv, sub_sub_sub_cancel_left] @@ -684,7 +684,7 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_le [FiniteDimensional ℝ F] have : (q : ℝ≥0∞) ≤ p' := by have H : (p' : ℝ)⁻¹ ≤ (↑q)⁻¹ := trans hp' hpq norm_cast at H ⊢ - rwa [inv_le_inv] at H + rwa [inv_le_inv₀] at H · dsimp have : 0 < p⁻¹ - (finrank ℝ E : ℝ≥0)⁻¹ := by simp only [tsub_pos_iff_lt] diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 96955c637d709..1e37689bab1f9 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -548,8 +548,30 @@ theorem inner_re_symm (x y : E) : re ⟪x, y⟫ = re ⟪y, x⟫ := by rw [← in theorem inner_im_symm (x y : E) : im ⟪x, y⟫ = -im ⟪y, x⟫ := by rw [← inner_conj_symm, conj_im] +section Algebra +variable {𝕝 : Type*} [CommSemiring 𝕝] [StarRing 𝕝] [Algebra 𝕝 𝕜] [Module 𝕝 E] + [IsScalarTower 𝕝 𝕜 E] [StarModule 𝕝 𝕜] + +/-- See `inner_smul_left` for the common special when `𝕜 = 𝕝`. -/ +lemma inner_smul_left_eq_star_smul (x y : E) (r : 𝕝) : ⟪r • x, y⟫ = r† • ⟪x, y⟫ := by + rw [← algebraMap_smul 𝕜 r, InnerProductSpace.smul_left, starRingEnd_apply, starRingEnd_apply, + ← algebraMap_star_comm, ← smul_eq_mul, algebraMap_smul] + +/-- Special case of `inner_smul_left_eq_star_smul` when the acting ring has a trivial star +(eg `ℕ`, `ℤ`, `ℚ≥0`, `ℚ`, `ℝ`). -/ +lemma inner_smul_left_eq_smul [TrivialStar 𝕝] (x y : E) (r : 𝕝) : ⟪r • x, y⟫ = r • ⟪x, y⟫ := by + rw [inner_smul_left_eq_star_smul, starRingEnd_apply, star_trivial] + +/-- See `inner_smul_right` for the common special when `𝕜 = 𝕝`. -/ +lemma inner_smul_right_eq_smul (x y : E) (r : 𝕝) : ⟪x, r • y⟫ = r • ⟪x, y⟫ := by + rw [← inner_conj_symm, inner_smul_left_eq_star_smul, starRingEnd_apply, starRingEnd_apply, + star_smul, star_star, ← starRingEnd_apply, inner_conj_symm] + +end Algebra + +/-- See `inner_smul_left_eq_star_smul` for the case of a general algebra action. -/ theorem inner_smul_left (x y : E) (r : 𝕜) : ⟪r • x, y⟫ = r† * ⟪x, y⟫ := - InnerProductSpace.smul_left _ _ _ + inner_smul_left_eq_star_smul .. theorem real_inner_smul_left (x y : F) (r : ℝ) : ⟪r • x, y⟫_ℝ = r * ⟪x, y⟫_ℝ := inner_smul_left _ _ _ @@ -557,8 +579,9 @@ theorem real_inner_smul_left (x y : F) (r : ℝ) : ⟪r • x, y⟫_ℝ = r * theorem inner_smul_real_left (x y : E) (r : ℝ) : ⟪(r : 𝕜) • x, y⟫ = r • ⟪x, y⟫ := by rw [inner_smul_left, conj_ofReal, Algebra.smul_def] -theorem inner_smul_right (x y : E) (r : 𝕜) : ⟪x, r • y⟫ = r * ⟪x, y⟫ := by - rw [← inner_conj_symm, inner_smul_left, RingHom.map_mul, conj_conj, inner_conj_symm] +/-- See `inner_smul_right_eq_smul` for the case of a general algebra action. -/ +theorem inner_smul_right (x y : E) (r : 𝕜) : ⟪x, r • y⟫ = r * ⟪x, y⟫ := + inner_smul_right_eq_smul .. theorem real_inner_smul_right (x y : F) (r : ℝ) : ⟪x, r • y⟫_ℝ = r * ⟪x, y⟫_ℝ := inner_smul_right _ _ _ @@ -2345,6 +2368,33 @@ theorem ContinuousLinearMap.reApplyInnerSelf_smul (T : E →L[𝕜] E) (x : E) { end ReApplyInnerSelf_Seminormed +section SeparationQuotient +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +theorem Inseparable.inner_eq_inner {x₁ x₂ y₁ y₂ : E} + (hx : Inseparable x₁ x₂) (hy : Inseparable y₁ y₂) : + inner x₁ y₁ = (inner x₂ y₂ : 𝕜) := + ((hx.prod hy).map continuous_inner).eq + +namespace SeparationQuotient + +instance : Inner 𝕜 (SeparationQuotient E) where + inner := SeparationQuotient.lift₂ Inner.inner fun _ _ _ _ => Inseparable.inner_eq_inner + +@[simp] +theorem inner_mk_mk (x y : E) : + inner (mk x) (mk y) = (inner x y : 𝕜) := rfl + +instance : InnerProductSpace 𝕜 (SeparationQuotient E) where + norm_sq_eq_inner := Quotient.ind norm_sq_eq_inner + conj_symm := Quotient.ind₂ inner_conj_symm + add_left := Quotient.ind fun x => Quotient.ind₂ <| inner_add_left x + smul_left := Quotient.ind₂ inner_smul_left + +end SeparationQuotient + +end SeparationQuotient + section UniformSpace.Completion variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] @@ -2361,11 +2411,11 @@ open UniformSpace Function instance toInner {𝕜' E' : Type*} [TopologicalSpace 𝕜'] [UniformSpace E'] [Inner 𝕜' E'] : Inner 𝕜' (Completion E') where - inner := curry <| (isDenseInducing_coe.prod isDenseInducing_coe).extend (uncurry inner) + inner := curry <| (isDenseInducing_coe.prodMap isDenseInducing_coe).extend (uncurry inner) @[simp] theorem inner_coe (a b : E) : inner (a : Completion E) (b : Completion E) = (inner a b : 𝕜) := - (isDenseInducing_coe.prod isDenseInducing_coe).extend_eq + (isDenseInducing_coe.prodMap isDenseInducing_coe).extend_eq (continuous_inner : Continuous (uncurry inner : E × E → 𝕜)) (a, b) protected theorem continuous_inner : @@ -2378,7 +2428,7 @@ protected theorem continuous_inner : rw [Completion.toInner, inner, uncurry_curry _] change Continuous - (((isDenseInducing_toCompl E).prod (isDenseInducing_toCompl E)).extend fun p : E × E => + (((isDenseInducing_toCompl E).prodMap (isDenseInducing_toCompl E)).extend fun p : E × E => inner' p.1 p.2) exact (isDenseInducing_toCompl E).extend_Z_bilin (isDenseInducing_toCompl E) this diff --git a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean index 868496a0b3462..049ba1532a238 100644 --- a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean +++ b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean @@ -49,7 +49,7 @@ variable {α : 𝕜} {A B : E →ₗ[𝕜] E} theorem eigenspace_invariant_of_commute (hAB : A ∘ₗ B = B ∘ₗ A) (α : 𝕜) : ∀ v ∈ (eigenspace A α), (B v ∈ eigenspace A α) := by intro v hv - rw [eigenspace, mem_ker, sub_apply, Module.algebraMap_end_apply, ← comp_apply A B v, hAB, + rw [eigenspace_def, mem_ker, sub_apply, smul_apply, one_apply, ← comp_apply A B v, hAB, comp_apply B A v, ← map_smul, ← map_sub, hv, map_zero] at * /--The simultaneous eigenspaces of a pair of commuting symmetric operators form an diff --git a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean index 4ab4ac3e7b9f6..13396afdb97d0 100644 --- a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean +++ b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean @@ -103,7 +103,7 @@ variable (hT : Dense (T.domain : Set E)) /-- The unique continuous extension of the operator `adjointDomainMkCLM` to `E`. -/ def adjointDomainMkCLMExtend (y : T.adjointDomain) : E →L[𝕜] 𝕜 := (T.adjointDomainMkCLM y).extend (Submodule.subtypeL T.domain) hT.denseRange_val - isUniformEmbedding_subtype_val.toUniformInducing + isUniformEmbedding_subtype_val.isUniformInducing @[simp] theorem adjointDomainMkCLMExtend_apply (y : T.adjointDomain) (x : T.domain) : diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index 17223b1faf998..28ef290eefc60 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -863,12 +863,13 @@ theorem orthogonalProjection_orthogonalProjection_of_le {U V : Submodule 𝕜 E} /-- Given a monotone family `U` of complete submodules of `E` and a fixed `x : E`, the orthogonal projection of `x` on `U i` tends to the orthogonal projection of `x` on `(⨆ i, U i).topologicalClosure` along `atTop`. -/ -theorem orthogonalProjection_tendsto_closure_iSup [CompleteSpace E] {ι : Type*} [SemilatticeSup ι] - (U : ι → Submodule 𝕜 E) [∀ i, CompleteSpace (U i)] (hU : Monotone U) (x : E) : +theorem orthogonalProjection_tendsto_closure_iSup {ι : Type*} [Preorder ι] + (U : ι → Submodule 𝕜 E) [∀ i, HasOrthogonalProjection (U i)] + [HasOrthogonalProjection (⨆ i, U i).topologicalClosure] (hU : Monotone U) (x : E) : Filter.Tendsto (fun i => (orthogonalProjection (U i) x : E)) atTop (𝓝 (orthogonalProjection (⨆ i, U i).topologicalClosure x : E)) := by - cases isEmpty_or_nonempty ι - · exact tendsto_of_isEmpty + refine .of_neBot_imp fun h ↦ ?_ + cases atTop_neBot_iff.mp h let y := (orthogonalProjection (⨆ i, U i).topologicalClosure x : E) have proj_x : ∀ i, orthogonalProjection (U i) x = orthogonalProjection (U i) y := fun i => (orthogonalProjection_orthogonalProjection_of_le @@ -890,14 +891,15 @@ theorem orthogonalProjection_tendsto_closure_iSup [CompleteSpace E] {ι : Type*} /-- Given a monotone family `U` of complete submodules of `E` with dense span supremum, and a fixed `x : E`, the orthogonal projection of `x` on `U i` tends to `x` along `at_top`. -/ -theorem orthogonalProjection_tendsto_self [CompleteSpace E] {ι : Type*} [SemilatticeSup ι] - (U : ι → Submodule 𝕜 E) [∀ t, CompleteSpace (U t)] (hU : Monotone U) (x : E) +theorem orthogonalProjection_tendsto_self {ι : Type*} [Preorder ι] + (U : ι → Submodule 𝕜 E) [∀ t, HasOrthogonalProjection (U t)] (hU : Monotone U) (x : E) (hU' : ⊤ ≤ (⨆ t, U t).topologicalClosure) : Filter.Tendsto (fun t => (orthogonalProjection (U t) x : E)) atTop (𝓝 x) := by - rw [← eq_top_iff] at hU' + have : HasOrthogonalProjection (⨆ i, U i).topologicalClosure := by + rw [top_unique hU'] + infer_instance convert orthogonalProjection_tendsto_closure_iSup U hU x - rw [orthogonalProjection_eq_self_iff.mpr _] - rw [hU'] + rw [eq_comm, orthogonalProjection_eq_self_iff, top_unique hU'] trivial /-- The orthogonal complement satisfies `Kᗮᗮᗮ = Kᗮ`. -/ diff --git a/Mathlib/Analysis/InnerProductSpace/l2Space.lean b/Mathlib/Analysis/InnerProductSpace/l2Space.lean index 9eb925f593b03..54a08fc3d190b 100644 --- a/Mathlib/Analysis/InnerProductSpace/l2Space.lean +++ b/Mathlib/Analysis/InnerProductSpace/l2Space.lean @@ -249,7 +249,7 @@ protected theorem range_linearIsometry [∀ i, CompleteSpace (G i)] : rintro i x ⟨x, rfl⟩ use lp.single 2 i x exact hV.linearIsometry_apply_single x - exact hV.linearIsometry.isometry.uniformInducing.isComplete_range.isClosed + exact hV.linearIsometry.isometry.isUniformInducing.isComplete_range.isClosed end OrthogonalFamily diff --git a/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean b/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean index 5667478982a21..08b82218840f4 100644 --- a/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean +++ b/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean @@ -131,7 +131,7 @@ theorem balancedHull.balanced (s : Set E) : Balanced 𝕜 (balancedHull 𝕜 s) simp_rw [balancedHull, smul_set_iUnion₂, subset_def, mem_iUnion₂] rintro x ⟨r, hr, hx⟩ rw [← smul_assoc] at hx - exact ⟨a • r, (SeminormedRing.norm_mul _ _).trans (mul_le_one ha (norm_nonneg r) hr), hx⟩ + exact ⟨a • r, (SeminormedRing.norm_mul _ _).trans (mul_le_one₀ ha (norm_nonneg r) hr), hx⟩ end Module @@ -158,7 +158,7 @@ theorem balancedCoreAux_balanced (h0 : (0 : E) ∈ balancedCoreAux 𝕜 s) : intro r hr have h'' : 1 ≤ ‖a⁻¹ • r‖ := by rw [norm_smul, norm_inv] - exact one_le_mul_of_one_le_of_one_le (one_le_inv (norm_pos_iff.mpr h) ha) hr + exact one_le_mul_of_one_le_of_one_le ((one_le_inv₀ (norm_pos_iff.mpr h)).2 ha) hr have h' := hy (a⁻¹ • r) h'' rwa [smul_assoc, mem_inv_smul_set_iff₀ h] at h' @@ -167,7 +167,7 @@ theorem balancedCoreAux_maximal (h : t ⊆ s) (ht : Balanced 𝕜 t) : t ⊆ bal rw [mem_smul_set_iff_inv_smul_mem₀ (norm_pos_iff.mp <| zero_lt_one.trans_le hr)] refine h (ht.smul_mem ?_ hx) rw [norm_inv] - exact inv_le_one hr + exact inv_le_one_of_one_le₀ hr theorem balancedCore_subset_balancedCoreAux : balancedCore 𝕜 s ⊆ balancedCoreAux 𝕜 s := balancedCoreAux_maximal (balancedCore_subset s) (balancedCore_balanced s) @@ -185,7 +185,7 @@ theorem subset_balancedCore (ht : (0 : E) ∈ t) (hst : ∀ a : 𝕜, ‖a‖ rw [subset_set_smul_iff₀ (norm_pos_iff.mp <| zero_lt_one.trans_le ha)] apply hst rw [norm_inv] - exact inv_le_one ha + exact inv_le_one_of_one_le₀ ha end NormedField diff --git a/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean b/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean index a61d1ac3cdcf5..d1b8d23fb2ca3 100644 --- a/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean +++ b/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean @@ -109,7 +109,7 @@ theorem LinearMap.continuousAt_zero_of_locally_bounded (f : E →ₛₗ[σ] F) refine (bE1 (n + 1)).2.smul_mem ?_ hx have h' : 0 < (n : ℝ) + 1 := n.cast_add_one_pos rw [norm_inv, ← Nat.cast_one, ← Nat.cast_add, RCLike.norm_natCast, Nat.cast_add, - Nat.cast_one, inv_le h' zero_lt_one] + Nat.cast_one, inv_le_comm₀ h' zero_lt_one] simp intro n hn -- The converse direction follows from continuity of the scalar multiplication diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index d8e1833217bd0..ebc7bb9056c0f 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -272,11 +272,11 @@ theorem harm_mean_le_geom_mean_weighted (w z : ι → ℝ) (hs : s.Nonempty) (hw have s_pos : 0 < ∑ i in s, w i * (z i)⁻¹ := sum_pos (fun i hi => mul_pos (hw i hi) (inv_pos.2 (hz i hi))) hs norm_num at this - rw [← inv_le_inv s_pos p_pos] at this + rw [← inv_le_inv₀ s_pos p_pos] at this apply le_trans this have p_pos₂ : 0 < (∏ i in s, (z i) ^ w i)⁻¹ := inv_pos.2 (prod_pos fun i hi => rpow_pos_of_pos ((hz i hi)) _ ) - rw [← inv_inv (∏ i in s, z i ^ w i), inv_le_inv p_pos p_pos₂, ← Finset.prod_inv_distrib] + rw [← inv_inv (∏ i in s, z i ^ w i), inv_le_inv₀ p_pos p_pos₂, ← Finset.prod_inv_distrib] gcongr · exact fun i hi ↦ inv_nonneg.mpr (Real.rpow_nonneg (le_of_lt (hz i hi)) _) · rw [Real.inv_rpow]; apply fun i hi ↦ le_of_lt (hz i hi); assumption @@ -821,7 +821,7 @@ lemma inner_le_weight_mul_Lp_of_nonneg (s : Finset ι) {p : ℝ} (hp : 1 ≤ p) obtain rfl | hp := hp.eq_or_lt · simp have hp₀ : 0 < p := by positivity - have hp₁ : p⁻¹ < 1 := inv_lt_one hp + have hp₁ : p⁻¹ < 1 := inv_lt_one_of_one_lt₀ hp by_cases H : (∑ i ∈ s, w i) ^ (1 - p⁻¹) = 0 ∨ (∑ i ∈ s, w i * f i ^ p) ^ p⁻¹ = 0 · replace H : (∀ i ∈ s, w i = 0) ∨ ∀ i ∈ s, w i = 0 ∨ f i = 0 := by simpa [hp₀, hp₁, hp₀.not_lt, hp₁.not_lt, sum_eq_zero_iff_of_nonneg] using H diff --git a/Mathlib/Analysis/Normed/Affine/Isometry.lean b/Mathlib/Analysis/Normed/Affine/Isometry.lean index f0545ea11cc87..67f6a72b44673 100644 --- a/Mathlib/Analysis/Normed/Affine/Isometry.lean +++ b/Mathlib/Analysis/Normed/Affine/Isometry.lean @@ -69,9 +69,9 @@ theorem linear_eq_linearIsometry : f.linear = f.linearIsometry.toLinearMap := by ext rfl -instance : FunLike (P →ᵃⁱ[𝕜] P₂) P P₂ := - { coe := fun f => f.toFun, - coe_injective' := fun f g => by cases f; cases g; simp } +instance : FunLike (P →ᵃⁱ[𝕜] P₂) P P₂ where + coe f := f.toFun + coe_injective' f g := by cases f; cases g; simp @[simp] theorem coe_toAffineMap : ⇑f.toAffineMap = f := by @@ -282,16 +282,16 @@ theorem linear_eq_linear_isometry : e.linear = e.linearIsometryEquiv.toLinearEqu ext rfl -instance : EquivLike (P ≃ᵃⁱ[𝕜] P₂) P P₂ := - { coe := fun f => f.toFun - inv := fun f => f.invFun - left_inv := fun f => f.left_inv - right_inv := fun f => f.right_inv, - coe_injective' := fun f g h _ => by - cases f - cases g - congr - simpa [DFunLike.coe_injective.eq_iff] using h } +instance : EquivLike (P ≃ᵃⁱ[𝕜] P₂) P P₂ where + coe f := f.toFun + inv f := f.invFun + left_inv f := f.left_inv + right_inv f := f.right_inv + coe_injective' f g h _ := by + cases f + cases g + congr + simpa [DFunLike.coe_injective.eq_iff] using h @[simp] theorem coe_mk (e : P ≃ᵃ[𝕜] P₂) (he : ∀ x, ‖e.linear x‖ = ‖x‖) : ⇑(mk e he) = e := diff --git a/Mathlib/Analysis/Normed/Algebra/Unitization.lean b/Mathlib/Analysis/Normed/Algebra/Unitization.lean index 19b11e9819409..de8b39cdbddce 100644 --- a/Mathlib/Analysis/Normed/Algebra/Unitization.lean +++ b/Mathlib/Analysis/Normed/Algebra/Unitization.lean @@ -184,8 +184,8 @@ open scoped Uniformity Topology theorem uniformity_eq_aux : 𝓤[instUniformSpaceProd.comap <| addEquiv 𝕜 A] = 𝓤 (Unitization 𝕜 A) := by - have key : UniformInducing (addEquiv 𝕜 A) := - antilipschitzWith_addEquiv.uniformInducing lipschitzWith_addEquiv.uniformContinuous + have key : IsUniformInducing (addEquiv 𝕜 A) := + antilipschitzWith_addEquiv.isUniformInducing lipschitzWith_addEquiv.uniformContinuous rw [← key.comap_uniformity] rfl @@ -202,7 +202,7 @@ instance instUniformSpace : UniformSpace (Unitization 𝕜 A) := /-- The natural equivalence between `Unitization 𝕜 A` and `𝕜 × A` as a uniform equivalence. -/ def uniformEquivProd : (Unitization 𝕜 A) ≃ᵤ (𝕜 × A) := - Equiv.toUniformEquivOfUniformInducing (addEquiv 𝕜 A) ⟨rfl⟩ + Equiv.toUniformEquivOfIsUniformInducing (addEquiv 𝕜 A) ⟨rfl⟩ /-- The bornology on `Unitization 𝕜 A` is inherited from `𝕜 × A`. -/ instance instBornology : Bornology (Unitization 𝕜 A) := diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index cb044efa45b52..405b9ae12e4fa 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -701,7 +701,7 @@ lemma norm_eq_one_iff_ne_zero_of_discrete {x : 𝕜} : ‖x‖ = 1 ↔ x ≠ 0 : · push_neg at h rcases h.eq_or_lt with h|h · rw [h] - replace h := norm_inv x ▸ inv_lt_one h + replace h := norm_inv x ▸ inv_lt_one_of_one_lt₀ h rw [← inv_inj, inv_one, ← norm_inv] exact H (by simpa) h' h obtain ⟨k, hk⟩ : ∃ k : ℕ, ‖x‖ ^ k < ε := exists_pow_lt_of_lt_one εpos h @@ -850,7 +850,7 @@ def NontriviallyNormedField.ofNormNeOne {𝕜 : Type*} [h' : NormedField 𝕜] rcases hx1.lt_or_lt with hlt | hlt · use x⁻¹ rw [norm_inv] - exact one_lt_inv (norm_pos_iff.2 hx) hlt + exact (one_lt_inv₀ (norm_pos_iff.2 hx)).2 hlt · exact ⟨x, hlt⟩ instance Real.normedCommRing : NormedCommRing ℝ := @@ -899,16 +899,16 @@ theorem nnnorm_norm [SeminormedAddCommGroup α] (a : α) : ‖‖a‖‖₊ = rw [Real.nnnorm_of_nonneg (norm_nonneg a)]; rfl /-- A restatement of `MetricSpace.tendsto_atTop` in terms of the norm. -/ -theorem NormedAddCommGroup.tendsto_atTop [Nonempty α] [SemilatticeSup α] {β : Type*} - [SeminormedAddCommGroup β] {f : α → β} {b : β} : +theorem NormedAddCommGroup.tendsto_atTop [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] + {β : Type*} [SeminormedAddCommGroup β] {f : α → β} {b : β} : Tendsto f atTop (𝓝 b) ↔ ∀ ε, 0 < ε → ∃ N, ∀ n, N ≤ n → ‖f n - b‖ < ε := (atTop_basis.tendsto_iff Metric.nhds_basis_ball).trans (by simp [dist_eq_norm]) /-- A variant of `NormedAddCommGroup.tendsto_atTop` that uses `∃ N, ∀ n > N, ...` rather than `∃ N, ∀ n ≥ N, ...` -/ -theorem NormedAddCommGroup.tendsto_atTop' [Nonempty α] [SemilatticeSup α] [NoMaxOrder α] - {β : Type*} [SeminormedAddCommGroup β] {f : α → β} {b : β} : +theorem NormedAddCommGroup.tendsto_atTop' [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] + [NoMaxOrder α] {β : Type*} [SeminormedAddCommGroup β] {f : α → β} {b : β} : Tendsto f atTop (𝓝 b) ↔ ∀ ε, 0 < ε → ∃ N, ∀ n, N < n → ‖f n - b‖ < ε := (atTop_basis_Ioi.tendsto_iff Metric.nhds_basis_ball).trans (by simp [dist_eq_norm]) diff --git a/Mathlib/Analysis/Normed/Field/Lemmas.lean b/Mathlib/Analysis/Normed/Field/Lemmas.lean index 7b97a37837eb7..a2236fabdefec 100644 --- a/Mathlib/Analysis/Normed/Field/Lemmas.lean +++ b/Mathlib/Analysis/Normed/Field/Lemmas.lean @@ -148,7 +148,7 @@ instance Pi.normedCommutativeRing {π : ι → Type*} [Fintype ι] [∀ i, Norme end NormedCommRing -- see Note [lower instance priority] -instance (priority := 100) semi_normed_ring_top_monoid [NonUnitalSeminormedRing α] : +instance (priority := 100) NonUnitalSeminormedRing.toContinuousMul [NonUnitalSeminormedRing α] : ContinuousMul α := ⟨continuous_iff_continuousAt.2 fun x => tendsto_iff_norm_sub_tendsto_zero.2 <| by @@ -174,9 +174,37 @@ instance (priority := 100) semi_normed_ring_top_monoid [NonUnitalSeminormedRing -- see Note [lower instance priority] /-- A seminormed ring is a topological ring. -/ -instance (priority := 100) semi_normed_top_ring [NonUnitalSeminormedRing α] : +instance (priority := 100) NonUnitalSeminormedRing.toTopologicalRing [NonUnitalSeminormedRing α] : TopologicalRing α where +namespace SeparationQuotient + +instance [NonUnitalSeminormedRing α] : NonUnitalNormedRing (SeparationQuotient α) where + __ : NonUnitalRing (SeparationQuotient α) := inferInstance + __ : NormedAddCommGroup (SeparationQuotient α) := inferInstance + norm_mul := Quotient.ind₂ norm_mul_le + +instance [NonUnitalSeminormedCommRing α] : NonUnitalNormedCommRing (SeparationQuotient α) where + __ : NonUnitalCommRing (SeparationQuotient α) := inferInstance + __ : NormedAddCommGroup (SeparationQuotient α) := inferInstance + norm_mul := Quotient.ind₂ norm_mul_le + +instance [SeminormedRing α] : NormedRing (SeparationQuotient α) where + __ : Ring (SeparationQuotient α) := inferInstance + __ : NormedAddCommGroup (SeparationQuotient α) := inferInstance + norm_mul := Quotient.ind₂ norm_mul_le + +instance [SeminormedCommRing α] : NormedCommRing (SeparationQuotient α) where + __ : CommRing (SeparationQuotient α) := inferInstance + __ : NormedAddCommGroup (SeparationQuotient α) := inferInstance + norm_mul := Quotient.ind₂ norm_mul_le + +instance [SeminormedAddCommGroup α] [One α] [NormOneClass α] : + NormOneClass (SeparationQuotient α) where + norm_one := norm_one (α := α) + +end SeparationQuotient + section NormedDivisionRing variable [NormedDivisionRing α] {a : α} diff --git a/Mathlib/Analysis/Normed/Field/UnitBall.lean b/Mathlib/Analysis/Normed/Field/UnitBall.lean index 2817b533b7761..146e37b0e5af2 100644 --- a/Mathlib/Analysis/Normed/Field/UnitBall.lean +++ b/Mathlib/Analysis/Normed/Field/UnitBall.lean @@ -51,7 +51,7 @@ def Subsemigroup.unitClosedBall (𝕜 : Type*) [NonUnitalSeminormedRing 𝕜] : carrier := closedBall 0 1 mul_mem' hx hy := by rw [mem_closedBall_zero_iff] at * - exact (norm_mul_le _ _).trans (mul_le_one hx (norm_nonneg _) hy) + exact (norm_mul_le _ _).trans (mul_le_one₀ hx (norm_nonneg _) hy) instance Metric.unitClosedBall.semigroup [NonUnitalSeminormedRing 𝕜] : Semigroup (closedBall (0 : 𝕜) 1) := diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 1eba1978859f4..04e1413e4bea8 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -770,6 +770,16 @@ theorem continuous_norm' : Continuous fun a : E => ‖a‖ := by theorem continuous_nnnorm' : Continuous fun a : E => ‖a‖₊ := continuous_norm'.subtype_mk _ +set_option linter.docPrime false in +@[to_additive Inseparable.norm_eq_norm] +theorem Inseparable.norm_eq_norm' {u v : E} (h : Inseparable u v) : ‖u‖ = ‖v‖ := + h.map continuous_norm' |>.eq + +set_option linter.docPrime false in +@[to_additive Inseparable.nnnorm_eq_nnnorm] +theorem Inseparable.nnnorm_eq_nnnorm' {u v : E} (h : Inseparable u v) : ‖u‖₊ = ‖v‖₊ := + h.map continuous_nnnorm' |>.eq + @[to_additive] theorem mem_closure_one_iff_norm {x : E} : x ∈ closure ({1} : Set E) ↔ ‖x‖ = 0 := by rw [← closedBall_zero', mem_closedBall_one_iff, (norm_nonneg' x).le_iff_eq] diff --git a/Mathlib/Analysis/Normed/Group/Hom.lean b/Mathlib/Analysis/Normed/Group/Hom.lean index 941289a09e507..36439765ebe21 100644 --- a/Mathlib/Analysis/Normed/Group/Hom.lean +++ b/Mathlib/Analysis/Normed/Group/Hom.lean @@ -82,7 +82,7 @@ def ofLipschitz (f : V₁ →+ V₂) {K : ℝ≥0} (h : LipschitzWith K f) : Nor instance funLike : FunLike (NormedAddGroupHom V₁ V₂) V₁ V₂ where coe := toFun - coe_injective' := fun f g h => by cases f; cases g; congr + coe_injective' f g h := by cases f; cases g; congr -- Porting note: moved this declaration up so we could get a `FunLike` instance sooner. instance toAddMonoidHomClass : AddMonoidHomClass (NormedAddGroupHom V₁ V₂) V₁ V₂ where diff --git a/Mathlib/Analysis/Normed/Group/SemiNormedGrp.lean b/Mathlib/Analysis/Normed/Group/SemiNormedGrp.lean index 2e56d519dc3b2..f63a09356056d 100644 --- a/Mathlib/Analysis/Normed/Group/SemiNormedGrp.lean +++ b/Mathlib/Analysis/Normed/Group/SemiNormedGrp.lean @@ -56,7 +56,7 @@ instance (M : SemiNormedGrp) : SeminormedAddCommGroup M := -- Porting note (#10754): added instance instance funLike {V W : SemiNormedGrp} : FunLike (V ⟶ W) V W where coe := (forget SemiNormedGrp).map - coe_injective' := fun f g h => by cases f; cases g; congr + coe_injective' f g h := by cases f; cases g; congr instance toAddMonoidHomClass {V W : SemiNormedGrp} : AddMonoidHomClass (V ⟶ W) V W where map_add f := f.map_add' diff --git a/Mathlib/Analysis/Normed/Group/Seminorm.lean b/Mathlib/Analysis/Normed/Group/Seminorm.lean index ca8445c51637b..06d80361d16b3 100644 --- a/Mathlib/Analysis/Normed/Group/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Group/Seminorm.lean @@ -48,7 +48,7 @@ open Set open NNReal -variable {ι R R' E F G : Type*} +variable {R R' E F G : Type*} /-- A seminorm on an additive group `G` is a function `f : G → ℝ` that preserves zero, is subadditive and such that `f (-x) = f x` for all `x`. -/ @@ -325,7 +325,7 @@ end Group section CommGroup -variable [CommGroup E] [CommGroup F] (p q : GroupSeminorm E) (x y : E) +variable [CommGroup E] [CommGroup F] (p q : GroupSeminorm E) (x : E) @[to_additive] theorem comp_mul_le (f g : F →* E) : p.comp (f * g) ≤ p.comp f + p.comp g := fun _ => @@ -381,7 +381,7 @@ end GroupSeminorm see that `SMul R ℝ` should be fixed because `ℝ` is fixed. -/ namespace AddGroupSeminorm -variable [AddGroup E] [SMul R ℝ] [SMul R ℝ≥0] [IsScalarTower R ℝ≥0 ℝ] (p : AddGroupSeminorm E) +variable [AddGroup E] [SMul R ℝ] [SMul R ℝ≥0] [IsScalarTower R ℝ≥0 ℝ] instance toOne [DecidableEq E] : One (AddGroupSeminorm E) := ⟨{ toFun := fun x => if x = 0 then 0 else 1 @@ -435,7 +435,7 @@ namespace NonarchAddGroupSeminorm section AddGroup -variable [AddGroup E] [AddGroup F] [AddGroup G] {p q : NonarchAddGroupSeminorm E} +variable [AddGroup E] {p q : NonarchAddGroupSeminorm E} instance funLike : FunLike (NonarchAddGroupSeminorm E) E ℝ where coe f := f.toFun @@ -477,7 +477,7 @@ theorem coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := theorem coe_lt_coe : (p : E → ℝ) < q ↔ p < q := Iff.rfl -variable (p q) (f : F →+ E) +variable (p q) instance : Zero (NonarchAddGroupSeminorm E) := ⟨{ toFun := 0 @@ -522,7 +522,7 @@ end AddGroup section AddCommGroup -variable [AddCommGroup E] [AddCommGroup F] (p q : NonarchAddGroupSeminorm E) (x y : E) +variable [AddCommGroup E] theorem add_bddBelow_range_add {p q : NonarchAddGroupSeminorm E} {x : E} : BddBelow (range fun y => p y + q (x - y)) := @@ -653,7 +653,7 @@ namespace GroupNorm section Group -variable [Group E] [Group F] [Group G] {p q : GroupNorm E} +variable [Group E] {p q : GroupNorm E} @[to_additive] instance funLike : FunLike (GroupNorm E) E ℝ where @@ -703,7 +703,7 @@ theorem coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := theorem coe_lt_coe : (p : E → ℝ) < q ↔ p < q := Iff.rfl -variable (p q) (f : F →* E) +variable (p q) @[to_additive] instance : Add (GroupNorm E) := @@ -787,7 +787,7 @@ namespace NonarchAddGroupNorm section AddGroup -variable [AddGroup E] [AddGroup F] {p q : NonarchAddGroupNorm E} +variable [AddGroup E] {p q : NonarchAddGroupNorm E} instance funLike : FunLike (NonarchAddGroupNorm E) E ℝ where coe f := f.toFun @@ -829,7 +829,7 @@ theorem coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := theorem coe_lt_coe : (p : E → ℝ) < q ↔ p < q := Iff.rfl -variable (p q) (f : F →+ E) +variable (p q) instance : Sup (NonarchAddGroupNorm E) := ⟨fun p q => diff --git a/Mathlib/Analysis/Normed/Group/Uniform.lean b/Mathlib/Analysis/Normed/Group/Uniform.lean index 7ec9757016b86..f8e72a4535b45 100644 --- a/Mathlib/Analysis/Normed/Group/Uniform.lean +++ b/Mathlib/Analysis/Normed/Group/Uniform.lean @@ -365,6 +365,29 @@ instance (priority := 100) SeminormedCommGroup.to_uniformGroup : UniformGroup E instance (priority := 100) SeminormedCommGroup.toTopologicalGroup : TopologicalGroup E := inferInstance +/-! ### SeparationQuotient -/ + +namespace SeparationQuotient + +@[to_additive instNorm] +instance instMulNorm : Norm (SeparationQuotient E) where + norm := lift Norm.norm fun _ _ h => h.norm_eq_norm' + +set_option linter.docPrime false in +@[to_additive (attr := simp) norm_mk] +theorem norm_mk' (p : E) : ‖mk p‖ = ‖p‖ := rfl + +@[to_additive] +instance : NormedCommGroup (SeparationQuotient E) where + __ : CommGroup (SeparationQuotient E) := instCommGroup + dist_eq := Quotient.ind₂ dist_eq_norm_div + +set_option linter.docPrime false in +@[to_additive (attr := simp) nnnorm_mk] +theorem nnnorm_mk' (p : E) : ‖mk p‖₊ = ‖p‖₊ := rfl + +end SeparationQuotient + @[to_additive] theorem cauchySeq_prod_of_eventually_eq {u v : ℕ → E} {N : ℕ} (huv : ∀ n ≥ N, u n = v n) (hv : CauchySeq fun n => ∏ k ∈ range (n + 1), v k) : diff --git a/Mathlib/Analysis/Normed/Lp/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean index cff87b17991f8..decc915289980 100644 --- a/Mathlib/Analysis/Normed/Lp/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -413,8 +413,8 @@ theorem antilipschitzWith_equiv_aux : rw [this, ENNReal.coe_rpow_of_nonneg _ nonneg] theorem aux_uniformity_eq : 𝓤 (PiLp p β) = 𝓤[Pi.uniformSpace _] := by - have A : UniformInducing (WithLp.equiv p (∀ i, β i)) := - (antilipschitzWith_equiv_aux p β).uniformInducing + have A : IsUniformInducing (WithLp.equiv p (∀ i, β i)) := + (antilipschitzWith_equiv_aux p β).isUniformInducing (lipschitzWith_equiv_aux p β).uniformContinuous have : (fun x : PiLp p β × PiLp p β => (WithLp.equiv p _ x.fst, WithLp.equiv p _ x.snd)) = id := by ext i <;> rfl diff --git a/Mathlib/Analysis/Normed/Lp/ProdLp.lean b/Mathlib/Analysis/Normed/Lp/ProdLp.lean index 9cc05f4783196..23d86ced09013 100644 --- a/Mathlib/Analysis/Normed/Lp/ProdLp.lean +++ b/Mathlib/Analysis/Normed/Lp/ProdLp.lean @@ -424,8 +424,8 @@ theorem prod_antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricS theorem prod_aux_uniformity_eq [PseudoEMetricSpace α] [PseudoEMetricSpace β] : 𝓤 (WithLp p (α × β)) = 𝓤[instUniformSpaceProd] := by - have A : UniformInducing (WithLp.equiv p (α × β)) := - (prod_antilipschitzWith_equiv_aux p α β).uniformInducing + have A : IsUniformInducing (WithLp.equiv p (α × β)) := + (prod_antilipschitzWith_equiv_aux p α β).isUniformInducing (prod_lipschitzWith_equiv_aux p α β).uniformContinuous have : (fun x : WithLp p (α × β) × WithLp p (α × β) => ((WithLp.equiv p (α × β)) x.fst, (WithLp.equiv p (α × β)) x.snd)) = id := by diff --git a/Mathlib/Analysis/Normed/Module/Basic.lean b/Mathlib/Analysis/Normed/Module/Basic.lean index b35a5c4e3baf1..d1bb26f18316d 100644 --- a/Mathlib/Analysis/Normed/Module/Basic.lean +++ b/Mathlib/Analysis/Normed/Module/Basic.lean @@ -119,6 +119,9 @@ instance Pi.normedSpace {ι : Type*} {E : ι → Type*} [Fintype ι] [∀ i, Sem NNReal.mul_finset_sup] exact Finset.sup_mono_fun fun _ _ => norm_smul_le a _ +instance SeparationQuotient.instNormedSpace : NormedSpace 𝕜 (SeparationQuotient E) where + norm_smul_le := norm_smul_le + instance MulOpposite.instNormedSpace : NormedSpace 𝕜 Eᵐᵒᵖ where norm_smul_le _ x := norm_smul_le _ x.unop @@ -348,6 +351,10 @@ instance Pi.normedAlgebra {ι : Type*} {E : ι → Type*} [Fintype ι] [∀ i, S variable [SeminormedRing E] [NormedAlgebra 𝕜 E] +instance SeparationQuotient.instNormedAlgebra : NormedAlgebra 𝕜 (SeparationQuotient E) where + __ : NormedSpace 𝕜 (SeparationQuotient E) := inferInstance + __ : Algebra 𝕜 (SeparationQuotient E) := inferInstance + instance MulOpposite.instNormedAlgebra {E : Type*} [SeminormedRing E] [NormedAlgebra 𝕜 E] : NormedAlgebra 𝕜 Eᵐᵒᵖ where __ := instAlgebra diff --git a/Mathlib/Analysis/Normed/Module/Completion.lean b/Mathlib/Analysis/Normed/Module/Completion.lean index c4377b5b2a68a..f0a8fdeed667d 100644 --- a/Mathlib/Analysis/Normed/Module/Completion.lean +++ b/Mathlib/Analysis/Normed/Module/Completion.lean @@ -74,10 +74,7 @@ instance [SeminormedRing A] : NormedRing (Completion A) where norm_mul x y := by induction x, y using induction_on₂ with | hp => - exact - isClosed_le (Continuous.comp continuous_norm continuous_mul) - (Continuous.comp _root_.continuous_mul - (Continuous.prod_map continuous_norm continuous_norm)) + apply isClosed_le <;> fun_prop | ih x y => simp only [← coe_mul, norm_coe] exact norm_mul_le x y diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index 77c5f4cf43a72..92b76035981f0 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -210,7 +210,7 @@ protected lemma embedding (f : F →ₛₗᵢ[σ₁₂] E₂) : Embedding f := f -- Should be `@[simp]` but it doesn't fire due to `lean4#3107`. theorem isComplete_image_iff [SemilinearIsometryClass 𝓕 σ₁₂ E E₂] (f : 𝓕) {s : Set E} : IsComplete (f '' s) ↔ IsComplete s := - _root_.isComplete_image_iff (SemilinearIsometryClass.isometry f).uniformInducing + _root_.isComplete_image_iff (SemilinearIsometryClass.isometry f).isUniformInducing @[simp] -- Should be replaced with `LinearIsometry.isComplete_image_iff` when `lean4#3107` is fixed. theorem isComplete_image_iff' (f : LinearIsometry σ₁₂ E E₂) {s : Set E} : diff --git a/Mathlib/Analysis/Normed/Ring/Units.lean b/Mathlib/Analysis/Normed/Ring/Units.lean index 6a9faca97ff80..86883d4c5d017 100644 --- a/Mathlib/Analysis/Normed/Ring/Units.lean +++ b/Mathlib/Analysis/Normed/Ring/Units.lean @@ -145,7 +145,7 @@ theorem inverse_one_sub_norm : (fun t : R => inverse (1 - t)) =O[𝓝 0] (fun _t have := tsum_geometric_le_of_norm_lt_one t ht' have : (1 - ‖t‖)⁻¹ ≤ 2 := by rw [← inv_inv (2 : ℝ)] - refine inv_le_inv_of_le (by norm_num) ?_ + refine inv_anti₀ (by norm_num) ?_ have : (2 : ℝ)⁻¹ + (2 : ℝ)⁻¹ = 1 := by ring linarith linarith diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 898e770696c72..4b87da345fe31 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -87,7 +87,7 @@ theorem ContinuousMultilinearMap.continuous_eval : let _ := TopologicalAddGroup.toUniformSpace F have := comm_topologicalAddGroup_is_uniform (G := F) refine (UniformOnFun.continuousOn_eval₂ fun m ↦ ?_).comp_continuous - (embedding_toUniformOnFun.continuous.prod_map continuous_id) fun (f, x) ↦ f.cont.continuousAt + (embedding_toUniformOnFun.continuous.prodMap continuous_id) fun (f, x) ↦ f.cont.continuousAt exact ⟨ball m 1, NormedSpace.isVonNBounded_of_isBounded _ isBounded_ball, ball_mem_nhds _ one_pos⟩ @@ -98,7 +98,7 @@ variable {G : Type*} [AddCommGroup G] [TopologicalSpace G] [Module 𝕜 G] [Cont lemma continuous_uncurry_of_multilinear : Continuous (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) := - ContinuousMultilinearMap.continuous_eval.comp <| .prod_map (map_continuous f) continuous_id + ContinuousMultilinearMap.continuous_eval.comp <| .prodMap (map_continuous f) continuous_id lemma continuousOn_uncurry_of_multilinear {s} : ContinuousOn (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) s := diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean index c0c4cc0b30b0d..5d787cb4490a6 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean @@ -264,7 +264,7 @@ theorem opNorm_le_of_shell' {f : E →SL[σ₁₂] F} {ε C : ℝ} (ε_pos : 0 < · refine opNorm_le_of_ball ε_pos hC fun x hx => hf x ?_ ?_ · simp [h0] · rwa [ball_zero_eq] at hx - · rw [← inv_inv c, norm_inv, inv_lt_one_iff_of_pos (norm_pos_iff.2 <| inv_ne_zero h0)] at hc + · rw [← inv_inv c, norm_inv, inv_lt_one₀ (norm_pos_iff.2 <| inv_ne_zero h0)] at hc refine opNorm_le_of_shell ε_pos hC hc ?_ rwa [norm_inv, div_eq_mul_inv, inv_inv] diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean index 9c9517852f03e..783b55750d8fb 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean @@ -181,7 +181,7 @@ variable [CompleteSpace F] (e : E →L[𝕜] Fₗ) (h_dense : DenseRange e) section -variable (h_e : UniformInducing e) +variable (h_e : IsUniformInducing e) /-- Extension of a continuous linear map `f : E →SL[σ₁₂] F`, with `E` a normed space and `F` a complete normed space, along a uniform and dense embedding `e : E →L[𝕜] Fₗ`. -/ @@ -230,7 +230,7 @@ variable {N : ℝ≥0} (h_e : ∀ x, ‖x‖ ≤ N * ‖e x‖) [RingHomIsometri /-- If a dense embedding `e : E →L[𝕜] G` expands the norm by a constant factor `N⁻¹`, then the norm of the extension of `f` along `e` is bounded by `N * ‖f‖`. -/ theorem opNorm_extend_le : - ‖f.extend e h_dense (isUniformEmbedding_of_bound _ h_e).toUniformInducing‖ ≤ N * ‖f‖ := by + ‖f.extend e h_dense (isUniformEmbedding_of_bound _ h_e).isUniformInducing‖ ≤ N * ‖f‖ := by -- Add `opNorm_le_of_dense`? refine opNorm_le_bound _ ?_ (isClosed_property h_dense (isClosed_le ?_ ?_) fun x ↦ ?_) · cases le_total 0 N with diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean index 9eae72e5fb45b..a42244f918fa7 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean @@ -121,7 +121,7 @@ examples. Any algebra with an approximate identity (e.g., $$L^1$$) is also regul This is a useful class because it gives rise to a nice norm on the unitization; in particular it is a C⋆-norm when the norm on `A` is a C⋆-norm. -/ -class _root_.RegularNormedAlgebra : Prop := +class _root_.RegularNormedAlgebra : Prop where /-- The left regular representation of the algebra on itself is an isometry. -/ isometry_mul' : Isometry (mul 𝕜 𝕜') diff --git a/Mathlib/Analysis/ODE/PicardLindelof.lean b/Mathlib/Analysis/ODE/PicardLindelof.lean index a1db140c6f1f2..9c46abf777613 100644 --- a/Mathlib/Analysis/ODE/PicardLindelof.lean +++ b/Mathlib/Analysis/ODE/PicardLindelof.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Winston Yin -/ import Mathlib.Analysis.SpecialFunctions.Integrals +import Mathlib.Topology.Algebra.Order.Floor import Mathlib.Topology.MetricSpace.Contracting /-! @@ -169,9 +170,12 @@ def toContinuousMap : v.FunSpace ↪ C(Icc v.tMin v.tMax, E) := instance : MetricSpace v.FunSpace := MetricSpace.induced toContinuousMap toContinuousMap.injective inferInstance -theorem uniformInducing_toContinuousMap : UniformInducing (@toContinuousMap _ _ _ v) := +theorem isUniformInducing_toContinuousMap : IsUniformInducing (@toContinuousMap _ _ _ v) := ⟨rfl⟩ +@[deprecated (since := "2024-10-05")] +alias uniformInducing_toContinuousMap := isUniformInducing_toContinuousMap + theorem range_toContinuousMap : range toContinuousMap = {f : C(Icc v.tMin v.tMax, E) | f v.t₀ = v.x₀ ∧ LipschitzWith v.C f} := by @@ -216,7 +220,7 @@ theorem dist_le_of_forall {f₁ f₂ : FunSpace v} {d : ℝ} (h : ∀ t, dist (f v.nonempty_Icc.to_subtype).2 h instance [CompleteSpace E] : CompleteSpace v.FunSpace := by - refine (completeSpace_iff_isComplete_range uniformInducing_toContinuousMap).2 + refine (completeSpace_iff_isComplete_range isUniformInducing_toContinuousMap).2 (IsClosed.isComplete ?_) rw [range_toContinuousMap, setOf_and] refine (isClosed_eq (ContinuousMap.continuous_eval_const _) continuous_const).inter ?_ @@ -301,7 +305,7 @@ section theorem exists_contracting_iterate : ∃ (N : ℕ) (K : _), ContractingWith K (FunSpace.next : v.FunSpace → v.FunSpace)^[N] := by - rcases ((Real.tendsto_pow_div_factorial_atTop (v.L * v.tDist)).eventually + rcases ((FloorSemiring.tendsto_pow_div_factorial_atTop (v.L * v.tDist)).eventually (gt_mem_nhds zero_lt_one)).exists with ⟨N, hN⟩ have : (0 : ℝ) ≤ (v.L * v.tDist) ^ N / N ! := div_nonneg (pow_nonneg (mul_nonneg v.L.2 v.tDist_nonneg) _) (Nat.cast_nonneg _) diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index aed28d9ca5915..a0365a83cfc7b 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -63,7 +63,7 @@ theorem le_sum_schlomilch' (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f theorem le_sum_condensed' (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f m) (n : ℕ) : (∑ k ∈ Ico 1 (2 ^ n), f k) ≤ ∑ k ∈ range n, 2 ^ k • f (2 ^ k) := by convert le_sum_schlomilch' hf (fun n => pow_pos zero_lt_two n) - (fun m n hm => pow_le_pow_right one_le_two hm) n using 2 + (fun m n hm => pow_right_mono₀ one_le_two hm) n using 2 simp [pow_succ, mul_two, two_mul] theorem le_sum_schlomilch (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f m) (h_pos : ∀ n, 0 < u n) @@ -98,7 +98,7 @@ theorem sum_schlomilch_le' (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f theorem sum_condensed_le' (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f m) (n : ℕ) : (∑ k ∈ range n, 2 ^ k • f (2 ^ (k + 1))) ≤ ∑ k ∈ Ico 2 (2 ^ n + 1), f k := by convert sum_schlomilch_le' hf (fun n => pow_pos zero_lt_two n) - (fun m n hm => pow_le_pow_right one_le_two hm) n using 2 + (fun m n hm => pow_right_mono₀ one_le_two hm) n using 2 simp [pow_succ, mul_two, two_mul] theorem sum_schlomilch_le {C : ℕ} (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f m) (h_pos : ∀ n, 0 < u n) @@ -283,7 +283,7 @@ theorem summable_nat_rpow_inv {p : ℝ} : (eventually_cofinite_ne 0)).exists apply hk₀ rw [← pos_iff_ne_zero, ← @Nat.cast_pos ℝ] at hk₀ - simpa [inv_lt_one_iff_of_pos (rpow_pos_of_pos hk₀ _), one_lt_rpow_iff_of_pos hk₀, hp, + simpa [inv_lt_one₀ (rpow_pos_of_pos hk₀ _), one_lt_rpow_iff_of_pos hk₀, hp, hp.not_lt, hk₀] using hk₁ @[simp] @@ -416,7 +416,7 @@ theorem sum_Ioo_inv_sq_le (k n : ℕ) : (∑ i ∈ Ioo k n, (i ^ 2 : α)⁻¹) have A : (1 : α) ≤ k + 1 := by simp only [le_add_iff_nonneg_left, Nat.cast_nonneg] simp_rw [← one_div] gcongr - simpa using pow_le_pow_right A one_le_two + simpa using pow_right_mono₀ A one_le_two _ = 2 / (k + 1) := by ring end diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 674e2c7b5df43..69097ee6513a0 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ import Mathlib.Algebra.Algebra.Field +import Mathlib.Algebra.BigOperators.Balance import Mathlib.Algebra.Order.BigOperators.Expect import Mathlib.Algebra.Order.Star.Basic import Mathlib.Analysis.CStarAlgebra.Basic @@ -40,6 +41,7 @@ their counterparts in `Mathlib/Analysis/Complex/Basic.lean` (which causes linter A few lemmas requiring heavier imports are in `Mathlib/Data/RCLike/Lemmas.lean`. -/ +open Fintype open scoped BigOperators ComplexConjugate section @@ -239,6 +241,13 @@ instance (priority := 100) charZero_rclike : CharZero K := lemma ofReal_expect {α : Type*} (s : Finset α) (f : α → ℝ) : 𝔼 i ∈ s, f i = 𝔼 i ∈ s, (f i : K) := map_expect (algebraMap ..) .. +@[norm_cast] +lemma ofReal_balance {ι : Type*} [Fintype ι] (f : ι → ℝ) (i : ι) : + ((balance f i : ℝ) : K) = balance ((↑) ∘ f) i := map_balance (algebraMap ..) .. + +@[simp] lemma ofReal_comp_balance {ι : Type*} [Fintype ι] (f : ι → ℝ) : + ofReal ∘ balance f = balance (ofReal ∘ f : ι → K) := funext <| ofReal_balance _ + /-! ### The imaginary unit, `I` -/ /-- The imaginary unit. -/ @@ -1077,7 +1086,7 @@ section /-- A mixin over a normed field, saying that the norm field structure is the same as `ℝ` or `ℂ`. To endow such a field with a compatible `RCLike` structure in a proof, use `letI := IsRCLikeNormedField.rclike 𝕜`.-/ -class IsRCLikeNormedField (𝕜 : Type*) [hk : NormedField 𝕜] : Prop := +class IsRCLikeNormedField (𝕜 : Type*) [hk : NormedField 𝕜] : Prop where out : ∃ h : RCLike 𝕜, hk = h.toNormedField instance (priority := 100) (𝕜 : Type*) [h : RCLike 𝕜] : IsRCLikeNormedField 𝕜 := ⟨⟨h, rfl⟩⟩ diff --git a/Mathlib/Analysis/SpecialFunctions/BinaryEntropy.lean b/Mathlib/Analysis/SpecialFunctions/BinaryEntropy.lean index 402c63a136cb7..c74cbff25c9e5 100644 --- a/Mathlib/Analysis/SpecialFunctions/BinaryEntropy.lean +++ b/Mathlib/Analysis/SpecialFunctions/BinaryEntropy.lean @@ -80,8 +80,8 @@ lemma binEntropy_two_inv_add (p : ℝ) : binEntropy (2⁻¹ + p) = binEntropy (2 lemma binEntropy_pos (hp₀ : 0 < p) (hp₁ : p < 1) : 0 < binEntropy p := by unfold binEntropy have : 0 < 1 - p := sub_pos.2 hp₁ - have : 0 < log p⁻¹ := log_pos <| one_lt_inv hp₀ hp₁ - have : 0 < log (1 - p)⁻¹ := log_pos <| one_lt_inv ‹_› (sub_lt_self _ hp₀) + have : 0 < log p⁻¹ := log_pos <| (one_lt_inv₀ hp₀).2 hp₁ + have : 0 < log (1 - p)⁻¹ := log_pos <| (one_lt_inv₀ ‹_›).2 (sub_lt_self _ hp₀) positivity lemma binEntropy_nonneg (hp₀ : 0 ≤ p) (hp₁ : p ≤ 1) : 0 ≤ binEntropy p := by @@ -397,7 +397,7 @@ lemma qaryEntropy_strictAntiOn (qLe2 : 2 ≤ q) : · exact qaryEntropy_continuous.continuousOn · intro p hp have : 2 ≤ (q : ℝ) := Nat.ofNat_le_cast.mpr qLe2 - have qinv_lt_1 : (q : ℝ)⁻¹ < 1 := inv_lt_one (by linarith) + have qinv_lt_1 : (q : ℝ)⁻¹ < 1 := inv_lt_one_of_one_lt₀ (by linarith) have zero_lt_1_sub_p : 0 < 1 - p := by simp_all only [sub_pos, hp.2, interior_Icc, mem_Ioo] simp only [one_div, interior_Icc, mem_Ioo] at hp rw [deriv_qaryEntropy (by linarith)] diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean b/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean index 7d994e5d897af..f5e05a1e202df 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean @@ -110,7 +110,7 @@ lemma norm_one_add_mul_inv_le {t : ℝ} (ht : t ∈ Set.Icc 0 1) {z : ℂ} (hz : ‖(1 + t * z)⁻¹‖ ≤ (1 - ‖z‖)⁻¹ := by rw [Set.mem_Icc] at ht rw [norm_inv, norm_eq_abs] - refine inv_le_inv_of_le (by linarith) ?_ + refine inv_anti₀ (by linarith) ?_ calc 1 - ‖z‖ _ ≤ 1 - t * ‖z‖ := by nlinarith [norm_nonneg z] diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index c5f7769a08f9d..ccc7141a90155 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -301,7 +301,7 @@ theorem abs_log_mul_self_lt (x : ℝ) (h1 : 0 < x) (h2 : x ≤ 1) : |log x * x| refine mul_nonneg ?_ h1.le rw [← log_inv] apply log_nonneg - rw [← le_inv h1 zero_lt_one, inv_one] + rw [← le_inv_comm₀ h1 zero_lt_one, inv_one] exact h2 rw [← abs_of_nonneg aux, neg_mul, abs_neg] at this exact this diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean index 529dea6298d19..841d4fc4de757 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean @@ -332,8 +332,7 @@ theorem continuousAt_ofReal_cpow (x : ℝ) (y : ℂ) (h : 0 < y.re ∨ x ≠ 0) ContinuousAt (fun p => (p.1 : ℂ) ^ p.2 : ℝ × ℂ → ℂ) (x, y) := by rcases lt_trichotomy (0 : ℝ) x with (hx | rfl | hx) · -- x > 0 : easy case - have : ContinuousAt (fun p => ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y) := - continuous_ofReal.continuousAt.prod_map continuousAt_id + have : ContinuousAt (fun p => ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y) := by fun_prop refine (continuousAt_cpow (Or.inl ?_)).comp this rwa [ofReal_re] · -- x = 0 : reduce to continuousAt_cpow_zero_of_re_pos @@ -341,15 +340,13 @@ theorem continuousAt_ofReal_cpow (x : ℝ) (y : ℂ) (h : 0 < y.re ∨ x ≠ 0) rw [ofReal_zero] apply continuousAt_cpow_zero_of_re_pos tauto - have B : ContinuousAt (fun p => ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) ⟨0, y⟩ := - continuous_ofReal.continuousAt.prod_map continuousAt_id + have B : ContinuousAt (fun p => ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) ⟨0, y⟩ := by fun_prop exact A.comp_of_eq B rfl · -- x < 0 : difficult case suffices ContinuousAt (fun p => (-(p.1 : ℂ)) ^ p.2 * exp (π * I * p.2) : ℝ × ℂ → ℂ) (x, y) by refine this.congr (eventually_of_mem (prod_mem_nhds (Iio_mem_nhds hx) univ_mem) ?_) exact fun p hp => (ofReal_cpow_of_nonpos (le_of_lt hp.1) p.2).symm - have A : ContinuousAt (fun p => ⟨-↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y) := - ContinuousAt.prod_map continuous_ofReal.continuousAt.neg continuousAt_id + have A : ContinuousAt (fun p => ⟨-↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y) := by fun_prop apply ContinuousAt.mul · refine (continuousAt_cpow (Or.inl ?_)).comp A rwa [neg_re, ofReal_re, neg_pos] @@ -394,7 +391,7 @@ theorem eventually_pow_one_div_le (x : ℝ≥0) {y : ℝ≥0} (hy : 1 < y) : refine eventually_atTop.2 ⟨m + 1, fun n hn => ?_⟩ simp only [one_div] simpa only [NNReal.rpow_inv_le_iff (Nat.cast_pos.2 <| m.succ_pos.trans_le hn), - NNReal.rpow_natCast] using hm.le.trans (pow_le_pow_right hy.le (m.le_succ.trans hn)) + NNReal.rpow_natCast] using hm.le.trans (pow_right_mono₀ hy.le (m.le_succ.trans hn)) end NNReal diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index 80400c08a6e5d..b712f2e3cd568 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -528,13 +528,13 @@ theorem monotoneOn_rpow_Ici_of_exponent_nonneg {r : ℝ} (hr : 0 ≤ r) : lemma rpow_lt_rpow_of_neg (hx : 0 < x) (hxy : x < y) (hz : z < 0) : y ^ z < x ^ z := by have := hx.trans hxy - rw [← inv_lt_inv, ← rpow_neg, ← rpow_neg] + rw [← inv_lt_inv₀, ← rpow_neg, ← rpow_neg] on_goal 1 => refine rpow_lt_rpow ?_ hxy (neg_pos.2 hz) all_goals positivity lemma rpow_le_rpow_of_nonpos (hx : 0 < x) (hxy : x ≤ y) (hz : z ≤ 0) : y ^ z ≤ x ^ z := by have := hx.trans_le hxy - rw [← inv_le_inv, ← rpow_neg, ← rpow_neg] + rw [← inv_le_inv₀, ← rpow_neg, ← rpow_neg] on_goal 1 => refine rpow_le_rpow ?_ hxy (neg_nonneg.2 hz) all_goals positivity @@ -592,7 +592,7 @@ theorem rpow_lt_rpow_of_exponent_neg {x y z : ℝ} (hy : 0 < y) (hxy : y < x) (h x ^ z < y ^ z := by have hx : 0 < x := hy.trans hxy rw [← neg_neg z, Real.rpow_neg (le_of_lt hx) (-z), Real.rpow_neg (le_of_lt hy) (-z), - inv_lt_inv (rpow_pos_of_pos hx _) (rpow_pos_of_pos hy _)] + inv_lt_inv₀ (rpow_pos_of_pos hx _) (rpow_pos_of_pos hy _)] exact Real.rpow_lt_rpow (by positivity) hxy <| neg_pos_of_neg hz theorem strictAntiOn_rpow_Ioi_of_exponent_neg {r : ℝ} (hr : r < 0) : diff --git a/Mathlib/Analysis/SpecialFunctions/Stirling.lean b/Mathlib/Analysis/SpecialFunctions/Stirling.lean index a1b4c1c5dd12d..35d9ac4c47d5b 100644 --- a/Mathlib/Analysis/SpecialFunctions/Stirling.lean +++ b/Mathlib/Analysis/SpecialFunctions/Stirling.lean @@ -102,12 +102,12 @@ theorem log_stirlingSeq_diff_le_geo_sum (n : ℕ) : · simp_rw [← _root_.pow_succ'] at this exact this rw [one_div, inv_pow] - exact inv_lt_one (one_lt_pow₀ ((lt_add_iff_pos_left 1).mpr <| by positivity) two_ne_zero) + exact inv_lt_one_of_one_lt₀ (one_lt_pow₀ (lt_add_of_pos_left _ <| by positivity) two_ne_zero) have hab (k : ℕ) : (1 : ℝ) / (2 * ↑(k + 1) + 1) * ((1 / (2 * ↑(n + 1) + 1)) ^ 2) ^ ↑(k + 1) ≤ (((1 : ℝ) / (2 * ↑(n + 1) + 1)) ^ 2) ^ ↑(k + 1) := by refine mul_le_of_le_one_left (pow_nonneg h_nonneg ↑(k + 1)) ?_ rw [one_div] - exact inv_le_one (le_add_of_nonneg_left <| by positivity) + exact inv_le_one_of_one_le₀ (le_add_of_nonneg_left <| by positivity) exact hasSum_le hab (log_stirlingSeq_diff_hasSum n) g /-- We have the bound `log (stirlingSeq n) - log (stirlingSeq (n+1))` ≤ 1/(4 n^2) diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean index ee71f3e92b789..6919f9dcbb728 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean @@ -228,7 +228,7 @@ theorem arctan_add_eq_add_pi {x y : ℝ} (h : 1 < x * y) (hx : 0 < x) : have hy : 0 < y := by have := mul_pos_iff.mp (zero_lt_one.trans h) simpa [hx, hx.asymm] - have k := arctan_add (mul_inv x y ▸ inv_lt_one h) + have k := arctan_add (mul_inv x y ▸ inv_lt_one_of_one_lt₀ h) rw [arctan_inv_of_pos hx, arctan_inv_of_pos hy, show _ + _ = π - (arctan x + arctan y) by ring, sub_eq_iff_eq_add, ← sub_eq_iff_eq_add', sub_eq_add_neg, ← arctan_neg, add_comm] at k convert k.symm using 3 diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean index e8e17232197f1..f3a9858bddf37 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean @@ -192,7 +192,7 @@ theorem lt_tan {x : ℝ} (h1 : 0 < x) (h2 : x < π / 2) : x < tan x := by apply lt_of_le_of_ne y.cos_sq_le_one rw [cos_sq'] simpa only [Ne, sub_eq_self, sq_eq_zero_iff] using (sin_pos hy).ne' - rwa [lt_inv, inv_one] + rwa [lt_inv_comm₀, inv_one] · exact zero_lt_one simpa only [sq, mul_self_pos] using this.ne' have mono := strictMonoOn_of_deriv_pos (convex_Ico 0 (π / 2)) tan_minus_id_cts deriv_pos diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean index e11d7973f3b2c..862ffd739dff4 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean @@ -45,8 +45,10 @@ theorem contDiff_sin {n} : ContDiff ℂ n sin := (((contDiff_neg.mul contDiff_const).cexp.sub (contDiff_id.mul contDiff_const).cexp).mul contDiff_const).div_const _ +@[fun_prop] theorem differentiable_sin : Differentiable ℂ sin := fun x => (hasDerivAt_sin x).differentiableAt +@[fun_prop] theorem differentiableAt_sin {x : ℂ} : DifferentiableAt ℂ sin x := differentiable_sin x @@ -70,8 +72,10 @@ theorem hasDerivAt_cos (x : ℂ) : HasDerivAt cos (-sin x) x := theorem contDiff_cos {n} : ContDiff ℂ n cos := ((contDiff_id.mul contDiff_const).cexp.add (contDiff_neg.mul contDiff_const).cexp).div_const _ +@[fun_prop] theorem differentiable_cos : Differentiable ℂ cos := fun x => (hasDerivAt_cos x).differentiableAt +@[fun_prop] theorem differentiableAt_cos {x : ℂ} : DifferentiableAt ℂ cos x := differentiable_cos x @@ -98,8 +102,10 @@ theorem hasDerivAt_sinh (x : ℂ) : HasDerivAt sinh (cosh x) x := theorem contDiff_sinh {n} : ContDiff ℂ n sinh := (contDiff_exp.sub contDiff_neg.cexp).div_const _ +@[fun_prop] theorem differentiable_sinh : Differentiable ℂ sinh := fun x => (hasDerivAt_sinh x).differentiableAt +@[fun_prop] theorem differentiableAt_sinh {x : ℂ} : DifferentiableAt ℂ sinh x := differentiable_sinh x @@ -123,8 +129,10 @@ theorem hasDerivAt_cosh (x : ℂ) : HasDerivAt cosh (sinh x) x := theorem contDiff_cosh {n} : ContDiff ℂ n cosh := (contDiff_exp.add contDiff_neg.cexp).div_const _ +@[fun_prop] theorem differentiable_cosh : Differentiable ℂ cosh := fun x => (hasDerivAt_cosh x).differentiableAt +@[fun_prop] theorem differentiableAt_cosh {x : ℂ} : DifferentiableAt ℂ cosh x := differentiable_cosh x @@ -482,8 +490,10 @@ theorem hasDerivAt_sin (x : ℝ) : HasDerivAt sin (cos x) x := theorem contDiff_sin {n} : ContDiff ℝ n sin := Complex.contDiff_sin.real_of_complex +@[fun_prop] theorem differentiable_sin : Differentiable ℝ sin := fun x => (hasDerivAt_sin x).differentiableAt +@[fun_prop] theorem differentiableAt_sin : DifferentiableAt ℝ sin x := differentiable_sin x @@ -500,8 +510,10 @@ theorem hasDerivAt_cos (x : ℝ) : HasDerivAt cos (-sin x) x := theorem contDiff_cos {n} : ContDiff ℝ n cos := Complex.contDiff_cos.real_of_complex +@[fun_prop] theorem differentiable_cos : Differentiable ℝ cos := fun x => (hasDerivAt_cos x).differentiableAt +@[fun_prop] theorem differentiableAt_cos : DifferentiableAt ℝ cos x := differentiable_cos x @@ -521,8 +533,10 @@ theorem hasDerivAt_sinh (x : ℝ) : HasDerivAt sinh (cosh x) x := theorem contDiff_sinh {n} : ContDiff ℝ n sinh := Complex.contDiff_sinh.real_of_complex +@[fun_prop] theorem differentiable_sinh : Differentiable ℝ sinh := fun x => (hasDerivAt_sinh x).differentiableAt +@[fun_prop] theorem differentiableAt_sinh : DifferentiableAt ℝ sinh x := differentiable_sinh x @@ -539,8 +553,10 @@ theorem hasDerivAt_cosh (x : ℝ) : HasDerivAt cosh (sinh x) x := theorem contDiff_cosh {n} : ContDiff ℝ n cosh := Complex.contDiff_cosh.real_of_complex +@[fun_prop] theorem differentiable_cosh : Differentiable ℝ cosh := fun x => (hasDerivAt_cosh x).differentiableAt +@[fun_prop] theorem differentiableAt_cosh : DifferentiableAt ℝ cosh x := differentiable_cosh x diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 370cf43a6f16a..c0a2df7d23e61 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -119,7 +119,7 @@ theorem tendsto_natCast_div_add_atTop {𝕜 : Type*} [DivisionRing 𝕜] [Topolo theorem tendsto_add_one_pow_atTop_atTop_of_pos [LinearOrderedSemiring α] [Archimedean α] {r : α} (h : 0 < r) : Tendsto (fun n : ℕ ↦ (r + 1) ^ n) atTop atTop := - tendsto_atTop_atTop_of_monotone' (fun _ _ ↦ pow_le_pow_right <| le_add_of_nonneg_left h.le) <| + tendsto_atTop_atTop_of_monotone' (pow_right_mono₀ <| le_add_of_nonneg_left h.le) <| not_bddAbove_iff.2 fun _ ↦ Set.exists_range_iff.2 <| add_one_pow_unbounded_of_pos _ h theorem tendsto_pow_atTop_atTop_of_one_lt [LinearOrderedRing α] [Archimedean α] {r : α} @@ -137,7 +137,7 @@ theorem tendsto_pow_atTop_nhds_zero_of_lt_one {𝕜 : Type*} [LinearOrderedField (fun hr ↦ (tendsto_add_atTop_iff_nat 1).mp <| by simp [_root_.pow_succ, ← hr, tendsto_const_nhds]) (fun hr ↦ - have := one_lt_inv hr h₂ |> tendsto_pow_atTop_atTop_of_one_lt + have := (one_lt_inv₀ hr).2 h₂ |> tendsto_pow_atTop_atTop_of_one_lt (tendsto_inv_atTop_zero.comp this).congr fun n ↦ by simp) @[deprecated (since := "2024-01-31")] alias tendsto_pow_atTop_nhds_0_of_lt_1 := tendsto_pow_atTop_nhds_zero_of_lt_one @@ -513,7 +513,7 @@ theorem summable_one_div_pow_of_le {m : ℝ} {f : ℕ → ℕ} (hm : 1 < m) (fi (summable_geometric_of_lt_one (one_div_nonneg.mpr (zero_le_one.trans hm.le)) ((one_div_lt (zero_lt_one.trans hm) zero_lt_one).mpr (one_div_one.le.trans_lt hm))) rw [div_pow, one_pow] - refine (one_div_le_one_div ?_ ?_).mpr (pow_le_pow_right hm.le (fi a)) <;> + refine (one_div_le_one_div ?_ ?_).mpr (pow_right_mono₀ hm.le (fi a)) <;> exact pow_pos (zero_lt_one.trans hm) _ /-! ### Positive sequences with small sums on countable types -/ diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index 68cd1103fc6e6..0d410a3b11645 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -222,12 +222,12 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc have B : c ^ 2 * ((1 : ℝ) - c⁻¹ ^ 2)⁻¹ ≤ c ^ 3 * (c - 1)⁻¹ := by rw [← div_eq_mul_inv, ← div_eq_mul_inv, div_le_div_iff _ (sub_pos.2 hc)] swap - · exact sub_pos.2 (pow_lt_one₀ (inv_nonneg.2 cpos.le) (inv_lt_one hc) two_ne_zero) + · exact sub_pos.2 (pow_lt_one₀ (inv_nonneg.2 cpos.le) (inv_lt_one_of_one_lt₀ hc) two_ne_zero) have : c ^ 3 = c ^ 2 * c := by ring simp only [mul_sub, this, mul_one, inv_pow, sub_le_sub_iff_left] rw [mul_assoc, mul_comm c, ← mul_assoc, mul_inv_cancel₀ (sq_pos_of_pos cpos).ne', one_mul] - simpa using pow_le_pow_right hc.le one_le_two - have C : c⁻¹ ^ 2 < 1 := pow_lt_one₀ (inv_nonneg.2 cpos.le) (inv_lt_one hc) two_ne_zero + simpa using pow_right_mono₀ hc.le one_le_two + have C : c⁻¹ ^ 2 < 1 := pow_lt_one₀ (inv_nonneg.2 cpos.le) (inv_lt_one_of_one_lt₀ hc) two_ne_zero calc (∑ i ∈ (range N).filter (j < c ^ ·), (1 : ℝ) / (c ^ i) ^ 2) ≤ ∑ i ∈ Ico ⌊Real.log j / Real.log c⌋₊ N, (1 : ℝ) / (c ^ i) ^ 2 := by @@ -270,7 +270,7 @@ theorem mul_pow_le_nat_floor_pow {c : ℝ} (hc : 1 < c) (i : ℕ) : (1 - c⁻¹) (1 - c⁻¹) * c ^ i = c ^ i - c ^ i * c⁻¹ := by ring _ ≤ c ^ i - 1 := by gcongr - simpa only [← div_eq_mul_inv, one_le_div cpos, pow_one] using le_self_pow hc.le hi + simpa only [← div_eq_mul_inv, one_le_div cpos, pow_one] using le_self_pow₀ hc.le hi _ ≤ ⌊c ^ i⌋₊ := (Nat.sub_one_lt_floor _).le /-- The sum of `1/⌊c^i⌋₊^2` above a threshold `j` is comparable to `1/j^2`, up to a multiplicative @@ -279,7 +279,7 @@ theorem sum_div_nat_floor_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : (∑ i ∈ (range N).filter (j < ⌊c ^ ·⌋₊), (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2) ≤ c ^ 5 * (c - 1)⁻¹ ^ 3 / j ^ 2 := by have cpos : 0 < c := zero_lt_one.trans hc - have A : 0 < 1 - c⁻¹ := sub_pos.2 (inv_lt_one hc) + have A : 0 < 1 - c⁻¹ := sub_pos.2 (inv_lt_one_of_one_lt₀ hc) calc (∑ i ∈ (range N).filter (j < ⌊c ^ ·⌋₊), (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2) ≤ ∑ i ∈ (range N).filter (j < c ^ ·), (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2 := by diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index a84b941a538ea..525b05e39a0cf 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -210,7 +210,7 @@ theorem tendsto_pow_const_mul_const_pow_of_abs_lt_one (k : ℕ) {r : ℝ} (hr : by_cases h0 : r = 0 · exact tendsto_const_nhds.congr' (mem_atTop_sets.2 ⟨1, fun n hn ↦ by simp [zero_lt_one.trans_le hn |>.ne', h0]⟩) - have hr' : 1 < |r|⁻¹ := one_lt_inv (abs_pos.2 h0) hr + have hr' : 1 < |r|⁻¹ := (one_lt_inv₀ (abs_pos.2 h0)).2 hr rw [tendsto_zero_iff_norm_tendsto_zero] simpa [div_eq_mul_inv] using tendsto_pow_const_div_const_pow_of_one_lt k hr' @@ -252,7 +252,7 @@ alias tendsto_pow_atTop_nhds_0_of_abs_lt_1 := tendsto_pow_atTop_nhds_zero_of_abs /-- A normed ring has summable geometric series if, for all `ξ` of norm `< 1`, the geometric series `∑ ξ ^ n` converges. This holds both in complete normed rings and in normed fields, providing a convenient abstraction of these two classes to avoid repeating the same proofs. -/ -class HasSummableGeomSeries (K : Type*) [NormedRing K] : Prop := +class HasSummableGeomSeries (K : Type*) [NormedRing K] : Prop where summable_geometric_of_norm_lt_one : ∀ (ξ : K), ‖ξ‖ < 1 → Summable (fun n ↦ ξ ^ n) lemma summable_geometric_of_norm_lt_one {K : Type*} [NormedRing K] [HasSummableGeomSeries K] @@ -908,6 +908,8 @@ theorem Real.summable_pow_div_factorial (x : ℝ) : Summable (fun n ↦ x ^ n / norm_div, Real.norm_natCast, Nat.cast_succ] _ ≤ ‖x‖ / (⌊‖x‖⌋₊ + 1) * ‖x ^ n / (n !)‖ := by gcongr +@[deprecated "`Real.tendsto_pow_div_factorial_atTop` has been deprecated, use +`FloorSemiring.tendsto_pow_div_factorial_atTop` instead" (since := "2024-10-05")] theorem Real.tendsto_pow_div_factorial_atTop (x : ℝ) : Tendsto (fun n ↦ x ^ n / n ! : ℕ → ℝ) atTop (𝓝 0) := (Real.summable_pow_div_factorial x).tendsto_atTop_zero diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 30d3ab3bb1913..3d9c1506b6275 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -14,7 +14,9 @@ import Mathlib.CategoryTheory.Yoneda We provide various useful constructors: * `mkOfHomEquiv` -* `mkOfUnitCounit` +* `mk'`: construct an adjunction from the data of a hom set equivalence, unit and counit natural + transformations together with proofs of the equalities `homEquiv_unit` and `homEquiv_counit` + relating them to each other. * `leftAdjointOfEquiv` / `rightAdjointOfEquiv` construct a left/right adjoint of a given functor given the action on objects and the relevant equivalence of morphism spaces. @@ -30,6 +32,44 @@ adjoint can be obtained as `F.rightAdjoint`. `toEquivalence` upgrades an adjunction to an equivalence, given witnesses that the unit and counit are pointwise isomorphisms. Conversely `Equivalence.toAdjunction` recovers the underlying adjunction from an equivalence. + +## Overview of the directory `CategoryTheory.Adjunction` + +* Adjoint lifting theorems are in the directory `Lifting`. +* The file `AdjointFunctorTheorems` proves the adjoint functor theorems. +* The file `Comma` shows that for a functor `G : D ⥤ C` the data of an initial object in each + `StructuredArrow` category on `G` is equivalent to a left adjoint to `G`, as well as the dual. +* The file `Evaluation` shows that products and coproducts are adjoint to evaluation of functors. +* The file `FullyFaithful` characterizes when adjoints are full or faithful in terms of the unit + and counit. +* The file `Limits` proves that left adjoints preserve colimits and right adjoints preserve limits. +* The file `Mates` establishes the bijection between the 2-cells + ``` + L₁ R₁ + C --→ D C ←-- D + G ↓ ↗ ↓ H G ↓ ↘ ↓ H + E --→ F E ←-- F + L₂ R₂ + ``` + where `L₁ ⊣ R₁` and `L₂ ⊣ R₂`. Specializing to a pair of adjoints `L₁ L₂ : C ⥤ D`, + `R₁ R₂ : D ⥤ C`, it provides equivalences `(L₂ ⟶ L₁) ≃ (R₁ ⟶ R₂)` and `(L₂ ≅ L₁) ≃ (R₁ ≅ R₂)`. +* The file `Opposites` contains constructions to relate adjunctions of functors to adjunctions of + their opposites. +* The file `Reflective` defines reflective functors, i.e. fully faithful right adjoints. Note that + many facts about reflective functors are proved in the earlier file `FullyFaithful`. +* The file `Restrict` defines the restriction of an adjunction along fully faithful functors. +* The file `Triple` proves that in an adjoint triple, the left adjoint is fully faithful if and + only if the right adjoint is. +* The file `Unique` proves uniqueness of adjoints. +* The file `Whiskering` proves that functors `F : D ⥤ E` and `G : E ⥤ D` with an adjunction + `F ⊣ G`, induce adjunctions between the functor categories `C ⥤ D` and `C ⥤ E`, + and the functor categories `E ⥤ C` and `D ⥤ C`. + +## Other files related to adjunctions + +* The file `CategoryTheory.Monad.Adjunction` develops the basic relationship between adjunctions + and (co)monads. There it is also shown that given an adjunction `L ⊣ R` and an isomorphism + `L ⋙ R ≅ 𝟭 C`, the unit is an isomorphism, and similarly for the counit. -/ @@ -59,8 +99,6 @@ hom set equivalence. To construct adjoints to a given functor, there are constructors `leftAdjointOfEquiv` and `adjunctionOfEquivLeft` (as well as their duals). -Uniqueness of adjoints is shown in `CategoryTheory.Adjunction.Unique`. - See . -/ structure Adjunction (F : C ⥤ D) (G : D ⥤ C) where diff --git a/Mathlib/CategoryTheory/Adjunction/Mates.lean b/Mathlib/CategoryTheory/Adjunction/Mates.lean index fa2243dda407c..ca344b3aa3796 100644 --- a/Mathlib/CategoryTheory/Adjunction/Mates.lean +++ b/Mathlib/CategoryTheory/Adjunction/Mates.lean @@ -13,11 +13,13 @@ import Mathlib.Tactic.ApplyFun This file establishes the bijection between the 2-cells +``` L₁ R₁ C --→ D C ←-- D G ↓ ↗ ↓ H G ↓ ↘ ↓ H E --→ F E ←-- F L₂ R₂ +``` where `L₁ ⊣ R₁` and `L₂ ⊣ R₂`. The corresponding natural transformations are called mates. diff --git a/Mathlib/CategoryTheory/Adjunction/Unique.lean b/Mathlib/CategoryTheory/Adjunction/Unique.lean index e911b3e4dcb59..9c4b07886fc7d 100644 --- a/Mathlib/CategoryTheory/Adjunction/Unique.lean +++ b/Mathlib/CategoryTheory/Adjunction/Unique.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta, Thomas Read, Andrew Yang, Dagur Asgeirsson, Joël Riou -/ -import Mathlib.CategoryTheory.Adjunction.Basic +import Mathlib.CategoryTheory.Adjunction.Mates /-! # Uniqueness of adjoints @@ -11,9 +11,6 @@ import Mathlib.CategoryTheory.Adjunction.Basic This file shows that adjoints are unique up to natural isomorphism. ## Main results -* `Adjunction.natTransEquiv` and `Adjunction.natIsoEquiv` If `F ⊣ G` and `F' ⊣ G'` are adjunctions, - then there are equivalences `(G ⟶ G') ≃ (F' ⟶ F)` and `(G ≅ G') ≃ (F' ≅ F)`. -Everything else is deduced from this: * `Adjunction.leftAdjointUniq` : If `F` and `F'` are both left adjoint to `G`, then they are naturally isomorphic. @@ -21,12 +18,6 @@ Everything else is deduced from this: * `Adjunction.rightAdjointUniq` : If `G` and `G'` are both right adjoint to `F`, then they are naturally isomorphic. -## TODO - -There some overlap with the file `Adjunction.Mates`. In particular, `natTransEquiv` is just a -special case of `mateEquiv`. However, before removing `natTransEquiv`, in favour of `mateEquiv`, -the latter needs some more API lemmas such as `natTransEquiv_apply_app`, `natTransEquiv_id`, etc. -in order to make automation work better in the rest of this file. -/ open CategoryTheory @@ -35,89 +26,9 @@ variable {C D : Type*} [Category C] [Category D] namespace CategoryTheory.Adjunction -/-- -If `F ⊣ G` and `F' ⊣ G'` are adjunctions, then giving a natural transformation `G ⟶ G'` is the -same as giving a natural transformation `F' ⟶ F`. --/ -@[simps] -def natTransEquiv {F F' : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G') : - (G ⟶ G') ≃ (F' ⟶ F) where - toFun f := { - app := fun X ↦ F'.map ((adj1.unit ≫ whiskerLeft F f).app X) ≫ adj2.counit.app _ - naturality := by - intro X Y g - simp only [← Category.assoc, ← Functor.map_comp] - erw [(adj1.unit ≫ (whiskerLeft F f)).naturality] - simp - } - invFun f := { - app := fun X ↦ adj2.unit.app (G.obj X) ≫ G'.map (f.app (G.obj X) ≫ adj1.counit.app X) - naturality := by - intro X Y g - erw [← adj2.unit_naturality_assoc] - simp only [← Functor.map_comp] - simp - } - left_inv f := by - ext X - simp only [Functor.comp_obj, NatTrans.comp_app, Functor.id_obj, whiskerLeft_app, - Functor.map_comp, Category.assoc, unit_naturality_assoc, right_triangle_components_assoc] - erw [← f.naturality (adj1.counit.app X), ← Category.assoc] - simp - right_inv f := by - ext - simp - -@[simp] -lemma natTransEquiv_id {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : - natTransEquiv adj adj (𝟙 _) = 𝟙 _ := by ext; simp - -@[simp] -lemma natTransEquiv_id_symm {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : - (natTransEquiv adj adj).symm (𝟙 _) = 𝟙 _ := by ext; simp - -@[simp] -lemma natTransEquiv_comp {F F' F'' : C ⥤ D} {G G' G'' : D ⥤ C} - (adj1 : F ⊣ G) (adj2 : F' ⊣ G') (adj3 : F'' ⊣ G'') (f : G ⟶ G') (g : G' ⟶ G'') : - natTransEquiv adj2 adj3 g ≫ natTransEquiv adj1 adj2 f = natTransEquiv adj1 adj3 (f ≫ g) := by - apply (natTransEquiv adj1 adj3).symm.injective - ext X - simp only [natTransEquiv_symm_apply_app, Functor.comp_obj, NatTrans.comp_app, - natTransEquiv_apply_app, Functor.id_obj, whiskerLeft_app, Functor.map_comp, Category.assoc, - unit_naturality_assoc, right_triangle_components_assoc, Equiv.symm_apply_apply, - ← g.naturality_assoc, ← g.naturality] - simp only [← Category.assoc, unit_naturality, Functor.comp_obj, right_triangle_components, - Category.comp_id, ← f.naturality, Category.id_comp] - -@[simp] -lemma natTransEquiv_comp_symm {F F' F'' : C ⥤ D} {G G' G'' : D ⥤ C} - (adj1 : F ⊣ G) (adj2 : F' ⊣ G') (adj3 : F'' ⊣ G'') (f : F' ⟶ F) (g : F'' ⟶ F') : - (natTransEquiv adj1 adj2).symm f ≫ (natTransEquiv adj2 adj3).symm g = - (natTransEquiv adj1 adj3).symm (g ≫ f) := by - apply (natTransEquiv adj1 adj3).injective - ext - simp - -/-- -If `F ⊣ G` and `F' ⊣ G'` are adjunctions, then giving a natural isomorphism `G ≅ G'` is the -same as giving a natural transformation `F' ≅ F`. --/ -@[simps] -def natIsoEquiv {F F' : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G') : - (G ≅ G') ≃ (F' ≅ F) where - toFun i := { - hom := natTransEquiv adj1 adj2 i.hom - inv := natTransEquiv adj2 adj1 i.inv - } - invFun i := { - hom := (natTransEquiv adj1 adj2).symm i.hom - inv := (natTransEquiv adj2 adj1).symm i.inv } - left_inv i := by simp - right_inv i := by simp - /-- If `F` and `F'` are both left adjoint to `G`, then they are naturally isomorphic. -/ def leftAdjointUniq {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G) : F ≅ F' := - (natIsoEquiv adj1 adj2 (Iso.refl _)).symm + ((conjugateIsoEquiv adj1 adj2).symm (Iso.refl G)).symm -- Porting note (#10618): removed simp as simp can prove this theorem homEquiv_leftAdjointUniq_hom_app {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G) @@ -141,9 +52,10 @@ theorem unit_leftAdjointUniq_hom_app theorem leftAdjointUniq_hom_counit {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G) : whiskerLeft G (leftAdjointUniq adj1 adj2).hom ≫ adj2.counit = adj1.counit := by ext x - simp only [Functor.comp_obj, Functor.id_obj, leftAdjointUniq, Iso.symm_hom, natIsoEquiv_apply_inv, - Iso.refl_inv, NatTrans.comp_app, whiskerLeft_app, natTransEquiv_apply_app, whiskerLeft_id', - Category.comp_id, Category.assoc] + simp only [Functor.comp_obj, Functor.id_obj, leftAdjointUniq, Iso.symm_hom, + conjugateIsoEquiv_symm_apply_inv, Iso.refl_inv, NatTrans.comp_app, whiskerLeft_app, + conjugateEquiv_symm_apply_app, NatTrans.id_app, Functor.map_id, Category.id_comp, + Category.assoc] rw [← adj1.counit_naturality, ← Category.assoc, ← F.map_comp] simp @@ -180,7 +92,7 @@ theorem leftAdjointUniq_refl {F : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) : /-- If `G` and `G'` are both right adjoint to `F`, then they are naturally isomorphic. -/ def rightAdjointUniq {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G') : G ≅ G' := - (natIsoEquiv adj1 adj2).symm (Iso.refl _) + conjugateIsoEquiv adj1 adj2 (Iso.refl _) -- Porting note (#10618): simp can prove this theorem homEquiv_symm_rightAdjointUniq_hom_app {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) @@ -192,8 +104,8 @@ theorem homEquiv_symm_rightAdjointUniq_hom_app {F : C ⥤ D} {G G' : D ⥤ C} (a theorem unit_rightAdjointUniq_hom_app {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G') (x : C) : adj1.unit.app x ≫ (rightAdjointUniq adj1 adj2).hom.app (F.obj x) = adj2.unit.app x := by - simp only [Functor.id_obj, Functor.comp_obj, rightAdjointUniq, natIsoEquiv_symm_apply_hom, - Iso.refl_hom, natTransEquiv_symm_apply_app, NatTrans.id_app, Category.id_comp] + simp only [Functor.id_obj, Functor.comp_obj, rightAdjointUniq, conjugateIsoEquiv_apply_hom, + Iso.refl_hom, conjugateEquiv_apply_app, NatTrans.id_app, Functor.map_id, Category.id_comp] rw [← adj2.unit_naturality_assoc, ← G'.map_comp] simp @@ -243,4 +155,7 @@ theorem rightAdjointUniq_refl {F : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) : end Adjunction +@[deprecated (since := "2024-10-07")] alias Adjunction.natTransEquiv := conjugateEquiv +@[deprecated (since := "2024-10-07")] alias Adjunction.natIsoEquiv := conjugateIsoEquiv + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Bicategory/Functor/Prelax.lean b/Mathlib/CategoryTheory/Bicategory/Functor/Prelax.lean index 5cace28c07436..f7bb524702e5c 100644 --- a/Mathlib/CategoryTheory/Bicategory/Functor/Prelax.lean +++ b/Mathlib/CategoryTheory/Bicategory/Functor/Prelax.lean @@ -76,8 +76,6 @@ def mkOfHomPrefunctors (F : B → C) (F' : (a : B) → (b : B) → Prefunctor (a map {a b} := (F' a b).obj map₂ {a b} := (F' a b).map -variable (F : PrelaxFunctorStruct B C) - -- Porting note: deleted syntactic tautologies `toPrefunctor_eq_coe : F.toPrefunctor = F` -- and `to_prefunctor_obj : (F : Prefunctor B C).obj = F.obj` -- and `to_prefunctor_map` diff --git a/Mathlib/CategoryTheory/Category/Bipointed.lean b/Mathlib/CategoryTheory/Category/Bipointed.lean index d6bd45ad2304f..e85c5549e371d 100644 --- a/Mathlib/CategoryTheory/Category/Bipointed.lean +++ b/Mathlib/CategoryTheory/Category/Bipointed.lean @@ -20,9 +20,6 @@ open CategoryTheory universe u -variable {α β : Type*} - - /-- The category of bipointed types. -/ structure Bipointed : Type (u + 1) where /-- The underlying type of a bipointed type. -/ diff --git a/Mathlib/CategoryTheory/Category/PartialFun.lean b/Mathlib/CategoryTheory/Category/PartialFun.lean index 0ed2b108fb523..b448e2cf2a62f 100644 --- a/Mathlib/CategoryTheory/Category/PartialFun.lean +++ b/Mathlib/CategoryTheory/Category/PartialFun.lean @@ -30,8 +30,6 @@ open CategoryTheory Option universe u -variable {α β : Type*} - /-- The category of types equipped with partial functions. -/ def PartialFun : Type _ := Type* diff --git a/Mathlib/CategoryTheory/Category/Pointed.lean b/Mathlib/CategoryTheory/Category/Pointed.lean index a1f720281f992..8e963c0fc6da9 100644 --- a/Mathlib/CategoryTheory/Category/Pointed.lean +++ b/Mathlib/CategoryTheory/Category/Pointed.lean @@ -22,8 +22,6 @@ open CategoryTheory universe u -variable {α β : Type*} - /-- The category of pointed types. -/ structure Pointed : Type (u + 1) where /-- the underlying type -/ diff --git a/Mathlib/CategoryTheory/ChosenFiniteProducts.lean b/Mathlib/CategoryTheory/ChosenFiniteProducts.lean index 94e7ecaf158cf..7ae5f97b5511b 100644 --- a/Mathlib/CategoryTheory/ChosenFiniteProducts.lean +++ b/Mathlib/CategoryTheory/ChosenFiniteProducts.lean @@ -103,6 +103,14 @@ lemma lift_fst {T X Y : C} (f : T ⟶ X) (g : T ⟶ Y) : lift f g ≫ fst _ _ = lemma lift_snd {T X Y : C} (f : T ⟶ X) (g : T ⟶ Y) : lift f g ≫ snd _ _ = g := by simp [lift, snd] +instance mono_lift_of_mono_left {W X Y : C} (f : W ⟶ X) (g : W ⟶ Y) + [Mono f] : Mono (lift f g) := + mono_of_mono_fac <| lift_fst _ _ + +instance mono_lift_of_mono_right {W X Y : C} (f : W ⟶ X) (g : W ⟶ Y) + [Mono g] : Mono (lift f g) := + mono_of_mono_fac <| lift_snd _ _ + @[ext 1050] lemma hom_ext {T X Y : C} (f g : T ⟶ X ⊗ Y) (h_fst : f ≫ fst _ _ = g ≫ fst _ _) diff --git a/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean b/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean index 0e42d549b245a..38882d4e9f448 100644 --- a/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean +++ b/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean @@ -76,7 +76,7 @@ instance : HasCoeToSort X := ConcreteCategory.hasCoeToSort X -/ def ConcreteCategory.hasCoeToSort (C : Type u) [Category.{v} C] [ConcreteCategory.{w} C] : CoeSort C (Type w) where - coe := fun X => (forget C).obj X + coe X := (forget C).obj X section diff --git a/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean b/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean index d2a6dab1aa44a..1fe01ebf2d14a 100644 --- a/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean +++ b/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean @@ -22,7 +22,7 @@ universe u v namespace CategoryTheory -variable {c d : Type u → Type v} {α : Type u} +variable {c d : Type u → Type v} /-- `Bundled` is a type bundled with a type class instance for that type. Only the type class is exposed as a parameter. -/ diff --git a/Mathlib/CategoryTheory/Functor/Currying.lean b/Mathlib/CategoryTheory/Functor/Currying.lean index edaba6e369347..e130aedb422c8 100644 --- a/Mathlib/CategoryTheory/Functor/Currying.lean +++ b/Mathlib/CategoryTheory/Functor/Currying.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ +import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Products.Basic /-! diff --git a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean index e848bbc5c6d75..fba4b027d12de 100644 --- a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean +++ b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean @@ -373,7 +373,7 @@ lemma endMulEquivAutGalois_pi (f : End F) (A : PointedGaloisObject F) : /-- Any endomorphism of a fiber functor is a unit. -/ theorem FibreFunctor.end_isUnit (f : End F) : IsUnit f := - (MulEquiv.map_isUnit_iff (endMulEquivAutGalois F)).mp + (isUnit_map_iff (endMulEquivAutGalois F) _).mp (Group.isUnit ((endMulEquivAutGalois F) f)) /-- Any endomorphism of a fiber functor is an isomorphism. -/ diff --git a/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean b/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean index 1c901601ac535..b0d783ab327b1 100644 --- a/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean +++ b/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean @@ -259,7 +259,7 @@ variable (F₁₂ : C₁ ⥤ C₂ ⥤ C₁₂) (G : C₁₂ ⥤ C₃ ⥤ C₄) /-- Given a map `r : I₁ × I₂ × I₃ → J`, a `BifunctorComp₁₂IndexData r` consists of the data of a type `I₁₂`, maps `p : I₁ × I₂ → I₁₂` and `q : I₁₂ × I₃ → J`, such that `r` is obtained by composition of `p` and `q`. -/ -structure BifunctorComp₁₂IndexData := +structure BifunctorComp₁₂IndexData where /-- an auxiliary type -/ I₁₂ : Type* /-- a map `I₁ × I₂ → I₁₂` -/ @@ -439,7 +439,7 @@ variable (F : C₁ ⥤ C₂₃ ⥤ C₄) (G₂₃ : C₂ ⥤ C₃ ⥤ C₂₃) /-- Given a map `r : I₁ × I₂ × I₃ → J`, a `BifunctorComp₂₃IndexData r` consists of the data of a type `I₂₃`, maps `p : I₂ × I₃ → I₂₃` and `q : I₁ × I₂₃ → J`, such that `r` is obtained by composition of `p` and `q`. -/ -structure BifunctorComp₂₃IndexData := +structure BifunctorComp₂₃IndexData where /-- an auxiliary type -/ I₂₃ : Type* /-- a map `I₂ × I₃ → I₂₃` -/ diff --git a/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean b/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean index 54aa24e65b239..79f597c240e2e 100644 --- a/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean @@ -132,7 +132,7 @@ theorem of_eq : section UniversalProperty -variable {V' : Type u'} [Groupoid V'] (φ : V ⥤q V') +variable {V' : Type u'} [Groupoid V'] /-- The lift of a prefunctor to a groupoid, to a functor from `FreeGroupoid V` -/ def lift (φ : V ⥤q V') : FreeGroupoid V ⥤ V' := diff --git a/Mathlib/CategoryTheory/Iso.lean b/Mathlib/CategoryTheory/Iso.lean index 8d7ad6c8ea53a..534b77b1210e3 100644 --- a/Mathlib/CategoryTheory/Iso.lean +++ b/Mathlib/CategoryTheory/Iso.lean @@ -332,7 +332,7 @@ 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 g : X ⟶ Y} {h : Y ⟶ Z} +variable {f : X ⟶ Y} {h : Y ⟶ Z} instance inv_isIso [IsIso f] : IsIso (inv f) := (asIso f).isIso_inv @@ -509,7 +509,7 @@ theorem cancel_iso_inv_right_assoc {W X X' Y Z : C} (f : W ⟶ X) (g : X ⟶ Y) section -variable {D E : Type*} [Category D] [Category E] {X Y : C} (e : X ≅ Y) +variable {D : Type*} [Category D] {X Y : C} (e : X ≅ Y) @[reassoc (attr := simp)] lemma map_hom_inv_id (F : C ⥤ D) : diff --git a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean index 305a3bbd5a648..0b2e0d72e2f6e 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean @@ -146,18 +146,23 @@ def combinedIsColimit (F : J ⥤ K ⥤ C) (c : ∀ k : K, ColimitCocone (F.flip. noncomputable section +instance functorCategoryHasLimit (F : J ⥤ K ⥤ C) [∀ k, HasLimit (F.flip.obj k)] : HasLimit F := + HasLimit.mk + { cone := combineCones F fun _ => getLimitCone _ + isLimit := combinedIsLimit _ _ } + instance functorCategoryHasLimitsOfShape [HasLimitsOfShape J C] : HasLimitsOfShape J (K ⥤ C) where - has_limit F := - HasLimit.mk - { cone := combineCones F fun _ => getLimitCone _ - isLimit := combinedIsLimit _ _ } + has_limit _ := inferInstance + +instance functorCategoryHasColimit (F : J ⥤ K ⥤ C) [∀ k, HasColimit (F.flip.obj k)] : + HasColimit F := + HasColimit.mk + { cocone := combineCocones F fun _ => getColimitCocone _ + isColimit := combinedIsColimit _ _ } instance functorCategoryHasColimitsOfShape [HasColimitsOfShape J C] : HasColimitsOfShape J (K ⥤ C) where - has_colimit _ := - HasColimit.mk - { cocone := combineCocones _ fun _ => getColimitCocone _ - isColimit := combinedIsColimit _ _ } + has_colimit _ := inferInstance -- Porting note: previously Lean could see through the binders and infer_instance sufficed instance functorCategoryHasLimitsOfSize [HasLimitsOfSize.{v₁, u₁} C] : @@ -169,14 +174,20 @@ instance functorCategoryHasColimitsOfSize [HasColimitsOfSize.{v₁, u₁} C] : HasColimitsOfSize.{v₁, u₁} (K ⥤ C) where has_colimits_of_shape := fun _ _ => inferInstance +instance hasLimitCompEvalution (F : J ⥤ K ⥤ C) (k : K) [HasLimit (F.flip.obj k)] : + HasLimit (F ⋙ (evaluation _ _).obj k) := + hasLimitOfIso (F := F.flip.obj k) (Iso.refl _) + +instance evaluationPreservesLimit (F : J ⥤ K ⥤ C) [∀ k, HasLimit (F.flip.obj k)] (k : K) : + PreservesLimit F ((evaluation K C).obj k) := + -- Porting note: added a let because X was not inferred + let X : (k : K) → LimitCone (F.flip.obj k) := fun k => getLimitCone (F.flip.obj k) + preservesLimitOfPreservesLimitCone (combinedIsLimit _ X) <| + IsLimit.ofIsoLimit (limit.isLimit _) (evaluateCombinedCones F X k).symm + instance evaluationPreservesLimitsOfShape [HasLimitsOfShape J C] (k : K) : PreservesLimitsOfShape J ((evaluation K C).obj k) where - preservesLimit {F} := by - -- Porting note: added a let because X was not inferred - let X : (k : K) → LimitCone (Prefunctor.obj (Functor.flip F).toPrefunctor k) := - fun k => getLimitCone (Prefunctor.obj (Functor.flip F).toPrefunctor k) - exact preservesLimitOfPreservesLimitCone (combinedIsLimit _ _) <| - IsLimit.ofIsoLimit (limit.isLimit _) (evaluateCombinedCones F X k).symm + preservesLimit := inferInstance /-- If `F : J ⥤ K ⥤ C` is a functor into a functor category which has a limit, then the evaluation of that limit at `k` is the limit of the evaluations of `F.obj j` at `k`. @@ -225,14 +236,20 @@ theorem limit_obj_ext {H : J ⥤ K ⥤ C} [HasLimitsOfShape J C] {k : K} {W : C} ext j simpa using w j +instance hasColimitCompEvaluation (F : J ⥤ K ⥤ C) (k : K) [HasColimit (F.flip.obj k)] : + HasColimit (F ⋙ (evaluation _ _).obj k) := + hasColimitOfIso (F := F.flip.obj k) (Iso.refl _) + +instance evaluationPreservesColimit (F : J ⥤ K ⥤ C) [∀ k, HasColimit (F.flip.obj k)] (k : K) : + PreservesColimit F ((evaluation K C).obj k) := + -- Porting note: added a let because X was not inferred + let X : (k : K) → ColimitCocone (F.flip.obj k) := fun k => getColimitCocone (F.flip.obj k) + preservesColimitOfPreservesColimitCocone (combinedIsColimit _ X) <| + IsColimit.ofIsoColimit (colimit.isColimit _) (evaluateCombinedCocones F X k).symm + instance evaluationPreservesColimitsOfShape [HasColimitsOfShape J C] (k : K) : PreservesColimitsOfShape J ((evaluation K C).obj k) where - preservesColimit {F} := by - -- Porting note: added a let because X was not inferred - let X : (k : K) → ColimitCocone (Prefunctor.obj (Functor.flip F).toPrefunctor k) := - fun k => getColimitCocone (Prefunctor.obj (Functor.flip F).toPrefunctor k) - refine preservesColimitOfPreservesColimitCocone (combinedIsColimit _ _) <| - IsColimit.ofIsoColimit (colimit.isColimit _) (evaluateCombinedCocones F X k).symm + preservesColimit := inferInstance /-- If `F : J ⥤ K ⥤ C` is a functor into a functor category which has a colimit, then the evaluation of that colimit at `k` is the colimit of the evaluations of `F.obj j` at `k`. diff --git a/Mathlib/CategoryTheory/Limits/FunctorToTypes.lean b/Mathlib/CategoryTheory/Limits/FunctorToTypes.lean index 62eb96c6be392..ba9abc78c25d6 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorToTypes.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorToTypes.lean @@ -20,15 +20,20 @@ open CategoryTheory.Limits universe w v₁ v₂ u₁ u₂ variable {J : Type u₁} [Category.{v₁} J] {K : Type u₂} [Category.{v₂} K] -variable (F : J ⥤ K ⥤ TypeMax.{u₁, w}) +variable (F : J ⥤ K ⥤ Type w) -theorem jointly_surjective (k : K) {t : Cocone F} (h : IsColimit t) (x : t.pt.obj k) : - ∃ j y, x = (t.ι.app j).app k y := by +theorem jointly_surjective (k : K) {t : Cocone F} (h : IsColimit t) (x : t.pt.obj k) + [∀ k, HasColimit (F.flip.obj k)] : ∃ j y, x = (t.ι.app j).app k y := by let hev := isColimitOfPreserves ((evaluation _ _).obj k) h obtain ⟨j, y, rfl⟩ := Types.jointly_surjective _ hev x exact ⟨j, y, by simp⟩ -theorem jointly_surjective' (k : K) (x : (colimit F).obj k) : ∃ j y, x = (colimit.ι F j).app k y := +theorem jointly_surjective' [∀ k, HasColimit (F.flip.obj k)] (k : K) (x : (colimit F).obj k) : + ∃ j y, x = (colimit.ι F j).app k y := jointly_surjective _ _ (colimit.isColimit _) x +theorem colimit.map_ι_apply [HasColimit F] (j : J) {k k' : K} {f : k ⟶ k'} {x} : + (colimit F).map f ((colimit.ι F j).app _ x) = (colimit.ι F j).app _ ((F.obj j).map f x) := + congrFun ((colimit.ι F j).naturality _).symm _ + end CategoryTheory.FunctorToTypes diff --git a/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean index 4ea865a6b1d0d..351236724ffa1 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean @@ -31,7 +31,7 @@ braiding and associating isomorphisms, and the product comparison morphism. noncomputable section -universe v u u₂ +universe v v₁ u u₁ u₂ open CategoryTheory @@ -160,7 +160,7 @@ def diagramIsoPair (F : Discrete WalkingPair ⥤ C) : section -variable {D : Type u} [Category.{v} D] +variable {D : Type u₁} [Category.{v₁} D] /-- The natural isomorphism between `pair X Y ⋙ F` and `pair (F.obj X) (F.obj Y)`. -/ def pairComp (X Y : C) (F : C ⥤ D) : pair X Y ⋙ F ≅ pair (F.obj X) (F.obj Y) := diff --git a/Mathlib/CategoryTheory/Linear/Yoneda.lean b/Mathlib/CategoryTheory/Linear/Yoneda.lean index 4836d3c76a695..ee468da7d7708 100644 --- a/Mathlib/CategoryTheory/Linear/Yoneda.lean +++ b/Mathlib/CategoryTheory/Linear/Yoneda.lean @@ -28,7 +28,7 @@ namespace CategoryTheory variable (R : Type w) [Ring R] {C : Type u} [Category.{v} C] [Preadditive C] [Linear R C] variable (C) --- Porting note: inserted specific `ModuleCat.ofHom` in the definition of `linearYoneda` +-- Porting note: inserted specific `ModuleCat.asHom` in the definition of `linearYoneda` -- and similarly in `linearCoyoneda`, otherwise many simp lemmas are not triggered automatically. -- Eventually, doing so allows more proofs to be automatic! /-- The Yoneda embedding for `R`-linear categories `C`, @@ -38,9 +38,9 @@ with value on `Y : Cᵒᵖ` given by `ModuleCat.of R (unop Y ⟶ X)`. -/ def linearYoneda : C ⥤ Cᵒᵖ ⥤ ModuleCat R where obj X := { obj := fun Y => ModuleCat.of R (unop Y ⟶ X) - map := fun f => ModuleCat.ofHom (Linear.leftComp R _ f.unop) } + map := fun f => ModuleCat.asHom (Linear.leftComp R _ f.unop) } map {X₁ X₂} f := - { app := fun Y => @ModuleCat.ofHom R _ (Y.unop ⟶ X₁) (Y.unop ⟶ X₂) _ _ _ _ + { app := fun Y => @ModuleCat.asHom R _ (Y.unop ⟶ X₁) (Y.unop ⟶ X₂) _ _ _ _ (Linear.rightComp R _ f) } /-- The Yoneda embedding for `R`-linear categories `C`, @@ -50,9 +50,9 @@ with value on `X : C` given by `ModuleCat.of R (unop Y ⟶ X)`. -/ def linearCoyoneda : Cᵒᵖ ⥤ C ⥤ ModuleCat R where obj Y := { obj := fun X => ModuleCat.of R (unop Y ⟶ X) - map := fun f => ModuleCat.ofHom (Linear.rightComp R _ f) } + map := fun f => ModuleCat.asHom (Linear.rightComp R _ f) } map {Y₁ Y₂} f := - { app := fun X => @ModuleCat.ofHom R _ (unop Y₁ ⟶ X) (unop Y₂ ⟶ X) _ _ _ _ + { app := fun X => @ModuleCat.asHom R _ (unop Y₁ ⟶ X) (unop Y₂ ⟶ X) _ _ _ _ (Linear.leftComp _ _ f.unop) } instance linearYoneda_obj_additive (X : C) : ((linearYoneda R C).obj X).Additive where diff --git a/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean b/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean index c0f59f13389ca..a0c71108f57aa 100644 --- a/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean +++ b/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean @@ -128,7 +128,7 @@ lemma isEquivalence_iff : G.IsEquivalence ↔ G'.IsEquivalence := end /-- Condition that a `LocalizerMorphism` induces an equivalence on the localized categories -/ -class IsLocalizedEquivalence : Prop := +class IsLocalizedEquivalence : Prop where /-- the induced functor on the constructed localized categories is an equivalence -/ isEquivalence : (Φ.localizedFunctor W₁.Q W₂.Q).IsEquivalence diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index c38de64d692f4..d80764aa71802 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Michael Jendrusch. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Jendrusch, Kim Morrison, Bhavik Mehta, Jakob von Raumer -/ +import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Functor.Trifunctor import Mathlib.CategoryTheory.Products.Basic diff --git a/Mathlib/CategoryTheory/MorphismProperty/Composition.lean b/Mathlib/CategoryTheory/MorphismProperty/Composition.lean index 3595312f8c115..aff2f87e7acc0 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Composition.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Composition.lean @@ -24,7 +24,7 @@ namespace MorphismProperty variable {C : Type u} [Category.{v} C] {D : Type u'} [Category.{v'} D] /-- Typeclass expressing that a morphism property contain identities. -/ -class ContainsIdentities (W : MorphismProperty C) : Prop := +class ContainsIdentities (W : MorphismProperty C) : Prop where /-- for all `X : C`, the identity of `X` satisfies the morphism property -/ id_mem : ∀ (X : C), W (𝟙 X) @@ -63,7 +63,7 @@ instance Pi.containsIdentities {J : Type w} {C : J → Type u} /-- A morphism property satisfies `IsStableUnderComposition` if the composition of two such morphisms still falls in the class. -/ -class IsStableUnderComposition (P : MorphismProperty C) : Prop := +class IsStableUnderComposition (P : MorphismProperty C) : Prop where comp_mem {X Y Z} (f : X ⟶ Y) (g : Y ⟶ Z) : P f → P g → P (f ≫ g) lemma comp_mem (W : MorphismProperty C) [W.IsStableUnderComposition] @@ -129,7 +129,7 @@ end naturalityProperty /-- A morphism property is multiplicative if it contains identities and is stable by composition. -/ class IsMultiplicative (W : MorphismProperty C) - extends W.ContainsIdentities, W.IsStableUnderComposition : Prop := + extends W.ContainsIdentities, W.IsStableUnderComposition : Prop namespace IsMultiplicative diff --git a/Mathlib/CategoryTheory/MorphismProperty/Limits.lean b/Mathlib/CategoryTheory/MorphismProperty/Limits.lean index e3c20a156e317..7352e3fdd289a 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Limits.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Limits.lean @@ -190,7 +190,7 @@ abbrev IsStableUnderProductsOfShape (J : Type*) := W.IsStableUnderLimitsOfShape lemma IsStableUnderProductsOfShape.mk (J : Type*) [W.RespectsIso] [HasProductsOfShape J C] (hW : ∀ (X₁ X₂ : J → C) (f : ∀ j, X₁ j ⟶ X₂ j) (_ : ∀ (j : J), W (f j)), - W (Pi.map f)) : W.IsStableUnderProductsOfShape J := by + W (Limits.Pi.map f)) : W.IsStableUnderProductsOfShape J := by intro X₁ X₂ c₁ c₂ hc₁ hc₂ f hf let φ := fun j => f.app (Discrete.mk j) have hf' := hW _ _ φ (fun j => hf (Discrete.mk j)) @@ -203,7 +203,7 @@ lemma IsStableUnderProductsOfShape.mk (J : Type*) simp /-- The condition that a property of morphisms is stable by finite products. -/ -class IsStableUnderFiniteProducts : Prop := +class IsStableUnderFiniteProducts : Prop where isStableUnderProductsOfShape (J : Type) [Finite J] : W.IsStableUnderProductsOfShape J lemma isStableUnderProductsOfShape_of_isStableUnderFiniteProducts diff --git a/Mathlib/CategoryTheory/NatTrans.lean b/Mathlib/CategoryTheory/NatTrans.lean index 0c12f372eaeac..af2c4079f9b36 100644 --- a/Mathlib/CategoryTheory/NatTrans.lean +++ b/Mathlib/CategoryTheory/NatTrans.lean @@ -72,7 +72,7 @@ open CategoryTheory.Functor section -variable {F G H I : C ⥤ D} +variable {F G H : C ⥤ D} /-- `vcomp α β` is the vertical compositions of natural transformations. -/ def vcomp (α : NatTrans F G) (β : NatTrans G H) : NatTrans F H where diff --git a/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean b/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean index 6770c026e762f..f7dd3eabd7f76 100644 --- a/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean +++ b/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean @@ -78,7 +78,6 @@ instance Monad.algebraPreadditive : Preadditive (Monad.Algebra T) where zsmul_succ' := by intros ext - dsimp simp only [natCast_zsmul, succ_nsmul] rfl zsmul_neg' := by @@ -159,7 +158,6 @@ instance Comonad.coalgebraPreadditive : Preadditive (Comonad.Coalgebra U) where zsmul_succ' := by intros ext - dsimp simp only [natCast_zsmul, succ_nsmul] rfl zsmul_neg' := by diff --git a/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean b/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean index 006286e5be91c..f7fd470a65628 100644 --- a/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean +++ b/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean @@ -78,7 +78,6 @@ instance Endofunctor.algebraPreadditive : Preadditive (Endofunctor.Algebra F) wh zsmul_succ' := by intros apply Algebra.Hom.ext - dsimp simp only [natCast_zsmul, succ_nsmul] rfl zsmul_neg' := by @@ -156,7 +155,6 @@ instance Endofunctor.coalgebraPreadditive : Preadditive (Endofunctor.Coalgebra F zsmul_succ' := by intros apply Coalgebra.Hom.ext - dsimp simp only [natCast_zsmul, succ_nsmul] rfl zsmul_neg' := by diff --git a/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean b/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean index 96b1f0270625c..3e19f976a5b18 100644 --- a/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean +++ b/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean @@ -39,7 +39,7 @@ object `X` to the `End Y`-module of morphisms `X ⟶ Y`. @[simps] def preadditiveYonedaObj (Y : C) : Cᵒᵖ ⥤ ModuleCat.{v} (End Y) where obj X := ModuleCat.of _ (X.unop ⟶ Y) - map f := ModuleCat.ofHom + map f := ModuleCat.asHom { toFun := fun g => f.unop ≫ g map_add' := fun g g' => comp_add _ _ _ _ _ _ map_smul' := fun r g => Eq.symm <| Category.assoc _ _ _ } @@ -66,7 +66,7 @@ object `Y` to the `End X`-module of morphisms `X ⟶ Y`. @[simps] def preadditiveCoyonedaObj (X : Cᵒᵖ) : C ⥤ ModuleCat.{v} (End X) where obj Y := ModuleCat.of _ (unop X ⟶ Y) - map f := ModuleCat.ofHom + map f := ModuleCat.asHom { toFun := fun g => g ≫ f map_add' := fun g g' => add_comp _ _ _ _ _ _ map_smul' := fun r g => Category.assoc _ _ _ } diff --git a/Mathlib/CategoryTheory/Products/Basic.lean b/Mathlib/CategoryTheory/Products/Basic.lean index c38b430629b3b..58958ff88a2de 100644 --- a/Mathlib/CategoryTheory/Products/Basic.lean +++ b/Mathlib/CategoryTheory/Products/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2017 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Stephen Morgan, Kim Morrison -/ -import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Functor.Const import Mathlib.CategoryTheory.Opposites import Mathlib.Data.Prod.Basic @@ -296,7 +295,18 @@ end Equivalence /-- `F.flip` composed with evaluation is the same as evaluating `F`. -/ @[simps!] def flipCompEvaluation (F : A ⥤ B ⥤ C) (a) : F.flip ⋙ (evaluation _ _).obj a ≅ F.obj a := - NatIso.ofComponents fun b => eqToIso rfl + NatIso.ofComponents fun b => Iso.refl _ + +theorem flip_comp_evaluation (F : A ⥤ B ⥤ C) (a) : F.flip ⋙ (evaluation _ _).obj a = F.obj a := + rfl + +/-- `F` composed with evaluation is the same as evaluating `F.flip`. -/ +@[simps!] +def compEvaluation (F : A ⥤ B ⥤ C) (b) : F ⋙ (evaluation _ _).obj b ≅ F.flip.obj b := + NatIso.ofComponents fun a => Iso.refl _ + +theorem comp_evaluation (F : A ⥤ B ⥤ C) (b) : F ⋙ (evaluation _ _).obj b = F.flip.obj b := + rfl variable (A B C) diff --git a/Mathlib/CategoryTheory/Shift/CommShift.lean b/Mathlib/CategoryTheory/Shift/CommShift.lean index 5f75457aa59db..e75c095dca271 100644 --- a/Mathlib/CategoryTheory/Shift/CommShift.lean +++ b/Mathlib/CategoryTheory/Shift/CommShift.lean @@ -246,7 +246,7 @@ variable {C D E J : Type*} [Category C] [Category D] [Category E] [Category J] /-- If `τ : F₁ ⟶ F₂` is a natural transformation between two functors which commute with a shift by an additive monoid `A`, this typeclass asserts a compatibility of `τ` with these shifts. -/ -class CommShift : Prop := +class CommShift : Prop where comm' (a : A) : (F₁.commShiftIso a).hom ≫ whiskerRight τ _ = whiskerLeft _ τ ≫ (F₂.commShiftIso a).hom diff --git a/Mathlib/CategoryTheory/Shift/Localization.lean b/Mathlib/CategoryTheory/Shift/Localization.lean index 5e9f9b7a925b5..1f74884924d44 100644 --- a/Mathlib/CategoryTheory/Shift/Localization.lean +++ b/Mathlib/CategoryTheory/Shift/Localization.lean @@ -32,7 +32,7 @@ namespace MorphismProperty /-- A morphism property `W` on a category `C` is compatible with the shift by a monoid `A` when for all `a : A`, a morphism `f` belongs to `W` if and only if `f⟦a⟧'` does. -/ -class IsCompatibleWithShift : Prop := +class IsCompatibleWithShift : Prop where /-- the condition that for all `a : A`, the morphism property `W` is not changed when we take its inverse image by the shift functor by `a` -/ condition : ∀ (a : A), W.inverseImage (shiftFunctor C a) = W diff --git a/Mathlib/CategoryTheory/Shift/Quotient.lean b/Mathlib/CategoryTheory/Shift/Quotient.lean index d029cd486c714..f789f82c2883c 100644 --- a/Mathlib/CategoryTheory/Shift/Quotient.lean +++ b/Mathlib/CategoryTheory/Shift/Quotient.lean @@ -32,7 +32,7 @@ namespace HomRel /-- A relation on morphisms is compatible with the shift by a monoid `A` when the relation if preserved by the shift. -/ -class IsCompatibleWithShift : Prop := +class IsCompatibleWithShift : Prop where /-- the condition that the relation is preserved by the shift -/ condition : ∀ (a : A) ⦃X Y : C⦄ (f g : X ⟶ Y), r f g → r (f⟦a⟧') (g⟦a⟧') diff --git a/Mathlib/CategoryTheory/Skeletal.lean b/Mathlib/CategoryTheory/Skeletal.lean index 2ff548def5bb3..af988c24022a5 100644 --- a/Mathlib/CategoryTheory/Skeletal.lean +++ b/Mathlib/CategoryTheory/Skeletal.lean @@ -67,7 +67,7 @@ variable (C D) /-- Construct the skeleton category as the induced category on the isomorphism classes, and derive its category structure. -/ -def Skeleton : Type u₁ := InducedCategory C Quotient.out +def Skeleton : Type u₁ := InducedCategory (C := Quotient (isIsomorphicSetoid C)) C Quotient.out instance [Inhabited C] : Inhabited (Skeleton C) := ⟨⟦default⟧⟩ diff --git a/Mathlib/CategoryTheory/SmallObject/Iteration.lean b/Mathlib/CategoryTheory/SmallObject/Iteration.lean index 4e9947245dd1b..a9470418b418d 100644 --- a/Mathlib/CategoryTheory/SmallObject/Iteration.lean +++ b/Mathlib/CategoryTheory/SmallObject/Iteration.lean @@ -123,7 +123,7 @@ lemma mapSucc_eq (i : J) (hi : i < j) : end -variable (iter₁ iter₂ iter₃ : Φ.Iteration ε j) +variable (iter₁ iter₂ : Φ.Iteration ε j) /-- A morphism between two objects `iter₁` and `iter₂` in the category `Φ.Iteration ε j` of `j`th iterations of a functor `Φ` @@ -148,7 +148,7 @@ attribute [simp, reassoc] natTrans_app_zero def id : Hom iter₁ iter₁ where natTrans := 𝟙 _ -variable {iter₁ iter₂ iter₃} +variable {iter₁ iter₂} -- Note: this is not made a global ext lemma because it is shown below -- that the type of morphisms is a subsingleton. diff --git a/Mathlib/CategoryTheory/Triangulated/Basic.lean b/Mathlib/CategoryTheory/Triangulated/Basic.lean index 308f5456a6a29..e76c664d5f1d3 100644 --- a/Mathlib/CategoryTheory/Triangulated/Basic.lean +++ b/Mathlib/CategoryTheory/Triangulated/Basic.lean @@ -260,9 +260,9 @@ variable {J : Type*} (T : J → Triangle C) /-- The product of a family of triangles. -/ @[simps!] def productTriangle : Triangle C := - Triangle.mk (Pi.map (fun j => (T j).mor₁)) - (Pi.map (fun j => (T j).mor₂)) - (Pi.map (fun j => (T j).mor₃) ≫ inv (piComparison _ _)) + Triangle.mk (Limits.Pi.map (fun j => (T j).mor₁)) + (Limits.Pi.map (fun j => (T j).mor₂)) + (Limits.Pi.map (fun j => (T j).mor₃) ≫ inv (piComparison _ _)) /-- A projection from the product of a family of triangles. -/ @[simps] diff --git a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean index 091de90c8c3c1..c3b954b379632 100644 --- a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean +++ b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean @@ -563,7 +563,7 @@ lemma productTriangle_distinguished {J : Type*} (T : J → Triangle C) `φ'.hom₁` and `φ'.hom₂` are identities. Then, it suffices to show that `φ'.hom₃` is an isomorphism, which is achieved by using Yoneda's lemma and diagram chases. -/ - let f₁ := Pi.map (fun j => (T j).mor₁) + let f₁ := Limits.Pi.map (fun j => (T j).mor₁) obtain ⟨Z, f₂, f₃, hT'⟩ := distinguished_cocone_triangle f₁ let T' := Triangle.mk f₁ f₂ f₃ change T' ∈ distTriang C at hT' diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 5fbeccb9b2ec8..f06afc7170c1c 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -417,6 +417,7 @@ theorem yonedaEquiv_symm_app_apply {X : C} {F : Cᵒᵖ ⥤ Type v₁} (x : F.ob (f : Y.unop ⟶ X) : (yonedaEquiv.symm x).app Y f = F.map f.op x := rfl +/-- See also `yonedaEquiv_naturality'` for a more general version. -/ lemma yonedaEquiv_naturality {X Y : C} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.obj X ⟶ F) (g : Y ⟶ X) : F.map g.op (yonedaEquiv f) = yonedaEquiv (yoneda.map g ≫ f) := by change (f.app (op X) ≫ F.map g.op) (𝟙 X) = f.app (op Y) (𝟙 Y ≫ g) @@ -424,6 +425,9 @@ lemma yonedaEquiv_naturality {X Y : C} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.o dsimp simp +/-- Variant of `yonedaEquiv_naturality` with general `g`. This is technically strictly more general + than `yonedaEquiv_naturality`, but `yonedaEquiv_naturality` is sometimes preferable because it + can avoid the "motive is not type correct" error. -/ lemma yonedaEquiv_naturality' {X Y : Cᵒᵖ} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.obj (unop X) ⟶ F) (g : X ⟶ Y) : F.map g (yonedaEquiv f) = yonedaEquiv (yoneda.map g.unop ≫ f) := yonedaEquiv_naturality _ _ @@ -436,6 +440,18 @@ lemma yonedaEquiv_yoneda_map {X Y : C} (f : X ⟶ Y) : yonedaEquiv (yoneda.map f rw [yonedaEquiv_apply] simp +/-- See also `map_yonedaEquiv'` for a more general version. -/ +lemma map_yonedaEquiv {X Y : C} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.obj X ⟶ F) + (g : Y ⟶ X) : F.map g.op (yonedaEquiv f) = f.app (op Y) g := by + rw [yonedaEquiv_naturality, yonedaEquiv_comp, yonedaEquiv_yoneda_map] + +/-- Variant of `map_yonedaEquiv` with general `g`. This is technically strictly more general + than `map_yonedaEquiv`, but `map_yonedaEquiv` is sometimes preferable because it + can avoid the "motive is not type correct" error. -/ +lemma map_yonedaEquiv' {X Y : Cᵒᵖ} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.obj (unop X) ⟶ F) + (g : X ⟶ Y) : F.map g (yonedaEquiv f) = f.app Y g.unop := by + rw [yonedaEquiv_naturality', yonedaEquiv_comp, yonedaEquiv_yoneda_map] + lemma yonedaEquiv_symm_map {X Y : Cᵒᵖ} (f : X ⟶ Y) {F : Cᵒᵖ ⥤ Type v₁} (t : F.obj X) : yonedaEquiv.symm (F.map f t) = yoneda.map f.unop ≫ yonedaEquiv.symm t := by obtain ⟨u, rfl⟩ := yonedaEquiv.surjective t @@ -612,6 +628,10 @@ lemma coyonedaEquiv_coyoneda_map {X Y : C} (f : X ⟶ Y) : rw [coyonedaEquiv_apply] simp +lemma map_coyonedaEquiv {X Y : C} {F : C ⥤ Type v₁} (f : coyoneda.obj (op X) ⟶ F) + (g : X ⟶ Y) : F.map g (coyonedaEquiv f) = f.app Y g := by + rw [coyonedaEquiv_naturality, coyonedaEquiv_comp, coyonedaEquiv_coyoneda_map] + lemma coyonedaEquiv_symm_map {X Y : C} (f : X ⟶ Y) {F : C ⥤ Type v₁} (t : F.obj X) : coyonedaEquiv.symm (F.map f t) = coyoneda.map f.op ≫ coyonedaEquiv.symm t := by obtain ⟨u, rfl⟩ := coyonedaEquiv.surjective t diff --git a/Mathlib/Combinatorics/Colex.lean b/Mathlib/Combinatorics/Colex.lean index 0d59516a64ae1..54e26b8a2cc4c 100644 --- a/Mathlib/Combinatorics/Colex.lean +++ b/Mathlib/Combinatorics/Colex.lean @@ -65,7 +65,7 @@ namespace Finset /-- Type synonym of `Finset α` equipped with the colexicographic order rather than the inclusion order. -/ @[ext] -structure Colex (α) := +structure Colex (α) where /-- `toColex` is the "identity" function between `Finset α` and `Finset.Colex α`. -/ toColex :: /-- `ofColex` is the "identity" function between `Finset.Colex α` and `Finset α`. -/ diff --git a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean index 542deff1911e1..0264391a00385 100644 --- a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean +++ b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean @@ -39,7 +39,7 @@ namespace Finset section Bipartite -variable (r : α → β → Prop) (s : Finset α) (t : Finset β) (a a' : α) (b b' : β) +variable (r : α → β → Prop) (s : Finset α) (t : Finset β) (a : α) (b : β) [DecidablePred (r a)] [∀ a, Decidable (r a b)] {m n : ℕ} /-- Elements of `s` which are "below" `b` according to relation `r`. -/ @@ -58,7 +58,7 @@ theorem coe_bipartiteBelow : s.bipartiteBelow r b = ({a ∈ s | r a b} : Set α) @[simp, norm_cast] theorem coe_bipartiteAbove : t.bipartiteAbove r a = ({b ∈ t | r a b} : Set β) := coe_filter _ _ -variable {s t a a' b b'} +variable {s t a b} @[simp] theorem mem_bipartiteBelow {a : α} : a ∈ s.bipartiteBelow r b ↔ a ∈ s ∧ r a b := mem_filter @@ -72,7 +72,7 @@ theorem sum_card_bipartiteAbove_eq_sum_card_bipartiteBelow [∀ a b, Decidable ( exact sum_comm section OrderedSemiring -variable [OrderedSemiring R] [DecidablePred (r a)] [∀ a, Decidable (r a b)] {m n : R} +variable [OrderedSemiring R] {m n : R} /-- **Double counting** argument. @@ -100,7 +100,7 @@ end OrderedSemiring section StrictOrderedSemiring variable [StrictOrderedSemiring R] (r : α → β → Prop) {s : Finset α} {t : Finset β} - (a a' : α) (b b' : β) [DecidablePred (r a)] [∀ a, Decidable (r a b)] {m n : R} + (a b) {m n : R} /-- **Double counting** argument. diff --git a/Mathlib/Combinatorics/SetFamily/Compression/Down.lean b/Mathlib/Combinatorics/SetFamily/Compression/Down.lean index 222c68d63fd87..a93982a2764cf 100644 --- a/Mathlib/Combinatorics/SetFamily/Compression/Down.lean +++ b/Mathlib/Combinatorics/SetFamily/Compression/Down.lean @@ -35,7 +35,7 @@ compression, down-compression -/ -variable {α : Type*} [DecidableEq α] {𝒜 ℬ : Finset (Finset α)} {s : Finset α} {a : α} +variable {α : Type*} [DecidableEq α] {𝒜 : Finset (Finset α)} {s : Finset α} {a : α} namespace Finset diff --git a/Mathlib/Combinatorics/SetFamily/Kleitman.lean b/Mathlib/Combinatorics/SetFamily/Kleitman.lean index 7fe6c323c2871..e534f30f51b45 100644 --- a/Mathlib/Combinatorics/SetFamily/Kleitman.lean +++ b/Mathlib/Combinatorics/SetFamily/Kleitman.lean @@ -79,5 +79,5 @@ theorem Finset.card_biUnion_le_of_intersecting (s : Finset ι) (f : ι → Finse (ih _ (fun i hi ↦ (hf₁ _ <| subset_cons _ hi).2.2) ((card_le_card <| subset_cons _).trans hs)) _).trans ?_ rw [mul_tsub, two_mul, ← pow_succ', - ← add_tsub_assoc_of_le (pow_le_pow_right' (one_le_two : (1 : ℕ) ≤ 2) tsub_le_self), + ← add_tsub_assoc_of_le (pow_right_mono₀ (one_le_two : (1 : ℕ) ≤ 2) tsub_le_self), tsub_add_eq_add_tsub hs, card_cons, add_tsub_add_eq_tsub_right] diff --git a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean index 29713f6399feb..8b5ce0135fec7 100644 --- a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean +++ b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean @@ -52,10 +52,6 @@ namespace Finset namespace Colex variable {α : Type*} [LinearOrder α] {𝒜 𝒜₁ 𝒜₂ : Finset (Finset α)} {s t : Finset α} {r : ℕ} -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- This is important for iterating Kruskal-Katona: the shadow of an initial segment is also an initial segment. -/ lemma shadow_initSeg [Fintype α] (hs : s.Nonempty) : @@ -71,7 +67,7 @@ lemma shadow_initSeg [Fintype α] (hs : s.Nonempty) : · simpa [ha] using erase_le_erase_min' hts hst.ge (mem_insert_self _ _) -- Now show that if t ≤ s - min s, there is j such that t ∪ j ≤ s -- We choose j as the smallest thing not in t - simp_rw [le_iff_eq_or_lt, lt_iff_exists_filter_lt, mem_sdiff, filter_inj, _root_.and_assoc] + simp_rw [le_iff_eq_or_lt, lt_iff_exists_filter_lt, mem_sdiff, filter_inj, and_assoc] simp only [toColex_inj, ofColex_toColex, ne_eq, and_imp] rintro cards' (rfl | ⟨k, hks, hkt, z⟩) -- If t = s - min s, then use j = min s so t ∪ j = s @@ -129,17 +125,13 @@ variable {α : Type*} [LinearOrder α] {s U V : Finset α} {n : ℕ} namespace UV -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- Applying the compression makes the set smaller in colex. This is intuitive since a portion of the set is being "shifted down" as `max U < max V`. -/ lemma toColex_compress_lt_toColex {hU : U.Nonempty} {hV : V.Nonempty} (h : max' U hU < max' V hV) (hA : compress U V s ≠ s) : toColex (compress U V s) < toColex s := by rw [compress, ite_ne_right_iff] at hA rw [compress, if_pos hA.1, lt_iff_exists_filter_lt] - simp_rw [mem_sdiff (s := s), filter_inj, _root_.and_assoc] + simp_rw [mem_sdiff (s := s), filter_inj, and_assoc] refine ⟨_, hA.1.2 <| max'_mem _ hV, not_mem_sdiff_of_mem_right <| max'_mem _ _, fun a ha ↦ ?_⟩ have : a ∉ V := fun H ↦ ha.not_le (le_max' _ _ H) have : a ∉ U := fun H ↦ ha.not_lt ((le_max' _ _ H).trans_lt h) diff --git a/Mathlib/Combinatorics/SimpleGraph/Density.lean b/Mathlib/Combinatorics/SimpleGraph/Density.lean index 0c78f1e2c23b3..f71bb81713dbb 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Density.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Density.lean @@ -175,7 +175,7 @@ theorem edgeDensity_sub_edgeDensity_le_one_sub_mul (hs : s₂ ⊆ s₁) (ht : t refine (sub_le_sub_left (mul_edgeDensity_le_edgeDensity r hs ht hs₂ ht₂) _).trans ?_ refine le_trans ?_ (mul_le_of_le_one_right ?_ (edgeDensity_le_one r s₂ t₂)) · rw [sub_mul, one_mul] - refine sub_nonneg_of_le (mul_le_one ?_ ?_ ?_) + refine sub_nonneg_of_le (mul_le_one₀ ?_ ?_ ?_) · exact div_le_one_of_le₀ ((@Nat.cast_le ℚ).2 (card_le_card hs)) (Nat.cast_nonneg _) · apply div_nonneg <;> exact mod_cast Nat.zero_le _ · exact div_le_one_of_le₀ ((@Nat.cast_le ℚ).2 (card_le_card ht)) (Nat.cast_nonneg _) diff --git a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean index 30662b99dcef9..c5f35dfab6814 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean @@ -66,7 +66,7 @@ lemma IsHamiltonian.length_eq (hp : p.IsHamiltonian) : p.length = Fintype.card end /-- A hamiltonian cycle is a cycle that visits every vertex once. -/ -structure IsHamiltonianCycle (p : G.Walk a a) extends p.IsCycle : Prop := +structure IsHamiltonianCycle (p : G.Walk a a) extends p.IsCycle : Prop where isHamiltonian_tail : p.tail.IsHamiltonian variable {p : G.Walk a a} diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean index d903c4c522373..14c2df9eb7249 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean @@ -114,7 +114,7 @@ private theorem card_nonuniformWitness_sdiff_biUnion_star (hV : V ∈ P.parts) ( rw [sum_const] refine mul_le_mul_right' ?_ _ have t := card_filter_atomise_le_two_pow (s := U) hX - refine t.trans (pow_le_pow_right (by norm_num) <| tsub_le_tsub_right ?_ _) + refine t.trans (pow_right_mono₀ (by norm_num) <| tsub_le_tsub_right ?_ _) exact card_image_le.trans (card_le_card <| filter_subset _ _) private theorem one_sub_eps_mul_card_nonuniformWitness_le_card_star (hV : V ∈ P.parts) @@ -379,7 +379,7 @@ private theorem eps_le_card_star_div [Nonempty α] (hPα : P.parts.card * 16 ^ P (hPε : ↑100 ≤ ↑4 ^ P.parts.card * ε ^ 5) (hε₁ : ε ≤ 1) (hU : U ∈ P.parts) (hV : V ∈ P.parts) (hUV : U ≠ V) (hunif : ¬G.IsUniform ε U V) : ↑4 / ↑5 * ε ≤ (star hP G ε hU V).card / ↑4 ^ P.parts.card := by - have hm : (0 : ℝ) ≤ 1 - (↑m)⁻¹ := sub_nonneg_of_le (inv_le_one <| one_le_m_coe hPα) + have hm : (0 : ℝ) ≤ 1 - (↑m)⁻¹ := sub_nonneg_of_le (inv_le_one_of_one_le₀ <| one_le_m_coe hPα) have hε : 0 ≤ 1 - ε / 10 := sub_nonneg_of_le (div_le_one_of_le₀ (hε₁.trans <| by norm_num) <| by norm_num) have hε₀ : 0 < ε := by sz_positivity diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean index e74fad0ddff83..96d140587f9cc 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean @@ -32,10 +32,6 @@ namespace Finpartition variable {α : Type*} [DecidableEq α] {s t : Finset α} {m n a b : ℕ} {P : Finpartition s} -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- Given a partition `P` of `s`, as well as a proof that `a * m + b * (m + 1) = s.card`, we can find a new partition `Q` of `s` where each part has size `m` or `m + 1`, every part of `P` is the union of parts of `Q` plus at most `m` extra elements, there are `b` parts of size `m + 1` and @@ -50,7 +46,7 @@ theorem equitabilise_aux (hs : a * m + b * (m + 1) = s.card) : obtain rfl | m_pos := m.eq_zero_or_pos · refine ⟨⊥, by simp, ?_, by simpa [Finset.filter_true_of_mem] using hs.symm⟩ simp only [le_zero_iff, card_eq_zero, mem_biUnion, exists_prop, mem_filter, id, - _root_.and_assoc, sdiff_eq_empty_iff_subset, subset_iff] + and_assoc, sdiff_eq_empty_iff_subset, subset_iff] exact fun x hx a ha => ⟨{a}, mem_map_of_mem _ (P.le hx ha), singleton_subset_iff.2 ha, mem_singleton_self _⟩ -- Prove the case `m > 0` by strong induction on `s` diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean index 5e120205aeeb8..bee7ea6b54e48 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean @@ -128,7 +128,7 @@ theorem szemeredi_regularity (hε : 0 < ε) (hl : l ≤ card α) : -- We gather a few numerical facts. have hεl' : 100 ≤ 4 ^ P.parts.card * ε ^ 5 := (hundred_lt_pow_initialBound_mul hε l).le.trans - (mul_le_mul_of_nonneg_right (pow_le_pow_right (by norm_num) hP₂) <| by positivity) + (mul_le_mul_of_nonneg_right (pow_right_mono₀ (by norm_num) hP₂) <| by positivity) have hi : (i : ℝ) ≤ 4 / ε ^ 5 := by have hi : ε ^ 5 / 4 * ↑i ≤ 1 := hP₄.trans (mod_cast P.energy_le_one G) rw [div_mul_eq_mul_div, div_le_iff₀ (show (0 : ℝ) < 4 by norm_num)] at hi diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean index a96ef7c723fb8..96cd101472ef9 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean @@ -83,10 +83,6 @@ lemma LocallyLinear.map (f : α ↪ β) (hG : G.LocallyLinear) : (G.map f).Local · rw [← Equiv.coe_toEmbedding, ← map_symm] exact LocallyLinear.map _ -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ lemma edgeDisjointTriangles_iff_mem_sym2_subsingleton : G.EdgeDisjointTriangles ↔ ∀ ⦃e : Sym2 α⦄, ¬ e.IsDiag → {s ∈ G.cliqueSet 3 | e ∈ (s : Finset α).sym2}.Subsingleton := by @@ -96,7 +92,7 @@ lemma edgeDisjointTriangles_iff_mem_sym2_subsingleton : ext s simp only [mem_sym2_iff, Sym2.mem_iff, forall_eq_or_imp, forall_eq, Set.sep_and, Set.mem_inter_iff, Set.mem_sep_iff, mem_cliqueSet_iff, Set.mem_setOf_eq, - and_and_and_comm (b := _ ∈ _), _root_.and_self, is3Clique_iff] + and_and_and_comm (b := _ ∈ _), and_self, is3Clique_iff] constructor · rintro ⟨⟨c, d, e, hcd, hce, hde, rfl⟩, hab⟩ simp only [mem_insert, mem_singleton] at hab @@ -280,7 +276,7 @@ lemma FarFromTriangleFree.lt_half (hG : G.FarFromTriangleFree ε) : ε < 2⁻¹ apply tsub_lt_self <;> positivity lemma FarFromTriangleFree.lt_one (hG : G.FarFromTriangleFree ε) : ε < 1 := - hG.lt_half.trans <| inv_lt_one one_lt_two + hG.lt_half.trans two_inv_lt_one theorem FarFromTriangleFree.nonpos (h₀ : G.FarFromTriangleFree ε) (h₁ : G.CliqueFree 3) : ε ≤ 0 := by diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index 4e120e2d50249..fad6a3949ea28 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -67,7 +67,7 @@ namespace YoungDiagram instance : SetLike YoungDiagram (ℕ × ℕ) where -- Porting note (#11215): TODO: figure out how to do this correctly - coe := fun y => y.cells + coe y := y.cells coe_injective' μ ν h := by rwa [YoungDiagram.ext_iff, ← Finset.coe_inj] @[simp] diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index 7017d99c1eb68..211e594b14670 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ -import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Computability.Primrec import Mathlib.Tactic.Ring import Mathlib.Tactic.Linarith @@ -94,7 +93,7 @@ theorem ack_three (n : ℕ) : ack 3 n = 2 ^ (n + 3) - 3 := by have H : 2 * 3 ≤ 2 * 2 ^ 3 := by norm_num apply H.trans rw [_root_.mul_le_mul_left two_pos] - exact pow_le_pow_right one_le_two (Nat.le_add_left 3 n) + exact pow_right_mono₀ one_le_two (Nat.le_add_left 3 n) theorem ack_pos : ∀ m n, 0 < ack m n | 0, n => by simp diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index 103424afc40a5..350ee120e6376 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -275,7 +275,7 @@ end Primcodable namespace Primrec -variable {α : Type*} {σ : Type*} [Primcodable α] [Primcodable σ] +variable {α : Type*} [Primcodable α] open Nat.Primrec @@ -457,8 +457,8 @@ end Primrec₂ namespace Primrec -variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} {σ : Type*} -variable [Primcodable α] [Primcodable β] [Primcodable γ] [Primcodable δ] [Primcodable σ] +variable {α : Type*} {β : Type*} {σ : Type*} +variable [Primcodable α] [Primcodable β] [Primcodable σ] theorem to₂ {f : α × β → σ} (hf : Primrec f) : Primrec₂ fun a b => f (a, b) := hf.of_eq fun _ => rfl @@ -1088,8 +1088,7 @@ end Primrec namespace Primcodable -variable {α : Type*} {β : Type*} -variable [Primcodable α] [Primcodable β] +variable {α : Type*} [Primcodable α] open Primrec @@ -1139,8 +1138,8 @@ end Primcodable namespace Primrec -variable {α : Type*} {β : Type*} {γ : Type*} {σ : Type*} -variable [Primcodable α] [Primcodable β] [Primcodable γ] [Primcodable σ] +variable {α : Type*} {β : Type*} {σ : Type*} +variable [Primcodable α] [Primcodable β] [Primcodable σ] theorem subtype_val {p : α → Prop} [DecidablePred p] {hp : PrimrecPred p} : haveI := Primcodable.subtype hp diff --git a/Mathlib/Condensed/Discrete/Module.lean b/Mathlib/Condensed/Discrete/Module.lean index 0c30c3516af79..3bcf129729c30 100644 --- a/Mathlib/Condensed/Discrete/Module.lean +++ b/Mathlib/Condensed/Discrete/Module.lean @@ -76,13 +76,13 @@ abbrev functor : ModuleCat R ⥤ CondensedMod.{u} R := CompHausLike.LocallyConstantModule.functor.{u+1, u} R (fun _ _ _ ↦ ((CompHaus.effectiveEpi_tfae _).out 0 2).mp) -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₁ (M : ModuleCat.{u+1} R) : M ≅ (ModuleCat.of R (LocallyConstant (CompHaus.of PUnit.{u+1}) M)) where hom := constₗ R inv := evalₗ R PUnit.unit -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₂ (M : ModuleCat R) : (discrete _).obj M ≅ (discrete _).obj (ModuleCat.of R (LocallyConstant (CompHaus.of PUnit.{u+1}) M)) := @@ -103,7 +103,7 @@ instance (M : ModuleCat R) : IsIso ((forget R).map rw [essImage_eq_of_natIso CondensedSet.LocallyConstant.iso.symm] exact obj_mem_essImage CondensedSet.LocallyConstant.functor M -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteComponents (M : ModuleCat R) : (discrete _).obj M ≅ (functor R).obj M := have : (Condensed.forget R).ReflectsIsomorphisms := @@ -183,13 +183,13 @@ abbrev functor : ModuleCat R ⥤ LightCondMod.{u} R := CompHausLike.LocallyConstantModule.functor.{u, u} R (fun _ _ _ ↦ (LightProfinite.effectiveEpi_iff_surjective _).mp) -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₁ (M : ModuleCat.{u} R) : M ≅ (ModuleCat.of R (LocallyConstant (LightProfinite.of PUnit.{u+1}) M)) where hom := constₗ R inv := evalₗ R PUnit.unit -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₂ (M : ModuleCat.{u} R) : (discrete _).obj M ≅ (discrete _).obj (ModuleCat.of R (LocallyConstant (LightProfinite.of PUnit.{u+1}) M)) := @@ -216,7 +216,7 @@ instance (M : ModuleCat R) : rw [essImage_eq_of_natIso LightCondSet.LocallyConstant.iso.symm] exact obj_mem_essImage LightCondSet.LocallyConstant.functor M -/-- Auxilary definition for `functorIsoDiscrete`. -/ +/-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteComponents (M : ModuleCat R) : (discrete _).obj M ≅ (functor R).obj M := have : (LightCondensed.forget R).ReflectsIsomorphisms := diff --git a/Mathlib/Control/Lawful.lean b/Mathlib/Control/Lawful.lean index ce4807b4cde66..f2cd84ab5da94 100644 --- a/Mathlib/Control/Lawful.lean +++ b/Mathlib/Control/Lawful.lean @@ -73,7 +73,6 @@ namespace ReaderT section -variable {ρ : Type u} variable {m : Type u → Type v} variable {α σ : Type u} diff --git a/Mathlib/Control/Monad/Cont.lean b/Mathlib/Control/Monad/Cont.lean index 4a719782d7e74..29ed930cc7cb5 100644 --- a/Mathlib/Control/Monad/Cont.lean +++ b/Mathlib/Control/Monad/Cont.lean @@ -49,7 +49,7 @@ namespace ContT export MonadCont (Label goto) -variable {r : Type u} {m : Type u → Type v} {α β γ ω : Type w} +variable {r : Type u} {m : Type u → Type v} {α β : Type w} def run : ContT r m α → (α → m r) → m r := id diff --git a/Mathlib/Control/Monad/Writer.lean b/Mathlib/Control/Monad/Writer.lean index fcbf590189c9f..a9b71af0c73a6 100644 --- a/Mathlib/Control/Monad/Writer.lean +++ b/Mathlib/Control/Monad/Writer.lean @@ -57,7 +57,7 @@ protected def runThe (ω : Type u) (cmd : WriterT ω M α) : M (α × ω) := cmd @[ext] protected theorem ext {ω : Type u} (x x' : WriterT ω M α) (h : x.run = x'.run) : x = x' := h -variable {ω : Type u} {α β : Type u} [Monad M] +variable [Monad M] /-- Creates an instance of `Monad`, with explicitly given `empty` and `append` operations. diff --git a/Mathlib/Control/Traversable/Basic.lean b/Mathlib/Control/Traversable/Basic.lean index 52d25c007718a..eff2d5afd7056 100644 --- a/Mathlib/Control/Traversable/Basic.lean +++ b/Mathlib/Control/Traversable/Basic.lean @@ -61,8 +61,8 @@ universe u v w section ApplicativeTransformation -variable (F : Type u → Type v) [Applicative F] [LawfulApplicative F] -variable (G : Type u → Type w) [Applicative G] [LawfulApplicative G] +variable (F : Type u → Type v) [Applicative F] +variable (G : Type u → Type w) [Applicative G] /-- A transformation between applicative functors. It is a natural transformation such that `app` preserves the `Pure.pure` and @@ -204,8 +204,7 @@ export Traversable (traverse) section Functions variable {t : Type u → Type u} -variable {m : Type u → Type v} [Applicative m] -variable {α β : Type u} +variable {α : Type u} variable {f : Type u → Type u} [Applicative f] /-- A traversable functor commutes with all applicative functors. -/ @@ -250,8 +249,6 @@ instance : LawfulTraversable Id where section -variable {F : Type u → Type v} [Applicative F] - instance : Traversable Option := ⟨Option.traverse⟩ diff --git a/Mathlib/Control/Traversable/Lemmas.lean b/Mathlib/Control/Traversable/Lemmas.lean index 25ada54df0f22..7b98b4bf635c3 100644 --- a/Mathlib/Control/Traversable/Lemmas.lean +++ b/Mathlib/Control/Traversable/Lemmas.lean @@ -39,7 +39,6 @@ variable [Applicative F] [LawfulApplicative F] variable [Applicative G] [LawfulApplicative G] variable {α β γ : Type u} variable (g : α → F β) -variable (h : β → G γ) variable (f : β → γ) /-- The natural applicative transformation from the identity functor @@ -56,7 +55,7 @@ def PureTransformation : theorem pureTransformation_apply {α} (x : id α) : PureTransformation F x = pure x := rfl -variable {F G} (x : t β) +variable {F G} -- Porting note: need to specify `m/F/G := Id` because `id` no longer has a `Monad` instance theorem map_eq_traverse_id : map (f := t) f = traverse (m := Id) (pure ∘ f) := diff --git a/Mathlib/Control/ULift.lean b/Mathlib/Control/ULift.lean index f294e51cf20de..2482e0095fe06 100644 --- a/Mathlib/Control/ULift.lean +++ b/Mathlib/Control/ULift.lean @@ -78,7 +78,7 @@ end PLift namespace ULift -variable {α : Type u} {β : Type v} {f : α → β} +variable {α : Type u} {β : Type v} /-- Functorial action. -/ protected def map (f : α → β) (a : ULift.{u'} α) : ULift.{v'} β := ULift.up.{v'} (f a.down) diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index bc2017ba047a1..bb39cc9638ee7 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -329,8 +329,7 @@ instance addCommGroup : AddCommGroup ℂ := intros; ext <;> simp [AddMonoid.nsmul_succ, add_mul, add_comm, smul_re, smul_im] zsmul_succ' := by - intros; ext <;> simp [SubNegMonoid.zsmul_succ', add_mul, add_comm, - smul_re, smul_im] + intros; ext <;> simp [add_mul, smul_re, smul_im] zsmul_neg' := by intros; ext <;> simp [zsmul_neg', add_mul, smul_re, smul_im] add_assoc := by intros; ext <;> simp [add_assoc] diff --git a/Mathlib/Data/Complex/BigOperators.lean b/Mathlib/Data/Complex/BigOperators.lean index cd9cb5f344fa0..df65eec1fe310 100644 --- a/Mathlib/Data/Complex/BigOperators.lean +++ b/Mathlib/Data/Complex/BigOperators.lean @@ -3,13 +3,14 @@ Copyright (c) 2017 Kevin Buzzard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Buzzard, Mario Carneiro -/ -import Mathlib.Algebra.BigOperators.Expect +import Mathlib.Algebra.BigOperators.Balance import Mathlib.Data.Complex.Basic /-! # Finite sums and products of complex numbers -/ +open Fintype open scoped BigOperators namespace Complex @@ -28,6 +29,13 @@ theorem ofReal_sum (f : α → ℝ) : ((∑ i ∈ s, f i : ℝ) : ℂ) = ∑ i lemma ofReal_expect (f : α → ℝ) : (𝔼 i ∈ s, f i : ℝ) = 𝔼 i ∈ s, (f i : ℂ) := map_expect ofReal .. +@[simp, norm_cast] +lemma ofReal_balance [Fintype α] (f : α → ℝ) (a : α) : + ((balance f a : ℝ) : ℂ) = balance ((↑) ∘ f) a := by simp [balance] + +@[simp] lemma ofReal_comp_balance {ι : Type*} [Fintype ι] (f : ι → ℝ) : + ofReal ∘ balance f = balance (ofReal ∘ f : ι → ℂ) := funext <| ofReal_balance _ + @[simp] theorem re_sum (f : α → ℂ) : (∑ i ∈ s, f i).re = ∑ i ∈ s, (f i).re := map_sum reAddGroupHom f s @@ -36,6 +44,13 @@ theorem re_sum (f : α → ℂ) : (∑ i ∈ s, f i).re = ∑ i ∈ s, (f i).re lemma re_expect (f : α → ℂ) : (𝔼 i ∈ s, f i).re = 𝔼 i ∈ s, (f i).re := map_expect (LinearMap.mk reAddGroupHom.toAddHom (by simp)) f s +@[simp] +lemma re_balance [Fintype α] (f : α → ℂ) (a : α) : re (balance f a) = balance (re ∘ f) a := by + simp [balance] + +@[simp] lemma re_comp_balance {ι : Type*} [Fintype ι] (f : ι → ℂ) : + re ∘ balance f = balance (re ∘ f) := funext <| re_balance _ + @[simp] theorem im_sum (f : α → ℂ) : (∑ i ∈ s, f i).im = ∑ i ∈ s, (f i).im := map_sum imAddGroupHom f s @@ -44,4 +59,11 @@ theorem im_sum (f : α → ℂ) : (∑ i ∈ s, f i).im = ∑ i ∈ s, (f i).im lemma im_expect (f : α → ℂ) : (𝔼 i ∈ s, f i).im = 𝔼 i ∈ s, (f i).im := map_expect (LinearMap.mk imAddGroupHom.toAddHom (by simp)) f s +@[simp] +lemma im_balance [Fintype α] (f : α → ℂ) (a : α) : im (balance f a) = balance (im ∘ f) a := by + simp [balance] + +@[simp] lemma im_comp_balance {ι : Type*} [Fintype ι] (f : ι → ℂ) : + im ∘ balance f = balance (im ∘ f) := funext <| im_balance _ + end Complex diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 9c2622348ebe2..990585b33e180 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1336,7 +1336,7 @@ theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := gcongr · exact pow_le_one₀ (abs_nonneg _) hx · rw [sq, ← abs_mul_self, abs_mul] - exact mul_le_one hx (abs_nonneg _) hx + exact mul_le_one₀ hx (abs_nonneg _) hx _ < 1 := by norm_num) _ ≤ cos x := sub_le_comm.1 (abs_sub_le_iff.1 (cos_bound hx)).2 @@ -1412,7 +1412,7 @@ theorem add_one_lt_exp {x : ℝ} (hx : x ≠ 0) : x + 1 < Real.exp x := by obtain h' | h' := le_or_lt 1 (-x) · linarith [x.exp_pos] have hx' : 0 < x + 1 := by linarith - simpa [add_comm, exp_neg, inv_lt_inv (exp_pos _) hx'] + simpa [add_comm, exp_neg, inv_lt_inv₀ (exp_pos _) hx'] using exp_bound_div_one_sub_of_interval' (neg_pos.2 hx) h' theorem add_one_le_exp (x : ℝ) : x + 1 ≤ Real.exp x := by diff --git a/Mathlib/Data/Complex/ExponentialBounds.lean b/Mathlib/Data/Complex/ExponentialBounds.lean index cb4022b3c3fe8..9531022b50196 100644 --- a/Mathlib/Data/Complex/ExponentialBounds.lean +++ b/Mathlib/Data/Complex/ExponentialBounds.lean @@ -36,16 +36,14 @@ theorem exp_one_lt_d9 : exp 1 < 2.7182818286 := lt_of_le_of_lt (sub_le_iff_le_add.1 (abs_sub_le_iff.1 exp_one_near_10).1) (by norm_num) theorem exp_neg_one_gt_d9 : 0.36787944116 < exp (-1) := by - rw [exp_neg, lt_inv _ (exp_pos _)] + rw [exp_neg, lt_inv_comm₀ _ (exp_pos _)] · refine lt_of_le_of_lt (sub_le_iff_le_add.1 (abs_sub_le_iff.1 exp_one_near_10).1) ?_ norm_num · norm_num theorem exp_neg_one_lt_d9 : exp (-1) < 0.3678794412 := by - rw [exp_neg, inv_lt (exp_pos _)] - · refine lt_of_lt_of_le ?_ (sub_le_comm.1 (abs_sub_le_iff.1 exp_one_near_10).2) - norm_num - · norm_num + rw [exp_neg, inv_lt_comm₀ (exp_pos _) (by norm_num)] + exact lt_of_lt_of_le (by norm_num) (sub_le_comm.1 (abs_sub_le_iff.1 exp_one_near_10).2) theorem log_two_near_10 : |log 2 - 287209 / 414355| ≤ 1 / 10 ^ 10 := by suffices |log 2 - 287209 / 414355| ≤ 1 / 17179869184 + (1 / 10 ^ 10 - 1 / 2 ^ 34) by diff --git a/Mathlib/Data/Countable/Basic.lean b/Mathlib/Data/Countable/Basic.lean index 9b80d33df3f88..89694dd120b07 100644 --- a/Mathlib/Data/Countable/Basic.lean +++ b/Mathlib/Data/Countable/Basic.lean @@ -66,7 +66,7 @@ instance Sum.uncountable_inr [Uncountable β] : Uncountable (α ⊕ β) := inr_injective.uncountable instance Option.instCountable [Countable α] : Countable (Option α) := - Countable.of_equiv _ (Equiv.optionEquivSumPUnit.{_, 0} α).symm + Countable.of_equiv _ (Equiv.optionEquivSumPUnit.{0, _} α).symm instance WithTop.instCountable [Countable α] : Countable (WithTop α) := Option.instCountable instance WithBot.instCountable [Countable α] : Countable (WithBot α) := Option.instCountable diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 5847cba382776..1fc1f2eb0733b 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -574,13 +574,13 @@ theorem Ioo_zero_top_eq_iUnion_Ico_zpow {y : ℝ≥0∞} (hy : 1 < y) (h'y : y theorem zpow_le_of_le {x : ℝ≥0∞} (hx : 1 ≤ x) {a b : ℤ} (h : a ≤ b) : x ^ a ≤ x ^ b := by induction' a with a a <;> induction' b with b b · simp only [Int.ofNat_eq_coe, zpow_natCast] - exact pow_le_pow_right hx (Int.le_of_ofNat_le_ofNat h) + exact pow_right_mono₀ hx (Int.le_of_ofNat_le_ofNat h) · apply absurd h (not_le_of_gt _) exact lt_of_lt_of_le (Int.negSucc_lt_zero _) (Int.ofNat_nonneg _) · simp only [zpow_negSucc, Int.ofNat_eq_coe, zpow_natCast] refine (ENNReal.inv_le_one.2 ?_).trans ?_ <;> exact one_le_pow_of_one_le' hx _ · simp only [zpow_negSucc, ENNReal.inv_le_inv] - apply pow_le_pow_right hx + apply pow_right_mono₀ hx simpa only [← Int.ofNat_le, neg_le_neg_iff, Int.ofNat_add, Int.ofNat_one, Int.negSucc_eq] using h diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index d875809cf1a4a..edc71f8ff4d06 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -680,13 +680,14 @@ def castSuccEmb : Fin n ↪ Fin (n + 1) := castAddEmb _ @[simp, norm_cast] lemma coe_castSuccEmb : (castSuccEmb : Fin n → Fin (n + 1)) = Fin.castSucc := rfl -@[simp] -theorem castSucc_le_castSucc_iff {a b : Fin n} : castSucc a ≤ castSucc b ↔ a ≤ b := Iff.rfl -@[simp] -theorem succ_le_castSucc_iff {a b : Fin n} : succ a ≤ castSucc b ↔ a < b := by +theorem castSucc_le_succ {n} (i : Fin n) : i.castSucc ≤ i.succ := Nat.le_succ i + +@[simp] theorem castSucc_le_castSucc_iff {a b : Fin n} : castSucc a ≤ castSucc b ↔ a ≤ b := .rfl + +@[simp] theorem succ_le_castSucc_iff {a b : Fin n} : succ a ≤ castSucc b ↔ a < b := by rw [le_castSucc_iff, succ_lt_succ_iff] -@[simp] -theorem castSucc_lt_succ_iff {a b : Fin n} : castSucc a < succ b ↔ a ≤ b := by + +@[simp] theorem castSucc_lt_succ_iff {a b : Fin n} : castSucc a < succ b ↔ a ≤ b := by rw [castSucc_lt_iff_succ_le, succ_le_succ_iff] theorem le_of_castSucc_lt_of_succ_lt {a b : Fin (n + 1)} {i : Fin n} @@ -1514,10 +1515,16 @@ lemma sub_succ_le_sub_of_le {n : ℕ} {u v : Fin (n + 2)} (h : u < v) : v - (u + end AddGroup @[simp] -theorem coe_ofNat_eq_mod (m n : ℕ) [NeZero m] : +theorem coe_natCast_eq_mod (m n : ℕ) [NeZero m] : ((n : Fin m) : ℕ) = n % m := rfl +-- See note [no_index around OfNat.ofNat] +@[simp] +theorem coe_ofNat_eq_mod (m n : ℕ) [NeZero m] : + ((no_index OfNat.ofNat n : Fin m) : ℕ) = OfNat.ofNat n % m := + rfl + section Mul /-! diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 3e1b0ec9fd4a7..e86ab424aa5fb 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -807,6 +807,11 @@ theorem ssubset_iff_exists_cons_subset : s ⊂ t ↔ ∃ (a : _) (h : a ∉ s), obtain ⟨a, hs, ht⟩ := not_subset.1 h.2 exact ⟨a, ht, cons_subset.2 ⟨hs, h.subset⟩⟩ +theorem cons_swap (hb : b ∉ s) (ha : a ∉ s.cons b hb) : + (s.cons b hb).cons a ha = (s.cons a fun h ↦ ha (mem_cons.mpr (.inr h))).cons b fun h ↦ + ha (mem_cons.mpr (.inl ((mem_cons.mp h).elim symm (fun h ↦ False.elim (hb h))))) := + eq_of_veq <| Multiset.cons_swap a b s.val + end Cons /-! ### disjoint -/ @@ -3093,4 +3098,4 @@ def proveFinsetNonempty {u : Level} {α : Q(Type u)} (s : Q(Finset $α)) : end Mathlib.Meta -set_option linter.style.longFile 3100 +set_option linter.style.longFile 3200 diff --git a/Mathlib/Data/Finset/Image.lean b/Mathlib/Data/Finset/Image.lean index 3fc36f5e38e47..88939058da20c 100644 --- a/Mathlib/Data/Finset/Image.lean +++ b/Mathlib/Data/Finset/Image.lean @@ -697,9 +697,23 @@ theorem fin_mono {n} : Monotone (Finset.fin n) := fun s t h x => by simpa using theorem fin_map {n} {s : Finset ℕ} : (s.fin n).map Fin.valEmbedding = s.filter (· < n) := by simp [Finset.fin, Finset.map_map] +/-- +If a finset `t` is a subset of the image of another finset `s` under `f`, then it is equal to the +image of a subset of `s`. + +For the version where `s` is a set, see `subset_set_image_iff`. +-/ +theorem subset_image_iff [DecidableEq β] {s : Finset α} {t : Finset β} {f : α → β} : + t ⊆ s.image f ↔ ∃ s' : Finset α, s' ⊆ s ∧ s'.image f = t := by + refine ⟨fun ht => ?_, fun ⟨s', hs', h⟩ => h ▸ image_subset_image hs'⟩ + refine ⟨s.filter (f · ∈ t), filter_subset _ _, le_antisymm (by simp [image_subset_iff]) ?_⟩ + intro x hx + specialize ht hx + aesop + /-- If a `Finset` is a subset of the image of a `Set` under `f`, then it is equal to the `Finset.image` of a `Finset` subset of that `Set`. -/ -theorem subset_image_iff [DecidableEq β] {s : Set α} {t : Finset β} {f : α → β} : +theorem subset_set_image_iff [DecidableEq β] {s : Set α} {t : Finset β} {f : α → β} : ↑t ⊆ f '' s ↔ ∃ s' : Finset α, ↑s' ⊆ s ∧ s'.image f = t := by constructor; swap · rintro ⟨t, ht, rfl⟩ diff --git a/Mathlib/Data/Finset/NAry.lean b/Mathlib/Data/Finset/NAry.lean index fa32e7a832096..4db70840b348d 100644 --- a/Mathlib/Data/Finset/NAry.lean +++ b/Mathlib/Data/Finset/NAry.lean @@ -454,9 +454,9 @@ theorem card_dvd_card_image₂_left (hf : ∀ b ∈ t, Injective fun a => f a b) /-- If a `Finset` is a subset of the image of two `Set`s under a binary operation, then it is a subset of the `Finset.image₂` of two `Finset` subsets of these `Set`s. -/ -theorem subset_image₂ {s : Set α} {t : Set β} (hu : ↑u ⊆ image2 f s t) : +theorem subset_set_image₂ {s : Set α} {t : Set β} (hu : ↑u ⊆ image2 f s t) : ∃ (s' : Finset α) (t' : Finset β), ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ image₂ f s' t' := by - rw [← Set.image_prod, subset_image_iff] at hu + rw [← Set.image_prod, subset_set_image_iff] at hu rcases hu with ⟨u, hu, rfl⟩ classical use u.image Prod.fst, u.image Prod.snd @@ -464,6 +464,8 @@ theorem subset_image₂ {s : Set α} {t : Set β} (hu : ↑u ⊆ image2 f s t) : image_subset_iff] exact ⟨fun _ h ↦ (hu h).1, fun _ h ↦ (hu h).2, fun x hx ↦ mem_image₂_of_mem hx hx⟩ +@[deprecated (since := "2024-09-22")] alias subset_image₂ := subset_set_image₂ + end section UnionInter diff --git a/Mathlib/Data/Finset/NatDivisors.lean b/Mathlib/Data/Finset/NatDivisors.lean index 03edf844dbd26..cbc7aa896c2f5 100644 --- a/Mathlib/Data/Finset/NatDivisors.lean +++ b/Mathlib/Data/Finset/NatDivisors.lean @@ -16,16 +16,12 @@ exhibiting `Nat.divisors` as a multiplicative homomorphism from `ℕ` to `Finset open Nat Finset open scoped Pointwise -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- The divisors of a product of natural numbers are the pointwise product of the divisors of the factors. -/ lemma Nat.divisors_mul (m n : ℕ) : divisors (m * n) = divisors m * divisors n := by ext k simp_rw [mem_mul, mem_divisors, dvd_mul, mul_ne_zero_iff, ← exists_and_left, ← exists_and_right] - simp only [_root_.and_assoc, _root_.and_comm, and_left_comm] + simp only [and_assoc, and_comm, and_left_comm] /-- `Nat.divisors` as a `MonoidHom`. -/ @[simps] diff --git a/Mathlib/Data/Finset/Preimage.lean b/Mathlib/Data/Finset/Preimage.lean index 45d7cdc8ae557..657acee065919 100644 --- a/Mathlib/Data/Finset/Preimage.lean +++ b/Mathlib/Data/Finset/Preimage.lean @@ -100,7 +100,7 @@ theorem preimage_subset {f : α ↪ β} {s : Finset β} {t : Finset α} (hs : s theorem subset_map_iff {f : α ↪ β} {s : Finset β} {t : Finset α} : s ⊆ t.map f ↔ ∃ u ⊆ t, s = u.map f := by classical - simp_rw [← coe_subset, coe_map, subset_image_iff, map_eq_image, eq_comm] + simp_rw [map_eq_image, subset_image_iff, eq_comm] theorem sigma_preimage_mk {β : α → Type*} [DecidableEq α] (s : Finset (Σa, β a)) (t : Finset α) : (t.sigma fun a => s.preimage (Sigma.mk a) sigma_mk_injective.injOn) = diff --git a/Mathlib/Data/Finset/Sups.lean b/Mathlib/Data/Finset/Sups.lean index 16e426c59c332..5188385616bb7 100644 --- a/Mathlib/Data/Finset/Sups.lean +++ b/Mathlib/Data/Finset/Sups.lean @@ -146,7 +146,7 @@ theorem sups_inter_subset_right : s ⊻ (t₁ ∩ t₂) ⊆ s ⊻ t₁ ∩ s ⊻ theorem subset_sups {s t : Set α} : ↑u ⊆ s ⊻ t → ∃ s' t' : Finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' ⊻ t' := - subset_image₂ + subset_set_image₂ lemma image_sups (f : F) (s t : Finset α) : image f (s ⊻ t) = image f s ⊻ image f t := image_image₂_distrib <| map_sup f @@ -293,7 +293,7 @@ theorem infs_inter_subset_right : s ⊼ (t₁ ∩ t₂) ⊆ s ⊼ t₁ ∩ s ⊼ theorem subset_infs {s t : Set α} : ↑u ⊆ s ⊼ t → ∃ s' t' : Finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' ⊼ t' := - subset_image₂ + subset_set_image₂ lemma image_infs (f : F) (s t : Finset α) : image f (s ⊼ t) = image f s ⊼ image f t := image_image₂_distrib <| map_inf f @@ -582,7 +582,7 @@ lemma diffs_inter_subset_right : s \\ (t₁ ∩ t₂) ⊆ s \\ t₁ ∩ s \\ t lemma subset_diffs {s t : Set α} : ↑u ⊆ Set.image2 (· \ ·) s t → ∃ s' t' : Finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' \\ t' := - subset_image₂ + subset_set_image₂ variable (s t u) diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 70c91059b3738..11717314f657f 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -1048,6 +1048,60 @@ theorem induction_linear {p : (α →₀ M) → Prop} (f : α →₀ M) (h0 : p (hadd : ∀ f g : α →₀ M, p f → p g → p (f + g)) (hsingle : ∀ a b, p (single a b)) : p f := induction₂ f h0 fun _a _b _f _ _ w => hadd _ _ w (hsingle _ _) +section LinearOrder + +variable [LinearOrder α] {p : (α →₀ M) → Prop} + +/-- A finitely supported function can be built by adding up `single a b` for increasing `a`. + +The theorem `induction_on_max₂` swaps the argument order in the sum. -/ +theorem induction_on_max (f : α →₀ M) (h0 : p 0) + (ha : ∀ (a b) (f : α →₀ M), (∀ c ∈ f.support, c < a) → b ≠ 0 → p f → p (single a b + f)) : + p f := by + suffices ∀ (s) (f : α →₀ M), f.support = s → p f from this _ _ rfl + refine fun s => s.induction_on_max (fun f h => ?_) (fun a s hm hf f hs => ?_) + · rwa [support_eq_empty.1 h] + · have hs' : (erase a f).support = s := by + rw [support_erase, hs, erase_insert (fun ha => (hm a ha).false)] + rw [← single_add_erase a f] + refine ha _ _ _ (fun c hc => hm _ <| hs'.symm ▸ hc) ?_ (hf _ hs') + rw [← mem_support_iff, hs] + exact mem_insert_self a s + +/-- A finitely supported function can be built by adding up `single a b` for decreasing `a`. + +The theorem `induction_on_min₂` swaps the argument order in the sum. -/ +theorem induction_on_min (f : α →₀ M) (h0 : p 0) + (ha : ∀ (a b) (f : α →₀ M), (∀ c ∈ f.support, a < c) → b ≠ 0 → p f → p (single a b + f)) : + p f := + induction_on_max (α := αᵒᵈ) f h0 ha + +/-- A finitely supported function can be built by adding up `single a b` for increasing `a`. + +The theorem `induction_on_max` swaps the argument order in the sum. -/ +theorem induction_on_max₂ (f : α →₀ M) (h0 : p 0) + (ha : ∀ (a b) (f : α →₀ M), (∀ c ∈ f.support, c < a) → b ≠ 0 → p f → p (f + single a b)) : + p f := by + suffices ∀ (s) (f : α →₀ M), f.support = s → p f from this _ _ rfl + refine fun s => s.induction_on_max (fun f h => ?_) (fun a s hm hf f hs => ?_) + · rwa [support_eq_empty.1 h] + · have hs' : (erase a f).support = s := by + rw [support_erase, hs, erase_insert (fun ha => (hm a ha).false)] + rw [← erase_add_single a f] + refine ha _ _ _ (fun c hc => hm _ <| hs'.symm ▸ hc) ?_ (hf _ hs') + rw [← mem_support_iff, hs] + exact mem_insert_self a s + +/-- A finitely supported function can be built by adding up `single a b` for decreasing `a`. + +The theorem `induction_on_min` swaps the argument order in the sum. -/ +theorem induction_on_min₂ (f : α →₀ M) (h0 : p 0) + (ha : ∀ (a b) (f : α →₀ M), (∀ c ∈ f.support, a < c) → b ≠ 0 → p f → p (f + single a b)) : + p f := + induction_on_max₂ (α := αᵒᵈ) f h0 ha + +end LinearOrder + @[simp] theorem add_closure_setOf_eq_single : AddSubmonoid.closure { f : α →₀ M | ∃ a b, f = single a b } = ⊤ := diff --git a/Mathlib/Data/Finsupp/Order.lean b/Mathlib/Data/Finsupp/Order.lean index f36bba9cf7a88..4726fc0f7bb64 100644 --- a/Mathlib/Data/Finsupp/Order.lean +++ b/Mathlib/Data/Finsupp/Order.lean @@ -3,7 +3,9 @@ Copyright (c) 2021 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Aaron Anderson -/ +import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.Module.Defs +import Mathlib.Algebra.Order.Pi import Mathlib.Data.Finsupp.Basic /-! @@ -21,7 +23,7 @@ noncomputable section open Finset -variable {ι α β : Type*} +variable {ι κ α β : Type*} namespace Finsupp @@ -32,6 +34,15 @@ section Zero variable [Zero α] +section OrderedAddCommMonoid +variable [OrderedAddCommMonoid β] {f : ι →₀ α} {h₁ h₂ : ι → α → β} + +@[gcongr] +lemma sum_le_sum (h : ∀ i ∈ f.support, h₁ i (f i) ≤ h₂ i (f i)) : f.sum h₁ ≤ f.sum h₂ := + Finset.sum_le_sum h + +end OrderedAddCommMonoid + section LE variable [LE α] {f g : ι →₀ α} @@ -58,7 +69,7 @@ theorem orderEmbeddingToFun_apply {f : ι →₀ α} {i : ι} : orderEmbeddingTo end LE section Preorder -variable [Preorder α] {f g : ι →₀ α} +variable [Preorder α] {f g : ι →₀ α} {i : ι} {a b : α} instance preorder : Preorder (ι →₀ α) := { Finsupp.instLEFinsupp with @@ -72,6 +83,26 @@ lemma coe_mono : Monotone (Finsupp.toFun : (ι →₀ α) → ι → α) := fun lemma coe_strictMono : Monotone (Finsupp.toFun : (ι →₀ α) → ι → α) := fun _ _ ↦ id +@[simp] lemma single_le_single : single i a ≤ single i b ↔ a ≤ b := by + classical exact Pi.single_le_single + +lemma single_mono : Monotone (single i : α → ι →₀ α) := fun _ _ ↦ single_le_single.2 + +@[gcongr] protected alias ⟨_, GCongr.single_mono⟩ := single_le_single + +@[simp] lemma single_nonneg : 0 ≤ single i a ↔ 0 ≤ a := by classical exact Pi.single_nonneg +@[simp] lemma single_nonpos : single i a ≤ 0 ↔ a ≤ 0 := by classical exact Pi.single_nonpos + +variable [OrderedAddCommMonoid β] + +lemma sum_le_sum_index [DecidableEq ι] {f₁ f₂ : ι →₀ α} {h : ι → α → β} (hf : f₁ ≤ f₂) + (hh : ∀ i ∈ f₁.support ∪ f₂.support, Monotone (h i)) + (hh₀ : ∀ i ∈ f₁.support ∪ f₂.support, h i 0 = 0) : f₁.sum h ≤ f₂.sum h := by + classical + rw [sum_of_support_subset _ Finset.subset_union_left _ hh₀, + sum_of_support_subset _ Finset.subset_union_right _ hh₀] + exact Finset.sum_le_sum fun i hi ↦ hh _ hi <| hf _ + end Preorder instance partialorder [PartialOrder α] : PartialOrder (ι →₀ α) := @@ -117,11 +148,24 @@ end Zero /-! ### Algebraic order structures -/ +section OrderedAddCommMonoid +variable [OrderedAddCommMonoid α] {i : ι} {f : ι → κ} {g g₁ g₂ : ι →₀ α} -instance orderedAddCommMonoid [OrderedAddCommMonoid α] : OrderedAddCommMonoid (ι →₀ α) := +instance orderedAddCommMonoid : OrderedAddCommMonoid (ι →₀ α) := { Finsupp.instAddCommMonoid, Finsupp.partialorder with add_le_add_left := fun _a _b h c s => add_le_add_left (h s) (c s) } +lemma mapDomain_mono : Monotone (mapDomain f : (ι →₀ α) → (κ →₀ α)) := by + classical exact fun g₁ g₂ h ↦ sum_le_sum_index h (fun _ _ ↦ single_mono) (by simp) + +@[gcongr] protected lemma GCongr.mapDomain_mono (hg : g₁ ≤ g₂) : g₁.mapDomain f ≤ g₂.mapDomain f := + mapDomain_mono hg + +lemma mapDomain_nonneg (hg : 0 ≤ g) : 0 ≤ g.mapDomain f := by simpa using mapDomain_mono hg +lemma mapDomain_nonpos (hg : g ≤ 0) : g.mapDomain f ≤ 0 := by simpa using mapDomain_mono hg + +end OrderedAddCommMonoid + instance orderedCancelAddCommMonoid [OrderedCancelAddCommMonoid α] : OrderedCancelAddCommMonoid (ι →₀ α) := { Finsupp.orderedAddCommMonoid with diff --git a/Mathlib/Data/Fintype/Sigma.lean b/Mathlib/Data/Fintype/Sigma.lean index 64f8de8f89ab0..6827beb23b927 100644 --- a/Mathlib/Data/Fintype/Sigma.lean +++ b/Mathlib/Data/Fintype/Sigma.lean @@ -17,7 +17,7 @@ open Nat universe u v -variable {ι α β γ : Type*} {κ : ι → Type*} [Π i, Fintype (κ i)] +variable {ι α : Type*} {κ : ι → Type*} [Π i, Fintype (κ i)] open Finset Function diff --git a/Mathlib/Data/FunLike/Basic.lean b/Mathlib/Data/FunLike/Basic.lean index f23481b549544..3a04435a85ca6 100644 --- a/Mathlib/Data/FunLike/Basic.lean +++ b/Mathlib/Data/FunLike/Basic.lean @@ -17,7 +17,7 @@ There is the "D"ependent version `DFunLike` and the non-dependent version `FunLi A typical type of morphisms should be declared as: ``` -structure MyHom (A B : Type*) [MyClass A] [MyClass B] := +structure MyHom (A B : Type*) [MyClass A] [MyClass B] where (toFun : A → B) (map_op' : ∀ (x y : A), toFun (MyClass.op x y) = MyClass.op (toFun x) (toFun y)) @@ -79,7 +79,7 @@ The second step is to add instances of your new `MyHomClass` for all types exten Typically, you can just declare a new class analogous to `MyHomClass`: ``` -structure CoolerHom (A B : Type*) [CoolClass A] [CoolClass B] extends MyHom A B := +structure CoolerHom (A B : Type*) [CoolClass A] [CoolClass B] extends MyHom A B where (map_cool' : toFun CoolClass.cool = CoolClass.cool) class CoolerHomClass (F : Type*) (A B : outParam Type*) [CoolClass A] [CoolClass B] diff --git a/Mathlib/Data/FunLike/Embedding.lean b/Mathlib/Data/FunLike/Embedding.lean index b65a26da5cc05..a8551af72e079 100644 --- a/Mathlib/Data/FunLike/Embedding.lean +++ b/Mathlib/Data/FunLike/Embedding.lean @@ -14,7 +14,7 @@ This typeclass is primarily for use by embeddings such as `RelEmbedding`. A typical type of embeddings should be declared as: ``` -structure MyEmbedding (A B : Type*) [MyClass A] [MyClass B] := +structure MyEmbedding (A B : Type*) [MyClass A] [MyClass B] where (toFun : A → B) (injective' : Function.Injective toFun) (map_op' : ∀ (x y : A), toFun (MyClass.op x y) = MyClass.op (toFun x) (toFun y)) @@ -58,8 +58,8 @@ Continuing the example above: You should extend this class when you extend `MyEmbedding`. -/ class MyEmbeddingClass (F : Type*) (A B : outParam Type*) [MyClass A] [MyClass B] [FunLike F A B] - extends EmbeddingLike F A B := - (map_op : ∀ (f : F) (x y : A), f (MyClass.op x y) = MyClass.op (f x) (f y)) + extends EmbeddingLike F A B where + map_op : ∀ (f : F) (x y : A), f (MyClass.op x y) = MyClass.op (f x) (f y) @[simp] lemma map_op {F A B : Type*} [MyClass A] [MyClass B] [FunLike F A B] [MyEmbeddingClass F A B] @@ -84,12 +84,12 @@ The second step is to add instances of your new `MyEmbeddingClass` for all types Typically, you can just declare a new class analogous to `MyEmbeddingClass`: ``` -structure CoolerEmbedding (A B : Type*) [CoolClass A] [CoolClass B] extends MyEmbedding A B := +structure CoolerEmbedding (A B : Type*) [CoolClass A] [CoolClass B] extends MyEmbedding A B where (map_cool' : toFun CoolClass.cool = CoolClass.cool) class CoolerEmbeddingClass (F : Type*) (A B : outParam Type*) [CoolClass A] [CoolClass B] [FunLike F A B] - extends MyEmbeddingClass F A B := + extends MyEmbeddingClass F A B where (map_cool : ∀ (f : F), f CoolClass.cool = CoolClass.cool) @[simp] diff --git a/Mathlib/Data/FunLike/Equiv.lean b/Mathlib/Data/FunLike/Equiv.lean index d2a640eddc593..92881dc5ec3a9 100644 --- a/Mathlib/Data/FunLike/Equiv.lean +++ b/Mathlib/Data/FunLike/Equiv.lean @@ -14,7 +14,7 @@ This typeclass is primarily for use by isomorphisms like `MonoidEquiv` and `Line A typical type of isomorphisms should be declared as: ``` -structure MyIso (A B : Type*) [MyClass A] [MyClass B] extends Equiv A B := +structure MyIso (A B : Type*) [MyClass A] [MyClass B] extends Equiv A B where (map_op' : ∀ (x y : A), toFun (MyClass.op x y) = MyClass.op (toFun x) (toFun y)) namespace MyIso @@ -77,12 +77,12 @@ The second step is to add instances of your new `MyIsoClass` for all types exten Typically, you can just declare a new class analogous to `MyIsoClass`: ``` -structure CoolerIso (A B : Type*) [CoolClass A] [CoolClass B] extends MyIso A B := +structure CoolerIso (A B : Type*) [CoolClass A] [CoolClass B] extends MyIso A B where (map_cool' : toFun CoolClass.cool = CoolClass.cool) class CoolerIsoClass (F : Type*) (A B : outParam Type*) [CoolClass A] [CoolClass B] [EquivLike F A B] - extends MyIsoClass F A B := + extends MyIsoClass F A B where (map_cool : ∀ (f : F), f CoolClass.cool = CoolClass.cool) @[simp] lemma map_cool {F A B : Type*} [CoolClass A] [CoolClass B] diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 4cb4e2f0e0ba6..6bf06c9ff023a 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -30,7 +30,6 @@ namespace Int variable {a b c d m n : ℤ} section Order -variable {a b c : ℤ} protected lemma le_rfl : a ≤ a := a.le_refl protected lemma lt_or_lt_of_ne : a ≠ b → a < b ∨ b < a := Int.lt_or_gt_of_ne @@ -44,12 +43,7 @@ protected lemma le_antisymm_iff : a = b ↔ a ≤ b ∧ b ≤ a := protected lemma le_iff_eq_or_lt : a ≤ b ↔ a = b ∨ a < b := by rw [Int.le_antisymm_iff, Int.lt_iff_le_not_le, ← and_or_left]; simp [em] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ -protected lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := by rw [Int.le_iff_eq_or_lt, - _root_.or_comm] +protected lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := by rw [Int.le_iff_eq_or_lt, or_comm] end Order diff --git a/Mathlib/Data/Int/Log.lean b/Mathlib/Data/Int/Log.lean index 8d9efb0ac2af0..ac78731f26da4 100644 --- a/Mathlib/Data/Int/Log.lean +++ b/Mathlib/Data/Int/Log.lean @@ -94,7 +94,7 @@ theorem zpow_log_le_self {b : ℕ} {r : R} (hb : 1 < b) (hr : 0 < r) : (b : R) ^ rw [zpow_natCast, ← Nat.cast_pow, ← Nat.le_floor_iff hr.le] exact Nat.pow_log_le_self b (Nat.floor_pos.mpr hr1).ne' · rw [log_of_right_le_one _ hr1, zpow_neg, zpow_natCast, ← Nat.cast_pow] - exact inv_le_of_inv_le hr (Nat.ceil_le.1 <| Nat.le_pow_clog hb _) + exact inv_le_of_inv_le₀ hr (Nat.ceil_le.1 <| Nat.le_pow_clog hb _) theorem lt_zpow_succ_log_self {b : ℕ} (hb : 1 < b) (r : R) : r < (b : R) ^ (log b r + 1) := by rcases le_or_lt r 0 with hr | hr @@ -106,11 +106,11 @@ theorem lt_zpow_succ_log_self {b : ℕ} (hb : 1 < b) (r : R) : r < (b : R) ^ (lo apply Nat.lt_of_floor_lt exact Nat.lt_pow_succ_log_self hb _ · rw [log_of_right_le_one _ hr1.le] - have hcri : 1 < r⁻¹ := one_lt_inv hr hr1 + have hcri : 1 < r⁻¹ := (one_lt_inv₀ hr).2 hr1 have : 1 ≤ Nat.clog b ⌈r⁻¹⌉₊ := Nat.succ_le_of_lt (Nat.clog_pos hb <| Nat.one_lt_cast.1 <| hcri.trans_le (Nat.le_ceil _)) rw [neg_add_eq_sub, ← neg_sub, ← Int.ofNat_one, ← Int.ofNat_sub this, zpow_neg, zpow_natCast, - lt_inv hr (pow_pos (Nat.cast_pos.mpr <| zero_lt_one.trans hb) _), ← Nat.cast_pow] + lt_inv_comm₀ hr (pow_pos (Nat.cast_pos.mpr <| zero_lt_one.trans hb) _), ← Nat.cast_pow] refine Nat.lt_ceil.1 ?_ exact Nat.pow_pred_clog_lt_self hb <| Nat.one_lt_cast.1 <| hcri.trans_le <| Nat.le_ceil _ @@ -146,7 +146,7 @@ theorem log_zpow {b : ℕ} (hb : 1 < b) (z : ℤ) : log b ((b : R) ^ z : R) = z theorem log_mono_right {b : ℕ} {r₁ r₂ : R} (h₀ : 0 < r₁) (h : r₁ ≤ r₂) : log b r₁ ≤ log b r₂ := by rcases le_total r₁ 1 with h₁ | h₁ <;> rcases le_total r₂ 1 with h₂ | h₂ · rw [log_of_right_le_one _ h₁, log_of_right_le_one _ h₂, neg_le_neg_iff, Int.ofNat_le] - exact Nat.clog_mono_right _ (Nat.ceil_mono <| inv_le_inv_of_le h₀ h) + exact Nat.clog_mono_right _ (Nat.ceil_mono <| inv_anti₀ h₀ h) · rw [log_of_right_le_one _ h₁, log_of_one_le_right _ h₂] exact (neg_nonpos.mpr (Int.natCast_nonneg _)).trans (Int.natCast_nonneg _) · obtain rfl := le_antisymm h (h₂.trans h₁) @@ -203,8 +203,8 @@ theorem clog_of_right_le_zero (b : ℕ) {r : R} (hr : r ≤ 0) : clog b r = 0 := theorem clog_inv (b : ℕ) (r : R) : clog b r⁻¹ = -log b r := by cases' lt_or_le 0 r with hrp hrp · obtain hr | hr := le_total 1 r - · rw [clog_of_right_le_one _ (inv_le_one hr), log_of_one_le_right _ hr, inv_inv] - · rw [clog_of_one_le_right _ (one_le_inv hrp hr), log_of_right_le_one _ hr, neg_neg] + · rw [clog_of_right_le_one _ (inv_le_one_of_one_le₀ hr), log_of_one_le_right _ hr, inv_inv] + · rw [clog_of_one_le_right _ ((one_le_inv₀ hrp).2 hr), log_of_right_le_one _ hr, neg_neg] · rw [clog_of_right_le_zero _ (inv_nonpos.mpr hrp), log_of_right_le_zero _ hrp, neg_zero] @[simp] @@ -235,13 +235,13 @@ theorem self_le_zpow_clog {b : ℕ} (hb : 1 < b) (r : R) : r ≤ (b : R) ^ clog rcases le_or_lt r 0 with hr | hr · rw [clog_of_right_le_zero _ hr, zpow_zero] exact hr.trans zero_le_one - rw [← neg_log_inv_eq_clog, zpow_neg, le_inv hr (zpow_pos_of_pos _ _)] + rw [← neg_log_inv_eq_clog, zpow_neg, le_inv_comm₀ hr (zpow_pos_of_pos _ _)] · exact zpow_log_le_self hb (inv_pos.mpr hr) · exact Nat.cast_pos.mpr (zero_le_one.trans_lt hb) theorem zpow_pred_clog_lt_self {b : ℕ} {r : R} (hb : 1 < b) (hr : 0 < r) : (b : R) ^ (clog b r - 1) < r := by - rw [← neg_log_inv_eq_clog, ← neg_add', zpow_neg, inv_lt _ hr] + rw [← neg_log_inv_eq_clog, ← neg_add', zpow_neg, inv_lt_comm₀ _ hr] · exact lt_zpow_succ_log_self hb _ · exact zpow_pos_of_pos (Nat.cast_pos.mpr <| zero_le_one.trans_lt hb) _ @@ -271,7 +271,7 @@ theorem clog_zpow {b : ℕ} (hb : 1 < b) (z : ℤ) : clog b ((b : R) ^ z : R) = theorem clog_mono_right {b : ℕ} {r₁ r₂ : R} (h₀ : 0 < r₁) (h : r₁ ≤ r₂) : clog b r₁ ≤ clog b r₂ := by rw [← neg_log_inv_eq_clog, ← neg_log_inv_eq_clog, neg_le_neg_iff] - exact log_mono_right (inv_pos.mpr <| h₀.trans_le h) (inv_le_inv_of_le h₀ h) + exact log_mono_right (inv_pos.mpr <| h₀.trans_le h) (inv_anti₀ h₀ h) variable (R) diff --git a/Mathlib/Data/Int/WithZero.lean b/Mathlib/Data/Int/WithZero.lean index 4e3294a68595e..3cf41e5f88e38 100644 --- a/Mathlib/Data/Int/WithZero.lean +++ b/Mathlib/Data/Int/WithZero.lean @@ -63,7 +63,9 @@ theorem toNNReal_pos_apply {e : NNReal} (he : e ≠ 0) {x : ℤₘ₀} (hx : x = theorem toNNReal_neg_apply {e : NNReal} (he : e ≠ 0) {x : ℤₘ₀} (hx : x ≠ 0) : toNNReal he x = e ^ Multiplicative.toAdd (WithZero.unzero hx) := by simp only [toNNReal, MonoidWithZeroHom.coe_mk, ZeroHom.coe_mk] - split_ifs; tauto; rfl + split_ifs + · tauto + · rfl /-- `toNNReal` sends nonzero elements to nonzero elements. -/ theorem toNNReal_ne_zero {e : NNReal} {m : ℤₘ₀} (he : e ≠ 0) (hm : m ≠ 0) : toNNReal he m ≠ 0 := by diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 9c2f87605db92..d3d29ac80d33d 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1358,7 +1358,7 @@ theorem intersperse_cons_cons (a b c : α) (tl : List α) : section SplitAtOn -variable (p : α → Bool) (xs ys : List α) (ls : List (List α)) (f : List α → List α) +variable (p : α → Bool) (xs : List α) (ls : List (List α)) attribute [simp] splitAt_eq @@ -1535,8 +1535,6 @@ theorem sizeOf_lt_sizeOf_of_mem [SizeOf α] {x : α} {l : List α} (hx : x ∈ l section find? -variable {p : α → Bool} {l : List α} {a : α} - @[deprecated (since := "2024-05-05")] alias find?_mem := mem_of_find?_eq_some end find? @@ -2179,18 +2177,13 @@ end Forall theorem get_attach (L : List α) (i) : (L.attach.get i).1 = L.get ⟨i, length_attach (L := L) ▸ i.2⟩ := by simp -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ @[simp 1100] theorem mem_map_swap (x : α) (y : β) (xs : List (α × β)) : (y, x) ∈ map Prod.swap xs ↔ (x, y) ∈ xs := by induction' xs with x xs xs_ih · simp only [not_mem_nil, map_nil] · cases' x with a b - simp only [mem_cons, Prod.mk.inj_iff, map, Prod.swap_prod_mk, Prod.exists, xs_ih, - _root_.and_comm] + simp only [mem_cons, Prod.mk.inj_iff, map, Prod.swap_prod_mk, Prod.exists, xs_ih, and_comm] theorem dropSlice_eq (xs : List α) (n m : ℕ) : dropSlice n m xs = xs.take n ++ xs.drop (n + m) := by induction n generalizing xs @@ -2278,8 +2271,7 @@ theorem disjoint_reverse_right {l₁ l₂ : List α} : Disjoint l₁ l₂.revers end Disjoint section lookup - -variable {α β : Type*} [BEq α] [LawfulBEq α] +variable [BEq α] [LawfulBEq α] lemma lookup_graph (f : α → β) {a : α} {as : List α} (h : a ∈ as) : lookup a (as.map fun x => (x, f x)) = some (f a) := by diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index 1213160f03d0a..965aee596b838 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -47,10 +47,6 @@ theorem Chain.iff_mem {a : α} {l : List α} : theorem chain_singleton {a b : α} : Chain R a [b] ↔ R a b := by simp only [chain_cons, Chain.nil, and_true] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem chain_split {a b : α} {l₁ l₂ : List α} : Chain R a (l₁ ++ b :: l₂) ↔ Chain R a (l₁ ++ [b]) ∧ Chain R b l₂ := by induction' l₁ with x l₁ IH generalizing a <;> @@ -142,6 +138,19 @@ theorem chain_iff_get {R} : ∀ {a : α} {l : List α}, Chain R a l ↔ intro i w exact h (i+1) (by simp only [length_cons]; omega) +theorem chain_replicate_of_rel (n : ℕ) {a : α} (h : r a a) : Chain r a (replicate n a) := + match n with + | 0 => Chain.nil + | n + 1 => Chain.cons h (chain_replicate_of_rel n h) + +theorem chain_eq_iff_eq_replicate {a : α} {l : List α} : + Chain (· = ·) a l ↔ l = replicate l.length a := + match l with + | [] => by simp + | b :: l => by + rw [chain_cons] + simp (config := {contextual := true}) [eq_comm, replicate_succ, chain_eq_iff_eq_replicate] + theorem Chain'.imp {S : α → α → Prop} (H : ∀ a b, R a b → S a b) {l : List α} (p : Chain' R l) : Chain' S l := by cases l <;> [trivial; exact Chain.imp H p] @@ -230,18 +239,13 @@ theorem Chain'.cons' {x} : ∀ {l : List α}, Chain' R l → (∀ y ∈ l.head?, theorem chain'_cons' {x l} : Chain' R (x :: l) ↔ (∀ y ∈ head? l, R x y) ∧ Chain' R l := ⟨fun h => ⟨h.rel_head?, h.tail⟩, fun ⟨h₁, h₂⟩ => h₂.cons' h₁⟩ -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefixes below. --/ theorem chain'_append : ∀ {l₁ l₂ : List α}, Chain' R (l₁ ++ l₂) ↔ Chain' R l₁ ∧ Chain' R l₂ ∧ ∀ x ∈ l₁.getLast?, ∀ y ∈ l₂.head?, R x y | [], l => by simp - | [a], l => by simp [chain'_cons', _root_.and_comm] + | [a], l => by simp [chain'_cons', and_comm] | a :: b :: l₁, l₂ => by - rw [cons_append, cons_append, chain'_cons, chain'_cons, ← cons_append, chain'_append, - _root_.and_assoc] + rw [cons_append, cons_append, chain'_cons, chain'_cons, ← cons_append, chain'_append, and_assoc] simp theorem Chain'.append (h₁ : Chain' R l₁) (h₂ : Chain' R l₂) @@ -280,16 +284,12 @@ theorem Chain'.imp_head {x y} (h : ∀ {z}, R x z → R y z) {l} (hl : Chain' R Chain' R (y :: l) := hl.tail.cons' fun _ hz => h <| hl.rel_head? hz -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem chain'_reverse : ∀ {l}, Chain' R (reverse l) ↔ Chain' (flip R) l | [] => Iff.rfl | [a] => by simp only [chain'_singleton, reverse_singleton] | a :: b :: l => by rw [chain'_cons, reverse_cons, reverse_cons, append_assoc, cons_append, nil_append, - chain'_split, ← reverse_cons, @chain'_reverse (b :: l), _root_.and_comm, chain'_pair, flip] + chain'_split, ← reverse_cons, @chain'_reverse (b :: l), and_comm, chain'_pair, flip] theorem chain'_iff_get {R} : ∀ {l : List α}, Chain' R l ↔ ∀ (i : ℕ) (h : i < length l - 1), @@ -307,10 +307,6 @@ theorem Chain'.append_overlap {l₁ l₂ l₃ : List α} (h₁ : Chain' R (l₁ h₁.append h₂.right_of_append <| by simpa only [getLast?_append_of_ne_nil _ hn] using (chain'_append.1 h₂).2.2 -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ lemma chain'_join : ∀ {L : List (List α)}, [] ∉ L → (Chain' R L.join ↔ (∀ l ∈ L, Chain' R l) ∧ L.Chain' (fun l₁ l₂ => ∀ᵉ (x ∈ l₁.getLast?) (y ∈ l₂.head?), R x y)) @@ -320,7 +316,7 @@ lemma chain'_join : ∀ {L : List (List α)}, [] ∉ L → rw [mem_cons, not_or, ← Ne] at hL rw [join, chain'_append, chain'_join hL.2, forall_mem_cons, chain'_cons] rw [mem_cons, not_or, ← Ne] at hL - simp only [forall_mem_cons, _root_.and_assoc, join, head?_append_of_ne_nil _ hL.2.1.symm] + simp only [forall_mem_cons, and_assoc, join, head?_append_of_ne_nil _ hL.2.1.symm] exact Iff.rfl.and (Iff.rfl.and <| Iff.rfl.and and_comm) /-- If `a` and `b` are related by the reflexive transitive closure of `r`, then there is an @@ -435,6 +431,17 @@ lemma Chain'.iterate_eq_of_apply_eq {α : Type*} {f : α → α} {l : List α} apply hl omega +theorem chain'_replicate_of_rel (n : ℕ) {a : α} (h : r a a) : Chain' r (replicate n a) := + match n with + | 0 => chain'_nil + | n + 1 => chain_replicate_of_rel n h + +theorem chain'_eq_iff_eq_replicate {l : List α} : + Chain' (· = ·) l ↔ ∀ a ∈ l.head?, l = replicate l.length a := + match l with + | [] => by simp + | a :: l => by simp [Chain', chain_eq_iff_eq_replicate, replicate_succ] + end List diff --git a/Mathlib/Data/List/Count.lean b/Mathlib/Data/List/Count.lean index 1b2789649d795..89475cec9fe1e 100644 --- a/Mathlib/Data/List/Count.lean +++ b/Mathlib/Data/List/Count.lean @@ -20,7 +20,7 @@ assert_not_exists Ring open Nat -variable {α : Type*} {l : List α} +variable {α : Type*} namespace List diff --git a/Mathlib/Data/List/Enum.lean b/Mathlib/Data/List/Enum.lean index d4b73fc6bec78..a574ea756a233 100644 --- a/Mathlib/Data/List/Enum.lean +++ b/Mathlib/Data/List/Enum.lean @@ -19,7 +19,7 @@ Any downstream users who can not easily adapt may remove the deprecations as nee namespace List -variable {α β : Type*} +variable {α : Type*} @[deprecated getElem?_enumFrom (since := "2024-08-15")] theorem get?_enumFrom (n) (l : List α) (m) : diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index 7e655efa8abee..b7bf8b65b8356 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -25,11 +25,11 @@ All those (except `insert`) are defined in `Mathlib.Data.List.Defs`. * `l₁ <:+: l₂`: `l₁` is an infix of `l₂`. -/ -variable {α β : Type*} +variable {α : Type*} namespace List -variable {l l₁ l₂ l₃ : List α} {a b : α} {m n : ℕ} +variable {l l₁ l₂ : List α} {a b : α} /-! ### prefix, suffix, infix -/ @@ -78,6 +78,14 @@ theorem concat_get_prefix {x y : List α} (h : x <+: y) (hl : x.length < y.lengt convert List.take_append_drop (x.length + 1) y using 2 rw [← List.take_concat_get, List.concat_eq_append]; rfl +instance decidableInfix [DecidableEq α] : ∀ l₁ l₂ : List α, Decidable (l₁ <:+: l₂) + | [], l₂ => isTrue ⟨[], l₂, rfl⟩ + | a :: l₁, [] => isFalse fun ⟨s, t, te⟩ => by simp at te + | l₁, b :: l₂ => + letI := l₁.decidableInfix l₂ + @decidable_of_decidable_of_iff (l₁ <+: b :: l₂ ∨ l₁ <:+: l₂) _ _ + infix_cons_iff.symm + @[deprecated cons_prefix_cons (since := "2024-08-14")] theorem cons_prefix_iff : a :: l₁ <+: b :: l₂ ↔ a = b ∧ l₁ <+: l₂ := by simp diff --git a/Mathlib/Data/List/InsertNth.lean b/Mathlib/Data/List/InsertNth.lean index b6b147595d8d8..00cef198f31e7 100644 --- a/Mathlib/Data/List/InsertNth.lean +++ b/Mathlib/Data/List/InsertNth.lean @@ -19,9 +19,9 @@ open Nat hiding one_pos namespace List -universe u v w +universe u -variable {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {l₁ l₂ : List α} +variable {α : Type u} section InsertNth diff --git a/Mathlib/Data/List/Lattice.lean b/Mathlib/Data/List/Lattice.lean index 75dc0c1e112bf..5a0ae92ccc6c0 100644 --- a/Mathlib/Data/List/Lattice.lean +++ b/Mathlib/Data/List/Lattice.lean @@ -29,7 +29,7 @@ open Nat namespace List -variable {α : Type*} {l l₁ l₂ : List α} {p : α → Prop} {a : α} +variable {α : Type*} {l₁ l₂ : List α} {p : α → Prop} {a : α} /-! ### `Disjoint` -/ diff --git a/Mathlib/Data/List/Monad.lean b/Mathlib/Data/List/Monad.lean index a89b5f05d5272..04d1477cbd14d 100644 --- a/Mathlib/Data/List/Monad.lean +++ b/Mathlib/Data/List/Monad.lean @@ -9,11 +9,11 @@ import Mathlib.Init # Monad instances for `List` -/ -universe u v w +universe u namespace List -variable {α : Type u} {β : Type v} {γ : Type w} +variable {α : Type u} instance instMonad : Monad List.{u} where pure := @List.pure diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index 9d86abe857e93..f1a4e94793afc 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -18,7 +18,7 @@ universe u v open Function -variable {α : Type u} {β : Type v} {l l₁ l₂ : List α} {r : α → α → Prop} {a b : α} +variable {α : Type u} {β : Type v} {l l₁ l₂ : List α} {r : α → α → Prop} {a : α} namespace List @@ -102,17 +102,13 @@ theorem nodup_iff_get?_ne_get? {l : List α} : l.Nodup ↔ ∀ i j : ℕ, i < j → j < l.length → l.get? i ≠ l.get? j := by simp [nodup_iff_getElem?_ne_getElem?] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem Nodup.ne_singleton_iff {l : List α} (h : Nodup l) (x : α) : l ≠ [x] ↔ l = [] ∨ ∃ y ∈ l, y ≠ x := by induction' l with hd tl hl · simp · specialize hl h.of_cons by_cases hx : tl = [x] - · simpa [hx, _root_.and_comm, and_or_left] using h + · simpa [hx, and_comm, and_or_left] using h · rw [← Ne, hl] at hx rcases hx with (rfl | ⟨y, hy, hx⟩) · simp @@ -178,13 +174,9 @@ theorem Nodup.append (d₁ : Nodup l₁) (d₂ : Nodup l₂) (dj : Disjoint l₁ theorem nodup_append_comm {l₁ l₂ : List α} : Nodup (l₁ ++ l₂) ↔ Nodup (l₂ ++ l₁) := by simp only [nodup_append, and_left_comm, disjoint_comm] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem nodup_middle {a : α} {l₁ l₂ : List α} : Nodup (l₁ ++ a :: l₂) ↔ Nodup (a :: (l₁ ++ l₂)) := by - simp only [nodup_append, not_or, and_left_comm, _root_.and_assoc, nodup_cons, mem_append, + simp only [nodup_append, not_or, and_left_comm, and_assoc, nodup_cons, mem_append, disjoint_cons_right] theorem Nodup.of_map (f : α → β) {l : List α} : Nodup (map f l) → Nodup l := @@ -267,14 +259,10 @@ theorem nodup_join {L : List (List α)} : Nodup (join L) ↔ (∀ l ∈ L, Nodup l) ∧ Pairwise Disjoint L := by simp only [Nodup, pairwise_join, disjoint_left.symm, forall_mem_ne] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem nodup_bind {l₁ : List α} {f : α → List β} : Nodup (l₁.bind f) ↔ (∀ x ∈ l₁, Nodup (f x)) ∧ Pairwise (fun a b : α => Disjoint (f a) (f b)) l₁ := by - simp only [List.bind, nodup_join, pairwise_map, _root_.and_comm, and_left_comm, mem_map, + simp only [List.bind, nodup_join, pairwise_map, and_comm, and_left_comm, mem_map, exists_imp, and_imp] rw [show (∀ (l : List β) (x : α), f x = l → x ∈ l₁ → Nodup l) ↔ ∀ x : α, x ∈ l₁ → Nodup (f x) from forall_swap.trans <| forall_congr' fun _ => forall_eq'] diff --git a/Mathlib/Data/List/Pairwise.lean b/Mathlib/Data/List/Pairwise.lean index 1a562e854c8c7..507a0a3195b58 100644 --- a/Mathlib/Data/List/Pairwise.lean +++ b/Mathlib/Data/List/Pairwise.lean @@ -29,7 +29,7 @@ open Nat Function namespace List -variable {α β : Type*} {R S T : α → α → Prop} {a : α} {l : List α} +variable {α β : Type*} {R : α → α → Prop} {l : List α} mk_iff_of_inductive_prop List.Pairwise List.pairwise_iff @@ -69,9 +69,6 @@ theorem pairwise_of_reflexive_of_forall_ne {l : List α} {r : α → α → Prop /-! ### Pairwise filtering -/ - -variable [DecidableRel R] - alias ⟨_, Pairwise.pwFilter⟩ := pwFilter_eq_self -- Porting note: commented out diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index f0e74dcc0d2b3..14611136f55f6 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -246,17 +246,13 @@ theorem Perm.bagInter {l₁ l₂ t₁ t₂ : List α} (hl : l₁ ~ l₂) (ht : t l₁.bagInter t₁ ~ l₂.bagInter t₂ := ht.bagInter_left l₂ ▸ hl.bagInter_right _ -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h : a ≠ b) : l ~ replicate m a ++ replicate n b ↔ count a l = m ∧ count b l = n ∧ l ⊆ [a, b] := by rw [perm_iff_count, ← Decidable.and_forall_ne a, ← Decidable.and_forall_ne b] suffices l ⊆ [a, b] ↔ ∀ c, c ≠ b → c ≠ a → c ∉ l by simp (config := { contextual := true }) [count_replicate, h, this, count_eq_zero, Ne.symm] trans ∀ c, c ∈ l → c = b ∨ c = a - · simp [subset_def, _root_.or_comm] + · simp [subset_def, or_comm] · exact forall_congr' fun _ => by rw [← and_imp, ← not_or, not_imp_not] theorem Perm.dedup {l₁ l₂ : List α} (p : l₁ ~ l₂) : dedup l₁ ~ dedup l₂ := diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 37fb13a041412..dfa4674956933 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -166,10 +166,6 @@ theorem foldr_permutationsAux2 (t : α) (ts : List α) (r L : List (List α)) : · rfl · simp_rw [foldr_cons, ih, bind_cons, append_assoc, permutationsAux2_append] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem mem_foldr_permutationsAux2 {t : α} {ts : List α} {r L : List (List α)} {l' : List α} : l' ∈ foldr (fun y r => (permutationsAux2 t ts r y id).2) r L ↔ l' ∈ r ∨ ∃ l₁ l₂, l₁ ++ l₂ ∈ L ∧ l₂ ≠ [] ∧ l' = l₁ ++ t :: l₂ ++ ts := by @@ -180,7 +176,7 @@ theorem mem_foldr_permutationsAux2 {t : α} {ts : List α} {r L : List (List α) ⟨fun ⟨_, aL, l₁, l₂, l0, e, h⟩ => ⟨l₁, l₂, l0, e ▸ aL, h⟩, fun ⟨l₁, l₂, l0, aL, h⟩ => ⟨_, aL, l₁, l₂, l0, rfl, h⟩⟩ rw [foldr_permutationsAux2] - simp only [mem_permutationsAux2', ← this, _root_.or_comm, and_left_comm, mem_append, mem_bind, + simp only [mem_permutationsAux2', ← this, or_comm, and_left_comm, mem_append, mem_bind, append_assoc, cons_append, exists_prop] theorem length_foldr_permutationsAux2 (t : α) (ts : List α) (r L : List (List α)) : diff --git a/Mathlib/Data/List/Rotate.lean b/Mathlib/Data/List/Rotate.lean index 923a8ba69b0ed..04072b6a698df 100644 --- a/Mathlib/Data/List/Rotate.lean +++ b/Mathlib/Data/List/Rotate.lean @@ -104,15 +104,11 @@ theorem rotate_cons_succ (l : List α) (a : α) (n : ℕ) : (a :: l : List α).rotate (n + 1) = (l ++ [a]).rotate n := by rw [rotate_eq_rotate', rotate_eq_rotate', rotate'_cons_succ] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ @[simp] theorem mem_rotate : ∀ {l : List α} {a : α} {n : ℕ}, a ∈ l.rotate n ↔ a ∈ l | [], _, n => by simp | a :: l, _, 0 => by simp - | a :: l, _, n + 1 => by simp [rotate_cons_succ, mem_rotate, _root_.or_comm] + | a :: l, _, n + 1 => by simp [rotate_cons_succ, mem_rotate, or_comm] @[simp] theorem length_rotate (l : List α) (n : ℕ) : (l.rotate n).length = l.length := by diff --git a/Mathlib/Data/Matrix/DMatrix.lean b/Mathlib/Data/Matrix/DMatrix.lean index d677b3ace646e..2b482b6e0ee61 100644 --- a/Mathlib/Data/Matrix/DMatrix.lean +++ b/Mathlib/Data/Matrix/DMatrix.lean @@ -20,7 +20,7 @@ In most applications `m` and `n` are finite types. -/ def DMatrix (m : Type u) (n : Type u') (α : m → n → Type v) : Type max u u' v := ∀ i j, α i j -variable {l m n o : Type*} +variable {m n : Type*} variable {α : m → n → Type v} namespace DMatrix diff --git a/Mathlib/Data/Matrix/DoublyStochastic.lean b/Mathlib/Data/Matrix/DoublyStochastic.lean new file mode 100644 index 0000000000000..8bc91f3996116 --- /dev/null +++ b/Mathlib/Data/Matrix/DoublyStochastic.lean @@ -0,0 +1,139 @@ +/- +Copyright (c) 2024 Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Bhavik Mehta +-/ + +import Mathlib.Analysis.Convex.Basic +import Mathlib.LinearAlgebra.Matrix.Permutation + +/-! +# Doubly stochastic matrices + +## Main definitions + +* `doublyStochastic`: a square matrix is doubly stochastic if all entries are nonnegative, and left + or right multiplication by the vector of all 1s gives the vector of all 1s. Equivalently, all + row and column sums are equal to 1. + +## Main statements + +* `convex_doublyStochastic`: The set of doubly stochastic matrices is convex. +* `permMatrix_mem_doublyStochastic`: Any permutation matrix is doubly stochastic. + +## TODO + +Define the submonoids of row-stochastic and column-stochastic matrices. +Show that the submonoid of doubly stochastic matrices is the meet of them, or redefine it as such. + +## Tags + +Doubly stochastic, Birkhoff's theorem, Birkhoff-von Neumann theorem +-/ + +open Finset Function Matrix + +variable {R n : Type*} [Fintype n] [DecidableEq n] + +section OrderedSemiring +variable [OrderedSemiring R] {M : Matrix n n R} + +/-- +A square matrix is doubly stochastic iff all entries are nonnegative, and left or right +multiplication by the vector of all 1s gives the vector of all 1s. +-/ +def doublyStochastic (R n : Type*) [Fintype n] [DecidableEq n] [OrderedSemiring R] : + Submonoid (Matrix n n R) where + carrier := {M | (∀ i j, 0 ≤ M i j) ∧ M *ᵥ 1 = 1 ∧ 1 ᵥ* M = 1 } + mul_mem' {M N} hM hN := by + refine ⟨fun i j => sum_nonneg fun i _ => mul_nonneg (hM.1 _ _) (hN.1 _ _), ?_, ?_⟩ + next => rw [← mulVec_mulVec, hN.2.1, hM.2.1] + next => rw [← vecMul_vecMul, hM.2.2, hN.2.2] + one_mem' := by simp [zero_le_one_elem] + +lemma mem_doublyStochastic : + M ∈ doublyStochastic R n ↔ (∀ i j, 0 ≤ M i j) ∧ M *ᵥ 1 = 1 ∧ 1 ᵥ* M = 1 := + Iff.rfl + +lemma mem_doublyStochastic_iff_sum : + M ∈ doublyStochastic R n ↔ + (∀ i j, 0 ≤ M i j) ∧ (∀ i, ∑ j, M i j = 1) ∧ ∀ j, ∑ i, M i j = 1 := by + simp [funext_iff, doublyStochastic, mulVec, vecMul, dotProduct] + +/-- Every entry of a doubly stochastic matrix is nonnegative. -/ +lemma nonneg_of_mem_doublyStochastic (hM : M ∈ doublyStochastic R n) {i j : n} : 0 ≤ M i j := + hM.1 _ _ + +/-- Each row sum of a doubly stochastic matrix is 1. -/ +lemma sum_row_of_mem_doublyStochastic (hM : M ∈ doublyStochastic R n) (i : n) : ∑ j, M i j = 1 := + (mem_doublyStochastic_iff_sum.1 hM).2.1 _ + +/-- Each column sum of a doubly stochastic matrix is 1. -/ +lemma sum_col_of_mem_doublyStochastic (hM : M ∈ doublyStochastic R n) (j : n) : ∑ i, M i j = 1 := + (mem_doublyStochastic_iff_sum.1 hM).2.2 _ + +/-- A doubly stochastic matrix multiplied with the all-ones column vector is 1. -/ +lemma mulVec_one_of_mem_doublyStochastic (hM : M ∈ doublyStochastic R n) : M *ᵥ 1 = 1 := + (mem_doublyStochastic.1 hM).2.1 + +/-- The all-ones row vector multiplied with a doubly stochastic matrix is 1. -/ +lemma one_vecMul_of_mem_doublyStochastic (hM : M ∈ doublyStochastic R n) : 1 ᵥ* M = 1 := + (mem_doublyStochastic.1 hM).2.2 + +/-- Every entry of a doubly stochastic matrix is less than or equal to 1. -/ +lemma le_one_of_mem_doublyStochastic (hM : M ∈ doublyStochastic R n) {i j : n} : + M i j ≤ 1 := by + rw [← sum_row_of_mem_doublyStochastic hM i] + exact single_le_sum (fun k _ => hM.1 _ k) (mem_univ j) + +/-- The set of doubly stochastic matrices is convex. -/ +lemma convex_doublyStochastic : Convex R (doublyStochastic R n : Set (Matrix n n R)) := by + intro x hx y hy a b ha hb h + simp only [SetLike.mem_coe, mem_doublyStochastic_iff_sum] at hx hy ⊢ + simp [add_nonneg, ha, hb, mul_nonneg, hx, hy, sum_add_distrib, ← mul_sum, h] + +/-- Any permutation matrix is doubly stochastic. -/ +lemma permMatrix_mem_doublyStochastic {σ : Equiv.Perm n} : + σ.permMatrix R ∈ doublyStochastic R n := by + rw [mem_doublyStochastic_iff_sum] + refine ⟨fun i j => ?g1, ?g2, ?g3⟩ + case g1 => aesop + case g2 => simp [Equiv.toPEquiv_apply] + case g3 => simp [Equiv.toPEquiv_apply, ← Equiv.eq_symm_apply] + +end OrderedSemiring + +section LinearOrderedSemifield + +variable [LinearOrderedSemifield R] {M : Matrix n n R} + +/-- +A matrix is `s` times a doubly stochastic matrix iff all entries are nonnegative, and all row and +column sums are equal to `s`. + +This lemma is useful for the proof of Birkhoff's theorem - in particular because it allows scaling +by nonnegative factors rather than positive ones only. +-/ +lemma exists_mem_doublyStochastic_eq_smul_iff {M : Matrix n n R} {s : R} (hs : 0 ≤ s) : + (∃ M' ∈ doublyStochastic R n, M = s • M') ↔ + (∀ i j, 0 ≤ M i j) ∧ (∀ i, ∑ j, M i j = s) ∧ (∀ j, ∑ i, M i j = s) := by + classical + constructor + case mp => + rintro ⟨M', hM', rfl⟩ + rw [mem_doublyStochastic_iff_sum] at hM' + simp only [smul_apply, smul_eq_mul, ← mul_sum] + exact ⟨fun i j => mul_nonneg hs (hM'.1 _ _), by simp [hM']⟩ + rcases eq_or_lt_of_le hs with rfl | hs + case inl => + simp only [zero_smul, exists_and_right, and_imp] + intro h₁ h₂ _ + refine ⟨⟨1, Submonoid.one_mem _⟩, ?_⟩ + ext i j + specialize h₂ i + rw [sum_eq_zero_iff_of_nonneg (by simp [h₁ i])] at h₂ + exact h₂ _ (by simp) + rintro ⟨hM₁, hM₂, hM₃⟩ + exact ⟨s⁻¹ • M, by simp [mem_doublyStochastic_iff_sum, ← mul_sum, hs.ne', inv_mul_cancel₀, *]⟩ + +end LinearOrderedSemifield diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index c8a856b4d002c..15129dea51895 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1832,25 +1832,17 @@ theorem filter_filter (q) [DecidablePred q] (s : Multiset α) : filter p (filter q s) = filter (fun a => p a ∧ q a) s := Quot.inductionOn s fun l => by simp -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ lemma filter_comm (q) [DecidablePred q] (s : Multiset α) : - filter p (filter q s) = filter q (filter p s) := by simp [_root_.and_comm] + filter p (filter q s) = filter q (filter p s) := by simp [and_comm] theorem filter_add_filter (q) [DecidablePred q] (s : Multiset α) : filter p s + filter q s = filter (fun a => p a ∨ q a) s + filter (fun a => p a ∧ q a) s := Multiset.induction_on s rfl fun a s IH => by by_cases p a <;> by_cases q a <;> simp [*] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem filter_add_not (s : Multiset α) : filter p s + filter (fun a => ¬p a) s = s := by rw [filter_add_filter, filter_eq_self.2, filter_eq_nil.2] · simp only [add_zero] - · simp [Decidable.em, -Bool.not_eq_true, -not_and, not_and_or, _root_.or_comm] + · simp [Decidable.em, -Bool.not_eq_true, -not_and, not_and_or, or_comm] · simp only [Bool.not_eq_true, decide_eq_true_eq, Bool.eq_false_or_eq_true, decide_True, implies_true, Decidable.em] @@ -2690,13 +2682,9 @@ lemma add_eq_union_left_of_le [DecidableEq α] {s t u : Multiset α} (h : t ≤ · rintro ⟨h0, rfl⟩ exact h0 -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ lemma add_eq_union_right_of_le [DecidableEq α] {x y z : Multiset α} (h : z ≤ y) : x + y = x ∪ z ↔ y = z ∧ x.Disjoint y := by - simpa only [_root_.and_comm] using add_eq_union_left_of_le h + simpa only [and_comm] using add_eq_union_left_of_le h theorem disjoint_map_map {f : α → γ} {g : β → γ} {s : Multiset α} {t : Multiset β} : Disjoint (s.map f) (t.map g) ↔ ∀ a ∈ s, ∀ b ∈ t, f a ≠ g b := by diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index b82787ebd3a91..6679702f69c42 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -907,7 +907,7 @@ theorem zpow_pos {x : ℝ≥0} (hx : x ≠ 0) (n : ℤ) : 0 < x ^ n := zpow_pos_of_pos hx.bot_lt _ theorem inv_lt_inv {x y : ℝ≥0} (hx : x ≠ 0) (h : x < y) : y⁻¹ < x⁻¹ := - inv_lt_inv_of_lt hx.bot_lt h + inv_strictAnti₀ hx.bot_lt h end Inv diff --git a/Mathlib/Data/Nat/Cast/Commute.lean b/Mathlib/Data/Nat/Cast/Commute.lean index 4d8306384d75f..9d1e9bc89094b 100644 --- a/Mathlib/Data/Nat/Cast/Commute.lean +++ b/Mathlib/Data/Nat/Cast/Commute.lean @@ -11,7 +11,7 @@ import Mathlib.Algebra.Ring.Commute -/ -variable {α β : Type*} +variable {α : Type*} namespace Nat diff --git a/Mathlib/Data/Nat/Cast/Order/Field.lean b/Mathlib/Data/Nat/Cast/Order/Field.lean index 8b16df559d51a..bf5cfd6d2e238 100644 --- a/Mathlib/Data/Nat/Cast/Order/Field.lean +++ b/Mathlib/Data/Nat/Cast/Order/Field.lean @@ -22,7 +22,7 @@ variable {α : Type*} [LinearOrderedSemifield α] lemma cast_inv_le_one : ∀ n : ℕ, (n⁻¹ : α) ≤ 1 | 0 => by simp - | n + 1 => inv_le_one <| by simp [Nat.cast_nonneg] + | n + 1 => inv_le_one_of_one_le₀ <| by simp [Nat.cast_nonneg] /-- Natural division is always less than division in the field. -/ theorem cast_div_le {m n : ℕ} : ((m / n : ℕ) : α) ≤ m / n := by diff --git a/Mathlib/Data/Nat/Cast/Synonym.lean b/Mathlib/Data/Nat/Cast/Synonym.lean index fad1703cc9e63..075d62d0d87cf 100644 --- a/Mathlib/Data/Nat/Cast/Synonym.lean +++ b/Mathlib/Data/Nat/Cast/Synonym.lean @@ -22,7 +22,7 @@ the natural numbers into an additive monoid with a one (`Nat.cast`). -- where `simp [map_zero]` should suffice. (Similarly for `map_one`.) -- See https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/simp.20regression.20with.20MonoidHomClass -variable {α β : Type*} +variable {α : Type*} /-! ### Order dual -/ diff --git a/Mathlib/Data/Nat/Choose/Factorization.lean b/Mathlib/Data/Nat/Choose/Factorization.lean index 71966a049eb3a..258c4eb897817 100644 --- a/Mathlib/Data/Nat/Choose/Factorization.lean +++ b/Mathlib/Data/Nat/Choose/Factorization.lean @@ -78,7 +78,7 @@ theorem factorization_choose_of_lt_three_mul (hp' : p ≠ 2) (hk : p ≤ k) (hk' n < 3 * p := hn _ ≤ p * p := mul_le_mul_right' this p _ = p ^ 2 := (sq p).symm - _ ≤ p ^ i := pow_le_pow_right hp.one_lt.le hi + _ ≤ p ^ i := pow_right_mono₀ hp.one_lt.le hi rwa [mod_eq_of_lt (lt_of_le_of_lt hkn hn), mod_eq_of_lt (lt_of_le_of_lt tsub_le_self hn), add_tsub_cancel_of_le hkn] diff --git a/Mathlib/Data/Nat/Choose/Sum.lean b/Mathlib/Data/Nat/Choose/Sum.lean index f339da88c6daf..d15893af1129d 100644 --- a/Mathlib/Data/Nat/Choose/Sum.lean +++ b/Mathlib/Data/Nat/Choose/Sum.lean @@ -192,7 +192,7 @@ theorem sum_powerset_neg_one_pow_card_of_nonempty {α : Type*} {x : Finset α} ( rw [sum_powerset_neg_one_pow_card] exact if_neg (nonempty_iff_ne_empty.mp h0) -variable {M R : Type*} [CommMonoid M] [NonAssocSemiring R] +variable [NonAssocSemiring R] @[to_additive sum_choose_succ_nsmul] theorem prod_pow_choose_succ {M : Type*} [CommMonoid M] (f : ℕ → ℕ → M) (n : ℕ) : diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index d7f4c31154a69..3bff75138a170 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -57,7 +57,7 @@ assert_not_exists Monoid open Function namespace Nat -variable {a b c d m n k : ℕ} {p q : ℕ → Prop} +variable {a b c d m n k : ℕ} {p : ℕ → Prop} -- TODO: Move the `LinearOrder ℕ` instance to `Order.Nat` (#13092). instance instLinearOrder : LinearOrder ℕ where @@ -223,8 +223,7 @@ attribute [simp] le_add_left le_add_right Nat.lt_add_left_iff_pos Nat.lt_add_rig -- 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. --- TODO: This is a duplicate of `Nat.add_eq` -@[simp] lemma add_def : Nat.add m n = m + n := rfl +@[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 diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 1285453521a3b..27ddcf0d3a424 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -455,10 +455,6 @@ theorem setOf_pow_dvd_eq_Icc_factorization {n p : ℕ} (pp : p.Prime) (hn : n ext simp [Nat.lt_succ_iff, one_le_iff_ne_zero, pp.pow_dvd_iff_le_factorization hn] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- The set of positive powers of prime `p` that divide `n` is exactly the set of positive natural numbers up to `n.factorization p`. -/ theorem Icc_factorization_eq_pow_dvd (n : ℕ) {p : ℕ} (pp : Prime p) : @@ -466,7 +462,7 @@ theorem Icc_factorization_eq_pow_dvd (n : ℕ) {p : ℕ} (pp : Prime p) : rcases eq_or_ne n 0 with (rfl | hn) · simp ext x - simp only [mem_Icc, Finset.mem_filter, mem_Ico, _root_.and_assoc, and_congr_right_iff, + simp only [mem_Icc, Finset.mem_filter, mem_Ico, and_assoc, and_congr_right_iff, pp.pow_dvd_iff_le_factorization hn, iff_and_self] exact fun _ H => lt_of_le_of_lt H (factorization_lt p hn) diff --git a/Mathlib/Data/Nat/Find.lean b/Mathlib/Data/Nat/Find.lean index e37f555306458..84ee3406a3e5a 100644 --- a/Mathlib/Data/Nat/Find.lean +++ b/Mathlib/Data/Nat/Find.lean @@ -11,7 +11,7 @@ import Batteries.WF # `Nat.find` and `Nat.findGreatest` -/ -variable {a b c d m n k : ℕ} {p q : ℕ → Prop} +variable {m n k : ℕ} {p q : ℕ → Prop} namespace Nat diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index 744b67ca07c10..f2cf3df6eca07 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -121,15 +121,11 @@ lemma log_le_self (b x : ℕ) : log b x ≤ x := theorem lt_pow_succ_log_self {b : ℕ} (hb : 1 < b) (x : ℕ) : x < b ^ (log b x).succ := lt_pow_of_log_lt hb (lt_succ_self _) -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem log_eq_iff {b m n : ℕ} (h : m ≠ 0 ∨ 1 < b ∧ n ≠ 0) : log b n = m ↔ b ^ m ≤ n ∧ n < b ^ (m + 1) := by rcases em (1 < b ∧ n ≠ 0) with (⟨hb, hn⟩ | hbn) · rw [le_antisymm_iff, ← Nat.lt_succ_iff, ← pow_le_iff_le_log, ← lt_pow_iff_log_lt, - _root_.and_comm] <;> assumption + and_comm] <;> assumption have hm : m ≠ 0 := h.resolve_right hbn rw [not_and_or, not_lt, Ne, not_not] at hbn rcases hbn with (hb | rfl) diff --git a/Mathlib/Data/Nat/Prime/Basic.lean b/Mathlib/Data/Nat/Prime/Basic.lean index d0fe2d6bad6ba..06a9ca4221721 100644 --- a/Mathlib/Data/Nat/Prime/Basic.lean +++ b/Mathlib/Data/Nat/Prime/Basic.lean @@ -287,37 +287,6 @@ lemma Prime.pow_inj {p q m n : ℕ} (hp : p.Prime) (hq : q.Prime) (Prime.dvd_of_dvd_pow hq <| h.symm ▸ dvd_pow_self q (succ_ne_zero n)) exact ⟨H, succ_inj'.mp <| Nat.pow_right_injective hq.two_le (H ▸ h)⟩ -theorem exists_pow_lt_factorial (c : ℕ) : ∃ n0 > 1, ∀ n ≥ n0, c ^ n < (n - 1)! := by - refine ⟨2 * (c ^ 2 + 1), ?_, ?_⟩ - · omega - intro n hn - obtain ⟨d, rfl⟩ := Nat.exists_eq_add_of_le hn - obtain (rfl | c0) := c.eq_zero_or_pos - · simp [Nat.factorial_pos] - refine (Nat.le_mul_of_pos_right _ (Nat.pow_pos (n := d) c0)).trans_lt ?_ - convert_to (c ^ 2) ^ (c ^ 2 + d + 1) < (c ^ 2 + (c ^ 2 + d + 1))! - · rw [← pow_mul, ← pow_add] - congr 1 - omega - · congr - omega - refine lt_of_lt_of_le ?_ Nat.factorial_mul_pow_le_factorial - rw [← one_mul (_ ^ _ : ℕ)] - exact Nat.mul_lt_mul_of_le_of_lt (Nat.one_le_of_lt (Nat.factorial_pos _)) - (Nat.pow_lt_pow_left (Nat.lt_succ_self _) (Nat.succ_ne_zero _)) (Nat.factorial_pos _) - -theorem exists_mul_pow_lt_factorial (a : ℕ) (c : ℕ) : ∃ n0, ∀ n ≥ n0, a * c ^ n < (n - 1)! := by - obtain ⟨n0, hn, h⟩ := Nat.exists_pow_lt_factorial (a * c) - refine ⟨n0, fun n hn => lt_of_le_of_lt ?_ (h n hn)⟩ - rw [mul_pow] - refine Nat.mul_le_mul_right _ (Nat.le_self_pow ?_ _) - omega - -theorem exists_prime_mul_pow_lt_factorial (n a c : ℕ) : ∃ p > n, p.Prime ∧ a * c ^ p < (p - 1)! := - have ⟨n0, h⟩ := Nat.exists_mul_pow_lt_factorial a c - have ⟨p, hp, prime_p⟩ := (max (n + 1) n0).exists_infinite_primes - ⟨p, (le_max_left _ _).trans hp, prime_p, h _ <| le_of_max_le_right hp⟩ - end Nat namespace Int diff --git a/Mathlib/Data/Nat/WithBot.lean b/Mathlib/Data/Nat/WithBot.lean index 418d1347bbc66..125eeae18ef2b 100644 --- a/Mathlib/Data/Nat/WithBot.lean +++ b/Mathlib/Data/Nat/WithBot.lean @@ -30,13 +30,9 @@ theorem add_eq_zero_iff {n m : WithBot ℕ} : n + m = 0 ↔ n = 0 ∧ m = 0 := b · simp [WithBot.add_bot] simp [← WithBot.coe_add, add_eq_zero_iff_of_nonneg] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem add_eq_one_iff {n m : WithBot ℕ} : n + m = 1 ↔ n = 0 ∧ m = 1 ∨ n = 1 ∧ m = 0 := by cases n - · simp only [WithBot.bot_add, WithBot.bot_ne_one, WithBot.bot_ne_zero, false_and, _root_.or_self] + · simp only [WithBot.bot_add, WithBot.bot_ne_one, WithBot.bot_ne_zero, false_and, or_self] cases m · simp [WithBot.add_bot] simp [← WithBot.coe_add, Nat.add_eq_one_iff] diff --git a/Mathlib/Data/Opposite.lean b/Mathlib/Data/Opposite.lean index 2482808fe2c64..c9f9e120603ca 100644 --- a/Mathlib/Data/Opposite.lean +++ b/Mathlib/Data/Opposite.lean @@ -30,7 +30,7 @@ variable (α : Sort u) both `unop (op X) = X` and `op (unop X) = X` are definitional equalities. -/ -structure Opposite := +structure Opposite where /-- The canonical map `α → αᵒᵖ`. -/ op :: /-- The canonical map `αᵒᵖ → α`. -/ diff --git a/Mathlib/Data/PFunctor/Univariate/M.lean b/Mathlib/Data/PFunctor/Univariate/M.lean index 699b8b76a816f..73fe251a4dda5 100644 --- a/Mathlib/Data/PFunctor/Univariate/M.lean +++ b/Mathlib/Data/PFunctor/Univariate/M.lean @@ -470,10 +470,6 @@ theorem corec_def {X} (f : X → F X) (x₀ : X) : M.corec f x₀ = M.mk (F.map dsimp only [PFunctor.map] congr -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem ext_aux [Inhabited (M F)] [DecidableEq F.A] {n : ℕ} (x y z : M F) (hx : Agree' n z x) (hy : Agree' n z y) (hrec : ∀ ps : Path F, n = ps.length → iselect ps x = iselect ps y) : x.approx (n + 1) = y.approx (n + 1) := by diff --git a/Mathlib/Data/PNat/Interval.lean b/Mathlib/Data/PNat/Interval.lean index 4a310ffe979ce..f10b5ea85500c 100644 --- a/Mathlib/Data/PNat/Interval.lean +++ b/Mathlib/Data/PNat/Interval.lean @@ -53,35 +53,19 @@ theorem map_subtype_embedding_uIcc : (uIcc a b).map (Embedding.subtype _) = uIcc @[simp] theorem card_Icc : (Icc a b).card = b + 1 - a := by - rw [← Nat.card_Icc] - -- Porting note: I had to change this to `erw` *and* provide the proof, yuck. - -- https://github.com/leanprover-community/mathlib4/issues/5164 - erw [← Finset.map_subtype_embedding_Icc _ a b (fun c x _ hx _ hc _ => hc.trans_le hx)] - rw [card_map] + rw [← Nat.card_Icc, ← map_subtype_embedding_Icc, card_map] @[simp] theorem card_Ico : (Ico a b).card = b - a := by - rw [← Nat.card_Ico] - -- Porting note: I had to change this to `erw` *and* provide the proof, yuck. - -- https://github.com/leanprover-community/mathlib4/issues/5164 - erw [← Finset.map_subtype_embedding_Ico _ a b (fun c x _ hx _ hc _ => hc.trans_le hx)] - rw [card_map] + rw [← Nat.card_Ico, ← map_subtype_embedding_Ico, card_map] @[simp] theorem card_Ioc : (Ioc a b).card = b - a := by - rw [← Nat.card_Ioc] - -- Porting note: I had to change this to `erw` *and* provide the proof, yuck. - -- https://github.com/leanprover-community/mathlib4/issues/5164 - erw [← Finset.map_subtype_embedding_Ioc _ a b (fun c x _ hx _ hc _ => hc.trans_le hx)] - rw [card_map] + rw [← Nat.card_Ioc, ← map_subtype_embedding_Ioc, card_map] @[simp] theorem card_Ioo : (Ioo a b).card = b - a - 1 := by - rw [← Nat.card_Ioo] - -- Porting note: I had to change this to `erw` *and* provide the proof, yuck. - -- https://github.com/leanprover-community/mathlib4/issues/5164 - erw [← Finset.map_subtype_embedding_Ioo _ a b (fun c x _ hx _ hc _ => hc.trans_le hx)] - rw [card_map] + rw [← Nat.card_Ioo, ← map_subtype_embedding_Ioo, card_map] @[simp] theorem card_uIcc : (uIcc a b).card = (b - a : ℤ).natAbs + 1 := by diff --git a/Mathlib/Data/Prod/Lex.lean b/Mathlib/Data/Prod/Lex.lean index f8116316921ee..b0271fecee021 100644 --- a/Mathlib/Data/Prod/Lex.lean +++ b/Mathlib/Data/Prod/Lex.lean @@ -29,17 +29,13 @@ Related files are: -/ -variable {α β γ : Type*} +variable {α β : Type*} namespace Prod.Lex -@[inherit_doc] notation:35 α " ×ₗ " β:34 => Lex (Prod α β) - -instance decidableEq (α β : Type*) [DecidableEq α] [DecidableEq β] : DecidableEq (α ×ₗ β) := - instDecidableEqProd +open Batteries -instance inhabited (α β : Type*) [Inhabited α] [Inhabited β] : Inhabited (α ×ₗ β) := - instInhabitedProd +@[inherit_doc] notation:35 α " ×ₗ " β:34 => Lex (Prod α β) /-- Dictionary / lexicographic ordering on pairs. -/ instance instLE (α β : Type*) [LT α] [LE β] : LE (α ×ₗ β) where le := Prod.Lex (· < ·) (· ≤ ·) @@ -124,16 +120,36 @@ instance partialOrder (α β : Type*) [PartialOrder α] [PartialOrder β] : Part haveI : IsAntisymm β (· ≤ ·) := ⟨fun _ _ => le_antisymm⟩ exact antisymm (r := Prod.Lex _ _) +instance instOrdLexProd [Ord α] [Ord β] : Ord (α ×ₗ β) := lexOrd + +theorem compare_def [Ord α] [Ord β] : @compare (α ×ₗ β) _ = + compareLex (compareOn fun x => (ofLex x).1) (compareOn fun x => (ofLex x).2) := rfl + +theorem _root_.lexOrd_eq [Ord α] [Ord β] : @lexOrd α β _ _ = instOrdLexProd := rfl + +theorem _root_.Ord.lex_eq [oα : Ord α] [oβ : Ord β] : Ord.lex oα oβ = instOrdLexProd := rfl + +instance [Ord α] [Ord β] [OrientedOrd α] [OrientedOrd β] : OrientedOrd (α ×ₗ β) := + inferInstanceAs (OrientedCmp (compareLex _ _)) + +instance [Ord α] [Ord β] [TransOrd α] [TransOrd β] : TransOrd (α ×ₗ β) := + inferInstanceAs (TransCmp (compareLex _ _)) + /-- Dictionary / lexicographic linear order for pairs. -/ instance linearOrder (α β : Type*) [LinearOrder α] [LinearOrder β] : LinearOrder (α ×ₗ β) := { Prod.Lex.partialOrder α β with - le_total := total_of (Prod.Lex _ _), - decidableLE := Prod.Lex.decidable _ _, - decidableLT := Prod.Lex.decidable _ _, - decidableEq := Lex.decidableEq _ _, } - -instance [Ord α] [Ord β] : Ord (α ×ₗ β) where - compare := compareLex (compareOn (·.1)) (compareOn (·.2)) + le_total := total_of (Prod.Lex _ _) + decidableLE := Prod.Lex.decidable _ _ + decidableLT := Prod.Lex.decidable _ _ + decidableEq := instDecidableEqLex _ + compare_eq_compareOfLessAndEq := fun a b => by + have : DecidableRel (· < · : α ×ₗ β → α ×ₗ β → Prop) := Prod.Lex.decidable _ _ + have : BEqOrd (α ×ₗ β) := ⟨by + simp [compare_def, compareLex, compareOn, Ordering.then_eq_eq, compare_eq_iff_eq]⟩ + have : LTOrd (α ×ₗ β) := ⟨by + simp [compare_def, compareLex, compareOn, Ordering.then_eq_lt, lt_iff, + compare_lt_iff_lt, compare_eq_iff_eq]⟩ + convert LTCmp.eq_compareOfLessAndEq (cmp := compare) a b } instance orderBot [PartialOrder α] [Preorder β] [OrderBot α] [OrderBot β] : OrderBot (α ×ₗ β) where bot := toLex ⊥ diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 6af294f0b9665..a0295e98e32df 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -9,8 +9,11 @@ import Mathlib.Util.Notation3 /-! # Quotient types + This module extends the core library's treatment of quotient types (`Init.Core`). + ## Tags + quotient -/ @@ -18,8 +21,15 @@ variable {α : Sort*} {β : Sort*} namespace Setoid -theorem ext {α : Sort*} : ∀ {s t : Setoid α}, - (∀ a b, @Setoid.r α s a b ↔ @Setoid.r α t a b) → s = t +-- Pretty print `@Setoid.r _ s a b` as `s a b`. +run_cmd Lean.Elab.Command.liftTermElabM do + Lean.Meta.registerCoercion ``Setoid.r + (some { numArgs := 2, coercee := 1, type := .coeFun }) + +instance : CoeFun (Setoid α) (fun _ ↦ α → α → Prop) where + coe := @Setoid.r _ + +theorem ext {α : Sort*} : ∀ {s t : Setoid α}, (∀ a b, s a b ↔ t a b) → s = t | ⟨r, _⟩, ⟨p, _⟩, Eq => by have : r = p := funext fun a ↦ funext fun b ↦ propext <| Eq a b subst this @@ -189,7 +199,7 @@ end Quot namespace Quotient -variable [sa : Setoid α] [sb : Setoid β] +variable {sa : Setoid α} {sb : Setoid β} variable {φ : Quotient sa → Quotient sb → Sort*} -- Porting note: in mathlib3 this notation took the Setoid as an instance-implicit argument, @@ -228,7 +238,7 @@ theorem map_mk (f : α → β) (h : ((· ≈ ·) ⇒ (· ≈ ·)) f f) (x : α) Quotient.map f h (⟦x⟧ : Quotient sa) = (⟦f x⟧ : Quotient sb) := rfl -variable {γ : Sort*} [sc : Setoid γ] +variable {γ : Sort*} {sc : Setoid γ} /-- Map a function `f : α → β → γ` that sends equivalent elements to equivalent elements to a function `f : Quotient sa → Quotient sb → Quotient sc`. @@ -269,7 +279,7 @@ theorem Quot.eq {α : Type*} {r : α → α → Prop} {x y : α} : ⟨Quot.eqvGen_exact, Quot.eqvGen_sound⟩ @[simp] -theorem Quotient.eq [r : Setoid α] {x y : α} : Quotient.mk r x = ⟦y⟧ ↔ x ≈ y := +theorem Quotient.eq {r : Setoid α} {x y : α} : Quotient.mk r x = ⟦y⟧ ↔ x ≈ y := ⟨Quotient.exact, Quotient.sound⟩ theorem Quotient.forall {α : Sort*} {s : Setoid α} {p : Quotient s → Prop} : @@ -281,29 +291,29 @@ theorem Quotient.exists {α : Sort*} {s : Setoid α} {p : Quotient s → Prop} : ⟨fun ⟨q, hq⟩ ↦ q.ind (motive := (p · → _)) .intro hq, fun ⟨a, ha⟩ ↦ ⟨⟦a⟧, ha⟩⟩ @[simp] -theorem Quotient.lift_mk [s : Setoid α] (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) (x : α) : +theorem Quotient.lift_mk {s : Setoid α} (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) (x : α) : Quotient.lift f h (Quotient.mk s x) = f x := rfl @[simp] -theorem Quotient.lift_comp_mk [Setoid α] (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) : +theorem Quotient.lift_comp_mk {_ : Setoid α} (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) : Quotient.lift f h ∘ Quotient.mk _ = f := rfl @[simp] -theorem Quotient.lift₂_mk {α : Sort*} {β : Sort*} {γ : Sort*} [Setoid α] [Setoid β] +theorem Quotient.lift₂_mk {α : Sort*} {β : Sort*} {γ : Sort*} {_ : Setoid α} {_ : Setoid β} (f : α → β → γ) (h : ∀ (a₁ : α) (a₂ : β) (b₁ : α) (b₂ : β), a₁ ≈ b₁ → a₂ ≈ b₂ → f a₁ a₂ = f b₁ b₂) (a : α) (b : β) : Quotient.lift₂ f h (Quotient.mk _ a) (Quotient.mk _ b) = f a b := rfl -theorem Quotient.liftOn_mk [s : Setoid α] (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) (x : α) : +theorem Quotient.liftOn_mk {s : Setoid α} (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) (x : α) : Quotient.liftOn (Quotient.mk s x) f h = f x := rfl @[simp] -theorem Quotient.liftOn₂_mk {α : Sort*} {β : Sort*} [Setoid α] (f : α → α → β) +theorem Quotient.liftOn₂_mk {α : Sort*} {β : Sort*} {_ : Setoid α} (f : α → α → β) (h : ∀ a₁ a₂ b₁ b₂ : α, a₁ ≈ b₁ → a₂ ≈ b₂ → f a₁ a₂ = f b₁ b₂) (x y : α) : Quotient.liftOn₂ (Quotient.mk _ x) (Quotient.mk _ y) f h = f x y := rfl @@ -338,22 +348,22 @@ theorem Quot.out_eq {r : α → α → Prop} (q : Quot r) : Quot.mk r q.out = q /-- Choose an element of the equivalence class using the axiom of choice. Sound but noncomputable. -/ -noncomputable def Quotient.out [s : Setoid α] : Quotient s → α := +noncomputable def Quotient.out {s : Setoid α} : Quotient s → α := Quot.out @[simp] -theorem Quotient.out_eq [s : Setoid α] (q : Quotient s) : ⟦q.out⟧ = q := +theorem Quotient.out_eq {s : Setoid α} (q : Quotient s) : ⟦q.out⟧ = q := Quot.out_eq q -theorem Quotient.mk_out [Setoid α] (a : α) : ⟦a⟧.out ≈ a := +theorem Quotient.mk_out {s : Setoid α} (a : α) : (⟦a⟧ : Quotient s).out ≈ a := Quotient.exact (Quotient.out_eq _) -theorem Quotient.mk_eq_iff_out [s : Setoid α] {x : α} {y : Quotient s} : +theorem Quotient.mk_eq_iff_out {s : Setoid α} {x : α} {y : Quotient s} : ⟦x⟧ = y ↔ x ≈ Quotient.out y := by refine Iff.trans ?_ Quotient.eq rw [Quotient.out_eq y] -theorem Quotient.eq_mk_iff_out [s : Setoid α] {x : Quotient s} {y : α} : +theorem Quotient.eq_mk_iff_out {s : Setoid α} {x : Quotient s} {y : α} : x = ⟦y⟧ ↔ Quotient.out x ≈ y := by refine Iff.trans ?_ Quotient.eq rw [Quotient.out_eq x] @@ -379,18 +389,18 @@ instance piSetoid {ι : Sort*} {α : ι → Sort*} [∀ i, Setoid (α i)] : Seto /-- Given a function `f : Π i, Quotient (S i)`, returns the class of functions `Π i, α i` sending each `i` to an element of the class `f i`. -/ -noncomputable def Quotient.choice {ι : Type*} {α : ι → Type*} [S : ∀ i, Setoid (α i)] +noncomputable def Quotient.choice {ι : Type*} {α : ι → Type*} {S : ∀ i, Setoid (α i)} (f : ∀ i, Quotient (S i)) : @Quotient (∀ i, α i) (by infer_instance) := ⟦fun i ↦ (f i).out⟧ @[simp] -theorem Quotient.choice_eq {ι : Type*} {α : ι → Type*} [∀ i, Setoid (α i)] (f : ∀ i, α i) : - (Quotient.choice fun i ↦ ⟦f i⟧) = ⟦f⟧ := +theorem Quotient.choice_eq {ι : Type*} {α : ι → Type*} {S : ∀ i, Setoid (α i)} (f : ∀ i, α i) : + (Quotient.choice (S := S) fun i ↦ ⟦f i⟧) = ⟦f⟧ := Quotient.sound fun _ ↦ Quotient.mk_out _ @[elab_as_elim] -theorem Quotient.induction_on_pi {ι : Type*} {α : ι → Sort*} [s : ∀ i, Setoid (α i)] +theorem Quotient.induction_on_pi {ι : Type*} {α : ι → Sort*} {s : ∀ i, Setoid (α i)} {p : (∀ i, Quotient (s i)) → Prop} (f : ∀ i, Quotient (s i)) (h : ∀ a : ∀ i, α i, p fun i ↦ ⟦a i⟧) : p f := by rw [← (funext fun i ↦ Quotient.out_eq (f i) : (fun i ↦ ⟦(f i).out⟧) = f)] @@ -689,7 +699,7 @@ theorem sound' {a b : α} : @Setoid.r _ s₁ a b → @Quotient.mk'' α s₁ a = Quotient.sound @[simp] -protected theorem eq' [s₁ : Setoid α] {a b : α} : +protected theorem eq' {s₁ : Setoid α} {a b : α} : @Quotient.mk' α s₁ a = @Quotient.mk' α s₁ b ↔ @Setoid.r _ s₁ a b := Quotient.eq @@ -711,7 +721,7 @@ theorem mk_out' (a : α) : @Setoid.r α s₁ (Quotient.mk'' a : Quotient s₁).o section -variable [s : Setoid α] +variable {s : Setoid α} protected theorem mk''_eq_mk : Quotient.mk'' = Quotient.mk s := rfl @@ -721,12 +731,12 @@ protected theorem liftOn'_mk (x : α) (f : α → β) (h) : (Quotient.mk s x).li rfl @[simp] -protected theorem liftOn₂'_mk [t : Setoid β] (f : α → β → γ) (h) (a : α) (b : β) : +protected theorem liftOn₂'_mk {t : Setoid β} (f : α → β → γ) (h) (a : α) (b : β) : Quotient.liftOn₂' (Quotient.mk s a) (Quotient.mk t b) f h = f a b := Quotient.liftOn₂'_mk'' _ _ _ _ @[simp] -theorem map'_mk [t : Setoid β] (f : α → β) (h) (x : α) : +theorem map'_mk {t : Setoid β} (f : α → β) (h) (x : α) : (Quotient.mk s x).map' f h = (Quotient.mk t (f x)) := rfl diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index f27de94d155e4..5c7543769551c 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -82,7 +82,7 @@ theorem exists_isLUB (hne : s.Nonempty) (hbdd : BddAbove s) : ∃ x, IsLUB s x : have j0 := Nat.cast_pos.1 ((inv_pos.2 ε0).trans_le ij) have k0 := Nat.cast_pos.1 ((inv_pos.2 ε0).trans_le ik) rcases hf₁ _ j0 with ⟨y, yS, hy⟩ - refine lt_of_lt_of_le ((Rat.cast_lt (K := ℝ)).1 ?_) ((inv_le ε0 (Nat.cast_pos.2 k0)).1 ik) + refine lt_of_lt_of_le ((Rat.cast_lt (K := ℝ)).1 ?_) ((inv_le_comm₀ ε0 (Nat.cast_pos.2 k0)).1 ik) simpa using sub_lt_iff_lt_add'.2 (lt_of_le_of_lt hy <| sub_lt_iff_lt_add.1 <| hf₂ _ k0 _ yS) let g : CauSeq ℚ abs := ⟨fun n => f n / n, hg⟩ refine ⟨mk g, ⟨fun x xS => ?_, fun y h => ?_⟩⟩ @@ -93,7 +93,7 @@ theorem exists_isLUB (hne : s.Nonempty) (hbdd : BddAbove s) : ∃ x, IsLUB s x : replace hK := hK.le.trans (Nat.cast_le.2 nK) have n0 : 0 < n := Nat.cast_pos.1 ((inv_pos.2 xz).trans_le hK) refine le_trans ?_ (hf₂ _ n0 _ xS).le - rwa [le_sub_comm, inv_le (Nat.cast_pos.2 n0 : (_ : ℝ) < _) xz] + rwa [le_sub_comm, inv_le_comm₀ (Nat.cast_pos.2 n0 : (_ : ℝ) < _) xz] · exact mk_le_of_forall_le ⟨1, fun n n1 => diff --git a/Mathlib/Data/Real/ConjExponents.lean b/Mathlib/Data/Real/ConjExponents.lean index 8ade2aebf3b1e..04bfc20da0ec1 100644 --- a/Mathlib/Data/Real/ConjExponents.lean +++ b/Mathlib/Data/Real/ConjExponents.lean @@ -108,7 +108,7 @@ theorem inv_add_inv_conj_ennreal : (ENNReal.ofReal p)⁻¹ + (ENNReal.ofReal q) end protected lemma inv_inv (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) : a⁻¹.IsConjExponent b⁻¹ := - ⟨one_lt_inv ha <| by linarith, by simpa only [inv_inv]⟩ + ⟨(one_lt_inv₀ ha).2 <| by linarith, by simpa only [inv_inv]⟩ lemma inv_one_sub_inv (ha₀ : 0 < a) (ha₁ : a < 1) : a⁻¹.IsConjExponent (1 - a)⁻¹ := .inv_inv ha₀ (sub_pos_of_lt ha₁) <| add_tsub_cancel_of_le ha₁.le @@ -199,7 +199,7 @@ end protected lemma inv_inv (ha : a ≠ 0) (hb : b ≠ 0) (hab : a + b = 1) : a⁻¹.IsConjExponent b⁻¹ := - ⟨one_lt_inv ha.bot_lt <| by rw [← hab]; exact lt_add_of_pos_right _ hb.bot_lt, by + ⟨(one_lt_inv₀ ha.bot_lt).2 <| by rw [← hab]; exact lt_add_of_pos_right _ hb.bot_lt, by simpa only [inv_inv] using hab⟩ lemma inv_one_sub_inv (ha₀ : a ≠ 0) (ha₁ : a < 1) : a⁻¹.IsConjExponent (1 - a)⁻¹ := diff --git a/Mathlib/Data/Real/GoldenRatio.lean b/Mathlib/Data/Real/GoldenRatio.lean index 7165e37237c30..42e9c2559ca88 100644 --- a/Mathlib/Data/Real/GoldenRatio.lean +++ b/Mathlib/Data/Real/GoldenRatio.lean @@ -109,7 +109,7 @@ theorem goldConj_ne_zero : ψ ≠ 0 := theorem neg_one_lt_goldConj : -1 < ψ := by rw [neg_lt, ← inv_gold] - exact inv_lt_one one_lt_gold + exact inv_lt_one_of_one_lt₀ one_lt_gold /-! ## Irrationality diff --git a/Mathlib/Data/Real/Hyperreal.lean b/Mathlib/Data/Real/Hyperreal.lean index 68aee623dbde7..6ffec46ad6c83 100644 --- a/Mathlib/Data/Real/Hyperreal.lean +++ b/Mathlib/Data/Real/Hyperreal.lean @@ -599,12 +599,12 @@ theorem infinitePos_iff_infinitesimal_inv_pos {x : ℝ*} : ⟨fun hip => ⟨infinitesimal_def.mpr fun r hr => ⟨lt_trans (coe_lt_coe.2 (neg_neg_of_pos hr)) (inv_pos.2 (hip 0)), - (inv_lt (coe_lt_coe.2 hr) (hip 0)).mp (by convert hip r⁻¹)⟩, + inv_lt_of_inv_lt₀ (coe_lt_coe.2 hr) (by convert hip r⁻¹)⟩, inv_pos.2 <| hip 0⟩, fun ⟨hi, hp⟩ r => @_root_.by_cases (r = 0) (↑r < x) (fun h => Eq.substr h (inv_pos.mp hp)) fun h => lt_of_le_of_lt (coe_le_coe.2 (le_abs_self r)) - ((inv_lt_inv (inv_pos.mp hp) (coe_lt_coe.2 (abs_pos.2 h))).mp + ((inv_lt_inv₀ (inv_pos.mp hp) (coe_lt_coe.2 (abs_pos.2 h))).mp ((infinitesimal_def.mp hi) |r|⁻¹ (inv_pos.2 (abs_pos.2 h))).2)⟩ theorem infiniteNeg_iff_infinitesimal_inv_neg {x : ℝ*} : diff --git a/Mathlib/Data/Real/Pi/Bounds.lean b/Mathlib/Data/Real/Pi/Bounds.lean index 9a77915281ee6..131af1b4136f2 100644 --- a/Mathlib/Data/Real/Pi/Bounds.lean +++ b/Mathlib/Data/Real/Pi/Bounds.lean @@ -40,7 +40,7 @@ theorem pi_lt_sqrtTwoAddSeries (n : ℕ) : · rw [div_le_iff₀'] · refine le_trans pi_le_four ?_ simp only [show (4 : ℝ) = (2 : ℝ) ^ 2 by norm_num, mul_one] - apply pow_le_pow_right (by norm_num) + apply pow_right_mono₀ (by norm_num) apply le_add_of_nonneg_left; apply Nat.zero_le · apply pow_pos; norm_num apply add_le_add_left; rw [div_le_div_right (by norm_num)] diff --git a/Mathlib/Data/Real/Pi/Irrational.lean b/Mathlib/Data/Real/Pi/Irrational.lean index 778010911a87f..e3884b774698d 100644 --- a/Mathlib/Data/Real/Pi/Irrational.lean +++ b/Mathlib/Data/Real/Pi/Irrational.lean @@ -5,6 +5,7 @@ Authors: Bhavik Mehta -/ import Mathlib.Analysis.SpecialFunctions.Integrals import Mathlib.Data.Real.Irrational +import Mathlib.Topology.Algebra.Order.Floor /-! # `Real.pi` is irrational @@ -250,7 +251,7 @@ private lemma I_le (n : ℕ) : I n (π / 2) ≤ 2 := by intros x hx simp only [uIoc_of_le, neg_le_self_iff, zero_le_one, mem_Ioc] at hx rw [norm_eq_abs, abs_mul, abs_pow] - refine mul_le_one (pow_le_one₀ (abs_nonneg _) ?_) (abs_nonneg _) (abs_cos_le_one _) + refine mul_le_one₀ (pow_le_one₀ (abs_nonneg _) ?_) (abs_nonneg _) (abs_cos_le_one _) rw [abs_le] constructor <;> nlinarith @@ -262,7 +263,7 @@ reformulation of tendsto_pow_div_factorial_atTop, which asserts the same for `a private lemma tendsto_pow_div_factorial_at_top_aux (a : ℝ) : Tendsto (fun n => (a : ℝ) ^ (2 * n + 1) / n !) atTop (nhds 0) := by rw [← mul_zero a] - refine ((tendsto_pow_div_factorial_atTop (a ^ 2)).const_mul a).congr (fun x => ?_) + refine ((FloorSemiring.tendsto_pow_div_factorial_atTop (a ^ 2)).const_mul a).congr (fun x => ?_) rw [← pow_mul, mul_div_assoc', _root_.pow_succ'] /-- If `x` is rational, it can be written as `a / b` with `a : ℤ` and `b : ℕ` satisfying `b > 0`. -/ diff --git a/Mathlib/Data/Seq/Seq.lean b/Mathlib/Data/Seq/Seq.lean index e1c08ea1f9bdc..d8dfa6e7ed037 100644 --- a/Mathlib/Data/Seq/Seq.lean +++ b/Mathlib/Data/Seq/Seq.lean @@ -512,8 +512,6 @@ def zipWith (f : α → β → γ) (s₁ : Seq α) (s₂ : Seq β) : Seq γ := ⟨fun n => Option.map₂ f (s₁.get? n) (s₂.get? n), fun {_} hn => Option.map₂_eq_none_iff.2 <| (Option.map₂_eq_none_iff.1 hn).imp s₁.2 s₂.2⟩ -variable {s : Seq α} {s' : Seq β} {n : ℕ} - @[simp] theorem get?_zipWith (f : α → β → γ) (s s' n) : (zipWith f s s').get? n = Option.map₂ f (s.get? n) (s'.get? n) := diff --git a/Mathlib/Data/Set/Accumulate.lean b/Mathlib/Data/Set/Accumulate.lean index fcbe0ea222997..ffa9845a9ad40 100644 --- a/Mathlib/Data/Set/Accumulate.lean +++ b/Mathlib/Data/Set/Accumulate.lean @@ -12,7 +12,7 @@ The function `Accumulate` takes a set `s` and returns `⋃ y ≤ x, s y`. -/ -variable {α β γ : Type*} {s : α → Set β} {t : α → Set γ} +variable {α β : Type*} {s : α → Set β} namespace Set diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 5a4257688c30d..691836860b538 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -1958,7 +1958,7 @@ open Set namespace Function -variable {ι : Sort*} {α : Type*} {β : Type*} {f : α → β} +variable {α : Type*} {β : Type*} theorem Injective.nonempty_apply_iff {f : Set α → Set β} (hf : Injective f) (h2 : f ∅ = ∅) {s : Set α} : (f s).Nonempty ↔ s.Nonempty := by @@ -2141,7 +2141,7 @@ end Monotone /-! ### Disjoint sets -/ -variable {α β : Type*} {s t u : Set α} {f : α → β} +variable {α : Type*} {s t u : Set α} namespace Disjoint diff --git a/Mathlib/Data/Set/Finite.lean b/Mathlib/Data/Set/Finite.lean index dcd9b53e9dcc2..cf7a2bfbb02d2 100644 --- a/Mathlib/Data/Set/Finite.lean +++ b/Mathlib/Data/Set/Finite.lean @@ -897,7 +897,7 @@ theorem exists_subset_image_finite_and {f : α → β} {s : Set α} {p : Set β (∃ t ⊆ f '' s, t.Finite ∧ p t) ↔ ∃ t ⊆ s, t.Finite ∧ p (f '' t) := by classical simp_rw [@and_comm (_ ⊆ _), and_assoc, exists_finite_iff_finset, @and_comm (p _), - Finset.subset_image_iff] + Finset.subset_set_image_iff] aesop section Pi diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 769796325c380..3810aaead4d4d 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -625,8 +625,15 @@ theorem InjOn.imageFactorization_injective (h : InjOn f s) : end injOn section graphOn +variable {x : α × β} + +@[simp] lemma mem_graphOn : x ∈ s.graphOn f ↔ x.1 ∈ s ∧ f x.1 = x.2 := by aesop (add simp graphOn) @[simp] lemma graphOn_empty (f : α → β) : graphOn f ∅ = ∅ := image_empty _ +@[simp] lemma graphOn_eq_empty : graphOn f s = ∅ ↔ s = ∅ := image_eq_empty +@[simp] lemma graphOn_nonempty : (s.graphOn f).Nonempty ↔ s.Nonempty := image_nonempty + +protected alias ⟨_, Nonempty.graphOn⟩ := graphOn_nonempty @[simp] lemma graphOn_union (f : α → β) (s t : Set α) : graphOn f (s ∪ t) = graphOn f s ∪ graphOn f t := @@ -645,6 +652,24 @@ lemma graphOn_insert (f : α → β) (x : α) (s : Set α) : lemma image_fst_graphOn (f : α → β) (s : Set α) : Prod.fst '' graphOn f s = s := by simp [graphOn, image_image] +@[simp] lemma image_snd_graphOn (f : α → β) : Prod.snd '' s.graphOn f = f '' s := by ext x; simp + +lemma fst_injOn_graph : (s.graphOn f).InjOn Prod.fst := by aesop (add simp InjOn) + +lemma graphOn_comp (s : Set α) (f : α → β) (g : β → γ) : + s.graphOn (g ∘ f) = (fun x ↦ (x.1, g x.2)) '' s.graphOn f := by + simpa using image_comp (fun x ↦ (x.1, g x.2)) (fun x ↦ (x, f x)) _ + +lemma graphOn_univ_eq_range : univ.graphOn f = range fun x ↦ (x, f x) := image_univ + +@[simp] lemma graphOn_inj {g : α → β} : s.graphOn f = s.graphOn g ↔ s.EqOn f g := by + simp [Set.ext_iff, funext_iff, forall_swap, EqOn] + +lemma graphOn_univ_inj {g : α → β} : univ.graphOn f = univ.graphOn g ↔ f = g := by simp + +lemma graphOn_univ_injective : Injective (univ.graphOn : (α → β) → Set (α × β)) := + fun _f _g ↦ graphOn_univ_inj.1 + lemma exists_eq_graphOn_image_fst [Nonempty β] {s : Set (α × β)} : (∃ f : α → β, s = graphOn f (Prod.fst '' s)) ↔ InjOn Prod.fst s := by refine ⟨?_, fun h ↦ ?_⟩ @@ -1310,11 +1335,14 @@ theorem le_piecewise {δ : α → Type*} [∀ i, Preorder (δ i)] {s : Set α} [ g ≤ s.piecewise f₁ f₂ := @piecewise_le α (fun i => (δ i)ᵒᵈ) _ s _ _ _ _ h₁ h₂ -theorem piecewise_le_piecewise {δ : α → Type*} [∀ i, Preorder (δ i)] {s : Set α} +@[gcongr] +theorem piecewise_mono {δ : α → Type*} [∀ i, Preorder (δ i)] {s : Set α} [∀ j, Decidable (j ∈ s)] {f₁ f₂ g₁ g₂ : ∀ i, δ i} (h₁ : ∀ i ∈ s, f₁ i ≤ g₁ i) (h₂ : ∀ i ∉ s, f₂ i ≤ g₂ i) : s.piecewise f₁ f₂ ≤ s.piecewise g₁ g₂ := by apply piecewise_le <;> intros <;> simp [*] +@[deprecated (since := "2024-10-06")] alias piecewise_le_piecewise := piecewise_mono + @[simp] theorem piecewise_insert_of_ne {i j : α} (h : i ≠ j) [∀ i, Decidable (i ∈ insert j s)] : (insert j s).piecewise f g i = s.piecewise f g i := by simp [piecewise, h] diff --git a/Mathlib/Data/Set/Monotone.lean b/Mathlib/Data/Set/Monotone.lean index f811b7ae15131..5bd60bb9fb1c5 100644 --- a/Mathlib/Data/Set/Monotone.lean +++ b/Mathlib/Data/Set/Monotone.lean @@ -85,6 +85,16 @@ protected theorem _root_.StrictAntiOn.strictAnti (h : StrictAntiOn f s) : StrictAnti (f ∘ Subtype.val : s → β) := fun x y hlt => h x.coe_prop y.coe_prop hlt +lemma MonotoneOn_insert_iff {a : α} : + MonotoneOn f (insert a s) ↔ + (∀ b ∈ s, b ≤ a → f b ≤ f a) ∧ (∀ b ∈ s, a ≤ b → f a ≤ f b) ∧ MonotoneOn f s := by + simp [MonotoneOn, forall_and] + +lemma AntitoneOn_insert_iff {a : α} : + AntitoneOn f (insert a s) ↔ + (∀ b ∈ s, b ≤ a → f a ≤ f b) ∧ (∀ b ∈ s, a ≤ b → f b ≤ f a) ∧ AntitoneOn f s := + @MonotoneOn_insert_iff α βᵒᵈ _ _ _ _ _ + end Mono end Set @@ -148,6 +158,16 @@ theorem StrictMono.codRestrict [Preorder α] [Preorder β] {f : α → β} (hf : {s : Set β} (hs : ∀ x, f x ∈ s) : StrictMono (Set.codRestrict f s hs) := hf +lemma strictMonoOn_insert_iff [Preorder α] [Preorder β] {f : α → β} {s : Set α} {a : α} : + StrictMonoOn f (insert a s) ↔ + (∀ b ∈ s, b < a → f b < f a) ∧ (∀ b ∈ s, a < b → f a < f b) ∧ StrictMonoOn f s := by + simp [StrictMonoOn, forall_and] + +lemma strictAntiOn_insert_iff [Preorder α] [Preorder β] {f : α → β} {s : Set α} {a : α} : + StrictAntiOn f (insert a s) ↔ + (∀ b ∈ s, b < a → f a < f b) ∧ (∀ b ∈ s, a < b → f b < f a) ∧ StrictAntiOn f s := + @strictMonoOn_insert_iff α βᵒᵈ _ _ _ _ _ + end strictMono namespace Function diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Data/Set/Pointwise/Interval.lean index 023d3be7a38d7..27d7659087f07 100644 --- a/Mathlib/Data/Set/Pointwise/Interval.lean +++ b/Mathlib/Data/Set/Pointwise/Interval.lean @@ -726,9 +726,9 @@ theorem image_mul_left_Ioo {a : α} (h : 0 < a) (b c : α) : theorem inv_Ioo_0_left {a : α} (ha : 0 < a) : (Ioo 0 a)⁻¹ = Ioi a⁻¹ := by ext x exact - ⟨fun h => inv_inv x ▸ (inv_lt_inv ha h.1).2 h.2, fun h => + ⟨fun h => inv_inv x ▸ (inv_lt_inv₀ ha h.1).2 h.2, fun h => ⟨inv_pos.2 <| (inv_pos.2 ha).trans h, - inv_inv a ▸ (inv_lt_inv ((inv_pos.2 ha).trans h) + inv_inv a ▸ (inv_lt_inv₀ ((inv_pos.2 ha).trans h) (inv_pos.2 ha)).2 h⟩⟩ theorem inv_Ioi {a : α} (ha : 0 < a) : (Ioi a)⁻¹ = Ioo 0 a⁻¹ := by diff --git a/Mathlib/Data/Set/Pointwise/ListOfFn.lean b/Mathlib/Data/Set/Pointwise/ListOfFn.lean index 21e7712a88181..120059da9dfca 100644 --- a/Mathlib/Data/Set/Pointwise/ListOfFn.lean +++ b/Mathlib/Data/Set/Pointwise/ListOfFn.lean @@ -15,8 +15,7 @@ This file proves some lemmas about pointwise algebraic operations with lists of namespace Set -variable {F α β γ : Type*} -variable [Monoid α] {s t : Set α} {a : α} {m n : ℕ} +variable {α : Type*} [Monoid α] {s : Set α} {n : ℕ} open Pointwise diff --git a/Mathlib/Data/Set/Prod.lean b/Mathlib/Data/Set/Prod.lean index b1c42bc17e180..d147798baa0d9 100644 --- a/Mathlib/Data/Set/Prod.lean +++ b/Mathlib/Data/Set/Prod.lean @@ -686,15 +686,6 @@ theorem disjoint_pi : Disjoint (s.pi t₁) (s.pi t₂) ↔ ∃ i ∈ s, Disjoint end Nonempty --- Porting note: Removing `simp` - LHS does not simplify -theorem range_dcomp (f : ∀ i, α i → β i) : - (range fun g : ∀ i, α i => fun i => f i (g i)) = pi univ fun i => range (f i) := by - refine Subset.antisymm ?_ fun x hx => ?_ - · rintro _ ⟨x, rfl⟩ i - - exact ⟨x i, rfl⟩ - · choose y hy using hx - exact ⟨fun i => y i trivial, funext fun i => hy i trivial⟩ - @[simp] theorem insert_pi (i : ι) (s : Set ι) (t : ∀ i, Set (α i)) : pi (insert i s) t = eval i ⁻¹' t i ∩ pi s t := by @@ -809,8 +800,8 @@ theorem eval_image_univ_pi (ht : (pi univ t).Nonempty) : (fun f : ∀ i, α i => f i) '' pi univ t = t i := eval_image_pi (mem_univ i) ht -theorem dcomp_image_pi {f : ∀ i, α i → β i} (hf : ∀ i ∉ s, Surjective (f i)) (t : ∀ i, Set (α i)) : - (f _ ∘' ·) '' s.pi t = s.pi fun i ↦ f i '' t i := by +theorem piMap_image_pi {f : ∀ i, α i → β i} (hf : ∀ i ∉ s, Surjective (f i)) (t : ∀ i, Set (α i)) : + Pi.map f '' s.pi t = s.pi fun i ↦ f i '' t i := by refine Subset.antisymm (image_subset_iff.2 fun a ha i hi ↦ mem_image_of_mem _ (ha _ hi)) ?_ intro b hb have : ∀ i, ∃ a, f i a = b i ∧ (i ∈ s → a ∈ t i) := by @@ -822,9 +813,19 @@ theorem dcomp_image_pi {f : ∀ i, α i → β i} (hf : ∀ i ∉ s, Surjective choose a hab hat using this exact ⟨a, hat, funext hab⟩ -theorem dcomp_image_univ_pi (f : ∀ i, α i → β i) (t : ∀ i, Set (α i)) : - (f _ ∘' ·) '' univ.pi t = univ.pi fun i ↦ f i '' t i := - dcomp_image_pi (by simp) t +@[deprecated (since := "2024-10-06")] alias dcomp_image_pi := piMap_image_pi + +theorem piMap_image_univ_pi (f : ∀ i, α i → β i) (t : ∀ i, Set (α i)) : + Pi.map f '' univ.pi t = univ.pi fun i ↦ f i '' t i := + piMap_image_pi (by simp) t + +@[deprecated (since := "2024-10-06")] alias dcomp_image_univ_pi := piMap_image_univ_pi + +@[simp] +theorem range_piMap (f : ∀ i, α i → β i) : range (Pi.map f) = pi univ fun i ↦ range (f i) := by + simp only [← image_univ, ← piMap_image_univ_pi, pi_univ] + +@[deprecated (since := "2024-10-06")] alias range_dcomp := range_piMap theorem pi_subset_pi_iff : pi s t₁ ⊆ pi s t₂ ↔ (∀ i ∈ s, t₁ i ⊆ t₂ i) ∨ pi s t₁ = ∅ := by refine diff --git a/Mathlib/Data/SetLike/Basic.lean b/Mathlib/Data/SetLike/Basic.lean index 9613ca363c6b8..2a07db5d06280 100644 --- a/Mathlib/Data/SetLike/Basic.lean +++ b/Mathlib/Data/SetLike/Basic.lean @@ -28,7 +28,7 @@ boilerplate for every `SetLike`: a `coe_sort`, a `coe` to set, a A typical subobject should be declared as: ``` -structure MySubobject (X : Type*) [ObjectTypeclass X] := +structure MySubobject (X : Type*) [ObjectTypeclass X] where (carrier : Set X) (op_mem' : ∀ {x : X}, x ∈ carrier → sorry ∈ carrier) @@ -60,7 +60,7 @@ end MySubobject An alternative to `SetLike` could have been an extensional `Membership` typeclass: ``` -class ExtMembership (α : out_param <| Type u) (β : Type v) extends Membership α β := +class ExtMembership (α : out_param <| Type u) (β : Type v) extends Membership α β where (ext_iff : ∀ {s t : β}, s = t ↔ ∀ (x : α), x ∈ s ↔ x ∈ t) ``` While this is equivalent, `SetLike` conveniently uses a carrier set projection directly. diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index 4f812be71f3cd..8734d5bf5b8d2 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -91,7 +91,7 @@ def ker (f : α → β) : Setoid α := theorem ker_mk_eq (r : Setoid α) : ker (@Quotient.mk'' _ r) = r := ext' fun _ _ => Quotient.eq -theorem ker_apply_mk_out {f : α → β} (a : α) : f (haveI := Setoid.ker f; ⟦a⟧.out) = f a := +theorem ker_apply_mk_out {f : α → β} (a : α) : f (⟦a⟧ : Quotient (Setoid.ker f)).out = f a := @Quotient.mk_out _ (Setoid.ker f) a theorem ker_apply_mk_out' {f : α → β} (a : α) : @@ -111,6 +111,49 @@ protected def prod (r : Setoid α) (s : Setoid β) : ⟨fun x => ⟨r.refl' x.1, s.refl' x.2⟩, fun h => ⟨r.symm' h.1, s.symm' h.2⟩, fun h₁ h₂ => ⟨r.trans' h₁.1 h₂.1, s.trans' h₁.2 h₂.2⟩⟩ +lemma prod_apply {r : Setoid α} {s : Setoid β} {x₁ x₂ : α} {y₁ y₂ : β} : + @Setoid.r _ (r.prod s) (x₁, y₁) (x₂, y₂) ↔ (@Setoid.r _ r x₁ x₂ ∧ @Setoid.r _ s y₁ y₂) := + Iff.rfl + +lemma piSetoid_apply {ι : Sort*} {α : ι → Sort*} {r : ∀ i, Setoid (α i)} {x y : ∀ i, α i} : + @Setoid.r _ (@piSetoid _ _ r) x y ↔ ∀ i, @Setoid.r _ (r i) (x i) (y i) := + Iff.rfl + +/-- A bijection between the product of two quotients and the quotient by the product of the +equivalence relations. -/ +@[simps] +def prodQuotientEquiv (r : Setoid α) (s : Setoid β) : + Quotient r × Quotient s ≃ Quotient (r.prod s) where + toFun := fun (x, y) ↦ Quotient.map₂' Prod.mk (fun _ _ hx _ _ hy ↦ ⟨hx, hy⟩) x y + invFun := fun q ↦ Quotient.liftOn' q (fun xy ↦ (Quotient.mk'' xy.1, Quotient.mk'' xy.2)) + fun x y hxy ↦ Prod.ext (by simpa using hxy.1) (by simpa using hxy.2) + left_inv := fun q ↦ by + rcases q with ⟨qa, qb⟩ + exact Quotient.inductionOn₂' qa qb fun _ _ ↦ rfl + right_inv := fun q ↦ by + simp only + refine Quotient.inductionOn' q fun _ ↦ rfl + +/-- A bijection between an indexed product of quotients and the quotient by the product of the +equivalence relations. -/ +@[simps] +noncomputable def piQuotientEquiv {ι : Sort*} {α : ι → Sort*} (r : ∀ i, Setoid (α i)) : + (∀ i, Quotient (r i)) ≃ Quotient (@piSetoid _ _ r) where + toFun := fun x ↦ Quotient.mk'' fun i ↦ (x i).out' + invFun := fun q ↦ Quotient.liftOn' q (fun x i ↦ Quotient.mk'' (x i)) fun x y hxy ↦ by + ext i + simpa using hxy i + left_inv := fun q ↦ by + ext i + simp + right_inv := fun q ↦ by + refine Quotient.inductionOn' q fun _ ↦ ?_ + simp only [Quotient.liftOn'_mk'', Quotient.eq''] + intro i + change Setoid.r _ _ + rw [← Quotient.eq''] + simp + /-- The infimum of two equivalence relations. -/ instance : Inf (Setoid α) := ⟨fun r s => diff --git a/Mathlib/Data/Sigma/Lex.lean b/Mathlib/Data/Sigma/Lex.lean index 6b5e6be4f0f6d..db66010da8c5d 100644 --- a/Mathlib/Data/Sigma/Lex.lean +++ b/Mathlib/Data/Sigma/Lex.lean @@ -137,7 +137,7 @@ end Sigma namespace PSigma -variable {ι : Sort*} {α : ι → Sort*} {r r₁ r₂ : ι → ι → Prop} {s s₁ s₂ : ∀ i, α i → α i → Prop} +variable {ι : Sort*} {α : ι → Sort*} {r : ι → ι → Prop} {s : ∀ i, α i → α i → Prop} theorem lex_iff {a b : Σ' i, α i} : Lex r s a b ↔ r a.1 b.1 ∨ ∃ h : a.1 = b.1, s b.1 (h.rec a.2) b.2 := by diff --git a/Mathlib/Data/Sum/Basic.lean b/Mathlib/Data/Sum/Basic.lean index 9da00e98e9adb..d65e9ce7874fe 100644 --- a/Mathlib/Data/Sum/Basic.lean +++ b/Mathlib/Data/Sum/Basic.lean @@ -35,7 +35,7 @@ theorem sum_rec_congr (P : α ⊕ β → Sort*) (f : ∀ i, P (inl i)) (g : ∀ section get -variable {x y : α ⊕ β} +variable {x : α ⊕ β} theorem eq_left_iff_getLeft_eq {a : α} : x = inl a ↔ ∃ h, x.getLeft h = a := by cases x <;> simp diff --git a/Mathlib/Deprecated/Group.lean b/Mathlib/Deprecated/Group.lean index d2d59cec9f5b0..539782f49e725 100644 --- a/Mathlib/Deprecated/Group.lean +++ b/Mathlib/Deprecated/Group.lean @@ -359,7 +359,7 @@ end Units namespace IsUnit -variable {M : Type*} {N : Type*} [Monoid M] [Monoid N] {x : M} +variable {M : Type*} {N : Type*} [Monoid M] [Monoid N] theorem map' {f : M → N} (hf : IsMonoidHom f) {x : M} (h : IsUnit x) : IsUnit (f x) := h.map (MonoidHom.of hf) diff --git a/Mathlib/Deprecated/Ring.lean b/Mathlib/Deprecated/Ring.lean index 5cc4f4a17e1ca..fc34ab12c07ee 100644 --- a/Mathlib/Deprecated/Ring.lean +++ b/Mathlib/Deprecated/Ring.lean @@ -27,7 +27,7 @@ IsSemiringHom, IsRingHom -/ -universe u v w +universe u v variable {α : Type u} @@ -45,7 +45,7 @@ structure IsSemiringHom {α : Type u} {β : Type v} [Semiring α] [Semiring β] namespace IsSemiringHom variable {β : Type v} [Semiring α] [Semiring β] -variable {f : α → β} (hf : IsSemiringHom f) {x y : α} +variable {f : α → β} /-- The identity map is a semiring homomorphism. -/ theorem id : IsSemiringHom (@id α) := by constructor <;> intros <;> rfl @@ -85,7 +85,7 @@ variable {β : Type v} [Ring α] [Ring β] theorem of_semiring {f : α → β} (H : IsSemiringHom f) : IsRingHom f := { H with } -variable {f : α → β} (hf : IsRingHom f) {x y : α} +variable {f : α → β} {x y : α} /-- Ring homomorphisms map zero to zero. -/ theorem map_zero (hf : IsRingHom f) : f 0 = 0 := @@ -122,7 +122,7 @@ theorem to_isAddGroupHom (hf : IsRingHom f) : IsAddGroupHom f := end IsRingHom -variable {β : Type v} {γ : Type w} {rα : Semiring α} {rβ : Semiring β} +variable {β : Type v} {rα : Semiring α} {rβ : Semiring β} namespace RingHom diff --git a/Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean b/Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean index ad6c52b6be339..08b6c75d0d4c0 100644 --- a/Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean +++ b/Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean @@ -88,7 +88,7 @@ lemma _root_.isOpen.dynEntourage [TopologicalSpace X] {T : X → X} (T_cont : Co IsOpen (dynEntourage T U n) := by rw [dynEntourage_eq_inter_Ico T U n] refine isOpen_iInter_of_finite fun k ↦ ?_ - exact continuous_def.1 ((T_cont.prod_map T_cont).iterate k) U U_open + exact U_open.preimage ((T_cont.prodMap T_cont).iterate k) lemma dynEntourage_monotone (T : X → X) (n : ℕ) : Monotone (fun U : Set (X × X) ↦ dynEntourage T U n) := diff --git a/Mathlib/FieldTheory/Minpoly/Field.lean b/Mathlib/FieldTheory/Minpoly/Field.lean index 3a33353b9af98..b129b8a0ecbc5 100644 --- a/Mathlib/FieldTheory/Minpoly/Field.lean +++ b/Mathlib/FieldTheory/Minpoly/Field.lean @@ -148,6 +148,27 @@ theorem sub_algebraMap {B : Type*} [CommRing B] [Algebra A B] {x : B} (a : A) : minpoly A (x - algebraMap A B a) = (minpoly A x).comp (X + C a) := by simpa [sub_eq_add_neg] using add_algebraMap x (-a) +theorem neg {B : Type*} [CommRing B] [Algebra A B] (x : B) : + minpoly A (- x) = (-1) ^ (natDegree (minpoly A x)) * (minpoly A x).comp (- X) := by + by_cases hx : IsIntegral A x + · refine (minpoly.unique _ _ ((minpoly.monic hx).neg_one_pow_natDegree_mul_comp_neg_X) + ?_ fun q qmo hq => ?_).symm + · simp [aeval_comp] + · have : (Polynomial.aeval x) ((-1) ^ q.natDegree * q.comp (- X)) = 0 := by + simpa [aeval_comp] using hq + have H := minpoly.min A x qmo.neg_one_pow_natDegree_mul_comp_neg_X this + have n1 := ((minpoly.monic hx).neg_one_pow_natDegree_mul_comp_neg_X).ne_zero + have n2 := qmo.neg_one_pow_natDegree_mul_comp_neg_X.ne_zero + rw [degree_eq_natDegree qmo.ne_zero, + degree_eq_natDegree n1, natDegree_mul (by simp) (right_ne_zero_of_mul n1), natDegree_comp] + rw [degree_eq_natDegree (minpoly.ne_zero hx), + degree_eq_natDegree qmo.neg_one_pow_natDegree_mul_comp_neg_X.ne_zero, + natDegree_mul (by simp) (right_ne_zero_of_mul n2), natDegree_comp] at H + simpa using H + · rw [minpoly.eq_zero hx, minpoly.eq_zero, zero_comp] + · simp only [natDegree_zero, pow_zero, mul_zero] + · exact IsIntegral.neg_iff.not.mpr hx + section AlgHomFintype /-- A technical finiteness result. -/ diff --git a/Mathlib/FieldTheory/RatFunc/Basic.lean b/Mathlib/FieldTheory/RatFunc/Basic.lean index 477f652402ff1..66d5c87abee7c 100644 --- a/Mathlib/FieldTheory/RatFunc/Basic.lean +++ b/Mathlib/FieldTheory/RatFunc/Basic.lean @@ -269,7 +269,7 @@ macro "smul_tac" : tactic => `(tactic| simp_rw [← ofFractionRing_smul] <;> simp only [add_comm, mul_comm, zero_smul, succ_nsmul, zsmul_eq_mul, mul_add, mul_one, mul_zero, neg_add, mul_neg, - Int.ofNat_eq_coe, Int.cast_zero, Int.cast_add, Int.cast_one, + Int.cast_zero, Int.cast_add, Int.cast_one, Int.cast_negSucc, Int.cast_natCast, Nat.cast_succ, Localization.mk_zero, Localization.add_mk_self, Localization.neg_mk, ofFractionRing_zero, ← ofFractionRing_add, ← ofFractionRing_neg]) diff --git a/Mathlib/FieldTheory/Separable.lean b/Mathlib/FieldTheory/Separable.lean index fd852c7cdf7d7..11f14b5edd03f 100644 --- a/Mathlib/FieldTheory/Separable.lean +++ b/Mathlib/FieldTheory/Separable.lean @@ -337,8 +337,7 @@ theorem separable_or {f : F[X]} (hf : Irreducible f) : exact Or.inr ⟨by rw [separable_iff_derivative_ne_zero hf, Classical.not_not, H], contract p f, - of_irreducible_map (expand F p : F[X] →+* F[X]) - (by rwa [← expand_contract p H hp.ne'] at hf), + Irreducible.of_map (by rwa [← expand_contract p H hp.ne'] at hf), expand_contract p H hp.ne'⟩ else Or.inl <| (separable_iff_derivative_ne_zero hf).2 H diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean index 9dcdde2ba5730..b58876fe73b9d 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean @@ -350,7 +350,7 @@ theorem angle_eq_zero_iff_ne_and_wbtw {p₁ p₂ p₃ : P} : · rw [angle, angle_eq_zero_iff] rintro ⟨hp₁p₂, r, hr0, hp₃p₂⟩ rcases le_or_lt 1 r with (hr1 | hr1) - · refine Or.inl ⟨vsub_ne_zero.1 hp₁p₂, r⁻¹, ⟨(inv_pos.2 hr0).le, inv_le_one hr1⟩, ?_⟩ + · refine Or.inl ⟨vsub_ne_zero.1 hp₁p₂, r⁻¹, ⟨(inv_pos.2 hr0).le, inv_le_one_of_one_le₀ hr1⟩, ?_⟩ rw [AffineMap.lineMap_apply, hp₃p₂, smul_smul, inv_mul_cancel₀ hr0.ne.symm, one_smul, vsub_vadd] · refine Or.inr ⟨?_, r, ⟨hr0.le, hr1.le⟩, ?_⟩ diff --git a/Mathlib/Geometry/Manifold/Complex.lean b/Mathlib/Geometry/Manifold/Complex.lean index a0110600277bb..a9ea94cdde69c 100644 --- a/Mathlib/Geometry/Manifold/Complex.lean +++ b/Mathlib/Geometry/Manifold/Complex.lean @@ -90,7 +90,7 @@ theorem norm_eqOn_of_isPreconnected_of_isMaxOn {f : M → F} {U : Set M} {c : M} replace hm : IsLocalMax (‖f ·‖) x := mem_of_superset (ho.mem_nhds hx.1) fun z hz ↦ (hm hz).out.trans_eq hx.2.symm replace hd : ∀ᶠ y in 𝓝 x, MDifferentiableAt I 𝓘(ℂ, F) f y := - (eventually_mem_nhds.2 (ho.mem_nhds hx.1)).mono fun z ↦ hd.mdifferentiableAt + (eventually_mem_nhds_iff.2 (ho.mem_nhds hx.1)).mono fun z ↦ hd.mdifferentiableAt exact (Complex.norm_eventually_eq_of_mdifferentiableAt_of_isLocalMax hd hm).mono fun _ ↦ (Eq.trans · hx.2) have hVne : (U ∩ V).Nonempty := ⟨c, hcU, hcU, rfl⟩ diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index 3028f6318c3cd..cde9ee6782c47 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -769,7 +769,7 @@ theorem contMDiffAt_iff_contMDiffAt_nhds refine ⟨?_, fun h => h.self_of_nhds⟩ rw [contMDiffAt_iff_contMDiffOn_nhds] rintro ⟨u, hu, h⟩ - refine (eventually_mem_nhds.mpr hu).mono fun x' hx' => ?_ + refine (eventually_mem_nhds_iff.mpr hu).mono fun x' hx' => ?_ exact (h x' <| mem_of_mem_nhds hx').contMDiffAt hx' /-! ### Congruence lemmas -/ diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 064e40fa40d71..76b954315da82 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -171,7 +171,7 @@ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (h lemma IsIntegralCurveAt.eventually_hasDerivAt (hγ : IsIntegralCurveAt γ v t₀) : ∀ᶠ t in 𝓝 t₀, HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t := by - apply eventually_mem_nhds.mpr + apply eventually_mem_nhds_iff.mpr (hγ.continuousAt.preimage_mem_nhds (extChartAt_source_mem_nhds I _)) |>.and hγ |>.mono rintro t ⟨ht1, ht2⟩ have hsrc := mem_of_mem_nhds ht1 @@ -325,7 +325,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt [CompleteSpace E] rw [continuousAt_def, hf1] at hcont have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ 𝓝 t₀ := hcont _ (isOpen_interior.mem_nhds ((I.isInteriorPoint_iff).mp hx)) - rw [← eventually_mem_nhds] at hnhds + rw [← eventually_mem_nhds_iff] at hnhds -- obtain a neighbourhood `s` so that the above conditions both hold in `s` obtain ⟨s, hs, haux⟩ := (hf2.and hnhds).exists_mem -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve @@ -391,7 +391,7 @@ theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt₀ : I.IsInteriorPoi have hlip (t : ℝ) : LipschitzOnWith K ((fun _ ↦ v') t) ((fun _ ↦ s) t) := hlip -- internal lemmas to reduce code duplication have hsrc {g} (hg : IsIntegralCurveAt g v t₀) : - ∀ᶠ t in 𝓝 t₀, g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t := eventually_mem_nhds.mpr <| + ∀ᶠ t in 𝓝 t₀, g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t := eventually_mem_nhds_iff.mpr <| continuousAt_def.mp hg.continuousAt _ <| extChartAt_source_mem_nhds I (g t₀) have hmem {g : ℝ → M} {t} (ht : g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t) : g t ∈ (extChartAt I (g t₀)).source := mem_preimage.mp <| mem_of_mem_nhds ht diff --git a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean index 9feae63ffc666..5ea4b4a42e6b8 100644 --- a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean +++ b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean @@ -68,7 +68,7 @@ structure LocalInvariantProp (P : (H → H') → Set H → H → Prop) : Prop wh left_invariance' : ∀ {s x f} {e' : PartialHomeomorph H' H'}, e' ∈ G' → s ⊆ f ⁻¹' e'.source → f x ∈ e'.source → P f s x → P (e' ∘ f) s x -variable {G G'} {P : (H → H') → Set H → H → Prop} {s t u : Set H} {x : H} +variable {G G'} {P : (H → H') → Set H → H → Prop} variable (hG : G.LocalInvariantProp G' P) include hG diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 9f6ab2d09bb29..1b1e09a291612 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -382,8 +382,8 @@ def ModelWithCorners.prod {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Ty source := { x | x.1 ∈ I.source ∧ x.2 ∈ I'.source } source_eq := by simp only [setOf_true, mfld_simps] uniqueDiffOn' := I.uniqueDiffOn'.prod I'.uniqueDiffOn' - continuous_toFun := I.continuous_toFun.prod_map I'.continuous_toFun - continuous_invFun := I.continuous_invFun.prod_map I'.continuous_invFun } + continuous_toFun := I.continuous_toFun.prodMap I'.continuous_toFun + continuous_invFun := I.continuous_invFun.prodMap I'.continuous_invFun } /-- Given a finite family of `ModelWithCorners` `I i` on `(E i, H i)`, we define the model with corners `pi I` on `(Π i, E i, ModelPi H)`. See note [Manifold type tags] for explanation about diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean index 75e798757aa2f..80753af5d8fda 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean @@ -89,10 +89,11 @@ noncomputable def Hom.stalkMap {X Y : LocallyRingedSpace.{u}} (f : Hom X Y) (x : Y.presheaf.stalk (f.1.1 x) ⟶ X.presheaf.stalk x := f.val.stalkMap x -instance {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : IsLocalRingHom (f.stalkMap x) := +instance isLocalRingHomStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : + IsLocalRingHom (f.stalkMap x) := f.2 x -instance {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : +instance isLocalRingHomValStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : IsLocalRingHom (f.val.stalkMap x) := f.2 x @@ -108,7 +109,7 @@ instance (X : LocallyRingedSpace.{u}) : Inhabited (Hom X X) := def comp {X Y Z : LocallyRingedSpace.{u}} (f : Hom X Y) (g : Hom Y Z) : Hom X Z := ⟨f.val ≫ g.val, fun x => by erw [PresheafedSpace.stalkMap.comp] - exact @isLocalRingHom_comp _ _ _ _ _ _ _ _ (f.2 _) (g.2 _)⟩ + exact @RingHom.isLocalRingHom_comp _ _ _ _ _ _ _ _ (f.2 _) (g.2 _)⟩ /-- The category of locally ringed spaces. -/ instance : Category LocallyRingedSpace.{u} where @@ -271,7 +272,7 @@ theorem preimage_basicOpen {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) {U : Ope · rintro ⟨y, hy : IsUnit _, rfl⟩ erw [RingedSpace.mem_basicOpen _ _ ⟨f.1.base y.1, y.2⟩] erw [← PresheafedSpace.stalkMap_germ_apply] at hy - exact (isUnit_map_iff (f.val.stalkMap _) _).mp hy + exact (isUnit_map_iff (f.stalkMap _) _).mp hy -- This actually holds for all ringed spaces with nontrivial stalks. theorem basicOpen_zero (X : LocallyRingedSpace.{u}) (U : Opens X.carrier) : diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index b94c92ea5a9a2..48bad4266f46f 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -247,6 +247,8 @@ noncomputable def coequalizer : LocallyRingedSpace where localRing x := by obtain ⟨y, rfl⟩ := (TopCat.epi_iff_surjective (coequalizer.π f.val g.val).base).mp inferInstance x + -- TODO: this instance was found automatically before #6045 + have _ : IsLocalRingHom ((coequalizer.π f.val g.val).stalkMap y) := inferInstance exact ((coequalizer.π f.val g.val : _).stalkMap y).domain_localRing /-- The explicit coequalizer cofork of locally ringed spaces. -/ @@ -275,9 +277,12 @@ noncomputable def coequalizerCoforkIsColimit : IsColimit (coequalizerCofork f g) -- but this is no longer possible set h := _ change IsLocalRingHom h - suffices IsLocalRingHom (((coequalizerCofork f g).π.val.stalkMap _).comp h) from - isLocalRingHom_of_comp _ ((coequalizerCofork f g).π.val.stalkMap _) - change IsLocalRingHom (_ ≫ (coequalizerCofork f g).π.val.stalkMap y) + suffices _ : IsLocalRingHom (((coequalizerCofork f g).π.1.stalkMap _).comp h) by + apply isLocalRingHom_of_comp _ ((coequalizerCofork f g).π.1.stalkMap _) + -- note to reviewers: this `change` is now more brittle because it now has to fully resolve + -- the type to be able to search for `MonoidHomClass`, even though of course all homs in + -- `CommRingCat` are clearly such + change IsLocalRingHom (h ≫ (coequalizerCofork f g).π.val.stalkMap y) erw [← PresheafedSpace.stalkMap.comp] apply isLocalRingHom_stalkMap_congr _ _ (coequalizer.π_desc s.π.1 e).symm y infer_instance diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean index e4395a78b6708..0baaf07ca911e 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean @@ -77,7 +77,15 @@ lemma Γevaluation_ne_zero_iff_mem_basicOpen (x : X) (f : X.presheaf.obj (op ⊤ X.Γevaluation x f ≠ 0 ↔ x ∈ X.toRingedSpace.basicOpen f := evaluation_ne_zero_iff_mem_basicOpen X ⟨x, show x ∈ ⊤ by trivial⟩ f -variable {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) +variable {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) + +-- TODO: This instance is found before #6045. +-- We need this strange instance for `residueFieldMap`, the type of `F` must be fixed +-- like this. The instance `IsLocalRingHom (f.stalkMap x)` already exists, but does not work for +-- `residueFieldMap`. +instance : IsLocalRingHom (F := Y.presheaf.stalk (f.val.base x) →+* X.presheaf.stalk x) + (f.stalkMap x) := + f.2 x /-- If `X ⟶ Y` is a morphism of locally ringed spaces and `x` a point of `X`, we obtain a morphism of residue fields in the other direction. -/ @@ -101,6 +109,8 @@ lemma residueFieldMap_comp {Z : LocallyRingedSpace.{u}} (g : Y ⟶ Z) (x : X) : simp only [comp_val, SheafedSpace.comp_base, Function.comp_apply, residueFieldMap] simp_rw [stalkMap_comp] haveI : IsLocalRingHom (g.stalkMap (f.val.base x)) := inferInstance + -- TODO: This instance is found before #6045. + haveI : IsLocalRingHom (f.stalkMap x) := inferInstance apply LocalRing.ResidueField.map_comp @[reassoc] diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index a966c23b3b918..011fbc8d6b3f1 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -946,6 +946,8 @@ instance mono : Mono f := instance : SheafedSpace.IsOpenImmersion (LocallyRingedSpace.forgetToSheafedSpace.map f) := H +-- note to reviewers: is there a `count_heartbeats` for this? +set_option synthInstance.maxHeartbeats 30000 in /-- An explicit pullback cone over `cospan f g` if `f` is an open immersion. -/ def pullbackConeOfLeft : PullbackCone f g := by refine PullbackCone.mk ?_ @@ -964,6 +966,7 @@ def pullbackConeOfLeft : PullbackCone f g := by instance : LocallyRingedSpace.IsOpenImmersion (pullbackConeOfLeft f g).snd := show PresheafedSpace.IsOpenImmersion (Y.toPresheafedSpace.ofRestrict _) by infer_instance +set_option synthInstance.maxHeartbeats 80000 in /-- The constructed `pullbackConeOfLeft` is indeed limiting. -/ def pullbackConeOfLeftIsLimit : IsLimit (pullbackConeOfLeft f g) := PullbackCone.isLimitAux' _ fun s => by @@ -982,7 +985,8 @@ def pullbackConeOfLeftIsLimit : IsLimit (pullbackConeOfLeft f g) := change _ = _ ≫ s.snd.1.stalkMap x at this rw [PresheafedSpace.stalkMap.comp, ← IsIso.eq_inv_comp] at this rw [this] - infer_instance + -- TODO: This instance is found by `infer_instance` before #6045. + apply CommRingCat.isLocalRingHom_comp · intro m _ h₂ rw [← cancel_mono (pullbackConeOfLeft f g).snd] exact h₂.trans <| LocallyRingedSpace.Hom.ext diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean index dd988d85c537d..43d40b52b040c 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean @@ -58,7 +58,7 @@ attribute [coe] PresheafedSpace.carrier -- Porting note: we add this instance, as Lean does not reliably use the `CoeOut` instance above -- in downstream files. -instance : CoeSort (PresheafedSpace C) Type* where coe := fun X => X.carrier +instance : CoeSort (PresheafedSpace C) Type* where coe X := X.carrier -- Porting note: the following lemma is removed because it is a syntactic tauto /-@[simp] diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index 4d9076090a048..2bd1fd993c02a 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -334,9 +334,9 @@ def ιInvAppπApp {i : D.J} (U : Opens (D.U i).carrier) (j) : rw [Set.preimage_preimage] change (D.f j k ≫ 𝖣.ι j).base ⁻¹' _ = _ -- Porting note: used to be `congr 3` - refine congr_arg (· ⁻¹' _) ?_ - convert congr_arg (ContinuousMap.toFun (α := D.V ⟨j, k⟩) (β := D.glued) ·) ?_ - refine congr_arg (PresheafedSpace.Hom.base (C := C) ·) ?_ + suffices D.f j k ≫ D.ι j = colimit.ι D.diagram.multispan (WalkingMultispan.left (j, k)) by + rw [this] + rfl exact colimit.w 𝖣.diagram.multispan (WalkingMultispan.Hom.fst (j, k)) · exact D.opensImagePreimageMap i j U @@ -439,23 +439,23 @@ theorem π_ιInvApp_π (i j : D.J) (U : Opens (D.U i).carrier) : rw [← @cancel_mono (f := (componentwiseDiagram 𝖣.diagram.multispan _).map (Quiver.Hom.op (WalkingMultispan.Hom.snd (i, j))) ≫ 𝟙 _) ..] - simp_rw [Category.assoc] - rw [limit.w_assoc] - erw [limit.lift_π_assoc] - rw [Category.comp_id, Category.comp_id] - change _ ≫ _ ≫ (_ ≫ _) ≫ _ = _ - rw [congr_app (D.t_id _), id_c_app] - simp_rw [Category.assoc] - rw [← Functor.map_comp_assoc] - -- Porting note (#11224): change `rw` to `erw` - erw [IsOpenImmersion.inv_naturality_assoc] - erw [IsOpenImmersion.app_invApp_assoc] - iterate 3 rw [← Functor.map_comp_assoc] - rw [NatTrans.naturality_assoc] - erw [← (D.V (i, j)).presheaf.map_comp] - convert - limit.w (componentwiseDiagram 𝖣.diagram.multispan _) - (Quiver.Hom.op (WalkingMultispan.Hom.fst (i, j))) + · simp_rw [Category.assoc] + rw [limit.w_assoc] + erw [limit.lift_π_assoc] + rw [Category.comp_id, Category.comp_id] + change _ ≫ _ ≫ (_ ≫ _) ≫ _ = _ + rw [congr_app (D.t_id _), id_c_app] + simp_rw [Category.assoc] + rw [← Functor.map_comp_assoc] + -- Porting note (#11224): change `rw` to `erw` + erw [IsOpenImmersion.inv_naturality_assoc] + erw [IsOpenImmersion.app_invApp_assoc] + iterate 3 rw [← Functor.map_comp_assoc] + rw [NatTrans.naturality_assoc] + erw [← (D.V (i, j)).presheaf.map_comp] + convert + limit.w (componentwiseDiagram 𝖣.diagram.multispan _) + (Quiver.Hom.op (WalkingMultispan.Hom.fst (i, j))) · rw [Category.comp_id] apply (config := { allowSynthFailures := true }) mono_comp change Mono ((_ ≫ D.f j i).c.app _) diff --git a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean index 637e73b1c7aaa..11a937b20ac78 100644 --- a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean @@ -47,7 +47,7 @@ namespace SheafedSpace instance coeCarrier : CoeOut (SheafedSpace C) TopCat where coe X := X.carrier instance coeSort : CoeSort (SheafedSpace C) Type* where - coe := fun X => X.1 + coe X := X.1 /-- Extract the `sheaf C (X : Top)` from a `SheafedSpace C`. -/ def sheaf (X : SheafedSpace C) : Sheaf C (X : TopCat) := diff --git a/Mathlib/GroupTheory/Congruence/Basic.lean b/Mathlib/GroupTheory/Congruence/Basic.lean index 59ad69ed64e0d..fe6cfe5261d0a 100644 --- a/Mathlib/GroupTheory/Congruence/Basic.lean +++ b/Mathlib/GroupTheory/Congruence/Basic.lean @@ -114,7 +114,7 @@ instance : Inhabited (Con M) := @[to_additive "A coercion from an additive congruence relation to its underlying binary relation."] instance : FunLike (Con M) M (M → Prop) where coe c := c.r - coe_injective' := fun x y h => by + coe_injective' x y h := by rcases x with ⟨⟨x, _⟩, _⟩ rcases y with ⟨⟨y, _⟩, _⟩ have : x = y := h diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index 6143225812daa..75ec037119a2a 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -402,6 +402,21 @@ theorem out_eq' (a : α ⧸ s) : mk a.out' = a := variable (s) +/-- Given a subgroup `s`, the function that sends a subgroup `t` to the pair consisting of +its intersection with `s` and its image in the quotient `α ⧸ s` is strictly monotone, even though +it is not injective in general. -/ +@[to_additive QuotientAddGroup.strictMono_comap_prod_image "Given an additive subgroup `s`, +the function that sends an additive subgroup `t` to the pair consisting of +its intersection with `s` and its image in the quotient `α ⧸ s` +is strictly monotone, even though it is not injective in general."] +theorem strictMono_comap_prod_image : + StrictMono fun t : Subgroup α ↦ (t.comap s.subtype, mk (s := s) '' t) := by + refine fun t₁ t₂ h ↦ ⟨⟨Subgroup.comap_mono h.1, Set.image_mono h.1⟩, + mt (fun ⟨le1, le2⟩ a ha ↦ ?_) h.2⟩ + obtain ⟨a', h', eq⟩ := le2 ⟨_, ha, rfl⟩ + convert ← t₁.mul_mem h' (@le1 ⟨_, QuotientGroup.eq.1 eq⟩ <| t₂.mul_mem (t₂.inv_mem <| h.1 h') ha) + apply mul_inv_cancel_left + /- It can be useful to write `obtain ⟨h, H⟩ := mk_out'_eq_mul ...`, and then `rw [H]` or `simp_rw [H]` or `simp only [H]`. In order for `simp_rw` and `simp only` to work, this lemma is stated in terms of an arbitrary `h : s`, rather than the specific `h = g⁻¹ * (mk g).out'`. -/ diff --git a/Mathlib/GroupTheory/CosetCover.lean b/Mathlib/GroupTheory/CosetCover.lean index c9cefdcb06558..03164c5973c5e 100644 --- a/Mathlib/GroupTheory/CosetCover.lean +++ b/Mathlib/GroupTheory/CosetCover.lean @@ -343,7 +343,7 @@ theorem exists_index_le_card_of_leftCoset_cover : | inl hindex => rwa [hindex, Nat.cast_zero, inv_zero, inv_pos, Nat.cast_pos] | inr hindex => - exact inv_lt_inv_of_lt (by exact_mod_cast hs') (by exact_mod_cast h i hi ⟨hindex⟩) + exact inv_strictAnti₀ (by exact_mod_cast hs') (by exact_mod_cast h i hi ⟨hindex⟩) apply (Finset.sum_lt_sum_of_nonempty hs hlt).trans_eq rw [Finset.sum_const, nsmul_eq_mul, mul_inv_cancel₀ (Nat.cast_ne_zero.mpr hs'.ne')] diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 7bf00ac53f9d3..3cd5bab40dafd 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -402,14 +402,8 @@ noncomputable def equivSubgroupOrbitsQuotientGroup [IsPretransitive α β] rw [Quotient.eq'', leftRel_eq] simp only convert one_mem H - rw [inv_mul_eq_one, eq_comm, ← inv_mul_eq_one, ← Subgroup.mem_bot, ← free (g⁻¹ • x), + · rw [inv_mul_eq_one, eq_comm, ← inv_mul_eq_one, ← Subgroup.mem_bot, ← free (g⁻¹ • x), mem_stabilizer_iff, mul_smul, (exists_smul_eq α (g⁻¹ • x) x).choose_spec] - #adaptation_note - /-- - After https://github.com/leanprover/lean4/pull/5376 we need to search for this instance explicitly. - TODO: change `convert` to more agressively solve such goals with `infer_instance` itself. - -/ - infer_instance end MulAction diff --git a/Mathlib/GroupTheory/HNNExtension.lean b/Mathlib/GroupTheory/HNNExtension.lean index d4c6a1677d248..2ba02b333a426 100644 --- a/Mathlib/GroupTheory/HNNExtension.lean +++ b/Mathlib/GroupTheory/HNNExtension.lean @@ -174,7 +174,7 @@ namespace NormalWord variable (G A B) /-- To put word in the HNN Extension into a normal form, we must choose an element of each right coset of both `A` and `B`, such that the chosen element of the subgroup itself is `1`. -/ -structure TransversalPair : Type _ := +structure TransversalPair : Type _ where /-- The transversal of each subgroup -/ set : ℤˣ → Set G /-- We have exactly one element of each coset of the subgroup -/ @@ -187,7 +187,7 @@ instance TransversalPair.nonempty : Nonempty (TransversalPair G A B) := by /-- A reduced word is a `head`, which is an element of `G`, followed by the product list of pairs. There should also be no sequences of the form `t^u * g * t^-u`, where `g` is in `toSubgroup A B u` This is a less strict condition than required for `NormalWord`. -/ -structure ReducedWord : Type _ := +structure ReducedWord : Type _ where /-- Every `ReducedWord` is the product of an element of the group and a word made up of letters each of which is in the transversal. `head` is that element of the base group. -/ head : G @@ -215,7 +215,7 @@ The normal form is a `head`, which is an element of `G`, followed by the product `toSubgroup A B u`. There should also be no sequences of the form `t^u * g * t^-u` where `g ∈ toSubgroup A B u` -/ structure _root_.HNNExtension.NormalWord (d : TransversalPair G A B) - extends ReducedWord G A B : Type _ := + extends ReducedWord G A B : Type _ where /-- Every element `g : G` in the list is the chosen element of its coset -/ mem_set : ∀ (u : ℤˣ) (g : G), (u, g) ∈ toList → g ∈ d.set u diff --git a/Mathlib/GroupTheory/Perm/DomMulAct.lean b/Mathlib/GroupTheory/Perm/DomMulAct.lean index 34ebef212fe7d..4bfd4c758ab0d 100644 --- a/Mathlib/GroupTheory/Perm/DomMulAct.lean +++ b/Mathlib/GroupTheory/Perm/DomMulAct.lean @@ -88,14 +88,14 @@ lemma stabilizerMulEquiv_apply (g : (stabilizer (Perm α)ᵈᵐᵃ f)ᵐᵒᵖ) section Fintype -variable [Fintype α] [DecidableEq α] [DecidableEq ι] +variable [Fintype α] open Nat variable (f) /-- The cardinality of the type of permutations preserving a function -/ -theorem stabilizer_card [Fintype ι] : +theorem stabilizer_card [DecidableEq α] [DecidableEq ι] [Fintype ι] : Fintype.card {g : Perm α // f ∘ g = f} = ∏ i, (Fintype.card {a // f a = i})! := by -- rewriting via Nat.card because Fintype instance is not found rw [← Nat.card_eq_fintype_card, @@ -108,9 +108,12 @@ theorem stabilizer_card [Fintype ι] : /-- The cardinality of the set of permutations preserving a function -/ theorem stabilizer_ncard [Fintype ι] : Set.ncard {g : Perm α | f ∘ g = f} = ∏ i, (Set.ncard {a | f a = i})! := by + classical simp only [← Set.Nat.card_coe_set_eq, Set.coe_setOf, card_eq_fintype_card] exact stabilizer_card f +variable [DecidableEq α] [DecidableEq ι] + /-- The cardinality of the type of permutations preserving a function (without the finiteness assumption on target)-/ theorem stabilizer_card': diff --git a/Mathlib/GroupTheory/QuotientGroup/Basic.lean b/Mathlib/GroupTheory/QuotientGroup/Basic.lean index 7857eb080a61d..ec446e552f0d8 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Basic.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Basic.lean @@ -167,6 +167,11 @@ theorem mk_prod {G ι : Type*} [CommGroup G] (N : Subgroup G) (s : Finset ι) {f @[to_additive (attr := simp)] lemma map_mk'_self : N.map (mk' N) = ⊥ := by aesop +@[to_additive QuotientAddGroup.strictMono_comap_prod_map] +theorem strictMono_comap_prod_map : + StrictMono fun H : Subgroup G ↦ (H.comap N.subtype, H.map (mk' N)) := + strictMono_comap_prod_image N + /-- A group homomorphism `φ : G →* M` with `N ⊆ ker(φ)` descends (i.e. `lift`s) to a group homomorphism `G/N →* M`. -/ @[to_additive "An `AddGroup` homomorphism `φ : G →+ M` with `N ⊆ ker(φ)` descends (i.e. `lift`s) diff --git a/Mathlib/GroupTheory/Torsion.lean b/Mathlib/GroupTheory/Torsion.lean index 1fe565875fe91..08d223f5db3ba 100644 --- a/Mathlib/GroupTheory/Torsion.lean +++ b/Mathlib/GroupTheory/Torsion.lean @@ -297,7 +297,7 @@ end CommGroup end CommGroup namespace Monoid - +section Monoid variable (G) [Monoid G] /-- A predicate on a monoid saying that only 1 is of finite order. -/ @@ -326,22 +326,17 @@ lemma isTorsionFree_iff_torsion_eq_bot {G} [CommGroup G] : end Monoid section Group - -open Monoid - variable [Group G] /-- A nontrivial torsion group is not torsion-free. -/ -@[to_additive AddMonoid.IsTorsion.not_torsion_free - "A nontrivial additive torsion group is not torsion-free."] +@[to_additive "A nontrivial additive torsion group is not torsion-free."] theorem IsTorsion.not_torsion_free [hN : Nontrivial G] : IsTorsion G → ¬IsTorsionFree G := fun tG => not_isTorsionFree_iff.mpr <| by obtain ⟨x, hx⟩ := (nontrivial_iff_exists_ne (1 : G)).mp hN exact ⟨x, hx, tG x⟩ /-- A nontrivial torsion-free group is not torsion. -/ -@[to_additive AddMonoid.IsTorsionFree.not_torsion - "A nontrivial torsion-free additive group is not torsion."] +@[to_additive "A nontrivial torsion-free additive group is not torsion."] theorem IsTorsionFree.not_torsion [hN : Nontrivial G] : IsTorsionFree G → ¬IsTorsion G := fun tfG => (not_isTorsion_iff _).mpr <| by obtain ⟨x, hx⟩ := (nontrivial_iff_exists_ne (1 : G)).mp hN @@ -371,8 +366,8 @@ open CommGroup (torsion) variable (G) [CommGroup G] /-- Quotienting a group by its torsion subgroup yields a torsion free group. -/ -@[to_additive AddIsTorsionFree.quotient_torsion - "Quotienting a group by its additive torsion subgroup yields an additive torsion free group."] +@[to_additive +"Quotienting a group by its additive torsion subgroup yields an additive torsion free group."] theorem IsTorsionFree.quotient_torsion : IsTorsionFree <| G ⧸ torsion G := fun g hne hfin => hne <| by induction' g using QuotientGroup.induction_on with g @@ -383,21 +378,26 @@ theorem IsTorsionFree.quotient_torsion : IsTorsionFree <| G ⧸ torsion G := fun (isOfFinOrder_iff_pow_eq_one.mpr ⟨m * n, mul_pos mpos npos, (pow_mul g m n).symm ▸ hn⟩) end CommGroup +end Monoid + +namespace AddMonoid lemma isTorsionFree_iff_noZeroSMulDivisors_nat {M : Type*} [AddMonoid M] : - AddMonoid.IsTorsionFree M ↔ NoZeroSMulDivisors ℕ M := by + IsTorsionFree M ↔ NoZeroSMulDivisors ℕ M := by simp_rw [AddMonoid.IsTorsionFree, isOfFinAddOrder_iff_nsmul_eq_zero, not_exists, not_and, pos_iff_ne_zero, noZeroSMulDivisors_iff, forall_swap (β := ℕ)] exact forall₂_congr fun _ _ ↦ by tauto lemma isTorsionFree_iff_noZeroSMulDivisors_int [AddGroup G] : - AddMonoid.IsTorsionFree G ↔ NoZeroSMulDivisors ℤ G := by + IsTorsionFree G ↔ NoZeroSMulDivisors ℤ G := by simp_rw [AddMonoid.IsTorsionFree, isOfFinAddOrder_iff_zsmul_eq_zero, not_exists, not_and, noZeroSMulDivisors_iff, forall_swap (β := ℤ)] exact forall₂_congr fun _ _ ↦ by tauto -@[deprecated (since := "2024-02-29")] -alias AddMonoid.IsTorsionFree_iff_noZeroSMulDivisors := isTorsionFree_iff_noZeroSMulDivisors_int - lemma IsTorsionFree.of_noZeroSMulDivisors {M : Type*} [AddMonoid M] [NoZeroSMulDivisors ℕ M] : - AddMonoid.IsTorsionFree M := isTorsionFree_iff_noZeroSMulDivisors_nat.2 ‹_› + IsTorsionFree M := isTorsionFree_iff_noZeroSMulDivisors_nat.2 ‹_› + +alias ⟨IsTorsionFree.noZeroSMulDivisors_nat, _⟩ := isTorsionFree_iff_noZeroSMulDivisors_nat +alias ⟨IsTorsionFree.noZeroSMulDivisors_int, _⟩ := isTorsionFree_iff_noZeroSMulDivisors_int + +end AddMonoid diff --git a/Mathlib/Init/Algebra/Classes.lean b/Mathlib/Init/Algebra/Classes.lean index c7be944592e5c..bb4475917a4fd 100644 --- a/Mathlib/Init/Algebra/Classes.lean +++ b/Mathlib/Init/Algebra/Classes.lean @@ -25,7 +25,7 @@ set_option linter.deprecated false universe u v -variable {α : Sort u} {β : Sort v} +variable {α : Sort u} @[deprecated (since := "2024-09-11")] class IsLeftCancel (α : Sort u) (op : α → α → α) : Prop where diff --git a/Mathlib/Lean/Expr/Basic.lean b/Mathlib/Lean/Expr/Basic.lean index add7ca2edafdb..5991b42748293 100644 --- a/Mathlib/Lean/Expr/Basic.lean +++ b/Mathlib/Lean/Expr/Basic.lean @@ -324,6 +324,13 @@ otherwise, it returns `none`. -/ let (type, _, lhs, rhs) ← p.app4? ``LE.le return (type, lhs, rhs) +/-- `Lean.Expr.lt? e` takes `e : Expr` as input. +If `e` represents `a < b`, then it returns `some (t, a, b)`, where `t` is the Type of `a`, +otherwise, it returns `none`. -/ +@[inline] def lt? (p : Expr) : Option (Expr × Expr × Expr) := do + let (type, _, lhs, rhs) ← p.app4? ``LT.lt + return (type, lhs, rhs) + /-- Given a proposition `ty` that is an `Eq`, `Iff`, or `HEq`, returns `(tyLhs, lhs, tyRhs, rhs)`, where `lhs : tyLhs` and `rhs : tyRhs`, and where `lhs` is related to `rhs` by the respective relation. diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean index b66ea39b78a5a..eb8a41bbe9312 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean @@ -565,6 +565,13 @@ theorem lineMap_vsub_lineMap (p₁ p₂ p₃ p₄ : P1) (c : k) : lineMap p₁ p₂ c -ᵥ lineMap p₃ p₄ c = lineMap (p₁ -ᵥ p₃) (p₂ -ᵥ p₄) c := ((fst : P1 × P1 →ᵃ[k] P1) -ᵥ (snd : P1 × P1 →ᵃ[k] P1)).apply_lineMap (_, _) (_, _) c +@[simp] lemma lineMap_lineMap_right (p₀ p₁ : P1) (c d : k) : + lineMap p₀ (lineMap p₀ p₁ c) d = lineMap p₀ p₁ (d * c) := by simp [lineMap_apply, mul_smul] + +@[simp] lemma lineMap_lineMap_left (p₀ p₁ : P1) (c d : k) : + lineMap (lineMap p₀ p₁ c) p₁ d = lineMap p₀ p₁ (1 - (1 - d) * (1 - c)) := by + simp_rw [lineMap_apply_one_sub, ← lineMap_apply_one_sub p₁, lineMap_lineMap_right] + /-- Decomposition of an affine map in the special case when the point space and vector space are the same. -/ theorem decomp (f : V1 →ᵃ[k] V2) : (f : V1 → V2) = ⇑f.linear + fun _ => f 0 := by diff --git a/Mathlib/LinearAlgebra/AffineSpace/Ordered.lean b/Mathlib/LinearAlgebra/AffineSpace/Ordered.lean index bc28ee54454cb..9f3c4c81c015f 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Ordered.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Ordered.lean @@ -134,15 +134,13 @@ theorem lineMap_le_right_iff_le (h : r < 1) : lineMap a b r ≤ b ↔ a ≤ b := Iff.trans (by rw [lineMap_apply_one]) (lineMap_le_lineMap_iff_of_lt h) @[simp] -theorem midpoint_le_right : midpoint k a b ≤ b ↔ a ≤ b := - lineMap_le_right_iff_le <| inv_lt_one one_lt_two +theorem midpoint_le_right : midpoint k a b ≤ b ↔ a ≤ b := lineMap_le_right_iff_le two_inv_lt_one theorem right_le_lineMap_iff_le (h : r < 1) : b ≤ lineMap a b r ↔ b ≤ a := lineMap_le_right_iff_le (E := Eᵒᵈ) h @[simp] -theorem right_le_midpoint : b ≤ midpoint k a b ↔ b ≤ a := - right_le_lineMap_iff_le <| inv_lt_one one_lt_two +theorem right_le_midpoint : b ≤ midpoint k a b ↔ b ≤ a := right_le_lineMap_iff_le two_inv_lt_one end diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index 1c884e7f3cd86..007a27d357368 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -88,7 +88,7 @@ section Coercions instance instFunLike : FunLike (M [⋀^ι]→ₗ[R] N) (ι → M) N where coe f := f.toFun - coe_injective' := fun f g h ↦ by + coe_injective' f g h := by rcases f with ⟨⟨_, _, _⟩, _⟩ rcases g with ⟨⟨_, _, _⟩, _⟩ congr diff --git a/Mathlib/LinearAlgebra/Basis/VectorSpace.lean b/Mathlib/LinearAlgebra/Basis/VectorSpace.lean index 5ffcc1cc36b43..e7a219e569dab 100644 --- a/Mathlib/LinearAlgebra/Basis/VectorSpace.lean +++ b/Mathlib/LinearAlgebra/Basis/VectorSpace.lean @@ -86,6 +86,61 @@ theorem subset_extend {s : Set V} (hs : LinearIndependent K ((↑) : s → V)) : s ⊆ hs.extend (Set.subset_univ _) := hs.subset_extend _ +/-- If `s` is a family of linearly independent vectors contained in a set `t` spanning `V`, +then one can get a basis of `V` containing `s` and contained in `t`. -/ +noncomputable def extendLe (hs : LinearIndependent K ((↑) : s → V)) + (hst : s ⊆ t) (ht : ⊤ ≤ span K t) : + Basis (hs.extend hst) K V := + Basis.mk + (@LinearIndependent.restrict_of_comp_subtype _ _ _ id _ _ _ _ (hs.linearIndependent_extend _)) + (le_trans ht <| Submodule.span_le.2 <| by simpa using hs.subset_span_extend hst) + +theorem extendLe_apply_self (hs : LinearIndependent K ((↑) : s → V)) + (hst : s ⊆ t) (ht : ⊤ ≤ span K t) (x : hs.extend hst) : + Basis.extendLe hs hst ht x = x := + Basis.mk_apply _ _ _ + +@[simp] +theorem coe_extendLe (hs : LinearIndependent K ((↑) : s → V)) + (hst : s ⊆ t) (ht : ⊤ ≤ span K t) : ⇑(Basis.extendLe hs hst ht) = ((↑) : _ → _) := + funext (extendLe_apply_self hs hst ht) + +theorem range_extendLe (hs : LinearIndependent K ((↑) : s → V)) + (hst : s ⊆ t) (ht : ⊤ ≤ span K t) : + range (Basis.extendLe hs hst ht) = hs.extend hst := by + rw [coe_extendLe, Subtype.range_coe_subtype, setOf_mem_eq] + +theorem subset_extendLe (hs : LinearIndependent K ((↑) : s → V)) + (hst : s ⊆ t) (ht : ⊤ ≤ span K t) : + s ⊆ range (Basis.extendLe hs hst ht) := + (range_extendLe hs hst ht).symm ▸ hs.subset_extend hst + +theorem extendLe_subset (hs : LinearIndependent K ((↑) : s → V)) + (hst : s ⊆ t) (ht : ⊤ ≤ span K t) : + range (Basis.extendLe hs hst ht) ⊆ t := + (range_extendLe hs hst ht).symm ▸ hs.extend_subset hst + +/-- If a set `s` spans the space, this is a basis contained in `s`. -/ +noncomputable def ofSpan (hs : ⊤ ≤ span K s) : + Basis ((linearIndependent_empty K V).extend (empty_subset s)) K V := + extendLe (linearIndependent_empty K V) (empty_subset s) hs + +theorem ofSpan_apply_self (hs : ⊤ ≤ span K s) + (x : (linearIndependent_empty K V).extend (empty_subset s)) : + Basis.ofSpan hs x = x := + extendLe_apply_self (linearIndependent_empty K V) (empty_subset s) hs x + +@[simp] +theorem coe_ofSpan (hs : ⊤ ≤ span K s) : ⇑(ofSpan hs) = ((↑) : _ → _) := + funext (ofSpan_apply_self hs) + +theorem range_ofSpan (hs : ⊤ ≤ span K s) : + range (ofSpan hs) = (linearIndependent_empty K V).extend (empty_subset s) := by + rw [coe_ofSpan, Subtype.range_coe_subtype, setOf_mem_eq] + +theorem ofSpan_subset (hs : ⊤ ≤ span K s) : range (ofSpan hs) ⊆ s := + extendLe_subset (linearIndependent_empty K V) (empty_subset s) hs + section variable (K V) diff --git a/Mathlib/LinearAlgebra/BilinearMap.lean b/Mathlib/LinearAlgebra/BilinearMap.lean index b9e49faf0456c..10c008423626e 100644 --- a/Mathlib/LinearAlgebra/BilinearMap.lean +++ b/Mathlib/LinearAlgebra/BilinearMap.lean @@ -37,16 +37,15 @@ variable {R : Type*} [Semiring R] {S : Type*} [Semiring S] variable {R₂ : Type*} [Semiring R₂] {S₂ : Type*} [Semiring S₂] variable {M : Type*} {N : Type*} {P : Type*} variable {M₂ : Type*} {N₂ : Type*} {P₂ : Type*} -variable {Nₗ : Type*} {Pₗ : Type*} -variable {M' : Type*} {N' : Type*} {P' : Type*} +variable {Pₗ : Type*} +variable {M' : Type*} {P' : Type*} variable [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid P] -variable [AddCommMonoid M₂] [AddCommMonoid N₂] [AddCommMonoid P₂] -variable [AddCommMonoid Nₗ] [AddCommMonoid Pₗ] -variable [AddCommGroup M'] [AddCommGroup N'] [AddCommGroup P'] +variable [AddCommMonoid M₂] [AddCommMonoid N₂] [AddCommMonoid P₂] [AddCommMonoid Pₗ] +variable [AddCommGroup M'] [AddCommGroup P'] variable [Module R M] [Module S N] [Module R₂ P] [Module S₂ P] variable [Module R M₂] [Module S N₂] [Module R P₂] [Module S₂ P₂] variable [Module R Pₗ] [Module S Pₗ] -variable [Module R M'] [Module S N'] [Module R₂ P'] [Module S₂ P'] +variable [Module R M'] [Module R₂ P'] [Module S₂ P'] variable [SMulCommClass S₂ R₂ P] [SMulCommClass S R Pₗ] [SMulCommClass S₂ R₂ P'] variable [SMulCommClass S₂ R P₂] variable {ρ₁₂ : R →+* R₂} {σ₁₂ : S →+* S₂} @@ -376,14 +375,11 @@ end CommSemiring section CommRing -variable {R R₂ S S₂ M N P : Type*} -variable {Mₗ Nₗ Pₗ : Type*} -variable [CommRing R] [CommRing S] [CommRing R₂] [CommRing S₂] +variable {R M : Type*} [CommRing R] section AddCommGroup -variable [AddCommGroup M] [AddCommGroup N] [AddCommGroup P] -variable [Module R M] [Module S N] [Module R₂ P] [Module S₂ P] +variable [AddCommGroup M] [Module R M] theorem lsmul_injective [NoZeroSMulDivisors R M] {x : R} (hx : x ≠ 0) : Function.Injective (lsmul R M x) := diff --git a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean index 12001f79238c5..26dc00965d15b 100644 --- a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean @@ -430,6 +430,14 @@ theorem rank_lt_aleph0 [Module.Finite R M] : Module.rank R M < ℵ₀ := by refine (ciSup_le' fun i => ?_).trans_lt (nat_lt_aleph0 S.card) exact linearIndependent_le_span_finset _ i.prop S hS +noncomputable instance {R M : Type*} [DivisionRing R] [AddCommGroup M] [Module R M] + {s t : Set M} [Module.Finite R (span R t)] + (hs : LinearIndependent R ((↑) : s → M)) (hst : s ⊆ t) : + Fintype (hs.extend hst) := by + refine Classical.choice (Cardinal.lt_aleph0_iff_fintype.1 ?_) + rw [← rank_span_set (hs.linearIndependent_extend hst), hs.span_extend_eq_span] + exact Module.rank_lt_aleph0 .. + /-- If `M` is finite, `finrank M = rank M`. -/ @[simp] theorem finrank_eq_rank [Module.Finite R M] : ↑(finrank R M) = Module.rank R M := by diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 2b1fa24c031b5..5d556906e958a 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -580,6 +580,7 @@ class IsReflexive : Prop where lemma bijective_dual_eval [IsReflexive R M] : Bijective (Dual.eval R M) := IsReflexive.bijective_dual_eval' +/-- See also `Module.instFiniteDimensionalOfIsReflexive` for the converse over a field. -/ instance IsReflexive.of_finite_of_free [Module.Finite R M] [Free R M] : IsReflexive R M where bijective_dual_eval' := ⟨LinearMap.ker_eq_bot.mp (Free.chooseBasis R M).eval_ker, LinearMap.range_eq_top.mp (Free.chooseBasis R M).eval_range⟩ @@ -650,6 +651,22 @@ instance _root_.MulOpposite.instModuleIsReflexive : IsReflexive R (MulOpposite M instance _root_.ULift.instModuleIsReflexive.{w} : IsReflexive R (ULift.{w} M) := equiv ULift.moduleEquiv.symm +instance instFiniteDimensionalOfIsReflexive (K V : Type*) + [Field K] [AddCommGroup V] [Module K V] [IsReflexive K V] : + FiniteDimensional K V := by + rw [FiniteDimensional, ← rank_lt_aleph0_iff] + by_contra! contra + suffices lift (Module.rank K V) < Module.rank K (Dual K (Dual K V)) by + have heq := lift_rank_eq_of_equiv_equiv (R := K) (R' := K) (M := V) (M' := Dual K (Dual K V)) + (ZeroHom.id K) (evalEquiv K V) bijective_id (fun r v ↦ (evalEquiv K V).map_smul _ _) + rw [← lift_umax, heq, lift_id'] at this + exact lt_irrefl _ this + have h₁ : lift (Module.rank K V) < Module.rank K (Dual K V) := lift_rank_lt_rank_dual contra + have h₂ : Module.rank K (Dual K V) < Module.rank K (Dual K (Dual K V)) := by + convert lift_rank_lt_rank_dual <| le_trans (by simpa) h₁.le + rw [lift_id'] + exact lt_trans h₁ h₂ + end IsReflexive end Module @@ -673,6 +690,47 @@ theorem exists_dual_map_eq_bot_of_lt_top (hp : p < ⊤) (hp' : Free R (M ⧸ p)) obtain ⟨f, hf, hf'⟩ := p.exists_dual_map_eq_bot_of_nmem hx hp' exact ⟨f, by aesop, hf'⟩ +variable {ι 𝕜 E : Type*} [Field 𝕜] [AddCommGroup E] [Module 𝕜 E] + +open LinearMap Set FiniteDimensional + +theorem _root_.FiniteDimensional.mem_span_of_iInf_ker_le_ker [FiniteDimensional 𝕜 E] + {L : ι → E →ₗ[𝕜] 𝕜} {K : E →ₗ[𝕜] 𝕜} + (h : ⨅ i, LinearMap.ker (L i) ≤ ker K) : K ∈ span 𝕜 (range L) := by + by_contra hK + rcases exists_dual_map_eq_bot_of_nmem hK inferInstance with ⟨φ, φne, hφ⟩ + let φs := (Module.evalEquiv 𝕜 E).symm φ + have : K φs = 0 := by + refine h <| (Submodule.mem_iInf _).2 fun i ↦ (mem_bot 𝕜).1 ?_ + rw [← hφ, Submodule.mem_map] + exact ⟨L i, Submodule.subset_span ⟨i, rfl⟩, (apply_evalEquiv_symm_apply 𝕜 E _ φ).symm⟩ + simp only [apply_evalEquiv_symm_apply, φs, φne] at this + +/-- Given some linear forms $L_1, ..., L_n, K$ over a vector space $E$, if +$\bigcap_{i=1}^n \mathrm{ker}(L_i) \subseteq \mathrm{ker}(K)$, then $K$ is in the space generated +by $L_1, ..., L_n$. -/ +theorem _root_.mem_span_of_iInf_ker_le_ker [Finite ι] {L : ι → E →ₗ[𝕜] 𝕜} {K : E →ₗ[𝕜] 𝕜} + (h : ⨅ i, ker (L i) ≤ ker K) : K ∈ span 𝕜 (range L) := by + have _ := Fintype.ofFinite ι + let φ : E →ₗ[𝕜] ι → 𝕜 := LinearMap.pi L + let p := ⨅ i, ker (L i) + have p_eq : p = ker φ := (ker_pi L).symm + let ψ : (E ⧸ p) →ₗ[𝕜] ι → 𝕜 := p.liftQ φ p_eq.le + have _ : FiniteDimensional 𝕜 (E ⧸ p) := of_injective ψ (ker_eq_bot.1 (ker_liftQ_eq_bot' p φ p_eq)) + let L' i : (E ⧸ p) →ₗ[𝕜] 𝕜 := p.liftQ (L i) (iInf_le _ i) + let K' : (E ⧸ p) →ₗ[𝕜] 𝕜 := p.liftQ K h + have : ⨅ i, ker (L' i) ≤ ker K' := by + simp_rw [← ker_pi, L', pi_liftQ_eq_liftQ_pi, ker_liftQ_eq_bot' p φ p_eq] + exact bot_le + obtain ⟨c, hK'⟩ := + (mem_span_range_iff_exists_fun 𝕜).1 (FiniteDimensional.mem_span_of_iInf_ker_le_ker this) + refine (mem_span_range_iff_exists_fun 𝕜).2 ⟨c, ?_⟩ + conv_lhs => enter [2]; intro i; rw [← p.liftQ_mkQ (L i) (iInf_le _ i)] + rw [← p.liftQ_mkQ K h] + ext x + convert LinearMap.congr_fun hK' (p.mkQ x) + simp only [coeFn_sum, Finset.sum_apply, smul_apply, coe_comp, Function.comp_apply, smul_eq_mul] + end Submodule section DualBases @@ -1166,12 +1224,12 @@ def dualCopairing (W : Submodule R M) : W.dualAnnihilator →ₗ[R] M ⧸ W → exact (mem_dualAnnihilator φ).mp hφ w hw) -- Porting note: helper instance -instance (W : Submodule R M) : FunLike (W.dualAnnihilator) M R := - { coe := fun φ => φ.val, - coe_injective' := fun φ ψ h => by - ext - simp only [Function.funext_iff] at h - exact h _ } +instance (W : Submodule R M) : FunLike (W.dualAnnihilator) M R where + coe φ := φ.val + coe_injective' φ ψ h := by + ext + simp only [Function.funext_iff] at h + exact h _ @[simp] theorem dualCopairing_apply {W : Submodule R M} (φ : W.dualAnnihilator) (x : M) : diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index 3effed50341ef..e2e7aa9fcb4c9 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -58,227 +58,525 @@ open Module Set variable {K R : Type v} {V M : Type w} [CommRing R] [AddCommGroup M] [Module R M] [Field K] [AddCommGroup V] [Module K V] +/-- The submodule `unifEigenspace f μ k` for a linear map `f`, a scalar `μ`, +and a number `k : ℕ∞` is the kernel of `(f - μ • id) ^ k` if `k` is a natural number, +or the union of all these kernels if `k = ∞`. -/ +def unifEigenspace (f : End R M) (μ : R) : ℕ∞ →o Submodule R M where + toFun k := ⨆ l : ℕ, ⨆ _ : l ≤ k, LinearMap.ker ((f - μ • 1) ^ l) + monotone' _ _ hkl := biSup_mono fun _ hi ↦ hi.trans hkl + +lemma mem_unifEigenspace {f : End R M} {μ : R} {k : ℕ∞} {x : M} : + x ∈ f.unifEigenspace μ k ↔ ∃ l : ℕ, l ≤ k ∧ x ∈ LinearMap.ker ((f - μ • 1) ^ l) := by + have : Nonempty {l : ℕ // l ≤ k} := ⟨⟨0, zero_le _⟩⟩ + have : Directed (ι := { i : ℕ // i ≤ k }) (· ≤ ·) fun i ↦ LinearMap.ker ((f - μ • 1) ^ (i : ℕ)) := + Monotone.directed_le fun m n h ↦ by simpa using (f - μ • 1).iterateKer.monotone h + simp_rw [unifEigenspace, OrderHom.coe_mk, LinearMap.mem_ker, iSup_subtype', + Submodule.mem_iSup_of_directed _ this, LinearMap.mem_ker, Subtype.exists, exists_prop] + +lemma unifEigenspace_directed {f : End R M} {μ : R} {k : ℕ∞} : + Directed (· ≤ ·) (fun l : {l : ℕ // l ≤ k} ↦ f.unifEigenspace μ l) := by + have aux : Monotone ((↑) : {l : ℕ // l ≤ k} → ℕ∞) := fun x y h ↦ by simpa using h + exact ((unifEigenspace f μ).monotone.comp aux).directed_le + +lemma mem_unifEigenspace_nat {f : End R M} {μ : R} {k : ℕ} {x : M} : + x ∈ f.unifEigenspace μ k ↔ x ∈ LinearMap.ker ((f - μ • 1) ^ k) := by + rw [mem_unifEigenspace] + constructor + · rintro ⟨l, hl, hx⟩ + simp only [Nat.cast_le] at hl + exact (f - μ • 1).iterateKer.monotone hl hx + · intro hx + exact ⟨k, le_rfl, hx⟩ + +lemma mem_unifEigenspace_top {f : End R M} {μ : R} {x : M} : + x ∈ f.unifEigenspace μ ⊤ ↔ ∃ k : ℕ, x ∈ LinearMap.ker ((f - μ • 1) ^ k) := by + simp [mem_unifEigenspace] + +lemma unifEigenspace_nat {f : End R M} {μ : R} {k : ℕ} : + f.unifEigenspace μ k = LinearMap.ker ((f - μ • 1) ^ k) := by + ext; simp [mem_unifEigenspace_nat] + +lemma unifEigenspace_eq_iSup_unifEigenspace_nat (f : End R M) (μ : R) (k : ℕ∞) : + f.unifEigenspace μ k = ⨆ l : {l : ℕ // l ≤ k}, f.unifEigenspace μ l := by + simp_rw [unifEigenspace_nat, unifEigenspace, OrderHom.coe_mk, iSup_subtype] + +lemma unifEigenspace_top (f : End R M) (μ : R) : + f.unifEigenspace μ ⊤ = ⨆ k : ℕ, f.unifEigenspace μ k := by + rw [unifEigenspace_eq_iSup_unifEigenspace_nat, iSup_subtype] + simp only [le_top, iSup_pos, OrderHom.coe_mk] + +lemma unifEigenspace_one {f : End R M} {μ : R} : + f.unifEigenspace μ 1 = LinearMap.ker (f - μ • 1) := by + rw [← Nat.cast_one, unifEigenspace_nat, pow_one] + +@[simp] +lemma mem_unifEigenspace_one {f : End R M} {μ : R} {x : M} : + x ∈ f.unifEigenspace μ 1 ↔ f x = μ • x := by + rw [unifEigenspace_one, LinearMap.mem_ker, LinearMap.sub_apply, + sub_eq_zero, LinearMap.smul_apply, LinearMap.one_apply] + +-- `simp` can prove this using `unifEigenspace_zero` +lemma mem_unifEigenspace_zero {f : End R M} {μ : R} {x : M} : + x ∈ f.unifEigenspace μ 0 ↔ x = 0 := by + rw [← Nat.cast_zero, mem_unifEigenspace_nat, pow_zero, LinearMap.mem_ker, LinearMap.one_apply] + +@[simp] +lemma unifEigenspace_zero {f : End R M} {μ : R} : + f.unifEigenspace μ 0 = ⊥ := by + ext; apply mem_unifEigenspace_zero + +@[simp] +lemma unifEigenspace_zero_nat (f : End R M) (k : ℕ) : + f.unifEigenspace 0 k = LinearMap.ker (f ^ k) := by + ext; simp [mem_unifEigenspace_nat] + +/-- Let `M` be an `R`-module, and `f` an `R`-linear endomorphism of `M`, +and let `μ : R` and `k : ℕ∞` be given. +Then `x : M` satisfies `HasUnifEigenvector f μ k x` if +`x ∈ f.unifEigenspace μ k` and `x ≠ 0`. + +For `k = 1`, this means that `x` is an eigenvector of `f` with eigenvalue `μ`. -/ +def HasUnifEigenvector (f : End R M) (μ : R) (k : ℕ∞) (x : M) : Prop := + x ∈ f.unifEigenspace μ k ∧ x ≠ 0 + +/-- Let `M` be an `R`-module, and `f` an `R`-linear endomorphism of `M`. +Then `μ : R` and `k : ℕ∞` satisfy `HasUnifEigenvalue f μ k` if +`f.unifEigenspace μ k ≠ ⊥`. + +For `k = 1`, this means that `μ` is an eigenvalue of `f`. -/ +def HasUnifEigenvalue (f : End R M) (μ : R) (k : ℕ∞) : Prop := + f.unifEigenspace μ k ≠ ⊥ + +/-- Let `M` be an `R`-module, and `f` an `R`-linear endomorphism of `M`. +For `k : ℕ∞`, we define `UnifEigenvalues f k` to be the type of all +`μ : R` that satisfy `f.HasUnifEigenvalue μ k`. + +For `k = 1` this is the type of all eigenvalues of `f`. -/ +def UnifEigenvalues (f : End R M) (k : ℕ∞) : Type _ := + { μ : R // f.HasUnifEigenvalue μ k } + +/-- The underlying value of a bundled eigenvalue. -/ +@[coe] +def UnifEigenvalues.val (f : Module.End R M) (k : ℕ∞) : UnifEigenvalues f k → R := Subtype.val + +instance UnifEigenvalues.instCoeOut {f : Module.End R M} (k : ℕ∞) : + CoeOut (UnifEigenvalues f k) R where + coe := UnifEigenvalues.val f k + +instance UnivEigenvalues.instDecidableEq [DecidableEq R] (f : Module.End R M) (k : ℕ∞) : + DecidableEq (UnifEigenvalues f k) := + inferInstanceAs (DecidableEq (Subtype (fun x : R ↦ f.HasUnifEigenvalue x k))) + +lemma HasUnifEigenvector.hasUnifEigenvalue {f : End R M} {μ : R} {k : ℕ∞} {x : M} + (h : f.HasUnifEigenvector μ k x) : f.HasUnifEigenvalue μ k := by + rw [HasUnifEigenvalue, Submodule.ne_bot_iff] + use x; exact h + +lemma HasUnifEigenvector.apply_eq_smul {f : End R M} {μ : R} {x : M} + (hx : f.HasUnifEigenvector μ 1 x) : f x = μ • x := + mem_unifEigenspace_one.mp hx.1 + +lemma HasUnifEigenvector.pow_apply {f : End R M} {μ : R} {v : M} (hv : f.HasUnifEigenvector μ 1 v) + (n : ℕ) : (f ^ n) v = μ ^ n • v := by + induction n <;> simp [*, pow_succ f, hv.apply_eq_smul, smul_smul, pow_succ' μ] + +theorem HasUnifEigenvalue.exists_hasUnifEigenvector + {f : End R M} {μ : R} {k : ℕ∞} (hμ : f.HasUnifEigenvalue μ k) : + ∃ v, f.HasUnifEigenvector μ k v := + Submodule.exists_mem_ne_zero_of_ne_bot hμ + +lemma HasUnifEigenvalue.pow {f : End R M} {μ : R} (h : f.HasUnifEigenvalue μ 1) (n : ℕ) : + (f ^ n).HasUnifEigenvalue (μ ^ n) 1 := by + rw [HasUnifEigenvalue, Submodule.ne_bot_iff] + obtain ⟨m : M, hm⟩ := h.exists_hasUnifEigenvector + exact ⟨m, by simpa [mem_unifEigenspace_one] using hm.pow_apply n, hm.2⟩ + +/-- A nilpotent endomorphism has nilpotent eigenvalues. + +See also `LinearMap.isNilpotent_trace_of_isNilpotent`. -/ +lemma HasUnifEigenvalue.isNilpotent_of_isNilpotent [NoZeroSMulDivisors R M] {f : End R M} + (hfn : IsNilpotent f) {μ : R} (hf : f.HasUnifEigenvalue μ 1) : + IsNilpotent μ := by + obtain ⟨m : M, hm⟩ := hf.exists_hasUnifEigenvector + obtain ⟨n : ℕ, hn : f ^ n = 0⟩ := hfn + exact ⟨n, by simpa [hn, hm.2, eq_comm (a := (0 : M))] using hm.pow_apply n⟩ + +lemma HasUnifEigenvalue.mem_spectrum {f : End R M} {μ : R} (hμ : HasUnifEigenvalue f μ 1) : + μ ∈ spectrum R f := by + refine spectrum.mem_iff.mpr fun h_unit ↦ ?_ + set f' := LinearMap.GeneralLinearGroup.toLinearEquiv h_unit.unit + rcases hμ.exists_hasUnifEigenvector with ⟨v, hv⟩ + refine hv.2 ((LinearMap.ker_eq_bot'.mp f'.ker) v (?_ : μ • v - f v = 0)) + rw [hv.apply_eq_smul, sub_self] + +lemma hasUnifEigenvalue_iff_mem_spectrum [FiniteDimensional K V] {f : End K V} {μ : K} : + f.HasUnifEigenvalue μ 1 ↔ μ ∈ spectrum K f := by + rw [spectrum.mem_iff, IsUnit.sub_iff, LinearMap.isUnit_iff_ker_eq_bot, + HasUnifEigenvalue, unifEigenspace_one, ne_eq, not_iff_not] + simp [Submodule.ext_iff, LinearMap.mem_ker] + +alias ⟨_, HasUnifEigenvalue.of_mem_spectrum⟩ := hasUnifEigenvalue_iff_mem_spectrum + +lemma unifEigenspace_div (f : End K V) (a b : K) (hb : b ≠ 0) : + unifEigenspace f (a / b) 1 = LinearMap.ker (b • f - a • 1) := + calc + unifEigenspace f (a / b) 1 = unifEigenspace f (b⁻¹ * a) 1 := by rw [div_eq_mul_inv, mul_comm] + _ = LinearMap.ker (f - (b⁻¹ * a) • 1) := by rw [unifEigenspace_one] + _ = LinearMap.ker (f - b⁻¹ • a • 1) := by rw [smul_smul] + _ = LinearMap.ker (b • (f - b⁻¹ • a • 1)) := by rw [LinearMap.ker_smul _ b hb] + _ = LinearMap.ker (b • f - a • 1) := by rw [smul_sub, smul_inv_smul₀ hb] + +/-- The generalized eigenrange for a linear map `f`, a scalar `μ`, and an exponent `k ∈ ℕ∞` +is the range of `(f - μ • id) ^ k` if `k` is a natural number, +or the infimum of these ranges if `k = ∞`. -/ +def unifEigenrange (f : End R M) (μ : R) (k : ℕ∞) : Submodule R M := + ⨅ l : ℕ, ⨅ (_ : l ≤ k), LinearMap.range ((f - μ • 1) ^ l) + +lemma unifEigenrange_nat {f : End R M} {μ : R} {k : ℕ} : + f.unifEigenrange μ k = LinearMap.range ((f - μ • 1) ^ k) := by + ext x + simp only [unifEigenrange, Nat.cast_le, Submodule.mem_iInf, LinearMap.mem_range] + constructor + · intro h + exact h _ le_rfl + · rintro ⟨x, rfl⟩ i hi + have : k = i + (k - i) := by omega + rw [this, pow_add] + exact ⟨_, rfl⟩ + +/-- The exponent of a generalized eigenvalue is never 0. -/ +lemma HasUnifEigenvalue.exp_ne_zero {f : End R M} {μ : R} {k : ℕ} + (h : f.HasUnifEigenvalue μ k) : k ≠ 0 := by + rintro rfl + simp [HasUnifEigenvalue, Nat.cast_zero, unifEigenspace_zero] at h + +/-- If there exists a natural number `k` such that the kernel of `(f - μ • id) ^ k` is the +maximal generalized eigenspace, then this value is the least such `k`. If not, this value is not +meaningful. -/ +noncomputable def maxUnifEigenspaceIndex (f : End R M) (μ : R) := + monotonicSequenceLimitIndex <| (f.unifEigenspace μ).comp <| WithTop.coeOrderHom.toOrderHom + +/-- For an endomorphism of a Noetherian module, the maximal eigenspace is always of the form kernel +`(f - μ • id) ^ k` for some `k`. -/ +lemma unifEigenspace_top_eq_maxUnifEigenspaceIndex [h : IsNoetherian R M] (f : End R M) (μ : R) : + unifEigenspace f μ ⊤ = f.unifEigenspace μ (maxUnifEigenspaceIndex f μ) := by + rw [isNoetherian_iff] at h + have := WellFounded.iSup_eq_monotonicSequenceLimit h <| + (f.unifEigenspace μ).comp <| WithTop.coeOrderHom.toOrderHom + convert this using 1 + simp only [unifEigenspace, OrderHom.coe_mk, le_top, iSup_pos, OrderHom.comp_coe, + Function.comp_def] + rw [iSup_prod', iSup_subtype', ← sSup_range, ← sSup_range] + congr + aesop + +lemma unifEigenspace_le_unifEigenspace_maxUnifEigenspaceIndex [IsNoetherian R M] (f : End R M) + (μ : R) (k : ℕ∞) : + f.unifEigenspace μ k ≤ f.unifEigenspace μ (maxUnifEigenspaceIndex f μ) := by + rw [← unifEigenspace_top_eq_maxUnifEigenspaceIndex] + exact (f.unifEigenspace μ).monotone le_top + +/-- Generalized eigenspaces for exponents at least `finrank K V` are equal to each other. -/ +theorem unifEigenspace_eq_unifEigenspace_maxUnifEigenspaceIndex_of_le [IsNoetherian R M] + (f : End R M) (μ : R) {k : ℕ} (hk : maxUnifEigenspaceIndex f μ ≤ k) : + f.unifEigenspace μ k = f.unifEigenspace μ (maxUnifEigenspaceIndex f μ) := + le_antisymm + (unifEigenspace_le_unifEigenspace_maxUnifEigenspaceIndex _ _ _) + ((f.unifEigenspace μ).monotone <| by simpa using hk) + +/-- A generalized eigenvalue for some exponent `k` is also + a generalized eigenvalue for exponents larger than `k`. -/ +lemma HasUnifEigenvalue.le {f : End R M} {μ : R} {k m : ℕ∞} + (hm : k ≤ m) (hk : f.HasUnifEigenvalue μ k) : + f.HasUnifEigenvalue μ m := by + unfold HasUnifEigenvalue at * + contrapose! hk + rw [← le_bot_iff, ← hk] + exact (f.unifEigenspace _).monotone hm + +/-- A generalized eigenvalue for some exponent `k` is also + a generalized eigenvalue for positive exponents. -/ +lemma HasUnifEigenvalue.lt {f : End R M} {μ : R} {k m : ℕ∞} + (hm : 0 < m) (hk : f.HasUnifEigenvalue μ k) : + f.HasUnifEigenvalue μ m := by + apply HasUnifEigenvalue.le (k := 1) (Order.one_le_iff_pos.mpr hm) + intro contra; apply hk + rw [unifEigenspace_one, LinearMap.ker_eq_bot] at contra + rw [eq_bot_iff] + intro x hx + rw [mem_unifEigenspace] at hx + rcases hx with ⟨l, -, hx⟩ + rwa [LinearMap.ker_eq_bot.mpr] at hx + rw [LinearMap.coe_pow (f - μ • 1) l] + exact Function.Injective.iterate contra l + +/-- Generalized eigenvalues are actually just eigenvalues. -/ +@[simp] +lemma hasUnifEigenvalue_iff_hasUnifEigenvalue_one {f : End R M} {μ : R} {k : ℕ∞} (hk : 0 < k) : + f.HasUnifEigenvalue μ k ↔ f.HasUnifEigenvalue μ 1 := + ⟨HasUnifEigenvalue.lt zero_lt_one, HasUnifEigenvalue.lt hk⟩ + +lemma maxUnifEigenspaceIndex_le_finrank [FiniteDimensional K V] (f : End K V) (μ : K) : + maxUnifEigenspaceIndex f μ ≤ finrank K V := by + apply Nat.sInf_le + intro n hn + apply le_antisymm + · exact (f.unifEigenspace μ).monotone <| WithTop.coeOrderHom.monotone hn + · show (f.unifEigenspace μ) n ≤ (f.unifEigenspace μ) (finrank K V) + rw [unifEigenspace_nat, unifEigenspace_nat] + apply ker_pow_le_ker_pow_finrank + +/-- Every generalized eigenvector is a generalized eigenvector for exponent `finrank K V`. + (Lemma 8.11 of [axler2015]) -/ +lemma unifEigenspace_le_unifEigenspace_finrank [FiniteDimensional K V] (f : End K V) + (μ : K) (k : ℕ∞) : f.unifEigenspace μ k ≤ f.unifEigenspace μ (finrank K V) := by + calc f.unifEigenspace μ k + ≤ f.unifEigenspace μ ⊤ := (f.unifEigenspace _).monotone le_top + _ ≤ f.unifEigenspace μ (finrank K V) := by + rw [unifEigenspace_top_eq_maxUnifEigenspaceIndex] + exact (f.unifEigenspace _).monotone <| by simpa using maxUnifEigenspaceIndex_le_finrank f μ + +/-- Generalized eigenspaces for exponents at least `finrank K V` are equal to each other. -/ +theorem unifEigenspace_eq_unifEigenspace_finrank_of_le [FiniteDimensional K V] + (f : End K V) (μ : K) {k : ℕ} (hk : finrank K V ≤ k) : + f.unifEigenspace μ k = f.unifEigenspace μ (finrank K V) := + le_antisymm + (unifEigenspace_le_unifEigenspace_finrank _ _ _) + ((f.unifEigenspace μ).monotone <| by simpa using hk) + +lemma mapsTo_unifEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) (k : ℕ∞) : + MapsTo g (f.unifEigenspace μ k) (f.unifEigenspace μ k) := by + intro x hx + simp only [SetLike.mem_coe, mem_unifEigenspace, LinearMap.mem_ker] at hx ⊢ + rcases hx with ⟨l, hl, hx⟩ + replace h : Commute ((f - μ • (1 : End R M)) ^ l) g := + (h.sub_left <| Algebra.commute_algebraMap_left μ g).pow_left l + use l, hl + rw [← LinearMap.comp_apply, ← LinearMap.mul_eq_comp, h.eq, LinearMap.mul_eq_comp, + LinearMap.comp_apply, hx, map_zero] + +/-- The restriction of `f - μ • 1` to the `k`-fold generalized `μ`-eigenspace is nilpotent. -/ +lemma isNilpotent_restrict_unifEigenspace_nat (f : End R M) (μ : R) (k : ℕ) + (h : MapsTo (f - μ • (1 : End R M)) + (f.unifEigenspace μ k) (f.unifEigenspace μ k) := + mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ k) : + IsNilpotent ((f - μ • 1).restrict h) := by + use k + ext ⟨x, hx⟩ + rw [mem_unifEigenspace_nat] at hx + rw [LinearMap.zero_apply, ZeroMemClass.coe_zero, ZeroMemClass.coe_eq_zero, + LinearMap.pow_restrict, LinearMap.restrict_apply] + ext + simpa + +/-- The restriction of `f - μ • 1` to the generalized `μ`-eigenspace is nilpotent. -/ +lemma isNilpotent_restrict_unifEigenspace_top [IsNoetherian R M] (f : End R M) (μ : R) + (h : MapsTo (f - μ • (1 : End R M)) + (f.unifEigenspace μ ⊤) (f.unifEigenspace μ ⊤) := + mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ _) : + IsNilpotent ((f - μ • 1).restrict h) := by + apply isNilpotent_restrict_of_le + swap; apply isNilpotent_restrict_unifEigenspace_nat f μ (maxUnifEigenspaceIndex f μ) + rw [unifEigenspace_top_eq_maxUnifEigenspaceIndex] + /-- The submodule `eigenspace f μ` for a linear map `f` and a scalar `μ` consists of all vectors `x` such that `f x = μ • x`. (Def 5.36 of [axler2015])-/ -def eigenspace (f : End R M) (μ : R) : Submodule R M := - LinearMap.ker (f - algebraMap R (End R M) μ) +abbrev eigenspace (f : End R M) (μ : R) : Submodule R M := + f.unifEigenspace μ 1 -lemma eigenspace_def (f : End R M) (μ : R) : - f.eigenspace μ = LinearMap.ker (f - algebraMap R (End R M) μ) := rfl +lemma eigenspace_def {f : End R M} {μ : R} : + f.eigenspace μ = LinearMap.ker (f - μ • 1) := by + rw [eigenspace, unifEigenspace_one] @[simp] -theorem eigenspace_zero (f : End R M) : f.eigenspace 0 = LinearMap.ker f := by simp [eigenspace] +theorem eigenspace_zero (f : End R M) : f.eigenspace 0 = LinearMap.ker f := by + simp only [eigenspace, ← Nat.cast_one (R := ℕ∞), unifEigenspace_zero_nat, pow_one] /-- A nonzero element of an eigenspace is an eigenvector. (Def 5.7 of [axler2015]) -/ -def HasEigenvector (f : End R M) (μ : R) (x : M) : Prop := - x ∈ eigenspace f μ ∧ x ≠ 0 +abbrev HasEigenvector (f : End R M) (μ : R) (x : M) : Prop := + HasUnifEigenvector f μ 1 x lemma hasEigenvector_iff {f : End R M} {μ : R} {x : M} : - f.HasEigenvector μ x ↔ x ∈ eigenspace f μ ∧ x ≠ 0 := Iff.rfl + f.HasEigenvector μ x ↔ x ∈ f.eigenspace μ ∧ x ≠ 0 := Iff.rfl /-- A scalar `μ` is an eigenvalue for a linear map `f` if there are nonzero vectors `x` such that `f x = μ • x`. (Def 5.5 of [axler2015]) -/ -def HasEigenvalue (f : End R M) (a : R) : Prop := - eigenspace f a ≠ ⊥ +abbrev HasEigenvalue (f : End R M) (a : R) : Prop := + HasUnifEigenvalue f a 1 -lemma hasEigenvalue_iff (f : End R M) (μ : R) : f.HasEigenvalue μ ↔ eigenspace f μ ≠ ⊥ := Iff.rfl +lemma hasEigenvalue_iff {f : End R M} {μ : R} : + f.HasEigenvalue μ ↔ f.eigenspace μ ≠ ⊥ := Iff.rfl /-- The eigenvalues of the endomorphism `f`, as a subtype of `R`. -/ -def Eigenvalues (f : End R M) : Type _ := - { μ : R // f.HasEigenvalue μ } +abbrev Eigenvalues (f : End R M) : Type _ := + UnifEigenvalues f 1 @[coe] -def Eigenvalues.val (f : Module.End R M) : Eigenvalues f → R := Subtype.val - -instance Eigenvalues.instCoeOut {f : Module.End R M} : CoeOut (Eigenvalues f) R where - coe := Eigenvalues.val f - -instance Eigenvalues.instDecidableEq [DecidableEq R] (f : Module.End R M) : - DecidableEq (Eigenvalues f) := - inferInstanceAs (DecidableEq (Subtype (fun x : R => HasEigenvalue f x))) +abbrev Eigenvalues.val (f : Module.End R M) : Eigenvalues f → R := UnifEigenvalues.val f 1 theorem hasEigenvalue_of_hasEigenvector {f : End R M} {μ : R} {x : M} (h : HasEigenvector f μ x) : - HasEigenvalue f μ := by - rw [HasEigenvalue, Submodule.ne_bot_iff] - use x; exact h + HasEigenvalue f μ := + h.hasUnifEigenvalue -theorem mem_eigenspace_iff {f : End R M} {μ : R} {x : M} : x ∈ eigenspace f μ ↔ f x = μ • x := by - rw [eigenspace, LinearMap.mem_ker, LinearMap.sub_apply, algebraMap_end_apply, sub_eq_zero] +theorem mem_eigenspace_iff {f : End R M} {μ : R} {x : M} : x ∈ eigenspace f μ ↔ f x = μ • x := + mem_unifEigenspace_one +nonrec theorem HasEigenvector.apply_eq_smul {f : End R M} {μ : R} {x : M} (hx : f.HasEigenvector μ x) : f x = μ • x := - mem_eigenspace_iff.mp hx.1 + hx.apply_eq_smul +nonrec theorem HasEigenvector.pow_apply {f : End R M} {μ : R} {v : M} (hv : f.HasEigenvector μ v) (n : ℕ) : - (f ^ n) v = μ ^ n • v := by - induction n <;> simp [*, pow_succ f, hv.apply_eq_smul, smul_smul, pow_succ' μ] + (f ^ n) v = μ ^ n • v := + hv.pow_apply n theorem HasEigenvalue.exists_hasEigenvector {f : End R M} {μ : R} (hμ : f.HasEigenvalue μ) : ∃ v, f.HasEigenvector μ v := Submodule.exists_mem_ne_zero_of_ne_bot hμ +nonrec lemma HasEigenvalue.pow {f : End R M} {μ : R} (h : f.HasEigenvalue μ) (n : ℕ) : - (f ^ n).HasEigenvalue (μ ^ n) := by - rw [HasEigenvalue, Submodule.ne_bot_iff] - obtain ⟨m : M, hm⟩ := h.exists_hasEigenvector - exact ⟨m, by simpa [mem_eigenspace_iff] using hm.pow_apply n, hm.2⟩ + (f ^ n).HasEigenvalue (μ ^ n) := + h.pow n /-- A nilpotent endomorphism has nilpotent eigenvalues. See also `LinearMap.isNilpotent_trace_of_isNilpotent`. -/ +nonrec lemma HasEigenvalue.isNilpotent_of_isNilpotent [NoZeroSMulDivisors R M] {f : End R M} (hfn : IsNilpotent f) {μ : R} (hf : f.HasEigenvalue μ) : - IsNilpotent μ := by - obtain ⟨m : M, hm⟩ := hf.exists_hasEigenvector - obtain ⟨n : ℕ, hn : f ^ n = 0⟩ := hfn - exact ⟨n, by simpa [hn, hm.2, eq_comm (a := (0 : M))] using hm.pow_apply n⟩ + IsNilpotent μ := + hf.isNilpotent_of_isNilpotent hfn +nonrec theorem HasEigenvalue.mem_spectrum {f : End R M} {μ : R} (hμ : HasEigenvalue f μ) : - μ ∈ spectrum R f := by - refine spectrum.mem_iff.mpr fun h_unit => ?_ - set f' := LinearMap.GeneralLinearGroup.toLinearEquiv h_unit.unit - rcases hμ.exists_hasEigenvector with ⟨v, hv⟩ - refine hv.2 ((LinearMap.ker_eq_bot'.mp f'.ker) v (?_ : μ • v - f v = 0)) - rw [hv.apply_eq_smul, sub_self] + μ ∈ spectrum R f := + hμ.mem_spectrum theorem hasEigenvalue_iff_mem_spectrum [FiniteDimensional K V] {f : End K V} {μ : K} : - f.HasEigenvalue μ ↔ μ ∈ spectrum K f := by - rw [spectrum.mem_iff, IsUnit.sub_iff, LinearMap.isUnit_iff_ker_eq_bot, HasEigenvalue, eigenspace] + f.HasEigenvalue μ ↔ μ ∈ spectrum K f := + hasUnifEigenvalue_iff_mem_spectrum alias ⟨_, HasEigenvalue.of_mem_spectrum⟩ := hasEigenvalue_iff_mem_spectrum theorem eigenspace_div (f : End K V) (a b : K) (hb : b ≠ 0) : eigenspace f (a / b) = LinearMap.ker (b • f - algebraMap K (End K V) a) := - calc - eigenspace f (a / b) = eigenspace f (b⁻¹ * a) := by rw [div_eq_mul_inv, mul_comm] - _ = LinearMap.ker (f - (b⁻¹ * a) • LinearMap.id) := by rw [eigenspace]; rfl - _ = LinearMap.ker (f - b⁻¹ • a • LinearMap.id) := by rw [smul_smul] - _ = LinearMap.ker (f - b⁻¹ • algebraMap K (End K V) a) := rfl - _ = LinearMap.ker (b • (f - b⁻¹ • algebraMap K (End K V) a)) := by - rw [LinearMap.ker_smul _ b hb] - _ = LinearMap.ker (b • f - algebraMap K (End K V) a) := by rw [smul_sub, smul_inv_smul₀ hb] + unifEigenspace_div f a b hb /-- The generalized eigenspace for a linear map `f`, a scalar `μ`, and an exponent `k ∈ ℕ` is the kernel of `(f - μ • id) ^ k`. (Def 8.10 of [axler2015]). Furthermore, a generalized eigenspace for some exponent `k` is contained in the generalized eigenspace for exponents larger than `k`. -/ def genEigenspace (f : End R M) (μ : R) : ℕ →o Submodule R M where - toFun k := LinearMap.ker ((f - algebraMap R (End R M) μ) ^ k) - monotone' k m hm := by - simp only [← pow_sub_mul_pow _ hm] - exact - LinearMap.ker_le_ker_comp ((f - algebraMap R (End R M) μ) ^ k) - ((f - algebraMap R (End R M) μ) ^ (m - k)) + toFun k := f.unifEigenspace μ k + monotone' k l hkl := (f.unifEigenspace μ).monotone <| by simpa lemma genEigenspace_def (f : End R M) (μ : R) (k : ℕ) : - f.genEigenspace μ k = LinearMap.ker ((f - algebraMap R (End R M) μ) ^ k) := rfl + f.genEigenspace μ k = LinearMap.ker ((f - μ • 1) ^ k) := by + rw [genEigenspace, OrderHom.coe_mk, unifEigenspace_nat] @[simp] theorem mem_genEigenspace (f : End R M) (μ : R) (k : ℕ) (m : M) : - m ∈ f.genEigenspace μ k ↔ ((f - μ • (1 : End R M)) ^ k) m = 0 := Iff.rfl + m ∈ f.genEigenspace μ k ↔ ((f - μ • (1 : End R M)) ^ k) m = 0 := + mem_unifEigenspace_nat @[simp] theorem genEigenspace_zero (f : End R M) (k : ℕ) : - f.genEigenspace 0 k = LinearMap.ker (f ^ k) := by - simp [Module.End.genEigenspace] + f.genEigenspace 0 k = LinearMap.ker (f ^ k) := + unifEigenspace_zero_nat _ _ /-- A nonzero element of a generalized eigenspace is a generalized eigenvector. (Def 8.9 of [axler2015])-/ -def HasGenEigenvector (f : End R M) (μ : R) (k : ℕ) (x : M) : Prop := - x ≠ 0 ∧ x ∈ genEigenspace f μ k +abbrev HasGenEigenvector (f : End R M) (μ : R) (k : ℕ) (x : M) : Prop := + HasUnifEigenvector f μ k x lemma hasGenEigenvector_iff {f : End R M} {μ : R} {k : ℕ} {x : M} : - f.HasGenEigenvector μ k x ↔ x ≠ 0 ∧ x ∈ f.genEigenspace μ k := Iff.rfl + f.HasGenEigenvector μ k x ↔ x ∈ f.genEigenspace μ k ∧ x ≠ 0 := Iff.rfl /-- A scalar `μ` is a generalized eigenvalue for a linear map `f` and an exponent `k ∈ ℕ` if there are generalized eigenvectors for `f`, `k`, and `μ`. -/ -def HasGenEigenvalue (f : End R M) (μ : R) (k : ℕ) : Prop := - genEigenspace f μ k ≠ ⊥ +abbrev HasGenEigenvalue (f : End R M) (μ : R) (k : ℕ) : Prop := + HasUnifEigenvalue f μ k -lemma hasGenEigenvalue_iff (f : End R M) (μ : R) (k : ℕ) : - f.HasGenEigenvalue μ k ↔ genEigenspace f μ k ≠ ⊥ := Iff.rfl +lemma hasGenEigenvalue_iff {f : End R M} {μ : R} {k : ℕ} : + f.HasGenEigenvalue μ k ↔ f.genEigenspace μ k ≠ ⊥ := Iff.rfl /-- The generalized eigenrange for a linear map `f`, a scalar `μ`, and an exponent `k ∈ ℕ` is the range of `(f - μ • id) ^ k`. -/ -def genEigenrange (f : End R M) (μ : R) (k : ℕ) : Submodule R M := - LinearMap.range ((f - algebraMap R (End R M) μ) ^ k) +abbrev genEigenrange (f : End R M) (μ : R) (k : ℕ) : Submodule R M := + unifEigenrange f μ k -lemma genEigenrange_def (f : End R M) (μ : R) (k : ℕ) : - f.genEigenrange μ k = LinearMap.range ((f - algebraMap R (End R M) μ) ^ k) := rfl +lemma genEigenrange_def {f : End R M} {μ : R} {k : ℕ} : + f.genEigenrange μ k = LinearMap.range ((f - μ • 1) ^ k) := by + rw [genEigenrange, unifEigenrange_nat] /-- The exponent of a generalized eigenvalue is never 0. -/ theorem exp_ne_zero_of_hasGenEigenvalue {f : End R M} {μ : R} {k : ℕ} - (h : f.HasGenEigenvalue μ k) : k ≠ 0 := by - rintro rfl - exact h LinearMap.ker_id + (h : f.HasGenEigenvalue μ k) : k ≠ 0 := + HasUnifEigenvalue.exp_ne_zero h /-- The union of the kernels of `(f - μ • id) ^ k` over all `k`. -/ -def maxGenEigenspace (f : End R M) (μ : R) : Submodule R M := - ⨆ k, f.genEigenspace μ k +abbrev maxGenEigenspace (f : End R M) (μ : R) : Submodule R M := + unifEigenspace f μ ⊤ lemma maxGenEigenspace_def (f : End R M) (μ : R) : - f.maxGenEigenspace μ = ⨆ k, f.genEigenspace μ k := rfl + f.maxGenEigenspace μ = ⨆ k, f.genEigenspace μ k := by + simp_rw [maxGenEigenspace, unifEigenspace_top, genEigenspace, OrderHom.coe_mk] theorem genEigenspace_le_maximal (f : End R M) (μ : R) (k : ℕ) : f.genEigenspace μ k ≤ f.maxGenEigenspace μ := - le_iSup _ _ + (f.unifEigenspace μ).monotone le_top @[simp] theorem mem_maxGenEigenspace (f : End R M) (μ : R) (m : M) : - m ∈ f.maxGenEigenspace μ ↔ ∃ k : ℕ, ((f - μ • (1 : End R M)) ^ k) m = 0 := by - simp only [maxGenEigenspace, ← mem_genEigenspace, Submodule.mem_iSup_of_chain] + m ∈ f.maxGenEigenspace μ ↔ ∃ k : ℕ, ((f - μ • (1 : End R M)) ^ k) m = 0 := + mem_unifEigenspace_top /-- If there exists a natural number `k` such that the kernel of `(f - μ • id) ^ k` is the maximal generalized eigenspace, then this value is the least such `k`. If not, this value is not meaningful. -/ -noncomputable def maxGenEigenspaceIndex (f : End R M) (μ : R) := - monotonicSequenceLimitIndex (f.genEigenspace μ) +noncomputable abbrev maxGenEigenspaceIndex (f : End R M) (μ : R) := + maxUnifEigenspaceIndex f μ /-- For an endomorphism of a Noetherian module, the maximal eigenspace is always of the form kernel `(f - μ • id) ^ k` for some `k`. -/ -theorem maxGenEigenspace_eq [h : IsNoetherian R M] (f : End R M) (μ : R) : +theorem maxGenEigenspace_eq [IsNoetherian R M] (f : End R M) (μ : R) : maxGenEigenspace f μ = f.genEigenspace μ (maxGenEigenspaceIndex f μ) := - h.wf.iSup_eq_monotonicSequenceLimit (f.genEigenspace μ) + unifEigenspace_top_eq_maxUnifEigenspaceIndex _ _ /-- A generalized eigenvalue for some exponent `k` is also a generalized eigenvalue for exponents larger than `k`. -/ theorem hasGenEigenvalue_of_hasGenEigenvalue_of_le {f : End R M} {μ : R} {k : ℕ} {m : ℕ} (hm : k ≤ m) (hk : f.HasGenEigenvalue μ k) : - f.HasGenEigenvalue μ m := by - unfold HasGenEigenvalue at * - contrapose! hk - rw [← le_bot_iff, ← hk] - exact (f.genEigenspace μ).monotone hm + f.HasGenEigenvalue μ m := + hk.le <| by simpa using hm /-- The eigenspace is a subspace of the generalized eigenspace. -/ theorem eigenspace_le_genEigenspace {f : End R M} {μ : R} {k : ℕ} (hk : 0 < k) : f.eigenspace μ ≤ f.genEigenspace μ k := - (f.genEigenspace μ).monotone (Nat.succ_le_of_lt hk) + (f.unifEigenspace _).monotone <| by simpa using Nat.succ_le_of_lt hk /-- All eigenvalues are generalized eigenvalues. -/ theorem hasGenEigenvalue_of_hasEigenvalue {f : End R M} {μ : R} {k : ℕ} (hk : 0 < k) - (hμ : f.HasEigenvalue μ) : f.HasGenEigenvalue μ k := by - apply hasGenEigenvalue_of_hasGenEigenvalue_of_le hk - rw [HasGenEigenvalue, genEigenspace, OrderHom.coe_mk, pow_one] - exact hμ + (hμ : f.HasEigenvalue μ) : f.HasGenEigenvalue μ k := + hμ.lt <| by simpa using hk /-- All generalized eigenvalues are eigenvalues. -/ theorem hasEigenvalue_of_hasGenEigenvalue {f : End R M} {μ : R} {k : ℕ} - (hμ : f.HasGenEigenvalue μ k) : f.HasEigenvalue μ := by - intro contra; apply hμ - erw [LinearMap.ker_eq_bot] at contra ⊢; rw [LinearMap.coe_pow] - exact Function.Injective.iterate contra k + (hμ : f.HasGenEigenvalue μ k) : f.HasEigenvalue μ := + hμ.lt zero_lt_one /-- Generalized eigenvalues are actually just eigenvalues. -/ @[simp] theorem hasGenEigenvalue_iff_hasEigenvalue {f : End R M} {μ : R} {k : ℕ} (hk : 0 < k) : f.HasGenEigenvalue μ k ↔ f.HasEigenvalue μ := - ⟨hasEigenvalue_of_hasGenEigenvalue, hasGenEigenvalue_of_hasEigenvalue hk⟩ + hasUnifEigenvalue_iff_hasUnifEigenvalue_one <| by simpa using hk /-- Every generalized eigenvector is a generalized eigenvector for exponent `finrank K V`. (Lemma 8.11 of [axler2015]) -/ theorem genEigenspace_le_genEigenspace_finrank [FiniteDimensional K V] (f : End K V) (μ : K) (k : ℕ) : f.genEigenspace μ k ≤ f.genEigenspace μ (finrank K V) := - ker_pow_le_ker_pow_finrank _ _ + unifEigenspace_le_unifEigenspace_finrank _ _ _ @[simp] theorem iSup_genEigenspace_eq_genEigenspace_finrank [FiniteDimensional K V] (f : End K V) (μ : K) : @@ -289,7 +587,7 @@ theorem genEigenspace_le_genEigenspace_finrank [FiniteDimensional K V] (f : End theorem genEigenspace_eq_genEigenspace_finrank_of_le [FiniteDimensional K V] (f : End K V) (μ : K) {k : ℕ} (hk : finrank K V ≤ k) : f.genEigenspace μ k = f.genEigenspace μ (finrank K V) := - ker_pow_eq_ker_pow_finrank_of_le hk + unifEigenspace_eq_unifEigenspace_finrank_of_le f μ hk lemma mapsTo_genEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) (k : ℕ) : MapsTo g (f.genEigenspace μ k) (f.genEigenspace μ k) := by @@ -300,21 +598,43 @@ lemma mapsTo_genEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) (k rw [← LinearMap.comp_apply, ← LinearMap.mul_eq_comp, h.eq, LinearMap.mul_eq_comp, LinearMap.comp_apply, hx, map_zero] -lemma mapsTo_iSup_genEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) : - MapsTo g ↑(⨆ k, f.genEigenspace μ k) ↑(⨆ k, f.genEigenspace μ k) := by +lemma iSup_genEigenspace_eq (f : End R M) (μ : R) : + ⨆ k, (f.genEigenspace μ) k = f.unifEigenspace μ ⊤ := by + rw [unifEigenspace_eq_iSup_unifEigenspace_nat] + ext + simp only [iSup_subtype, le_top, iSup_pos] + rfl + +lemma mapsTo_maxGenEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) : + MapsTo g ↑(f.maxGenEigenspace μ) ↑(f.maxGenEigenspace μ) := by + rw [maxGenEigenspace_def] simp only [MapsTo, Submodule.coe_iSup_of_chain, mem_iUnion, SetLike.mem_coe] rintro x ⟨k, hk⟩ exact ⟨k, f.mapsTo_genEigenspace_of_comm h μ k hk⟩ +lemma mapsTo_iSup_genEigenspace_of_comm {f g : End R M} (h : Commute f g) (μ : R) : + MapsTo g ↑(⨆ k, f.genEigenspace μ k) ↑(⨆ k, f.genEigenspace μ k) := by + rw [← maxGenEigenspace_def] + apply mapsTo_maxGenEigenspace_of_comm h + /-- The restriction of `f - μ • 1` to the `k`-fold generalized `μ`-eigenspace is nilpotent. -/ lemma isNilpotent_restrict_sub_algebraMap (f : End R M) (μ : R) (k : ℕ) (h : MapsTo (f - algebraMap R (End R M) μ) (f.genEigenspace μ k) (f.genEigenspace μ k) := mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ k) : + IsNilpotent ((f - algebraMap R (End R M) μ).restrict h) := + isNilpotent_restrict_unifEigenspace_nat _ _ _ + +/-- The restriction of `f - μ • 1` to the generalized `μ`-eigenspace is nilpotent. -/ +lemma isNilpotent_restrict_maxGenEigenspace_sub_algebraMap [IsNoetherian R M] (f : End R M) (μ : R) + (h : MapsTo (f - algebraMap R (End R M) μ) + ↑(f.maxGenEigenspace μ) ↑(f.maxGenEigenspace μ) := + mapsTo_maxGenEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ) : IsNilpotent ((f - algebraMap R (End R M) μ).restrict h) := by - use k - ext - simp [LinearMap.restrict_apply, LinearMap.pow_restrict _] + apply isNilpotent_restrict_of_le (q := f.unifEigenspace μ (maxUnifEigenspaceIndex f μ)) + _ (isNilpotent_restrict_unifEigenspace_nat f μ (maxUnifEigenspaceIndex f μ)) + rw [maxGenEigenspace_eq] + exact le_rfl /-- The restriction of `f - μ • 1` to the generalized `μ`-eigenspace is nilpotent. -/ lemma isNilpotent_restrict_iSup_sub_algebraMap [IsNoetherian R M] (f : End R M) (μ : R) @@ -322,42 +642,53 @@ lemma isNilpotent_restrict_iSup_sub_algebraMap [IsNoetherian R M] (f : End R M) ↑(⨆ k, f.genEigenspace μ k) ↑(⨆ k, f.genEigenspace μ k) := mapsTo_iSup_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ) μ) : IsNilpotent ((f - algebraMap R (End R M) μ).restrict h) := by - obtain ⟨l, hl⟩ : ∃ l, ⨆ k, f.genEigenspace μ k = f.genEigenspace μ l := - ⟨_, maxGenEigenspace_eq f μ⟩ - use l - ext ⟨x, hx⟩ - simpa [hl, LinearMap.restrict_apply, LinearMap.pow_restrict _] using hx - -lemma disjoint_genEigenspace [NoZeroSMulDivisors R M] - (f : End R M) {μ₁ μ₂ : R} (hμ : μ₁ ≠ μ₂) (k l : ℕ) : - Disjoint (f.genEigenspace μ₁ k) (f.genEigenspace μ₂ l) := by + apply isNilpotent_restrict_of_le (q := f.unifEigenspace μ (maxUnifEigenspaceIndex f μ)) + _ (isNilpotent_restrict_unifEigenspace_nat f μ (maxUnifEigenspaceIndex f μ)) + apply iSup_le + intro k + apply unifEigenspace_le_unifEigenspace_maxUnifEigenspaceIndex + +lemma disjoint_unifEigenspace [NoZeroSMulDivisors R M] + (f : End R M) {μ₁ μ₂ : R} (hμ : μ₁ ≠ μ₂) (k l : ℕ∞) : + Disjoint (f.unifEigenspace μ₁ k) (f.unifEigenspace μ₂ l) := by + rw [unifEigenspace_eq_iSup_unifEigenspace_nat, unifEigenspace_eq_iSup_unifEigenspace_nat] + simp_rw [unifEigenspace_directed.disjoint_iSup_left, unifEigenspace_directed.disjoint_iSup_right] + rintro ⟨k, -⟩ ⟨l, -⟩ nontriviality M have := NoZeroSMulDivisors.isReduced R M rw [disjoint_iff] - set p := f.genEigenspace μ₁ k ⊓ f.genEigenspace μ₂ l + set p := f.unifEigenspace μ₁ k ⊓ f.unifEigenspace μ₂ l by_contra hp replace hp : Nontrivial p := Submodule.nontrivial_iff_ne_bot.mpr hp let f₁ : End R p := (f - algebraMap R (End R M) μ₁).restrict <| MapsTo.inter_inter - (mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₁) μ₁ k) - (mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₁) μ₂ l) + (mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₁) μ₁ k) + (mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₁) μ₂ l) let f₂ : End R p := (f - algebraMap R (End R M) μ₂).restrict <| MapsTo.inter_inter - (mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₁ k) - (mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₂ l) + (mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₁ k) + (mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f μ₂) μ₂ l) have : IsNilpotent (f₂ - f₁) := by - apply Commute.isNilpotent_sub (x := f₂) (y := f₁) _ ⟨l, ?_⟩ ⟨k, ?_⟩ + apply Commute.isNilpotent_sub (x := f₂) (y := f₁) _ + (isNilpotent_restrict_of_le inf_le_right _) + (isNilpotent_restrict_of_le inf_le_left _) · ext; simp [f₁, f₂, smul_sub, sub_sub, smul_comm μ₁, add_sub_left_comm] - all_goals ext ⟨x, _, _⟩; simpa [LinearMap.restrict_apply, LinearMap.pow_restrict _] using ‹_› + apply mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f _) + apply isNilpotent_restrict_unifEigenspace_nat + apply mapsTo_unifEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f _) + apply isNilpotent_restrict_unifEigenspace_nat have hf₁₂ : f₂ - f₁ = algebraMap R (End R p) (μ₁ - μ₂) := by ext; simp [f₁, f₂, sub_smul] rw [hf₁₂, IsNilpotent.map_iff (NoZeroSMulDivisors.algebraMap_injective R (End R p)), isNilpotent_iff_eq_zero, sub_eq_zero] at this contradiction +lemma disjoint_genEigenspace [NoZeroSMulDivisors R M] + (f : End R M) {μ₁ μ₂ : R} (hμ : μ₁ ≠ μ₂) (k l : ℕ) : + Disjoint (f.genEigenspace μ₁ k) (f.genEigenspace μ₂ l) := + disjoint_unifEigenspace f hμ k l + lemma disjoint_iSup_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) {μ₁ μ₂ : R} (hμ : μ₁ ≠ μ₂) : Disjoint (⨆ k, f.genEigenspace μ₁ k) (⨆ k, f.genEigenspace μ₂ k) := by - simp_rw [(f.genEigenspace μ₁).mono.directed_le.disjoint_iSup_left, - (f.genEigenspace μ₂).mono.directed_le.disjoint_iSup_right] - exact disjoint_genEigenspace f hμ + simpa only [iSup_genEigenspace_eq] using disjoint_unifEigenspace f hμ ⊤ ⊤ lemma injOn_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : InjOn (⨆ k, f.genEigenspace · k) {μ | ⨆ k, f.genEigenspace μ k ≠ ⊥} := by @@ -366,12 +697,14 @@ lemma injOn_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : apply hμ₂ simpa only [hμ₁₂, disjoint_self] using f.disjoint_iSup_genEigenspace contra -theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : - CompleteLattice.Independent (fun μ ↦ ⨆ k, f.genEigenspace μ k) := by +theorem independent_maxGenEigenspace [NoZeroSMulDivisors R M] (f : End R M) : + CompleteLattice.Independent f.maxGenEigenspace := by classical suffices ∀ μ (s : Finset R), μ ∉ s → Disjoint (⨆ k, f.genEigenspace μ k) (s.sup fun μ ↦ ⨆ k, f.genEigenspace μ k) by - simp_rw [CompleteLattice.independent_iff_supIndep_of_injOn f.injOn_genEigenspace, + show CompleteLattice.Independent (f.maxGenEigenspace ·) + simp_rw [maxGenEigenspace_def, + CompleteLattice.independent_iff_supIndep_of_injOn f.injOn_genEigenspace, Finset.supIndep_iff_disjoint_erase] exact fun s μ _ ↦ this _ _ (s.not_mem_erase μ) intro μ₁ s @@ -405,6 +738,11 @@ theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : simp_rw [Submodule.mem_iSup_of_chain, mem_genEigenspace] exact ⟨k, hyz⟩ +theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : + CompleteLattice.Independent (fun μ ↦ ⨆ k, f.genEigenspace μ k) := by + simp_rw [← maxGenEigenspace_def] + apply independent_maxGenEigenspace + /-- The eigenspaces of a linear operator form an independent family of subspaces of `M`. That is, any eigenspace has trivial intersection with the span of all the other eigenspaces. -/ theorem eigenspaces_independent [NoZeroSMulDivisors R M] (f : End R M) : @@ -417,7 +755,7 @@ theorem eigenvectors_linearIndependent' {ι : Type*} [NoZeroSMulDivisors R M] (f : End R M) (μ : ι → R) (hμ : Function.Injective μ) (v : ι → M) (h_eigenvec : ∀ i, f.HasEigenvector (μ i) (v i)) : LinearIndependent R v := f.eigenspaces_independent.comp hμ |>.linearIndependent _ - (fun i => h_eigenvec i |>.left) (fun i => h_eigenvec i |>.right) + (fun i ↦ h_eigenvec i |>.left) (fun i ↦ h_eigenvec i |>.right) /-- Eigenvectors corresponding to distinct eigenvalues of a linear operator are linearly independent. (Lemma 5.10 of [axler2015]) @@ -428,7 +766,7 @@ theorem eigenvectors_linearIndependent' {ι : Type*} [NoZeroSMulDivisors R M] theorem eigenvectors_linearIndependent [NoZeroSMulDivisors R M] (f : End R M) (μs : Set R) (xs : μs → M) (h_eigenvec : ∀ μ : μs, f.HasEigenvector μ (xs μ)) : LinearIndependent R xs := - f.eigenvectors_linearIndependent' (fun μ : μs => μ) Subtype.coe_injective _ h_eigenvec + f.eigenvectors_linearIndependent' (fun μ : μs ↦ μ) Subtype.coe_injective _ h_eigenvec /-- If `f` maps a subspace `p` into itself, then the generalized eigenspace of the restriction of `f` to `p` is the part of the generalized eigenspace of `f` that lies in `p`. -/ @@ -436,7 +774,7 @@ theorem genEigenspace_restrict (f : End R M) (p : Submodule R M) (k : ℕ) (μ : (hfp : ∀ x : M, x ∈ p → f x ∈ p) : genEigenspace (LinearMap.restrict f hfp) μ k = Submodule.comap p.subtype (f.genEigenspace μ k) := by - simp only [genEigenspace, OrderHom.coe_mk, ← LinearMap.ker_comp] + simp only [genEigenspace_def, OrderHom.coe_mk, ← LinearMap.ker_comp] induction' k with k ih · rw [pow_zero, pow_zero, LinearMap.one_eq_id] apply (Submodule.ker_subtype _).symm @@ -459,7 +797,7 @@ lemma _root_.Submodule.inf_iInf_maxGenEigenspace_of_forall_mapsTo {ι : Type*} { (⨅ i, maxGenEigenspace ((f i).restrict (hfp i)) (μ i)).map p.subtype := by cases isEmpty_or_nonempty ι · simp [iInf_of_isEmpty] - · simp_rw [inf_iInf, maxGenEigenspace, ((f _).genEigenspace _).mono.directed_le.inf_iSup_eq, + · simp_rw [inf_iInf, maxGenEigenspace_def, ((f _).genEigenspace _).mono.directed_le.inf_iSup_eq, p.inf_genEigenspace _ (hfp _), ← Submodule.map_iSup, Submodule.map_iInf _ p.injective_subtype] /-- Given a family of endomorphisms `i ↦ f i`, a family of candidate eigenvalues `i ↦ μ i`, and a @@ -480,7 +818,6 @@ lemma iInf_maxGenEigenspace_restrict_map_subtype_eq rw [Submodule.map_iInf _ p.injective_subtype, this, Submodule.inf_iInf] simp_rw [maxGenEigenspace_def, Submodule.map_iSup, ((f _).genEigenspace _).mono.directed_le.inf_iSup_eq, p.inf_genEigenspace (f _) (h _)] - rfl lemma mapsTo_restrict_maxGenEigenspace_restrict_of_mapsTo {p : Submodule R M} (f g : End R M) (hf : MapsTo f p p) (hg : MapsTo g p p) {μ₁ μ₂ : R} @@ -508,16 +845,18 @@ theorem generalized_eigenvec_disjoint_range_ker [FiniteDimensional K V] (f : End (f.genEigenspace μ (finrank K V)) := by have h := calc - Submodule.comap ((f - algebraMap _ _ μ) ^ finrank K V) + Submodule.comap ((f - μ • 1) ^ finrank K V) (f.genEigenspace μ (finrank K V)) = LinearMap.ker ((f - algebraMap _ _ μ) ^ finrank K V * (f - algebraMap K (End K V) μ) ^ finrank K V) := by - rw [genEigenspace, OrderHom.coe_mk, ← LinearMap.ker_comp]; rfl - _ = f.genEigenspace μ (finrank K V + finrank K V) := by rw [← pow_add]; rfl + rw [genEigenspace, OrderHom.coe_mk, unifEigenspace_nat, ← LinearMap.ker_comp]; rfl + _ = f.genEigenspace μ (finrank K V + finrank K V) := by + rw [← pow_add, genEigenspace, OrderHom.coe_mk, unifEigenspace_nat]; rfl _ = f.genEigenspace μ (finrank K V) := by - rw [genEigenspace_eq_genEigenspace_finrank_of_le]; omega - rw [disjoint_iff_inf_le, genEigenrange, LinearMap.range_eq_map, - Submodule.map_inf_eq_map_inf_comap, top_inf_eq, h] + rw [genEigenspace_eq_genEigenspace_finrank_of_le]; omega + rw [disjoint_iff_inf_le, genEigenrange, unifEigenrange_nat, LinearMap.range_eq_map, + Submodule.map_inf_eq_map_inf_comap, top_inf_eq, h, + genEigenspace, OrderHom.coe_mk, unifEigenspace_nat] apply Submodule.map_comap_le /-- If an invariant subspace `p` of an endomorphism `f` is disjoint from the `μ`-eigenspace of `f`, @@ -544,11 +883,13 @@ theorem map_genEigenrange_le {f : End K V} {μ : K} {n : ℕ} : calc Submodule.map f (f.genEigenrange μ n) = LinearMap.range (f * (f - algebraMap _ _ μ) ^ n) := by - rw [genEigenrange]; exact (LinearMap.range_comp _ _).symm + rw [genEigenrange, unifEigenrange_nat]; exact (LinearMap.range_comp _ _).symm _ = LinearMap.range ((f - algebraMap _ _ μ) ^ n * f) := by rw [Algebra.mul_sub_algebraMap_pow_commutes] _ = Submodule.map ((f - algebraMap _ _ μ) ^ n) (LinearMap.range f) := LinearMap.range_comp _ _ - _ ≤ f.genEigenrange μ n := LinearMap.map_le_range + _ ≤ f.genEigenrange μ n := by + rw [genEigenrange, unifEigenrange_nat] + apply LinearMap.map_le_range lemma iSup_genEigenspace_le_smul (f : Module.End R M) (μ t : R) : (⨆ k, f.genEigenspace μ k) ≤ ⨆ k, (t • f).genEigenspace (t * μ) k := by diff --git a/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean b/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean index e6cf66f99d9ea..70c17ee7b2097 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean @@ -78,7 +78,7 @@ theorem hasEigenvalue_of_isRoot (h : (minpoly K f).IsRoot μ) : f.HasEigenvalue have : (aeval f) p = 0 := by have h_aeval := minpoly.aeval K f revert h_aeval - simp [hp, ← hu] + simp [hp, ← hu, Algebra.algebraMap_eq_smul_one] have h_deg := minpoly.degree_le_of_ne_zero K f p_ne_0 this rw [hp, degree_mul, degree_X_sub_C, Polynomial.degree_eq_natDegree p_ne_0] at h_deg norm_cast at h_deg diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index c0f53eccf53f8..03e4ac0132b66 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -250,7 +250,8 @@ lemma Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo apply ih _ (hy φ) · intro j k μ exact mapsTo_restrict_maxGenEigenspace_restrict_of_mapsTo (f j) (f k) _ _ (h j k μ) - · exact fun j ↦ Module.End.iSup_genEigenspace_restrict_eq_top _ (h' j) + · simp_rw [maxGenEigenspace_def] at h' ⊢ + exact fun j ↦ Module.End.iSup_genEigenspace_restrict_eq_top _ (h' j) · rfl replace ih (φ : K) : ⨆ (χ : ι → K) (_ : χ i = φ), ⨅ j, maxGenEigenspace ((f j).restrict (hi j φ)) (χ j) = ⊤ := by @@ -260,6 +261,7 @@ lemma Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo rw [eq_bot_iff, ← ((f i).maxGenEigenspace φ).ker_subtype, LinearMap.ker, ← Submodule.map_le_iff_le_comap, ← Submodule.inf_iInf_maxGenEigenspace_of_forall_mapsTo, ← disjoint_iff_inf_le] + simp_rw [maxGenEigenspace_def] exact ((f i).disjoint_iSup_genEigenspace hχ.symm).mono_right (iInf_le _ i) replace ih (φ : K) : ⨆ (χ : ι → K) (_ : χ i = φ), ⨅ j, maxGenEigenspace (f j) (χ j) = diff --git a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean index 8992078b835ca..cd48d0d7611d4 100644 --- a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean @@ -164,6 +164,9 @@ theorem algebraMap_eq_zero_iff (x : R) : algebraMap R (ExteriorAlgebra R M) x = theorem algebraMap_eq_one_iff (x : R) : algebraMap R (ExteriorAlgebra R M) x = 1 ↔ x = 1 := map_eq_one_iff (algebraMap _ _) (algebraMap_leftInverse _).injective +instance isLocalRingHom_algebraMap : IsLocalRingHom (algebraMap R (ExteriorAlgebra R M)) := + isLocalRingHom_of_leftInverse _ (algebraMap_leftInverse M) + theorem isUnit_algebraMap (r : R) : IsUnit (algebraMap R (ExteriorAlgebra R M) r) ↔ IsUnit r := isUnit_map_of_leftInverse _ (algebraMap_leftInverse M) diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index f045983a13cd3..9ce4ca2f473b7 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -1252,7 +1252,7 @@ theorem linearIndependent_option' : LinearIndependent K (fun o => Option.casesOn' o x v : Option ι → V) ↔ LinearIndependent K v ∧ x ∉ Submodule.span K (range v) := by -- Porting note: Explicit universe level is required in `Equiv.optionEquivSumPUnit`. - rw [← linearIndependent_equiv (Equiv.optionEquivSumPUnit.{_, u'} ι).symm, linearIndependent_sum, + rw [← linearIndependent_equiv (Equiv.optionEquivSumPUnit.{u', _} ι).symm, linearIndependent_sum, @range_unique _ PUnit, @linearIndependent_unique_iff PUnit, disjoint_span_singleton] dsimp [(· ∘ ·)] refine ⟨fun h => ⟨h.1, fun hx => h.2.1 <| h.2.2 hx⟩, fun h => ⟨h.1, ?_, fun hx => (h.2 hx).elim⟩⟩ @@ -1401,6 +1401,10 @@ theorem LinearIndependent.subset_span_extend (hs : LinearIndependent K (fun x => let ⟨_hbt, _hsb, htb, _hli⟩ := Classical.choose_spec (exists_linearIndependent_extension hs hst) htb +theorem LinearIndependent.span_extend_eq_span (hs : LinearIndependent K (fun x => x : s → V)) + (hst : s ⊆ t) : span K (hs.extend hst) = span K t := + le_antisymm (span_mono (hs.extend_subset hst)) (span_le.2 (hs.subset_span_extend hst)) + theorem LinearIndependent.linearIndependent_extend (hs : LinearIndependent K (fun x => x : s → V)) (hst : s ⊆ t) : LinearIndependent K ((↑) : hs.extend hst → V) := let ⟨_hbt, _hsb, _htb, hli⟩ := Classical.choose_spec (exists_linearIndependent_extension hs hst) diff --git a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean index 50c1c2dfc8bb0..5ae49bf1b77c9 100644 --- a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean +++ b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean @@ -137,6 +137,10 @@ def map (f : R →+* S) : GL n R →* GL n S := Units.map <| (RingHom.mapMatrix theorem map_id : map (RingHom.id R) = MonoidHom.id (GL n R) := rfl +@[simp] +protected lemma map_apply (f : R →+* S) (i j : n) (g : GL n R) : map f g i j = f (g i j) := by + rfl + @[simp] theorem map_comp (f : T →+* R) (g : R →+* S) : map (g.comp f) = (map g).comp (map (n := n) f) := @@ -147,6 +151,44 @@ theorem map_comp_apply (f : T →+* R) (g : R →+* S) (x : GL n T) : (map g).comp (map f) x = map g (map f x) := rfl +variable (f : R →+* S) + +@[simp] +protected lemma map_one : map f (1 : GL n R) = 1 := by + ext + simp only [_root_.map_one, Units.val_one] + +protected lemma map_mul (g h : GL n R) : map f (g * h) = map f g * map f h := by + ext + simp only [_root_.map_mul, Units.val_mul] + +protected lemma map_inv (g : GL n R) : map f g⁻¹ = (map f g)⁻¹ := by + ext + simp only [_root_.map_inv, coe_units_inv] + +protected lemma map_det (g : GL n R) : Matrix.GeneralLinearGroup.det (map f g) = + Units.map f (Matrix.GeneralLinearGroup.det g) := by + ext + simp only [map, RingHom.mapMatrix_apply, Units.inv_eq_val_inv, Matrix.coe_units_inv, + Matrix.GeneralLinearGroup.val_det_apply, Units.coe_map, MonoidHom.coe_coe] + exact Eq.symm (RingHom.map_det f g.1) + +lemma map_mul_map_inv (g : GL n R) : map f g * map f g⁻¹ = 1 := by + simp only [map_inv, mul_inv_cancel] + +lemma map_inv_mul_map (g : GL n R) : map f g⁻¹ * map f g = 1 := by + simp only [map_inv, inv_mul_cancel] + +@[simp] +lemma coe_map_mul_map_inv (g : GL n R) : g.val.map f * g.val⁻¹.map f = 1 := by + rw [← Matrix.map_mul] + simp only [isUnits_det_units, mul_nonsing_inv, map_zero, _root_.map_one, Matrix.map_one] + +@[simp] +lemma coe_map_inv_mul_map (g : GL n R) : g.val⁻¹.map f * g.val.map f = 1 := by + rw [← Matrix.map_mul] + simp only [isUnits_det_units, nonsing_inv_mul, map_zero, _root_.map_one, Matrix.map_one] + end GeneralLinearGroup namespace SpecialLinearGroup diff --git a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean index 6ef3b1d0b480f..41e80c3f6b369 100644 --- a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean +++ b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean @@ -776,4 +776,21 @@ theorem det_conj' {M : Matrix m m α} (h : IsUnit M) (N : Matrix m m α) : end Det +/-! ### More results about traces -/ + + +section trace + +variable [Fintype m] [DecidableEq m] + +/-- A variant of `Matrix.trace_units_conj`. -/ +theorem trace_conj {M : Matrix m m α} (h : IsUnit M) (N : Matrix m m α) : + trace (M * N * M⁻¹) = trace N := by rw [← h.unit_spec, ← coe_units_inv, trace_units_conj] + +/-- A variant of `Matrix.trace_units_conj'`. -/ +theorem trace_conj' {M : Matrix m m α} (h : IsUnit M) (N : Matrix m m α) : + trace (M⁻¹ * N * M) = trace N := by rw [← h.unit_spec, ← coe_units_inv, trace_units_conj'] + +end trace + end Matrix diff --git a/Mathlib/LinearAlgebra/Matrix/Trace.lean b/Mathlib/LinearAlgebra/Matrix/Trace.lean index fbc571b846cfc..40da0be3f44c0 100644 --- a/Mathlib/LinearAlgebra/Matrix/Trace.lean +++ b/Mathlib/LinearAlgebra/Matrix/Trace.lean @@ -175,6 +175,23 @@ lemma trace_submatrix_succ {n : ℕ} [NonUnitalNonAssocSemiring R] rw [← (finSuccEquiv n).symm.sum_comp] simp +section CommSemiring + +variable [DecidableEq m] [CommSemiring R] + +-- TODO(mathlib4#6607): fix elaboration so that the ascription isn't needed +theorem trace_units_conj (M : (Matrix m m R)ˣ) (N : Matrix m m R) : + trace ((M : Matrix _ _ _) * N * (↑M⁻¹ : Matrix _ _ _)) = trace N := by + rw [trace_mul_cycle, Units.inv_mul, one_mul] + +set_option linter.docPrime false in +-- TODO(mathlib4#6607): fix elaboration so that the ascription isn't needed +theorem trace_units_conj' (M : (Matrix m m R)ˣ) (N : Matrix m m R) : + trace ((↑M⁻¹ : Matrix _ _ _) * N * (↑M : Matrix _ _ _)) = trace N := + trace_units_conj M⁻¹ N + +end CommSemiring + section Fin variable [AddCommMonoid R] diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 2a78e4ae591af..25022902c14a5 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -109,7 +109,7 @@ variable [Semiring R] [∀ i, AddCommMonoid (M i)] [∀ i, AddCommMonoid (M₁ i -- Porting note: Replaced CoeFun with FunLike instance instance : FunLike (MultilinearMap R M₁ M₂) (∀ i, M₁ i) M₂ where coe f := f.toFun - coe_injective' := fun f g h ↦ by cases f; cases g; cases h; rfl + coe_injective' f g h := by cases f; cases g; cases h; rfl initialize_simps_projections MultilinearMap (toFun → apply) diff --git a/Mathlib/LinearAlgebra/PerfectPairing.lean b/Mathlib/LinearAlgebra/PerfectPairing.lean index d5e2dd945411b..dfec7b49779cd 100644 --- a/Mathlib/LinearAlgebra/PerfectPairing.lean +++ b/Mathlib/LinearAlgebra/PerfectPairing.lean @@ -34,7 +34,7 @@ open Function Module variable (R M N : Type*) [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] /-- A perfect pairing of two (left) modules over a commutative ring. -/ -structure PerfectPairing := +structure PerfectPairing where toLin : M →ₗ[R] N →ₗ[R] R bijectiveLeft : Bijective toLin bijectiveRight : Bijective toLin.flip diff --git a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean index c05f820fefc9c..c0d2eb6ac8a7a 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean @@ -43,7 +43,7 @@ def of {X : Type v} [AddCommGroup X] [Module R X] (Q : QuadraticForm R X) : /-- A type alias for `QuadraticForm.LinearIsometry` to avoid confusion between the categorical and algebraic spellings of composition. -/ @[ext] -structure Hom (V W : QuadraticModuleCat.{v} R) := +structure Hom (V W : QuadraticModuleCat.{v} R) where /-- The underlying isometry -/ toIsometry : V.form →qᵢ W.form diff --git a/Mathlib/LinearAlgebra/Quotient.lean b/Mathlib/LinearAlgebra/Quotient.lean index 61d6fd9576a2a..c512d5ed3b658 100644 --- a/Mathlib/LinearAlgebra/Quotient.lean +++ b/Mathlib/LinearAlgebra/Quotient.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov -/ import Mathlib.LinearAlgebra.Span +import Mathlib.LinearAlgebra.Pi import Mathlib.Algebra.Module.Equiv.Basic import Mathlib.GroupTheory.QuotientGroup.Basic import Mathlib.SetTheory.Cardinal.Finite @@ -282,9 +283,14 @@ def mkQ : M →ₗ[R] M ⧸ p where theorem mkQ_apply (x : M) : p.mkQ x = (Quotient.mk x : M ⧸ p) := rfl -theorem mkQ_surjective (A : Submodule R M) : Function.Surjective A.mkQ := by +theorem mkQ_surjective : Function.Surjective p.mkQ := by rintro ⟨x⟩; exact ⟨x, rfl⟩ +theorem strictMono_comap_prod_map : + StrictMono fun m : Submodule R M ↦ (m.comap p.subtype, m.map p.mkQ) := + fun m₁ m₂ ↦ QuotientAddGroup.strictMono_comap_prod_map + p.toAddSubgroup (a := m₁.toAddSubgroup) (b := m₂.toAddSubgroup) + end variable {R₂ M₂ : Type*} [Ring R₂] [AddCommGroup M₂] [Module R₂ M₂] {τ₁₂ : R →+* R₂} @@ -310,6 +316,14 @@ theorem liftQ_apply (f : M →ₛₗ[τ₁₂] M₂) {h} (x : M) : p.liftQ f h ( @[simp] theorem liftQ_mkQ (f : M →ₛₗ[τ₁₂] M₂) (h) : (p.liftQ f h).comp p.mkQ = f := by ext; rfl +theorem pi_liftQ_eq_liftQ_pi {ι : Type*} {N : ι → Type*} + [∀ i, AddCommGroup (N i)] [∀ i, Module R (N i)] + (f : (i : ι) → M →ₗ[R] (N i)) {p : Submodule R M} (h : ∀ i, p ≤ ker (f i)) : + LinearMap.pi (fun i ↦ p.liftQ (f i) (h i)) = + p.liftQ (LinearMap.pi f) (LinearMap.ker_pi f ▸ le_iInf h) := by + ext x i + simp + /-- Special case of `submodule.liftQ` when `p` is the span of `x`. In this case, the condition on `f` simply becomes vanishing at `x`. -/ def liftQSpanSingleton (x : M) (f : M →ₛₗ[τ₁₂] M₂) (h : f x = 0) : (M ⧸ R ∙ x) →ₛₗ[τ₁₂] M₂ := diff --git a/Mathlib/LinearAlgebra/RootSystem/Defs.lean b/Mathlib/LinearAlgebra/RootSystem/Defs.lean index 607057bd4e391..eac40333d52fc 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Defs.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Defs.lean @@ -86,7 +86,7 @@ evaluates to `2`, and the permutation attached to each element of `ι` is compat reflections on the corresponding roots and coroots. It exists to allow for a convenient unification of the theories of root systems and root data. -/ -structure RootPairing extends PerfectPairing R M N := +structure RootPairing extends PerfectPairing R M N where /-- A parametrized family of vectors, called roots. -/ root : ι ↪ M /-- A parametrized family of dual vectors, called coroots. -/ @@ -110,7 +110,7 @@ abbrev RootDatum (X₁ X₂ : Type*) [AddCommGroup X₁] [AddCommGroup X₂] := Note that this is slightly more general than the usual definition in the sense that `N` is not required to be the dual of `M`. -/ -structure RootSystem extends RootPairing ι R M N := +structure RootSystem extends RootPairing ι R M N where span_eq_top : span R (range root) = ⊤ attribute [simp] RootSystem.span_eq_top diff --git a/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean index 026b7b23f30d7..b339fa3a131cc 100644 --- a/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean @@ -33,7 +33,7 @@ namespace LinearMap injective, and for any vector `y`, the norm of `x` divides twice the inner product of `x` and `y`. These conditions are what we need when describing reflection as a map taking `y` to `y - 2 • (B x y) / (B x x) • x`. -/ -structure IsReflective (B : M →ₗ[R] M →ₗ[R] R) (x : M) : Prop := +structure IsReflective (B : M →ₗ[R] M →ₗ[R] R) (x : M) : Prop where regular : IsRegular (B x x) dvd_two_mul : ∀ y, B x x ∣ 2 * B x y diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 2bdc1a10ded1d..538530f266a43 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -670,7 +670,7 @@ namespace Classical /-- Any prop `p` is decidable classically. A shorthand for `Classical.propDecidable`. -/ noncomputable def dec (p : Prop) : Decidable p := by infer_instance -variable {α : Sort*} {p : α → Prop} +variable {α : Sort*} /-- Any predicate `p` is decidable classically. -/ noncomputable def decPred (p : α → Prop) : DecidablePred p := by infer_instance diff --git a/Mathlib/Logic/Encodable/Basic.lean b/Mathlib/Logic/Encodable/Basic.lean index c111c09a26737..60412ec201292 100644 --- a/Mathlib/Logic/Encodable/Basic.lean +++ b/Mathlib/Logic/Encodable/Basic.lean @@ -555,13 +555,9 @@ theorem sequence_mono_nat {r : β → β → Prop} {f : α → β} (hf : Directe · exact (Classical.choose_spec (hf p p)).1 · exact (Classical.choose_spec (hf p a)).1 -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem rel_sequence {r : β → β → Prop} {f : α → β} (hf : Directed r f) (a : α) : r (f a) (f (hf.sequence f (encode a + 1))) := by - simp only [Directed.sequence, add_eq, add_zero, encodek, _root_.and_self] + simp only [Directed.sequence, add_eq, add_zero, encodek, and_self] exact (Classical.choose_spec (hf _ a)).2 variable [Preorder β] {f : α → β} diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index d8c4acbd7d587..14135e7e720df 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -52,7 +52,7 @@ universe u v w z open Function -- Unless required to be `Type*`, all variables in this file are `Sort*` -variable {α α₁ α₂ β β₁ β₂ γ γ₁ γ₂ δ : Sort*} +variable {α α₁ α₂ β β₁ β₂ γ δ : Sort*} namespace Equiv @@ -629,7 +629,7 @@ section /-- A family of equivalences `∀ a, β₁ a ≃ β₂ a` generates an equivalence between `∀ a, β₁ a` and `∀ a, β₂ a`. -/ def piCongrRight {β₁ β₂ : α → Sort*} (F : ∀ a, β₁ a ≃ β₂ a) : (∀ a, β₁ a) ≃ (∀ a, β₂ a) := - ⟨fun H a => F a (H a), fun H a => (F a).symm (H a), fun H => funext <| by simp, + ⟨Pi.map fun a ↦ F a, Pi.map fun a ↦ (F a).symm, fun H => funext <| by simp, fun H => funext <| by simp⟩ /-- Given `φ : α → β → Sort*`, we have an equivalence between `∀ a b, φ a b` and `∀ b a, φ a b`. @@ -1688,7 +1688,7 @@ theorem piCongr_symm_apply (f : ∀ b, Z b) : @[simp] theorem piCongr_apply_apply (f : ∀ a, W a) (a : α) : h₁.piCongr h₂ f (h₁ a) = h₂ a (f a) := by - simp only [piCongr, piCongrRight, trans_apply, coe_fn_mk, piCongrLeft_apply_apply] + simp only [piCongr, piCongrRight, trans_apply, coe_fn_mk, piCongrLeft_apply_apply, Pi.map_apply] end diff --git a/Mathlib/Logic/Equiv/Fin.lean b/Mathlib/Logic/Equiv/Fin.lean index 1d07d2d1b2b25..3453d19eac020 100644 --- a/Mathlib/Logic/Equiv/Fin.lean +++ b/Mathlib/Logic/Equiv/Fin.lean @@ -470,7 +470,7 @@ def Int.divModEquiv (n : ℕ) [NeZero n] : ℤ ≃ ℤ × Fin n where toFun a := (a / n, ↑(a.natMod n)) invFun p := p.1 * n + ↑p.2 left_inv a := by - simp_rw [Fin.coe_ofNat_eq_mod, natCast_mod, natMod, + simp_rw [Fin.coe_natCast_eq_mod, natCast_mod, natMod, toNat_of_nonneg (emod_nonneg _ <| natCast_eq_zero.not.2 (NeZero.ne n)), emod_emod, ediv_add_emod'] right_inv := fun ⟨q, r, hrn⟩ => by diff --git a/Mathlib/Logic/Equiv/PartialEquiv.lean b/Mathlib/Logic/Equiv/PartialEquiv.lean index 7d93cc552f446..9dd50be972f0c 100644 --- a/Mathlib/Logic/Equiv/PartialEquiv.lean +++ b/Mathlib/Logic/Equiv/PartialEquiv.lean @@ -316,7 +316,7 @@ def IsImage (s : Set α) (t : Set β) : Prop := namespace IsImage -variable {e} {s : Set α} {t : Set β} {x : α} {y : β} +variable {e} {s : Set α} {t : Set β} {x : α} theorem apply_mem_iff (h : e.IsImage s t) (hx : x ∈ e.source) : e x ∈ t ↔ x ∈ s := h hx @@ -843,8 +843,8 @@ variable {ι : Type*} {αi βi γi : ι → Type*} /-- The product of a family of partial equivalences, as a partial equivalence on the pi type. -/ @[simps (config := mfld_cfg) apply source target] protected def pi (ei : ∀ i, PartialEquiv (αi i) (βi i)) : PartialEquiv (∀ i, αi i) (∀ i, βi i) where - toFun f i := ei i (f i) - invFun f i := (ei i).symm (f i) + toFun := Pi.map fun i ↦ ei i + invFun := Pi.map fun i ↦ (ei i).symm source := pi univ fun i => (ei i).source target := pi univ fun i => (ei i).target map_source' _ hf i hi := (ei i).map_source (hf i hi) diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index 21435518020b1..1af46bf00e7fa 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -118,10 +118,15 @@ theorem Injective.of_comp_iff' (f : α → β) {g : γ → α} (hg : Bijective g Injective (f ∘ g) ↔ Injective f := ⟨fun I ↦ I.of_comp_right hg.2, fun h ↦ h.comp hg.injective⟩ +theorem Injective.piMap {ι : Sort*} {α β : ι → Sort*} {f : ∀ i, α i → β i} + (hf : ∀ i, Injective (f i)) : Injective (Pi.map f) := fun _ _ h ↦ + funext fun i ↦ hf i <| congrFun h _ + +@[deprecated (since := "2024-10-06")] alias injective_pi_map := Injective.piMap + /-- Composition by an injective function on the left is itself injective. -/ -theorem Injective.comp_left {g : β → γ} (hg : Function.Injective g) : - Function.Injective (g ∘ · : (α → β) → α → γ) := - fun _ _ hgf ↦ funext fun i ↦ hg <| (congr_fun hgf i : _) +theorem Injective.comp_left {g : β → γ} (hg : Injective g) : Injective (g ∘ · : (α → β) → α → γ) := + .piMap fun _ ↦ hg theorem injective_of_subsingleton [Subsingleton α] (f : α → β) : Injective f := fun _ _ _ ↦ Subsingleton.elim _ _ @@ -359,7 +364,7 @@ end section InvFun -variable {α β : Sort*} [Nonempty α] {f : α → β} {a : α} {b : β} +variable {α β : Sort*} [Nonempty α] {f : α → β} {b : β} attribute [local instance] Classical.propDecidable @@ -442,10 +447,22 @@ theorem surjective_to_subsingleton [na : Nonempty α] [Subsingleton β] (f : α Surjective f := fun _ ↦ let ⟨a⟩ := na; ⟨a, Subsingleton.elim _ _⟩ +theorem Surjective.piMap {ι : Sort*} {α β : ι → Sort*} {f : ∀ i, α i → β i} + (hf : ∀ i, Surjective (f i)) : Surjective (Pi.map f) := fun g ↦ + ⟨fun i ↦ surjInv (hf i) (g i), funext fun _ ↦ rightInverse_surjInv _ _⟩ + +@[deprecated (since := "2024-10-06")] alias surjective_pi_map := Surjective.piMap + /-- Composition by a surjective function on the left is itself surjective. -/ theorem Surjective.comp_left {g : β → γ} (hg : Surjective g) : - Surjective (g ∘ · : (α → β) → α → γ) := fun f ↦ - ⟨surjInv hg ∘ f, funext fun _ ↦ rightInverse_surjInv _ _⟩ + Surjective (g ∘ · : (α → β) → α → γ) := + .piMap fun _ ↦ hg + +theorem Bijective.piMap {ι : Sort*} {α β : ι → Sort*} {f : ∀ i, α i → β i} + (hf : ∀ i, Bijective (f i)) : Bijective (Pi.map f) := + ⟨.piMap fun i ↦ (hf i).1, .piMap fun i ↦ (hf i).2⟩ + +@[deprecated (since := "2024-10-06")] alias bijective_pi_map := Bijective.piMap /-- Composition by a bijective function on the left is itself bijective. -/ theorem Bijective.comp_left {g : β → γ} (hg : Bijective g) : @@ -457,7 +474,7 @@ end SurjInv section Update variable {α : Sort u} {β : α → Sort v} {α' : Sort w} [DecidableEq α] - {f g : (a : α) → β a} {a : α} {b : β a} + {f : (a : α) → β a} {a : α} {b : β a} /-- Replacing the value of a function at a given point by a given value. -/ diff --git a/Mathlib/Logic/Function/CompTypeclasses.lean b/Mathlib/Logic/Function/CompTypeclasses.lean index 7fcd8ca32286c..69d8ee1e0294f 100644 --- a/Mathlib/Logic/Function/CompTypeclasses.lean +++ b/Mathlib/Logic/Function/CompTypeclasses.lean @@ -24,7 +24,7 @@ TODO : section CompTriple /-- Class of composing triples -/ -class CompTriple {M N P : Type*} (φ : M → N) (ψ : N → P) (χ : outParam (M → P)) : Prop where +class CompTriple {M N P : Type*} (φ : M → N) (ψ : N → P) (χ : outParam (M → P)) : Prop where /-- The maps form a commuting triangle -/ comp_eq : ψ.comp φ = χ diff --git a/Mathlib/Logic/Function/Defs.lean b/Mathlib/Logic/Function/Defs.lean index 78f83b38c66c0..ab1505baf76fa 100644 --- a/Mathlib/Logic/Function/Defs.lean +++ b/Mathlib/Logic/Function/Defs.lean @@ -203,3 +203,16 @@ protected theorem RightInverse.id {g : β → α} {f : α → β} (h : RightInve def IsFixedPt (f : α → α) (x : α) := f x = x end Function + +namespace Pi + +variable {ι : Sort*} {α β : ι → Sort*} + +/-- Sends a dependent function `a : ∀ i, α i` to a dependent function `Pi.map f a : ∀ i, β i` +by applying `f i` to `i`-th component. -/ +protected def map (f : ∀ i, α i → β i) : (∀ i, α i) → (∀ i, β i) := fun a i ↦ f i (a i) + +@[simp] +lemma map_apply (f : ∀ i, α i → β i) (a : ∀ i, α i) (i : ι) : Pi.map f a i = f i (a i) := rfl + +end Pi diff --git a/Mathlib/Logic/Lemmas.lean b/Mathlib/Logic/Lemmas.lean index 1056a7ffc156f..3e92a9dc21060 100644 --- a/Mathlib/Logic/Lemmas.lean +++ b/Mathlib/Logic/Lemmas.lean @@ -24,7 +24,7 @@ theorem iff_right_comm {a b c : Prop} : ((a ↔ b) ↔ c) ↔ ((a ↔ c) ↔ b) protected alias ⟨HEq.eq, Eq.heq⟩ := heq_iff_eq -variable {α : Sort*} {p q r : Prop} [Decidable p] [Decidable q] {a b c : α} +variable {α : Sort*} {p q : Prop} [Decidable p] [Decidable q] {a b c : α} theorem dite_dite_distrib_left {a : p → α} {b : ¬p → q → α} {c : ¬p → ¬q → α} : (dite p a fun hp ↦ dite q (b hp) (c hp)) = diff --git a/Mathlib/Logic/Pairwise.lean b/Mathlib/Logic/Pairwise.lean index 26dd58b7b3c9e..13ae772203bbe 100644 --- a/Mathlib/Logic/Pairwise.lean +++ b/Mathlib/Logic/Pairwise.lean @@ -21,11 +21,11 @@ This file defines pairwise relations. open Set Function -variable {α β γ ι ι' : Type*} {r p q : α → α → Prop} +variable {α β ι : Type*} {r p : α → α → Prop} section Pairwise -variable {f g : ι → α} {s t u : Set α} {a b : α} +variable {f : ι → α} {s : Set α} {a b : α} /-- A relation `r` holds pairwise if `r i j` for all `i ≠ j`. -/ def Pairwise (r : α → α → Prop) := diff --git a/Mathlib/Logic/Relation.lean b/Mathlib/Logic/Relation.lean index bbe9ad200f0ca..2a90c042f467f 100644 --- a/Mathlib/Logic/Relation.lean +++ b/Mathlib/Logic/Relation.lean @@ -213,7 +213,7 @@ instance [Decidable (∃ a b, r a b ∧ f a = c ∧ g b = d)] : Decidable (Relat end Map -variable {r : α → α → Prop} {a b c d : α} +variable {r : α → α → Prop} {a b c : α} /-- `ReflTransGen r`: reflexive transitive closure of `r` -/ @[mk_iff ReflTransGen.cases_tail_iff] diff --git a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean index f1677d8c92db5..89db09097d3cb 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean @@ -343,7 +343,7 @@ protected lemma AnalyticSet.preimage {X Y : Type*} [TopologicalSpace X] [Topolog [PolishSpace X] [T2Space Y] {s : Set Y} (hs : AnalyticSet s) {f : X → Y} (hf : Continuous f) : AnalyticSet (f ⁻¹' s) := by rcases analyticSet_iff_exists_polishSpace_range.1 hs with ⟨Z, _, _, g, hg, rfl⟩ - have : IsClosed {x : X × Z | f x.1 = g x.2} := isClosed_diagonal.preimage (hf.prod_map hg) + have : IsClosed {x : X × Z | f x.1 = g x.2} := isClosed_eq hf.fst' hg.snd' convert this.analyticSet.image_of_continuous continuous_fst ext x simp [eq_comm] diff --git a/Mathlib/MeasureTheory/Constructions/Projective.lean b/Mathlib/MeasureTheory/Constructions/Projective.lean index 1754b6934c461..37ef05c547f45 100644 --- a/Mathlib/MeasureTheory/Constructions/Projective.lean +++ b/Mathlib/MeasureTheory/Constructions/Projective.lean @@ -46,6 +46,16 @@ namespace IsProjectiveMeasureFamily variable {I J : Finset ι} +lemma eq_zero_of_isEmpty [h : IsEmpty (Π i, α i)] + (hP : IsProjectiveMeasureFamily P) (I : Finset ι) : + P I = 0 := by + classical + obtain ⟨i, hi⟩ := isEmpty_pi.mp h + rw [hP (insert i I) I (I.subset_insert i)] + have : IsEmpty (Π j : ↑(insert i I), α j) := by simp [hi] + rw [(P (insert i I)).eq_zero_of_isEmpty] + simp + /-- Auxiliary lemma for `measure_univ_eq`. -/ lemma measure_univ_eq_of_subset (hP : IsProjectiveMeasureFamily P) (hJI : J ⊆ I) : P I univ = P J univ := by diff --git a/Mathlib/MeasureTheory/Covering/Besicovitch.lean b/Mathlib/MeasureTheory/Covering/Besicovitch.lean index 466e9d9155efd..91cc3db5e2be0 100644 --- a/Mathlib/MeasureTheory/Covering/Besicovitch.lean +++ b/Mathlib/MeasureTheory/Covering/Besicovitch.lean @@ -309,7 +309,7 @@ theorem mem_iUnionUpTo_lastStep (x : β) : p.c x ∈ p.iUnionUpTo p.lastStep := apply lt_trans (mul_pos (_root_.zero_lt_one.trans p.one_lt_tau) (p.rpos _)) H have B : p.τ⁻¹ * p.R p.lastStep < p.R p.lastStep := by conv_rhs => rw [← one_mul (p.R p.lastStep)] - exact mul_lt_mul (inv_lt_one p.one_lt_tau) le_rfl Rpos zero_le_one + exact mul_lt_mul (inv_lt_one_of_one_lt₀ p.one_lt_tau) le_rfl Rpos zero_le_one obtain ⟨y, hy1, hy2⟩ : ∃ y, p.c y ∉ p.iUnionUpTo p.lastStep ∧ p.τ⁻¹ * p.R p.lastStep < p.r y := by have := exists_lt_of_lt_csSup ?_ B · simpa only [exists_prop, mem_range, exists_exists_and_eq_and, Subtype.exists, diff --git a/Mathlib/MeasureTheory/Covering/Differentiation.lean b/Mathlib/MeasureTheory/Covering/Differentiation.lean index 4af0418c40a8e..803d605ca1b7f 100644 --- a/Mathlib/MeasureTheory/Covering/Differentiation.lean +++ b/Mathlib/MeasureTheory/Covering/Differentiation.lean @@ -573,7 +573,7 @@ theorem withDensity_le_mul {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht conv_rhs => rw [← mul_one (t ^ n)] gcongr rw [zpow_neg_one] - exact inv_lt_one ht + exact inv_lt_one_of_one_lt₀ ht calc ν s = ν (s ∩ f ⁻¹' {0}) + ν (s ∩ f ⁻¹' {∞}) + diff --git a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean index 947cbabbc4212..77cd7b8a79759 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean @@ -312,7 +312,7 @@ theorem exists_subset_restrict_nonpos (hi : s i < 0) : · have : 1 / s E < bdd k := by linarith only [le_of_max_le_left (hk k le_rfl)] rw [one_div] at this ⊢ - rwa [inv_lt (lt_trans (inv_pos.2 hE₃) this) hE₃] + exact inv_lt_of_inv_lt₀ hE₃ this obtain ⟨k, hk₁, hk₂⟩ := this have hA' : A ⊆ i \ ⋃ l ≤ k, restrictNonposSeq s i l := by apply Set.diff_subset_diff_right diff --git a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean index 2d82b8b7e4d4e..f9a7daf12fa48 100644 --- a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean +++ b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean @@ -208,7 +208,8 @@ theorem TendstoInMeasure.exists_seq_tendsto_ae (hfg : TendstoInMeasure μ f atTo congr rw [← pow_one (2 : ℝ)⁻¹] rw [← pow_add, add_comm] - exact pow_le_pow_of_le_one (one_div (2 : ℝ) ▸ one_half_pos.le) (inv_le_one one_le_two) + exact pow_le_pow_of_le_one (one_div (2 : ℝ) ▸ one_half_pos.le) + (inv_le_one_of_one_le₀ one_le_two) ((le_tsub_add.trans (add_le_add_right (le_max_right _ _) 1)).trans (add_le_add_right hn_ge 1)) exact le_trans hNx.le h_inv_n_le_k diff --git a/Mathlib/MeasureTheory/Function/EssSup.lean b/Mathlib/MeasureTheory/Function/EssSup.lean index 926d7c0f68080..503c49f68e8df 100644 --- a/Mathlib/MeasureTheory/Function/EssSup.lean +++ b/Mathlib/MeasureTheory/Function/EssSup.lean @@ -36,7 +36,7 @@ variable {α β : Type*} {m : MeasurableSpace α} {μ ν : Measure α} section ConditionallyCompleteLattice -variable [ConditionallyCompleteLattice β] +variable [ConditionallyCompleteLattice β] {f : α → β} /-- Essential supremum of `f` with respect to measure `μ`: the smallest `c : β` such that `f x ≤ c` a.e. -/ @@ -68,6 +68,32 @@ theorem essSup_const (c : β) (hμ : μ ≠ 0) : essSup (fun _ : α => c) μ = c theorem essInf_const (c : β) (hμ : μ ≠ 0) : essInf (fun _ : α => c) μ = c := have := NeZero.mk hμ; essInf_const' _ +section SMul +variable {R : Type*} [Zero R] [SMulWithZero R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] + [NoZeroSMulDivisors R ℝ≥0∞] {c : R} + +@[simp] +lemma essSup_smul_measure (hc : c ≠ 0) (f : α → β) : essSup f (c • μ) = essSup f μ := by + simp_rw [essSup, Measure.ae_smul_measure_eq hc] + +end SMul + +variable [Nonempty α] + +lemma essSup_eq_ciSup (hμ : ∀ a, μ {a} ≠ 0) (hf : BddAbove (Set.range f)) : + essSup f μ = ⨆ a, f a := by rw [essSup, ae_eq_top.2 hμ, limsup_top_eq_ciSup hf] + +lemma essInf_eq_ciInf (hμ : ∀ a, μ {a} ≠ 0) (hf : BddBelow (Set.range f)) : + essInf f μ = ⨅ a, f a := by rw [essInf, ae_eq_top.2 hμ, liminf_top_eq_ciInf hf] + +variable [MeasurableSingletonClass α] + +@[simp] lemma essSup_count_eq_ciSup (hf : BddAbove (Set.range f)) : + essSup f .count = ⨆ a, f a := essSup_eq_ciSup (by simp) hf + +@[simp] lemma essInf_count_eq_ciInf (hf : BddBelow (Set.range f)) : + essInf f .count = ⨅ a, f a := essInf_eq_ciInf (by simp) hf + end ConditionallyCompleteLattice section ConditionallyCompleteLinearOrder @@ -172,10 +198,6 @@ theorem essInf_antitone_measure {f : α → β} (hμν : μ ≪ ν) : essInf f refine liminf_le_liminf_of_le (Measure.ae_le_iff_absolutelyContinuous.mpr hμν) ?_ ?_ all_goals isBoundedDefault -theorem essSup_smul_measure {f : α → β} {c : ℝ≥0∞} (hc : c ≠ 0) : - essSup f (c • μ) = essSup f μ := by - simp_rw [essSup, Measure.ae_smul_measure_eq hc] - lemma essSup_eq_iSup (hμ : ∀ a, μ {a} ≠ 0) (f : α → β) : essSup f μ = ⨆ i, f i := by rw [essSup, ae_eq_top.2 hμ, limsup_top_eq_iSup] diff --git a/Mathlib/MeasureTheory/Function/Intersectivity.lean b/Mathlib/MeasureTheory/Function/Intersectivity.lean index 5ea828b2fcf5b..c571061c05018 100644 --- a/Mathlib/MeasureTheory/Function/Intersectivity.lean +++ b/Mathlib/MeasureTheory/Function/Intersectivity.lean @@ -61,8 +61,7 @@ lemma bergelson' {s : ℕ → Set α} (hs : ∀ n, MeasurableSet (s n)) (hr₀ : have hfapp : ∀ n a, f n a = (↑(n + 1))⁻¹ * ∑ k in Finset.range (n + 1), (s k).indicator 1 a := by simp only [f, Pi.natCast_def, Pi.smul_apply, Pi.inv_apply, Finset.sum_apply, eq_self_iff_true, forall_const, imp_true_iff, smul_eq_mul] - have hf n : Measurable (f n) := Measurable.mul' (@measurable_const ℝ≥0∞ _ _ _ (↑(n + 1))⁻¹) - (Finset.measurable_sum' _ fun i _ ↦ measurable_one.indicator <| hs i) + have hf n : Measurable (f n) := by fun_prop (disch := exact hs _) have hf₁ n : f n ≤ 1 := by rintro a rw [hfapp, ← ENNReal.div_eq_inv_mul] diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 75f2bb1b23b55..fdf5f19ca96be 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -794,10 +794,16 @@ theorem eLpNorm'_smul_measure {p : ℝ} (hp : 0 ≤ p) {f : α → F} (c : ℝ @[deprecated (since := "2024-07-27")] alias snorm'_smul_measure := eLpNorm'_smul_measure -theorem eLpNormEssSup_smul_measure {f : α → F} {c : ℝ≥0∞} (hc : c ≠ 0) : +section SMul +variable {R : Type*} [Zero R] [SMulWithZero R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] + [NoZeroSMulDivisors R ℝ≥0∞] {c : R} + +@[simp] lemma eLpNormEssSup_smul_measure (hc : c ≠ 0) (f : α → F) : eLpNormEssSup f (c • μ) = eLpNormEssSup f μ := by simp_rw [eLpNormEssSup] - exact essSup_smul_measure hc + exact essSup_smul_measure hc _ + +end SMul @[deprecated (since := "2024-07-27")] alias snormEssSup_smul_measure := eLpNormEssSup_smul_measure diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean index dcafb2aeecbec..e7723bfa145de 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean @@ -126,7 +126,7 @@ theorem eLpNorm'_lt_top_of_eLpNorm'_lt_top_of_exponent_le {p q : ℝ} [IsFiniteM _ < ∞ := by rw [ENNReal.mul_lt_top_iff] 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] + 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 := diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean index 42f9edb4bd2b4..cc56d954aa492 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean @@ -88,9 +88,8 @@ theorem LpAddConst_lt_top (p : ℝ≥0∞) : LpAddConst p < ∞ := by rw [LpAddConst] split_ifs with h · apply ENNReal.rpow_lt_top_of_nonneg _ ENNReal.two_ne_top - simp only [one_div, sub_nonneg] - apply one_le_inv (ENNReal.toReal_pos h.1.ne' (h.2.trans ENNReal.one_lt_top).ne) - simpa using ENNReal.toReal_mono ENNReal.one_ne_top h.2.le + rw [one_div, sub_nonneg, ← ENNReal.toReal_inv, ← ENNReal.one_toReal] + exact ENNReal.toReal_mono (by simpa using h.1.ne') (ENNReal.one_le_inv.2 h.2.le) · exact ENNReal.one_lt_top theorem eLpNorm_add_le' (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasurable g μ) diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 5615eaebce091..0935ab9a20775 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -161,7 +161,7 @@ theorem measurableSet_preimage (f : α →ₛ β) (s) : MeasurableSet (f ⁻¹' measurableSet_cut (fun _ b => b ∈ s) f fun b => MeasurableSet.const (b ∈ s) /-- A simple function is measurable -/ -@[measurability] +@[measurability, fun_prop] protected theorem measurable [MeasurableSpace β] (f : α →ₛ β) : Measurable f := fun s _ => measurableSet_preimage f s diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index eee3b2db0726d..44a87eb7d5c81 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -684,8 +684,11 @@ lemma isUniformEmbedding : IsUniformEmbedding ((↑) : Lp.simpleFunc E p μ → @[deprecated (since := "2024-10-01")] alias uniformEmbedding := isUniformEmbedding -protected theorem uniformInducing : UniformInducing ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := - simpleFunc.isUniformEmbedding.toUniformInducing +theorem isUniformInducing : IsUniformInducing ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := + simpleFunc.isUniformEmbedding.isUniformInducing + +@[deprecated (since := "2024-10-05")] +alias uniformInducing := isUniformInducing lemma isDenseEmbedding (hp_ne_top : p ≠ ∞) : IsDenseEmbedding ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := by diff --git a/Mathlib/MeasureTheory/Group/Arithmetic.lean b/Mathlib/MeasureTheory/Group/Arithmetic.lean index 1893764f613b0..a265438e4ebd7 100644 --- a/Mathlib/MeasureTheory/Group/Arithmetic.lean +++ b/Mathlib/MeasureTheory/Group/Arithmetic.lean @@ -87,8 +87,8 @@ export MeasurableMul₂ (measurable_mul) section Mul -variable {M α : Type*} [MeasurableSpace M] [Mul M] {m : MeasurableSpace α} {f g : α → M} - {μ : Measure α} +variable {M α β : Type*} [MeasurableSpace M] [Mul M] {m : MeasurableSpace α} + {mβ : MeasurableSpace β} {f g : α → M} {μ : Measure α} @[to_additive (attr := fun_prop, measurability)] theorem Measurable.const_mul [MeasurableMul M] (hf : Measurable f) (c : M) : @@ -110,17 +110,19 @@ theorem AEMeasurable.mul_const [MeasurableMul M] (hf : AEMeasurable f μ) (c : M AEMeasurable (fun x => f x * c) μ := (measurable_mul_const c).comp_aemeasurable hf -@[to_additive (attr := aesop safe 20 apply (rule_sets := [Measurable]))] -theorem Measurable.mul' [MeasurableMul₂ M] (hf : Measurable f) (hg : Measurable g) : - Measurable (f * g) := - measurable_mul.comp (hf.prod_mk hg) - @[to_additive (attr := fun_prop, aesop safe 20 apply (rule_sets := [Measurable]))] theorem Measurable.mul [MeasurableMul₂ M] (hf : Measurable f) (hg : Measurable g) : Measurable fun a => f a * g a := measurable_mul.comp (hf.prod_mk hg) -@[to_additive (attr := aesop safe 20 apply (rule_sets := [Measurable]))] +/-- Compositional version of `Measurable.mul` for use by `fun_prop`. -/ +@[to_additive (attr := fun_prop, aesop safe 20 apply (rule_sets := [Measurable])) +"Compositional version of `Measurable.add` for use by `fun_prop`."] +lemma Measurable.mul' [MeasurableMul₂ M] {f g : α → β → M} {h : α → β} (hf : Measurable ↿f) + (hg : Measurable ↿g) (hh : Measurable h) : Measurable fun a ↦ (f a * g a) (h a) := by + simp; fun_prop + +@[to_additive (attr := fun_prop, aesop safe 20 apply (rule_sets := [Measurable]))] theorem AEMeasurable.mul' [MeasurableMul₂ M] (hf : AEMeasurable f μ) (hg : AEMeasurable g μ) : AEMeasurable (f * g) μ := measurable_mul.comp_aemeasurable (hf.prod_mk hg) @@ -238,8 +240,8 @@ export MeasurableDiv₂ (measurable_div) section Div -variable {G α : Type*} [MeasurableSpace G] [Div G] {m : MeasurableSpace α} {f g : α → G} - {μ : Measure α} +variable {G α β : Type*} [MeasurableSpace G] [Div G] {m : MeasurableSpace α} + {mβ : MeasurableSpace β} {f g : α → G} {μ : Measure α} @[to_additive (attr := measurability)] theorem Measurable.const_div [MeasurableDiv G] (hf : Measurable f) (c : G) : @@ -261,17 +263,17 @@ theorem AEMeasurable.div_const [MeasurableDiv G] (hf : AEMeasurable f μ) (c : G AEMeasurable (fun x => f x / c) μ := (MeasurableDiv.measurable_div_const c).comp_aemeasurable hf -@[to_additive (attr := aesop safe 20 apply (rule_sets := [Measurable]))] -theorem Measurable.div' [MeasurableDiv₂ G] (hf : Measurable f) (hg : Measurable g) : - Measurable (f / g) := - measurable_div.comp (hf.prod_mk hg) - @[to_additive (attr := fun_prop, aesop safe 20 apply (rule_sets := [Measurable]))] theorem Measurable.div [MeasurableDiv₂ G] (hf : Measurable f) (hg : Measurable g) : Measurable fun a => f a / g a := measurable_div.comp (hf.prod_mk hg) -@[to_additive (attr := aesop safe 20 apply (rule_sets := [Measurable]))] +@[to_additive (attr := fun_prop, aesop safe 20 apply (rule_sets := [Measurable]))] +lemma Measurable.div' [MeasurableDiv₂ G] {f g : α → β → G} {h : α → β} (hf : Measurable ↿f) + (hg : Measurable ↿g) (hh : Measurable h) : Measurable fun a ↦ (f a / g a) (h a) := by + simp; fun_prop + +@[to_additive (attr := fun_prop, aesop safe 20 apply (rule_sets := [Measurable]))] theorem AEMeasurable.div' [MeasurableDiv₂ G] (hf : AEMeasurable f μ) (hg : AEMeasurable g μ) : AEMeasurable (f / g) μ := measurable_div.comp_aemeasurable (hf.prod_mk hg) @@ -528,42 +530,52 @@ instance Subgroup.measurableSMul {G α} [MeasurableSpace G] [MeasurableSpace α] section SMul -variable {M β α : Type*} [MeasurableSpace M] [MeasurableSpace β] [_root_.SMul M β] - {m : MeasurableSpace α} {f : α → M} {g : α → β} +variable {M X α β : Type*} [MeasurableSpace M] [MeasurableSpace X] [SMul M X] + {m : MeasurableSpace α} {mβ : MeasurableSpace β} {f : α → M} {g : α → X} @[to_additive (attr := fun_prop, aesop safe 20 apply (rule_sets := [Measurable]))] -theorem Measurable.smul [MeasurableSMul₂ M β] (hf : Measurable f) (hg : Measurable g) : +theorem Measurable.smul [MeasurableSMul₂ M X] (hf : Measurable f) (hg : Measurable g) : Measurable fun x => f x • g x := measurable_smul.comp (hf.prod_mk hg) +/-- Compositional version of `Measurable.smul` for use by `fun_prop`. -/ +@[to_additive (attr := fun_prop) +"Compositional version of `Measurable.vadd` for use by `fun_prop`."] +lemma Measurable.smul' [MeasurableSMul₂ M X] {f : α → β → M} {g : α → β → X} {h : α → β} + (hf : Measurable ↿f) (hg : Measurable ↿g) (hh : Measurable h) : + Measurable fun a ↦ (f a • g a) (h a) := by simp; fun_prop + @[to_additive (attr := fun_prop, aesop safe 20 apply (rule_sets := [Measurable]))] -theorem AEMeasurable.smul [MeasurableSMul₂ M β] {μ : Measure α} (hf : AEMeasurable f μ) +theorem AEMeasurable.smul [MeasurableSMul₂ M X] {μ : Measure α} (hf : AEMeasurable f μ) (hg : AEMeasurable g μ) : AEMeasurable (fun x => f x • g x) μ := MeasurableSMul₂.measurable_smul.comp_aemeasurable (hf.prod_mk hg) @[to_additive] -instance (priority := 100) MeasurableSMul₂.toMeasurableSMul [MeasurableSMul₂ M β] : - MeasurableSMul M β := +instance (priority := 100) MeasurableSMul₂.toMeasurableSMul [MeasurableSMul₂ M X] : + MeasurableSMul M X := ⟨fun _ => measurable_const.smul measurable_id, fun _ => measurable_id.smul measurable_const⟩ -variable [MeasurableSMul M β] {μ : Measure α} +variable [MeasurableSMul M X] {μ : Measure α} @[to_additive (attr := measurability)] -theorem Measurable.smul_const (hf : Measurable f) (y : β) : Measurable fun x => f x • y := +theorem Measurable.smul_const (hf : Measurable f) (y : X) : Measurable fun x => f x • y := (MeasurableSMul.measurable_smul_const y).comp hf @[to_additive (attr := measurability)] -theorem AEMeasurable.smul_const (hf : AEMeasurable f μ) (y : β) : +theorem AEMeasurable.smul_const (hf : AEMeasurable f μ) (y : X) : AEMeasurable (fun x => f x • y) μ := (MeasurableSMul.measurable_smul_const y).comp_aemeasurable hf -@[to_additive (attr := measurability)] -theorem Measurable.const_smul' (hg : Measurable g) (c : M) : Measurable fun x => c • g x := +@[to_additive (attr := fun_prop, measurability)] +theorem Measurable.const_smul (hg : Measurable g) (c : M) : Measurable (c • g) := (MeasurableSMul.measurable_const_smul c).comp hg -@[to_additive (attr := measurability)] -theorem Measurable.const_smul (hg : Measurable g) (c : M) : Measurable (c • g) := - hg.const_smul' c +/-- Compositional version of `Measurable.const_smul` for use by `fun_prop`. -/ +@[to_additive (attr := fun_prop) +"Compositional version of `Measurable.const_vadd` for use by `fun_prop`."] +lemma Measurable.const_smul' {g : α → β → X} {h : α → β} (hg : Measurable ↿g) (hh : Measurable h) + (c : M) : Measurable fun a ↦ (c • g a) (h a) := + (hg.comp <| measurable_id.prod_mk hh).const_smul _ @[to_additive (attr := measurability)] theorem AEMeasurable.const_smul' (hg : AEMeasurable g μ) (c : M) : @@ -638,12 +650,12 @@ variable {G : Type*} [Group G] [MeasurableSpace G] [MulAction G β] [MeasurableS @[to_additive] theorem measurable_const_smul_iff (c : G) : (Measurable fun x => c • f x) ↔ Measurable f := - ⟨fun h => by simpa only [inv_smul_smul] using h.const_smul' c⁻¹, fun h => h.const_smul c⟩ + ⟨fun h => by simpa [inv_smul_smul, Pi.smul_def] using h.const_smul c⁻¹, fun h => h.const_smul c⟩ @[to_additive] theorem aemeasurable_const_smul_iff (c : G) : AEMeasurable (fun x => c • f x) μ ↔ AEMeasurable f μ := - ⟨fun h => by simpa only [inv_smul_smul] using h.const_smul' c⁻¹, fun h => h.const_smul c⟩ + ⟨fun h => by simpa [inv_smul_smul, Pi.smul_def] using h.const_smul c⁻¹, fun h => h.const_smul c⟩ @[to_additive] instance Units.instMeasurableSpace : MeasurableSpace Mˣ := MeasurableSpace.comap ((↑) : Mˣ → M) ‹_› @@ -782,8 +794,8 @@ end Monoid section CommMonoid -variable {M ι α : Type*} [CommMonoid M] [MeasurableSpace M] [MeasurableMul₂ M] - {m : MeasurableSpace α} {μ : Measure α} {f : ι → α → M} +variable {M ι α β : Type*} [CommMonoid M] [MeasurableSpace M] [MeasurableMul₂ M] + {m : MeasurableSpace α} {mβ : MeasurableSpace β} {μ : Measure α} {f : ι → α → M} @[to_additive (attr := measurability)] theorem Multiset.measurable_prod' (l : Multiset (α → M)) (hl : ∀ f ∈ l, Measurable f) : @@ -807,15 +819,18 @@ theorem Multiset.aemeasurable_prod (s : Multiset (α → M)) (hs : ∀ f ∈ s, AEMeasurable (fun x => (s.map fun f : α → M => f x).prod) μ := by simpa only [← Pi.multiset_prod_apply] using s.aemeasurable_prod' hs -@[to_additive (attr := measurability)] -theorem Finset.measurable_prod' (s : Finset ι) (hf : ∀ i ∈ s, Measurable (f i)) : - Measurable (∏ i ∈ s, f i) := - Finset.prod_induction _ _ (fun _ _ => Measurable.mul) (@measurable_one M _ _ _ _) hf - -@[to_additive (attr := measurability)] +@[to_additive (attr := fun_prop, measurability)] theorem Finset.measurable_prod (s : Finset ι) (hf : ∀ i ∈ s, Measurable (f i)) : - Measurable fun a => ∏ i ∈ s, f i a := by - simpa only [← Finset.prod_apply] using s.measurable_prod' hf + Measurable fun a ↦ ∏ i ∈ s, f i a := by + simp_rw [← Finset.prod_apply] + exact Finset.prod_induction _ _ (fun _ _ => Measurable.mul) (@measurable_one M _ _ _ _) hf + +/-- Compositional version of `Finset.measurable_prod` for use by `fun_prop`. -/ +@[to_additive (attr := measurability, fun_prop) +"Compositional version of `Finset.measurable_sum` for use by `fun_prop`."] +lemma Finset.measurable_prod' {f : ι → α → β → M} {g : α → β} {s : Finset ι} + (hf : ∀ i, Measurable ↿(f i)) (hg : Measurable g) : + Measurable fun a ↦ (∏ i ∈ s, f i a) (g a) := by simp; fun_prop @[to_additive (attr := measurability)] theorem Finset.aemeasurable_prod' (s : Finset ι) (hf : ∀ i ∈ s, AEMeasurable (f i) μ) : diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 1ea3d369e9733..411b231ffbadc 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -588,7 +588,7 @@ variable (𝕜) /-- The Bochner integral in L1 space as a continuous linear map. -/ nonrec def integralCLM' : (α →₁[μ] E) →L[𝕜] E := (integralCLM' α E 𝕜 μ).extend (coeToLp α E 𝕜) (simpleFunc.denseRange one_ne_top) - simpleFunc.uniformInducing + simpleFunc.isUniformInducing variable {𝕜} diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index ac2ffb0b0ce8f..737247f5c8918 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -1075,7 +1075,7 @@ theorem integral_comp_rpow_Ioi (g : ℝ → E) {p : ℝ} (hp : p ≠ 0) : rcases lt_or_gt_of_ne hp with (h | h) · apply StrictAntiOn.injOn intro x hx y hy hxy - rw [← inv_lt_inv (rpow_pos_of_pos hx p) (rpow_pos_of_pos hy p), ← rpow_neg (le_of_lt hx), + rw [← inv_lt_inv₀ (rpow_pos_of_pos hx p) (rpow_pos_of_pos hy p), ← rpow_neg (le_of_lt hx), ← rpow_neg (le_of_lt hy)] exact rpow_lt_rpow (le_of_lt hx) hxy (neg_pos.mpr h) exact StrictMonoOn.injOn fun x hx y _ hxy => rpow_lt_rpow (mem_Ioi.mp hx).le hxy h @@ -1129,7 +1129,7 @@ theorem integrableOn_Ioi_comp_rpow_iff [NormedSpace ℝ E] (f : ℝ → E) {p : rcases lt_or_gt_of_ne hp with (h | h) · apply StrictAntiOn.injOn intro x hx y hy hxy - rw [← inv_lt_inv (rpow_pos_of_pos hx p) (rpow_pos_of_pos hy p), ← rpow_neg (le_of_lt hx), ← + rw [← inv_lt_inv₀ (rpow_pos_of_pos hx p) (rpow_pos_of_pos hy p), ← rpow_neg (le_of_lt hx), ← rpow_neg (le_of_lt hy)] exact rpow_lt_rpow (le_of_lt hx) hxy (neg_pos.mpr h) exact StrictMonoOn.injOn fun x hx y _hy hxy => rpow_lt_rpow (mem_Ioi.mp hx).le hxy h diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index 9fced5cac44e7..675449008740c 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -530,7 +530,7 @@ theorem lintegral_add_aux {f g : α → ℝ≥0∞} (hf : Measurable f) (hg : Me funext n rw [← SimpleFunc.add_lintegral, ← SimpleFunc.lintegral_eq_lintegral] simp only [Pi.add_apply, SimpleFunc.coe_add] - · measurability + · fun_prop · intro i j h a dsimp gcongr <;> exact monotone_eapprox _ h _ diff --git a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean index 7b324585c63a5..424952fb5cb92 100644 --- a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean +++ b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean @@ -180,8 +180,7 @@ theorem lintegral_mul_norm_pow_le {α} [MeasurableSpace α] {μ : Measure α} · rw [add_zero] at hpq simp [hpq] have h2p : 1 < 1 / p := by - rw [one_div] - apply one_lt_inv hp + rw [one_div, one_lt_inv₀ hp] linarith have h2pq : (1 / p)⁻¹ + (1 / q)⁻¹ = 1 := by simp [hp.ne', hq.ne', hpq] have := ENNReal.lintegral_mul_le_Lp_mul_Lq μ ⟨h2p, h2pq⟩ (hf.pow_const p) (hg.pow_const q) diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index 02dbf1cdd368b..b1c2e325add44 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -920,18 +920,18 @@ variable (𝕜) [NontriviallyNormedField 𝕜] [NormedSpace 𝕜 E] [NormedSpace def setToL1' (hT : DominatedFinMeasAdditive μ T C) (h_smul : ∀ c : 𝕜, ∀ s x, T s (c • x) = c • T s x) : (α →₁[μ] E) →L[𝕜] F := (setToL1SCLM' α E 𝕜 μ hT h_smul).extend (coeToLp α E 𝕜) (simpleFunc.denseRange one_ne_top) - simpleFunc.uniformInducing + simpleFunc.isUniformInducing variable {𝕜} /-- Extend `Set α → E →L[ℝ] F` to `(α →₁[μ] E) →L[ℝ] F`. -/ def setToL1 (hT : DominatedFinMeasAdditive μ T C) : (α →₁[μ] E) →L[ℝ] F := (setToL1SCLM α E μ hT).extend (coeToLp α E ℝ) (simpleFunc.denseRange one_ne_top) - simpleFunc.uniformInducing + simpleFunc.isUniformInducing theorem setToL1_eq_setToL1SCLM (hT : DominatedFinMeasAdditive μ T C) (f : α →₁ₛ[μ] E) : setToL1 hT f = setToL1SCLM α E μ hT f := - uniformly_extend_of_ind simpleFunc.uniformInducing (simpleFunc.denseRange one_ne_top) + uniformly_extend_of_ind simpleFunc.isUniformInducing (simpleFunc.denseRange one_ne_top) (setToL1SCLM α E μ hT).uniformContinuous _ theorem setToL1_eq_setToL1' (hT : DominatedFinMeasAdditive μ T C) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index b47c60a0d8a7f..61c55dee0d648 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -231,7 +231,7 @@ theorem Subsingleton.measurable [Subsingleton α] : Measurable f := fun _ _ => theorem measurable_of_subsingleton_codomain [Subsingleton β] (f : α → β) : Measurable f := fun s _ => Subsingleton.set_cases MeasurableSet.empty MeasurableSet.univ s -@[to_additive (attr := measurability)] +@[to_additive (attr := measurability, fun_prop)] theorem measurable_one [One α] : Measurable (1 : β → α) := @measurable_const _ _ _ _ 1 diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 19be50fbff257..9ff18dc3605be 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -516,7 +516,8 @@ variable [MeasurableSpace β] /-- A class registering that either `α` is countable or `β` is a countably generated measurable space. -/ -class CountableOrCountablyGenerated (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] : Prop := +class CountableOrCountablyGenerated (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] : + Prop where countableOrCountablyGenerated : Countable α ∨ MeasurableSpace.CountablyGenerated β instance instCountableOrCountablyGeneratedOfCountable [h1 : Countable α] : diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean b/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean index 6345012319658..5c9f7d9d195dc 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean @@ -534,7 +534,7 @@ variable [MeasurableSpace α] [MeasurableSpace β] [DiscreteMeasurableSpace α] @[measurability] lemma MeasurableSet.of_discrete : MeasurableSet s := DiscreteMeasurableSpace.forall_measurableSet _ -@[measurability] lemma Measurable.of_discrete : Measurable f := fun _ _ ↦ .of_discrete +@[measurability, fun_prop] lemma Measurable.of_discrete : Measurable f := fun _ _ ↦ .of_discrete @[deprecated MeasurableSet.of_discrete (since := "2024-08-25")] lemma measurableSet_discrete (s : Set α) : MeasurableSet s := .of_discrete diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean index 539df69ae463e..5cb861c7dd95a 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean @@ -60,6 +60,8 @@ structure MeasurableEmbedding [MeasurableSpace α] [MeasurableSpace β] (f : α /-- The image of a measurable set under a measurable embedding is a measurable set. -/ protected measurableSet_image' : ∀ ⦃s⦄, MeasurableSet s → MeasurableSet (f '' s) +attribute [fun_prop] MeasurableEmbedding.measurable + namespace MeasurableEmbedding variable {mα : MeasurableSpace α} [MeasurableSpace β] [MeasurableSpace γ] {f : α → β} {g : β → γ} @@ -155,7 +157,7 @@ instance instEquivLike : EquivLike (α ≃ᵐ β) α β where theorem coe_toEquiv (e : α ≃ᵐ β) : (e.toEquiv : α → β) = e := rfl -@[measurability] +@[measurability, fun_prop] protected theorem measurable (e : α ≃ᵐ β) : Measurable (e : α → β) := e.measurable_toFun diff --git a/Mathlib/MeasureTheory/Measure/AddContent.lean b/Mathlib/MeasureTheory/Measure/AddContent.lean index b2e51584933b0..e07a40a78c24f 100644 --- a/Mathlib/MeasureTheory/Measure/AddContent.lean +++ b/Mathlib/MeasureTheory/Measure/AddContent.lean @@ -63,9 +63,8 @@ instance : Inhabited (AddContent C) := sUnion' := by simp }⟩ instance : DFunLike (AddContent C) (Set α) (fun _ ↦ ℝ≥0∞) where - coe := fun m s ↦ m.toFun s - coe_injective' := by - intro m m' h + coe m s := m.toFun s + coe_injective' m m' _ := by cases m cases m' congr diff --git a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean index 5b7eb39f8469e..fd684e18cafc6 100644 --- a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean +++ b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean @@ -28,10 +28,6 @@ import Mathlib.MeasureTheory.Integral.BoundedContinuousFunction probability measures on a separable space coincides with the topology of convergence in distribution, and in particular convergence in distribution is then pseudometrizable. -## TODO - -* Show that in Borel spaces, the Lévy-Prokhorov distance is a metric; not just a pseudometric. - ## Tags finite measure, probability measure, weak convergence, convergence in distribution, metrizability @@ -198,6 +194,9 @@ lemma levyProkhorovDist_triangle [OpensMeasurableSpace Ω] (μ ν κ : Measure when they are to be equipped with the Lévy-Prokhorov distance. -/ def LevyProkhorov (α : Type*) := α +/-- The "identity" equivalence between the type synonym `LevyProkhorov α` and `α`. -/ +def LevyProkhorov.equiv (α : Type*) : LevyProkhorov α ≃ α := Equiv.refl _ + variable [OpensMeasurableSpace Ω] /-- The Lévy-Prokhorov distance `levyProkhorovEDist` makes `Measure Ω` a pseudoemetric @@ -219,9 +218,43 @@ noncomputable instance levyProkhorovDist_pseudoMetricSpace_finiteMeasure : dist_triangle μ ν κ := levyProkhorovDist_triangle _ _ _ edist_dist μ ν := by simp [← ENNReal.ofReal_coe_nnreal] -/-- The Lévy-Prokhorov distance `levyProkhorovDist` makes `ProbabilityMeasure Ω` a pseudoemetric +lemma measure_le_measure_closure_of_levyProkhorovEDist_eq_zero {μ ν : Measure Ω} + (hLP : levyProkhorovEDist μ ν = 0) {s : Set Ω} (s_mble : MeasurableSet s) + (h_finite : ∃ δ > 0, ν (thickening δ s) ≠ ∞) : + μ s ≤ ν (closure s) := by + have key : Tendsto (fun ε ↦ ν (thickening ε.toReal s)) (𝓝[>] (0 : ℝ≥0∞)) (𝓝 (ν (closure s))) := by + have aux : Tendsto ENNReal.toReal (𝓝[>] 0) (𝓝[>] 0) := by + apply tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within (s := Ioi 0) ENNReal.toReal + · exact tendsto_nhdsWithin_of_tendsto_nhds (continuousAt_toReal zero_ne_top).tendsto + · filter_upwards [Ioo_mem_nhdsWithin_Ioi ⟨le_rfl, zero_lt_one⟩] with x hx + exact toReal_pos hx.1.ne.symm <| ne_top_of_lt hx.2 + exact (tendsto_measure_thickening h_finite).comp aux + have obs := Tendsto.add key (tendsto_nhdsWithin_of_tendsto_nhds tendsto_id) + simp only [id_eq, add_zero] at obs + apply ge_of_tendsto (b := μ s) obs + filter_upwards [self_mem_nhdsWithin] with ε ε_pos + exact left_measure_le_of_levyProkhorovEDist_lt (B_mble := s_mble) (hLP ▸ ε_pos) + +/-- Two measures at vanishing Lévy-Prokhorov distance from each other assign the same values to all +closed sets. -/ +lemma measure_eq_measure_of_levyProkhorovEDist_eq_zero_of_isClosed {μ ν : Measure Ω} + (hLP : levyProkhorovEDist μ ν = 0) {s : Set Ω} (s_closed : IsClosed s) + (hμs : ∃ δ > 0, μ (thickening δ s) ≠ ∞) (hνs : ∃ δ > 0, ν (thickening δ s) ≠ ∞) : + μ s = ν s := by + apply le_antisymm + · exact measure_le_measure_closure_of_levyProkhorovEDist_eq_zero + hLP s_closed.measurableSet hνs |>.trans <| + le_of_eq (congr_arg _ s_closed.closure_eq) + · exact measure_le_measure_closure_of_levyProkhorovEDist_eq_zero + (levyProkhorovEDist_comm μ ν ▸ hLP) s_closed.measurableSet hμs |>.trans <| + le_of_eq (congr_arg _ s_closed.closure_eq) + +/-- The Lévy-Prokhorov distance `levyProkhorovDist` makes `ProbabilityMeasure Ω` a pseudometric space. The instance is recorded on the type synonym -`LevyProkhorov (ProbabilityMeasure Ω) := ProbabilityMeasure Ω`. -/ +`LevyProkhorov (ProbabilityMeasure Ω) := ProbabilityMeasure Ω`. + +Note: For this pseudometric to give the topology of convergence in distribution, one must +furthermore assume that `Ω` is separable. -/ noncomputable instance levyProkhorovDist_pseudoMetricSpace_probabilityMeasure : PseudoMetricSpace (LevyProkhorov (ProbabilityMeasure Ω)) where dist μ ν := levyProkhorovDist μ.toMeasure ν.toMeasure @@ -233,6 +266,29 @@ noncomputable instance levyProkhorovDist_pseudoMetricSpace_probabilityMeasure : lemma LevyProkhorov.dist_def (μ ν : LevyProkhorov (ProbabilityMeasure Ω)) : dist μ ν = levyProkhorovDist μ.toMeasure ν.toMeasure := rfl +/-- If `Ω` is a Borel space, then the Lévy-Prokhorov distance `levyProkhorovDist` makes +`ProbabilityMeasure Ω` a metric space. The instance is recorded on the type synonym +`LevyProkhorov (ProbabilityMeasure Ω) := ProbabilityMeasure Ω`. + +Note: For this metric to give the topology of convergence in distribution, one must +furthermore assume that `Ω` is separable. -/ +noncomputable instance levyProkhorovDist_metricSpace_probabilityMeasure [BorelSpace Ω] : + MetricSpace (LevyProkhorov (ProbabilityMeasure Ω)) where + eq_of_dist_eq_zero := by + intro μ ν h + apply (LevyProkhorov.equiv _).injective + apply ProbabilityMeasure.toMeasure_injective + apply ext_of_generate_finite _ ?_ isPiSystem_isClosed ?_ (by simp) + · rw [BorelSpace.measurable_eq (α := Ω), borel_eq_generateFrom_isClosed] + · intro A A_closed + apply measure_eq_measure_of_levyProkhorovEDist_eq_zero_of_isClosed + · simpa only [levyProkhorovEDist_ne_top μ.toMeasure ν.toMeasure, mem_setOf_eq, + or_false, ne_eq, zero_ne_top, not_false_eq_true, zero_toReal] + using (toReal_eq_zero_iff _).mp h + · exact A_closed + · exact ⟨1, Real.zero_lt_one, measure_ne_top _ _⟩ + · exact ⟨1, Real.zero_lt_one, measure_ne_top _ _⟩ + /-- A simple sufficient condition for bounding `levyProkhorovEDist` between probability measures from above. The condition involves only one of two natural bounds, the other bound is for free. -/ lemma levyProkhorovEDist_le_of_forall_le @@ -279,20 +335,6 @@ open BoundedContinuousFunction variable {ι : Type*} {Ω : Type*} [MeasurableSpace Ω] -/-- Coercion from the type synonym `LevyProkhorov (ProbabilityMeasure Ω)` -to `ProbabilityMeasure Ω`. -/ -def LevyProkhorov.toProbabilityMeasure (μ : LevyProkhorov (ProbabilityMeasure Ω)) : - ProbabilityMeasure Ω := μ - -/-- Coercion to the type synonym `LevyProkhorov (ProbabilityMeasure Ω)` -from `ProbabilityMeasure Ω`. -/ -def ProbabilityMeasure.toLevyProkhorov (μ : ProbabilityMeasure Ω) : - LevyProkhorov (ProbabilityMeasure Ω) := μ - -/-- Coercion from the type synonym `LevyProkhorov (FiniteMeasure Ω)` to `FiniteMeasure Ω`. -/ -def LevyProkhorov.finiteMeasure (μ : LevyProkhorov (FiniteMeasure Ω)) : - FiniteMeasure Ω := μ - variable [PseudoMetricSpace Ω] [OpensMeasurableSpace Ω] /-- A version of the layer cake formula for bounded continuous functions which have finite integral: @@ -384,13 +426,13 @@ lemma tendsto_integral_meas_thickening_le (f : Ω →ᵇ ℝ) · exact isClosed_le continuous_const f.continuous · exact measure_ne_top _ _ -/-- The coercion `LevyProkhorov (ProbabilityMeasure Ω) → ProbabilityMeasure Ω` is continuous. -/ -lemma LevyProkhorov.continuous_toProbabilityMeasure : - Continuous (LevyProkhorov.toProbabilityMeasure (Ω := Ω)) := by +/-- The identity map `LevyProkhorov (ProbabilityMeasure Ω) → ProbabilityMeasure Ω` is continuous. -/ +lemma LevyProkhorov.continuous_equiv_probabilityMeasure : + Continuous (LevyProkhorov.equiv (α := ProbabilityMeasure Ω)) := by refine SeqContinuous.continuous ?_ intro μs ν hμs - set P := ν.toProbabilityMeasure -- more palatable notation - set Ps := fun n ↦ (μs n).toProbabilityMeasure -- more palatable notation + set P := LevyProkhorov.equiv _ ν -- more palatable notation + set Ps := fun n ↦ LevyProkhorov.equiv _ (μs n) -- more palatable notation rw [ProbabilityMeasure.tendsto_iff_forall_integral_tendsto] refine fun f ↦ tendsto_integral_of_forall_limsup_integral_le_integral ?_ f intro f f_nn @@ -433,9 +475,8 @@ lemma LevyProkhorov.continuous_toProbabilityMeasure : · rw [ENNReal.ofReal_add (by positivity) (by positivity), ← add_zero (levyProkhorovEDist _ _)] apply ENNReal.add_lt_add_of_le_of_lt (levyProkhorovEDist_ne_top _ _) (le_of_eq ?_) (ofReal_pos.mpr εs_pos) - rw [LevyProkhorov.dist_def, levyProkhorovDist, - ofReal_toReal (levyProkhorovEDist_ne_top _ _)] - simp only [Ps, P, LevyProkhorov.toProbabilityMeasure] + rw [LevyProkhorov.dist_def, levyProkhorovDist, ofReal_toReal (levyProkhorovEDist_ne_top _ _)] + rfl · exact Eventually.of_forall f_nn · simp only [IsCoboundedUnder, IsCobounded, eventually_map, eventually_atTop, forall_exists_index] @@ -444,9 +485,9 @@ lemma LevyProkhorov.continuous_toProbabilityMeasure : /-- The topology of the Lévy-Prokhorov metric is at least as fine as the topology of convergence in distribution. -/ theorem levyProkhorov_le_convergenceInDistribution : - TopologicalSpace.coinduced (LevyProkhorov.toProbabilityMeasure (Ω := Ω)) inferInstance + TopologicalSpace.coinduced (LevyProkhorov.equiv (α := ProbabilityMeasure Ω)) inferInstance ≤ (inferInstance : TopologicalSpace (ProbabilityMeasure Ω)) := - (LevyProkhorov.continuous_toProbabilityMeasure).coinduced_le + (LevyProkhorov.continuous_equiv_probabilityMeasure).coinduced_le end Levy_Prokhorov_is_finer @@ -456,13 +497,34 @@ section Levy_Prokhorov_metrizes_convergence_in_distribution open BoundedContinuousFunction TopologicalSpace -variable {ι : Type*} (Ω : Type*) [PseudoMetricSpace Ω] +variable {ι : Type*} {Ω : Type*} [PseudoMetricSpace Ω] variable [MeasurableSpace Ω] [OpensMeasurableSpace Ω] +lemma ProbabilityMeasure.toMeasure_add_pos_gt_mem_nhds (P : ProbabilityMeasure Ω) + {G : Set Ω} (G_open : IsOpen G) {ε : ℝ≥0∞} (ε_pos : 0 < ε) : + {Q | P.toMeasure G < Q.toMeasure G + ε} ∈ 𝓝 P := by + by_cases easy : P.toMeasure G < ε + · exact Eventually.of_forall (fun _ ↦ lt_of_lt_of_le easy le_add_self) + by_cases ε_top : ε = ∞ + · simp [ε_top, measure_lt_top] + simp only [not_lt] at easy + have aux : P.toMeasure G - ε < liminf (fun Q ↦ Q.toMeasure G) (𝓝 P) := by + apply lt_of_lt_of_le (ENNReal.sub_lt_self (measure_lt_top _ _).ne _ _) + <| ProbabilityMeasure.le_liminf_measure_open_of_tendsto tendsto_id G_open + · exact (lt_of_lt_of_le ε_pos easy).ne.symm + · exact ε_pos.ne.symm + filter_upwards [gt_mem_sets_of_limsInf_gt (α := ℝ≥0∞) isBounded_ge_of_bot + (show P.toMeasure G - ε < limsInf ((𝓝 P).map (fun Q ↦ Q.toMeasure G)) from aux)] with Q hQ + simp only [preimage_setOf_eq, mem_setOf_eq] at hQ + convert ENNReal.add_lt_add_right ε_top hQ + exact (tsub_add_cancel_of_le easy).symm + +variable [SeparableSpace Ω] + +variable (Ω) in /-- In a separable pseudometric space, for any ε > 0 there exists a countable collection of disjoint Borel measurable subsets of diameter at most ε that cover the whole space. -/ -lemma SeparableSpace.exists_measurable_partition_diam_le [SeparableSpace Ω] - {ε : ℝ} (ε_pos : 0 < ε) : +lemma SeparableSpace.exists_measurable_partition_diam_le {ε : ℝ} (ε_pos : 0 < ε) : ∃ (As : ℕ → Set Ω), (∀ n, MeasurableSet (As n)) ∧ (∀ n, Bornology.IsBounded (As n)) ∧ (∀ n, diam (As n) ≤ ε) ∧ (⋃ n, As n = univ) ∧ (Pairwise (fun (n m : ℕ) ↦ Disjoint (As n) (As m))) := by @@ -492,29 +554,8 @@ lemma SeparableSpace.exists_measurable_partition_diam_le [SeparableSpace Ω] simpa only [← aux] using iUnion_disjointed · exact disjoint_disjointed Bs -variable {Ω} - -lemma ProbabilityMeasure.toMeasure_add_pos_gt_mem_nhds (P : ProbabilityMeasure Ω) - {G : Set Ω} (G_open : IsOpen G) {ε : ℝ≥0∞} (ε_pos : 0 < ε) : - {Q | P.toMeasure G < Q.toMeasure G + ε} ∈ 𝓝 P := by - by_cases easy : P.toMeasure G < ε - · exact Eventually.of_forall (fun _ ↦ lt_of_lt_of_le easy le_add_self) - by_cases ε_top : ε = ∞ - · simp [ε_top, measure_lt_top] - simp only [not_lt] at easy - have aux : P.toMeasure G - ε < liminf (fun Q ↦ Q.toMeasure G) (𝓝 P) := by - apply lt_of_lt_of_le (ENNReal.sub_lt_self (measure_lt_top _ _).ne _ _) - <| ProbabilityMeasure.le_liminf_measure_open_of_tendsto tendsto_id G_open - · exact (lt_of_lt_of_le ε_pos easy).ne.symm - · exact ε_pos.ne.symm - filter_upwards [gt_mem_sets_of_limsInf_gt (α := ℝ≥0∞) isBounded_ge_of_bot - (show P.toMeasure G - ε < limsInf ((𝓝 P).map (fun Q ↦ Q.toMeasure G)) from aux)] with Q hQ - simp only [preimage_setOf_eq, mem_setOf_eq] at hQ - convert ENNReal.add_lt_add_right ε_top hQ - exact (tsub_add_cancel_of_le easy).symm - -lemma ProbabilityMeasure.continuous_toLevyProkhorov [SeparableSpace Ω] : - Continuous (ProbabilityMeasure.toLevyProkhorov (Ω := Ω)) := by +lemma LevyProkhorov.continuous_equiv_symm_probabilityMeasure : + Continuous (LevyProkhorov.equiv (α := ProbabilityMeasure Ω)).symm := by -- We check continuity of `id : ProbabilityMeasure Ω → LevyProkhorov (ProbabilityMeasure Ω)` at -- each point `P : ProbabilityMeasure Ω`. rw [continuous_iff_continuousAt] @@ -612,22 +653,22 @@ lemma ProbabilityMeasure.continuous_toLevyProkhorov [SeparableSpace Ω] : /-- The topology of the Lévy-Prokhorov metric on probability measures on a separable space coincides with the topology of convergence in distribution. -/ -theorem levyProkhorov_eq_convergenceInDistribution [SeparableSpace Ω] : +theorem levyProkhorov_eq_convergenceInDistribution : (inferInstance : TopologicalSpace (ProbabilityMeasure Ω)) - = TopologicalSpace.coinduced (LevyProkhorov.toProbabilityMeasure (Ω := Ω)) inferInstance := - le_antisymm (ProbabilityMeasure.continuous_toLevyProkhorov (Ω := Ω)).coinduced_le + = TopologicalSpace.coinduced (LevyProkhorov.equiv _) inferInstance := + le_antisymm (LevyProkhorov.continuous_equiv_symm_probabilityMeasure (Ω := Ω)).coinduced_le levyProkhorov_le_convergenceInDistribution /-- The identity map is a homeomorphism from `ProbabilityMeasure Ω` with the topology of convergence in distribution to `ProbabilityMeasure Ω` with the Lévy-Prokhorov (pseudo)metric. -/ -def homeomorph_probabilityMeasure_levyProkhorov [SeparableSpace Ω] : +def homeomorph_probabilityMeasure_levyProkhorov : ProbabilityMeasure Ω ≃ₜ LevyProkhorov (ProbabilityMeasure Ω) where - toFun := ProbabilityMeasure.toLevyProkhorov (Ω := Ω) - invFun := LevyProkhorov.toProbabilityMeasure (Ω := Ω) + toFun := LevyProkhorov.equiv _ + invFun := (LevyProkhorov.equiv _).symm left_inv := congrFun rfl right_inv := congrFun rfl - continuous_toFun := ProbabilityMeasure.continuous_toLevyProkhorov - continuous_invFun := LevyProkhorov.continuous_toProbabilityMeasure + continuous_toFun := LevyProkhorov.continuous_equiv_symm_probabilityMeasure + continuous_invFun := LevyProkhorov.continuous_equiv_probabilityMeasure /-- The topology of convergence in distribution on a separable space is pseudo-metrizable. -/ instance (X : Type*) [TopologicalSpace X] [PseudoMetrizableSpace X] [SeparableSpace X] @@ -636,6 +677,13 @@ instance (X : Type*) [TopologicalSpace X] [PseudoMetrizableSpace X] [SeparableSp letI : PseudoMetricSpace X := TopologicalSpace.pseudoMetrizableSpacePseudoMetric X (homeomorph_probabilityMeasure_levyProkhorov (Ω := X)).inducing.pseudoMetrizableSpace +/-- The topology of convergence in distribution on a separable Borel space is metrizable. -/ +instance instMetrizableSpaceProbabilityMeasure (X : Type*) [TopologicalSpace X] + [PseudoMetrizableSpace X] [SeparableSpace X] [MeasurableSpace X] [BorelSpace X] : + MetrizableSpace (ProbabilityMeasure X) := by + letI : PseudoMetricSpace X := TopologicalSpace.pseudoMetrizableSpacePseudoMetric X + exact homeomorph_probabilityMeasure_levyProkhorov.embedding.metrizableSpace + end Levy_Prokhorov_metrizes_convergence_in_distribution end MeasureTheory -- namespace diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 8ce344e3a9ca8..e3b46dcc67079 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -543,25 +543,27 @@ theorem measure_iInter_eq_iInf' {α ι : Type*} [MeasurableSpace α] {μ : Measu /-- Continuity from below: the measure of the union of an increasing sequence of (not necessarily measurable) sets is the limit of the measures. -/ -theorem tendsto_measure_iUnion_atTop [Preorder ι] [IsDirected ι (· ≤ ·)] - [IsCountablyGenerated (atTop : Filter ι)] {s : ι → Set α} (hm : Monotone s) : - Tendsto (μ ∘ s) atTop (𝓝 (μ (⋃ n, s n))) := by +theorem tendsto_measure_iUnion_atTop [Preorder ι] [IsCountablyGenerated (atTop : Filter ι)] + {s : ι → Set α} (hm : Monotone s) : Tendsto (μ ∘ s) atTop (𝓝 (μ (⋃ n, s n))) := by + refine .of_neBot_imp fun h ↦ ?_ + have := (atTop_neBot_iff.1 h).2 rw [hm.measure_iUnion] exact tendsto_atTop_iSup fun n m hnm => measure_mono <| hm hnm @[deprecated (since := "2024-09-01")] alias tendsto_measure_iUnion := tendsto_measure_iUnion_atTop -theorem tendsto_measure_iUnion_atBot [Preorder ι] [IsDirected ι (· ≥ ·)] - [IsCountablyGenerated (atBot : Filter ι)] {s : ι → Set α} (hm : Antitone s) : - Tendsto (μ ∘ s) atBot (𝓝 (μ (⋃ n, s n))) := +theorem tendsto_measure_iUnion_atBot [Preorder ι] [IsCountablyGenerated (atBot : Filter ι)] + {s : ι → Set α} (hm : Antitone s) : Tendsto (μ ∘ s) atBot (𝓝 (μ (⋃ n, s n))) := tendsto_measure_iUnion_atTop (ι := ιᵒᵈ) hm.dual_left /-- Continuity from below: the measure of the union of a sequence of (not necessarily measurable) sets is the limit of the measures of the partial unions. -/ theorem tendsto_measure_iUnion_accumulate {α ι : Type*} - [Preorder ι] [IsDirected ι (· ≤ ·)] [IsCountablyGenerated (atTop : Filter ι)] + [Preorder ι] [IsCountablyGenerated (atTop : Filter ι)] [MeasurableSpace α] {μ : Measure α} {f : ι → Set α} : Tendsto (fun i ↦ μ (Accumulate f i)) atTop (𝓝 (μ (⋃ i, f i))) := by + refine .of_neBot_imp fun h ↦ ?_ + have := (atTop_neBot_iff.1 h).2 rw [measure_iUnion_eq_iSup_accumulate] exact tendsto_atTop_iSup fun i j hij ↦ by gcongr @@ -570,9 +572,11 @@ alias tendsto_measure_iUnion' := tendsto_measure_iUnion_accumulate /-- Continuity from above: the measure of the intersection of a decreasing sequence of measurable sets is the limit of the measures. -/ -theorem tendsto_measure_iInter [Countable ι] [Preorder ι] [IsDirected ι (· ≤ ·)] {s : ι → Set α} +theorem tendsto_measure_iInter [Countable ι] [Preorder ι] {s : ι → Set α} (hs : ∀ n, NullMeasurableSet (s n) μ) (hm : Antitone s) (hf : ∃ i, μ (s i) ≠ ∞) : Tendsto (μ ∘ s) atTop (𝓝 (μ (⋂ n, s n))) := by + refine .of_neBot_imp fun h ↦ ?_ + have := (atTop_neBot_iff.1 h).2 rw [measure_iInter_eq_iInf hs hm.directed_ge hf] exact tendsto_atTop_iInf fun n m hnm => measure_mono <| hm hnm @@ -1830,27 +1834,25 @@ theorem biSup_measure_Iic [Preorder α] {s : Set α} (hsc : s.Countable) exact iUnion₂_eq_univ_iff.2 hst · exact directedOn_iff_directed.2 (hdir.directed_val.mono_comp _ fun x y => Iic_subset_Iic.2) -theorem tendsto_measure_Ico_atTop [Preorder α] [IsDirected α (· ≤ ·)] [NoMaxOrder α] +theorem tendsto_measure_Ico_atTop [Preorder α] [NoMaxOrder α] [(atTop : Filter α).IsCountablyGenerated] (μ : Measure α) (a : α) : Tendsto (fun x => μ (Ico a x)) atTop (𝓝 (μ (Ici a))) := by rw [← iUnion_Ico_right] exact tendsto_measure_iUnion_atTop (antitone_const.Ico monotone_id) -theorem tendsto_measure_Ioc_atBot [Preorder α] [IsDirected α (· ≥ ·)] [NoMinOrder α] +theorem tendsto_measure_Ioc_atBot [Preorder α] [NoMinOrder α] [(atBot : Filter α).IsCountablyGenerated] (μ : Measure α) (a : α) : Tendsto (fun x => μ (Ioc x a)) atBot (𝓝 (μ (Iic a))) := by rw [← iUnion_Ioc_left] exact tendsto_measure_iUnion_atBot (monotone_id.Ioc antitone_const) -theorem tendsto_measure_Iic_atTop [Preorder α] [IsDirected α (· ≤ ·)] - [(atTop : Filter α).IsCountablyGenerated] (μ : Measure α) : - Tendsto (fun x => μ (Iic x)) atTop (𝓝 (μ univ)) := by +theorem tendsto_measure_Iic_atTop [Preorder α] [(atTop : Filter α).IsCountablyGenerated] + (μ : Measure α) : Tendsto (fun x => μ (Iic x)) atTop (𝓝 (μ univ)) := by rw [← iUnion_Iic] exact tendsto_measure_iUnion_atTop monotone_Iic -theorem tendsto_measure_Ici_atBot [Preorder α] [IsDirected α (· ≥ ·)] - [(atBot : Filter α).IsCountablyGenerated] (μ : Measure α) : - Tendsto (fun x => μ (Ici x)) atBot (𝓝 (μ univ)) := +theorem tendsto_measure_Ici_atBot [Preorder α] [(atBot : Filter α).IsCountablyGenerated] + (μ : Measure α) : Tendsto (fun x => μ (Ici x)) atBot (𝓝 (μ univ)) := tendsto_measure_Iic_atTop (α := αᵒᵈ) μ variable [PartialOrder α] {a b : α} diff --git a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean index ec687e378d87f..f40b6cba3186b 100644 --- a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean @@ -83,7 +83,7 @@ sets with finite measures. The term "measure-dense" is justified by the fact that the approximating condition translates to the usual notion of density in the metric space made by constant indicators of measurable sets equipped with the `Lᵖ` norm. -/ -structure Measure.MeasureDense (μ : Measure X) (𝒜 : Set (Set X)) : Prop := +structure Measure.MeasureDense (μ : Measure X) (𝒜 : Set (Set X)) : Prop where /-- Each set has to be measurable. -/ measurable : ∀ s ∈ 𝒜, MeasurableSet s /-- Any measurable set can be approximated by sets in the family. -/ @@ -319,7 +319,7 @@ section IsSeparable The term "separable" is justified by the fact that the definition translates to the usual notion of separability in the metric space made by constant indicators equipped with the `Lᵖ` norm. -/ -class IsSeparable (μ : Measure X) : Prop := +class IsSeparable (μ : Measure X) : Prop where exists_countable_measureDense : ∃ 𝒜, 𝒜.Countable ∧ μ.MeasureDense 𝒜 /-- By definition, a separable measure admits a countable and measure-dense family of sets. -/ diff --git a/Mathlib/NumberTheory/ADEInequality.lean b/Mathlib/NumberTheory/ADEInequality.lean index aabf630647bb3..f6cf4a0e5c35e 100644 --- a/Mathlib/NumberTheory/ADEInequality.lean +++ b/Mathlib/NumberTheory/ADEInequality.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Ring.Rat import Mathlib.Data.Multiset.Sort import Mathlib.Data.PNat.Basic @@ -157,15 +156,15 @@ theorem lt_three {p q r : ℕ+} (hpq : p ≤ q) (hqr : q ≤ r) (H : 1 < sumInv have h3q := H.trans hpq have h3r := h3q.trans hqr have hp : (p : ℚ)⁻¹ ≤ 3⁻¹ := by - rw [inv_le_inv _ h3] + rw [inv_le_inv₀ _ h3] · assumption_mod_cast · norm_num have hq : (q : ℚ)⁻¹ ≤ 3⁻¹ := by - rw [inv_le_inv _ h3] + rw [inv_le_inv₀ _ h3] · assumption_mod_cast · norm_num have hr : (r : ℚ)⁻¹ ≤ 3⁻¹ := by - rw [inv_le_inv _ h3] + rw [inv_le_inv₀ _ h3] · assumption_mod_cast · norm_num calc @@ -178,11 +177,11 @@ theorem lt_four {q r : ℕ+} (hqr : q ≤ r) (H : 1 < sumInv {2, q, r}) : q < 4 rw [sumInv_pqr] have h4r := H.trans hqr have hq : (q : ℚ)⁻¹ ≤ 4⁻¹ := by - rw [inv_le_inv _ h4] + rw [inv_le_inv₀ _ h4] · assumption_mod_cast · norm_num have hr : (r : ℚ)⁻¹ ≤ 4⁻¹ := by - rw [inv_le_inv _ h4] + rw [inv_le_inv₀ _ h4] · assumption_mod_cast · norm_num calc @@ -194,7 +193,7 @@ theorem lt_six {r : ℕ+} (H : 1 < sumInv {2, 3, r}) : r < 6 := by contrapose! H rw [sumInv_pqr] have hr : (r : ℚ)⁻¹ ≤ 6⁻¹ := by - rw [inv_le_inv _ h6] + rw [inv_le_inv₀ _ h6] · assumption_mod_cast · norm_num calc diff --git a/Mathlib/NumberTheory/Bertrand.lean b/Mathlib/NumberTheory/Bertrand.lean index 7f3a1d190e6f7..f35408c70ebe0 100644 --- a/Mathlib/NumberTheory/Bertrand.lean +++ b/Mathlib/NumberTheory/Bertrand.lean @@ -121,10 +121,6 @@ theorem bertrand_main_inequality {n : ℕ} (n_large : 512 ≤ n) : · norm_num1 · exact cast_div_le.trans (by norm_cast) -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- A lemma that tells us that, in the case where Bertrand's postulate does not hold, the prime factorization of the central binomial coefficient only has factors at most `2 * n / 3 + 1`. -/ @@ -138,7 +134,7 @@ theorem centralBinom_factorization_small (n : ℕ) (n_large : 2 < n) rw [Finset.mem_range, Nat.lt_succ_iff] at hx h2x rw [not_le, div_lt_iff_lt_mul' three_pos, mul_comm x] at h2x replace no_prime := not_exists.mp no_prime x - rw [← _root_.and_assoc, not_and', not_and_or, not_lt] at no_prime + rw [← and_assoc, not_and', not_and_or, not_lt] at no_prime cases' no_prime hx with h h · rw [factorization_eq_zero_of_non_prime n.centralBinom h, Nat.pow_zero] · rw [factorization_centralBinom_of_two_mul_self_lt_three_mul n_large h h2x, Nat.pow_zero] @@ -169,13 +165,13 @@ theorem centralBinom_le_of_no_bertrand_prime (n : ℕ) (n_large : 2 < n) · exact pow_factorization_choose_le (mul_pos two_pos n_pos) have : (Finset.Icc 1 (sqrt (2 * n))).card = sqrt (2 * n) := by rw [card_Icc, Nat.add_sub_cancel] rw [Finset.prod_const] - refine pow_le_pow_right n2_pos ((Finset.card_le_card fun x hx => ?_).trans this.le) + refine pow_right_mono₀ n2_pos ((Finset.card_le_card fun x hx => ?_).trans this.le) obtain ⟨h1, h2⟩ := Finset.mem_filter.1 hx exact Finset.mem_Icc.mpr ⟨(Finset.mem_filter.1 h1).2.one_lt.le, h2⟩ · refine le_trans ?_ (primorial_le_4_pow (2 * n / 3)) refine (Finset.prod_le_prod' fun p hp => (?_ : f p ≤ p)).trans ?_ · obtain ⟨h1, h2⟩ := Finset.mem_filter.1 hp - refine (pow_le_pow_right (Finset.mem_filter.1 h1).2.one_lt.le ?_).trans (pow_one p).le + refine (pow_right_mono₀ (Finset.mem_filter.1 h1).2.one_lt.le ?_).trans (pow_one p).le exact Nat.factorization_choose_le_one (sqrt_lt'.mp <| not_le.1 h2) refine Finset.prod_le_prod_of_subset_of_one_le' (Finset.filter_subset _ _) ?_ exact fun p hp _ => (Finset.mem_filter.1 hp).2.one_lt.le diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean index c1fd34862bfb8..fabb26655f44b 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean @@ -160,7 +160,7 @@ theorem cardPowDegree_anti_archimedean {x y z : Fq[X]} {a : ℤ} (hxy : cardPowD cardPowDegree_nonzero _ hyz'] have : (1 : ℤ) ≤ Fintype.card Fq := mod_cast (@Fintype.one_lt_card Fq _ _).le simp only [Int.cast_pow, Int.cast_natCast, le_max_iff] - refine Or.imp (pow_le_pow_right this) (pow_le_pow_right this) ?_ + refine Or.imp (pow_le_pow_right₀ this) (pow_le_pow_right₀ this) ?_ rw [natDegree_le_iff_degree_le, natDegree_le_iff_degree_le, ← le_max_iff, ← degree_eq_natDegree hxy', ← degree_eq_natDegree hyz'] convert degree_add_le (x - y) (y - z) using 2 diff --git a/Mathlib/NumberTheory/ClassNumber/FunctionField.lean b/Mathlib/NumberTheory/ClassNumber/FunctionField.lean index e45f913fd3b2e..04ead5e691263 100644 --- a/Mathlib/NumberTheory/ClassNumber/FunctionField.lean +++ b/Mathlib/NumberTheory/ClassNumber/FunctionField.lean @@ -24,7 +24,7 @@ namespace FunctionField open scoped Polynomial -variable (Fq F : Type) [Field Fq] [Fintype Fq] [Field F] +variable (Fq F : Type*) [Field Fq] [Fintype Fq] [Field F] variable [Algebra Fq[X] F] [Algebra (RatFunc Fq) F] variable [IsScalarTower Fq[X] (RatFunc Fq) F] variable [FunctionField Fq F] [Algebra.IsSeparable (RatFunc Fq) F] diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index a9c07d9408500..e4a82b6a04edb 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -113,9 +113,10 @@ variable {C} /-- The `PowerBasis` given by a primitive root `η`. -/ @[simps!] protected noncomputable def powerBasis : PowerBasis K L := - PowerBasis.map (Algebra.adjoin.powerBasis <| (integral {n} K L).isIntegral ζ) <| - (Subalgebra.equivOfEq _ _ (IsCyclotomicExtension.adjoin_primitive_root_eq_top hζ)).trans - Subalgebra.topEquiv + -- this is purely an optimization + letI pb := Algebra.adjoin.powerBasis <| (integral {n} K L).isIntegral ζ + pb.map <| (Subalgebra.equivOfEq _ _ (IsCyclotomicExtension.adjoin_primitive_root_eq_top hζ)).trans + Subalgebra.topEquiv theorem powerBasis_gen_mem_adjoin_zeta_sub_one : (hζ.powerBasis K).gen ∈ adjoin K ({ζ - 1} : Set L) := by @@ -392,9 +393,8 @@ theorem norm_pow_sub_one_of_prime_pow_ne_two {k s : ℕ} (hζ : IsPrimitiveRoot [hpri : Fact (p : ℕ).Prime] [IsCyclotomicExtension {p ^ (k + 1)} K L] (hirr : Irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) (hs : s ≤ k) (htwo : p ^ (k - s + 1) ≠ 2) : norm K (ζ ^ (p : ℕ) ^ s - 1) = (p : K) ^ (p : ℕ) ^ s := by --- Porting note: `by simp` was `by linarith` that now fails. have hirr₁ : Irreducible (cyclotomic ((p : ℕ) ^ (k - s + 1)) K) := - cyclotomic_irreducible_pow_of_irreducible_pow hpri.1 (by simp) hirr + cyclotomic_irreducible_pow_of_irreducible_pow hpri.1 (by omega) hirr rw [← PNat.pow_coe] at hirr₁ set η := ζ ^ (p : ℕ) ^ s - 1 let η₁ : K⟮η⟯ := IntermediateField.AdjoinSimple.gen K η @@ -403,22 +403,17 @@ theorem norm_pow_sub_one_of_prime_pow_ne_two {k s : ℕ} (hζ : IsPrimitiveRoot refine IsPrimitiveRoot.pow (p ^ (k + 1)).pos hζ ?_ rw [PNat.pow_coe, ← pow_add, add_comm s, Nat.sub_add_cancel (le_trans hs (Nat.le_succ k))] have : IsCyclotomicExtension {p ^ (k - s + 1)} K K⟮η⟯ := by - suffices IsCyclotomicExtension {p ^ (k - s + 1)} K K⟮η + 1⟯.toSubalgebra by - have H : K⟮η + 1⟯.toSubalgebra = K⟮η⟯.toSubalgebra := by - simp only [IntermediateField.adjoin_simple_toSubalgebra_of_integral - ((integral {p ^ (k + 1)} K L).isIntegral _)] - refine Subalgebra.ext fun x => ⟨fun hx => adjoin_le ?_ hx, fun hx => adjoin_le ?_ hx⟩ - · simp only [Set.singleton_subset_iff, SetLike.mem_coe] - exact Subalgebra.add_mem _ (subset_adjoin (mem_singleton η)) (Subalgebra.one_mem _) - · simp only [Set.singleton_subset_iff, SetLike.mem_coe] - nth_rw 2 [← add_sub_cancel_right η 1] - exact Subalgebra.sub_mem _ (subset_adjoin (mem_singleton _)) (Subalgebra.one_mem _) --- Porting note: the previous proof was `rw [H] at this; exact this` but it now fails. - exact IsCyclotomicExtension.equiv _ _ _ (Subalgebra.equivOfEq _ _ H) --- Porting note: the next `refine` was `rw [H]`, abusing defeq, and it now fails. + have HKη : K⟮η⟯ = K⟮η + 1⟯ := by + refine le_antisymm ?_ ?_ + all_goals rw [IntermediateField.adjoin_simple_le_iff] + · nth_rw 2 [← add_sub_cancel_right η 1] + exact sub_mem (IntermediateField.mem_adjoin_simple_self K (η + 1)) (one_mem _) + · exact add_mem (IntermediateField.mem_adjoin_simple_self K η) (one_mem _) + rw [HKη] have H := IntermediateField.adjoin_simple_toSubalgebra_of_integral - ((integral {p ^ (k + 1)} K L).isIntegral (η + 1)) - refine @IsCyclotomicExtension.equiv _ _ _ _ _ _ _ _ _ ?_ (Subalgebra.equivOfEq _ _ H).symm + ((integral {p ^ (k + 1)} K L).isIntegral (η + 1)) + refine IsCyclotomicExtension.equiv _ _ _ (h := ?_) (.refl : K⟮η + 1⟯.toSubalgebra ≃ₐ[K] _) + rw [H] have hη' : IsPrimitiveRoot (η + 1) ↑(p ^ (k + 1 - s)) := by simpa using hη -- Porting note: `using 1` was not needed. convert hη'.adjoin_isCyclotomicExtension K using 1 @@ -427,10 +422,10 @@ theorem norm_pow_sub_one_of_prime_pow_ne_two {k s : ℕ} (hζ : IsPrimitiveRoot apply coe_submonoidClass_iff.1 convert hη using 1 rw [Nat.sub_add_comm hs, pow_coe] --- Porting note: the following `haveI` were not needed because the locale `cyclotomic` set them +-- Porting note: the following `have` were not needed because the locale `cyclotomic` set them -- as instances. - haveI := IsCyclotomicExtension.finiteDimensional {p ^ (k + 1)} K L - haveI := IsCyclotomicExtension.isGalois (p ^ (k + 1)) K L + have := IsCyclotomicExtension.finiteDimensional {p ^ (k + 1)} K L + have := IsCyclotomicExtension.isGalois (p ^ (k + 1)) K L rw [norm_eq_norm_adjoin K] have H := hη.sub_one_norm_isPrimePow ?_ hirr₁ htwo swap; · rw [PNat.pow_coe]; exact hpri.1.isPrimePow.pow (Nat.succ_ne_zero _) diff --git a/Mathlib/NumberTheory/DiophantineApproximation.lean b/Mathlib/NumberTheory/DiophantineApproximation.lean index 946ac3af1f703..1d052c7a9f81e 100644 --- a/Mathlib/NumberTheory/DiophantineApproximation.lean +++ b/Mathlib/NumberTheory/DiophantineApproximation.lean @@ -406,8 +406,7 @@ private theorem aux₁ : 0 < fract ξ := by refine fract_pos.mpr fun hf => ?_ rw [hf] at h have H : (2 * v - 1 : ℝ) < 1 := by - refine - (mul_lt_iff_lt_one_right hv₀).mp ((inv_lt_inv hv₀ (mul_pos hv₁ hv₂)).mp (lt_of_le_of_lt ?_ h)) + refine (mul_lt_iff_lt_one_right hv₀).1 ((inv_lt_inv₀ hv₀ (mul_pos hv₁ hv₂)).1 (h.trans_le' ?_)) have h' : (⌊ξ⌋ : ℝ) - u / v = (⌊ξ⌋ * v - u) / v := by field_simp rw [h', abs_div, abs_of_pos hv₀, ← one_div, div_le_div_right hv₀] norm_cast @@ -490,7 +489,7 @@ private theorem aux₃ : _ < ((v : ℝ) * (2 * v - 1))⁻¹ * (v / u' / fract ξ) := (mul_lt_mul_right H₁).mpr h' _ = (u' * (2 * v - 1) * fract ξ)⁻¹ := help₂ hξ₀.ne' Hv.ne' Hv'.ne' Hu.ne' _ ≤ ((u' : ℝ) * (2 * u' - 1))⁻¹ := by - rwa [inv_le_inv (mul_pos (mul_pos Hu Hv') hξ₀) <| mul_pos Hu Hu', mul_assoc, + rwa [inv_le_inv₀ (mul_pos (mul_pos Hu Hv') hξ₀) <| mul_pos Hu Hu', mul_assoc, mul_le_mul_left Hu] -- The conditions `ass ξ u v` persist in the inductive step. @@ -506,7 +505,7 @@ private theorem invariant : ContfracLegendre.Ass (fract ξ)⁻¹ v (u - ⌊ξ⌋ have h' := (abs_sub_lt_iff.mp h.2.2).1 rw [Huv, ← sub_sub, sub_lt_iff_lt_add, self_sub_floor, Hv] at h' rwa [lt_sub_iff_add_lt', (by ring : (v : ℝ) + -(1 / 2) = (2 * v - 1) / 2), - lt_inv (div_pos hv₀' zero_lt_two) (aux₁ hv h), inv_div] + lt_inv_comm₀ (div_pos hv₀' zero_lt_two) (aux₁ hv h), inv_div] end @@ -538,8 +537,8 @@ theorem exists_rat_eq_convergent' {v : ℕ} (h : ContfracLegendre.Ass ξ u v) : · rw [Hξ, hξ₁, cast_sub, cast_one, ← sub_eq_add_neg, sub_lt_sub_iff_left] at h₁ exact False.elim (lt_irrefl _ <| h₁.trans one_half_lt_one) · have hξ₂ : ⌊(fract ξ)⁻¹⌋ = 1 := by - rw [floor_eq_iff, cast_one, le_inv zero_lt_one (fract_pos.mpr Hξ), inv_one, - one_add_one_eq_two, inv_lt (fract_pos.mpr Hξ) zero_lt_two] + rw [floor_eq_iff, cast_one, le_inv_comm₀ zero_lt_one (fract_pos.mpr Hξ), inv_one, + one_add_one_eq_two, inv_lt_comm₀ (fract_pos.mpr Hξ) zero_lt_two] refine ⟨(fract_lt_one ξ).le, ?_⟩ rw [fract, hξ₁, cast_sub, cast_one, lt_sub_iff_add_lt', sub_add] convert h₁ using 1 diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 43de39c1d54b0..0c98cefad6ecd 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -66,14 +66,10 @@ theorem filter_dvd_eq_properDivisors (h : n ≠ 0) : theorem properDivisors.not_self_mem : ¬n ∈ properDivisors n := by simp [properDivisors] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ @[simp] theorem mem_properDivisors {m : ℕ} : n ∈ properDivisors m ↔ n ∣ m ∧ n < m := by rcases eq_or_ne m 0 with (rfl | hm); · simp [properDivisors] - simp only [_root_.and_comm, ← filter_dvd_eq_properDivisors hm, mem_filter, mem_range] + simp only [and_comm, ← filter_dvd_eq_properDivisors hm, mem_filter, mem_range] theorem insert_self_properDivisors (h : n ≠ 0) : insert n (properDivisors n) = divisors n := by rw [divisors, properDivisors, Ico_succ_right_eq_insert_Ico (one_le_iff_ne_zero.2 h), @@ -100,15 +96,11 @@ theorem dvd_of_mem_divisors {m : ℕ} (h : n ∈ divisors m) : n ∣ m := by · apply dvd_zero · simp [mem_divisors.1 h] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ @[simp] theorem mem_divisorsAntidiagonal {x : ℕ × ℕ} : x ∈ divisorsAntidiagonal n ↔ x.fst * x.snd = n ∧ n ≠ 0 := by simp only [divisorsAntidiagonal, Finset.mem_Ico, Ne, Finset.mem_filter, Finset.mem_product] - rw [_root_.and_comm] + rw [and_comm] apply and_congr_right rintro rfl constructor <;> intro h @@ -429,14 +421,10 @@ theorem sum_properDivisors_eq_one_iff_prime : ∑ x ∈ n.properDivisors, x = 1 (one_mem_properDivisors_iff_one_lt.2 (succ_lt_succ (Nat.succ_pos _)))) ((sum_singleton _ _).trans h.symm) -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem mem_properDivisors_prime_pow {p : ℕ} (pp : p.Prime) (k : ℕ) {x : ℕ} : x ∈ properDivisors (p ^ k) ↔ ∃ (j : ℕ) (_ : j < k), x = p ^ j := by rw [mem_properDivisors, Nat.dvd_prime_pow pp, ← exists_and_right] - simp only [exists_prop, _root_.and_assoc] + simp only [exists_prop, and_assoc] apply exists_congr intro a constructor <;> intro h @@ -481,17 +469,13 @@ theorem prod_divisorsAntidiagonal' {M : Type*} [CommMonoid M] (f : ℕ → ℕ rw [← map_swap_divisorsAntidiagonal, Finset.prod_map] exact prod_divisorsAntidiagonal fun i j => f j i -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- The factors of `n` are the prime divisors -/ theorem primeFactors_eq_to_filter_divisors_prime (n : ℕ) : n.primeFactors = (divisors n).filter Prime := by rcases n.eq_zero_or_pos with (rfl | hn) · simp · ext q - simpa [hn, hn.ne', mem_primeFactorsList] using _root_.and_comm + 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 diff --git a/Mathlib/NumberTheory/FermatPsp.lean b/Mathlib/NumberTheory/FermatPsp.lean index f077508bbee07..113b74cc7f237 100644 --- a/Mathlib/NumberTheory/FermatPsp.lean +++ b/Mathlib/NumberTheory/FermatPsp.lean @@ -135,7 +135,7 @@ private theorem b_id_helper {a b : ℕ} (ha : 2 ≤ a) (hb : 2 < b) : 2 ≤ (a ^ calc 2 * a + 1 ≤ a ^ 2 * a := by nlinarith _ = a ^ 3 := by rw [Nat.pow_succ a 2] - _ ≤ a ^ b := pow_le_pow_right (Nat.le_of_succ_le ha) hb + _ ≤ a ^ b := pow_right_mono₀ (Nat.le_of_succ_le ha) hb private theorem AB_id_helper (b p : ℕ) (_ : 2 ≤ b) (hp : Odd p) : (b ^ p - 1) / (b - 1) * ((b ^ p + 1) / (b + 1)) = (b ^ (2 * p) - 1) / (b ^ 2 - 1) := by diff --git a/Mathlib/NumberTheory/FunctionField.lean b/Mathlib/NumberTheory/FunctionField.lean index 425189bfa3c7b..69740e0014839 100644 --- a/Mathlib/NumberTheory/FunctionField.lean +++ b/Mathlib/NumberTheory/FunctionField.lean @@ -44,7 +44,7 @@ noncomputable section open scoped nonZeroDivisors Polynomial Multiplicative -variable (Fq F : Type) [Field Fq] [Field F] +variable (Fq F : Type*) [Field Fq] [Field F] /-- `F` is a function field over the finite field `Fq` if it is a finite extension of the field of rational functions in one variable over `Fq`. diff --git a/Mathlib/NumberTheory/JacobiSum/Basic.lean b/Mathlib/NumberTheory/JacobiSum/Basic.lean index f9a173578aa98..9088bfbe869f8 100644 --- a/Mathlib/NumberTheory/JacobiSum/Basic.lean +++ b/Mathlib/NumberTheory/JacobiSum/Basic.lean @@ -270,7 +270,7 @@ lemma MulChar.exists_apply_sub_one_mul_apply_sub_one {n : ℕ} (hn : n ≠ 0) { with values in an integral domain `R` and `μ` is a primitive `n`th root of unity in `R`, then `J(χ,ψ) = -1 + z*(μ - 1)^2` for some `z ∈ ℤ[μ] ⊆ R`. (We assume that `#F ≡ 1 mod n`.) Note that we do not state this as a divisbility in `R`, as this would give a weaker statement. -/ -lemma exists_jacobiSum_eq_neg_one_add [DecidableEq F] {n : ℕ} (hn : 2 < n) {χ ψ : MulChar F R} +lemma exists_jacobiSum_eq_neg_one_add {n : ℕ} (hn : 2 < n) {χ ψ : MulChar F R} {μ : R} (hχ : χ ^ n = 1) (hψ : ψ ^ n = 1) (hn' : n ∣ Fintype.card F - 1) (hμ : IsPrimitiveRoot μ n) : ∃ z ∈ Algebra.adjoin ℤ {μ}, jacobiSum χ ψ = -1 + z * (μ - 1) ^ 2 := by @@ -286,7 +286,8 @@ lemma exists_jacobiSum_eq_neg_one_add [DecidableEq F] {n : ℕ} (hn : 2 < n) {χ rw [hχ₀, jacobiSum_one_nontrivial hψ₀, zero_mul, add_zero] · refine ⟨0, Subalgebra.zero_mem _, ?_⟩ rw [jacobiSum_comm, hψ₀, jacobiSum_one_nontrivial hχ₀, zero_mul, add_zero] - · rw [jacobiSum_eq_aux, MulChar.sum_eq_zero_of_ne_one hχ₀, MulChar.sum_eq_zero_of_ne_one hψ₀, hq] + · classical + rw [jacobiSum_eq_aux, MulChar.sum_eq_zero_of_ne_one hχ₀, MulChar.sum_eq_zero_of_ne_one hψ₀, hq] have H := MulChar.exists_apply_sub_one_mul_apply_sub_one (by omega) hχ hψ hμ have Hcs x := (H x).choose_spec refine ⟨-q * z₁ + ∑ x ∈ (univ \ {0, 1} : Finset F), (H x).choose, ?_, ?_⟩ diff --git a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean index 6c8bfa2572edc..64d09847b3f9e 100644 --- a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean +++ b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean @@ -167,7 +167,7 @@ lemma hf_zero' (P : WeakFEPair E) : filter_upwards [eventually_le_nhds zero_lt_one] with x hx' (hx : 0 < x) apply le_mul_of_one_le_right (norm_nonneg _) rw [norm_of_nonneg (rpow_pos_of_pos hx _).le, rpow_neg hx.le] - exact one_le_inv (rpow_pos_of_pos hx _) (rpow_le_one hx.le hx' P.hk.le) + exact (one_le_inv₀ (rpow_pos_of_pos hx _)).2 (rpow_le_one hx.le hx' P.hk.le) end WeakFEPair diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean index b2cb30daa0d3e..392736e26084d 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean @@ -180,10 +180,6 @@ private theorem sum_Ico_eq_card_lt {p q : ℕ} : (by simp (config := { contextual := true }) only [mem_filter, mem_sigma, and_self_iff, forall_true_iff, mem_product]) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- Each of the sums in this lemma is the cardinality of the set of integer points in each of the two triangles formed by the diagonal of the rectangle `(0, p/2) × (0, q/2)`. Adding them gives the number of points in the rectangle. -/ @@ -197,7 +193,7 @@ theorem sum_mul_div_add_sum_mul_div_eq_mul (p q : ℕ) [hp : Fact p.Prime] (hq0 card_equiv (Equiv.prodComm _ _) (fun ⟨_, _⟩ => by simp (config := { contextual := true }) only [mem_filter, and_self_iff, Prod.swap_prod_mk, - forall_true_iff, mem_product, Equiv.prodComm_apply, _root_.and_assoc, and_left_comm]) + forall_true_iff, mem_product, Equiv.prodComm_apply, and_assoc, and_left_comm]) have hdisj : Disjoint ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => x.2 * p ≤ x.1 * q) ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => x.1 * q ≤ x.2 * p) := by diff --git a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean b/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean index 3fb42ea6a0f27..49a8325713112 100644 --- a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean +++ b/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean @@ -145,7 +145,7 @@ theorem aux_calc (n : ℕ) {m : ℝ} (hm : 2 ≤ m) : any_goals exact pow_pos (zero_lt_two.trans_le hm) _ -- `2 ≤ m ^ n!` is a consequence of monotonicity of exponentiation at `2 ≤ m`. exact _root_.trans (_root_.trans hm (pow_one _).symm.le) - (pow_right_mono (one_le_two.trans hm) n.factorial_pos) + (pow_right_mono₀ (one_le_two.trans hm) n.factorial_pos) _ = 1 / (m ^ n !) ^ n := congr_arg (1 / ·) (pow_mul m n ! n) /-- An upper estimate on the remainder. This estimate works with `m ∈ ℝ` satisfying `2 ≤ m` and is diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 7ee6f2f7441bb..d27a687df5826 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -351,7 +351,7 @@ theorem g_eq_of_c_eq_one (hc : (↑ₘg) 1 0 = 1) : g = T ^ (↑ₘg) 0 0 * S * /-- If `1 < |z|`, then `|S • z| < 1`. -/ theorem normSq_S_smul_lt_one (h : 1 < normSq z) : normSq ↑(S • z) < 1 := by - simpa [coe_S, num, denom] using (inv_lt_inv z.normSq_pos zero_lt_one).mpr h + simpa [coe_S, num, denom] using (inv_lt_inv₀ z.normSq_pos zero_lt_one).mpr h /-- If `|z| < 1`, then applying `S` strictly decreases `im`. -/ theorem im_lt_im_S_smul (h : normSq z < 1) : z.im < (S • z).im := by diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 2ff96d352b04b..7268d68dd4275 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -78,7 +78,7 @@ lemma r_lower_bound_on_verticalStrip {A B : ℝ} (h : 0 < B) (hz : z ∈ vertica apply min_le_min hz.2 rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] simp only [r1_eq, div_pow, one_div] - rw [inv_le_inv (by positivity) (by positivity), add_le_add_iff_right] + rw [inv_le_inv₀ (by positivity) (by positivity), add_le_add_iff_right] apply div_le_div (sq_nonneg _) _ (by positivity) (pow_le_pow_left h.le hz.2 2) simpa only [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index b56a1d4722194..52a08d5551c4a 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -337,6 +337,8 @@ theorem finite_of_discr_bdd_of_isReal : (Set.finite_Icc (-C : ℤ) C)) (fun ⟨K, hK₀⟩ ⟨hK₁, hK₂⟩ ↦ ?_) -- We now need to prove that each field is generated by an element of the union of the rootset simp_rw [Set.mem_iUnion] + -- this is purely an optimization + have : CharZero K := SubsemiringClass.instCharZero K haveI : NumberField K := @NumberField.mk _ _ inferInstance hK₀ obtain ⟨w₀, hw₀⟩ := hK₁ suffices minkowskiBound K ↑1 < (convexBodyLTFactor K) * B by @@ -360,12 +362,15 @@ theorem finite_of_discr_bdd_of_isReal : · refine mem_rootSet.mpr ⟨minpoly.ne_zero hx, ?_⟩ exact (aeval_algebraMap_eq_zero_iff _ _ _).mpr (minpoly.aeval ℤ (x : K)) · rw [← (IntermediateField.lift_injective _).eq_iff, eq_comm] at hx₁ - convert hx₁ <;> simp + convert hx₁ + · simp only [IntermediateField.lift_top] + · simp only [IntermediateField.lift_adjoin, Set.image_singleton] have := one_le_convexBodyLTFactor K convert lt_of_le_of_lt (mul_right_mono (coe_le_coe.mpr this)) (ENNReal.mul_lt_mul_left' (by positivity) coe_ne_top (minkowskiBound_lt_boundOfDiscBdd hK₂)) simp_rw [ENNReal.coe_one, one_mul] + theorem finite_of_discr_bdd_of_isComplex : {K : { F : IntermediateField ℚ A // FiniteDimensional ℚ F} | haveI : NumberField K := @NumberField.mk _ _ inferInstance K.prop @@ -380,6 +385,8 @@ theorem finite_of_discr_bdd_of_isComplex : (Set.finite_Icc (-C : ℤ) C)) (fun ⟨K, hK₀⟩ ⟨hK₁, hK₂⟩ ↦ ?_) -- We now need to prove that each field is generated by an element of the union of the rootset simp_rw [Set.mem_iUnion] + -- this is purely an optimization + have : CharZero K := SubsemiringClass.instCharZero K haveI : NumberField K := @NumberField.mk _ _ inferInstance hK₀ obtain ⟨w₀, hw₀⟩ := hK₁ suffices minkowskiBound K ↑1 < (convexBodyLT'Factor K) * boundOfDiscBdd N by @@ -404,7 +411,9 @@ theorem finite_of_discr_bdd_of_isComplex : · refine mem_rootSet.mpr ⟨minpoly.ne_zero hx, ?_⟩ exact (aeval_algebraMap_eq_zero_iff _ _ _).mpr (minpoly.aeval ℤ (x : K)) · rw [← (IntermediateField.lift_injective _).eq_iff, eq_comm] at hx₁ - convert hx₁ <;> simp + convert hx₁ + · simp only [IntermediateField.lift_top] + · simp only [IntermediateField.lift_adjoin, Set.image_singleton] have := one_le_convexBodyLT'Factor K convert lt_of_le_of_lt (mul_right_mono (coe_le_coe.mpr this)) (ENNReal.mul_lt_mul_left' (by positivity) coe_ne_top (minkowskiBound_lt_boundOfDiscBdd hK₂)) @@ -419,6 +428,8 @@ theorem _root_.NumberField.finite_of_discr_bdd : refine Set.Finite.subset (Set.Finite.union (finite_of_discr_bdd_of_isReal A N) (finite_of_discr_bdd_of_isComplex A N)) ?_ rintro ⟨K, hK₀⟩ hK₁ + -- this is purely an optimization + have : CharZero K := SubsemiringClass.instCharZero K haveI : NumberField K := @NumberField.mk _ _ inferInstance hK₀ obtain ⟨w₀⟩ := (inferInstance : Nonempty (InfinitePlace K)) by_cases hw₀ : IsReal w₀ diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index c30dda51f0a43..9e1f2fc310559 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -259,7 +259,7 @@ open NumberField instance {K : Type*} [Field K] : FunLike (InfinitePlace K) K ℝ where coe w x := w.1 x - coe_injective' := fun _ _ h => Subtype.eq (AbsoluteValue.ext fun x => congr_fun h x) + coe_injective' _ _ h := Subtype.eq (AbsoluteValue.ext fun x => congr_fun h x) instance : MonoidWithZeroHomClass (InfinitePlace K) K ℝ where map_mul w _ _ := w.1.map_mul _ _ diff --git a/Mathlib/NumberTheory/NumberField/Units/Basic.lean b/Mathlib/NumberTheory/NumberField/Units/Basic.lean index eaabbc19a44f8..7907293a0c950 100644 --- a/Mathlib/NumberTheory/NumberField/Units/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Units/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ import Mathlib.NumberTheory.NumberField.Embeddings +import Mathlib.RingTheory.LocalRing.RingHom.Basic /-! # Units of a number field diff --git a/Mathlib/NumberTheory/Padics/Hensel.lean b/Mathlib/NumberTheory/Padics/Hensel.lean index 0de565647f9ee..05f97027ce268 100644 --- a/Mathlib/NumberTheory/Padics/Hensel.lean +++ b/Mathlib/NumberTheory/Padics/Hensel.lean @@ -156,7 +156,7 @@ private theorem calc_norm_le_one {n : ℕ} {z : ℤ_[p]} (hz : ih n z) : gcongr apply hz.2 _ = ‖F.derivative.eval a‖ * T ^ 2 ^ n := div_sq_cancel _ _ - _ ≤ 1 := mul_le_one (PadicInt.norm_le_one _) (T_pow_nonneg _) (le_of_lt (T_pow' hnorm _)) + _ ≤ 1 := mul_le_one₀ (PadicInt.norm_le_one _) (T_pow_nonneg _) (le_of_lt (T_pow' hnorm _)) private theorem calc_deriv_dist {z z' z1 : ℤ_[p]} (hz' : z' = z - z1) @@ -183,7 +183,7 @@ private def calc_eval_z' {z z' z1 : ℤ_[p]} (hz' : z' = z - z1) {n} (hz : ih n obtain ⟨q, hq⟩ := F.binomExpansion z (-z1) have : ‖(↑(F.derivative.eval z) * (↑(F.eval z) / ↑(F.derivative.eval z)) : ℚ_[p])‖ ≤ 1 := by rw [padicNormE.mul] - exact mul_le_one (PadicInt.norm_le_one _) (norm_nonneg _) h1 + exact mul_le_one₀ (PadicInt.norm_le_one _) (norm_nonneg _) h1 have : F.derivative.eval z * -z1 = -F.eval z := by calc F.derivative.eval z * -z1 = @@ -275,7 +275,7 @@ private theorem newton_seq_dist_aux (n : ℕ) : | 0 => by simp [T_pow_nonneg, mul_nonneg] | k + 1 => have : 2 ^ n ≤ 2 ^ (n + k) := by - apply pow_le_pow_right + apply pow_right_mono₀ · norm_num · apply Nat.le_add_right calc @@ -356,7 +356,7 @@ private theorem T_pos : T > 0 := by private theorem newton_seq_succ_dist_weak (n : ℕ) : ‖newton_seq (n + 2) - newton_seq (n + 1)‖ < ‖F.eval a‖ / ‖F.derivative.eval a‖ := have : 2 ≤ 2 ^ (n + 1) := by - have := pow_le_pow_right (by norm_num : 1 ≤ 2) (Nat.le_add_left _ _ : 1 ≤ n + 1) + have := pow_right_mono₀ (by norm_num : 1 ≤ 2) (Nat.le_add_left _ _ : 1 ≤ n + 1) simpa using this calc ‖newton_seq (n + 2) - newton_seq (n + 1)‖ ≤ ‖F.derivative.eval a‖ * T ^ 2 ^ (n + 1) := diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 02820e33e2483..a7add4280a7e7 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -78,7 +78,7 @@ def subring : Subring ℚ_[p] where zero_mem' := by norm_num one_mem' := by norm_num add_mem' hx hy := (padicNormE.nonarchimedean _ _).trans <| max_le_iff.2 ⟨hx, hy⟩ - mul_mem' hx hy := (padicNormE.mul _ _).trans_le <| mul_le_one hx (norm_nonneg _) hy + mul_mem' hx hy := (padicNormE.mul _ _).trans_le <| mul_le_one₀ hx (norm_nonneg _) hy neg_mem' hx := (norm_neg _).trans_le hx @[simp] @@ -304,7 +304,7 @@ variable (p : ℕ) [hp : Fact p.Prime] theorem exists_pow_neg_lt {ε : ℝ} (hε : 0 < ε) : ∃ k : ℕ, (p : ℝ) ^ (-(k : ℤ)) < ε := by obtain ⟨k, hk⟩ := exists_nat_gt ε⁻¹ use k - rw [← inv_lt_inv hε (_root_.zpow_pos_of_pos _ _)] + rw [← inv_lt_inv₀ hε (_root_.zpow_pos_of_pos _ _)] · rw [zpow_neg, inv_inv, zpow_natCast] apply lt_of_lt_of_le hk norm_cast @@ -468,7 +468,7 @@ theorem norm_le_pow_iff_le_valuation (x : ℤ_[p]) (hx : x ≠ 0) (n : ℕ) : intro m refine pow_pos ?_ m exact mod_cast hp.1.pos - rw [inv_le_inv (aux _) (aux _)] + rw [inv_le_inv₀ (aux _) (aux _)] have : p ^ n ≤ p ^ k ↔ n ≤ k := (pow_right_strictMono hp.1.one_lt).le_iff_le rw [← this] norm_cast @@ -528,7 +528,7 @@ instance : LocalRing ℤ_[p] := LocalRing.of_nonunits_add <| by simp only [mem_nonunits]; exact fun x y => norm_lt_one_add theorem p_nonnunit : (p : ℤ_[p]) ∈ nonunits ℤ_[p] := by - have : (p : ℝ)⁻¹ < 1 := inv_lt_one <| mod_cast hp.1.one_lt + have : (p : ℝ)⁻¹ < 1 := inv_lt_one_of_one_lt₀ <| mod_cast hp.1.one_lt rwa [← norm_p, ← mem_nonunits] at this theorem maximalIdeal_eq_span_p : maximalIdeal ℤ_[p] = Ideal.span {(p : ℤ_[p])} := by diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index 9a46b58826307..c61676a7816b8 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -92,7 +92,7 @@ theorem padicNorm_of_prime_of_ne {q : ℕ} [p_prime : Fact p.Prime] [q_prime : F See also `padicNorm.padicNorm_p_lt_one_of_prime` for a version assuming `p` is prime. -/ theorem padicNorm_p_lt_one (hp : 1 < p) : padicNorm p p < 1 := by - rw [padicNorm_p hp, inv_lt_one_iff] + rw [padicNorm_p hp, inv_lt_one_iff₀] exact mod_cast Or.inr hp /-- The `p`-adic norm of `p` is less than `1` if `p` is prime. @@ -246,7 +246,7 @@ theorem int_eq_one_iff (m : ℤ) : padicNorm p m = 1 ↔ ¬(p : ℤ) ∣ m := by simp only [dvd_iff_norm_le, Int.cast_natCast, Nat.cast_one, zpow_neg, zpow_one, not_le] constructor · intro h - rw [h, inv_lt_one_iff_of_pos] <;> norm_cast + rw [h, inv_lt_one₀] <;> norm_cast · exact Nat.Prime.one_lt Fact.out · exact Nat.Prime.pos Fact.out · simp only [padicNorm] diff --git a/Mathlib/NumberTheory/Padics/PadicNumbers.lean b/Mathlib/NumberTheory/Padics/PadicNumbers.lean index abccff73fbb92..ff2f78e73d60b 100644 --- a/Mathlib/NumberTheory/Padics/PadicNumbers.lean +++ b/Mathlib/NumberTheory/Padics/PadicNumbers.lean @@ -779,8 +779,7 @@ theorem norm_p : ‖(p : ℚ_[p])‖ = (p : ℝ)⁻¹ := by theorem norm_p_lt_one : ‖(p : ℚ_[p])‖ < 1 := by rw [norm_p] - apply inv_lt_one - exact mod_cast hp.1.one_lt + exact inv_lt_one_of_one_lt₀ <| mod_cast hp.1.one_lt -- Porting note: Linter thinks this is a duplicate simp lemma, so `priority` is assigned @[simp (high)] @@ -833,7 +832,7 @@ theorem norm_rat_le_one : ∀ {q : ℚ} (_ : ¬p ∣ q.den), ‖(q : ℚ_[p])‖ -- Porting note: `Nat.cast_zero` instead of another `norm_cast` call rw [padicNorm.eq_zpow_of_nonzero hnz', padicValRat, neg_sub, padicValNat.eq_zero_of_not_dvd hq, Nat.cast_zero, zero_sub, zpow_neg, zpow_natCast] - apply inv_le_one + apply inv_le_one_of_one_le₀ norm_cast apply one_le_pow exact hp.1.pos @@ -860,8 +859,7 @@ theorem norm_int_lt_one_iff_dvd (k : ℤ) : ‖(k : ℚ_[p])‖ < 1 ↔ ↑p ∣ mul_le_mul le_rfl (by simpa using norm_int_le_one _) (norm_nonneg _) (norm_nonneg _) _ < 1 := by rw [mul_one, padicNormE.norm_p] - apply inv_lt_one - exact mod_cast hp.1.one_lt + exact inv_lt_one_of_one_lt₀ <| mod_cast hp.1.one_lt theorem norm_int_le_pow_iff_dvd (k : ℤ) (n : ℕ) : ‖(k : ℚ_[p])‖ ≤ (p : ℝ) ^ (-n : ℤ) ↔ (p ^ n : ℤ) ∣ k := by diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index 31ba8f2ff384a..8f65a9cbc788e 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -256,7 +256,7 @@ theorem toZMod_spec : x - (ZMod.cast (toZMod x) : ℤ_[p]) ∈ maximalIdeal ℤ_ dsimp [toZMod, toZModHom] rcases Nat.exists_eq_add_of_lt hp_prime.1.pos with ⟨p', rfl⟩ change ↑((_ : ZMod (0 + p' + 1)).val) = (_ : ℤ_[0 + p' + 1]) - simp only [ZMod.val_natCast, add_zero, add_def, Nat.cast_inj, zero_add] + rw [Nat.cast_inj] apply mod_eq_of_lt simpa only [zero_add] using zmodRepr_lt_p x diff --git a/Mathlib/NumberTheory/RamificationInertia.lean b/Mathlib/NumberTheory/RamificationInertia.lean index 738d4accca7ef..670a2c3ba24a8 100644 --- a/Mathlib/NumberTheory/RamificationInertia.lean +++ b/Mathlib/NumberTheory/RamificationInertia.lean @@ -797,11 +797,11 @@ section tower variable {R S T : Type*} [CommRing R] [CommRing S] [CommRing T] -theorem ramificationIdx_tower [IsDedekindDomain S] [DecidableEq (Ideal S)] - [IsDedekindDomain T] [DecidableEq (Ideal T)]{f : R →+* S} {g : S →+* T} +theorem ramificationIdx_tower [IsDedekindDomain S] [IsDedekindDomain T] {f : R →+* S} {g : S →+* T} {p : Ideal R} {P : Ideal S} {Q : Ideal T} [hpm : P.IsPrime] [hqm : Q.IsPrime] (hg0 : map g P ≠ ⊥) (hfg : map (g.comp f) p ≠ ⊥) (hg : map g P ≤ Q) : ramificationIdx (g.comp f) p Q = ramificationIdx f p P * ramificationIdx g P Q := by + classical have hf0 : map f p ≠ ⊥ := ne_bot_of_map_ne_bot (Eq.mp (congrArg (fun I ↦ I ≠ ⊥) (map_map f g).symm) hfg) have hp0 : P ≠ ⊥ := ne_bot_of_map_ne_bot hg0 @@ -836,12 +836,13 @@ variable [Algebra R S] [Algebra S T] [Algebra R T] [IsScalarTower R S T] /-- Let `T / S / R` be a tower of algebras, `p, P, Q` be ideals in `R, S, T` respectively, and `P` and `Q` are prime. If `P = Q ∩ S`, then `e (Q | p) = e (P | p) * e (Q | P)`. -/ -theorem ramificationIdx_algebra_tower [IsDedekindDomain S] [DecidableEq (Ideal S)] - [IsDedekindDomain T] [DecidableEq (Ideal T)] {p : Ideal R} {P : Ideal S} {Q : Ideal T} - [hpm : P.IsPrime] [hqm : Q.IsPrime] (hg0 : map (algebraMap S T) P ≠ ⊥) +theorem ramificationIdx_algebra_tower [IsDedekindDomain S] [IsDedekindDomain T] + {p : Ideal R} {P : Ideal S} {Q : Ideal T} [hpm : P.IsPrime] [hqm : Q.IsPrime] + (hg0 : map (algebraMap S T) P ≠ ⊥) (hfg : map (algebraMap R T) p ≠ ⊥) (hg : map (algebraMap S T) P ≤ Q) : ramificationIdx (algebraMap R T) p Q = ramificationIdx (algebraMap R S) p P * ramificationIdx (algebraMap S T) P Q := by + classical rw [IsScalarTower.algebraMap_eq R S T] at hfg ⊢ exact ramificationIdx_tower hg0 hfg hg diff --git a/Mathlib/NumberTheory/SumFourSquares.lean b/Mathlib/NumberTheory/SumFourSquares.lean index a08b8c479c8ee..5d7a765adee86 100644 --- a/Mathlib/NumberTheory/SumFourSquares.lean +++ b/Mathlib/NumberTheory/SumFourSquares.lean @@ -121,10 +121,6 @@ private theorem sum_four_squares_of_two_mul_sum_four_squares {m a b c d : ℤ} have : (∑ x, f (σ x) ^ 2) = ∑ x, f x ^ 2 := Equiv.sum_comp σ (f · ^ 2) simpa only [← hx, ← hy, Fin.sum_univ_four, add_assoc] using this -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- Lagrange's **four squares theorem** for a prime number. Use `Nat.sum_four_squares` instead. -/ protected theorem Prime.sum_four_squares {p : ℕ} (hp : p.Prime) : ∃ a b c d : ℕ, a ^ 2 + b ^ 2 + c ^ 2 + d ^ 2 = p := by @@ -174,10 +170,9 @@ protected theorem Prime.sum_four_squares {p : ℕ} (hp : p.Prime) : -- The quotient `r` is not zero, because otherwise `f a = f b = f c = f d = 0`, hence -- `m` divides each `a`, `b`, `c`, `d`, thus `m ∣ p` which is impossible. rcases (zero_le r).eq_or_gt with rfl | hr₀ - · replace hr : f a = 0 ∧ f b = 0 ∧ f c = 0 ∧ f d = 0 := by simpa [_root_.and_assoc] using hr + · replace hr : f a = 0 ∧ f b = 0 ∧ f c = 0 ∧ f d = 0 := by simpa [and_assoc] using hr obtain ⟨⟨a, rfl⟩, ⟨b, rfl⟩, ⟨c, rfl⟩, ⟨d, rfl⟩⟩ : m ∣ a ∧ m ∣ b ∧ m ∣ c ∧ m ∣ d := by - simp only [← ZMod.natCast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero, - _root_.and_self] + simp only [← ZMod.natCast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero, and_self] have : m * m ∣ m * p := habcd ▸ ⟨a ^ 2 + b ^ 2 + c ^ 2 + d ^ 2, by ring⟩ rw [mul_dvd_mul_iff_left hm₀.ne'] at this exact (hp.eq_one_or_self_of_dvd _ this).elim hm₁.ne' hmp.ne diff --git a/Mathlib/Order/Atoms/Finite.lean b/Mathlib/Order/Atoms/Finite.lean index c17782a493d32..93cca39127a7d 100644 --- a/Mathlib/Order/Atoms/Finite.lean +++ b/Mathlib/Order/Atoms/Finite.lean @@ -22,23 +22,26 @@ variable {α β : Type*} namespace IsSimpleOrder +variable [LE α] [BoundedOrder α] [IsSimpleOrder α] + section DecidableEq /- It is important that `IsSimpleOrder` is the last type-class argument of this instance, so that type-class inference fails quickly if it doesn't apply. -/ -instance (priority := 200) {α} [DecidableEq α] [LE α] [BoundedOrder α] [IsSimpleOrder α] : - Fintype α := +instance (priority := 200) [DecidableEq α] : Fintype α := Fintype.ofEquiv Bool equivBool.symm end DecidableEq +instance (priority := 200) : Finite α := by classical infer_instance + end IsSimpleOrder namespace Fintype namespace IsSimpleOrder -variable [PartialOrder α] [BoundedOrder α] [IsSimpleOrder α] [DecidableEq α] +variable [LE α] [BoundedOrder α] [IsSimpleOrder α] [DecidableEq α] theorem univ : (Finset.univ : Finset α) = {⊤, ⊥} := by change Finset.map _ (Finset.univ : Finset Bool) = _ diff --git a/Mathlib/Order/Booleanisation.lean b/Mathlib/Order/Booleanisation.lean index dd3aa408d0c38..aa7f7370581b4 100644 --- a/Mathlib/Order/Booleanisation.lean +++ b/Mathlib/Order/Booleanisation.lean @@ -58,7 +58,7 @@ instance instCompl : HasCompl (Booleanisation α) where @[simp] lemma compl_lift (a : α) : (lift a)ᶜ = comp a := rfl @[simp] lemma compl_comp (a : α) : (comp a)ᶜ = lift a := rfl -variable [GeneralizedBooleanAlgebra α] {x y : Booleanisation α} {a b : α} +variable [GeneralizedBooleanAlgebra α] {a b : α} /-- The order on `Booleanisation α` is as follows: For `a b : α`, * `a ≤ b` iff `a ≤ b` in `α` diff --git a/Mathlib/Order/BoundedOrder.lean b/Mathlib/Order/BoundedOrder.lean index 7850b239ce06a..a28c66b0685a2 100644 --- a/Mathlib/Order/BoundedOrder.lean +++ b/Mathlib/Order/BoundedOrder.lean @@ -33,7 +33,7 @@ open Function OrderDual universe u v -variable {α : Type u} {β : Type v} {γ δ : Type*} +variable {α : Type u} {β : Type v} /-! ### Top, bottom element -/ @@ -353,7 +353,7 @@ theorem OrderBot.ext_bot {α} {hA : PartialOrder α} (A : OrderBot α) {hB : Par section SemilatticeSupTop -variable [SemilatticeSup α] [OrderTop α] {a : α} +variable [SemilatticeSup α] [OrderTop α] -- Porting note: Not simp because simp can prove it theorem top_sup_eq (a : α) : ⊤ ⊔ a = ⊤ := @@ -400,7 +400,7 @@ end SemilatticeInfTop section SemilatticeInfBot -variable [SemilatticeInf α] [OrderBot α] {a : α} +variable [SemilatticeInf α] [OrderBot α] -- Porting note: Not simp because simp can prove it lemma bot_inf_eq (a : α) : ⊥ ⊓ a = ⊥ := inf_of_le_left bot_le diff --git a/Mathlib/Order/Category/NonemptyFinLinOrd.lean b/Mathlib/Order/Category/NonemptyFinLinOrd.lean index 02b1bfe0adc42..3c6d13dec2f73 100644 --- a/Mathlib/Order/Category/NonemptyFinLinOrd.lean +++ b/Mathlib/Order/Category/NonemptyFinLinOrd.lean @@ -235,3 +235,6 @@ def nonemptyFinLinOrdDualCompForgetToFinPartOrd : forget₂ NonemptyFinLinOrd FinPartOrd ⋙ FinPartOrd.dual where hom := { app := fun X => OrderHom.id } inv := { app := fun X => OrderHom.id } + +/-- The generating arrow `i ⟶ i+1` in the category `Fin n`.-/ +def Fin.hom_succ {n} (i : Fin n) : i.castSucc ⟶ i.succ := homOfLE (Fin.castSucc_le_succ i) diff --git a/Mathlib/Order/Chain.lean b/Mathlib/Order/Chain.lean index 2f086f7f4f276..f8f5ce2631f67 100644 --- a/Mathlib/Order/Chain.lean +++ b/Mathlib/Order/Chain.lean @@ -52,7 +52,7 @@ def SuperChain (s t : Set α) : Prop := def IsMaxChain (s : Set α) : Prop := IsChain r s ∧ ∀ ⦃t⦄, IsChain r t → s ⊆ t → s = t -variable {r} {c c₁ c₂ c₃ s t : Set α} {a b x y : α} +variable {r} {c c₁ c₂ s t : Set α} {a x y : α} theorem isChain_empty : IsChain r ∅ := Set.pairwise_empty _ diff --git a/Mathlib/Order/Closure.lean b/Mathlib/Order/Closure.lean index 4005f3290f47d..53f361b143e1e 100644 --- a/Mathlib/Order/Closure.lean +++ b/Mathlib/Order/Closure.lean @@ -248,7 +248,7 @@ end SemilatticeSup section CompleteLattice -variable [CompleteLattice α] (c : ClosureOperator α) {p : α → Prop} +variable [CompleteLattice α] (c : ClosureOperator α) /-- Define a closure operator from a predicate that's preserved under infima. -/ @[simps!] diff --git a/Mathlib/Order/CompactlyGenerated/Basic.lean b/Mathlib/Order/CompactlyGenerated/Basic.lean index 60620cc7187cf..4fb6ec86d8fa8 100644 --- a/Mathlib/Order/CompactlyGenerated/Basic.lean +++ b/Mathlib/Order/CompactlyGenerated/Basic.lean @@ -189,10 +189,10 @@ theorem isCompactElement_finsetSup {α β : Type*} [CompleteLattice α] {f : β specialize h d hemp hdir (le_trans (Finset.le_sup hps) hsup) simpa only [exists_prop] -theorem WellFounded.isSupFiniteCompact (h : WellFounded ((· > ·) : α → α → Prop)) : +theorem WellFoundedGT.isSupFiniteCompact [WellFoundedGT α] : IsSupFiniteCompact α := fun s => by let S := { x | ∃ t : Finset α, ↑t ⊆ s ∧ t.sup id = x } - obtain ⟨m, ⟨t, ⟨ht₁, rfl⟩⟩, hm⟩ := h.has_min S ⟨⊥, ∅, by simp⟩ + obtain ⟨m, ⟨t, ⟨ht₁, rfl⟩⟩, hm⟩ := wellFounded_gt.has_min S ⟨⊥, ∅, by simp⟩ refine ⟨t, ht₁, (sSup_le _ _ fun y hy => ?_).antisymm ?_⟩ · classical rw [eq_of_le_of_not_lt (Finset.sup_mono (t.subset_insert y)) @@ -212,25 +212,26 @@ theorem IsSupFiniteCompact.isSupClosedCompact (h : IsSupFiniteCompact α) : · rw [ht₂] exact hsc.finsetSup_mem h ht₁ -theorem IsSupClosedCompact.wellFounded (h : IsSupClosedCompact α) : - WellFounded ((· > ·) : α → α → Prop) := by - refine RelEmbedding.wellFounded_iff_no_descending_seq.mpr ⟨fun a => ?_⟩ - suffices sSup (Set.range a) ∈ Set.range a by - obtain ⟨n, hn⟩ := Set.mem_range.mp this - have h' : sSup (Set.range a) < a (n + 1) := by - change _ > _ - simp [← hn, a.map_rel_iff] - apply lt_irrefl (a (n + 1)) - apply lt_of_le_of_lt _ h' - apply le_sSup - apply Set.mem_range_self - apply h (Set.range a) - · use a 37 - apply Set.mem_range_self - · rintro x ⟨m, hm⟩ y ⟨n, hn⟩ - use m ⊔ n - rw [← hm, ← hn] - apply RelHomClass.map_sup a +theorem IsSupClosedCompact.wellFoundedGT (h : IsSupClosedCompact α) : + WellFoundedGT α where + wf := by + refine RelEmbedding.wellFounded_iff_no_descending_seq.mpr ⟨fun a => ?_⟩ + suffices sSup (Set.range a) ∈ Set.range a by + obtain ⟨n, hn⟩ := Set.mem_range.mp this + have h' : sSup (Set.range a) < a (n + 1) := by + change _ > _ + simp [← hn, a.map_rel_iff] + apply lt_irrefl (a (n + 1)) + apply lt_of_le_of_lt _ h' + apply le_sSup + apply Set.mem_range_self + apply h (Set.range a) + · use a 37 + apply Set.mem_range_self + · rintro x ⟨m, hm⟩ y ⟨n, hn⟩ + use m ⊔ n + rw [← hm, ← hn] + apply RelHomClass.map_sup a theorem isSupFiniteCompact_iff_all_elements_compact : IsSupFiniteCompact α ↔ ∀ k : α, IsCompactElement k := by @@ -247,39 +248,38 @@ theorem isSupFiniteCompact_iff_all_elements_compact : exact ⟨t, hts, this⟩ open List in -theorem wellFounded_characterisations : List.TFAE - [WellFounded ((· > ·) : α → α → Prop), - IsSupFiniteCompact α, IsSupClosedCompact α, ∀ k : α, IsCompactElement k] := by - tfae_have 1 → 2 := WellFounded.isSupFiniteCompact α +theorem wellFoundedGT_characterisations : List.TFAE + [WellFoundedGT α, IsSupFiniteCompact α, IsSupClosedCompact α, ∀ k : α, IsCompactElement k] := by + tfae_have 1 → 2 := @WellFoundedGT.isSupFiniteCompact α _ tfae_have 2 → 3 := IsSupFiniteCompact.isSupClosedCompact α - tfae_have 3 → 1 := IsSupClosedCompact.wellFounded α + tfae_have 3 → 1 := IsSupClosedCompact.wellFoundedGT α tfae_have 2 ↔ 4 := isSupFiniteCompact_iff_all_elements_compact α tfae_finish -theorem wellFounded_iff_isSupFiniteCompact : - WellFounded ((· > ·) : α → α → Prop) ↔ IsSupFiniteCompact α := - (wellFounded_characterisations α).out 0 1 +theorem wellFoundedGT_iff_isSupFiniteCompact : + WellFoundedGT α ↔ IsSupFiniteCompact α := + (wellFoundedGT_characterisations α).out 0 1 theorem isSupFiniteCompact_iff_isSupClosedCompact : IsSupFiniteCompact α ↔ IsSupClosedCompact α := - (wellFounded_characterisations α).out 1 2 + (wellFoundedGT_characterisations α).out 1 2 -theorem isSupClosedCompact_iff_wellFounded : - IsSupClosedCompact α ↔ WellFounded ((· > ·) : α → α → Prop) := - (wellFounded_characterisations α).out 2 0 +theorem isSupClosedCompact_iff_wellFoundedGT : + IsSupClosedCompact α ↔ WellFoundedGT α := + (wellFoundedGT_characterisations α).out 2 0 -alias ⟨_, IsSupFiniteCompact.wellFounded⟩ := wellFounded_iff_isSupFiniteCompact +alias ⟨_, IsSupFiniteCompact.wellFoundedGT⟩ := wellFoundedGT_iff_isSupFiniteCompact alias ⟨_, IsSupClosedCompact.isSupFiniteCompact⟩ := isSupFiniteCompact_iff_isSupClosedCompact -alias ⟨_, _root_.WellFounded.isSupClosedCompact⟩ := isSupClosedCompact_iff_wellFounded +alias ⟨_, WellFoundedGT.isSupClosedCompact⟩ := isSupClosedCompact_iff_wellFoundedGT variable {α} -theorem WellFounded.finite_of_setIndependent (h : WellFounded ((· > ·) : α → α → Prop)) {s : Set α} +theorem WellFoundedGT.finite_of_setIndependent [WellFoundedGT α] {s : Set α} (hs : SetIndependent s) : s.Finite := by classical refine Set.not_infinite.mp fun contra => ?_ - obtain ⟨t, ht₁, ht₂⟩ := WellFounded.isSupFiniteCompact α h s + obtain ⟨t, ht₁, ht₂⟩ := WellFoundedGT.isSupFiniteCompact α s replace contra : ∃ x : α, x ∈ s ∧ x ≠ ⊥ ∧ x ∉ t := by have : (s \ (insert ⊥ t : Finset α)).Infinite := contra.diff (Finset.finite_toSet _) obtain ⟨x, hx₁, hx₂⟩ := this.nonempty @@ -292,14 +292,36 @@ theorem WellFounded.finite_of_setIndependent (h : WellFounded ((· > ·) : α rw [← hs, eq_comm, inf_eq_left] exact le_sSup _ _ hx₀ -theorem WellFounded.finite_ne_bot_of_independent (hwf : WellFounded ((· > ·) : α → α → Prop)) +theorem WellFoundedGT.finite_ne_bot_of_independent [WellFoundedGT α] {ι : Type*} {t : ι → α} (ht : Independent t) : Set.Finite {i | t i ≠ ⊥} := by refine Finite.of_finite_image (Finite.subset ?_ (image_subset_range t _)) ht.injOn - exact WellFounded.finite_of_setIndependent hwf ht.setIndependent_range + exact WellFoundedGT.finite_of_setIndependent ht.setIndependent_range -theorem WellFounded.finite_of_independent (hwf : WellFounded ((· > ·) : α → α → Prop)) {ι : Type*} +theorem WellFoundedGT.finite_of_independent [WellFoundedGT α] {ι : Type*} {t : ι → α} (ht : Independent t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Finite ι := - haveI := (WellFounded.finite_of_setIndependent hwf ht.setIndependent_range).to_subtype + haveI := (WellFoundedGT.finite_of_setIndependent ht.setIndependent_range).to_subtype + Finite.of_injective_finite_range (ht.injective h_ne_bot) + +theorem WellFoundedLT.finite_of_setIndependent [WellFoundedLT α] {s : Set α} + (hs : SetIndependent s) : s.Finite := by + by_contra inf + let e := (Infinite.diff inf <| finite_singleton ⊥).to_subtype.natEmbedding + let a n := ⨆ i ≥ n, (e i).1 + have sup_le n : (e n).1 ⊔ a (n + 1) ≤ a n := sup_le_iff.mpr ⟨le_iSup₂_of_le n le_rfl le_rfl, + iSup₂_le fun i hi ↦ le_iSup₂_of_le i (n.le_succ.trans hi) le_rfl⟩ + have lt n : a (n + 1) < a n := (Disjoint.right_lt_sup_of_left_ne_bot + ((hs (e n).2.1).mono_right <| iSup₂_le fun i hi ↦ le_sSup _ _ ?_) (e n).2.2).trans_le (sup_le n) + · exact (RelEmbedding.natGT a lt).not_wellFounded_of_decreasing_seq wellFounded_lt + exact ⟨(e i).2.1, fun h ↦ n.lt_succ_self.not_le <| hi.trans_eq <| e.2 <| Subtype.val_injective h⟩ + +theorem WellFoundedLT.finite_ne_bot_of_independent [WellFoundedLT α] + {ι : Type*} {t : ι → α} (ht : Independent t) : Set.Finite {i | t i ≠ ⊥} := by + refine Finite.of_finite_image (Finite.subset ?_ (image_subset_range t _)) ht.injOn + exact WellFoundedLT.finite_of_setIndependent ht.setIndependent_range + +theorem WellFoundedLT.finite_of_independent [WellFoundedLT α] {ι : Type*} + {t : ι → α} (ht : Independent t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Finite ι := + haveI := (WellFoundedLT.finite_of_setIndependent ht.setIndependent_range).to_subtype Finite.of_injective_finite_range (ht.injective h_ne_bot) end CompleteLattice @@ -313,7 +335,7 @@ class IsCompactlyGenerated (α : Type*) [CompleteLattice α] : Prop where section -variable [IsCompactlyGenerated α] {a b : α} {s : Set α} +variable [IsCompactlyGenerated α] {a : α} {s : Set α} @[simp] theorem sSup_compact_le_eq (b) : @@ -452,12 +474,35 @@ end namespace CompleteLattice -theorem isCompactlyGenerated_of_wellFounded (h : WellFounded ((· > ·) : α → α → Prop)) : +theorem isCompactlyGenerated_of_wellFoundedGT [h : WellFoundedGT α] : IsCompactlyGenerated α := by - rw [wellFounded_iff_isSupFiniteCompact, isSupFiniteCompact_iff_all_elements_compact] at h + rw [wellFoundedGT_iff_isSupFiniteCompact, isSupFiniteCompact_iff_all_elements_compact] at h -- x is the join of the set of compact elements {x} exact ⟨fun x => ⟨{x}, ⟨fun x _ => h x, sSup_singleton⟩⟩⟩ +@[deprecated (since := "2024-10-07")] +alias WellFounded.isSupFiniteCompact := WellFoundedGT.isSupFiniteCompact +@[deprecated (since := "2024-10-07")] +alias IsSupClosedCompact.wellFounded := IsSupClosedCompact.wellFoundedGT +@[deprecated (since := "2024-10-07")] +alias wellFounded_characterisations := wellFoundedGT_characterisations +@[deprecated (since := "2024-10-07")] +alias wellFounded_iff_isSupFiniteCompact := wellFoundedGT_iff_isSupFiniteCompact +@[deprecated (since := "2024-10-07")] +alias isSupClosedCompact_iff_wellFounded := isSupClosedCompact_iff_wellFoundedGT +@[deprecated (since := "2024-10-07")] +alias IsSupFiniteCompact.wellFounded := IsSupFiniteCompact.wellFoundedGT +@[deprecated (since := "2024-10-07")] +alias _root_.WellFounded.isSupClosedCompact := WellFoundedGT.isSupClosedCompact +@[deprecated (since := "2024-10-07")] +alias WellFounded.finite_of_setIndependent := WellFoundedGT.finite_of_setIndependent +@[deprecated (since := "2024-10-07")] +alias WellFounded.finite_ne_bot_of_independent := WellFoundedGT.finite_ne_bot_of_independent +@[deprecated (since := "2024-10-07")] +alias WellFounded.finite_of_independent := WellFoundedGT.finite_of_independent +@[deprecated (since := "2024-10-07")] +alias isCompactlyGenerated_of_wellFounded := isCompactlyGenerated_of_wellFoundedGT + /-- A compact element `k` has the property that any `b < k` lies below a "maximal element below `k`", which is to say `[⊥, k]` is coatomic. -/ theorem Iic_coatomic_of_compact_element {k : α} (h : IsCompactElement k) : diff --git a/Mathlib/Order/Concept.lean b/Mathlib/Order/Concept.lean index b4818c18d6fc2..ea202e9b33def 100644 --- a/Mathlib/Order/Concept.lean +++ b/Mathlib/Order/Concept.lean @@ -37,8 +37,7 @@ concept, formal concept analysis, intent, extend, attribute open Function OrderDual Set -variable {ι : Sort*} {α β γ : Type*} {κ : ι → Sort*} (r : α → β → Prop) {s s₁ s₂ : Set α} - {t t₁ t₂ : Set β} +variable {ι : Sort*} {α β : Type*} {κ : ι → Sort*} (r : α → β → Prop) {s : Set α} {t : Set β} /-! ### Intent and extent -/ diff --git a/Mathlib/Order/Cover.lean b/Mathlib/Order/Cover.lean index 17ba059d015c4..cf13cbab95f7f 100644 --- a/Mathlib/Order/Cover.lean +++ b/Mathlib/Order/Cover.lean @@ -591,3 +591,19 @@ variable [Preorder α] {a b : α} simp only [wcovBy_iff_Ioo_eq, ← image_coe_Iio, bot_le, image_eq_empty, true_and, Iio_eq_empty_iff] end WithBot + +section WellFounded + +variable [Preorder α] + +lemma exists_covBy_of_wellFoundedLT [wf : WellFoundedLT α] ⦃a : α⦄ (h : ¬ IsMax a) : + ∃ a', a ⋖ a' := by + rw [not_isMax_iff] at h + exact ⟨_, wellFounded_lt.min_mem _ h, fun a' ↦ wf.wf.not_lt_min _ h⟩ + +lemma exists_covBy_of_wellFoundedGT [wf : WellFoundedGT α] ⦃a : α⦄ (h : ¬ IsMin a) : + ∃ a', a' ⋖ a := by + rw [not_isMin_iff] at h + exact ⟨_, wf.wf.min_mem _ h, fun a' h₁ h₂ ↦ wf.wf.not_lt_min _ h h₂ h₁⟩ + +end WellFounded diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index 1bdaefb9a2491..a6a6ffa250fe6 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -357,7 +357,7 @@ macro "compareOfLessAndEq_rfl" : tactic => /-- A linear order is reflexive, transitive, antisymmetric and total relation `≤`. We assume that every linear ordered type has decidable `(≤)`, `(<)`, and `(=)`. -/ -class LinearOrder (α : Type*) extends PartialOrder α, Min α, Max α, Ord α := +class LinearOrder (α : Type*) extends PartialOrder α, Min α, Max α, Ord α where /-- A linear order is total. -/ le_total (a b : α) : a ≤ b ∨ b ≤ a /-- In a linearly ordered type, we assume the order relations are all decidable. -/ @@ -595,7 +595,17 @@ lemma compare_iff (a b : α) {o : Ordering} : compare a b = o ↔ o.Compares a b · exact compare_eq_iff_eq · exact compare_gt_iff_gt -instance : Batteries.TransCmp (compare (α := α)) where +theorem cmp_eq_compare (a b : α) : cmp a b = compare a b := by + refine ((compare_iff ..).2 ?_).symm + unfold cmp cmpUsing; split_ifs with h1 h2 + · exact h1 + · exact h2 + · exact le_antisymm (not_lt.1 h2) (not_lt.1 h1) + +theorem cmp_eq_compareOfLessAndEq (a b : α) : cmp a b = compareOfLessAndEq a b := + (cmp_eq_compare ..).trans (LinearOrder.compare_eq_compareOfLessAndEq ..) + +instance : Batteries.LawfulCmp (compare (α := α)) where symm a b := by cases h : compare a b <;> simp only [Ordering.swap] <;> symm @@ -604,6 +614,9 @@ instance : Batteries.TransCmp (compare (α := α)) where · exact compare_lt_iff_lt.2 <| compare_gt_iff_gt.1 h le_trans := fun h₁ h₂ ↦ compare_le_iff_le.2 <| le_trans (compare_le_iff_le.1 h₁) (compare_le_iff_le.1 h₂) + cmp_iff_beq := by simp [compare_eq_iff_eq] + cmp_iff_lt := by simp [compare_lt_iff_lt] + cmp_iff_le := by simp [compare_le_iff_le] end Ord diff --git a/Mathlib/Order/Disjoint.lean b/Mathlib/Order/Disjoint.lean index 29678782f984f..959e072e1b1b6 100644 --- a/Mathlib/Order/Disjoint.lean +++ b/Mathlib/Order/Disjoint.lean @@ -108,7 +108,7 @@ end PartialBoundedOrder section SemilatticeInfBot -variable [SemilatticeInf α] [OrderBot α] {a b c d : α} +variable [SemilatticeInf α] [OrderBot α] {a b c : α} theorem disjoint_iff_inf_le : Disjoint a b ↔ a ⊓ b ≤ ⊥ := ⟨fun hd ↦ hd inf_le_left inf_le_right, fun h _ ha hb ↦ (le_inf ha hb).trans h⟩ @@ -155,6 +155,10 @@ theorem Disjoint.of_disjoint_inf_of_le' (h : Disjoint (a ⊓ b) c) (hle : b ≤ end SemilatticeInfBot +theorem Disjoint.right_lt_sup_of_left_ne_bot [SemilatticeSup α] [OrderBot α] {a b : α} + (h : Disjoint a b) (ha : a ≠ ⊥) : b < a ⊔ b := + le_sup_right.lt_of_ne fun eq ↦ ha (le_bot_iff.mp <| h le_rfl <| sup_eq_right.mp eq.symm) + section DistribLatticeBot variable [DistribLattice α] [OrderBot α] {a b c : α} @@ -267,7 +271,7 @@ end PartialBoundedOrder section SemilatticeSupTop -variable [SemilatticeSup α] [OrderTop α] {a b c d : α} +variable [SemilatticeSup α] [OrderTop α] {a b c : α} theorem codisjoint_iff_le_sup : Codisjoint a b ↔ ⊤ ≤ a ⊔ b := @disjoint_iff_inf_le αᵒᵈ _ _ _ _ @@ -401,7 +405,7 @@ namespace IsCompl section BoundedPartialOrder -variable [PartialOrder α] [BoundedOrder α] {x y z : α} +variable [PartialOrder α] [BoundedOrder α] {x y : α} @[symm] protected theorem symm (h : IsCompl x y) : IsCompl y x := @@ -419,7 +423,7 @@ end BoundedPartialOrder section BoundedLattice -variable [Lattice α] [BoundedOrder α] {x y z : α} +variable [Lattice α] [BoundedOrder α] {x y : α} theorem of_le (h₁ : x ⊓ y ≤ ⊥) (h₂ : ⊤ ≤ x ⊔ y) : IsCompl x y := ⟨disjoint_iff_inf_le.mpr h₁, codisjoint_iff_le_sup.mpr h₂⟩ diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index eda822fe73f47..f393265bd856c 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Jeremy Avigad, Yury Kudryashov, Patrick Massot import Mathlib.Data.Finset.Preimage import Mathlib.Order.ConditionallyCompleteLattice.Basic import Mathlib.Order.Filter.Bases +import Mathlib.Order.Filter.Prod import Mathlib.Order.Interval.Set.Disjoint /-! @@ -257,6 +258,15 @@ variable [Nonempty α] @[instance] lemma atTop_neBot : NeBot (atTop : Filter α) := atTop_basis.neBot_iff.2 fun _ => nonempty_Ici +theorem atTop_neBot_iff {α : Type*} [Preorder α] : + (atTop : Filter α).NeBot ↔ Nonempty α ∧ IsDirected α (· ≤ ·) := by + refine ⟨fun h ↦ ⟨nonempty_of_neBot atTop, ⟨fun x y ↦ ?_⟩⟩, fun ⟨h₁, h₂⟩ ↦ atTop_neBot⟩ + exact ((eventually_ge_atTop x).and (eventually_ge_atTop y)).exists + +theorem atBot_neBot_iff {α : Type*} [Preorder α] : + (atBot : Filter α).NeBot ↔ Nonempty α ∧ IsDirected α (· ≥ ·) := + atTop_neBot_iff (α := αᵒᵈ) + @[simp] lemma mem_atTop_sets {s : Set α} : s ∈ (atTop : Filter α) ↔ ∃ a : α, ∀ b ≥ a, b ∈ s := atTop_basis.mem_iff.trans <| exists_congr fun _ => iff_of_eq (true_and _) @@ -1215,3 +1225,45 @@ theorem Antitone.piecewise_eventually_eq_iInter {β : α → Type*} [Preorder ι convert ← (compl_anti.comp hs).piecewise_eventually_eq_iUnion g f a using 3 · convert congr_fun (Set.piecewise_compl (s _) g f) a · simp only [(· ∘ ·), ← compl_iInter, Set.piecewise_compl] + +namespace Nat + +theorem eventually_pow_lt_factorial_sub (c d : ℕ) : ∀ᶠ n in atTop, c ^ n < (n - d)! := by + rw [eventually_atTop] + refine ⟨2 * (c ^ 2 + d + 1), ?_⟩ + intro n hn + obtain ⟨d', rfl⟩ := Nat.exists_eq_add_of_le hn + obtain (rfl | c0) := c.eq_zero_or_pos + · simp [Nat.two_mul, ← Nat.add_assoc, Nat.add_right_comm _ 1, Nat.factorial_pos] + refine (Nat.le_mul_of_pos_right _ (Nat.pow_pos (n := d') c0)).trans_lt ?_ + convert_to (c ^ 2) ^ (c ^ 2 + d' + d + 1) < (c ^ 2 + (c ^ 2 + d' + d + 1) + 1)! + · rw [← pow_mul, ← pow_add] + congr 1 + omega + · congr 1 + omega + refine (lt_of_lt_of_le ?_ Nat.factorial_mul_pow_le_factorial).trans_le <| + (factorial_le (Nat.le_succ _)) + rw [← one_mul (_ ^ _ : ℕ)] + apply Nat.mul_lt_mul_of_le_of_lt + · exact Nat.one_le_of_lt (Nat.factorial_pos _) + · exact Nat.pow_lt_pow_left (Nat.lt_succ_self _) (Nat.succ_ne_zero _) + · exact (Nat.factorial_pos _) + +theorem eventually_mul_pow_lt_factorial_sub (a c d : ℕ) : + ∀ᶠ n in atTop, a * c ^ n < (n - d)! := by + filter_upwards [Nat.eventually_pow_lt_factorial_sub (a * c) d, Filter.eventually_gt_atTop 0] + with n hn hn0 + rw [mul_pow] at hn + exact (Nat.mul_le_mul_right _ (Nat.le_self_pow hn0.ne' _)).trans_lt hn + +@[deprecated eventually_pow_lt_factorial_sub (since := "2024-09-25")] +theorem exists_pow_lt_factorial (c : ℕ) : ∃ n0 > 1, ∀ n ≥ n0, c ^ n < (n - 1)! := + let ⟨n0, h⟩ := (eventually_pow_lt_factorial_sub c 1).exists_forall_of_atTop + ⟨max n0 2, by omega, fun n hn ↦ h n (by omega)⟩ + +@[deprecated eventually_mul_pow_lt_factorial_sub (since := "2024-09-25")] +theorem exists_mul_pow_lt_factorial (a : ℕ) (c : ℕ) : ∃ n0, ∀ n ≥ n0, a * c ^ n < (n - 1)! := + (eventually_mul_pow_lt_factorial_sub a c 1).exists_forall_of_atTop + +end Nat diff --git a/Mathlib/Order/Filter/AtTopBot/Floor.lean b/Mathlib/Order/Filter/AtTopBot/Floor.lean new file mode 100644 index 0000000000000..32e3e3aad7c86 --- /dev/null +++ b/Mathlib/Order/Filter/AtTopBot/Floor.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2022 Yuyang Zhao. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yuyang Zhao +-/ +import Mathlib.Algebra.Order.Floor +import Mathlib.Order.Filter.AtTopBot + +/-! +# `a * c ^ n < (n - d)!` holds true for sufficiently large `n`. +-/ + +open Filter +open scoped Nat + +variable {K : Type*} [LinearOrderedRing K] [FloorSemiring K] + +theorem FloorSemiring.eventually_mul_pow_lt_factorial_sub (a c : K) (d : ℕ) : + ∀ᶠ n in atTop, a * c ^ n < (n - d)! := by + filter_upwards [Nat.eventually_mul_pow_lt_factorial_sub ⌈|a|⌉₊ ⌈|c|⌉₊ d] with n h + calc a * c ^ n + _ ≤ |a * c ^ n| := le_abs_self _ + _ ≤ ⌈|a|⌉₊ * (⌈|c|⌉₊ : K) ^ n := ?_ + _ = ↑(⌈|a|⌉₊ * ⌈|c|⌉₊ ^ n) := ?_ + _ < (n - d)! := Nat.cast_lt.mpr h + · rw [abs_mul, abs_pow] + gcongr <;> try first | positivity | apply Nat.le_ceil + · simp_rw [Nat.cast_mul, Nat.cast_pow] diff --git a/Mathlib/Order/Filter/AtTopBot/Ring.lean b/Mathlib/Order/Filter/AtTopBot/Ring.lean index 874630711de2c..74207b0e55f14 100644 --- a/Mathlib/Order/Filter/AtTopBot/Ring.lean +++ b/Mathlib/Order/Filter/AtTopBot/Ring.lean @@ -3,8 +3,8 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Order.Filter.AtTopBot.Group -import Mathlib.Algebra.Order.Ring.Basic /-! # Convergence to ±infinity in ordered rings @@ -30,7 +30,7 @@ theorem tendsto_mul_self_atTop : Tendsto (fun x : α => x * x) atTop atTop := /-- The monomial function `x^n` tends to `+∞` at `+∞` for any positive natural `n`. A version for positive real powers exists as `tendsto_rpow_atTop`. -/ theorem tendsto_pow_atTop {n : ℕ} (hn : n ≠ 0) : Tendsto (fun x : α => x ^ n) atTop atTop := - tendsto_atTop_mono' _ ((eventually_ge_atTop 1).mono fun _x hx => le_self_pow hx hn) tendsto_id + tendsto_atTop_mono' _ ((eventually_ge_atTop 1).mono fun _x hx => le_self_pow₀ hx hn) tendsto_id end OrderedSemiring diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index 66d78953d6538..ffd5fe2a6c9e4 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -5,8 +5,7 @@ Authors: Yury Kudryashov, Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import Mathlib.Data.Prod.PProd import Mathlib.Data.Set.Countable -import Mathlib.Order.Filter.Prod -import Mathlib.Order.Filter.Ker +import Mathlib.Order.Filter.Basic /-! # Filter bases @@ -117,7 +116,7 @@ def Filter.asBasis (f : Filter α) : FilterBasis α := ⟨f.sets, ⟨univ, univ_mem⟩, fun {x y} hx hy => ⟨x ∩ y, inter_mem hx hy, subset_rfl⟩⟩ -- Porting note: was `protected` in Lean 3 but `protected` didn't work; removed -/-- `is_basis p s` means the image of `s` bounded by `p` is a filter basis. -/ +/-- `IsBasis p s` means the image of `s` bounded by `p` is a filter basis. -/ structure Filter.IsBasis (p : ι → Prop) (s : ι → Set α) : Prop where /-- There exists at least one `i` that satisfies `p`. -/ nonempty : ∃ i, p i @@ -632,10 +631,6 @@ alias ⟨_, _root_.Disjoint.filter_principal⟩ := disjoint_principal_principal theorem disjoint_pure_pure {x y : α} : Disjoint (pure x : Filter α) (pure y) ↔ x ≠ y := by simp only [← principal_singleton, disjoint_principal_principal, disjoint_singleton] -@[simp] -theorem compl_diagonal_mem_prod {l₁ l₂ : Filter α} : (diagonal α)ᶜ ∈ l₁ ×ˢ l₂ ↔ Disjoint l₁ l₂ := by - simp only [mem_prod_iff, Filter.disjoint_iff, prod_subset_compl_diagonal_iff_disjoint] - -- Porting note: use `∃ i, p i ∧ _` instead of `∃ i (hi : p i), _`. theorem HasBasis.disjoint_iff_left (h : l.HasBasis p s) : Disjoint l l' ↔ ∃ i, p i ∧ (s i)ᶜ ∈ l' := by @@ -722,7 +717,7 @@ protected theorem HasBasis.biInter_mem {f : Set α → Set β} (h : HasBasis l p h.biInf_mem hf protected theorem HasBasis.ker (h : HasBasis l p s) : l.ker = ⋂ (i) (_ : p i), s i := - l.ker_def.trans <| h.biInter_mem monotone_id + sInter_eq_biInter.trans <| h.biInter_mem monotone_id variable {ι'' : Type*} [Preorder ι''] (l) (s'' : ι'' → Set α) diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index d73d46a740880..85fed348f7035 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -4,24 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jeremy Avigad -/ import Mathlib.Data.Set.Finite +import Mathlib.Order.Filter.Defs /-! # Theory of filters on sets -## Main definitions - -* `Filter` : filters on a set; -* `Filter.principal` : filter of all sets containing a given set; -* `Filter.map`, `Filter.comap` : operations on filters; -* `Filter.Tendsto` : limit with respect to filters; -* `Filter.Eventually` : `f.eventually p` means `{x | p x} ∈ f`; -* `Filter.Frequently` : `f.frequently p` means `{x | ¬p x} ∉ f`; -* `filter_upwards [h₁, ..., hₙ]` : - a tactic that takes a list of proofs `hᵢ : sᵢ ∈ f`, - and replaces a goal `s ∈ f` with `∀ x, x ∈ s₁ → ... → x ∈ sₙ → x ∈ s`; -* `Filter.NeBot f` : a utility class stating that `f` is a non-trivial filter. - -Filters on a type `X` are sets of sets of `X` satisfying three conditions. They are mostly used to +A *filter* on a type `α` is a collection of sets of `α` which contains the whole `α`, +is upwards-closed, and is stable under intersection. They are mostly used to abstract two related kinds of ideas: * *limits*, including finite or infinite limits of sequences, finite or infinite limits of functions at a point or at infinity, etc... @@ -30,8 +19,10 @@ abstract two related kinds of ideas: sense of measure theory. Dually, filters can also express the idea of *things happening often*: for arbitrarily large `n`, or at a point in any neighborhood of given a point etc... -In this file, we define the type `Filter X` of filters on `X`, and endow it with a complete lattice -structure. This structure is lifted from the lattice structure on `Set (Set X)` using the Galois +## Main definitions + +In this file, we endow `Filter α` it with a complete lattice structure. +This structure is lifted from the lattice structure on `Set (Set X)` using the Galois insertion which maps a filter to its elements in one direction, and an arbitrary set of sets to the smallest filter containing it in the other direction. We also prove `Filter` is a monadic functor, with a push-forward operation @@ -82,69 +73,27 @@ open scoped symmDiff universe u v w x y -/-- A filter `F` on a type `α` is a collection of sets of `α` which contains the whole `α`, -is upwards-closed, and is stable under intersection. We do not forbid this collection to be -all sets of `α`. -/ -structure Filter (α : Type*) where - /-- The set of sets that belong to the filter. -/ - sets : Set (Set α) - /-- The set `Set.univ` belongs to any filter. -/ - univ_sets : Set.univ ∈ sets - /-- If a set belongs to a filter, then its superset belongs to the filter as well. -/ - sets_of_superset {x y} : x ∈ sets → x ⊆ y → y ∈ sets - /-- If two sets belong to a filter, then their intersection belongs to the filter as well. -/ - inter_sets {x y} : x ∈ sets → y ∈ sets → x ∩ y ∈ sets - -/-- If `F` is a filter on `α`, and `U` a subset of `α` then we can write `U ∈ F` as on paper. -/ -instance {α : Type*} : Membership (Set α) (Filter α) := - ⟨fun F U => U ∈ F.sets⟩ - namespace Filter variable {α : Type u} {f g : Filter α} {s t : Set α} -@[simp] -protected theorem mem_mk {t : Set (Set α)} {h₁ h₂ h₃} : s ∈ mk t h₁ h₂ h₃ ↔ s ∈ t := - Iff.rfl - -@[simp] -protected theorem mem_sets : s ∈ f.sets ↔ s ∈ f := - Iff.rfl - instance inhabitedMem : Inhabited { s : Set α // s ∈ f } := ⟨⟨univ, f.univ_sets⟩⟩ -theorem filter_eq : ∀ {f g : Filter α}, f.sets = g.sets → f = g - | ⟨_, _, _, _⟩, ⟨_, _, _, _⟩, rfl => rfl - theorem filter_eq_iff : f = g ↔ f.sets = g.sets := ⟨congr_arg _, filter_eq⟩ -@[ext] -protected theorem ext (h : ∀ s, s ∈ f ↔ s ∈ g) : f = g := by - simpa [filter_eq_iff, Set.ext_iff, Filter.mem_sets] - /-- An extensionality lemma that is useful for filters with good lemmas about `sᶜ ∈ f` (e.g., `Filter.comap`, `Filter.coprod`, `Filter.Coprod`, `Filter.cofinite`). -/ protected theorem coext (h : ∀ s, sᶜ ∈ f ↔ sᶜ ∈ g) : f = g := Filter.ext <| compl_surjective.forall.2 h -@[simp] -theorem univ_mem : univ ∈ f := - f.univ_sets - -theorem mem_of_superset {x y : Set α} (hx : x ∈ f) (hxy : x ⊆ y) : y ∈ f := - f.sets_of_superset hx hxy - instance : Trans (· ⊇ ·) ((· ∈ ·) : Set α → Filter α → Prop) (· ∈ ·) where trans h₁ h₂ := mem_of_superset h₂ h₁ instance : Trans Membership.mem (· ⊆ ·) (Membership.mem : Filter α → Set α → Prop) where trans h₁ h₂ := mem_of_superset h₁ h₂ -theorem inter_mem {s t : Set α} (hs : s ∈ f) (ht : t ∈ f) : s ∩ t ∈ f := - f.inter_sets hs ht - @[simp] theorem inter_mem_iff {s t : Set α} : s ∩ t ∈ f ↔ s ∈ f ∧ t ∈ f := ⟨fun h => ⟨mem_of_superset h inter_subset_left, mem_of_superset h inter_subset_right⟩, @@ -153,27 +102,12 @@ theorem inter_mem_iff {s t : Set α} : s ∩ t ∈ f ↔ s ∈ f ∧ t ∈ f := theorem diff_mem {s t : Set α} (hs : s ∈ f) (ht : tᶜ ∈ f) : s \ t ∈ f := inter_mem hs ht -theorem univ_mem' (h : ∀ a, a ∈ s) : s ∈ f := - mem_of_superset univ_mem fun x _ => h x - -theorem mp_mem (hs : s ∈ f) (h : { x | x ∈ s → x ∈ t } ∈ f) : t ∈ f := - mem_of_superset (inter_mem hs h) fun _ ⟨h₁, h₂⟩ => h₂ h₁ - theorem congr_sets (h : { x | x ∈ s ↔ x ∈ t } ∈ f) : s ∈ f ↔ t ∈ f := ⟨fun hs => mp_mem hs (mem_of_superset h fun _ => Iff.mp), fun hs => mp_mem hs (mem_of_superset h fun _ => Iff.mpr)⟩ -/-- Override `sets` field of a filter to provide better definitional equality. -/ -protected def copy (f : Filter α) (S : Set (Set α)) (hmem : ∀ s, s ∈ S ↔ s ∈ f) : Filter α where - sets := S - univ_sets := (hmem _).2 univ_mem - sets_of_superset h hsub := (hmem _).2 <| mem_of_superset ((hmem _).1 h) hsub - inter_sets h₁ h₂ := (hmem _).2 <| inter_mem ((hmem _).1 h₁) ((hmem _).1 h₂) - lemma copy_eq {S} (hmem : ∀ s, s ∈ S ↔ s ∈ f) : f.copy S hmem = f := Filter.ext hmem -@[simp] lemma mem_copy {S hmem} : s ∈ f.copy S hmem ↔ s ∈ S := Iff.rfl - @[simp] theorem biInter_mem {β : Type v} {s : β → Set α} {is : Set β} (hf : is.Finite) : (⋂ i ∈ is, s i) ∈ f ↔ ∀ i ∈ is, s i ∈ f := @@ -217,100 +151,17 @@ theorem forall_in_swap {β : Type*} {p : Set α → β → Prop} : end Filter -namespace Mathlib.Tactic - -open Lean Meta Elab Tactic - -/-- -`filter_upwards [h₁, ⋯, hₙ]` replaces a goal of the form `s ∈ f` and terms -`h₁ : t₁ ∈ f, ⋯, hₙ : tₙ ∈ f` with `∀ x, x ∈ t₁ → ⋯ → x ∈ tₙ → x ∈ s`. -The list is an optional parameter, `[]` being its default value. - -`filter_upwards [h₁, ⋯, hₙ] with a₁ a₂ ⋯ aₖ` is a short form for -`{ filter_upwards [h₁, ⋯, hₙ], intros a₁ a₂ ⋯ aₖ }`. - -`filter_upwards [h₁, ⋯, hₙ] using e` is a short form for -`{ filter_upwards [h1, ⋯, hn], exact e }`. - -Combining both shortcuts is done by writing `filter_upwards [h₁, ⋯, hₙ] with a₁ a₂ ⋯ aₖ using e`. -Note that in this case, the `aᵢ` terms can be used in `e`. --/ -syntax (name := filterUpwards) "filter_upwards" (" [" term,* "]")? - (" with" (ppSpace colGt term:max)*)? (" using " term)? : tactic - -elab_rules : tactic -| `(tactic| filter_upwards $[[$[$args],*]]? $[with $wth*]? $[using $usingArg]?) => do - let config : ApplyConfig := {newGoals := ApplyNewGoals.nonDependentOnly} - for e in args.getD #[] |>.reverse do - let goal ← getMainGoal - replaceMainGoal <| ← goal.withContext <| runTermElab do - let m ← mkFreshExprMVar none - let lem ← Term.elabTermEnsuringType - (← ``(Filter.mp_mem $e $(← Term.exprToSyntax m))) (← goal.getType) - goal.assign lem - return [m.mvarId!] - liftMetaTactic fun goal => do - goal.apply (← mkConstWithFreshMVarLevels ``Filter.univ_mem') config - evalTactic <|← `(tactic| dsimp (config := {zeta := false}) only [Set.mem_setOf_eq]) - if let some l := wth then - evalTactic <|← `(tactic| intro $[$l]*) - if let some e := usingArg then - evalTactic <|← `(tactic| exact $e) - -end Mathlib.Tactic namespace Filter variable {α : Type u} {β : Type v} {γ : Type w} {δ : Type*} {ι : Sort x} -section Principal - -/-- The principal filter of `s` is the collection of all supersets of `s`. -/ -def principal (s : Set α) : Filter α where - sets := { t | s ⊆ t } - univ_sets := subset_univ s - sets_of_superset hx := Subset.trans hx - inter_sets := subset_inter - -@[inherit_doc] -scoped notation "𝓟" => Filter.principal - -@[simp] theorem mem_principal {s t : Set α} : s ∈ 𝓟 t ↔ t ⊆ s := Iff.rfl - theorem mem_principal_self (s : Set α) : s ∈ 𝓟 s := Subset.rfl -end Principal - -open Filter - -section Join - -/-- The join of a filter of filters is defined by the relation `s ∈ join f ↔ {t | s ∈ t} ∈ f`. -/ -def join (f : Filter (Filter α)) : Filter α where - sets := { s | { t : Filter α | s ∈ t } ∈ f } - univ_sets := by simp only [mem_setOf_eq, univ_sets, ← Filter.mem_sets, setOf_true] - sets_of_superset hx xy := mem_of_superset hx fun f h => mem_of_superset h xy - inter_sets hx hy := mem_of_superset (inter_mem hx hy) fun f ⟨h₁, h₂⟩ => inter_mem h₁ h₂ - -@[simp] -theorem mem_join {s : Set α} {f : Filter (Filter α)} : s ∈ join f ↔ { t | s ∈ t } ∈ f := - Iff.rfl - -end Join - section Lattice variable {f g : Filter α} {s t : Set α} -instance : PartialOrder (Filter α) where - le f g := ∀ ⦃U : Set α⦄, U ∈ g → U ∈ f - le_antisymm a b h₁ h₂ := filter_eq <| Subset.antisymm h₂ h₁ - le_refl a := Subset.rfl - le_trans a b c h₁ h₂ := Subset.trans h₂ h₁ - -theorem le_def : f ≤ g ↔ ∀ x ∈ g, x ∈ f := - Iff.rfl - protected theorem not_le : ¬f ≤ g ↔ ∃ s ∈ g, s ∉ f := by simp_rw [le_def, not_forall, exists_prop] /-- `GenerateSets g s`: `s` is in the filter closure of `g`. -/ @@ -378,23 +229,6 @@ def giGenerate (α : Type*) : choice s hs := Filter.mkOfClosure s (le_antisymm hs <| le_generate_iff.1 <| le_rfl) choice_eq _ _ := mkOfClosure_sets -/-- The infimum of filters is the filter generated by intersections - of elements of the two filters. -/ -instance : Inf (Filter α) := - ⟨fun f g : Filter α => - { sets := { s | ∃ a ∈ f, ∃ b ∈ g, s = a ∩ b } - univ_sets := ⟨_, univ_mem, _, univ_mem, by simp⟩ - sets_of_superset := by - rintro x y ⟨a, ha, b, hb, rfl⟩ xy - refine - ⟨a ∪ y, mem_of_superset ha subset_union_left, b ∪ y, - mem_of_superset hb subset_union_left, ?_⟩ - rw [← inter_union_distrib_right, union_eq_self_of_subset_left xy] - inter_sets := by - rintro x y ⟨a, ha, b, hb, rfl⟩ ⟨c, hc, d, hd, rfl⟩ - refine ⟨a ∩ c, inter_mem ha hc, b ∩ d, inter_mem hb hd, ?_⟩ - ac_rfl }⟩ - theorem mem_inf_iff {f g : Filter α} {s : Set α} : s ∈ f ⊓ g ↔ ∃ t₁ ∈ f, ∃ t₂ ∈ g, s = t₁ ∩ t₂ := Iff.rfl @@ -417,19 +251,6 @@ theorem mem_inf_iff_superset {f g : Filter α} {s : Set α} : ⟨fun ⟨t₁, h₁, t₂, h₂, Eq⟩ => ⟨t₁, h₁, t₂, h₂, Eq ▸ Subset.rfl⟩, fun ⟨_, h₁, _, h₂, sub⟩ => mem_inf_of_inter h₁ h₂ sub⟩ -instance : Top (Filter α) := - ⟨{ sets := { s | ∀ x, x ∈ s } - univ_sets := fun x => mem_univ x - sets_of_superset := fun hx hxy a => hxy (hx a) - inter_sets := fun hx hy _ => mem_inter (hx _) (hy _) }⟩ - -theorem mem_top_iff_forall {s : Set α} : s ∈ (⊤ : Filter α) ↔ ∀ x, x ∈ s := - Iff.rfl - -@[simp] -theorem mem_top {s : Set α} : s ∈ (⊤ : Filter α) ↔ s = univ := by - rw [mem_top_iff_forall, eq_univ_iff_forall] - section CompleteLattice /- We lift the complete lattice along the Galois connection `generate` / `sets`. Unfortunately, @@ -452,16 +273,6 @@ instance : Inhabited (Filter α) := ⟨⊥⟩ end CompleteLattice -/-- A filter is `NeBot` if it is not equal to `⊥`, or equivalently the empty set does not belong to -the filter. Bourbaki include this assumption in the definition of a filter but we prefer to have a -`CompleteLattice` structure on `Filter _`, so we use a typeclass argument in lemmas instead. -/ -class NeBot (f : Filter α) : Prop where - /-- The filter is nontrivial: `f ≠ ⊥` or equivalently, `∅ ∉ f`. -/ - ne' : f ≠ ⊥ - -theorem neBot_iff {f : Filter α} : NeBot f ↔ f ≠ ⊥ := - ⟨fun h => h.1, fun h => ⟨h⟩⟩ - theorem NeBot.ne {f : Filter α} (hf : NeBot f) : f ≠ ⊥ := hf.ne' @[simp] theorem not_neBot {f : Filter α} : ¬f.NeBot ↔ f = ⊥ := neBot_iff.not_left @@ -506,10 +317,6 @@ theorem generate_iUnion {s : ι → Set (Set α)} : Filter.generate (⋃ i, s i) = ⨅ i, Filter.generate (s i) := (giGenerate α).gc.l_iSup -@[simp] -theorem mem_bot {s : Set α} : s ∈ (⊥ : Filter α) := - trivial - @[simp] theorem mem_sup {f g : Filter α} {s : Set α} : s ∈ f ⊔ g ↔ s ∈ f ∧ s ∈ g := Iff.rfl @@ -950,14 +757,6 @@ theorem join_mono {f₁ f₂ : Filter (Filter α)} (h : f₁ ≤ f₂) : join f /-! ### Eventually -/ -/-- `f.Eventually p` or `∀ᶠ x in f, p x` mean that `{x | p x} ∈ f`. E.g., `∀ᶠ x in atTop, p x` -means that `p` holds true for sufficiently large `x`. -/ -protected def Eventually (p : α → Prop) (f : Filter α) : Prop := - { x | p x } ∈ f - -@[inherit_doc Filter.Eventually] -notation3 "∀ᶠ "(...)" in "f", "r:(scoped p => Filter.Eventually p f) => r - theorem eventually_iff {f : Filter α} {P : α → Prop} : (∀ᶠ x in f, P x) ↔ { x | P x } ∈ f := Iff.rfl @@ -1106,14 +905,6 @@ theorem eventually_inf_principal {f : Filter α} {p : α → Prop} {s : Set α} /-! ### Frequently -/ -/-- `f.Frequently p` or `∃ᶠ x in f, p x` mean that `{x | ¬p x} ∉ f`. E.g., `∃ᶠ x in atTop, p x` -means that there exist arbitrarily large `x` for which `p` holds true. -/ -protected def Frequently (p : α → Prop) (f : Filter α) : Prop := - ¬∀ᶠ x in f, ¬p x - -@[inherit_doc Filter.Frequently] -notation3 "∃ᶠ "(...)" in "f", "r:(scoped p => Filter.Frequently p f) => r - theorem Eventually.frequently {f : Filter α} [NeBot f] {p : α → Prop} (h : ∀ᶠ x in f, p x) : ∃ᶠ x in f, p x := compl_not_mem h @@ -1266,17 +1057,9 @@ theorem Eventually.choice {r : α → β → Prop} {l : Filter α} [l.NeBot] (h ### Relation “eventually equal” -/ -/-- Two functions `f` and `g` are *eventually equal* along a filter `l` if the set of `x` such that -`f x = g x` belongs to `l`. -/ -def EventuallyEq (l : Filter α) (f g : α → β) : Prop := - ∀ᶠ x in l, f x = g x - section EventuallyEq variable {l : Filter α} {f g : α → β} -@[inherit_doc] -notation:50 f " =ᶠ[" l:50 "] " g:50 => EventuallyEq l f g - theorem EventuallyEq.eventually (h : f =ᶠ[l] g) : ∀ᶠ x in l, f x = g x := h @[simp] lemma eventuallyEq_top : f =ᶠ[⊤] g ↔ f = g := by simp [EventuallyEq, funext_iff] @@ -1449,13 +1232,6 @@ section LE variable [LE β] {l : Filter α} -/-- A function `f` is eventually less than or equal to a function `g` at a filter `l`. -/ -def EventuallyLE (l : Filter α) (f g : α → β) : Prop := - ∀ᶠ x in l, f x ≤ g x - -@[inherit_doc] -notation:50 f " ≤ᶠ[" l:50 "] " g:50 => EventuallyLE l f g - theorem EventuallyLE.congr {f f' g g' : α → β} (H : f ≤ᶠ[l] g) (hf : f =ᶠ[l] f') (hg : g =ᶠ[l] g') : f' ≤ᶠ[l] g' := H.mp <| hg.mp <| hf.mono fun x hf hg H => by rwa [hf, hg] at H @@ -1654,13 +1430,6 @@ end EventuallyEq section Map -/-- The forward map of a filter -/ -def map (m : α → β) (f : Filter α) : Filter β where - sets := preimage m ⁻¹' f.sets - univ_sets := univ_mem - sets_of_superset hs st := mem_of_superset hs <| preimage_mono st - inter_sets hs ht := inter_mem hs ht - @[simp] theorem map_principal {s : Set α} {f : α → β} : map f (𝓟 s) = 𝓟 (Set.image f s) := Filter.ext fun _ => image_subset_iff.symm @@ -1725,20 +1494,6 @@ end Map section Comap -/-- The inverse map of a filter. A set `s` belongs to `Filter.comap m f` if either of the following -equivalent conditions hold. - -1. There exists a set `t ∈ f` such that `m ⁻¹' t ⊆ s`. This is used as a definition. -2. The set `kernImage m s = {y | ∀ x, m x = y → x ∈ s}` belongs to `f`, see `Filter.mem_comap'`. -3. The set `(m '' sᶜ)ᶜ` belongs to `f`, see `Filter.mem_comap_iff_compl` and -`Filter.compl_mem_comap`. -/ -def comap (m : α → β) (f : Filter β) : Filter α where - sets := { s | ∃ t ∈ f, m ⁻¹' t ⊆ s } - univ_sets := ⟨univ, univ_mem, by simp only [subset_univ, preimage_univ]⟩ - sets_of_superset := fun ⟨a', ha', ma'a⟩ ab => ⟨a', ha', ma'a.trans ab⟩ - inter_sets := fun ⟨a', ha₁, ha₂⟩ ⟨b', hb₁, hb₂⟩ => - ⟨a' ∩ b', inter_mem ha₁ hb₁, inter_subset_inter ha₂ hb₂⟩ - variable {f : α → β} {l : Filter β} {p : α → Prop} {s : Set α} theorem mem_comap' : s ∈ comap f l ↔ { y | ∀ ⦃x⦄, f x = y → x ∈ s } ∈ l := @@ -1809,37 +1564,6 @@ theorem compl_mem_kernMap {s : Set β} : sᶜ ∈ kernMap m f ↔ ∃ t, tᶜ end KernMap -/-- The monadic bind operation on filter is defined the usual way in terms of `map` and `join`. - -Unfortunately, this `bind` does not result in the expected applicative. See `Filter.seq` for the -applicative instance. -/ -def bind (f : Filter α) (m : α → Filter β) : Filter β := - join (map m f) - -/-- The applicative sequentiation operation. This is not induced by the bind operation. -/ -def seq (f : Filter (α → β)) (g : Filter α) : Filter β where - sets := { s | ∃ u ∈ f, ∃ t ∈ g, ∀ m ∈ u, ∀ x ∈ t, (m : α → β) x ∈ s } - univ_sets := ⟨univ, univ_mem, univ, univ_mem, fun _ _ _ _ => trivial⟩ - sets_of_superset := fun ⟨t₀, t₁, h₀, h₁, h⟩ hst => - ⟨t₀, t₁, h₀, h₁, fun _ hx _ hy => hst <| h _ hx _ hy⟩ - inter_sets := fun ⟨t₀, ht₀, t₁, ht₁, ht⟩ ⟨u₀, hu₀, u₁, hu₁, hu⟩ => - ⟨t₀ ∩ u₀, inter_mem ht₀ hu₀, t₁ ∩ u₁, inter_mem ht₁ hu₁, fun _ ⟨hx₀, hx₁⟩ _ ⟨hy₀, hy₁⟩ => - ⟨ht _ hx₀ _ hy₀, hu _ hx₁ _ hy₁⟩⟩ - -/-- `pure x` is the set of sets that contain `x`. It is equal to `𝓟 {x}` but -with this definition we have `s ∈ pure a` defeq `a ∈ s`. -/ -instance : Pure Filter := - ⟨fun x => - { sets := { s | x ∈ s } - inter_sets := And.intro - sets_of_superset := fun hs hst => hst hs - univ_sets := trivial }⟩ - -instance : Bind Filter := - ⟨@Filter.bind⟩ - -instance : Functor Filter where map := @Filter.map - instance : LawfulFunctor (Filter : Type u → Type u) where id_map _ := map_id comp_map _ _ _ := map_map.symm @@ -2596,12 +2320,6 @@ end Bind /-! ### Limits -/ -/-- `Filter.Tendsto` is the generic "limit of a function" predicate. - `Tendsto f l₁ l₂` asserts that for every `l₂` neighborhood `a`, - the `f`-preimage of `a` is an `l₁` neighborhood. -/ -def Tendsto (f : α → β) (l₁ : Filter α) (l₂ : Filter β) := - l₁.map f ≤ l₂ - theorem tendsto_def {f : α → β} {l₁ : Filter α} {l₂ : Filter β} : Tendsto f l₁ l₂ ↔ ∀ s ∈ l₂, f ⁻¹' s ∈ l₁ := Iff.rfl @@ -2638,6 +2356,12 @@ theorem Tendsto.frequently_map {l₁ : Filter α} {l₂ : Filter β} {p : α → @[simp] theorem tendsto_bot {f : α → β} {l : Filter β} : Tendsto f ⊥ l := by simp [Tendsto] +theorem Tendsto.of_neBot_imp {f : α → β} {la : Filter α} {lb : Filter β} + (h : NeBot la → Tendsto f la lb) : Tendsto f la lb := by + rcases eq_or_neBot la with rfl | hla + · exact tendsto_bot + · exact h hla + @[simp] theorem tendsto_top {f : α → β} {l : Filter α} : Tendsto f l ⊤ := le_top theorem le_map_of_right_inverse {mab : α → β} {mba : β → α} {f : Filter α} {g : Filter β} @@ -2915,21 +2639,6 @@ alias ⟨_, Set.InjOn.filter_map_Iic⟩ := Filter.filter_injOn_Iic_iff_injOn namespace Filter -/-- Construct a filter from a property that is stable under finite unions. -A set `s` belongs to `Filter.comk p _ _ _` iff its complement satisfies the predicate `p`. -This constructor is useful to define filters like `Filter.cofinite`. -/ -def comk (p : Set α → Prop) (he : p ∅) (hmono : ∀ t, p t → ∀ s ⊆ t, p s) - (hunion : ∀ s, p s → ∀ t, p t → p (s ∪ t)) : Filter α where - sets := {t | p tᶜ} - univ_sets := by simpa - sets_of_superset := fun ht₁ ht => hmono _ ht₁ _ (compl_subset_compl.2 ht) - inter_sets := fun ht₁ ht₂ => by simp [compl_inter, hunion _ ht₁ _ ht₂] - -@[simp] -lemma mem_comk {p : Set α → Prop} {he hmono hunion s} : - s ∈ comk p he hmono hunion ↔ p sᶜ := - .rfl - lemma compl_mem_comk {p : Set α → Prop} {he hmono hunion s} : sᶜ ∈ comk p he hmono hunion ↔ p s := by simp diff --git a/Mathlib/Order/Filter/Cofinite.lean b/Mathlib/Order/Filter/Cofinite.lean index 03cc4bdf98814..94c269c447ffa 100644 --- a/Mathlib/Order/Filter/Cofinite.lean +++ b/Mathlib/Order/Filter/Cofinite.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jeremy Avigad, Yury Kudryashov -/ import Mathlib.Order.Filter.AtTopBot +import Mathlib.Order.Filter.Ker import Mathlib.Order.Filter.Pi /-! diff --git a/Mathlib/Order/Filter/Curry.lean b/Mathlib/Order/Filter/Curry.lean index 561f48bd033b3..7f91d3cc76ef9 100644 --- a/Mathlib/Order/Filter/Curry.lean +++ b/Mathlib/Order/Filter/Curry.lean @@ -49,13 +49,6 @@ namespace Filter variable {α β γ : Type*} -/-- This filter is characterized by `Filter.eventually_curry_iff`: -`(∀ᶠ (x : α × β) in f.curry g, p x) ↔ ∀ᶠ (x : α) in f, ∀ᶠ (y : β) in g, p (x, y)`. Useful -in adding quantifiers to the middle of `Tendsto`s. See -`hasFDerivAt_of_tendstoUniformlyOnFilter`. -/ -def curry (f : Filter α) (g : Filter β) : Filter (α × β) := - bind f fun a ↦ map (a, ·) g - theorem eventually_curry_iff {f : Filter α} {g : Filter β} {p : α × β → Prop} : (∀ᶠ x : α × β in f.curry g, p x) ↔ ∀ᶠ x : α in f, ∀ᶠ y : β in g, p (x, y) := Iff.rfl diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean new file mode 100644 index 0000000000000..cd23371bb3f2f --- /dev/null +++ b/Mathlib/Order/Filter/Defs.lean @@ -0,0 +1,378 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jeremy Avigad +-/ +import Mathlib.Data.Set.Basic +import Mathlib.Order.SetNotation + +/-! +# Definitions about filters + +A *filter* on a type `α` is a collection of sets of `α` which contains the whole `α`, +is upwards-closed, and is stable under intersection. Filters are mostly used to +abstract two related kinds of ideas: +* *limits*, including finite or infinite limits of sequences, finite or infinite limits of functions + at a point or at infinity, etc... +* *things happening eventually*, including things happening for large enough `n : ℕ`, or near enough + a point `x`, or for close enough pairs of points, or things happening almost everywhere in the + sense of measure theory. Dually, filters can also express the idea of *things happening often*: + for arbitrarily large `n`, or at a point in any neighborhood of given a point etc... + +## Main definitions + +* `Filter` : filters on a set; +* `Filter.principal`, `𝓟 s` : filter of all sets containing a given set; +* `Filter.map`, `Filter.comap` : operations on filters; +* `Filter.Tendsto` : limit with respect to filters; +* `Filter.Eventually` : `f.Eventually p` means `{x | p x} ∈ f`; +* `Filter.Frequently` : `f.Frequently p` means `{x | ¬p x} ∉ f`; +* `filter_upwards [h₁, ..., hₙ]` : + a tactic that takes a list of proofs `hᵢ : sᵢ ∈ f`, + and replaces a goal `s ∈ f` with `∀ x, x ∈ s₁ → ... → x ∈ sₙ → x ∈ s`; +* `Filter.NeBot f` : a utility class stating that `f` is a non-trivial filter. + +## Notations + +* `∀ᶠ x in f, p x` : `f.Eventually p`; +* `∃ᶠ x in f, p x` : `f.Frequently p`; +* `f =ᶠ[l] g` : `∀ᶠ x in l, f x = g x`; +* `f ≤ᶠ[l] g` : `∀ᶠ x in l, f x ≤ g x`; +* `𝓟 s` : `Filter.Principal s`, localized in `Filter`. + +## Implementation Notes + +Important note: Bourbaki requires that a filter on `X` cannot contain all sets of `X`, +which we do *not* require. +This gives `Filter X` better formal properties, +in particular a bottom element `⊥` for its lattice structure, +at the cost of including the assumption `[NeBot f]` in a number of lemmas and definitions. + +## References + +* [N. Bourbaki, *General Topology*][bourbaki1966] +-/ + +open Set + +/-- A filter `F` on a type `α` is a collection of sets of `α` which contains the whole `α`, +is upwards-closed, and is stable under intersection. We do not forbid this collection to be +all sets of `α`. -/ +structure Filter (α : Type*) where + /-- The set of sets that belong to the filter. -/ + sets : Set (Set α) + /-- The set `Set.univ` belongs to any filter. -/ + univ_sets : Set.univ ∈ sets + /-- If a set belongs to a filter, then its superset belongs to the filter as well. -/ + sets_of_superset {x y} : x ∈ sets → x ⊆ y → y ∈ sets + /-- If two sets belong to a filter, then their intersection belongs to the filter as well. -/ + inter_sets {x y} : x ∈ sets → y ∈ sets → x ∩ y ∈ sets + +namespace Filter + +variable {α β : Type*} {f g : Filter α} {s t : Set α} + +theorem filter_eq : ∀ {f g : Filter α}, f.sets = g.sets → f = g + | ⟨_, _, _, _⟩, ⟨_, _, _, _⟩, rfl => rfl + +/-- If `F` is a filter on `α`, and `U` a subset of `α` then we can write `U ∈ F` as on paper. -/ +instance instMembership : Membership (Set α) (Filter α) := ⟨fun F U => U ∈ F.sets⟩ + +@[ext] +protected theorem ext (h : ∀ s, s ∈ f ↔ s ∈ g) : f = g := filter_eq <| Set.ext h + +@[simp] +protected theorem mem_mk {t : Set (Set α)} {h₁ h₂ h₃} : s ∈ mk t h₁ h₂ h₃ ↔ s ∈ t := + Iff.rfl + +@[simp] +protected theorem mem_sets : s ∈ f.sets ↔ s ∈ f := + Iff.rfl + +@[simp] +theorem univ_mem : univ ∈ f := + f.univ_sets + +theorem mem_of_superset {x y : Set α} (hx : x ∈ f) (hxy : x ⊆ y) : y ∈ f := + f.sets_of_superset hx hxy + +theorem univ_mem' (h : ∀ a, a ∈ s) : s ∈ f := + mem_of_superset univ_mem fun x _ => h x + +theorem inter_mem (hs : s ∈ f) (ht : t ∈ f) : s ∩ t ∈ f := + f.inter_sets hs ht + +theorem mp_mem (hs : s ∈ f) (h : { x | x ∈ s → x ∈ t } ∈ f) : t ∈ f := + mem_of_superset (inter_mem hs h) fun _ ⟨h₁, h₂⟩ => h₂ h₁ + +/-- Override `sets` field of a filter to provide better definitional equality. -/ +protected def copy (f : Filter α) (S : Set (Set α)) (hmem : ∀ s, s ∈ S ↔ s ∈ f) : Filter α where + sets := S + univ_sets := (hmem _).2 univ_mem + sets_of_superset h hsub := (hmem _).2 <| mem_of_superset ((hmem _).1 h) hsub + inter_sets h₁ h₂ := (hmem _).2 <| inter_mem ((hmem _).1 h₁) ((hmem _).1 h₂) + +@[simp] theorem mem_copy {S hmem} : s ∈ f.copy S hmem ↔ s ∈ S := Iff.rfl + +/-- Construct a filter from a property that is stable under finite unions. +A set `s` belongs to `Filter.comk p _ _ _` iff its complement satisfies the predicate `p`. +This constructor is useful to define filters like `Filter.cofinite`. -/ +def comk (p : Set α → Prop) (he : p ∅) (hmono : ∀ t, p t → ∀ s ⊆ t, p s) + (hunion : ∀ s, p s → ∀ t, p t → p (s ∪ t)) : Filter α where + sets := {t | p tᶜ} + univ_sets := by simpa + sets_of_superset := fun ht₁ ht => hmono _ ht₁ _ (compl_subset_compl.2 ht) + inter_sets := fun ht₁ ht₂ => by simp [compl_inter, hunion _ ht₁ _ ht₂] + +@[simp] +lemma mem_comk {p : Set α → Prop} {he hmono hunion s} : + s ∈ comk p he hmono hunion ↔ p sᶜ := + .rfl + +/-- The principal filter of `s` is the collection of all supersets of `s`. -/ +def principal (s : Set α) : Filter α where + sets := { t | s ⊆ t } + univ_sets := subset_univ s + sets_of_superset hx := Subset.trans hx + inter_sets := subset_inter + +@[inherit_doc] +scoped notation "𝓟" => Filter.principal + +@[simp] theorem mem_principal : s ∈ 𝓟 t ↔ t ⊆ s := Iff.rfl + +/-- The *kernel* of a filter is the intersection of all its sets. -/ +def ker (f : Filter α) : Set α := ⋂₀ f.sets + +/-- The join of a filter of filters is defined by the relation `s ∈ join f ↔ {t | s ∈ t} ∈ f`. -/ +def join (f : Filter (Filter α)) : Filter α where + sets := { s | { t : Filter α | s ∈ t } ∈ f } + univ_sets := by simp only [mem_setOf_eq, univ_mem, setOf_true] + sets_of_superset hx xy := mem_of_superset hx fun f h => mem_of_superset h xy + inter_sets hx hy := mem_of_superset (inter_mem hx hy) fun f ⟨h₁, h₂⟩ => inter_mem h₁ h₂ + +@[simp] +theorem mem_join {s : Set α} {f : Filter (Filter α)} : s ∈ join f ↔ { t | s ∈ t } ∈ f := + Iff.rfl + +instance : PartialOrder (Filter α) where + le f g := ∀ ⦃U : Set α⦄, U ∈ g → U ∈ f + le_antisymm a b h₁ h₂ := filter_eq <| Subset.antisymm h₂ h₁ + le_refl a := Subset.rfl + le_trans a b c h₁ h₂ := Subset.trans h₂ h₁ + +theorem le_def : f ≤ g ↔ ∀ x ∈ g, x ∈ f := + Iff.rfl + +instance : Top (Filter α) := + ⟨{ sets := { s | ∀ x, x ∈ s } + univ_sets := fun x => mem_univ x + sets_of_superset := fun hx hxy a => hxy (hx a) + inter_sets := fun hx hy _ => mem_inter (hx _) (hy _) }⟩ + +theorem mem_top_iff_forall {s : Set α} : s ∈ (⊤ : Filter α) ↔ ∀ x, x ∈ s := + Iff.rfl + +@[simp] +theorem mem_top {s : Set α} : s ∈ (⊤ : Filter α) ↔ s = univ := by + rw [mem_top_iff_forall, eq_univ_iff_forall] + +instance : Bot (Filter α) where + bot := + { sets := univ + univ_sets := trivial + sets_of_superset := fun _ _ ↦ trivial + inter_sets := fun _ _ ↦ trivial } + +@[simp] +theorem mem_bot {s : Set α} : s ∈ (⊥ : Filter α) := + trivial + +/-- The infimum of filters is the filter generated by intersections + of elements of the two filters. -/ +instance : Inf (Filter α) := + ⟨fun f g : Filter α => + { sets := { s | ∃ a ∈ f, ∃ b ∈ g, s = a ∩ b } + univ_sets := ⟨_, univ_mem, _, univ_mem, by simp⟩ + sets_of_superset := by + rintro x y ⟨a, ha, b, hb, rfl⟩ xy + refine ⟨a ∪ y, mem_of_superset ha subset_union_left, b ∪ y, + mem_of_superset hb subset_union_left, ?_⟩ + rw [← inter_union_distrib_right, union_eq_self_of_subset_left xy] + inter_sets := by + rintro x y ⟨a, ha, b, hb, rfl⟩ ⟨c, hc, d, hd, rfl⟩ + refine ⟨a ∩ c, inter_mem ha hc, b ∩ d, inter_mem hb hd, ?_⟩ + ac_rfl }⟩ + +/-- The supremum of two filters is the filter that contains sets that belong to both filters. -/ +instance : Sup (Filter α) where + sup f g := + { sets := {s | s ∈ f ∧ s ∈ g} + univ_sets := ⟨univ_mem, univ_mem⟩ + sets_of_superset := fun h₁ h₂ ↦ ⟨mem_of_superset h₁.1 h₂, mem_of_superset h₁.2 h₂⟩ + inter_sets := fun h₁ h₂ ↦ ⟨inter_mem h₁.1 h₂.1, inter_mem h₁.2 h₂.2⟩ } + +/-- A filter is `NeBot` if it is not equal to `⊥`, or equivalently the empty set does not belong to +the filter. Bourbaki include this assumption in the definition of a filter but we prefer to have a +`CompleteLattice` structure on `Filter _`, so we use a typeclass argument in lemmas instead. -/ +class NeBot (f : Filter α) : Prop where + /-- The filter is nontrivial: `f ≠ ⊥` or equivalently, `∅ ∉ f`. -/ + ne' : f ≠ ⊥ + +theorem neBot_iff {f : Filter α} : NeBot f ↔ f ≠ ⊥ := + ⟨fun h => h.1, fun h => ⟨h⟩⟩ + +/-- `f.Eventually p` or `∀ᶠ x in f, p x` mean that `{x | p x} ∈ f`. E.g., `∀ᶠ x in atTop, p x` +means that `p` holds true for sufficiently large `x`. -/ +protected def Eventually (p : α → Prop) (f : Filter α) : Prop := + { x | p x } ∈ f + +@[inherit_doc Filter.Eventually] +notation3 "∀ᶠ "(...)" in "f", "r:(scoped p => Filter.Eventually p f) => r + +/-- `f.Frequently p` or `∃ᶠ x in f, p x` mean that `{x | ¬p x} ∉ f`. E.g., `∃ᶠ x in atTop, p x` +means that there exist arbitrarily large `x` for which `p` holds true. -/ +protected def Frequently (p : α → Prop) (f : Filter α) : Prop := + ¬∀ᶠ x in f, ¬p x + +@[inherit_doc Filter.Frequently] +notation3 "∃ᶠ "(...)" in "f", "r:(scoped p => Filter.Frequently p f) => r + +/-- Two functions `f` and `g` are *eventually equal* along a filter `l` if the set of `x` such that +`f x = g x` belongs to `l`. -/ +def EventuallyEq (l : Filter α) (f g : α → β) : Prop := + ∀ᶠ x in l, f x = g x + +@[inherit_doc] +notation:50 f " =ᶠ[" l:50 "] " g:50 => EventuallyEq l f g + +/-- A function `f` is eventually less than or equal to a function `g` at a filter `l`. -/ +def EventuallyLE [LE β] (l : Filter α) (f g : α → β) : Prop := + ∀ᶠ x in l, f x ≤ g x + +@[inherit_doc] +notation:50 f " ≤ᶠ[" l:50 "] " g:50 => EventuallyLE l f g + +/-- The forward map of a filter -/ +def map (m : α → β) (f : Filter α) : Filter β where + sets := preimage m ⁻¹' f.sets + univ_sets := univ_mem + sets_of_superset hs st := mem_of_superset hs fun _x hx ↦ st hx + inter_sets hs ht := inter_mem hs ht + +/-- `Filter.Tendsto` is the generic "limit of a function" predicate. + `Tendsto f l₁ l₂` asserts that for every `l₂` neighborhood `a`, + the `f`-preimage of `a` is an `l₁` neighborhood. -/ +def Tendsto (f : α → β) (l₁ : Filter α) (l₂ : Filter β) := + l₁.map f ≤ l₂ + +/-- The inverse map of a filter. A set `s` belongs to `Filter.comap m f` if either of the following +equivalent conditions hold. + +1. There exists a set `t ∈ f` such that `m ⁻¹' t ⊆ s`. This is used as a definition. +2. The set `kernImage m s = {y | ∀ x, m x = y → x ∈ s}` belongs to `f`, see `Filter.mem_comap'`. +3. The set `(m '' sᶜ)ᶜ` belongs to `f`, see `Filter.mem_comap_iff_compl` and +`Filter.compl_mem_comap`. -/ +def comap (m : α → β) (f : Filter β) : Filter α where + sets := { s | ∃ t ∈ f, m ⁻¹' t ⊆ s } + univ_sets := ⟨univ, univ_mem, subset_univ _⟩ + sets_of_superset := fun ⟨a', ha', ma'a⟩ ab => ⟨a', ha', ma'a.trans ab⟩ + inter_sets := fun ⟨a', ha₁, ha₂⟩ ⟨b', hb₁, hb₂⟩ => + ⟨a' ∩ b', inter_mem ha₁ hb₁, inter_subset_inter ha₂ hb₂⟩ + +/-- Product of filters. This is the filter generated by cartesian products +of elements of the component filters. -/ +protected def prod (f : Filter α) (g : Filter β) : Filter (α × β) := + f.comap Prod.fst ⊓ g.comap Prod.snd + +/-- Coproduct of filters. -/ +protected def coprod (f : Filter α) (g : Filter β) : Filter (α × β) := + f.comap Prod.fst ⊔ g.comap Prod.snd + +instance instSProd : SProd (Filter α) (Filter β) (Filter (α × β)) where + sprod := Filter.prod + +theorem prod_eq_inf (f : Filter α) (g : Filter β) : f ×ˢ g = f.comap Prod.fst ⊓ g.comap Prod.snd := + rfl + +/-- The monadic bind operation on filter is defined the usual way in terms of `map` and `join`. + +Unfortunately, this `bind` does not result in the expected applicative. See `Filter.seq` for the +applicative instance. -/ +def bind (f : Filter α) (m : α → Filter β) : Filter β := + join (map m f) + +/-- The applicative sequentiation operation. This is not induced by the bind operation. -/ +def seq (f : Filter (α → β)) (g : Filter α) : Filter β where + sets := { s | ∃ u ∈ f, ∃ t ∈ g, ∀ m ∈ u, ∀ x ∈ t, (m : α → β) x ∈ s } + univ_sets := ⟨univ, univ_mem, univ, univ_mem, fun _ _ _ _ => trivial⟩ + sets_of_superset := fun ⟨t₀, t₁, h₀, h₁, h⟩ hst => + ⟨t₀, t₁, h₀, h₁, fun _ hx _ hy => hst <| h _ hx _ hy⟩ + inter_sets := fun ⟨t₀, ht₀, t₁, ht₁, ht⟩ ⟨u₀, hu₀, u₁, hu₁, hu⟩ => + ⟨t₀ ∩ u₀, inter_mem ht₀ hu₀, t₁ ∩ u₁, inter_mem ht₁ hu₁, fun _ ⟨hx₀, hx₁⟩ _ ⟨hy₀, hy₁⟩ => + ⟨ht _ hx₀ _ hy₀, hu _ hx₁ _ hy₁⟩⟩ + +/-- This filter is characterized by `Filter.eventually_curry_iff`: +`(∀ᶠ (x : α × β) in f.curry g, p x) ↔ ∀ᶠ (x : α) in f, ∀ᶠ (y : β) in g, p (x, y)`. Useful +in adding quantifiers to the middle of `Tendsto`s. See +`hasFDerivAt_of_tendstoUniformlyOnFilter`. -/ +def curry (f : Filter α) (g : Filter β) : Filter (α × β) := + bind f fun a ↦ map (a, ·) g + +/-- `pure x` is the set of sets that contain `x`. It is equal to `𝓟 {x}` but +with this definition we have `s ∈ pure a` defeq `a ∈ s`. -/ +instance : Pure Filter := + ⟨fun x => + { sets := { s | x ∈ s } + inter_sets := And.intro + sets_of_superset := fun hs hst => hst hs + univ_sets := trivial }⟩ + +instance : Bind Filter := + ⟨@Filter.bind⟩ + +instance : Functor Filter where map := @Filter.map + +end Filter + +namespace Mathlib.Tactic + +open Lean Meta Elab Tactic + +/-- +`filter_upwards [h₁, ⋯, hₙ]` replaces a goal of the form `s ∈ f` and terms +`h₁ : t₁ ∈ f, ⋯, hₙ : tₙ ∈ f` with `∀ x, x ∈ t₁ → ⋯ → x ∈ tₙ → x ∈ s`. +The list is an optional parameter, `[]` being its default value. + +`filter_upwards [h₁, ⋯, hₙ] with a₁ a₂ ⋯ aₖ` is a short form for +`{ filter_upwards [h₁, ⋯, hₙ], intros a₁ a₂ ⋯ aₖ }`. + +`filter_upwards [h₁, ⋯, hₙ] using e` is a short form for +`{ filter_upwards [h1, ⋯, hn], exact e }`. + +Combining both shortcuts is done by writing `filter_upwards [h₁, ⋯, hₙ] with a₁ a₂ ⋯ aₖ using e`. +Note that in this case, the `aᵢ` terms can be used in `e`. +-/ +syntax (name := filterUpwards) "filter_upwards" (" [" term,* "]")? + (" with" (ppSpace colGt term:max)*)? (" using " term)? : tactic + +elab_rules : tactic +| `(tactic| filter_upwards $[[$[$args],*]]? $[with $wth*]? $[using $usingArg]?) => do + let config : ApplyConfig := {newGoals := ApplyNewGoals.nonDependentOnly} + for e in args.getD #[] |>.reverse do + let goal ← getMainGoal + replaceMainGoal <| ← goal.withContext <| runTermElab do + let m ← mkFreshExprMVar none + let lem ← Term.elabTermEnsuringType + (← ``(Filter.mp_mem $e $(← Term.exprToSyntax m))) (← goal.getType) + goal.assign lem + return [m.mvarId!] + liftMetaTactic fun goal => do + goal.apply (← mkConstWithFreshMVarLevels ``Filter.univ_mem') config + evalTactic <|← `(tactic| dsimp (config := {zeta := false}) only [Set.mem_setOf_eq]) + if let some l := wth then + evalTactic <|← `(tactic| intro $[$l]*) + if let some e := usingArg then + evalTactic <|← `(tactic| exact $e) + +end Mathlib.Tactic diff --git a/Mathlib/Order/Filter/Ker.lean b/Mathlib/Order/Filter/Ker.lean index 29f0a9d1d3129..84ab8db20e0cf 100644 --- a/Mathlib/Order/Filter/Ker.lean +++ b/Mathlib/Order/Filter/Ker.lean @@ -21,9 +21,6 @@ namespace Filter variable {ι : Sort*} {α β : Type*} {f g : Filter α} {s : Set α} {a : α} -/-- The *kernel* of a filter is the intersection of all its sets. -/ -def ker (f : Filter α) : Set α := ⋂₀ f.sets - lemma ker_def (f : Filter α) : f.ker = ⋂ s ∈ f, s := sInter_eq_biInter @[simp] lemma mem_ker : a ∈ f.ker ↔ ∀ s ∈ f, a ∈ s := mem_sInter diff --git a/Mathlib/Order/Filter/Lift.lean b/Mathlib/Order/Filter/Lift.lean index d292896c54051..95e8d4708a1be 100644 --- a/Mathlib/Order/Filter/Lift.lean +++ b/Mathlib/Order/Filter/Lift.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Order.Filter.Bases +import Mathlib.Order.Filter.Prod import Mathlib.Order.ConditionallyCompleteLattice.Basic /-! diff --git a/Mathlib/Order/Filter/NAry.lean b/Mathlib/Order/Filter/NAry.lean index 881c3b3ed2461..747639ad50753 100644 --- a/Mathlib/Order/Filter/NAry.lean +++ b/Mathlib/Order/Filter/NAry.lean @@ -29,8 +29,8 @@ open Filter namespace Filter variable {α α' β β' γ γ' δ δ' ε ε' : Type*} {m : α → β → γ} {f f₁ f₂ : Filter α} - {g g₁ g₂ : Filter β} {h h₁ h₂ : Filter γ} {s s₁ s₂ : Set α} {t t₁ t₂ : Set β} {u : Set γ} - {v : Set δ} {a : α} {b : β} {c : γ} + {g g₁ g₂ : Filter β} {h : Filter γ} {s : Set α} {t : Set β} {u : Set γ} + {a : α} {b : β} /-- The image of a binary function `m : α → β → γ` as a function `Filter α → Filter β → Filter γ`. Mathematically this should be thought of as the image of the corresponding function `α × β → γ`. -/ @@ -51,7 +51,7 @@ theorem map_prod_eq_map₂ (m : α → β → γ) (f : Filter α) (g : Filter β theorem map_prod_eq_map₂' (m : α × β → γ) (f : Filter α) (g : Filter β) : Filter.map m (f ×ˢ g) = map₂ (fun a b => m (a, b)) f g := - map_prod_eq_map₂ (curry m) f g + map_prod_eq_map₂ m.curry f g @[simp] theorem map₂_mk_eq_prod (f : Filter α) (g : Filter β) : map₂ Prod.mk f g = f ×ˢ g := by @@ -145,7 +145,7 @@ theorem map₂_map_right (m : α → γ → δ) (n : β → γ) : @[simp] theorem map₂_curry (m : α × β → γ) (f : Filter α) (g : Filter β) : - map₂ (curry m) f g = (f ×ˢ g).map m := + map₂ m.curry f g = (f ×ˢ g).map m := (map_prod_eq_map₂' _ _ _).symm @[simp] diff --git a/Mathlib/Order/Filter/Prod.lean b/Mathlib/Order/Filter/Prod.lean index 4cf3233718cf7..b2af14e42295d 100644 --- a/Mathlib/Order/Filter/Prod.lean +++ b/Mathlib/Order/Filter/Prod.lean @@ -46,17 +46,6 @@ section Prod variable {s : Set α} {t : Set β} {f : Filter α} {g : Filter β} -/-- Product of filters. This is the filter generated by cartesian products -of elements of the component filters. -/ -protected def prod (f : Filter α) (g : Filter β) : Filter (α × β) := - f.comap Prod.fst ⊓ g.comap Prod.snd - -instance instSProd : SProd (Filter α) (Filter β) (Filter (α × β)) where - sprod := Filter.prod - -theorem prod_eq_inf (f : Filter α) (g : Filter β) : f ×ˢ g = f.comap Prod.fst ⊓ g.comap Prod.snd := - rfl - theorem prod_mem_prod (hs : s ∈ f) (ht : t ∈ g) : s ×ˢ t ∈ f ×ˢ g := inter_mem_inf (preimage_mem_comap hs) (preimage_mem_comap ht) @@ -69,6 +58,10 @@ theorem mem_prod_iff {s : Set (α × β)} {f : Filter α} {g : Filter β} : · rintro ⟨t₁, ht₁, t₂, ht₂, h⟩ exact mem_inf_of_inter (preimage_mem_comap ht₁) (preimage_mem_comap ht₂) h +@[simp] +theorem compl_diagonal_mem_prod {l₁ l₂ : Filter α} : (diagonal α)ᶜ ∈ l₁ ×ˢ l₂ ↔ Disjoint l₁ l₂ := by + simp only [mem_prod_iff, Filter.disjoint_iff, prod_subset_compl_diagonal_iff_disjoint] + @[simp] theorem prod_mem_prod_iff [f.NeBot] [g.NeBot] : s ×ˢ t ∈ f ×ˢ g ↔ s ∈ f ∧ t ∈ g := ⟨fun h => @@ -436,10 +429,6 @@ section Coprod variable {f : Filter α} {g : Filter β} -/-- Coproduct of filters. -/ -protected def coprod (f : Filter α) (g : Filter β) : Filter (α × β) := - f.comap Prod.fst ⊔ g.comap Prod.snd - theorem coprod_eq_prod_top_sup_top_prod (f : Filter α) (g : Filter β) : Filter.coprod f g = f ×ˢ ⊤ ⊔ ⊤ ×ˢ g := by rw [prod_top, top_prod] diff --git a/Mathlib/Order/Heyting/Basic.lean b/Mathlib/Order/Heyting/Basic.lean index 913803105f8c0..f12c306e19f10 100644 --- a/Mathlib/Order/Heyting/Basic.lean +++ b/Mathlib/Order/Heyting/Basic.lean @@ -589,7 +589,7 @@ end GeneralizedCoheytingAlgebra section HeytingAlgebra -variable [HeytingAlgebra α] {a b c : α} +variable [HeytingAlgebra α] {a b : α} @[simp] theorem himp_bot (a : α) : a ⇨ ⊥ = aᶜ := @@ -760,7 +760,7 @@ end HeytingAlgebra section CoheytingAlgebra -variable [CoheytingAlgebra α] {a b c : α} +variable [CoheytingAlgebra α] {a b : α} @[simp] theorem top_sdiff' (a : α) : ⊤ \ a = ¬a := diff --git a/Mathlib/Order/Heyting/Hom.lean b/Mathlib/Order/Heyting/Hom.lean index 358a7b6215806..6b0940817f6d3 100644 --- a/Mathlib/Order/Heyting/Hom.lean +++ b/Mathlib/Order/Heyting/Hom.lean @@ -190,7 +190,6 @@ theorem map_compl (a : α) : f aᶜ = (f a)ᶜ := by rw [← himp_bot, ← himp_ @[simp] theorem map_bihimp (a b : α) : f (a ⇔ b) = f a ⇔ f b := by simp_rw [bihimp, map_inf, map_himp] --- TODO: `map_bihimp` end HeytingAlgebra section CoheytingAlgebra diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index f77c690521b89..96b44c3f6a264 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -915,7 +915,7 @@ open Set section LE -variable [LE α] [LE β] [LE γ] +variable [LE α] [LE β] --@[simp] Porting note (#10618): simp can prove it theorem le_iff_le (e : α ≃o β) {x y : α} : e x ≤ e y ↔ x ≤ y := @@ -929,7 +929,7 @@ theorem symm_apply_le (e : α ≃o β) {x : α} {y : β} : e.symm y ≤ x ↔ y end LE -variable [Preorder α] [Preorder β] [Preorder γ] +variable [Preorder α] [Preorder β] protected theorem monotone (e : α ≃o β) : Monotone e := e.toOrderEmbedding.monotone @@ -1053,7 +1053,7 @@ end Equiv namespace StrictMono variable [LinearOrder α] [Preorder β] -variable (f : α → β) (h_mono : StrictMono f) (h_surj : Function.Surjective f) +variable (f : α → β) (h_mono : StrictMono f) /-- A strictly monotone function with a right inverse is an order isomorphism. -/ @[simps (config := .asFn)] @@ -1172,6 +1172,13 @@ theorem coe_toDualTopEquiv_eq [LE α] : (WithBot.toDualTopEquiv : WithBot αᵒᵈ → (WithTop α)ᵒᵈ) = toDual ∘ WithBot.ofDual := funext fun _ => rfl +/-- The coercion `α → WithBot α` bundled as monotone map. -/ +@[simps] +def coeOrderHom {α : Type*} [Preorder α] : α ↪o WithBot α where + toFun := (↑) + inj' := WithBot.coe_injective + map_rel_iff' := WithBot.coe_le_coe + end WithBot namespace WithTop @@ -1203,6 +1210,13 @@ theorem coe_toDualBotEquiv [LE α] : (WithTop.toDualBotEquiv : WithTop αᵒᵈ → (WithBot α)ᵒᵈ) = toDual ∘ WithTop.ofDual := funext fun _ => rfl +/-- The coercion `α → WithTop α` bundled as monotone map. -/ +@[simps] +def coeOrderHom {α : Type*} [Preorder α] : α ↪o WithTop α where + toFun := (↑) + inj' := WithTop.coe_injective + map_rel_iff' := WithTop.coe_le_coe + end WithTop namespace OrderIso diff --git a/Mathlib/Order/Hom/Set.lean b/Mathlib/Order/Hom/Set.lean index 96809de232fb1..382a2354e4f5d 100644 --- a/Mathlib/Order/Hom/Set.lean +++ b/Mathlib/Order/Hom/Set.lean @@ -16,13 +16,13 @@ import Mathlib.Order.WellFounded open OrderDual -variable {F α β γ δ : Type*} +variable {α β : Type*} namespace OrderIso section LE -variable [LE α] [LE β] [LE γ] +variable [LE α] [LE β] theorem range_eq (e : α ≃o β) : Set.range e = Set.univ := e.surjective.range_eq @@ -58,7 +58,7 @@ end LE open Set -variable [Preorder α] [Preorder β] [Preorder γ] +variable [Preorder α] /-- Order isomorphism between two equal sets. -/ def setCongr (s t : Set α) (h : s = t) : diff --git a/Mathlib/Order/InitialSeg.lean b/Mathlib/Order/InitialSeg.lean index 83ab64a3e5d24..f2e7f48acdac1 100644 --- a/Mathlib/Order/InitialSeg.lean +++ b/Mathlib/Order/InitialSeg.lean @@ -158,16 +158,19 @@ theorem antisymm_symm [IsWellOrder α r] [IsWellOrder β s] (f : r ≼i s) (g : RelIso.coe_fn_injective rfl theorem eq_or_principal [IsWellOrder β s] (f : r ≼i s) : - Surjective f ∨ ∃ b, ∀ x, s x b ↔ ∃ y, f y = x := - or_iff_not_imp_right.2 fun h b => - Acc.recOn (IsWellFounded.wf.apply b : Acc s b) fun x _ IH => - not_forall_not.1 fun hn => - h - ⟨x, fun y => - ⟨IH _, fun ⟨a, e⟩ => by - rw [← e] - exact (trichotomous _ _).resolve_right - (not_or_intro (hn a) fun hl => not_exists.2 hn (f.mem_range_of_rel hl))⟩⟩ + Surjective f ∨ ∃ b, ∀ x, x ∈ Set.range f ↔ s x b := by + apply or_iff_not_imp_right.2 + intro h b + push_neg at h + apply IsWellFounded.induction s b + intro x IH + obtain ⟨y, ⟨hy, hs⟩ | ⟨hy, hs⟩⟩ := h x + · obtain (rfl | h) := (trichotomous y x).resolve_left hs + · exact hy + · obtain ⟨z, rfl⟩ := hy + exact f.mem_range_of_rel h + · obtain ⟨z, rfl⟩ := IH y hs + cases hy (Set.mem_range_self z) /-- Restrict the codomain of an initial segment -/ def codRestrict (p : Set β) (f : r ≼i s) (H : ∀ a, f a ∈ p) : r ≼i Subrel s p := @@ -216,8 +219,8 @@ embeddings are called principal segments -/ structure PrincipalSeg {α β : Type*} (r : α → α → Prop) (s : β → β → Prop) extends r ↪r s where /-- The supremum of the principal segment -/ top : β - /-- The image of the order embedding is the set of elements `b` such that `s b top` -/ - down' : ∀ b, s b top ↔ ∃ a, toRelEmbedding a = b + /-- The range of the order embedding is the set of elements `b` such that `s b top` -/ + mem_range_iff_rel' : ∀ b, b ∈ Set.range toRelEmbedding ↔ s b top -- Porting note: deleted `scoped[InitialSeg]` /-- If `r` is a relation on `α` and `s` in a relation on `β`, then `f : r ≺i s` is an order @@ -237,15 +240,22 @@ instance : CoeFun (r ≺i s) fun _ => α → β := theorem coe_fn_mk (f : r ↪r s) (t o) : (@PrincipalSeg.mk _ _ r s f t o : α → β) = f := rfl +theorem mem_range_iff_rel (f : r ≺i s) : ∀ {b : β}, b ∈ Set.range f ↔ s b f.top := + f.mem_range_iff_rel' _ + +@[deprecated mem_range_iff_rel (since := "2024-10-07")] theorem down (f : r ≺i s) : ∀ {b : β}, s b f.top ↔ ∃ a, f a = b := - f.down' _ + f.mem_range_iff_rel.symm theorem lt_top (f : r ≺i s) (a : α) : s (f a) f.top := - f.down.2 ⟨_, rfl⟩ + f.mem_range_iff_rel.1 ⟨_, rfl⟩ + +theorem mem_range_of_rel_top (f : r ≺i s) {b : β} (h : s b f.top) : b ∈ Set.range f := + f.mem_range_iff_rel.2 h theorem mem_range_of_rel [IsTrans β s] (f : r ≺i s) {a : α} {b : β} (h : s b (f a)) : b ∈ Set.range f := - f.down.1 <| _root_.trans h <| f.lt_top _ + f.mem_range_of_rel_top <| _root_.trans h <| f.lt_top _ @[deprecated mem_range_of_rel (since := "2024-09-21")] alias init := mem_range_of_rel @@ -267,8 +277,7 @@ alias init_iff := exists_eq_iff_rel /-- A principal segment is the same as a non-surjective initial segment. -/ noncomputable def _root_.InitialSeg.toPrincipalSeg [IsWellOrder β s] (f : r ≼i s) (hf : ¬ Surjective f) : r ≺i s := - letI H := f.eq_or_principal.resolve_left hf - ⟨f, Classical.choose H, Classical.choose_spec H⟩ + ⟨f, _, Classical.choose_spec (f.eq_or_principal.resolve_left hf)⟩ @[simp] theorem _root_.InitialSeg.toPrincipalSeg_apply [IsWellOrder β s] (f : r ≼i s) @@ -286,8 +295,7 @@ instance (r : α → α → Prop) [IsWellOrder α r] : IsEmpty (r ≺i r) := /-- Composition of a principal segment with an initial segment, as a principal segment -/ def ltLe (f : r ≺i s) (g : s ≼i t) : r ≺i t := ⟨@RelEmbedding.trans _ _ _ r s t f g, g f.top, fun a => by - simp only [g.exists_eq_iff_rel, PrincipalSeg.down, exists_and_left.symm, exists_swap, - RelEmbedding.trans_apply, exists_eq_right', InitialSeg.coe_coe_fn]⟩ + simp [g.exists_eq_iff_rel, ← PrincipalSeg.mem_range_iff_rel, exists_swap, ← exists_and_left]⟩ @[simp] theorem lt_le_apply (f : r ≺i s) (g : s ≼i t) (a : α) : (f.ltLe g) a = g (f a) := @@ -313,7 +321,7 @@ theorem trans_top [IsTrans γ t] (f : r ≺i s) (g : s ≺i t) : (f.trans g).top /-- Composition of an order isomorphism with a principal segment, as a principal segment -/ def equivLT (f : r ≃r s) (g : s ≺i t) : r ≺i t := ⟨@RelEmbedding.trans _ _ _ r s t f g, g.top, fun c => - suffices (∃ a : β, g a = c) ↔ ∃ a : α, g (f a) = c by simpa [PrincipalSeg.down] + suffices (∃ a, g a = c) ↔ ∃ a, g (f a) = c by simp [← PrincipalSeg.mem_range_iff_rel] ⟨fun ⟨b, h⟩ => ⟨f.symm b, by simp only [h, RelIso.apply_symm_apply]⟩, fun ⟨a, h⟩ => ⟨f a, h⟩⟩⟩ @@ -322,8 +330,8 @@ def ltEquiv {r : α → α → Prop} {s : β → β → Prop} {t : γ → γ → (g : s ≃r t) : PrincipalSeg r t := ⟨@RelEmbedding.trans _ _ _ r s t f g, g f.top, by intro x - rw [← g.apply_symm_apply x, g.map_rel_iff, f.down', exists_congr] - intro y; exact ⟨congr_arg g, fun h => g.toEquiv.bijective.1 h⟩⟩ + rw [← g.apply_symm_apply x, g.map_rel_iff, ← f.mem_range_iff_rel] + exact exists_congr <| fun _ ↦ ⟨fun h => g.toEquiv.bijective.1 h, congr_arg g⟩⟩ @[simp] theorem equivLT_apply (f : r ≃r s) (g : s ≺i t) (a : α) : (equivLT f g) a = g (f a) := @@ -341,7 +349,7 @@ instance [IsWellOrder β s] : Subsingleton (r ≺i s) := rw [@Subsingleton.elim _ _ (f : r ≼i s) g] have et : f.top = g.top := by refine extensional_of_trichotomous_of_irrefl s fun x => ?_ - simp only [PrincipalSeg.down, ef] + simp only [← PrincipalSeg.mem_range_iff_rel, ef] cases f cases g have := RelEmbedding.coe_fn_injective ef; congr ⟩ @@ -356,7 +364,7 @@ theorem topLTTop {r : α → α → Prop} {s : β → β → Prop} {t : γ → /-- Any element of a well order yields a principal segment -/ def ofElement {α : Type*} (r : α → α → Prop) (a : α) : Subrel r { b | r b a } ≺i r := - ⟨Subrel.relEmbedding _ _, a, fun _ => ⟨fun h => ⟨⟨_, h⟩, rfl⟩, fun ⟨⟨_, h⟩, rfl⟩ => h⟩⟩ + ⟨Subrel.relEmbedding _ _, a, fun _ => ⟨fun ⟨⟨_, h⟩, rfl⟩ => h, fun h => ⟨⟨_, h⟩, rfl⟩⟩⟩ -- This lemma was always bad, but the linter only noticed after lean4#2644 @[simp, nolint simpNF] @@ -372,7 +380,7 @@ theorem ofElement_top {α : Type*} (r : α → α → Prop) (a : α) : (ofElemen noncomputable def subrelIso (f : r ≺i s) : Subrel s {b | s b f.top} ≃r r := RelIso.symm { toEquiv := ((Equiv.ofInjective f f.injective).trans (Equiv.setCongr - (funext fun _ ↦ propext f.down.symm))), + (funext fun _ ↦ propext f.mem_range_iff_rel))), map_rel_iff' := f.map_rel_iff } -- This lemma was always bad, but the linter only noticed after lean4#2644 @@ -386,15 +394,12 @@ theorem apply_subrelIso (f : r ≺i s) (b : {b | s b f.top}) : -- This lemma was always bad, but the linter only noticed after lean4#2644 @[simp, nolint simpNF] -theorem subrelIso_apply (f : r ≺i s) (a : α) : - f.subrelIso ⟨f a, f.down.mpr ⟨a, rfl⟩⟩ = a := +theorem subrelIso_apply (f : r ≺i s) (a : α) : f.subrelIso ⟨f a, f.lt_top a⟩ = a := Equiv.ofInjective_symm_apply f.injective _ /-- Restrict the codomain of a principal segment -/ def codRestrict (p : Set β) (f : r ≺i s) (H : ∀ a, f a ∈ p) (H₂ : f.top ∈ p) : r ≺i Subrel s p := - ⟨RelEmbedding.codRestrict p f H, ⟨f.top, H₂⟩, fun ⟨_, _⟩ => - f.down.trans <| - exists_congr fun a => show (⟨f a, H a⟩ : p).1 = _ ↔ _ from ⟨Subtype.eq, congr_arg _⟩⟩ + ⟨RelEmbedding.codRestrict p f H, ⟨f.top, H₂⟩, fun ⟨_, _⟩ => by simp [← f.mem_range_iff_rel]⟩ @[simp] theorem codRestrict_apply (p) (f : r ≺i s) (H H₂ a) : codRestrict p f H H₂ a = ⟨f a, H a⟩ := @@ -408,7 +413,7 @@ theorem codRestrict_top (p) (f : r ≺i s) (H H₂) : (codRestrict p f H H₂).t def ofIsEmpty (r : α → α → Prop) [IsEmpty α] {b : β} (H : ∀ b', ¬s b' b) : r ≺i s := { RelEmbedding.ofIsEmpty r s with top := b - down' := by simp [H] } + mem_range_iff_rel' := by simp [H] } @[simp] theorem ofIsEmpty_top (r : α → α → Prop) [IsEmpty α] {b : β} (H : ∀ b', ¬s b' b) : @@ -435,7 +440,7 @@ theorem wellFounded_iff_wellFounded_subrel {β : Type*} {s : β → β → Prop} ⟨fun wf b => ⟨fun b' => ((PrincipalSeg.ofElement _ b).acc b').mpr (wf.apply b')⟩, fun wf => ⟨fun b => Acc.intro _ fun b' hb' => ?_⟩⟩ let f := PrincipalSeg.ofElement s b - obtain ⟨b', rfl⟩ := f.down.mp ((PrincipalSeg.ofElement_top s b).symm ▸ hb' : s b' f.top) + obtain ⟨b', rfl⟩ := f.mem_range_of_rel_top ((PrincipalSeg.ofElement_top s b).symm ▸ hb') exact (f.acc b').mp ((wf b).apply b') theorem wellFounded_iff_principalSeg.{u} {β : Type u} {s : β → β → Prop} [IsTrans β s] : diff --git a/Mathlib/Order/Interval/Basic.lean b/Mathlib/Order/Interval/Basic.lean index a8d64a2279fcd..b2c3ac1ca1c31 100644 --- a/Mathlib/Order/Interval/Basic.lean +++ b/Mathlib/Order/Interval/Basic.lean @@ -22,7 +22,7 @@ interval arithmetic. open Function OrderDual Set -variable {α β γ δ : Type*} {ι : Sort*} {κ : ι → Sort*} +variable {α β γ : Type*} {ι : Sort*} {κ : ι → Sort*} /-- The nonempty closed intervals in an order. @@ -91,8 +91,7 @@ end LE section Preorder -variable [Preorder α] [Preorder β] [Preorder γ] [Preorder δ] {s : NonemptyInterval α} {x : α × α} - {a : α} +variable [Preorder α] [Preorder β] [Preorder γ] {s : NonemptyInterval α} {x : α × α} {a : α} instance : Preorder (NonemptyInterval α) := Preorder.lift toDualProd @@ -189,7 +188,7 @@ end Preorder section PartialOrder -variable [PartialOrder α] [PartialOrder β] {s t : NonemptyInterval α} {x : α × α} {a b : α} +variable [PartialOrder α] [PartialOrder β] {s t : NonemptyInterval α} {a b : α} instance : PartialOrder (NonemptyInterval α) := PartialOrder.lift _ toDualProd_injective @@ -272,7 +271,7 @@ namespace Interval section LE -variable [LE α] {s t : Interval α} +variable [LE α] -- Porting note: previously found using `deriving` instance : Inhabited (Interval α) := WithBot.inhabited diff --git a/Mathlib/Order/Interval/Finset/Defs.lean b/Mathlib/Order/Interval/Finset/Defs.lean index 0a306c3345be5..d81c2dec0066c 100644 --- a/Mathlib/Order/Interval/Finset/Defs.lean +++ b/Mathlib/Order/Interval/Finset/Defs.lean @@ -271,23 +271,23 @@ section LocallyFiniteOrder variable [LocallyFiniteOrder α] {a b x : α} -/-- The finset of elements `x` such that `a ≤ x` and `x ≤ b`. Basically `Set.Icc a b` as a finset. --/ +/-- The finset $[a, b]$ of elements `x` such that `a ≤ x` and `x ≤ b`. Basically `Set.Icc a b` as a +finset. -/ def Icc (a b : α) : Finset α := LocallyFiniteOrder.finsetIcc a b -/-- The finset of elements `x` such that `a ≤ x` and `x < b`. Basically `Set.Ico a b` as a finset. --/ +/-- The finset $[a, b)$ of elements `x` such that `a ≤ x` and `x < b`. Basically `Set.Ico a b` as a +finset. -/ def Ico (a b : α) : Finset α := LocallyFiniteOrder.finsetIco a b -/-- The finset of elements `x` such that `a < x` and `x ≤ b`. Basically `Set.Ioc a b` as a finset. --/ +/-- The finset $(a, b]$ of elements `x` such that `a < x` and `x ≤ b`. Basically `Set.Ioc a b` as a +finset. -/ def Ioc (a b : α) : Finset α := LocallyFiniteOrder.finsetIoc a b -/-- The finset of elements `x` such that `a < x` and `x < b`. Basically `Set.Ioo a b` as a finset. --/ +/-- The finset $(a, b)$ of elements `x` such that `a < x` and `x < b`. Basically `Set.Ioo a b` as a +finset. -/ def Ioo (a b : α) : Finset α := LocallyFiniteOrder.finsetIoo a b @@ -329,11 +329,11 @@ section LocallyFiniteOrderTop variable [LocallyFiniteOrderTop α] {a x : α} -/-- The finset of elements `x` such that `a ≤ x`. Basically `Set.Ici a` as a finset. -/ +/-- The finset $[a, ∞)$ of elements `x` such that `a ≤ x`. Basically `Set.Ici a` as a finset. -/ def Ici (a : α) : Finset α := LocallyFiniteOrderTop.finsetIci a -/-- The finset of elements `x` such that `a < x`. Basically `Set.Ioi a` as a finset. -/ +/-- The finset $(a, ∞)$ of elements `x` such that `a < x`. Basically `Set.Ioi a` as a finset. -/ def Ioi (a : α) : Finset α := LocallyFiniteOrderTop.finsetIoi a @@ -359,13 +359,13 @@ section LocallyFiniteOrderBot variable [LocallyFiniteOrderBot α] {a x : α} -/-- The finset of elements `x` such that `a ≤ x`. Basically `Set.Iic a` as a finset. -/ -def Iic (a : α) : Finset α := - LocallyFiniteOrderBot.finsetIic a +/-- The finset $(-∞, b]$ of elements `x` such that `x ≤ b`. Basically `Set.Iic b` as a finset. -/ +def Iic (b : α) : Finset α := + LocallyFiniteOrderBot.finsetIic b -/-- The finset of elements `x` such that `a < x`. Basically `Set.Iio a` as a finset. -/ -def Iio (a : α) : Finset α := - LocallyFiniteOrderBot.finsetIio a +/-- The finset $(-∞, b)$ of elements `x` such that `x < b`. Basically `Set.Iio b` as a finset. -/ +def Iio (b : α) : Finset α := + LocallyFiniteOrderBot.finsetIio b @[simp] theorem mem_Iic : x ∈ Iic a ↔ x ≤ a := diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index d59c44eaa2139..8be56194d6481 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -185,14 +185,10 @@ theorem Ico_image_const_sub_eq_Ico (hac : a ≤ c) : rintro ⟨x, hx, rfl⟩ omega -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem Ico_succ_left_eq_erase_Ico : Ico a.succ b = erase (Ico a b) a := by ext x - rw [Ico_succ_left, mem_erase, mem_Ico, mem_Ioo, ← _root_.and_assoc, ne_comm, - _root_.and_comm (a := a ≠ x), lt_iff_le_and_ne] + rw [Ico_succ_left, mem_erase, mem_Ico, mem_Ioo, ← and_assoc, ne_comm, + and_comm (a := a ≠ x), lt_iff_le_and_ne] theorem mod_injOn_Ico (n a : ℕ) : Set.InjOn (· % a) (Finset.Ico n (n + a)) := by induction' n with n ih diff --git a/Mathlib/Order/Interval/Set/Basic.lean b/Mathlib/Order/Interval/Set/Basic.lean index 66feeca1fd5ef..0be1e96eff231 100644 --- a/Mathlib/Order/Interval/Set/Basic.lean +++ b/Mathlib/Order/Interval/Set/Basic.lean @@ -38,35 +38,35 @@ section Preorder variable [Preorder α] {a a₁ a₂ b b₁ b₂ c x : α} -/-- Left-open right-open interval -/ +/-- `Ioo a b` is the left-open right-open interval $(a, b)$. -/ def Ioo (a b : α) := { x | a < x ∧ x < b } -/-- Left-closed right-open interval -/ +/-- `Ico a b` is the left-closed right-open interval $[a, b)$. -/ def Ico (a b : α) := { x | a ≤ x ∧ x < b } -/-- Left-infinite right-open interval -/ -def Iio (a : α) := - { x | x < a } +/-- `Iio b` is the left-infinite right-open interval $(-∞, b)$. -/ +def Iio (b : α) := + { x | x < b } -/-- Left-closed right-closed interval -/ +/-- `Icc a b` is the left-closed right-closed interval $[a, b]$. -/ def Icc (a b : α) := { x | a ≤ x ∧ x ≤ b } -/-- Left-infinite right-closed interval -/ +/-- `Iic b` is the left-infinite right-closed interval $(-∞, b]$. -/ def Iic (b : α) := { x | x ≤ b } -/-- Left-open right-closed interval -/ +/-- `Ioc a b` is the left-open right-closed interval $(a, b]$. -/ def Ioc (a b : α) := { x | a < x ∧ x ≤ b } -/-- Left-closed right-infinite interval -/ +/-- `Ici a` is the left-closed right-infinite interval $[a, ∞)$. -/ def Ici (a : α) := { x | a ≤ x } -/-- Left-open right-infinite interval -/ +/-- `Ioi a` is the left-open right-infinite interval $(a, ∞)$. -/ def Ioi (a : α) := { x | a < x } diff --git a/Mathlib/Order/Interval/Set/UnorderedInterval.lean b/Mathlib/Order/Interval/Set/UnorderedInterval.lean index 8c14071911f3d..ea241a7fbf3dc 100644 --- a/Mathlib/Order/Interval/Set/UnorderedInterval.lean +++ b/Mathlib/Order/Interval/Set/UnorderedInterval.lean @@ -44,7 +44,7 @@ namespace Set section Lattice -variable [Lattice α] [Lattice β] {a a₁ a₂ b b₁ b₂ c x : α} +variable [Lattice α] [Lattice β] {a a₁ a₂ b b₁ b₂ x : α} /-- `uIcc a b` is the set of elements lying between `a` and `b`, with `a` and `b` included. Note that we define it more generally in a lattice as `Set.Icc (a ⊓ b) (a ⊔ b)`. In a product type, @@ -134,7 +134,7 @@ open Interval section DistribLattice -variable [DistribLattice α] {a a₁ a₂ b b₁ b₂ c x : α} +variable [DistribLattice α] {a b c : α} lemma eq_of_mem_uIcc_of_mem_uIcc (ha : a ∈ [[b, c]]) (hb : b ∈ [[a, c]]) : a = b := eq_of_inf_eq_sup_eq (inf_congr_right ha.1 hb.1) <| sup_congr_right ha.2 hb.2 @@ -155,7 +155,7 @@ section LinearOrder variable [LinearOrder α] section Lattice -variable [Lattice β] {f : α → β} {s : Set α} {a b : α} +variable [Lattice β] {f : α → β} {a b : α} lemma _root_.MonotoneOn.mapsTo_uIcc (hf : MonotoneOn f (uIcc a b)) : MapsTo f (uIcc a b) (uIcc (f a) (f b)) := by @@ -187,7 +187,7 @@ lemma _root_.Antitone.image_uIcc_subset (hf : Antitone f) : f '' uIcc a b ⊆ uI end Lattice -variable [LinearOrder β] {f : α → β} {s : Set α} {a a₁ a₂ b b₁ b₂ c d x : α} +variable [LinearOrder β] {f : α → β} {s : Set α} {a a₁ a₂ b b₁ b₂ c : α} theorem Icc_min_max : Icc (min a b) (max a b) = [[a, b]] := rfl diff --git a/Mathlib/Order/Irreducible.lean b/Mathlib/Order/Irreducible.lean index ecc3744d96ae4..12ebc35f11d80 100644 --- a/Mathlib/Order/Irreducible.lean +++ b/Mathlib/Order/Irreducible.lean @@ -251,7 +251,7 @@ end SemilatticeInf section DistribLattice -variable [DistribLattice α] {a b c : α} +variable [DistribLattice α] {a : α} @[simp] theorem supPrime_iff_supIrred : SupPrime a ↔ SupIrred a := diff --git a/Mathlib/Order/Lattice.lean b/Mathlib/Order/Lattice.lean index 76070c79b4151..8383d090d9c12 100644 --- a/Mathlib/Order/Lattice.lean +++ b/Mathlib/Order/Lattice.lean @@ -528,7 +528,7 @@ def Lattice.mk' {α : Type*} [Sup α] [Inf α] (sup_comm : ∀ a b : α, a ⊔ b section Lattice -variable [Lattice α] {a b c d : α} +variable [Lattice α] {a b c : α} theorem inf_le_sup : a ⊓ b ≤ a ⊔ b := inf_le_left.trans le_sup_left diff --git a/Mathlib/Order/Max.lean b/Mathlib/Order/Max.lean index 61f5b1037d505..36518a3cb4aa5 100644 --- a/Mathlib/Order/Max.lean +++ b/Mathlib/Order/Max.lean @@ -158,7 +158,7 @@ theorem NoMaxOrder.not_acc [LT α] [NoMaxOrder α] (a : α) : ¬Acc (· > ·) a section LE -variable [LE α] {a b : α} +variable [LE α] {a : α} /-- `a : α` is a bottom element of `α` if it is less than or equal to any other element of `α`. This predicate is roughly an unbundled version of `OrderBot`, except that a preorder may have diff --git a/Mathlib/Order/ModularLattice.lean b/Mathlib/Order/ModularLattice.lean index 6dfc3e2db07ba..7b2b7f77ae7f8 100644 --- a/Mathlib/Order/ModularLattice.lean +++ b/Mathlib/Order/ModularLattice.lean @@ -218,6 +218,10 @@ theorem sup_lt_sup_of_lt_of_inf_le_inf (hxy : x < y) (hinf : y ⊓ z ≤ x ⊓ z theorem inf_lt_inf_of_lt_of_sup_le_sup (hxy : x < y) (hinf : y ⊔ z ≤ x ⊔ z) : x ⊓ z < y ⊓ z := sup_lt_sup_of_lt_of_inf_le_inf (α := αᵒᵈ) hxy hinf +theorem strictMono_inf_prod_sup : StrictMono fun x ↦ (x ⊓ z, x ⊔ z) := fun _x _y hxy ↦ + ⟨⟨inf_le_inf_right _ hxy.le, sup_le_sup_right hxy.le _⟩, + fun ⟨inf_le, sup_le⟩ ↦ (sup_lt_sup_of_lt_of_inf_le_inf hxy inf_le).not_le sup_le⟩ + /-- A generalization of the theorem that if `N` is a submodule of `M` and `N` and `M / N` are both Artinian, then `M` is Artinian. -/ theorem wellFounded_lt_exact_sequence {β γ : Type*} [PartialOrder β] [Preorder γ] @@ -225,16 +229,9 @@ theorem wellFounded_lt_exact_sequence {β γ : Type*} [PartialOrder β] [Preorde (f₁ : β → α) (f₂ : α → β) (g₁ : γ → α) (g₂ : α → γ) (gci : GaloisCoinsertion f₁ f₂) (gi : GaloisInsertion g₂ g₁) (hf : ∀ a, f₁ (f₂ a) = a ⊓ K) (hg : ∀ a, g₁ (g₂ a) = a ⊔ K) : WellFoundedLT α := - ⟨Subrelation.wf - (@fun A B hAB => - show Prod.Lex (· < ·) (· < ·) (f₂ A, g₂ A) (f₂ B, g₂ B) by - simp only [Prod.lex_def, lt_iff_le_not_le, ← gci.l_le_l_iff, ← gi.u_le_u_iff, hf, hg, - le_antisymm_iff] - simp only [gci.l_le_l_iff, gi.u_le_u_iff, ← lt_iff_le_not_le, ← le_antisymm_iff] - rcases lt_or_eq_of_le (inf_le_inf_right K (le_of_lt hAB)) with h | h - · exact Or.inl h - · exact Or.inr ⟨h, sup_lt_sup_of_lt_of_inf_le_inf hAB (le_of_eq h.symm)⟩) - (InvImage.wf _ (h₁.wf.prod_lex h₂.wf))⟩ + StrictMono.wellFoundedLT (f := fun A ↦ (f₂ A, g₂ A)) fun A B hAB ↦ by + simp only [Prod.le_def, lt_iff_le_not_le, ← gci.l_le_l_iff, ← gi.u_le_u_iff, hf, hg] + exact strictMono_inf_prod_sup hAB /-- A generalization of the theorem that if `N` is a submodule of `M` and `N` and `M / N` are both Noetherian, then `M` is Noetherian. -/ @@ -261,7 +258,7 @@ def infIccOrderIsoIccSup (a b : α) : Set.Icc (a ⊓ b) a ≃o Set.Icc b (a ⊔ (by change a ⊓ ↑x ⊔ b = ↑x rw [inf_comm, inf_sup_assoc_of_le _ x.prop.1, inf_eq_left.2 x.prop.2]) - map_rel_iff' := @fun x y => by + map_rel_iff' {x y} := by simp only [Subtype.mk_le_mk, Equiv.coe_fn_mk, le_sup_right] rw [← Subtype.coe_le_coe] refine ⟨fun h => ?_, fun h => sup_le_sup_right h _⟩ diff --git a/Mathlib/Order/Monotone/Basic.lean b/Mathlib/Order/Monotone/Basic.lean index 3292cce1aeeec..bafac4b6ac367 100644 --- a/Mathlib/Order/Monotone/Basic.lean +++ b/Mathlib/Order/Monotone/Basic.lean @@ -66,7 +66,6 @@ open Function OrderDual universe u v w variable {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {δ : Type*} {π : ι → Type*} - {r : α → α → Prop} section MonotoneDef @@ -1025,7 +1024,7 @@ theorem Subtype.strictMono_coe [Preorder α] (t : Set α) : section Preorder -variable [Preorder α] [Preorder β] [Preorder γ] [Preorder δ] {f : α → γ} {g : β → δ} {a b : α} +variable [Preorder α] [Preorder β] [Preorder γ] [Preorder δ] {f : α → γ} {g : β → δ} theorem monotone_fst : Monotone (@Prod.fst α β) := fun _ _ ↦ And.left @@ -1074,7 +1073,7 @@ theorem const_strictMono [Nonempty β] : StrictMono (const β : α → β → α end Function section apply -variable {ι α : Type*} {β : ι → Type*} [∀ i, Preorder (β i)] [Preorder α] {f : α → ∀ i, β i} +variable {β : ι → Type*} [∀ i, Preorder (β i)] [Preorder α] {f : α → ∀ i, β i} lemma monotone_iff_apply₂ : Monotone f ↔ ∀ i, Monotone (f · i) := by simp [Monotone, Pi.le_def, @forall_swap ι] diff --git a/Mathlib/Order/Monotone/Extension.lean b/Mathlib/Order/Monotone/Extension.lean index 7d56c6bba761f..6576bde8e49da 100644 --- a/Mathlib/Order/Monotone/Extension.lean +++ b/Mathlib/Order/Monotone/Extension.lean @@ -16,7 +16,6 @@ monotone extension to the whole space. open Set variable {α β : Type*} [LinearOrder α] [ConditionallyCompleteLinearOrder β] {f : α → β} {s : Set α} - {a b : α} /-- If a function is monotone and is bounded on a set `s`, then it admits a monotone extension to the whole space. -/ diff --git a/Mathlib/Order/Monotone/Monovary.lean b/Mathlib/Order/Monotone/Monovary.lean index 4c25c914b2698..9dcfecb90a91f 100644 --- a/Mathlib/Order/Monotone/Monovary.lean +++ b/Mathlib/Order/Monotone/Monovary.lean @@ -31,7 +31,7 @@ variable {ι ι' α β γ : Type*} section Preorder -variable [Preorder α] [Preorder β] [Preorder γ] {f : ι → α} {f' : α → γ} {g : ι → β} {g' : β → γ} +variable [Preorder α] [Preorder β] [Preorder γ] {f : ι → α} {f' : α → γ} {g : ι → β} {s t : Set ι} /-- `f` monovaries with `g` if `g i < g j` implies `f i ≤ f j`. -/ @@ -282,7 +282,7 @@ end Preorder section LinearOrder -variable [Preorder α] [LinearOrder β] [Preorder γ] {f : ι → α} {f' : α → γ} {g : ι → β} {g' : β → γ} +variable [Preorder α] [LinearOrder β] [Preorder γ] {f : ι → α} {g : ι → β} {g' : β → γ} {s : Set ι} theorem MonovaryOn.comp_monotoneOn_right (h : MonovaryOn f g s) (hg : MonotoneOn g' (g '' s)) : diff --git a/Mathlib/Order/OrdContinuous.lean b/Mathlib/Order/OrdContinuous.lean index 1af016fa24c89..ac0e1eab6eac1 100644 --- a/Mathlib/Order/OrdContinuous.lean +++ b/Mathlib/Order/OrdContinuous.lean @@ -231,7 +231,7 @@ namespace OrderIso section Preorder -variable [Preorder α] [Preorder β] (e : α ≃o β) {s : Set α} {x : α} +variable [Preorder α] [Preorder β] (e : α ≃o β) protected theorem leftOrdContinuous : LeftOrdContinuous e := fun _ _ hx => ⟨Monotone.mem_upperBounds_image (fun _ _ => e.map_rel_iff.2) hx.1, fun _ hy => diff --git a/Mathlib/Order/OrderIsoNat.lean b/Mathlib/Order/OrderIsoNat.lean index 5eddb07567f64..da0fabb452ad3 100644 --- a/Mathlib/Order/OrderIsoNat.lean +++ b/Mathlib/Order/OrderIsoNat.lean @@ -232,3 +232,20 @@ theorem WellFounded.iSup_eq_monotonicSequenceLimit [CompleteLattice α] · cases' WellFounded.monotone_chain_condition'.1 h a with n hn have : n ∈ {n | ∀ m, n ≤ m → a n = a m} := fun k hk => (a.mono hk).eq_of_not_lt (hn k hk) exact (Nat.sInf_mem ⟨n, this⟩ m hm.le).ge + +theorem exists_covBy_seq_of_wellFoundedLT_wellFoundedGT (α) [Preorder α] + [Nonempty α] [wfl : WellFoundedLT α] [wfg : WellFoundedGT α] : + ∃ a : ℕ → α, IsMin (a 0) ∧ ∃ n, IsMax (a n) ∧ ∀ i < n, a i ⋖ a (i + 1) := by + choose next hnext using exists_covBy_of_wellFoundedLT (α := α) + have hα := Set.nonempty_iff_univ_nonempty.mp ‹_› + classical + let a : ℕ → α := Nat.rec (wfl.wf.min _ hα) fun _n a ↦ if ha : IsMax a then a else next ha + refine ⟨a, isMin_iff_forall_not_lt.mpr fun _ ↦ wfl.wf.not_lt_min _ hα trivial, ?_⟩ + have cov n (hn : ¬ IsMax (a n)) : a n ⋖ a (n + 1) := by + change a n ⋖ if ha : IsMax (a n) then a n else _ + rw [dif_neg hn] + exact hnext hn + have H : ∃ n, IsMax (a n) := by + by_contra! + exact (RelEmbedding.natGT a fun n ↦ (cov n (this n)).1).not_wellFounded_of_decreasing_seq wfg.wf + exact ⟨_, wellFounded_lt.min_mem _ H, fun i h ↦ cov _ fun h' ↦ wellFounded_lt.not_lt_min _ H h' h⟩ diff --git a/Mathlib/Order/RelIso/Basic.lean b/Mathlib/Order/RelIso/Basic.lean index e801c4d599937..4c5eb975fb749 100644 --- a/Mathlib/Order/RelIso/Basic.lean +++ b/Mathlib/Order/RelIso/Basic.lean @@ -201,7 +201,7 @@ instance : Coe (r ↪r s) (r →r s) := -- TODO: define and instantiate a `RelEmbeddingClass` when `EmbeddingLike` is defined instance : FunLike (r ↪r s) α β where - coe := fun x => x.toFun + coe x := x.toFun coe_injective' f g h := by rcases f with ⟨⟨⟩⟩ rcases g with ⟨⟨⟩⟩ @@ -536,7 +536,7 @@ instance : CoeOut (r ≃r s) (r ↪r s) := -- TODO: define and instantiate a `RelIsoClass` when `EquivLike` is defined instance : FunLike (r ≃r s) α β where - coe := fun x => x + coe x := x coe_injective' := Equiv.coe_fn_injective.comp toEquiv_injective -- TODO: define and instantiate a `RelIsoClass` when `EquivLike` is defined diff --git a/Mathlib/Order/RelIso/Set.lean b/Mathlib/Order/RelIso/Set.lean index fbf5dd9ac208e..961b9773f5912 100644 --- a/Mathlib/Order/RelIso/Set.lean +++ b/Mathlib/Order/RelIso/Set.lean @@ -19,8 +19,7 @@ open Function universe u v w -variable {α β γ δ : Type*} {r : α → α → Prop} {s : β → β → Prop} {t : γ → γ → Prop} - {u : δ → δ → Prop} +variable {α β : Type*} {r : α → α → Prop} {s : β → β → Prop} namespace RelHomClass @@ -92,8 +91,6 @@ theorem RelEmbedding.codRestrict_apply (p) (f : r ↪r s) (H a) : section image -variable {α β : Type*} {r : α → α → Prop} {s : β → β → Prop} - theorem RelIso.image_eq_preimage_symm (e : r ≃r s) (t : Set α) : e '' t = e.symm ⁻¹' t := e.toEquiv.image_eq_preimage t diff --git a/Mathlib/Order/SuccPred/Archimedean.lean b/Mathlib/Order/SuccPred/Archimedean.lean index 6a225c713afdc..0db5ba2b12320 100644 --- a/Mathlib/Order/SuccPred/Archimedean.lean +++ b/Mathlib/Order/SuccPred/Archimedean.lean @@ -234,6 +234,39 @@ lemma SuccOrder.forall_ne_bot_iff simp only [Function.iterate_succ', Function.comp_apply] apply h +section IsLeast + +-- TODO: generalize to PartialOrder and `DirectedOn` after #16272 +lemma BddAbove.exists_isGreatest_of_nonempty {X : Type*} [LinearOrder X] [SuccOrder X] + [IsSuccArchimedean X] {S : Set X} (hS : BddAbove S) (hS' : S.Nonempty) : + ∃ x, IsGreatest S x := by + obtain ⟨m, hm⟩ := hS + obtain ⟨n, hn⟩ := hS' + by_cases hm' : m ∈ S + · exact ⟨_, hm', hm⟩ + have hn' := hm hn + revert hn hm hm' + refine Succ.rec ?_ ?_ hn' + · simp (config := {contextual := true}) + intro m _ IH hm hn hm' + rw [mem_upperBounds] at IH hm + simp_rw [Order.le_succ_iff_eq_or_le] at hm + replace hm : ∀ x ∈ S, x ≤ m := by + intro x hx + refine (hm x hx).resolve_left ?_ + rintro rfl + exact hm' hx + by_cases hmS : m ∈ S + · exact ⟨m, hmS, hm⟩ + · exact IH hm hn hmS + +lemma BddBelow.exists_isLeast_of_nonempty {X : Type*} [LinearOrder X] [PredOrder X] + [IsPredArchimedean X] {S : Set X} (hS : BddBelow S) (hS' : S.Nonempty) : + ∃ x, IsLeast S x := + hS.dual.exists_isGreatest_of_nonempty hS' + +end IsLeast + section OrderIso variable {X Y : Type*} [PartialOrder X] [PartialOrder Y] diff --git a/Mathlib/Order/SuccPred/Basic.lean b/Mathlib/Order/SuccPred/Basic.lean index aef647a6f8d31..43142cc9da386 100644 --- a/Mathlib/Order/SuccPred/Basic.lean +++ b/Mathlib/Order/SuccPred/Basic.lean @@ -24,7 +24,7 @@ order... Maximal elements don't have a sensible successor. Thus the naïve typeclass ```lean -class NaiveSuccOrder (α : Type*) [Preorder α] := +class NaiveSuccOrder (α : Type*) [Preorder α] where (succ : α → α) (succ_le_iff : ∀ {a b}, succ a ≤ b ↔ a < b) (lt_succ_iff : ∀ {a b}, a < succ b ↔ a ≤ b) diff --git a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean index f721b4e776979..e758ae8c0f2f4 100644 --- a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean +++ b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean @@ -185,7 +185,7 @@ end LinearLocallyFiniteOrder section toZ -- Requiring either of `IsSuccArchimedean` or `IsPredArchimedean` is equivalent. -variable [SuccOrder ι] [IsSuccArchimedean ι] [P : PredOrder ι] {i0 i : ι} +variable [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 i : ι} -- For "to_Z" @@ -200,7 +200,7 @@ theorem toZ_of_ge (hi : i0 ≤ i) : toZ i0 i = Nat.find (exists_succ_iterate_of_ dif_pos hi theorem toZ_of_lt (hi : i < i0) : - toZ i0 i = -Nat.find (@exists_pred_iterate_of_le _ _ P _ _ _ hi.le) := + toZ i0 i = -Nat.find (exists_pred_iterate_of_le (α := ι) hi.le) := dif_neg (not_le.mpr hi) @[simp] @@ -311,8 +311,8 @@ theorem toZ_mono {i j : ι} (h_le : i ≤ j) : toZ i0 i ≤ toZ i0 j := by · exact le_of_not_le h · exact absurd h_le (not_le.mpr (hj.trans_le hi)) · exact (toZ_neg hi).le.trans (toZ_nonneg hj) - · let m := Nat.find (@exists_pred_iterate_of_le _ _ P _ _ _ h_le) - have hm : pred^[m] j = i := Nat.find_spec (exists_pred_iterate_of_le h_le) + · let m := Nat.find (exists_pred_iterate_of_le (α := ι) h_le) + have hm : pred^[m] j = i := Nat.find_spec (exists_pred_iterate_of_le (α := ι) h_le) have hj_eq : i = pred^[(-toZ i0 j).toNat + m] i0 := by rw [← hm, add_comm] nth_rw 1 [← iterate_pred_toZ j hj] diff --git a/Mathlib/Order/Synonym.lean b/Mathlib/Order/Synonym.lean index 1ff3b4dca5241..35decae0e1c87 100644 --- a/Mathlib/Order/Synonym.lean +++ b/Mathlib/Order/Synonym.lean @@ -34,7 +34,7 @@ This file is similar to `Algebra.Group.TypeTags`. -/ -variable {α β γ : Type*} +variable {α : Type*} /-! ### Order dual -/ @@ -170,6 +170,18 @@ theorem toLex_inj {a b : α} : toLex a = toLex b ↔ a = b := theorem ofLex_inj {a b : Lex α} : ofLex a = ofLex b ↔ a = b := Iff.rfl +instance (α : Type*) [BEq α] : BEq (Lex α) where + beq a b := ofLex a == ofLex b + +instance (α : Type*) [BEq α] [LawfulBEq α] : LawfulBEq (Lex α) := + inferInstanceAs (LawfulBEq α) + +instance (α : Type*) [DecidableEq α] : DecidableEq (Lex α) := + inferInstanceAs (DecidableEq α) + +instance (α : Type*) [Inhabited α] : Inhabited (Lex α) := + inferInstanceAs (Inhabited α) + /-- A recursor for `Lex`. Use as `induction x`. -/ @[elab_as_elim, induction_eliminator, cases_eliminator] protected def Lex.rec {β : Lex α → Sort*} (h : ∀ a, β (toLex a)) : ∀ a, β a := fun a => h (ofLex a) diff --git a/Mathlib/Order/ULift.lean b/Mathlib/Order/ULift.lean index b3fa9d2fcbc4c..05f7ece9dd38b 100644 --- a/Mathlib/Order/ULift.lean +++ b/Mathlib/Order/ULift.lean @@ -14,6 +14,8 @@ the corresponding `Prod` instances. namespace ULift +open Batteries + universe v u variable {α : Type u} @@ -28,6 +30,11 @@ instance [LT α] : LT (ULift.{v} α) where lt x y := x.down < y.down @[simp] theorem up_lt [LT α] {a b : α} : up a < up b ↔ a < b := Iff.rfl @[simp] theorem down_lt [LT α] {a b : ULift α} : down a < down b ↔ a < b := Iff.rfl +instance [BEq α] : BEq (ULift.{v} α) where beq x y := x.down == y.down + +@[simp] theorem up_beq [BEq α] (a b : α) : (up a == up b) = (a == b) := rfl +@[simp] theorem down_beq [BEq α] (a b : ULift α) : (down a == down b) = (a == b) := rfl + instance [Ord α] : Ord (ULift.{v} α) where compare x y := compare x.down y.down @[simp] theorem up_compare [Ord α] (a b : α) : compare (up a) (up b) = compare a b := rfl @@ -54,6 +61,25 @@ instance [HasCompl α] : HasCompl (ULift.{v} α) where compl x := up <| x.down @[simp] theorem up_compl [HasCompl α] (a : α) : up (aᶜ) = (up a)ᶜ := rfl @[simp] theorem down_compl [HasCompl α] (a : ULift α) : down aᶜ = (down a)ᶜ := rfl +instance [Ord α] [inst : OrientedOrd α] : OrientedOrd (ULift.{v} α) where + symm _ _ := inst.symm .. + +instance [Ord α] [inst : TransOrd α] : TransOrd (ULift.{v} α) where + le_trans := inst.le_trans + +instance [BEq α] [Ord α] [inst : BEqOrd α] : BEqOrd (ULift.{v} α) where + cmp_iff_beq := inst.cmp_iff_beq + +instance [LT α] [Ord α] [inst : LTOrd α] : LTOrd (ULift.{v} α) where + cmp_iff_lt := inst.cmp_iff_lt + +instance [LE α] [Ord α] [inst : LEOrd α] : LEOrd (ULift.{v} α) where + cmp_iff_le := inst.cmp_iff_le + +instance [LE α] [LT α] [BEq α] [Ord α] [inst : LawfulOrd α] : LawfulOrd (ULift.{v} α) where + cmp_iff_lt := inst.cmp_iff_lt + cmp_iff_le := inst.cmp_iff_le + instance [Preorder α] : Preorder (ULift.{v} α) := Preorder.lift ULift.down diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index e9c5896a6d456..81f3b53ee346a 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -933,7 +933,7 @@ lemma ge_of_forall_gt_iff_ge [LinearOrder α] [DenselyOrdered α] [NoMinOrder α section LE -variable [LE α] {a b : α} +variable [LE α] theorem toDual_le_iff {a : WithBot α} {b : WithTop αᵒᵈ} : WithBot.toDual a ≤ b ↔ WithTop.ofDual b ≤ a := @@ -963,7 +963,7 @@ end LE section LT -variable [LT α] {a b : α} +variable [LT α] theorem toDual_lt_iff {a : WithBot α} {b : WithTop αᵒᵈ} : WithBot.toDual a < b ↔ WithTop.ofDual b < a := diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 2a16eb0159a62..513325918f94f 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -234,7 +234,7 @@ lemma compProd_zero_left (κ : Kernel (α × β) γ) : · rw [Kernel.compProd_of_not_isSFiniteKernel_right _ _ h] @[simp] -lemma compProd_zero_right (κ : Kernel α β) (γ : Type*) [MeasurableSpace γ] : +lemma compProd_zero_right (κ : Kernel α β) (γ : Type*) {mγ : MeasurableSpace γ} : κ ⊗ₖ (0 : Kernel (α × β) γ) = 0 := by by_cases h : IsSFiniteKernel κ · ext a s hs @@ -566,7 +566,7 @@ section MapComap /-! ### map, comap -/ -variable {γ δ : Type*} [MeasurableSpace γ] {mδ : MeasurableSpace δ} {f : β → γ} {g : γ → α} +variable {γ δ : Type*} {mγ : MeasurableSpace γ} {mδ : MeasurableSpace δ} {f : β → γ} {g : γ → α} /-- The pushforward of a kernel along a measurable function. This is an implementation detail, use `map κ f` instead. -/ @@ -580,7 +580,7 @@ open Classical in If the function is not measurable, we use zero instead. This choice of junk value ensures that typeclass inference can infer that the `map` of a kernel satisfying `IsZeroOrMarkovKernel` again satisfies this property. -/ -noncomputable def map (κ : Kernel α β) (f : β → γ) : Kernel α γ := +noncomputable def map [MeasurableSpace γ] (κ : Kernel α β) (f : β → γ) : Kernel α γ := if hf : Measurable f then mapOfMeasurable κ f hf else 0 theorem map_of_not_measurable (κ : Kernel α β) {f : β → γ} (hf : ¬(Measurable f)) : @@ -802,7 +802,7 @@ lemma map_prodMkLeft (γ : Type*) [MeasurableSpace γ] (κ : Kernel α β) (f : rfl · simp [map_of_not_measurable _ hf] -lemma map_prodMkRight (κ : Kernel α β) (γ : Type*) [MeasurableSpace γ] (f : β → δ) : +lemma map_prodMkRight (κ : Kernel α β) (γ : Type*) {mγ : MeasurableSpace γ} (f : β → δ) : map (prodMkRight γ κ) f = prodMkRight γ (map κ f) := by by_cases hf : Measurable f · simp only [map, hf, ↓reduceDIte] @@ -832,10 +832,10 @@ instance IsFiniteKernel.swapLeft (κ : Kernel (α × β) γ) [IsFiniteKernel κ] instance IsSFiniteKernel.swapLeft (κ : Kernel (α × β) γ) [IsSFiniteKernel κ] : IsSFiniteKernel (swapLeft κ) := by rw [Kernel.swapLeft]; infer_instance -@[simp] lemma swapLeft_prodMkLeft (κ : Kernel α β) (γ : Type*) [MeasurableSpace γ] : +@[simp] lemma swapLeft_prodMkLeft (κ : Kernel α β) (γ : Type*) {_ : MeasurableSpace γ} : swapLeft (prodMkLeft γ κ) = prodMkRight γ κ := rfl -@[simp] lemma swapLeft_prodMkRight (κ : Kernel α β) (γ : Type*) [MeasurableSpace γ] : +@[simp] lemma swapLeft_prodMkRight (κ : Kernel α β) (γ : Type*) {_ : MeasurableSpace γ} : swapLeft (prodMkRight γ κ) = prodMkLeft γ κ := rfl /-- Define a `Kernel α (γ × β)` from a `Kernel α (β × γ)` by taking the map of `Prod.swap`. diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index 5b8e36ae7d2f4..cd5425670e316 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -63,7 +63,8 @@ to `ν` if is measurable, if `fun b ↦ f (a, b) x` is `(ν a)`-integrable for a and for all measurable sets `s : Set β`, `∫ b in s, f (a, b) x ∂(ν a) = (κ a (s ×ˢ Iic x)).toReal`. Also the `ℚ → ℝ` function `f (a, b)` should satisfy the properties of a Sieltjes function for `(ν a)`-almost all `b : β`. -/ -structure IsRatCondKernelCDF (f : α × β → ℚ → ℝ) (κ : Kernel α (β × ℝ)) (ν : Kernel α β) : Prop := +structure IsRatCondKernelCDF (f : α × β → ℚ → ℝ) (κ : Kernel α (β × ℝ)) (ν : Kernel α β) : + Prop where measurable : Measurable f isRatStieltjesPoint_ae (a : α) : ∀ᵐ b ∂(ν a), IsRatStieltjesPoint f (a, b) integrable (a : α) (q : ℚ) : Integrable (fun b ↦ f (a, b) q) (ν a) diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 94bcc41f1f6db..34d9233a9272f 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -396,7 +396,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : set Y := fun n : ℕ => truncation (X n) n set S := fun n => ∑ i ∈ range n, Y i with hS let u : ℕ → ℕ := fun n => ⌊c ^ n⌋₊ - have u_mono : Monotone u := fun i j hij => Nat.floor_mono (pow_le_pow_right c_one.le hij) + have u_mono : Monotone u := fun i j hij => Nat.floor_mono (pow_right_mono₀ c_one.le hij) have I1 : ∀ K, ∑ j ∈ range K, ((j : ℝ) ^ 2)⁻¹ * Var[Y j] ≤ 2 * 𝔼[X 0] := by intro K calc diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index 2b83e2348e8b1..7cd53f1275514 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -202,9 +202,9 @@ theorem dOne_comp_dZero : dOne A ∘ₗ dZero A = 0 := by rfl theorem dTwo_comp_dOne : dTwo A ∘ₗ dOne A = 0 := by - show ModuleCat.ofHom (dOne A) ≫ ModuleCat.ofHom (dTwo A) = _ - have h1 : _ ≫ ModuleCat.ofHom (dOne A) = _ ≫ _ := congr_arg ModuleCat.ofHom (dOne_comp_eq A) - have h2 : _ ≫ ModuleCat.ofHom (dTwo A) = _ ≫ _ := congr_arg ModuleCat.ofHom (dTwo_comp_eq A) + show ModuleCat.asHom (dOne A) ≫ ModuleCat.asHom (dTwo A) = _ + have h1 : _ ≫ ModuleCat.asHom (dOne A) = _ ≫ _ := congr_arg ModuleCat.asHom (dOne_comp_eq A) + have h2 : _ ≫ ModuleCat.asHom (dTwo A) = _ ≫ _ := congr_arg ModuleCat.asHom (dTwo_comp_eq A) simp only [← LinearEquiv.toModuleIso_hom] at h1 h2 simp only [(Iso.eq_inv_comp _).2 h2, (Iso.eq_inv_comp _).2 h1, Category.assoc, Iso.hom_inv_id_assoc, HomologicalComplex.d_comp_d_assoc, zero_comp, comp_zero] @@ -716,7 +716,7 @@ lemma shortComplexH0_exact : (shortComplexH0 A).Exact := by `(inhomogeneousCochains A).d 0 1` of the complex of inhomogeneous cochains of `A`. -/ @[simps! hom_left hom_right inv_left inv_right] def dZeroArrowIso : Arrow.mk ((inhomogeneousCochains A).d 0 1) ≅ - Arrow.mk (ModuleCat.ofHom (dZero A)) := + Arrow.mk (ModuleCat.asHom (dZero A)) := Arrow.isoMk (zeroCochainsLequiv A).toModuleIso (oneCochainsLequiv A).toModuleIso (dZero_comp_eq A) @@ -764,7 +764,7 @@ def isoOneCocycles : cocycles A 1 ≅ ModuleCat.of k (oneCocycles A) := cyclesMapIso (shortComplexH1Iso A) ≪≫ (shortComplexH1 A).moduleCatCyclesIso lemma isoOneCocycles_hom_comp_subtype : - (isoOneCocycles A).hom ≫ ModuleCat.ofHom (oneCocycles A).subtype = + (isoOneCocycles A).hom ≫ ModuleCat.asHom (oneCocycles A).subtype = iCocycles A 1 ≫ (oneCochainsLequiv A).toModuleIso.hom := by dsimp [isoOneCocycles] rw [Category.assoc, Category.assoc] @@ -774,7 +774,7 @@ lemma isoOneCocycles_hom_comp_subtype : lemma toCocycles_comp_isoOneCocycles_hom : toCocycles A 0 1 ≫ (isoOneCocycles A).hom = (zeroCochainsLequiv A).toModuleIso.hom ≫ - ModuleCat.ofHom (shortComplexH1 A).moduleCatToCycles := by + ModuleCat.asHom (shortComplexH1 A).moduleCatToCycles := by simp [isoOneCocycles] rfl @@ -812,7 +812,7 @@ def isoTwoCocycles : cocycles A 2 ≅ ModuleCat.of k (twoCocycles A) := cyclesMapIso (shortComplexH2Iso A) ≪≫ (shortComplexH2 A).moduleCatCyclesIso lemma isoTwoCocycles_hom_comp_subtype : - (isoTwoCocycles A).hom ≫ ModuleCat.ofHom (twoCocycles A).subtype = + (isoTwoCocycles A).hom ≫ ModuleCat.asHom (twoCocycles A).subtype = iCocycles A 2 ≫ (twoCochainsLequiv A).toModuleIso.hom := by dsimp [isoTwoCocycles] rw [Category.assoc, Category.assoc] @@ -822,7 +822,7 @@ lemma isoTwoCocycles_hom_comp_subtype : lemma toCocycles_comp_isoTwoCocycles_hom : toCocycles A 1 2 ≫ (isoTwoCocycles A).hom = (oneCochainsLequiv A).toModuleIso.hom ≫ - ModuleCat.ofHom (shortComplexH2 A).moduleCatToCycles := by + ModuleCat.asHom (shortComplexH2 A).moduleCatToCycles := by simp [isoTwoCocycles] rfl diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 223f8229607a1..fc63a62f92cea 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -332,7 +332,7 @@ variable [Group G] (A B C : Rep k G) protected def ihom (A : Rep k G) : Rep k G ⥤ Rep k G where obj B := Rep.of (Representation.linHom A.ρ B.ρ) map := fun {X} {Y} f => - { hom := ModuleCat.ofHom (LinearMap.llcomp k _ _ _ f.hom) + { hom := ModuleCat.asHom (LinearMap.llcomp k _ _ _ f.hom) comm := fun g => LinearMap.ext fun x => LinearMap.ext fun y => by show f.hom (X.ρ g _) = _ simp only [hom_comm_apply]; rfl } diff --git a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean index afc24e101812b..1c4dfed5cf15a 100644 --- a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean +++ b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean @@ -241,17 +241,17 @@ private instance : AddCommGroup (AdicCompletion I R ⊗[R] (LinearMap.ker f)) := private def firstRow : ComposableArrows (ModuleCat (AdicCompletion I R)) 4 := ComposableArrows.mk₄ - (ModuleCat.ofHom <| lTensorKerIncl I M f) - (ModuleCat.ofHom <| lTensorf I M f) - (ModuleCat.ofHom (0 : AdicCompletion I R ⊗[R] M →ₗ[AdicCompletion I R] PUnit)) - (ModuleCat.ofHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.asHom <| lTensorKerIncl I M f) + (ModuleCat.asHom <| lTensorf I M f) + (ModuleCat.asHom (0 : AdicCompletion I R ⊗[R] M →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.asHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) private def secondRow : ComposableArrows (ModuleCat (AdicCompletion I R)) 4 := ComposableArrows.mk₄ - (ModuleCat.ofHom (map I <| (LinearMap.ker f).subtype)) - (ModuleCat.ofHom (map I f)) - (ModuleCat.ofHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) - (ModuleCat.ofHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.asHom (map I <| (LinearMap.ker f).subtype)) + (ModuleCat.asHom (map I f)) + (ModuleCat.asHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.asHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) include hf @@ -282,25 +282,25 @@ private lemma secondRow_exact [Fintype ι] [IsNoetherianRing R] : (secondRow I M /- The compatible vertical maps between the first and the second row. -/ private def firstRowToSecondRow : firstRow I M f ⟶ secondRow I M f := ComposableArrows.homMk₄ - (ModuleCat.ofHom (ofTensorProduct I (LinearMap.ker f))) - (ModuleCat.ofHom (ofTensorProduct I (ι → R))) - (ModuleCat.ofHom (ofTensorProduct I M)) - (ModuleCat.ofHom 0) - (ModuleCat.ofHom 0) + (ModuleCat.asHom (ofTensorProduct I (LinearMap.ker f))) + (ModuleCat.asHom (ofTensorProduct I (ι → R))) + (ModuleCat.asHom (ofTensorProduct I M)) + (ModuleCat.asHom 0) + (ModuleCat.asHom 0) (ofTensorProduct_naturality I <| (LinearMap.ker f).subtype).symm (ofTensorProduct_naturality I f).symm rfl rfl private lemma ofTensorProduct_iso [Fintype ι] [IsNoetherianRing R] : - IsIso (ModuleCat.ofHom (ofTensorProduct I M)) := by + IsIso (ModuleCat.asHom (ofTensorProduct I M)) := by refine Abelian.isIso_of_epi_of_isIso_of_isIso_of_mono (firstRow_exact I M f hf) (secondRow_exact I M f hf) (firstRowToSecondRow I M f) ?_ ?_ ?_ ?_ · apply ConcreteCategory.epi_of_surjective exact ofTensorProduct_surjective_of_finite I (LinearMap.ker f) · apply (ConcreteCategory.isIso_iff_bijective _).mpr exact ofTensorProduct_bijective_of_pi_of_fintype I ι - · show IsIso (ModuleCat.ofHom 0) + · show IsIso (ModuleCat.asHom 0) apply Limits.isIso_of_isTerminal <;> exact Limits.IsZero.isTerminal (ModuleCat.isZero_of_subsingleton _) · apply ConcreteCategory.mono_of_injective @@ -310,9 +310,9 @@ private lemma ofTensorProduct_iso [Fintype ι] [IsNoetherianRing R] : private lemma ofTensorProduct_bijective_of_map_from_fin [Fintype ι] [IsNoetherianRing R] : Function.Bijective (ofTensorProduct I M) := by - have : IsIso (ModuleCat.ofHom (ofTensorProduct I M)) := + have : IsIso (ModuleCat.asHom (ofTensorProduct I M)) := ofTensorProduct_iso I M f hf - exact ConcreteCategory.bijective_of_isIso (ModuleCat.ofHom (ofTensorProduct I M)) + exact ConcreteCategory.bijective_of_isIso (ModuleCat.asHom (ofTensorProduct I M)) end diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index 2f80a41c0e3cb..67c966cf31c26 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -5,6 +5,7 @@ Authors: Johan Commelin -/ import Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Basic import Mathlib.RingTheory.Polynomial.IntegralNormalization +import Mathlib.RingTheory.LocalRing.Basic /-! # Algebraic elements and algebraic extensions diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian.lean index 60965f6c689e4..9e43f7cbc62b9 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian.lean @@ -91,22 +91,33 @@ theorem isArtinian_of_surjective (f : M →ₗ[R] P) (hf : Function.Surjective f show A.comap f < B.comap f from Submodule.comap_strictMono_of_surjective hf hAB) (InvImage.wf (Submodule.comap f) IsWellFounded.wf)⟩ +instance isArtinian_of_quotient_of_artinian + (N : Submodule R M) [IsArtinian R M] : IsArtinian R (M ⧸ N) := + isArtinian_of_surjective M (Submodule.mkQ N) (Submodule.Quotient.mk_surjective N) + variable {M} theorem isArtinian_of_linearEquiv (f : M ≃ₗ[R] P) [IsArtinian R M] : IsArtinian R P := isArtinian_of_surjective _ f.toLinearMap f.toEquiv.surjective theorem isArtinian_of_range_eq_ker [IsArtinian R M] [IsArtinian R P] (f : M →ₗ[R] N) (g : N →ₗ[R] P) - (hf : Function.Injective f) (hg : Function.Surjective g) (h : LinearMap.range f = LinearMap.ker g) : IsArtinian R N := - wellFounded_lt_exact_sequence (LinearMap.range f) (Submodule.map f) - (Submodule.comap f) (Submodule.comap g) (Submodule.map g) (Submodule.gciMapComap hf) - (Submodule.giMapComap hg) - (by simp [Submodule.map_comap_eq, inf_comm]) (by simp [Submodule.comap_map_eq, h]) + wellFounded_lt_exact_sequence (LinearMap.range f) (Submodule.map (f.ker.liftQ f le_rfl)) + (Submodule.comap (f.ker.liftQ f le_rfl)) + (Submodule.comap g.rangeRestrict) (Submodule.map g.rangeRestrict) + (Submodule.gciMapComap <| LinearMap.ker_eq_bot.mp <| Submodule.ker_liftQ_eq_bot _ _ _ le_rfl) + (Submodule.giMapComap g.surjective_rangeRestrict) + (by simp [Submodule.map_comap_eq, inf_comm, Submodule.range_liftQ]) + (by simp [Submodule.comap_map_eq, h]) + +theorem isArtinian_iff_submodule_quotient (S : Submodule R P) : + IsArtinian R P ↔ IsArtinian R S ∧ IsArtinian R (P ⧸ S) := by + refine ⟨fun h ↦ ⟨inferInstance, inferInstance⟩, fun ⟨_, _⟩ ↦ ?_⟩ + apply isArtinian_of_range_eq_ker S.subtype S.mkQ + rw [Submodule.ker_mkQ, Submodule.range_subtype] instance isArtinian_prod [IsArtinian R M] [IsArtinian R P] : IsArtinian R (M × P) := - isArtinian_of_range_eq_ker (LinearMap.inl R M P) (LinearMap.snd R M P) LinearMap.inl_injective - LinearMap.snd_surjective (LinearMap.range_inl R M P) + isArtinian_of_range_eq_ker (LinearMap.inl R M P) (LinearMap.snd R M P) (LinearMap.range_inl R M P) instance (priority := 100) isArtinian_of_finite [Finite M] : IsArtinian R M := ⟨Finite.wellFounded_of_trans_of_irrefl _⟩ @@ -115,17 +126,12 @@ instance (priority := 100) isArtinian_of_finite [Finite M] : IsArtinian R M := -- attribute [local elab_as_elim] Finite.induction_empty_option instance isArtinian_pi {R ι : Type*} [Finite ι] : - ∀ {M : ι → Type*} [Ring R] [∀ i, AddCommGroup (M i)], - ∀ [∀ i, Module R (M i)], ∀ [∀ i, IsArtinian R (M i)], IsArtinian R (∀ i, M i) := by + ∀ {M : ι → Type*} [Ring R] [∀ i, AddCommGroup (M i)] + [∀ i, Module R (M i)] [∀ i, IsArtinian R (M i)], IsArtinian R (∀ i, M i) := by apply Finite.induction_empty_option _ _ _ ι - · intro α β e hα M _ _ _ _ - have := @hα - exact isArtinian_of_linearEquiv (LinearEquiv.piCongrLeft R M e) - · intro M _ _ _ _ - infer_instance - · intro α _ ih M _ _ _ _ - have := @ih - exact isArtinian_of_linearEquiv (LinearEquiv.piOptionEquivProd R).symm + · exact fun e h ↦ isArtinian_of_linearEquiv (LinearEquiv.piCongrLeft R _ e) + · infer_instance + · exact fun ih ↦ isArtinian_of_linearEquiv (LinearEquiv.piOptionEquivProd R).symm /-- A version of `isArtinian_pi` for non-dependent functions. We need this instance because sometimes Lean fails to apply the dependent version in non-dependent settings (e.g., it fails to @@ -336,10 +342,6 @@ theorem Ring.isArtinian_of_zero_eq_one {R} [Ring R] (h01 : (0 : R) = 1) : IsArti theorem isArtinian_of_submodule_of_artinian (R M) [Ring R] [AddCommGroup M] [Module R M] (N : Submodule R M) (_ : IsArtinian R M) : IsArtinian R N := inferInstance -instance isArtinian_of_quotient_of_artinian (R) [Ring R] (M) [AddCommGroup M] [Module R M] - (N : Submodule R M) [IsArtinian R M] : IsArtinian R (M ⧸ N) := - isArtinian_of_surjective M (Submodule.mkQ N) (Submodule.Quotient.mk_surjective N) - /-- If `M / S / R` is a scalar tower, and `M / R` is Artinian, then `M / S` is also Artinian. -/ theorem isArtinian_of_tower (R) {S M} [CommRing R] [Ring S] [AddCommGroup M] [Algebra R S] [Module S M] [Module R M] [IsScalarTower R S M] (h : IsArtinian R M) : IsArtinian S M := @@ -520,7 +522,7 @@ instance [IsReduced R] : DecompositionMonoid (Polynomial R) := theorem isSemisimpleRing_of_isReduced [IsReduced R] : IsSemisimpleRing R := (equivPi R).symm.isSemisimpleRing -proof_wanted IsSemisimpleRing.isArtinianRing (R : Type*) [CommRing R] [IsSemisimpleRing R] : +proof_wanted IsSemisimpleRing.isArtinianRing (R : Type*) [Ring R] [IsSemisimpleRing R] : IsArtinianRing R end IsArtinianRing diff --git a/Mathlib/RingTheory/Binomial.lean b/Mathlib/RingTheory/Binomial.lean index 45bbba6501dad..659a95f8fb226 100644 --- a/Mathlib/RingTheory/Binomial.lean +++ b/Mathlib/RingTheory/Binomial.lean @@ -175,7 +175,7 @@ theorem ascPochhammer_smeval_cast (R : Type*) [Semiring R] {S : Type*} [NonAssoc simp only [← C_eq_natCast, smeval_C_mul, hn, Nat.cast_smul_eq_nsmul R n] simp only [nsmul_eq_mul, Nat.cast_id] -variable {R S : Type*} +variable {R : Type*} theorem ascPochhammer_smeval_eq_eval [Semiring R] (r : R) (n : ℕ) : (ascPochhammer ℕ n).smeval r = (ascPochhammer R n).eval r := by diff --git a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean index dcd8d8d13df2e..5226e76ef5d0e 100644 --- a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean +++ b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean @@ -58,7 +58,7 @@ noncomputable instance TensorProduct.instCoalgebra : Coalgebra R (M ⊗[R] N) := rw [CoalgebraCat.ofComonObjCoalgebraStruct_comul] simp [-Mon_.monMonoidalStruct_tensorObj_X, ModuleCat.MonoidalCategory.instMonoidalCategoryStruct_tensorHom, - ModuleCat.comp_def, ModuleCat.of, ModuleCat.ofHom, + ModuleCat.comp_def, ModuleCat.of, ModuleCat.asHom, ModuleCat.MonoidalCategory.tensor_μ_eq_tensorTensorTensorComm] } end diff --git a/Mathlib/RingTheory/Congruence/Basic.lean b/Mathlib/RingTheory/Congruence/Basic.lean index 6d8463a131488..131575077c37b 100644 --- a/Mathlib/RingTheory/Congruence/Basic.lean +++ b/Mathlib/RingTheory/Congruence/Basic.lean @@ -69,7 +69,7 @@ variable [Add R] [Mul R] (c : RingCon R) /-- A coercion from a congruence relation to its underlying binary relation. -/ instance : FunLike (RingCon R) R (R → Prop) where - coe := fun c => c.r + coe c := c.r coe_injective' x y h := by rcases x with ⟨⟨x, _⟩, _⟩ rcases y with ⟨⟨y, _⟩, _⟩ diff --git a/Mathlib/RingTheory/DedekindDomain/Basic.lean b/Mathlib/RingTheory/DedekindDomain/Basic.lean index 29a6df2a644da..aae1db626fae1 100644 --- a/Mathlib/RingTheory/DedekindDomain/Basic.lean +++ b/Mathlib/RingTheory/DedekindDomain/Basic.lean @@ -49,7 +49,7 @@ variable (R A K : Type*) [CommRing R] [CommRing A] [Field K] open scoped nonZeroDivisors Polynomial /-- A ring `R` has Krull dimension at most one if all nonzero prime ideals are maximal. -/ -class Ring.DimensionLEOne : Prop := +class Ring.DimensionLEOne : Prop where (maximalOfPrime : ∀ {p : Ideal R}, p ≠ ⊥ → p.IsPrime → p.IsMaximal) open Ideal Ring @@ -122,7 +122,6 @@ use `isDedekindDomain_iff` to prove `IsDedekindDomain` for a given `fraction_map This is the default implementation, but there are equivalent definitions, `IsDedekindDomainDvr` and `IsDedekindDomainInv`. -TODO: Prove that these are actually equivalent definitions. -/ class IsDedekindDomain extends IsDomain A, IsDedekindRing A : Prop diff --git a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean index 395fe6de21fed..9cd04bf371ef8 100644 --- a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean +++ b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean @@ -337,7 +337,7 @@ instance : IsScalarTower R K (FiniteAdeleRing R K) := IsScalarTower.of_algebraMap_eq' rfl instance : Coe (FiniteAdeleRing R K) (K_hat R K) where - coe := fun x ↦ x.1 + coe x := x.1 @[ext] lemma ext {a₁ a₂ : FiniteAdeleRing R K} (h : (a₁ : K_hat R K) = a₂) : a₁ = a₂ := diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 0c6332b7d5a5b..426d939612987 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -1191,26 +1191,6 @@ theorem Ideal.le_mul_of_no_prime_factors {I J K : Ideal R} (UniqueFactorizationMonoid.dvd_of_dvd_mul_right_of_no_prime_factors (b := K) hJ0 ?_ hJ) exact fun hPJ hPK => mt Ideal.isPrime_of_prime (coprime _ hPJ hPK) -theorem Ideal.le_of_pow_le_prime {I P : Ideal R} [hP : P.IsPrime] {n : ℕ} (h : I ^ n ≤ P) : - I ≤ P := by - by_cases hP0 : P = ⊥ - · simp only [hP0, le_bot_iff] at h ⊢ - exact pow_eq_zero h - rw [← Ideal.dvd_iff_le] at h ⊢ - exact ((Ideal.prime_iff_isPrime hP0).mpr hP).dvd_of_dvd_pow h - -theorem Ideal.pow_le_prime_iff {I P : Ideal R} [_hP : P.IsPrime] {n : ℕ} (hn : n ≠ 0) : - I ^ n ≤ P ↔ I ≤ P := - ⟨Ideal.le_of_pow_le_prime, fun h => _root_.trans (Ideal.pow_le_self hn) h⟩ - -theorem Ideal.prod_le_prime {ι : Type*} {s : Finset ι} {f : ι → Ideal R} {P : Ideal R} - [hP : P.IsPrime] : ∏ i ∈ s, f i ≤ P ↔ ∃ i ∈ s, f i ≤ P := by - by_cases hP0 : P = ⊥ - · simp only [hP0, le_bot_iff] - rw [← Ideal.zero_eq_bot, Finset.prod_eq_zero_iff] - simp only [← Ideal.dvd_iff_le] - exact ((Ideal.prime_iff_isPrime hP0).mpr hP).dvd_finset_prod_iff _ - /-- The intersection of distinct prime powers in a Dedekind domain is the product of these prime powers. -/ theorem IsDedekindDomain.inf_prime_pow_eq_prod {ι : Type*} (s : Finset ι) (f : ι → Ideal R) @@ -1228,15 +1208,13 @@ theorem IsDedekindDomain.inf_prime_pow_eq_prod {ι : Type*} (s : Finset ι) (f : rw [Finset.inf_insert, Finset.prod_insert ha, ih] refine le_antisymm (Ideal.le_mul_of_no_prime_factors ?_ inf_le_left inf_le_right) Ideal.mul_le_inf intro P hPa hPs hPp - obtain ⟨b, hb, hPb⟩ := Ideal.prod_le_prime.mp hPs + obtain ⟨b, hb, hPb⟩ := hPp.prod_le.mp hPs haveI := Ideal.isPrime_of_prime (prime a (Finset.mem_insert_self a s)) haveI := Ideal.isPrime_of_prime (prime b (Finset.mem_insert_of_mem hb)) refine coprime a (Finset.mem_insert_self a s) b (Finset.mem_insert_of_mem hb) ?_ ?_ · exact (ne_of_mem_of_not_mem hb ha).symm - · refine ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp - (Ideal.le_of_pow_le_prime hPa)).trans - ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp - (Ideal.le_of_pow_le_prime hPb)).symm + · refine ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp (hPp.le_of_pow_le hPa)).trans + ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp (hPp.le_of_pow_le hPb)).symm · exact (prime a (Finset.mem_insert_self a s)).ne_zero · exact (prime b (Finset.mem_insert_of_mem hb)).ne_zero @@ -1251,17 +1229,13 @@ noncomputable def IsDedekindDomain.quotientEquivPiOfProdEq {ι : Type*} [Fintype simp only [← prod_eq, Finset.inf_eq_iInf, Finset.mem_univ, ciInf_pos, ← IsDedekindDomain.inf_prime_pow_eq_prod _ _ _ (fun i _ => prime i) (coprime.set_pairwise _)])).trans <| - Ideal.quotientInfRingEquivPiQuotient _ fun i j hij => Ideal.coprime_of_no_prime_ge (by + Ideal.quotientInfRingEquivPiQuotient _ fun i j hij => Ideal.coprime_of_no_prime_ge <| by intro P hPi hPj hPp haveI := Ideal.isPrime_of_prime (prime i) haveI := Ideal.isPrime_of_prime (prime j) - refine coprime hij ?_ - refine ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp - (Ideal.le_of_pow_le_prime hPi)).trans - ((Ring.DimensionLeOne.prime_le_prime_iff_eq ?_).mp - (Ideal.le_of_pow_le_prime hPj)).symm - · exact (prime i).ne_zero - · exact (prime j).ne_zero) + exact coprime hij <| ((Ring.DimensionLeOne.prime_le_prime_iff_eq (prime i).ne_zero).mp + (hPp.le_of_pow_le hPi)).trans <| Eq.symm <| + (Ring.DimensionLeOne.prime_le_prime_iff_eq (prime j).ne_zero).mp (hPp.le_of_pow_le hPj) open scoped Classical diff --git a/Mathlib/RingTheory/Derivation/MapCoeffs.lean b/Mathlib/RingTheory/Derivation/MapCoeffs.lean index 01ab5a9134a4a..ec1b10e49cb27 100644 --- a/Mathlib/RingTheory/Derivation/MapCoeffs.lean +++ b/Mathlib/RingTheory/Derivation/MapCoeffs.lean @@ -22,7 +22,7 @@ open Polynomial Module namespace Derivation variable {R A M : Type*} [CommRing R] [CommRing A] [Algebra R A] [AddCommGroup M] - [Module A M] [Module R M] (d : Derivation R A M) (a : A) + [Module A M] [Module R M] (d : Derivation R A M) /-- The `R`-derivation from `A[X]` to `M[X]` which applies the derivative to each diff --git a/Mathlib/RingTheory/DualNumber.lean b/Mathlib/RingTheory/DualNumber.lean new file mode 100644 index 0000000000000..cf0ac64836b99 --- /dev/null +++ b/Mathlib/RingTheory/DualNumber.lean @@ -0,0 +1,168 @@ +/- +Copyright (c) 2024 Yakov Pechersky. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yakov Pechersky +-/ +import Mathlib.Algebra.DualNumber +import Mathlib.RingTheory.LocalRing.MaximalIdeal.Basic + +/-! +# Algebraic properties of dual numbers + +## Main results + +* `DualNumber.instLocalRing`: The dual numbers over a field `K` form a local ring. +* `DualNumber.instPrincipalIdealRing`: The dual numbers over a field `K` form a principal ideal + ring. + +-/ + +namespace TrivSqZeroExt + +variable {R M : Type*} + +section Semiring +variable [Semiring R] [AddCommMonoid M] [Module R M] [Module Rᵐᵒᵖ M] [SMulCommClass R Rᵐᵒᵖ M] + +lemma isNilpotent_iff_isNilpotent_fst {x : TrivSqZeroExt R M} : + IsNilpotent x ↔ IsNilpotent x.fst := by + constructor <;> rintro ⟨n, hn⟩ + · refine ⟨n, ?_⟩ + rw [← fst_pow, hn, fst_zero] + · refine ⟨n * 2, ?_⟩ + rw [pow_mul] + ext + · rw [fst_pow, fst_pow, hn, zero_pow two_ne_zero, fst_zero] + · rw [pow_two, snd_mul, fst_pow, hn, MulOpposite.op_zero, zero_smul, zero_smul, zero_add, + snd_zero] + +@[simp] +lemma isNilpotent_inl_iff (r : R) : IsNilpotent (.inl r : TrivSqZeroExt R M) ↔ IsNilpotent r := by + rw [isNilpotent_iff_isNilpotent_fst, fst_inl] + +@[simp] +lemma isNilpotent_inr (x : M) : IsNilpotent (.inr x : TrivSqZeroExt R M) := by + refine ⟨2, by simp [pow_two]⟩ + +end Semiring + +lemma isUnit_or_isNilpotent_of_isMaximal_isNilpotent [CommSemiring R] [AddCommGroup M] + [Module R M] [Module Rᵐᵒᵖ M] [IsCentralScalar R M] + (h : ∀ I : Ideal R, I.IsMaximal → IsNilpotent I) + (a : TrivSqZeroExt R M) : + IsUnit a ∨ IsNilpotent a := by + rw [isUnit_iff_isUnit_fst, isNilpotent_iff_isNilpotent_fst] + refine (em _).imp_right fun ha ↦ ?_ + obtain ⟨I, hI, haI⟩ := exists_max_ideal_of_mem_nonunits (mem_nonunits_iff.mpr ha) + refine (h _ hI).imp fun n hn ↦ ?_ + exact hn.le (Ideal.pow_mem_pow haI _) + +lemma isUnit_or_isNilpotent [DivisionSemiring R] [AddCommGroup M] + [Module R M] [Module Rᵐᵒᵖ M] [SMulCommClass R Rᵐᵒᵖ M] + (a : TrivSqZeroExt R M) : + IsUnit a ∨ IsNilpotent a := by + simp [isUnit_iff_isUnit_fst, isNilpotent_iff_isNilpotent_fst, em'] + +end TrivSqZeroExt + +namespace DualNumber +variable {R : Type*} + +lemma fst_eq_zero_iff_eps_dvd [Semiring R] {x : R[ε]} : + x.fst = 0 ↔ ε ∣ x := by + simp_rw [dvd_def, TrivSqZeroExt.ext_iff, TrivSqZeroExt.fst_mul, TrivSqZeroExt.snd_mul, + fst_eps, snd_eps, zero_mul, zero_smul, zero_add, MulOpposite.smul_eq_mul_unop, + MulOpposite.unop_op, one_mul, exists_and_left, iff_self_and] + intro + exact ⟨.inl x.snd, rfl⟩ + +lemma isNilpotent_eps [Semiring R] : + IsNilpotent (ε : R[ε]) := + TrivSqZeroExt.isNilpotent_inr 1 + +open TrivSqZeroExt + +lemma isNilpotent_iff_eps_dvd [DivisionSemiring R] {x : R[ε]} : + IsNilpotent x ↔ ε ∣ x := by + simp only [isNilpotent_iff_isNilpotent_fst, isNilpotent_iff_eq_zero, fst_eq_zero_iff_eps_dvd] + +section Field + +variable {K : Type*} + +instance [DivisionRing K] : LocalRing K[ε] where + isUnit_or_isUnit_of_add_one {a b} h := by + rw [add_comm, ← eq_sub_iff_add_eq] at h + rcases eq_or_ne (fst a) 0 with ha|ha <;> + simp [isUnit_iff_isUnit_fst, h, ha] + +lemma ideal_trichotomy [DivisionRing K] (I : Ideal K[ε]) : + I = ⊥ ∨ I = .span {ε} ∨ I = ⊤ := by + refine (eq_or_ne I ⊥).imp_right fun hb ↦ ?_ + refine (eq_or_ne I ⊤).symm.imp_left fun ht ↦ ?_ + have hd : ∀ x ∈ I, ε ∣ x := by + intro x hxI + rcases isUnit_or_isNilpotent x with hx|hx + · exact absurd (Ideal.eq_top_of_isUnit_mem _ hxI hx) ht + · rwa [← isNilpotent_iff_eps_dvd] + have hd' : ∀ x ∈ I, x ≠ 0 → ∃ r, ε = r * x := by + intro x hxI hx0 + obtain ⟨r, rfl⟩ := hd _ hxI + have : ε * r = (fst r) • ε := by ext <;> simp + rw [this] at hxI hx0 ⊢ + have hr : fst r ≠ 0 := by + contrapose! hx0 + simp [hx0] + refine ⟨r⁻¹, ?_⟩ + simp [TrivSqZeroExt.ext_iff, inv_mul_cancel₀ hr] + refine le_antisymm ?_ ?_ <;> intro x <;> + simp_rw [Ideal.mem_span_singleton', (commute_eps_right _).eq, eq_comm, ← dvd_def] + · intro hx + simp_rw [hd _ hx] + · intro hx + obtain ⟨p, rfl⟩ := hx + obtain ⟨y, hyI, hy0⟩ := Submodule.exists_mem_ne_zero_of_ne_bot hb + obtain ⟨r, hr⟩ := hd' _ hyI hy0 + rw [(commute_eps_left _).eq, hr, ← mul_assoc] + exact Ideal.mul_mem_left _ _ hyI + +lemma isMaximal_span_singleton_eps [DivisionRing K] : + (Ideal.span {ε} : Ideal K[ε]).IsMaximal := by + refine ⟨?_, fun I hI ↦ ?_⟩ + · simp [ne_eq, Ideal.eq_top_iff_one, Ideal.mem_span_singleton', TrivSqZeroExt.ext_iff] + · rcases ideal_trichotomy I with rfl|rfl|rfl <;> + first | simp at hI | simp + +lemma maximalIdeal_eq_span_singleton_eps [Field K] : + LocalRing.maximalIdeal K[ε] = Ideal.span {ε} := + (LocalRing.eq_maximalIdeal isMaximal_span_singleton_eps).symm + +instance [DivisionRing K] : IsPrincipalIdealRing K[ε] where + principal I := by + rcases ideal_trichotomy I with rfl|rfl|rfl + · exact bot_isPrincipal + · exact ⟨_, rfl⟩ + · exact top_isPrincipal + +lemma exists_mul_left_or_mul_right [DivisionRing K] (a b : K[ε]) : + ∃ c, a * c = b ∨ b * c = a := by + rcases isUnit_or_isNilpotent a with ha|ha + · lift a to K[ε]ˣ using ha + exact ⟨a⁻¹ * b, by simp⟩ + rcases isUnit_or_isNilpotent b with hb|hb + · lift b to K[ε]ˣ using hb + exact ⟨b⁻¹ * a, by simp⟩ + rw [isNilpotent_iff_eps_dvd] at ha hb + obtain ⟨x, rfl⟩ := ha + obtain ⟨y, rfl⟩ := hb + suffices ∃ c, fst x * fst c = fst y ∨ fst y * fst c = fst x by + simpa [TrivSqZeroExt.ext_iff] using this + rcases eq_or_ne (fst x) 0 with hx|hx + · refine ⟨ε, Or.inr ?_⟩ + simp [hx] + refine ⟨inl ((fst x)⁻¹ * fst y), ?_⟩ + simp [← inl_mul, ← mul_assoc, mul_inv_cancel₀ hx] + +end Field + +end DualNumber diff --git a/Mathlib/RingTheory/Etale/Basic.lean b/Mathlib/RingTheory/Etale/Basic.lean index 171f57f41c31c..70ab05f56b544 100644 --- a/Mathlib/RingTheory/Etale/Basic.lean +++ b/Mathlib/RingTheory/Etale/Basic.lean @@ -56,7 +56,6 @@ section variable {R : Type u} [CommSemiring R] variable {A : Type u} [Semiring A] [Algebra R A] -variable {B : Type u} [CommRing B] [Algebra R B] (I : Ideal B) theorem iff_unramified_and_smooth : FormallyEtale R A ↔ FormallyUnramified R A ∧ FormallySmooth R A := by diff --git a/Mathlib/RingTheory/FiniteLength.lean b/Mathlib/RingTheory/FiniteLength.lean new file mode 100644 index 0000000000000..d0305806bd193 --- /dev/null +++ b/Mathlib/RingTheory/FiniteLength.lean @@ -0,0 +1,78 @@ +/- +Copyright (c) 2024 Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ +import Mathlib.Order.Atoms.Finite +import Mathlib.RingTheory.Artinian + +/-! +# Modules of finite length + +We define modules of finite length (`IsFiniteLength`) to be finite iterated extensions of +simple modules, and show that a module is of finite length iff it is both Noetherian and Artinian, +iff it admits a composition series. +We do not make `IsFiniteLength` a class, instead we use `[IsNoetherian R M] [IsArtinian R M]`. + +## Tag + +Finite length, Composition series +-/ + +universe u + +variable (R : Type*) [Ring R] + +/-- A module of finite length is either trivial or a simple extension of a module known +to be of finite length. -/ +inductive IsFiniteLength : ∀ (M : Type u) [AddCommGroup M] [Module R M], Prop + | of_subsingleton {M} [AddCommGroup M] [Module R M] [Subsingleton M] : IsFiniteLength M + | of_simple_quotient {M} [AddCommGroup M] [Module R M] {N : Submodule R M} + [IsSimpleModule R (M ⧸ N)] : IsFiniteLength N → IsFiniteLength M + +variable {R} {M N : Type*} [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] + +theorem LinearEquiv.isFiniteLength (e : M ≃ₗ[R] N) + (h : IsFiniteLength R M) : IsFiniteLength R N := by + induction' h with M _ _ _ M _ _ S _ _ ih generalizing N + · have := e.symm.toEquiv.subsingleton; exact .of_subsingleton + · have : IsSimpleModule R (N ⧸ Submodule.map (e : M →ₗ[R] N) S) := + IsSimpleModule.congr (Submodule.Quotient.equiv S _ e rfl).symm + exact .of_simple_quotient (ih <| e.submoduleMap S) + +variable (R M) in +theorem exists_compositionSeries_of_isNoetherian_isArtinian [IsNoetherian R M] [IsArtinian R M] : + ∃ s : CompositionSeries (Submodule R M), s.head = ⊥ ∧ s.last = ⊤ := by + obtain ⟨f, f0, n, hn⟩ := exists_covBy_seq_of_wellFoundedLT_wellFoundedGT (Submodule R M) + exact ⟨⟨n, fun i ↦ f i, fun i ↦ hn.2 i i.2⟩, f0.eq_bot, hn.1.eq_top⟩ + +theorem isFiniteLength_of_exists_compositionSeries + (h : ∃ s : CompositionSeries (Submodule R M), s.head = ⊥ ∧ s.last = ⊤) : + IsFiniteLength R M := + Submodule.topEquiv.isFiniteLength <| by + obtain ⟨s, s_head, s_last⟩ := h + rw [← s_last] + suffices ∀ i, IsFiniteLength R (s i) from this (Fin.last _) + intro i + induction' i using Fin.induction with i ih + · change IsFiniteLength R s.head; rw [s_head]; exact .of_subsingleton + let cov := s.step i + have := (covBy_iff_quot_is_simple cov.le).mp cov + have := ((s i.castSucc).comap (s i.succ).subtype).equivMapOfInjective + _ (Submodule.injective_subtype _) + rw [Submodule.map_comap_subtype, inf_of_le_right cov.le] at this + exact .of_simple_quotient (this.symm.isFiniteLength ih) + +theorem isFiniteLength_iff_isNoetherian_isArtinian : + IsFiniteLength R M ↔ IsNoetherian R M ∧ IsArtinian R M := + ⟨fun h ↦ h.rec (fun {M} _ _ _ ↦ ⟨inferInstance, inferInstance⟩) fun M _ _ {N} _ _ ⟨_, _⟩ ↦ + ⟨(isNoetherian_iff_submodule_quotient N).mpr ⟨‹_›, isNoetherian_iff'.mpr inferInstance⟩, + (isArtinian_iff_submodule_quotient N).mpr ⟨‹_›, inferInstance⟩⟩, + fun ⟨_, _⟩ ↦ isFiniteLength_of_exists_compositionSeries + (exists_compositionSeries_of_isNoetherian_isArtinian R M)⟩ + +theorem isFiniteLength_iff_exists_compositionSeries : + IsFiniteLength R M ↔ ∃ s : CompositionSeries (Submodule R M), s.head = ⊥ ∧ s.last = ⊤ := + ⟨fun h ↦ have ⟨_, _⟩ := isFiniteLength_iff_isNoetherian_isArtinian.mp h + exists_compositionSeries_of_isNoetherian_isArtinian R M, + isFiniteLength_of_exists_compositionSeries⟩ diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index 19651b5ecbdf0..0039e1437e417 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -381,9 +381,7 @@ theorem fg_iff_compact (s : Submodule R M) : s.FG ↔ CompleteLattice.IsCompactE suffices u.sup id ≤ s from le_antisymm husup this rw [sSup', Finset.sup_id_eq_sSup] exact sSup_le_sSup huspan - -- Porting note: had to split this out of the `obtain` - have := Finset.subset_image_iff.mp huspan - obtain ⟨t, ⟨-, rfl⟩⟩ := this + obtain ⟨t, -, rfl⟩ := Finset.subset_set_image_iff.mp huspan rw [Finset.sup_image, Function.id_comp, Finset.sup_eq_iSup, supr_rw, ← span_eq_iSup_of_singleton_spans, eq_comm] at ssup exact ⟨t, ssup⟩ diff --git a/Mathlib/RingTheory/Flat/CategoryTheory.lean b/Mathlib/RingTheory/Flat/CategoryTheory.lean index ea7bc816d2f57..ea6d1442dad38 100644 --- a/Mathlib/RingTheory/Flat/CategoryTheory.lean +++ b/Mathlib/RingTheory/Flat/CategoryTheory.lean @@ -52,7 +52,7 @@ lemma iff_lTensor_preserves_shortComplex_exact : ⟨fun _ _ ↦ lTensor_shortComplex_exact _ _, fun H ↦ iff_lTensor_exact.2 <| fun _ _ _ _ _ _ _ _ _ f g h ↦ moduleCat_exact_iff_function_exact _ |>.1 <| - H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) + H (.mk (ModuleCat.asHom f) (ModuleCat.asHom g) (DFunLike.ext _ _ h.apply_apply_eq_zero)) (moduleCat_exact_iff_function_exact _ |>.2 h)⟩ @@ -62,7 +62,7 @@ lemma iff_rTensor_preserves_shortComplex_exact : ⟨fun _ _ ↦ rTensor_shortComplex_exact _ _, fun H ↦ iff_rTensor_exact.2 <| fun _ _ _ _ _ _ _ _ _ f g h ↦ moduleCat_exact_iff_function_exact _ |>.1 <| - H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) + H (.mk (ModuleCat.asHom f) (ModuleCat.asHom g) (DFunLike.ext _ _ h.apply_apply_eq_zero)) (moduleCat_exact_iff_function_exact _ |>.2 h)⟩ diff --git a/Mathlib/RingTheory/FractionalIdeal/Operations.lean b/Mathlib/RingTheory/FractionalIdeal/Operations.lean index 50c15f91c6b62..d6ab280bcde60 100644 --- a/Mathlib/RingTheory/FractionalIdeal/Operations.lean +++ b/Mathlib/RingTheory/FractionalIdeal/Operations.lean @@ -5,6 +5,7 @@ Authors: Anne Baanen, Filippo A. E. Nuccio -/ import Mathlib.RingTheory.FractionalIdeal.Basic import Mathlib.RingTheory.IntegralClosure.IsIntegral.Basic +import Mathlib.RingTheory.LocalRing.Basic /-! # More operations on fractional ideals diff --git a/Mathlib/RingTheory/Henselian.lean b/Mathlib/RingTheory/Henselian.lean index c439259451ff0..aa9b54b1479e5 100644 --- a/Mathlib/RingTheory/Henselian.lean +++ b/Mathlib/RingTheory/Henselian.lean @@ -194,7 +194,7 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] have hf'c : ∀ n, IsUnit (f'.eval (c n)) := by intro n haveI := isLocalRingHom_of_le_jacobson_bot I (IsAdicComplete.le_jacobson_bot I) - apply isUnit_of_map_unit (Ideal.Quotient.mk I) + apply IsUnit.of_map (Ideal.Quotient.mk I) convert h₂ using 1 exact SModEq.def.mp ((hc_mod n).eval _) have hfcI : ∀ n, f.eval (c n) ∈ I ^ (n + 1) := by diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 36081a3d2b091..69a42212b475e 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -28,7 +28,7 @@ Support right ideals, and two-sided ideals over non-commutative rings. universe u v w -variable {α : Type u} {β : Type v} +variable {α : Type u} {β : Type v} {F : Type w} open Set Function @@ -805,6 +805,12 @@ theorem zero_mem_nonunits [Semiring α] : 0 ∈ nonunits α ↔ (0 : α) ≠ 1 : theorem one_not_mem_nonunits [Monoid α] : (1 : α) ∉ nonunits α := not_not_intro isUnit_one +-- Porting note : as this can be proved by other `simp` lemmas, this is marked as high priority. +@[simp (high)] +theorem map_mem_nonunits_iff [Monoid α] [Monoid β] [FunLike F α β] [MonoidHomClass F α β] (f : F) + [IsLocalRingHom f] (a) : f a ∈ nonunits β ↔ a ∈ nonunits α := + ⟨fun h ha => h <| ha.map f, fun h ha => h <| ha.of_map⟩ + theorem coe_subset_nonunits [Semiring α] {I : Ideal α} (h : I ≠ ⊤) : (I : Set α) ⊆ nonunits α := fun _x hx hu => h <| I.eq_top_of_isUnit_mem hx hu diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 35b79eb232f6d..952cf0b1f19a1 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -980,10 +980,37 @@ theorem IsPrime.multiset_prod_mem_iff_exists_mem {I : Ideal R} (hI : I.IsPrime) s.prod ∈ I ↔ ∃ p ∈ s, p ∈ I := by simpa [span_singleton_le_iff_mem] using (hI.multiset_prod_map_le (span {·})) +theorem IsPrime.pow_le_iff {I P : Ideal R} [hP : P.IsPrime] {n : ℕ} (hn : n ≠ 0) : + I ^ n ≤ P ↔ I ≤ P := by + have h : (Multiset.replicate n I).prod ≤ P ↔ _ := hP.multiset_prod_le + simp_rw [Multiset.prod_replicate, Multiset.mem_replicate, ne_eq, hn, not_false_eq_true, + true_and, exists_eq_left] at h + exact h + +@[deprecated (since := "2024-10-06")] alias pow_le_prime_iff := IsPrime.pow_le_iff + +theorem IsPrime.le_of_pow_le {I P : Ideal R} [hP : P.IsPrime] {n : ℕ} (h : I ^ n ≤ P) : + I ≤ P := by + by_cases hn : n = 0 + · rw [hn, pow_zero, one_eq_top] at h + exact fun ⦃_⦄ _ ↦ h Submodule.mem_top + · exact (pow_le_iff hn).mp h + +@[deprecated (since := "2024-10-06")] alias le_of_pow_le_prime := IsPrime.le_of_pow_le + theorem IsPrime.prod_le {s : Finset ι} {f : ι → Ideal R} {P : Ideal R} (hp : IsPrime P) : s.prod f ≤ P ↔ ∃ i ∈ s, f i ≤ P := hp.multiset_prod_map_le f +@[deprecated (since := "2024-10-06")] alias prod_le_prime := IsPrime.prod_le + +/-- The product of a finite number of elements in the commutative semiring `R` lies in the + prime ideal `p` if and only if at least one of those elements is in `p`. -/ +theorem IsPrime.prod_mem_iff {s : Finset ι} {x : ι → R} {p : Ideal R} [hp : p.IsPrime] : + ∏ i in s, x i ∈ p ↔ ∃ i ∈ s, x i ∈ p := by + simp_rw [← span_singleton_le_iff_mem, ← prod_span_singleton] + exact hp.prod_le + theorem IsPrime.prod_mem_iff_exists_mem {I : Ideal R} (hI : I.IsPrime) (s : Finset R) : s.prod (fun x ↦ x) ∈ I ↔ ∃ p ∈ s, p ∈ I := by rw [Finset.prod_eq_multiset_prod, Multiset.map_id'] diff --git a/Mathlib/RingTheory/Idempotents.lean b/Mathlib/RingTheory/Idempotents.lean index 2e8db654b863c..c2dd1724dad6f 100644 --- a/Mathlib/RingTheory/Idempotents.lean +++ b/Mathlib/RingTheory/Idempotents.lean @@ -214,8 +214,6 @@ A family `{ eᵢ }` of idempotent elements is complete orthogonal if structure CompleteOrthogonalIdempotents (e : I → R) extends OrthogonalIdempotents e : Prop where complete : ∑ i, e i = 1 -variable (he : CompleteOrthogonalIdempotents e) - lemma CompleteOrthogonalIdempotents.unique_iff [Unique I] : CompleteOrthogonalIdempotents e ↔ e default = 1 := by rw [completeOrthogonalIdempotents_iff, OrthogonalIdempotents.unique, Fintype.sum_unique, diff --git a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean index ac6ad06cc3419..37343ddf2c611 100644 --- a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean @@ -148,9 +148,23 @@ variable (f : R →+* S) theorem RingHom.IsIntegralElem.neg {x : S} (hx : f.IsIntegralElem x) : f.IsIntegralElem (-x) := hx.of_mem_closure f hx (Subring.neg_mem _ (Subring.subset_closure (Or.inl rfl))) +theorem RingHom.IsIntegralElem.of_neg {x : S} (h : f.IsIntegralElem (-x)) : f.IsIntegralElem x := + neg_neg x ▸ h.neg + +@[simp] +theorem RingHom.IsIntegralElem.neg_iff {x : S} : f.IsIntegralElem (-x) ↔ f.IsIntegralElem x := + ⟨fun h => h.of_neg, fun h => h.neg⟩ + theorem IsIntegral.neg {x : B} (hx : IsIntegral R x) : IsIntegral R (-x) := .of_mem_of_fg _ hx.fg_adjoin_singleton _ (Subalgebra.neg_mem _ <| Algebra.subset_adjoin rfl) +theorem IsIntegral.of_neg {x : B} (hx : IsIntegral R (-x)) : IsIntegral R x := + neg_neg x ▸ hx.neg + +@[simp] +theorem IsIntegral.neg_iff {x : B} : IsIntegral R (-x) ↔ IsIntegral R x := + ⟨IsIntegral.of_neg, IsIntegral.neg⟩ + theorem RingHom.IsIntegralElem.sub {x y : S} (hx : f.IsIntegralElem x) (hy : f.IsIntegralElem y) : f.IsIntegralElem (x - y) := by simpa only [sub_eq_add_neg] using hx.add f (hy.neg f) diff --git a/Mathlib/RingTheory/Jacobson.lean b/Mathlib/RingTheory/Jacobson.lean index 29cd5339dde32..93ed9e32a751f 100644 --- a/Mathlib/RingTheory/Jacobson.lean +++ b/Mathlib/RingTheory/Jacobson.lean @@ -144,7 +144,7 @@ section Localization open IsLocalization Submonoid -variable {R S : Type*} [CommRing R] [CommRing S] {I : Ideal R} +variable {R S : Type*} [CommRing R] [CommRing S] variable (y : R) [Algebra R S] [IsLocalization.Away y S] variable (S) diff --git a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean index bcdc42a93b266..75e250102f8fe 100644 --- a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean +++ b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean @@ -157,10 +157,6 @@ lemma map_comp_cotangentComplex (f : Hom P P') : end CotangentSpace -universe uT - -variable {T : Type uT} [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] - lemma Hom.sub_aux (f g : Hom P P') (x y) : letI := ((algebraMap S S').comp (algebraMap P.Ring S)).toAlgebra f.toAlgHom (x * y) - g.toAlgHom (x * y) - diff --git a/Mathlib/RingTheory/Kaehler/Polynomial.lean b/Mathlib/RingTheory/Kaehler/Polynomial.lean index c5e94eaca9de6..688733e7fd0c5 100644 --- a/Mathlib/RingTheory/Kaehler/Polynomial.lean +++ b/Mathlib/RingTheory/Kaehler/Polynomial.lean @@ -16,7 +16,7 @@ open Algebra universe u v -variable (R : Type u) (S : Type v) [CommRing R] [CommRing S] [Algebra R S] +variable (R : Type u) [CommRing R] suppress_compilation diff --git a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean index 7dd4c8dade367..4b44e62f3ee66 100644 --- a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Chris Hughes, Mario Carneiro -/ -import Mathlib.RingTheory.LocalRing.RingHom.Defs +import Mathlib.Algebra.Group.Units.Hom import Mathlib.RingTheory.LocalRing.MaximalIdeal.Defs import Mathlib.RingTheory.Ideal.Maps import Mathlib.Logic.Equiv.TransferInstance @@ -24,34 +24,14 @@ variable [Semiring R] [Semiring S] [Semiring T] instance isLocalRingHom_id (R : Type*) [Semiring R] : IsLocalRingHom (RingHom.id R) where map_nonunit _ := id -@[simp] -theorem isUnit_map_iff (f : R →+* S) [IsLocalRingHom f] (a) : IsUnit (f a) ↔ IsUnit a := - ⟨IsLocalRingHom.map_nonunit a, f.isUnit_map⟩ - --- Porting note: as this can be proved by other `simp` lemmas, this is marked as high priority. -@[simp (high)] -theorem map_mem_nonunits_iff (f : R →+* S) [IsLocalRingHom f] (a) : - f a ∈ nonunits S ↔ a ∈ nonunits R := - ⟨fun h ha => h <| (isUnit_map_iff f a).mpr ha, fun h ha => h <| (isUnit_map_iff f a).mp ha⟩ - -instance isLocalRingHom_comp (g : S →+* T) (f : R →+* S) [IsLocalRingHom g] [IsLocalRingHom f] : - IsLocalRingHom (g.comp f) where - map_nonunit a := IsLocalRingHom.map_nonunit a ∘ IsLocalRingHom.map_nonunit (f a) - -instance isLocalRingHom_equiv (f : R ≃+* S) : IsLocalRingHom (f : R →+* S) where - map_nonunit a ha := by - convert RingHom.isUnit_map (f.symm : S →+* R) ha - exact (RingEquiv.symm_apply_apply f a).symm - -@[simp] -theorem isUnit_of_map_unit (f : R →+* S) [IsLocalRingHom f] (a) (h : IsUnit (f a)) : IsUnit a := - IsLocalRingHom.map_nonunit a h - -theorem of_irreducible_map (f : R →+* S) [h : IsLocalRingHom f] {x} (hfx : Irreducible (f x)) : - Irreducible x := - ⟨fun h => hfx.not_unit <| IsUnit.map f h, fun p q hx => - let ⟨H⟩ := h - Or.imp (H p) (H q) <| hfx.isUnit_or_isUnit <| f.map_mul p q ▸ congr_arg f hx⟩ +-- see note [lower instance priority] +instance (priority := 100) isLocalRingHom_toRingHom {F : Type*} [FunLike F R S] + [RingHomClass F R S] (f : F) [IsLocalRingHom f] : IsLocalRingHom (f : R →+* S) := + ⟨IsLocalRingHom.map_nonunit (f := f)⟩ + +instance RingHom.isLocalRingHom_comp (g : S →+* T) (f : R →+* S) [IsLocalRingHom g] + [IsLocalRingHom f] : IsLocalRingHom (g.comp f) where + map_nonunit a := IsLocalRingHom.map_nonunit a ∘ IsLocalRingHom.map_nonunit (f := g) (f a) theorem isLocalRingHom_of_comp (f : R →+* S) (g : S →+* T) [IsLocalRingHom (g.comp f)] : IsLocalRingHom f := diff --git a/Mathlib/RingTheory/LocalRing/RingHom/Defs.lean b/Mathlib/RingTheory/LocalRing/RingHom/Defs.lean deleted file mode 100644 index 5949c6dcab0ff..0000000000000 --- a/Mathlib/RingTheory/LocalRing/RingHom/Defs.lean +++ /dev/null @@ -1,28 +0,0 @@ -/- -Copyright (c) 2018 Kenny Lau. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kenny Lau, Chris Hughes, Mario Carneiro --/ -import Mathlib.Algebra.Group.Units -import Mathlib.Algebra.Ring.Hom.Defs - -/-! - -# Local rings homomorphisms - -We Define local ring homomorhpisms. - -## Main definitions - -* `IsLocalRingHom`: A predicate on semiring homomorphisms, requiring that it maps nonunits - to nonunits. For local rings, this means that the image of the unique maximal ideal is again - contained in the unique maximal ideal. - --/ - -/-- A local ring homomorphism is a homomorphism `f` between local rings such that `a` in the domain - is a unit if `f a` is a unit for any `a`. See `LocalRing.local_hom_TFAE` for other equivalent - definitions. -/ -class IsLocalRingHom {R S : Type*} [Semiring R] [Semiring S] (f : R →+* S) : Prop where - /-- A local ring homomorphism `f : R ⟶ S` will send nonunits of `R` to nonunits of `S`. -/ - map_nonunit : ∀ a, IsUnit (f a) → IsUnit a diff --git a/Mathlib/RingTheory/Localization/AtPrime.lean b/Mathlib/RingTheory/Localization/AtPrime.lean index a52c0df699d51..5b122fb2d696c 100644 --- a/Mathlib/RingTheory/Localization/AtPrime.lean +++ b/Mathlib/RingTheory/Localization/AtPrime.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baan -/ import Mathlib.RingTheory.Localization.Ideal import Mathlib.RingTheory.LocalRing.MaximalIdeal.Basic -import Mathlib.RingTheory.LocalRing.RingHom.Defs +import Mathlib.Algebra.Group.Units.Hom /-! # Localizations of commutative rings at the complement of a prime ideal diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean index b003422369bc5..717b6f8ac0bf2 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean @@ -146,7 +146,7 @@ variable (σ) in noncomputable def esymmAlgHomMonomial (t : Fin n →₀ ℕ) (r : R) : MvPolynomial σ R := (esymmAlgHom σ R n <| monomial t r).val -variable {i : Fin n} {j : Fin m} {r : R} +variable {i : Fin n} {r : R} lemma isSymmetric_esymmAlgHomMonomial (t : Fin n →₀ ℕ) (r : R) : (esymmAlgHomMonomial σ t r).IsSymmetric := (esymmAlgHom _ _ _ _).2 diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean index ddadb79adcf83..3aa09ead7b4ab 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean @@ -250,17 +250,13 @@ theorem sum_antidiagonal_card_esymm_psum_eq_zero : simp [← sum_filter_add_sum_filter_not (antidiagonal k) (fun a ↦ a.fst < k), ← mul_esymm_eq_sum, mul_add, ← mul_assoc, ← pow_add, mul_comm ↑k (esymm σ R k)] -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ /-- A version of Newton's identities which may be more useful in the case that we know the values of the elementary symmetric polynomials and would like to calculate the values of the power sums. -/ theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : psum σ R k = (-1) ^ (k + 1) * k * esymm σ R k - ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst ∈ Set.Ioo 0 k), (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by - simp only [Set.Ioo, Set.mem_setOf_eq, _root_.and_comm] + simp only [Set.Ioo, Set.mem_setOf_eq, and_comm] have hesymm := mul_esymm_eq_sum σ R k rw [← (sum_filter_add_sum_filter_not ((antidiagonal k).filter (fun a ↦ a.fst < k)) (fun a ↦ 0 < a.fst) (fun a ↦ (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd))] at hesymm diff --git a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean index a7e076eb389ea..d53136ff960c2 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean @@ -7,7 +7,7 @@ Authors: Johan Commelin, Kenny Lau import Mathlib.Algebra.Group.Units import Mathlib.RingTheory.MvPowerSeries.Basic import Mathlib.RingTheory.MvPowerSeries.NoZeroDivisors -import Mathlib.RingTheory.LocalRing.RingHom.Basic +import Mathlib.RingTheory.LocalRing.Basic /-! # Formal (multivariate) power series - Inverses diff --git a/Mathlib/RingTheory/Nilpotent/Basic.lean b/Mathlib/RingTheory/Nilpotent/Basic.lean index 0080e768a4300..71e196c394db9 100644 --- a/Mathlib/RingTheory/Nilpotent/Basic.lean +++ b/Mathlib/RingTheory/Nilpotent/Basic.lean @@ -82,6 +82,15 @@ theorem IsNilpotent.isUnit_add_right_of_commute [Ring R] {r u : R} IsUnit (r + u) := add_comm r u ▸ hnil.isUnit_add_left_of_commute hu h_comm +lemma IsUnit.not_isNilpotent [Ring R] [Nontrivial R] {x : R} (hx : IsUnit x) : + ¬ IsNilpotent x := by + intro H + simpa using H.isUnit_add_right_of_commute hx.neg (by simp) + +lemma IsNilpotent.not_isUnit [Ring R] [Nontrivial R] {x : R} (hx : IsNilpotent x) : + ¬ IsUnit x := + mt IsUnit.not_isNilpotent (by simpa only [not_not] using hx) + instance [Zero R] [Pow R ℕ] [Zero S] [Pow S ℕ] [IsReduced R] [IsReduced S] : IsReduced (R × S) where eq_zero _ := fun ⟨n, hn⟩ ↦ have hn := Prod.ext_iff.1 hn Prod.ext (IsReduced.eq_zero _ ⟨n, hn.1⟩) (IsReduced.eq_zero _ ⟨n, hn.2⟩) diff --git a/Mathlib/RingTheory/Nilpotent/Lemmas.lean b/Mathlib/RingTheory/Nilpotent/Lemmas.lean index d5d61f66838de..fdbae3b577172 100644 --- a/Mathlib/RingTheory/Nilpotent/Lemmas.lean +++ b/Mathlib/RingTheory/Nilpotent/Lemmas.lean @@ -31,16 +31,6 @@ theorem isRadical_iff_span_singleton [CommSemiring R] : simp_rw [IsRadical, ← Ideal.mem_span_singleton] exact forall_swap.trans (forall_congr' fun r => exists_imp.symm) -namespace Commute - -section Semiring - -variable [Semiring R] (h_comm : Commute x y) - -end Semiring - -end Commute - section CommSemiring variable [CommSemiring R] {x y : R} @@ -99,13 +89,31 @@ end LinearMap namespace Module.End -lemma isNilpotent.restrict {R M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] +section + +variable {M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] + +lemma isNilpotent_restrict_of_le {f : End R M} {p q : Submodule R M} + {hp : MapsTo f p p} {hq : MapsTo f q q} (h : p ≤ q) (hf : IsNilpotent (f.restrict hq)) : + IsNilpotent (f.restrict hp) := by + obtain ⟨n, hn⟩ := hf + use n + ext ⟨x, hx⟩ + replace hn := DFunLike.congr_fun hn ⟨x, h hx⟩ + simp_rw [LinearMap.zero_apply, ZeroMemClass.coe_zero, ZeroMemClass.coe_eq_zero] at hn ⊢ + rw [LinearMap.pow_restrict, LinearMap.restrict_apply] at hn ⊢ + ext + exact (congr_arg Subtype.val hn : _) + +lemma isNilpotent.restrict {f : M →ₗ[R] M} {p : Submodule R M} (hf : MapsTo f p p) (hnil : IsNilpotent f) : IsNilpotent (f.restrict hf) := by obtain ⟨n, hn⟩ := hnil exact ⟨n, LinearMap.ext fun m ↦ by simp only [LinearMap.pow_restrict n, hn, LinearMap.restrict_apply, LinearMap.zero_apply]; rfl⟩ +end + variable {M : Type v} [Ring R] [AddCommGroup M] [Module R M] variable {f : Module.End R M} {p : Submodule R M} (hp : p ≤ p.comap f) diff --git a/Mathlib/RingTheory/Noetherian.lean b/Mathlib/RingTheory/Noetherian.lean index d89b6584db58f..903c8862e2b78 100644 --- a/Mathlib/RingTheory/Noetherian.lean +++ b/Mathlib/RingTheory/Noetherian.lean @@ -144,6 +144,10 @@ variable {R M N : Type*} variable [Semiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] variable (R M) +-- see Note [lower instance priority] +instance (priority := 80) _root_.isNoetherian_of_finite [Finite M] : IsNoetherian R M := + ⟨fun s => ⟨(s : Set M).toFinite.toFinset, by rw [Set.Finite.coe_toFinset, Submodule.span_eq]⟩⟩ + -- see Note [lower instance priority] instance (priority := 100) IsNoetherian.finite [IsNoetherian R M] : Module.Finite R M := ⟨IsNoetherian.noetherian ⊤⟩ @@ -184,77 +188,13 @@ instance isNoetherian_prod [IsNoetherian R M] [IsNoetherian R P] : IsNoetherian fun x ⟨_, hx2⟩ => ⟨x.1, Prod.ext rfl <| Eq.symm <| LinearMap.mem_ker.1 hx2⟩ Submodule.map_comap_eq_self this ▸ (noetherian _).map _⟩ -instance isNoetherian_pi {R ι : Type*} {M : ι → Type*} - [Ring R] [∀ i, AddCommGroup (M i)] [∀ i, Module R (M i)] [Finite ι] - [∀ i, IsNoetherian R (M i)] : IsNoetherian R (∀ i, M i) := by - cases nonempty_fintype ι - haveI := Classical.decEq ι - suffices on_finset : ∀ s : Finset ι, IsNoetherian R (∀ i : s, M i) by - let coe_e := Equiv.subtypeUnivEquiv <| @Finset.mem_univ ι _ - letI : IsNoetherian R (∀ i : Finset.univ, M (coe_e i)) := on_finset Finset.univ - exact isNoetherian_of_linearEquiv (LinearEquiv.piCongrLeft R M coe_e) - intro s - induction' s using Finset.induction with a s has ih - · exact ⟨fun s => by - have : s = ⊥ := by simp only [eq_iff_true_of_subsingleton] - rw [this] - apply Submodule.fg_bot⟩ - refine - @isNoetherian_of_linearEquiv R (M a × ((i : s) → M i)) _ _ _ _ _ _ ?_ <| - @isNoetherian_prod R (M a) _ _ _ _ _ _ _ ih - refine - { toFun := fun f i => - (Finset.mem_insert.1 i.2).by_cases - (fun h : i.1 = a => show M i.1 from Eq.recOn h.symm f.1) - (fun h : i.1 ∈ s => show M i.1 from f.2 ⟨i.1, h⟩), - invFun := fun f => - (f ⟨a, Finset.mem_insert_self _ _⟩, fun i => f ⟨i.1, Finset.mem_insert_of_mem i.2⟩), - map_add' := ?_, - map_smul' := ?_ - left_inv := ?_, - right_inv := ?_ } - · intro f g - ext i - unfold Or.by_cases - cases' i with i hi - rcases Finset.mem_insert.1 hi with (rfl | h) - · change _ = _ + _ - simp only [dif_pos] - rfl - · change _ = _ + _ - have : ¬i = a := by - rintro rfl - exact has h - simp only [dif_neg this, dif_pos h] - rfl - · intro c f - ext i - unfold Or.by_cases - cases' i with i hi - rcases Finset.mem_insert.1 hi with (rfl | h) - · dsimp - simp only [dif_pos] - · dsimp - have : ¬i = a := by - rintro rfl - exact has h - simp only [dif_neg this, dif_pos h] - · intro f - apply Prod.ext - · simp only [Or.by_cases, dif_pos] - · ext ⟨i, his⟩ - have : ¬i = a := by - rintro rfl - exact has his - simp only [Or.by_cases, this, not_false_iff, dif_neg] - · intro f - ext ⟨i, hi⟩ - rcases Finset.mem_insert.1 hi with (rfl | h) - · simp only [Or.by_cases, dif_pos] - · have : ¬i = a := by - rintro rfl - exact has h - simp only [Or.by_cases, dif_neg this, dif_pos h] +instance isNoetherian_pi {R ι : Type*} [Finite ι] : + ∀ {M : ι → Type*} [Ring R] [∀ i, AddCommGroup (M i)] + [∀ i, Module R (M i)] [∀ i, IsNoetherian R (M i)], IsNoetherian R (∀ i, M i) := by + apply Finite.induction_empty_option _ _ _ ι + · exact fun e h ↦ isNoetherian_of_linearEquiv (LinearEquiv.piCongrLeft R _ e) + · infer_instance + · exact fun ih ↦ isNoetherian_of_linearEquiv (LinearEquiv.piOptionEquivProd R).symm /-- A version of `isNoetherian_pi` for non-dependent functions. We need this instance because sometimes Lean fails to apply the dependent version in non-dependent settings (e.g., it fails to @@ -289,19 +229,19 @@ universe w variable {R M P : Type*} {N : Type w} [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid N] [Module R N] [AddCommMonoid P] [Module R P] -theorem isNoetherian_iff : - IsNoetherian R M ↔ WellFounded ((· > ·) : Submodule R M → Submodule R M → Prop) := by - have := (CompleteLattice.wellFounded_characterisations <| Submodule R M).out 0 3 +theorem isNoetherian_iff' : IsNoetherian R M ↔ WellFoundedGT (Submodule R M) := by + have := (CompleteLattice.wellFoundedGT_characterisations <| Submodule R M).out 0 3 -- Porting note: inlining this makes rw complain about it being a metavariable rw [this] exact ⟨fun ⟨h⟩ => fun k => (fg_iff_compact k).mp (h k), fun h => ⟨fun k => (fg_iff_compact k).mpr (h k)⟩⟩ -alias ⟨IsNoetherian.wf, _⟩ := isNoetherian_iff +theorem isNoetherian_iff : + IsNoetherian R M ↔ WellFounded ((· > ·) : Submodule R M → Submodule R M → Prop) := by + rw [isNoetherian_iff', ← isWellFounded_iff] -theorem isNoetherian_iff' : IsNoetherian R M ↔ WellFoundedGT (Submodule R M) := by - rw [isNoetherian_iff, ← isWellFounded_iff] +alias ⟨IsNoetherian.wf, _⟩ := isNoetherian_iff alias ⟨IsNoetherian.wellFoundedGT, isNoetherian_mk⟩ := isNoetherian_iff' @@ -363,14 +303,13 @@ variable {R M P : Type*} {N : Type w} [Ring R] [AddCommGroup M] [Module R M] [Ad lemma Submodule.finite_ne_bot_of_independent {ι : Type*} {N : ι → Submodule R M} (h : CompleteLattice.Independent N) : Set.Finite {i | N i ≠ ⊥} := - CompleteLattice.WellFounded.finite_ne_bot_of_independent - (IsWellFounded.wf) h + CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent h /-- A linearly-independent family of vectors in a module over a non-trivial ring must be finite if the module is Noetherian. -/ theorem LinearIndependent.finite_of_isNoetherian [Nontrivial R] {ι} {v : ι → M} (hv : LinearIndependent R v) : Finite ι := by - refine CompleteLattice.WellFounded.finite_of_independent IsWellFounded.wf + refine CompleteLattice.WellFoundedGT.finite_of_independent hv.independent_span_singleton fun i contra => ?_ apply hv.ne_zero i @@ -389,15 +328,20 @@ theorem isNoetherian_of_range_eq_ker [IsNoetherian R P] isNoetherian_mk <| wellFounded_gt_exact_sequence (LinearMap.range f) - (Submodule.map (f.ker.liftQ f <| le_rfl)) - (Submodule.comap (f.ker.liftQ f <| le_rfl)) + (Submodule.map (f.ker.liftQ f le_rfl)) + (Submodule.comap (f.ker.liftQ f le_rfl)) (Submodule.comap g.rangeRestrict) (Submodule.map g.rangeRestrict) - (Submodule.gciMapComap <| LinearMap.ker_eq_bot.mp <| - Submodule.ker_liftQ_eq_bot _ _ _ (le_refl _)) + (Submodule.gciMapComap <| LinearMap.ker_eq_bot.mp <| Submodule.ker_liftQ_eq_bot _ _ _ le_rfl) (Submodule.giMapComap g.surjective_rangeRestrict) (by simp [Submodule.map_comap_eq, inf_comm, Submodule.range_liftQ]) (by simp [Submodule.comap_map_eq, h]) +theorem isNoetherian_iff_submodule_quotient (S : Submodule R P) : + IsNoetherian R P ↔ IsNoetherian R S ∧ IsNoetherian R (P ⧸ S) := by + refine ⟨fun _ ↦ ⟨inferInstance, inferInstance⟩, fun ⟨_, _⟩ ↦ ?_⟩ + apply isNoetherian_of_range_eq_ker S.subtype S.mkQ + rw [Submodule.ker_mkQ, Submodule.range_subtype] + /-- For an endomorphism of a Noetherian module, any sufficiently large iterate has disjoint kernel and range. -/ theorem LinearMap.eventually_disjoint_ker_pow_range_pow (f : M →ₗ[R] M) : @@ -500,11 +444,6 @@ theorem isNoetherianRing_iff_ideal_fg (R : Type*) [Semiring R] : IsNoetherianRing R ↔ ∀ I : Ideal R, I.FG := isNoetherianRing_iff.trans isNoetherian_def --- see Note [lower instance priority] -instance (priority := 80) isNoetherian_of_finite (R M) [Finite M] [Semiring R] [AddCommMonoid M] - [Module R M] : IsNoetherian R M := - ⟨fun s => ⟨(s : Set M).toFinite.toFinset, by rw [Set.Finite.coe_toFinset, Submodule.span_eq]⟩⟩ - -- see Note [lower instance priority] /-- Modules over the trivial ring are Noetherian. -/ instance (priority := 100) isNoetherian_of_subsingleton (R M) [Subsingleton R] [Semiring R] diff --git a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean index 69c9fcd2931b4..f71dbda3f4e85 100644 --- a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean @@ -71,7 +71,7 @@ universe u v w section Basic -variable {R : Type u} {S : Type v} {T : Type w} [NonUnitalNonAssocRing R] +variable {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] section NonUnitalSubringClass @@ -124,8 +124,6 @@ end NonUnitalSubringClass end NonUnitalSubringClass -variable [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] - /-- `NonUnitalSubring R` is the type of non-unital subrings of `R`. A non-unital subring of `R` is a subset `s` that is a multiplicative subsemigroup and an additive subgroup. Note in particular that it shares the same 0 as R. -/ @@ -387,7 +385,7 @@ section Hom namespace NonUnitalSubring -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} {SR : Type*} +variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] [FunLike F R S] [NonUnitalRingHomClass F R S] (s : NonUnitalSubring R) @@ -503,10 +501,7 @@ namespace NonUnitalSubring section Order -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} - [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] - [FunLike F R S] [NonUnitalRingHomClass F R S] - (g : S →ₙ+* T) (f : R →ₙ+* S) +variable {R : Type u} [NonUnitalNonAssocRing R] /-! ## bot -/ @@ -640,10 +635,9 @@ end Center /-! ## `NonUnitalSubring` closure of a subset -/ -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} - [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] +variable {F : Type w} {R : Type u} {S : Type v} + [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [FunLike F R S] [NonUnitalRingHomClass F R S] - (g : S →ₙ+* T) (f : R →ₙ+* S) /-- The `NonUnitalSubring` generated by a set. -/ def closure (s : Set R) : NonUnitalSubring R := @@ -910,11 +904,8 @@ end NonUnitalSubring namespace NonUnitalRingHom -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} - [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] - [FunLike F R S] [NonUnitalRingHomClass F R S] - (g : S →ₙ+* T) (f : R →ₙ+* S) - {s : NonUnitalSubring R} +variable {R : Type u} {S : Type v} + [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] open NonUnitalSubring @@ -980,11 +971,8 @@ end NonUnitalRingHom namespace NonUnitalSubring -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} - [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] - [FunLike F R S] [NonUnitalRingHomClass F R S] - (g : S →ₙ+* T) (f : R →ₙ+* S) - {s : NonUnitalSubring R} +variable {R : Type u} {S : Type v} + [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] open NonUnitalRingHom @@ -1006,11 +994,7 @@ end NonUnitalSubring namespace RingEquiv -variable {F : Type w} {R : Type u} {S : Type v} {T : Type*} - [NonUnitalRing R] [NonUnitalRing S] [NonUnitalRing T] - [FunLike F R S] [NonUnitalRingHomClass F R S] - (g : S →ₙ+* T) (f : R →ₙ+* S) - {s t : NonUnitalSubring R} +variable {R : Type u} {S : Type v} [NonUnitalRing R] [NonUnitalRing S] {s t : NonUnitalSubring R} /-- Makes the identity isomorphism from a proof two `NonUnitalSubring`s of a multiplicative monoid are equal. -/ diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean index 43e050ec7489d..4a1ad15664ee8 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean @@ -104,17 +104,12 @@ theorem cyclotomic_pos {n : ℕ} (hn : 2 < n) {R} [LinearOrderedCommRing R] (x : exact hn'.ne' hi.2.2.1 · simpa only [eval_X, eval_one, cyclotomic_two, eval_add] using h.right.le -#adaptation_note -/-- -After nightly-2024-09-06 we can remove the `_root_` prefix below. --/ theorem cyclotomic_pos_and_nonneg (n : ℕ) {R} [LinearOrderedCommRing R] (x : R) : (1 < x → 0 < eval x (cyclotomic n R)) ∧ (1 ≤ x → 0 ≤ eval x (cyclotomic n R)) := by rcases n with (_ | _ | _ | n) - · simp only [cyclotomic_zero, eval_one, zero_lt_one, implies_true, zero_le_one, - _root_.and_self] + · simp only [cyclotomic_zero, eval_one, zero_lt_one, implies_true, zero_le_one, and_self] · simp only [zero_add, cyclotomic_one, eval_sub, eval_X, eval_one, sub_pos, imp_self, sub_nonneg, - _root_.and_self] + and_self] · simp only [zero_add, reduceAdd, cyclotomic_two, eval_add, eval_X, eval_one] constructor <;> intro <;> linarith · constructor <;> intro <;> [skip; apply le_of_lt] <;> apply cyclotomic_pos (by omega) diff --git a/Mathlib/RingTheory/SimpleModule.lean b/Mathlib/RingTheory/SimpleModule.lean index 86e5d7e807588..9a27c4b72d304 100644 --- a/Mathlib/RingTheory/SimpleModule.lean +++ b/Mathlib/RingTheory/SimpleModule.lean @@ -338,7 +338,7 @@ variable (ι R) proof_wanted IsSemisimpleRing.mulOpposite [IsSemisimpleRing R] : IsSemisimpleRing Rᵐᵒᵖ -proof_wanted IsSemisimpleRing.module_end [IsSemisimpleRing R] [Module.Finite R M] : +proof_wanted IsSemisimpleRing.module_end [IsSemisimpleModule R M] [Module.Finite R M] : IsSemisimpleRing (Module.End R M) proof_wanted IsSemisimpleRing.matrix [Fintype ι] [DecidableEq ι] [IsSemisimpleRing R] : diff --git a/Mathlib/RingTheory/Smooth/Basic.lean b/Mathlib/RingTheory/Smooth/Basic.lean index d61b49519ff71..179e34269a3af 100644 --- a/Mathlib/RingTheory/Smooth/Basic.lean +++ b/Mathlib/RingTheory/Smooth/Basic.lean @@ -60,7 +60,7 @@ section variable {R : Type u} [CommSemiring R] variable {A : Type u} [Semiring A] [Algebra R A] -variable {B : Type u} [CommRing B] [Algebra R B] (I : Ideal B) +variable {B : Type u} [CommRing B] [Algebra R B] theorem exists_lift {B : Type u} [CommRing B] [_RB : Algebra R B] [FormallySmooth R A] (I : Ideal B) (hI : IsNilpotent I) (g : A →ₐ[R] B ⧸ I) : @@ -187,9 +187,9 @@ end Comp section OfSurjective -variable {R S : Type u} [CommRing R] [CommSemiring S] +variable {R : Type u} [CommRing R] variable {P A : Type u} [CommRing A] [Algebra R A] [CommRing P] [Algebra R P] -variable (I : Ideal P) (f : P →ₐ[R] A) +variable (f : P →ₐ[R] A) theorem of_split [FormallySmooth R P] (g : A →ₐ[R] P ⧸ (RingHom.ker f.toRingHom) ^ 2) (hg : f.kerSquareLift.comp g = AlgHom.id R A) : FormallySmooth R A := by diff --git a/Mathlib/RingTheory/SurjectiveOnStalks.lean b/Mathlib/RingTheory/SurjectiveOnStalks.lean index 67f0451472874..94c5443047b71 100644 --- a/Mathlib/RingTheory/SurjectiveOnStalks.lean +++ b/Mathlib/RingTheory/SurjectiveOnStalks.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.RingTheory.Localization.AtPrime -import Mathlib.RingTheory.LocalRing.RingHom.Basic import Mathlib.RingTheory.TensorProduct.Basic /-! diff --git a/Mathlib/RingTheory/Valuation/Basic.lean b/Mathlib/RingTheory/Valuation/Basic.lean index 83c9f045e79f8..eec13a2b17f5f 100644 --- a/Mathlib/RingTheory/Valuation/Basic.lean +++ b/Mathlib/RingTheory/Valuation/Basic.lean @@ -207,6 +207,9 @@ theorem zero_iff [Nontrivial Γ₀] (v : Valuation K Γ₀) {x : K} : v x = 0 theorem ne_zero_iff [Nontrivial Γ₀] (v : Valuation K Γ₀) {x : K} : v x ≠ 0 ↔ x ≠ 0 := map_ne_zero v +lemma pos_iff [Nontrivial Γ₀] (v : Valuation K Γ₀) {x : K} : 0 < v x ↔ x ≠ 0 := by + rw [zero_lt_iff, ne_zero_iff] + theorem unit_map_eq (u : Rˣ) : (Units.map (v : R →* Γ₀) u : Γ₀) = v u := rfl @@ -322,17 +325,16 @@ theorem map_one_sub_of_lt (h : v x < 1) : v (1 - x) = 1 := by simpa only [v.map_one, v.map_neg] using v.map_add_eq_of_lt_left h theorem one_lt_val_iff (v : Valuation K Γ₀) {x : K} (h : x ≠ 0) : 1 < v x ↔ v x⁻¹ < 1 := by - simpa using (inv_lt_inv₀ (v.ne_zero_iff.2 h) one_ne_zero).symm + simp [inv_lt_one₀ (v.pos_iff.2 h)] theorem one_le_val_iff (v : Valuation K Γ₀) {x : K} (h : x ≠ 0) : 1 ≤ v x ↔ v x⁻¹ ≤ 1 := by - convert (one_lt_val_iff v (inv_ne_zero h)).symm.not <;> - push_neg <;> simp only [inv_inv] + simp [inv_le_one₀ (v.pos_iff.2 h)] theorem val_lt_one_iff (v : Valuation K Γ₀) {x : K} (h : x ≠ 0) : v x < 1 ↔ 1 < v x⁻¹ := by - simpa only [inv_inv] using (one_lt_val_iff v (inv_ne_zero h)).symm + simp [one_lt_inv₀ (v.pos_iff.2 h)] theorem val_le_one_iff (v : Valuation K Γ₀) {x : K} (h : x ≠ 0) : v x ≤ 1 ↔ 1 ≤ v x⁻¹ := by - simpa [inv_inv] using (one_le_val_iff v (inv_ne_zero h)).symm + simp [one_le_inv₀ (v.pos_iff.2 h)] theorem val_eq_one_iff (v : Valuation K Γ₀) {x : K} : v x = 1 ↔ v x⁻¹ = 1 := by by_cases h : x = 0 diff --git a/Mathlib/RingTheory/Valuation/ValExtension.lean b/Mathlib/RingTheory/Valuation/ValExtension.lean index 888369fc04bb9..68c03f9397d28 100644 --- a/Mathlib/RingTheory/Valuation/ValExtension.lean +++ b/Mathlib/RingTheory/Valuation/ValExtension.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jiedong Jiang, Bichang Lei -/ import Mathlib.RingTheory.Valuation.Integers -import Mathlib.RingTheory.LocalRing.RingHom.Basic -import Mathlib.RingTheory.LocalRing.RingHom.Defs +import Mathlib.Algebra.Group.Units.Hom /-! # Extension of Valuation diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index 066bed0604120..a66cc1e821359 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -300,9 +300,8 @@ theorem ofPrime_idealOfLE (R S : ValuationSubring K) (h : R ≤ S) : · rintro ⟨a, r, hr, rfl⟩; apply mul_mem; · exact h a.2 · rw [← valuation_le_one_iff, map_inv₀, ← inv_one, inv_le_inv₀] · exact not_lt.1 ((not_iff_not.2 <| valuation_lt_one_iff S _).1 hr) - · intro hh; erw [Valuation.zero_iff, Subring.coe_eq_zero_iff] at hh - apply hr; rw [hh]; apply Ideal.zero_mem (R.idealOfLE S h) - · exact one_ne_zero + · simpa [Valuation.pos_iff] using fun hr₀ ↦ hr₀ ▸ hr <| Ideal.zero_mem (R.idealOfLE S h) + · exact zero_lt_one · intro hx; by_cases hr : x ∈ R; · exact R.le_ofPrime _ hr have : x ≠ 0 := fun h => hr (by rw [h]; exact R.zero_mem) replace hr := (R.mem_or_inv_mem x).resolve_left hr diff --git a/Mathlib/RingTheory/WittVector/Basic.lean b/Mathlib/RingTheory/WittVector/Basic.lean index 5630d6a00254e..e2c5bfed3057c 100644 --- a/Mathlib/RingTheory/WittVector/Basic.lean +++ b/Mathlib/RingTheory/WittVector/Basic.lean @@ -49,7 +49,7 @@ noncomputable section open MvPolynomial Function -variable {p : ℕ} {R S T : Type*} [CommRing R] [CommRing S] [CommRing T] +variable {p : ℕ} {R S : Type*} [CommRing R] [CommRing S] variable {α : Type*} {β : Type*} local notation "𝕎" => WittVector p diff --git a/Mathlib/RingTheory/WittVector/Frobenius.lean b/Mathlib/RingTheory/WittVector/Frobenius.lean index 0d989e4ecdebe..3a15c88d7c9d5 100644 --- a/Mathlib/RingTheory/WittVector/Frobenius.lean +++ b/Mathlib/RingTheory/WittVector/Frobenius.lean @@ -48,7 +48,7 @@ and bundle it into `WittVector.frobenius`. namespace WittVector -variable {p : ℕ} {R S : Type*} [hp : Fact p.Prime] [CommRing R] [CommRing S] +variable {p : ℕ} {R : Type*} [hp : Fact p.Prime] [CommRing R] local notation "𝕎" => WittVector p -- type as `\bbW` diff --git a/Mathlib/RingTheory/WittVector/IsPoly.lean b/Mathlib/RingTheory/WittVector/IsPoly.lean index ef5430b614c18..2e6c8619172e2 100644 --- a/Mathlib/RingTheory/WittVector/IsPoly.lean +++ b/Mathlib/RingTheory/WittVector/IsPoly.lean @@ -92,7 +92,7 @@ namespace WittVector universe u -variable {p : ℕ} {R S : Type u} {σ idx : Type*} [CommRing R] [CommRing S] +variable {p : ℕ} {R S : Type u} {idx : Type*} [CommRing R] [CommRing S] local notation "𝕎" => WittVector p -- type as `\bbW` diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 67f6202f93e56..e315b75b75233 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1370,6 +1370,9 @@ lemma two_le_iff_one_lt {c : Cardinal} : 2 ≤ c ↔ 1 < c := by @[simp] theorem succ_zero : succ (0 : Cardinal) = 1 := by norm_cast +-- This works generally to prove inequalities between numeric cardinals. +theorem one_lt_two : (1 : Cardinal) < 2 := by norm_cast + theorem exists_finset_le_card (α : Type*) (n : ℕ) (h : n ≤ #α) : ∃ s : Finset α, n ≤ s.card := by obtain hα|hα := finite_or_infinite α diff --git a/Mathlib/SetTheory/Cardinal/Ordinal.lean b/Mathlib/SetTheory/Cardinal/Ordinal.lean index fe8a2ac903837..2126055537d6c 100644 --- a/Mathlib/SetTheory/Cardinal/Ordinal.lean +++ b/Mathlib/SetTheory/Cardinal/Ordinal.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn import Mathlib.Order.Bounded import Mathlib.SetTheory.Cardinal.PartENat import Mathlib.SetTheory.Ordinal.Principal +import Mathlib.SetTheory.Ordinal.Enum import Mathlib.Tactic.Linarith /-! @@ -79,7 +80,7 @@ def aleph' : Ordinal.{u} ≃o Cardinal.{u} := by refine (OrderIso.ofRelIsoLT <| RelIso.ofSurjective f ?_).symm apply f.eq_or_principal.resolve_right rintro ⟨o, e⟩ - have : ∀ c, f c < o := fun c => (e _).2 ⟨_, rfl⟩ + have : ∀ c, f c < o := fun c => (e _).1 ⟨_, rfl⟩ refine Ordinal.inductionOn o ?_ this intro α r _ h let s := ⨆ a, invFun f (Ordinal.typein r a) @@ -325,7 +326,17 @@ theorem aleph0_lt_aleph_one : ℵ₀ < aleph 1 := by theorem countable_iff_lt_aleph_one {α : Type*} (s : Set α) : s.Countable ↔ #s < aleph 1 := by rw [← succ_aleph0, lt_succ_iff, le_aleph0_iff_set_countable] +section deprecated + +set_option linter.deprecated false + +-- TODO: these lemmas should be stated in terms of the `ω` function and of an `IsInitial` predicate, +-- neither of which currently exist. +-- +-- They should also use `¬ BddAbove` instead of `Unbounded (· < ·)`. + /-- Ordinals that are cardinals are unbounded. -/ +@[deprecated (since := "2024-09-24")] theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b } := unbounded_lt_iff.2 fun a => ⟨_, @@ -333,22 +344,16 @@ theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b dsimp rw [card_ord], (lt_ord_succ_card a).le⟩⟩ +@[deprecated (since := "2024-09-24")] theorem eq_aleph'_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) : ∃ a, (aleph' a).ord = o := ⟨aleph'.symm o.card, by simpa using ho⟩ -/-- `ord ∘ aleph'` enumerates the ordinals that are cardinals. -/ -theorem ord_aleph'_eq_enum_card : ord ∘ aleph' = enumOrd { b : Ordinal | b.card.ord = b } := by - rw [← eq_enumOrd _ ord_card_unbounded, range_eq_iff] - exact - ⟨aleph'_isNormal.strictMono, - ⟨fun a => by - dsimp - rw [card_ord], fun b hb => eq_aleph'_of_eq_card_ord hb⟩⟩ - /-- Infinite ordinals that are cardinals are unbounded. -/ +@[deprecated (since := "2024-09-24")] theorem ord_card_unbounded' : Unbounded (· < ·) { b : Ordinal | b.card.ord = b ∧ ω ≤ b } := (unbounded_lt_inter_le ω).2 ord_card_unbounded +@[deprecated (since := "2024-09-24")] theorem eq_aleph_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) (ho' : ω ≤ o) : ∃ a, (aleph a).ord = o := by cases' eq_aleph'_of_eq_card_ord ho with a ha @@ -356,16 +361,7 @@ theorem eq_aleph_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) (ho' : ω rwa [aleph_eq_aleph', Ordinal.add_sub_cancel_of_le] rwa [← aleph0_le_aleph', ← ord_le_ord, ha, ord_aleph0] -/-- `ord ∘ aleph` enumerates the infinite ordinals that are cardinals. -/ -theorem ord_aleph_eq_enum_card : - ord ∘ aleph = enumOrd { b : Ordinal | b.card.ord = b ∧ ω ≤ b } := by - rw [← eq_enumOrd _ ord_card_unbounded'] - use aleph_isNormal.strictMono - rw [range_eq_iff] - refine ⟨fun a => ⟨?_, ?_⟩, fun b hb => eq_aleph_of_eq_card_ord hb.1 hb.2⟩ - · rw [Function.comp_apply, card_ord] - · rw [← ord_aleph0, Function.comp_apply, ord_le_ord] - exact aleph0_le_aleph _ +end deprecated end aleph @@ -1225,209 +1221,6 @@ theorem extend_function_of_lt {α β : Type*} {s : Set α} (f : s ↪ β) (hs : rwa [mk_compl_of_infinite s hs, mk_compl_of_infinite] rwa [← lift_lt, mk_range_eq_of_injective f.injective, ← h, lift_lt] - --- Porting note: we no longer express literals as `bit0` and `bit1` in Lean 4, so we can't use this --- section Bit - --- /-! --- This section proves inequalities for `bit0` and `bit1`, enabling `simp` to solve inequalities --- for numeral cardinals. The complexity of the resulting algorithm is not good, as in some cases --- `simp` reduces an inequality to a disjunction of two situations, depending on whether a cardinal --- is finite or infinite. Since the evaluation of the branches is not lazy, this is bad. It is good --- enough for practical situations, though. - --- For specific numbers, these inequalities could also be deduced from the corresponding --- inequalities of natural numbers using `norm_cast`: --- ``` --- example : (37 : cardinal) < 42 := --- by { norm_cast, norm_num } --- ``` --- -/ - - --- theorem bit0_ne_zero (a : Cardinal) : ¬bit0 a = 0 ↔ ¬a = 0 := by simp [bit0] - --- @[simp] --- theorem bit1_ne_zero (a : Cardinal) : ¬bit1 a = 0 := by simp [bit1] - --- @[simp] --- theorem zero_lt_bit0 (a : Cardinal) : 0 < bit0 a ↔ 0 < a := by --- rw [← not_iff_not] --- simp [bit0] - --- @[simp] --- theorem zero_lt_bit1 (a : Cardinal) : 0 < bit1 a := --- zero_lt_one.trans_le (self_le_add_left _ _) - --- @[simp] --- theorem one_le_bit0 (a : Cardinal) : 1 ≤ bit0 a ↔ 0 < a := --- ⟨fun h => (zero_lt_bit0 a).mp (zero_lt_one.trans_le h), fun h => --- (one_le_iff_pos.mpr h).trans (self_le_add_left a a)⟩ - --- @[simp] --- theorem one_le_bit1 (a : Cardinal) : 1 ≤ bit1 a := --- self_le_add_left _ _ - --- theorem bit0_eq_self {c : Cardinal} (h : ℵ₀ ≤ c) : bit0 c = c := --- add_eq_self h - --- @[simp] --- theorem bit0_lt_aleph0 {c : Cardinal} : bit0 c < ℵ₀ ↔ c < ℵ₀ := --- by simp [bit0, add_lt_aleph_0_iff] - --- @[simp] --- theorem aleph0_le_bit0 {c : Cardinal} : ℵ₀ ≤ bit0 c ↔ ℵ₀ ≤ c := by --- rw [← not_iff_not] --- simp - --- @[simp] --- theorem bit1_eq_self_iff {c : Cardinal} : bit1 c = c ↔ ℵ₀ ≤ c := by --- by_cases h : ℵ₀ ≤ c --- · simp only [bit1, bit0_eq_self h, h, eq_self_iff_true, add_one_of_aleph_0_le] --- · refine' iff_of_false (ne_of_gt _) h --- rcases lt_aleph_0.1 (not_le.1 h) with ⟨n, rfl⟩ --- norm_cast --- dsimp [bit1, bit0] --- linarith - --- @[simp] --- theorem bit1_lt_aleph0 {c : Cardinal} : bit1 c < ℵ₀ ↔ c < ℵ₀ := by --- simp [bit1, bit0, add_lt_aleph_0_iff, one_lt_aleph_0] - --- @[simp] --- theorem aleph0_le_bit1 {c : Cardinal} : ℵ₀ ≤ bit1 c ↔ ℵ₀ ≤ c := by --- rw [← not_iff_not] --- simp - --- @[simp] --- theorem bit0_le_bit0 {a b : Cardinal} : bit0 a ≤ bit0 b ↔ a ≤ b := by --- rcases le_or_lt ℵ₀ a with ha | ha <;> rcases le_or_lt ℵ₀ b with hb | hb --- · rw [bit0_eq_self ha, bit0_eq_self hb] --- · rw [bit0_eq_self ha] --- refine' iff_of_false (fun h => _) (hb.trans_le ha).not_le --- have A : bit0 b < ℵ₀ := by simpa using hb --- exact lt_irrefl _ ((A.trans_le ha).trans_le h) --- · rw [bit0_eq_self hb] --- exact iff_of_true ((bit0_lt_aleph_0.2 ha).le.trans hb) (ha.le.trans hb) --- · rcases lt_aleph_0.1 ha with ⟨m, rfl⟩ --- rcases lt_aleph_0.1 hb with ⟨n, rfl⟩ --- norm_cast --- exact bit0_le_bit0 - --- @[simp] --- theorem bit0_le_bit1 {a b : Cardinal} : bit0 a ≤ bit1 b ↔ a ≤ b := by --- rcases le_or_lt ℵ₀ a with ha | ha <;> rcases le_or_lt ℵ₀ b with hb | hb --- · rw [bit0_eq_self ha, bit1_eq_self_iff.2 hb] --- · rw [bit0_eq_self ha] --- refine' iff_of_false (fun h => _) (hb.trans_le ha).not_le --- have A : bit1 b < ℵ₀ := by simpa using hb --- exact lt_irrefl _ ((A.trans_le ha).trans_le h) --- · rw [bit1_eq_self_iff.2 hb] --- exact iff_of_true ((bit0_lt_aleph_0.2 ha).le.trans hb) (ha.le.trans hb) --- · rcases lt_aleph_0.1 ha with ⟨m, rfl⟩ --- rcases lt_aleph_0.1 hb with ⟨n, rfl⟩ --- norm_cast --- exact Nat.bit0_le_bit1_iff - --- @[simp] --- theorem bit1_le_bit1 {a b : Cardinal} : bit1 a ≤ bit1 b ↔ a ≤ b := --- ⟨fun h => bit0_le_bit1.1 ((self_le_add_right (bit0 a) 1).trans h), fun h => --- (add_le_add_right (add_le_add_left h a) 1).trans (add_le_add_right (add_le_add_right h b) 1)⟩ - --- @[simp] --- theorem bit1_le_bit0 {a b : Cardinal} : bit1 a ≤ bit0 b ↔ a < b ∨ a ≤ b ∧ ℵ₀ ≤ a := by --- rcases le_or_lt ℵ₀ a with ha | ha <;> rcases le_or_lt ℵ₀ b with hb | hb --- · simp only [bit1_eq_self_iff.mpr ha, bit0_eq_self hb, ha, and_true_iff] --- refine' ⟨fun h => Or.inr h, fun h => _⟩ --- cases h --- · exact le_of_lt h --- · exact h --- · rw [bit1_eq_self_iff.2 ha] --- refine' iff_of_false (fun h => _) fun h => _ --- · have A : bit0 b < ℵ₀ := by simpa using hb --- exact lt_irrefl _ ((A.trans_le ha).trans_le h) --- · exact not_le_of_lt (hb.trans_le ha) (h.elim le_of_lt And.left) --- · rw [bit0_eq_self hb] --- exact iff_of_true ((bit1_lt_aleph_0.2 ha).le.trans hb) (Or.inl <| ha.trans_le hb) --- · rcases lt_aleph_0.1 ha with ⟨m, rfl⟩ --- rcases lt_aleph_0.1 hb with ⟨n, rfl⟩ --- norm_cast --- simp [not_le.mpr ha] - --- @[simp] --- theorem bit0_lt_bit0 {a b : Cardinal} : bit0 a < bit0 b ↔ a < b := by --- rcases le_or_lt ℵ₀ a with ha | ha <;> rcases le_or_lt ℵ₀ b with hb | hb --- · rw [bit0_eq_self ha, bit0_eq_self hb] --- · rw [bit0_eq_self ha] --- refine' iff_of_false (fun h => _) (hb.le.trans ha).not_lt --- have A : bit0 b < ℵ₀ := by simpa using hb --- exact lt_irrefl _ ((A.trans_le ha).trans h) --- · rw [bit0_eq_self hb] --- exact iff_of_true ((bit0_lt_aleph_0.2 ha).trans_le hb) (ha.trans_le hb) --- · rcases lt_aleph_0.1 ha with ⟨m, rfl⟩ --- rcases lt_aleph_0.1 hb with ⟨n, rfl⟩ --- norm_cast --- exact bit0_lt_bit0 - --- @[simp] --- theorem bit1_lt_bit0 {a b : Cardinal} : bit1 a < bit0 b ↔ a < b := by --- rcases le_or_lt ℵ₀ a with ha | ha <;> rcases le_or_lt ℵ₀ b with hb | hb --- · rw [bit1_eq_self_iff.2 ha, bit0_eq_self hb] --- · rw [bit1_eq_self_iff.2 ha] --- refine' iff_of_false (fun h => _) (hb.le.trans ha).not_lt --- have A : bit0 b < ℵ₀ := by simpa using hb --- exact lt_irrefl _ ((A.trans_le ha).trans h) --- · rw [bit0_eq_self hb] --- exact iff_of_true ((bit1_lt_aleph_0.2 ha).trans_le hb) (ha.trans_le hb) --- · rcases lt_aleph_0.1 ha with ⟨m, rfl⟩ --- rcases lt_aleph_0.1 hb with ⟨n, rfl⟩ --- norm_cast --- exact Nat.bit1_lt_bit0_iff - --- @[simp] --- theorem bit1_lt_bit1 {a b : Cardinal} : bit1 a < bit1 b ↔ a < b := by --- rcases le_or_lt ℵ₀ a with ha | ha <;> rcases le_or_lt ℵ₀ b with hb | hb --- · rw [bit1_eq_self_iff.2 ha, bit1_eq_self_iff.2 hb] --- · rw [bit1_eq_self_iff.2 ha] --- refine' iff_of_false (fun h => _) (hb.le.trans ha).not_lt --- have A : bit1 b < ℵ₀ := by simpa using hb --- exact lt_irrefl _ ((A.trans_le ha).trans h) --- · rw [bit1_eq_self_iff.2 hb] --- exact iff_of_true ((bit1_lt_aleph_0.2 ha).trans_le hb) (ha.trans_le hb) --- · rcases lt_aleph_0.1 ha with ⟨m, rfl⟩ --- rcases lt_aleph_0.1 hb with ⟨n, rfl⟩ --- norm_cast --- exact bit1_lt_bit1 - --- @[simp] --- theorem bit0_lt_bit1 {a b : Cardinal} : bit0 a < bit1 b ↔ a < b ∨ a ≤ b ∧ a < ℵ₀ := by --- rcases le_or_lt ℵ₀ a with ha | ha <;> rcases le_or_lt ℵ₀ b with hb | hb --- · simp [bit0_eq_self ha, bit1_eq_self_iff.2 hb, not_lt.mpr ha] --- · rw [bit0_eq_self ha] --- refine' iff_of_false (fun h => _) fun h => _ --- · have A : bit1 b < ℵ₀ := by simpa using hb --- exact lt_irrefl _ ((A.trans_le ha).trans h) --- · exact (hb.trans_le ha).not_le (h.elim le_of_lt And.left) --- · rw [bit1_eq_self_iff.2 hb] --- exact iff_of_true ((bit0_lt_aleph_0.2 ha).trans_le hb) (Or.inl <| ha.trans_le hb) --- · rcases lt_aleph_0.1 ha with ⟨m, rfl⟩ --- rcases lt_aleph_0.1 hb with ⟨n, rfl⟩ --- norm_cast --- simp only [ha, and_true_iff, Nat.bit0_lt_bit1_iff, or_iff_right_of_imp le_of_lt] - --- theorem one_lt_two : (1 : Cardinal) < 2 := by --- -- This strategy works generally to prove inequalities between numerals in `cardinality`. --- norm_cast --- norm_num - --- @[simp] --- theorem one_lt_bit0 {a : Cardinal} : 1 < bit0 a ↔ 0 < a := by simp [← bit1_zero] - --- @[simp] --- theorem one_lt_bit1 (a : Cardinal) : 1 < bit1 a ↔ 0 < a := by simp [← bit1_zero] - --- end Bit - end Cardinal section Initial diff --git a/Mathlib/SetTheory/Game/Birthday.lean b/Mathlib/SetTheory/Game/Birthday.lean index 016fd8366eb3a..c211b6f8b75e4 100644 --- a/Mathlib/SetTheory/Game/Birthday.lean +++ b/Mathlib/SetTheory/Game/Birthday.lean @@ -126,7 +126,7 @@ theorem le_birthday : ∀ x : PGame, x ≤ x.birthday.toPGame Or.inl ⟨toLeftMovesToPGame ⟨_, birthday_moveLeft_lt i⟩, by simp [le_birthday (xL i)]⟩, isEmptyElim⟩ -variable (a b x : PGame.{u}) +variable (x : PGame.{u}) theorem neg_birthday_le : -x.birthday.toPGame ≤ x := by simpa only [birthday_neg, ← neg_le_iff] using le_birthday (-x) diff --git a/Mathlib/SetTheory/Game/Nim.lean b/Mathlib/SetTheory/Game/Nim.lean index 46bd9ae9c794e..f563fafa77d4b 100644 --- a/Mathlib/SetTheory/Game/Nim.lean +++ b/Mathlib/SetTheory/Game/Nim.lean @@ -218,6 +218,7 @@ theorem grundyValue_eq_sInf_moveLeft (G : PGame) : grundyValue G = sInf (Set.range (grundyValue ∘ G.moveLeft))ᶜ := by rw [grundyValue]; rfl +set_option linter.deprecated false in @[deprecated grundyValue_eq_sInf_moveLeft (since := "2024-09-16")] theorem grundyValue_eq_mex_left (G : PGame) : grundyValue G = Ordinal.mex fun i => grundyValue (G.moveLeft i) := @@ -302,6 +303,7 @@ theorem grundyValue_eq_sInf_moveRight (G : PGame) [G.Impartial] : ext i exact @grundyValue_neg _ (@Impartial.moveRight_impartial ⟨l, r, L, R⟩ _ _) +set_option linter.deprecated false in @[deprecated grundyValue_eq_sInf_moveRight (since := "2024-09-16")] theorem grundyValue_eq_mex_right (G : PGame) [G.Impartial] : grundyValue G = Ordinal.mex.{u, u} fun i => grundyValue (G.moveRight i) := diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index d17770942775a..83004bbd598e2 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -360,16 +360,28 @@ theorem bounded_singleton {r : α → α → Prop} [IsWellOrder α r] (hr : (typ rw [@enum_lt_enum _ r, Subtype.mk_lt_mk] apply lt_succ --- Porting note: `· < ·` requires a type ascription for an `IsWellOrder` instance. -theorem type_subrel_lt (o : Ordinal.{u}) : - type (@Subrel Ordinal (· < ·) { o' : Ordinal | o' < o }) = Ordinal.lift.{u + 1} o := by +@[simp] +theorem typein_ordinal (o : Ordinal.{u}) : + @typein Ordinal (· < ·) _ o = Ordinal.lift.{u + 1} o := by refine Quotient.inductionOn o ?_ rintro ⟨α, r, wo⟩; apply Quotient.sound constructor; refine ((RelIso.preimage Equiv.ulift r).trans (enum r).symm).symm +-- Porting note: `· < ·` requires a type ascription for an `IsWellOrder` instance. +@[deprecated typein_ordinal (since := "2024-09-19")] +theorem type_subrel_lt (o : Ordinal.{u}) : + type (@Subrel Ordinal (· < ·) { o' : Ordinal | o' < o }) = Ordinal.lift.{u + 1} o := + typein_ordinal o + +theorem mk_Iio_ordinal (o : Ordinal.{u}) : + #(Iio o) = Cardinal.lift.{u + 1} o.card := by + rw [lift_card, ← typein_ordinal] + rfl + +@[deprecated mk_Iio_ordinal (since := "2024-09-19")] theorem mk_initialSeg (o : Ordinal.{u}) : - #{ o' : Ordinal | o' < o } = Cardinal.lift.{u + 1} o.card := by - rw [lift_card, ← type_subrel_lt, card_type] + #{ o' : Ordinal | o' < o } = Cardinal.lift.{u + 1} o.card := mk_Iio_ordinal o + /-! ### Normal ordinal functions -/ @@ -1373,6 +1385,22 @@ theorem iSup_ord {ι} {f : ι → Cardinal} (hf : BddAbove (range f)) : conv_lhs => change range (ord ∘ f) rw [range_comp] +theorem sInf_compl_lt_lift_ord_succ {ι : Type u} (f : ι → Ordinal.{max u v}) : + sInf (range f)ᶜ < lift.{v} (succ #ι).ord := by + by_contra! h + have : Iio (lift.{v} (succ #ι).ord) ⊆ range f := by + intro o ho + have := not_mem_of_lt_csInf' (ho.trans_le h) + rwa [not_mem_compl_iff] at this + have := mk_le_mk_of_subset this + rw [mk_Iio_ordinal, ← lift_card, Cardinal.lift_lift, card_ord, Cardinal.lift_succ, + succ_le_iff, ← Cardinal.lift_id'.{u, max (u + 1) (v + 1)} #_] at this + exact this.not_le mk_range_le_lift + +theorem sInf_compl_lt_ord_succ {ι : Type u} (f : ι → Ordinal.{u}) : + sInf (range f)ᶜ < (succ #ι).ord := + lift_id (succ #ι).ord ▸ sInf_compl_lt_lift_ord_succ f + -- TODO: remove `bsup` in favor of `iSup` in a future refactor. section bsup @@ -1918,30 +1946,38 @@ set_option linter.deprecated false /-- The minimum excluded ordinal in a family of ordinals. -/ +@[deprecated "use sInf sᶜ instead" (since := "2024-09-20")] def mex {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal := sInf (Set.range f)ᶜ +@[deprecated (since := "2024-09-20")] theorem mex_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ∉ Set.range f := csInf_mem (nonempty_compl_range.{_, v} f) +@[deprecated (since := "2024-09-20")] theorem le_mex_of_forall {ι : Type u} {f : ι → Ordinal.{max u v}} {a : Ordinal} (H : ∀ b < a, ∃ i, f i = b) : a ≤ mex.{_, v} f := by by_contra! h exact mex_not_mem_range f (H _ h) +@[deprecated (since := "2024-09-20")] theorem ne_mex {ι : Type u} (f : ι → Ordinal.{max u v}) : ∀ i, f i ≠ mex.{_, v} f := by simpa using mex_not_mem_range.{_, v} f +@[deprecated (since := "2024-09-20")] theorem mex_le_of_ne {ι} {f : ι → Ordinal} {a} (ha : ∀ i, f i ≠ a) : mex f ≤ a := csInf_le' (by simp [ha]) +@[deprecated (since := "2024-09-20")] theorem exists_of_lt_mex {ι} {f : ι → Ordinal} {a} (ha : a < mex f) : ∃ i, f i = a := by by_contra! ha' exact ha.not_le (mex_le_of_ne ha') +@[deprecated (since := "2024-09-20")] theorem mex_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ≤ lsub.{_, v} f := csInf_le' (lsub_not_mem_range f) +@[deprecated (since := "2024-09-20")] theorem mex_monotone {α β : Type u} {f : α → Ordinal.{max u v}} {g : β → Ordinal.{max u v}} (h : Set.range f ⊆ Set.range g) : mex.{_, v} f ≤ mex.{_, v} g := by refine mex_le_of_ne fun i hi => ?_ @@ -1949,6 +1985,7 @@ theorem mex_monotone {α β : Type u} {f : α → Ordinal.{max u v}} {g : β → rw [← hj] at hi exact ne_mex g j hi +@[deprecated sInf_compl_lt_ord_succ (since := "2024-09-20")] theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : mex.{_, u} f < (succ #ι).ord := by by_contra! h @@ -1969,18 +2006,22 @@ theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : `familyOfBFamily`. This is to `mex` as `bsup` is to `sup`. -/ +@[deprecated "use sInf sᶜ instead" (since := "2024-09-20")] def bmex (o : Ordinal) (f : ∀ a < o, Ordinal) : Ordinal := mex (familyOfBFamily o f) +@[deprecated (since := "2024-09-20")] theorem bmex_not_mem_brange {o : Ordinal} (f : ∀ a < o, Ordinal) : bmex o f ∉ brange o f := by rw [← range_familyOfBFamily] apply mex_not_mem_range +@[deprecated (since := "2024-09-20")] theorem le_bmex_of_forall {o : Ordinal} (f : ∀ a < o, Ordinal) {a : Ordinal} (H : ∀ b < a, ∃ i hi, f i hi = b) : a ≤ bmex o f := by by_contra! h exact bmex_not_mem_brange f (H _ h) +@[deprecated (since := "2024-09-20")] theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : f i hi ≠ bmex.{_, v} o f := by convert (config := {transparency := .default}) @@ -1988,24 +2029,29 @@ theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : -- Porting note: `familyOfBFamily_enum` → `typein_enum` rw [typein_enum] +@[deprecated (since := "2024-09-20")] theorem bmex_le_of_ne {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : ∀ i hi, f i hi ≠ a) : bmex o f ≤ a := mex_le_of_ne fun _i => ha _ _ +@[deprecated (since := "2024-09-20")] theorem exists_of_lt_bmex {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : a < bmex o f) : ∃ i hi, f i hi = a := by cases' exists_of_lt_mex ha with i hi exact ⟨_, typein_lt_self i, hi⟩ +@[deprecated (since := "2024-09-20")] theorem bmex_le_blsub {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : bmex.{_, v} o f ≤ blsub.{_, v} o f := mex_le_lsub _ +@[deprecated (since := "2024-09-20")] theorem bmex_monotone {o o' : Ordinal.{u}} {f : ∀ a < o, Ordinal.{max u v}} {g : ∀ a < o', Ordinal.{max u v}} (h : brange o f ⊆ brange o' g) : bmex.{_, v} o f ≤ bmex.{_, v} o' g := mex_monotone (by rwa [range_familyOfBFamily, range_familyOfBFamily]) +@[deprecated sInf_compl_lt_ord_succ (since := "2024-09-20")] theorem bmex_lt_ord_succ_card {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{u}) : bmex.{_, u} o f < (succ o.card).ord := by rw [← mk_toType] @@ -2043,134 +2089,11 @@ theorem Ordinal.not_bddAbove_compl_of_small (s : Set Ordinal.{u}) [hs : Small.{u rw [union_compl_self, small_univ_iff] at this exact not_small_ordinal this -/-! ### Enumerating unbounded sets of ordinals with ordinals -/ +/-! ### Casting naturals into ordinals, compatibility with operations -/ namespace Ordinal -section - -/-- Enumerator function for an unbounded set of ordinals. -/ -def enumOrd (S : Set Ordinal.{u}) : Ordinal → Ordinal := - lt_wf.fix fun o f => sInf (S ∩ Set.Ici (blsub.{u, u} o f)) - -variable {S : Set Ordinal.{u}} - -/-- The equation that characterizes `enumOrd` definitionally. This isn't the nicest expression to - work with, so consider using `enumOrd_def` instead. -/ -theorem enumOrd_def' (o) : - enumOrd S o = sInf (S ∩ Set.Ici (blsub.{u, u} o fun a _ => enumOrd S a)) := - lt_wf.fix_eq _ _ - -/-- The set in `enumOrd_def'` is nonempty. -/ -theorem enumOrd_def'_nonempty (hS : Unbounded (· < ·) S) (a) : (S ∩ Set.Ici a).Nonempty := - let ⟨b, hb, hb'⟩ := hS a - ⟨b, hb, le_of_not_gt hb'⟩ - -private theorem enumOrd_mem_aux (hS : Unbounded (· < ·) S) (o) : - enumOrd S o ∈ S ∩ Set.Ici (blsub.{u, u} o fun c _ => enumOrd S c) := by - rw [enumOrd_def'] - exact csInf_mem (enumOrd_def'_nonempty hS _) - -theorem enumOrd_mem (hS : Unbounded (· < ·) S) (o) : enumOrd S o ∈ S := - (enumOrd_mem_aux hS o).left - -theorem blsub_le_enumOrd (hS : Unbounded (· < ·) S) (o) : - (blsub.{u, u} o fun c _ => enumOrd S c) ≤ enumOrd S o := - (enumOrd_mem_aux hS o).right - -theorem enumOrd_strictMono (hS : Unbounded (· < ·) S) : StrictMono (enumOrd S) := fun _ _ h => - (lt_blsub.{u, u} _ _ h).trans_le (blsub_le_enumOrd hS _) - -/-- A more workable definition for `enumOrd`. -/ -theorem enumOrd_def (o) : enumOrd S o = sInf (S ∩ { b | ∀ c, c < o → enumOrd S c < b }) := by - rw [enumOrd_def'] - congr; ext - exact ⟨fun h a hao => (lt_blsub.{u, u} _ _ hao).trans_le h, blsub_le⟩ - -/-- The set in `enumOrd_def` is nonempty. -/ -theorem enumOrd_def_nonempty (hS : Unbounded (· < ·) S) {o} : - { x | x ∈ S ∧ ∀ c, c < o → enumOrd S c < x }.Nonempty := - ⟨_, enumOrd_mem hS o, fun _ b => enumOrd_strictMono hS b⟩ - -@[simp] -theorem enumOrd_range {f : Ordinal → Ordinal} (hf : StrictMono f) : enumOrd (range f) = f := - funext fun o => by - apply Ordinal.induction o - intro a H - rw [enumOrd_def a] - have Hfa : f a ∈ range f ∩ { b | ∀ c, c < a → enumOrd (range f) c < b } := - ⟨mem_range_self a, fun b hb => by - rw [H b hb] - exact hf hb⟩ - refine (csInf_le' Hfa).antisymm ((le_csInf_iff'' ⟨_, Hfa⟩).2 ?_) - rintro _ ⟨⟨c, rfl⟩, hc : ∀ b < a, enumOrd (range f) b < f c⟩ - rw [hf.le_iff_le] - contrapose! hc - exact ⟨c, hc, (H c hc).ge⟩ - -@[simp] -theorem enumOrd_univ : enumOrd Set.univ = id := by - rw [← range_id] - exact enumOrd_range strictMono_id - -@[simp] -theorem enumOrd_zero : enumOrd S 0 = sInf S := by - rw [enumOrd_def] - simp [Ordinal.not_lt_zero] - -theorem enumOrd_succ_le {a b} (hS : Unbounded (· < ·) S) (ha : a ∈ S) (hb : enumOrd S b < a) : - enumOrd S (succ b) ≤ a := by - rw [enumOrd_def] - exact - csInf_le' ⟨ha, fun c hc => ((enumOrd_strictMono hS).monotone (le_of_lt_succ hc)).trans_lt hb⟩ - -theorem enumOrd_le_of_subset {S T : Set Ordinal} (hS : Unbounded (· < ·) S) (hST : S ⊆ T) (a) : - enumOrd T a ≤ enumOrd S a := by - apply Ordinal.induction a - intro b H - rw [enumOrd_def] - exact csInf_le' ⟨hST (enumOrd_mem hS b), fun c h => (H c h).trans_lt (enumOrd_strictMono hS h)⟩ - -theorem enumOrd_surjective (hS : Unbounded (· < ·) S) : ∀ s ∈ S, ∃ a, enumOrd S a = s := fun s hs => - ⟨sSup { a | enumOrd S a ≤ s }, by - apply le_antisymm - · rw [enumOrd_def] - refine csInf_le' ⟨hs, fun a ha => ?_⟩ - have : enumOrd S 0 ≤ s := by - rw [enumOrd_zero] - exact csInf_le' hs - -- Porting note: `flip` is required to infer a metavariable. - rcases flip exists_lt_of_lt_csSup ha ⟨0, this⟩ with ⟨b, hb, hab⟩ - exact (enumOrd_strictMono hS hab).trans_le hb - · by_contra! h - exact (le_csSup ⟨s, fun a => ((enumOrd_strictMono hS).id_le a).trans⟩ - (enumOrd_succ_le hS hs h)).not_lt (lt_succ _)⟩ - -/-- An order isomorphism between an unbounded set of ordinals and the ordinals. -/ -def enumOrdOrderIso (hS : Unbounded (· < ·) S) : Ordinal ≃o S := - StrictMono.orderIsoOfSurjective (fun o => ⟨_, enumOrd_mem hS o⟩) (enumOrd_strictMono hS) fun s => - let ⟨a, ha⟩ := enumOrd_surjective hS s s.prop - ⟨a, Subtype.eq ha⟩ - -theorem range_enumOrd (hS : Unbounded (· < ·) S) : range (enumOrd S) = S := by - rw [range_eq_iff] - exact ⟨enumOrd_mem hS, enumOrd_surjective hS⟩ - -/-- A characterization of `enumOrd`: it is the unique strict monotonic function with range `S`. -/ -theorem eq_enumOrd (f : Ordinal → Ordinal) (hS : Unbounded (· < ·) S) : - StrictMono f ∧ range f = S ↔ f = enumOrd S := by - constructor - · rintro ⟨h₁, h₂⟩ - rwa [← h₁.range_inj (enumOrd_strictMono hS), range_enumOrd hS] - · rintro rfl - exact ⟨enumOrd_strictMono hS, range_enumOrd hS⟩ - -end - -/-! ### Casting naturals into ordinals, compatibility with operations -/ - - @[simp] theorem one_add_natCast (m : ℕ) : 1 + (m : Ordinal) = succ m := by rw [← Nat.cast_one, ← Nat.cast_add, add_comm] diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 64a510e18dd89..89c4ad1600e69 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -382,7 +382,7 @@ theorem typein_top {α β} {r : α → α → Prop} {s : β → β → Prop} [Is Eq.symm <| Quot.sound ⟨RelIso.ofSurjective (RelEmbedding.codRestrict _ f f.lt_top) fun ⟨a, h⟩ => by - rcases f.down.1 h with ⟨b, rfl⟩; exact ⟨b, rfl⟩⟩ + rcases f.mem_range_of_rel_top h with ⟨b, rfl⟩; exact ⟨b, rfl⟩⟩ @[simp] theorem typein_apply {α β} {r : α → α → Prop} {s : β → β → Prop} [IsWellOrder α r] [IsWellOrder β s] @@ -426,7 +426,7 @@ principal segment. -/ def typein.principalSeg {α : Type u} (r : α → α → Prop) [IsWellOrder α r] : @PrincipalSeg α Ordinal.{u} r (· < ·) := ⟨⟨⟨typein r, typein_injective r⟩, typein_lt_typein r⟩, type r, - fun _ ↦ ⟨typein_surj r, fun ⟨a, h⟩ ↦ h ▸ typein_lt_type r a⟩⟩ + fun _ ↦ ⟨fun ⟨a, h⟩ ↦ h ▸ typein_lt_type r a, typein_surj r⟩⟩ @[simp] theorem typein.principalSeg_coe (r : α → α → Prop) [IsWellOrder α r] : @@ -856,8 +856,8 @@ private theorem succ_le_iff' {a b : Ordinal} : a + 1 ≤ b ↔ a < b := (inductionOn a fun α r _ => ⟨⟨⟨⟨fun x => Sum.inl x, fun _ _ => Sum.inl.inj⟩, Sum.lex_inl_inl⟩, Sum.inr PUnit.unit, fun b => - Sum.recOn b (fun x => ⟨fun _ => ⟨x, rfl⟩, fun _ => Sum.Lex.sep _ _⟩) fun x => - Sum.lex_inr_inr.trans ⟨False.elim, fun ⟨x, H⟩ => Sum.inl_ne_inr H⟩⟩⟩), + Sum.recOn b (fun x => ⟨fun _ => Sum.Lex.sep _ _, fun _ => ⟨x, rfl⟩⟩) fun x => + (Sum.lex_inr_inr.trans ⟨False.elim, fun ⟨x, H⟩ => Sum.inl_ne_inr H⟩).symm⟩⟩), inductionOn a fun α r hr => inductionOn b fun β s hs ⟨⟨f, t, hf⟩⟩ => by haveI := hs @@ -865,7 +865,7 @@ private theorem succ_le_iff' {a b : Ordinal} : a + 1 ≤ b ↔ a < b := · rcases a with (a | _) <;> rcases b with (b | _) · simpa only [Sum.lex_inl_inl] using f.map_rel_iff.2 · intro - rw [hf] + rw [← hf] exact ⟨_, rfl⟩ · exact False.elim ∘ Sum.lex_inr_inl · exact False.elim ∘ Sum.lex_inr_inr.1 @@ -875,7 +875,7 @@ private theorem succ_le_iff' {a b : Ordinal} : a + 1 ≤ b ↔ a < b := cases' this with w h exact ⟨Sum.inl w, h⟩ · intro h - cases' (hf b).1 h with w h + cases' (hf b).2 h with w h exact ⟨Sum.inl w, h⟩⟩ instance noMaxOrder : NoMaxOrder Ordinal := @@ -1001,7 +1001,7 @@ theorem enum_inj {r : α → α → Prop} [IsWellOrder α r] {o₁ o₂ : {o // rw [EmbeddingLike.apply_eq_iff_eq, Subtype.mk.injEq] /-- The order isomorphism between ordinals less than `o` and `o.toType`. -/ -@[simps!] +@[simps! (config := .lemmasOnly)] noncomputable def enumIsoToType (o : Ordinal) : Set.Iio o ≃o o.toType where toFun x := enum (α := o.toType) (· < ·) ⟨x.1, by @@ -1068,6 +1068,11 @@ def liftPrincipalSeg : @PrincipalSeg Ordinal.{u} Ordinal.{max (u + 1) v} (· < ⟨↑liftInitialSeg.{u, max (u + 1) v}, univ.{u, v}, by refine fun b => inductionOn b ?_; intro β s _ rw [univ, ← lift_umax]; constructor <;> intro h + · cases' h with a e + rw [← e] + refine inductionOn a ?_ + intro α r _ + exact lift_type_lt.{u, u + 1, max (u + 1) v}.2 ⟨typein.principalSeg r⟩ · rw [← lift_id (type s)] at h ⊢ cases' lift_type_lt.{_,_,v}.1 h with f cases' f with f a hf @@ -1076,23 +1081,17 @@ def liftPrincipalSeg : @PrincipalSeg Ordinal.{u} Ordinal.{max (u + 1) v} (· < -- Porting note: apply inductionOn does not work, refine does refine inductionOn a ?_ intro α r _ hf - refine - lift_type_eq.{u, max (u + 1) v, max (u + 1) v}.2 - ⟨(RelIso.ofSurjective (RelEmbedding.ofMonotone ?_ ?_) ?_).symm⟩ - · exact fun b => enum r ⟨f b, (hf _).2 ⟨_, rfl⟩⟩ + refine lift_type_eq.{u, max (u + 1) v, max (u + 1) v}.2 + ⟨(RelIso.ofSurjective (RelEmbedding.ofMonotone ?_ ?_) ?_).symm⟩ + · exact fun b => enum r ⟨f b, (hf _).1 ⟨_, rfl⟩⟩ · refine fun a b h => (typein_lt_typein r).1 ?_ rw [typein_enum, typein_enum] exact f.map_rel_iff.2 h · intro a' - cases' (hf _).1 (typein_lt_type _ a') with b e + cases' (hf _).2 (typein_lt_type _ a') with b e exists b simp only [RelEmbedding.ofMonotone_coe] - simp [e] - · cases' h with a e - rw [← e] - refine inductionOn a ?_ - intro α r _ - exact lift_type_lt.{u, u + 1, max (u + 1) v}.2 ⟨typein.principalSeg r⟩⟩ + simp [e]⟩ @[deprecated liftPrincipalSeg (since := "2024-09-21")] alias lift.principalSeg := liftPrincipalSeg @@ -1194,6 +1193,9 @@ theorem card_ord (c) : (ord c).card = c := -- Porting note: cardinal.mk_def is now Cardinal.mk'_def, not sure why simp only [mk'_def, e, card_type] +theorem card_surjective : Function.Surjective card := + fun c ↦ ⟨_, card_ord c⟩ + /-- Galois coinsertion between `Cardinal.ord` and `Ordinal.card`. -/ def gciOrdCard : GaloisCoinsertion ord card := gc_ord_card.toGaloisCoinsertion fun c => c.card_ord.le @@ -1341,7 +1343,7 @@ theorem lift_lt_univ' (c : Cardinal) : lift.{max (u + 1) v, u} c < univ.{u, v} : @[simp] theorem ord_univ : ord univ.{u, v} = Ordinal.univ.{u, v} := by refine le_antisymm (ord_card_le _) <| le_of_forall_lt fun o h => lt_ord.2 ?_ - have := liftPrincipalSeg.{u, v}.down.1 (by simpa only [liftPrincipalSeg_coe] using h) + have := liftPrincipalSeg.mem_range_of_rel_top (by simpa only [liftPrincipalSeg_coe] using h) rcases this with ⟨o, h'⟩ rw [← h', liftPrincipalSeg_coe, ← lift_card] apply lift_lt_univ' @@ -1350,7 +1352,7 @@ theorem lt_univ {c} : c < univ.{u, u + 1} ↔ ∃ c', c = lift.{u + 1, u} c' := ⟨fun h => by have := ord_lt_ord.2 h rw [ord_univ] at this - cases' liftPrincipalSeg.{u, u + 1}.down.1 (by simpa only [liftPrincipalSeg_top] ) with o e + cases' liftPrincipalSeg.mem_range_of_rel_top (by simpa only [liftPrincipalSeg_top]) with o e have := card_ord c rw [← e, liftPrincipalSeg_coe, ← lift_card] at this exact ⟨_, this.symm⟩, fun ⟨c', e⟩ => e.symm ▸ lift_lt_univ _⟩ diff --git a/Mathlib/SetTheory/Ordinal/Enum.lean b/Mathlib/SetTheory/Ordinal/Enum.lean new file mode 100644 index 0000000000000..c5a4c3b7442d6 --- /dev/null +++ b/Mathlib/SetTheory/Ordinal/Enum.lean @@ -0,0 +1,125 @@ +/- +Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Violeta Hernández Palacios +-/ +import Mathlib.SetTheory.Ordinal.Arithmetic + +/-! +# Enumerating sets of ordinals by ordinals + +The ordinals have the peculiar property that every subset bounded above is a small type, while +themselves not being small. As a consequence of this, every unbounded subset of `Ordinal` is order +isomorphic to `Ordinal`. + +We define this correspondence as `enumOrd`, and use it to then define an order isomorphism +`enumOrdOrderIso`. + +This can be thought of as an ordinal analog of `Nat.nth`. +-/ + +universe u + +open Order Set + +namespace Ordinal + +variable {o a b : Ordinal.{u}} + +/-- Enumerator function for an unbounded set of ordinals. -/ +noncomputable def enumOrd (s : Set Ordinal.{u}) (o : Ordinal.{u}) : Ordinal.{u} := + sInf (s ∩ { b | ∀ c, c < o → enumOrd s c < b }) +termination_by o + +variable {s : Set Ordinal.{u}} + +@[deprecated (since := "2024-09-20")] +theorem enumOrd_def (o : Ordinal.{u}) : + enumOrd s o = sInf (s ∩ { b | ∀ c, c < o → enumOrd s c < b }) := by + rw [enumOrd] + +theorem enumOrd_le_of_forall_lt (ha : a ∈ s) (H : ∀ b < o, enumOrd s b < a) : enumOrd s o ≤ a := by + rw [enumOrd] + exact csInf_le' ⟨ha, H⟩ + +/-- The set in the definition of `enumOrd` is nonempty. -/ +private theorem enumOrd_nonempty (hs : ¬ BddAbove s) (o : Ordinal) : + (s ∩ { b | ∀ c, c < o → enumOrd s c < b }).Nonempty := by + rw [not_bddAbove_iff] at hs + obtain ⟨a, ha⟩ := bddAbove_of_small (enumOrd s '' Iio o) + obtain ⟨b, hb, hba⟩ := hs a + exact ⟨b, hb, fun c hc ↦ (ha (mem_image_of_mem _ hc)).trans_lt hba⟩ + +private theorem enumOrd_mem_aux (hs : ¬ BddAbove s) (o : Ordinal) : + enumOrd s o ∈ s ∩ { b | ∀ c, c < o → enumOrd s c < b } := by + rw [enumOrd] + exact csInf_mem (enumOrd_nonempty hs o) + +theorem enumOrd_mem (hs : ¬ BddAbove s) (o : Ordinal) : enumOrd s o ∈ s := + (enumOrd_mem_aux hs o).1 + +theorem enumOrd_strictMono (hs : ¬ BddAbove s) : StrictMono (enumOrd s) := + fun a b ↦ (enumOrd_mem_aux hs b).2 a + +theorem enumOrd_succ_le (hs : ¬ BddAbove s) (ha : a ∈ s) (hb : enumOrd s b < a) : + enumOrd s (succ b) ≤ a := by + apply enumOrd_le_of_forall_lt ha + intro c hc + rw [lt_succ_iff] at hc + exact ((enumOrd_strictMono hs).monotone hc).trans_lt hb + +theorem range_enumOrd (hs : ¬ BddAbove s) : range (enumOrd s) = s := by + ext a + let t := { b | a ≤ enumOrd s b } + constructor + · rintro ⟨b, rfl⟩ + exact enumOrd_mem hs b + · intro ha + refine ⟨sInf t, (enumOrd_le_of_forall_lt ha ?_).antisymm ?_⟩ + · intro b hb + by_contra! hb' + exact hb.not_le (csInf_le' hb') + · exact csInf_mem (s := t) ⟨a, (enumOrd_strictMono hs).id_le a⟩ + +theorem enumOrd_surjective (hs : ¬ BddAbove s) {b : Ordinal} (hb : b ∈ s) : + ∃ a, enumOrd s a = b := by + rwa [← range_enumOrd hs] at hb + +theorem enumOrd_le_of_subset {t : Set Ordinal} (hs : ¬ BddAbove s) (hst : s ⊆ t) : + enumOrd t ≤ enumOrd s := by + intro a + rw [enumOrd, enumOrd] + apply csInf_le_csInf' (enumOrd_nonempty hs a) (inter_subset_inter hst _) + intro b hb c hc + exact (enumOrd_le_of_subset hs hst c).trans_lt <| hb c hc +termination_by a => a + +/-- A characterization of `enumOrd`: it is the unique strict monotonic function with range `s`. -/ +theorem eq_enumOrd (f : Ordinal → Ordinal) (hs : ¬ BddAbove s) : + enumOrd s = f ↔ StrictMono f ∧ range f = s := by + constructor + · rintro rfl + exact ⟨enumOrd_strictMono hs, range_enumOrd hs⟩ + · rintro ⟨h₁, h₂⟩ + rwa [← (enumOrd_strictMono hs).range_inj h₁, range_enumOrd hs, eq_comm] + +theorem enumOrd_range {f : Ordinal → Ordinal} (hf : StrictMono f) : enumOrd (range f) = f := + (eq_enumOrd _ hf.not_bddAbove_range_of_wellFoundedLT).2 ⟨hf, rfl⟩ + +@[simp] +theorem enumOrd_univ : enumOrd Set.univ = id := by + rw [← range_id] + exact enumOrd_range strictMono_id + +@[simp] +theorem enumOrd_zero : enumOrd s 0 = sInf s := by + rw [enumOrd] + simp [Ordinal.not_lt_zero] + +/-- An order isomorphism between an unbounded set of ordinals and the ordinals. -/ +noncomputable def enumOrdOrderIso (s : Set Ordinal) (hs : ¬ BddAbove s) : Ordinal ≃o s := + StrictMono.orderIsoOfSurjective (fun o => ⟨_, enumOrd_mem hs o⟩) (enumOrd_strictMono hs) fun s => + let ⟨a, ha⟩ := enumOrd_surjective hs s.prop + ⟨a, Subtype.eq ha⟩ + +end Ordinal diff --git a/Mathlib/SetTheory/Ordinal/Exponential.lean b/Mathlib/SetTheory/Ordinal/Exponential.lean index f5d3ed03ea35f..ca373660f3136 100644 --- a/Mathlib/SetTheory/Ordinal/Exponential.lean +++ b/Mathlib/SetTheory/Ordinal/Exponential.lean @@ -86,6 +86,14 @@ theorem opow_pos {a : Ordinal} (b : Ordinal) (a0 : 0 < a) : 0 < a ^ b := by theorem opow_ne_zero {a : Ordinal} (b : Ordinal) (a0 : a ≠ 0) : a ^ b ≠ 0 := Ordinal.pos_iff_ne_zero.1 <| opow_pos b <| Ordinal.pos_iff_ne_zero.2 a0 +@[simp] +theorem opow_eq_zero {a b : Ordinal} : a ^ b = 0 ↔ a = 0 ∧ b ≠ 0 := by + obtain rfl | ha := eq_or_ne a 0 + · obtain rfl | hb := eq_or_ne b 0 + · simp + · simp [hb] + · simp [opow_ne_zero b ha, ha] + @[simp, norm_cast] theorem opow_natCast (a : Ordinal) (n : ℕ) : a ^ (n : Ordinal) = a ^ n := by induction n with @@ -286,17 +294,67 @@ theorem opow_log_le_self (b : Ordinal) {x : Ordinal} (hx : x ≠ 0) : b ^ log b rwa [← succ_log_def hb hx] at this · rwa [one_opow, one_le_iff_ne_zero] -/-- `opow b` and `log b` (almost) form a Galois connection. -/ -theorem opow_le_iff_le_log {b x c : Ordinal} (hb : 1 < b) (hx : x ≠ 0) : b ^ c ≤ x ↔ c ≤ log b x := - ⟨fun h => - le_of_not_lt fun hn => - (lt_opow_succ_log_self hb x).not_le <| - ((opow_le_opow_iff_right hb).2 (succ_le_of_lt hn)).trans h, - fun h => ((opow_le_opow_iff_right hb).2 h).trans (opow_log_le_self b hx)⟩ - +/-- `opow b` and `log b` (almost) form a Galois connection. + +See `opow_le_iff_le_log'` for a variant assuming `c ≠ 0` rather than `x ≠ 0`. See also +`le_log_of_opow_le` and `opow_le_of_le_log`, which are both separate implications under weaker +assumptions. -/ +theorem opow_le_iff_le_log {b x c : Ordinal} (hb : 1 < b) (hx : x ≠ 0) : + b ^ c ≤ x ↔ c ≤ log b x := by + constructor <;> + intro h + · apply le_of_not_lt + intro hn + apply (lt_opow_succ_log_self hb x).not_le <| + ((opow_le_opow_iff_right hb).2 <| succ_le_of_lt hn).trans h + · exact ((opow_le_opow_iff_right hb).2 h).trans <| opow_log_le_self b hx + +/-- `opow b` and `log b` (almost) form a Galois connection. + +See `opow_le_iff_le_log` for a variant assuming `x ≠ 0` rather than `c ≠ 0`. See also +`le_log_of_opow_le` and `opow_le_of_le_log`, which are both separate implications under weaker +assumptions. -/ +theorem opow_le_iff_le_log' {b x c : Ordinal} (hb : 1 < b) (hc : c ≠ 0) : + b ^ c ≤ x ↔ c ≤ log b x := by + obtain rfl | hx := eq_or_ne x 0 + · rw [log_zero_right, Ordinal.le_zero, Ordinal.le_zero, opow_eq_zero] + simp [hc, (zero_lt_one.trans hb).ne'] + · exact opow_le_iff_le_log hb hx + +theorem le_log_of_opow_le {b x c : Ordinal} (hb : 1 < b) (h : b ^ c ≤ x) : c ≤ log b x := by + obtain rfl | hx := eq_or_ne x 0 + · rw [Ordinal.le_zero, opow_eq_zero] at h + exact (zero_lt_one.asymm <| h.1 ▸ hb).elim + · exact (opow_le_iff_le_log hb hx).1 h + +theorem opow_le_of_le_log {b x c : Ordinal} (hc : c ≠ 0) (h : c ≤ log b x) : b ^ c ≤ x := by + obtain hb | hb := le_or_lt b 1 + · rw [log_of_left_le_one hb] at h + exact (h.not_lt (Ordinal.pos_iff_ne_zero.2 hc)).elim + · rwa [opow_le_iff_le_log' hb hc] + +/-- `opow b` and `log b` (almost) form a Galois connection. + +See `lt_opow_iff_log_lt'` for a variant assuming `c ≠ 0` rather than `x ≠ 0`. See also +`lt_opow_of_log_lt` and `lt_log_of_lt_opow`, which are both separate implications under weaker +assumptions. -/ theorem lt_opow_iff_log_lt {b x c : Ordinal} (hb : 1 < b) (hx : x ≠ 0) : x < b ^ c ↔ log b x < c := lt_iff_lt_of_le_iff_le (opow_le_iff_le_log hb hx) +/-- `opow b` and `log b` (almost) form a Galois connection. + +See `lt_opow_iff_log_lt` for a variant assuming `x ≠ 0` rather than `c ≠ 0`. See also +`lt_opow_of_log_lt` and `lt_log_of_lt_opow`, which are both separate implications under weaker +assumptions. -/ +theorem lt_opow_iff_log_lt' {b x c : Ordinal} (hb : 1 < b) (hc : c ≠ 0) : x < b ^ c ↔ log b x < c := + lt_iff_lt_of_le_iff_le (opow_le_iff_le_log' hb hc) + +theorem lt_opow_of_log_lt {b x c : Ordinal} (hb : 1 < b) : log b x < c → x < b ^ c := + lt_imp_lt_of_le_imp_le <| le_log_of_opow_le hb + +theorem lt_log_of_lt_opow {b x c : Ordinal} (hc : c ≠ 0) : x < b ^ c → log b x < c := + lt_imp_lt_of_le_imp_le <| opow_le_of_le_log hc + theorem log_pos {b o : Ordinal} (hb : 1 < b) (ho : o ≠ 0) (hbo : b ≤ o) : 0 < log b o := by rwa [← succ_le_iff, succ_zero, ← opow_le_iff_le_log hb ho, opow_one] diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index 4fe7f81c09218..cedf18985846e 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Violeta Hernández Palacios, Mario Carneiro. All rights reser Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios, Mario Carneiro -/ -import Mathlib.SetTheory.Ordinal.Arithmetic +import Mathlib.SetTheory.Ordinal.Enum import Mathlib.SetTheory.Ordinal.Exponential /-! @@ -129,6 +129,14 @@ theorem nfpFamily_eq_self {f : ι → Ordinal → Ordinal} {a} (h : ∀ i, f i a -- Todo: This is actually a special case of the fact the intersection of club sets is a club set. /-- A generalization of the fixed point lemma for normal functions: any family of normal functions has an unbounded set of common fixed points. -/ +theorem not_bddAbove_fp_family (H : ∀ i, IsNormal (f i)) : + ¬ BddAbove (⋂ i, Function.fixedPoints (f i)) := by + rw [not_bddAbove_iff] + refine fun a ↦ ⟨nfpFamily f (succ a), ?_, (lt_succ a).trans_le (le_nfpFamily f _)⟩ + rintro _ ⟨i, rfl⟩ + exact nfpFamily_fp (H i) _ + +@[deprecated not_bddAbove_fp_family (since := "2024-09-20")] theorem fp_family_unbounded (H : ∀ i, IsNormal (f i)) : (⋂ i, Function.fixedPoints (f i)).Unbounded (· < ·) := fun a => ⟨nfpFamily.{u, v} f a, fun s ⟨i, hi⟩ => by @@ -206,7 +214,7 @@ theorem fp_iff_derivFamily (H : ∀ i, IsNormal (f i)) {a} : /-- For a family of normal functions, `Ordinal.derivFamily` enumerates the common fixed points. -/ theorem derivFamily_eq_enumOrd (H : ∀ i, IsNormal (f i)) : derivFamily.{u, v} f = enumOrd (⋂ i, Function.fixedPoints (f i)) := by - rw [← eq_enumOrd _ (fp_family_unbounded.{u, v} H)] + rw [eq_comm, eq_enumOrd _ (not_bddAbove_fp_family H)] use (derivFamily_isNormal f).strictMono rw [Set.range_eq_iff] refine ⟨?_, fun a ha => ?_⟩ @@ -301,6 +309,16 @@ theorem nfpBFamily_eq_self {a} (h : ∀ i hi, f i hi a = a) : nfpBFamily.{u, v} /-- A generalization of the fixed point lemma for normal functions: any family of normal functions has an unbounded set of common fixed points. -/ +theorem not_bddAbove_fp_bfamily (H : ∀ i hi, IsNormal (f i hi)) : + ¬ BddAbove (⋂ (i) (hi), Function.fixedPoints (f i hi)) := by + rw [not_bddAbove_iff] + refine fun a ↦ ⟨nfpBFamily _ f (succ a), ?_, (lt_succ a).trans_le (le_nfpBFamily f _)⟩ + rw [Set.mem_iInter₂] + exact fun i hi ↦ nfpBFamily_fp (H i hi) _ + +/-- A generalization of the fixed point lemma for normal functions: any family of normal functions + has an unbounded set of common fixed points. -/ +@[deprecated not_bddAbove_fp_bfamily (since := "2024-09-20")] theorem fp_bfamily_unbounded (H : ∀ i hi, IsNormal (f i hi)) : (⋂ (i) (hi), Function.fixedPoints (f i hi)).Unbounded (· < ·) := fun a => ⟨nfpBFamily.{u, v} _ f a, by @@ -347,7 +365,7 @@ theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : /-- For a family of normal functions, `Ordinal.derivBFamily` enumerates the common fixed points. -/ theorem derivBFamily_eq_enumOrd (H : ∀ i hi, IsNormal (f i hi)) : derivBFamily.{u, v} o f = enumOrd (⋂ (i) (hi), Function.fixedPoints (f i hi)) := by - rw [← eq_enumOrd _ (fp_bfamily_unbounded.{u, v} H)] + rw [eq_comm, eq_enumOrd _ (not_bddAbove_fp_bfamily H)] use (derivBFamily_isNormal f).strictMono rw [Set.range_eq_iff] refine ⟨fun a => Set.mem_iInter₂.2 fun i hi => derivBFamily_fp (H i hi) a, fun a ha => ?_⟩ @@ -445,6 +463,14 @@ theorem nfp_eq_self {f : Ordinal → Ordinal} {a} (h : f a = a) : nfp f a = a := /-- The fixed point lemma for normal functions: any normal function has an unbounded set of fixed points. -/ +theorem not_bddAbove_fp (H : IsNormal f) : ¬ BddAbove (Function.fixedPoints f) := by + convert not_bddAbove_fp_family fun _ : Unit => H + exact (Set.iInter_const _).symm + +set_option linter.deprecated false in +/-- The fixed point lemma for normal functions: any normal function has an unbounded set of +fixed points. -/ +@[deprecated not_bddAbove_fp (since := "2024-09-20")] theorem fp_unbounded (H : IsNormal f) : (Function.fixedPoints f).Unbounded (· < ·) := by convert fp_family_unbounded fun _ : Unit => H exact (Set.iInter_const _).symm diff --git a/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean b/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean index ddec8ac7b3058..97272de665f4f 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean @@ -50,7 +50,7 @@ theorem not_injective_limitation_set : ¬ InjOn g (Iio (ord <| succ #α)) := by have h := lift_mk_le_lift_mk_of_injective <| injOn_iff_injective.1 h_inj have mk_initialSeg_subtype : #(Iio (ord <| succ #α)) = lift.{u + 1} (succ #α) := by - simpa only [coe_setOf, card_typein, card_ord] using mk_initialSeg (ord <| succ #α) + simpa only [coe_setOf, card_typein, card_ord] using mk_Iio_ordinal (ord <| succ #α) rw [mk_initialSeg_subtype, lift_lift, lift_le] at h exact not_le_of_lt (Order.lt_succ #α) h diff --git a/Mathlib/SetTheory/Ordinal/Nimber.lean b/Mathlib/SetTheory/Ordinal/Nimber.lean index 538d61ec4400d..92bb773226c18 100644 --- a/Mathlib/SetTheory/Ordinal/Nimber.lean +++ b/Mathlib/SetTheory/Ordinal/Nimber.lean @@ -152,8 +152,6 @@ theorem not_small_nimber : ¬ Small.{u} Nimber.{max u v} := namespace Ordinal -variable {a b c : Ordinal.{u}} - @[simp] theorem toNimber_symm_eq : toNimber.symm = Nimber.toOrdinal := rfl diff --git a/Mathlib/SetTheory/Ordinal/Topology.lean b/Mathlib/SetTheory/Ordinal/Topology.lean index 73e5ab814961e..6993d517742e4 100644 --- a/Mathlib/SetTheory/Ordinal/Topology.lean +++ b/Mathlib/SetTheory/Ordinal/Topology.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ -import Mathlib.SetTheory.Ordinal.Arithmetic +import Mathlib.SetTheory.Ordinal.Enum import Mathlib.Tactic.TFAE import Mathlib.Topology.Order.Monotone @@ -219,27 +219,26 @@ theorem isNormal_iff_strictMono_and_continuous (f : Ordinal.{u} → Ordinal.{u}) ⟨_, toType_nonempty_iff_ne_zero.2 ho.1, typein (· < ·), fun i => h _ (typein_lt_self i), sup_typein_limit ho.2⟩ -theorem enumOrd_isNormal_iff_isClosed (hs : s.Unbounded (· < ·)) : +theorem enumOrd_isNormal_iff_isClosed (hs : ¬ BddAbove s) : IsNormal (enumOrd s) ↔ IsClosed s := by have Hs := enumOrd_strictMono hs refine ⟨fun h => isClosed_iff_iSup.2 fun {ι} hι f hf => ?_, fun h => (isNormal_iff_strictMono_limit _).2 ⟨Hs, fun a ha o H => ?_⟩⟩ - · let g : ι → Ordinal.{u} := fun i => (enumOrdOrderIso hs).symm ⟨_, hf i⟩ + · let g : ι → Ordinal.{u} := fun i => (enumOrdOrderIso s hs).symm ⟨_, hf i⟩ suffices enumOrd s (⨆ i, g i) = ⨆ i, f i by rw [← this] exact enumOrd_mem hs _ rw [IsNormal.map_iSup h g] congr ext x - change ((enumOrdOrderIso hs) _).val = f x + change (enumOrdOrderIso s hs _).val = f x rw [OrderIso.apply_symm_apply] · rw [isClosed_iff_bsup] at h suffices enumOrd s a ≤ bsup.{u, u} a fun b (_ : b < a) => enumOrd s b from this.trans (bsup_le H) - cases' enumOrd_surjective hs _ - (h ha.1 (fun b _ => enumOrd s b) fun b _ => enumOrd_mem hs b) with - b hb + obtain ⟨b, hb⟩ := enumOrd_surjective hs (h ha.1 (fun b _ => enumOrd s b) + fun b _ => enumOrd_mem hs b) rw [← hb] apply Hs.monotone by_contra! hba diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index a7379f3dc58f9..46952152d3576 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -143,6 +143,7 @@ import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Linter.Lint import Mathlib.Tactic.Linter.MinImports import Mathlib.Tactic.Linter.OldObtain +import Mathlib.Tactic.Linter.PPRoundtrip import Mathlib.Tactic.Linter.RefineLinter import Mathlib.Tactic.Linter.Style import Mathlib.Tactic.Linter.TextBased @@ -208,6 +209,7 @@ import Mathlib.Tactic.RewriteSearch import Mathlib.Tactic.Rify import Mathlib.Tactic.Ring import Mathlib.Tactic.Ring.Basic +import Mathlib.Tactic.Ring.Compare import Mathlib.Tactic.Ring.PNat import Mathlib.Tactic.Ring.RingNF import Mathlib.Tactic.Sat.FromLRAT diff --git a/Mathlib/Tactic/CategoryTheory/Bicategory/Datatypes.lean b/Mathlib/Tactic/CategoryTheory/Bicategory/Datatypes.lean index e7c2dd37d4f41..42225947c25da 100644 --- a/Mathlib/Tactic/CategoryTheory/Bicategory/Datatypes.lean +++ b/Mathlib/Tactic/CategoryTheory/Bicategory/Datatypes.lean @@ -103,7 +103,7 @@ instance : MonadMor₁ BicategoryM where section universe w v u -variable {B : Type u} [Bicategory.{w, v} B] {a b c d e : B} +variable {B : Type u} [Bicategory.{w, v} B] {a b c : B} theorem structuralIso_inv {f g : a ⟶ b} (η : f ≅ g) : η.symm.hom = η.inv := by diff --git a/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean b/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean index 2e58492c79e90..7104020508294 100644 --- a/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean +++ b/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean @@ -27,7 +27,7 @@ open CategoryTheory CategoryTheory.FreeBicategory open scoped Bicategory -variable {B : Type u} [Bicategory.{w, v} B] {a b c d e : B} +variable {B : Type u} [Bicategory.{w, v} B] {a b c d : B} namespace Mathlib.Tactic.BicategoryCoherence diff --git a/Mathlib/Tactic/DeprecateMe.lean b/Mathlib/Tactic/DeprecateMe.lean index 3fb56313bec00..b84b97f0b4c2c 100644 --- a/Mathlib/Tactic/DeprecateMe.lean +++ b/Mathlib/Tactic/DeprecateMe.lean @@ -107,7 +107,7 @@ Technically, the command also take an optional `String` argument to fill in the However, its use is mostly intended for debugging purposes, where having a variable date would make tests time-dependent. -/ -elab tk:"deprecate to " id:ident* dat:(str)? ppLine cmd:command : command => do +elab tk:"deprecate to " id:ident* dat:(ppSpace str ppSpace)? ppLine cmd:command : command => do let oldEnv ← getEnv try elabCommand cmd diff --git a/Mathlib/Tactic/FunProp/ContDiff.lean b/Mathlib/Tactic/FunProp/ContDiff.lean index adc20fffe2b95..037d1812296a5 100644 --- a/Mathlib/Tactic/FunProp/ContDiff.lean +++ b/Mathlib/Tactic/FunProp/ContDiff.lean @@ -23,8 +23,7 @@ variable {K : Type*} [NontriviallyNormedField K] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace K E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace K F] variable {G : Type*} [NormedAddCommGroup G] [NormedSpace K G] -variable {G' : Type*} [NormedAddCommGroup G'] [NormedSpace K G'] -variable {f f₀ f₁ g : E → F} {x} {s t} {n} +variable {f : E → F} {x} {s} {n} theorem contDiff_id' : ContDiff K n (fun x : E => x) := contDiff_id @@ -42,8 +41,8 @@ theorem ContDiffAt.comp' {f : E → F} {g : F → G} (hg : ContDiffAt K n g (f x -- theorem ContDiffOn.comp'' {g : F → G} {t : Set F} (hg : ContDiffOn K n g t) -- (hf : ContDiffOn K n f s) (st : Set.MapsTo f s t) : ContDiffOn K n (fun x => g (f x)) s := -variable {ι ι' : Type*} [Fintype ι] [Fintype ι'] {F' : ι → Type*} [∀ i, NormedAddCommGroup (F' i)] - [∀ i, NormedSpace K (F' i)] {φ : ∀ i, E → F' i} {Φ : E → ∀ i, F' i} +variable {ι : Type*} [Fintype ι] {F' : ι → Type*} [∀ i, NormedAddCommGroup (F' i)] + [∀ i, NormedSpace K (F' i)] {Φ : E → ∀ i, F' i} theorem contDiff_pi' (hΦ : ∀ i, ContDiff K n fun x => Φ x i) : ContDiff K n Φ := contDiff_pi.2 hΦ @@ -60,8 +59,7 @@ section div variable {K : Type*} [NontriviallyNormedField K] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace K E] -variable {F : Type*} [NormedAddCommGroup F] [NormedSpace K F] -variable {f f₀ f₁ g : E → F} {x} {s t} {n} +variable {s} theorem ContDiffOn.div' [CompleteSpace K] {f g : E → K} {n} (hf : ContDiffOn K n f s) (hg : ContDiffOn K n g s) (h₀ : ∀ x ∈ s, g x ≠ 0) : ContDiffOn K n (fun x => f x / g x) s := @@ -74,7 +72,6 @@ end div section deriv variable {K : Type*} [NontriviallyNormedField K] -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace K E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace K F] /-- Original version `ContDiff.differentiable_iteratedDeriv` introduces a new variable `(n:ℕ∞)` diff --git a/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean b/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean index da93ffa29a9af..141a5bb819107 100644 --- a/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean +++ b/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean @@ -449,7 +449,7 @@ partial def reduce (e : Expr) (config : WhnfCoreConfig) : MetaM Expr := do /-- Repeatedly apply reduce while stripping lambda binders and introducing their variables -/ @[specialize] partial def lambdaTelescopeReduce {m} [Monad m] [MonadLiftT MetaM m] [MonadControlT MetaM m] - [Inhabited α] (e : Expr) (fvars : List FVarId) (config : WhnfCoreConfig) + [Nonempty α] (e : Expr) (fvars : List FVarId) (config : WhnfCoreConfig) (k : Expr → List FVarId → m α) : m α := do match ← reduce e config with | .lam n d b bi => diff --git a/Mathlib/Tactic/Lemma.lean b/Mathlib/Tactic/Lemma.lean index 901874a9f585a..10d49cd8d8a8f 100644 --- a/Mathlib/Tactic/Lemma.lean +++ b/Mathlib/Tactic/Lemma.lean @@ -12,8 +12,9 @@ import Lean.Parser.Command open Lean +-- higher priority to override the one in Batteries /-- `lemma` means the same as `theorem`. It is used to denote "less important" theorems -/ -syntax (name := lemma) declModifiers +syntax (name := lemma) (priority := default + 1) declModifiers group("lemma " declId ppIndent(declSig) declVal) : command /-- Implementation of the `lemma` command, by macro expansion to `theorem`. -/ diff --git a/Mathlib/Tactic/Linter.lean b/Mathlib/Tactic/Linter.lean index e5c4da4423d3c..2f67da1bea192 100644 --- a/Mathlib/Tactic/Linter.lean +++ b/Mathlib/Tactic/Linter.lean @@ -11,3 +11,4 @@ This file is ignored by `shake`: import Mathlib.Tactic.Linter.FlexibleLinter import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Linter.MinImports +import Mathlib.Tactic.Linter.PPRoundtrip diff --git a/Mathlib/Tactic/Linter/PPRoundtrip.lean b/Mathlib/Tactic/Linter/PPRoundtrip.lean new file mode 100644 index 0000000000000..f6f25280cc338 --- /dev/null +++ b/Mathlib/Tactic/Linter/PPRoundtrip.lean @@ -0,0 +1,144 @@ +/- +Copyright (c) 2024 Damiano Testa. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa +-/ + +import Lean.Elab.Command +import Mathlib.Init + +/-! +# The "ppRoundtrip" linter + +The "ppRoundtrip" linter emits a warning when the syntax of a command differs substantially +from the pretty-printed version of itself. +-/ +open Lean Elab Command + +namespace Mathlib.Linter + +/-- +The "ppRoundtrip" linter emits a warning when the syntax of a command differs substantially +from the pretty-printed version of itself. + +The linter makes an effort to start the highlighting at the first difference. +However, it may not always be successful. +It also prints both the source code and the "expected code" in a 5-character radius from +the first difference. +-/ +register_option linter.ppRoundtrip : Bool := { + defValue := false + descr := "enable the ppRoundtrip linter" +} + +/-- `polishPP s` takes as input a `String` `s`, assuming that it is the output of +pretty-printing a lean command. +The main intent is to convert `s` to a reasonable candidate for a desirable source code format. +The function first replaces consecutive whitespace sequences into a single space (` `), in an +attempt to side-step line-break differences. +After that, it applies some pre-emptive changes: +* doc-module beginnings tend to have some whitespace following them, so we add a space back in; +* name quotations such as ``` ``Nat``` get pretty-printed as ``` `` Nat```, so we remove a space + after double back-ticks, but take care of adding one more for triple (or more) back-ticks; +* `notation3` is not followed by a pretty-printer space, so we add it here (#15515). +-/ +def polishPP (s : String) : String := + let s := s.split (·.isWhitespace) + (" ".intercalate (s.filter (!·.isEmpty))) + |>.replace "/-!" "/-! " + |>.replace "``` " "``` " -- avoid losing an existing space after the triple back-ticks + -- as a consequence of the following replacement + |>.replace "`` " "``" -- weird pp ```#eval ``«Nat»``` pretty-prints as ```#eval `` «Nat»``` + |>.replace "notation3(" "notation3 (" + |>.replace "notation3\"" "notation3 \"" + +/-- `polishSource s` is similar to `polishPP s`, but expects the input to be actual source code. +For this reason, `polishSource s` performs more conservative changes: +it only replace all whitespace starting from a linebreak (`\n`) with a single whitespace. -/ +def polishSource (s : String) : String × Array Nat := + let split := s.split (· == '\n') + let preWS := split.foldl (init := #[]) fun p q => + let txt := q.trimLeft.length + (p.push (q.length - txt)).push txt + let preWS := preWS.eraseIdx 0 + let s := (split.map .trimLeft).filter (· != "") + (" ".intercalate (s.filter (!·.isEmpty)), preWS) + +/-- `posToShiftedPos lths diff` takes as input an array `lths` of natural numbers, +and one further natural number `diff`. +It adds up the elements of `lths` occupying the odd positions, as long as the sum of the +elements in the even positions does not exceed `diff`. +It returns the sum of the accumulated odds and `diff`. +This is useful to figure out the difference between the output of `polishSource s` and `s` itself. +It plays a role similar to the `fileMap`. -/ +def posToShiftedPos (lths : Array Nat) (diff : Nat) : Nat := Id.run do + let mut (ws, noWS) := (diff, 0) + for con in [:lths.size / 2] do + let curr := lths[2 * con]! + if noWS + curr < diff then + noWS := noWS + curr + ws := ws + lths[2 * con + 1]! + else + break + return ws + +/-- `zoomString str centre offset` returns the substring of `str` consisting of the `offset` +characters around the `centre`th character. -/ +def zoomString (str : String) (centre offset : Nat) : Substring := + { str := str, startPos := ⟨centre - offset⟩, stopPos := ⟨centre + offset⟩ } + +/-- `capSourceInfo s p` "shortens" all end-position information in the `SourceInfo` `s` to be +at most `p`, trimming down also the relevant substrings. -/ +def capSourceInfo (s : SourceInfo) (p : Nat) : SourceInfo := + match s with + | .original leading pos trailing endPos => + .original leading pos {trailing with stopPos := ⟨min endPos.1 p⟩} ⟨min endPos.1 p⟩ + | .synthetic pos endPos canonical => + .synthetic pos ⟨min endPos.1 p⟩ canonical + | .none => s + +/-- `capSyntax stx p` applies `capSourceInfo · s` to all `SourceInfo`s in all +`node`s, `atom`s and `ident`s contained in `stx`. + +This is used to trim away all "fluff" that follows a command: comments and whitespace after +a command get removed with `capSyntax stx stx.getTailPos?.get!`. +-/ +partial +def capSyntax (stx : Syntax) (p : Nat) : Syntax := + match stx with + | .node si k args => .node (capSourceInfo si p) k (args.map (capSyntax · p)) + | .atom si val => .atom (capSourceInfo si p) (val.take p) + | .ident si r v pr => .ident (capSourceInfo si p) { r with stopPos := ⟨min r.stopPos.1 p⟩ } v pr + | s => s + +namespace PPRoundtrip + +@[inherit_doc Mathlib.Linter.linter.ppRoundtrip] +def ppRoundtrip : Linter where run := withSetOptionIn fun stx ↦ do + unless Linter.getLinterValue linter.ppRoundtrip (← getOptions) do + return + if (← MonadState.get).messages.hasErrors then + return + let stx := capSyntax stx (stx.getTailPos?.getD default).1 + let origSubstring := stx.getSubstring?.getD default + let (real, lths) := polishSource origSubstring.toString + let fmt ← (liftCoreM do PrettyPrinter.ppCategory `command stx <|> (do + Linter.logLint linter.ppRoundtrip stx + m!"The ppRoundtrip linter had some parsing issues: \ + feel free to silence it with `set_option linter.ppRoundtrip false in` \ + and report this error!" + return real)) + let st := polishPP fmt.pretty + if st != real then + let diff := real.firstDiffPos st + let pos := posToShiftedPos lths diff.1 + origSubstring.startPos.1 + let f := origSubstring.str.drop (pos) + let extraLth := (f.takeWhile (· != st.get diff)).length + let srcCtxt := zoomString real diff.1 5 + let ppCtxt := zoomString st diff.1 5 + Linter.logLint linter.ppRoundtrip (.ofRange ⟨⟨pos⟩, ⟨pos + extraLth + 1⟩⟩) + m!"source context\n'{srcCtxt}'\n'{ppCtxt}'\npretty-printed context" + +initialize addLinter ppRoundtrip + +end Mathlib.Linter.PPRoundtrip diff --git a/Mathlib/Tactic/NormNum/Ineq.lean b/Mathlib/Tactic/NormNum/Ineq.lean index 987fbd80ced59..9c5a20797bcc2 100644 --- a/Mathlib/Tactic/NormNum/Ineq.lean +++ b/Mathlib/Tactic/NormNum/Ineq.lean @@ -108,12 +108,21 @@ such that `norm_num` successfully recognises both `a` and `b`. -/ let ⟨u, α, a⟩ ← inferTypeQ' a have b : Q($α) := b let ra ← derive a; let rb ← derive b + let lα ← synthInstanceQ q(LE $α) + guard <|← withNewMCtxDepth <| isDefEq f q(LE.le (α := $α)) + core lα ra rb +where + /-- Identify (as `true` or `false`) expressions of the form `a ≤ b`, where `a` and `b` are numeric + expressions whose evaluations to `NormNum.Result` have already been computed. -/ + core {u : Level} {α : Q(Type u)} (lα : Q(LE $α)) {a b : Q($α)} + (ra : NormNum.Result a) (rb : NormNum.Result b) : MetaM (NormNum.Result q($a ≤ $b)) := do + let e := q($a ≤ $b) let rec intArm : MetaM (Result e) := do let _i ← inferOrderedRing α - guard <|← withNewMCtxDepth <| isDefEq f q(LE.le (α := $α)) haveI' : $e =Q ($a ≤ $b) := ⟨⟩ let ⟨za, na, pa⟩ ← ra.toInt q(OrderedRing.toRing) let ⟨zb, nb, pb⟩ ← rb.toInt q(OrderedRing.toRing) + assumeInstancesCommute if decide (za ≤ zb) then let r : Q(decide ($na ≤ $nb) = true) := (q(Eq.refl true) : Expr) return .isTrue q(isInt_le_true $pa $pb $r) @@ -125,10 +134,10 @@ such that `norm_num` successfully recognises both `a` and `b`. -/ let rec ratArm : MetaM (Result e) := do -- We need a division ring with an order, and `LinearOrderedField` is the closest mathlib has. let _i ← inferLinearOrderedField α - guard <|← withNewMCtxDepth <| isDefEq f q(LE.le (α := $α)) haveI' : $e =Q ($a ≤ $b) := ⟨⟩ let ⟨qa, na, da, pa⟩ ← ra.toRat' q(Field.toDivisionRing) let ⟨qb, nb, db, pb⟩ ← rb.toRat' q(Field.toDivisionRing) + assumeInstancesCommute if decide (qa ≤ qb) then let r : Q(decide ($na * $db ≤ $nb * $da) = true) := (q(Eq.refl true) : Expr) return (.isTrue q(isRat_le_true $pa $pb $r)) @@ -144,8 +153,8 @@ such that `norm_num` successfully recognises both `a` and `b`. -/ let _i ← inferOrderedSemiring α haveI' : $ra =Q by clear! $ra $rb; infer_instance := ⟨⟩ haveI' : $rb =Q by clear! $ra $rb; infer_instance := ⟨⟩ - guard <|← withNewMCtxDepth <| isDefEq f q(LE.le (α := $α)) haveI' : $e =Q ($a ≤ $b) := ⟨⟩ + assumeInstancesCommute if na.natLit! ≤ nb.natLit! then let r : Q(Nat.ble $na $nb = true) := (q(Eq.refl true) : Expr) return .isTrue q(isNat_le_true $pa $pb $r) @@ -163,13 +172,21 @@ such that `norm_num` successfully recognises both `a` and `b`. -/ let ⟨u, α, a⟩ ← inferTypeQ' a have b : Q($α) := b let ra ← derive a; let rb ← derive b + let lα ← synthInstanceQ q(LT $α) + guard <|← withNewMCtxDepth <| isDefEq f q(LT.lt (α := $α)) + core lα ra rb +where + /-- Identify (as `true` or `false`) expressions of the form `a < b`, where `a` and `b` are numeric + expressions whose evaluations to `NormNum.Result` have already been computed. -/ + core {u : Level} {α : Q(Type u)} (lα : Q(LT $α)) {a b : Q($α)} + (ra : NormNum.Result a) (rb : NormNum.Result b) : MetaM (NormNum.Result q($a < $b)) := do + let e := q($a < $b) let rec intArm : MetaM (Result e) := do let _i ← inferOrderedRing α - assumeInstancesCommute - guard <|← withNewMCtxDepth <| isDefEq f q(LT.lt (α := $α)) haveI' : $e =Q ($a < $b) := ⟨⟩ let ⟨za, na, pa⟩ ← ra.toInt q(OrderedRing.toRing) let ⟨zb, nb, pb⟩ ← rb.toInt q(OrderedRing.toRing) + assumeInstancesCommute if za < zb then if let .some _i ← trySynthInstanceQ (q(@Nontrivial $α) : Q(Prop)) then let r : Q(decide ($na < $nb) = true) := (q(Eq.refl true) : Expr) @@ -184,7 +201,6 @@ such that `norm_num` successfully recognises both `a` and `b`. -/ let _i ← inferLinearOrderedField α assumeInstancesCommute haveI' : $e =Q ($a < $b) := ⟨⟩ - guard <|← withNewMCtxDepth <| isDefEq f q(LT.lt (α := $α)) let ⟨qa, na, da, pa⟩ ← ra.toRat' q(Field.toDivisionRing) let ⟨qb, nb, db, pb⟩ ← rb.toRat' q(Field.toDivisionRing) if qa < qb then @@ -202,7 +218,7 @@ such that `norm_num` successfully recognises both `a` and `b`. -/ haveI' : $ra =Q by clear! $ra $rb; infer_instance := ⟨⟩ haveI' : $rb =Q by clear! $ra $rb; infer_instance := ⟨⟩ haveI' : $e =Q ($a < $b) := ⟨⟩ - guard <|← withNewMCtxDepth <| isDefEq f q(LT.lt (α := $α)) + assumeInstancesCommute if na.natLit! < nb.natLit! then if let .some _i ← trySynthInstanceQ q(CharZero $α) then let r : Q(Nat.ble $nb $na = false) := (q(Eq.refl false) : Expr) diff --git a/Mathlib/Tactic/Ring/Basic.lean b/Mathlib/Tactic/Ring/Basic.lean index dcd4f4ddf5f6f..c5ac7fd426cd8 100644 --- a/Mathlib/Tactic/Ring/Basic.lean +++ b/Mathlib/Tactic/Ring/Basic.lean @@ -885,7 +885,7 @@ def evalPow {a : Q($α)} {b : Q(ℕ)} (va : ExSum sα a) (vb : ExSum sℕ b) : return ⟨_, vd, q(pow_add $pc₁ $pc₂ $pd)⟩ /-- This cache contains data required by the `ring` tactic during execution. -/ -structure Cache {α : Q(Type u)} (sα : Q(CommSemiring $α)) := +structure Cache {α : Q(Type u)} (sα : Q(CommSemiring $α)) where /-- A ring instance on `α`, if available. -/ rα : Option Q(Ring $α) /-- A division ring instance on `α`, if available. -/ diff --git a/Mathlib/Tactic/Ring/Compare.lean b/Mathlib/Tactic/Ring/Compare.lean new file mode 100644 index 0000000000000..5868ee2ee1706 --- /dev/null +++ b/Mathlib/Tactic/Ring/Compare.lean @@ -0,0 +1,239 @@ +/- +Copyright (c) 2024 Heather Macbeth. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Heather Macbeth +-/ +import Mathlib.Tactic.Ring.Basic +import Mathlib.Tactic.NormNum.Ineq + +/-! +# Automation for proving inequalities in commutative (semi)rings + +This file provides automation for proving certain kinds of inequalities in commutative semirings: +goals of the form `A ≤ B` and `A < B` for which the ring-normal forms of `A` and `B` differ by a +nonnegative (resp. positive) constant. + +For example, `⊢ x + 3 + y < y + x + 4` is in scope because the normal forms of the LHS and RHS are, +respectively, `3 + (x + y)` and `4 + (x + y)`, which differ by an additive constant. + +## Main declarations + +* `Mathlib.Tactic.Ring.proveLE`: prove goals of the form `A ≤ B` (subject to the scope constraints + described) +* `Mathlib.Tactic.Ring.proveLT`: prove goals of the form `A < B` (subject to the scope constraints + described) + +## Implementation notes + +The automation is provided in the `MetaM` monad; that is, these functions are not user-facing. It +would not be hard to provide user-facing versions (see the test file), but the scope of this +automation is rather specialized and might be confusing to users. It is also subsumed by `linarith`. +-/ + +namespace Mathlib.Tactic.Ring + +open Lean Qq Meta + +/-! Rather than having the metaprograms `Mathlib.Tactic.Ring.evalLE` and +`Mathlib.Tactic.Ring.evalLT` perform all type class inference at the point of use, we record in +advance, as `abbrev`s, a few type class deductions which will certainly be necessary. They add no +new information (they can already be proved by `inferInstance`). + +This helps in speeding up the metaprograms in this file substantially -- about a 50% reduction in +heartbeat count in representative test cases -- since otherwise a substantial fraction of their +runtime is devoted to type class inference. -/ + +section Typeclass + +/-- `OrderedCommSemiring` implies `CommSemiring`. -/ +abbrev cs_of_ocs (α : Type*) [OrderedCommSemiring α] : CommSemiring α := inferInstance + +/-- `OrderedCommSemiring` implies `AddMonoidWithOne`. -/ +abbrev amwo_of_ocs (α : Type*) [OrderedCommSemiring α] : AddMonoidWithOne α := inferInstance + +/-- `OrderedCommSemiring` implies `LE`. -/ +abbrev le_of_ocs (α : Type*) [OrderedCommSemiring α] : LE α := inferInstance + +/-- `StrictOrderedCommSemiring` implies `CommSemiring`. -/ +abbrev cs_of_socs (α : Type*) [StrictOrderedCommSemiring α] : CommSemiring α := inferInstance + +/-- `StrictOrderedCommSemiring` implies `AddMonoidWithOne`. -/ +abbrev amwo_of_socs (α : Type*) [StrictOrderedCommSemiring α] : AddMonoidWithOne α := inferInstance + +/-- `StrictOrderedCommSemiring` implies `LT`. -/ +abbrev lt_of_socs (α : Type*) [StrictOrderedCommSemiring α] : LT α := inferInstance + +end Typeclass + +/-! The lemmas like `add_le_add_right` in the root namespace are stated under minimal type classes, +typically just `[CovariantClass α α (swap (· + ·)) (· ≤ ·)]` or similar. Here we restate these +lemmas under stronger type class assumptions (`[OrderedCommSemiring α]` or similar), which helps in +speeding up the metaprograms in this file (`Mathlib.Tactic.Ring.proveLT` and +`Mathlib.Tactic.Ring.proveLE`) substantially -- about a 50% reduction in heartbeat count in +representative test cases -- since otherwise a substantial fraction of their runtime is devoted to +type class inference. + +These metaprograms at least require `CommSemiring`, `LE`/`LT`, and all +`CovariantClass`/`ContravariantClass` permutations for addition, and in their main use case (in +`linear_combination`) the `Preorder` type class is also required, so it is rather little loss of +generality simply to require `OrderedCommSemiring`/`StrictOrderedCommSemiring`. -/ + +section Lemma + +theorem add_le_add_right {α : Type*} [OrderedCommSemiring α] {b c : α} (bc : b ≤ c) (a : α) : + b + a ≤ c + a := + _root_.add_le_add_right bc a + +theorem add_le_of_nonpos_left {α : Type*} [OrderedCommSemiring α] (a : α) {b : α} (h : b ≤ 0) : + b + a ≤ a := + _root_.add_le_of_nonpos_left h + +theorem le_add_of_nonneg_left {α : Type*} [OrderedCommSemiring α] (a : α) {b : α} (h : 0 ≤ b) : + a ≤ b + a := + _root_.le_add_of_nonneg_left h + +theorem add_lt_add_right {α : Type*} [StrictOrderedCommSemiring α] {b c : α} (bc : b < c) (a : α) : + b + a < c + a := + _root_.add_lt_add_right bc a + +theorem add_lt_of_neg_left {α : Type*} [StrictOrderedCommSemiring α] (a : α) {b : α} (h : b < 0) : + b + a < a := + _root_.add_lt_of_neg_left a h + +theorem lt_add_of_pos_left {α : Type*} [StrictOrderedCommSemiring α] (a : α) {b : α} (h : 0 < b) : + a < b + a := + _root_.lt_add_of_pos_left a h + +end Lemma + +/-- Inductive type carrying the two kinds of errors which can arise in the metaprograms +`Mathlib.Tactic.Ring.evalLE` and `Mathlib.Tactic.Ring.evalLT`. -/ +inductive ExceptType | tooSmall | notComparable +export ExceptType (tooSmall notComparable) + +/-- In a commutative semiring, given `Ring.ExSum` objects `va`, `vb` which differ by a positive +(additive) constant, construct a proof of `$a < $b`, where `a` (resp. `b`) is the expression in the +semiring to which `va` (resp. `vb`) evaluates. -/ +def evalLE {v : Level} {α : Q(Type v)} (_ : Q(OrderedCommSemiring $α)) {a b : Q($α)} + (va : Ring.ExSum q(cs_of_ocs $α) a) (vb : Ring.ExSum q(cs_of_ocs $α) b) : + MetaM (Except ExceptType Q($a ≤ $b)) := do + let lα : Q(LE $α) := q(le_of_ocs $α) + assumeInstancesCommute + let ⟨_, pz⟩ ← NormNum.mkOfNat α q(amwo_of_ocs $α) (mkRawNatLit 0) + let rz : NormNum.Result q((0:$α)) := + NormNum.Result.isNat q(amwo_of_ocs $α) (mkRawNatLit 0) (q(NormNum.isNat_ofNat $α $pz):) + match va, vb with + /- `0 ≤ 0` -/ + | .zero, .zero => pure <| .ok (q(le_refl (0:$α)):) + /- For numerals `ca` and `cb`, `ca + x ≤ cb + x` if `ca ≤ cb` -/ + | .add (b := a') (.const (e := xa) ca hypa) va', .add (.const (e := xb) cb hypb) vb' => do + unless va'.eq vb' do return .error notComparable + let rxa := NormNum.Result.ofRawRat ca xa hypa + let rxb := NormNum.Result.ofRawRat cb xb hypb + let NormNum.Result.isTrue pf ← NormNum.evalLE.core lα rxa rxb | return .error tooSmall + pure <| .ok (q(add_le_add_right (a := $a') $pf):) + /- For a numeral `ca ≤ 0`, `ca + x ≤ x` -/ + | .add (.const (e := xa) ca hypa) va', _ => do + unless va'.eq vb do return .error notComparable + let rxa := NormNum.Result.ofRawRat ca xa hypa + let NormNum.Result.isTrue pf ← NormNum.evalLE.core lα rxa rz | return .error tooSmall + pure <| .ok (q(add_le_of_nonpos_left (a := $b) $pf):) + /- For a numeral `0 ≤ cb`, `x ≤ cb + x` -/ + | _, .add (.const (e := xb) cb hypb) vb' => do + unless va.eq vb' do return .error notComparable + let rxb := NormNum.Result.ofRawRat cb xb hypb + let NormNum.Result.isTrue pf ← NormNum.evalLE.core lα rz rxb | return .error tooSmall + pure <| .ok (q(le_add_of_nonneg_left (a := $a) $pf):) + | _, _ => return .error notComparable + +/-- In a commutative semiring, given `Ring.ExSum` objects `va`, `vb` which differ by a positive +(additive) constant, construct a proof of `$a < $b`, where `a` (resp. `b`) is the expression in the +semiring to which `va` (resp. `vb`) evaluates. -/ +def evalLT {v : Level} {α : Q(Type v)} (_ : Q(StrictOrderedCommSemiring $α)) {a b : Q($α)} + (va : Ring.ExSum q(cs_of_socs $α) a) (vb : Ring.ExSum q(cs_of_socs $α) b) : + MetaM (Except ExceptType Q($a < $b)) := do + let lα : Q(LT $α) := q(lt_of_socs $α) + assumeInstancesCommute + let ⟨_, pz⟩ ← NormNum.mkOfNat α q(amwo_of_socs $α) (mkRawNatLit 0) + let rz : NormNum.Result q((0:$α)) := + NormNum.Result.isNat q(amwo_of_socs $α) (mkRawNatLit 0) (q(NormNum.isNat_ofNat $α $pz):) + match va, vb with + /- `0 < 0` -/ + | .zero, .zero => return .error tooSmall + /- For numerals `ca` and `cb`, `ca + x < cb + x` if `ca < cb` -/ + | .add (b := a') (.const (e := xa) ca hypa) va', .add (.const (e := xb) cb hypb) vb' => do + unless va'.eq vb' do return .error notComparable + let rxa := NormNum.Result.ofRawRat ca xa hypa + let rxb := NormNum.Result.ofRawRat cb xb hypb + let NormNum.Result.isTrue pf ← NormNum.evalLT.core lα rxa rxb | return .error tooSmall + pure <| .ok (q(add_lt_add_right $pf $a'):) + /- For a numeral `ca < 0`, `ca + x < x` -/ + | .add (.const (e := xa) ca hypa) va', _ => do + unless va'.eq vb do return .error notComparable + let rxa := NormNum.Result.ofRawRat ca xa hypa + let NormNum.Result.isTrue pf ← NormNum.evalLT.core lα rxa rz | return .error tooSmall + have pf : Q($xa < 0) := pf + pure <| .ok (q(add_lt_of_neg_left $b $pf):) + /- For a numeral `0 < cb`, `x < cb + x` -/ + | _, .add (.const (e := xb) cb hypb) vb' => do + unless va.eq vb' do return .error notComparable + let rxb := NormNum.Result.ofRawRat cb xb hypb + let NormNum.Result.isTrue pf ← NormNum.evalLT.core lα rz rxb | return .error tooSmall + pure <| .ok (q(lt_add_of_pos_left $a $pf):) + | _, _ => return .error notComparable + +theorem le_congr {α : Type*} [LE α] {a b c d : α} (h1 : a = b) (h2 : b ≤ c) (h3 : d = c) : + a ≤ d := by + rwa [h1, h3] + +theorem lt_congr {α : Type*} [LT α] {a b c d : α} (h1 : a = b) (h2 : b < c) (h3 : d = c) : + a < d := by + rwa [h1, h3] + +/-- Prove goals of the form `A ≤ B` in an ordered commutative semiring, if the ring-normal forms of +`A` and `B` differ by a nonnegative (additive) constant. -/ +def proveLE (g : MVarId) : MetaM Unit := do + let some (α, e₁, e₂) := (← whnfR <|← instantiateMVars <|← g.getType).le? + | throwError "ring failed: not of the form `A ≤ B`" + let .sort u ← whnf (← inferType α) | unreachable! + let v ← try u.dec catch _ => throwError "not a type{indentExpr α}" + have α : Q(Type v) := α + let sα ← synthInstanceQ q(OrderedCommSemiring $α) + assumeInstancesCommute + have e₁ : Q($α) := e₁; have e₂ : Q($α) := e₂ + let c ← mkCache q(cs_of_ocs $α) + let (⟨a, va, pa⟩, ⟨b, vb, pb⟩) + ← AtomM.run .instances do pure (← eval q(cs_of_ocs $α) c e₁, ← eval q(cs_of_ocs $α) c e₂) + match ← evalLE sα va vb with + | .ok p => g.assign q(le_congr $pa $p $pb) + | .error e => + let g' ← mkFreshExprMVar (← (← ringCleanupRef.get) q($a ≤ $b)) + match e with + | notComparable => + throwError "ring failed, ring expressions not equal up to an additive constant\n{g'.mvarId!}" + | tooSmall => throwError "comparison failed, LHS is larger\n{g'.mvarId!}" + +/-- Prove goals of the form `A < B` in an ordered commutative semiring, if the ring-normal forms of +`A` and `B` differ by a positive (additive) constant. -/ +def proveLT (g : MVarId) : MetaM Unit := do + let some (α, e₁, e₂) := (← whnfR <|← instantiateMVars <|← g.getType).lt? + | throwError "ring failed: not of the form `A < B`" + let .sort u ← whnf (← inferType α) | unreachable! + let v ← try u.dec catch _ => throwError "not a type{indentExpr α}" + have α : Q(Type v) := α + let sα ← synthInstanceQ q(StrictOrderedCommSemiring $α) + assumeInstancesCommute + have e₁ : Q($α) := e₁; have e₂ : Q($α) := e₂ + let c ← mkCache q(cs_of_socs $α) + let (⟨a, va, pa⟩, ⟨b, vb, pb⟩) + ← AtomM.run .instances do pure (← eval q(cs_of_socs $α) c e₁, ← eval q(cs_of_socs $α) c e₂) + match ← evalLT sα va vb with + | .ok p => g.assign q(lt_congr $pa $p $pb) + | .error e => + let g' ← mkFreshExprMVar (← (← ringCleanupRef.get) q($a < $b)) + match e with + | notComparable => + throwError "ring failed, ring expressions not equal up to an additive constant\n{g'.mvarId!}" + | tooSmall => throwError "comparison failed, LHS is at least as large\n{g'.mvarId!}" + +end Mathlib.Tactic.Ring diff --git a/Mathlib/Tactic/Says.lean b/Mathlib/Tactic/Says.lean index 87f93ae495ba1..29bd33ed183cf 100644 --- a/Mathlib/Tactic/Says.lean +++ b/Mathlib/Tactic/Says.lean @@ -40,7 +40,7 @@ register_option says.verify : Bool := register_option says.no_verify_in_CI : Bool := { defValue := false group := "says" - descr := "Disable reverification, even if `the `CI` environment variable is set." } + descr := "Disable reverification, even if the `CI` environment variable is set." } open Parser Tactic diff --git a/Mathlib/Tactic/Widget/InteractiveUnfold.lean b/Mathlib/Tactic/Widget/InteractiveUnfold.lean index b88ea0416d05d..7d3eae135546f 100644 --- a/Mathlib/Tactic/Widget/InteractiveUnfold.lean +++ b/Mathlib/Tactic/Widget/InteractiveUnfold.lean @@ -225,7 +225,7 @@ elab stx:"unfold?" : tactic => do /-- `#unfold? e` gives all unfolds of `e`. In tactic mode, use `unfold?` instead. -/ -syntax (name := unfoldCommand) "#unfold?" term : command +syntax (name := unfoldCommand) "#unfold? " term : command open Elab /-- Elaborate a `#unfold?` command. -/ diff --git a/Mathlib/Testing/SlimCheck/Testable.lean b/Mathlib/Testing/SlimCheck/Testable.lean index e4ece33e569d8..af3278d4fc34d 100644 --- a/Mathlib/Testing/SlimCheck/Testable.lean +++ b/Mathlib/Testing/SlimCheck/Testable.lean @@ -423,7 +423,7 @@ end Testable section PrintableProp -variable {α : Type*} {x y : α} +variable {α : Type*} instance Eq.printableProp [Repr α] {x y : α} : PrintableProp (x = y) where printProp := s!"{repr x} = {repr y}" diff --git a/Mathlib/Topology/AlexandrovDiscrete.lean b/Mathlib/Topology/AlexandrovDiscrete.lean index 517c0d910bfeb..455d1c1792743 100644 --- a/Mathlib/Topology/AlexandrovDiscrete.lean +++ b/Mathlib/Topology/AlexandrovDiscrete.lean @@ -139,7 +139,7 @@ lemma alexandrovDiscrete_iSup {t : ι → TopologicalSpace α} (_ : ∀ i, @Alex section variable [TopologicalSpace α] [TopologicalSpace β] [AlexandrovDiscrete α] [AlexandrovDiscrete β] - {s t : Set α} {a x y : α} + {s t : Set α} {a : α} @[simp] lemma isOpen_exterior : IsOpen (exterior s) := by rw [exterior_def]; exact isOpen_sInter fun _ ↦ And.left diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index 39bd780335969..7f721b2b8dbdd 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -79,7 +79,7 @@ theorem Filter.Tendsto.const_smul {f : β → α} {l : Filter β} {a : α} (hf : (c : M) : Tendsto (fun x => c • f x) l (𝓝 (c • a)) := ((continuous_const_smul _).tendsto _).comp hf -variable [TopologicalSpace β] {f : β → M} {g : β → α} {b : β} {s : Set β} +variable [TopologicalSpace β] {g : β → α} {b : β} {s : Set β} @[to_additive] nonrec theorem ContinuousWithinAt.const_smul (hg : ContinuousWithinAt g s b) (c : M) : diff --git a/Mathlib/Topology/Algebra/ContinuousAffineMap.lean b/Mathlib/Topology/Algebra/ContinuousAffineMap.lean index b2d7477d2d987..0c94ea673275e 100644 --- a/Mathlib/Topology/Algebra/ContinuousAffineMap.lean +++ b/Mathlib/Topology/Algebra/ContinuousAffineMap.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import Mathlib.LinearAlgebra.AffineSpace.AffineMap -import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Topology.Algebra.Module.Basic /-! diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index 00b7d7e857345..4e852f064519b 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -30,49 +30,45 @@ variable (F A B C D E : Type*) [Monoid A] [Monoid B] [Monoid C] [Monoid D] [Comm /-- The type of continuous additive monoid homomorphisms from `A` to `B`. When possible, instead of parametrizing results over `(f : ContinuousAddMonoidHom A B)`, -you should parametrize over `(F : Type*) [ContinuousAddMonoidHomClass F A B] (f : F)`. +you should parametrize +over `(F : Type*) [FunLike F A B] [ContinuousMapClass F A B] [AddMonoidHomClass F A B] (f : F)`. -When you extend this structure, make sure to extend `ContinuousAddMonoidHomClass`. -/ +When you extend this structure, +make sure to extend `ContinuousMapClass` and/or `AddMonoidHomClass`, if needed. -/ structure ContinuousAddMonoidHom (A B : Type*) [AddMonoid A] [AddMonoid B] [TopologicalSpace A] - [TopologicalSpace B] extends A →+ B where - /-- Proof of continuity of the Hom. -/ - continuous_toFun : @Continuous A B _ _ toFun + [TopologicalSpace B] extends A →+ B, C(A, B) /-- The type of continuous monoid homomorphisms from `A` to `B`. When possible, instead of parametrizing results over `(f : ContinuousMonoidHom A B)`, -you should parametrize over `(F : Type*) [ContinuousMonoidHomClass F A B] (f : F)`. +you should parametrize +over `(F : Type*) [FunLike F A B] [ContinuousMapClass F A B] [MonoidHomClass F A B] (f : F)`. -When you extend this structure, make sure to extend `ContinuousAddMonoidHomClass`. -/ +When you extend this structure, +make sure to extend `ContinuousMapClass` and/or `MonoidHomClass`, if needed. -/ @[to_additive "The type of continuous additive monoid homomorphisms from `A` to `B`."] -structure ContinuousMonoidHom extends A →* B where - /-- Proof of continuity of the Hom. -/ - continuous_toFun : @Continuous A B _ _ toFun +structure ContinuousMonoidHom extends A →* B, C(A, B) section /-- `ContinuousAddMonoidHomClass F A B` states that `F` is a type of continuous additive monoid homomorphisms. -You should also extend this typeclass when you extend `ContinuousAddMonoidHom`. -/ --- Porting note: Changed A B to outParam to help synthesizing order -class ContinuousAddMonoidHomClass (A B : outParam Type*) [AddMonoid A] [AddMonoid B] +Deprecated and changed from a `class` to a `structure`. +Use `[AddMonoidHomClass F A B] [ContinuousMapClass F A B]` instead. -/ +structure ContinuousAddMonoidHomClass (A B : outParam Type*) [AddMonoid A] [AddMonoid B] [TopologicalSpace A] [TopologicalSpace B] [FunLike F A B] - extends AddMonoidHomClass F A B : Prop where - /-- Proof of the continuity of the map. -/ - map_continuous (f : F) : Continuous f + extends AddMonoidHomClass F A B, ContinuousMapClass F A B : Prop /-- `ContinuousMonoidHomClass F A B` states that `F` is a type of continuous monoid homomorphisms. -You should also extend this typeclass when you extend `ContinuousMonoidHom`. -/ --- Porting note: Changed A B to outParam to help synthesizing order -@[to_additive] -class ContinuousMonoidHomClass (A B : outParam Type*) [Monoid A] [Monoid B] +Deprecated and changed from a `class` to a `structure`. +Use `[MonoidHomClass F A B] [ContinuousMapClass F A B]` instead. -/ +@[to_additive (attr := deprecated (since := "2024-10-08"))] +structure ContinuousMonoidHomClass (A B : outParam Type*) [Monoid A] [Monoid B] [TopologicalSpace A] [TopologicalSpace B] [FunLike F A B] - extends MonoidHomClass F A B : Prop where - /-- Proof of the continuity of the map. -/ - map_continuous (f : F) : Continuous f + extends MonoidHomClass F A B, ContinuousMapClass F A B : Prop end @@ -82,18 +78,18 @@ add_decl_doc ContinuousMonoidHom.toMonoidHom /-- Reinterpret a `ContinuousAddMonoidHom` as an `AddMonoidHom`. -/ add_decl_doc ContinuousAddMonoidHom.toAddMonoidHom --- See note [lower instance priority] -@[to_additive] -instance (priority := 100) ContinuousMonoidHomClass.toContinuousMapClass - [FunLike F A B] [ContinuousMonoidHomClass F A B] : ContinuousMapClass F A B := - { ‹ContinuousMonoidHomClass F A B› with } +/-- Reinterpret a `ContinuousMonoidHom` as a `ContinuousMap`. -/ +add_decl_doc ContinuousMonoidHom.toContinuousMap + +/-- Reinterpret a `ContinuousAddMonoidHom` as a `ContinuousMap`. -/ +add_decl_doc ContinuousAddMonoidHom.toContinuousMap namespace ContinuousMonoidHom variable {A B C D E} @[to_additive] -instance funLike : FunLike (ContinuousMonoidHom A B) A B where +instance instFunLike : FunLike (ContinuousMonoidHom A B) A B where coe f := f.toFun coe_injective' f g h := by obtain ⟨⟨⟨ _ , _ ⟩, _⟩, _⟩ := f @@ -101,53 +97,61 @@ instance funLike : FunLike (ContinuousMonoidHom A B) A B where congr @[to_additive] -instance ContinuousMonoidHomClass : ContinuousMonoidHomClass (ContinuousMonoidHom A B) A B where +instance instMonoidHomClass : MonoidHomClass (ContinuousMonoidHom A B) A B where map_mul f := f.map_mul' map_one f := f.map_one' + +@[to_additive] +instance instContinuousMapClass : ContinuousMapClass (ContinuousMonoidHom A B) A B where map_continuous f := f.continuous_toFun @[to_additive (attr := ext)] theorem ext {f g : ContinuousMonoidHom A B} (h : ∀ x, f x = g x) : f = g := DFunLike.ext _ _ h -/-- Reinterpret a `ContinuousMonoidHom` as a `ContinuousMap`. -/ -@[to_additive "Reinterpret a `ContinuousAddMonoidHom` as a `ContinuousMap`."] -def toContinuousMap (f : ContinuousMonoidHom A B) : C(A, B) := - { f with } - @[to_additive] theorem toContinuousMap_injective : Injective (toContinuousMap : _ → C(A, B)) := fun f g h => ext <| by convert DFunLike.ext_iff.1 h --- Porting note: Removed simps because given definition is not a constructor application -/-- Construct a `ContinuousMonoidHom` from a `Continuous` `MonoidHom`. -/ -@[to_additive "Construct a `ContinuousAddMonoidHom` from a `Continuous` `AddMonoidHom`."] -def mk' (f : A →* B) (hf : Continuous f) : ContinuousMonoidHom A B := - { f with continuous_toFun := (hf : Continuous f.toFun)} +@[deprecated (since := "2024-10-08")] protected alias mk' := mk + +@[deprecated (since := "2024-10-08")] +protected alias _root_.ContinuousAddMonoidHom.mk' := ContinuousAddMonoidHom.mk + +set_option linter.existingAttributeWarning false in +attribute [to_additive existing] ContinuousMonoidHom.mk' /-- Composition of two continuous homomorphisms. -/ @[to_additive (attr := simps!) "Composition of two continuous homomorphisms."] def comp (g : ContinuousMonoidHom B C) (f : ContinuousMonoidHom A B) : ContinuousMonoidHom A C := - mk' (g.toMonoidHom.comp f.toMonoidHom) (g.continuous_toFun.comp f.continuous_toFun) + ⟨g.toMonoidHom.comp f.toMonoidHom, (map_continuous g).comp (map_continuous f)⟩ /-- Product of two continuous homomorphisms on the same space. -/ -@[to_additive (attr := simps!) "Product of two continuous homomorphisms on the same space."] +@[to_additive (attr := simps!) prod "Product of two continuous homomorphisms on the same space."] def prod (f : ContinuousMonoidHom A B) (g : ContinuousMonoidHom A C) : ContinuousMonoidHom A (B × C) := - mk' (f.toMonoidHom.prod g.toMonoidHom) (f.continuous_toFun.prod_mk g.continuous_toFun) + ⟨f.toMonoidHom.prod g.toMonoidHom, f.continuous_toFun.prod_mk g.continuous_toFun⟩ /-- Product of two continuous homomorphisms on different spaces. -/ -@[to_additive (attr := simps!) "Product of two continuous homomorphisms on different spaces."] -def prod_map (f : ContinuousMonoidHom A C) (g : ContinuousMonoidHom B D) : +@[to_additive (attr := simps!) prodMap + "Product of two continuous homomorphisms on different spaces."] +def prodMap (f : ContinuousMonoidHom A C) (g : ContinuousMonoidHom B D) : ContinuousMonoidHom (A × B) (C × D) := - mk' (f.toMonoidHom.prodMap g.toMonoidHom) (f.continuous_toFun.prod_map g.continuous_toFun) + ⟨f.toMonoidHom.prodMap g.toMonoidHom, f.continuous_toFun.prodMap g.continuous_toFun⟩ + +@[deprecated (since := "2024-10-05")] alias prod_map := prodMap +@[deprecated (since := "2024-10-05")] +alias _root_.ContinuousAddMonoidHom.sum_map := ContinuousAddMonoidHom.prodMap + +set_option linter.existingAttributeWarning false in +attribute [to_additive existing] prod_map variable (A B C D E) /-- The trivial continuous homomorphism. -/ @[to_additive (attr := simps!) "The trivial continuous homomorphism."] def one : ContinuousMonoidHom A B := - mk' 1 continuous_const + ⟨1, continuous_const⟩ @[to_additive] instance : Inhabited (ContinuousMonoidHom A B) := @@ -156,19 +160,19 @@ instance : Inhabited (ContinuousMonoidHom A B) := /-- The identity continuous homomorphism. -/ @[to_additive (attr := simps!) "The identity continuous homomorphism."] def id : ContinuousMonoidHom A A := - mk' (MonoidHom.id A) continuous_id + ⟨.id A, continuous_id⟩ /-- The continuous homomorphism given by projection onto the first factor. -/ @[to_additive (attr := simps!) "The continuous homomorphism given by projection onto the first factor."] def fst : ContinuousMonoidHom (A × B) A := - mk' (MonoidHom.fst A B) continuous_fst + ⟨MonoidHom.fst A B, continuous_fst⟩ /-- The continuous homomorphism given by projection onto the second factor. -/ @[to_additive (attr := simps!) "The continuous homomorphism given by projection onto the second factor."] def snd : ContinuousMonoidHom (A × B) B := - mk' (MonoidHom.snd A B) continuous_snd + ⟨MonoidHom.snd A B, continuous_snd⟩ /-- The continuous homomorphism given by inclusion of the first factor. -/ @[to_additive (attr := simps!) @@ -195,12 +199,12 @@ def swap : ContinuousMonoidHom (A × B) (B × A) := /-- The continuous homomorphism given by multiplication. -/ @[to_additive (attr := simps!) "The continuous homomorphism given by addition."] def mul : ContinuousMonoidHom (E × E) E := - mk' mulMonoidHom continuous_mul + ⟨mulMonoidHom, continuous_mul⟩ /-- The continuous homomorphism given by inversion. -/ @[to_additive (attr := simps!) "The continuous homomorphism given by negation."] def inv : ContinuousMonoidHom E E := - mk' invMonoidHom continuous_inv + ⟨invMonoidHom, continuous_inv⟩ variable {A B C D E} @@ -208,7 +212,7 @@ variable {A B C D E} @[to_additive (attr := simps!) "Coproduct of two continuous homomorphisms to the same space."] def coprod (f : ContinuousMonoidHom A E) (g : ContinuousMonoidHom B E) : ContinuousMonoidHom (A × B) E := - (mul E).comp (f.prod_map g) + (mul E).comp (f.prodMap g) @[to_additive] instance : CommGroup (ContinuousMonoidHom A E) where @@ -265,7 +269,7 @@ instance [T2Space B] : T2Space (ContinuousMonoidHom A B) := instance : TopologicalGroup (ContinuousMonoidHom A E) := let hi := inducing_toContinuousMap A E let hc := hi.continuous - { continuous_mul := hi.continuous_iff.mpr (continuous_mul.comp (Continuous.prod_map hc hc)) + { continuous_mul := hi.continuous_iff.mpr (continuous_mul.comp (Continuous.prodMap hc hc)) continuous_inv := hi.continuous_iff.mpr (continuous_inv.comp hc) } @[to_additive] @@ -280,7 +284,7 @@ theorem continuous_comp [LocallyCompactSpace B] : Continuous fun f : ContinuousMonoidHom A B × ContinuousMonoidHom B C => f.2.comp f.1 := (inducing_toContinuousMap A C).continuous_iff.2 <| ContinuousMap.continuous_comp'.comp - ((inducing_toContinuousMap A B).prod_map (inducing_toContinuousMap B C)).continuous + ((inducing_toContinuousMap A B).prodMap (inducing_toContinuousMap B C)).continuous @[to_additive] theorem continuous_comp_left (f : ContinuousMonoidHom A B) : diff --git a/Mathlib/Topology/Algebra/Field.lean b/Mathlib/Topology/Algebra/Field.lean index 0b7ccfdded133..c897f0bb8fa18 100644 --- a/Mathlib/Topology/Algebra/Field.lean +++ b/Mathlib/Topology/Algebra/Field.lean @@ -101,7 +101,7 @@ open Topology theorem IsLocalMin.inv {f : α → β} {a : α} (h1 : IsLocalMin f a) (h2 : ∀ᶠ z in 𝓝 a, 0 < f z) : IsLocalMax f⁻¹ a := by - filter_upwards [h1, h2] with z h3 h4 using(inv_le_inv h4 h2.self_of_nhds).mpr h3 + filter_upwards [h1, h2] with z h3 h4 using(inv_le_inv₀ h4 h2.self_of_nhds).mpr h3 end LocalExtr diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 99386fcf912e5..49c268a901a0c 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -536,7 +536,7 @@ end OrderedCommGroup @[to_additive] instance [TopologicalSpace H] [Group H] [TopologicalGroup H] : TopologicalGroup (G × H) where - continuous_inv := continuous_inv.prod_map continuous_inv + continuous_inv := continuous_inv.prodMap continuous_inv @[to_additive] instance Pi.topologicalGroup {C : β → Type*} [∀ b, TopologicalSpace (C b)] [∀ b, Group (C b)] @@ -1426,7 +1426,7 @@ theorem Subgroup.properlyDiscontinuousSMul_opposite_of_tendsto_cofinite (S : Sub (hS : Tendsto S.subtype cofinite (cocompact G)) : ProperlyDiscontinuousSMul S.op G := { finite_disjoint_inter_image := by intro K L hK hL - have : Continuous fun p : G × G => (p.1⁻¹, p.2) := continuous_inv.prod_map continuous_id + have : Continuous fun p : G × G => (p.1⁻¹, p.2) := continuous_inv.prodMap continuous_id have H : Set.Finite _ := hS ((hK.prod hL).image (continuous_mul.comp this)).compl_mem_cocompact simp only [preimage_compl, compl_compl, coeSubtype, comp_apply] at H diff --git a/Mathlib/Topology/Algebra/GroupCompletion.lean b/Mathlib/Topology/Algebra/GroupCompletion.lean index cadc887e6ae6e..5935b46b343e8 100644 --- a/Mathlib/Topology/Algebra/GroupCompletion.lean +++ b/Mathlib/Topology/Algebra/GroupCompletion.lean @@ -136,8 +136,8 @@ instance : SubNegMonoid (Completion α) := zsmul_succ' := fun n a ↦ Completion.induction_on a (isClosed_eq continuous_map <| continuous_map₂ continuous_map continuous_id) fun a ↦ - show Int.ofNat n.succ • (a : Completion α) = _ by - rw [← coe_smul, show Int.ofNat n.succ • a = Int.ofNat n • a + a from + show (n.succ : ℤ) • (a : Completion α) = _ by + rw [← coe_smul, show (n.succ : ℤ) • a = (n : ℤ) • a + a from SubNegMonoid.zsmul_succ' n a, coe_add, coe_smul] zsmul_neg' := fun n a ↦ Completion.induction_on a diff --git a/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean b/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean index 0391d939b2f46..1b43517bdbae1 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean @@ -447,9 +447,8 @@ end Semiring section Ring -variable {R M M' N N' ι : Type*} [Ring R] [AddCommGroup M] [Module R M] [TopologicalSpace M] - [AddCommGroup M'] [Module R M'] [TopologicalSpace M'] [AddCommGroup N] [Module R N] - [TopologicalSpace N] [AddCommGroup N'] [Module R N'] [TopologicalSpace N'] {n : ℕ} +variable {R M N ι : Type*} [Ring R] [AddCommGroup M] [Module R M] [TopologicalSpace M] + [AddCommGroup N] [Module R N] [TopologicalSpace N] (f g : M [⋀^ι]→L[R] N) @[simp] @@ -489,10 +488,9 @@ end Ring section CommSemiring -variable {R M M' N N' ι : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] - [TopologicalSpace M] [AddCommMonoid M'] [Module R M'] [TopologicalSpace M'] [AddCommMonoid N] - [Module R N] [TopologicalSpace N] [AddCommMonoid N'] [Module R N'] [TopologicalSpace N'] {n : ℕ} - (f g : M [⋀^ι]→L[R] N) +variable {R M N ι : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + [TopologicalSpace M] [AddCommMonoid N] [Module R N] [TopologicalSpace N] + (f : M [⋀^ι]→L[R] N) theorem map_piecewise_smul [DecidableEq ι] (c : ι → R) (m : ι → M) (s : Finset ι) : f (s.piecewise (fun i => c i • m i) m) = (∏ i ∈ s, c i) • f m := @@ -561,7 +559,7 @@ end Module section SMulRight -variable {R A M N ι : Type*} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] +variable {R M N ι : Type*} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [TopologicalSpace R] [TopologicalSpace M] [TopologicalSpace N] [ContinuousSMul R N] (f : M [⋀^ι]→L[R] R) (z : N) @@ -603,7 +601,7 @@ namespace ContinuousMultilinearMap variable {R M N ι : Type*} [Semiring R] [AddCommMonoid M] [Module R M] [TopologicalSpace M] [AddCommGroup N] [Module R N] [TopologicalSpace N] [TopologicalAddGroup N] [Fintype ι] - [DecidableEq ι] (f g : ContinuousMultilinearMap R (fun _ : ι => M) N) + [DecidableEq ι] (f : ContinuousMultilinearMap R (fun _ : ι => M) N) /-- Alternatization of a continuous multilinear map. -/ @[simps (config := .lemmasOnly) apply_toContinuousMultilinearMap] diff --git a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean index 9f613d081db1f..32efbbd5a39c7 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean @@ -86,7 +86,7 @@ theorem completeSpace (h : RestrictGenTopology {s : Set (ι → E) | IsVonNBound CompleteSpace (E [⋀^ι]→L[𝕜] F) := by have := ContinuousMultilinearMap.completeSpace (F := F) h rw [completeSpace_iff_isComplete_range - isUniformEmbedding_toContinuousMultilinearMap.toUniformInducing] + isUniformEmbedding_toContinuousMultilinearMap.isUniformInducing] apply isClosed_range_toContinuousMultilinearMap.isComplete instance instCompleteSpace [TopologicalAddGroup E] [SequentialSpace (ι → E)] : diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 76d6e898d399b..a1479380641d5 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -7,7 +7,6 @@ Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo, Yury Kudryashov, Fréd import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Algebra.MulAction import Mathlib.Topology.Algebra.UniformGroup -import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Topology.UniformSpace.UniformEmbedding import Mathlib.Algebra.Algebra.Defs import Mathlib.LinearAlgebra.Projection @@ -1817,8 +1816,8 @@ theorem trans_toLinearEquiv (e₁ : M₁ ≃SL[σ₁₂] M₂) (e₂ : M₂ ≃S def prod [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) (e' : M₃ ≃L[R₁] M₄) : (M₁ × M₃) ≃L[R₁] M₂ × M₄ := { e.toLinearEquiv.prod e'.toLinearEquiv with - continuous_toFun := e.continuous_toFun.prod_map e'.continuous_toFun - continuous_invFun := e.continuous_invFun.prod_map e'.continuous_invFun } + continuous_toFun := e.continuous_toFun.prodMap e'.continuous_toFun + continuous_invFun := e.continuous_invFun.prodMap e'.continuous_invFun } @[simp, norm_cast] theorem prod_apply [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) diff --git a/Mathlib/Topology/Algebra/Module/Cardinality.lean b/Mathlib/Topology/Algebra/Module/Cardinality.lean index c45fb8205558f..b5c7c62b830cd 100644 --- a/Mathlib/Topology/Algebra/Module/Cardinality.lean +++ b/Mathlib/Topology/Algebra/Module/Cardinality.lean @@ -77,7 +77,7 @@ lemma cardinal_eq_of_mem_nhds_zero simp_rw [← inv_pow] apply tendsto_pow_atTop_nhds_zero_of_norm_lt_one rw [norm_inv] - exact inv_lt_one hc + exact inv_lt_one_of_one_lt₀ hc exact Tendsto.smul_const this x rw [zero_smul] at this filter_upwards [this hs] with n (hn : (c ^ n)⁻¹ • x ∈ s) diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean index c73eb8c4be001..baf4b78201f9a 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean @@ -107,7 +107,7 @@ theorem completeSpace (h : RestrictGenTopology {s : Set (Π i, E i) | IsVonNBoun have H : ∀ {m : Π i, E i}, Continuous fun f : (Π i, E i) →ᵤ[{s | IsVonNBounded 𝕜 s}] F ↦ toFun _ f m := (uniformContinuous_eval (isVonNBounded_covers) _).continuous - rw [completeSpace_iff_isComplete_range isUniformEmbedding_toUniformOnFun.toUniformInducing, + rw [completeSpace_iff_isComplete_range isUniformEmbedding_toUniformOnFun.isUniformInducing, range_toUniformOnFun] simp only [setOf_and, setOf_forall] apply_rules [IsClosed.isComplete, IsClosed.inter] diff --git a/Mathlib/Topology/Algebra/Module/StrongTopology.lean b/Mathlib/Topology/Algebra/Module/StrongTopology.lean index 6bb9968ff157e..8e988265fd745 100644 --- a/Mathlib/Topology/Algebra/Module/StrongTopology.lean +++ b/Mathlib/Topology/Algebra/Module/StrongTopology.lean @@ -55,7 +55,8 @@ uniform convergence, bounded convergence -/ -open scoped Topology UniformConvergence +open scoped Topology UniformConvergence Uniformity +open Filter Set Function Bornology section General @@ -92,7 +93,7 @@ instance instTopologicalSpace [TopologicalSpace F] [TopologicalAddGroup F] (𝔖 (DFunLike.coe : (UniformConvergenceCLM σ F 𝔖) → (E →ᵤ[𝔖] F)) theorem topologicalSpace_eq [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (Set E)) : - instTopologicalSpace σ F 𝔖 = TopologicalSpace.induced DFunLike.coe + instTopologicalSpace σ F 𝔖 = TopologicalSpace.induced (UniformOnFun.ofFun 𝔖 ∘ DFunLike.coe) (UniformOnFun.topologicalSpace E F 𝔖) := by rw [instTopologicalSpace] congr @@ -103,12 +104,13 @@ that this has nice definitional properties. -/ instance instUniformSpace [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (Set E)) : UniformSpace (UniformConvergenceCLM σ F 𝔖) := UniformSpace.replaceTopology - ((UniformOnFun.uniformSpace E F 𝔖).comap - (DFunLike.coe : (UniformConvergenceCLM σ F 𝔖) → (E →ᵤ[𝔖] F))) + ((UniformOnFun.uniformSpace E F 𝔖).comap (UniformOnFun.ofFun 𝔖 ∘ DFunLike.coe)) (by rw [UniformConvergenceCLM.instTopologicalSpace, UniformAddGroup.toUniformSpace_eq]; rfl) theorem uniformSpace_eq [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (Set E)) : - instUniformSpace σ F 𝔖 = UniformSpace.comap DFunLike.coe (UniformOnFun.uniformSpace E F 𝔖) := by + instUniformSpace σ F 𝔖 = + UniformSpace.comap (UniformOnFun.ofFun 𝔖 ∘ DFunLike.coe) + (UniformOnFun.uniformSpace E F 𝔖) := by rw [instUniformSpace, UniformSpace.replaceTopology_eq] @[simp] @@ -118,7 +120,7 @@ theorem uniformity_toTopologicalSpace_eq [UniformSpace F] [UniformAddGroup F] ( rfl theorem isUniformEmbedding_coeFn [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (Set E)) : - IsUniformEmbedding (α := UniformConvergenceCLM σ F 𝔖) (β := E →ᵤ[𝔖] F) DFunLike.coe := + IsUniformEmbedding (α := UniformConvergenceCLM σ F 𝔖) (UniformOnFun.ofFun 𝔖 ∘ DFunLike.coe) := ⟨⟨rfl⟩, DFunLike.coe_injective⟩ @[deprecated (since := "2024-10-01")] @@ -132,6 +134,11 @@ theorem embedding_coeFn [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (Set E) instance instAddCommGroup [TopologicalSpace F] [TopologicalAddGroup F] (𝔖 : Set (Set E)) : AddCommGroup (UniformConvergenceCLM σ F 𝔖) := ContinuousLinearMap.addCommGroup +@[simp] +theorem coe_zero [TopologicalSpace F] [TopologicalAddGroup F] (𝔖 : Set (Set E)) : + ⇑(0 : UniformConvergenceCLM σ F 𝔖) = 0 := + rfl + instance instUniformAddGroup [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (Set E)) : UniformAddGroup (UniformConvergenceCLM σ F 𝔖) := by let φ : (UniformConvergenceCLM σ F 𝔖) →+ E →ᵤ[𝔖] F := @@ -145,7 +152,7 @@ instance instTopologicalAddGroup [TopologicalSpace F] [TopologicalAddGroup F] infer_instance theorem t2Space [TopologicalSpace F] [TopologicalAddGroup F] [T2Space F] - (𝔖 : Set (Set E)) (h𝔖 : ⋃₀ 𝔖 = Set.univ) : T2Space (UniformConvergenceCLM σ F 𝔖) := by + (𝔖 : Set (Set E)) (h𝔖 : ⋃₀ 𝔖 = univ) : T2Space (UniformConvergenceCLM σ F 𝔖) := by letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform haveI : T2Space (E →ᵤ[𝔖] F) := UniformOnFun.t2Space_of_covering h𝔖 @@ -161,7 +168,7 @@ instance instModule (R : Type*) [Semiring R] [Module R F] [SMulCommClass 𝕜₂ theorem continuousSMul [RingHomSurjective σ] [RingHomIsometric σ] [TopologicalSpace F] [TopologicalAddGroup F] [ContinuousSMul 𝕜₂ F] (𝔖 : Set (Set E)) - (h𝔖₃ : ∀ S ∈ 𝔖, Bornology.IsVonNBounded 𝕜₁ S) : + (h𝔖₃ : ∀ S ∈ 𝔖, IsVonNBounded 𝕜₁ S) : ContinuousSMul 𝕜₂ (UniformConvergenceCLM σ F 𝔖) := by letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform @@ -188,11 +195,66 @@ theorem hasBasis_nhds_zero [TopologicalSpace F] [TopologicalAddGroup F] { f : UniformConvergenceCLM σ F 𝔖 | ∀ x ∈ SV.1, f x ∈ SV.2 } := hasBasis_nhds_zero_of_basis σ F 𝔖 h𝔖₁ h𝔖₂ (𝓝 0).basis_sets +theorem nhds_zero_eq_of_basis [TopologicalSpace F] [TopologicalAddGroup F] (𝔖 : Set (Set E)) + {ι : Type*} {p : ι → Prop} {b : ι → Set F} (h : (𝓝 0 : Filter F).HasBasis p b) : + 𝓝 (0 : UniformConvergenceCLM σ F 𝔖) = + ⨅ (s : Set E) (_ : s ∈ 𝔖) (i : ι) (_ : p i), + 𝓟 {f : UniformConvergenceCLM σ F 𝔖 | MapsTo f s (b i)} := by + letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F + haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform + rw [(embedding_coeFn σ F 𝔖).toInducing.nhds_eq_comap, + UniformOnFun.nhds_eq_of_basis _ _ h.uniformity_of_nhds_zero] + simp [MapsTo] + +theorem nhds_zero_eq [TopologicalSpace F] [TopologicalAddGroup F] (𝔖 : Set (Set E)) : + 𝓝 (0 : UniformConvergenceCLM σ F 𝔖) = + ⨅ s ∈ 𝔖, ⨅ t ∈ 𝓝 (0 : F), + 𝓟 {f : UniformConvergenceCLM σ F 𝔖 | MapsTo f s t} := + nhds_zero_eq_of_basis _ _ _ (𝓝 0).basis_sets + +variable {F} in +theorem eventually_nhds_zero_mapsTo [TopologicalSpace F] [TopologicalAddGroup F] + {𝔖 : Set (Set E)} {s : Set E} (hs : s ∈ 𝔖) {U : Set F} (hu : U ∈ 𝓝 0) : + ∀ᶠ f : UniformConvergenceCLM σ F 𝔖 in 𝓝 0, MapsTo f s U := by + rw [nhds_zero_eq] + apply_rules [mem_iInf_of_mem, mem_principal_self] + +variable {σ F} in +theorem isVonNBounded_image2_apply {R : Type*} [SeminormedRing R] + [TopologicalSpace F] [TopologicalAddGroup F] + [Module R F] [ContinuousConstSMul R F] [SMulCommClass 𝕜₂ R F] + {𝔖 : Set (Set E)} {S : Set (UniformConvergenceCLM σ F 𝔖)} (hS : IsVonNBounded R S) + {s : Set E} (hs : s ∈ 𝔖) : IsVonNBounded R (Set.image2 (fun f x ↦ f x) S s) := by + intro U hU + filter_upwards [hS (eventually_nhds_zero_mapsTo σ hs hU)] with c hc + rw [image2_subset_iff] + intro f hf x hx + rcases hc hf with ⟨g, hg, rfl⟩ + exact smul_mem_smul_set (hg hx) + +variable {σ F} in +/-- A set `S` of continuous linear maps with topology of uniform convergence on sets `s ∈ 𝔖` +is von Neumann bounded iff for any `s ∈ 𝔖`, +the set `{f x | (f ∈ S) (x ∈ s)}` is von Neumann bounded. -/ +theorem isVonNBounded_iff {R : Type*} [NormedDivisionRing R] + [TopologicalSpace F] [TopologicalAddGroup F] + [Module R F] [ContinuousConstSMul R F] [SMulCommClass 𝕜₂ R F] + {𝔖 : Set (Set E)} {S : Set (UniformConvergenceCLM σ F 𝔖)} : + IsVonNBounded R S ↔ ∀ s ∈ 𝔖, IsVonNBounded R (Set.image2 (fun f x ↦ f x) S s) := by + refine ⟨fun hS s hs ↦ isVonNBounded_image2_apply hS hs, fun h ↦ ?_⟩ + simp_rw [isVonNBounded_iff_absorbing_le, nhds_zero_eq, le_iInf_iff, le_principal_iff] + intro s hs U hU + rw [Filter.mem_absorbing, Absorbs] + filter_upwards [h s hs hU, eventually_ne_cobounded 0] with c hc hc₀ f hf + rw [mem_smul_set_iff_inv_smul_mem₀ hc₀] + intro x hx + simpa only [mem_smul_set_iff_inv_smul_mem₀ hc₀] using hc (mem_image2_of_mem hf hx) + instance instUniformContinuousConstSMul (M : Type*) [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜₂ M F] [UniformSpace F] [UniformAddGroup F] [UniformContinuousConstSMul M F] (𝔖 : Set (Set E)) : UniformContinuousConstSMul M (UniformConvergenceCLM σ F 𝔖) := - (isUniformEmbedding_coeFn σ F 𝔖).toUniformInducing.uniformContinuousConstSMul fun _ _ ↦ by rfl + (isUniformEmbedding_coeFn σ F 𝔖).isUniformInducing.uniformContinuousConstSMul fun _ _ ↦ by rfl instance instContinuousConstSMul (M : Type*) [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜₂ M F] @@ -243,7 +305,7 @@ variable {𝕜₁ 𝕜₂ 𝕜₃ : Type*} [NormedField 𝕜₁] [NormedField the operator norm when `E` and `F` are normed spaces. -/ instance topologicalSpace [TopologicalSpace F] [TopologicalAddGroup F] : TopologicalSpace (E →SL[σ] F) := - UniformConvergenceCLM.instTopologicalSpace σ F { S | Bornology.IsVonNBounded 𝕜₁ S } + UniformConvergenceCLM.instTopologicalSpace σ F { S | IsVonNBounded 𝕜₁ S } instance topologicalAddGroup [TopologicalSpace F] [TopologicalAddGroup F] : TopologicalAddGroup (E →SL[σ] F) := @@ -251,10 +313,10 @@ instance topologicalAddGroup [TopologicalSpace F] [TopologicalAddGroup F] : instance continuousSMul [RingHomSurjective σ] [RingHomIsometric σ] [TopologicalSpace F] [TopologicalAddGroup F] [ContinuousSMul 𝕜₂ F] : ContinuousSMul 𝕜₂ (E →SL[σ] F) := - UniformConvergenceCLM.continuousSMul σ F { S | Bornology.IsVonNBounded 𝕜₁ S } fun _ hs => hs + UniformConvergenceCLM.continuousSMul σ F { S | IsVonNBounded 𝕜₁ S } fun _ hs => hs instance uniformSpace [UniformSpace F] [UniformAddGroup F] : UniformSpace (E →SL[σ] F) := - UniformConvergenceCLM.instUniformSpace σ F { S | Bornology.IsVonNBounded 𝕜₁ S } + UniformConvergenceCLM.instUniformSpace σ F { S | IsVonNBounded 𝕜₁ S } instance uniformAddGroup [UniformSpace F] [UniformAddGroup F] : UniformAddGroup (E →SL[σ] F) := UniformConvergenceCLM.instUniformAddGroup σ F _ @@ -263,19 +325,19 @@ instance [TopologicalSpace F] [TopologicalAddGroup F] [ContinuousSMul 𝕜₁ E] T2Space (E →SL[σ] F) := UniformConvergenceCLM.t2Space σ F _ (Set.eq_univ_of_forall fun x => - Set.mem_sUnion_of_mem (Set.mem_singleton x) (Bornology.isVonNBounded_singleton x)) + Set.mem_sUnion_of_mem (Set.mem_singleton x) (isVonNBounded_singleton x)) protected theorem hasBasis_nhds_zero_of_basis [TopologicalSpace F] [TopologicalAddGroup F] {ι : Type*} {p : ι → Prop} {b : ι → Set F} (h : (𝓝 0 : Filter F).HasBasis p b) : - (𝓝 (0 : E →SL[σ] F)).HasBasis (fun Si : Set E × ι => Bornology.IsVonNBounded 𝕜₁ Si.1 ∧ p Si.2) + (𝓝 (0 : E →SL[σ] F)).HasBasis (fun Si : Set E × ι => IsVonNBounded 𝕜₁ Si.1 ∧ p Si.2) fun Si => { f : E →SL[σ] F | ∀ x ∈ Si.1, f x ∈ b Si.2 } := - UniformConvergenceCLM.hasBasis_nhds_zero_of_basis σ F { S | Bornology.IsVonNBounded 𝕜₁ S } - ⟨∅, Bornology.isVonNBounded_empty 𝕜₁ E⟩ - (directedOn_of_sup_mem fun _ _ => Bornology.IsVonNBounded.union) h + UniformConvergenceCLM.hasBasis_nhds_zero_of_basis σ F { S | IsVonNBounded 𝕜₁ S } + ⟨∅, isVonNBounded_empty 𝕜₁ E⟩ + (directedOn_of_sup_mem fun _ _ => IsVonNBounded.union) h protected theorem hasBasis_nhds_zero [TopologicalSpace F] [TopologicalAddGroup F] : (𝓝 (0 : E →SL[σ] F)).HasBasis - (fun SV : Set E × Set F => Bornology.IsVonNBounded 𝕜₁ SV.1 ∧ SV.2 ∈ (𝓝 0 : Filter F)) + (fun SV : Set E × Set F => IsVonNBounded 𝕜₁ SV.1 ∧ SV.2 ∈ (𝓝 0 : Filter F)) fun SV => { f : E →SL[σ] F | ∀ x ∈ SV.1, f x ∈ SV.2 } := ContinuousLinearMap.hasBasis_nhds_zero_of_basis (𝓝 0).basis_sets @@ -298,6 +360,51 @@ instance continuousConstSMul {M : Type*} [Monoid M] [DistribMulAction M F] [SMul ContinuousConstSMul M (E →SL[σ] F) := UniformConvergenceCLM.instContinuousConstSMul σ F _ _ +protected theorem nhds_zero_eq_of_basis [TopologicalSpace F] [TopologicalAddGroup F] + {ι : Type*} {p : ι → Prop} {b : ι → Set F} (h : (𝓝 0 : Filter F).HasBasis p b) : + 𝓝 (0 : E →SL[σ] F) = + ⨅ (s : Set E) (_ : IsVonNBounded 𝕜₁ s) (i : ι) (_ : p i), + 𝓟 {f : E →SL[σ] F | MapsTo f s (b i)} := + UniformConvergenceCLM.nhds_zero_eq_of_basis _ _ _ h + +protected theorem nhds_zero_eq [TopologicalSpace F] [TopologicalAddGroup F] : + 𝓝 (0 : E →SL[σ] F) = + ⨅ (s : Set E) (_ : IsVonNBounded 𝕜₁ s) (U : Set F) (_ : U ∈ 𝓝 0), + 𝓟 {f : E →SL[σ] F | MapsTo f s U} := + UniformConvergenceCLM.nhds_zero_eq .. + +/-- If `s` is a von Neumann bounded set and `U` is a neighbourhood of zero, +then sufficiently small continuous linear maps map `s` to `U`. -/ +theorem eventually_nhds_zero_mapsTo [TopologicalSpace F] [TopologicalAddGroup F] + {s : Set E} (hs : IsVonNBounded 𝕜₁ s) {U : Set F} (hu : U ∈ 𝓝 0) : + ∀ᶠ f : E →SL[σ] F in 𝓝 0, MapsTo f s U := + UniformConvergenceCLM.eventually_nhds_zero_mapsTo _ hs hu + +/-- If `S` is a von Neumann bounded set of continuous linear maps `f : E →SL[σ] F` +and `s` is a von Neumann bounded set in the domain, +then the set `{f x | (f ∈ S) (x ∈ s)}` is von Neumann bounded. + +See also `isVonNBounded_iff` for an `Iff` version with stronger typeclass assumptions. -/ +theorem isVonNBounded_image2_apply {R : Type*} [SeminormedRing R] + [TopologicalSpace F] [TopologicalAddGroup F] + [Module R F] [ContinuousConstSMul R F] [SMulCommClass 𝕜₂ R F] + {S : Set (E →SL[σ] F)} (hS : IsVonNBounded R S) {s : Set E} (hs : IsVonNBounded 𝕜₁ s) : + IsVonNBounded R (Set.image2 (fun f x ↦ f x) S s) := + UniformConvergenceCLM.isVonNBounded_image2_apply hS hs + +/-- A set `S` of continuous linear maps is von Neumann bounded +iff for any von Neumann bounded set `s`, +the set `{f x | (f ∈ S) (x ∈ s)}` is von Neumann bounded. + +For the forward implication with weaker typeclass assumptions, see `isVonNBounded_image2_apply`. -/ +theorem isVonNBounded_iff {R : Type*} [NormedDivisionRing R] + [TopologicalSpace F] [TopologicalAddGroup F] + [Module R F] [ContinuousConstSMul R F] [SMulCommClass 𝕜₂ R F] + {S : Set (E →SL[σ] F)} : + IsVonNBounded R S ↔ + ∀ s, IsVonNBounded 𝕜₁ s → IsVonNBounded R (Set.image2 (fun f x ↦ f x) S s) := + UniformConvergenceCLM.isVonNBounded_iff + variable (G) [TopologicalSpace F] [TopologicalSpace G] /-- Pre-composition by a *fixed* continuous linear map as a continuous linear map. diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index a8001b461554b..126c6636dd313 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -7,8 +7,8 @@ import Mathlib.Algebra.BigOperators.Finprod import Mathlib.Order.Filter.Pointwise import Mathlib.Topology.Algebra.MulAction import Mathlib.Algebra.BigOperators.Pi -import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Algebra.Group.ULift +import Mathlib.Topology.ContinuousMap.Defs /-! # Theory of topological monoids @@ -75,7 +75,7 @@ instance ContinuousMul.to_continuousSMul : ContinuousSMul M M := instance ContinuousMul.to_continuousSMul_op : ContinuousSMul Mᵐᵒᵖ M := ⟨show Continuous ((fun p : M × M => p.1 * p.2) ∘ Prod.swap ∘ Prod.map MulOpposite.unop id) from continuous_mul.comp <| - continuous_swap.comp <| Continuous.prod_map MulOpposite.continuous_unop continuous_id⟩ + continuous_swap.comp <| Continuous.prodMap MulOpposite.continuous_unop continuous_id⟩ @[to_additive] theorem ContinuousMul.induced {α : Type*} {β : Type*} {F : Type*} [FunLike F α β] [MulOneClass α] diff --git a/Mathlib/Topology/Algebra/MulAction.lean b/Mathlib/Topology/Algebra/MulAction.lean index b17c9a9e4bb1a..91dc843aa0507 100644 --- a/Mathlib/Topology/Algebra/MulAction.lean +++ b/Mathlib/Topology/Algebra/MulAction.lean @@ -132,13 +132,13 @@ action is."] instance ContinuousSMul.op [SMul Mᵐᵒᵖ X] [IsCentralScalar M X] : ContinuousSMul Mᵐᵒᵖ X := ⟨by suffices Continuous fun p : M × X => MulOpposite.op p.fst • p.snd from - this.comp (MulOpposite.continuous_unop.prod_map continuous_id) + this.comp (MulOpposite.continuous_unop.prodMap continuous_id) simpa only [op_smul_eq_smul] using (continuous_smul : Continuous fun p : M × X => _)⟩ @[to_additive] instance MulOpposite.continuousSMul : ContinuousSMul M Xᵐᵒᵖ := ⟨MulOpposite.continuous_op.comp <| - continuous_smul.comp <| continuous_id.prod_map MulOpposite.continuous_unop⟩ + continuous_smul.comp <| continuous_id.prodMap MulOpposite.continuous_unop⟩ @[to_additive] protected theorem Specializes.smul {a b : M} {x y : X} (h₁ : a ⤳ b) (h₂ : x ⤳ y) : diff --git a/Mathlib/Topology/Algebra/OpenSubgroup.lean b/Mathlib/Topology/Algebra/OpenSubgroup.lean index 2b3a59d1cf9c5..99c3a4c276f10 100644 --- a/Mathlib/Topology/Algebra/OpenSubgroup.lean +++ b/Mathlib/Topology/Algebra/OpenSubgroup.lean @@ -419,7 +419,7 @@ instance : SubgroupClass (OpenNormalSubgroup G) G where @[to_additive] instance : Coe (OpenNormalSubgroup G) (Subgroup G) where - coe := fun H ↦ H.toOpenSubgroup.toSubgroup + coe H := H.toOpenSubgroup.toSubgroup @[to_additive] instance instPartialOrderOpenNormalSubgroup : PartialOrder (OpenNormalSubgroup G) := inferInstance diff --git a/Mathlib/Topology/Algebra/Order/Field.lean b/Mathlib/Topology/Algebra/Order/Field.lean index 9779411cedf57..b80ccdfcef320 100644 --- a/Mathlib/Topology/Algebra/Order/Field.lean +++ b/Mathlib/Topology/Algebra/Order/Field.lean @@ -195,10 +195,10 @@ instance (priority := 100) LinearOrderedSemifield.toHasContinuousInv₀ {𝕜} · obtain ⟨x', h₀, hxx', h₁⟩ : ∃ x', 0 < x' ∧ x ≤ x' ∧ x' < 1 := ⟨max x (1 / 2), one_half_pos.trans_le (le_max_right _ _), le_max_left _ _, max_lt hx one_half_lt_one⟩ - filter_upwards [Ioo_mem_nhds one_pos (one_lt_inv h₀ h₁)] with y hy - exact hxx'.trans_lt <| inv_inv x' ▸ inv_lt_inv_of_lt hy.1 hy.2 - · filter_upwards [Ioi_mem_nhds (inv_lt_one hx)] with y hy - simpa only [inv_inv] using inv_lt_inv_of_lt (inv_pos.2 <| one_pos.trans hx) hy + filter_upwards [Ioo_mem_nhds one_pos ((one_lt_inv₀ h₀).2 h₁)] with y hy + exact hxx'.trans_lt <| lt_inv_of_lt_inv₀ hy.1 hy.2 + · filter_upwards [Ioi_mem_nhds (inv_lt_one_of_one_lt₀ hx)] with y hy + exact inv_lt_of_inv_lt₀ (by positivity) hy instance (priority := 100) LinearOrderedField.toTopologicalDivisionRing : TopologicalDivisionRing 𝕜 := ⟨⟩ diff --git a/Mathlib/Topology/Algebra/Order/Floor.lean b/Mathlib/Topology/Algebra/Order/Floor.lean index 3877c5331e7bf..bde8ae91b5715 100644 --- a/Mathlib/Topology/Algebra/Order/Floor.lean +++ b/Mathlib/Topology/Algebra/Order/Floor.lean @@ -3,9 +3,8 @@ Copyright (c) 2020 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ -import Mathlib.Algebra.Order.Floor +import Mathlib.Order.Filter.AtTopBot.Floor import Mathlib.Topology.Algebra.Order.Group -import Mathlib.Topology.Order.Basic /-! # Topological facts about `Int.floor`, `Int.ceil` and `Int.fract` @@ -27,8 +26,36 @@ This file proves statements about limits and continuity of functions involving ` open Filter Function Int Set Topology +namespace FloorSemiring + +open scoped Nat + +variable {K : Type*} [LinearOrderedField K] [FloorSemiring K] [TopologicalSpace K] [OrderTopology K] + +theorem tendsto_mul_pow_div_factorial_sub_atTop (a c : K) (d : ℕ) : + Tendsto (fun n ↦ a * c ^ n / (n - d)!) atTop (𝓝 0) := by + rw [tendsto_order] + constructor + all_goals + intro ε hε + filter_upwards [eventually_mul_pow_lt_factorial_sub (a * ε⁻¹) c d] with n h + rw [mul_right_comm, ← div_eq_mul_inv] at h + · rw [div_lt_iff_of_neg hε] at h + rwa [lt_div_iff₀' (Nat.cast_pos.mpr (Nat.factorial_pos _))] + · rw [div_lt_iff₀ hε] at h + rwa [div_lt_iff₀' (Nat.cast_pos.mpr (Nat.factorial_pos _))] + +theorem tendsto_pow_div_factorial_atTop (c : K) : + Tendsto (fun n ↦ c ^ n / n !) atTop (𝓝 0) := by + convert tendsto_mul_pow_div_factorial_sub_atTop 1 c 0 + rw [one_mul] + +end FloorSemiring + variable {α β γ : Type*} [LinearOrderedRing α] [FloorRing α] +-- TODO: move to `Mathlib.Order.Filter.AtTopBot.Floor` + theorem tendsto_floor_atTop : Tendsto (floor : α → ℤ) atTop atTop := floor_mono.tendsto_atTop_atTop fun b => ⟨(b + 1 : ℤ), by rw [floor_intCast]; exact (lt_add_one _).le⟩ @@ -179,7 +206,7 @@ theorem ContinuousOn.comp_fract' {f : β → α → γ} (h : ContinuousOn (uncur (tendsto_id.prod_map (tendsto_fract_right _))).mono_right (le_of_eq ?_) <;> simp [nhdsWithin_prod_eq, nhdsWithin_univ] · replace ht : t ≠ ⌊t⌋ := fun ht' => ht ⟨_, ht'⟩ - refine (h.continuousAt ?_).comp (continuousAt_id.prod_map (continuousAt_fract ht)) + refine (h.continuousAt ?_).comp (continuousAt_id.prodMap (continuousAt_fract ht)) exact prod_mem_nhds univ_mem (Icc_mem_nhds (fract_pos.2 ht) (fract_lt_one _)) theorem ContinuousOn.comp_fract {s : β → α} {f : β → α → γ} diff --git a/Mathlib/Topology/Algebra/PontryaginDual.lean b/Mathlib/Topology/Algebra/PontryaginDual.lean index fe184b35a5fa5..4229049cf7321 100644 --- a/Mathlib/Topology/Algebra/PontryaginDual.lean +++ b/Mathlib/Topology/Algebra/PontryaginDual.lean @@ -22,9 +22,9 @@ isomorphic to its double dual. open Pointwise Function -variable (A B C D E G : Type*) [Monoid A] [Monoid B] [Monoid C] [Monoid D] [CommGroup E] [Group G] - [TopologicalSpace A] [TopologicalSpace B] [TopologicalSpace C] [TopologicalSpace D] - [TopologicalSpace E] [TopologicalSpace G] [TopologicalGroup E] [TopologicalGroup G] +variable (A B C G H : Type*) [Monoid A] [Monoid B] [Monoid C] [CommGroup G] [Group H] + [TopologicalSpace A] [TopologicalSpace B] [TopologicalSpace C] + [TopologicalSpace G] [TopologicalSpace H] [TopologicalGroup G] [TopologicalGroup H] /-- The Pontryagin dual of `A` is the group of continuous homomorphism `A → Circle`. -/ def PontryaginDual := @@ -48,7 +48,7 @@ instance : TopologicalGroup (PontryaginDual A) := noncomputable instance : Inhabited (PontryaginDual A) := (inferInstance : Inhabited (ContinuousMonoidHom A Circle)) -instance [LocallyCompactSpace G] : LocallyCompactSpace (PontryaginDual G) := by +instance [LocallyCompactSpace H] : LocallyCompactSpace (PontryaginDual H) := by let Vn : ℕ → Set Circle := fun n ↦ Circle.exp '' { x | |x| < Real.pi / 2 ^ (n + 1)} have hVn : ∀ n x, x ∈ Vn n ↔ |Complex.arg x| < Real.pi / 2 ^ (n + 1) := by @@ -64,7 +64,7 @@ instance [LocallyCompactSpace G] : LocallyCompactSpace (PontryaginDual G) := by apply Set.Ioo_subset_Ioc_self rw [← two_mul, Set.mem_Ioo, ← abs_lt, abs_mul, abs_two, ← lt_div_iff₀' two_pos] exact h1.trans_le - (div_le_div_of_nonneg_left Real.pi_nonneg two_pos (le_self_pow one_le_two n.succ_ne_zero)) + (div_le_div_of_nonneg_left Real.pi_nonneg two_pos (le_self_pow₀ one_le_two n.succ_ne_zero)) · rw [← Circle.exp_zero, ← isLocalHomeomorph_circleExp.map_nhds_eq 0] refine ((nhds_basis_zero_abs_sub_lt ℝ).to_hasBasis (fun x hx ↦ ⟨Nat.ceil (Real.pi / x), trivial, fun t ht ↦ ?_⟩) @@ -75,17 +75,20 @@ instance [LocallyCompactSpace G] : LocallyCompactSpace (PontryaginDual G) := by refine (Nat.le_ceil (Real.pi / x)).trans ?_ exact_mod_cast (Nat.le_succ _).trans (Nat.lt_two_pow _).le -variable {A B C D E} +variable {A B C G} namespace PontryaginDual open ContinuousMonoidHom instance : FunLike (PontryaginDual A) A Circle := - ContinuousMonoidHom.funLike + ContinuousMonoidHom.instFunLike -noncomputable instance : ContinuousMonoidHomClass (PontryaginDual A) A Circle := - ContinuousMonoidHom.ContinuousMonoidHomClass +noncomputable instance instContinuousMapClass : ContinuousMapClass (PontryaginDual A) A Circle := + ContinuousMonoidHom.instContinuousMapClass + +noncomputable instance instMonoidHomClass : MonoidHomClass (PontryaginDual A) A Circle := + ContinuousMonoidHom.instMonoidHomClass /-- `PontryaginDual` is a contravariant functor. -/ noncomputable def map (f : ContinuousMonoidHom A B) : @@ -107,15 +110,15 @@ theorem map_comp (g : ContinuousMonoidHom B C) (f : ContinuousMonoidHom A B) : ext fun _x => ext fun _y => rfl @[simp] -nonrec theorem map_mul (f g : ContinuousMonoidHom A E) : map (f * g) = map f * map g := +nonrec theorem map_mul (f g : ContinuousMonoidHom A G) : map (f * g) = map f * map g := ext fun x => ext fun y => map_mul x (f y) (g y) -variable (A B C D E) +variable (A B C G) /-- `ContinuousMonoidHom.dual` as a `ContinuousMonoidHom`. -/ -noncomputable def mapHom [LocallyCompactSpace E] : - ContinuousMonoidHom (ContinuousMonoidHom A E) - (ContinuousMonoidHom (PontryaginDual E) (PontryaginDual A)) where +noncomputable def mapHom [LocallyCompactSpace G] : + ContinuousMonoidHom (ContinuousMonoidHom A G) + (ContinuousMonoidHom (PontryaginDual G) (PontryaginDual A)) where toFun := map map_one' := map_one map_mul' := map_mul diff --git a/Mathlib/Topology/Algebra/ProperAction.lean b/Mathlib/Topology/Algebra/ProperAction.lean index b937385d62d95..4a337be7878ae 100644 --- a/Mathlib/Topology/Algebra/ProperAction.lean +++ b/Mathlib/Topology/Algebra/ProperAction.lean @@ -73,9 +73,8 @@ class ProperSMul (G X : Type*) [TopologicalSpace G] [TopologicalSpace X] [Group attribute [to_additive existing] properSMul_iff -variable {G X Y Z : Type*} [Group G] [MulAction G X] [MulAction G Y] -variable [TopologicalSpace G] [TopologicalSpace X] [TopologicalSpace Y] -variable [TopologicalSpace Z] +variable {G X : Type*} [Group G] [MulAction G X] +variable [TopologicalSpace G] [TopologicalSpace X] /-- If a group acts properly then in particular it acts continuously. -/ @[to_additive "If a group acts properly then in particular it acts continuously."] @@ -173,7 +172,7 @@ theorem properSMul_of_closedEmbedding {H : Type*} [Group H] [MulAction H X] [Top (f_compat : ∀ (h : H) (x : X), f h • x = h • x) : ProperSMul H X where isProperMap_smul_pair := by have := isProperMap_of_closedEmbedding f_clemb - have h : IsProperMap (Prod.map f (fun x : X ↦ x)) := IsProperMap.prod_map this isProperMap_id + have h : IsProperMap (Prod.map f (fun x : X ↦ x)) := this.prodMap isProperMap_id have : (fun hx : H × X ↦ (hx.1 • hx.2, hx.2)) = (fun hx ↦ (f hx.1 • hx.2, hx.2)) := by simp [f_compat] rw [this] @@ -245,7 +244,7 @@ theorem properlyDiscontinuousSMul_iff_properSMul [T2Space X] [DiscreteTopology G apply IsCompact.finite_of_discrete -- Now set `h : (g, x) ↦ (g⁻¹ • x, x)`, because `f` is proper by hypothesis, so is `h`. have : IsProperMap (fun gx : G × X ↦ (gx.1⁻¹ • gx.2, gx.2)) := - (IsProperMap.prod_map (Homeomorph.isProperMap (Homeomorph.inv G)) isProperMap_id).comp <| + (IsProperMap.prodMap (Homeomorph.isProperMap (Homeomorph.inv G)) isProperMap_id).comp <| ProperSMul.isProperMap_smul_pair --But we also have that `{g | Set.Nonempty ((g • ·) '' K ∩ L)} = h ⁻¹ (K × L)`, which -- concludes the proof. diff --git a/Mathlib/Topology/Algebra/ProperConstSMul.lean b/Mathlib/Topology/Algebra/ProperConstSMul.lean index 30fbbb4b83bf7..7e3167890f5ee 100644 --- a/Mathlib/Topology/Algebra/ProperConstSMul.lean +++ b/Mathlib/Topology/Algebra/ProperConstSMul.lean @@ -61,7 +61,7 @@ instance {M X Y : Type*} [SMul M X] [TopologicalSpace X] [ProperConstSMul M X] [SMul M Y] [TopologicalSpace Y] [ProperConstSMul M Y] : ProperConstSMul M (X × Y) := - ⟨fun c ↦ (isProperMap_smul c X).prod_map (isProperMap_smul c Y)⟩ + ⟨fun c ↦ (isProperMap_smul c X).prodMap (isProperMap_smul c Y)⟩ instance {M ι : Type*} {X : ι → Type*} [∀ i, SMul M (X i)] [∀ i, TopologicalSpace (X i)] [∀ i, ProperConstSMul M (X i)] : diff --git a/Mathlib/Topology/Algebra/SeparationQuotient.lean b/Mathlib/Topology/Algebra/SeparationQuotient.lean index 070cdf516d506..83775bba6cbe3 100644 --- a/Mathlib/Topology/Algebra/SeparationQuotient.lean +++ b/Mathlib/Topology/Algebra/SeparationQuotient.lean @@ -333,6 +333,12 @@ instance instCommRing [CommRing R] [TopologicalRing R] : surjective_mk.commRing mk mk_zero mk_one mk_add mk_mul mk_neg mk_sub mk_smul mk_smul mk_pow mk_natCast mk_intCast +/-- `SeparationQuotient.mk` as a `RingHom`. -/ +@[simps] +def mkRingHom [NonAssocSemiring R] [TopologicalSemiring R] : R →+* SeparationQuotient R where + toFun := mk + map_one' := mk_one; map_zero' := mk_zero; map_add' := mk_add; map_mul' := mk_mul + end Ring section DistribSMul @@ -375,6 +381,21 @@ def mkCLM : M →L[R] SeparationQuotient M where end Module +section Algebra +variable {R A : Type*} [CommSemiring R] [Semiring A] [Algebra R A] + [TopologicalSpace A] [TopologicalSemiring A] [ContinuousConstSMul R A] + +instance instAlgebra : Algebra R (SeparationQuotient A) where + toRingHom := mkRingHom.comp (algebraMap R A) + commutes' r := Quotient.ind fun a => congrArg _ <| Algebra.commutes r a + smul_def' r := Quotient.ind fun a => congrArg _ <| Algebra.smul_def r a + +@[simp] +theorem mk_algebraMap (r : R) : mk (algebraMap R A r) = algebraMap R (SeparationQuotient A) r := + rfl + +end Algebra + section VectorSpace variable (K E : Type*) [DivisionRing K] [AddCommGroup E] [Module K E] @@ -431,19 +452,22 @@ section VectorSpaceUniform variable (K E : Type*) [DivisionRing K] [AddCommGroup E] [Module K E] [UniformSpace E] [UniformAddGroup E] [ContinuousConstSMul K E] -theorem outCLM_uniformInducing : UniformInducing (outCLM K E) := by - rw [← uniformInducing_mk.uniformInducing_comp_iff, mk_comp_outCLM] - exact uniformInducing_id +theorem outCLM_isUniformInducing : IsUniformInducing (outCLM K E) := by + rw [← isUniformInducing_mk.isUniformInducing_comp_iff, mk_comp_outCLM] + exact .id + +@[deprecated (since := "2024-10-05")] +alias outCLM_uniformInducing := outCLM_isUniformInducing theorem outCLM_isUniformEmbedding : IsUniformEmbedding (outCLM K E) where inj := outCLM_injective K E - toUniformInducing := outCLM_uniformInducing K E + toIsUniformInducing := outCLM_isUniformInducing K E @[deprecated (since := "2024-10-01")] alias outCLM_uniformEmbedding := outCLM_isUniformEmbedding theorem outCLM_uniformContinuous : UniformContinuous (outCLM K E) := - (outCLM_uniformInducing K E).uniformContinuous + (outCLM_isUniformInducing K E).uniformContinuous end VectorSpaceUniform diff --git a/Mathlib/Topology/Algebra/Star.lean b/Mathlib/Topology/Algebra/Star.lean index abbca7eae5b32..88af814e80c6e 100644 --- a/Mathlib/Topology/Algebra/Star.lean +++ b/Mathlib/Topology/Algebra/Star.lean @@ -6,7 +6,7 @@ Authors: Eric Wieser import Mathlib.Algebra.Star.Pi import Mathlib.Algebra.Star.Prod import Mathlib.Topology.Algebra.Constructions -import Mathlib.Topology.ContinuousMap.Basic +import Mathlib.Topology.ContinuousMap.Defs /-! # Continuity of `star` diff --git a/Mathlib/Topology/Algebra/UniformField.lean b/Mathlib/Topology/Algebra/UniformField.lean index 6746d27d74d8b..da3fb454151fa 100644 --- a/Mathlib/Topology/Algebra/UniformField.lean +++ b/Mathlib/Topology/Algebra/UniformField.lean @@ -176,7 +176,7 @@ variable (L : Type*) [Field L] [UniformSpace L] [CompletableTopField L] instance Subfield.completableTopField (K : Subfield L) : CompletableTopField K where nice F F_cau inf_F := by let i : K →+* L := K.subtype - have hi : UniformInducing i := isUniformEmbedding_subtype_val.toUniformInducing + have hi : IsUniformInducing i := isUniformEmbedding_subtype_val.isUniformInducing rw [← hi.cauchy_map_iff] at F_cau ⊢ rw [map_comm (show (i ∘ fun x => x⁻¹) = (fun x => x⁻¹) ∘ i by ext; rfl)] apply CompletableTopField.nice _ F_cau @@ -201,12 +201,12 @@ variable {α β : Type*} [Field β] [b : UniformSpace β] [CompletableTopField /-- The pullback of a completable topological field along a uniform inducing ring homomorphism is a completable topological field. -/ -theorem UniformInducing.completableTopField +theorem IsUniformInducing.completableTopField [UniformSpace α] [T0Space α] - {f : α →+* β} (hf : UniformInducing f) : + {f : α →+* β} (hf : IsUniformInducing f) : CompletableTopField α := by refine CompletableTopField.mk (fun F F_cau inf_F => ?_) - rw [← UniformInducing.cauchy_map_iff hf] at F_cau ⊢ + rw [← IsUniformInducing.cauchy_map_iff hf] at F_cau ⊢ have h_comm : (f ∘ fun x => x⁻¹) = (fun x => x⁻¹) ∘ f := by ext; simp only [Function.comp_apply, map_inv₀, Subfield.coe_inv] rw [Filter.map_comm h_comm] diff --git a/Mathlib/Topology/Algebra/UniformGroup.lean b/Mathlib/Topology/Algebra/UniformGroup.lean index daf8079c33a30..c05bd7340d340 100644 --- a/Mathlib/Topology/Algebra/UniformGroup.lean +++ b/Mathlib/Topology/Algebra/UniformGroup.lean @@ -225,18 +225,21 @@ theorem uniformGroup_inf {u₁ u₂ : UniformSpace β} (h₁ : @UniformGroup β cases b <;> assumption @[to_additive] -lemma UniformInducing.uniformGroup {γ : Type*} [Group γ] [UniformSpace γ] [UniformGroup γ] +lemma IsUniformInducing.uniformGroup {γ : Type*} [Group γ] [UniformSpace γ] [UniformGroup γ] [UniformSpace β] {F : Type*} [FunLike F β γ] [MonoidHomClass F β γ] - (f : F) (hf : UniformInducing f) : + (f : F) (hf : IsUniformInducing f) : UniformGroup β where uniformContinuous_div := by simp_rw [hf.uniformContinuous_iff, Function.comp_def, map_div] - exact uniformContinuous_div.comp (hf.uniformContinuous.prod_map hf.uniformContinuous) + exact uniformContinuous_div.comp (hf.uniformContinuous.prodMap hf.uniformContinuous) + +@[deprecated (since := "2024-10-05")] +alias UniformInducing.uniformGroup := IsUniformInducing.uniformGroup @[to_additive] protected theorem UniformGroup.comap {γ : Type*} [Group γ] {u : UniformSpace γ} [UniformGroup γ] {F : Type*} [FunLike F β γ] [MonoidHomClass F β γ] (f : F) : @UniformGroup β (u.comap f) _ := - letI : UniformSpace β := u.comap f; UniformInducing.uniformGroup f ⟨rfl⟩ + letI : UniformSpace β := u.comap f; IsUniformInducing.uniformGroup f ⟨rfl⟩ end LatticeOps @@ -735,14 +738,14 @@ variable [T0Space G] [CompleteSpace G] /-- Bourbaki GT III.6.5 Theorem I: ℤ-bilinear continuous maps from dense images into a complete Hausdorff group extend by continuity. Note: Bourbaki assumes that α and β are also complete Hausdorff, but this is not necessary. -/ -theorem extend_Z_bilin : Continuous (extend (de.prod df) (fun p : β × δ => φ p.1 p.2)) := by +theorem extend_Z_bilin : Continuous (extend (de.prodMap df) (fun p : β × δ => φ p.1 p.2)) := by refine continuous_extend_of_cauchy _ ?_ rintro ⟨x₀, y₀⟩ constructor · apply NeBot.map apply comap_neBot intro U h - rcases mem_closure_iff_nhds.1 ((de.prod df).dense (x₀, y₀)) U h with ⟨x, x_in, ⟨z, z_x⟩⟩ + rcases mem_closure_iff_nhds.1 ((de.prodMap df).dense (x₀, y₀)) U h with ⟨x, x_in, ⟨z, z_x⟩⟩ exists z aesop · suffices map (fun p : (β × δ) × β × δ => (fun p : β × δ => φ p.1 p.2) p.2 - diff --git a/Mathlib/Topology/Algebra/UniformMulAction.lean b/Mathlib/Topology/Algebra/UniformMulAction.lean index d98b4c082accb..3f0de8b6107d6 100644 --- a/Mathlib/Topology/Algebra/UniformMulAction.lean +++ b/Mathlib/Topology/Algebra/UniformMulAction.lean @@ -88,13 +88,16 @@ theorem UniformContinuous.const_smul [UniformContinuousConstSMul M X] {f : Y → (uniformContinuous_const_smul c).comp hf @[to_additive] -lemma UniformInducing.uniformContinuousConstSMul [SMul M Y] [UniformContinuousConstSMul M Y] - {f : X → Y} (hf : UniformInducing f) (hsmul : ∀ (c : M) x, f (c • x) = c • f x) : +lemma IsUniformInducing.uniformContinuousConstSMul [SMul M Y] [UniformContinuousConstSMul M Y] + {f : X → Y} (hf : IsUniformInducing f) (hsmul : ∀ (c : M) x, f (c • x) = c • f x) : UniformContinuousConstSMul M X where uniformContinuous_const_smul c := by simpa only [hf.uniformContinuous_iff, Function.comp_def, hsmul] using hf.uniformContinuous.const_smul c +@[deprecated (since := "2024-10-05")] +alias UniformInducing.uniformContinuousConstSMul := IsUniformInducing.uniformContinuousConstSMul + /-- If a scalar action is central, then its right action is uniform continuous when its left action is. -/ @[to_additive "If an additive action is central, then its right action is uniform diff --git a/Mathlib/Topology/Algebra/UniformRing.lean b/Mathlib/Topology/Algebra/UniformRing.lean index f4799994dceaf..f315780ad96d3 100644 --- a/Mathlib/Topology/Algebra/UniformRing.lean +++ b/Mathlib/Topology/Algebra/UniformRing.lean @@ -47,7 +47,7 @@ instance one : One (Completion α) := ⟨(1 : α)⟩ instance mul : Mul (Completion α) := - ⟨curry <| (isDenseInducing_coe.prod isDenseInducing_coe).extend ((↑) ∘ uncurry (· * ·))⟩ + ⟨curry <| (isDenseInducing_coe.prodMap isDenseInducing_coe).extend ((↑) ∘ uncurry (· * ·))⟩ @[norm_cast] theorem coe_one : ((1 : α) : Completion α) = 1 := @@ -59,7 +59,7 @@ variable {α : Type*} [Ring α] [UniformSpace α] [TopologicalRing α] @[norm_cast] theorem coe_mul (a b : α) : ((a * b : α) : Completion α) = a * b := - ((isDenseInducing_coe.prod isDenseInducing_coe).extend_eq + ((isDenseInducing_coe.prodMap isDenseInducing_coe).extend_eq ((continuous_coe α).comp (@continuous_mul α _ _ _)) (a, b)).symm variable [UniformAddGroup α] @@ -276,8 +276,8 @@ variable {γ : Type*} [UniformSpace γ] [Semiring γ] [TopologicalSemiring γ] variable [T2Space γ] [CompleteSpace γ] /-- The dense inducing extension as a ring homomorphism. -/ -noncomputable def IsDenseInducing.extendRingHom {i : α →+* β} {f : α →+* γ} (ue : UniformInducing i) - (dr : DenseRange i) (hf : UniformContinuous f) : β →+* γ where +noncomputable def IsDenseInducing.extendRingHom {i : α →+* β} {f : α →+* γ} + (ue : IsUniformInducing i) (dr : DenseRange i) (hf : UniformContinuous f) : β →+* γ where toFun := (ue.isDenseInducing dr).extend f map_one' := by convert IsDenseInducing.extend_eq (ue.isDenseInducing dr) hf.continuous 1 diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index 388fd5adcf054..55a8aa4aec4ec 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -799,9 +799,11 @@ theorem frequently_frequently_nhds {p : X → Prop} : simp only [not_frequently, eventually_eventually_nhds] @[simp] -theorem eventually_mem_nhds : (∀ᶠ x' in 𝓝 x, s ∈ 𝓝 x') ↔ s ∈ 𝓝 x := +theorem eventually_mem_nhds_iff : (∀ᶠ x' in 𝓝 x, s ∈ 𝓝 x') ↔ s ∈ 𝓝 x := eventually_eventually_nhds +@[deprecated (since := "2024-10-04")] alias eventually_mem_nhds := eventually_mem_nhds_iff + @[simp] theorem nhds_bind_nhds : (𝓝 x).bind 𝓝 = 𝓝 x := Filter.ext fun _ => eventually_eventually_nhds @@ -851,13 +853,13 @@ theorem tendsto_atTop_nhds [Nonempty α] [SemilatticeSup α] {f : α → X} : theorem tendsto_const_nhds {f : Filter α} : Tendsto (fun _ : α => x) f (𝓝 x) := tendsto_nhds.mpr fun _ _ ha => univ_mem' fun _ => ha -theorem tendsto_atTop_of_eventually_const {ι : Type*} [SemilatticeSup ι] [Nonempty ι] +theorem tendsto_atTop_of_eventually_const {ι : Type*} [Preorder ι] {u : ι → X} {i₀ : ι} (h : ∀ i ≥ i₀, u i = x) : Tendsto u atTop (𝓝 x) := - Tendsto.congr' (EventuallyEq.symm (eventually_atTop.mpr ⟨i₀, h⟩)) tendsto_const_nhds + Tendsto.congr' (EventuallyEq.symm ((eventually_ge_atTop i₀).mono h)) tendsto_const_nhds -theorem tendsto_atBot_of_eventually_const {ι : Type*} [SemilatticeInf ι] [Nonempty ι] +theorem tendsto_atBot_of_eventually_const {ι : Type*} [Preorder ι] {u : ι → X} {i₀ : ι} (h : ∀ i ≤ i₀, u i = x) : Tendsto u atBot (𝓝 x) := - Tendsto.congr' (EventuallyEq.symm (eventually_atBot.mpr ⟨i₀, h⟩)) tendsto_const_nhds + tendsto_atTop_of_eventually_const (ι := ιᵒᵈ) h theorem pure_le_nhds : pure ≤ (𝓝 : X → Filter X) := fun _ _ hs => mem_pure.2 <| mem_of_mem_nhds hs @@ -1400,7 +1402,7 @@ theorem ContinuousAt.eventually_mem {f : X → Y} {x : X} (hf : ContinuousAt f x (hs : s ∈ 𝓝 (f x)) : ∀ᶠ y in 𝓝 x, f y ∈ s := hf hs -/-- If a function ``f` tends to somewhere other than `𝓝 (f x)` at `x`, +/-- If a function `f` tends to somewhere other than `𝓝 (f x)` at `x`, then `f` is not continuous at `x` -/ lemma not_continuousAt_of_tendsto {f : X → Y} {l₁ : Filter X} {l₂ : Filter Y} {x : X} diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index d9f4b3c3fea60..3807d5656a53a 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -1354,13 +1354,13 @@ theorem CC_exact {f : LocallyConstant C ℤ} (hf : Linear_CC' C hsC ho f = 0) : exact C1_projOrd C hsC ho hx₁ variable (o) in -theorem succ_mono : CategoryTheory.Mono (ModuleCat.ofHom (πs C o)) := by +theorem succ_mono : CategoryTheory.Mono (ModuleCat.asHom (πs C o)) := by rw [ModuleCat.mono_iff_injective] exact injective_πs _ _ include hC in theorem succ_exact : - (ShortComplex.mk (ModuleCat.ofHom (πs C o)) (ModuleCat.ofHom (Linear_CC' C hsC ho)) + (ShortComplex.mk (ModuleCat.asHom (πs C o)) (ModuleCat.asHom (Linear_CC' C hsC ho)) (by ext; apply CC_comp_zero)).Exact := by rw [ShortComplex.moduleCat_exact_iff] intro f @@ -1478,7 +1478,7 @@ theorem span_sum : Set.range (eval C) = Set.range (Sum.elim theorem square_commutes : SumEval C ho ∘ Sum.inl = - ModuleCat.ofHom (πs C o) ∘ eval (π C (ord I · < o)) := by + ModuleCat.asHom (πs C o) ∘ eval (π C (ord I · < o)) := by ext l dsimp [SumEval] rw [← Products.eval_πs C (Products.prop_of_isGood _ _ l.prop)] @@ -1644,7 +1644,7 @@ theorem maxTail_isGood (l : MaxProducts C ho) rfl have hse := succ_exact C hC hsC ho rw [ShortComplex.moduleCat_exact_iff_range_eq_ker] at hse - dsimp [ModuleCat.ofHom] at hse + dsimp [ModuleCat.asHom] at hse -- Rewrite `this` using exact sequence manipulations to conclude that a term is in the range of -- the linear map `πs`: @@ -1701,8 +1701,8 @@ include hC in theorem linearIndependent_comp_of_eval (h₁ : ⊤ ≤ Submodule.span ℤ (Set.range (eval (π C (ord I · < o))))) : LinearIndependent ℤ (eval (C' C ho)) → - LinearIndependent ℤ (ModuleCat.ofHom (Linear_CC' C hsC ho) ∘ SumEval C ho ∘ Sum.inr) := by - dsimp [SumEval, ModuleCat.ofHom] + LinearIndependent ℤ (ModuleCat.asHom (Linear_CC' C hsC ho) ∘ SumEval C ho ∘ Sum.inr) := by + dsimp [SumEval, ModuleCat.asHom] erw [max_eq_eval_unapply C hsC ho] intro h let f := MaxToGood C hC hsC ho h₁ diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 9addb98e719ef..ab44aa713d2eb 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -56,7 +56,7 @@ instance topologicalSpaceUnbundled (X : TopCat) : TopologicalSpace X := instance instFunLike (X Y : TopCat) : FunLike (X ⟶ Y) X Y := inferInstanceAs <| FunLike C(X, Y) X Y -instance instMonoidHomClass (X Y : TopCat) : ContinuousMapClass (X ⟶ Y) X Y := +instance instContinuousMapClass (X Y : TopCat) : ContinuousMapClass (X ⟶ Y) X Y := inferInstanceAs <| ContinuousMapClass C(X, Y) X Y -- Porting note (#10618): simp can prove this; removed simp diff --git a/Mathlib/Topology/Clopen.lean b/Mathlib/Topology/Clopen.lean index c560e10009cd9..ff10a04698b33 100644 --- a/Mathlib/Topology/Clopen.lean +++ b/Mathlib/Topology/Clopen.lean @@ -102,6 +102,10 @@ theorem isClopen_inter_of_disjoint_cover_clopen {s a b : Set X} (h : IsClopen s) rintro x ⟨hx₁, hx₂⟩ exact ⟨hx₁, by simpa [not_mem_of_mem_compl hx₂] using cover hx₁⟩ +theorem isClopen_of_disjoint_cover_open {a b : Set X} (cover : univ ⊆ a ∪ b) + (ha : IsOpen a) (hb : IsOpen b) (hab : Disjoint a b) : IsClopen a := + univ_inter a ▸ isClopen_inter_of_disjoint_cover_clopen isClopen_univ cover ha hb hab + @[simp] theorem isClopen_discrete [DiscreteTopology X] (s : Set X) : IsClopen s := ⟨isClosed_discrete _, isOpen_discrete _⟩ diff --git a/Mathlib/Topology/CompactOpen.lean b/Mathlib/Topology/CompactOpen.lean index 6462eef116241..0338dd8e2d361 100644 --- a/Mathlib/Topology/CompactOpen.lean +++ b/Mathlib/Topology/CompactOpen.lean @@ -374,7 +374,7 @@ theorem continuous_curry [LocallyCompactSpace (X × Y)] : /-- The uncurried form of a continuous map `X → C(Y, Z)` is a continuous map `X × Y → Z`. -/ theorem continuous_uncurry_of_continuous [LocallyCompactSpace Y] (f : C(X, C(Y, Z))) : Continuous (Function.uncurry fun x y => f x y) := - continuous_eval.comp <| f.continuous.prod_map continuous_id + continuous_eval.comp <| f.continuous.prodMap continuous_id /-- The uncurried form of a continuous map `X → C(Y, Z)` as a continuous map `X × Y → Z` (if `Y` is locally compact). If `X` is also locally compact, then this is a homeomorphism between the two @@ -388,7 +388,7 @@ theorem continuous_uncurry [LocallyCompactSpace X] [LocallyCompactSpace Y] : Continuous (uncurry : C(X, C(Y, Z)) → C(X × Y, Z)) := by apply continuous_of_continuous_uncurry rw [← (Homeomorph.prodAssoc _ _ _).comp_continuous_iff'] - apply continuous_eval.comp (continuous_eval.prod_map continuous_id) + apply continuous_eval.comp (continuous_eval.prodMap continuous_id) /-- The family of constant maps: `Y → C(X, Y)` as a continuous map. -/ def const' : C(Y, C(X, Y)) := diff --git a/Mathlib/Topology/Compactification/OnePoint.lean b/Mathlib/Topology/Compactification/OnePoint.lean index ffe89b1a9a2a8..7552539c3ada1 100644 --- a/Mathlib/Topology/Compactification/OnePoint.lean +++ b/Mathlib/Topology/Compactification/OnePoint.lean @@ -187,7 +187,7 @@ instance : TopologicalSpace (OnePoint X) where rw [preimage_sUnion] exact isOpen_biUnion fun s hs => (ho s hs).2 -variable {s : Set (OnePoint X)} {t : Set X} +variable {s : Set (OnePoint X)} theorem isOpen_def : IsOpen s ↔ (∞ ∈ s → IsCompact ((↑) ⁻¹' s : Set X)ᶜ) ∧ IsOpen ((↑) ⁻¹' s : Set X) := diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 1cad1a9499c1e..644fcb694b3a8 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -75,8 +75,8 @@ structure Path (x y : X) extends C(I, X) where target' : toFun 1 = y instance Path.funLike : FunLike (Path x y) I X where - coe := fun γ ↦ ⇑γ.toContinuousMap - coe_injective' := fun γ₁ γ₂ h => by + coe γ := ⇑γ.toContinuousMap + coe_injective' γ₁ γ₂ h := by simp only [DFunLike.coe_fn_eq] at h cases γ₁; cases γ₂; congr @@ -200,7 +200,7 @@ instance topologicalSpace : TopologicalSpace (Path x y) := theorem continuous_eval : Continuous fun p : Path x y × I => p.1 p.2 := ContinuousMap.continuous_eval.comp <| - (continuous_induced_dom (α := Path x y)).prod_map continuous_id + (continuous_induced_dom (α := Path x y)).prodMap continuous_id @[continuity] theorem _root_.Continuous.path_eval {Y} [TopologicalSpace Y] {f : Y → Path x y} {g : Y → I} @@ -423,7 +423,7 @@ theorem cast_coe (γ : Path x y) {x' y'} (hx : x' = x) (hy : y' = y) : (γ.cast theorem symm_continuous_family {ι : Type*} [TopologicalSpace ι] {a b : ι → X} (γ : ∀ t : ι, Path (a t) (b t)) (h : Continuous ↿γ) : Continuous ↿fun t => (γ t).symm := - h.comp (continuous_id.prod_map continuous_symm) + h.comp (continuous_id.prodMap continuous_symm) @[continuity] theorem continuous_symm : Continuous (symm : Path x y → Path y x) := @@ -433,7 +433,7 @@ theorem continuous_symm : Continuous (symm : Path x y → Path y x) := theorem continuous_uncurry_extend_of_continuous_family {ι : Type*} [TopologicalSpace ι] {a b : ι → X} (γ : ∀ t : ι, Path (a t) (b t)) (h : Continuous ↿γ) : Continuous ↿fun t => (γ t).extend := by - apply h.comp (continuous_id.prod_map continuous_projIcc) + apply h.comp (continuous_id.prodMap continuous_projIcc) exact zero_le_one @[continuity] @@ -447,12 +447,12 @@ theorem trans_continuous_family {ι : Type*} [TopologicalSpace ι] refine Continuous.if_le ?_ ?_ (continuous_subtype_val.comp continuous_snd) continuous_const ?_ · change Continuous ((fun p : ι × ℝ => (γ₁ p.1).extend p.2) ∘ Prod.map id (fun x => 2 * x : I → ℝ)) - exact h₁'.comp (continuous_id.prod_map <| continuous_const.mul continuous_subtype_val) + exact h₁'.comp (continuous_id.prodMap <| continuous_const.mul continuous_subtype_val) · change Continuous ((fun p : ι × ℝ => (γ₂ p.1).extend p.2) ∘ Prod.map id (fun x => 2 * x - 1 : I → ℝ)) exact h₂'.comp - (continuous_id.prod_map <| + (continuous_id.prodMap <| (continuous_const.mul continuous_subtype_val).sub continuous_const) · rintro st hst simp [hst, mul_inv_cancel₀ (two_ne_zero' ℝ)] diff --git a/Mathlib/Topology/Connected/TotallyDisconnected.lean b/Mathlib/Topology/Connected/TotallyDisconnected.lean index 0da1004c037f4..b7f1ff02b4fb4 100644 --- a/Mathlib/Topology/Connected/TotallyDisconnected.lean +++ b/Mathlib/Topology/Connected/TotallyDisconnected.lean @@ -195,7 +195,7 @@ alias IsTotallySeparated.isTotallyDisconnected := isTotallyDisconnected_of_isTot /-- A space is totally separated if any two points can be separated by two disjoint open sets covering the whole space. -/ -class TotallySeparatedSpace (α : Type u) [TopologicalSpace α] : Prop where +@[mk_iff] class TotallySeparatedSpace (α : Type u) [TopologicalSpace α] : Prop where /-- The universal set `Set.univ` in a totally separated space is totally separated. -/ isTotallySeparated_univ : IsTotallySeparated (univ : Set α) @@ -210,15 +210,19 @@ instance (priority := 100) TotallySeparatedSpace.of_discrete (α : Type*) [Topol ⟨fun _ _ b _ h => ⟨{b}ᶜ, {b}, isOpen_discrete _, isOpen_discrete _, h, rfl, (compl_union_self _).symm.subset, disjoint_compl_left⟩⟩ +theorem totallySeparatedSpace_iff_exists_isClopen {α : Type*} [TopologicalSpace α] : + TotallySeparatedSpace α ↔ ∀ x y : α, x ≠ y → ∃ U : Set α, IsClopen U ∧ x ∈ U ∧ y ∈ Uᶜ := by + simp only [totallySeparatedSpace_iff, IsTotallySeparated, Set.Pairwise, mem_univ, true_implies] + refine forall₃_congr fun x y _ ↦ + ⟨fun ⟨U, V, hU, hV, Ux, Vy, f, disj⟩ ↦ ?_, fun ⟨U, hU, Ux, Ucy⟩ ↦ ?_⟩ + · exact ⟨U, isClopen_of_disjoint_cover_open f hU hV disj, + Ux, fun Uy ↦ Set.disjoint_iff.mp disj ⟨Uy, Vy⟩⟩ + · exact ⟨U, Uᶜ, hU.2, hU.compl.2, Ux, Ucy, (Set.union_compl_self U).ge, disjoint_compl_right⟩ + theorem exists_isClopen_of_totally_separated {α : Type*} [TopologicalSpace α] [TotallySeparatedSpace α] {x y : α} (hxy : x ≠ y) : - ∃ U : Set α, IsClopen U ∧ x ∈ U ∧ y ∈ Uᶜ := by - obtain ⟨U, V, hU, hV, Ux, Vy, f, disj⟩ := - TotallySeparatedSpace.isTotallySeparated_univ (Set.mem_univ x) (Set.mem_univ y) hxy - have hU := isClopen_inter_of_disjoint_cover_clopen isClopen_univ f hU hV disj - rw [univ_inter _] at hU - rw [← Set.subset_compl_iff_disjoint_right, subset_compl_comm] at disj - exact ⟨U, hU, Ux, disj Vy⟩ + ∃ U : Set α, IsClopen U ∧ x ∈ U ∧ y ∈ Uᶜ := + totallySeparatedSpace_iff_exists_isClopen.mp ‹_› _ _ hxy end TotallySeparated @@ -260,7 +264,6 @@ theorem Continuous.connectedComponentsLift_unique (h : Continuous f) (g : Connec (hg : g ∘ (↑) = f) : g = h.connectedComponentsLift := connectedComponents_lift_unique' <| hg.trans h.connectedComponentsLift_comp_coe.symm - instance ConnectedComponents.totallyDisconnectedSpace : TotallyDisconnectedSpace (ConnectedComponents α) := by rw [totallyDisconnectedSpace_iff_connectedComponent_singleton] diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index f23b56ce022b3..655e49f24bd70 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -400,10 +400,12 @@ theorem Continuous.comp₄ {g : X × Y × Z × ζ → ε} (hg : Continuous g) {e hg.comp₃ he hf <| hk.prod_mk hl @[continuity] -theorem Continuous.prod_map {f : Z → X} {g : W → Y} (hf : Continuous f) (hg : Continuous g) : - Continuous fun p : Z × W => (f p.1, g p.2) := +theorem Continuous.prodMap {f : Z → X} {g : W → Y} (hf : Continuous f) (hg : Continuous g) : + Continuous (Prod.map f g) := hf.fst'.prod_mk hg.snd' +@[deprecated (since := "2024-10-05")] alias Continuous.prod_map := Continuous.prodMap + /-- A version of `continuous_inf_dom_left` for binary functions -/ theorem continuous_inf_dom_left₂ {X Y Z} {f : X → Y → Z} {ta1 ta2 : TopologicalSpace X} {tb1 tb2 : TopologicalSpace Y} {tc1 : TopologicalSpace Z} @@ -411,7 +413,7 @@ theorem continuous_inf_dom_left₂ {X Y Z} {f : X → Y → Z} {ta1 ta2 : Topolo haveI := ta1 ⊓ ta2; haveI := tb1 ⊓ tb2; exact Continuous fun p : X × Y => f p.1 p.2 := by have ha := @continuous_inf_dom_left _ _ id ta1 ta2 ta1 (@continuous_id _ (id _)) have hb := @continuous_inf_dom_left _ _ id tb1 tb2 tb1 (@continuous_id _ (id _)) - have h_continuous_id := @Continuous.prod_map _ _ _ _ ta1 tb1 (ta1 ⊓ ta2) (tb1 ⊓ tb2) _ _ ha hb + have h_continuous_id := @Continuous.prodMap _ _ _ _ ta1 tb1 (ta1 ⊓ ta2) (tb1 ⊓ tb2) _ _ ha hb exact @Continuous.comp _ _ _ (id _) (id _) _ _ _ h h_continuous_id /-- A version of `continuous_inf_dom_right` for binary functions -/ @@ -421,7 +423,7 @@ theorem continuous_inf_dom_right₂ {X Y Z} {f : X → Y → Z} {ta1 ta2 : Topol haveI := ta1 ⊓ ta2; haveI := tb1 ⊓ tb2; exact Continuous fun p : X × Y => f p.1 p.2 := by have ha := @continuous_inf_dom_right _ _ id ta1 ta2 ta2 (@continuous_id _ (id _)) have hb := @continuous_inf_dom_right _ _ id tb1 tb2 tb2 (@continuous_id _ (id _)) - have h_continuous_id := @Continuous.prod_map _ _ _ _ ta2 tb2 (ta1 ⊓ ta2) (tb1 ⊓ tb2) _ _ ha hb + have h_continuous_id := @Continuous.prodMap _ _ _ _ ta2 tb2 (ta1 ⊓ ta2) (tb1 ⊓ tb2) _ _ ha hb exact @Continuous.comp _ _ _ (id _) (id _) _ _ _ h h_continuous_id /-- A version of `continuous_sInf_dom` for binary functions -/ @@ -433,7 +435,7 @@ theorem continuous_sInf_dom₂ {X Y Z} {f : X → Y → Z} {tas : Set (Topologic exact @Continuous _ _ _ tc fun p : X × Y => f p.1 p.2 := by have hX := continuous_sInf_dom hX continuous_id have hY := continuous_sInf_dom hY continuous_id - have h_continuous_id := @Continuous.prod_map _ _ _ _ tX tY (sInf tas) (sInf tbs) _ _ hX hY + have h_continuous_id := @Continuous.prodMap _ _ _ _ tX tY (sInf tas) (sInf tbs) _ _ hX hY exact @Continuous.comp _ _ _ (id _) (id _) _ _ _ hf h_continuous_id theorem Filter.Eventually.prod_inl_nhds {p : X → Prop} {x : X} (h : ∀ᶠ x in 𝓝 x, p x) (y : Y) : @@ -582,13 +584,19 @@ theorem ContinuousAt.prod {f : X → Y} {g : X → Z} {x : X} (hf : ContinuousAt (hg : ContinuousAt g x) : ContinuousAt (fun x => (f x, g x)) x := hf.prod_mk_nhds hg -theorem ContinuousAt.prod_map {f : X → Z} {g : Y → W} {p : X × Y} (hf : ContinuousAt f p.fst) - (hg : ContinuousAt g p.snd) : ContinuousAt (fun p : X × Y => (f p.1, g p.2)) p := +theorem ContinuousAt.prodMap {f : X → Z} {g : Y → W} {p : X × Y} (hf : ContinuousAt f p.fst) + (hg : ContinuousAt g p.snd) : ContinuousAt (Prod.map f g) p := hf.fst''.prod hg.snd'' -theorem ContinuousAt.prod_map' {f : X → Z} {g : Y → W} {x : X} {y : Y} (hf : ContinuousAt f x) - (hg : ContinuousAt g y) : ContinuousAt (fun p : X × Y => (f p.1, g p.2)) (x, y) := - hf.fst'.prod hg.snd' +@[deprecated (since := "2024-10-05")] alias ContinuousAt.prod_map := ContinuousAt.prodMap + +/-- A version of `ContinuousAt.prodMap` that avoids `Prod.fst`/`Prod.snd` +by assuming that the point is `(x, y)`. -/ +theorem ContinuousAt.prodMap' {f : X → Z} {g : Y → W} {x : X} {y : Y} (hf : ContinuousAt f x) + (hg : ContinuousAt g y) : ContinuousAt (Prod.map f g) (x, y) := + hf.prodMap hg + +@[deprecated (since := "2024-10-05")] alias ContinuousAt.prod_map' := ContinuousAt.prodMap' theorem ContinuousAt.comp₂ {f : Y × Z → W} {g : X → Y} {h : X → Z} {x : X} (hf : ContinuousAt f (g x, h x)) (hg : ContinuousAt g x) (hh : ContinuousAt h x) : @@ -776,15 +784,19 @@ theorem Dense.prod {s : Set X} {t : Set Y} (hs : Dense s) (ht : Dense t) : Dense exact ⟨hs x.1, ht x.2⟩ /-- If `f` and `g` are maps with dense range, then `Prod.map f g` has dense range. -/ -theorem DenseRange.prod_map {ι : Type*} {κ : Type*} {f : ι → Y} {g : κ → Z} (hf : DenseRange f) +theorem DenseRange.prodMap {ι : Type*} {κ : Type*} {f : ι → Y} {g : κ → Z} (hf : DenseRange f) (hg : DenseRange g) : DenseRange (Prod.map f g) := by simpa only [DenseRange, prod_range_range_eq] using hf.prod hg -theorem Inducing.prod_map {f : X → Y} {g : Z → W} (hf : Inducing f) (hg : Inducing g) : +@[deprecated (since := "2024-10-05")] alias DenseRange.prod_map := DenseRange.prodMap + +theorem Inducing.prodMap {f : X → Y} {g : Z → W} (hf : Inducing f) (hg : Inducing g) : Inducing (Prod.map f g) := inducing_iff_nhds.2 fun (x, z) => by simp_rw [Prod.map_def, nhds_prod_eq, hf.nhds_eq_comap, hg.nhds_eq_comap, prod_comap_comap_eq] +@[deprecated (since := "2024-10-05")] alias Inducing.prod_map := Inducing.prodMap + @[simp] theorem inducing_const_prod {x : X} {f : Y → Z} : (Inducing fun x' => (x, f x')) ↔ Inducing f := by simp_rw [inducing_iff, instTopologicalSpaceProd, induced_inf, induced_compose, Function.comp_def, @@ -795,21 +807,27 @@ theorem inducing_prod_const {y : Y} {f : X → Z} : (Inducing fun x => (f x, y)) simp_rw [inducing_iff, instTopologicalSpaceProd, induced_inf, induced_compose, Function.comp_def, induced_const, inf_top_eq] -theorem Embedding.prod_map {f : X → Y} {g : Z → W} (hf : Embedding f) (hg : Embedding g) : +theorem Embedding.prodMap {f : X → Y} {g : Z → W} (hf : Embedding f) (hg : Embedding g) : Embedding (Prod.map f g) := - { hf.toInducing.prod_map hg.toInducing with + { hf.toInducing.prodMap hg.toInducing with inj := fun ⟨x₁, z₁⟩ ⟨x₂, z₂⟩ => by simp [hf.inj.eq_iff, hg.inj.eq_iff] } -protected theorem IsOpenMap.prod {f : X → Y} {g : Z → W} (hf : IsOpenMap f) (hg : IsOpenMap g) : - IsOpenMap fun p : X × Z => (f p.1, g p.2) := by +@[deprecated (since := "2024-10-05")] alias Embedding.prod_map := Embedding.prodMap + +protected theorem IsOpenMap.prodMap {f : X → Y} {g : Z → W} (hf : IsOpenMap f) (hg : IsOpenMap g) : + IsOpenMap (Prod.map f g) := by rw [isOpenMap_iff_nhds_le] rintro ⟨a, b⟩ - rw [nhds_prod_eq, nhds_prod_eq, ← Filter.prod_map_map_eq] + rw [nhds_prod_eq, nhds_prod_eq, ← Filter.prod_map_map_eq'] exact Filter.prod_mono (hf.nhds_le a) (hg.nhds_le b) -protected theorem OpenEmbedding.prod {f : X → Y} {g : Z → W} (hf : OpenEmbedding f) - (hg : OpenEmbedding g) : OpenEmbedding fun x : X × Z => (f x.1, g x.2) := - openEmbedding_of_embedding_open (hf.1.prod_map hg.1) (hf.isOpenMap.prod hg.isOpenMap) +@[deprecated (since := "2024-10-05")] alias IsOpenMap.prod := IsOpenMap.prodMap + +protected theorem OpenEmbedding.prodMap {f : X → Y} {g : Z → W} (hf : OpenEmbedding f) + (hg : OpenEmbedding g) : OpenEmbedding (Prod.map f g) := + openEmbedding_of_embedding_open (hf.1.prodMap hg.1) (hf.isOpenMap.prodMap hg.isOpenMap) + +@[deprecated (since := "2024-10-05")] alias OpenEmbedding.prod := OpenEmbedding.prodMap theorem embedding_graph {f : X → Y} (hf : Continuous f) : Embedding fun x => (x, f x) := embedding_of_embedding_compose (continuous_id.prod_mk hf) continuous_fst embedding_id @@ -819,7 +837,7 @@ theorem embedding_prod_mk (x : X) : Embedding (Prod.mk x : Y → X × Y) := theorem IsOpenQuotientMap.prodMap {f : X → Y} {g : Z → W} (hf : IsOpenQuotientMap f) (hg : IsOpenQuotientMap g) : IsOpenQuotientMap (Prod.map f g) := - ⟨.prodMap hf.1 hg.1, .prod_map hf.2 hg.2, .prod hf.3 hg.3⟩ + ⟨.prodMap hf.1 hg.1, .prodMap hf.2 hg.2, .prodMap hf.3 hg.3⟩ end Prod @@ -967,7 +985,7 @@ end Sum section Subtype -variable [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {p : X → Prop} +variable [TopologicalSpace X] [TopologicalSpace Y] {p : X → Prop} theorem inducing_subtype_val {t : Set Y} : Inducing ((↑) : t → Y) := ⟨rfl⟩ @@ -1122,7 +1140,7 @@ end Subtype section Quotient -variable [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] +variable [TopologicalSpace X] [TopologicalSpace Y] variable {r : X → X → Prop} {s : Setoid X} theorem quotientMap_quot_mk : QuotientMap (@Quot.mk X r) := @@ -1248,7 +1266,7 @@ theorem Finset.continuous_restrict₂ {s t : Finset ι} (hst : s ⊆ t) : Continuous (Finset.restrict₂ (π := π) hst) := continuous_pi fun _ ↦ continuous_apply _ -variable {Z : Type*} [TopologicalSpace Z] +variable [TopologicalSpace Z] @[continuity, fun_prop] theorem Pi.continuous_restrict_apply (s : Set X) {f : X → Z} (hf : Continuous f) : @@ -1642,8 +1660,8 @@ theorem IsClosed.trans (ht : IsClosed t) (hs : IsClosed s) : IsClosed (t : Set X end Monad section NhdsSet -variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {f : Filter X} - {s : Set X} {t : Set Y} {x : X} +variable [TopologicalSpace X] [TopologicalSpace Y] + {s : Set X} {t : Set Y} /-- The product of a neighborhood of `s` and a neighborhood of `t` is a neighborhood of `s ×ˢ t`, formulated in terms of a filter inequality. -/ diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index 46166bf24c1ea..4b056cc3d11f1 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -295,9 +295,9 @@ instance [LocallyCompactSpace α] [Mul β] [ContinuousMul β] : ContinuousMul C( ⟨by refine continuous_of_continuous_uncurry _ ?_ have h1 : Continuous fun x : (C(α, β) × C(α, β)) × α => x.fst.fst x.snd := - continuous_eval.comp (continuous_fst.prod_map continuous_id) + continuous_eval.comp (continuous_fst.prodMap continuous_id) have h2 : Continuous fun x : (C(α, β) × C(α, β)) × α => x.fst.snd x.snd := - continuous_eval.comp (continuous_snd.prod_map continuous_id) + continuous_eval.comp (continuous_snd.prodMap continuous_id) exact h1.mul h2⟩ /-- Coercion to a function as a `MonoidHom`. Similar to `MonoidHom.coeFn`. -/ @@ -550,7 +550,7 @@ instance [LocallyCompactSpace α] [TopologicalSpace R] [SMul R M] [ContinuousSMu ⟨by refine continuous_of_continuous_uncurry _ ?_ have h : Continuous fun x : (R × C(α, M)) × α => x.fst.snd x.snd := - continuous_eval.comp (continuous_snd.prod_map continuous_id) + continuous_eval.comp (continuous_snd.prodMap continuous_id) exact (continuous_fst.comp continuous_fst).smul h⟩ @[to_additive (attr := simp, norm_cast)] diff --git a/Mathlib/Topology/ContinuousMap/Basic.lean b/Mathlib/Topology/ContinuousMap/Basic.lean index 34583c8f81047..bd8aca0de948b 100644 --- a/Mathlib/Topology/ContinuousMap/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nicolò Cavalleri -/ import Mathlib.Data.Set.UnionLift +import Mathlib.Topology.ContinuousMap.Defs import Mathlib.Topology.Homeomorph /-! @@ -19,37 +20,6 @@ be satisfied by itself and all stricter types. open Function open scoped Topology -/-- The type of continuous maps from `α` to `β`. - -When possible, instead of parametrizing results over `(f : C(α, β))`, -you should parametrize over `{F : Type*} [ContinuousMapClass F α β] (f : F)`. - -When you extend this structure, make sure to extend `ContinuousMapClass`. -/ -structure ContinuousMap (α β : Type*) [TopologicalSpace α] [TopologicalSpace β] where - /-- The function `α → β` -/ - protected toFun : α → β - /-- Proposition that `toFun` is continuous -/ - protected continuous_toFun : Continuous toFun := by continuity - -/-- The type of continuous maps from `α` to `β`. -/ -notation "C(" α ", " β ")" => ContinuousMap α β - -section - -/-- `ContinuousMapClass F α β` states that `F` is a type of continuous maps. - -You should extend this class when you extend `ContinuousMap`. -/ -class ContinuousMapClass (F : Type*) (α β : outParam Type*) - [TopologicalSpace α] [TopologicalSpace β] [FunLike F α β] : Prop where - /-- Continuity -/ - map_continuous (f : F) : Continuous f - -end - -export ContinuousMapClass (map_continuous) - -attribute [continuity, fun_prop] map_continuous - section ContinuousMapClass variable {F α β : Type*} [TopologicalSpace α] [TopologicalSpace β] [FunLike F α β] @@ -61,11 +31,6 @@ theorem map_continuousAt (f : F) (a : α) : ContinuousAt f a := theorem map_continuousWithinAt (f : F) (s : Set α) (a : α) : ContinuousWithinAt f s a := (map_continuous f).continuousWithinAt -/-- Coerce a bundled morphism with a `ContinuousMapClass` instance to a `ContinuousMap`. -/ -@[coe] def toContinuousMap (f : F) : C(α, β) := ⟨f, map_continuous f⟩ - -instance : CoeTC F C(α, β) := ⟨toContinuousMap⟩ - end ContinuousMapClass /-! ### Continuous maps -/ @@ -76,75 +41,11 @@ namespace ContinuousMap variable {α β γ δ : Type*} [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] [TopologicalSpace δ] -instance funLike : FunLike C(α, β) α β where - coe := ContinuousMap.toFun - coe_injective' f g h := by cases f; cases g; congr - -instance toContinuousMapClass : ContinuousMapClass C(α, β) α β where - map_continuous := ContinuousMap.continuous_toFun - -@[simp] -theorem toFun_eq_coe {f : C(α, β)} : f.toFun = (f : α → β) := - rfl - -instance : CanLift (α → β) C(α, β) DFunLike.coe Continuous := ⟨fun f hf ↦ ⟨⟨f, hf⟩, rfl⟩⟩ - -/-- See note [custom simps projection]. -/ -def Simps.apply (f : C(α, β)) : α → β := f - --- this must come after the coe_to_fun definition -initialize_simps_projections ContinuousMap (toFun → apply) - -@[simp] -- Porting note: removed `norm_cast` attribute -protected theorem coe_coe {F : Type*} [FunLike F α β] [ContinuousMapClass F α β] (f : F) : - ⇑(f : C(α, β)) = f := - rfl - -@[ext] -theorem ext {f g : C(α, β)} (h : ∀ a, f a = g a) : f = g := - DFunLike.ext _ _ h - -/-- Copy of a `ContinuousMap` with a new `toFun` equal to the old one. Useful to fix definitional -equalities. -/ -protected def copy (f : C(α, β)) (f' : α → β) (h : f' = f) : C(α, β) where - toFun := f' - continuous_toFun := h.symm ▸ f.continuous_toFun - -@[simp] -theorem coe_copy (f : C(α, β)) (f' : α → β) (h : f' = f) : ⇑(f.copy f' h) = f' := - rfl - -theorem copy_eq (f : C(α, β)) (f' : α → β) (h : f' = f) : f.copy f' h = f := - DFunLike.ext' h - variable {f g : C(α, β)} -/-- Deprecated. Use `map_continuous` instead. -/ -protected theorem continuous (f : C(α, β)) : Continuous f := - f.continuous_toFun - -@[continuity] -theorem continuous_set_coe (s : Set C(α, β)) (f : s) : Continuous (f : α → β) := - f.1.continuous - /-- Deprecated. Use `map_continuousAt` instead. -/ protected theorem continuousAt (f : C(α, β)) (x : α) : ContinuousAt f x := - f.continuous.continuousAt - -/-- Deprecated. Use `DFunLike.congr_fun` instead. -/ -protected theorem congr_fun {f g : C(α, β)} (H : f = g) (x : α) : f x = g x := - H ▸ rfl - -/-- Deprecated. Use `DFunLike.congr_arg` instead. -/ -protected theorem congr_arg (f : C(α, β)) {x y : α} (h : x = y) : f x = f y := - h ▸ rfl - -theorem coe_injective : @Function.Injective C(α, β) (α → β) (↑) := fun f g h => by - cases f; cases g; congr - -@[simp] -theorem coe_mk (f : α → β) (h : Continuous f) : ⇑(⟨f, h⟩ : C(α, β)) = f := - rfl + map_continuousAt f x theorem map_specializes (f : C(α, β)) {x y : α} (h : x ⤳ y) : f x ⤳ f y := h.map f.2 @@ -271,8 +172,6 @@ def prodMk (f : C(α, β₁)) (g : C(α, β₂)) : C(α, β₁ × β₂) where @[simps] def prodMap (f : C(α₁, α₂)) (g : C(β₁, β₂)) : C(α₁ × β₁, α₂ × β₂) where toFun := Prod.map f g - continuous_toFun := f.continuous.prod_map g.continuous - -- Porting note: proof was `continuity` @[simp] theorem prod_eval (f : C(α, β₁)) (g : C(α, β₂)) (a : α) : (prodMk f g) a = (f a, g a) := @@ -380,13 +279,14 @@ theorem restrict_apply_mk (f : C(α, β)) (s : Set α) (x : α) (hx : x ∈ s) : theorem injective_restrict [T2Space β] {s : Set α} (hs : Dense s) : Injective (restrict s : C(α, β) → C(s, β)) := fun f g h ↦ - DFunLike.ext' <| f.continuous.ext_on hs g.continuous <| Set.restrict_eq_restrict_iff.1 <| - congr_arg DFunLike.coe h + DFunLike.ext' <| (map_continuous f).ext_on hs (map_continuous g) <| + Set.restrict_eq_restrict_iff.1 <| congr_arg DFunLike.coe h /-- The restriction of a continuous map to the preimage of a set. -/ @[simps] def restrictPreimage (f : C(α, β)) (s : Set β) : C(f ⁻¹' s, s) := - ⟨s.restrictPreimage f, continuous_iff_continuousAt.mpr fun _ => f.2.continuousAt.restrictPreimage⟩ + ⟨s.restrictPreimage f, continuous_iff_continuousAt.mpr fun _ ↦ + (map_continuousAt f _).restrictPreimage⟩ end Restrict @@ -404,8 +304,8 @@ noncomputable def liftCover : C(α, β) := Set.iUnion_eq_univ_iff.2 fun x ↦ (hS x).imp fun _ ↦ mem_of_mem_nhds mk (Set.liftCover S (fun i ↦ φ i) hφ H) <| continuous_of_cover_nhds hS fun i ↦ by rw [continuousOn_iff_continuous_restrict] - simpa (config := { unfoldPartialApp := true }) only [Set.restrict, Set.liftCover_coe] using - (φ i).continuous + simpa (config := { unfoldPartialApp := true }) only [Set.restrict, Set.liftCover_coe] + using map_continuous (φ i) variable {S φ hφ hS} @@ -463,8 +363,8 @@ variable {X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalS def Function.RightInverse.homeomorph {f' : C(Y, X)} (hf : Function.RightInverse f' f) : Quotient (Setoid.ker f) ≃ₜ Y where toEquiv := Setoid.quotientKerEquivOfRightInverse _ _ hf - continuous_toFun := quotientMap_quot_mk.continuous_iff.mpr f.continuous - continuous_invFun := continuous_quotient_mk'.comp f'.continuous + continuous_toFun := quotientMap_quot_mk.continuous_iff.mpr (map_continuous f) + continuous_invFun := continuous_quotient_mk'.comp (map_continuous f') namespace QuotientMap diff --git a/Mathlib/Topology/ContinuousMap/Bounded.lean b/Mathlib/Topology/ContinuousMap/Bounded.lean index a1fa188b5cc16..01840e85e6d01 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded.lean @@ -524,7 +524,7 @@ theorem arzela_ascoli₂ (s : Set β) (hs : IsCompact s) (A : Set (α →ᵇ β) fun f hf => ?_ · haveI : CompactSpace s := isCompact_iff_compactSpace.1 hs refine arzela_ascoli₁ _ (continuous_iff_isClosed.1 (continuous_comp M) _ closed) ?_ - rw [isUniformEmbedding_subtype_val.toUniformInducing.equicontinuous_iff] + rw [isUniformEmbedding_subtype_val.isUniformInducing.equicontinuous_iff] exact H.comp (A.restrictPreimage F) · let g := codRestrict s f fun x => in_s f x hf rw [show f = F g by ext; rfl] at hf ⊢ diff --git a/Mathlib/Topology/ContinuousMap/Compact.lean b/Mathlib/Topology/ContinuousMap/Compact.lean index f68ff46c4011d..ca3851110e939 100644 --- a/Mathlib/Topology/ContinuousMap/Compact.lean +++ b/Mathlib/Topology/ContinuousMap/Compact.lean @@ -29,8 +29,8 @@ open NNReal BoundedContinuousFunction Set Metric namespace ContinuousMap -variable {α β E : Type*} [TopologicalSpace α] [CompactSpace α] [MetricSpace β] - [NormedAddCommGroup E] +variable {α β E : Type*} +variable [TopologicalSpace α] [CompactSpace α] [PseudoMetricSpace β] [SeminormedAddCommGroup E] section @@ -47,8 +47,8 @@ def equivBoundedOfCompact : C(α, β) ≃ (α →ᵇ β) := ext rfl⟩ -theorem uniformInducing_equivBoundedOfCompact : UniformInducing (equivBoundedOfCompact α β) := - UniformInducing.mk' +theorem isUniformInducing_equivBoundedOfCompact : IsUniformInducing (equivBoundedOfCompact α β) := + IsUniformInducing.mk' (by simp only [hasBasis_compactConvergenceUniformity.mem_iff, uniformity_basis_dist_le.mem_iff] exact fun s => @@ -59,8 +59,12 @@ theorem uniformInducing_equivBoundedOfCompact : UniformInducing (equivBoundedOfC ⟨⟨Set.univ, { p | dist p.1 p.2 ≤ ε }⟩, ⟨isCompact_univ, ⟨ε, hε, fun _ h => h⟩⟩, fun ⟨f, g⟩ h => hs _ _ (ht ((dist_le hε.le).mpr fun x => h x (mem_univ x)))⟩⟩) +@[deprecated (since := "2024-10-05")] +alias uniformInducing_equivBoundedOfCompact := isUniformInducing_equivBoundedOfCompact + theorem isUniformEmbedding_equivBoundedOfCompact : IsUniformEmbedding (equivBoundedOfCompact α β) := - { uniformInducing_equivBoundedOfCompact α β with inj := (equivBoundedOfCompact α β).injective } + { isUniformInducing_equivBoundedOfCompact α β with + inj := (equivBoundedOfCompact α β).injective } @[deprecated (since := "2024-10-01")] alias uniformEmbedding_equivBoundedOfCompact := isUniformEmbedding_equivBoundedOfCompact @@ -85,9 +89,14 @@ theorem addEquivBoundedOfCompact_apply [AddMonoid β] [LipschitzAdd β] : ⇑(addEquivBoundedOfCompact α β) = mkOfCompact := rfl -instance metricSpace : MetricSpace C(α, β) := +instance instPseudoMetricSpace : PseudoMetricSpace C(α, β) := + (isUniformEmbedding_equivBoundedOfCompact α β).comapPseudoMetricSpace _ + +instance instMetricSpace {β : Type*} [MetricSpace β] : + MetricSpace C(α, β) := (isUniformEmbedding_equivBoundedOfCompact α β).comapMetricSpace _ + /-- When `α` is compact, and `β` is a metric space, the bounded continuous maps `α →ᵇ β` are isometric to `C(α, β)`. -/ @@ -136,6 +145,13 @@ theorem dist_lt_iff (C0 : (0 : ℝ) < C) : dist f g < C ↔ ∀ x : α, dist (f rw [← dist_mkOfCompact, dist_lt_iff_of_compact C0] simp only [mkOfCompact_apply] +instance {R} [Zero R] [Zero β] [PseudoMetricSpace R] [SMul R β] [BoundedSMul R β] : + BoundedSMul R C(α, β) where + dist_smul_pair' r f g := by + simpa only [← dist_mkOfCompact] using dist_smul_pair r (mkOfCompact f) (mkOfCompact g) + dist_pair_smul' r₁ r₂ f := by + simpa only [← dist_mkOfCompact] using dist_pair_smul r₁ r₂ (mkOfCompact f) + end -- TODO at some point we will need lemmas characterising this norm! @@ -153,13 +169,17 @@ theorem _root_.BoundedContinuousFunction.norm_toContinuousMap_eq (f : α →ᵇ open BoundedContinuousFunction -instance : NormedAddCommGroup C(α, E) := - { ContinuousMap.metricSpace _ _, - ContinuousMap.instAddCommGroupContinuousMap with - dist_eq := fun x y => by - rw [← norm_mkOfCompact, ← dist_mkOfCompact, dist_eq_norm, mkOfCompact_sub] - dist := dist - norm := norm } +instance : SeminormedAddCommGroup C(α, E) where + __ := ContinuousMap.instPseudoMetricSpace _ _ + __ := ContinuousMap.instAddCommGroupContinuousMap + dist_eq x y := by + rw [← norm_mkOfCompact, ← dist_mkOfCompact, dist_eq_norm, mkOfCompact_sub] + dist := dist + norm := norm + +instance {E : Type*} [NormedAddCommGroup E] : NormedAddCommGroup C(α, E) where + __ : SeminormedAddCommGroup C(α, E) := inferInstance + __ : MetricSpace C(α, E) := inferInstance instance [Nonempty α] [One E] [NormOneClass E] : NormOneClass C(α, E) where norm_one := by simp only [← norm_mkOfCompact, mkOfCompact_one, norm_one] @@ -218,11 +238,40 @@ end section -variable {R : Type*} [NormedRing R] +variable {R : Type*} + +instance [NonUnitalSeminormedRing R] : NonUnitalSeminormedRing C(α, R) where + __ : SeminormedAddCommGroup C(α, R) := inferInstance + __ : NonUnitalRing C(α, R) := inferInstance + norm_mul f g := norm_mul_le (mkOfCompact f) (mkOfCompact g) + +instance [NonUnitalSeminormedCommRing R] : NonUnitalSeminormedCommRing C(α, R) where + __ : NonUnitalSeminormedRing C(α, R) := inferInstance + __ : NonUnitalCommRing C(α, R) := inferInstance + +instance [SeminormedRing R] : SeminormedRing C(α, R) where + __ : NonUnitalSeminormedRing C(α, R) := inferInstance + __ : Ring C(α, R) := inferInstance + +instance [SeminormedCommRing R] : SeminormedCommRing C(α, R) where + __ : SeminormedRing C(α, R) := inferInstance + __ : CommRing C(α, R) := inferInstance + +instance [NonUnitalNormedRing R] : NonUnitalNormedRing C(α, R) where + __ : NormedAddCommGroup C(α, R) := inferInstance + __ : NonUnitalSeminormedRing C(α, R) := inferInstance + +instance [NonUnitalNormedCommRing R] : NonUnitalNormedCommRing C(α, R) where + __ : NonUnitalNormedRing C(α, R) := inferInstance + __ : NonUnitalCommRing C(α, R) := inferInstance + +instance [NormedRing R] : NormedRing C(α, R) where + __ : NormedAddCommGroup C(α, R) := inferInstance + __ : SeminormedRing C(α, R) := inferInstance -instance : NormedRing C(α, R) := - { (inferInstance : NormedAddCommGroup C(α, R)), ContinuousMap.instRing with - norm_mul := fun f g => norm_mul_le (mkOfCompact f) (mkOfCompact g) } +instance [NormedCommRing R] : NormedCommRing C(α, R) where + __ : NormedRing C(α, R) := inferInstance + __ : CommRing C(α, R) := inferInstance end @@ -231,7 +280,7 @@ section variable {𝕜 : Type*} [NormedField 𝕜] [NormedSpace 𝕜 E] instance normedSpace : NormedSpace 𝕜 C(α, E) where - norm_smul_le c f := (norm_smul_le c (mkOfCompact f) : _) + norm_smul_le := norm_smul_le section @@ -290,7 +339,7 @@ end section -variable {𝕜 : Type*} {γ : Type*} [NormedField 𝕜] [NormedRing γ] [NormedAlgebra 𝕜 γ] +variable {𝕜 : Type*} {γ : Type*} [NormedField 𝕜] [SeminormedRing γ] [NormedAlgebra 𝕜 γ] instance : NormedAlgebra 𝕜 C(α, γ) := { ContinuousMap.normedSpace, ContinuousMap.algebra with } @@ -304,7 +353,7 @@ namespace ContinuousMap section UniformContinuity variable {α β : Type*} -variable [MetricSpace α] [CompactSpace α] [MetricSpace β] +variable [PseudoMetricSpace α] [CompactSpace α] [PseudoMetricSpace β] /-! We now set up some declarations making it convenient to use uniform continuity. @@ -338,7 +387,7 @@ section CompLeft variable (X : Type*) {𝕜 β γ : Type*} [TopologicalSpace X] [CompactSpace X] [NontriviallyNormedField 𝕜] -variable [NormedAddCommGroup β] [NormedSpace 𝕜 β] [NormedAddCommGroup γ] [NormedSpace 𝕜 γ] +variable [SeminormedAddCommGroup β] [NormedSpace 𝕜 β] [SeminormedAddCommGroup γ] [NormedSpace 𝕜 γ] open ContinuousMap @@ -385,7 +434,8 @@ section CompRight /-- Precomposition by a continuous map is itself a continuous map between spaces of continuous maps. -/ def compRightContinuousMap {X Y : Type*} (T : Type*) [TopologicalSpace X] [CompactSpace X] - [TopologicalSpace Y] [CompactSpace Y] [MetricSpace T] (f : C(X, Y)) : C(C(Y, T), C(X, T)) where + [TopologicalSpace Y] [CompactSpace Y] [PseudoMetricSpace T] (f : C(X, Y)) : + C(C(Y, T), C(X, T)) where toFun g := g.comp f continuous_toFun := by refine Metric.continuous_iff.mpr ?_ @@ -396,14 +446,15 @@ def compRightContinuousMap {X Y : Type*} (T : Type*) [TopologicalSpace X] [Compa @[simp] theorem compRightContinuousMap_apply {X Y : Type*} (T : Type*) [TopologicalSpace X] - [CompactSpace X] [TopologicalSpace Y] [CompactSpace Y] [MetricSpace T] (f : C(X, Y)) + [CompactSpace X] [TopologicalSpace Y] [CompactSpace Y] [PseudoMetricSpace T] (f : C(X, Y)) (g : C(Y, T)) : (compRightContinuousMap T f) g = g.comp f := rfl /-- Precomposition by a homeomorphism is itself a homeomorphism between spaces of continuous maps. -/ def compRightHomeomorph {X Y : Type*} (T : Type*) [TopologicalSpace X] [CompactSpace X] - [TopologicalSpace Y] [CompactSpace Y] [MetricSpace T] (f : X ≃ₜ Y) : C(Y, T) ≃ₜ C(X, T) where + [TopologicalSpace Y] [CompactSpace Y] [PseudoMetricSpace T] (f : X ≃ₜ Y) : + C(Y, T) ≃ₜ C(X, T) where toFun := compRightContinuousMap T f.toContinuousMap invFun := compRightContinuousMap T f.symm.toContinuousMap left_inv g := ext fun _ => congr_arg g (f.apply_symm_apply _) @@ -411,7 +462,7 @@ def compRightHomeomorph {X Y : Type*} (T : Type*) [TopologicalSpace X] [CompactS theorem compRightAlgHom_continuous {X Y : Type*} (R A : Type*) [TopologicalSpace X] [CompactSpace X] [TopologicalSpace Y] [CompactSpace Y] [CommSemiring R] [Semiring A] - [MetricSpace A] [TopologicalSemiring A] [Algebra R A] (f : C(X, Y)) : + [PseudoMetricSpace A] [TopologicalSemiring A] [Algebra R A] (f : C(X, Y)) : Continuous (compRightAlgHom R A f) := map_continuous (compRightContinuousMap A f) @@ -460,7 +511,7 @@ Furthermore, if `α` is compact and `β` is a C⋆-ring, then `C(α, β)` is a C section NormedSpace variable {α : Type*} {β : Type*} -variable [TopologicalSpace α] [NormedAddCommGroup β] [StarAddMonoid β] [NormedStarGroup β] +variable [TopologicalSpace α] [SeminormedAddCommGroup β] [StarAddMonoid β] [NormedStarGroup β] theorem _root_.BoundedContinuousFunction.mkOfCompact_star [CompactSpace α] (f : C(α, β)) : mkOfCompact (star f) = star (mkOfCompact f) := @@ -476,7 +527,7 @@ end NormedSpace section CStarRing variable {α : Type*} {β : Type*} -variable [TopologicalSpace α] [NormedRing β] [StarRing β] +variable [TopologicalSpace α] [NonUnitalNormedRing β] [StarRing β] instance [CompactSpace α] [CStarRing β] : CStarRing C(α, β) where norm_mul_self_le f := by diff --git a/Mathlib/Topology/ContinuousMap/CompactlySupported.lean b/Mathlib/Topology/ContinuousMap/CompactlySupported.lean index 2b51d3ea2f2d0..cea48f4937a02 100644 --- a/Mathlib/Topology/ContinuousMap/CompactlySupported.lean +++ b/Mathlib/Topology/ContinuousMap/CompactlySupported.lean @@ -119,7 +119,6 @@ theorem eq_of_empty [IsEmpty α] (f g : C_c(α, β)) : f = g := def ContinuousMap.liftCompactlySupported [CompactSpace α] : C(α, β) ≃ C_c(α, β) where toFun f := { toFun := f - continuous_toFun := f.continuous hasCompactSupport' := HasCompactSupport.of_compactSpace f } invFun f := f left_inv _ := rfl @@ -168,7 +167,7 @@ theorem mul_apply [MulZeroClass β] [ContinuousMul β] (f g : C_c(α, β)) : (f instance [Zero β] [TopologicalSpace γ] [SMulZeroClass γ β] [ContinuousSMul γ β] {F : Type*} [FunLike F α γ] [ContinuousMapClass F α γ] : SMul F C_c(α, β) where smul f g := - ⟨⟨fun x ↦ f x • g x, (map_continuous f).smul g.continuous⟩, g.hasCompactSupport'.smul_left⟩ + ⟨⟨fun x ↦ f x • g x, (map_continuous f).smul (map_continuous g)⟩, g.hasCompactSupport.smul_left⟩ @[simp] theorem coe_smulc [Zero β] [TopologicalSpace γ] [SMulZeroClass γ β] [ContinuousSMul γ β] @@ -209,7 +208,7 @@ def coeFnMonoidHom [AddMonoid β] [ContinuousAdd β] : C_c(α, β) →+ α → instance [Zero β] {R : Type*} [SMulZeroClass R β] [ContinuousConstSMul R β] : SMul R C_c(α, β) := - ⟨fun r f => ⟨⟨r • ⇑f, Continuous.const_smul f.continuous r⟩, HasCompactSupport.smul_left f.2⟩⟩ + ⟨fun r f => ⟨⟨r • ⇑f, (map_continuous f).const_smul r⟩, HasCompactSupport.smul_left f.2⟩⟩ @[simp, norm_cast] theorem coe_smul [Zero β] {R : Type*} [SMulZeroClass R β] [ContinuousConstSMul R β] (r : R) diff --git a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean index 7dcfc9a8fff16..5fa65e8b7c342 100644 --- a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean @@ -277,7 +277,7 @@ lemma isUniformEmbedding_toContinuousMap : alias uniformEmbedding_toContinuousMap := isUniformEmbedding_toContinuousMap instance [T1Space R] [CompleteSpace C(X, R)] : CompleteSpace C(X, R)₀ := - completeSpace_iff_isComplete_range isUniformEmbedding_toContinuousMap.toUniformInducing + completeSpace_iff_isComplete_range isUniformEmbedding_toContinuousMap.isUniformInducing |>.mpr closedEmbedding_toContinuousMap.isClosed_range.isComplete lemma isUniformEmbedding_comp {Y : Type*} [UniformSpace Y] [Zero Y] (g : C(Y, R)₀) diff --git a/Mathlib/Topology/ContinuousMap/Defs.lean b/Mathlib/Topology/ContinuousMap/Defs.lean new file mode 100644 index 0000000000000..f576214a27758 --- /dev/null +++ b/Mathlib/Topology/ContinuousMap/Defs.lean @@ -0,0 +1,136 @@ +/- +Copyright (c) 2020 Nicolò Cavalleri. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nicolò Cavalleri, Yury Kudryashov +-/ +import Mathlib.Tactic.Continuity +import Mathlib.Tactic.Lift +import Mathlib.Topology.Defs.Basic + +/-! +# Continuous bundled maps + +In this file we define the type `ContinuousMap` of continuous bundled maps. + +We use the `DFunLike` design, so each type of morphisms has a companion typeclass +which is meant to be satisfied by itself and all stricter types. +-/ + +open Function +open scoped Topology + +/-- The type of continuous maps from `X` to `Y`. + +When possible, instead of parametrizing results over `(f : C(X, Y))`, +you should parametrize over `{F : Type*} [ContinuousMapClass F X Y] (f : F)`. + +When you extend this structure, make sure to extend `ContinuousMapClass`. -/ +structure ContinuousMap (X Y : Type*) [TopologicalSpace X] [TopologicalSpace Y] where + /-- The function `X → Y` -/ + protected toFun : X → Y + /-- Proposition that `toFun` is continuous -/ + protected continuous_toFun : Continuous toFun := by continuity + +/-- The type of continuous maps from `X` to `Y`. -/ +notation "C(" X ", " Y ")" => ContinuousMap X Y + +section + +/-- `ContinuousMapClass F X Y` states that `F` is a type of continuous maps. + +You should extend this class when you extend `ContinuousMap`. -/ +class ContinuousMapClass (F : Type*) (X Y : outParam Type*) + [TopologicalSpace X] [TopologicalSpace Y] [FunLike F X Y] : Prop where + /-- Continuity -/ + map_continuous (f : F) : Continuous f + +end + +export ContinuousMapClass (map_continuous) + +attribute [continuity, fun_prop] map_continuous + +section ContinuousMapClass + +variable {F X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [FunLike F X Y] +variable [ContinuousMapClass F X Y] + +/-- Coerce a bundled morphism with a `ContinuousMapClass` instance to a `ContinuousMap`. -/ +@[coe] def toContinuousMap (f : F) : C(X, Y) := ⟨f, map_continuous f⟩ + +instance : CoeTC F C(X, Y) := ⟨toContinuousMap⟩ + +end ContinuousMapClass + +/-! ### Continuous maps -/ + + +namespace ContinuousMap + +variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] + +instance instFunLike : FunLike C(X, Y) X Y where + coe := ContinuousMap.toFun + coe_injective' f g h := by cases f; cases g; congr + +instance instContinuousMapClass : ContinuousMapClass C(X, Y) X Y where + map_continuous := ContinuousMap.continuous_toFun + +@[simp] +theorem toFun_eq_coe {f : C(X, Y)} : f.toFun = (f : X → Y) := + rfl + +instance : CanLift (X → Y) C(X, Y) DFunLike.coe Continuous := ⟨fun f hf ↦ ⟨⟨f, hf⟩, rfl⟩⟩ + +/-- See note [custom simps projection]. -/ +def Simps.apply (f : C(X, Y)) : X → Y := f + +-- this must come after the coe_to_fun definition +initialize_simps_projections ContinuousMap (toFun → apply) + +@[simp] -- Porting note: removed `norm_cast` attribute +protected theorem coe_coe {F : Type*} [FunLike F X Y] [ContinuousMapClass F X Y] (f : F) : + ⇑(f : C(X, Y)) = f := + rfl + +@[ext] +theorem ext {f g : C(X, Y)} (h : ∀ a, f a = g a) : f = g := + DFunLike.ext _ _ h + +/-- Copy of a `ContinuousMap` with a new `toFun` equal to the old one. Useful to fix definitional +equalities. -/ +protected def copy (f : C(X, Y)) (f' : X → Y) (h : f' = f) : C(X, Y) where + toFun := f' + continuous_toFun := h.symm ▸ f.continuous_toFun + +@[simp] +theorem coe_copy (f : C(X, Y)) (f' : X → Y) (h : f' = f) : ⇑(f.copy f' h) = f' := + rfl + +theorem copy_eq (f : C(X, Y)) (f' : X → Y) (h : f' = f) : f.copy f' h = f := + DFunLike.ext' h + +/-- Deprecated. Use `map_continuous` instead. -/ +protected theorem continuous (f : C(X, Y)) : Continuous f := + f.continuous_toFun + +@[deprecated map_continuous (since := "2024-09-29")] +theorem continuous_set_coe (s : Set C(X, Y)) (f : s) : Continuous (f : X → Y) := + map_continuous _ + +/-- Deprecated. Use `DFunLike.congr_fun` instead. -/ +protected theorem congr_fun {f g : C(X, Y)} (H : f = g) (x : X) : f x = g x := + H ▸ rfl + +/-- Deprecated. Use `DFunLike.congr_arg` instead. -/ +protected theorem congr_arg (f : C(X, Y)) {x y : X} (h : x = y) : f x = f y := + h ▸ rfl + +theorem coe_injective : Function.Injective (DFunLike.coe : C(X, Y) → (X → Y)) := + DFunLike.coe_injective + +@[simp] +theorem coe_mk (f : X → Y) (h : Continuous f) : ⇑(⟨f, h⟩ : C(X, Y)) = f := + rfl + +end ContinuousMap diff --git a/Mathlib/Topology/ContinuousMap/LocallyConstant.lean b/Mathlib/Topology/ContinuousMap/LocallyConstant.lean index dce7a9a92d359..bbe26fe3ba211 100644 --- a/Mathlib/Topology/ContinuousMap/LocallyConstant.lean +++ b/Mathlib/Topology/ContinuousMap/LocallyConstant.lean @@ -15,7 +15,7 @@ import Mathlib.Topology.ContinuousMap.Algebra namespace LocallyConstant -variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] (f : LocallyConstant X Y) +variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] /-- The inclusion of locally-constant functions into continuous functions as a multiplicative monoid hom. -/ diff --git a/Mathlib/Topology/ContinuousMap/Ordered.lean b/Mathlib/Topology/ContinuousMap/Ordered.lean index 24c45724b82b7..14a4185687a12 100644 --- a/Mathlib/Topology/ContinuousMap/Ordered.lean +++ b/Mathlib/Topology/ContinuousMap/Ordered.lean @@ -3,9 +3,9 @@ Copyright (c) 2021 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Shing Tak Lam -/ -import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Topology.Order.Lattice import Mathlib.Topology.Order.ProjIcc +import Mathlib.Topology.ContinuousMap.Defs /-! # Bundled continuous maps into orders, with order-compatible topology diff --git a/Mathlib/Topology/ContinuousMap/Sigma.lean b/Mathlib/Topology/ContinuousMap/Sigma.lean index a444781b712f5..320fd5cc4ab1f 100644 --- a/Mathlib/Topology/ContinuousMap/Sigma.lean +++ b/Mathlib/Topology/ContinuousMap/Sigma.lean @@ -63,7 +63,7 @@ some `i` and a continuous map `g : C(X, Y i)`. See also `Continuous.exists_lift_ with unbundled functions and `ContinuousMap.sigmaCodHomeomorph` for a homeomorphism defined using this fact. -/ theorem exists_lift_sigma (f : C(X, Σ i, Y i)) : ∃ i g, f = (sigmaMk i).comp g := - let ⟨i, g, hg, hfg⟩ := f.continuous.exists_lift_sigma + let ⟨i, g, hg, hfg⟩ := (map_continuous f).exists_lift_sigma ⟨i, ⟨g, hg⟩, DFunLike.ext' hfg⟩ variable (X Y) diff --git a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean index ff47e32977e60..643b6ad6595d6 100644 --- a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean @@ -222,9 +222,7 @@ theorem sublattice_closure_eq_top (L : Set C(X, ℝ)) (nA : L.Nonempty) have W_nhd : ∀ x, W x ∈ 𝓝 x := by intro x refine IsOpen.mem_nhds ?_ ?_ - · -- Porting note: mathlib3 `continuity` found `continuous_set_coe` - apply isOpen_lt (continuous_set_coe _ _) - continuity + · apply isOpen_lt <;> fun_prop · dsimp only [W, Set.mem_setOf_eq] rw [h_eq] exact lt_add_of_pos_right _ pos diff --git a/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean b/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean index c8c2dd96d5d02..b9c5713bf06fd 100644 --- a/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean +++ b/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean @@ -428,7 +428,7 @@ theorem isClosed_range_toBCF : IsClosed (range (toBCF : C₀(α, β) → α → /-- Continuous functions vanishing at infinity taking values in a complete space form a complete space. -/ instance instCompleteSpace [CompleteSpace β] : CompleteSpace C₀(α, β) := - (completeSpace_iff_isComplete_range isometry_toBCF.uniformInducing).mpr + (completeSpace_iff_isComplete_range isometry_toBCF.isUniformInducing).mpr isClosed_range_toBCF.isComplete end Metric diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index b526794778cbb..7608fd4158dbb 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -54,12 +54,20 @@ theorem mem_closure_ne_iff_frequently_within {z : α} {s : Set α} : simp [mem_closure_iff_frequently, frequently_nhdsWithin_iff] @[simp] -theorem eventually_nhdsWithin_nhdsWithin {a : α} {s : Set α} {p : α → Prop} : +theorem eventually_eventually_nhdsWithin {a : α} {s : Set α} {p : α → Prop} : (∀ᶠ y in 𝓝[s] a, ∀ᶠ x in 𝓝[s] y, p x) ↔ ∀ᶠ x in 𝓝[s] a, p x := by refine ⟨fun h => ?_, fun h => (eventually_nhds_nhdsWithin.2 h).filter_mono inf_le_left⟩ simp only [eventually_nhdsWithin_iff] at h ⊢ exact h.mono fun x hx hxs => (hx hxs).self_of_nhds hxs +@[deprecated (since := "2024-10-04")] +alias eventually_nhdsWithin_nhdsWithin := eventually_eventually_nhdsWithin + +@[simp] +theorem eventually_mem_nhdsWithin_iff {x : α} {s t : Set α} : + (∀ᶠ x' in 𝓝[s] x, t ∈ 𝓝[s] x') ↔ t ∈ 𝓝[s] x := + eventually_eventually_nhdsWithin + theorem nhdsWithin_eq (a : α) (s : Set α) : 𝓝[s] a = ⨅ t ∈ { t : Set α | a ∈ t ∧ IsOpen t }, 𝓟 (t ∩ s) := ((nhds_basis_opens a).inf_principal s).eq_biInf diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index 84bfa9ec96c78..210caa3dca529 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.lean @@ -81,11 +81,14 @@ theorem interior_compact_eq_empty [T2Space β] (di : IsDenseInducing i) (hd : De exact hyi (image_subset_range _ _ hys) /-- The product of two dense inducings is a dense inducing -/ -protected theorem prod [TopologicalSpace γ] [TopologicalSpace δ] {e₁ : α → β} {e₂ : γ → δ} +protected theorem prodMap [TopologicalSpace γ] [TopologicalSpace δ] {e₁ : α → β} {e₂ : γ → δ} (de₁ : IsDenseInducing e₁) (de₂ : IsDenseInducing e₂) : - IsDenseInducing fun p : α × γ => (e₁ p.1, e₂ p.2) where - toInducing := de₁.toInducing.prod_map de₂.toInducing - dense := de₁.dense.prod_map de₂.dense + IsDenseInducing (Prod.map e₁ e₂) where + toInducing := de₁.toInducing.prodMap de₂.toInducing + dense := de₁.dense.prodMap de₂.dense + +@[deprecated (since := "2024-10-06")] +protected alias prod := IsDenseInducing.prodMap open TopologicalSpace @@ -234,11 +237,13 @@ protected theorem separableSpace [SeparableSpace α] (de : IsDenseEmbedding e) : de.toIsDenseInducing.separableSpace /-- The product of two dense embeddings is a dense embedding. -/ -protected theorem prod {e₁ : α → β} {e₂ : γ → δ} (de₁ : IsDenseEmbedding e₁) +protected theorem prodMap {e₁ : α → β} {e₂ : γ → δ} (de₁ : IsDenseEmbedding e₁) (de₂ : IsDenseEmbedding e₂) : IsDenseEmbedding fun p : α × γ => (e₁ p.1, e₂ p.2) := - { de₁.toIsDenseInducing.prod de₂.toIsDenseInducing with + { de₁.toIsDenseInducing.prodMap de₂.toIsDenseInducing with inj := de₁.inj.prodMap de₂.inj } +@[deprecated (since := "2024-10-06")] protected alias prod := IsDenseEmbedding.prodMap + /-- The dense embedding of a subtype inside its closure. -/ @[simps] def subtypeEmb {α : Type*} (p : α → Prop) (e : α → β) (x : { x // p x }) : @@ -288,14 +293,14 @@ theorem isClosed_property [TopologicalSpace β] {e : α → β} {p : β → Prop theorem isClosed_property2 [TopologicalSpace β] {e : α → β} {p : β → β → Prop} (he : DenseRange e) (hp : IsClosed { q : β × β | p q.1 q.2 }) (h : ∀ a₁ a₂, p (e a₁) (e a₂)) : ∀ b₁ b₂, p b₁ b₂ := - have : ∀ q : β × β, p q.1 q.2 := isClosed_property (he.prod_map he) hp fun _ => h _ _ + have : ∀ q : β × β, p q.1 q.2 := isClosed_property (he.prodMap he) hp fun _ => h _ _ fun b₁ b₂ => this ⟨b₁, b₂⟩ theorem isClosed_property3 [TopologicalSpace β] {e : α → β} {p : β → β → β → Prop} (he : DenseRange e) (hp : IsClosed { q : β × β × β | p q.1 q.2.1 q.2.2 }) (h : ∀ a₁ a₂ a₃, p (e a₁) (e a₂) (e a₃)) : ∀ b₁ b₂ b₃, p b₁ b₂ b₃ := have : ∀ q : β × β × β, p q.1 q.2.1 q.2.2 := - isClosed_property (he.prod_map <| he.prod_map he) hp fun _ => h _ _ _ + isClosed_property (he.prodMap <| he.prodMap he) hp fun _ => h _ _ _ fun b₁ b₂ b₃ => this ⟨b₁, b₂, b₃⟩ @[elab_as_elim] diff --git a/Mathlib/Topology/EMetricSpace/Basic.lean b/Mathlib/Topology/EMetricSpace/Basic.lean index deca954dd5665..0aba00ba4c0d7 100644 --- a/Mathlib/Topology/EMetricSpace/Basic.lean +++ b/Mathlib/Topology/EMetricSpace/Basic.lean @@ -58,18 +58,21 @@ theorem edist_le_range_sum_of_edist_le {f : ℕ → α} (n : ℕ) {d : ℕ → namespace EMetric -theorem uniformInducing_iff [PseudoEMetricSpace β] {f : α → β} : - UniformInducing f ↔ UniformContinuous f ∧ +theorem isUniformInducing_iff [PseudoEMetricSpace β] {f : α → β} : + IsUniformInducing f ↔ UniformContinuous f ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, edist (f a) (f b) < ε → edist a b < δ := - uniformInducing_iff'.trans <| Iff.rfl.and <| + isUniformInducing_iff'.trans <| Iff.rfl.and <| ((uniformity_basis_edist.comap _).le_basis_iff uniformity_basis_edist).trans <| by simp only [subset_def, Prod.forall]; rfl +@[deprecated (since := "2024-10-05")] +alias uniformInducing_iff := isUniformInducing_iff + /-- ε-δ characterization of uniform embeddings on pseudoemetric spaces -/ nonrec theorem isUniformEmbedding_iff [PseudoEMetricSpace β] {f : α → β} : IsUniformEmbedding f ↔ Function.Injective f ∧ UniformContinuous f ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, edist (f a) (f b) < ε → edist a b < δ := - (isUniformEmbedding_iff _).trans <| and_comm.trans <| Iff.rfl.and uniformInducing_iff + (isUniformEmbedding_iff _).trans <| and_comm.trans <| Iff.rfl.and isUniformInducing_iff @[deprecated (since := "2024-10-01")] alias uniformEmbedding_iff := isUniformEmbedding_iff @@ -77,7 +80,7 @@ alias uniformEmbedding_iff := isUniformEmbedding_iff /-- If a map between pseudoemetric spaces is a uniform embedding then the edistance between `f x` and `f y` is controlled in terms of the distance between `x` and `y`. -In fact, this lemma holds for a `UniformInducing` map. +In fact, this lemma holds for a `IsUniformInducing` map. TODO: generalize? -/ theorem controlled_of_isUniformEmbedding [PseudoEMetricSpace β] {f : α → β} (h : IsUniformEmbedding f) : @@ -242,7 +245,7 @@ theorem EMetric.isUniformEmbedding_iff' [EMetricSpace β] {f : γ → β} : IsUniformEmbedding f ↔ (∀ ε > 0, ∃ δ > 0, ∀ {a b : γ}, edist a b < δ → edist (f a) (f b) < ε) ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : γ}, edist (f a) (f b) < ε → edist a b < δ := by - rw [isUniformEmbedding_iff_uniformInducing, uniformInducing_iff, uniformContinuous_iff] + rw [isUniformEmbedding_iff_isUniformInducing, isUniformInducing_iff, uniformContinuous_iff] @[deprecated (since := "2024-10-01")] alias EMetric.uniformEmbedding_iff' := EMetric.isUniformEmbedding_iff' diff --git a/Mathlib/Topology/FiberBundle/Constructions.lean b/Mathlib/Topology/FiberBundle/Constructions.lean index 8a7ea9e221f3c..b5cac2629b8cf 100644 --- a/Mathlib/Topology/FiberBundle/Constructions.lean +++ b/Mathlib/Topology/FiberBundle/Constructions.lean @@ -138,8 +138,7 @@ theorem Prod.continuous_to_fun : ContinuousOn (Prod.toFun' e₁ e₂) have hf₁ : Continuous f₁ := (Prod.inducing_diag F₁ E₁ F₂ E₂).continuous have hf₂ : ContinuousOn f₂ (e₁.source ×ˢ e₂.source) := e₁.toPartialHomeomorph.continuousOn.prod_map e₂.toPartialHomeomorph.continuousOn - have hf₃ : Continuous f₃ := - (continuous_fst.comp continuous_fst).prod_mk (continuous_snd.prod_map continuous_snd) + have hf₃ : Continuous f₃ := by fun_prop refine ((hf₃.comp_continuousOn hf₂).comp hf₁.continuousOn ?_).congr ?_ · rw [e₁.source_eq, e₂.source_eq] exact mapsTo_preimage _ _ @@ -176,8 +175,7 @@ theorem Prod.right_inv {x : B × F₁ × F₂} theorem Prod.continuous_inv_fun : ContinuousOn (Prod.invFun' e₁ e₂) ((e₁.baseSet ∩ e₂.baseSet) ×ˢ univ) := by rw [(Prod.inducing_diag F₁ E₁ F₂ E₂).continuousOn_iff] - have H₁ : Continuous fun p : B × F₁ × F₂ ↦ ((p.1, p.2.1), (p.1, p.2.2)) := - (continuous_id.prod_map continuous_fst).prod_mk (continuous_id.prod_map continuous_snd) + have H₁ : Continuous fun p : B × F₁ × F₂ ↦ ((p.1, p.2.1), (p.1, p.2.2)) := by fun_prop refine (e₁.continuousOn_symm.prod_map e₂.continuousOn_symm).comp H₁.continuousOn ?_ exact fun x h ↦ ⟨⟨h.1.1, mem_univ _⟩, ⟨h.1.2, mem_univ _⟩⟩ @@ -226,7 +224,7 @@ variable [∀ x, Zero (E₁ x)] [∀ x, Zero (E₂ x)] [∀ x : B, TopologicalSp noncomputable instance FiberBundle.prod : FiberBundle (F₁ × F₂) (E₁ ×ᵇ E₂) where totalSpaceMk_inducing' b := by rw [← (Prod.inducing_diag F₁ E₁ F₂ E₂).of_comp_iff] - exact (totalSpaceMk_inducing F₁ E₁ b).prod_map (totalSpaceMk_inducing F₂ E₂ b) + exact (totalSpaceMk_inducing F₁ E₁ b).prodMap (totalSpaceMk_inducing F₂ E₂ b) trivializationAtlas' := { e | ∃ (e₁ : Trivialization F₁ (π F₁ E₁)) (e₂ : Trivialization F₂ (π F₂ E₂)) (_ : MemTrivializationAtlas e₁) (_ : MemTrivializationAtlas e₂), @@ -334,7 +332,7 @@ noncomputable def Trivialization.pullback (e : Trivialization F (π F E)) (f : K pullbackTotalSpaceEmbedding] refine continuousOn_fst.prod - (e.continuousOn_symm.comp ((map_continuous f).prod_map continuous_id).continuousOn + (e.continuousOn_symm.comp ((map_continuous f).prodMap continuous_id).continuousOn Subset.rfl) source_eq := by dsimp only diff --git a/Mathlib/Topology/Germ.lean b/Mathlib/Topology/Germ.lean index 5b1e7856f5d62..35d47bcfa9b8a 100644 --- a/Mathlib/Topology/Germ.lean +++ b/Mathlib/Topology/Germ.lean @@ -30,9 +30,6 @@ to the corresponding germ of functions `X → Z` at `x ∈ X` resp. `Y → Z` at `f` is constant. -/ -variable {F G : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] - [NormedAddCommGroup G] [NormedSpace ℝ G] - open scoped Topology open Filter Set diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index d620e28015df4..6dc61fbb47e77 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -349,10 +349,7 @@ def MkCore.t' (h : MkCore.{u}) (i j k : h.J) : refine ⟨⟨⟨(h.t i j x.1.1).1, ?_⟩, h.t i j x.1.1⟩, rfl⟩ rcases x with ⟨⟨⟨x, hx⟩, ⟨x', hx'⟩⟩, rfl : x = x'⟩ exact h.t_inter _ ⟨x, hx⟩ hx' - -- Porting note: was `continuity`, see https://github.com/leanprover-community/mathlib4/issues/5030 - have : Continuous (h.t i j) := map_continuous (self := ContinuousMap.toContinuousMapClass) _ - set_option tactic.skipAssignedInstances false in - exact ((Continuous.subtype_mk (by fun_prop) _).prod_mk (by fun_prop)).subtype_mk _ + fun_prop /-- This is a constructor of `TopCat.GlueData` whose arguments are in terms of elements and intersections rather than subobjects and pullbacks. Please refer to `TopCat.GlueData.MkCore` for diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index 228d23d3112d8..6d33a57317001 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -54,11 +54,11 @@ theorem toEquiv_injective : Function.Injective (toEquiv : X ≃ₜ Y → X ≃ Y | ⟨_, _, _⟩, ⟨_, _, _⟩, rfl => rfl instance : EquivLike (X ≃ₜ Y) X Y where - coe := fun h => h.toEquiv - inv := fun h => h.toEquiv.symm - left_inv := fun h => h.left_inv - right_inv := fun h => h.right_inv - coe_injective' := fun _ _ H _ => toEquiv_injective <| DFunLike.ext' H + coe h := h.toEquiv + inv h := h.toEquiv.symm + left_inv h := h.left_inv + right_inv h := h.right_inv + coe_injective' _ _ H _ := toEquiv_injective <| DFunLike.ext' H instance : CoeFun (X ≃ₜ Y) fun _ ↦ X → Y := ⟨DFunLike.coe⟩ @@ -498,8 +498,6 @@ def sumCongr (h₁ : X ≃ₜ X') (h₂ : Y ≃ₜ Y') : X ⊕ Y ≃ₜ X' ⊕ Y /-- Product of two homeomorphisms. -/ def prodCongr (h₁ : X ≃ₜ X') (h₂ : Y ≃ₜ Y') : X × Y ≃ₜ X' × Y' where - continuous_toFun := h₁.continuous.prod_map h₂.continuous - continuous_invFun := h₁.symm.continuous.prod_map h₂.symm.continuous toEquiv := h₁.toEquiv.prodCongr h₂.toEquiv @[simp] @@ -696,9 +694,9 @@ section Distrib def sumProdDistrib : (X ⊕ Y) × Z ≃ₜ (X × Z) ⊕ (Y × Z) := Homeomorph.symm <| homeomorphOfContinuousOpen (Equiv.sumProdDistrib X Y Z).symm - ((continuous_inl.prod_map continuous_id).sum_elim - (continuous_inr.prod_map continuous_id)) <| - (isOpenMap_inl.prod IsOpenMap.id).sum_elim (isOpenMap_inr.prod IsOpenMap.id) + ((continuous_inl.prodMap continuous_id).sum_elim + (continuous_inr.prodMap continuous_id)) <| + (isOpenMap_inl.prodMap IsOpenMap.id).sum_elim (isOpenMap_inr.prodMap IsOpenMap.id) /-- `X × (Y ⊕ Z)` is homeomorphic to `X × Y ⊕ X × Z`. -/ def prodSumDistrib : X × (Y ⊕ Z) ≃ₜ (X × Y) ⊕ (X × Z) := @@ -712,7 +710,7 @@ def sigmaProdDistrib : (Σ i, X i) × Y ≃ₜ Σ i, X i × Y := Homeomorph.symm <| homeomorphOfContinuousOpen (Equiv.sigmaProdDistrib X Y).symm (continuous_sigma fun _ => continuous_sigmaMk.fst'.prod_mk continuous_snd) - (isOpenMap_sigma.2 fun _ => isOpenMap_sigmaMk.prod IsOpenMap.id) + (isOpenMap_sigma.2 fun _ => isOpenMap_sigmaMk.prodMap IsOpenMap.id) end Distrib @@ -952,7 +950,7 @@ lemma IsHomeomorph.sumMap {g : Z → W} (hf : IsHomeomorph f) (hg : IsHomeomorph IsHomeomorph (Sum.map f g) := ⟨hf.1.sum_map hg.1, hf.2.sumMap hg.2, hf.3.sum_map hg.3⟩ lemma IsHomeomorph.prodMap {g : Z → W} (hf : IsHomeomorph f) (hg : IsHomeomorph g) : - IsHomeomorph (Prod.map f g) := ⟨hf.1.prod_map hg.1, hf.2.prod hg.2, hf.3.prodMap hg.3⟩ + IsHomeomorph (Prod.map f g) := ⟨hf.1.prodMap hg.1, hf.2.prodMap hg.2, hf.3.prodMap hg.3⟩ lemma IsHomeomorph.sigmaMap {ι κ : Type*} {X : ι → Type*} {Y : κ → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, TopologicalSpace (Y i)] {f : ι → κ} diff --git a/Mathlib/Topology/Homotopy/HomotopyGroup.lean b/Mathlib/Topology/Homotopy/HomotopyGroup.lean index 93fad3a1199db..daaa41db02500 100644 --- a/Mathlib/Topology/Homotopy/HomotopyGroup.lean +++ b/Mathlib/Topology/Homotopy/HomotopyGroup.lean @@ -185,7 +185,7 @@ theorem continuous_toLoop (i : N) : Continuous (@toLoop N X _ x _ i) := Path.continuous_uncurry_iff.1 <| Continuous.subtype_mk (ContinuousMap.continuous_eval.comp <| - Continuous.prod_map + Continuous.prodMap (ContinuousMap.continuous_curry.comp <| (ContinuousMap.continuous_comp_left _).comp continuous_subtype_val) continuous_id) diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index f122cb8067af0..a179072f0d072 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang, Yury Kudryashov -/ import Mathlib.Tactic.TFAE import Mathlib.Topology.ContinuousOn +import Mathlib.Topology.Maps.OpenQuotient /-! # Inseparable points in a topological space @@ -602,14 +603,8 @@ theorem map_mk_nhdsWithin_preimage (s : Set (SeparationQuotient X)) (x : X) : rw [nhdsWithin, ← comap_principal, Filter.push_pull, nhdsWithin, map_mk_nhds] /-- The map `(x, y) ↦ (mk x, mk y)` is a quotient map. -/ -theorem quotientMap_prodMap_mk : QuotientMap (Prod.map mk mk : X × Y → _) := by - have hsurj : Surjective (Prod.map mk mk : X × Y → _) := surjective_mk.prodMap surjective_mk - refine quotientMap_iff.2 ⟨hsurj, fun s ↦ ?_⟩ - refine ⟨fun hs ↦ hs.preimage (continuous_mk.prod_map continuous_mk), fun hs ↦ ?_⟩ - refine isOpen_iff_mem_nhds.2 <| hsurj.forall.2 fun (x, y) h ↦ ?_ - rw [Prod.map_mk, nhds_prod_eq, ← map_mk_nhds, ← map_mk_nhds, Filter.prod_map_map_eq', - ← nhds_prod_eq, Filter.mem_map] - exact hs.mem_nhds h +theorem quotientMap_prodMap_mk : QuotientMap (Prod.map mk mk : X × Y → _) := + (isOpenQuotientMap_mk.prodMap isOpenQuotientMap_mk).quotientMap /-- Lift a map `f : X → α` such that `Inseparable x y → f x = f y` to a map `SeparationQuotient X → α`. -/ diff --git a/Mathlib/Topology/Instances/Complex.lean b/Mathlib/Topology/Instances/Complex.lean index 6b8e3c0862a14..4693c143277c8 100644 --- a/Mathlib/Topology/Instances/Complex.lean +++ b/Mathlib/Topology/Instances/Complex.lean @@ -52,7 +52,7 @@ theorem Complex.uniformContinuous_ringHom_eq_id_or_conj (K : Subfield ℂ) {ψ : letI : TopologicalRing K.topologicalClosure := Subring.instTopologicalRing K.topologicalClosure.toSubring set ι : K → K.topologicalClosure := ⇑(Subfield.inclusion K.le_topologicalClosure) - have ui : UniformInducing ι := + have ui : IsUniformInducing ι := ⟨by erw [uniformity_subtype, uniformity_subtype, Filter.comap_comap] congr ⟩ diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 609c01f853877..5623966b95601 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -91,7 +91,7 @@ theorem continuousAt_coe_iff {α : Type*} [TopologicalSpace α] {x : ℝ≥0} {f theorem nhds_coe_coe {r p : ℝ≥0} : 𝓝 ((r : ℝ≥0∞), (p : ℝ≥0∞)) = (𝓝 (r, p)).map fun p : ℝ≥0 × ℝ≥0 => (↑p.1, ↑p.2) := - ((openEmbedding_coe.prod openEmbedding_coe).map_nhds_eq (r, p)).symm + ((openEmbedding_coe.prodMap openEmbedding_coe).map_nhds_eq (r, p)).symm theorem continuous_ofReal : Continuous ENNReal.ofReal := (continuous_coe_iff.2 continuous_id).comp continuous_real_toNNReal diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index 8b4ae160970e4..4d970aa153cb8 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -70,7 +70,7 @@ theorem nhds_coe {r : ℝ} : 𝓝 (r : EReal) = (𝓝 r).map (↑) := theorem nhds_coe_coe {r p : ℝ} : 𝓝 ((r : EReal), (p : EReal)) = (𝓝 (r, p)).map fun p : ℝ × ℝ => (↑p.1, ↑p.2) := - ((openEmbedding_coe.prod openEmbedding_coe).map_nhds_eq (r, p)).symm + ((openEmbedding_coe.prodMap openEmbedding_coe).map_nhds_eq (r, p)).symm theorem tendsto_toReal {a : EReal} (ha : a ≠ ⊤) (h'a : a ≠ ⊥) : Tendsto EReal.toReal (𝓝 a) (𝓝 a.toReal) := by @@ -358,7 +358,7 @@ private lemma continuousAt_mul_symm1 {a b : EReal} simp rw [this] apply ContinuousAt.comp (Continuous.continuousAt continuous_neg) - <| ContinuousAt.comp _ (ContinuousAt.prod_map (Continuous.continuousAt continuous_neg) + <| ContinuousAt.comp _ (ContinuousAt.prodMap (Continuous.continuousAt continuous_neg) (Continuous.continuousAt continuous_id)) simp [h] diff --git a/Mathlib/Topology/Instances/NNReal.lean b/Mathlib/Topology/Instances/NNReal.lean index d6ef63c1b5f66..6f0a140493505 100644 --- a/Mathlib/Topology/Instances/NNReal.lean +++ b/Mathlib/Topology/Instances/NNReal.lean @@ -6,6 +6,7 @@ Authors: Johan Commelin import Mathlib.Data.NNReal.Star import Mathlib.Topology.Algebra.InfiniteSum.Order import Mathlib.Topology.Algebra.InfiniteSum.Ring +import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Topology.Instances.Real import Mathlib.Topology.MetricSpace.Isometry diff --git a/Mathlib/Topology/Instances/Rat.lean b/Mathlib/Topology/Instances/Rat.lean index a5585543b41fe..715361b5de0d9 100644 --- a/Mathlib/Topology/Instances/Rat.lean +++ b/Mathlib/Topology/Instances/Rat.lean @@ -81,10 +81,10 @@ namespace Rat instance : NoncompactSpace ℚ := Int.closedEmbedding_coe_rat.noncompactSpace theorem uniformContinuous_add : UniformContinuous fun p : ℚ × ℚ => p.1 + p.2 := - Rat.isUniformEmbedding_coe_real.toUniformInducing.uniformContinuous_iff.2 <| by + Rat.isUniformEmbedding_coe_real.isUniformInducing.uniformContinuous_iff.2 <| by simp only [Function.comp_def, Rat.cast_add] exact Real.uniformContinuous_add.comp - (Rat.uniformContinuous_coe_real.prod_map Rat.uniformContinuous_coe_real) + (Rat.uniformContinuous_coe_real.prodMap Rat.uniformContinuous_coe_real) theorem uniformContinuous_neg : UniformContinuous (@Neg.neg ℚ _) := Metric.uniformContinuous_iff.2 fun ε ε0 => @@ -106,7 +106,7 @@ instance : TopologicalRing ℚ := inferInstance nonrec theorem totallyBounded_Icc (a b : ℚ) : TotallyBounded (Icc a b) := by simpa only [preimage_cast_Icc] - using totallyBounded_preimage Rat.isUniformEmbedding_coe_real.toUniformInducing + using totallyBounded_preimage Rat.isUniformEmbedding_coe_real.isUniformInducing (totallyBounded_Icc (a : ℝ) b) end Rat diff --git a/Mathlib/Topology/Instances/RatLemmas.lean b/Mathlib/Topology/Instances/RatLemmas.lean index 69619e025121f..1f633201490d0 100644 --- a/Mathlib/Topology/Instances/RatLemmas.lean +++ b/Mathlib/Topology/Instances/RatLemmas.lean @@ -36,7 +36,7 @@ local notation "ℚ∞" => OnePoint ℚ namespace Rat -variable {p q : ℚ} {s t : Set ℚ} +variable {p : ℚ} {s : Set ℚ} theorem interior_compact_eq_empty (hs : IsCompact s) : interior s = ∅ := isDenseEmbedding_coe_real.toIsDenseInducing.interior_compact_eq_empty dense_irrational hs @@ -72,7 +72,7 @@ theorem not_secondCountableTopology_opc : ¬SecondCountableTopology ℚ∞ := by exact not_firstCountableTopology_opc inferInstance instance : TotallyDisconnectedSpace ℚ := by - clear p q s t + clear p s refine ⟨fun s hsu hs x hx y hy => ?_⟩; clear hsu by_contra! H : x ≠ y wlog hlt : x < y diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 5970438a142e8..6bac140e12e95 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -23,7 +23,7 @@ open TopologicalSpace Set Filter open Topology Filter variable {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] {f : α → β} -variable {s : Set β} {ι : Type*} {U : ι → Opens β} +variable {ι : Type*} {U : ι → Opens β} theorem Set.restrictPreimage_inducing (s : Set β) (h : Inducing f) : Inducing (s.restrictPreimage f) := by diff --git a/Mathlib/Topology/LocallyFinite.lean b/Mathlib/Topology/LocallyFinite.lean index f8fbbee18a97d..fef472c111a08 100644 --- a/Mathlib/Topology/LocallyFinite.lean +++ b/Mathlib/Topology/LocallyFinite.lean @@ -197,7 +197,7 @@ theorem LocallyFinite.sum_elim {g : ι' → Set X} (hf : LocallyFinite f) (hg : theorem locallyFinite_option {f : Option ι → Set X} : LocallyFinite f ↔ LocallyFinite (f ∘ some) := by - rw [← (Equiv.optionEquivSumPUnit.{_, 0} ι).symm.locallyFinite_comp_iff, locallyFinite_sum] + rw [← (Equiv.optionEquivSumPUnit.{0, _} ι).symm.locallyFinite_comp_iff, locallyFinite_sum] simp only [locallyFinite_of_finite, and_true] rfl diff --git a/Mathlib/Topology/Maps/Proper/Basic.lean b/Mathlib/Topology/Maps/Proper/Basic.lean index 4132dd33262f6..00bd0a3bd8677 100644 --- a/Mathlib/Topology/Maps/Proper/Basic.lean +++ b/Mathlib/Topology/Maps/Proper/Basic.lean @@ -174,12 +174,12 @@ lemma isProperMap_of_comp_of_t2 [T2Space Y] (hf : Continuous f) (hg : Continuous exact ⟨x, hx⟩ /-- A binary product of proper maps is proper. -/ -lemma IsProperMap.prod_map {g : Z → W} (hf : IsProperMap f) (hg : IsProperMap g) : +lemma IsProperMap.prodMap {g : Z → W} (hf : IsProperMap f) (hg : IsProperMap g) : IsProperMap (Prod.map f g) := by simp_rw [isProperMap_iff_ultrafilter] at hf hg ⊢ constructor -- Continuity is clear. - · exact hf.1.prod_map hg.1 + · exact hf.1.prodMap hg.1 -- Let `𝒰 : Ultrafilter (X × Z)`, and assume that `f × g` tends to some `(y, w) : Y × W` -- along `𝒰`. · intro 𝒰 ⟨y, w⟩ hyw @@ -197,6 +197,8 @@ lemma IsProperMap.prod_map {g : Z → W} (hf : IsProperMap f) (hg : IsProperMap rw [nhds_prod_eq, le_prod] exact ⟨hx, hz⟩ +@[deprecated (since := "2024-10-06")] alias IsProperMap.prod_map := IsProperMap.prodMap + /-- Any product of proper maps is proper. -/ lemma IsProperMap.pi_map {X Y : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, TopologicalSpace (Y i)] {f : (i : ι) → X i → Y i} (h : ∀ i, IsProperMap (f i)) : @@ -416,7 +418,7 @@ easier to use because it allows `Z` to live in any universe. -/ theorem IsProperMap.universally_closed (Z) [TopologicalSpace Z] (h : IsProperMap f) : IsClosedMap (Prod.map f id : X × Z → Y × Z) := -- `f × id` is proper as a product of proper maps, hence closed. - (h.prod_map isProperMap_id).isClosedMap + (h.prodMap isProperMap_id).isClosedMap /-- A map `f : X → Y` is proper if and only if it is continuous and the map `(Prod.map f id : X × Filter X → Y × Filter X)` is closed. This is stronger than diff --git a/Mathlib/Topology/MetricSpace/Algebra.lean b/Mathlib/Topology/MetricSpace/Algebra.lean index 724512392f31b..e185b9abdc3ba 100644 --- a/Mathlib/Topology/MetricSpace/Algebra.lean +++ b/Mathlib/Topology/MetricSpace/Algebra.lean @@ -6,6 +6,7 @@ Authors: Heather Macbeth import Mathlib.Topology.Algebra.MulAction import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Topology.MetricSpace.Lipschitz +import Mathlib.Topology.Algebra.SeparationQuotient /-! # Compatibility of algebraic operations with metric space structures @@ -213,5 +214,11 @@ instance Prod.instBoundedSMul {α β γ : Type*} [PseudoMetricSpace α] [PseudoM max_le ((dist_pair_smul _ _ _).trans <| mul_le_mul_of_nonneg_left (le_max_left _ _) dist_nonneg) ((dist_pair_smul _ _ _).trans <| mul_le_mul_of_nonneg_left (le_max_right _ _) dist_nonneg) +instance {α β : Type*} + [PseudoMetricSpace α] [PseudoMetricSpace β] [Zero α] [Zero β] [SMul α β] [BoundedSMul α β] : + BoundedSMul α (SeparationQuotient β) where + dist_smul_pair' _ := Quotient.ind₂ <| dist_smul_pair _ + dist_pair_smul' _ _ := Quotient.ind <| dist_pair_smul _ _ + -- We don't have the `SMul α γ → SMul β δ → SMul (α × β) (γ × δ)` instance, but if we did, then -- `BoundedSMul α γ → BoundedSMul β δ → BoundedSMul (α × β) (γ × δ)` would hold diff --git a/Mathlib/Topology/MetricSpace/Antilipschitz.lean b/Mathlib/Topology/MetricSpace/Antilipschitz.lean index ac87aa46237ea..67c88ebcb3216 100644 --- a/Mathlib/Topology/MetricSpace/Antilipschitz.lean +++ b/Mathlib/Topology/MetricSpace/Antilipschitz.lean @@ -143,19 +143,22 @@ theorem comap_uniformity_le (hf : AntilipschitzWith K f) : (𝓤 β).comap (Prod rw [mul_comm] exact ENNReal.mul_lt_of_lt_div hx -protected theorem uniformInducing (hf : AntilipschitzWith K f) (hfc : UniformContinuous f) : - UniformInducing f := +theorem isUniformInducing (hf : AntilipschitzWith K f) (hfc : UniformContinuous f) : + IsUniformInducing f := ⟨le_antisymm hf.comap_uniformity_le hfc.le_comap⟩ +@[deprecated (since := "2024-10-05")] +alias uniformInducing := isUniformInducing + lemma isUniformEmbedding {α β : Type*} [EMetricSpace α] [PseudoEMetricSpace β] {K : ℝ≥0} {f : α → β} (hf : AntilipschitzWith K f) (hfc : UniformContinuous f) : IsUniformEmbedding f := - ⟨hf.uniformInducing hfc, hf.injective⟩ + ⟨hf.isUniformInducing hfc, hf.injective⟩ @[deprecated (since := "2024-10-01")] alias uniformEmbedding := isUniformEmbedding theorem isComplete_range [CompleteSpace α] (hf : AntilipschitzWith K f) (hfc : UniformContinuous f) : IsComplete (range f) := - (hf.uniformInducing hfc).isComplete_range + (hf.isUniformInducing hfc).isComplete_range theorem isClosed_range {α β : Type*} [PseudoEMetricSpace α] [EMetricSpace β] [CompleteSpace α] {f : α → β} {K : ℝ≥0} (hf : AntilipschitzWith K f) (hfc : UniformContinuous f) : diff --git a/Mathlib/Topology/MetricSpace/Basic.lean b/Mathlib/Topology/MetricSpace/Basic.lean index 028c1ab133161..185ec4912deba 100644 --- a/Mathlib/Topology/MetricSpace/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Basic.lean @@ -36,7 +36,7 @@ theorem isUniformEmbedding_iff' [MetricSpace β] {f : γ → β} : IsUniformEmbedding f ↔ (∀ ε > 0, ∃ δ > 0, ∀ {a b : γ}, dist a b < δ → dist (f a) (f b) < ε) ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : γ}, dist (f a) (f b) < ε → dist a b < δ := by - rw [isUniformEmbedding_iff_uniformInducing, uniformInducing_iff, uniformContinuous_iff] + rw [isUniformEmbedding_iff_isUniformInducing, isUniformInducing_iff, uniformContinuous_iff] @[deprecated (since := "2024-10-01")] alias uniformEmbedding_iff' := isUniformEmbedding_iff' diff --git a/Mathlib/Topology/MetricSpace/Bilipschitz.lean b/Mathlib/Topology/MetricSpace/Bilipschitz.lean index 432fad174bbe3..0e69f60fc5dbc 100644 --- a/Mathlib/Topology/MetricSpace/Bilipschitz.lean +++ b/Mathlib/Topology/MetricSpace/Bilipschitz.lean @@ -55,7 +55,7 @@ instance : UniformSpace α := (inferInstance : UniformSpace β).comap f in order to avoid abuse of the definitional equality `α := β`. -/ lemma uniformity_eq_of_bilipschitz (hf₁ : AntilipschitzWith K₁ f) (hf₂ : LipschitzWith K₂ f) : 𝓤[(inferInstance : UniformSpace β).comap f] = 𝓤 α := - hf₁.uniformInducing hf₂.uniformContinuous |>.comap_uniformity + hf₁.isUniformInducing hf₂.uniformContinuous |>.comap_uniformity end Uniformity diff --git a/Mathlib/Topology/MetricSpace/Closeds.lean b/Mathlib/Topology/MetricSpace/Closeds.lean index cf5140b7380c9..2dd3375fbcd8e 100644 --- a/Mathlib/Topology/MetricSpace/Closeds.lean +++ b/Mathlib/Topology/MetricSpace/Closeds.lean @@ -281,7 +281,7 @@ theorem NonemptyCompacts.isClosed_in_closeds [CompleteSpace α] : from the same statement for closed subsets -/ instance NonemptyCompacts.completeSpace [CompleteSpace α] : CompleteSpace (NonemptyCompacts α) := (completeSpace_iff_isComplete_range - NonemptyCompacts.ToCloseds.isUniformEmbedding.toUniformInducing).2 <| + NonemptyCompacts.ToCloseds.isUniformEmbedding.isUniformInducing).2 <| NonemptyCompacts.isClosed_in_closeds.isComplete /-- In a compact space, the type of nonempty compact subsets is compact. This follows from diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index 0741095653d32..a2b542947d55f 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -370,12 +370,15 @@ theorem cancel_left {g : β →ᵈ γ} {f₁ f₂ : α →ᵈ β} (hg : Injectiv ⟨fun h => Dilation.ext fun x => hg <| by rw [← comp_apply, h, comp_apply], fun h => h ▸ rfl⟩ /-- A dilation from a metric space is a uniform inducing map -/ -protected theorem uniformInducing : UniformInducing (f : α → β) := - (antilipschitz f).uniformInducing (lipschitz f).uniformContinuous +theorem isUniformInducing : IsUniformInducing (f : α → β) := + (antilipschitz f).isUniformInducing (lipschitz f).uniformContinuous + +@[deprecated (since := "2024-10-05")] +alias uniformInducing := isUniformInducing theorem tendsto_nhds_iff {ι : Type*} {g : ι → α} {a : Filter ι} {b : α} : Filter.Tendsto g a (𝓝 b) ↔ Filter.Tendsto ((f : α → β) ∘ g) a (𝓝 (f b)) := - (Dilation.uniformInducing f).inducing.tendsto_nhds_iff + (Dilation.isUniformInducing f).inducing.tendsto_nhds_iff /-- A dilation is continuous. -/ theorem toContinuous : Continuous (f : α → β) := @@ -406,11 +409,11 @@ theorem mapsTo_emetric_closedBall (x : α) (r' : ℝ≥0∞) : theorem comp_continuousOn_iff {γ} [TopologicalSpace γ] {g : γ → α} {s : Set γ} : ContinuousOn ((f : α → β) ∘ g) s ↔ ContinuousOn g s := - (Dilation.uniformInducing f).inducing.continuousOn_iff.symm + (Dilation.isUniformInducing f).inducing.continuousOn_iff.symm theorem comp_continuous_iff {γ} [TopologicalSpace γ] {g : γ → α} : Continuous ((f : α → β) ∘ g) ↔ Continuous g := - (Dilation.uniformInducing f).inducing.continuous_iff.symm + (Dilation.isUniformInducing f).inducing.continuous_iff.symm end PseudoEmetricDilation diff --git a/Mathlib/Topology/MetricSpace/Gluing.lean b/Mathlib/Topology/MetricSpace/Gluing.lean index 53a7579e013fa..e3f3f1233ca43 100644 --- a/Mathlib/Topology/MetricSpace/Gluing.lean +++ b/Mathlib/Topology/MetricSpace/Gluing.lean @@ -243,8 +243,7 @@ private theorem Sum.mem_uniformity (s : Set ((X ⊕ Y) × (X ⊕ Y))) : · cases not_le_of_lt (lt_of_lt_of_le h (min_le_right _ _)) Sum.one_le_dist_inr_inl · exact hY (lt_of_lt_of_le h (le_trans (min_le_left _ _) (min_le_right _ _))) · rintro ⟨ε, ε0, H⟩ - constructor <;> rw [Filter.mem_sets, Filter.mem_map, mem_uniformity_dist] <;> - exact ⟨ε, ε0, fun h => H _ _ h⟩ + constructor <;> rw [Filter.mem_map, mem_uniformity_dist] <;> exact ⟨ε, ε0, fun h => H _ _ h⟩ /-- The distance on the disjoint union indeed defines a metric space. All the distance properties follow from our choice of the distance. The harder work is to show that the uniform structure @@ -432,7 +431,7 @@ protected theorem completeSpace [∀ i, CompleteSpace (E i)] : CompleteSpace (Σ set U := { p : (Σk, E k) × Σk, E k | dist p.1 p.2 < 1 } have hc : ∀ i, IsComplete (s i) := fun i => by simp only [s, ← range_sigmaMk] - exact (isometry_mk i).uniformInducing.isComplete_range + exact (isometry_mk i).isUniformInducing.isComplete_range have hd : ∀ (i j), ∀ x ∈ s i, ∀ y ∈ s j, (x, y) ∈ U → i = j := fun i j x hx y hy hxy => (Eq.symm hx).trans ((fst_eq_of_dist_lt_one _ _ hxy).trans hy) refine completeSpace_of_isComplete_univ ?_ diff --git a/Mathlib/Topology/MetricSpace/IsometricSMul.lean b/Mathlib/Topology/MetricSpace/IsometricSMul.lean index 1d8b4a236200b..8bbc25f34f0ab 100644 --- a/Mathlib/Topology/MetricSpace/IsometricSMul.lean +++ b/Mathlib/Topology/MetricSpace/IsometricSMul.lean @@ -405,19 +405,19 @@ instance ULift.isometricSMul' : IsometricSMul M (ULift X) := @[to_additive] instance {ι} {X : ι → Type*} [Fintype ι] [∀ i, SMul M (X i)] [∀ i, PseudoEMetricSpace (X i)] [∀ i, IsometricSMul M (X i)] : IsometricSMul M (∀ i, X i) := - ⟨fun c => isometry_dcomp (fun _ => (c • ·)) fun i => isometry_smul (X i) c⟩ + ⟨fun c => .piMap (fun _ => (c • ·)) fun i => isometry_smul (X i) c⟩ @[to_additive] instance Pi.isometricSMul' {ι} {M X : ι → Type*} [Fintype ι] [∀ i, SMul (M i) (X i)] [∀ i, PseudoEMetricSpace (X i)] [∀ i, IsometricSMul (M i) (X i)] : IsometricSMul (∀ i, M i) (∀ i, X i) := - ⟨fun c => isometry_dcomp (fun i => (c i • ·)) fun _ => isometry_smul _ _⟩ + ⟨fun c => .piMap (fun i => (c i • ·)) fun _ => isometry_smul _ _⟩ @[to_additive] instance Pi.isometricSMul'' {ι} {M : ι → Type*} [Fintype ι] [∀ i, Mul (M i)] [∀ i, PseudoEMetricSpace (M i)] [∀ i, IsometricSMul (M i)ᵐᵒᵖ (M i)] : IsometricSMul (∀ i, M i)ᵐᵒᵖ (∀ i, M i) := - ⟨fun c => isometry_dcomp (fun i (x : M i) => x * c.unop i) fun _ => isometry_mul_right _⟩ + ⟨fun c => .piMap (fun i (x : M i) => x * c.unop i) fun _ => isometry_mul_right _⟩ instance Additive.isometricVAdd : IsometricVAdd (Additive M) X := ⟨fun c => isometry_smul X (toMul c)⟩ diff --git a/Mathlib/Topology/MetricSpace/Isometry.lean b/Mathlib/Topology/MetricSpace/Isometry.lean index 8b05b577ab5bf..0640a68f619aa 100644 --- a/Mathlib/Topology/MetricSpace/Isometry.lean +++ b/Mathlib/Topology/MetricSpace/Isometry.lean @@ -84,10 +84,12 @@ theorem prod_map {δ} [PseudoEMetricSpace δ] {f : α → β} {g : γ → δ} (h (hg : Isometry g) : Isometry (Prod.map f g) := fun x y => by simp only [Prod.edist_eq, Prod.map_fst, hf.edist_eq, Prod.map_snd, hg.edist_eq] -theorem _root_.isometry_dcomp {ι} [Fintype ι] {α β : ι → Type*} [∀ i, PseudoEMetricSpace (α i)] +protected theorem piMap {ι} [Fintype ι] {α β : ι → Type*} [∀ i, PseudoEMetricSpace (α i)] [∀ i, PseudoEMetricSpace (β i)] (f : ∀ i, α i → β i) (hf : ∀ i, Isometry (f i)) : - Isometry (fun g : (i : ι) → α i => fun i => f i (g i)) := fun x y => by - simp only [edist_pi_def, (hf _).edist_eq] + Isometry (Pi.map f) := fun x y => by + simp only [edist_pi_def, (hf _).edist_eq, Pi.map_apply] + +@[deprecated (since := "2024-10-06")] alias _root_.isometry_dcomp := Isometry.piMap /-- The composition of isometries is an isometry. -/ theorem comp {g : β → γ} {f : α → β} (hg : Isometry g) (hf : Isometry f) : Isometry (g ∘ f) := @@ -98,12 +100,15 @@ protected theorem uniformContinuous (hf : Isometry f) : UniformContinuous f := hf.lipschitz.uniformContinuous /-- An isometry from a metric space is a uniform inducing map -/ -protected theorem uniformInducing (hf : Isometry f) : UniformInducing f := - hf.antilipschitz.uniformInducing hf.uniformContinuous +theorem isUniformInducing (hf : Isometry f) : IsUniformInducing f := + hf.antilipschitz.isUniformInducing hf.uniformContinuous + +@[deprecated (since := "2024-10-05")] +alias uniformInducing := isUniformInducing theorem tendsto_nhds_iff {ι : Type*} {f : α → β} {g : ι → α} {a : Filter ι} {b : α} (hf : Isometry f) : Filter.Tendsto g a (𝓝 b) ↔ Filter.Tendsto (f ∘ g) a (𝓝 (f b)) := - hf.uniformInducing.inducing.tendsto_nhds_iff + hf.isUniformInducing.inducing.tendsto_nhds_iff /-- An isometry is continuous. -/ protected theorem continuous (hf : Isometry f) : Continuous f := @@ -144,11 +149,11 @@ theorem _root_.isometry_subtype_coe {s : Set α} : Isometry ((↑) : s → α) : theorem comp_continuousOn_iff {γ} [TopologicalSpace γ] (hf : Isometry f) {g : γ → α} {s : Set γ} : ContinuousOn (f ∘ g) s ↔ ContinuousOn g s := - hf.uniformInducing.inducing.continuousOn_iff.symm + hf.isUniformInducing.inducing.continuousOn_iff.symm theorem comp_continuous_iff {γ} [TopologicalSpace γ] (hf : Isometry f) {g : γ → α} : Continuous (f ∘ g) ↔ Continuous g := - hf.uniformInducing.inducing.continuous_iff.symm + hf.isUniformInducing.inducing.continuous_iff.symm end PseudoEmetricIsometry @@ -470,7 +475,7 @@ theorem mul_apply (e₁ e₂ : α ≃ᵢ α) (x : α) : (e₁ * e₂) x = e₁ ( theorem completeSpace_iff (e : α ≃ᵢ β) : CompleteSpace α ↔ CompleteSpace β := by simp only [completeSpace_iff_isComplete_univ, ← e.range_eq_univ, ← image_univ, - isComplete_image_iff e.isometry.uniformInducing] + isComplete_image_iff e.isometry.isUniformInducing] protected theorem completeSpace [CompleteSpace β] (e : α ≃ᵢ β) : CompleteSpace α := e.completeSpace_iff.2 ‹_› diff --git a/Mathlib/Topology/MetricSpace/Kuratowski.lean b/Mathlib/Topology/MetricSpace/Kuratowski.lean index 8063585c24d8e..5a21609d09845 100644 --- a/Mathlib/Topology/MetricSpace/Kuratowski.lean +++ b/Mathlib/Topology/MetricSpace/Kuratowski.lean @@ -19,16 +19,16 @@ noncomputable section open Set Metric TopologicalSpace NNReal ENNReal lp Function -universe u v w +universe u -variable {α : Type u} {β : Type v} {γ : Type w} +variable {α : Type u} namespace KuratowskiEmbedding /-! ### Any separable metric space can be embedded isometrically in ℓ^∞(ℕ, ℝ) -/ -variable {f g : ℓ^∞(ℕ)} {n : ℕ} {C : ℝ} [MetricSpace α] (x : ℕ → α) (a b : α) +variable {n : ℕ} [MetricSpace α] (x : ℕ → α) (a : α) /-- A metric space can be embedded in `l^∞(ℝ)` via the distances to points in a fixed countable set, if this set is dense. This map is given in `kuratowskiEmbedding`, diff --git a/Mathlib/Topology/MetricSpace/PiNat.lean b/Mathlib/Topology/MetricSpace/PiNat.lean index 8c9627c5ed747..bfd470b9972da 100644 --- a/Mathlib/Topology/MetricSpace/PiNat.lean +++ b/Mathlib/Topology/MetricSpace/PiNat.lean @@ -51,7 +51,7 @@ noncomputable section open Topology TopologicalSpace Set Metric Filter Function -attribute [local simp] pow_le_pow_iff_right one_lt_two inv_le_inv zero_le_two zero_lt_two +attribute [local simp] pow_le_pow_iff_right one_lt_two inv_le_inv₀ zero_le_two zero_lt_two variable {E : ℕ → Type*} @@ -264,7 +264,7 @@ theorem dist_triangle_nonarch (x y z : ∀ n, E n) : dist x z ≤ max (dist x y) · simp rcases eq_or_ne y z with (rfl | hyz) · simp - simp only [dist_eq_of_ne, hxz, hxy, hyz, inv_le_inv, one_div, inv_pow, zero_lt_two, Ne, + simp only [dist_eq_of_ne, hxz, hxy, hyz, inv_le_inv₀, one_div, inv_pow, zero_lt_two, Ne, not_false_iff, le_max_iff, pow_le_pow_iff_right, one_lt_two, pow_pos, min_le_iff.1 (min_firstDiff_le x y z hxz)] @@ -294,7 +294,7 @@ theorem apply_eq_of_dist_lt {x y : ∀ n, E n} {n : ℕ} (h : dist x y < (1 / 2) rcases eq_or_ne x y with (rfl | hne) · rfl have : n < firstDiff x y := by - simpa [dist_eq_of_ne hne, inv_lt_inv, pow_lt_pow_iff_right, one_lt_two] using h + simpa [dist_eq_of_ne hne, inv_lt_inv₀, pow_lt_pow_iff_right, one_lt_two] using h exact apply_eq_of_lt_firstDiff (hi.trans_lt this) /-- A function to a pseudo-metric-space is `1`-Lipschitz if and only if points in the same cylinder diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index 66175bc090fb4..6db429aedccbf 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -146,7 +146,7 @@ theorem _root_.ClosedEmbedding.polishSpace [TopologicalSpace α] [TopologicalSpa letI : MetricSpace α := hf.toEmbedding.comapMetricSpace f haveI : SecondCountableTopology α := hf.toEmbedding.secondCountableTopology have : CompleteSpace α := by - rw [completeSpace_iff_isComplete_range hf.toEmbedding.to_isometry.uniformInducing] + rw [completeSpace_iff_isComplete_range hf.toEmbedding.to_isometry.isUniformInducing] exact hf.isClosed_range.isComplete infer_instance diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean b/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean index 346e0e97d3e23..550dee230352d 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean @@ -61,17 +61,20 @@ namespace Metric -- instantiate pseudometric space as a topology variable {x y z : α} {δ ε ε₁ ε₂ : ℝ} {s : Set α} -nonrec theorem uniformInducing_iff [PseudoMetricSpace β] {f : α → β} : - UniformInducing f ↔ UniformContinuous f ∧ +nonrec theorem isUniformInducing_iff [PseudoMetricSpace β] {f : α → β} : + IsUniformInducing f ↔ UniformContinuous f ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := - uniformInducing_iff'.trans <| Iff.rfl.and <| + isUniformInducing_iff'.trans <| Iff.rfl.and <| ((uniformity_basis_dist.comap _).le_basis_iff uniformity_basis_dist).trans <| by simp only [subset_def, Prod.forall, gt_iff_lt, preimage_setOf_eq, Prod.map_apply, mem_setOf] +@[deprecated (since := "2024-10-05")] +alias uniformInducing_iff := isUniformInducing_iff + nonrec theorem isUniformEmbedding_iff [PseudoMetricSpace β] {f : α → β} : IsUniformEmbedding f ↔ Function.Injective f ∧ UniformContinuous f ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := by - rw [isUniformEmbedding_iff, and_comm, uniformInducing_iff] + rw [isUniformEmbedding_iff, and_comm, isUniformInducing_iff] @[deprecated (since := "2024-10-01")] alias uniformEmbedding_iff := isUniformEmbedding_iff diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean b/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean index 3ad6637ff4565..764f2c2b7f1f1 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean @@ -44,10 +44,13 @@ def Inducing.comapPseudoMetricSpace {α β} [TopologicalSpace α] [m : PseudoMet /-- Pull back a pseudometric space structure by a uniform inducing map. This is a version of `PseudoMetricSpace.induced` useful in case if the domain already has a `UniformSpace` structure. -/ -def UniformInducing.comapPseudoMetricSpace {α β} [UniformSpace α] [m : PseudoMetricSpace β] - (f : α → β) (h : UniformInducing f) : PseudoMetricSpace α := +def IsUniformInducing.comapPseudoMetricSpace {α β} [UniformSpace α] [m : PseudoMetricSpace β] + (f : α → β) (h : IsUniformInducing f) : PseudoMetricSpace α := .replaceUniformity (.induced f m) h.comap_uniformity.symm +@[deprecated (since := "2024-10-08")] alias UniformInducing.comapPseudoMetricSpace := + IsUniformInducing.comapPseudoMetricSpace + instance Subtype.pseudoMetricSpace {p : α → Prop} : PseudoMetricSpace (Subtype p) := PseudoMetricSpace.induced Subtype.val ‹_› diff --git a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean index 37a5faee8383a..50b3c349f536c 100644 --- a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean +++ b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean @@ -128,7 +128,7 @@ theorem thickenedIndicatorAux_tendsto_indicator_closure {δseq : ℕ → ℝ} specialize δseq_lim ε ε_pos simp only [dist_zero_right, Real.norm_eq_abs, eventually_atTop] at δseq_lim rcases δseq_lim with ⟨N, hN⟩ - apply @tendsto_atTop_of_eventually_const _ _ _ _ _ _ _ N + apply tendsto_atTop_of_eventually_const (i₀ := N) intro n n_large have key : x ∉ thickening ε E := by simpa only [thickening, mem_setOf_eq, not_lt] using ε_lt.le refine le_antisymm ?_ bot_le diff --git a/Mathlib/Topology/MetricSpace/Ultra/TotallyDisconnected.lean b/Mathlib/Topology/MetricSpace/Ultra/TotallyDisconnected.lean deleted file mode 100644 index 161f861120e27..0000000000000 --- a/Mathlib/Topology/MetricSpace/Ultra/TotallyDisconnected.lean +++ /dev/null @@ -1,25 +0,0 @@ -/- -Copyright (c) 2024 Yakov Pechersky. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yakov Pechersky, David Loeffler --/ -import Mathlib.Topology.MetricSpace.Defs -import Mathlib.Topology.MetricSpace.Ultra.Basic - -/-! -# Ultrametric spaces are totally disconnected - -In a metric space with an ultrametric, the space is totally disconnected. - -## Tags - -ultrametric, nonarchimedean, totally disconnected --/ -open Metric IsUltrametricDist - -instance {X : Type*} [MetricSpace X] [IsUltrametricDist X] : TotallyDisconnectedSpace X := by - refine (totallyDisconnectedSpace_iff X).mpr (isTotallyDisconnected_of_isClopen_set fun x y h ↦ ?_) - obtain ⟨r, hr, hr'⟩ := exists_between (dist_pos.mpr h) - refine ⟨_, IsUltrametricDist.isClopen_ball x r, ?_, ?_⟩ - · simp only [mem_ball, dist_self, hr] - · simp only [mem_ball, dist_comm, not_lt, hr'.le] diff --git a/Mathlib/Topology/MetricSpace/Ultra/TotallySeparated.lean b/Mathlib/Topology/MetricSpace/Ultra/TotallySeparated.lean new file mode 100644 index 0000000000000..09edc0330ec3e --- /dev/null +++ b/Mathlib/Topology/MetricSpace/Ultra/TotallySeparated.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2024 Yakov Pechersky. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yakov Pechersky, David Loeffler +-/ +import Mathlib.Topology.MetricSpace.Defs +import Mathlib.Topology.MetricSpace.Ultra.Basic + +/-! +# Ultrametric spaces are totally separated + +In a metric space with an ultrametric, the space is totally separated, hence totally disconnected. + +## Tags + +ultrametric, nonarchimedean, totally separated, totally disconnected +-/ +open Metric IsUltrametricDist + +instance {X : Type*} [MetricSpace X] [IsUltrametricDist X] : TotallySeparatedSpace X := + totallySeparatedSpace_iff_exists_isClopen.mpr fun x y h ↦ by + obtain ⟨r, hr, hr'⟩ := exists_between (dist_pos.mpr h) + refine ⟨_, IsUltrametricDist.isClopen_ball x r, ?_, ?_⟩ + · simp only [mem_ball, dist_self, hr] + · simp only [Set.mem_compl, mem_ball, dist_comm, not_lt, hr'.le] + +example {X : Type*} [MetricSpace X] [IsUltrametricDist X] : TotallyDisconnectedSpace X := + inferInstance diff --git a/Mathlib/Topology/NoetherianSpace.lean b/Mathlib/Topology/NoetherianSpace.lean index 88d07a7493deb..9c2e714b18849 100644 --- a/Mathlib/Topology/NoetherianSpace.lean +++ b/Mathlib/Topology/NoetherianSpace.lean @@ -43,12 +43,10 @@ variable (α β : Type*) [TopologicalSpace α] [TopologicalSpace β] namespace TopologicalSpace /-- Type class for noetherian spaces. It is defined to be spaces whose open sets satisfies ACC. -/ -@[mk_iff] -class NoetherianSpace : Prop where - wellFounded_opens : WellFounded ((· > ·) : Opens α → Opens α → Prop) +abbrev NoetherianSpace : Prop := WellFoundedGT (Opens α) theorem noetherianSpace_iff_opens : NoetherianSpace α ↔ ∀ s : Opens α, IsCompact (s : Set α) := by - rw [noetherianSpace_iff, CompleteLattice.wellFounded_iff_isSupFiniteCompact, + rw [NoetherianSpace, CompleteLattice.wellFoundedGT_iff_isSupFiniteCompact, CompleteLattice.isSupFiniteCompact_iff_all_elements_compact] exact forall_congr' Opens.isCompactElement_iff @@ -78,12 +76,12 @@ variable (α) open List in theorem noetherianSpace_TFAE : TFAE [NoetherianSpace α, - WellFounded fun s t : Closeds α => s < t, + WellFoundedLT (Closeds α), ∀ s : Set α, IsCompact s, ∀ s : Opens α, IsCompact (s : Set α)] := by tfae_have 1 ↔ 2 := by - refine (noetherianSpace_iff α).trans (Opens.compl_bijective.2.wellFounded_iff ?_) - exact (@OrderIso.compl (Set α)).lt_iff_lt.symm + simp_rw [isWellFounded_iff] + exact Opens.compl_bijective.2.wellFounded_iff (@OrderIso.compl (Set α)).lt_iff_lt.symm tfae_have 1 ↔ 4 := noetherianSpace_iff_opens α tfae_have 1 → 3 := @NoetherianSpace.isCompact α _ tfae_have 3 → 4 := fun h s => h s @@ -94,9 +92,13 @@ variable {α} theorem noetherianSpace_iff_isCompact : NoetherianSpace α ↔ ∀ s : Set α, IsCompact s := (noetherianSpace_TFAE α).out 0 2 +instance [NoetherianSpace α] : WellFoundedLT (Closeds α) := + Iff.mp ((noetherianSpace_TFAE α).out 0 1) ‹_› + +@[deprecated (since := "2024-10-07")] theorem NoetherianSpace.wellFounded_closeds [NoetherianSpace α] : WellFounded fun s t : Closeds α => s < t := - Iff.mp ((noetherianSpace_TFAE α).out 0 1) ‹_› + wellFounded_lt instance {α} : NoetherianSpace (CofiniteTopology α) := by simp only [noetherianSpace_iff_isCompact, isCompact_iff_ultrafilter_le_nhds, @@ -153,7 +155,7 @@ instance (priority := 100) Finite.to_noetherianSpace [Finite α] : NoetherianSpa /-- In a Noetherian space, every closed set is a finite union of irreducible closed sets. -/ theorem NoetherianSpace.exists_finite_set_closeds_irreducible [NoetherianSpace α] (s : Closeds α) : ∃ S : Set (Closeds α), S.Finite ∧ (∀ t ∈ S, IsIrreducible (t : Set α)) ∧ s = sSup S := by - apply wellFounded_closeds.induction s; clear s + apply wellFounded_lt.induction s; clear s intro s H rcases eq_or_ne s ⊥ with rfl | h₀ · use ∅; simp diff --git a/Mathlib/Topology/Order/Hom/Basic.lean b/Mathlib/Topology/Order/Hom/Basic.lean index 94dd1b5acb371..b0e27200893c5 100644 --- a/Mathlib/Topology/Order/Hom/Basic.lean +++ b/Mathlib/Topology/Order/Hom/Basic.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Order.Hom.Basic -import Mathlib.Topology.ContinuousMap.Basic +import Mathlib.Topology.Basic +import Mathlib.Topology.ContinuousMap.Defs /-! # Continuous order homomorphisms diff --git a/Mathlib/Topology/Order/LawsonTopology.lean b/Mathlib/Topology/Order/LawsonTopology.lean index 61c08e2a6fe21..0a3d37bb82005 100644 --- a/Mathlib/Topology/Order/LawsonTopology.lean +++ b/Mathlib/Topology/Order/LawsonTopology.lean @@ -49,7 +49,7 @@ Lawson topology, preorder open Set TopologicalSpace -variable {α β : Type*} +variable {α : Type*} namespace Topology diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 7a476daeaf513..e2763d592bc24 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -543,7 +543,7 @@ namespace Subtype -- todo: add `OrderEmbedding.orderClosedTopology` instance {p : α → Prop} : OrderClosedTopology (Subtype p) := have this : Continuous fun p : Subtype p × Subtype p => ((p.fst : α), (p.snd : α)) := - continuous_subtype_val.prod_map continuous_subtype_val + continuous_subtype_val.prodMap continuous_subtype_val OrderClosedTopology.mk (t.isClosed_le'.preimage this) end Subtype diff --git a/Mathlib/Topology/Order/ScottTopology.lean b/Mathlib/Topology/Order/ScottTopology.lean index 6f26fd5939149..11867ff99b61a 100644 --- a/Mathlib/Topology/Order/ScottTopology.lean +++ b/Mathlib/Topology/Order/ScottTopology.lean @@ -107,7 +107,7 @@ lemma dirSupClosed_Iic (a : α) : DirSupClosed (Iic a) := fun _d _ _ _a ha ↦ ( end Preorder section CompleteLattice -variable [CompleteLattice α] {s t : Set α} +variable [CompleteLattice α] {s : Set α} lemma dirSupInacc_iff_forall_sSup : DirSupInacc s ↔ ∀ ⦃d⦄, d.Nonempty → DirectedOn (· ≤ ·) d → sSup d ∈ s → (d ∩ s).Nonempty := by @@ -124,7 +124,7 @@ namespace Topology /-! ### Scott-Hausdorff topology -/ section ScottHausdorff -variable [Preorder α] {s : Set α} +variable [Preorder α] /-- The Scott-Hausdorff topology. diff --git a/Mathlib/Topology/PartitionOfUnity.lean b/Mathlib/Topology/PartitionOfUnity.lean index e44e945ac68f0..661a124cbc7d4 100644 --- a/Mathlib/Topology/PartitionOfUnity.lean +++ b/Mathlib/Topology/PartitionOfUnity.lean @@ -139,7 +139,7 @@ variable {E : Type*} [AddCommMonoid E] [SMulWithZero ℝ E] [TopologicalSpace E] instance : FunLike (PartitionOfUnity ι X s) ι C(X, ℝ) where coe := toFun - coe_injective' := fun f g h ↦ by cases f; cases g; congr + coe_injective' f g h := by cases f; cases g; congr protected theorem locallyFinite : LocallyFinite fun i => support (f i) := f.locallyFinite' @@ -311,7 +311,7 @@ variable {s : Set X} (f : BumpCovering ι X s) instance : FunLike (BumpCovering ι X s) ι C(X, ℝ) where coe := toFun - coe_injective' := fun f g h ↦ by cases f; cases g; congr + coe_injective' f g h := by cases f; cases g; congr protected theorem locallyFinite : LocallyFinite fun i => support (f i) := f.locallyFinite' diff --git a/Mathlib/Topology/Perfect.lean b/Mathlib/Topology/Perfect.lean index 31622ec28846e..c390399234943 100644 --- a/Mathlib/Topology/Perfect.lean +++ b/Mathlib/Topology/Perfect.lean @@ -93,7 +93,7 @@ A topological space `X` is said to be perfect if its universe is a perfect set. Equivalently, this means that `𝓝[≠] x ≠ ⊥` for every point `x : X`. -/ @[mk_iff perfectSpace_def] -class PerfectSpace : Prop := +class PerfectSpace : Prop where univ_preperfect : Preperfect (Set.univ : Set α) theorem PerfectSpace.univ_perfect [PerfectSpace α] : Perfect (Set.univ : Set α) := diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 2f7f8cfdcd660..c2de323af42d9 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -1451,7 +1451,7 @@ theorem Set.InjOn.exists_mem_nhdsSet {X Y : Type*} [TopologicalSpace X] [Topolog · rcases loc x hx with ⟨u, hu, hf⟩ exact Filter.mem_of_superset (prod_mem_nhds hu hu) <| forall_prod_set.2 hf · suffices ∀ᶠ z in 𝓝 (x, y), f z.1 ≠ f z.2 from this.mono fun _ hne h ↦ absurd h hne - refine (fc x hx).prod_map' (fc y hy) <| isClosed_diagonal.isOpen_compl.mem_nhds ?_ + refine (fc x hx).prodMap' (fc y hy) <| isClosed_diagonal.isOpen_compl.mem_nhds ?_ exact inj.ne hx hy hne rw [← eventually_nhdsSet_iff_forall, sc.nhdsSet_prod_eq sc] at this exact eventually_prod_self_iff.1 this diff --git a/Mathlib/Topology/Sets/Order.lean b/Mathlib/Topology/Sets/Order.lean index f787206f3f3ad..1a139b07f9633 100644 --- a/Mathlib/Topology/Sets/Order.lean +++ b/Mathlib/Topology/Sets/Order.lean @@ -15,7 +15,7 @@ In this file we define the type of clopen upper sets. open Set TopologicalSpace -variable {α β : Type*} [TopologicalSpace α] [LE α] [TopologicalSpace β] [LE β] +variable {α : Type*} [TopologicalSpace α] [LE α] /-! ### Compact open sets -/ diff --git a/Mathlib/Topology/Specialization.lean b/Mathlib/Topology/Specialization.lean index 0cec4f5ed17b5..1c7ce45b4bb4b 100644 --- a/Mathlib/Topology/Specialization.lean +++ b/Mathlib/Topology/Specialization.lean @@ -63,7 +63,7 @@ instance instPartialOrder [T0Space α] : PartialOrder (Specialization α) := spe orders. -/ def map (f : C(α, β)) : Specialization α →o Specialization β where toFun := toEquiv ∘ f ∘ ofEquiv - monotone' := f.continuous.specialization_monotone + monotone' := (map_continuous f).specialization_monotone @[simp] lemma map_id : map (ContinuousMap.id α) = OrderHom.id := rfl @[simp] lemma map_comp (g : C(β, γ)) (f : C(α, β)) : map (g.comp f) = (map g).comp (map f) := rfl diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index cefda1d88701e..a2740461f605d 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -32,7 +32,7 @@ Furthermore, we say that `f` has compact support if the topological support of ` open Function Set Filter Topology -variable {X α α' β γ δ M E R : Type*} +variable {X α α' β γ δ M R : Type*} section One @@ -98,9 +98,9 @@ theorem mulTSupport_mul [TopologicalSpace X] [Monoid α] {f g : X → α} : section -variable [TopologicalSpace α] [TopologicalSpace α'] -variable [One β] [One γ] [One δ] -variable {g : β → γ} {f : α → β} {f₂ : α → γ} {m : β → γ → δ} {x : α} +variable [TopologicalSpace α] +variable [One β] +variable {f : α → β} {x : α} @[to_additive] theorem not_mem_mulTSupport_iff_eventuallyEq : x ∉ mulTSupport f ↔ f =ᶠ[𝓝 x] 1 := by @@ -119,7 +119,7 @@ end section CompactSupport variable [TopologicalSpace α] [TopologicalSpace α'] variable [One β] [One γ] [One δ] -variable {g : β → γ} {f : α → β} {f₂ : α → γ} {m : β → γ → δ} {x : α} +variable {g : β → γ} {f : α → β} {f₂ : α → γ} {m : β → γ → δ} /-- A function `f` *has compact multiplicative support* or is *compactly supported* if the closure of the multiplicative support of `f` is compact. In a T₂ space this is equivalent to `f` being equal @@ -294,7 +294,7 @@ section CompactSupport2 section Monoid variable [TopologicalSpace α] [MulOneClass β] -variable {f f' : α → β} {x : α} +variable {f f' : α → β} @[to_additive] theorem HasCompactMulSupport.mul (hf : HasCompactMulSupport f) (hf' : HasCompactMulSupport f') : @@ -320,7 +320,7 @@ end DivisionMonoid section SMulZeroClass variable [TopologicalSpace α] [Zero M] [SMulZeroClass R M] -variable {f : α → R} {f' : α → M} {x : α} +variable {f : α → R} {f' : α → M} theorem HasCompactSupport.smul_left (hf : HasCompactSupport f') : HasCompactSupport (f • f') := by rw [hasCompactSupport_iff_eventuallyEq] at hf ⊢ @@ -331,7 +331,7 @@ end SMulZeroClass section SMulWithZero variable [TopologicalSpace α] [Zero R] [Zero M] [SMulWithZero R M] -variable {f : α → R} {f' : α → M} {x : α} +variable {f : α → R} {f' : α → M} theorem HasCompactSupport.smul_right (hf : HasCompactSupport f) : HasCompactSupport (f • f') := by rw [hasCompactSupport_iff_eventuallyEq] at hf ⊢ @@ -345,7 +345,7 @@ end SMulWithZero section MulZeroClass variable [TopologicalSpace α] [MulZeroClass β] -variable {f f' : α → β} {x : α} +variable {f f' : α → β} theorem HasCompactSupport.mul_right (hf : HasCompactSupport f) : HasCompactSupport (f * f') := by rw [hasCompactSupport_iff_eventuallyEq] at hf ⊢ @@ -359,7 +359,7 @@ end MulZeroClass section OrderedAddGroup -variable {α β : Type*} [TopologicalSpace α] [AddGroup β] [Lattice β] +variable [TopologicalSpace α] [AddGroup β] [Lattice β] [AddLeftMono β] protected theorem HasCompactSupport.abs {f : α → β} (hf : HasCompactSupport f) : diff --git a/Mathlib/Topology/UniformSpace/AbstractCompletion.lean b/Mathlib/Topology/UniformSpace/AbstractCompletion.lean index 9b5cc359931bb..ce9024be56637 100644 --- a/Mathlib/Topology/UniformSpace/AbstractCompletion.lean +++ b/Mathlib/Topology/UniformSpace/AbstractCompletion.lean @@ -66,7 +66,7 @@ structure AbstractCompletion (α : Type u) [UniformSpace α] where /-- The completion is a T₀ space. -/ separation : T0Space space /-- The map into the completion is uniform-inducing. -/ - uniformInducing : UniformInducing coe + isUniformInducing : IsUniformInducing coe /-- The map into the completion has dense range. -/ dense : DenseRange coe @@ -81,18 +81,20 @@ local notation "hatα" => pkg.space local notation "ι" => pkg.coe +@[deprecated (since := "2024-10-08")] alias uniformInducing := isUniformInducing + /-- If `α` is complete, then it is an abstract completion of itself. -/ def ofComplete [T0Space α] [CompleteSpace α] : AbstractCompletion α := - mk α id inferInstance inferInstance inferInstance uniformInducing_id denseRange_id + mk α id inferInstance inferInstance inferInstance .id denseRange_id theorem closure_range : closure (range ι) = univ := pkg.dense.closure_range theorem isDenseInducing : IsDenseInducing ι := - ⟨pkg.uniformInducing.inducing, pkg.dense⟩ + ⟨pkg.isUniformInducing.inducing, pkg.dense⟩ theorem uniformContinuous_coe : UniformContinuous ι := - UniformInducing.uniformContinuous pkg.uniformInducing + IsUniformInducing.uniformContinuous pkg.isUniformInducing theorem continuous_coe : Continuous ι := pkg.uniformContinuous_coe.continuous @@ -130,7 +132,7 @@ variable [CompleteSpace β] theorem uniformContinuous_extend : UniformContinuous (pkg.extend f) := by by_cases hf : UniformContinuous f · rw [pkg.extend_def hf] - exact uniformContinuous_uniformly_extend pkg.uniformInducing pkg.dense hf + exact uniformContinuous_uniformly_extend pkg.isUniformInducing pkg.dense hf · change UniformContinuous (ite _ _ _) rw [if_neg hf] exact uniformContinuous_of_const fun a b => by congr 1 @@ -305,8 +307,8 @@ protected def prod : AbstractCompletion (α × β) where uniformStruct := inferInstance complete := inferInstance separation := inferInstance - uniformInducing := UniformInducing.prod pkg.uniformInducing pkg'.uniformInducing - dense := DenseRange.prod_map pkg.dense pkg'.dense + isUniformInducing := IsUniformInducing.prod pkg.isUniformInducing pkg'.isUniformInducing + dense := pkg.dense.prodMap pkg'.dense end Prod diff --git a/Mathlib/Topology/UniformSpace/Ascoli.lean b/Mathlib/Topology/UniformSpace/Ascoli.lean index 0b702ae3a22f7..351606b63c56c 100644 --- a/Mathlib/Topology/UniformSpace/Ascoli.lean +++ b/Mathlib/Topology/UniformSpace/Ascoli.lean @@ -20,14 +20,14 @@ a family of compact subsets of `X`, and `α` is a uniform space. convergence coincide on equicontinuous subsets. This is the key fact that makes equicontinuity important in functional analysis. We state various versions of it: - as an equality of `UniformSpace`s: `Equicontinuous.comap_uniformFun_eq` - - in terms of `UniformInducing`: `Equicontinuous.uniformInducing_uniformFun_iff_pi` + - in terms of `IsUniformInducing`: `Equicontinuous.isUniformInducing_uniformFun_iff_pi` - in terms of `Inducing`: `Equicontinuous.inducing_uniformFun_iff_pi` - in terms of convergence along a filter: `Equicontinuous.tendsto_uniformFun_iff_pi` * As a consequence, if `𝔖` is a family of compact subsets of `X`, then the uniform structures of uniform convergence on `𝔖` and pointwise convergence on `⋃₀ 𝔖` coincide on equicontinuous subsets. Again, we prove multiple variations: - as an equality of `UniformSpace`s: `EquicontinuousOn.comap_uniformOnFun_eq` - - in terms of `UniformInducing`: `EquicontinuousOn.uniformInducing_uniformOnFun_iff_pi'` + - in terms of `IsUniformInducing`: `EquicontinuousOn.isUniformInducing_uniformOnFun_iff_pi'` - in terms of `Inducing`: `EquicontinuousOn.inducing_uniformOnFun_iff_pi'` - in terms of convergence along a filter: `EquicontinuousOn.tendsto_uniformOnFun_iff_pi'` * The **Arzela-Ascoli theorem** follows from the previous fact and Tykhonov's theorem. @@ -70,8 +70,7 @@ equicontinuity, uniform convergence, ascoli open Set Filter Uniformity Topology Function UniformConvergence -variable {ι X Y α β : Type*} [TopologicalSpace X] [UniformSpace α] [UniformSpace β] -variable {F : ι → X → α} {G : ι → β → α} +variable {ι X α : Type*} [TopologicalSpace X] [UniformSpace α] {F : ι → X → α} /-- Let `X` be a compact topological space, `α` a uniform space, and `F : ι → (X → α)` an equicontinuous family. Then, the uniform structures of uniform convergence and pointwise @@ -80,7 +79,7 @@ convergence induce the same uniform structure on `ι`. In other words, pointwise convergence and uniform convergence coincide on an equicontinuous subset of `X → α`. -Consider using `Equicontinuous.uniformInducing_uniformFun_iff_pi` and +Consider using `Equicontinuous.isUniformInducing_uniformFun_iff_pi` and `Equicontinuous.inducing_uniformFun_iff_pi` instead, to avoid rewriting instances. -/ theorem Equicontinuous.comap_uniformFun_eq [CompactSpace X] (F_eqcont : Equicontinuous F) : (UniformFun.uniformSpace X α).comap F = @@ -131,15 +130,19 @@ convergence induce the same uniform structure on `ι`. In other words, pointwise convergence and uniform convergence coincide on an equicontinuous subset of `X → α`. -This is a version of `Equicontinuous.comap_uniformFun_eq` stated in terms of `UniformInducing` +This is a version of `Equicontinuous.comap_uniformFun_eq` stated in terms of `IsUniformInducing` for convenuence. -/ -lemma Equicontinuous.uniformInducing_uniformFun_iff_pi [UniformSpace ι] [CompactSpace X] +lemma Equicontinuous.isUniformInducing_uniformFun_iff_pi [UniformSpace ι] [CompactSpace X] (F_eqcont : Equicontinuous F) : - UniformInducing (UniformFun.ofFun ∘ F) ↔ UniformInducing F := by - rw [uniformInducing_iff_uniformSpace, uniformInducing_iff_uniformSpace, + IsUniformInducing (UniformFun.ofFun ∘ F) ↔ IsUniformInducing F := by + rw [isUniformInducing_iff_uniformSpace, isUniformInducing_iff_uniformSpace, ← F_eqcont.comap_uniformFun_eq] rfl +@[deprecated (since := "2024-10-05")] +alias Equicontinuous.uniformInducing_uniformFun_iff_pi := + Equicontinuous.isUniformInducing_uniformFun_iff_pi + /-- Let `X` be a compact topological space, `α` a uniform space, and `F : ι → (X → α)` an equicontinuous family. Then, the topologies of uniform convergence and pointwise convergence induce the same topology on `ι`. @@ -206,7 +209,7 @@ uniform structure on `ι`. In particular, pointwise convergence and compact convergence coincide on an equicontinuous subset of `X → α`. -Consider using `EquicontinuousOn.uniformInducing_uniformOnFun_iff_pi'` and +Consider using `EquicontinuousOn.isUniformInducing_uniformOnFun_iff_pi'` and `EquicontinuousOn.inducing_uniformOnFun_iff_pi'` instead to avoid rewriting instances, as well as their unprimed versions in case `𝔖` covers `X`. -/ theorem EquicontinuousOn.comap_uniformOnFun_eq {𝔖 : Set (Set X)} (𝔖_compact : ∀ K ∈ 𝔖, IsCompact K) @@ -247,37 +250,45 @@ uniform structure on `ι`. In particular, pointwise convergence and compact convergence coincide on an equicontinuous subset of `X → α`. -This is a version of `EquicontinuousOn.comap_uniformOnFun_eq` stated in terms of `UniformInducing` +This is a version of `EquicontinuousOn.comap_uniformOnFun_eq` stated in terms of `IsUniformInducing` for convenuence. -/ -lemma EquicontinuousOn.uniformInducing_uniformOnFun_iff_pi' [UniformSpace ι] +lemma EquicontinuousOn.isUniformInducing_uniformOnFun_iff_pi' [UniformSpace ι] {𝔖 : Set (Set X)} (𝔖_compact : ∀ K ∈ 𝔖, IsCompact K) (F_eqcont : ∀ K ∈ 𝔖, EquicontinuousOn F K) : - UniformInducing (UniformOnFun.ofFun 𝔖 ∘ F) ↔ - UniformInducing ((⋃₀ 𝔖).restrict ∘ F) := by - rw [uniformInducing_iff_uniformSpace, uniformInducing_iff_uniformSpace, + IsUniformInducing (UniformOnFun.ofFun 𝔖 ∘ F) ↔ + IsUniformInducing ((⋃₀ 𝔖).restrict ∘ F) := by + rw [isUniformInducing_iff_uniformSpace, isUniformInducing_iff_uniformSpace, ← EquicontinuousOn.comap_uniformOnFun_eq 𝔖_compact F_eqcont] rfl +@[deprecated (since := "2024-10-05")] +alias EquicontinuousOn.uniformInducing_uniformOnFun_iff_pi' := + EquicontinuousOn.isUniformInducing_uniformOnFun_iff_pi' + /-- Let `X` be a topological space, `𝔖` a covering of `X` by compact subsets, `α` a uniform space, and `F : ι → (X → α)` a family which is equicontinuous on each `K ∈ 𝔖`. Then, the uniform structures of uniform convergence on `𝔖` and pointwise convergence induce the same uniform structure on `ι`. -This is a specialization of `EquicontinuousOn.uniformInducing_uniformOnFun_iff_pi'` to +This is a specialization of `EquicontinuousOn.isUniformInducing_uniformOnFun_iff_pi'` to the case where `𝔖` covers `X`. -/ -lemma EquicontinuousOn.uniformInducing_uniformOnFun_iff_pi [UniformSpace ι] +lemma EquicontinuousOn.isUniformInducing_uniformOnFun_iff_pi [UniformSpace ι] {𝔖 : Set (Set X)} (𝔖_covers : ⋃₀ 𝔖 = univ) (𝔖_compact : ∀ K ∈ 𝔖, IsCompact K) (F_eqcont : ∀ K ∈ 𝔖, EquicontinuousOn F K) : - UniformInducing (UniformOnFun.ofFun 𝔖 ∘ F) ↔ - UniformInducing F := by + IsUniformInducing (UniformOnFun.ofFun 𝔖 ∘ F) ↔ + IsUniformInducing F := by rw [eq_univ_iff_forall] at 𝔖_covers -- This obviously follows from the previous lemma, we formalize it by going through the -- isomorphism of uniform spaces between `(⋃₀ 𝔖) → α` and `X → α`. let φ : ((⋃₀ 𝔖) → α) ≃ᵤ (X → α) := UniformEquiv.piCongrLeft (β := fun _ ↦ α) (Equiv.subtypeUnivEquiv 𝔖_covers) - rw [EquicontinuousOn.uniformInducing_uniformOnFun_iff_pi' 𝔖_compact F_eqcont, + rw [EquicontinuousOn.isUniformInducing_uniformOnFun_iff_pi' 𝔖_compact F_eqcont, show restrict (⋃₀ 𝔖) ∘ F = φ.symm ∘ F by rfl] - exact ⟨fun H ↦ φ.uniformInducing.comp H, fun H ↦ φ.symm.uniformInducing.comp H⟩ + exact ⟨fun H ↦ φ.isUniformInducing.comp H, fun H ↦ φ.symm.isUniformInducing.comp H⟩ + +@[deprecated (since := "2024-10-05")] +alias EquicontinuousOn.uniformInducing_uniformOnFun_iff_pi := + EquicontinuousOn.isUniformInducing_uniformOnFun_iff_pi /-- Let `X` be a topological space, `𝔖` a family of compact subsets of `X`, `α` a uniform space, and `F : ι → (X → α)` a family which is equicontinuous on each `K ∈ 𝔖`. Then, the topologies diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index d2dc3a7e09b3f..caab9ab1db8cb 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -1464,10 +1464,12 @@ theorem UniformContinuous.prod_mk_right {f : α × β → γ} (h : UniformContin UniformContinuous fun b => f (a, b) := h.comp (uniformContinuous_const.prod_mk uniformContinuous_id) -theorem UniformContinuous.prod_map [UniformSpace δ] {f : α → γ} {g : β → δ} +theorem UniformContinuous.prodMap [UniformSpace δ] {f : α → γ} {g : β → δ} (hf : UniformContinuous f) (hg : UniformContinuous g) : UniformContinuous (Prod.map f g) := (hf.comp uniformContinuous_fst).prod_mk (hg.comp uniformContinuous_snd) +@[deprecated (since := "2024-10-06")] alias UniformContinuous.prod_map := UniformContinuous.prodMap + theorem toTopologicalSpace_prod {α} {β} [u : UniformSpace α] [v : UniformSpace β] : @UniformSpace.toTopologicalSpace (α × β) instUniformSpaceProd = @instTopologicalSpaceProd α β u.toTopologicalSpace v.toTopologicalSpace := @@ -1483,7 +1485,7 @@ theorem uniformContinuous_inf_dom_left₂ {α β γ} {f : α → β → γ} {ua1 have ha := @UniformContinuous.inf_dom_left _ _ id ua1 ua2 ua1 (@uniformContinuous_id _ (id _)) have hb := @UniformContinuous.inf_dom_left _ _ id ub1 ub2 ub1 (@uniformContinuous_id _ (id _)) have h_unif_cont_id := - @UniformContinuous.prod_map _ _ _ _ (ua1 ⊓ ua2) (ub1 ⊓ ub2) ua1 ub1 _ _ ha hb + @UniformContinuous.prodMap _ _ _ _ (ua1 ⊓ ua2) (ub1 ⊓ ub2) ua1 ub1 _ _ ha hb exact @UniformContinuous.comp _ _ _ (id _) (id _) _ _ _ h h_unif_cont_id /-- A version of `UniformContinuous.inf_dom_right` for binary functions -/ @@ -1496,7 +1498,7 @@ theorem uniformContinuous_inf_dom_right₂ {α β γ} {f : α → β → γ} {ua have ha := @UniformContinuous.inf_dom_right _ _ id ua1 ua2 ua2 (@uniformContinuous_id _ (id _)) have hb := @UniformContinuous.inf_dom_right _ _ id ub1 ub2 ub2 (@uniformContinuous_id _ (id _)) have h_unif_cont_id := - @UniformContinuous.prod_map _ _ _ _ (ua1 ⊓ ua2) (ub1 ⊓ ub2) ua2 ub2 _ _ ha hb + @UniformContinuous.prodMap _ _ _ _ (ua1 ⊓ ua2) (ub1 ⊓ ub2) ua2 ub2 _ _ ha hb exact @UniformContinuous.comp _ _ _ (id _) (id _) _ _ _ h h_unif_cont_id /-- A version of `uniformContinuous_sInf_dom` for binary functions -/ @@ -1509,7 +1511,7 @@ theorem uniformContinuous_sInf_dom₂ {α β γ} {f : α → β → γ} {uas : S let _ : UniformSpace (α × β) := instUniformSpaceProd have ha := uniformContinuous_sInf_dom ha uniformContinuous_id have hb := uniformContinuous_sInf_dom hb uniformContinuous_id - have h_unif_cont_id := @UniformContinuous.prod_map _ _ _ _ (sInf uas) (sInf ubs) ua ub _ _ ha hb + have h_unif_cont_id := @UniformContinuous.prodMap _ _ _ _ (sInf uas) (sInf ubs) ua ub _ _ ha hb exact @UniformContinuous.comp _ _ _ (id _) (id _) _ _ _ hf h_unif_cont_id end Prod @@ -1545,7 +1547,7 @@ theorem UniformContinuous₂.comp {f : α → β → γ} {g : γ → δ} (hg : U theorem UniformContinuous₂.bicompl {f : α → β → γ} {ga : δ → α} {gb : δ' → β} (hf : UniformContinuous₂ f) (hga : UniformContinuous ga) (hgb : UniformContinuous gb) : UniformContinuous₂ (bicompl f ga gb) := - hf.uniformContinuous.comp (hga.prod_map hgb) + hf.uniformContinuous.comp (hga.prodMap hgb) end @@ -1737,7 +1739,7 @@ theorem continuousAt_iff'_left [TopologicalSpace β] {f : β → α} {b : β} : theorem continuousAt_iff_prod [TopologicalSpace β] {f : β → α} {b : β} : ContinuousAt f b ↔ Tendsto (fun x : β × β => (f x.1, f x.2)) (𝓝 (b, b)) (𝓤 α) := - ⟨fun H => le_trans (H.prod_map' H) (nhds_le_uniformity _), fun H => + ⟨fun H => le_trans (H.prodMap' H) (nhds_le_uniformity _), fun H => continuousAt_iff'_left.2 <| H.comp <| tendsto_id.prod_mk_nhds tendsto_const_nhds⟩ theorem continuousWithinAt_iff'_right [TopologicalSpace β] {f : β → α} {b : β} {s : Set β} : diff --git a/Mathlib/Topology/UniformSpace/Compact.lean b/Mathlib/Topology/UniformSpace/Compact.lean index 8f4de0a668e02..175d8d75ed946 100644 --- a/Mathlib/Topology/UniformSpace/Compact.lean +++ b/Mathlib/Topology/UniformSpace/Compact.lean @@ -151,7 +151,7 @@ theorem CompactSpace.uniformContinuous_of_continuous [CompactSpace α] {f : α (h : Continuous f) : UniformContinuous f := calc map (Prod.map f f) (𝓤 α) = map (Prod.map f f) (𝓝ˢ (diagonal α)) := by rw [nhdsSet_diagonal_eq_uniformity] - _ ≤ 𝓝ˢ (diagonal β) := (h.prod_map h).tendsto_nhdsSet mapsTo_prod_map_diagonal + _ ≤ 𝓝ˢ (diagonal β) := (h.prodMap h).tendsto_nhdsSet mapsTo_prod_map_diagonal _ ≤ 𝓤 β := nhdsSet_diagonal_le_uniformity /-- Heine-Cantor: a continuous function on a compact set of a uniform space is uniformly diff --git a/Mathlib/Topology/UniformSpace/CompactConvergence.lean b/Mathlib/Topology/UniformSpace/CompactConvergence.lean index c316d5316d094..f13aaeff613af 100644 --- a/Mathlib/Topology/UniformSpace/CompactConvergence.lean +++ b/Mathlib/Topology/UniformSpace/CompactConvergence.lean @@ -267,11 +267,14 @@ theorem uniformContinuous_comp (g : C(β, δ)) (hg : UniformContinuous g) : UniformOnFun.postcomp_uniformContinuous hg |>.comp isUniformEmbedding_toUniformOnFunIsCompact.uniformContinuous -theorem uniformInducing_comp (g : C(β, δ)) (hg : UniformInducing g) : - UniformInducing (ContinuousMap.comp g : C(α, β) → C(α, δ)) := - isUniformEmbedding_toUniformOnFunIsCompact.toUniformInducing.of_comp_iff.mp <| - UniformOnFun.postcomp_uniformInducing hg |>.comp - isUniformEmbedding_toUniformOnFunIsCompact.toUniformInducing +theorem isUniformInducing_comp (g : C(β, δ)) (hg : IsUniformInducing g) : + IsUniformInducing (ContinuousMap.comp g : C(α, β) → C(α, δ)) := + isUniformEmbedding_toUniformOnFunIsCompact.isUniformInducing.of_comp_iff.mp <| + UniformOnFun.postcomp_isUniformInducing hg |>.comp + isUniformEmbedding_toUniformOnFunIsCompact.isUniformInducing + +@[deprecated (since := "2024-10-05")] +alias uniformInducing_comp := isUniformInducing_comp theorem isUniformEmbedding_comp (g : C(β, δ)) (hg : IsUniformEmbedding g) : IsUniformEmbedding (ContinuousMap.comp g : C(α, β) → C(α, δ)) := @@ -379,7 +382,7 @@ Sufficient conditions on `α` to satisfy this condition are (weak) local compact lemma completeSpace_of_restrictGenTopology (h : RestrictGenTopology {K : Set α | IsCompact K}) : CompleteSpace C(α, β) := by rw [completeSpace_iff_isComplete_range - isUniformEmbedding_toUniformOnFunIsCompact.toUniformInducing, + isUniformEmbedding_toUniformOnFunIsCompact.isUniformInducing, range_toUniformOnFunIsCompact, ← completeSpace_coe_iff_isComplete] exact (UniformOnFun.isClosed_setOf_continuous h).completeSpace_coe diff --git a/Mathlib/Topology/UniformSpace/CompareReals.lean b/Mathlib/Topology/UniformSpace/CompareReals.lean index 7f1c135164444..c443d0a470a42 100644 --- a/Mathlib/Topology/UniformSpace/CompareReals.lean +++ b/Mathlib/Topology/UniformSpace/CompareReals.lean @@ -70,9 +70,9 @@ def rationalCauSeqPkg : @AbstractCompletion ℚ <| (@AbsoluteValue.abs ℚ _).un (uniformStruct := by infer_instance) (complete := by infer_instance) (separation := by infer_instance) - (uniformInducing := by + (isUniformInducing := by rw [Rat.uniformSpace_eq] - exact Rat.isUniformEmbedding_coe_real.toUniformInducing) + exact Rat.isUniformEmbedding_coe_real.isUniformInducing) (dense := Rat.isDenseEmbedding_coe_real.dense) namespace CompareReals diff --git a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean index 2fcb34197b931..7444f5a467be3 100644 --- a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean +++ b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean @@ -30,7 +30,7 @@ theorem IsComplete.isClosed [UniformSpace α] [T0Space α] {s : Set α} (h : IsC theorem IsUniformEmbedding.toClosedEmbedding [UniformSpace α] [UniformSpace β] [CompleteSpace α] [T0Space β] {f : α → β} (hf : IsUniformEmbedding f) : ClosedEmbedding f := - ⟨hf.embedding, hf.toUniformInducing.isComplete_range.isClosed⟩ + ⟨hf.embedding, hf.isUniformInducing.isComplete_range.isClosed⟩ @[deprecated (since := "2024-10-01")] alias UniformEmbedding.toClosedEmbedding := IsUniformEmbedding.toClosedEmbedding diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index 2ca3f3f72a834..4d1365e733ea3 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -143,7 +143,7 @@ theorem mem_uniformity' {s : Set (CauchyFilter α × CauchyFilter α)} : def pureCauchy (a : α) : CauchyFilter α := ⟨pure a, cauchy_pure⟩ -theorem uniformInducing_pureCauchy : UniformInducing (pureCauchy : α → CauchyFilter α) := +theorem isUniformInducing_pureCauchy : IsUniformInducing (pureCauchy : α → CauchyFilter α) := ⟨have : (preimage fun x : α × α => (pureCauchy x.fst, pureCauchy x.snd)) ∘ gen = id := funext fun s => Set.ext fun ⟨a₁, a₂⟩ => by simp [preimage, gen, pureCauchy, prod_principal_principal] @@ -154,8 +154,11 @@ theorem uniformInducing_pureCauchy : UniformInducing (pureCauchy : α → Cauchy _ = 𝓤 α := by simp [this] ⟩ +@[deprecated (since := "2024-10-05")] +alias uniformInducing_pureCauchy := isUniformInducing_pureCauchy + theorem isUniformEmbedding_pureCauchy : IsUniformEmbedding (pureCauchy : α → CauchyFilter α) := - { uniformInducing_pureCauchy with + { isUniformInducing_pureCauchy with inj := fun _a₁ _a₂ h => pure_injective <| Subtype.ext_iff_val.1 h } @[deprecated (since := "2024-10-01")] @@ -184,7 +187,7 @@ theorem denseRange_pureCauchy : DenseRange (pureCauchy : α → CauchyFilter α) exact ⟨_, this⟩ theorem isDenseInducing_pureCauchy : IsDenseInducing (pureCauchy : α → CauchyFilter α) := - uniformInducing_pureCauchy.isDenseInducing denseRange_pureCauchy + isUniformInducing_pureCauchy.isDenseInducing denseRange_pureCauchy theorem isDenseEmbedding_pureCauchy : IsDenseEmbedding (pureCauchy : α → CauchyFilter α) := isUniformEmbedding_pureCauchy.isDenseEmbedding denseRange_pureCauchy @@ -205,7 +208,7 @@ section -- set_option eqn_compiler.zeta true instance : CompleteSpace (CauchyFilter α) := - completeSpace_extension uniformInducing_pureCauchy denseRange_pureCauchy fun f hf => + completeSpace_extension isUniformInducing_pureCauchy denseRange_pureCauchy fun f hf => let f' : CauchyFilter α := ⟨f, hf⟩ have : map pureCauchy f ≤ (𝓤 <| CauchyFilter α).lift' (preimage (Prod.mk f')) := le_lift'.2 fun s hs => @@ -240,7 +243,7 @@ variable [T0Space β] theorem extend_pureCauchy {f : α → β} (hf : UniformContinuous f) (a : α) : extend f (pureCauchy a) = f a := by rw [extend, if_pos hf] - exact uniformly_extend_of_ind uniformInducing_pureCauchy denseRange_pureCauchy hf _ + exact uniformly_extend_of_ind isUniformInducing_pureCauchy denseRange_pureCauchy hf _ end T0Space @@ -249,7 +252,7 @@ variable [CompleteSpace β] theorem uniformContinuous_extend {f : α → β} : UniformContinuous (extend f) := by by_cases hf : UniformContinuous f · rw [extend, if_pos hf] - exact uniformContinuous_uniformly_extend uniformInducing_pureCauchy denseRange_pureCauchy hf + exact uniformContinuous_uniformly_extend isUniformInducing_pureCauchy denseRange_pureCauchy hf · rw [extend, if_neg hf] exact uniformContinuous_of_const fun a _b => by congr @@ -322,12 +325,15 @@ instance : Coe α (Completion α) := -- note [use has_coe_t] protected theorem coe_eq : ((↑) : α → Completion α) = SeparationQuotient.mk ∘ pureCauchy := rfl -theorem uniformInducing_coe : UniformInducing ((↑) : α → Completion α) := - SeparationQuotient.uniformInducing_mk.comp uniformInducing_pureCauchy +theorem isUniformInducing_coe : IsUniformInducing ((↑) : α → Completion α) := + SeparationQuotient.isUniformInducing_mk.comp isUniformInducing_pureCauchy + +@[deprecated (since := "2024-10-05")] +alias uniformInducing_coe := isUniformInducing_coe theorem comap_coe_eq_uniformity : ((𝓤 _).comap fun p : α × α => ((p.1 : Completion α), (p.2 : Completion α))) = 𝓤 α := - (uniformInducing_coe _).1 + (isUniformInducing_coe _).1 variable {α} @@ -344,7 +350,7 @@ def cPkg {α : Type*} [UniformSpace α] : AbstractCompletion α where uniformStruct := by infer_instance complete := by infer_instance separation := by infer_instance - uniformInducing := Completion.uniformInducing_coe α + isUniformInducing := Completion.isUniformInducing_coe α dense := Completion.denseRange_coe instance AbstractCompletion.inhabited : Inhabited (AbstractCompletion α) := @@ -375,7 +381,7 @@ theorem coe_injective [T0Space α] : Function.Injective ((↑) : α → Completi variable {α} theorem isDenseInducing_coe : IsDenseInducing ((↑) : α → Completion α) := - { (uniformInducing_coe α).inducing with dense := denseRange_coe } + { (isUniformInducing_coe α).inducing with dense := denseRange_coe } /-- The uniform bijection between a complete space and its uniform completion. -/ def UniformCompletion.completeEquivSelf [CompleteSpace α] [T0Space α] : Completion α ≃ᵤ α := @@ -394,12 +400,12 @@ alias denseEmbedding_coe := isDenseEmbedding_coe theorem denseRange_coe₂ : DenseRange fun x : α × β => ((x.1 : Completion α), (x.2 : Completion β)) := - denseRange_coe.prod_map denseRange_coe + denseRange_coe.prodMap denseRange_coe theorem denseRange_coe₃ : DenseRange fun x : α × β × γ => ((x.1 : Completion α), ((x.2.1 : Completion β), (x.2.2 : Completion γ))) := - denseRange_coe.prod_map denseRange_coe₂ + denseRange_coe.prodMap denseRange_coe₂ @[elab_as_elim] theorem induction_on {p : Completion α → Prop} (a : Completion α) (hp : IsClosed { a | p a }) diff --git a/Mathlib/Topology/UniformSpace/Equicontinuity.lean b/Mathlib/Topology/UniformSpace/Equicontinuity.lean index 5fdd4c239966d..807c30865e06d 100644 --- a/Mathlib/Topology/UniformSpace/Equicontinuity.lean +++ b/Mathlib/Topology/UniformSpace/Equicontinuity.lean @@ -704,56 +704,74 @@ theorem Filter.HasBasis.uniformEquicontinuousOn_iff {κ₁ κ₂ : Type*} {p₁ /-- Given `u : α → β` a uniform inducing map, a family `𝓕 : ι → X → α` is equicontinuous at a point `x₀ : X` iff the family `𝓕'`, obtained by composing each function of `𝓕` by `u`, is equicontinuous at `x₀`. -/ -theorem UniformInducing.equicontinuousAt_iff {F : ι → X → α} {x₀ : X} {u : α → β} - (hu : UniformInducing u) : EquicontinuousAt F x₀ ↔ EquicontinuousAt ((u ∘ ·) ∘ F) x₀ := by - have := (UniformFun.postcomp_uniformInducing (α := ι) hu).inducing +theorem IsUniformInducing.equicontinuousAt_iff {F : ι → X → α} {x₀ : X} {u : α → β} + (hu : IsUniformInducing u) : EquicontinuousAt F x₀ ↔ EquicontinuousAt ((u ∘ ·) ∘ F) x₀ := by + have := (UniformFun.postcomp_isUniformInducing (α := ι) hu).inducing rw [equicontinuousAt_iff_continuousAt, equicontinuousAt_iff_continuousAt, this.continuousAt_iff] rfl +@[deprecated (since := "2024-10-05")] +alias UniformInducing.equicontinuousAt_iff := IsUniformInducing.equicontinuousAt_iff + /-- Given `u : α → β` a uniform inducing map, a family `𝓕 : ι → X → α` is equicontinuous at a point `x₀ : X` within a subset `S : Set X` iff the family `𝓕'`, obtained by composing each function of `𝓕` by `u`, is equicontinuous at `x₀` within `S`. -/ -theorem UniformInducing.equicontinuousWithinAt_iff {F : ι → X → α} {S : Set X} {x₀ : X} {u : α → β} - (hu : UniformInducing u) : EquicontinuousWithinAt F S x₀ ↔ +lemma IsUniformInducing.equicontinuousWithinAt_iff {F : ι → X → α} {S : Set X} {x₀ : X} {u : α → β} + (hu : IsUniformInducing u) : EquicontinuousWithinAt F S x₀ ↔ EquicontinuousWithinAt ((u ∘ ·) ∘ F) S x₀ := by - have := (UniformFun.postcomp_uniformInducing (α := ι) hu).inducing + have := (UniformFun.postcomp_isUniformInducing (α := ι) hu).inducing simp only [equicontinuousWithinAt_iff_continuousWithinAt, this.continuousWithinAt_iff] rfl +@[deprecated (since := "2024-10-05")] +alias UniformInducing.equicontinuousWithinAt_iff := IsUniformInducing.equicontinuousWithinAt_iff + /-- Given `u : α → β` a uniform inducing map, a family `𝓕 : ι → X → α` is equicontinuous iff the family `𝓕'`, obtained by composing each function of `𝓕` by `u`, is equicontinuous. -/ -theorem UniformInducing.equicontinuous_iff {F : ι → X → α} {u : α → β} (hu : UniformInducing u) : +lemma IsUniformInducing.equicontinuous_iff {F : ι → X → α} {u : α → β} (hu : IsUniformInducing u) : Equicontinuous F ↔ Equicontinuous ((u ∘ ·) ∘ F) := by congrm ∀ x, ?_ rw [hu.equicontinuousAt_iff] +@[deprecated (since := "2024-10-05")] +alias UniformInducing.equicontinuous_iff := IsUniformInducing.equicontinuous_iff + /-- Given `u : α → β` a uniform inducing map, a family `𝓕 : ι → X → α` is equicontinuous on a subset `S : Set X` iff the family `𝓕'`, obtained by composing each function of `𝓕` by `u`, is equicontinuous on `S`. -/ -theorem UniformInducing.equicontinuousOn_iff {F : ι → X → α} {S : Set X} {u : α → β} - (hu : UniformInducing u) : EquicontinuousOn F S ↔ EquicontinuousOn ((u ∘ ·) ∘ F) S := by +theorem IsUniformInducing.equicontinuousOn_iff {F : ι → X → α} {S : Set X} {u : α → β} + (hu : IsUniformInducing u) : EquicontinuousOn F S ↔ EquicontinuousOn ((u ∘ ·) ∘ F) S := by congrm ∀ x ∈ S, ?_ rw [hu.equicontinuousWithinAt_iff] +@[deprecated (since := "2024-10-05")] +alias UniformInducing.equicontinuousOn_iff := IsUniformInducing.equicontinuousOn_iff + /-- Given `u : α → γ` a uniform inducing map, a family `𝓕 : ι → β → α` is uniformly equicontinuous iff the family `𝓕'`, obtained by composing each function of `𝓕` by `u`, is uniformly equicontinuous. -/ -theorem UniformInducing.uniformEquicontinuous_iff {F : ι → β → α} {u : α → γ} - (hu : UniformInducing u) : UniformEquicontinuous F ↔ UniformEquicontinuous ((u ∘ ·) ∘ F) := by - have := UniformFun.postcomp_uniformInducing (α := ι) hu +theorem IsUniformInducing.uniformEquicontinuous_iff {F : ι → β → α} {u : α → γ} + (hu : IsUniformInducing u) : UniformEquicontinuous F ↔ UniformEquicontinuous ((u ∘ ·) ∘ F) := by + have := UniformFun.postcomp_isUniformInducing (α := ι) hu simp only [uniformEquicontinuous_iff_uniformContinuous, this.uniformContinuous_iff] rfl +@[deprecated (since := "2024-10-05")] +alias UniformInducing.uniformEquicontinuous_iff := IsUniformInducing.uniformEquicontinuous_iff + /-- Given `u : α → γ` a uniform inducing map, a family `𝓕 : ι → β → α` is uniformly equicontinuous on a subset `S : Set β` iff the family `𝓕'`, obtained by composing each function of `𝓕` by `u`, is uniformly equicontinuous on `S`. -/ -theorem UniformInducing.uniformEquicontinuousOn_iff {F : ι → β → α} {S : Set β} {u : α → γ} - (hu : UniformInducing u) : +theorem IsUniformInducing.uniformEquicontinuousOn_iff {F : ι → β → α} {S : Set β} {u : α → γ} + (hu : IsUniformInducing u) : UniformEquicontinuousOn F S ↔ UniformEquicontinuousOn ((u ∘ ·) ∘ F) S := by - have := UniformFun.postcomp_uniformInducing (α := ι) hu + have := UniformFun.postcomp_isUniformInducing (α := ι) hu simp only [uniformEquicontinuousOn_iff_uniformContinuousOn, this.uniformContinuousOn_iff] rfl +@[deprecated (since := "2024-10-05")] +alias UniformInducing.uniformEquicontinuousOn_iff := IsUniformInducing.uniformEquicontinuousOn_iff + /-- If a set of functions is equicontinuous at some `x₀` within a set `S`, the same is true for its closure in *any* topology for which evaluation at any `x ∈ S ∪ {x₀}` is continuous. Since this will be applied to `DFunLike` types, we state it for any topological space with a map diff --git a/Mathlib/Topology/UniformSpace/Equiv.lean b/Mathlib/Topology/UniformSpace/Equiv.lean index a52b28256431d..6fb396f7b0004 100644 --- a/Mathlib/Topology/UniformSpace/Equiv.lean +++ b/Mathlib/Topology/UniformSpace/Equiv.lean @@ -49,11 +49,11 @@ theorem toEquiv_injective : Function.Injective (toEquiv : α ≃ᵤ β → α | ⟨e, h₁, h₂⟩, ⟨e', h₁', h₂'⟩, h => by simpa only [mk.injEq] instance : EquivLike (α ≃ᵤ β) α β where - coe := fun h => h.toEquiv - inv := fun h => h.toEquiv.symm - left_inv := fun h => h.left_inv - right_inv := fun h => h.right_inv - coe_injective' := fun _ _ H _ => toEquiv_injective <| DFunLike.ext' H + coe h := h.toEquiv + inv h := h.toEquiv.symm + left_inv h := h.left_inv + right_inv h := h.right_inv + coe_injective' _ _ H _ := toEquiv_injective <| DFunLike.ext' H @[simp] theorem uniformEquiv_mk_coe (a : Equiv α β) (b c) : (UniformEquiv.mk a b c : α → β) = a := @@ -196,14 +196,17 @@ theorem image_preimage (h : α ≃ᵤ β) (s : Set β) : h '' (h ⁻¹' s) = s : theorem preimage_image (h : α ≃ᵤ β) (s : Set α) : h ⁻¹' (h '' s) = s := h.toEquiv.preimage_image s -protected theorem uniformInducing (h : α ≃ᵤ β) : UniformInducing h := - uniformInducing_of_compose h.uniformContinuous h.symm.uniformContinuous <| by - simp only [symm_comp_self, uniformInducing_id] +theorem isUniformInducing (h : α ≃ᵤ β) : IsUniformInducing h := + IsUniformInducing.of_comp h.uniformContinuous h.symm.uniformContinuous <| by + simp only [symm_comp_self, IsUniformInducing.id] + +@[deprecated (since := "2024-10-05")] +alias uniformInducing := isUniformInducing theorem comap_eq (h : α ≃ᵤ β) : UniformSpace.comap h ‹_› = ‹_› := - h.uniformInducing.comap_uniformSpace + h.isUniformInducing.comap_uniformSpace -theorem isUniformEmbedding (h : α ≃ᵤ β) : IsUniformEmbedding h := ⟨h.uniformInducing, h.injective⟩ +lemma isUniformEmbedding (h : α ≃ᵤ β) : IsUniformEmbedding h := ⟨h.isUniformInducing, h.injective⟩ @[deprecated (since := "2024-10-01")] alias uniformEmbedding := isUniformEmbedding @@ -213,9 +216,9 @@ theorem completeSpace_iff (h : α ≃ᵤ β) : CompleteSpace α ↔ CompleteSpac /-- Uniform equiv given a uniform embedding. -/ noncomputable def ofIsUniformEmbedding (f : α → β) (hf : IsUniformEmbedding f) : α ≃ᵤ Set.range f where - uniformContinuous_toFun := hf.toUniformInducing.uniformContinuous.subtype_mk _ + uniformContinuous_toFun := hf.isUniformInducing.uniformContinuous.subtype_mk _ uniformContinuous_invFun := by - rw [hf.toUniformInducing.uniformContinuous_iff, Equiv.invFun_as_coe, + rw [hf.isUniformInducing.uniformContinuous_iff, Equiv.invFun_as_coe, Equiv.self_comp_ofInjective_symm] exact uniformContinuous_subtype_val toEquiv := Equiv.ofInjective f hf.inj @@ -331,7 +334,7 @@ def ulift : ULift.{v, u} α ≃ᵤ α := { Equiv.ulift with uniformContinuous_toFun := uniformContinuous_comap uniformContinuous_invFun := by - have hf : UniformInducing (@Equiv.ulift.{v, u} α).toFun := ⟨rfl⟩ + have hf : IsUniformInducing (@Equiv.ulift.{v, u} α).toFun := ⟨rfl⟩ simp_rw [hf.uniformContinuous_iff] exact uniformContinuous_id } @@ -370,8 +373,8 @@ end UniformEquiv /-- A uniform inducing equiv between uniform spaces is a uniform isomorphism. -/ -- @[simps] -- Porting note: removed, `simps?` produced no `simp` lemmas -def Equiv.toUniformEquivOfUniformInducing [UniformSpace α] [UniformSpace β] (f : α ≃ β) - (hf : UniformInducing f) : α ≃ᵤ β := +def Equiv.toUniformEquivOfIsUniformInducing [UniformSpace α] [UniformSpace β] (f : α ≃ β) + (hf : IsUniformInducing f) : α ≃ᵤ β := { f with uniformContinuous_toFun := hf.uniformContinuous uniformContinuous_invFun := hf.uniformContinuous_iff.2 <| by simpa using uniformContinuous_id } diff --git a/Mathlib/Topology/UniformSpace/Pi.lean b/Mathlib/Topology/UniformSpace/Pi.lean index 5c35d1206a3ad..783300913d8d4 100644 --- a/Mathlib/Topology/UniformSpace/Pi.lean +++ b/Mathlib/Topology/UniformSpace/Pi.lean @@ -122,8 +122,8 @@ protected theorem CompleteSpace.iInf {ι X : Type*} {u : ι → UniformSpace X} nontriviality X rcases ht with ⟨t, ht, hut⟩ -- The diagonal map `(X, ⨅ i, u i) → ∀ i, (X, u i)` is a uniform embedding. - have : @UniformInducing X (ι → X) (⨅ i, u i) (Pi.uniformSpace (U := u)) (const ι) := by - simp_rw [uniformInducing_iff, iInf_uniformity, Pi.uniformity, Filter.comap_iInf, + have : @IsUniformInducing X (ι → X) (⨅ i, u i) (Pi.uniformSpace (U := u)) (const ι) := by + simp_rw [isUniformInducing_iff, iInf_uniformity, Pi.uniformity, Filter.comap_iInf, Filter.comap_comap, comp_def, const, Prod.eta, comap_id'] -- Hence, it suffices to show that its range, the diagonal, is closed in `Π i, (X, u i)`. simp_rw [@completeSpace_iff_isComplete_range _ _ (_) (_) _ this, range_const_eq_diagonal, diff --git a/Mathlib/Topology/UniformSpace/UniformConvergence.lean b/Mathlib/Topology/UniformSpace/UniformConvergence.lean index 4d216ca50a3ca..3d5593d33b3b1 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergence.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergence.lean @@ -196,6 +196,12 @@ theorem TendstoUniformlyOn.congr {F' : ι → α → β} (hf : TendstoUniformlyO simp only [Set.EqOn] at hff' simp only [mem_prod_principal, hff', mem_setOf_eq] +lemma tendstoUniformly_congr {F F' : ι → α → β} {f : α → β} (hF : F =ᶠ[p] F') : + TendstoUniformly F f p ↔ TendstoUniformly F' f p := by + simp_rw [← tendstoUniformlyOn_univ] at * + have HF := EventuallyEq.exists_mem hF + exact ⟨fun h => h.congr (by aesop), fun h => h.congr (by simp_rw [eqOn_comm]; aesop)⟩ + theorem TendstoUniformlyOn.congr_right {g : α → β} (hf : TendstoUniformlyOn F f p s) (hfg : EqOn f g s) : TendstoUniformlyOn F g p s := fun u hu => by filter_upwards [hf u hu] with i hi a ha using hfg ha ▸ hi a ha diff --git a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean index 851e5d68e405e..00cd2deafccea 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean @@ -88,7 +88,7 @@ connection API to do most of the work. * `UniformOnFun.postcomp_uniformContinuous`: if `f : γ → β` is uniformly continuous, then `(fun g ↦ f ∘ g) : (α →ᵤ[𝔖] γ) → (α →ᵤ[𝔖] β)` is uniformly continuous. -* `UniformOnFun.postcomp_uniformInducing`: if `f : γ → β` is a uniform +* `UniformOnFun.postcomp_isUniformInducing`: if `f : γ → β` is a uniform inducing, then `(fun g ↦ f ∘ g) : (α →ᵤ[𝔖] γ) → (α →ᵤ[𝔖] β)` is a uniform inducing. * `UniformOnFun.precomp_uniformContinuous`: let `f : γ → α`, `𝔖 : Set (Set α)`, `𝔗 : Set (Set γ)`, and assume that `∀ T ∈ 𝔗, f '' T ∈ 𝔖`. Then, the function @@ -367,11 +367,14 @@ a uniform inducing function for the uniform structures of uniform convergence. More precisely, if `f : γ → β` is uniform inducing, then `(f ∘ ·) : (α →ᵤ γ) → (α →ᵤ β)` is uniform inducing. -/ -protected theorem postcomp_uniformInducing [UniformSpace γ] {f : γ → β} (hf : UniformInducing f) : - UniformInducing (ofFun ∘ (f ∘ ·) ∘ toFun : (α →ᵤ γ) → α →ᵤ β) := +lemma postcomp_isUniformInducing [UniformSpace γ] {f : γ → β} + (hf : IsUniformInducing f) : IsUniformInducing (ofFun ∘ (f ∘ ·) ∘ toFun : (α →ᵤ γ) → α →ᵤ β) := ⟨((UniformFun.hasBasis_uniformity _ _).comap _).eq_of_same_basis <| UniformFun.hasBasis_uniformity_of_basis _ _ (hf.basis_uniformity (𝓤 β).basis_sets)⟩ +@[deprecated (since := "2024-10-05")] +alias postcomp_uniformInducing := postcomp_isUniformInducing + /-- Post-composition by a uniform embedding is a uniform embedding for the uniform structures of uniform convergence. @@ -380,7 +383,7 @@ then `(f ∘ ·) : (α →ᵤ γ) → (α →ᵤ β)` is a uniform embedding. -/ protected theorem postcomp_isUniformEmbedding [UniformSpace γ] {f : γ → β} (hf : IsUniformEmbedding f) : IsUniformEmbedding (ofFun ∘ (f ∘ ·) ∘ toFun : (α →ᵤ γ) → α →ᵤ β) where - toUniformInducing := UniformFun.postcomp_uniformInducing hf.toUniformInducing + toIsUniformInducing := UniformFun.postcomp_isUniformInducing hf.isUniformInducing inj _ _ H := funext fun _ ↦ hf.inj (congrFun H _) @[deprecated (since := "2024-10-01")] @@ -392,7 +395,7 @@ alias postcomp_uniformEmbedding := UniformFun.postcomp_isUniformEmbedding protected theorem comap_eq {f : γ → β} : 𝒰(α, γ, ‹UniformSpace β›.comap f) = 𝒰(α, β, _).comap (f ∘ ·) := by letI : UniformSpace γ := .comap f ‹_› - exact (UniformFun.postcomp_uniformInducing (f := f) ⟨rfl⟩).comap_uniformSpace.symm + exact (UniformFun.postcomp_isUniformInducing (f := f) ⟨rfl⟩).comap_uniformSpace.symm /-- Post-composition by a uniformly continuous function is uniformly continuous on `α →ᵤ β`. @@ -465,7 +468,7 @@ protected def uniformEquivProdArrow [UniformSpace γ] : (α →ᵤ β × γ) ≃ -- But `uβ × uγ` is defined as `comap fst uβ ⊓ comap snd uγ`, so we just have to apply -- `UniformFun.inf_eq` and `UniformFun.comap_eq`, which leaves us to check -- that some square commutes. - Equiv.toUniformEquivOfUniformInducing (Equiv.arrowProdEquivProdArrow _ _ _) <| by + Equiv.toUniformEquivOfIsUniformInducing (Equiv.arrowProdEquivProdArrow _ _ _) <| by constructor change comap (Prod.map (Equiv.arrowProdEquivProdArrow _ _ _) (Equiv.arrowProdEquivProdArrow _ _ _)) @@ -490,10 +493,10 @@ protected def uniformEquivPiComm : UniformEquiv (α →ᵤ ∀ i, δ i) (∀ i, -- But `Π i, uδ i` is defined as `⨅ i, comap (eval i) (uδ i)`, so we just have to apply -- `UniformFun.iInf_eq` and `UniformFun.comap_eq`, which leaves us to check -- that some square commutes. - @Equiv.toUniformEquivOfUniformInducing + @Equiv.toUniformEquivOfIsUniformInducing _ _ 𝒰(α, ∀ i, δ i, Pi.uniformSpace δ) (@Pi.uniformSpace ι (fun i => α → δ i) fun i => 𝒰(α, δ i, _)) (Equiv.piComm _) <| by - refine @UniformInducing.mk ?_ ?_ ?_ ?_ ?_ ?_ + refine @IsUniformInducing.mk ?_ ?_ ?_ ?_ ?_ ?_ change comap (Prod.map Function.swap Function.swap) _ = _ rw [← uniformity_comap] congr @@ -865,8 +868,8 @@ uniform structures of `𝔖`-convergence. More precisely, if `f : γ → β` is a uniform inducing, then `(fun g ↦ f ∘ g) : (α →ᵤ[𝔖] γ) → (α →ᵤ[𝔖] β)` is a uniform inducing. -/ -protected theorem postcomp_uniformInducing [UniformSpace γ] {f : γ → β} (hf : UniformInducing f) : - UniformInducing (ofFun 𝔖 ∘ (f ∘ ·) ∘ toFun 𝔖) := by +lemma postcomp_isUniformInducing [UniformSpace γ] {f : γ → β} + (hf : IsUniformInducing f) : IsUniformInducing (ofFun 𝔖 ∘ (f ∘ ·) ∘ toFun 𝔖) := by -- This is a direct consequence of `UniformOnFun.comap_eq` constructor replace hf : (𝓤 β).comap (Prod.map f f) = _ := hf.comap_uniformity @@ -876,6 +879,9 @@ protected theorem postcomp_uniformInducing [UniformSpace γ] {f : γ → β} (hf rw [← UniformSpace.ext hf, UniformOnFun.comap_eq] rfl +@[deprecated (since := "2024-10-05")] +alias postcomp_uniformInducing := postcomp_isUniformInducing + /-- Post-composition by a uniform embedding is a uniform embedding for the uniform structures of `𝔖`-convergence. @@ -883,7 +889,7 @@ More precisely, if `f : γ → β` is a uniform embedding, then `(fun g ↦ f ∘ g) : (α →ᵤ[𝔖] γ) → (α →ᵤ[𝔖] β)` is a uniform embedding. -/ protected theorem postcomp_isUniformEmbedding [UniformSpace γ] {f : γ → β} (hf : IsUniformEmbedding f) : IsUniformEmbedding (ofFun 𝔖 ∘ (f ∘ ·) ∘ toFun 𝔖) where - toUniformInducing := UniformOnFun.postcomp_uniformInducing hf.toUniformInducing + toIsUniformInducing := UniformOnFun.postcomp_isUniformInducing hf.isUniformInducing inj _ _ H := funext fun _ ↦ hf.inj (congrFun H _) @[deprecated (since := "2024-10-01")] @@ -1012,9 +1018,8 @@ protected def uniformEquivProdArrow [UniformSpace γ] : -- which leaves us to check that some square commutes. -- We could also deduce this from `UniformFun.uniformEquivProdArrow`, -- but it turns out to be more annoying. - ((UniformOnFun.ofFun 𝔖).symm.trans <| - (Equiv.arrowProdEquivProdArrow _ _ _).trans <| - (UniformOnFun.ofFun 𝔖).prodCongr (UniformOnFun.ofFun 𝔖)).toUniformEquivOfUniformInducing <| by + ((UniformOnFun.ofFun 𝔖).symm.trans <| (Equiv.arrowProdEquivProdArrow _ _ _).trans <| + (UniformOnFun.ofFun 𝔖).prodCongr (UniformOnFun.ofFun 𝔖)).toUniformEquivOfIsUniformInducing <| by constructor rw [uniformity_prod, comap_inf, comap_comap, comap_comap] have H := @UniformOnFun.inf_eq α (β × γ) 𝔖 @@ -1038,7 +1043,7 @@ protected def uniformEquivPiComm : (α →ᵤ[𝔖] ((i : ι) → δ i)) ≃ᵤ -- which leaves us to check that some square commutes. -- We could also deduce this from `UniformFun.uniformEquivPiComm`, but it turns out -- to be more annoying. - @Equiv.toUniformEquivOfUniformInducing (α →ᵤ[𝔖] ((i : ι) → δ i)) ((i : ι) → α →ᵤ[𝔖] δ i) + @Equiv.toUniformEquivOfIsUniformInducing (α →ᵤ[𝔖] ((i : ι) → δ i)) ((i : ι) → α →ᵤ[𝔖] δ i) _ _ (Equiv.piComm _) <| by constructor change comap (Prod.map Function.swap Function.swap) _ = _ @@ -1124,3 +1129,41 @@ instance {α β : Type*} [UniformSpace β] [CompleteSpace β] : CompleteSpace ( (UniformOnFun.uniformEquivUniformFun β {univ} (mem_singleton _)).completeSpace_iff.1 inferInstance end UniformFun + +section UniformComposition + +variable {α β γ ι : Type*} [UniformSpace β] [UniformSpace γ] {p : Filter ι} + +/-- Composing on the left by a uniformly continuous function preserves uniform convergence -/ +theorem UniformContinuousOn.comp_tendstoUniformly (s : Set β) (F : ι → α → β) (f : α → β) + (hF : ∀ i x, F i x ∈ s) (hf : ∀ x, f x ∈ s) + {g : β → γ} (hg : UniformContinuousOn g s) (h : TendstoUniformly F f p) : + TendstoUniformly (fun i x => g (F i x)) (fun x => g (f x)) p := by + rw [uniformContinuousOn_iff_restrict] at hg + lift F to ι → α → s using hF with F' hF' + lift f to α → s using hf with f' hf' + rw [tendstoUniformly_iff_tendsto] at h + have : Tendsto (fun q : ι × α ↦ (f' q.2, (F' q.1 q.2))) (p ×ˢ ⊤) (𝓤 s) := + h.of_tendsto_comp isUniformEmbedding_subtype_val.comap_uniformity.le + apply UniformContinuous.comp_tendstoUniformly hg ?_ + rwa [← tendstoUniformly_iff_tendsto] at this + +theorem UniformContinuousOn.comp_tendstoUniformly_eventually (s : Set β) (F : ι → α → β) (f : α → β) + (hF : ∀ᶠ i in p, ∀ x, F i x ∈ s) (hf : ∀ x, f x ∈ s) + {g : β → γ} (hg : UniformContinuousOn g s) (h : TendstoUniformly F f p) : + TendstoUniformly (fun i => fun x => g (F i x)) (fun x => g (f x)) p := by + classical + rw [eventually_iff_exists_mem] at hF + obtain ⟨s', hs', hs⟩ := hF + let F' : ι → α → β := fun (i : ι) x => if i ∈ s' then F i x else f x + have hF : F =ᶠ[p] F' := by + rw [eventuallyEq_iff_exists_mem] + refine ⟨s', hs', fun y hy => by aesop⟩ + have h' : TendstoUniformly F' f p := by + rwa [tendstoUniformly_congr hF] at h + apply (tendstoUniformly_congr _).mpr + (UniformContinuousOn.comp_tendstoUniformly s F' f (by aesop) hf hg h') + rw [eventuallyEq_iff_exists_mem] + refine ⟨s', hs', fun i hi => by aesop⟩ + +end UniformComposition diff --git a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean index adad8bcab4c26..d70f7fb955c8f 100644 --- a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean +++ b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean @@ -20,6 +20,7 @@ section universe u v w variable {α : Type u} {β : Type v} {γ : Type w} [UniformSpace α] [UniformSpace β] [UniformSpace γ] + {f : α → β} /-! ### Uniform inducing maps @@ -29,100 +30,163 @@ variable {α : Type u} {β : Type v} {γ : Type w} [UniformSpace α] [UniformSpa on `α` is the pullback of the uniformity filter on `β` under `Prod.map f f`. If `α` is a separated space, then this implies that `f` is injective, hence it is a `IsUniformEmbedding`. -/ @[mk_iff] -structure UniformInducing (f : α → β) : Prop where +structure IsUniformInducing (f : α → β) : Prop where /-- The uniformity filter on the domain is the pullback of the uniformity filter on the codomain under `Prod.map f f`. -/ comap_uniformity : comap (fun x : α × α => (f x.1, f x.2)) (𝓤 β) = 𝓤 α -lemma uniformInducing_iff_uniformSpace {f : α → β} : - UniformInducing f ↔ ‹UniformSpace β›.comap f = ‹UniformSpace α› := by - rw [uniformInducing_iff, UniformSpace.ext_iff, Filter.ext_iff] +@[deprecated (since := "2024-10-08")] alias UniformInducing := IsUniformInducing + +lemma isUniformInducing_iff_uniformSpace {f : α → β} : + IsUniformInducing f ↔ ‹UniformSpace β›.comap f = ‹UniformSpace α› := by + rw [isUniformInducing_iff, UniformSpace.ext_iff, Filter.ext_iff] rfl -protected alias ⟨UniformInducing.comap_uniformSpace, _⟩ := uniformInducing_iff_uniformSpace +@[deprecated (since := "2024-10-05")] +alias uniformInducing_iff_uniformSpace := isUniformInducing_iff_uniformSpace + +protected alias ⟨IsUniformInducing.comap_uniformSpace, _⟩ := isUniformInducing_iff_uniformSpace + +@[deprecated (since := "2024-10-08")] alias UniformInducing.comap_uniformSpace := + IsUniformInducing.comap_uniformSpace + +lemma isUniformInducing_iff' {f : α → β} : + IsUniformInducing f ↔ UniformContinuous f ∧ comap (Prod.map f f) (𝓤 β) ≤ 𝓤 α := by + rw [isUniformInducing_iff, UniformContinuous, tendsto_iff_comap, le_antisymm_iff, and_comm]; rfl -lemma uniformInducing_iff' {f : α → β} : - UniformInducing f ↔ UniformContinuous f ∧ comap (Prod.map f f) (𝓤 β) ≤ 𝓤 α := by - rw [uniformInducing_iff, UniformContinuous, tendsto_iff_comap, le_antisymm_iff, and_comm]; rfl +@[deprecated (since := "2024-10-05")] +alias uniformInducing_iff' := isUniformInducing_iff' -protected lemma Filter.HasBasis.uniformInducing_iff {ι ι'} {p : ι → Prop} {p' : ι' → Prop} {s s'} +protected lemma Filter.HasBasis.isUniformInducing_iff {ι ι'} {p : ι → Prop} {p' : ι' → Prop} {s s'} (h : (𝓤 α).HasBasis p s) (h' : (𝓤 β).HasBasis p' s') {f : α → β} : - UniformInducing f ↔ + IsUniformInducing f ↔ (∀ i, p' i → ∃ j, p j ∧ ∀ x y, (x, y) ∈ s j → (f x, f y) ∈ s' i) ∧ (∀ j, p j → ∃ i, p' i ∧ ∀ x y, (f x, f y) ∈ s' i → (x, y) ∈ s j) := by - simp [uniformInducing_iff', h.uniformContinuous_iff h', (h'.comap _).le_basis_iff h, subset_def] + simp [isUniformInducing_iff', h.uniformContinuous_iff h', (h'.comap _).le_basis_iff h, subset_def] -theorem UniformInducing.mk' {f : α → β} - (h : ∀ s, s ∈ 𝓤 α ↔ ∃ t ∈ 𝓤 β, ∀ x y : α, (f x, f y) ∈ t → (x, y) ∈ s) : UniformInducing f := +@[deprecated (since := "2024-10-05")] +alias Filter.HasBasis.uniformInducing_iff := Filter.HasBasis.isUniformInducing_iff + +theorem IsUniformInducing.mk' {f : α → β} + (h : ∀ s, s ∈ 𝓤 α ↔ ∃ t ∈ 𝓤 β, ∀ x y : α, (f x, f y) ∈ t → (x, y) ∈ s) : IsUniformInducing f := ⟨by simp [eq_comm, Filter.ext_iff, subset_def, h]⟩ -theorem uniformInducing_id : UniformInducing (@id α) := +@[deprecated (since := "2024-10-05")] +alias UniformInducing.mk' := IsUniformInducing.mk' + +theorem IsUniformInducing.id : IsUniformInducing (@id α) := ⟨by rw [← Prod.map_def, Prod.map_id, comap_id]⟩ -theorem UniformInducing.comp {g : β → γ} (hg : UniformInducing g) {f : α → β} - (hf : UniformInducing f) : UniformInducing (g ∘ f) := +@[deprecated (since := "2024-10-05")] +alias uniformInducing_id := IsUniformInducing.id + +theorem IsUniformInducing.comp {g : β → γ} (hg : IsUniformInducing g) {f : α → β} + (hf : IsUniformInducing f) : IsUniformInducing (g ∘ f) := ⟨by rw [← hf.1, ← hg.1, comap_comap]; rfl⟩ -theorem UniformInducing.of_comp_iff {g : β → γ} (hg : UniformInducing g) {f : α → β} : - UniformInducing (g ∘ f) ↔ UniformInducing f := by +@[deprecated (since := "2024-10-05")] +alias UniformInducing.comp := IsUniformInducing.comp + +theorem IsUniformInducing.of_comp_iff {g : β → γ} (hg : IsUniformInducing g) {f : α → β} : + IsUniformInducing (g ∘ f) ↔ IsUniformInducing f := by refine ⟨fun h ↦ ?_, hg.comp⟩ - rw [uniformInducing_iff, ← hg.comap_uniformity, comap_comap, ← h.comap_uniformity, + rw [isUniformInducing_iff, ← hg.comap_uniformity, comap_comap, ← h.comap_uniformity, Function.comp_def, Function.comp_def] -theorem UniformInducing.basis_uniformity {f : α → β} (hf : UniformInducing f) {ι : Sort*} +@[deprecated (since := "2024-10-05")] +alias UniformInducing.of_comp_iff := IsUniformInducing.of_comp_iff + +theorem IsUniformInducing.basis_uniformity {f : α → β} (hf : IsUniformInducing f) {ι : Sort*} {p : ι → Prop} {s : ι → Set (β × β)} (H : (𝓤 β).HasBasis p s) : (𝓤 α).HasBasis p fun i => Prod.map f f ⁻¹' s i := hf.1 ▸ H.comap _ -theorem UniformInducing.cauchy_map_iff {f : α → β} (hf : UniformInducing f) {F : Filter α} : +@[deprecated (since := "2024-10-05")] +alias UniformInducing.basis_uniformity := IsUniformInducing.basis_uniformity + +theorem IsUniformInducing.cauchy_map_iff {f : α → β} (hf : IsUniformInducing f) {F : Filter α} : Cauchy (map f F) ↔ Cauchy F := by simp only [Cauchy, map_neBot_iff, prod_map_map_eq, map_le_iff_le_comap, ← hf.comap_uniformity] -theorem uniformInducing_of_compose {f : α → β} {g : β → γ} (hf : UniformContinuous f) - (hg : UniformContinuous g) (hgf : UniformInducing (g ∘ f)) : UniformInducing f := by +@[deprecated (since := "2024-10-05")] +alias UniformInducing.cauchy_map_iff := IsUniformInducing.cauchy_map_iff + +theorem IsUniformInducing.of_comp {f : α → β} {g : β → γ} (hf : UniformContinuous f) + (hg : UniformContinuous g) (hgf : IsUniformInducing (g ∘ f)) : IsUniformInducing f := by refine ⟨le_antisymm ?_ hf.le_comap⟩ rw [← hgf.1, ← Prod.map_def, ← Prod.map_def, ← Prod.map_comp_map f f g g, ← comap_comap] exact comap_mono hg.le_comap -theorem UniformInducing.uniformContinuous {f : α → β} (hf : UniformInducing f) : - UniformContinuous f := (uniformInducing_iff'.1 hf).1 +@[deprecated (since := "2024-10-05")] +alias uniformInducing_of_compose := IsUniformInducing.of_comp + +theorem IsUniformInducing.uniformContinuous {f : α → β} (hf : IsUniformInducing f) : + UniformContinuous f := (isUniformInducing_iff'.1 hf).1 -theorem UniformInducing.uniformContinuous_iff {f : α → β} {g : β → γ} (hg : UniformInducing g) : +@[deprecated (since := "2024-10-05")] +alias UniformInducing.uniformContinuous := IsUniformInducing.uniformContinuous + +theorem IsUniformInducing.uniformContinuous_iff {f : α → β} {g : β → γ} (hg : IsUniformInducing g) : UniformContinuous f ↔ UniformContinuous (g ∘ f) := by dsimp only [UniformContinuous, Tendsto] simp only [← hg.comap_uniformity, ← map_le_iff_le_comap, Filter.map_map, Function.comp_def] -protected theorem UniformInducing.uniformInducing_comp_iff {f : α → β} {g : β → γ} - (hg : UniformInducing g) : UniformInducing (g ∘ f) ↔ UniformInducing f := by - simp only [uniformInducing_iff, ← hg.comap_uniformity, comap_comap, Function.comp_def] +@[deprecated (since := "2024-10-05")] +alias UniformInducing.uniformContinuous_iff := IsUniformInducing.uniformContinuous_iff + +protected theorem IsUniformInducing.isUniformInducing_comp_iff {f : α → β} {g : β → γ} + (hg : IsUniformInducing g) : IsUniformInducing (g ∘ f) ↔ IsUniformInducing f := by + simp only [isUniformInducing_iff, ← hg.comap_uniformity, comap_comap, Function.comp_def] + +@[deprecated (since := "2024-10-05")] +alias UniformInducing.uniformInducing_comp_iff := IsUniformInducing.isUniformInducing_comp_iff -theorem UniformInducing.uniformContinuousOn_iff {f : α → β} {g : β → γ} {S : Set α} - (hg : UniformInducing g) : +theorem IsUniformInducing.uniformContinuousOn_iff {f : α → β} {g : β → γ} {S : Set α} + (hg : IsUniformInducing g) : UniformContinuousOn f S ↔ UniformContinuousOn (g ∘ f) S := by dsimp only [UniformContinuousOn, Tendsto] rw [← hg.comap_uniformity, ← map_le_iff_le_comap, Filter.map_map, comp_def, comp_def] -theorem UniformInducing.inducing {f : α → β} (h : UniformInducing f) : Inducing f := by +@[deprecated (since := "2024-10-05")] +alias UniformInducing.uniformContinuousOn_iff := IsUniformInducing.uniformContinuousOn_iff + +theorem IsUniformInducing.inducing {f : α → β} (h : IsUniformInducing f) : Inducing f := by obtain rfl := h.comap_uniformSpace exact inducing_induced f -theorem UniformInducing.prod {α' : Type*} {β' : Type*} [UniformSpace α'] [UniformSpace β'] - {e₁ : α → α'} {e₂ : β → β'} (h₁ : UniformInducing e₁) (h₂ : UniformInducing e₂) : - UniformInducing fun p : α × β => (e₁ p.1, e₂ p.2) := +@[deprecated (since := "2024-10-05")] +alias UniformInducing.inducing := IsUniformInducing.inducing + +theorem IsUniformInducing.prod {α' : Type*} {β' : Type*} [UniformSpace α'] [UniformSpace β'] + {e₁ : α → α'} {e₂ : β → β'} (h₁ : IsUniformInducing e₁) (h₂ : IsUniformInducing e₂) : + IsUniformInducing fun p : α × β => (e₁ p.1, e₂ p.2) := ⟨by simp [Function.comp_def, uniformity_prod, ← h₁.1, ← h₂.1, comap_inf, comap_comap]⟩ -theorem UniformInducing.isDenseInducing {f : α → β} (h : UniformInducing f) (hd : DenseRange f) : +@[deprecated (since := "2024-10-05")] +alias UniformInducing.prod := IsUniformInducing.prod + +lemma IsUniformInducing.isDenseInducing (h : IsUniformInducing f) (hd : DenseRange f) : IsDenseInducing f := { dense := hd induced := h.inducing.induced } -theorem SeparationQuotient.uniformInducing_mk : UniformInducing (mk : α → SeparationQuotient α) := +@[deprecated (since := "2024-10-05")] +alias UniformInducing.isDenseInducing := IsUniformInducing.isDenseInducing + +lemma SeparationQuotient.isUniformInducing_mk : + IsUniformInducing (mk : α → SeparationQuotient α) := ⟨comap_mk_uniformity⟩ -protected theorem UniformInducing.injective [T0Space α] {f : α → β} (h : UniformInducing f) : +@[deprecated (since := "2024-10-05")] +alias SeparationQuotient.uniformInducing_mk := SeparationQuotient.isUniformInducing_mk + +protected theorem IsUniformInducing.injective [T0Space α] {f : α → β} (h : IsUniformInducing f) : Injective f := h.inducing.injective +@[deprecated (since := "2024-10-05")] +alias UniformInducing.injective := IsUniformInducing.injective + /-! ### Uniform embeddings -/ @@ -130,16 +194,19 @@ protected theorem UniformInducing.injective [T0Space α] {f : α → β} (h : Un /-- A map `f : α → β` between uniform spaces is a *uniform embedding* if it is uniform inducing and injective. If `α` is a separated space, then the latter assumption follows from the former. -/ @[mk_iff] -structure IsUniformEmbedding (f : α → β) extends UniformInducing f : Prop where +structure IsUniformEmbedding (f : α → β) extends IsUniformInducing f : Prop where /-- A uniform embedding is injective. -/ inj : Function.Injective f +lemma IsUniformEmbedding.isUniformInducing (hf : IsUniformEmbedding f) : IsUniformInducing f := + hf.toIsUniformInducing + @[deprecated (since := "2024-10-03")] alias UniformEmbedding := IsUniformEmbedding theorem isUniformEmbedding_iff' {f : α → β} : IsUniformEmbedding f ↔ Injective f ∧ UniformContinuous f ∧ comap (Prod.map f f) (𝓤 β) ≤ 𝓤 α := by - rw [isUniformEmbedding_iff, and_comm, uniformInducing_iff'] + rw [isUniformEmbedding_iff, and_comm, isUniformInducing_iff'] @[deprecated (since := "2024-10-01")] alias uniformEmbedding_iff' := isUniformEmbedding_iff' @@ -149,7 +216,7 @@ theorem Filter.HasBasis.isUniformEmbedding_iff' {ι ι'} {p : ι → Prop} {p' : IsUniformEmbedding f ↔ Injective f ∧ (∀ i, p' i → ∃ j, p j ∧ ∀ x y, (x, y) ∈ s j → (f x, f y) ∈ s' i) ∧ (∀ j, p j → ∃ i, p' i ∧ ∀ x y, (f x, f y) ∈ s' i → (x, y) ∈ s j) := by - rw [isUniformEmbedding_iff, and_comm, h.uniformInducing_iff h'] + rw [isUniformEmbedding_iff, and_comm, h.isUniformInducing_iff h'] @[deprecated (since := "2024-10-01")] alias Filter.HasBasis.uniformEmbedding_iff' := Filter.HasBasis.isUniformEmbedding_iff' @@ -181,14 +248,14 @@ alias uniformEmbedding_set_inclusion := isUniformEmbedding_set_inclusion theorem IsUniformEmbedding.comp {g : β → γ} (hg : IsUniformEmbedding g) {f : α → β} (hf : IsUniformEmbedding f) : IsUniformEmbedding (g ∘ f) := - { hg.toUniformInducing.comp hf.toUniformInducing with inj := hg.inj.comp hf.inj } + { hg.isUniformInducing.comp hf.isUniformInducing with inj := hg.inj.comp hf.inj } @[deprecated (since := "2024-10-01")] alias UniformEmbedding.comp := IsUniformEmbedding.comp theorem IsUniformEmbedding.of_comp_iff {g : β → γ} (hg : IsUniformEmbedding g) {f : α → β} : IsUniformEmbedding (g ∘ f) ↔ IsUniformEmbedding f := by - simp_rw [isUniformEmbedding_iff, hg.toUniformInducing.of_comp_iff, hg.inj.of_comp_iff f] + simp_rw [isUniformEmbedding_iff, hg.isUniformInducing.of_comp_iff, hg.inj.of_comp_iff f] @[deprecated (since := "2024-10-01")] alias UniformEmbedding.of_comp_iff := IsUniformEmbedding.of_comp_iff @@ -218,21 +285,27 @@ theorem isUniformEmbedding_inr : IsUniformEmbedding (Sum.inr : β → α ⊕ β) @[deprecated (since := "2024-10-01")] alias uniformEmbedding_inr := isUniformEmbedding_inr -/-- If the domain of a `UniformInducing` map `f` is a T₀ space, then `f` is injective, +/-- If the domain of a `IsUniformInducing` map `f` is a T₀ space, then `f` is injective, hence it is a `IsUniformEmbedding`. -/ -protected theorem UniformInducing.isUniformEmbedding [T0Space α] {f : α → β} - (hf : UniformInducing f) : IsUniformEmbedding f := +protected theorem IsUniformInducing.isUniformEmbedding [T0Space α] {f : α → β} + (hf : IsUniformInducing f) : IsUniformEmbedding f := ⟨hf, hf.inducing.injective⟩ +@[deprecated (since := "2024-10-05")] +alias UniformInducing.isUniformEmbedding := IsUniformInducing.isUniformEmbedding + @[deprecated (since := "2024-10-01")] -alias UniformInducing.uniformEmbedding := UniformInducing.isUniformEmbedding +alias IsUniformInducing.uniformEmbedding := IsUniformInducing.isUniformEmbedding + +theorem isUniformEmbedding_iff_isUniformInducing [T0Space α] {f : α → β} : + IsUniformEmbedding f ↔ IsUniformInducing f := + ⟨IsUniformEmbedding.isUniformInducing, IsUniformInducing.isUniformEmbedding⟩ -theorem isUniformEmbedding_iff_uniformInducing [T0Space α] {f : α → β} : - IsUniformEmbedding f ↔ UniformInducing f := - ⟨IsUniformEmbedding.toUniformInducing, UniformInducing.isUniformEmbedding⟩ +@[deprecated (since := "2024-10-05")] +alias isUniformEmbedding_iff_uniformInducing := isUniformEmbedding_iff_isUniformInducing @[deprecated (since := "2024-10-01")] -alias uniformEmbedding_iff_uniformInducing := isUniformEmbedding_iff_uniformInducing +alias uniformEmbedding_iff_isUniformInducing := isUniformEmbedding_iff_isUniformInducing /-- If a map `f : α → β` sends any two distinct points to point that are **not** related by a fixed `s ∈ 𝓤 β`, then `f` is uniform inducing with respect to the discrete uniformity on `α`: @@ -252,13 +325,13 @@ theorem comap_uniformity_of_spaced_out {α} {f : α → β} {s : Set (β × β)} theorem isUniformEmbedding_of_spaced_out {α} {f : α → β} {s : Set (β × β)} (hs : s ∈ 𝓤 β) (hf : Pairwise fun x y => (f x, f y) ∉ s) : @IsUniformEmbedding α β ⊥ ‹_› f := by let _ : UniformSpace α := ⊥; have := discreteTopology_bot α - exact UniformInducing.isUniformEmbedding ⟨comap_uniformity_of_spaced_out hs hf⟩ + exact IsUniformInducing.isUniformEmbedding ⟨comap_uniformity_of_spaced_out hs hf⟩ @[deprecated (since := "2024-10-01")] alias uniformEmbedding_of_spaced_out := isUniformEmbedding_of_spaced_out protected lemma IsUniformEmbedding.embedding {f : α → β} (h : IsUniformEmbedding f) : Embedding f := - { toInducing := h.toUniformInducing.inducing + { toInducing := h.isUniformInducing.inducing inj := h.inj } @[deprecated (since := "2024-10-01")] @@ -282,8 +355,8 @@ theorem closedEmbedding_of_spaced_out {α} [TopologicalSpace α] [DiscreteTopolo { (isUniformEmbedding_of_spaced_out hs hf).embedding with isClosed_range := isClosed_range_of_spaced_out hs hf } -theorem closure_image_mem_nhds_of_uniformInducing {s : Set (α × α)} {e : α → β} (b : β) - (he₁ : UniformInducing e) (he₂ : IsDenseInducing e) (hs : s ∈ 𝓤 α) : +theorem closure_image_mem_nhds_of_isUniformInducing {s : Set (α × α)} {e : α → β} (b : β) + (he₁ : IsUniformInducing e) (he₂ : IsDenseInducing e) (hs : s ∈ 𝓤 α) : ∃ a, closure (e '' { a' | (a, a') ∈ s }) ∈ 𝓝 b := by obtain ⟨U, ⟨hU, hUo, hsymm⟩, hs⟩ : ∃ U, (U ∈ 𝓤 β ∧ IsOpen U ∧ SymmetricRel U) ∧ Prod.map e e ⁻¹' U ⊆ s := by @@ -296,6 +369,9 @@ theorem closure_image_mem_nhds_of_uniformInducing {s : Set (α × α)} {e : α rcases he₂.dense.mem_nhds (inter_mem hV (ho.mem_nhds hy)) with ⟨x, hxV, hxU⟩ exact ⟨e x, hxV, mem_image_of_mem e hxU⟩ +@[deprecated (since := "2024-10-05")] +alias closure_image_mem_nhds_of_uniformInducing := closure_image_mem_nhds_of_isUniformInducing + theorem isUniformEmbedding_subtypeEmb (p : α → Prop) {e : α → β} (ue : IsUniformEmbedding e) (de : IsDenseEmbedding e) : IsUniformEmbedding (IsDenseEmbedding.subtypeEmb p e) := { comap_uniformity := by @@ -309,28 +385,31 @@ alias uniformEmbedding_subtypeEmb := isUniformEmbedding_subtypeEmb theorem IsUniformEmbedding.prod {α' : Type*} {β' : Type*} [UniformSpace α'] [UniformSpace β'] {e₁ : α → α'} {e₂ : β → β'} (h₁ : IsUniformEmbedding e₁) (h₂ : IsUniformEmbedding e₂) : IsUniformEmbedding fun p : α × β => (e₁ p.1, e₂ p.2) := - { h₁.toUniformInducing.prod h₂.toUniformInducing with inj := h₁.inj.prodMap h₂.inj } + { h₁.isUniformInducing.prod h₂.isUniformInducing with inj := h₁.inj.prodMap h₂.inj } @[deprecated (since := "2024-10-01")] alias UniformEmbedding.prod := IsUniformEmbedding.prod /-- A set is complete iff its image under a uniform inducing map is complete. -/ -theorem isComplete_image_iff {m : α → β} {s : Set α} (hm : UniformInducing m) : +theorem isComplete_image_iff {m : α → β} {s : Set α} (hm : IsUniformInducing m) : IsComplete (m '' s) ↔ IsComplete s := by have fact1 : SurjOn (map m) (Iic <| 𝓟 s) (Iic <| 𝓟 <| m '' s) := surjOn_image .. |>.filter_map_Iic have fact2 : MapsTo (map m) (Iic <| 𝓟 s) (Iic <| 𝓟 <| m '' s) := mapsTo_image .. |>.filter_map_Iic simp_rw [IsComplete, imp.swap (a := Cauchy _), ← mem_Iic (b := 𝓟 _), fact1.forall fact2, hm.cauchy_map_iff, exists_mem_image, map_le_iff_le_comap, hm.inducing.nhds_eq_comap] -/-- If `f : X → Y` is an `UniformInducing` map, the image `f '' s` of a set `s` is complete +/-- If `f : X → Y` is an `IsUniformInducing` map, the image `f '' s` of a set `s` is complete if and only if `s` is complete. -/ -theorem UniformInducing.isComplete_iff {f : α → β} {s : Set α} (hf : UniformInducing f) : +theorem IsUniformInducing.isComplete_iff {f : α → β} {s : Set α} (hf : IsUniformInducing f) : IsComplete (f '' s) ↔ IsComplete s := isComplete_image_iff hf +@[deprecated (since := "2024-10-05")] +alias UniformInducing.isComplete_iff := IsUniformInducing.isComplete_iff + /-- If `f : X → Y` is an `IsUniformEmbedding`, the image `f '' s` of a set `s` is complete if and only if `s` is complete. -/ theorem IsUniformEmbedding.isComplete_iff {f : α → β} {s : Set α} (hf : IsUniformEmbedding f) : - IsComplete (f '' s) ↔ IsComplete s := hf.toUniformInducing.isComplete_iff + IsComplete (f '' s) ↔ IsComplete s := hf.isUniformInducing.isComplete_iff @[deprecated (since := "2024-10-01")] alias UniformEmbedding.isComplete_iff := IsUniformEmbedding.isComplete_iff @@ -342,39 +421,48 @@ theorem Subtype.isComplete_iff {p : α → Prop} {s : Set { x // p x }} : alias ⟨isComplete_of_complete_image, _⟩ := isComplete_image_iff -theorem completeSpace_iff_isComplete_range {f : α → β} (hf : UniformInducing f) : +theorem completeSpace_iff_isComplete_range {f : α → β} (hf : IsUniformInducing f) : CompleteSpace α ↔ IsComplete (range f) := by rw [completeSpace_iff_isComplete_univ, ← isComplete_image_iff hf, image_univ] -alias ⟨_, UniformInducing.completeSpace⟩ := completeSpace_iff_isComplete_range +alias ⟨_, IsUniformInducing.completeSpace⟩ := completeSpace_iff_isComplete_range -theorem UniformInducing.isComplete_range [CompleteSpace α] {f : α → β} (hf : UniformInducing f) : +@[deprecated (since := "2024-10-08")] alias UniformInducing.completeSpace := + IsUniformInducing.completeSpace + +lemma IsUniformInducing.isComplete_range [CompleteSpace α] (hf : IsUniformInducing f) : IsComplete (range f) := (completeSpace_iff_isComplete_range hf).1 ‹_› +@[deprecated (since := "2024-10-05")] +alias UniformInducing.isComplete_range := IsUniformInducing.isComplete_range + /-- If `f` is a surjective uniform inducing map, then its domain is a complete space iff its codomain is a complete space. See also `_root_.completeSpace_congr` for a version that assumes `f` to be an equivalence. -/ -theorem UniformInducing.completeSpace_congr {f : α → β} (hf : UniformInducing f) +theorem IsUniformInducing.completeSpace_congr {f : α → β} (hf : IsUniformInducing f) (hsurj : f.Surjective) : CompleteSpace α ↔ CompleteSpace β := by rw [completeSpace_iff_isComplete_range hf, hsurj.range_eq, completeSpace_iff_isComplete_univ] +@[deprecated (since := "2024-10-05")] +alias UniformInducing.completeSpace_congr := IsUniformInducing.completeSpace_congr + theorem SeparationQuotient.completeSpace_iff : CompleteSpace (SeparationQuotient α) ↔ CompleteSpace α := - .symm <| uniformInducing_mk.completeSpace_congr surjective_mk + .symm <| isUniformInducing_mk.completeSpace_congr surjective_mk instance SeparationQuotient.instCompleteSpace [CompleteSpace α] : CompleteSpace (SeparationQuotient α) := completeSpace_iff.2 ‹_› -/-- See also `UniformInducing.completeSpace_congr` +/-- See also `IsUniformInducing.completeSpace_congr` for a version that works for non-injective maps. -/ theorem completeSpace_congr {e : α ≃ β} (he : IsUniformEmbedding e) : CompleteSpace α ↔ CompleteSpace β := he.completeSpace_congr e.surjective theorem completeSpace_coe_iff_isComplete {s : Set α} : CompleteSpace s ↔ IsComplete s := by - rw [completeSpace_iff_isComplete_range isUniformEmbedding_subtype_val.toUniformInducing, + rw [completeSpace_iff_isComplete_range isUniformEmbedding_subtype_val.isUniformInducing, Subtype.range_coe] alias ⟨_, IsComplete.completeSpace_coe⟩ := completeSpace_coe_iff_isComplete @@ -384,13 +472,13 @@ theorem IsClosed.completeSpace_coe [CompleteSpace α] {s : Set α} (hs : IsClose hs.isComplete.completeSpace_coe theorem completeSpace_ulift_iff : CompleteSpace (ULift α) ↔ CompleteSpace α := - UniformInducing.completeSpace_congr ⟨rfl⟩ ULift.down_surjective + IsUniformInducing.completeSpace_congr ⟨rfl⟩ ULift.down_surjective /-- The lift of a complete space to another universe is still complete. -/ instance ULift.instCompleteSpace [CompleteSpace α] : CompleteSpace (ULift α) := completeSpace_ulift_iff.2 ‹_› -theorem completeSpace_extension {m : β → α} (hm : UniformInducing m) (dense : DenseRange m) +theorem completeSpace_extension {m : β → α} (hm : IsUniformInducing m) (dense : DenseRange m) (h : ∀ f : Filter β, Cauchy f → ∃ x : α, map m f ≤ 𝓝 x) : CompleteSpace α := ⟨fun {f : Filter α} (hf : Cauchy f) => let p : Set (α × α) → Set α → Set α := fun s t => { y : α | ∃ x : α, x ∈ t ∧ (x, y) ∈ s } @@ -436,7 +524,7 @@ theorem completeSpace_extension {m : β → α} (hm : UniformInducing m) (dense _ ≤ 𝓝 x := le_nhds_of_cauchy_adhp ‹Cauchy g› this ⟩⟩ -lemma totallyBounded_image_iff {f : α → β} {s : Set α} (hf : UniformInducing f) : +lemma totallyBounded_image_iff {f : α → β} {s : Set α} (hf : IsUniformInducing f) : TotallyBounded (f '' s) ↔ TotallyBounded s := by refine ⟨fun hs ↦ ?_, fun h ↦ h.image hf.uniformContinuous⟩ simp_rw [(hf.basis_uniformity (basis_sets _)).totallyBounded_iff] @@ -445,21 +533,21 @@ lemma totallyBounded_image_iff {f : α → β} {s : Set α} (hf : UniformInducin use u, hfin rwa [biUnion_image, image_subset_iff, preimage_iUnion₂] at h -theorem totallyBounded_preimage {f : α → β} {s : Set β} (hf : UniformInducing f) +theorem totallyBounded_preimage {f : α → β} {s : Set β} (hf : IsUniformInducing f) (hs : TotallyBounded s) : TotallyBounded (f ⁻¹' s) := (totallyBounded_image_iff hf).1 <| hs.subset <| image_preimage_subset .. instance CompleteSpace.sum [CompleteSpace α] [CompleteSpace β] : CompleteSpace (α ⊕ β) := by rw [completeSpace_iff_isComplete_univ, ← range_inl_union_range_inr] - exact isUniformEmbedding_inl.toUniformInducing.isComplete_range.union - isUniformEmbedding_inr.toUniformInducing.isComplete_range + exact isUniformEmbedding_inl.isUniformInducing.isComplete_range.union + isUniformEmbedding_inr.isUniformInducing.isComplete_range end theorem isUniformEmbedding_comap {α : Type*} {β : Type*} {f : α → β} [u : UniformSpace β] (hf : Function.Injective f) : @IsUniformEmbedding α β (UniformSpace.comap f u) u f := @IsUniformEmbedding.mk _ _ (UniformSpace.comap f u) _ _ - (@UniformInducing.mk _ _ (UniformSpace.comap f u) _ _ rfl) hf + (@IsUniformInducing.mk _ _ (UniformSpace.comap f u) _ _ rfl) hf @[deprecated (since := "2024-10-01")] alias uniformEmbedding_comap := isUniformEmbedding_comap @@ -482,10 +570,10 @@ alias Embedding.to_uniformEmbedding := Embedding.to_isUniformEmbedding section UniformExtension variable {α : Type*} {β : Type*} {γ : Type*} [UniformSpace α] [UniformSpace β] [UniformSpace γ] - {e : β → α} (h_e : UniformInducing e) (h_dense : DenseRange e) {f : β → γ} + {e : β → α} (h_e : IsUniformInducing e) (h_dense : DenseRange e) {f : β → γ} (h_f : UniformContinuous f) -local notation "ψ" => IsDenseInducing.extend (UniformInducing.isDenseInducing h_e h_dense) f +local notation "ψ" => IsDenseInducing.extend (IsUniformInducing.isDenseInducing h_e h_dense) f include h_e h_dense h_f in theorem uniformly_extend_exists [CompleteSpace γ] (a : α) : ∃ c, Tendsto f (comap e (𝓝 a)) (𝓝 c) := @@ -506,7 +594,7 @@ theorem uniform_extend_subtype [CompleteSpace γ] {p : α → Prop} {e : α → isUniformEmbedding_subtypeEmb _ he de have : b ∈ closure (e '' { x | p x }) := (closure_mono <| monotone_image <| hp) (mem_of_mem_nhds hb) - let ⟨c, hc⟩ := uniformly_extend_exists ue'.toUniformInducing de'.dense hf ⟨b, this⟩ + let ⟨c, hc⟩ := uniformly_extend_exists ue'.isUniformInducing de'.dense hf ⟨b, this⟩ replace hc : Tendsto (f ∘ Subtype.val (p := p)) (((𝓝 b).comap e).comap Subtype.val) (𝓝 c) := by simpa only [nhds_subtype_eq_comap, comap_comap, IsDenseEmbedding.subtypeEmb_coe] using hc refine ⟨c, (tendsto_comap'_iff ?_).1 hc⟩ diff --git a/Mathlib/Topology/UnitInterval.lean b/Mathlib/Topology/UnitInterval.lean index c1aff48fbc13d..55c6405dda1b6 100644 --- a/Mathlib/Topology/UnitInterval.lean +++ b/Mathlib/Topology/UnitInterval.lean @@ -40,7 +40,7 @@ theorem one_mem : (1 : ℝ) ∈ I := ⟨zero_le_one, le_rfl⟩ theorem mul_mem {x y : ℝ} (hx : x ∈ I) (hy : y ∈ I) : x * y ∈ I := - ⟨mul_nonneg hx.1 hy.1, mul_le_one hx.2 hy.1 hy.2⟩ + ⟨mul_nonneg hx.1 hy.1, mul_le_one₀ hx.2 hy.1 hy.2⟩ theorem div_mem {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) (hxy : x ≤ y) : x / y ∈ I := ⟨div_nonneg hx hy, div_le_one_of_le₀ hxy hy⟩ diff --git a/Mathlib/Util/AtomM.lean b/Mathlib/Util/AtomM.lean index 2c411702a829b..0b829028c0b11 100644 --- a/Mathlib/Util/AtomM.lean +++ b/Mathlib/Util/AtomM.lean @@ -22,7 +22,7 @@ namespace Mathlib.Tactic open Lean Meta /-- The context (read-only state) of the `AtomM` monad. -/ -structure AtomM.Context := +structure AtomM.Context where /-- The reducibility setting for definitional equality of atoms -/ red : TransparencyMode /-- A simplification to apply to atomic expressions when they are encountered, @@ -31,7 +31,7 @@ structure AtomM.Context := deriving Inhabited /-- The mutable state of the `AtomM` monad. -/ -structure AtomM.State := +structure AtomM.State where /-- The list of atoms-up-to-defeq encountered thus far, used for atom sorting. -/ atoms : Array Expr := #[] diff --git a/Mathlib/Util/Superscript.lean b/Mathlib/Util/Superscript.lean index a696144b7f9fe..85d3b2fecae7f 100644 --- a/Mathlib/Util/Superscript.lean +++ b/Mathlib/Util/Superscript.lean @@ -95,7 +95,8 @@ partial def satisfyTokensFn (p : Char → Bool) (errorMsg : String) (many := tru variable {α : Type u} [Inhabited α] (as : Array α) (leftOfPartition : α → Bool) in /-- Given a predicate `leftOfPartition` which is true for indexes `< i` and false for `≥ i`, returns `i`, by binary search. -/ -@[specialize] partial def partitionPoint (lo := 0) (hi := as.size) : Nat := +@[specialize] +def partitionPoint (lo := 0) (hi := as.size) : Nat := if lo < hi then let m := (lo + hi)/2 let a := as.get! m @@ -104,6 +105,7 @@ returns `i`, by binary search. -/ else partitionPoint lo m else lo + termination_by hi - lo /-- The core function for super/subscript parsing. It consists of three stages: diff --git a/README.md b/README.md index 53d266e917989..dfe49c658f2b5 100644 --- a/README.md +++ b/README.md @@ -30,7 +30,7 @@ For more pointers, see [Learning Lean](https://leanprover-community.github.io/le ## Documentation Besides the installation guides above and [Lean's general -documentation](https://leanprover.github.io/documentation/), the documentation +documentation](https://docs.lean-lang.org/lean4/doc/), the documentation of mathlib consists of: - [The mathlib4 docs](https://leanprover-community.github.io/mathlib4_docs/index.html): documentation [generated diff --git a/lake-manifest.json b/lake-manifest.json index 33799d2be2391..4a49ec817f12a 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "13f9b00769bdac2c0041406a6c2524a361e8d660", + "rev": "daf1ed91789811cf6bbb7bf2f4dad6b3bad8fbf4", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -15,7 +15,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "2c8ae451ce9ffc83554322b14437159c1a9703f9", + "rev": "2b2f6d7fbe9d917fc010e9054c1ce11774c9088b", "name": "Qq", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -44,7 +44,7 @@ {"url": "https://github.com/leanprover/lean4-cli", "type": "git", "subDir": null, - "scope": "", + "scope": "leanprover", "rev": "2cf1030dc2ae6b3632c84a09350b675ef3e347d0", "name": "Cli", "manifestFile": "lake-manifest.json", @@ -55,7 +55,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "63a7d4a353f48f6c5f1bc19d0f018b0513cb370a", + "rev": "7376ac07aa2b0492372c056b7a2c3163b3026d1e", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/scripts/init_creation.sh b/scripts/init_creation.sh index 5fc09772f9670..7abca30241c43 100644 --- a/scripts/init_creation.sh +++ b/scripts/init_creation.sh @@ -2,8 +2,8 @@ : <<'BASH_MODULE_DOC' -These are the commands to generate a "root" `Mathlib/Init.lean` file, imported by all the -`Mathlib` files that do not import any `Mathlib` file. +These are the commands to add an import of `Mathlib/Init.lean` to all `Mathlib` files +that do not import any `Mathlib` file. BASH_MODULE_DOC @@ -35,11 +35,3 @@ printf 'Adding `import Mathlib.Init` to all file that import no Mathlib file.\n' # The `sed` command appends the line `import Mathlib.Init` after the first # `-/[linebreaks]*` of each file printed by `mathlibNonImportingFiles`. sed -i -z 's=-/\n*=&import Mathlib.Init\n=' $(mathlibNonImportingFiles) - -printf 'Creating `Mathlib/Init.lean`.\n' - -# Creates the `Mathlib/Init.lean` files. -echo '-- This is the root file in Mathlib: it is imported by virtually *all* Mathlib files' > Mathlib/Init.lean - -printf $'Don\'t forget to add `Mathlib.Init` to the `ignoreImport` field of `scripts/noshake.json` -This ensures that `import Mathlib.Init` does not trigger a `shake` exception.\n' diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index de0b42c05c8e5..16a703c2e1780 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -4729,8 +4729,8 @@ uniformContinuous_nnnorm' uniformContinuous_norm' isUniformEmbedding_iff' UniformGroup.mk' -uniformInducing_iff' -UniformInducing.mk' +isUniformInducing_iff' +IsUniformInducing.mk' uniformity_basis_edist' uniformity_basis_edist_le' uniformity_eq_comap_nhds_one' diff --git a/scripts/nolints.json b/scripts/nolints.json index d897d5a3db174..8ea41fbf362a9 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -23,7 +23,6 @@ ["docBlame", "RightInverse"], ["docBlame", "Writer"], ["docBlame", "WriterT"], - ["docBlame", "Zero"], ["docBlame", "cancelDenominators"], ["docBlame", "cancelDenominatorsAt"], ["docBlame", "cancelDenominatorsTarget"], @@ -374,7 +373,6 @@ ["docBlame", "WriterT.mkLabel'"], ["docBlame", "WriterT.run"], ["docBlame", "WriterT.runThe"], - ["docBlame", "Zero.zero"], ["docBlame", "Zsqrtd.im"], ["docBlame", "Zsqrtd.re"], ["docBlame", "algebraMap.coeHTCT"], @@ -714,10 +712,10 @@ ["docBlame", "Mathlib.Command.Variable.variable?.maxSteps"], ["docBlame", "Mathlib.Meta.NormNum.evalEq.intArm"], ["docBlame", "Mathlib.Meta.NormNum.evalEq.ratArm"], - ["docBlame", "Mathlib.Meta.NormNum.evalLE.intArm"], - ["docBlame", "Mathlib.Meta.NormNum.evalLE.ratArm"], - ["docBlame", "Mathlib.Meta.NormNum.evalLT.intArm"], - ["docBlame", "Mathlib.Meta.NormNum.evalLT.ratArm"], + ["docBlame", "Mathlib.Meta.NormNum.evalLE.core.intArm"], + ["docBlame", "Mathlib.Meta.NormNum.evalLE.core.ratArm"], + ["docBlame", "Mathlib.Meta.NormNum.evalLT.core.intArm"], + ["docBlame", "Mathlib.Meta.NormNum.evalLT.core.ratArm"], ["docBlame", "Mathlib.Meta.NormNum.evalMinFac.aux"], ["docBlame", "Mathlib.Meta.NormNum.evalMinFac.core"], ["docBlame", "Mathlib.Meta.NormNum.evalNatPrime.core"], diff --git a/scripts/technical-debt-metrics.sh b/scripts/technical-debt-metrics.sh index 20c4bc55f3a0f..ce9a3a5202b16 100755 --- a/scripts/technical-debt-metrics.sh +++ b/scripts/technical-debt-metrics.sh @@ -53,6 +53,8 @@ printf '%s|%s\n' "$(grep -c 'docBlame' scripts/nolints.json)" "documentation nol printf '%s|%s\n' "$(git grep '^set_option linter.style.longFile [0-9]*' Mathlib | wc -l)" "large files" printf '%s|%s\n' "$(git grep "^open .*Classical" | grep -v " in$" -c)" "bare open (scoped) Classical" +printf '%s|%s\n' "$(wc -l < scripts/no_lints_prime_decls.txt)" "exceptions for the docPrime linter" + deprecatedFiles="$(git ls-files '**/Deprecated/*.lean' | xargs wc -l | sed 's=^ *==')" printf '%s|%s\n' "$(printf '%s' "${deprecatedFiles}" | wc -l)" "\`Deprecated\` files" diff --git a/test/Continuity.lean b/test/Continuity.lean index 70da467d763ee..bf0d389a95e2a 100644 --- a/test/Continuity.lean +++ b/test/Continuity.lean @@ -53,8 +53,7 @@ example (f : C(X, Y)) (g : C(Y, Z)) : Continuous (g ∘ f) := by continuity example (f : C(X, Y)) (g : C(X, Z)) : Continuous (fun x => (f x, g x)) := by continuity -example (f : C(X, Y)) (g : C(W, Z)) : Continuous (Prod.map f g) := --by continuity - f.continuous.prod_map g.continuous +example (f : C(X, Y)) (g : C(W, Z)) : Continuous (Prod.map f g) := by continuity example (f : ∀ i, C(X, X' i)) : Continuous (fun a i => f i a) := by continuity diff --git a/test/MfldSetTac.lean b/test/MfldSetTac.lean index fd50fd05c2ca3..ad921934de097 100644 --- a/test/MfldSetTac.lean +++ b/test/MfldSetTac.lean @@ -47,7 +47,7 @@ test_sorry (e.toPartialEquiv.symm : β → α) = (e.symm : β → α) := test_sorry -structure ModelWithCorners (𝕜 E H : Type u) extends PartialEquiv H E := +structure ModelWithCorners (𝕜 E H : Type u) extends PartialEquiv H E where (source_eq : source = Set.univ) attribute [mfld_simps] ModelWithCorners.source_eq diff --git a/test/NthRewrite.lean b/test/NthRewrite.lean index 608d18fdb6bf8..287039007dce6 100644 --- a/test/NthRewrite.lean +++ b/test/NthRewrite.lean @@ -14,7 +14,7 @@ example [AddZeroClass G] {a : G} (h : a = a): a = (a + 0) := by example [AddZeroClass G] {a : G} : a + a = a + (a + 0) := by nth_rw 2 [← add_zero a] -structure F := +structure F where (a : ℕ) (v : Vector ℕ a) (p : v.val = []) @@ -22,7 +22,7 @@ structure F := example (f : F) : f.v.val = [] := by nth_rw 1 [f.p] -structure Cat := +structure Cat where (O : Type) (H : O → O → Type) (i : (o : O) → H o o) diff --git a/test/PPRoundtrip.lean b/test/PPRoundtrip.lean new file mode 100644 index 0000000000000..752a26277923b --- /dev/null +++ b/test/PPRoundtrip.lean @@ -0,0 +1,66 @@ +import Mathlib.Tactic.Linter.PPRoundtrip + +/-- +info: "a a" +--- +warning: source context +'al " a ' +'al " a a\n' +pretty-printed context +note: this linter can be disabled with `set_option linter.ppRoundtrip false` +-/ +#guard_msgs in +set_option linter.ppRoundtrip true in +#eval " a a\n " |>.trim + +/-- +warning: source context +'rd ¬ fa' +'rd ¬false' +pretty-printed context +note: this linter can be disabled with `set_option linter.ppRoundtrip false` +-/ +#guard_msgs in +set_option linter.ppRoundtrip true in +#guard ¬ false + +/-- +warning: source context +'le {a: Nat' +'le {a : Na' +pretty-printed context +note: this linter can be disabled with `set_option linter.ppRoundtrip false` +-/ +#guard_msgs in +set_option linter.ppRoundtrip true in +variable {a: Nat} + +/-- +warning: source context +' {a :Nat}' +' {a : Nat}' +pretty-printed context +note: this linter can be disabled with `set_option linter.ppRoundtrip false` +-/ +#guard_msgs in +set_option linter.ppRoundtrip true in +variable {a :Nat} + +/-- +info: (fun x1 x2 => x1 + x2) 0 1 : Nat +--- +warning: source context +'k (·+·) ' +'k (· + ·' +pretty-printed context +note: this linter can be disabled with `set_option linter.ppRoundtrip false` +-/ +#guard_msgs in +set_option linter.ppRoundtrip true in +#check (·+·) 0 1 + +#guard_msgs in +set_option linter.ppRoundtrip true in +-- check that trailing comments do not trigger the linter +example : 0 = 0 := by + rw [] -- this goal is closed by the `rfl` implied by `rw` diff --git a/test/Simps.lean b/test/Simps.lean index 4fb569642e899..af1bc4b6f4a5e 100644 --- a/test/Simps.lean +++ b/test/Simps.lean @@ -91,7 +91,7 @@ initialize_simps_projections Something universe v u w -structure Equiv' (α : Sort _) (β : Sort _) := +structure Equiv' (α : Sort _) (β : Sort _) where (toFun : α → β) (invFun : β → α) (left_inv : invFun.LeftInverse toFun) @@ -262,7 +262,7 @@ run_cmd liftTermElabM <| do guard <| env.find? `rflWithData'_toEquiv_toFun |>.isNone guard <| env.find? `test_sneaky_extra |>.isNone -structure PartiallyAppliedStr := +structure PartiallyAppliedStr where (data : ℕ → MyProd ℕ ℕ) /- if we have a partially applied constructor, we treat it as if it were eta-expanded -/ @@ -279,7 +279,7 @@ run_cmd liftTermElabM <| do guard <| simpsAttr.getParam? env `partially_applied_term == #[`partially_applied_term_data_fst, `partially_applied_term_data_snd] -structure VeryPartiallyAppliedStr := +structure VeryPartiallyAppliedStr where (data : ∀β, ℕ → β → MyProd ℕ β) /- if we have a partially applied constructor, we treat it as if it were eta-expanded. @@ -424,12 +424,12 @@ run_cmd liftTermElabM <| do guard <| env.find? `pprodEquivProd22_invFun_snd |>.isSome /- Tests with universe levels -/ -class has_hom (obj : Type u) : Type (max u (v+1)) := +class has_hom (obj : Type u) : Type (max u (v+1)) where (hom : obj → obj → Type v) infixr:10 " ⟶ " => has_hom.hom -- type as \h -class CategoryStruct (obj : Type u) extends has_hom.{v} obj : Type (max u (v+1)) := +class CategoryStruct (obj : Type u) extends has_hom.{v} obj : Type (max u (v+1)) where (id : ∀ X : obj, hom X X) (comp : ∀ {X Y Z : obj}, (X ⟶ Y) → (Y ⟶ Z) → (X ⟶ Z)) @@ -450,7 +450,7 @@ example (X Y Z : Type u) (f : X ⟶ Y) (g : Y ⟶ Z) {k : X → Z} (h : ∀ x, g namespace coercing -structure FooStr := +structure FooStr where (c : Type) (x : c) @@ -462,7 +462,7 @@ instance : CoeSort FooStr Type := ⟨FooStr.c⟩ example {x : Type} (h : ℕ = x) : foo = x := by simp only [foo_c]; rw [h] example {x : ℕ} (h : (3 : ℕ) = x) : foo.x = x := by simp only [foo_x]; rw [h] -structure VooStr (n : ℕ) := +structure VooStr (n : ℕ) where (c : Type) (x : c) @@ -474,7 +474,7 @@ instance (n : ℕ) : CoeSort (VooStr n) Type := ⟨VooStr.c⟩ example {x : Type} (h : ℕ = x) : voo = x := by simp only [voo_c]; rw [h] example {x : ℕ} (h : (3 : ℕ) = x) : voo.x = x := by simp only [voo_x]; rw [h] -structure Equiv2 (α : Sort _) (β : Sort _) := +structure Equiv2 (α : Sort _) (β : Sort _) where (toFun : α → β) (invFun : β → α) (left_inv : invFun.LeftInverse toFun) @@ -515,7 +515,7 @@ class Semigroup (G : Type u) extends Mul G where example {α β} [Semigroup α] [Semigroup β] (x y : α × β) : x * y = (x.1 * y.1, x.2 * y.2) := by simp example {α β} [Semigroup α] [Semigroup β] (x y : α × β) : (x * y).1 = x.1 * y.1 := by simp -structure BSemigroup := +structure BSemigroup where (G : Type _) (op : G → G → G) -- (infix:60 " * " => op) -- this seems to be removed @@ -535,8 +535,8 @@ protected def prod (G H : BSemigroup) : BSemigroup := end BSemigroup -class ExtendingStuff (G : Type u) extends Mul G, Zero G, Neg G, HasSubset G := - (new_axiom : ∀ x : G, x * - 0 ⊆ - x) +class ExtendingStuff (G : Type u) extends Mul G, Zero G, Neg G, HasSubset G where + new_axiom : ∀ x : G, x * - 0 ⊆ - x @[simps] def bar : ExtendingStuff ℕ := { mul := (·*·) @@ -550,8 +550,8 @@ attribute [local instance] bar example (x : ℕ) : x * - 0 ⊆ - x := by simp end -class new_ExtendingStuff (G : Type u) extends Mul G, Zero G, Neg G, HasSubset G := - (new_axiom : ∀ x : G, x * - 0 ⊆ - x) +class new_ExtendingStuff (G : Type u) extends Mul G, Zero G, Neg G, HasSubset G where + new_axiom : ∀ x : G, x * - 0 ⊆ - x @[simps] def new_bar : new_ExtendingStuff ℕ := { mul := (·*·) @@ -570,7 +570,7 @@ end coercing namespace ManualCoercion -structure Equiv (α : Sort _) (β : Sort _) := +structure Equiv (α : Sort _) (β : Sort _) where (toFun : α → β) (invFun : β → α) @@ -598,7 +598,7 @@ end ManualCoercion namespace FaultyManualCoercion -structure Equiv (α : Sort _) (β : Sort _) := +structure Equiv (α : Sort _) (β : Sort _) where (toFun : α → β) (invFun : β → α) @@ -622,7 +622,7 @@ namespace ManualInitialize /- defining a manual coercion. -/ variable {α β γ : Sort _} -structure Equiv (α : Sort _) (β : Sort _) := +structure Equiv (α : Sort _) (β : Sort _) where (toFun : α → β) (invFun : β → α) @@ -654,7 +654,7 @@ namespace FaultyUniverses variable {α β γ : Sort _} -structure Equiv (α : Sort u) (β : Sort v) := +structure Equiv (α : Sort u) (β : Sort v) where (toFun : α → β) (invFun : β → α) @@ -683,7 +683,7 @@ namespace ManualUniverses variable {α β γ : Sort _} -structure Equiv (α : Sort u) (β : Sort v) := +structure Equiv (α : Sort u) (β : Sort v) where (toFun : α → β) (invFun : β → α) @@ -704,7 +704,7 @@ end ManualUniverses namespace ManualProjectionNames -structure Equiv (α : Sort _) (β : Sort _) := +structure Equiv (α : Sort _) (β : Sort _) where (toFun : α → β) (invFun : β → α) @@ -744,7 +744,7 @@ end ManualProjectionNames namespace PrefixProjectionNames -structure Equiv (α : Sort _) (β : Sort _) := +structure Equiv (α : Sort _) (β : Sort _) where (toFun : α → β) (invFun : β → α) @@ -791,7 +791,7 @@ end PrefixProjectionNames -- test transparency setting -structure SetPlus (α : Type) := +structure SetPlus (α : Type) where (s : Set α) (x : α) (h : x ∈ s) @@ -818,7 +818,7 @@ example {x : Set ℕ} (h : Set.univ = x) : Nat.SetPlus3.s = x := by namespace NestedNonFullyApplied -structure Equiv (α : Sort _) (β : Sort _) := +structure Equiv (α : Sort _) (β : Sort _) where (toFun : α → β) (invFun : β → α) @@ -854,19 +854,19 @@ example (e : α ≃ β) {x : β → α} (h : e.invFun = x) : (Equiv.symm2.invFun end NestedNonFullyApplied -- test that type classes which are props work -class PropClass (n : ℕ) : Prop := - (has_true : True) +class PropClass (n : ℕ) : Prop where + has_true : True instance has_PropClass (n : ℕ) : PropClass n := ⟨trivial⟩ -structure NeedsPropClass (n : ℕ) [PropClass n] := +structure NeedsPropClass (n : ℕ) [PropClass n] where (t : True) @[simps] def test_PropClass : NeedsPropClass 1 := { t := trivial } /- check that when the coercion is given in eta-expanded form, we can also find the coercion. -/ -structure AlgHom (R A B : Type _) := +structure AlgHom (R A B : Type _) where (toFun : A → B) instance (R A B : Type _) : CoeFun (AlgHom R A B) (fun _ ↦ A → B) := ⟨fun f ↦ f.toFun⟩ @@ -931,7 +931,7 @@ section attribute [local simp] Nat.add -structure MyType := +structure MyType where (A : Type) @[simps (config := {simpRhs := true})] def myTypeDef : MyType := @@ -972,7 +972,7 @@ instance {α β} : CoeFun (α ≃ β) (fun _ ↦ α → β) := ⟨Equiv'.toFun @[simps] protected def Equiv'.symm {α β} (f : α ≃ β) : β ≃ α := ⟨f.invFun, f, f.right_inv, f.left_inv⟩ -structure DecoratedEquiv (α : Sort _) (β : Sort _) extends Equiv' α β := +structure DecoratedEquiv (α : Sort _) (β : Sort _) extends Equiv' α β where (P_toFun : Function.Injective toFun ) (P_invFun : Function.Injective invFun) @@ -1022,7 +1022,7 @@ example {α : Type} (x z : α) (h : x = z) : foo2 α x = z := by guard_target = x = z rw [h] -structure FurtherDecoratedEquiv (α : Sort _) (β : Sort _) extends DecoratedEquiv α β := +structure FurtherDecoratedEquiv (α : Sort _) (β : Sort _) extends DecoratedEquiv α β where (Q_toFun : Function.Surjective toFun ) (Q_invFun : Function.Surjective invFun ) @@ -1097,11 +1097,11 @@ def fffoo2 (α : Type) : OneMore α α := fffoo α /- test the case where a projection takes additional arguments. -/ variable {ι : Type _} [DecidableEq ι] (A : ι → Type _) -structure ZeroHom (M N : Type _) [Zero M] [Zero N] := +structure ZeroHom (M N : Type _) [Zero M] [Zero N] where (toFun : M → N) (map_zero' : toFun 0 = 0) -structure AddHom (M N : Type _) [Add M] [Add N] := +structure AddHom (M N : Type _) [Add M] [Add N] where (toFun : M → N) (map_add' : ∀ x y, toFun (x + y) = toFun x + toFun y) @@ -1112,7 +1112,7 @@ infixr:25 " →+ " => AddMonoidHom instance (M N : Type _) [AddMonoid M] [AddMonoid N] : CoeFun (M →+ N) (fun _ ↦ M → N) := ⟨(·.toFun)⟩ -class AddHomPlus [Add ι] [∀ i, AddCommMonoid (A i)] := +class AddHomPlus [Add ι] [∀ i, AddCommMonoid (A i)] where (myMul {i} : A i →+ A i) def AddHomPlus.Simps.apply [Add ι] [∀ i, AddCommMonoid (A i)] [AddHomPlus A] {i : ι} (x : A i) : @@ -1121,7 +1121,7 @@ def AddHomPlus.Simps.apply [Add ι] [∀ i, AddCommMonoid (A i)] [AddHomPlus A] initialize_simps_projections AddHomPlus (myMul_toFun → apply, -myMul) -class AddHomPlus2 [Add ι] := +class AddHomPlus2 [Add ι] where (myMul {i j} : A i ≃ (A j ≃ A (i + j))) def AddHomPlus2.Simps.mul [Add ι] [AddHomPlus2 A] {i j : ι} (x : A i) (y : A j) : A (i + j) := @@ -1153,7 +1153,7 @@ end comp_projs section /-! Check that the tactic also works if the elaborated type of `type` reduces to `Sort _`, but is not `Sort _` itself. -/ -structure MyFunctor (C D : Type _) := +structure MyFunctor (C D : Type _) where (obj : C → D) local infixr:26 " ⥤ " => MyFunctor diff --git a/test/fun_prop_dev.lean b/test/fun_prop_dev.lean index f3f139b1913a3..f78abae43617a 100644 --- a/test/fun_prop_dev.lean +++ b/test/fun_prop_dev.lean @@ -107,10 +107,10 @@ structure LinHom (α β) where infixr:25 " -o " => LinHom instance : CoeFun (α ->> β) (fun _ => α → β) where - coe := fun f => f.toFun + coe f := f.toFun instance : FunLike (α -o β) α β where - coe := fun f => f.toFun + coe f := f.toFun coe_injective' := silentSorry #eval Lean.Elab.Command.liftTermElabM do diff --git a/test/interactiveUnfold.lean b/test/interactiveUnfold.lean index 7972cf5866271..b17d3daa3cae8 100644 --- a/test/interactiveUnfold.lean +++ b/test/interactiveUnfold.lean @@ -57,7 +57,8 @@ info: Unfolds for 5 / 3: info: Unfolds for 1 + 1: · Ordinal.type (Sum.Lex EmptyRelation EmptyRelation) · ⟦{ α := PUnit.{u_1 + 1} ⊕ PUnit.{u_1 + 1}, r := Sum.Lex EmptyRelation EmptyRelation, wo := ⋯ }⟧ -· Quot.mk Setoid.r { α := PUnit.{u_1 + 1} ⊕ PUnit.{u_1 + 1}, r := Sum.Lex EmptyRelation EmptyRelation, wo := ⋯ } +· Quot.mk ⇑Ordinal.isEquivalent + { α := PUnit.{u_1 + 1} ⊕ PUnit.{u_1 + 1}, r := Sum.Lex EmptyRelation EmptyRelation, wo := ⋯ } -/ #guard_msgs in #unfold? (1 : Ordinal) + 1 diff --git a/test/linear_combination.lean b/test/linear_combination.lean index ac2eaed2036b2..9f652a8186a63 100644 --- a/test/linear_combination.lean +++ b/test/linear_combination.lean @@ -229,6 +229,15 @@ example (x y : ℤ) (h1 : x * y + 2 * x = 1) (h2 : x = y) : x * y = -2 * y + 1 : /-! ### Cases that should fail -/ +/-- +error: ring failed, ring expressions not equal +a : ℤ +ha : a = 1 +⊢ -1 = 0 +-/ +#guard_msgs in +example (a : ℤ) (ha : a = 1) : a = 2 := by linear_combination ha + /-- error: ring failed, ring expressions not equal a : ℚ diff --git a/test/measurability.lean b/test/measurability.lean index 2ebf1a799ee68..d373a1d2d8576 100644 --- a/test/measurability.lean +++ b/test/measurability.lean @@ -82,7 +82,7 @@ example [Div β] [MeasurableDiv₂ β] (hf : Measurable f) (hg : Measurable g) example [AddCommMonoid β] [MeasurableAdd₂ β] {s : Finset ℕ} {F : ℕ → α → β} (hF : ∀ i, Measurable (F i)) : Measurable (∑ i ∈ s, (fun x => F (i+1) x + F i x)) := by - measurability + fun_prop example [AddCommMonoid β] [MeasurableAdd₂ β] {s : Finset ℕ} {F : ℕ → α → β} (hF : ∀ i, AEMeasurable (F i) μ) : AEMeasurable (∑ i ∈ s, (fun x => F (i+1) x + F i x)) μ := by diff --git a/test/ring_compare.lean b/test/ring_compare.lean new file mode 100644 index 0000000000000..3e282d9fbfe30 --- /dev/null +++ b/test/ring_compare.lean @@ -0,0 +1,114 @@ +import Mathlib.Tactic.NormNum.OfScientific +import Mathlib.Tactic.Ring.Compare +import Mathlib.Tactic.Ring.RingNF + +open Lean Elab Tactic + +elab "ring_le" : tactic => liftMetaFinishingTactic Mathlib.Tactic.Ring.proveLE +elab "ring_lt" : tactic => liftMetaFinishingTactic Mathlib.Tactic.Ring.proveLT + +section Nat +variable {x y : ℕ} + +example : 3 ≤ (3:ℕ) := by ring_le +example : 1 ≤ (3:ℕ) := by ring_le +example : 0 ≤ (3:ℕ) + 1 := by ring_le +example : x ≤ x + 3 := by ring_le +example : x ≤ 1 + x := by ring_le +example : x + y + 1 ≤ y + x + 3 := by ring_le +example : x + y ≤ y + x + 3 := by ring_le +example : x + y + 1 ≤ y + 4 + x := by ring_le + +example : 1 < (3:ℕ) := by ring_lt +example : 0 < (3:ℕ) + 1 := by ring_lt +example : x < x + 3 := by ring_lt +example : x < 1 + x := by ring_lt +example : x + y + 1 < y + x + 3 := by ring_lt +example : x + y < y + x + 3 := by ring_lt +example : x + y + 1 < y + 4 + x := by ring_lt + +end Nat + +section LinearOrderedField +variable {K : Type*} [LinearOrderedField K] {x y : K} + +example : (0:K) ≤ 0 := by ring_le +example : 3 ≤ (3:K) := by ring_le +example : 1 ≤ (3:K) := by ring_le +example : -1 ≤ (3:K) := by ring_le +example : 1.5 ≤ (3:K) := by ring_le +example : 0 ≤ x + 3 - x := by ring_le +example : -1 ≤ x - x := by ring_le +example : x + y + 1 ≤ y + x + 3 := by ring_le +example : x + y + 1 ≤ y + x + 1 := by ring_le +example : x + y ≤ y + x + 3 := by ring_le +example : x + y - 3 ≤ y + x := by ring_le +example : x + y - x + 1 ≤ y + (4:K) := by ring_le + +example : 1 < (3:K) := by ring_lt +example : -1 < (3:K) := by ring_lt +example : 1.5 < (3:K) := by ring_lt +example : 0 < x + 3 - x := by ring_lt +example : -1 < x - x := by ring_lt +example : x + y + 1 < y + x + 3 := by ring_lt +example : x + y < y + x + 3 := by ring_lt +example : x + y - 3 < y + x := by ring_lt +example : x + y - x + 1 < y + (4:K) := by ring_lt + +/- The speed of `Mathlib.Tactic.Ring.proveLE` is very sensitive to how much typeclass inference is +demanded by the lemmas it orchestrates. This example took 1112 heartbeats (and 40 ms on a good +laptop) on an implementation with "minimal" typeclasses everywhere, e.g. lots of +`CovariantClass`/`ContravariantClass`, and takes 662 heartbeats (28 ms on a good laptop) on the +implementation at the time of joining Mathlib (October 2024). -/ +set_option maxHeartbeats 750 in +example : x + y - x + 1 ≤ y + (4:K) := by ring_le + +/- The speed of `Mathlib.Tactic.Ring.proveLT` is very sensitive to how much typeclass inference is +demanded by the lemmas it orchestrates. This example took 1410 heartbeats (and 48 ms on a good +laptop) on an implementation with "minimal" typeclasses everywhere, e.g. lots of +`CovariantClass`/`ContravariantClass`, and takes 676 heartbeats (28 ms on a good laptop) on the +implementation at the time of joining Mathlib (October 2024). -/ +set_option maxHeartbeats 750 in +example : x + y - x + 1 < y + (4:K) := by ring_lt + +/-- +error: ring failed, ring expressions not equal up to an additive constant +K : Type u_1 +inst✝ : LinearOrderedField K +x y : K +⊢ 1 + x + y ≤ 3 + y +-/ +#guard_msgs in +example : x + y + 1 ≤ y + 3 := by ring_le + +/-- +error: comparison failed, LHS is larger +K : Type u_1 +inst✝ : LinearOrderedField K +x y : K +⊢ 4 + x + y ≤ 3 + x + y +-/ +#guard_msgs in +example : x + y + 4 ≤ y + x + 3 := by ring_le + +/-- +error: ring failed, ring expressions not equal up to an additive constant +K : Type u_1 +inst✝ : LinearOrderedField K +x y : K +⊢ 1 + x + y < 3 + y +-/ +#guard_msgs in +example : x + y + 1 < y + 3 := by ring_lt + +/-- +error: comparison failed, LHS is at least as large +K : Type u_1 +inst✝ : LinearOrderedField K +x y : K +⊢ 4 + x + y < 4 + x + y +-/ +#guard_msgs in +example : x + y + 4 < y + x + 4 := by ring_lt + +end LinearOrderedField diff --git a/test/slow_simp.lean b/test/slow_simp.lean index fa6de3b23f5b9..651d0e10e6014 100644 --- a/test/slow_simp.lean +++ b/test/slow_simp.lean @@ -61,7 +61,7 @@ def PointedSpaceEquiv_inverse : Under (TopCat.of Unit) ⥤ PointedSpace where base := by have := f.w replace this := DFunLike.congr_fun this () - simp [- Under.w] at this + simp [-Under.w] at this simp exact this.symm } map_comp := by intros; simp_all; rfl -- This is the slow step. diff --git a/test/toAdditive.lean b/test/toAdditive.lean index cb1938438e035..bd2ad86c8e087 100644 --- a/test/toAdditive.lean +++ b/test/toAdditive.lean @@ -19,12 +19,12 @@ def foo0 {α} [Mul α] [One α] (x y : α) : α := x * y * 1 theorem bar0_works : bar0 3 4 = 7 := by decide -class my_has_pow (α : Type u) (β : Type v) := +class my_has_pow (α : Type u) (β : Type v) where (pow : α → β → α) instance : my_has_pow Nat Nat := ⟨fun a b => a ^ b⟩ -class my_has_scalar (M : Type u) (α : Type v) := +class my_has_scalar (M : Type u) (α : Type v) where (smul : M → α → α) instance : my_has_scalar Nat Nat := ⟨fun a b => a * b⟩