From d82cdbff79fdde07da9d4edc86a3d9d4fbc67c13 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 17 Oct 2024 14:30:25 +0000 Subject: [PATCH 001/131] feat(RingTheory/Smooth): Smoothness of product rings. (#15131) --- Mathlib.lean | 1 + Mathlib/Algebra/Algebra/Pi.lean | 9 ++ .../RingTheory/Ideal/QuotientOperations.lean | 6 + Mathlib/RingTheory/Smooth/Pi.lean | 121 ++++++++++++++++++ 4 files changed, 137 insertions(+) create mode 100644 Mathlib/RingTheory/Smooth/Pi.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8c55dc3701c47..155e528def04f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4133,6 +4133,7 @@ import Mathlib.RingTheory.SimpleRing.Basic import Mathlib.RingTheory.SimpleRing.Defs import Mathlib.RingTheory.Smooth.Basic import Mathlib.RingTheory.Smooth.Kaehler +import Mathlib.RingTheory.Smooth.Pi import Mathlib.RingTheory.Smooth.StandardSmooth import Mathlib.RingTheory.Support import Mathlib.RingTheory.SurjectiveOnStalks diff --git a/Mathlib/Algebra/Algebra/Pi.lean b/Mathlib/Algebra/Algebra/Pi.lean index db4ea723211c1..fbdfdb6ca061d 100644 --- a/Mathlib/Algebra/Algebra/Pi.lean +++ b/Mathlib/Algebra/Algebra/Pi.lean @@ -52,6 +52,15 @@ theorem algebraMap_apply {_ : CommSemiring R} [_s : ∀ i, Semiring (f i)] [∀ -- when each `A i` is an `R i`-algebra, although I'm not sure that it's useful. variable {I} (R) +/-- A family of algebra homomorphisms `g i : A →ₐ[R] f i` defines a ring homomorphism +`Pi.algHom g : A →ₐ[R] Π i, f i` given by `Pi.algHom g x i = f i x`. -/ +@[simps!] +def algHom [CommSemiring R] [s : ∀ i, Semiring (f i)] [∀ i, Algebra R (f i)] + {A : Type*} [Semiring A] [Algebra R A] (g : ∀ i, A →ₐ[R] f i) : + A →ₐ[R] ∀ i, f i where + __ := Pi.ringHom fun i ↦ (g i).toRingHom + commutes' r := by ext; simp + /-- `Function.eval` as an `AlgHom`. The name matches `Pi.evalRingHom`, `Pi.evalMonoidHom`, etc. -/ @[simps] diff --git a/Mathlib/RingTheory/Ideal/QuotientOperations.lean b/Mathlib/RingTheory/Ideal/QuotientOperations.lean index 5fb624f6f2143..efc97f56b8993 100644 --- a/Mathlib/RingTheory/Ideal/QuotientOperations.lean +++ b/Mathlib/RingTheory/Ideal/QuotientOperations.lean @@ -479,6 +479,12 @@ theorem quotientMap_comp_mk {J : Ideal R} {I : Ideal S} {f : R →+* S} (H : J (quotientMap I f H).comp (Quotient.mk J) = (Quotient.mk I).comp f := RingHom.ext fun x => by simp only [Function.comp_apply, RingHom.coe_comp, Ideal.quotientMap_mk] +lemma ker_quotientMap_mk {I J : Ideal R} : + RingHom.ker (quotientMap (J.map _) (Quotient.mk I) le_comap_map) = I.map (Quotient.mk J) := by + rw [Ideal.quotientMap, Ideal.ker_quotient_lift, ← RingHom.comap_ker, Ideal.mk_ker, + Ideal.comap_map_of_surjective _ Ideal.Quotient.mk_surjective, + ← RingHom.ker_eq_comap_bot, Ideal.mk_ker, Ideal.map_sup, Ideal.map_quotient_self, bot_sup_eq] + /-- The ring equiv `R/I ≃+* S/J` induced by a ring equiv `f : R ≃+* S`, where `J = f(I)`. -/ @[simps] def quotientEquiv (I : Ideal R) (J : Ideal S) (f : R ≃+* S) (hIJ : J = I.map (f : R →+* S)) : diff --git a/Mathlib/RingTheory/Smooth/Pi.lean b/Mathlib/RingTheory/Smooth/Pi.lean new file mode 100644 index 0000000000000..acebf93edbce4 --- /dev/null +++ b/Mathlib/RingTheory/Smooth/Pi.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.Idempotents +import Mathlib.RingTheory.Smooth.Basic + +/-! + +# Formal-smoothness of finite products of rings + +## Main result + +- `Algebra.FormallySmooth.pi_iff`: If `I` is finite, `Π i : I, A i` is `R`-formally-smooth + if and only if each `A i` is `R`-formally-smooth. + +-/ + +universe u v + +namespace Algebra.FormallySmooth + +variable {R : Type (max u v)} {I : Type u} (A : I → Type (max u v)) +variable [CommRing R] [∀ i, CommRing (A i)] [∀ i, Algebra R (A i)] + +theorem of_pi [FormallySmooth R (Π i, A i)] (i) : + FormallySmooth R (A i) := by + classical + fapply FormallySmooth.of_split (Pi.evalAlgHom R A i) + apply AlgHom.ofLinearMap + ((Ideal.Quotient.mkₐ R _).toLinearMap.comp (LinearMap.single _ _ i)) + · show Ideal.Quotient.mk _ (Pi.single i 1) = 1 + rw [← (Ideal.Quotient.mk _).map_one, ← sub_eq_zero, ← map_sub, + Ideal.Quotient.eq_zero_iff_mem] + have : Pi.single i 1 - 1 ∈ RingHom.ker (Pi.evalAlgHom R A i).toRingHom := by + simp [RingHom.mem_ker] + convert neg_mem (Ideal.pow_mem_pow this 2) using 1 + simp [pow_two, sub_mul, mul_sub, ← Pi.single_mul] + · intro x y + show Ideal.Quotient.mk _ _ = Ideal.Quotient.mk _ _ * Ideal.Quotient.mk _ _ + simp only [AlgHom.toRingHom_eq_coe, LinearMap.coe_single, Pi.single_mul, map_mul] + · ext x + show (Pi.single i x) i = x + simp + +theorem pi_iff [Finite I] : + FormallySmooth R (Π i, A i) ↔ ∀ i, FormallySmooth R (A i) := by + classical + cases nonempty_fintype I + constructor + · exact fun _ ↦ of_pi A + · intro H + constructor + intros B _ _ J hJ g + have hJ' (x) (hx : x ∈ RingHom.ker (Ideal.Quotient.mk J)) : IsNilpotent x := by + refine ⟨2, show x ^ 2 ∈ (⊥ : Ideal B) from ?_⟩ + rw [← hJ] + exact Ideal.pow_mem_pow (by simpa using hx) 2 + obtain ⟨e, he, he'⟩ := ((CompleteOrthogonalIdempotents.single A).map + g.toRingHom).lift_of_isNilpotent_ker (Ideal.Quotient.mk J) hJ' + fun _ ↦ Ideal.Quotient.mk_surjective _ + replace he' : ∀ i, Ideal.Quotient.mk J (e i) = g (Pi.single i 1) := congr_fun he' + let iso : B ≃ₐ[R] ∀ i, B ⧸ Ideal.span {1 - e i} := + { __ := Pi.algHom _ _ fun i ↦ Ideal.Quotient.mkₐ R _ + __ := Equiv.ofBijective _ he.bijective_pi } + let J' := fun i ↦ J.map (Ideal.Quotient.mk (Ideal.span {1 - e i})) + let ι : ∀ i, (B ⧸ J →ₐ[R] (B ⧸ _) ⧸ J' i) := fun i ↦ Ideal.quotientMapₐ _ + (IsScalarTower.toAlgHom R B _) Ideal.le_comap_map + have hι : ∀ i x, ι i x = 0 → (e i) * x = 0 := by + intros i x hix + have : x ∈ (Ideal.span {1 - e i}).map (Ideal.Quotient.mk J) := by + rw [← Ideal.ker_quotientMap_mk]; exact hix + rw [Ideal.map_span, Set.image_singleton, Ideal.mem_span_singleton] at this + obtain ⟨c, rfl⟩ := this + rw [← mul_assoc, ← map_mul, mul_sub, mul_one, (he.idem i).eq, sub_self, map_zero, zero_mul] + have : ∀ i : I, ∃ a : A i →ₐ[R] B ⧸ Ideal.span {1 - e i}, ∀ x, + Ideal.Quotient.mk (J' i) (a x) = ι i (g (Pi.single i x)) := by + intro i + let g' : A i →ₐ[R] (B ⧸ _) ⧸ (J' i) := by + apply AlgHom.ofLinearMap (((ι i).comp g).toLinearMap ∘ₗ LinearMap.single _ _ i) + · suffices Ideal.Quotient.mk (Ideal.span {1 - e i}) (e i) = 1 by simp [ι, ← he', this] + rw [← (Ideal.Quotient.mk _).map_one, eq_comm, Ideal.Quotient.mk_eq_mk_iff_sub_mem, + Ideal.mem_span_singleton] + · intros x y; simp [Pi.single_mul] + obtain ⟨a, ha⟩ := FormallySmooth.comp_surjective (I := J' i) + (by rw [← Ideal.map_pow, hJ, Ideal.map_bot]) g' + exact ⟨a, AlgHom.congr_fun ha⟩ + choose a ha using this + use iso.symm.toAlgHom.comp (Pi.algHom _ _ fun i ↦ (a i).comp (Pi.evalAlgHom R A i)) + ext x; rw [← AlgHom.toLinearMap_apply, ← AlgHom.toLinearMap_apply]; congr 1 + ext i x + simp only [AlgEquiv.toAlgHom_eq_coe, AlgHom.comp_toLinearMap, AlgEquiv.toAlgHom_toLinearMap, + LinearMap.coe_comp, LinearMap.coe_single, Function.comp_apply, AlgHom.toLinearMap_apply, + AlgEquiv.toLinearMap_apply, Ideal.Quotient.mkₐ_eq_mk] + obtain ⟨y, hy⟩ := Ideal.Quotient.mk_surjective (a i x) + have hy' : Ideal.Quotient.mk (Ideal.span {1 - e i}) (y * e i) = a i x := by + have : Ideal.Quotient.mk (Ideal.span {1 - e i}) (e i) = 1 := by + rw [← (Ideal.Quotient.mk _).map_one, eq_comm, Ideal.Quotient.mk_eq_mk_iff_sub_mem, + Ideal.mem_span_singleton] + rw [map_mul, this, hy, mul_one] + trans Ideal.Quotient.mk J (y * e i) + · congr 1; apply iso.injective; ext j + suffices a j (Pi.single i x j) = Ideal.Quotient.mk _ (y * e i) by simpa using this + by_cases hij : i = j + · subst hij + rw [Pi.single_eq_same, hy'] + · have : Ideal.Quotient.mk (Ideal.span {1 - e j}) (e i) = 0 := by + rw [Ideal.Quotient.eq_zero_iff_mem, Ideal.mem_span_singleton] + refine ⟨e i, by simp [he.ortho (Ne.symm hij), sub_mul]⟩ + rw [Pi.single_eq_of_ne (Ne.symm hij), map_zero, map_mul, this, mul_zero] + · have : ι i (Ideal.Quotient.mk J (y * e i)) = ι i (g (Pi.single i x)) := by + rw [← ha, ← hy'] + simp only [Ideal.quotient_map_mkₐ, IsScalarTower.coe_toAlgHom', + Ideal.Quotient.algebraMap_eq, Ideal.Quotient.mkₐ_eq_mk, ι] + rw [← sub_eq_zero, ← map_sub] at this + replace this := hι _ _ this + rwa [mul_sub, ← map_mul, mul_comm, mul_assoc, (he.idem i).eq, he', ← map_mul, ← Pi.single_mul, + one_mul, sub_eq_zero] at this + +end Algebra.FormallySmooth From e2b89093a1a30606a1de57c265dd8d8d33d3ae2b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 14:30:27 +0000 Subject: [PATCH 002/131] chore: don't export `GroupWithZero.inv_zero` (#17465) Exporting makes it not show as `inv_zero` in the docs search results, meaning it's very non-discoverable. See https://github.com/leanprover/doc-gen4/issues/222. --- Mathlib/Algebra/GroupWithZero/Defs.lean | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Defs.lean b/Mathlib/Algebra/GroupWithZero/Defs.lean index 14ace092e132d..221f060ebda9a 100644 --- a/Mathlib/Algebra/GroupWithZero/Defs.lean +++ b/Mathlib/Algebra/GroupWithZero/Defs.lean @@ -182,16 +182,15 @@ Examples include division rings and the ordered monoids that are the target of valuations in general valuation theory. -/ class GroupWithZero (G₀ : Type u) extends MonoidWithZero G₀, DivInvMonoid G₀, Nontrivial G₀ where /-- The inverse of `0` in a group with zero is `0`. -/ - inv_zero : (0 : G₀)⁻¹ = 0 + protected inv_zero : (0 : G₀)⁻¹ = 0 /-- Every nonzero element of a group with zero is invertible. -/ protected mul_inv_cancel (a : G₀) : a ≠ 0 → a * a⁻¹ = 1 -export GroupWithZero (inv_zero) -attribute [simp] inv_zero - section GroupWithZero variable [GroupWithZero G₀] {a : G₀} +@[simp] lemma inv_zero : (0 : G₀)⁻¹ = 0 := GroupWithZero.inv_zero + @[simp] lemma mul_inv_cancel₀ (h : a ≠ 0) : a * a⁻¹ = 1 := GroupWithZero.mul_inv_cancel a h -- See note [lower instance priority] From 58282561d2e928698a6727596a4ad96d74f64766 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 14:30:28 +0000 Subject: [PATCH 003/131] feat: Lagrange's theorem for the pointwise product of a subgroup with a set (#17680) From LeanCamCombi --- Mathlib/GroupTheory/Coset/Card.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/GroupTheory/Coset/Card.lean b/Mathlib/GroupTheory/Coset/Card.lean index 67a686a06e9b1..df9e2e1b9a667 100644 --- a/Mathlib/GroupTheory/Coset/Card.lean +++ b/Mathlib/GroupTheory/Coset/Card.lean @@ -10,6 +10,8 @@ import Mathlib.SetTheory.Cardinal.Finite # Lagrange's theorem: the order of a subgroup divides the order of the group. -/ +open scoped Pointwise + namespace Subgroup variable {α : Type*} [Group α] @@ -19,6 +21,15 @@ theorem card_eq_card_quotient_mul_card_subgroup (s : Subgroup α) : Nat.card α = Nat.card (α ⧸ s) * Nat.card s := by rw [← Nat.card_prod]; exact Nat.card_congr Subgroup.groupEquivQuotientProdSubgroup +@[to_additive] +lemma card_mul_eq_card_subgroup_mul_card_quotient (s : Subgroup α) (t : Set α) : + Nat.card (t * s : Set α) = Nat.card s * Nat.card (t.image (↑) : Set (α ⧸ s)) := by + rw [← Nat.card_prod, Nat.card_congr] + apply Equiv.trans _ (QuotientGroup.preimageMkEquivSubgroupProdSet _ _) + rw [QuotientGroup.preimage_image_mk] + convert Equiv.refl ↑(t * s) + aesop (add simp [Set.mem_mul]) + /-- **Lagrange's Theorem**: The order of a subgroup divides the order of its ambient group. -/ @[to_additive "**Lagrange's Theorem**: The order of an additive subgroup divides the order of its ambient additive group."] From 436ba3afd94da896fa1d20b511a1469ea5d018c3 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 14:30:30 +0000 Subject: [PATCH 004/131] chore: remove "instance was not necessary" porting notes (issue #10670) (#17866) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Discussed with @kim-em. These instances are now necessary because instance synthesis now searches with the full discrimination tree, so mismatches in the arguments to instances are scrutinized much harder. This is not likely to be reverted in the future, so the instances remain required. There were a few further occurrences of #10670 for Grp, see e.g. [Grp.instCoeFunHomForallαGroup](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Algebra/Category/Grp/Basic.html#Grp.instCoeFunHomForall%CE%B1Group). These do look like they are unnecessary but removing them breaks a lot of downstream proofs. (I think it's something to do with `coe_of` unfolding `of` causing a mismatch between the expected types, but it's a bit too subtle to deal with easily.) I'll leave those porting notes for the meantime. We need a big refactor cleaning up the concrete category bits of the library anyway. --- Mathlib/Algebra/Category/Grp/Basic.lean | 2 -- Mathlib/Algebra/Category/Grp/Preadditive.lean | 1 - Mathlib/Geometry/RingedSpace/Basic.lean | 1 - Mathlib/Order/Category/NonemptyFinLinOrd.lean | 1 - 4 files changed, 5 deletions(-) diff --git a/Mathlib/Algebra/Category/Grp/Basic.lean b/Mathlib/Algebra/Category/Grp/Basic.lean index 67d5908ffff48..1c03372df70b3 100644 --- a/Mathlib/Algebra/Category/Grp/Basic.lean +++ b/Mathlib/Algebra/Category/Grp/Basic.lean @@ -112,7 +112,6 @@ instance hasForgetToMonCat : HasForget₂ Grp MonCat := @[to_additive] instance : Coe Grp.{u} MonCat.{u} where coe := (forget₂ Grp MonCat).obj --- porting note (#10670): this instance was not necessary in mathlib @[to_additive] instance (G H : Grp) : One (G ⟶ H) := (inferInstance : One (MonoidHom G H)) @@ -257,7 +256,6 @@ instance hasForgetToCommMonCat : HasForget₂ CommGrp CommMonCat := @[to_additive] instance : Coe CommGrp.{u} CommMonCat.{u} where coe := (forget₂ CommGrp CommMonCat).obj --- porting note (#10670): this instance was not necessary in mathlib @[to_additive] instance (G H : CommGrp) : One (G ⟶ H) := (inferInstance : One (MonoidHom G H)) diff --git a/Mathlib/Algebra/Category/Grp/Preadditive.lean b/Mathlib/Algebra/Category/Grp/Preadditive.lean index 5866733e35b85..fd7ca0d120ef3 100644 --- a/Mathlib/Algebra/Category/Grp/Preadditive.lean +++ b/Mathlib/Algebra/Category/Grp/Preadditive.lean @@ -16,7 +16,6 @@ universe u namespace AddCommGrp --- porting note (#10670): this instance was not necessary in mathlib instance (P Q : AddCommGrp) : AddCommGroup (P ⟶ Q) := (inferInstance : AddCommGroup (AddMonoidHom P Q)) diff --git a/Mathlib/Geometry/RingedSpace/Basic.lean b/Mathlib/Geometry/RingedSpace/Basic.lean index 4a89dde97d9c9..a11af973bc890 100644 --- a/Mathlib/Geometry/RingedSpace/Basic.lean +++ b/Mathlib/Geometry/RingedSpace/Basic.lean @@ -48,7 +48,6 @@ lemma res_zero {X : RingedSpace.{u}} {U V : TopologicalSpace.Opens X} variable (X : RingedSpace) --- Porting note (#10670): this was not necessary in mathlib3 instance : CoeSort RingedSpace Type* where coe X := X.carrier diff --git a/Mathlib/Order/Category/NonemptyFinLinOrd.lean b/Mathlib/Order/Category/NonemptyFinLinOrd.lean index cfa1dab425200..29ef2f1dc61d2 100644 --- a/Mathlib/Order/Category/NonemptyFinLinOrd.lean +++ b/Mathlib/Order/Category/NonemptyFinLinOrd.lean @@ -121,7 +121,6 @@ instance {A B : NonemptyFinLinOrd.{u}} : FunLike (A ⟶ B) A B where ext x exact congr_fun h x --- porting note (#10670): this instance was not necessary in mathlib instance {A B : NonemptyFinLinOrd.{u}} : OrderHomClass (A ⟶ B) A B where map_rel f _ _ h := f.monotone h From 12a15447c89cc5bbbfa031af07a954ae31b3bbdf Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 14:30:31 +0000 Subject: [PATCH 005/131] chore: remove "CoeFun -> FunLike" porting/adaptation notes (#17871) We prefer `FunLike` over `CoeFun` if possible, so we don't care that we made that change during the port. See also #17870 for applying the change that was porting noted consistently across the entirety of Mathlib. --- Mathlib/Algebra/Module/LinearMap/Defs.lean | 2 -- Mathlib/Algebra/MvPolynomial/Basic.lean | 2 -- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 1 - Mathlib/LinearAlgebra/Multilinear/Basic.lean | 1 - Mathlib/NumberTheory/ArithmeticFunction.lean | 1 - Mathlib/Order/Heyting/Hom.lean | 7 ------- Mathlib/Order/Hom/CompleteLattice.lean | 6 ------ Mathlib/RingTheory/RingInvo.lean | 4 ---- Mathlib/Topology/Algebra/Module/Basic.lean | 9 --------- Mathlib/Topology/Category/TopCat/Basic.lean | 6 ------ 10 files changed, 39 deletions(-) diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index eceefd6b766ab..0a75f33f5bf9c 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -216,8 +216,6 @@ lemma coe_coe {F : Type*} [FunLike F M M₃] [SemilinearMapClass F σ M M₃] {f ⇑(f : M →ₛₗ[σ] M₃) = f := rfl --- Porting note: we don't port specialized `CoeFun` instances if there is `DFunLike` instead - /-- The `DistribMulActionHom` underlying a `LinearMap`. -/ def toDistribMulActionHom (f : M →ₛₗ[σ] M₃) : DistribMulActionHom σ.toMonoidHom M M₃ := { f with map_zero' := show f 0 = 0 from map_zero f } diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index c5e9e3036cee4..de4f2a12c09bb 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -509,8 +509,6 @@ section Coeff /-- The coefficient of the monomial `m` in the multi-variable polynomial `p`. -/ def coeff (m : σ →₀ ℕ) (p : MvPolynomial σ R) : R := @DFunLike.coe ((σ →₀ ℕ) →₀ R) _ _ _ p m - -- Porting note: I changed this from `@CoeFun.coe _ _ (MonoidAlgebra.coeFun _ _) p m` because - -- I think it should work better syntactically. They are defeq. @[simp] theorem mem_support_iff {p : MvPolynomial σ R} {m : σ →₀ ℕ} : m ∈ p.support ↔ p.coeff m ≠ 0 := by diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 5fb4aa1c20293..2c14b91ede9f1 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -306,7 +306,6 @@ theorem repr_injective : cases g congr --- Porting note: `CoeFun` → `FunLike` /-- `b i` is the `i`th basis vector. -/ instance instFunLike : FunLike (OrthonormalBasis ι 𝕜 E) ι E where coe b i := by classical exact b.repr.symm (EuclideanSpace.single i (1 : 𝕜)) diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 21a0244972b27..5fa2ebd1c7ff4 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -106,7 +106,6 @@ variable [Semiring R] [∀ i, AddCommMonoid (M i)] [∀ i, AddCommMonoid (M₁ i [AddCommMonoid M₃] [AddCommMonoid M'] [∀ i, Module R (M i)] [∀ i, Module R (M₁ i)] [Module R M₂] [Module R M₃] [Module R M'] (f f' : MultilinearMap R M₁ M₂) --- Porting note: Replaced CoeFun with FunLike instance instance : FunLike (MultilinearMap R M₁ M₂) (∀ i, M₁ i) M₂ where coe f := f.toFun coe_injective' f g h := by cases f; cases g; cases h; rfl diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 9e6e2e67edaef..71c09ddd1baad 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -92,7 +92,6 @@ section Zero variable [Zero R] --- porting note: used to be `CoeFun` instance : FunLike (ArithmeticFunction R) ℕ R := inferInstanceAs (FunLike (ZeroHom ℕ R) ℕ R) diff --git a/Mathlib/Order/Heyting/Hom.lean b/Mathlib/Order/Heyting/Hom.lean index 6b0940817f6d3..ab4677d3ab3cb 100644 --- a/Mathlib/Order/Heyting/Hom.lean +++ b/Mathlib/Order/Heyting/Hom.lean @@ -246,13 +246,6 @@ instance instHeytingHomClass : HeytingHomClass (HeytingHom α β) α β where map_bot f := f.map_bot' map_himp := HeytingHom.map_himp' - --- Porting note: CoeFun undesired here in lean 4 --- /-- Helper instance for when there's too many metavariables to apply `DFunLike.CoeFun` --- directly. -/ --- instance : CoeFun (HeytingHom α β) fun _ => α → β := --- DFunLike.hasCoeToFun - -- @[simp] -- Porting note: not in simp-nf, simp can simplify lhs. Added aux simp lemma theorem toFun_eq_coe {f : HeytingHom α β} : f.toFun = ⇑f := rfl diff --git a/Mathlib/Order/Hom/CompleteLattice.lean b/Mathlib/Order/Hom/CompleteLattice.lean index 8fecfd2875b9b..125753699728c 100644 --- a/Mathlib/Order/Hom/CompleteLattice.lean +++ b/Mathlib/Order/Hom/CompleteLattice.lean @@ -592,12 +592,6 @@ def tosSupHom (f : CompleteLatticeHom α β) : sSupHom α β := def toBoundedLatticeHom (f : CompleteLatticeHom α β) : BoundedLatticeHom α β := f --- Porting note: We do not want CoeFun for this in lean 4 --- /-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_toFun` --- directly. -/ --- instance : CoeFun (CompleteLatticeHom α β) fun _ => α → β := --- DFunLike.hasCoeToFun - lemma toFun_eq_coe (f : CompleteLatticeHom α β) : f.toFun = f := rfl @[simp] lemma coe_tosInfHom (f : CompleteLatticeHom α β) : ⇑f.tosInfHom = f := rfl diff --git a/Mathlib/RingTheory/RingInvo.lean b/Mathlib/RingTheory/RingInvo.lean index 0ebf763f2a467..86cbcf2eb18f4 100644 --- a/Mathlib/RingTheory/RingInvo.lean +++ b/Mathlib/RingTheory/RingInvo.lean @@ -83,10 +83,6 @@ def mk' (f : R →+* Rᵐᵒᵖ) (involution : ∀ r, (f (f r).unop).unop = r) : right_inv := fun _ => MulOpposite.unop_injective <| involution _ involution' := involution } --- Porting note: removed CoeFun instance, undesired in lean4 --- instance : CoeFun (RingInvo R) fun _ => R → Rᵐᵒᵖ := --- ⟨fun f => f.toRingEquiv.toFun⟩ - @[simp] theorem involution (f : RingInvo R) (x : R) : (f (f x).unop).unop = x := f.involution' x diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 451dff66e30df..df022af891410 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -378,10 +378,6 @@ instance continuousSemilinearMapClass : map_continuous f := f.2 map_smulₛₗ f := f.toLinearMap.map_smul' --- see Note [function coercion] -/-- Coerce continuous linear maps to functions. -/ ---instance toFun' : CoeFun (M₁ →SL[σ₁₂] M₂) fun _ => M₁ → M₂ := ⟨DFunLike.coe⟩ - -- porting note (#10618): was `simp`, now `simp only` proves it theorem coe_mk (f : M₁ →ₛₗ[σ₁₂] M₂) (h) : (mk f h : M₁ →ₛₗ[σ₁₂] M₂) = f := rfl @@ -1634,11 +1630,6 @@ instance continuousSemilinearEquivClass : map_continuous := continuous_toFun inv_continuous := continuous_invFun --- see Note [function coercion] --- /-- Coerce continuous linear equivs to maps. -/ --- instance : CoeFun (M₁ ≃SL[σ₁₂] M₂) fun _ => M₁ → M₂ := --- ⟨fun f => f⟩ - theorem coe_apply (e : M₁ ≃SL[σ₁₂] M₂) (b : M₁) : (e : M₁ →SL[σ₁₂] M₂) b = e b := rfl diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 468168d5c7fb5..8cf13a2f355f4 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -47,12 +47,6 @@ instance : CoeSort TopCat Type* where instance topologicalSpaceUnbundled (X : TopCat) : TopologicalSpace X := X.str --- We leave this temporarily as a reminder of the downstream instances #13170 --- -- Porting note: cannot find a coercion to function otherwise --- -- attribute [instance] ConcreteCategory.instFunLike in --- instance (X Y : TopCat.{u}) : CoeFun (X ⟶ Y) fun _ => X → Y where --- coe (f : C(X, Y)) := f - instance instFunLike (X Y : TopCat) : FunLike (X ⟶ Y) X Y := inferInstanceAs <| FunLike C(X, Y) X Y From d0df6356048affd1a0c86f935b0d04bf36a402d4 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Thu, 17 Oct 2024 15:03:58 +0000 Subject: [PATCH 006/131] feat(QuotientGroup,Ideal/Quotient): group is finite if quotient is finite (#15425) Does not actually require the subgroup to be normal Also provides a breakdown of the group as the disjoint indexed union of cosets Co-authored-by: Yakov Pechersky --- Mathlib/GroupTheory/Coset/Basic.lean | 21 +++++++++++++++++ Mathlib/GroupTheory/GroupAction/Basic.lean | 23 +++++++++++++++++++ Mathlib/GroupTheory/QuotientGroup/Finite.lean | 5 ++++ Mathlib/RingTheory/Ideal/Quotient.lean | 11 +++++++++ 4 files changed, 60 insertions(+) diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index d9ede0a36ed67..ab5461c4022dc 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -458,6 +458,18 @@ theorem preimage_image_mk_eq_mul (N : Subgroup α) (s : Set α) : rw [preimage_image_mk_eq_iUnion_image, iUnion_subtype, ← image2_mul, ← iUnion_image_right] simp only [SetLike.mem_coe] +open MulAction in +@[to_additive] +lemma orbit_mk_eq_smul (x : α) : MulAction.orbitRel.Quotient.orbit (x : α ⧸ s) = x • s := by + ext + rw [orbitRel.Quotient.mem_orbit] + simpa [mem_smul_set_iff_inv_smul_mem, ← leftRel_apply] using Setoid.comm' _ + +@[to_additive] +lemma orbit_eq_out'_smul (x : α ⧸ s) : MulAction.orbitRel.Quotient.orbit x = x.out' • s := by + induction x using QuotientGroup.induction_on + simp only [orbit_mk_eq_smul, ← eq_class_eq_leftCoset, Quotient.out_eq'] + end QuotientGroup namespace Subgroup @@ -706,4 +718,13 @@ noncomputable def preimageMkEquivSubgroupProdSet (s : Subgroup α) (t : Set (α left_inv := fun ⟨a, _⟩ => Subtype.eq <| show _ * _ = a by simp right_inv := fun ⟨⟨a, ha⟩, ⟨x, hx⟩⟩ => by ext <;> simp [ha] +open MulAction in +/-- A group is made up of a disjoint union of cosets of a subgroup. -/ +@[to_additive "An additive group is made up of a disjoint union of cosets of an additive +subgroup."] +lemma univ_eq_iUnion_smul (H : Subgroup α) : + (Set.univ (α := α)) = ⋃ x : α ⧸ H, x.out' • (H : Set _) := by + simp_rw [univ_eq_iUnion_orbit H.op, orbit_eq_out'_smul] + rfl + end QuotientGroup diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index cc847556ec7f8..9f99144a7769a 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -108,6 +108,10 @@ lemma snd_mem_orbit_of_mem_orbit {x y : α × β} (h : x ∈ MulAction.orbit M y rcases h with ⟨g, rfl⟩ exact mem_orbit _ _ +@[to_additive] +lemma _root_.Finite.finite_mulAction_orbit [Finite M] (a : α) : Set.Finite (orbit M a) := + Set.finite_range _ + variable (M) @[to_additive] @@ -625,6 +629,25 @@ def selfEquivSigmaOrbits : α ≃ Σω : Ω, orbit G ω.out' := Equiv.sigmaCongrRight fun _ => Equiv.Set.ofEq <| orbitRel.Quotient.orbit_eq_orbit_out _ Quotient.out_eq' +/-- Decomposition of a type `X` as a disjoint union of its orbits under a group action. +Phrased as a set union. See `MulAction.selfEquivSigmaOrbits` for the type isomorphism. -/ +@[to_additive "Decomposition of a type `X` as a disjoint union of its orbits under an additive group +action. Phrased as a set union. See `AddAction.selfEquivSigmaOrbits` for the type isomorphism."] +lemma univ_eq_iUnion_orbit : + Set.univ (α := α) = ⋃ x : Ω, x.orbit := by + ext x + simp only [Set.mem_univ, Set.mem_iUnion, true_iff] + exact ⟨Quotient.mk'' x, by simp⟩ + +@[to_additive] +lemma _root_.Finite.of_finite_mulAction_orbitRel_quotient [Finite G] [Finite Ω] : Finite α := by + rw [(selfEquivSigmaOrbits' G _).finite_iff] + have : ∀ g : Ω, Finite g.orbit := by + intro g + induction g using Quotient.inductionOn' + simpa [Set.finite_coe_iff] using Finite.finite_mulAction_orbit _ + exact Finite.instSigma + variable (β) @[to_additive] diff --git a/Mathlib/GroupTheory/QuotientGroup/Finite.lean b/Mathlib/GroupTheory/QuotientGroup/Finite.lean index 76877f51f87b7..6d234fdd50c16 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Finite.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Finite.lean @@ -50,4 +50,9 @@ noncomputable def fintypeOfKerOfCodom [Fintype g.ker] : Fintype G := noncomputable def fintypeOfDomOfCoker [Normal f.range] [Fintype <| G ⧸ f.range] : Fintype G := fintypeOfKerLeRange _ (mk' f.range) fun x => (eq_one_iff x).mp +@[to_additive] +lemma _root_.Finite.of_finite_quot_finite_subgroup {H : Subgroup G} [Finite H] [Finite (G ⧸ H)] : + Finite G := + Finite.of_equiv _ (groupEquivQuotientProdSubgroup (s := H)).symm + end Group diff --git a/Mathlib/RingTheory/Ideal/Quotient.lean b/Mathlib/RingTheory/Ideal/Quotient.lean index 3889f7377adb0..182a09a1dac95 100644 --- a/Mathlib/RingTheory/Ideal/Quotient.lean +++ b/Mathlib/RingTheory/Ideal/Quotient.lean @@ -3,6 +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, Anne Baanen -/ +import Mathlib.GroupTheory.QuotientGroup.Finite import Mathlib.LinearAlgebra.Quotient import Mathlib.RingTheory.Congruence.Basic import Mathlib.RingTheory.Ideal.Basic @@ -348,4 +349,14 @@ theorem map_pi {ι : Type*} [Finite ι] {ι' : Type w} (x : ι → R) (hi : ∀ end Pi +open scoped Pointwise in +/-- A ring is made up of a disjoint union of cosets of an ideal. -/ +lemma univ_eq_iUnion_image_add {R : Type*} [Ring R] (I : Ideal R) : + (Set.univ (α := R)) = ⋃ x : R ⧸ I, x.out' +ᵥ (I : Set R) := + QuotientAddGroup.univ_eq_iUnion_vadd I.toAddSubgroup + +lemma _root_.Finite.of_finite_quot_finite_ideal {R : Type*} [Ring R] {I : Ideal R} + [hI : Finite I] [h : Finite (R ⧸ I)] : Finite R := + @Finite.of_finite_quot_finite_addSubgroup _ _ _ hI h + end Ideal From 0b86556786ec68a8d3c198f66952df4d3b771cb2 Mon Sep 17 00:00:00 2001 From: FR Date: Thu, 17 Oct 2024 15:03:59 +0000 Subject: [PATCH 007/131] chore(SetTheory/Game/Basic): fix docs (#17121) --- Mathlib/SetTheory/Game/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/SetTheory/Game/Basic.lean b/Mathlib/SetTheory/Game/Basic.lean index 13900670705bd..3504ce49ac81a 100644 --- a/Mathlib/SetTheory/Game/Basic.lean +++ b/Mathlib/SetTheory/Game/Basic.lean @@ -247,7 +247,7 @@ Hence we define them here. -/ /-- The product of `x = {xL | xR}` and `y = {yL | yR}` is -`{xL*y + x*yL - xL*yL, xR*y + x*yR - xR*yR | xL*y + x*yR - xL*yR, x*yL + xR*y - xR*yL }`. -/ +`{xL*y + x*yL - xL*yL, xR*y + x*yR - xR*yR | xL*y + x*yR - xL*yR, xR*y + x*yL - xR*yL}`. -/ instance : Mul PGame.{u} := ⟨fun x y => by induction x generalizing y with | mk xl xr _ _ IHxl IHxr => _ From cff08d545656d8cabf918cdcbde50f4f6105ae0b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:04:01 +0000 Subject: [PATCH 008/131] chore(Order/Hom): let `gcongr` know about `OrderEmbedding` and `OrderIso` (#17224) From LeanAPAP --- Mathlib/NumberTheory/NumberField/Discriminant.lean | 2 +- Mathlib/Order/Hom/Basic.lean | 7 +++++++ 2 files changed, 8 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index 758d8a99b48ed..07199e222dcb4 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -308,7 +308,7 @@ theorem minkowskiBound_lt_boundOfDiscBdd : minkowskiBound K ↑1 < boundOfDiscBd coe_mul, ENNReal.coe_pow, coe_ofNat, show sqrt N = (1 : ℝ≥0∞) * sqrt N by rw [one_mul]] gcongr · exact pow_le_one₀ (by positivity) (by norm_num) - · rwa [sqrt_le_sqrt, ← NNReal.coe_le_coe, coe_nnnorm, Int.norm_eq_abs, ← Int.cast_abs, + · rwa [← NNReal.coe_le_coe, coe_nnnorm, Int.norm_eq_abs, ← Int.cast_abs, NNReal.coe_natCast, ← Int.cast_natCast, Int.cast_le] · exact one_le_two · exact rank_le_rankOfDiscrBdd hK diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index 96b44c3f6a264..c9ab0123359da 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -148,6 +148,9 @@ protected theorem monotone (f : F) : Monotone f := fun _ _ => map_rel f protected theorem mono (f : F) : Monotone f := fun _ _ => map_rel f +@[gcongr] protected lemma GCongr.mono (f : F) {a b : α} (hab : a ≤ b) : f a ≤ f b := + OrderHomClass.mono f hab + /-- Turn an element of a type `F` satisfying `OrderHomClass F α β` into an actual `OrderHom`. This is declared as the default coercion from `F` to `α →o β`. -/ @[coe] @@ -921,6 +924,8 @@ variable [LE α] [LE β] theorem le_iff_le (e : α ≃o β) {x y : α} : e x ≤ e y ↔ x ≤ y := e.map_rel_iff +@[gcongr] protected alias ⟨_, GCongr.orderIso_apply_le_apply⟩ := le_iff_le + theorem le_symm_apply (e : α ≃o β) {x : α} {y : β} : x ≤ e.symm y ↔ e x ≤ y := e.rel_symm_apply @@ -941,6 +946,8 @@ protected theorem strictMono (e : α ≃o β) : StrictMono e := theorem lt_iff_lt (e : α ≃o β) {x y : α} : e x < e y ↔ x < y := e.toOrderEmbedding.lt_iff_lt +@[gcongr] protected alias ⟨_, GCongr.orderIso_apply_lt_apply⟩ := lt_iff_lt + /-- Converts an `OrderIso` into a `RelIso (<) (<)`. -/ def toRelIsoLT (e : α ≃o β) : ((· < ·) : α → α → Prop) ≃r ((· < ·) : β → β → Prop) := ⟨e.toEquiv, lt_iff_lt e⟩ From cc8cf2ffff4f61b00d30cd82f66d1728e95dc4c3 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Thu, 17 Oct 2024 15:04:03 +0000 Subject: [PATCH 009/131] feat: add `--force` flag to shake (#17453) to disable the `lake build --no-build` step. There have been [reports](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/.60lake.20build.20--no-build.60.20at.20.60v4.2E13.2E0-rc3.60.20is.20broken.3F/near/474875999) that this step is somewhat flaky, so this option gives people the ability to skip it without turning off `shake` entirely. Co-authored-by: Mario Carneiro --- Shake/Main.lean | 13 ++++++++++--- 1 file changed, 10 insertions(+), 3 deletions(-) diff --git a/Shake/Main.lean b/Shake/Main.lean index e5b7645a31d81..20a8f24e116f0 100644 --- a/Shake/Main.lean +++ b/Shake/Main.lean @@ -35,6 +35,9 @@ Arguments: provided module(s) will be checked. Options: + --force + Skips the `lake build --no-build` sanity check + --fix Apply the suggested fixes directly. Make sure you have a clean checkout before running this, so you can review the changes. @@ -380,6 +383,8 @@ def toBitset (s : State) (ns : List Name) : Bitset := structure Args where /-- `--help`: shows the help -/ help : Bool := false + /-- `--force`: skips the `lake build --no-build` sanity check -/ + force : Bool := false /-- `--no-downstream`: disables downstream mode -/ downstream : Bool := true /-- `--gh-style`: output messages that can be parsed by `gh-problem-matcher-wrap` -/ @@ -422,6 +427,7 @@ def main (args : List String) : IO UInt32 := do let rec parseArgs (args : Args) : List String → Args | [] => args | "--help" :: rest => parseArgs { args with help := true } rest + | "--force" :: rest => parseArgs { args with force := true } rest | "--no-downstream" :: rest => parseArgs { args with downstream := false } rest | "--fix" :: rest => parseArgs { args with fix := true } rest | "--explain" :: rest => parseArgs { args with explain := true } rest @@ -438,9 +444,10 @@ def main (args : List String) : IO UInt32 := do IO.println help IO.Process.exit 0 - if (← IO.Process.output { cmd := "lake", args := #["build", "--no-build"] }).exitCode != 0 then - IO.println "There are out of date oleans. Run `lake build` or `lake exe cache get` first" - IO.Process.exit 1 + if !args.force then + if (← IO.Process.output { cmd := "lake", args := #["build", "--no-build"] }).exitCode != 0 then + IO.println "There are out of date oleans. Run `lake build` or `lake exe cache get` first" + IO.Process.exit 1 -- Parse the `--cfg` argument let srcSearchPath ← initSrcSearchPath From 10e0d7deb5528d3e0eb8331b42f15f24b0559767 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 17 Oct 2024 15:04:04 +0000 Subject: [PATCH 010/131] feat: seminorms for the product topology (#17493) --- Mathlib/Analysis/LocallyConvex/WithSeminorms.lean | 13 +++++++++++-- 1 file changed, 11 insertions(+), 2 deletions(-) diff --git a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean index dff476482e167..5238beff83992 100644 --- a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean +++ b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean @@ -877,13 +877,22 @@ protected def SeminormFamily.sigma {κ : ι → Type*} (p : (i : ι) → Seminor theorem withSeminorms_iInf {κ : ι → Type*} [Nonempty ((i : ι) × κ i)] [∀ i, Nonempty (κ i)] {p : (i : ι) → SeminormFamily 𝕜 E (κ i)} {t : ι → TopologicalSpace E} - [∀ i, @TopologicalAddGroup E (t i) _] (hp : ∀ i, WithSeminorms (topology := t i) (p i)) : + (hp : ∀ i, WithSeminorms (topology := t i) (p i)) : WithSeminorms (topology := ⨅ i, t i) (SeminormFamily.sigma p) := by - have : @TopologicalAddGroup E (⨅ i, t i) _ := topologicalAddGroup_iInf (fun i ↦ inferInstance) + have : ∀ i, @TopologicalAddGroup E (t i) _ := + fun i ↦ @WithSeminorms.topologicalAddGroup _ _ _ _ _ _ _ (t i) _ (hp i) + have : @TopologicalAddGroup E (⨅ i, t i) _ := topologicalAddGroup_iInf inferInstance simp_rw [@SeminormFamily.withSeminorms_iff_topologicalSpace_eq_iInf _ _ _ _ _ _ _ (_)] at hp ⊢ rw [iInf_sigma] exact iInf_congr hp +theorem withSeminorms_pi {κ : ι → Type*} {E : ι → Type*} + [∀ i, AddCommGroup (E i)] [∀ i, Module 𝕜 (E i)] [∀ i, TopologicalSpace (E i)] + [Nonempty ((i : ι) × κ i)] [∀ i, Nonempty (κ i)] {p : (i : ι) → SeminormFamily 𝕜 (E i) (κ i)} + (hp : ∀ i, WithSeminorms (p i)) : + WithSeminorms (SeminormFamily.sigma (fun i ↦ (p i).comp (LinearMap.proj i))) := + withSeminorms_iInf fun i ↦ (LinearMap.proj i).withSeminorms_induced (hp i) + end TopologicalConstructions section TopologicalProperties From 9dd0528bdc2e9987bc51e623517606202d56f143 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 17 Oct 2024 15:04:05 +0000 Subject: [PATCH 011/131] feat(GroupTheory/Coset/Basic): products, `leftRel` and `rightRel` (#17585) Add lemmas relating `leftRel` and `rightRel` for a (pairwise or indexed) product of subgroups to the corresponding products of setoids. From AperiodicMonotilesLean. --- Mathlib/GroupTheory/Coset/Basic.lean | 28 ++++++++++++++++++++++++++++ 1 file changed, 28 insertions(+) diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index ab5461c4022dc..e7aa809a354ae 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -262,6 +262,20 @@ instance leftRelDecidable [DecidablePred (· ∈ s)] : DecidableRel (leftRel s). rw [leftRel_eq] exact ‹DecidablePred (· ∈ s)› _ +@[to_additive] +lemma leftRel_prod {β : Type*} [Group β] (s' : Subgroup β) : + leftRel (s.prod s') = (leftRel s).prod (leftRel s') := by + refine Setoid.ext fun x y ↦ ?_ + rw [Setoid.prod_apply] + simp_rw [leftRel_apply] + rfl + +@[to_additive] +lemma leftRel_pi {ι : Type*} {β : ι → Type*} [∀ i, Group (β i)] (s' : ∀ i, Subgroup (β i)) : + leftRel (Subgroup.pi Set.univ s') = @piSetoid _ _ fun i ↦ leftRel (s' i) := by + refine Setoid.ext fun x y ↦ ?_ + simp [Setoid.piSetoid_apply, leftRel_apply, Subgroup.mem_pi] + /-- `α ⧸ s` is the quotient type representing the left cosets of `s`. If `s` is a normal subgroup, `α ⧸ s` is a group -/ @[to_additive "`α ⧸ s` is the quotient type representing the left cosets of `s`. If `s` is a normal @@ -304,6 +318,20 @@ instance rightRelDecidable [DecidablePred (· ∈ s)] : DecidableRel (rightRel s rw [rightRel_eq] exact ‹DecidablePred (· ∈ s)› _ +@[to_additive] +lemma rightRel_prod {β : Type*} [Group β] (s' : Subgroup β) : + rightRel (s.prod s') = (rightRel s).prod (rightRel s') := by + refine Setoid.ext fun x y ↦ ?_ + rw [Setoid.prod_apply] + simp_rw [rightRel_apply] + rfl + +@[to_additive] +lemma rightRel_pi {ι : Type*} {β : ι → Type*} [∀ i, Group (β i)] (s' : ∀ i, Subgroup (β i)) : + rightRel (Subgroup.pi Set.univ s') = @piSetoid _ _ fun i ↦ rightRel (s' i) := by + refine Setoid.ext fun x y ↦ ?_ + simp [Setoid.piSetoid_apply, rightRel_apply, Subgroup.mem_pi] + /-- Right cosets are in bijection with left cosets. -/ @[to_additive "Right cosets are in bijection with left cosets."] def quotientRightRelEquivQuotientLeftRel : Quotient (QuotientGroup.rightRel s) ≃ α ⧸ s where From f66852471cf4c2c6cb3b707dc1e9be6929360662 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:04:07 +0000 Subject: [PATCH 012/131] feat: translating the graph of a homomorphism (#17647) From PFR --- Mathlib/Data/Set/Pointwise/SMul.lean | 26 ++++++++++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index e3ed79c944b46..00fe29ebfea63 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -406,6 +406,32 @@ lemma smul_set_disjoint_iff : Disjoint (a • s) (a • t) ↔ Disjoint s t := b end Group +section Group +variable [Group α] [CommGroup β] [FunLike F α β] [MonoidHomClass F α β] + +@[to_additive] +lemma smul_graphOn (x : α × β) (s : Set α) (f : F) : + x • s.graphOn f = (x.1 • s).graphOn fun a ↦ x.2 / f x.1 * f a := by + ext ⟨a, b⟩ + simp [mem_smul_set_iff_inv_smul_mem, Prod.ext_iff, and_comm (a := _ = a), inv_mul_eq_iff_eq_mul, + mul_left_comm _ _⁻¹, eq_inv_mul_iff_mul_eq, ← mul_div_right_comm, div_eq_iff_eq_mul, mul_comm b] + +@[to_additive] +lemma smul_graphOn_univ (x : α × β) (f : F) : + x • univ.graphOn f = univ.graphOn fun a ↦ x.2 / f x.1 * f a := by simp [smul_graphOn] + +end Group + +section CommGroup +variable [CommGroup α] + +@[to_additive] lemma smul_div_smul_comm (a : α) (s : Set α) (b : α) (t : Set α) : + a • s / b • t = (a / b) • (s / t) := by + simp_rw [← image_smul, smul_eq_mul, ← singleton_mul, mul_div_mul_comm _ s, + singleton_div_singleton] + +end CommGroup + section GroupWithZero variable [GroupWithZero α] [MulAction α β] {s t : Set β} {a : α} From 1cfdb12acbfe9bb8e0c42aa6a32029b0ecb80b84 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:04:08 +0000 Subject: [PATCH 013/131] feat: Over a finite ring, a module is finite iff it is finite dimensional (#17707) From PFR --- Mathlib/FieldTheory/AxGrothendieck.lean | 3 +-- Mathlib/FieldTheory/PrimitiveElement.lean | 2 +- .../LinearAlgebra/FiniteDimensional/Defs.lean | 5 ----- Mathlib/RingTheory/Finiteness.lean | 22 +++++++++++++++++++ .../LocalRing/ResidueField/Basic.lean | 6 ++--- 5 files changed, 27 insertions(+), 11 deletions(-) diff --git a/Mathlib/FieldTheory/AxGrothendieck.lean b/Mathlib/FieldTheory/AxGrothendieck.lean index e386e095dad39..9ddfa55c77de3 100644 --- a/Mathlib/FieldTheory/AxGrothendieck.lean +++ b/Mathlib/FieldTheory/AxGrothendieck.lean @@ -48,8 +48,7 @@ theorem ax_grothendieck_of_locally_finite {ι K R : Type*} [Field K] [Finite K] (mem_union_left _ (mem_biUnion.2 ⟨i, mem_univ _, mem_image_of_mem _ hk⟩)) letI := isNoetherian_adjoin_finset s fun x _ => Algebra.IsIntegral.isIntegral (R := K) x letI := Module.IsNoetherian.finite K (Algebra.adjoin K (s : Set R)) - letI : Finite (Algebra.adjoin K (s : Set R)) := - FiniteDimensional.finite_of_finite K (Algebra.adjoin K (s : Set R)) + letI : Finite (Algebra.adjoin K (s : Set R)) := Module.finite_of_finite K -- The restriction of the polynomial map, `ps`, to the subalgebra generated by `s` let res : (ι → Algebra.adjoin K (s : Set R)) → ι → Algebra.adjoin K (s : Set R) := fun x i => ⟨eval (fun j : ι => (x j : R)) (ps i), eval_mem (hs₁ _) fun i => (x i).2⟩ diff --git a/Mathlib/FieldTheory/PrimitiveElement.lean b/Mathlib/FieldTheory/PrimitiveElement.lean index da215b5d8dbf4..09b73357bf4f2 100644 --- a/Mathlib/FieldTheory/PrimitiveElement.lean +++ b/Mathlib/FieldTheory/PrimitiveElement.lean @@ -63,7 +63,7 @@ theorem exists_primitive_element_of_finite_top [Finite E] : ∃ α : E, F⟮α /-- Primitive element theorem for finite dimensional extension of a finite field. -/ theorem exists_primitive_element_of_finite_bot [Finite F] [FiniteDimensional F E] : ∃ α : E, F⟮α⟯ = ⊤ := - haveI : Finite E := FiniteDimensional.finite_of_finite F E + haveI : Finite E := Module.finite_of_finite F exists_primitive_element_of_finite_top F E end PrimitiveElementFinite diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 2bce02c0a50fc..14b5f34c2a2b3 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -107,11 +107,6 @@ instance finiteDimensional_pi' {ι : Type*} [Finite ι] (M : ι → Type*) [∀ noncomputable def fintypeOfFintype [Fintype K] [FiniteDimensional K V] : Fintype V := Module.fintypeOfFintype (@finsetBasis K V _ _ _ (iff_fg.2 inferInstance)) -theorem finite_of_finite [Finite K] [FiniteDimensional K V] : Finite V := by - cases nonempty_fintype K - haveI := fintypeOfFintype K V - infer_instance - variable {K V} /-- If a vector space has a finite basis, then it is finite-dimensional. -/ diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index 1425279d2a674..264b1bb8d9350 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -537,6 +537,22 @@ lemma exists_fin' [Module.Finite R M] : ∃ (n : ℕ) (f : (Fin n → R) →ₗ[ refine ⟨n, Basis.constr (Pi.basisFun R _) ℕ s, ?_⟩ rw [← LinearMap.range_eq_top, Basis.constr_range, hs] +variable (R) in +lemma _root_.Module.finite_of_finite [Finite R] [Module.Finite R M] : Finite M := by + obtain ⟨n, f, hf⟩ := exists_fin' R M; exact .of_surjective f hf + +@[deprecated (since := "2024-10-13")] +alias _root_.FiniteDimensional.finite_of_finite := finite_of_finite + +-- See note [lower instance priority] +instance (priority := 100) of_finite [Finite M] : Module.Finite R M := by + cases nonempty_fintype M + exact ⟨⟨Finset.univ, by rw [Finset.coe_univ]; exact Submodule.span_univ⟩⟩ + +/-- A module over a finite ring has finite dimension iff it is finite. -/ +lemma _root_.Module.finite_iff_finite [Finite R] : Module.Finite R M ↔ Finite M := + ⟨fun _ ↦ finite_of_finite R, fun _ ↦ .of_finite⟩ + theorem of_surjective [hM : Module.Finite R M] (f : M →ₗ[R] N) (hf : Surjective f) : Module.Finite R N := ⟨by @@ -620,6 +636,12 @@ instance span_singleton (x : M) : Module.Finite R (R ∙ x) := instance span_finset (s : Finset M) : Module.Finite R (span R (s : Set M)) := ⟨(Submodule.fg_top _).mpr ⟨s, rfl⟩⟩ +lemma _root_.Set.Finite.submoduleSpan [Finite R] {s : Set M} (hs : s.Finite) : + (Submodule.span R s : Set M).Finite := by + lift s to Finset M using hs + rw [Set.Finite, ← Module.finite_iff_finite (R := R)] + dsimp + infer_instance theorem Module.End.isNilpotent_iff_of_finite {R M : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] [Module.Finite R M] {f : End R M} : diff --git a/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean b/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean index 885d727ff53f4..ca82438849309 100644 --- a/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean @@ -166,10 +166,10 @@ instance finiteDimensional_of_noetherian [IsNoetherian R S] : (LinearMap.range_eq_top.mpr Ideal.Quotient.mk_surjective) exact Algebra.algebra_ext _ _ (fun r => rfl) +-- We want to be able to refer to `hfin` +set_option linter.unusedVariables false in lemma finite_of_finite [IsNoetherian R S] (hfin : Finite (ResidueField R)) : - Finite (ResidueField S) := by - have := @finiteDimensional_of_noetherian R S - exact FiniteDimensional.finite_of_finite (ResidueField R) (ResidueField S) + Finite (ResidueField S) := Module.finite_of_finite (ResidueField R) end FiniteDimensional From 0d150bcbee58d7d46b6d72a2ee0f113c00e3ec34 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 17 Oct 2024 15:04:10 +0000 Subject: [PATCH 014/131] chore(CategoryTheory/Sites): add some tests for sheafcompose instances (#17817) This PR adds a few tests to make sure that there are `HasSheafCompose` instances for the forgetful functor from modules to sets, for different size restrictions on the defining site. --- test/CategoryTheory/Sites/Whiskering.lean | 28 +++++++++++++++++++++++ 1 file changed, 28 insertions(+) create mode 100644 test/CategoryTheory/Sites/Whiskering.lean diff --git a/test/CategoryTheory/Sites/Whiskering.lean b/test/CategoryTheory/Sites/Whiskering.lean new file mode 100644 index 0000000000000..af2efea006a9d --- /dev/null +++ b/test/CategoryTheory/Sites/Whiskering.lean @@ -0,0 +1,28 @@ +import Mathlib.Algebra.Category.ModuleCat.Colimits +import Mathlib.Algebra.Category.ModuleCat.Limits +import Mathlib.Algebra.Category.ModuleCat.FilteredColimits +import Mathlib.CategoryTheory.Sites.Equivalence + +universe u + +open CategoryTheory GrothendieckTopology + +section Small + +variable {C : Type u} [SmallCategory C] (J : GrothendieckTopology C) (R : Type u) [Ring R] + +example : HasSheafCompose J (forget (ModuleCat.{u} R)) := inferInstance + +end Small + +section Large + +variable {C : Type (u+1)} [LargeCategory C] (J : GrothendieckTopology C) + +example (R : Type (u+1)) [Ring R] : HasSheafCompose J (forget (ModuleCat.{u+1} R)) := inferInstance + +variable [EssentiallySmall.{u} C] + +example (R : Type u) [Ring R] : HasSheafCompose J (forget (ModuleCat.{u} R)) := inferInstance + +end Large From c9b69692d65d8b44e452ca4692ed9b04eb5e463f Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 17 Oct 2024 15:04:11 +0000 Subject: [PATCH 015/131] chore(CategoryTheory/Sites): add some tests for preserving sheafification (#17818) This PR adds a few tests to make sure that there are `PreservesSheafification` instances for the forgetful functor from modules to sets, for different size restrictions on the defining site. --- .../Sites/PreservesSheafification.lean | 30 +++++++++++++++++++ 1 file changed, 30 insertions(+) create mode 100644 test/CategoryTheory/Sites/PreservesSheafification.lean diff --git a/test/CategoryTheory/Sites/PreservesSheafification.lean b/test/CategoryTheory/Sites/PreservesSheafification.lean new file mode 100644 index 0000000000000..813746eb759e5 --- /dev/null +++ b/test/CategoryTheory/Sites/PreservesSheafification.lean @@ -0,0 +1,30 @@ +import Mathlib.Algebra.Category.ModuleCat.Colimits +import Mathlib.Algebra.Category.ModuleCat.Limits +import Mathlib.Algebra.Category.ModuleCat.FilteredColimits +import Mathlib.CategoryTheory.Sites.Equivalence + +universe u + +open CategoryTheory GrothendieckTopology + +section Small + +variable {C : Type u} [SmallCategory C] (J : GrothendieckTopology C) (R : Type u) [Ring R] + +example : PreservesSheafification J (forget (ModuleCat.{u} R)) := inferInstance + +end Small + +section Large + +variable {C : Type (u+1)} [LargeCategory C] (J : GrothendieckTopology C) + +example (R : Type (u+1)) [Ring R] : PreservesSheafification J (forget (ModuleCat.{u+1} R)) := + inferInstance + +variable [EssentiallySmall.{u} C] + +example (R : Type u) [Ring R] : PreservesSheafification J (forget (ModuleCat.{u} R)) := + inferInstance + +end Large From a70b960b28617ef85de0b08c7cb24b3f92dcc964 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:04:12 +0000 Subject: [PATCH 016/131] feat: the filtration of cylinder events (#17858) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The exterior σ-algebras of finite sets of `α` form a cofiltration indexed by `Finset α` From GibbsMeasure Co-authored-by: Kin Yau James Wong --- Mathlib/Probability/Process/Filtration.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/Probability/Process/Filtration.lean b/Mathlib/Probability/Process/Filtration.lean index cc95366031c1b..853bae3ead907 100644 --- a/Mathlib/Probability/Process/Filtration.lean +++ b/Mathlib/Probability/Process/Filtration.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying, Rémy Degenne -/ +import Mathlib.MeasureTheory.Constructions.Cylinders import Mathlib.MeasureTheory.Function.ConditionalExpectation.Real /-! @@ -328,6 +329,14 @@ alias memℒp_limitProcess_of_snorm_bdd := memℒp_limitProcess_of_eLpNorm_bdd end Limit +variable {α : Type*} + +/-- The exterior σ-algebras of finite sets of `α` form a cofiltration indexed by `Finset α`. -/ +def cylinderEventsCompl : Filtration (Finset α)ᵒᵈ (.pi (π := fun _ : α ↦ Ω)) where + seq Λ := cylinderEvents (↑(OrderDual.ofDual Λ))ᶜ + mono' _ _ h := cylinderEvents_mono <| Set.compl_subset_compl_of_subset h + le' _ := cylinderEvents_le_pi + end Filtration end MeasureTheory From 3f297f53816d8a85d03f8c08b3a7d475dfa51a6b Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 15:04:13 +0000 Subject: [PATCH 017/131] chore: remove porting notes about adding/removing `@[ext]` attributes (#17873) Discussed with @jcommelin and @kim-em. All of these porting notes are about the change in behaviour of the `ext` tactic and attribute between Lean 3 and Lean 4: in Lean 4 it only accepts lemmas with conclusion `x = y` with `x` and `y` free variables, and only applies these if the types are reducibly equal. Since those behaviours are not likely to change in the future, we don't need to keep these porting notes. I propose that we should keep #11041 since it indicates where these `ext` lemmas have not yet been added. Closes: #5229 Closes: #11182 --- Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean | 1 - Mathlib/Algebra/Group/AddChar.lean | 1 - Mathlib/Algebra/Homology/HomologicalComplex.lean | 1 - Mathlib/Algebra/Polynomial/Laurent.lean | 1 - Mathlib/AlgebraicTopology/SimplexCategory.lean | 1 - Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean | 2 -- Mathlib/AlgebraicTopology/SplitSimplicialObject.lean | 1 - Mathlib/CategoryTheory/Abelian/Pseudoelements.lean | 6 ------ .../Bicategory/NaturalTransformation/Oplax.lean | 1 - Mathlib/CategoryTheory/Comma/Basic.lean | 2 -- Mathlib/CategoryTheory/Comma/StructuredArrow.lean | 4 ---- Mathlib/CategoryTheory/Elements.lean | 1 - Mathlib/CategoryTheory/Endomorphism.lean | 1 - Mathlib/CategoryTheory/Functor/Category.lean | 4 ---- Mathlib/CategoryTheory/Idempotents/Karoubi.lean | 1 - Mathlib/CategoryTheory/Iso.lean | 4 ---- Mathlib/CategoryTheory/Limits/Cones.lean | 2 -- Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean | 2 -- Mathlib/CategoryTheory/Monoidal/Bimod.lean | 1 - Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean | 2 -- Mathlib/CategoryTheory/Monoidal/CommMon_.lean | 2 -- Mathlib/CategoryTheory/Monoidal/Mod_.lean | 1 - Mathlib/CategoryTheory/Monoidal/Mon_.lean | 1 - Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean | 1 - Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean | 2 -- Mathlib/CategoryTheory/Pi/Basic.lean | 1 - Mathlib/CategoryTheory/Preadditive/Mat.lean | 3 --- Mathlib/CategoryTheory/Sites/Sheaf.lean | 1 - Mathlib/CategoryTheory/Subobject/Basic.lean | 3 --- Mathlib/CategoryTheory/Yoneda.lean | 4 ---- Mathlib/Geometry/RingedSpace/PresheafedSpace.lean | 1 - Mathlib/Geometry/RingedSpace/SheafedSpace.lean | 1 - Mathlib/RepresentationTheory/Action/Basic.lean | 1 - 33 files changed, 61 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index c3f9a38ef087e..7d7e0e34c73af 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -365,7 +365,6 @@ def embeddingLiftIso (F : C ⥤ D) : embedding R C ⋙ lift R F ≅ F := /-- Two `R`-linear functors out of the `R`-linear completion are isomorphic iff their compositions with the embedding functor are isomorphic. -/ --- Porting note (#11182): used to be @[ext] def ext {F G : Free R C ⥤ D} [F.Additive] [F.Linear R] [G.Additive] [G.Linear R] (α : embedding R C ⋙ F ≅ embedding R C ⋙ G) : F ≅ G := NatIso.ofComponents (fun X => α.app X) diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index f9a1d383c185c..23df07dbc7f89 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -88,7 +88,6 @@ instance instFunLike : FunLike (AddChar A M) A M where coe := AddChar.toFun coe_injective' φ ψ h := by cases φ; cases ψ; congr --- Porting note (#5229): added. @[ext] lemma ext (f g : AddChar A M) (h : ∀ x : A, f x = g x) : f = g := DFunLike.ext f g h diff --git a/Mathlib/Algebra/Homology/HomologicalComplex.lean b/Mathlib/Algebra/Homology/HomologicalComplex.lean index 6ea13712672bd..441c75acfbc02 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplex.lean @@ -244,7 +244,6 @@ instance : Category (HomologicalComplex V c) where end --- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {C D : HomologicalComplex V c} (f g : C ⟶ D) (h : ∀ i, f.f i = g.f i) : f = g := by diff --git a/Mathlib/Algebra/Polynomial/Laurent.lean b/Mathlib/Algebra/Polynomial/Laurent.lean index 949d81ecba031..c231267f42556 100644 --- a/Mathlib/Algebra/Polynomial/Laurent.lean +++ b/Mathlib/Algebra/Polynomial/Laurent.lean @@ -87,7 +87,6 @@ scoped[LaurentPolynomial] notation:9000 R "[T;T⁻¹]" => LaurentPolynomial R open LaurentPolynomial --- Porting note (#5229): `ext` no longer applies `Finsupp.ext` automatically @[ext] theorem LaurentPolynomial.ext [Semiring R] {p q : R[T;T⁻¹]} (h : ∀ a, p a = q a) : p = q := Finsupp.ext h diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index ac2101ce20054..72bc3cc192913 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -142,7 +142,6 @@ lemma id_toOrderHom (a : SimplexCategory) : lemma comp_toOrderHom {a b c : SimplexCategory} (f : a ⟶ b) (g : b ⟶ c) : (f ≫ g).toOrderHom = g.toOrderHom.comp f.toOrderHom := rfl --- Porting note (#5229): added because `Hom.ext'` is not triggered automatically @[ext] theorem Hom.ext {a b : SimplexCategory} (f g : a ⟶ b) : f.toOrderHom = g.toOrderHom → f = g := diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index b11a417a297e7..32b22e7b91ce5 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -57,7 +57,6 @@ instance hasColimits : HasColimits SSet := by dsimp only [SSet] infer_instance --- Porting note (#5229): added an `ext` lemma. @[ext] lemma hom_ext {X Y : SSet} {f g : X ⟶ Y} (w : ∀ n, f.app n = g.app n) : f = g := SimplicialObject.hom_ext _ _ w @@ -354,7 +353,6 @@ 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 := diff --git a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean index 48c6023dd1053..7d8574e79627f 100644 --- a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean +++ b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean @@ -351,7 +351,6 @@ variable {C} namespace Split --- Porting note (#5229): added as `Hom.ext` is not triggered automatically @[ext] theorem hom_ext {S₁ S₂ : Split C} (Φ₁ Φ₂ : S₁ ⟶ S₂) (h : ∀ n : ℕ, Φ₁.f n = Φ₂.f n) : Φ₁ = Φ₂ := Hom.ext _ _ h diff --git a/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean b/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean index 7c2ed53010962..7bb3ceaced8e1 100644 --- a/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean +++ b/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean @@ -258,12 +258,6 @@ theorem zero_morphism_ext {P Q : C} (f : P ⟶ Q) : (∀ a, f a = 0) → f = 0 : theorem zero_morphism_ext' {P Q : C} (f : P ⟶ Q) : (∀ a, f a = 0) → 0 = f := Eq.symm ∘ zero_morphism_ext f --- Porting note (#11182): these are no longer valid as `ext` lemmas. --- scoped[Pseudoelement] --- attribute [ext] --- CategoryTheory.Abelian.Pseudoelement.zero_morphism_ext --- CategoryTheory.Abelian.Pseudoelement.zero_morphism_ext' - theorem eq_zero_iff {P Q : C} (f : P ⟶ Q) : f = 0 ↔ ∀ a, f a = 0 := ⟨fun h a => by simp [h], zero_morphism_ext _⟩ diff --git a/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean b/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean index 2d6217112669a..ee6b3cb99b45e 100644 --- a/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean +++ b/Mathlib/CategoryTheory/Bicategory/NaturalTransformation/Oplax.lean @@ -245,7 +245,6 @@ instance category (F G : OplaxFunctor B C) : Category (F ⟶ G) where id := Modification.id comp := Modification.vcomp --- Porting note (#5229): duplicating the `ext` lemma. @[ext] lemma ext {F G : OplaxFunctor B C} {α β : F ⟶ G} {m n : α ⟶ β} (w : ∀ b, m.app b = n.app b) : m = n := by diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index 74eebe08f5cf9..192ef0e4606e0 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -106,8 +106,6 @@ section variable {X Y Z : Comma L R} {f : X ⟶ Y} {g : Y ⟶ Z} --- Porting note (#5229): this lemma was added because `CommaMorphism.ext` --- was not triggered automatically @[ext] lemma hom_ext (f g : X ⟶ Y) (h₁ : f.left = g.left) (h₂ : f.right = g.right) : f = g := CommaMorphism.ext h₁ h₂ diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean index 3ee2b719536a8..d86b41669eacf 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean @@ -51,8 +51,6 @@ def proj (S : D) (T : C ⥤ D) : StructuredArrow S T ⥤ C := variable {S S' S'' : D} {Y Y' Y'' : C} {T T' : C ⥤ D} --- Porting note (#5229): this lemma was added because `Comma.hom_ext` --- was not triggered automatically @[ext] lemma hom_ext {X Y : StructuredArrow S T} (f g : X ⟶ Y) (h : f.right = g.right) : f = g := CommaMorphism.ext (Subsingleton.elim _ _) h @@ -400,8 +398,6 @@ def proj (S : C ⥤ D) (T : D) : CostructuredArrow S T ⥤ C := variable {T T' T'' : D} {Y Y' Y'' : C} {S S' : C ⥤ D} --- Porting note (#5229): this lemma was added because `Comma.hom_ext` --- was not triggered automatically @[ext] lemma hom_ext {X Y : CostructuredArrow S T} (f g : X ⟶ Y) (h : f.left = g.left) : f = g := CommaMorphism.ext h (Subsingleton.elim _ _) diff --git a/Mathlib/CategoryTheory/Elements.lean b/Mathlib/CategoryTheory/Elements.lean index 588561a7941f8..cf261e1d8d7fc 100644 --- a/Mathlib/CategoryTheory/Elements.lean +++ b/Mathlib/CategoryTheory/Elements.lean @@ -47,7 +47,6 @@ def Functor.Elements (F : C ⥤ Type w) := /-- Constructor for the type `F.Elements` when `F` is a functor to types. -/ abbrev Functor.elementsMk (F : C ⥤ Type w) (X : C) (x : F.obj X) : F.Elements := ⟨X, x⟩ --- Porting note (#5229): added because Sigma.ext would be triggered automatically lemma Functor.Elements.ext {F : C ⥤ Type w} (x y : F.Elements) (h₁ : x.fst = y.fst) (h₂ : F.map (eqToHom h₁) x.snd = y.snd) : x = y := by cases x diff --git a/Mathlib/CategoryTheory/Endomorphism.lean b/Mathlib/CategoryTheory/Endomorphism.lean index 4845f0a202a03..0ab626814ae99 100644 --- a/Mathlib/CategoryTheory/Endomorphism.lean +++ b/Mathlib/CategoryTheory/Endomorphism.lean @@ -114,7 +114,6 @@ def Aut (X : C) := X ≅ X namespace Aut --- Porting note (#5229): added because `Iso.ext` is not triggered automatically @[ext] lemma ext {X : C} {φ₁ φ₂ : Aut X} (h : φ₁.hom = φ₂.hom) : φ₁ = φ₂ := Iso.ext h diff --git a/Mathlib/CategoryTheory/Functor/Category.lean b/Mathlib/CategoryTheory/Functor/Category.lean index e1bd5dc1cfab1..7a1687ae5667d 100644 --- a/Mathlib/CategoryTheory/Functor/Category.lean +++ b/Mathlib/CategoryTheory/Functor/Category.lean @@ -48,10 +48,6 @@ instance Functor.category : Category.{max u₁ v₂} (C ⥤ D) where namespace NatTrans --- Porting note (#5229): the behaviour of `ext` has changed here. --- We need to provide a copy of the `NatTrans.ext` lemma, --- written in terms of `F ⟶ G` rather than `NatTrans F G`, --- or `ext` will not retrieve it from the cache. @[ext] theorem ext' {α β : F ⟶ G} (w : α.app = β.app) : α = β := NatTrans.ext w diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index abcded134084e..b08afa205a73a 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -101,7 +101,6 @@ theorem hom_ext_iff {P Q : Karoubi C} {f g : P ⟶ Q} : f = g ↔ f.f = g.f := b rw [h] · apply Hom.ext --- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] theorem hom_ext {P Q : Karoubi C} (f g : P ⟶ Q) (h : f.f = g.f) : f = g := by simpa [hom_ext_iff] using h diff --git a/Mathlib/CategoryTheory/Iso.lean b/Mathlib/CategoryTheory/Iso.lean index 7ce50b850ef49..6e0dd598bc818 100644 --- a/Mathlib/CategoryTheory/Iso.lean +++ b/Mathlib/CategoryTheory/Iso.lean @@ -305,7 +305,6 @@ instance (priority := 100) mono_of_iso (f : X ⟶ Y) [IsIso f] : Mono f where rw [← Category.comp_id g, ← Category.comp_id h, ← IsIso.hom_inv_id f, ← Category.assoc, w, ← Category.assoc] --- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory])] theorem inv_eq_of_hom_inv_id {f : X ⟶ Y} [IsIso f] {g : Y ⟶ X} (hom_inv_id : f ≫ g = 𝟙 X) : inv f = g := by @@ -317,7 +316,6 @@ theorem inv_eq_of_inv_hom_id {f : X ⟶ Y} [IsIso f] {g : Y ⟶ X} (inv_hom_id : apply (cancel_mono f).mp simp [inv_hom_id] --- Porting note (#11182): `@[ext]` used to accept lemmas like this. @[aesop apply safe (rule_sets := [CategoryTheory])] theorem eq_inv_of_hom_inv_id {f : X ⟶ Y} [IsIso f] {g : Y ⟶ X} (hom_inv_id : f ≫ g = 𝟙 X) : g = inv f := @@ -446,12 +444,10 @@ theorem isIso_of_comp_hom_eq_id (g : X ⟶ Y) [IsIso g] {f : Y ⟶ X} (h : f ≫ namespace Iso --- Porting note (#11182): `@[ext]` used to accept lemmas like this. @[aesop apply safe (rule_sets := [CategoryTheory])] theorem inv_ext {f : X ≅ Y} {g : Y ⟶ X} (hom_inv_id : f.hom ≫ g = 𝟙 X) : f.inv = g := ((hom_comp_eq_id f).1 hom_inv_id).symm --- Porting note (#11182): `@[ext]` used to accept lemmas like this. @[aesop apply safe (rule_sets := [CategoryTheory])] theorem inv_ext' {f : X ≅ Y} {g : Y ⟶ X} (hom_inv_id : f.hom ≫ g = 𝟙 X) : g = f.inv := (hom_comp_eq_id f).1 hom_inv_id diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index 098f9f7adcb6a..36d559c2f7911 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -283,7 +283,6 @@ namespace Cones /-- To give an isomorphism between cones, it suffices to give an isomorphism between their vertices which commutes with the cone maps. -/ --- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {c c' : Cone F} (φ : c.pt ≅ c'.pt) (w : ∀ j, c.π.app j = φ.hom ≫ c'.π.app j := by aesop_cat) : c ≅ c' where @@ -484,7 +483,6 @@ namespace Cocones /-- To give an isomorphism between cocones, it suffices to give an isomorphism between their vertices which commutes with the cocone maps. -/ --- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {c c' : Cocone F} (φ : c.pt ≅ c'.pt) (w : ∀ j, c.ι.app j ≫ φ.hom = c'.ι.app j := by aesop_cat) : c ≅ c' where diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index 2d16499600981..f290e6be94303 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -122,7 +122,6 @@ namespace Bicones /-- To give an isomorphism between cocones, it suffices to give an isomorphism between their vertices which commutes with the cocone maps. -/ --- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {c c' : Bicone F} (φ : c.pt ≅ c'.pt) (wι : ∀ j, c.ι j ≫ φ.hom = c'.ι j := by aesop_cat) @@ -1119,7 +1118,6 @@ namespace BinaryBicones /-- To give an isomorphism between cocones, it suffices to give an isomorphism between their vertices which commutes with the cocone maps. -/ --- Porting note (#11182): `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets := [CategoryTheory]), simps] def ext {P Q : C} {c c' : BinaryBicone P Q} (φ : c.pt ≅ c'.pt) (winl : c.inl ≫ φ.hom = c'.inl := by aesop_cat) diff --git a/Mathlib/CategoryTheory/Monoidal/Bimod.lean b/Mathlib/CategoryTheory/Monoidal/Bimod.lean index b6d566f92f32a..6182c9de23c17 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimod.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimod.lean @@ -120,7 +120,6 @@ instance : Category (Bimod A B) where id := id' comp f g := comp f g --- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {M N : Bimod A B} (f g : M ⟶ N) (h : f.hom = g.hom) : f = g := Hom.ext h diff --git a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean index c0dfcace9f9a6..ee91313dea7f3 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean @@ -399,7 +399,6 @@ def comp (F : LaxBraidedFunctor C D) (G : LaxBraidedFunctor D E) : LaxBraidedFun instance categoryLaxBraidedFunctor : Category (LaxBraidedFunctor C D) := InducedCategory.category LaxBraidedFunctor.toLaxMonoidalFunctor --- Porting note (#5229): added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : LaxBraidedFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := MonoidalNatTrans.ext (funext w) @@ -463,7 +462,6 @@ def comp (F : BraidedFunctor C D) (G : BraidedFunctor D E) : BraidedFunctor C E instance categoryBraidedFunctor : Category (BraidedFunctor C D) := InducedCategory.category BraidedFunctor.toMonoidalFunctor --- Porting note (#5229): added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : BraidedFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := MonoidalNatTrans.ext (funext w) diff --git a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean index b8a3c549ea911..95ff7663a6476 100644 --- a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean @@ -50,8 +50,6 @@ theorem comp_hom {R S T : CommMon_ C} (f : R ⟶ S) (g : S ⟶ T) : Mon_.Hom.hom (f ≫ g) = f.hom ≫ g.hom := rfl --- Porting note (#5229): added because `Mon_.Hom.ext` is not triggered automatically --- for morphisms in `CommMon_ C` @[ext] lemma hom_ext {A B : CommMon_ C} (f g : A ⟶ B) (h : f.hom = g.hom) : f = g := Mon_.Hom.ext h diff --git a/Mathlib/CategoryTheory/Monoidal/Mod_.lean b/Mathlib/CategoryTheory/Monoidal/Mod_.lean index bee5d1a324c2f..ec6327ca22917 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mod_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mod_.lean @@ -57,7 +57,6 @@ instance : Category (Mod_ A) where id := id comp f g := comp f g --- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {M N : Mod_ A} (f₁ f₂ : M ⟶ N) (h : f₁.hom = f₂.hom) : f₁ = f₂ := Hom.ext h diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index 82f033d04a045..fc0cec7c2faf7 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -104,7 +104,6 @@ instance : Category (Mon_ C) where id := id comp f g := comp f g --- Porting note (#5229): added, as `Hom.ext` does not apply to a morphism. @[ext] lemma ext {X Y : Mon_ C} {f g : X ⟶ Y} (w : f.hom = g.hom) : f = g := Hom.ext w diff --git a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean index 4d73ab1b96ac5..d41970b459890 100644 --- a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean +++ b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean @@ -81,7 +81,6 @@ theorem comp_toNatTrans_lax {F G H : LaxMonoidalFunctor C D} {α : F ⟶ G} {β instance categoryMonoidalFunctor : Category (MonoidalFunctor C D) := InducedCategory.category MonoidalFunctor.toLaxMonoidalFunctor --- Porting note (#5229): added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : LaxMonoidalFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := MonoidalNatTrans.ext (funext w) diff --git a/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean b/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean index e4d78ec3864b8..a20da408bc49c 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean @@ -104,8 +104,6 @@ lemma FunctorsInverting.ext {W : MorphismProperty C} {F₁ F₂ : FunctorsInvert instance (W : MorphismProperty C) (D : Type*) [Category D] : Category (FunctorsInverting W D) := FullSubcategory.category _ --- Porting note (#5229): add another `@[ext]` lemma --- since `ext` can't see through the definition to use `NatTrans.ext`. @[ext] lemma FunctorsInverting.hom_ext {W : MorphismProperty C} {F₁ F₂ : FunctorsInverting W D} {α β : F₁ ⟶ F₂} (h : α.app = β.app) : α = β := diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index 8515247f7d5ba..477e2fe2a21e5 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -49,7 +49,6 @@ theorem comp_apply {X Y Z : ∀ i, C i} (f : X ⟶ Y) (g : Y ⟶ Z) (i) : (f ≫ g : ∀ i, X i ⟶ Z i) i = f i ≫ g i := rfl --- Porting note (#5229): need to add an additional `ext` lemma. @[ext] lemma ext {X Y : ∀ i, C i} {f g : X ⟶ Y} (w : ∀ i, f i = g i) : f = g := funext (w ·) diff --git a/Mathlib/CategoryTheory/Preadditive/Mat.lean b/Mathlib/CategoryTheory/Preadditive/Mat.lean index 9d44442e07dd3..8ab0b6f6a6063 100644 --- a/Mathlib/CategoryTheory/Preadditive/Mat.lean +++ b/Mathlib/CategoryTheory/Preadditive/Mat.lean @@ -105,7 +105,6 @@ instance : Category.{v₁} (Mat_ C) where simp_rw [Hom.comp, sum_comp, comp_sum, Category.assoc] rw [Finset.sum_comm] --- Porting note (#5229): added because `DMatrix.ext` is not triggered automatically @[ext] theorem hom_ext {M N : Mat_ C} (f g : M ⟶ N) (H : ∀ i j, f i j = g i j) : f = g := DMatrix.ext_iff.mp H @@ -432,7 +431,6 @@ def liftUnique (F : C ⥤ D) [Functor.Additive F] (L : Mat_ C ⥤ D) [Functor.Ad dsimp simpa using α.hom.naturality (f j k) --- Porting note (#11182): removed @[ext] as the statement is not an equality -- TODO is there some uniqueness statement for the natural isomorphism in `liftUnique`? /-- Two additive functors `Mat_ C ⥤ D` are naturally isomorphic if their precompositions with `embedding C` are naturally isomorphic as functors `C ⥤ D`. -/ @@ -505,7 +503,6 @@ section variable {R : Type u} [Semiring R] --- Porting note (#5229): added because `Matrix.ext` is not triggered automatically @[ext] theorem hom_ext {X Y : Mat R} (f g : X ⟶ Y) (h : ∀ i j, f i j = g i j) : f = g := Matrix.ext_iff.mp h diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index 99d958c7ef686..10da297aaf6d1 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -332,7 +332,6 @@ instance instCategorySheaf : Category (Sheaf J A) where instance (X : Sheaf J A) : Inhabited (Hom X X) := ⟨𝟙 X⟩ --- Porting note (#5229): added because `Sheaf.Hom.ext` was not triggered automatically @[ext] lemma hom_ext {X Y : Sheaf J A} (x y : X ⟶ Y) (h : x.val = y.val) : x = y := Sheaf.Hom.ext h diff --git a/Mathlib/CategoryTheory/Subobject/Basic.lean b/Mathlib/CategoryTheory/Subobject/Basic.lean index 34c23433816f0..3afc05fb08654 100644 --- a/Mathlib/CategoryTheory/Subobject/Basic.lean +++ b/Mathlib/CategoryTheory/Subobject/Basic.lean @@ -255,21 +255,18 @@ theorem eq_of_comm {B : C} {X Y : Subobject B} (f : (X : C) ≅ (Y : C)) (w : f.hom ≫ Y.arrow = X.arrow) : X = Y := le_antisymm (le_of_comm f.hom w) <| le_of_comm f.inv <| f.inv_comp_eq.2 w.symm --- Porting note (#11182): removed @[ext] /-- To show that two subobjects are equal, it suffices to exhibit an isomorphism commuting with the arrows. -/ theorem eq_mk_of_comm {B A : C} {X : Subobject B} (f : A ⟶ B) [Mono f] (i : (X : C) ≅ A) (w : i.hom ≫ f = X.arrow) : X = mk f := eq_of_comm (i.trans (underlyingIso f).symm) <| by simp [w] --- Porting note (#11182): removed @[ext] /-- To show that two subobjects are equal, it suffices to exhibit an isomorphism commuting with the arrows. -/ theorem mk_eq_of_comm {B A : C} {X : Subobject B} (f : A ⟶ B) [Mono f] (i : A ≅ (X : C)) (w : i.hom ≫ X.arrow = f) : mk f = X := Eq.symm <| eq_mk_of_comm _ i.symm <| by rw [Iso.symm_hom, Iso.inv_comp_eq, w] --- Porting note (#11182): removed @[ext] /-- To show that two subobjects are equal, it suffices to exhibit an isomorphism commuting with the arrows. -/ theorem mk_eq_mk_of_comm {B A₁ A₂ : C} (f : A₁ ⟶ B) (g : A₂ ⟶ B) [Mono f] [Mono g] (i : A₁ ≅ A₂) diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index d056292dbff45..61b37dc22ad22 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -486,8 +486,6 @@ to `yoneda.op.obj X ⟶ F`, functorially in both `X` and `F`. def yonedaPairing : Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁) ⥤ Type max u₁ v₁ := Functor.prod yoneda.op (𝟭 (Cᵒᵖ ⥤ Type v₁)) ⋙ Functor.hom (Cᵒᵖ ⥤ Type v₁) --- Porting note (#5229): we need to provide this `@[ext]` lemma separately, --- as `ext` will not look through the definition. @[ext] lemma yonedaPairingExt {X : Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁)} {x y : (yonedaPairing C).obj X} (w : ∀ Y, x.app Y = y.app Y) : x = y := @@ -657,8 +655,6 @@ to `coyoneda.rightOp.obj X ⟶ F`, functorially in both `X` and `F`. def coyonedaPairing : C × (C ⥤ Type v₁) ⥤ Type max u₁ v₁ := Functor.prod coyoneda.rightOp (𝟭 (C ⥤ Type v₁)) ⋙ Functor.hom (C ⥤ Type v₁) --- Porting note (#5229): we need to provide this `@[ext]` lemma separately, --- as `ext` will not look through the definition. @[ext] lemma coyonedaPairingExt {X : C × (C ⥤ Type v₁)} {x y : (coyonedaPairing C).obj X} (w : ∀ Y, x.app Y = y.app Y) : x = y := diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean index f4762b2bced5c..12c2d277fbd66 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean @@ -160,7 +160,6 @@ variable {C} /-- Cast `Hom X Y` as an arrow `X ⟶ Y` of presheaves. -/ abbrev Hom.toPshHom {X Y : PresheafedSpace C} (f : Hom X Y) : X ⟶ Y := f --- Porting note (#5229): adding an `ext` lemma. @[ext (iff := false)] theorem ext {X Y : PresheafedSpace C} (α β : X ⟶ Y) (w : α.base = β.base) (h : α.c ≫ whiskerRight (eqToHom (by rw [w])) _ = β.c) : α = β := diff --git a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean index 7b65ebaffd7eb..3af5f2d7e91d0 100644 --- a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean @@ -80,7 +80,6 @@ instance : Category (SheafedSpace C) := show Category (InducedCategory (PresheafedSpace C) SheafedSpace.toPresheafedSpace) by infer_instance --- Porting note (#5229): adding an `ext` lemma. @[ext (iff := false)] theorem ext {X Y : SheafedSpace C} (α β : X ⟶ Y) (w : α.base = β.base) (h : α.c ≫ whiskerRight (eqToHom (by rw [w])) _ = β.c) : α = β := diff --git a/Mathlib/RepresentationTheory/Action/Basic.lean b/Mathlib/RepresentationTheory/Action/Basic.lean index f70ce1fb09e29..2f15639455bdf 100644 --- a/Mathlib/RepresentationTheory/Action/Basic.lean +++ b/Mathlib/RepresentationTheory/Action/Basic.lean @@ -111,7 +111,6 @@ instance : Category (Action V G) where id M := Hom.id M comp f g := Hom.comp f g --- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {M N : Action V G} (φ₁ φ₂ : M ⟶ N) (h : φ₁.hom = φ₂.hom) : φ₁ = φ₂ := Hom.ext h From 3ffd5e037b602dd5939615c5556c0464cae46492 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Calle=20S=C3=B6nne?= Date: Thu, 17 Oct 2024 15:41:44 +0000 Subject: [PATCH 018/131] feat(MorphismProperty/Presheaf): Add properties of the `relative` morphism property (#16142) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds some basic properties about the "presheaf" morphism property. For example, that it is stable under base change. Co-authored-by: Joël Riou --- .../MorphismProperty/Representable.lean | 58 ++++++++++++++++++- 1 file changed, 55 insertions(+), 3 deletions(-) diff --git a/Mathlib/CategoryTheory/MorphismProperty/Representable.lean b/Mathlib/CategoryTheory/MorphismProperty/Representable.lean index e2ff92296f5c9..df96aa39de010 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Representable.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Representable.lean @@ -25,9 +25,10 @@ Classically, a morphism `f : F ⟶ G` of presheaves is said to be representable In this file, we define a notion of relative representability which works with respect to any functor, and not just `yoneda`. The fact that a morphism `f : F ⟶ G` between presheaves is -representable in the classical case will then be given by `F.relativelyRepresentable f`. -` +representable in the classical case will then be given by `yoneda.relativelyRepresentable f`. + ## Main definitions + Throughout this file, `F : C ⥤ D` is a functor between categories `C` and `D`. * `Functor.relativelyRepresentable`: A morphism `f : X ⟶ Y` in `D` is said to be relatively @@ -312,7 +313,6 @@ category `Cᵒᵖ ⥤ Type v` satisfies the morphism property `P.presheaf` iff: * The morphism is representable. * For any morphism `g : F.obj a ⟶ G`, the property `P` holds for any represented pullback of `f` by `g`. - This is implemented as a special case of the more general notion of `P.relative`, to the case when the functor `F` is `yoneda`. -/ abbrev presheaf : MorphismProperty (Cᵒᵖ ⥤ Type v₁) := P.relative yoneda @@ -372,6 +372,58 @@ lemma relative_map_iff [F.Faithful] [F.Full] [PreservesLimitsOfShape WalkingCosp P.relative F (F.map f) ↔ P f := ⟨fun hf ↦ of_relative_map hf, fun hf ↦ relative_map hP hf⟩ +/-- If `P' : MorphismProperty C` is satisfied whenever `P` is, then also `P'.relative` is +satisfied whenever `P.relative` is. -/ +lemma relative_monotone {P' : MorphismProperty C} (h : P ≤ P') : + P.relative F ≤ P'.relative F := fun _ _ _ hf ↦ + ⟨hf.rep, fun _ _ g fst snd BC ↦ h _ (hf.property g fst snd BC)⟩ + +section + +variable (P) + +lemma relative_stableUnderBaseChange : StableUnderBaseChange (P.relative F) := + fun _ _ _ _ _ _ _ _ hfBC hg ↦ + ⟨stableUnderBaseChange F hfBC hg.rep, + fun _ _ _ _ _ BC ↦ hg.property _ _ _ (IsPullback.paste_horiz BC hfBC)⟩ + +instance relative_isStableUnderComposition [F.Faithful] [F.Full] [P.IsStableUnderComposition] : + IsStableUnderComposition (P.relative F) where + comp_mem {F G H} f g hf hg := by + refine ⟨comp_mem _ _ _ hf.1 hg.1, fun Z X p fst snd h ↦ ?_⟩ + rw [← hg.1.lift_snd (fst ≫ f) snd (by simpa using h.w)] + refine comp_mem _ _ _ (hf.property (hg.1.fst p) fst _ + (IsPullback.of_bot ?_ ?_ (hg.1.isPullback p))) (hg.property_snd p) + · rw [← Functor.map_comp, lift_snd] + exact h + · symm + apply hg.1.lift_fst + +instance relative_respectsIso : RespectsIso (P.relative F) := + (relative_stableUnderBaseChange P).respectsIso + +instance relative_isMultiplicative [F.Faithful] [F.Full] [P.IsMultiplicative] [P.RespectsIso] : + IsMultiplicative (P.relative F) where + id_mem X := relative.of_exists + (fun Y g ↦ ⟨Y, g, 𝟙 Y, by simpa using IsPullback.of_id_snd, id_mem _ _⟩) + +end + +/-- Morphisms satisfying `(monomorphism C).presheaf` are in particular monomorphisms. -/ +lemma presheaf_monomorphisms_le_monomorphisms : + (monomorphisms C).presheaf ≤ monomorphisms _ := fun F G f hf ↦ by + suffices ∀ {X : C} {a b : yoneda.obj X ⟶ F}, a ≫ f = b ≫ f → a = b from + ⟨fun _ _ h ↦ hom_ext_yoneda (fun _ _ ↦ this (by simp only [assoc, h]))⟩ + intro X a b h + /- It suffices to show that the lifts of `a` and `b` to morphisms + `X ⟶ hf.rep.pullback g` are equal, where `g = a ≫ f = a ≫ f`. -/ + suffices hf.rep.lift (g := a ≫ f) a (𝟙 X) (by simp) = + hf.rep.lift b (𝟙 X) (by simp [← h]) by + simpa using yoneda.congr_map this =≫ (hf.rep.fst (a ≫ f)) + -- This follows from the fact that the induced maps `hf.rep.pullback g ⟶ X` are mono. + have : Mono (hf.rep.snd (a ≫ f)) := hf.property_snd (a ≫ f) + simp only [← cancel_mono (hf.rep.snd (a ≫ f)), lift_snd] + end MorphismProperty end CategoryTheory From 27616824e4ad7714586263171a565683e6a6d7c9 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 17 Oct 2024 15:41:47 +0000 Subject: [PATCH 019/131] chore(Profinite): remove simp lemma `toProfinite_obj` (#17155) --- Mathlib/Condensed/Discrete/Colimit.lean | 41 ++++++++----------- .../Category/LightProfinite/AsLimit.lean | 3 +- .../Category/LightProfinite/Basic.lean | 2 +- .../Topology/Category/Profinite/Basic.lean | 2 +- 4 files changed, 21 insertions(+), 27 deletions(-) diff --git a/Mathlib/Condensed/Discrete/Colimit.lean b/Mathlib/Condensed/Discrete/Colimit.lean index 2e2de66d0a095..75f3c6ee7dfeb 100644 --- a/Mathlib/Condensed/Discrete/Colimit.lean +++ b/Mathlib/Condensed/Discrete/Colimit.lean @@ -48,10 +48,7 @@ noncomputable def isColimitLocallyConstantPresheaf (hc : IsLimit c) [∀ i, Epi (h : fi.comap (c.π.app i) = fj.comap (c.π.app j)) obtain ⟨k, ki, kj, _⟩ := IsCofilteredOrEmpty.cone_objs i j refine ⟨⟨k⟩, ki.op, kj.op, ?_⟩ - dsimp only [comp_obj, op_obj, functorToPresheaves_obj_obj, CompHausLike.coe_of, - Functor.comp_map, op_map, Quiver.Hom.unop_op, functorToPresheaves_obj_map] - -- Note: we might want to remove the `simps` attribute from `FintypeCat.toProfinite`; keeping - -- `toProfinite_obj` in the `dsimp` block above causes the following `ext` to fail. + dsimp ext x obtain ⟨x, hx⟩ := ((Profinite.epi_iff_surjective (c.π.app k)).mp inferInstance) x rw [← hx] @@ -237,7 +234,7 @@ def isoFinYoneda : toProfinite.op ⋙ F ≅ finYoneda F := NatIso.ofComponents (fun X ↦ isoFinYonedaComponents F (toProfinite.obj X.unop)) fun _ ↦ by simp only [comp_obj, op_obj, finYoneda_obj, Functor.comp_map, op_map] ext - simp only [toProfinite_obj, types_comp_apply, isoFinYonedaComponents_hom_apply, finYoneda_map, + simp only [types_comp_apply, isoFinYonedaComponents_hom_apply, finYoneda_map, op_obj, Function.comp_apply, ← FunctorToTypes.map_comp_apply] rfl @@ -264,14 +261,14 @@ lemma isoLocallyConstantOfIsColimit_inv (X : Profinite.{u}ᵒᵖ ⥤ Type (u+1)) apply colimit.hom_ext intro ⟨Y, _, g⟩ simp? [locallyConstantIsoFinYoneda, isoFinYoneda, counitApp] says - simp only [comp_obj, CostructuredArrow.proj_obj, op_obj, toProfinite_obj, - functorToPresheaves_obj_obj, CompHausLike.coe_of, isoFinYoneda, locallyConstantIsoFinYoneda, - finYoneda_obj, LocallyConstant.toFun_eq_coe, NatTrans.comp_app, pointwiseLeftKanExtension_obj, - lanPresheafExt_inv, Iso.trans_inv, Iso.symm_inv, whiskerLeft_comp, lanPresheafNatIso_hom_app, - Opposite.op_unop, colimit.map_desc, id_eq, Functor.comp_map, op_map, colimit.ι_desc, - Cocones.precompose_obj_pt, Profinite.Extend.cocone_pt, Cocones.precompose_obj_ι, - Category.assoc, const_obj_obj, whiskerLeft_app, NatIso.ofComponents_hom_app, - NatIso.ofComponents_inv_app, Profinite.Extend.cocone_ι_app, counitApp, colimit.ι_desc_assoc] + simp only [comp_obj, CostructuredArrow.proj_obj, op_obj, functorToPresheaves_obj_obj, + isoFinYoneda, locallyConstantIsoFinYoneda, finYoneda_obj, LocallyConstant.toFun_eq_coe, + NatTrans.comp_app, pointwiseLeftKanExtension_obj, lanPresheafExt_inv, Iso.trans_inv, + Iso.symm_inv, whiskerLeft_comp, lanPresheafNatIso_hom_app, Opposite.op_unop, colimit.map_desc, + id_eq, Functor.comp_map, op_map, colimit.ι_desc, Cocones.precompose_obj_pt, + Profinite.Extend.cocone_pt, Cocones.precompose_obj_ι, Category.assoc, const_obj_obj, + whiskerLeft_app, NatIso.ofComponents_hom_app, NatIso.ofComponents_inv_app, + Profinite.Extend.cocone_ι_app, counitApp, colimit.ι_desc_assoc] erw [(counitApp.{u, u+1} X).naturality] simp only [← Category.assoc] congr @@ -318,8 +315,7 @@ noncomputable def isColimitLocallyConstantPresheaf (hc : IsLimit c) [∀ i, Epi (h : fi.comap (c.π.app i) = fj.comap (c.π.app j)) obtain ⟨k, ki, kj, _⟩ := IsCofilteredOrEmpty.cone_objs i j refine ⟨⟨k⟩, ki.op, kj.op, ?_⟩ - dsimp only [comp_obj, op_obj, functorToPresheaves_obj_obj, CompHausLike.coe_of, - Functor.comp_map, op_map, Quiver.Hom.unop_op, functorToPresheaves_obj_map] + dsimp ext x obtain ⟨x, hx⟩ := ((LightProfinite.epi_iff_surjective (c.π.app k)).mp inferInstance) x rw [← hx] @@ -541,14 +537,13 @@ lemma isoLocallyConstantOfIsColimit_inv (X : LightProfinite.{u}ᵒᵖ ⥤ Type u intro ⟨Y, _, g⟩ simp? [locallyConstantIsoFinYoneda, isoFinYoneda, counitApp] says simp only [comp_obj, CostructuredArrow.proj_obj, op_obj, functorToPresheaves_obj_obj, - toLightProfinite_obj_toTop_α, isoFinYoneda, locallyConstantIsoFinYoneda, finYoneda_obj, - LocallyConstant.toFun_eq_coe, NatTrans.comp_app, pointwiseLeftKanExtension_obj, - lanPresheafExt_inv, Iso.trans_inv, Iso.symm_inv, whiskerLeft_comp, lanPresheafNatIso_hom_app, - Opposite.op_unop, colimit.map_desc, id_eq, Functor.comp_map, op_map, colimit.ι_desc, - Cocones.precompose_obj_pt, LightProfinite.Extend.cocone_pt, Cocones.precompose_obj_ι, - Category.assoc, const_obj_obj, whiskerLeft_app, NatIso.ofComponents_hom_app, - NatIso.ofComponents_inv_app, LightProfinite.Extend.cocone_ι_app, counitApp, - colimit.ι_desc_assoc] + isoFinYoneda, locallyConstantIsoFinYoneda, finYoneda_obj, LocallyConstant.toFun_eq_coe, + NatTrans.comp_app, pointwiseLeftKanExtension_obj, lanPresheafExt_inv, Iso.trans_inv, + Iso.symm_inv, whiskerLeft_comp, lanPresheafNatIso_hom_app, Opposite.op_unop, colimit.map_desc, + id_eq, Functor.comp_map, op_map, colimit.ι_desc, Cocones.precompose_obj_pt, + LightProfinite.Extend.cocone_pt, Cocones.precompose_obj_ι, Category.assoc, const_obj_obj, + whiskerLeft_app, NatIso.ofComponents_hom_app, NatIso.ofComponents_inv_app, + LightProfinite.Extend.cocone_ι_app, counitApp, colimit.ι_desc_assoc] erw [(counitApp.{u, u} X).naturality] simp only [← Category.assoc] congr diff --git a/Mathlib/Topology/Category/LightProfinite/AsLimit.lean b/Mathlib/Topology/Category/LightProfinite/AsLimit.lean index 48104258a6815..bb092cf45b600 100644 --- a/Mathlib/Topology/Category/LightProfinite/AsLimit.lean +++ b/Mathlib/Topology/Category/LightProfinite/AsLimit.lean @@ -79,8 +79,7 @@ abbrev proj (n : ℕ) : S ⟶ S.diagram.obj ⟨n⟩ := S.asLimitCone.π.app ⟨n lemma lightToProfinite_map_proj_eq (n : ℕ) : lightToProfinite.map (S.proj n) = (lightToProfinite.obj S).asLimitCone.π.app _ := by - simp? says simp only [toCompHausLike_obj, Functor.comp_obj, - FintypeCat.toLightProfinite_obj_toTop_α, toCompHausLike_map, coe_of] + simp only [toCompHausLike_obj, Functor.comp_obj, toCompHausLike_map, coe_of] let c : Cone (S.diagram ⋙ lightToProfinite) := S.toLightDiagram.cone let hc : IsLimit c := S.toLightDiagram.isLimit exact liftedLimitMapsToOriginal_inv_map_π hc _ diff --git a/Mathlib/Topology/Category/LightProfinite/Basic.lean b/Mathlib/Topology/Category/LightProfinite/Basic.lean index 73dbb2fa55ef6..7d300f3905088 100644 --- a/Mathlib/Topology/Category/LightProfinite/Basic.lean +++ b/Mathlib/Topology/Category/LightProfinite/Basic.lean @@ -100,7 +100,7 @@ attribute [local instance] FintypeCat.discreteTopology /-- The natural functor from `Fintype` to `LightProfinite`, endowing a finite type with the discrete topology. -/ -@[simps!] +@[simps! map_apply] def FintypeCat.toLightProfinite : FintypeCat ⥤ LightProfinite where obj A := LightProfinite.of A map f := ⟨f, by continuity⟩ diff --git a/Mathlib/Topology/Category/Profinite/Basic.lean b/Mathlib/Topology/Category/Profinite/Basic.lean index 1faa3653c9a74..45aaebd8b6c14 100644 --- a/Mathlib/Topology/Category/Profinite/Basic.lean +++ b/Mathlib/Topology/Category/Profinite/Basic.lean @@ -140,7 +140,7 @@ attribute [local instance] FintypeCat.discreteTopology /-- The natural functor from `Fintype` to `Profinite`, endowing a finite type with the discrete topology. -/ -@[simps] +@[simps map_apply] def FintypeCat.toProfinite : FintypeCat ⥤ Profinite where obj A := Profinite.of A map f := ⟨f, by continuity⟩ From d9aad45391e594b51dfd27fdf9c579e105a0183c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 17 Oct 2024 15:41:48 +0000 Subject: [PATCH 020/131] feat(CategoryTheory/Shift): shifted morphisms in the opposite category (#17546) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit When `C` is a triangulated category, we introduce a bijection `ShiftedHom.opEquiv` which identifies `X ⟶ Y⟦n⟧` and `op Y ⟶ (op X)⟦n⟧` for any objects `X` and `Y` in `C`, and `n : ℤ`. Along with the compatibilities that are obtained, this shall be the main ingredient in the construction of the (contravariant) long exact sequence of `Ext` in abelian categories #15092. --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Shift/Basic.lean | 65 +++++++- .../Shift/ShiftedHomOpposite.lean | 155 ++++++++++++++++++ .../CategoryTheory/Triangulated/Opposite.lean | 55 ++++++- .../CategoryTheory/Triangulated/Yoneda.lean | 40 ++++- Mathlib/Combinatorics/Quiver/Basic.lean | 9 + 6 files changed, 320 insertions(+), 5 deletions(-) create mode 100644 Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean diff --git a/Mathlib.lean b/Mathlib.lean index 155e528def04f..b43069a00a054 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1858,6 +1858,7 @@ import Mathlib.CategoryTheory.Shift.Pullback import Mathlib.CategoryTheory.Shift.Quotient import Mathlib.CategoryTheory.Shift.ShiftSequence import Mathlib.CategoryTheory.Shift.ShiftedHom +import Mathlib.CategoryTheory.Shift.ShiftedHomOpposite import Mathlib.CategoryTheory.Shift.SingleFunctors import Mathlib.CategoryTheory.Sigma.Basic import Mathlib.CategoryTheory.Simple diff --git a/Mathlib/CategoryTheory/Shift/Basic.lean b/Mathlib/CategoryTheory/Shift/Basic.lean index c8b2f4ddfaec9..7acb86cbee29d 100644 --- a/Mathlib/CategoryTheory/Shift/Basic.lean +++ b/Mathlib/CategoryTheory/Shift/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kim Morrison, Johan Commelin, Andrew Yang +Authors: Kim Morrison, Johan Commelin, Andrew Yang, Joël Riou -/ import Mathlib.Algebra.Group.Basic import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Zero @@ -487,6 +487,69 @@ theorem shift_shiftFunctorCompIsoId_neg_add_cancel_inv_app (n : A) (X : C) : end +section + +variable (A) + +lemma shiftFunctorCompIsoId_zero_zero_hom_app (X : C) : + (shiftFunctorCompIsoId C 0 0 (add_zero 0)).hom.app X = + ((shiftFunctorZero C A).hom.app X)⟦0⟧' ≫ (shiftFunctorZero C A).hom.app X := by + simp [shiftFunctorCompIsoId, shiftFunctorAdd'_zero_add_inv_app] + +lemma shiftFunctorCompIsoId_zero_zero_inv_app (X : C) : + (shiftFunctorCompIsoId C 0 0 (add_zero 0)).inv.app X = + (shiftFunctorZero C A).inv.app X ≫ ((shiftFunctorZero C A).inv.app X)⟦0⟧' := by + simp [shiftFunctorCompIsoId, shiftFunctorAdd'_zero_add_hom_app] + +end + +section + +variable (m n p m' n' p' : A) (hm : m' + m = 0) (hn : n' + n = 0) (hp : p' + p = 0) + (h : m + n = p) + +lemma shiftFunctorCompIsoId_add'_inv_app : + (shiftFunctorCompIsoId C p' p hp).inv.app X = + (shiftFunctorCompIsoId C n' n hn).inv.app X ≫ + (shiftFunctorCompIsoId C m' m hm).inv.app (X⟦n'⟧)⟦n⟧' ≫ + (shiftFunctorAdd' C m n p h).inv.app (X⟦n'⟧⟦m'⟧) ≫ + ((shiftFunctorAdd' C n' m' p' + (by rw [← add_left_inj p, hp, ← h, add_assoc, + ← add_assoc m', hm, zero_add, hn])).inv.app X)⟦p⟧' := by + dsimp [shiftFunctorCompIsoId] + simp only [Functor.map_comp, Category.assoc] + congr 1 + erw [← NatTrans.naturality] + dsimp + rw [← cancel_mono ((shiftFunctorAdd' C p' p 0 hp).inv.app X), Iso.hom_inv_id_app, + Category.assoc, Category.assoc, Category.assoc, Category.assoc, + ← shiftFunctorAdd'_assoc_inv_app p' m n n' p 0 + (by rw [← add_left_inj n, hn, add_assoc, h, hp]) h (by rw [add_assoc, h, hp]), + ← Functor.map_comp_assoc, ← Functor.map_comp_assoc, ← Functor.map_comp_assoc, + Category.assoc, Category.assoc, + shiftFunctorAdd'_assoc_inv_app n' m' m p' 0 n' _ _ + (by rw [add_assoc, hm, add_zero]), Iso.hom_inv_id_app_assoc, + ← shiftFunctorAdd'_add_zero_hom_app, Iso.hom_inv_id_app, + Functor.map_id, Category.id_comp, Iso.hom_inv_id_app] + +lemma shiftFunctorCompIsoId_add'_hom_app : + (shiftFunctorCompIsoId C p' p hp).hom.app X = + ((shiftFunctorAdd' C n' m' p' + (by rw [← add_left_inj p, hp, ← h, add_assoc, + ← add_assoc m', hm, zero_add, hn])).hom.app X)⟦p⟧' ≫ + (shiftFunctorAdd' C m n p h).hom.app (X⟦n'⟧⟦m'⟧) ≫ + (shiftFunctorCompIsoId C m' m hm).hom.app (X⟦n'⟧)⟦n⟧' ≫ + (shiftFunctorCompIsoId C n' n hn).hom.app X := by + rw [← cancel_mono ((shiftFunctorCompIsoId C p' p hp).inv.app X), Iso.hom_inv_id_app, + shiftFunctorCompIsoId_add'_inv_app m n p m' n' p' hm hn hp h, + Category.assoc, Category.assoc, Category.assoc, Iso.hom_inv_id_app_assoc, + ← Functor.map_comp_assoc, Iso.hom_inv_id_app] + dsimp + rw [Functor.map_id, Category.id_comp, Iso.hom_inv_id_app_assoc, + ← Functor.map_comp, Iso.hom_inv_id_app, Functor.map_id] + +end + open CategoryTheory.Limits variable [HasZeroMorphisms C] diff --git a/Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean b/Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean new file mode 100644 index 0000000000000..5f21658547de3 --- /dev/null +++ b/Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import Mathlib.CategoryTheory.Triangulated.Opposite +import Mathlib.CategoryTheory.Shift.ShiftedHom + +/-! Shifted morphisms in the opposite category + +If `C` is a category equipped with a shift by `ℤ`, `X` and `Y` are objects +of `C`, and `n : ℤ`, we define a bijection +`ShiftedHom.opEquiv : ShiftedHom X Y n ≃ ShiftedHom (Opposite.op Y) (Opposite.op X) n`. +We also introduce `ShiftedHom.opEquiv'` which produces a bijection +`ShiftedHom X Y a' ≃ (Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦n⟧)` when `n + a = a'`. +The compatibilities that are obtained shall be used in order to study +the homological functor `preadditiveYoneda.obj B : Cᵒᵖ ⥤ Type _` when `B` is an object +in a pretriangulated category `C`. + +-/ + +namespace CategoryTheory + +open Category Pretriangulated.Opposite Pretriangulated + +variable {C : Type*} [Category C] [HasShift C ℤ] {X Y Z : C} + +namespace ShiftedHom + +/-- The bijection `ShiftedHom X Y n ≃ ShiftedHom (Opposite.op Y) (Opposite.op X) n` when +`n : ℤ`, and `X` and `Y` are objects of a category equipped with a shift by `ℤ`. -/ +noncomputable def opEquiv (n : ℤ) : + ShiftedHom X Y n ≃ ShiftedHom (Opposite.op Y) (Opposite.op X) n := + Quiver.Hom.opEquiv.trans + ((opShiftFunctorEquivalence C n).symm.toAdjunction.homEquiv (Opposite.op Y) (Opposite.op X)) + +lemma opEquiv_symm_apply {n : ℤ} (f : ShiftedHom (Opposite.op Y) (Opposite.op X) n) : + (opEquiv n).symm f = + ((opShiftFunctorEquivalence C n).unitIso.inv.app (Opposite.op X)).unop ≫ f.unop⟦n⟧' := by + rfl + +lemma opEquiv_symm_apply_comp {X Y : C} {a : ℤ} + (f : ShiftedHom (Opposite.op X) (Opposite.op Y) a) {b : ℤ} {Z : C} + (z : ShiftedHom X Z b) {c : ℤ} (h : b + a = c) : + ((ShiftedHom.opEquiv a).symm f).comp z h = + (ShiftedHom.opEquiv a).symm (z.op ≫ f) ≫ + (shiftFunctorAdd' C b a c h).inv.app Z := by + rw [ShiftedHom.opEquiv_symm_apply, ShiftedHom.opEquiv_symm_apply, + ShiftedHom.comp] + dsimp + simp only [assoc, Functor.map_comp] + +lemma opEquiv_symm_comp {a b : ℤ} + (f : ShiftedHom (Opposite.op Z) (Opposite.op Y) a) + (g : ShiftedHom (Opposite.op Y) (Opposite.op X) b) + {c : ℤ} (h : b + a = c) : + (opEquiv _).symm (f.comp g h) = + ((opEquiv _).symm g).comp ((opEquiv _).symm f) (by omega) := by + rw [opEquiv_symm_apply, opEquiv_symm_apply, + opShiftFunctorEquivalence_unitIso_inv_app_eq _ _ _ _ (show a + b = c by omega), comp, comp] + dsimp + rw [assoc, assoc, assoc, assoc, ← Functor.map_comp, ← unop_comp_assoc, + Iso.inv_hom_id_app] + dsimp + rw [assoc, id_comp, Functor.map_comp, ← NatTrans.naturality_assoc, + ← NatTrans.naturality, opEquiv_symm_apply] + dsimp + rw [← Functor.map_comp_assoc, ← Functor.map_comp_assoc, + ← Functor.map_comp_assoc] + rw [← unop_comp_assoc] + erw [← NatTrans.naturality] + rfl + +/-- The bijection `ShiftedHom X Y a' ≃ (Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦n⟧)` +when integers `n`, `a` and `a'` satisfy `n + a = a'`, and `X` and `Y` are objects +of a category equipped with a shift by `ℤ`. -/ +noncomputable def opEquiv' (n a a' : ℤ) (h : n + a = a') : + ShiftedHom X Y a' ≃ (Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦n⟧) := + ((shiftFunctorAdd' C a n a' (by omega)).symm.app Y).homToEquiv.symm.trans (opEquiv n) + +lemma opEquiv'_symm_apply {n a : ℤ} (f : Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦n⟧) + (a' : ℤ) (h : n + a = a') : + (opEquiv' n a a' h).symm f = + (opEquiv n).symm f ≫ (shiftFunctorAdd' C a n a' (by omega)).inv.app _ := + rfl + +lemma opEquiv'_apply {a' : ℤ} (f : ShiftedHom X Y a') (n a : ℤ) (h : n + a = a') : + opEquiv' n a a' h f = + opEquiv n (f ≫ (shiftFunctorAdd' C a n a' (by omega)).hom.app Y) := by + rfl + +lemma opEquiv'_symm_op_opShiftFunctorEquivalence_counitIso_inv_app_op_shift + {n m : ℤ} (f : ShiftedHom X Y n) (g : ShiftedHom Y Z m) + (q : ℤ) (hq : n + m = q) : + (opEquiv' n m q hq).symm + (g.op ≫ (opShiftFunctorEquivalence C n).counitIso.inv.app _ ≫ f.op⟦n⟧') = + f.comp g (by omega) := by + rw [opEquiv'_symm_apply, opEquiv_symm_apply] + dsimp [comp] + apply Quiver.Hom.op_inj + simp only [assoc, Functor.map_comp, op_comp, Quiver.Hom.op_unop, + opShiftFunctorEquivalence_unitIso_inv_naturality] + erw [(opShiftFunctorEquivalence C n).inverse_counitInv_comp_assoc (Opposite.op Y)] + +lemma opEquiv'_symm_comp (f : Y ⟶ X) {n a : ℤ} (x : Opposite.op (Z⟦a⟧) ⟶ (Opposite.op X⟦n⟧)) + (a' : ℤ) (h : n + a = a') : + (opEquiv' n a a' h).symm (x ≫ f.op⟦n⟧') = f ≫ (opEquiv' n a a' h).symm x := + Quiver.Hom.op_inj (by simp [opEquiv'_symm_apply, opEquiv_symm_apply]) + +lemma opEquiv'_zero_add_symm (a : ℤ) (f : Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦(0 : ℤ)⟧) : + (opEquiv' 0 a a (zero_add a)).symm f = + ((shiftFunctorZero Cᵒᵖ ℤ).hom.app _).unop ≫ f.unop := by + simp [opEquiv'_symm_apply, opEquiv_symm_apply, shiftFunctorAdd'_add_zero, + opShiftFunctorEquivalence_zero_unitIso_inv_app] + +lemma opEquiv'_add_symm (n m a a' a'' : ℤ) (ha' : n + a = a') (ha'' : m + a' = a'') + (x : (Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦m + n⟧)) : + (opEquiv' (m + n) a a'' (by omega)).symm x = + (opEquiv' m a' a'' ha'').symm ((opEquiv' n a a' ha').symm + (x ≫ (shiftFunctorAdd Cᵒᵖ m n).hom.app _)).op := by + simp only [opEquiv'_symm_apply, opEquiv_symm_apply, + opShiftFunctorEquivalence_unitIso_inv_app_eq _ _ _ _ (add_comm n m)] + dsimp + simp only [assoc, Functor.map_comp, ← shiftFunctorAdd'_eq_shiftFunctorAdd, + ← NatTrans.naturality_assoc, + shiftFunctorAdd'_assoc_inv_app a n m a' (m + n) a'' (by omega) (by omega) (by omega)] + rfl + +section Preadditive + +variable [Preadditive C] [∀ (n : ℤ), (shiftFunctor C n).Additive] + +@[simp] +lemma opEquiv_symm_add {n : ℤ} (x y : ShiftedHom (Opposite.op Y) (Opposite.op X) n) : + (opEquiv n).symm (x + y) = (opEquiv n).symm x + (opEquiv n).symm y := by + dsimp [opEquiv_symm_apply] + rw [← Preadditive.comp_add, ← Functor.map_add] + rfl + +@[simp] +lemma opEquiv'_symm_add {n a : ℤ} (x y : (Opposite.op (Y⟦a⟧) ⟶ (Opposite.op X)⟦n⟧)) + (a' : ℤ) (h : n + a = a') : + (opEquiv' n a a' h).symm (x + y) = + (opEquiv' n a a' h).symm x + (opEquiv' n a a' h).symm y := by + dsimp [opEquiv'] + erw [opEquiv_symm_add, Iso.homToEquiv_apply, Iso.homToEquiv_apply, Iso.homToEquiv_apply] + rw [Preadditive.add_comp] + rfl + +end Preadditive + +end ShiftedHom + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite.lean b/Mathlib/CategoryTheory/Triangulated/Opposite.lean index 3b6c664323762..9e34b1cc2e37f 100644 --- a/Mathlib/CategoryTheory/Triangulated/Opposite.lean +++ b/Mathlib/CategoryTheory/Triangulated/Opposite.lean @@ -130,8 +130,7 @@ lemma shiftFunctor_op_map (n m : ℤ) (hnm : n + m = 0) {K L : Cᵒᵖ} (φ : K (shiftFunctorOpIso C n m hnm).inv.app L := (NatIso.naturality_2 (shiftFunctorOpIso C n m hnm) φ).symm -variable (C) - +variable (C) in /-- The autoequivalence `Cᵒᵖ ≌ Cᵒᵖ` whose functor is `shiftFunctor Cᵒᵖ n` and whose inverse functor is `(shiftFunctor C n).op`. Do not unfold the definitions of the unit and counit isomorphisms: the compatibilities they satisfy are stated as separate lemmas. -/ @@ -177,7 +176,57 @@ lemma opShiftFunctorEquivalence_counitIso_inv_naturality (n : ℤ) {X Y : Cᵒ (opShiftFunctorEquivalence C n).counitIso.inv.app X ≫ f.unop⟦n⟧'.op⟦n⟧' := (opShiftFunctorEquivalence C n).counitIso.inv.naturality f -variable {C} +lemma opShiftFunctorEquivalence_zero_unitIso_hom_app (X : Cᵒᵖ) : + (opShiftFunctorEquivalence C 0).unitIso.hom.app X = + ((shiftFunctorZero C ℤ).hom.app X.unop).op ≫ + (((shiftFunctorZero Cᵒᵖ ℤ).inv.app X).unop⟦(0 : ℤ)⟧').op := by + apply Quiver.Hom.unop_inj + dsimp [opShiftFunctorEquivalence] + rw [shiftFunctorZero_op_inv_app, unop_comp, Quiver.Hom.unop_op, Functor.map_comp, + shiftFunctorCompIsoId_zero_zero_hom_app, assoc] + +lemma opShiftFunctorEquivalence_zero_unitIso_inv_app (X : Cᵒᵖ) : + (opShiftFunctorEquivalence C 0).unitIso.inv.app X = + (((shiftFunctorZero Cᵒᵖ ℤ).hom.app X).unop⟦(0 : ℤ)⟧').op ≫ + ((shiftFunctorZero C ℤ).inv.app X.unop).op := by + apply Quiver.Hom.unop_inj + dsimp [opShiftFunctorEquivalence] + rw [shiftFunctorZero_op_hom_app, unop_comp, Quiver.Hom.unop_op, Functor.map_comp, + shiftFunctorCompIsoId_zero_zero_inv_app, assoc] + +lemma opShiftFunctorEquivalence_unitIso_hom_app_eq (X : Cᵒᵖ) (m n p : ℤ) (h : m + n = p) : + (opShiftFunctorEquivalence C p).unitIso.hom.app X = + (opShiftFunctorEquivalence C n).unitIso.hom.app X ≫ + (((opShiftFunctorEquivalence C m).unitIso.hom.app (X⟦n⟧)).unop⟦n⟧').op ≫ + ((shiftFunctorAdd' C m n p h).hom.app _).op ≫ + (((shiftFunctorAdd' Cᵒᵖ n m p (by omega)).inv.app X).unop⟦p⟧').op := by + dsimp [opShiftFunctorEquivalence] + simp only [shiftFunctorAdd'_op_inv_app _ n m p (by omega) _ _ _ (add_neg_cancel n) + (add_neg_cancel m) (add_neg_cancel p), shiftFunctor_op_map _ _ (add_neg_cancel m), + Category.assoc, Iso.inv_hom_id_app_assoc] + erw [Functor.map_id, Functor.map_id, Functor.map_id, Functor.map_id, + id_comp, id_comp, id_comp, comp_id, comp_id] + dsimp + rw [comp_id, shiftFunctorCompIsoId_add'_hom_app _ _ _ _ _ _ + (neg_add_cancel m) (neg_add_cancel n) (neg_add_cancel p) h] + dsimp + rw [Category.assoc, Category.assoc] + rfl + +lemma opShiftFunctorEquivalence_unitIso_inv_app_eq (X : Cᵒᵖ) (m n p : ℤ) (h : m + n = p) : + (opShiftFunctorEquivalence C p).unitIso.inv.app X = + (((shiftFunctorAdd' Cᵒᵖ n m p (by omega)).hom.app X).unop⟦p⟧').op ≫ + ((shiftFunctorAdd' C m n p h).inv.app _).op ≫ + (((opShiftFunctorEquivalence C m).unitIso.inv.app (X⟦n⟧)).unop⟦n⟧').op ≫ + (opShiftFunctorEquivalence C n).unitIso.inv.app X := by + rw [← cancel_mono ((opShiftFunctorEquivalence C p).unitIso.hom.app X), Iso.inv_hom_id_app, + opShiftFunctorEquivalence_unitIso_hom_app_eq _ _ _ _ h, + Category.assoc, Category.assoc, Category.assoc, Iso.inv_hom_id_app_assoc] + apply Quiver.Hom.unop_inj + dsimp + simp only [Category.assoc, ← Functor.map_comp_assoc, Iso.hom_inv_id_app_assoc, + ← unop_comp, Iso.inv_hom_id_app, Functor.comp_obj, Functor.op_obj, unop_id, + Functor.map_id, id_comp, ← Functor.map_comp, Iso.hom_inv_id_app] lemma shift_unop_opShiftFunctorEquivalence_counitIso_inv_app (X : Cᵒᵖ) (n : ℤ) : ((opShiftFunctorEquivalence C n).counitIso.inv.app X).unop⟦n⟧' = diff --git a/Mathlib/CategoryTheory/Triangulated/Yoneda.lean b/Mathlib/CategoryTheory/Triangulated/Yoneda.lean index 39ec3c60ebdb1..e136effd690d8 100644 --- a/Mathlib/CategoryTheory/Triangulated/Yoneda.lean +++ b/Mathlib/CategoryTheory/Triangulated/Yoneda.lean @@ -5,6 +5,7 @@ Authors: Joël Riou -/ import Mathlib.Algebra.Homology.ShortComplex.Ab import Mathlib.CategoryTheory.Preadditive.Yoneda.Basic +import Mathlib.CategoryTheory.Shift.ShiftedHomOpposite import Mathlib.CategoryTheory.Triangulated.HomologicalFunctor import Mathlib.CategoryTheory.Triangulated.Opposite @@ -23,7 +24,7 @@ variable {C : Type*} [Category C] [Preadditive C] [HasShift C ℤ] namespace CategoryTheory -open Limits Pretriangulated.Opposite +open Limits Opposite Pretriangulated.Opposite namespace Pretriangulated @@ -62,6 +63,43 @@ lemma preadditiveCoyoneda_homologySequenceδ_apply x ≫ T.mor₃⟦n₀⟧' ≫ (shiftFunctorAdd' C 1 n₀ n₁ (by omega)).inv.app _ := by apply Category.assoc +section + +variable [∀ (n : ℤ), (shiftFunctor C n).Additive] + +noncomputable instance (B : C) : (preadditiveYoneda.obj B).ShiftSequence ℤ where + sequence n := preadditiveYoneda.obj (B⟦n⟧) + isoZero := preadditiveYoneda.mapIso ((shiftFunctorZero C ℤ).app B) + shiftIso n a a' h := NatIso.ofComponents (fun A ↦ AddEquiv.toAddCommGrpIso + { toEquiv := Quiver.Hom.opEquiv.trans (ShiftedHom.opEquiv' n a a' h).symm + map_add' := fun _ _ ↦ ShiftedHom.opEquiv'_symm_add _ _ _ h }) + (by intros; ext; apply ShiftedHom.opEquiv'_symm_comp _ _ _ h) + shiftIso_zero a := by ext; apply ShiftedHom.opEquiv'_zero_add_symm + shiftIso_add n m a a' a'' ha' ha'' := by + ext _ x + exact ShiftedHom.opEquiv'_add_symm n m a a' a'' ha' ha'' x.op + +lemma preadditiveYoneda_shiftMap_apply (B : C) {X Y : Cᵒᵖ} (n : ℤ) (f : X ⟶ Y⟦n⟧) + (a a' : ℤ) (h : n + a = a') (z : X.unop ⟶ B⟦a⟧) : + (preadditiveYoneda.obj B).shiftMap f a a' h z = + ((ShiftedHom.opEquiv _).symm f).comp z (show a + n = a' by omega) := by + symm + apply ShiftedHom.opEquiv_symm_apply_comp + +lemma preadditiveYoneda_homologySequenceδ_apply + (T : Triangle C) (n₀ n₁ : ℤ) (h : n₀ + 1 = n₁) {B : C} (x : T.obj₁ ⟶ B⟦n₀⟧) : + (preadditiveYoneda.obj B).homologySequenceδ + ((triangleOpEquivalence _).functor.obj (op T)) n₀ n₁ h x = + T.mor₃ ≫ x⟦(1 : ℤ)⟧' ≫ (shiftFunctorAdd' C n₀ 1 n₁ h).inv.app B := by + simp only [Functor.homologySequenceδ, preadditiveYoneda_shiftMap_apply, + ShiftedHom.comp, ← Category.assoc] + congr 2 + apply (ShiftedHom.opEquiv _).injective + rw [Equiv.apply_symm_apply] + rfl + +end + end Pretriangulated end CategoryTheory diff --git a/Mathlib/Combinatorics/Quiver/Basic.lean b/Mathlib/Combinatorics/Quiver/Basic.lean index f0f77f70aa4b7..78d2fbdeef06a 100644 --- a/Mathlib/Combinatorics/Quiver/Basic.lean +++ b/Mathlib/Combinatorics/Quiver/Basic.lean @@ -137,6 +137,15 @@ def Hom.op {V} [Quiver V] {X Y : V} (f : X ⟶ Y) : op Y ⟶ op X := ⟨f⟩ /-- Given an arrow in `Vᵒᵖ`, we can take the "unopposite" back in `V`. -/ def Hom.unop {V} [Quiver V] {X Y : Vᵒᵖ} (f : X ⟶ Y) : unop Y ⟶ unop X := Opposite.unop f +/-- The bijection `(X ⟶ Y) ≃ (op Y ⟶ op X)`. -/ +@[simps] +def Hom.opEquiv {V} [Quiver V] {X Y : V} : + (X ⟶ Y) ≃ (Opposite.op Y ⟶ Opposite.op X) where + toFun := Opposite.op + invFun := Opposite.unop + left_inv _ := rfl + right_inv _ := rfl + /-- A type synonym for a quiver with no arrows. -/ -- Porting note(#5171): this linter isn't ported yet. -- @[nolint has_nonempty_instance] From cbcf4ea1cb4620e09549662684cb12de35d81299 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 17 Oct 2024 15:41:50 +0000 Subject: [PATCH 021/131] feat(Analysis/Normed/Group/Ultra): triangle ineq for tsum & tprod (#17584) --- Mathlib/Analysis/Normed/Group/Ultra.lean | 42 ++++++++++++++++++++++++ 1 file changed, 42 insertions(+) diff --git a/Mathlib/Analysis/Normed/Group/Ultra.lean b/Mathlib/Analysis/Normed/Group/Ultra.lean index 8e83b3535e430..2b8c5f341691f 100644 --- a/Mathlib/Analysis/Normed/Group/Ultra.lean +++ b/Mathlib/Analysis/Normed/Group/Ultra.lean @@ -6,6 +6,8 @@ Authors: Yakov Pechersky, David Loeffler import Mathlib.Analysis.Normed.Group.Uniform import Mathlib.Topology.Algebra.Nonarchimedean.Basic import Mathlib.Topology.MetricSpace.Ultra.Basic +import Mathlib.Topology.Algebra.InfiniteSum.Group +import Mathlib.Topology.Algebra.Order.LiminfLimsup /-! # Ultrametric norms @@ -31,6 +33,8 @@ ultrametric, nonarchimedean -/ open Metric NNReal +open scoped BigOperators + namespace IsUltrametricDist section Group @@ -235,6 +239,44 @@ lemma norm_prod_le_of_forall_le_of_nonneg {s : Finset ι} {f : ι → M} {C : lift C to NNReal using h_nonneg exact nnnorm_prod_le_of_forall_le hC +@[to_additive] +lemma norm_tprod_le (f : ι → M) : ‖∏' i, f i‖ ≤ ⨆ i, ‖f i‖ := by + rcases isEmpty_or_nonempty ι with hι | hι + · -- Silly case #1 : the index type is empty + simp only [tprod_empty, norm_one', Real.iSup_of_isEmpty, le_refl] + by_cases h : Multipliable f; swap + · -- Silly case #2 : the product is divergent + rw [tprod_eq_one_of_not_multipliable h, norm_one'] + by_cases h_bd : BddAbove (Set.range fun i ↦ ‖f i‖) + · exact le_ciSup_of_le h_bd hι.some (norm_nonneg' _) + · rw [Real.iSup_of_not_bddAbove h_bd] + -- now the interesting case + have h_bd : BddAbove (Set.range fun i ↦ ‖f i‖) := + h.tendsto_cofinite_one.norm'.bddAbove_range_of_cofinite + refine le_of_tendsto' h.hasProd.norm' (fun s ↦ norm_prod_le_of_forall_le_of_nonneg ?_ ?_) + · exact le_ciSup_of_le h_bd hι.some (norm_nonneg' _) + · exact fun i _ ↦ le_ciSup h_bd i + +@[to_additive] +lemma nnnorm_tprod_le (f : ι → M) : ‖∏' i, f i‖₊ ≤ ⨆ i, ‖f i‖₊ := by + simpa only [← NNReal.coe_le_coe, coe_nnnorm', coe_iSup] using norm_tprod_le f + +@[to_additive] +lemma norm_tprod_le_of_forall_le [Nonempty ι] {f : ι → M} {C : ℝ} (h : ∀ i, ‖f i‖ ≤ C) : + ‖∏' i, f i‖ ≤ C := + (norm_tprod_le f).trans (ciSup_le h) + +@[to_additive] +lemma norm_tprod_le_of_forall_le_of_nonneg {f : ι → M} {C : ℝ} (hC : 0 ≤ C) (h : ∀ i, ‖f i‖ ≤ C) : + ‖∏' i, f i‖ ≤ C := by + rcases isEmpty_or_nonempty ι + · simpa only [tprod_empty, norm_one'] using hC + · exact norm_tprod_le_of_forall_le h + +@[to_additive] +lemma nnnorm_tprod_le_of_forall_le {f : ι → M} {C : ℝ≥0} (h : ∀ i, ‖f i‖₊ ≤ C) : ‖∏' i, f i‖₊ ≤ C := + (nnnorm_tprod_le f).trans (ciSup_le' h) + end CommGroup end IsUltrametricDist From 0456b9a130212be2501255a28f73dbc146e05de2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:41:52 +0000 Subject: [PATCH 022/131] feat: `#s` as notation for `Finset.card s` (#17646) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I believe it is time that we finally introduce this notation. This makes serious finset calculations significantly nicer (see eg #5297) and is standard notation. `|s|` would be an alternative, but absolute value elaboration is very complicated like that, so I would rather not, and `#{x | p x}` looks nicer than `|{x | p x}|` (although in return `|s ∩ t|` looks nicer than `#(s ∩ t)`). I am scoping the notation to the `Finset` namespace to avoid clashes with `Cardinal.mk`, which also uses that notation. --- Mathlib/Data/Finset/Card.lean | 316 +++++++++++++++++----------------- 1 file changed, 161 insertions(+), 155 deletions(-) diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 653bbfc7a2fe2..376e7fe70747f 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -12,7 +12,7 @@ This defines the cardinality of a `Finset` and provides induction principles for ## Main declarations -* `Finset.card`: `s.card : ℕ` returns the cardinality of `s : Finset α`. +* `Finset.card`: `#s : ℕ` returns the cardinality of `s : Finset α`. ### Induction principles @@ -36,52 +36,56 @@ namespace Finset variable {s t : Finset α} {a b : α} -/-- `s.card` is the number of elements of `s`, aka its cardinality. -/ +/-- `s.card` is the number of elements of `s`, aka its cardinality. + +The notation `#s` can be accessed in the `Finset` locale. -/ def card (s : Finset α) : ℕ := Multiset.card s.1 -theorem card_def (s : Finset α) : s.card = Multiset.card s.1 := +@[inherit_doc] scoped prefix:arg "#" => Finset.card + +theorem card_def (s : Finset α) : #s = Multiset.card s.1 := rfl -@[simp] lemma card_val (s : Finset α) : Multiset.card s.1 = s.card := rfl +@[simp] lemma card_val (s : Finset α) : Multiset.card s.1 = #s := rfl @[simp] -theorem card_mk {m nodup} : (⟨m, nodup⟩ : Finset α).card = Multiset.card m := +theorem card_mk {m nodup} : #(⟨m, nodup⟩ : Finset α) = Multiset.card m := rfl @[simp] -theorem card_empty : card (∅ : Finset α) = 0 := +theorem card_empty : #(∅ : Finset α) = 0 := rfl @[gcongr] -theorem card_le_card : s ⊆ t → s.card ≤ t.card := +theorem card_le_card : s ⊆ t → #s ≤ #t := Multiset.card_le_card ∘ val_le_iff.mpr @[mono] theorem card_mono : Monotone (@card α) := by apply card_le_card -@[simp] lemma card_eq_zero : s.card = 0 ↔ s = ∅ := Multiset.card_eq_zero.trans val_eq_zero -lemma card_ne_zero : s.card ≠ 0 ↔ s.Nonempty := card_eq_zero.ne.trans nonempty_iff_ne_empty.symm -@[simp] lemma card_pos : 0 < s.card ↔ s.Nonempty := Nat.pos_iff_ne_zero.trans card_ne_zero -@[simp] lemma one_le_card : 1 ≤ s.card ↔ s.Nonempty := card_pos +@[simp] lemma card_eq_zero : #s = 0 ↔ s = ∅ := Multiset.card_eq_zero.trans val_eq_zero +lemma card_ne_zero : #s ≠ 0 ↔ s.Nonempty := card_eq_zero.ne.trans nonempty_iff_ne_empty.symm +@[simp] lemma card_pos : 0 < #s ↔ s.Nonempty := Nat.pos_iff_ne_zero.trans card_ne_zero +@[simp] lemma one_le_card : 1 ≤ #s ↔ s.Nonempty := card_pos alias ⟨_, Nonempty.card_pos⟩ := card_pos alias ⟨_, Nonempty.card_ne_zero⟩ := card_ne_zero -theorem card_ne_zero_of_mem (h : a ∈ s) : s.card ≠ 0 := +theorem card_ne_zero_of_mem (h : a ∈ s) : #s ≠ 0 := (not_congr card_eq_zero).2 <| ne_empty_of_mem h @[simp] -theorem card_singleton (a : α) : card ({a} : Finset α) = 1 := +theorem card_singleton (a : α) : #{a} = 1 := Multiset.card_singleton _ -theorem card_singleton_inter [DecidableEq α] : ({a} ∩ s).card ≤ 1 := by +theorem card_singleton_inter [DecidableEq α] : #({a} ∩ s) ≤ 1 := by cases' Finset.decidableMem a s with h h · simp [Finset.singleton_inter_of_not_mem h] · simp [Finset.singleton_inter_of_mem h] @[simp] -theorem card_cons (h : a ∉ s) : (s.cons a h).card = s.card + 1 := +theorem card_cons (h : a ∉ s) : #(s.cons a h) = #s + 1 := Multiset.card_cons _ _ section InsertErase @@ -89,12 +93,12 @@ section InsertErase variable [DecidableEq α] @[simp] -theorem card_insert_of_not_mem (h : a ∉ s) : (insert a s).card = s.card + 1 := by +theorem card_insert_of_not_mem (h : a ∉ s) : #(insert a s) = #s + 1 := by rw [← cons_eq_insert _ _ h, card_cons] -theorem card_insert_of_mem (h : a ∈ s) : card (insert a s) = s.card := by rw [insert_eq_of_mem h] +theorem card_insert_of_mem (h : a ∈ s) : #(insert a s) = #s := by rw [insert_eq_of_mem h] -theorem card_insert_le (a : α) (s : Finset α) : card (insert a s) ≤ s.card + 1 := by +theorem card_insert_le (a : α) (s : Finset α) : #(insert a s) ≤ #s + 1 := by by_cases h : a ∈ s · rw [insert_eq_of_mem h] exact Nat.le_succ _ @@ -104,43 +108,43 @@ section variable {a b c d e f : α} -theorem card_le_two : card {a, b} ≤ 2 := card_insert_le _ _ +theorem card_le_two : #{a, b} ≤ 2 := card_insert_le _ _ -theorem card_le_three : card {a, b, c} ≤ 3 := +theorem card_le_three : #{a, b, c} ≤ 3 := (card_insert_le _ _).trans (Nat.succ_le_succ card_le_two) -theorem card_le_four : card {a, b, c, d} ≤ 4 := +theorem card_le_four : #{a, b, c, d} ≤ 4 := (card_insert_le _ _).trans (Nat.succ_le_succ card_le_three) -theorem card_le_five : card {a, b, c, d, e} ≤ 5 := +theorem card_le_five : #{a, b, c, d, e} ≤ 5 := (card_insert_le _ _).trans (Nat.succ_le_succ card_le_four) -theorem card_le_six : card {a, b, c, d, e, f} ≤ 6 := +theorem card_le_six : #{a, b, c, d, e, f} ≤ 6 := (card_insert_le _ _).trans (Nat.succ_le_succ card_le_five) end /-- If `a ∈ s` is known, see also `Finset.card_insert_of_mem` and `Finset.card_insert_of_not_mem`. -/ -theorem card_insert_eq_ite : card (insert a s) = if a ∈ s then s.card else s.card + 1 := by +theorem card_insert_eq_ite : #(insert a s) = if a ∈ s then #s else #s + 1 := by by_cases h : a ∈ s · rw [card_insert_of_mem h, if_pos h] · rw [card_insert_of_not_mem h, if_neg h] @[simp] -theorem card_pair_eq_one_or_two : ({a,b} : Finset α).card = 1 ∨ ({a,b} : Finset α).card = 2 := by +theorem card_pair_eq_one_or_two : #{a, b} = 1 ∨ #{a, b} = 2 := by simp [card_insert_eq_ite] tauto @[simp] -theorem card_pair (h : a ≠ b) : ({a, b} : Finset α).card = 2 := by +theorem card_pair (h : a ≠ b) : #{a, b} = 2 := by rw [card_insert_of_not_mem (not_mem_singleton.2 h), card_singleton] @[deprecated (since := "2024-01-04")] alias card_doubleton := Finset.card_pair /-- $\#(s \setminus \{a\}) = \#s - 1$ if $a \in s$. -/ @[simp] -theorem card_erase_of_mem : a ∈ s → (s.erase a).card = s.card - 1 := +theorem card_erase_of_mem : a ∈ s → #(s.erase a) = #s - 1 := Multiset.card_erase_of_mem /-- $\#(s \setminus \{a\}) = \#s - 1$ if $a \in s$. @@ -148,55 +152,57 @@ theorem card_erase_of_mem : a ∈ s → (s.erase a).card = s.card - 1 := so that we don't have to work with `ℕ`-subtraction. -/ @[simp] theorem cast_card_erase_of_mem {R} [AddGroupWithOne R] {s : Finset α} (hs : a ∈ s) : - ((s.erase a).card : R) = s.card - 1 := by + (#(s.erase a) : R) = #s - 1 := by rw [card_erase_of_mem hs, Nat.cast_sub, Nat.cast_one] rw [Nat.add_one_le_iff, Finset.card_pos] exact ⟨a, hs⟩ @[simp] -theorem card_erase_add_one : a ∈ s → (s.erase a).card + 1 = s.card := +theorem card_erase_add_one : a ∈ s → #(s.erase a) + 1 = #s := Multiset.card_erase_add_one -theorem card_erase_lt_of_mem : a ∈ s → (s.erase a).card < s.card := +theorem card_erase_lt_of_mem : a ∈ s → #(s.erase a) < #s := Multiset.card_erase_lt_of_mem -theorem card_erase_le : (s.erase a).card ≤ s.card := +theorem card_erase_le : #(s.erase a) ≤ #s := Multiset.card_erase_le -theorem pred_card_le_card_erase : s.card - 1 ≤ (s.erase a).card := by +theorem pred_card_le_card_erase : #s - 1 ≤ #(s.erase a) := by by_cases h : a ∈ s · exact (card_erase_of_mem h).ge · rw [erase_eq_of_not_mem h] exact Nat.sub_le _ _ /-- If `a ∈ s` is known, see also `Finset.card_erase_of_mem` and `Finset.erase_eq_of_not_mem`. -/ -theorem card_erase_eq_ite : (s.erase a).card = if a ∈ s then s.card - 1 else s.card := +theorem card_erase_eq_ite : #(s.erase a) = if a ∈ s then #s - 1 else #s := Multiset.card_erase_eq_ite end InsertErase @[simp] -theorem card_range (n : ℕ) : (range n).card = n := +theorem card_range (n : ℕ) : #(range n) = n := Multiset.card_range n @[simp] -theorem card_attach : s.attach.card = s.card := +theorem card_attach : #s.attach = #s := Multiset.card_attach end Finset +open scoped Finset + section ToMLListultiset variable [DecidableEq α] (m : Multiset α) (l : List α) -theorem Multiset.card_toFinset : m.toFinset.card = Multiset.card m.dedup := +theorem Multiset.card_toFinset : #m.toFinset = Multiset.card m.dedup := rfl -theorem Multiset.toFinset_card_le : m.toFinset.card ≤ Multiset.card m := +theorem Multiset.toFinset_card_le : #m.toFinset ≤ Multiset.card m := card_le_card <| dedup_le _ theorem Multiset.toFinset_card_of_nodup {m : Multiset α} (h : m.Nodup) : - m.toFinset.card = Multiset.card m := + #m.toFinset = Multiset.card m := congr_arg card <| Multiset.dedup_eq_self.mpr h theorem Multiset.dedup_card_eq_card_iff_nodup {m : Multiset α} : @@ -204,15 +210,15 @@ theorem Multiset.dedup_card_eq_card_iff_nodup {m : Multiset α} : .trans ⟨fun h ↦ eq_of_le_of_card_le (dedup_le m) h.ge, congr_arg _⟩ dedup_eq_self theorem Multiset.toFinset_card_eq_card_iff_nodup {m : Multiset α} : - m.toFinset.card = card m ↔ m.Nodup := dedup_card_eq_card_iff_nodup + #m.toFinset = card m ↔ m.Nodup := dedup_card_eq_card_iff_nodup -theorem List.card_toFinset : l.toFinset.card = l.dedup.length := +theorem List.card_toFinset : #l.toFinset = l.dedup.length := rfl -theorem List.toFinset_card_le : l.toFinset.card ≤ l.length := +theorem List.toFinset_card_le : #l.toFinset ≤ l.length := Multiset.toFinset_card_le ⟦l⟧ -theorem List.toFinset_card_of_nodup {l : List α} (h : l.Nodup) : l.toFinset.card = l.length := +theorem List.toFinset_card_of_nodup {l : List α} (h : l.Nodup) : #l.toFinset = l.length := Multiset.toFinset_card_of_nodup h end ToMLListultiset @@ -222,16 +228,16 @@ namespace Finset variable {s t u : Finset α} {f : α → β} {n : ℕ} @[simp] -theorem length_toList (s : Finset α) : s.toList.length = s.card := by +theorem length_toList (s : Finset α) : s.toList.length = #s := by rw [toList, ← Multiset.coe_card, Multiset.coe_toList, card_def] -theorem card_image_le [DecidableEq β] : (s.image f).card ≤ s.card := by +theorem card_image_le [DecidableEq β] : #(s.image f) ≤ #s := by simpa only [card_map] using (s.1.map f).toFinset_card_le -theorem card_image_of_injOn [DecidableEq β] (H : Set.InjOn f s) : (s.image f).card = s.card := by +theorem card_image_of_injOn [DecidableEq β] (H : Set.InjOn f s) : #(s.image f) = #s := by simp only [card, image_val_of_injOn H, card_map] -theorem injOn_of_card_image_eq [DecidableEq β] (H : (s.image f).card = s.card) : Set.InjOn f s := by +theorem injOn_of_card_image_eq [DecidableEq β] (H : #(s.image f) = #s) : Set.InjOn f s := by rw [card_def, card_def, image, toFinset] at H dsimp only at H have : (s.1.map f).dedup = s.1.map f := by @@ -240,58 +246,58 @@ theorem injOn_of_card_image_eq [DecidableEq β] (H : (s.image f).card = s.card) rw [Multiset.dedup_eq_self] at this exact inj_on_of_nodup_map this -theorem card_image_iff [DecidableEq β] : (s.image f).card = s.card ↔ Set.InjOn f s := +theorem card_image_iff [DecidableEq β] : #(s.image f) = #s ↔ Set.InjOn f s := ⟨injOn_of_card_image_eq, card_image_of_injOn⟩ theorem card_image_of_injective [DecidableEq β] (s : Finset α) (H : Injective f) : - (s.image f).card = s.card := + #(s.image f) = #s := card_image_of_injOn fun _ _ _ _ h => H h theorem fiber_card_ne_zero_iff_mem_image (s : Finset α) (f : α → β) [DecidableEq β] (y : β) : - (s.filter fun x => f x = y).card ≠ 0 ↔ y ∈ s.image f := by + #(s.filter fun x ↦ f x = y) ≠ 0 ↔ y ∈ s.image f := by rw [← Nat.pos_iff_ne_zero, card_pos, fiber_nonempty_iff_mem_image] lemma card_filter_le_iff (s : Finset α) (P : α → Prop) [DecidablePred P] (n : ℕ) : - (s.filter P).card ≤ n ↔ ∀ s' ⊆ s, n < s'.card → ∃ a ∈ s', ¬ P a := + #(s.filter P) ≤ n ↔ ∀ s' ⊆ s, n < #s' → ∃ a ∈ s', ¬ P a := (s.1.card_filter_le_iff P n).trans ⟨fun H s' hs' h ↦ H s'.1 (by aesop) h, fun H s' hs' h ↦ H ⟨s', nodup_of_le hs' s.2⟩ (fun _ hx ↦ subset_of_le hs' hx) h⟩ @[simp] -theorem card_map (f : α ↪ β) : (s.map f).card = s.card := +theorem card_map (f : α ↪ β) : #(s.map f) = #s := Multiset.card_map _ _ @[simp] theorem card_subtype (p : α → Prop) [DecidablePred p] (s : Finset α) : - (s.subtype p).card = (s.filter p).card := by simp [Finset.subtype] + #(s.subtype p) = #(s.filter p) := by simp [Finset.subtype] theorem card_filter_le (s : Finset α) (p : α → Prop) [DecidablePred p] : - (s.filter p).card ≤ s.card := + #(s.filter p) ≤ #s := card_le_card <| filter_subset _ _ -theorem eq_of_subset_of_card_le {s t : Finset α} (h : s ⊆ t) (h₂ : t.card ≤ s.card) : s = t := +theorem eq_of_subset_of_card_le {s t : Finset α} (h : s ⊆ t) (h₂ : #t ≤ #s) : s = t := eq_of_veq <| Multiset.eq_of_le_of_card_le (val_le_iff.mpr h) h₂ -theorem eq_of_superset_of_card_ge (hst : s ⊆ t) (hts : t.card ≤ s.card) : t = s := +theorem eq_of_superset_of_card_ge (hst : s ⊆ t) (hts : #t ≤ #s) : t = s := (eq_of_subset_of_card_le hst hts).symm -theorem subset_iff_eq_of_card_le (h : t.card ≤ s.card) : s ⊆ t ↔ s = t := +theorem subset_iff_eq_of_card_le (h : #t ≤ #s) : s ⊆ t ↔ s = t := ⟨fun hst => eq_of_subset_of_card_le hst h, Eq.subset'⟩ theorem map_eq_of_subset {f : α ↪ α} (hs : s.map f ⊆ s) : s.map f = s := eq_of_subset_of_card_le hs (card_map _).ge -theorem filter_card_eq {p : α → Prop} [DecidablePred p] (h : (s.filter p).card = s.card) (x : α) +theorem filter_card_eq {p : α → Prop} [DecidablePred p] (h : #(s.filter p) = #s) (x : α) (hx : x ∈ s) : p x := by rw [← eq_of_subset_of_card_le (s.filter_subset p) h.ge, mem_filter] at hx exact hx.2 -nonrec lemma card_lt_card (h : s ⊂ t) : s.card < t.card := card_lt_card <| val_lt_iff.2 h +nonrec lemma card_lt_card (h : s ⊂ t) : #s < #t := card_lt_card <| val_lt_iff.2 h lemma card_strictMono : StrictMono (card : Finset α → ℕ) := fun _ _ ↦ card_lt_card theorem card_eq_of_bijective (f : ∀ i, i < n → α) (hf : ∀ a ∈ s, ∃ i, ∃ h : i < n, f i h = a) (hf' : ∀ i (h : i < n), f i h ∈ s) - (f_inj : ∀ i j (hi : i < n) (hj : j < n), f i hi = f j hj → i = j) : s.card = n := by + (f_inj : ∀ i j (hi : i < n) (hj : j < n), f i hi = f j hj → i = j) : #s = n := by classical have : s = (range n).attach.image fun i => f i.1 (mem_range.1 i.2) := by ext a @@ -301,10 +307,10 @@ theorem card_eq_of_bijective (f : ∀ i, i < n → α) (hf : ∀ a ∈ s, ∃ i, · intro ha; obtain ⟨i, hi, rfl⟩ := hf a ha; use i, mem_range.2 hi · rintro ⟨i, hi, rfl⟩; apply hf' calc - s.card = ((range n).attach.image fun i => f i.1 (mem_range.1 i.2)).card := by rw [this] - _ = (range n).attach.card := ?_ - _ = (range n).card := card_attach - _ = n := card_range n + #s = #((range n).attach.image fun i => f i.1 (mem_range.1 i.2)) := by rw [this] + _ = #(range n).attach := ?_ + _ = #(range n) := card_attach + _ = n := card_range n apply card_image_of_injective intro ⟨i, hi⟩ ⟨j, hj⟩ eq exact Subtype.eq <| f_inj i j (mem_range.1 hi) (mem_range.1 hj) eq @@ -321,12 +327,12 @@ The difference with `Finset.card_nbij` is that the bijection is allowed to use m domain, rather than being a non-dependent function. -/ lemma card_bij (i : ∀ a ∈ s, β) (hi : ∀ a ha, i a ha ∈ t) (i_inj : ∀ a₁ ha₁ a₂ ha₂, i a₁ ha₁ = i a₂ ha₂ → a₁ = a₂) - (i_surj : ∀ b ∈ t, ∃ a ha, i a ha = b) : s.card = t.card := by + (i_surj : ∀ b ∈ t, ∃ a ha, i a ha = b) : #s = #t := by classical calc - s.card = s.attach.card := card_attach.symm - _ = (s.attach.image fun a : { a // a ∈ s } => i a.1 a.2).card := Eq.symm ?_ - _ = t.card := ?_ + #s = #s.attach := card_attach.symm + _ = #(s.attach.image fun a ↦ i a.1 a.2) := Eq.symm ?_ + _ = #t := ?_ · apply card_image_of_injective intro ⟨_, _⟩ ⟨_, _⟩ h simpa using i_inj _ _ _ _ h @@ -347,7 +353,7 @@ The difference with `Finset.card_nbij'` is that the bijection and its inverse ar membership of the domains, rather than being non-dependent functions. -/ lemma card_bij' (i : ∀ a ∈ s, β) (j : ∀ a ∈ t, α) (hi : ∀ a ha, i a ha ∈ t) (hj : ∀ a ha, j a ha ∈ s) (left_inv : ∀ a ha, j (i a ha) (hi a ha) = a) - (right_inv : ∀ a ha, i (j a ha) (hj a ha) = a) : s.card = t.card := by + (right_inv : ∀ a ha, i (j a ha) (hj a ha) = a) : #s = #t := by refine card_bij i hi (fun a1 h1 a2 h2 eq ↦ ?_) (fun b hb ↦ ⟨_, hj b hb, right_inv b hb⟩) rw [← left_inv a1 h1, ← left_inv a2 h2] simp only [eq] @@ -360,7 +366,7 @@ injection, rather than by an inverse function. The difference with `Finset.card_bij` is that the bijection is a non-dependent function, rather than being allowed to use membership of the domain. -/ lemma card_nbij (i : α → β) (hi : ∀ a ∈ s, i a ∈ t) (i_inj : (s : Set α).InjOn i) - (i_surj : (s : Set α).SurjOn i t) : s.card = t.card := + (i_surj : (s : Set α).SurjOn i t) : #s = #t := card_bij (fun a _ ↦ i a) hi i_inj (by simpa using i_surj) /-- Reorder a finset. @@ -374,35 +380,35 @@ functions, rather than being allowed to use membership of the domains. The difference with `Finset.card_equiv` is that bijectivity is only required to hold on the domains, rather than on the entire types. -/ lemma card_nbij' (i : α → β) (j : β → α) (hi : ∀ a ∈ s, i a ∈ t) (hj : ∀ a ∈ t, j a ∈ s) - (left_inv : ∀ a ∈ s, j (i a) = a) (right_inv : ∀ a ∈ t, i (j a) = a) : s.card = t.card := + (left_inv : ∀ a ∈ s, j (i a) = a) (right_inv : ∀ a ∈ t, i (j a) = a) : #s = #t := card_bij' (fun a _ ↦ i a) (fun b _ ↦ j b) hi hj left_inv right_inv /-- Specialization of `Finset.card_nbij'` that automatically fills in most arguments. See `Fintype.card_equiv` for the version where `s` and `t` are `univ`. -/ -lemma card_equiv (e : α ≃ β) (hst : ∀ i, i ∈ s ↔ e i ∈ t) : s.card = t.card := by +lemma card_equiv (e : α ≃ β) (hst : ∀ i, i ∈ s ↔ e i ∈ t) : #s = #t := by refine card_nbij' e e.symm ?_ ?_ ?_ ?_ <;> simp [hst] /-- Specialization of `Finset.card_nbij` that automatically fills in most arguments. See `Fintype.card_bijective` for the version where `s` and `t` are `univ`. -/ lemma card_bijective (e : α → β) (he : e.Bijective) (hst : ∀ i, i ∈ s ↔ e i ∈ t) : - s.card = t.card := card_equiv (.ofBijective e he) hst + #s = #t := card_equiv (.ofBijective e he) hst lemma card_le_card_of_injOn (f : α → β) (hf : ∀ a ∈ s, f a ∈ t) (f_inj : (s : Set α).InjOn f) : - s.card ≤ t.card := by + #s ≤ #t := by classical calc - s.card = (s.image f).card := (card_image_of_injOn f_inj).symm - _ ≤ t.card := card_le_card <| image_subset_iff.2 hf + #s = #(s.image f) := (card_image_of_injOn f_inj).symm + _ ≤ #t := card_le_card <| image_subset_iff.2 hf @[deprecated (since := "2024-06-01")] alias card_le_card_of_inj_on := card_le_card_of_injOn -lemma card_le_card_of_surjOn (f : α → β) (hf : Set.SurjOn f s t) : t.card ≤ s.card := by +lemma card_le_card_of_surjOn (f : α → β) (hf : Set.SurjOn f s t) : #t ≤ #s := by classical unfold Set.SurjOn at hf; exact (card_le_card (mod_cast hf)).trans card_image_le /-- If there are more pigeons than pigeonholes, then there are two pigeons in the same pigeonhole. -/ -theorem exists_ne_map_eq_of_card_lt_of_maps_to {t : Finset β} (hc : t.card < s.card) {f : α → β} +theorem exists_ne_map_eq_of_card_lt_of_maps_to {t : Finset β} (hc : #t < #s) {f : α → β} (hf : ∀ a ∈ s, f a ∈ t) : ∃ x ∈ s, ∃ y ∈ s, x ≠ y ∧ f x = f y := by classical by_contra! hz @@ -412,16 +418,16 @@ theorem exists_ne_map_eq_of_card_lt_of_maps_to {t : Finset β} (hc : t.card < s. exact hz x hx y hy lemma le_card_of_inj_on_range (f : ℕ → α) (hf : ∀ i < n, f i ∈ s) - (f_inj : ∀ i < n, ∀ j < n, f i = f j → i = j) : n ≤ s.card := + (f_inj : ∀ i < n, ∀ j < n, f i = f j → i = j) : n ≤ #s := calc - n = card (range n) := (card_range n).symm - _ ≤ s.card := card_le_card_of_injOn f (by simpa only [mem_range]) (by simpa) + n = #(range n) := (card_range n).symm + _ ≤ #s := card_le_card_of_injOn f (by simpa only [mem_range]) (by simpa) lemma surj_on_of_inj_on_of_card_le (f : ∀ a ∈ s, β) (hf : ∀ a ha, f a ha ∈ t) - (hinj : ∀ a₁ a₂ ha₁ ha₂, f a₁ ha₁ = f a₂ ha₂ → a₁ = a₂) (hst : t.card ≤ s.card) : + (hinj : ∀ a₁ a₂ ha₁ ha₂, f a₁ ha₁ = f a₂ ha₂ → a₁ = a₂) (hst : #t ≤ #s) : ∀ b ∈ t, ∃ a ha, b = f a ha := by classical - have h : (s.attach.image fun a : { a // a ∈ s } => f a a.prop).card = s.card := by + have h : #(s.attach.image fun a : s ↦ f a a.2) = #s := by rw [← @card_attach _ s, card_image_of_injective] intro ⟨_, _⟩ ⟨_, _⟩ h exact Subtype.eq <| hinj _ _ _ _ h @@ -431,7 +437,7 @@ lemma surj_on_of_inj_on_of_card_le (f : ∀ a ∈ s, β) (hf : ∀ a ha, f a ha exact fun b a ha hb ↦ ⟨a, ha, hb.symm⟩ theorem inj_on_of_surj_on_of_card_le (f : ∀ a ∈ s, β) (hf : ∀ a ha, f a ha ∈ t) - (hsurj : ∀ b ∈ t, ∃ a ha, f a ha = b) (hst : s.card ≤ t.card) ⦃a₁⦄ (ha₁ : a₁ ∈ s) ⦃a₂⦄ + (hsurj : ∀ b ∈ t, ∃ a ha, f a ha = b) (hst : #s ≤ #t) ⦃a₁⦄ (ha₁ : a₁ ∈ s) ⦃a₂⦄ (ha₂ : a₂ ∈ s) (ha₁a₂ : f a₁ ha₁ = f a₂ ha₂) : a₁ = a₂ := haveI : Inhabited { x // x ∈ s } := ⟨⟨a₁, ha₁⟩⟩ let f' : { x // x ∈ s } → { x // x ∈ t } := fun x => ⟨f x.1 x.2, hf x.1 x.2⟩ @@ -453,7 +459,7 @@ theorem inj_on_of_surj_on_of_card_le (f : ∀ a ∈ s, β) (hf : ∀ a ha, f a h end bij @[simp] -theorem card_disjUnion (s t : Finset α) (h) : (s.disjUnion t h).card = s.card + t.card := +theorem card_disjUnion (s t : Finset α) (h) : #(s.disjUnion t h) = #s + #t := Multiset.card_add _ _ /-! ### Lattice structure -/ @@ -464,23 +470,23 @@ section Lattice variable [DecidableEq α] theorem card_union_add_card_inter (s t : Finset α) : - (s ∪ t).card + (s ∩ t).card = s.card + t.card := + #(s ∪ t) + #(s ∩ t) = #s + #t := Finset.induction_on t (by simp) fun a r har h => by by_cases a ∈ s <;> simp [*, ← add_assoc, add_right_comm _ 1] theorem card_inter_add_card_union (s t : Finset α) : - (s ∩ t).card + (s ∪ t).card = s.card + t.card := by rw [add_comm, card_union_add_card_inter] + #(s ∩ t) + #(s ∪ t) = #s + #t := by rw [add_comm, card_union_add_card_inter] -lemma card_union (s t : Finset α) : (s ∪ t).card = s.card + t.card - (s ∩ t).card := by +lemma card_union (s t : Finset α) : #(s ∪ t) = #s + #t - #(s ∩ t) := by rw [← card_union_add_card_inter, Nat.add_sub_cancel] -lemma card_inter (s t : Finset α) : (s ∩ t).card = s.card + t.card - (s ∪ t).card := by +lemma card_inter (s t : Finset α) : #(s ∩ t) = #s + #t - #(s ∪ t) := by rw [← card_inter_add_card_union, Nat.add_sub_cancel] -theorem card_union_le (s t : Finset α) : (s ∪ t).card ≤ s.card + t.card := +theorem card_union_le (s t : Finset α) : #(s ∪ t) ≤ #s + #t := card_union_add_card_inter s t ▸ Nat.le_add_right _ _ -lemma card_union_eq_card_add_card : (s ∪ t).card = s.card + t.card ↔ Disjoint s t := by +lemma card_union_eq_card_add_card : #(s ∪ t) = #s + #t ↔ Disjoint s t := by rw [← card_union_add_card_inter]; simp [disjoint_iff_inter_eq_empty] @[simp] alias ⟨_, card_union_of_disjoint⟩ := card_union_eq_card_add_card @@ -489,70 +495,70 @@ lemma card_union_eq_card_add_card : (s ∪ t).card = s.card + t.card ↔ Disjoin @[deprecated (since := "2024-02-09")] alias card_disjoint_union := card_union_of_disjoint lemma cast_card_inter [AddGroupWithOne R] : - ((s ∩ t).card : R) = s.card + t.card - (s ∪ t).card := by + (#(s ∩ t) : R) = #s + #t - #(s ∪ t) := by rw [eq_sub_iff_add_eq, ← cast_add, card_inter_add_card_union, cast_add] lemma cast_card_union [AddGroupWithOne R] : - ((s ∪ t).card : R) = s.card + t.card - (s ∩ t).card := by + (#(s ∪ t) : R) = #s + #t - #(s ∩ t) := by rw [eq_sub_iff_add_eq, ← cast_add, card_union_add_card_inter, cast_add] -theorem card_sdiff (h : s ⊆ t) : card (t \ s) = t.card - s.card := by - suffices card (t \ s) = card (t \ s ∪ s) - s.card by rwa [sdiff_union_of_subset h] at this +theorem card_sdiff (h : s ⊆ t) : #(t \ s) = #t - #s := by + suffices #(t \ s) = #(t \ s ∪ s) - #s by rwa [sdiff_union_of_subset h] at this rw [card_union_of_disjoint sdiff_disjoint, Nat.add_sub_cancel_right] -lemma cast_card_sdiff [AddGroupWithOne R] (h : s ⊆ t) : ((t \ s).card : R) = t.card - s.card := by +lemma cast_card_sdiff [AddGroupWithOne R] (h : s ⊆ t) : (#(t \ s) : R) = #t - #s := by rw [card_sdiff h, Nat.cast_sub (card_mono h)] -theorem card_sdiff_add_card_eq_card {s t : Finset α} (h : s ⊆ t) : card (t \ s) + card s = card t := +theorem card_sdiff_add_card_eq_card {s t : Finset α} (h : s ⊆ t) : #(t \ s) + #s = #t := ((Nat.sub_eq_iff_eq_add (card_le_card h)).mp (card_sdiff h).symm).symm -theorem le_card_sdiff (s t : Finset α) : t.card - s.card ≤ card (t \ s) := +theorem le_card_sdiff (s t : Finset α) : #t - #s ≤ #(t \ s) := calc - card t - card s ≤ card t - card (s ∩ t) := + #t - #s ≤ #t - #(s ∩ t) := Nat.sub_le_sub_left (card_le_card inter_subset_left) _ - _ = card (t \ (s ∩ t)) := (card_sdiff inter_subset_right).symm - _ ≤ card (t \ s) := by rw [sdiff_inter_self_right t s] + _ = #(t \ (s ∩ t)) := (card_sdiff inter_subset_right).symm + _ ≤ #(t \ s) := by rw [sdiff_inter_self_right t s] -theorem card_le_card_sdiff_add_card : s.card ≤ (s \ t).card + t.card := +theorem card_le_card_sdiff_add_card : #s ≤ #(s \ t) + #t := Nat.sub_le_iff_le_add.1 <| le_card_sdiff _ _ -theorem card_sdiff_add_card (s t : Finset α) : (s \ t).card + t.card = (s ∪ t).card := by +theorem card_sdiff_add_card (s t : Finset α) : #(s \ t) + #t = #(s ∪ t) := by rw [← card_union_of_disjoint sdiff_disjoint, sdiff_union_self_eq_union] -lemma card_sdiff_comm (h : s.card = t.card) : (s \ t).card = (t \ s).card := - add_left_injective t.card <| by +lemma card_sdiff_comm (h : #s = #t) : #(s \ t) = #(t \ s) := + add_left_injective #t <| by simp_rw [card_sdiff_add_card, ← h, card_sdiff_add_card, union_comm] @[simp] lemma card_sdiff_add_card_inter (s t : Finset α) : - (s \ t).card + (s ∩ t).card = s.card := by + #(s \ t) + #(s ∩ t) = #s := by rw [← card_union_of_disjoint (disjoint_sdiff_inter _ _), sdiff_union_inter] @[simp] lemma card_inter_add_card_sdiff (s t : Finset α) : - (s ∩ t).card + (s \ t).card = s.card := by + #(s ∩ t) + #(s \ t) = #s := by rw [add_comm, card_sdiff_add_card_inter] /-- **Pigeonhole principle** for two finsets inside an ambient finset. -/ theorem inter_nonempty_of_card_lt_card_add_card (hts : t ⊆ s) (hus : u ⊆ s) - (hstu : s.card < t.card + u.card) : (t ∩ u).Nonempty := by + (hstu : #s < #t + #u) : (t ∩ u).Nonempty := by contrapose! hstu calc - _ = (t ∪ u).card := by simp [← card_union_add_card_inter, not_nonempty_iff_eq_empty.1 hstu] - _ ≤ s.card := by gcongr; exact union_subset hts hus + _ = #(t ∪ u) := by simp [← card_union_add_card_inter, not_nonempty_iff_eq_empty.1 hstu] + _ ≤ #s := by gcongr; exact union_subset hts hus end Lattice theorem filter_card_add_filter_neg_card_eq_card (p : α → Prop) [DecidablePred p] [∀ x, Decidable (¬p x)] : - (s.filter p).card + (s.filter (fun a => ¬ p a)).card = s.card := by + #(s.filter p) + #(s.filter fun a ↦ ¬ p a) = #s := by classical rw [← card_union_of_disjoint (disjoint_filter_filter_neg _ _ _), filter_union_filter_neg_eq] /-- Given a subset `s` of a set `t`, of sizes at most and at least `n` respectively, there exists a set `u` of size `n` which is both a superset of `s` and a subset of `t`. -/ -lemma exists_subsuperset_card_eq (hst : s ⊆ t) (hsn : s.card ≤ n) (hnt : n ≤ t.card) : - ∃ u, s ⊆ u ∧ u ⊆ t ∧ card u = n := by +lemma exists_subsuperset_card_eq (hst : s ⊆ t) (hsn : #s ≤ n) (hnt : n ≤ #t) : + ∃ u, s ⊆ u ∧ u ⊆ t ∧ #u = n := by classical refine Nat.decreasingInduction' ?_ hnt ⟨t, by simp [hst]⟩ intro k _ hnk ⟨u, hu₁, hu₂, hu₃⟩ @@ -561,48 +567,48 @@ lemma exists_subsuperset_card_eq (hst : s ⊆ t) (hsn : s.card ≤ n) (hnt : n exact ⟨u.erase a, by simp [subset_erase, erase_subset_iff_of_mem (hu₂ _), *]⟩ /-- We can shrink a set to any smaller size. -/ -lemma exists_subset_card_eq (hns : n ≤ s.card) : ∃ t ⊆ s, t.card = n := by +lemma exists_subset_card_eq (hns : n ≤ #s) : ∃ t ⊆ s, #t = n := by simpa using exists_subsuperset_card_eq s.empty_subset (by simp) hns /-- Given a set `A` and a set `B` inside it, we can shrink `A` to any appropriate size, and keep `B` inside it. -/ @[deprecated exists_subsuperset_card_eq (since := "2024-06-23")] -theorem exists_intermediate_set {A B : Finset α} (i : ℕ) (h₁ : i + card B ≤ card A) (h₂ : B ⊆ A) : - ∃ C : Finset α, B ⊆ C ∧ C ⊆ A ∧ card C = i + card B := +theorem exists_intermediate_set {A B : Finset α} (i : ℕ) (h₁ : i + #B ≤ #A) (h₂ : B ⊆ A) : + ∃ C : Finset α, B ⊆ C ∧ C ⊆ A ∧ #C = i + #B := exists_subsuperset_card_eq h₂ (Nat.le_add_left ..) h₁ /-- We can shrink `A` to any smaller size. -/ @[deprecated exists_subset_card_eq (since := "2024-06-23")] -theorem exists_smaller_set (A : Finset α) (i : ℕ) (h₁ : i ≤ card A) : - ∃ B : Finset α, B ⊆ A ∧ card B = i := exists_subset_card_eq h₁ +theorem exists_smaller_set (A : Finset α) (i : ℕ) (h₁ : i ≤ #A) : + ∃ B : Finset α, B ⊆ A ∧ #B = i := exists_subset_card_eq h₁ -theorem le_card_iff_exists_subset_card : n ≤ s.card ↔ ∃ t ⊆ s, t.card = n := by +theorem le_card_iff_exists_subset_card : n ≤ #s ↔ ∃ t ⊆ s, #t = n := by refine ⟨fun h => ?_, fun ⟨t, hst, ht⟩ => ht ▸ card_le_card hst⟩ exact exists_subset_card_eq h theorem exists_subset_or_subset_of_two_mul_lt_card [DecidableEq α] {X Y : Finset α} {n : ℕ} - (hXY : 2 * n < (X ∪ Y).card) : ∃ C : Finset α, n < C.card ∧ (C ⊆ X ∨ C ⊆ Y) := by - have h₁ : (X ∩ (Y \ X)).card = 0 := Finset.card_eq_zero.mpr (Finset.inter_sdiff_self X Y) - have h₂ : (X ∪ Y).card = X.card + (Y \ X).card := by + (hXY : 2 * n < #(X ∪ Y)) : ∃ C : Finset α, n < #C ∧ (C ⊆ X ∨ C ⊆ Y) := by + have h₁ : #(X ∩ (Y \ X)) = 0 := Finset.card_eq_zero.mpr (Finset.inter_sdiff_self X Y) + have h₂ : #(X ∪ Y) = #X + #(Y \ X) := by rw [← card_union_add_card_inter X (Y \ X), Finset.union_sdiff_self_eq_union, h₁, add_zero] rw [h₂, Nat.two_mul] at hXY - obtain h | h : n < X.card ∨ n < (Y \ X).card := by contrapose! hXY; omega + obtain h | h : n < #X ∨ n < #(Y \ X) := by contrapose! hXY; omega · exact ⟨X, h, Or.inl (Finset.Subset.refl X)⟩ · exact ⟨Y \ X, h, Or.inr sdiff_subset⟩ /-! ### Explicit description of a finset from its card -/ -theorem card_eq_one : s.card = 1 ↔ ∃ a, s = {a} := by +theorem card_eq_one : #s = 1 ↔ ∃ a, s = {a} := by cases s simp only [Multiset.card_eq_one, Finset.card, ← val_inj, singleton_val] theorem _root_.Multiset.toFinset_card_eq_one_iff [DecidableEq α] (s : Multiset α) : - s.toFinset.card = 1 ↔ Multiset.card s ≠ 0 ∧ ∃ a : α, s = Multiset.card s • {a} := by + #s.toFinset = 1 ↔ Multiset.card s ≠ 0 ∧ ∃ a : α, s = Multiset.card s • {a} := by simp_rw [card_eq_one, Multiset.toFinset_eq_singleton_iff, exists_and_left] theorem exists_eq_insert_iff [DecidableEq α] {s t : Finset α} : - (∃ a ∉ s, insert a s = t) ↔ s ⊆ t ∧ s.card + 1 = t.card := by + (∃ a ∉ s, insert a s = t) ↔ s ⊆ t ∧ #s + 1 = #t := by constructor · rintro ⟨a, ha, rfl⟩ exact ⟨subset_insert _ _, (card_insert_of_not_mem ha).symm⟩ @@ -615,7 +621,7 @@ theorem exists_eq_insert_iff [DecidableEq α] {s t : Finset α} : rw [← ha] exact not_mem_sdiff_of_mem_right hs -theorem card_le_one : s.card ≤ 1 ↔ ∀ a ∈ s, ∀ b ∈ s, a = b := by +theorem card_le_one : #s ≤ 1 ↔ ∀ a ∈ s, ∀ b ∈ s, a = b := by obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty · simp refine (Nat.succ_le_of_lt (card_pos.2 ⟨x, hx⟩)).le_iff_eq.trans (card_eq_one.trans ⟨?_, ?_⟩) @@ -623,14 +629,14 @@ theorem card_le_one : s.card ≤ 1 ↔ ∀ a ∈ s, ∀ b ∈ s, a = b := by simp · exact fun h => ⟨x, eq_singleton_iff_unique_mem.2 ⟨hx, fun y hy => h _ hy _ hx⟩⟩ -theorem card_le_one_iff : s.card ≤ 1 ↔ ∀ {a b}, a ∈ s → b ∈ s → a = b := by +theorem card_le_one_iff : #s ≤ 1 ↔ ∀ {a b}, a ∈ s → b ∈ s → a = b := by rw [card_le_one] tauto -theorem card_le_one_iff_subsingleton_coe : s.card ≤ 1 ↔ Subsingleton (s : Type _) := +theorem card_le_one_iff_subsingleton_coe : #s ≤ 1 ↔ Subsingleton (s : Type _) := card_le_one.trans (s : Set α).subsingleton_coe.symm -theorem card_le_one_iff_subset_singleton [Nonempty α] : s.card ≤ 1 ↔ ∃ x : α, s ⊆ {x} := by +theorem card_le_one_iff_subset_singleton [Nonempty α] : #s ≤ 1 ↔ ∃ x : α, s ⊆ {x} := by refine ⟨fun H => ?_, ?_⟩ · obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty · exact ⟨Classical.arbitrary α, empty_subset _⟩ @@ -639,32 +645,32 @@ theorem card_le_one_iff_subset_singleton [Nonempty α] : s.card ≤ 1 ↔ ∃ x rw [← card_singleton x] exact card_le_card hx -lemma exists_mem_ne (hs : 1 < s.card) (a : α) : ∃ b ∈ s, b ≠ a := by +lemma exists_mem_ne (hs : 1 < #s) (a : α) : ∃ b ∈ s, b ≠ a := by have : Nonempty α := ⟨a⟩ by_contra! exact hs.not_le (card_le_one_iff_subset_singleton.2 ⟨a, subset_singleton_iff'.2 this⟩) /-- A `Finset` of a subsingleton type has cardinality at most one. -/ -theorem card_le_one_of_subsingleton [Subsingleton α] (s : Finset α) : s.card ≤ 1 := +theorem card_le_one_of_subsingleton [Subsingleton α] (s : Finset α) : #s ≤ 1 := Finset.card_le_one_iff.2 fun {_ _ _ _} => Subsingleton.elim _ _ -theorem one_lt_card : 1 < s.card ↔ ∃ a ∈ s, ∃ b ∈ s, a ≠ b := by +theorem one_lt_card : 1 < #s ↔ ∃ a ∈ s, ∃ b ∈ s, a ≠ b := by rw [← not_iff_not] push_neg exact card_le_one -theorem one_lt_card_iff : 1 < s.card ↔ ∃ a b, a ∈ s ∧ b ∈ s ∧ a ≠ b := by +theorem one_lt_card_iff : 1 < #s ↔ ∃ a b, a ∈ s ∧ b ∈ s ∧ a ≠ b := by rw [one_lt_card] simp only [exists_prop, exists_and_left] -theorem one_lt_card_iff_nontrivial : 1 < s.card ↔ s.Nontrivial := by +theorem one_lt_card_iff_nontrivial : 1 < #s ↔ s.Nontrivial := by rw [← not_iff_not, not_lt, Finset.Nontrivial, ← Set.nontrivial_coe_sort, not_nontrivial_iff_subsingleton, card_le_one_iff_subsingleton_coe, coe_sort_coe] @[deprecated (since := "2024-02-05")] alias one_lt_card_iff_nontrivial_coe := one_lt_card_iff_nontrivial -theorem exists_ne_of_one_lt_card (hs : 1 < s.card) (a : α) : ∃ b, b ∈ s ∧ b ≠ a := by +theorem exists_ne_of_one_lt_card (hs : 1 < #s) (a : α) : ∃ b, b ∈ s ∧ b ≠ a := by obtain ⟨x, hx, y, hy, hxy⟩ := Finset.one_lt_card.mp hs by_cases ha : y = a · exact ⟨x, hx, ne_of_ne_of_eq hxy ha⟩ @@ -674,8 +680,8 @@ theorem exists_ne_of_one_lt_card (hs : 1 < s.card) (a : α) : ∃ b, b ∈ s ∧ its projection to some factor is nontrivial, and the fibers of the projection are proper subsets. -/ lemma exists_of_one_lt_card_pi {ι : Type*} {α : ι → Type*} [∀ i, DecidableEq (α i)] - {s : Finset (∀ i, α i)} (h : 1 < s.card) : - ∃ i, 1 < (s.image (· i)).card ∧ ∀ ai, s.filter (· i = ai) ⊂ s := by + {s : Finset (∀ i, α i)} (h : 1 < #s) : + ∃ i, 1 < #(s.image (· i)) ∧ ∀ ai, s.filter (· i = ai) ⊂ s := by simp_rw [one_lt_card_iff, Function.ne_iff] at h ⊢ obtain ⟨a1, a2, h1, h2, i, hne⟩ := h refine ⟨i, ⟨_, _, mem_image_of_mem _ h1, mem_image_of_mem _ h2, hne⟩, fun ai => ?_⟩ @@ -684,21 +690,21 @@ lemma exists_of_one_lt_card_pi {ι : Type*} {α : ι → Type*} [∀ i, Decidabl exacts [⟨a1, h1, hne⟩, ⟨a2, h2, hne⟩] theorem card_eq_succ_iff_cons : - s.card = n + 1 ↔ ∃ a t, ∃ (h : a ∉ t), cons a t h = s ∧ t.card = n := + #s = n + 1 ↔ ∃ a t, ∃ (h : a ∉ t), cons a t h = s ∧ #t = n := ⟨cons_induction_on s (by simp) fun a s _ _ _ => ⟨a, s, by simp_all⟩, fun ⟨a, t, _, hs, _⟩ => by simpa [← hs]⟩ section DecidableEq variable [DecidableEq α] -theorem card_eq_succ : s.card = n + 1 ↔ ∃ a t, a ∉ t ∧ insert a t = s ∧ t.card = n := +theorem card_eq_succ : #s = n + 1 ↔ ∃ a t, a ∉ t ∧ insert a t = s ∧ #t = n := ⟨fun h => - let ⟨a, has⟩ := card_pos.mp (h.symm ▸ Nat.zero_lt_succ _ : 0 < s.card) + let ⟨a, has⟩ := card_pos.mp (h.symm ▸ Nat.zero_lt_succ _ : 0 < #s) ⟨a, s.erase a, s.not_mem_erase a, insert_erase has, by simp only [h, card_erase_of_mem has, Nat.add_sub_cancel_right]⟩, fun ⟨_, _, hat, s_eq, n_eq⟩ => s_eq ▸ n_eq ▸ card_insert_of_not_mem hat⟩ -theorem card_eq_two : s.card = 2 ↔ ∃ x y, x ≠ y ∧ s = {x, y} := by +theorem card_eq_two : #s = 2 ↔ ∃ x y, x ≠ y ∧ s = {x, y} := by constructor · rw [card_eq_succ] simp_rw [card_eq_one] @@ -707,7 +713,7 @@ theorem card_eq_two : s.card = 2 ↔ ∃ x y, x ≠ y ∧ s = {x, y} := by · rintro ⟨x, y, h, rfl⟩ exact card_pair h -theorem card_eq_three : s.card = 3 ↔ ∃ x y z, x ≠ y ∧ x ≠ z ∧ y ≠ z ∧ s = {x, y, z} := by +theorem card_eq_three : #s = 3 ↔ ∃ x y z, x ≠ y ∧ x ≠ z ∧ y ≠ z ∧ s = {x, y, z} := by constructor · rw [card_eq_succ] simp_rw [card_eq_two] @@ -720,7 +726,7 @@ theorem card_eq_three : s.card = 3 ↔ ∃ x y z, x ≠ y ∧ x ≠ z ∧ y ≠ end DecidableEq -theorem two_lt_card_iff : 2 < s.card ↔ ∃ a b c, a ∈ s ∧ b ∈ s ∧ c ∈ s ∧ a ≠ b ∧ a ≠ c ∧ b ≠ c := by +theorem two_lt_card_iff : 2 < #s ↔ ∃ a b c, a ∈ s ∧ b ∈ s ∧ c ∈ s ∧ a ≠ b ∧ a ≠ c ∧ b ≠ c := by classical simp_rw [lt_iff_add_one_le, le_card_iff_exists_subset_card, reduceAdd, card_eq_three, ← exists_and_left, exists_comm (α := Finset α)] @@ -730,7 +736,7 @@ theorem two_lt_card_iff : 2 < s.card ↔ ∃ a b c, a ∈ s ∧ b ∈ s ∧ c · rintro ⟨a, b, c, ha, hb, hc, hab, hac, hbc⟩ exact ⟨a, b, c, {a, b, c}, by simp_all [insert_subset_iff]⟩ -theorem two_lt_card : 2 < s.card ↔ ∃ a ∈ s, ∃ b ∈ s, ∃ c ∈ s, a ≠ b ∧ a ≠ c ∧ b ≠ c := by +theorem two_lt_card : 2 < #s ↔ ∃ a ∈ s, ∃ b ∈ s, ∃ c ∈ s, a ≠ b ∧ a ≠ c ∧ b ≠ c := by simp_rw [two_lt_card_iff, exists_and_left] /-! ### Inductions -/ @@ -743,9 +749,9 @@ def strongInduction {p : Finset α → Sort*} (H : ∀ s, (∀ t ⊂ s, p t) → ∀ s : Finset α, p s | s => H s fun t h => - have : t.card < s.card := card_lt_card h + have : #t < #s := card_lt_card h strongInduction H t - termination_by s => Finset.card s + termination_by s => #s @[nolint unusedHavesSuffices] -- Porting note: false positive theorem strongInduction_eq {p : Finset α → Sort*} (H : ∀ s, (∀ t ⊂ s, p t) → p s) @@ -789,25 +795,25 @@ protected lemma Nonempty.strong_induction {p : ∀ s, s.Nonempty → Prop} · refine h₁ hs fun t ht hts ↦ ?_ have := card_lt_card hts exact ht.strong_induction h₀ h₁ -termination_by s => Finset.card s +termination_by s => #s /-- Suppose that, given that `p t` can be defined on all supersets of `s` of cardinality less than `n`, one knows how to define `p s`. Then one can inductively define `p s` for all finsets `s` of cardinality less than `n`, starting from finsets of card `n` and iterating. This can be used either to define data, or to prove properties. -/ def strongDownwardInduction {p : Finset α → Sort*} {n : ℕ} - (H : ∀ t₁, (∀ {t₂ : Finset α}, t₂.card ≤ n → t₁ ⊂ t₂ → p t₂) → t₁.card ≤ n → p t₁) : - ∀ s : Finset α, s.card ≤ n → p s + (H : ∀ t₁, (∀ {t₂ : Finset α}, #t₂ ≤ n → t₁ ⊂ t₂ → p t₂) → #t₁ ≤ n → p t₁) : + ∀ s : Finset α, #s ≤ n → p s | s => H s fun {t} ht h => have := Finset.card_lt_card h - have : n - t.card < n - s.card := by omega + have : n - #t < n - #s := by omega strongDownwardInduction H t ht - termination_by s => n - s.card + termination_by s => n - #s @[nolint unusedHavesSuffices] -- Porting note: false positive theorem strongDownwardInduction_eq {p : Finset α → Sort*} - (H : ∀ t₁, (∀ {t₂ : Finset α}, t₂.card ≤ n → t₁ ⊂ t₂ → p t₂) → t₁.card ≤ n → p t₁) + (H : ∀ t₁, (∀ {t₂ : Finset α}, #t₂ ≤ n → t₁ ⊂ t₂ → p t₂) → #t₁ ≤ n → p t₁) (s : Finset α) : strongDownwardInduction H s = H s fun {t} ht _ => strongDownwardInduction H t ht := by rw [strongDownwardInduction] @@ -815,13 +821,13 @@ theorem strongDownwardInduction_eq {p : Finset α → Sort*} /-- Analogue of `strongDownwardInduction` with order of arguments swapped. -/ @[elab_as_elim] def strongDownwardInductionOn {p : Finset α → Sort*} (s : Finset α) - (H : ∀ t₁, (∀ {t₂ : Finset α}, t₂.card ≤ n → t₁ ⊂ t₂ → p t₂) → t₁.card ≤ n → p t₁) : - s.card ≤ n → p s := + (H : ∀ t₁, (∀ {t₂ : Finset α}, #t₂ ≤ n → t₁ ⊂ t₂ → p t₂) → #t₁ ≤ n → p t₁) : + #s ≤ n → p s := strongDownwardInduction H s @[nolint unusedHavesSuffices] -- Porting note: false positive theorem strongDownwardInductionOn_eq {p : Finset α → Sort*} (s : Finset α) - (H : ∀ t₁, (∀ {t₂ : Finset α}, t₂.card ≤ n → t₁ ⊂ t₂ → p t₂) → t₁.card ≤ n → p t₁) : + (H : ∀ t₁, (∀ {t₂ : Finset α}, #t₂ ≤ n → t₁ ⊂ t₂ → p t₂) → #t₁ ≤ n → p t₁) : s.strongDownwardInductionOn H = H s fun {t} ht _ => t.strongDownwardInductionOn H ht := by dsimp only [strongDownwardInductionOn] rw [strongDownwardInduction] From e007f8f7404b132837d23260b6733539080a0511 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 15:41:53 +0000 Subject: [PATCH 023/131] =?UTF-8?q?feat(LinearAlgebra/Quotient):=20`(?= =?UTF-8?q?=E2=88=80=20a=20:=20=20M=20=E2=A7=B8=20p,=20P=20a)=20=E2=86=94?= =?UTF-8?q?=20=E2=88=80=20a=20:=20M,=20P=20=E2=9F=A6a=E2=9F=A7`=20(#17808)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From PFR --- Mathlib/LinearAlgebra/Quotient.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/LinearAlgebra/Quotient.lean b/Mathlib/LinearAlgebra/Quotient.lean index 9f3e755dd20eb..ca687f953d024 100644 --- a/Mathlib/LinearAlgebra/Quotient.lean +++ b/Mathlib/LinearAlgebra/Quotient.lean @@ -104,6 +104,8 @@ theorem mk_neg : (mk (-x) : M ⧸ p) = -(mk x : M ⧸ p) := theorem mk_sub : (mk (x - y) : M ⧸ p) = (mk x : M ⧸ p) - (mk y : M ⧸ p) := rfl +protected nonrec lemma «forall» {P : M ⧸ p → Prop} : (∀ a, P a) ↔ ∀ a, P (mk a) := Quotient.forall + section SMul variable {S : Type*} [SMul S R] [SMul S M] [IsScalarTower S R M] (P : Submodule R M) From d2ce9970d101778fc9d13ea498ab229461ed3c07 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 15:41:55 +0000 Subject: [PATCH 024/131] chore(Analysis): remove unused variables (#17874) These were already unused on the master branch but became exposed to the linter after #17870 changed some `CoeFun` instances. --- Mathlib/Analysis/Analytic/Basic.lean | 2 +- Mathlib/Analysis/Analytic/Inverse.lean | 6 +++--- Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean | 2 +- Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean | 2 +- 4 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index 04e001a4a25bc..35af733c58d75 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -638,7 +638,7 @@ theorem HasFPowerSeriesWithinOnBall.coeff_zero (hf : HasFPowerSeriesWithinOnBall (v : Fin 0 → E) : pf 0 v = f x := by have v_eq : v = fun i => 0 := Subsingleton.elim _ _ have zero_mem : (0 : E) ∈ EMetric.ball (0 : E) r := by simp [hf.r_pos] - have : ∀ i, i ≠ 0 → (pf i fun j => 0) = 0 := by + have : ∀ i, i ≠ 0 → (pf i fun _ => 0) = 0 := by intro i hi have : 0 < i := pos_iff_ne_zero.2 hi exact ContinuousMultilinearMap.map_coord_zero _ (⟨0, this⟩ : Fin i) rfl diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index eadc577c91124..34ad1130c225c 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -272,7 +272,7 @@ theorem rightInv_coeff (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] congr (config := { closePost := false }) 1 ext v have N : 0 < n + 2 := by norm_num - have : ((p 1) fun i : Fin 1 => 0) = 0 := ContinuousMultilinearMap.map_zero _ + have : ((p 1) fun _ : Fin 1 => 0) = 0 := ContinuousMultilinearMap.map_zero _ simp [comp_rightInv_aux1 N, lt_irrefl n, this, comp_rightInv_aux2, -Set.toFinset_setOf] /-! ### Coincidence of the left and the right inverse -/ @@ -630,10 +630,10 @@ lemma HasFPowerSeriesAt.eventually_hasSum_of_comp {f : E → F} {g : F → G} simp only [id_eq, eventually_atTop, ge_iff_le] rcases mem_nhds_iff.1 hu with ⟨v, vu, v_open, hv⟩ obtain ⟨a₀, b₀, hab⟩ : ∃ a₀ b₀, ∀ (a b : ℕ), a₀ ≤ a → b₀ ≤ b → - q.partialSum a (p.partialSum b y - (p 0) fun x ↦ 0) ∈ v := by + q.partialSum a (p.partialSum b y - (p 0) fun _ ↦ 0) ∈ v := by simpa using hy (v_open.mem_nhds hv) refine ⟨a₀, fun a ha ↦ ?_⟩ - have : Tendsto (fun b ↦ q.partialSum a (p.partialSum b y - (p 0) fun x ↦ 0)) atTop + have : Tendsto (fun b ↦ q.partialSum a (p.partialSum b y - (p 0) fun _ ↦ 0)) atTop (𝓝 (q.partialSum a (f (x + y) - f x))) := by have : ContinuousAt (q.partialSum a) (f (x + y) - f x) := (partialSum_continuous q a).continuousAt diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean index f5790c73a631f..cf2ae12780d4c 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean @@ -166,7 +166,7 @@ theorem iteratedDerivWithin_succ {x : 𝕜} (hxs : UniqueDiffWithinAt 𝕜 s x) rw [iteratedDerivWithin_eq_iteratedFDerivWithin, iteratedFDerivWithin_succ_apply_left, iteratedFDerivWithin_eq_equiv_comp, LinearIsometryEquiv.comp_fderivWithin _ hxs, derivWithin] change ((ContinuousMultilinearMap.mkPiRing 𝕜 (Fin n) ((fderivWithin 𝕜 - (iteratedDerivWithin n f s) s x : 𝕜 → F) 1) : (Fin n → 𝕜) → F) fun i : Fin n => 1) = + (iteratedDerivWithin n f s) s x : 𝕜 → F) 1) : (Fin n → 𝕜) → F) fun _ : Fin n => 1) = (fderivWithin 𝕜 (iteratedDerivWithin n f s) s x : 𝕜 → F) 1 simp diff --git a/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean index 64ad667b3190a..c3e4233f09ac9 100644 --- a/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean @@ -115,7 +115,7 @@ variable [T2Space R] [T2Space M] theorem exp_def_of_smul_comm (x : tsze R M) (hx : MulOpposite.op x.fst • x.snd = x.fst • x.snd) : exp 𝕜 x = inl (exp 𝕜 x.fst) + inr (exp 𝕜 x.fst • x.snd) := by simp_rw [exp, FormalMultilinearSeries.sum] - by_cases h : Summable (fun (n : ℕ) => (expSeries 𝕜 R n) fun x_1 ↦ fst x) + by_cases h : Summable (fun (n : ℕ) => (expSeries 𝕜 R n) fun _ ↦ fst x) · refine (hasSum_expSeries_of_smul_comm 𝕜 x hx ?_).tsum_eq exact h.hasSum · rw [tsum_eq_zero_of_not_summable h, zero_smul, inr_zero, inl_zero, zero_add, From d3038070bcab78ca09317951523576e264e9cf84 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 17 Oct 2024 16:19:08 +0000 Subject: [PATCH 025/131] chore: another removal of unused variables (#17790) --- Mathlib/Algebra/Field/ULift.lean | 4 +-- Mathlib/Algebra/Group/Action/Basic.lean | 2 +- Mathlib/Algebra/Group/Indicator.lean | 6 ++-- Mathlib/Algebra/Group/Pi/Lemmas.lean | 3 +- Mathlib/Algebra/Group/Submonoid/Defs.lean | 4 --- Mathlib/Algebra/Group/Subsemigroup/Defs.lean | 4 +-- .../Algebra/GroupWithZero/Action/Basic.lean | 2 +- Mathlib/Algebra/GroupWithZero/Action/Pi.lean | 5 +--- Mathlib/Algebra/Module/Defs.lean | 10 +++---- Mathlib/Algebra/Module/Equiv/Defs.lean | 12 +++----- Mathlib/Algebra/Module/LinearMap/Defs.lean | 28 ++++++++----------- Mathlib/Algebra/Module/Pi.lean | 3 -- Mathlib/Algebra/Module/Submodule/Range.lean | 19 ++++++------- Mathlib/Algebra/Order/Interval/Set/Group.lean | 2 +- Mathlib/Algebra/Order/Ring/Abs.lean | 8 +++--- Mathlib/Algebra/Regular/SMul.lean | 3 +- Mathlib/Algebra/Ring/Pi.lean | 3 +- Mathlib/Algebra/Ring/Prod.lean | 2 +- .../Analysis/InnerProductSpace/Symmetric.lean | 10 ++----- .../OperatorNorm/Completeness.lean | 11 ++++---- Mathlib/Data/Fin/VecNotation.lean | 2 +- Mathlib/Data/Int/Cast/Lemmas.lean | 6 ++-- Mathlib/Data/Int/Lemmas.lean | 2 -- Mathlib/Data/Int/Order/Lemmas.lean | 3 -- Mathlib/Data/List/Perm.lean | 2 +- Mathlib/Data/List/Sort.lean | 2 +- Mathlib/Data/Nat/Cast/Order/Ring.lean | 2 +- Mathlib/Dynamics/FixedPoints/Basic.lean | 2 +- .../Geometry/Manifold/Algebra/LieGroup.lean | 7 ++--- .../Manifold/VectorBundle/SmoothSection.lean | 10 ++----- Mathlib/GroupTheory/Congruence/Defs.lean | 4 +-- Mathlib/GroupTheory/CosetCover.lean | 2 +- Mathlib/GroupTheory/DoubleCoset.lean | 2 +- Mathlib/GroupTheory/GroupAction/Hom.lean | 5 +--- Mathlib/GroupTheory/NoncommPiCoprod.lean | 9 ------ Mathlib/GroupTheory/Sylow.lean | 4 +-- Mathlib/LinearAlgebra/Basis/Basic.lean | 15 ++++------ .../LinearAlgebra/FiniteDimensional/Defs.lean | 9 ++---- Mathlib/MeasureTheory/Integral/Bochner.lean | 21 ++++++-------- Mathlib/Order/Bounds/Basic.lean | 2 +- Mathlib/Order/CompleteLattice.lean | 2 +- Mathlib/Order/GaloisConnection.lean | 6 ++-- Mathlib/Order/PiLex.lean | 2 +- .../RingTheory/Localization/Submodule.lean | 9 ++---- Mathlib/Topology/Algebra/Module/Basic.lean | 14 ++++------ 45 files changed, 106 insertions(+), 179 deletions(-) diff --git a/Mathlib/Algebra/Field/ULift.lean b/Mathlib/Algebra/Field/ULift.lean index ccb6baccbc575..215b218b67831 100644 --- a/Mathlib/Algebra/Field/ULift.lean +++ b/Mathlib/Algebra/Field/ULift.lean @@ -15,8 +15,8 @@ This file defines instances for field, semifield and related structures on `ULif (Recall `ULift α` is just a "copy" of a type `α` in a higher universe.) -/ -universe u v -variable {α : Type u} {x y : ULift.{v} α} +universe u +variable {α : Type u} namespace ULift diff --git a/Mathlib/Algebra/Group/Action/Basic.lean b/Mathlib/Algebra/Group/Action/Basic.lean index 8515f84bac94f..0077326ec936b 100644 --- a/Mathlib/Algebra/Group/Action/Basic.lean +++ b/Mathlib/Algebra/Group/Action/Basic.lean @@ -16,7 +16,7 @@ This file contains lemmas about group actions that require more imports than assert_not_exists MonoidWithZero -variable {α β γ : Type*} +variable {α β : Type*} section MulAction diff --git a/Mathlib/Algebra/Group/Indicator.lean b/Mathlib/Algebra/Group/Indicator.lean index fc2d14cdfb140..38b9e8ce30021 100644 --- a/Mathlib/Algebra/Group/Indicator.lean +++ b/Mathlib/Algebra/Group/Indicator.lean @@ -33,7 +33,7 @@ assert_not_exists MonoidWithZero open Function -variable {α β ι M N : Type*} +variable {α β M N : Type*} namespace Set @@ -262,7 +262,7 @@ end One section Monoid -variable [MulOneClass M] {s t : Set α} {f g : α → M} {a : α} +variable [MulOneClass M] {s t : Set α} {a : α} @[to_additive] theorem mulIndicator_union_mul_inter_apply (f : α → M) (s t : Set α) (a : α) : @@ -360,7 +360,7 @@ end Monoid section Group -variable {G : Type*} [Group G] {s t : Set α} {f g : α → G} {a : α} +variable {G : Type*} [Group G] {s t : Set α} @[to_additive] theorem mulIndicator_inv' (s : Set α) (f : α → G) : mulIndicator s f⁻¹ = (mulIndicator s f)⁻¹ := diff --git a/Mathlib/Algebra/Group/Pi/Lemmas.lean b/Mathlib/Algebra/Group/Pi/Lemmas.lean index 1dee66ec01d1b..0f412767df4f9 100644 --- a/Mathlib/Algebra/Group/Pi/Lemmas.lean +++ b/Mathlib/Algebra/Group/Pi/Lemmas.lean @@ -26,8 +26,7 @@ variable {I : Type u} -- The indexing type variable {f : I → Type v} --- The family of types already equipped with instances -variable (x y : ∀ i, f i) (i j : I) +variable (i : I) @[to_additive (attr := simp)] theorem Set.range_one {α β : Type*} [One β] [Nonempty α] : Set.range (1 : α → β) = {1} := diff --git a/Mathlib/Algebra/Group/Submonoid/Defs.lean b/Mathlib/Algebra/Group/Submonoid/Defs.lean index 4f005ea640695..d1a62b639d86a 100644 --- a/Mathlib/Algebra/Group/Submonoid/Defs.lean +++ b/Mathlib/Algebra/Group/Submonoid/Defs.lean @@ -47,15 +47,11 @@ submonoid, submonoids assert_not_exists CompleteLattice assert_not_exists MonoidWithZero --- Only needed for notation --- Only needed for notation variable {M : Type*} {N : Type*} -variable {A : Type*} section NonAssoc 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 : outParam Type*) [One M] [SetLike S M] : Prop where diff --git a/Mathlib/Algebra/Group/Subsemigroup/Defs.lean b/Mathlib/Algebra/Group/Subsemigroup/Defs.lean index 946d325924172..aa0008eee16ba 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Defs.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Defs.lean @@ -47,13 +47,11 @@ subsemigroup, subsemigroups assert_not_exists CompleteLattice assert_not_exists MonoidWithZero --- Only needed for notation -variable {M : Type*} {N : Type*} {A : Type*} +variable {M : Type*} {N : Type*} section NonAssoc 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 : outParam Type*) [Mul M] [SetLike S M] : Prop where diff --git a/Mathlib/Algebra/GroupWithZero/Action/Basic.lean b/Mathlib/Algebra/GroupWithZero/Action/Basic.lean index f61edb220be6b..39b682e9f1e61 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Basic.lean @@ -68,7 +68,7 @@ protected lemma MulAction.surjective₀ (ha : a ≠ 0) : Surjective (a • · : end GroupWithZero section DistribMulAction -variable [Group G] [Monoid M] [AddMonoid A] [DistribMulAction M A] +variable [Group G] [Monoid M] [AddMonoid A] variable (A) /-- Each element of the group defines an additive monoid isomorphism. diff --git a/Mathlib/Algebra/GroupWithZero/Action/Pi.lean b/Mathlib/Algebra/GroupWithZero/Action/Pi.lean index a68e1a97b1d8a..f465f4b4caa56 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Pi.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Pi.lean @@ -21,16 +21,13 @@ This file defines instances for `MulActionWithZero` and related structures on `P -/ -universe u v w +universe u v variable {I : Type u} -- The indexing type variable {f : I → Type v} --- The family of types already equipped with instances -variable (x y : ∀ i, f i) (i : I) - namespace Pi instance smulZeroClass (α) {n : ∀ i, Zero <| f i} [∀ i, SMulZeroClass α <| f i] : diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index f810b34644f19..a024554de986a 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -46,7 +46,7 @@ open Function Set universe u v -variable {α R k S M M₂ M₃ ι : Type*} +variable {R S M M₂ : Type*} /-- A module is a generalization of vector spaces to a scalar semiring. It consists of a scalar semiring `R` and an additive monoid of "vectors" `M`, @@ -63,7 +63,7 @@ class Module (R : Type u) (M : Type v) [Semiring R] [AddCommMonoid M] extends section AddCommMonoid -variable [Semiring R] [AddCommMonoid M] [Module R M] (r s : R) (x y : M) +variable [Semiring R] [AddCommMonoid M] [Module R M] (r s : R) (x : M) -- see Note [lower instance priority] /-- A module over a semiring automatically inherits a `MulActionWithZero` structure. -/ @@ -218,7 +218,7 @@ theorem Module.ext' {R : Type*} [Semiring R] {M : Type*} [AddCommMonoid M] (P Q section Module -variable [Ring R] [AddCommGroup M] [Module R M] (r s : R) (x y : M) +variable [Ring R] [AddCommGroup M] [Module R M] (r : R) (x : M) @[simp] theorem neg_smul : -r • x = -(r • x) := @@ -360,7 +360,7 @@ end AddCommMonoid section AddCommGroup -variable [Semiring S] [Ring R] [AddCommGroup M] [Module S M] [Module R M] +variable [Ring R] [AddCommGroup M] [Module R M] section @@ -483,7 +483,7 @@ theorem two_nsmul_eq_zero end Nat -variable [Semiring R] [AddCommMonoid M] [Module R M] +variable [Semiring R] variable (R M) /-- If `M` is an `R`-module with one and `M` has characteristic zero, then `R` has characteristic diff --git a/Mathlib/Algebra/Module/Equiv/Defs.lean b/Mathlib/Algebra/Module/Equiv/Defs.lean index 9a5d7dcbdfd0d..7d08836e2d79f 100644 --- a/Mathlib/Algebra/Module/Equiv/Defs.lean +++ b/Mathlib/Algebra/Module/Equiv/Defs.lean @@ -38,11 +38,9 @@ assert_not_exists Pi.module open Function -universe u u' v w x y z - variable {R : Type*} {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} -variable {k : Type*} {K : Type*} {S : Type*} {M : Type*} {M₁ : Type*} {M₂ : Type*} {M₃ : Type*} -variable {N₁ : Type*} {N₂ : Type*} {N₃ : Type*} {N₄ : Type*} {ι : Type*} +variable {S : Type*} {M : Type*} {M₁ : Type*} {M₂ : Type*} {M₃ : Type*} +variable {N₁ : Type*} {N₂ : Type*} section @@ -133,7 +131,6 @@ namespace LinearEquiv section AddCommMonoid -variable {M₄ : Type*} variable [Semiring R] [Semiring S] section @@ -198,7 +195,7 @@ section variable [Semiring R₁] [Semiring R₂] [Semiring R₃] variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] -variable [AddCommMonoid M₃] [AddCommMonoid M₄] +variable [AddCommMonoid M₃] variable [AddCommMonoid N₁] [AddCommMonoid N₂] variable {module_M : Module R M} {module_S_M₂ : Module S M₂} {σ : R →+* S} {σ' : S →+* R} variable {re₁ : RingHomInvPair σ σ'} {re₂ : RingHomInvPair σ' σ} @@ -506,8 +503,7 @@ def _root_.RingEquiv.toSemilinearEquiv (f : R ≃+* S) : toFun := f map_smul' := f.map_mul } -variable [Semiring R₁] [Semiring R₂] [Semiring R₃] -variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] +variable [AddCommMonoid M] /-- An involutive linear map is a linear equivalence. -/ def ofInvolutive {σ σ' : R →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index 0a75f33f5bf9c..21a6e98c5445a 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -57,9 +57,9 @@ assert_not_exists Field open Function -universe u u' v w x y z +universe u u' v w -variable {R R₁ R₂ R₃ k S S₃ T M M₁ M₂ M₃ N₁ N₂ N₃ ι : Type*} +variable {R R₁ R₂ R₃ S S₃ T M M₁ M₂ M₃ N₂ N₃ : Type*} /-- A map `f` between modules over a semiring is linear if it satisfies the two properties `f (x + y) = f x + f y` and `f (c • x) = c • f x`. The predicate `IsLinearMap R f` asserts this @@ -194,7 +194,6 @@ variable [Semiring R] [Semiring S] section variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] -variable [AddCommMonoid N₁] [AddCommMonoid N₂] [AddCommMonoid N₃] variable [Module R M] [Module R M₂] [Module S M₃] variable {σ : R →+* S} @@ -299,15 +298,14 @@ end section variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] -variable [AddCommMonoid N₁] [AddCommMonoid N₂] [AddCommMonoid N₃] variable [Module R M] [Module R M₂] [Module S M₃] variable (σ : R →+* S) -variable (fₗ gₗ : M →ₗ[R] M₂) (f g : M →ₛₗ[σ] M₃) +variable (fₗ : M →ₗ[R] M₂) (f g : M →ₛₗ[σ] M₃) theorem isLinear : IsLinearMap R fₗ := ⟨fₗ.map_add', fₗ.map_smul'⟩ -variable {fₗ gₗ f g σ} +variable {fₗ f g σ} theorem coe_injective : Injective (DFunLike.coe : (M →ₛₗ[σ] M₃) → _) := DFunLike.coe_injective @@ -323,7 +321,7 @@ protected theorem congr_fun (h : f = g) (x : M) : f x = g x := theorem mk_coe (f : M →ₛₗ[σ] M₃) (h) : (LinearMap.mk f h : M →ₛₗ[σ] M₃) = f := rfl -variable (fₗ gₗ f g) +variable (fₗ f g) protected theorem map_add (x y : M) : f (x + y) = f x + f y := map_add f x y @@ -537,7 +535,7 @@ theorem cancel_left (hf : Injective f) : f.comp g = f.comp g' ↔ g = g' := end -variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] +variable [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] /-- If a function `g` is a left and right inverse of a linear map `f`, then `g` is linear itself. -/ def inverse [Module R M] [Module S M₂] {σ : R →+* S} {σ' : S →+* R} [RingHomInvPair σ σ'] @@ -721,12 +719,11 @@ namespace LinearMap section SMul -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₂} 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 : SMul S (M →ₛₗ[σ₁₂] M₂) := @@ -762,9 +759,9 @@ section Arithmetic variable [Semiring R₁] [Semiring R₂] [Semiring R₃] variable [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] -variable [AddCommGroup N₁] [AddCommGroup N₂] [AddCommGroup N₃] +variable [AddCommGroup N₂] [AddCommGroup N₃] variable [Module R₁ M] [Module R₂ M₂] [Module R₃ M₃] -variable [Module R₁ N₁] [Module R₂ N₂] [Module R₃ N₃] +variable [Module R₂ N₂] [Module R₃ N₃] variable {σ₁₂ : R₁ →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R₁ →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] /-- The constant 0 map is linear. -/ @@ -897,7 +894,6 @@ 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 : DistribMulAction S (M →ₛₗ[σ₁₂] M₂) where one_smul _ := ext fun _ ↦ one_smul _ _ diff --git a/Mathlib/Algebra/Module/Pi.lean b/Mathlib/Algebra/Module/Pi.lean index 002ec529cf852..1d6a967b8db91 100644 --- a/Mathlib/Algebra/Module/Pi.lean +++ b/Mathlib/Algebra/Module/Pi.lean @@ -22,9 +22,6 @@ variable {I : Type u} -- The indexing type variable {f : I → Type v} --- The family of types already equipped with instances -variable (x y : ∀ i, f i) (i : I) - namespace Pi theorem _root_.IsSMulRegular.pi {α : Type*} [∀ i, SMul α <| f i] {k : α} diff --git a/Mathlib/Algebra/Module/Submodule/Range.lean b/Mathlib/Algebra/Module/Submodule/Range.lean index 40e7e2b452789..933cc427c94e9 100644 --- a/Mathlib/Algebra/Module/Submodule/Range.lean +++ b/Mathlib/Algebra/Module/Submodule/Range.lean @@ -28,7 +28,7 @@ linear algebra, vector space, module, range open Function variable {R : Type*} {R₂ : Type*} {R₃ : Type*} -variable {K : Type*} {K₂ : Type*} +variable {K : Type*} variable {M : Type*} {M₂ : Type*} {M₃ : Type*} variable {V : Type*} {V₂ : Type*} @@ -38,13 +38,11 @@ section AddCommMonoid variable [Semiring R] [Semiring R₂] [Semiring R₃] variable [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] -variable {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} -variable [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] variable [Module R M] [Module R₂ M₂] [Module R₃ M₃] open Submodule -variable {σ₂₁ : R₂ →+* R} {τ₁₂ : R →+* R₂} {τ₂₃ : R₂ →+* R₃} {τ₁₃ : R →+* R₃} +variable {τ₁₂ : R →+* R₂} {τ₂₃ : R₂ →+* R₃} {τ₁₃ : R →+* R₃} variable [RingHomCompTriple τ₁₂ τ₂₃ τ₁₃] section @@ -190,11 +188,10 @@ end AddCommMonoid section Ring -variable [Ring R] [Ring R₂] [Ring R₃] -variable [AddCommGroup M] [AddCommGroup M₂] [AddCommGroup M₃] -variable [Module R M] [Module R₂ M₂] [Module R₃ M₃] -variable {τ₁₂ : R →+* R₂} {τ₂₃ : R₂ →+* R₃} {τ₁₃ : R →+* R₃} -variable [RingHomCompTriple τ₁₂ τ₂₃ τ₁₃] +variable [Ring R] [Ring R₂] +variable [AddCommGroup M] [AddCommGroup M₂] +variable [Module R M] [Module R₂ M₂] +variable {τ₁₂ : R →+* R₂} variable {F : Type*} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] variable {f : F} @@ -228,7 +225,7 @@ end Ring section Semifield -variable [Semifield K] [Semifield K₂] +variable [Semifield K] variable [AddCommMonoid V] [Module K V] variable [AddCommMonoid V₂] [Module K V₂] @@ -249,7 +246,7 @@ section AddCommMonoid variable [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] variable [Module R M] [Module R₂ M₂] -variable (p p' : Submodule R M) (q : Submodule R₂ M₂) +variable (p : Submodule R M) variable {τ₁₂ : R →+* R₂} variable {F : Type*} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] diff --git a/Mathlib/Algebra/Order/Interval/Set/Group.lean b/Mathlib/Algebra/Order/Interval/Set/Group.lean index a95c3a44af7a1..d9684e57e4eb9 100644 --- a/Mathlib/Algebra/Order/Interval/Set/Group.lean +++ b/Mathlib/Algebra/Order/Interval/Set/Group.lean @@ -19,7 +19,7 @@ namespace Set section OrderedCommGroup -variable [OrderedCommGroup α] {a b c d : α} +variable [OrderedCommGroup α] {a c d : α} /-! `inv_mem_Ixx_iff`, `sub_mem_Ixx_iff` -/ diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index e61b8d4297a31..25cd393bf7434 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -17,7 +17,7 @@ import Mathlib.Data.Nat.Cast.Order.Ring variable {α : Type*} section LinearOrderedAddCommGroup -variable [LinearOrderedCommGroup α] {a b : α} +variable [LinearOrderedCommGroup α] @[to_additive] lemma mabs_zpow (n : ℤ) (a : α) : |a ^ n|ₘ = |a|ₘ ^ |n| := by obtain n0 | n0 := le_total 0 n @@ -34,7 +34,7 @@ lemma odd_abs [LinearOrder α] [Ring α] {a : α} : Odd (abs a) ↔ Odd a := by section LinearOrderedRing -variable [LinearOrderedRing α] {n : ℕ} {a b c : α} +variable [LinearOrderedRing α] {n : ℕ} {a b : α} @[simp] lemma abs_one : |(1 : α)| = 1 := abs_of_pos zero_lt_one @@ -141,7 +141,7 @@ end LinearOrderedRing section LinearOrderedCommRing -variable [LinearOrderedCommRing α] {a b c d : α} +variable [LinearOrderedCommRing α] theorem abs_sub_sq (a b : α) : |a - b| * |a - b| = a * a + b * b - (1 + 1) * a * b := by rw [abs_mul_abs_self] @@ -152,7 +152,7 @@ end LinearOrderedCommRing section -variable [Ring α] [LinearOrder α] {a b : α} +variable [Ring α] [LinearOrder α] @[simp] theorem abs_dvd (a b : α) : |a| ∣ b ↔ a ∣ b := by diff --git a/Mathlib/Algebra/Regular/SMul.lean b/Mathlib/Algebra/Regular/SMul.lean index 5895185e0774c..3dfb2cff80cf0 100644 --- a/Mathlib/Algebra/Regular/SMul.lean +++ b/Mathlib/Algebra/Regular/SMul.lean @@ -162,8 +162,7 @@ end MonoidSMul section MonoidWithZero -variable [MonoidWithZero R] [MonoidWithZero S] [Zero M] [MulActionWithZero R M] - [MulActionWithZero R S] [MulActionWithZero S M] [IsScalarTower R S M] +variable [MonoidWithZero R] [Zero M] [MulActionWithZero R M] /-- The element `0` is `M`-regular if and only if `M` is trivial. -/ protected theorem subsingleton (h : IsSMulRegular M (0 : R)) : Subsingleton M := diff --git a/Mathlib/Algebra/Ring/Pi.lean b/Mathlib/Algebra/Ring/Pi.lean index 394051c18c7d9..4c20d456a41b7 100644 --- a/Mathlib/Algebra/Ring/Pi.lean +++ b/Mathlib/Algebra/Ring/Pi.lean @@ -25,8 +25,7 @@ variable {I : Type u} -- The indexing type variable {f : I → Type v} --- The family of types already equipped with instances -variable (x y : ∀ i, f i) (i : I) +variable (i : I) instance distrib [∀ i, Distrib <| f i] : Distrib (∀ i : I, f i) := { add := (· + ·) diff --git a/Mathlib/Algebra/Ring/Prod.lean b/Mathlib/Algebra/Ring/Prod.lean index 2883cf4e33ea0..67ee0ebf5b125 100644 --- a/Mathlib/Algebra/Ring/Prod.lean +++ b/Mathlib/Algebra/Ring/Prod.lean @@ -23,7 +23,7 @@ trivial `simp` lemmas, and define the following operations on `RingHom`s and sim -/ -variable {α β R R' S S' T T' : Type*} +variable {R R' S S' T : Type*} namespace Prod diff --git a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean index 9acf7b33a0635..2fcd08ccaa66d 100644 --- a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean +++ b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean @@ -38,11 +38,8 @@ open ComplexConjugate section Seminormed -variable {𝕜 E E' F G : Type*} [RCLike 𝕜] +variable {𝕜 E : Type*} [RCLike 𝕜] variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] -variable [SeminormedAddCommGroup F] [InnerProductSpace 𝕜 F] -variable [SeminormedAddCommGroup G] [InnerProductSpace 𝕜 G] -variable [SeminormedAddCommGroup E'] [InnerProductSpace ℝ E'] local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y @@ -189,11 +186,8 @@ end Seminormed section Normed -variable {𝕜 E E' F G : Type*} [RCLike 𝕜] +variable {𝕜 E : Type*} [RCLike 𝕜] variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] -variable [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] -variable [NormedAddCommGroup G] [InnerProductSpace 𝕜 G] -variable [NormedAddCommGroup E'] [InnerProductSpace ℝ E'] local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean index 783b55750d8fb..4043b6833c43d 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean @@ -20,13 +20,12 @@ open Filter hiding map_smul open scoped NNReal Topology Uniformity -- the `ₗ` subscript variables are for special cases about linear (as opposed to semilinear) maps -variable {𝕜 𝕜₂ 𝕜₃ E Eₗ F Fₗ G Gₗ 𝓕 : Type*} -variable [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedAddCommGroup G] - [NormedAddCommGroup Fₗ] +variable {𝕜 𝕜₂ E F Fₗ : Type*} +variable [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedAddCommGroup Fₗ] -variable [NontriviallyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] [NontriviallyNormedField 𝕜₃] - [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [NormedSpace 𝕜₃ G] [NormedSpace 𝕜 Fₗ] (c : 𝕜) - {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₃ : 𝕜₂ →+* 𝕜₃} (f g : E →SL[σ₁₂] F) (x y z : E) +variable [NontriviallyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] + [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [NormedSpace 𝕜 Fₗ] + {σ₁₂ : 𝕜 →+* 𝕜₂} (f g : E →SL[σ₁₂] F) namespace ContinuousLinearMap diff --git a/Mathlib/Data/Fin/VecNotation.lean b/Mathlib/Data/Fin/VecNotation.lean index d3b561640f0ae..c931078da5ce1 100644 --- a/Mathlib/Data/Fin/VecNotation.lean +++ b/Mathlib/Data/Fin/VecNotation.lean @@ -108,7 +108,7 @@ instance _root_.PiFin.hasRepr [Repr α] : Repr (Fin n → α) where end MatrixNotation -variable {m n o : ℕ} {m' n' o' : Type*} +variable {m n o : ℕ} theorem empty_eq (v : Fin 0 → α) : v = ![] := Subsingleton.elim _ _ diff --git a/Mathlib/Data/Int/Cast/Lemmas.lean b/Mathlib/Data/Int/Cast/Lemmas.lean index da5f6ffae2ed7..822449e1e3d3b 100644 --- a/Mathlib/Data/Int/Cast/Lemmas.lean +++ b/Mathlib/Data/Int/Cast/Lemmas.lean @@ -77,7 +77,7 @@ lemma cast_ne_one : (n : α) ≠ 1 ↔ n ≠ 1 := cast_eq_one.not end AddGroupWithOne section NonAssocRing -variable [NonAssocRing α] {a b : α} {n : ℤ} +variable [NonAssocRing α] variable (α) in /-- `coe : ℤ → α` as a `RingHom`. -/ @@ -147,7 +147,7 @@ end SemiconjBy namespace Commute section NonAssocRing -variable [NonAssocRing α] {a b : α} {n : ℤ} +variable [NonAssocRing α] {a : α} {n : ℤ} @[simp] lemma intCast_left : Commute (n : α) a := Int.cast_commute _ _ @@ -159,7 +159,7 @@ variable [NonAssocRing α] {a b : α} {n : ℤ} end NonAssocRing section Ring -variable [Ring α] {a b : α} {n : ℤ} +variable [Ring α] {a b : α} @[simp] lemma intCast_mul_right (h : Commute a b) (m : ℤ) : Commute a (m * b) := SemiconjBy.intCast_mul_right h m diff --git a/Mathlib/Data/Int/Lemmas.lean b/Mathlib/Data/Int/Lemmas.lean index e989bd71419b0..3842fbb1c3222 100644 --- a/Mathlib/Data/Int/Lemmas.lean +++ b/Mathlib/Data/Int/Lemmas.lean @@ -37,8 +37,6 @@ theorem succ_natCast_pos (n : ℕ) : 0 < (n : ℤ) + 1 := /-! ### `natAbs` -/ -variable {a b : ℤ} {n : ℕ} - theorem natAbs_eq_iff_sq_eq {a b : ℤ} : a.natAbs = b.natAbs ↔ a ^ 2 = b ^ 2 := by rw [sq, sq] exact natAbs_eq_iff_mul_self_eq diff --git a/Mathlib/Data/Int/Order/Lemmas.lean b/Mathlib/Data/Int/Order/Lemmas.lean index d664a5b284aca..3e9566c7910ff 100644 --- a/Mathlib/Data/Int/Order/Lemmas.lean +++ b/Mathlib/Data/Int/Order/Lemmas.lean @@ -19,9 +19,6 @@ namespace Int /-! ### nat abs -/ - -variable {a b : ℤ} {n : ℕ} - theorem natAbs_eq_iff_mul_self_eq {a b : ℤ} : a.natAbs = b.natAbs ↔ a * a = b * b := by rw [← abs_eq_iff_mul_self_eq, abs_eq_natAbs, abs_eq_natAbs] exact Int.natCast_inj.symm diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index 8f5cb1997dd64..efeb8dd594a06 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -45,7 +45,7 @@ section Rel open Relator -variable {γ : Type*} {δ : Type*} {r : α → β → Prop} {p : γ → δ → Prop} +variable {r : α → β → Prop} local infixr:80 " ∘r " => Relation.Comp diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 2b327d5d70ca3..513e81df6f0a3 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -49,7 +49,7 @@ namespace List section Sorted -variable {α : Type u} {β : Type v} {r : α → α → Prop} {s : β → β → Prop} {a : α} {l : List α} +variable {α : Type u} {r : α → α → Prop} {a : α} {l : List α} /-- `Sorted r l` is the same as `List.Pairwise r l`, preferred in the case that `r` is a `<` or `≤`-like relation (transitive and antisymmetric or asymmetric) -/ diff --git a/Mathlib/Data/Nat/Cast/Order/Ring.lean b/Mathlib/Data/Nat/Cast/Order/Ring.lean index b03a53d8607f4..2f6687f56c7ac 100644 --- a/Mathlib/Data/Nat/Cast/Order/Ring.lean +++ b/Mathlib/Data/Nat/Cast/Order/Ring.lean @@ -12,7 +12,7 @@ import Mathlib.Data.Nat.Cast.Order.Basic -/ -variable {α β : Type*} +variable {α : Type*} namespace Nat diff --git a/Mathlib/Dynamics/FixedPoints/Basic.lean b/Mathlib/Dynamics/FixedPoints/Basic.lean index 9f13fb19bc6cf..78081cb348764 100644 --- a/Mathlib/Dynamics/FixedPoints/Basic.lean +++ b/Mathlib/Dynamics/FixedPoints/Basic.lean @@ -27,7 +27,7 @@ open Equiv universe u v -variable {α : Type u} {β : Type v} {f fa g : α → α} {x y : α} {fb : β → β} {m n k : ℕ} {e : Perm α} +variable {α : Type u} {β : Type v} {f fa g : α → α} {x : α} {fb : β → β} {e : Perm α} namespace Function diff --git a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean index 8bec2b7fa5b4f..c054be8efa60b 100644 --- a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean +++ b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean @@ -79,13 +79,10 @@ class LieGroup {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [Topolo section PointwiseDivision variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalSpace H] {E : Type*} - [NormedAddCommGroup E] [NormedSpace 𝕜 E] {I : ModelWithCorners 𝕜 E H} {F : Type*} - [NormedAddCommGroup F] [NormedSpace 𝕜 F] {J : ModelWithCorners 𝕜 F F} {G : Type*} + [NormedAddCommGroup E] [NormedSpace 𝕜 E] {I : ModelWithCorners 𝕜 E H} {G : Type*} [TopologicalSpace G] [ChartedSpace H G] [Group G] [LieGroup I G] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M : Type*} [TopologicalSpace M] [ChartedSpace H' M] - {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] - {I'' : ModelWithCorners 𝕜 E'' H''} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H'' M'] {n : ℕ∞} section @@ -236,7 +233,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalS [TopologicalSpace G] [ChartedSpace H G] [Inv G] [Zero G] [SmoothInv₀ I G] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M : Type*} [TopologicalSpace M] [ChartedSpace H' M] - {n : ℕ∞} {f g : M → G} + {n : ℕ∞} {f : M → G} theorem smoothAt_inv₀ {x : G} (hx : x ≠ 0) : SmoothAt I I (fun y ↦ y⁻¹) x := SmoothInv₀.smoothAt_inv₀ hx diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index 10e4fe0a92fe3..198e5170ad73e 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -21,12 +21,8 @@ open Bundle Filter Function open scoped Bundle Manifold variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H : Type*} - [TopologicalSpace H] {H' : Type*} [TopologicalSpace H'] (I : ModelWithCorners 𝕜 E H) - (I' : ModelWithCorners 𝕜 E' H') {M : Type*} [TopologicalSpace M] [ChartedSpace H M] {M' : Type*} - [TopologicalSpace M'] [ChartedSpace H' M'] {E'' : Type*} [NormedAddCommGroup E''] - [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] {I'' : ModelWithCorners 𝕜 E'' H''} - {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] + [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] variable (F : Type*) [NormedAddCommGroup F] [NormedSpace 𝕜 F] -- `F` model fiber @@ -51,7 +47,7 @@ abbrev SmoothSection := namespace ContMDiffSection -variable {I} {I'} {n} {F} {V} +variable {I} {n} {F} {V} instance : DFunLike Cₛ^n⟮I; F, V⟯ M V where coe := ContMDiffSection.toFun diff --git a/Mathlib/GroupTheory/Congruence/Defs.lean b/Mathlib/GroupTheory/Congruence/Defs.lean index c8c38842eee77..ced9c0f1660d1 100644 --- a/Mathlib/GroupTheory/Congruence/Defs.lean +++ b/Mathlib/GroupTheory/Congruence/Defs.lean @@ -584,7 +584,7 @@ end section MulOneClass -variable [MulOneClass M] [MulOneClass N] [MulOneClass P] (c : Con M) +variable [MulOneClass M] (c : Con M) /-- The quotient of a monoid by a congruence relation is a monoid. -/ @[to_additive "The quotient of an `AddMonoid` by an additive congruence relation is @@ -709,7 +709,7 @@ end Monoids section Groups -variable [Group M] [Group N] [Group P] (c : Con M) +variable [Group M] (c : Con M) /-- Multiplicative congruence relations preserve inversion. -/ @[to_additive "Additive congruence relations preserve negation."] diff --git a/Mathlib/GroupTheory/CosetCover.lean b/Mathlib/GroupTheory/CosetCover.lean index 03164c5973c5e..a8ba591868ffa 100644 --- a/Mathlib/GroupTheory/CosetCover.lean +++ b/Mathlib/GroupTheory/CosetCover.lean @@ -366,7 +366,7 @@ end Submodule section Subspace -variable {k E ι : Type*} [DivisionRing k] [Infinite k] [AddCommGroup E] [Module k E] +variable {k E : Type*} [DivisionRing k] [Infinite k] [AddCommGroup E] [Module k E] {s : Finset (Subspace k E)} /- A vector space over an infinite field cannot be a finite union of proper subspaces. -/ diff --git a/Mathlib/GroupTheory/DoubleCoset.lean b/Mathlib/GroupTheory/DoubleCoset.lean index 771be71a00bb9..a7d7e29638510 100644 --- a/Mathlib/GroupTheory/DoubleCoset.lean +++ b/Mathlib/GroupTheory/DoubleCoset.lean @@ -22,7 +22,7 @@ this is the usual left or right quotient of a group by a subgroup. -- Porting note: removed import -- import Mathlib.Tactic.Group -variable {G : Type*} [Group G] {α : Type*} [Mul α] (J : Subgroup G) (g : G) +variable {G : Type*} [Group G] {α : Type*} [Mul α] open MulOpposite open scoped Pointwise diff --git a/Mathlib/GroupTheory/GroupAction/Hom.lean b/Mathlib/GroupTheory/GroupAction/Hom.lean index e2fc5be8f79d8..fdebbaf6ec0f9 100644 --- a/Mathlib/GroupTheory/GroupAction/Hom.lean +++ b/Mathlib/GroupTheory/GroupAction/Hom.lean @@ -524,13 +524,10 @@ def inverse (f : A →+[M] B₁) (g : B₁ → A) (h₁ : Function.LeftInverse g section Semiring variable (R : Type*) [Semiring R] [MulSemiringAction M R] -variable (R' : Type*) [Ring R'] [MulSemiringAction M R'] variable (S : Type*) [Semiring S] [MulSemiringAction N S] -variable (S' : Type*) [Ring S'] [MulSemiringAction N S'] variable (T : Type*) [Semiring T] [MulSemiringAction P T] -variable {R S M' N'} -variable [AddMonoid M'] [DistribMulAction R M'] +variable {R S N'} variable [AddMonoid N'] [DistribMulAction S N'] variable {σ : R →* S} diff --git a/Mathlib/GroupTheory/NoncommPiCoprod.lean b/Mathlib/GroupTheory/NoncommPiCoprod.lean index 70be3543599df..d675c2b73a7f5 100644 --- a/Mathlib/GroupTheory/NoncommPiCoprod.lean +++ b/Mathlib/GroupTheory/NoncommPiCoprod.lean @@ -90,9 +90,6 @@ variable (ϕ : ∀ i : ι, N i →* M) -- We assume that the elements of different morphism commute variable (hcomm : Pairwise fun i j => ∀ x y, Commute (ϕ i x) (ϕ j y)) --- We use `f` and `g` to denote elements of `Π (i : ι), N i` -variable (f g : ∀ i : ι, N i) - namespace MonoidHom /-- The canonical homomorphism from a family of monoids. -/ @@ -168,9 +165,6 @@ variable {ι : Type*} variable {H : ι → Type*} [∀ i, Group (H i)] variable (ϕ : ∀ i : ι, H i →* G) --- We use `f` and `g` to denote elements of `Π (i : ι), H i` -variable (f g : ∀ i : ι, H i) - namespace MonoidHom -- The subgroup version of `MonoidHom.noncommPiCoprod_mrange` @[to_additive] @@ -251,9 +245,6 @@ namespace Subgroup variable {G : Type*} [Group G] variable {ι : Type*} {H : ι → Subgroup G} --- Elements of `Π (i : ι), H i` are called `f` and `g` here -variable (f g : ∀ i : ι, H i) - section CommutingSubgroups -- We assume that the elements of different subgroups commute diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index 7792a01f89afb..6609de9811ac4 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -414,9 +414,9 @@ end InfiniteSylow open Equiv Equiv.Perm Finset Function List QuotientGroup -universe u v w +universe u -variable {G : Type u} {α : Type v} {β : Type w} [Group G] +variable {G : Type u} [Group G] theorem QuotientGroup.card_preimage_mk (s : Subgroup G) (t : Set (G ⧸ s)) : Nat.card (QuotientGroup.mk ⁻¹' t) = Nat.card s * Nat.card t := by diff --git a/Mathlib/LinearAlgebra/Basis/Basic.lean b/Mathlib/LinearAlgebra/Basis/Basic.lean index 46ed8282bcd9a..3e294300107a0 100644 --- a/Mathlib/LinearAlgebra/Basis/Basic.lean +++ b/Mathlib/LinearAlgebra/Basis/Basic.lean @@ -24,18 +24,15 @@ universe u open Function Set Submodule Finsupp -variable {ι : Type*} {ι' : Type*} {R : Type*} {R₂ : Type*} {K : Type*} -variable {M : Type*} {M' M'' : Type*} {V : Type u} {V' : Type*} +variable {ι : Type*} {ι' : Type*} {R : Type*} {R₂ : Type*} {M : Type*} {M' : Type*} section Module -variable [Semiring R] -variable [AddCommMonoid M] [Module R M] [AddCommMonoid M'] [Module R M'] - +variable [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M'] [Module R M'] namespace Basis -variable (b b₁ : Basis ι R M) (i : ι) (c : R) (x : M) +variable (b : Basis ι R M) section Coord @@ -187,9 +184,9 @@ section Module open LinearMap variable {v : ι → M} -variable [Ring R] [CommRing R₂] [AddCommGroup M] [AddCommGroup M'] [AddCommGroup M''] -variable [Module R M] [Module R₂ M] [Module R M'] [Module R M''] -variable {c d : R} {x y : M} +variable [Ring R] [CommRing R₂] [AddCommGroup M] +variable [Module R M] [Module R₂ M] +variable {x y : M} variable (b : Basis ι R M) namespace Basis diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 14b5f34c2a2b3..172f5f1c0d87a 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -162,8 +162,7 @@ end FiniteDimensional namespace Module variable (K V) -variable [DivisionRing K] [AddCommGroup V] [Module K V] {V₂ : Type v'} [AddCommGroup V₂] - [Module K V₂] +variable [DivisionRing K] [AddCommGroup V] [Module K V] /-- In a finite-dimensional space, its dimension (seen as a cardinal) coincides with its `finrank`. This is a copy of `finrank_eq_rank _ _` which creates easier typeclass searches. -/ @@ -426,9 +425,6 @@ protected theorem finiteDimensional (f : V ≃ₗ[K] V₂) [FiniteDimensional K FiniteDimensional K V₂ := Module.Finite.equiv f -variable {R M M₂ : Type*} [Ring R] [AddCommGroup M] [AddCommGroup M₂] -variable [Module R M] [Module R M₂] - end LinearEquiv section @@ -442,8 +438,7 @@ instance finiteDimensional_finsupp {ι : Type*} [Finite ι] [FiniteDimensional K end namespace Submodule -variable [DivisionRing K] [AddCommGroup V] [Module K V] {V₂ : Type v'} [AddCommGroup V₂] - [Module K V₂] +variable [DivisionRing K] [AddCommGroup V] [Module K V] /-- If a submodule is contained in a finite-dimensional submodule with the same or smaller dimension, they are equal. -/ diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 411b231ffbadc..b4eab8d959db6 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -279,9 +279,8 @@ and prove basic property of this integral. open Finset -variable [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedSpace ℝ F] {p : ℝ≥0∞} {G F' : Type*} - [NormedAddCommGroup G] [NormedAddCommGroup F'] [NormedSpace ℝ F'] {m : MeasurableSpace α} - {μ : Measure α} +variable [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedSpace ℝ F] + {m : MeasurableSpace α} {μ : Measure α} /-- Bochner integral of simple functions whose codomain is a real `NormedSpace`. This is equal to `∑ x ∈ f.range, (μ (f ⁻¹' {x})).toReal • x` (see `integral_eq`). -/ @@ -423,7 +422,7 @@ namespace L1 open AEEqFun Lp.simpleFunc Lp -variable [NormedAddCommGroup E] [NormedAddCommGroup F] {m : MeasurableSpace α} {μ : Measure α} +variable [NormedAddCommGroup E] {m : MeasurableSpace α} {μ : Measure α} namespace SimpleFunc @@ -464,8 +463,7 @@ Define the Bochner integral on `α →₁ₛ[μ] E` by extension from the simple and prove basic properties of this integral. -/ -variable [NormedField 𝕜] [NormedSpace 𝕜 E] [NormedSpace ℝ E] [SMulCommClass ℝ 𝕜 E] {F' : Type*} - [NormedAddCommGroup F'] [NormedSpace ℝ F'] +variable [NormedField 𝕜] [NormedSpace 𝕜 E] [NormedSpace ℝ E] [SMulCommClass ℝ 𝕜 E] attribute [local instance] simpleFunc.normedSpace @@ -495,7 +493,6 @@ theorem norm_integral_le_norm (f : α →₁ₛ[μ] E) : ‖integral f‖ ≤ rw [integral, norm_eq_integral] exact (toSimpleFunc f).norm_integral_le_integral_norm (SimpleFunc.integrable f) -variable {E' : Type*} [NormedAddCommGroup E'] [NormedSpace ℝ E'] [NormedSpace 𝕜 E'] variable (α E μ 𝕜) /-- The Bochner integral over simple functions in L1 space as a continuous linear map. -/ @@ -575,7 +572,7 @@ open SimpleFunc local notation "Integral" => @integralCLM α E _ _ _ _ _ μ _ variable [NormedSpace ℝ E] [NontriviallyNormedField 𝕜] [NormedSpace 𝕜 E] [SMulCommClass ℝ 𝕜 E] - [NormedSpace ℝ F] [CompleteSpace E] + [CompleteSpace E] section IntegrationInL1 @@ -700,7 +697,7 @@ functions, and 0 otherwise; prove its basic properties. -/ variable [NormedAddCommGroup E] [hE : CompleteSpace E] [NontriviallyNormedField 𝕜] - [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpace F] + [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpace F] {G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] open Classical in @@ -730,8 +727,8 @@ section Properties open ContinuousLinearMap MeasureTheory.SimpleFunc -variable [NormedSpace ℝ E] [SMulCommClass ℝ 𝕜 E] -variable {f g : α → E} {m : MeasurableSpace α} {μ : Measure α} +variable [NormedSpace ℝ E] +variable {f : α → E} {m : MeasurableSpace α} {μ : Measure α} theorem integral_eq (f : α → E) (hf : Integrable f μ) : ∫ a, f a ∂μ = L1.integral (hf.toL1 f) := by simp [integral, hE, hf] @@ -1788,7 +1785,7 @@ end Properties section IntegralTrim -variable {H β γ : Type*} [NormedAddCommGroup H] {m m0 : MeasurableSpace β} {μ : Measure β} +variable {β γ : Type*} {m m0 : MeasurableSpace β} {μ : Measure β} /-- Simple function seen as simple function of a larger `MeasurableSpace`. -/ def SimpleFunc.toLargerSpace (hm : m ≤ m0) (f : @SimpleFunc β m γ) : SimpleFunc β γ := diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index eaee5e41d1330..fb6e31a954dff 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -26,7 +26,7 @@ variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} section -variable [Preorder α] [Preorder β] {s t : Set α} {a b : α} +variable [Preorder α] {s t : Set α} {a b : α} theorem mem_upperBounds : a ∈ upperBounds s ↔ ∀ x ∈ s, x ≤ a := Iff.rfl diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index dc192f4a528ec..2d47e7b97ad07 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -45,7 +45,7 @@ In lemma names, open Function OrderDual Set -variable {α β β₂ γ : Type*} {ι ι' : Sort*} {κ : ι → Sort*} {κ' : ι' → Sort*} +variable {α β γ : Type*} {ι ι' : Sort*} {κ : ι → Sort*} {κ' : ι' → Sort*} @[simp] lemma iSup_ulift {ι : Type*} [SupSet α] (f : ULift ι → α) : ⨆ i : ULift ι, f i = ⨆ i, f (.up i) := by simp [iSup]; congr with x; simp diff --git a/Mathlib/Order/GaloisConnection.lean b/Mathlib/Order/GaloisConnection.lean index d4ea55c673a62..7420bcaf290d1 100644 --- a/Mathlib/Order/GaloisConnection.lean +++ b/Mathlib/Order/GaloisConnection.lean @@ -50,8 +50,8 @@ open Function OrderDual Set universe u v w x -variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {κ : ι → Sort*} {a a₁ a₂ : α} - {b b₁ b₂ : β} +variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {κ : ι → Sort*} {a₁ a₂ : α} + {b₁ b₂ : β} /-- A Galois connection is a pair of functions `l` and `u` satisfying `l a ≤ b ↔ a ≤ u b`. They are special cases of adjoint functors in category theory, @@ -337,7 +337,7 @@ theorem gc_Ici_sInf [CompleteSemilatticeInf α] : GaloisConnection (toDual ∘ Ici : α → (Set α)ᵒᵈ) (sInf ∘ ofDual : (Set α)ᵒᵈ → α) := fun _ _ ↦ le_sInf_iff.symm -variable [CompleteLattice α] [CompleteLattice β] [CompleteLattice γ] {f : α → β → γ} {s : Set α} +variable [CompleteLattice α] [CompleteLattice β] [CompleteLattice γ] {s : Set α} {t : Set β} {l u : α → β → γ} {l₁ u₁ : β → γ → α} {l₂ u₂ : α → γ → β} theorem sSup_image2_eq_sSup_sSup (h₁ : ∀ b, GaloisConnection (swap l b) (u₁ b)) diff --git a/Mathlib/Order/PiLex.lean b/Mathlib/Order/PiLex.lean index bcad74a66ca1d..4fd435168435e 100644 --- a/Mathlib/Order/PiLex.lean +++ b/Mathlib/Order/PiLex.lean @@ -101,7 +101,7 @@ noncomputable instance [LinearOrder ι] [IsWellOrder ι (· < ·)] [∀ a, Linea section PartialOrder -variable [LinearOrder ι] [IsWellOrder ι (· < ·)] [∀ i, PartialOrder (β i)] {x y : ∀ i, β i} {i : ι} +variable [LinearOrder ι] [IsWellOrder ι (· < ·)] [∀ i, PartialOrder (β i)] {x : ∀ i, β i} {i : ι} {a : β i} open Function diff --git a/Mathlib/RingTheory/Localization/Submodule.lean b/Mathlib/RingTheory/Localization/Submodule.lean index 64d723f0fd50d..1e3d73e6d705b 100644 --- a/Mathlib/RingTheory/Localization/Submodule.lean +++ b/Mathlib/RingTheory/Localization/Submodule.lean @@ -21,7 +21,7 @@ commutative ring, field of fractions variable {R : Type*} [CommRing R] (M : Submonoid R) (S : Type*) [CommRing S] -variable [Algebra R S] {P : Type*} [CommRing P] +variable [Algebra R S] namespace IsLocalization @@ -72,9 +72,6 @@ theorem coeSubmodule_span_singleton (x : R) : coeSubmodule S (Ideal.span {x}) = Submodule.span R {(algebraMap R S) x} := by rw [coeSubmodule_span, Set.image_singleton] -variable {g : R →+* P} -variable {T : Submonoid P} (hy : M ≤ T.comap g) {Q : Type*} [CommRing Q] -variable [Algebra P Q] [IsLocalization T Q] variable [IsLocalization M S] include M in @@ -159,11 +156,11 @@ namespace IsFractionRing open IsLocalization -variable {A K : Type*} [CommRing A] +variable {K : Type*} section CommRing -variable [CommRing K] [Algebra R K] [IsFractionRing R K] [Algebra A K] [IsFractionRing A K] +variable [CommRing K] [Algebra R K] [IsFractionRing R K] @[simp, mono] theorem coeSubmodule_le_coeSubmodule {I J : Ideal R} : diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index df022af891410..637ef9b981a96 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -130,9 +130,8 @@ end Submodule section closure -variable {R R' : Type u} {M M' : Type v} [Semiring R] [Ring R'] - [TopologicalSpace M] [AddCommMonoid M] [TopologicalSpace M'] [AddCommGroup M'] [Module R M] - [ContinuousConstSMul R M] [Module R' M'] [ContinuousConstSMul R' M'] +variable {R : Type u} {M : Type v} [Semiring R] [TopologicalSpace M] [AddCommMonoid M] [Module R M] + [ContinuousConstSMul R M] theorem Submodule.mapsTo_smul_closure (s : Submodule R M) (c : R) : Set.MapsTo (c • ·) (closure s : Set M) (closure s) := @@ -1418,7 +1417,7 @@ variable {R R₂ R₃ S S₃ : Type*} [Semiring R] [Semiring R₂] [Semiring R [SMulCommClass R₃ S₃ M₃] [ContinuousConstSMul S₃ M₃] [Module S N₂] [ContinuousConstSMul S N₂] [SMulCommClass R S N₂] [Module S N₃] [SMulCommClass R S N₃] [ContinuousConstSMul S N₃] {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] (c : S) - (h : M₂ →SL[σ₂₃] M₃) (f g : M →SL[σ₁₂] M₂) (x y z : M) + (h : M₂ →SL[σ₂₃] M₃) (f : M →SL[σ₁₂] M₂) /-- `ContinuousLinearMap.prod` as an `Equiv`. -/ @[simps apply] @@ -1505,7 +1504,7 @@ section CommRing variable {R : Type*} [CommRing R] {M : Type*} [TopologicalSpace M] [AddCommGroup M] {M₂ : Type*} [TopologicalSpace M₂] [AddCommGroup M₂] {M₃ : Type*} [TopologicalSpace M₃] [AddCommGroup M₃] - [Module R M] [Module R M₂] [Module R M₃] [ContinuousConstSMul R M₃] + [Module R M] [Module R M₂] [Module R M₃] variable [TopologicalAddGroup M₂] [ContinuousConstSMul R M₂] @@ -1596,10 +1595,10 @@ variable {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} [Semiring R₁] [Semiring {σ₂₃ : R₂ →+* R₃} {σ₃₂ : R₃ →+* R₂} [RingHomInvPair σ₂₃ σ₃₂] [RingHomInvPair σ₃₂ σ₂₃] {σ₁₃ : R₁ →+* R₃} {σ₃₁ : R₃ →+* R₁} [RingHomInvPair σ₁₃ σ₃₁] [RingHomInvPair σ₃₁ σ₁₃] [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [RingHomCompTriple σ₃₂ σ₂₁ σ₃₁] {M₁ : Type*} - [TopologicalSpace M₁] [AddCommMonoid M₁] {M'₁ : Type*} [TopologicalSpace M'₁] [AddCommMonoid M'₁] + [TopologicalSpace M₁] [AddCommMonoid M₁] {M₂ : Type*} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type*} [TopologicalSpace M₃] [AddCommMonoid M₃] {M₄ : Type*} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] - [Module R₁ M'₁] [Module R₂ M₂] [Module R₃ M₃] + [Module R₂ M₂] [Module R₃ M₃] /-- A continuous linear equivalence induces a continuous linear map. -/ @[coe] @@ -2302,7 +2301,6 @@ end ContinuousLinearMap namespace Submodule variable {R : Type*} [Ring R] {M : Type*} [TopologicalSpace M] [AddCommGroup M] [Module R M] - {M₂ : Type*} [TopologicalSpace M₂] [AddCommGroup M₂] [Module R M₂] open ContinuousLinearMap From c744def559485980adb40b991247176a127345b7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 17 Oct 2024 16:19:09 +0000 Subject: [PATCH 026/131] feat(Algebra/Module/Submodule): lemmas about `domRestrict` (#17806) From PFR --- Mathlib/Algebra/Module/Submodule/Ker.lean | 10 +++++++--- Mathlib/Algebra/Module/Submodule/Map.lean | 5 ++++- Mathlib/Algebra/Module/Submodule/Range.lean | 3 +++ 3 files changed, 14 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Module/Submodule/Ker.lean b/Mathlib/Algebra/Module/Submodule/Ker.lean index 23ce2537e64cc..b291dd847675d 100644 --- a/Mathlib/Algebra/Module/Submodule/Ker.lean +++ b/Mathlib/Algebra/Module/Submodule/Ker.lean @@ -109,10 +109,13 @@ theorem le_ker_iff_map [RingHomSurjective τ₁₂] {f : F} {p : Submodule R M} theorem ker_codRestrict {τ₂₁ : R₂ →+* R} (p : Submodule R M) (f : M₂ →ₛₗ[τ₂₁] M) (hf) : ker (codRestrict p f hf) = ker f := by rw [ker, comap_codRestrict, Submodule.map_bot]; rfl +lemma ker_domRestrict [AddCommMonoid M₁] [Module R M₁] (p : Submodule R M) (f : M →ₗ[R] M₁) : + ker (domRestrict f p) = (ker f).comap p.subtype := ker_comp .. + theorem ker_restrict [AddCommMonoid M₁] [Module R M₁] {p : Submodule R M} {q : Submodule R M₁} {f : M →ₗ[R] M₁} (hf : ∀ x : M, x ∈ p → f x ∈ q) : - ker (f.restrict hf) = LinearMap.ker (f.domRestrict p) := by - rw [restrict_eq_codRestrict_domRestrict, ker_codRestrict] + ker (f.restrict hf) = (ker f).comap p.subtype := by + rw [restrict_eq_codRestrict_domRestrict, ker_codRestrict, ker_domRestrict] @[simp] theorem ker_zero : ker (0 : M →ₛₗ[τ₁₂] M₂) = ⊤ := @@ -198,7 +201,8 @@ theorem ker_eq_bot {f : M →ₛₗ[τ₁₂] M₂} : ker f = ⊥ ↔ Injective @[simp] theorem injective_restrict_iff_disjoint {p : Submodule R M} {f : M →ₗ[R] M} (hf : ∀ x ∈ p, f x ∈ p) : Injective (f.restrict hf) ↔ Disjoint p (ker f) := by - rw [← ker_eq_bot, ker_restrict hf, ker_eq_bot, injective_domRestrict_iff, disjoint_iff] + rw [← ker_eq_bot, ker_restrict hf, ← ker_domRestrict, ker_eq_bot, injective_domRestrict_iff, + disjoint_iff] end Ring diff --git a/Mathlib/Algebra/Module/Submodule/Map.lean b/Mathlib/Algebra/Module/Submodule/Map.lean index 8e6b50ba4e465..84ea00362c6d2 100644 --- a/Mathlib/Algebra/Module/Submodule/Map.lean +++ b/Mathlib/Algebra/Module/Submodule/Map.lean @@ -290,9 +290,12 @@ theorem map_iInf_comap_of_surjective {ι : Sort*} (S : ι → Submodule R₂ M (⨅ i, (S i).comap f).map f = iInf S := (giMapComap hf).l_iInf_u _ -theorem comap_le_comap_iff_of_surjective (p q : Submodule R₂ M₂) : p.comap f ≤ q.comap f ↔ p ≤ q := +theorem comap_le_comap_iff_of_surjective {p q : Submodule R₂ M₂} : p.comap f ≤ q.comap f ↔ p ≤ q := (giMapComap hf).u_le_u_iff +lemma comap_lt_comap_iff_of_surjective {p q : Submodule R₂ M₂} : p.comap f < q.comap f ↔ p < q := by + apply lt_iff_lt_of_le_iff_le' <;> exact comap_le_comap_iff_of_surjective hf + theorem comap_strictMono_of_surjective : StrictMono (comap f) := (giMapComap hf).strictMono_u diff --git a/Mathlib/Algebra/Module/Submodule/Range.lean b/Mathlib/Algebra/Module/Submodule/Range.lean index 933cc427c94e9..1053c1cf8c54c 100644 --- a/Mathlib/Algebra/Module/Submodule/Range.lean +++ b/Mathlib/Algebra/Module/Submodule/Range.lean @@ -100,6 +100,9 @@ theorem range_neg {R : Type*} {R₂ : Type*} {M : Type*} {M₂ : Type*} [Semirin change range ((-LinearMap.id : M₂ →ₗ[R₂] M₂).comp f) = _ rw [range_comp, Submodule.map_neg, Submodule.map_id] +@[simp] lemma range_domRestrict [Module R M₂] (K : Submodule R M) (f : M →ₗ[R] M₂) : + range (domRestrict f K) = K.map f := by ext; simp + lemma range_domRestrict_le_range [RingHomSurjective τ₁₂] (f : M →ₛₗ[τ₁₂] M₂) (S : Submodule R M) : LinearMap.range (f.domRestrict S) ≤ LinearMap.range f := by rintro x ⟨⟨y, hy⟩, rfl⟩ From 9d883e2896b1b16d8e646c336c647da4fc1cacd5 Mon Sep 17 00:00:00 2001 From: FMLJohn <417963616@qq.com> Date: Thu, 17 Oct 2024 17:56:40 +0000 Subject: [PATCH 027/131] feat(RingTheory/PowerSeries/WellKnown): changed `PowerSeries.invOneSubPow` (#17360) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The definition of `PowerSeries.invOneSubPow` has been changed into the following: Given a commutative ring `S` and a number `d : ℕ`, `PowerSeries.invOneSubPow S d` is the multiplicative inverse of `(1 - X) ^ d` in `S⟦X⟧ˣ`. When `d` is `0`, `PowerSeries.invOneSubPow S d` will just be `1`. When `d` is positive, `PowerSeries.invOneSubPow S d` will be `∑ n, Nat.choose (d - 1 + n) (d - 1)`. Co-authored-by: FMLJohn <128468682+FMLJohn@users.noreply.github.com> --- Mathlib/RingTheory/PowerSeries/WellKnown.lean | 82 +++++++++++++------ 1 file changed, 56 insertions(+), 26 deletions(-) diff --git a/Mathlib/RingTheory/PowerSeries/WellKnown.lean b/Mathlib/RingTheory/PowerSeries/WellKnown.lean index 8b3ce3c34a0ce..292659bb8862f 100644 --- a/Mathlib/RingTheory/PowerSeries/WellKnown.lean +++ b/Mathlib/RingTheory/PowerSeries/WellKnown.lean @@ -18,8 +18,9 @@ In this file we define the following power series: It is given by `∑ n, x ^ n /ₚ u ^ (n + 1)`. * `PowerSeries.invOneSubPow`: given a commutative ring `S` and a number `d : ℕ`, - `PowerSeries.invOneSubPow d : S⟦X⟧ˣ` is the power series `∑ n, Nat.choose (d + n) d` - whose multiplicative inverse is `(1 - X) ^ (d + 1)`. + `PowerSeries.invOneSubPow S d` is the multiplicative inverse of `(1 - X) ^ d` in `S⟦X⟧ˣ`. + When `d` is `0`, `PowerSeries.invOneSubPow S d` will just be `1`. When `d` is positive, + `PowerSeries.invOneSubPow S d` will be `∑ n, Nat.choose (d - 1 + n) (d - 1)`. * `PowerSeries.sin`, `PowerSeries.cos`, `PowerSeries.exp` : power series for sin, cosine, and exponential functions. @@ -64,7 +65,7 @@ end Ring section invOneSubPow -variable {S : Type*} [CommRing S] (d : ℕ) +variable (S : Type*) [CommRing S] (d : ℕ) /-- (1 + X + X^2 + ...) * (1 - X) = 1. @@ -97,39 +98,68 @@ theorem mk_one_pow_eq_mk_choose_add : add_right_comm] /-- -The power series `mk fun n => Nat.choose (d + n) d`, whose multiplicative inverse is -`(1 - X) ^ (d + 1)`. +Given a natural number `d : ℕ` and a commutative ring `S`, `PowerSeries.invOneSubPow S d` is the +multiplicative inverse of `(1 - X) ^ d` in `S⟦X⟧ˣ`. When `d` is `0`, `PowerSeries.invOneSubPow S d` +will just be `1`. When `d` is positive, `PowerSeries.invOneSubPow S d` will be the power series +`mk fun n => Nat.choose (d - 1 + n) (d - 1)`. -/ -noncomputable def invOneSubPow : S⟦X⟧ˣ where - val := mk fun n => Nat.choose (d + n) d - inv := (1 - X) ^ (d + 1) - val_inv := by - rw [← mk_one_pow_eq_mk_choose_add, ← mul_pow, mk_one_mul_one_sub_eq_one, one_pow] - inv_val := by - rw [← mk_one_pow_eq_mk_choose_add, ← mul_pow, mul_comm, mk_one_mul_one_sub_eq_one, one_pow] - -theorem invOneSubPow_val_eq_mk_choose_add : - (invOneSubPow d).val = (mk fun n => Nat.choose (d + n) d : S⟦X⟧) := rfl - -theorem invOneSubPow_val_zero_eq_invUnitSub_one : - (invOneSubPow 0).val = invUnitsSub (1 : Sˣ) := by +noncomputable def invOneSubPow : ℕ → S⟦X⟧ˣ + | 0 => 1 + | d + 1 => { + val := mk fun n => Nat.choose (d + n) d + inv := (1 - X) ^ (d + 1) + val_inv := by + rw [← mk_one_pow_eq_mk_choose_add, ← mul_pow, mk_one_mul_one_sub_eq_one, one_pow] + inv_val := by + rw [← mk_one_pow_eq_mk_choose_add, ← mul_pow, mul_comm, mk_one_mul_one_sub_eq_one, one_pow] + } + +theorem invOneSubPow_zero : invOneSubPow S 0 = 1 := by + delta invOneSubPow + simp only [Units.val_one] + +theorem invOneSubPow_val_eq_mk_sub_one_add_choose_of_pos (h : 0 < d) : + (invOneSubPow S d).val = (mk fun n => Nat.choose (d - 1 + n) (d - 1) : S⟦X⟧) := by + rw [← Nat.sub_one_add_one_eq_of_pos h, invOneSubPow, add_tsub_cancel_right] + +theorem invOneSubPow_val_succ_eq_mk_add_choose : + (invOneSubPow S (d + 1)).val = (mk fun n => Nat.choose (d + n) d : S⟦X⟧) := rfl + +theorem invOneSubPow_val_one_eq_invUnitSub_one : + (invOneSubPow S 1).val = invUnitsSub (1 : Sˣ) := by simp [invOneSubPow, invUnitsSub] /-- The theorem `PowerSeries.mk_one_mul_one_sub_eq_one` implies that `1 - X` is a unit in `S⟦X⟧` whose inverse is the power series `1 + X + X^2 + ...`. This theorem states that for any `d : ℕ`, -`PowerSeries.invOneSubPow d` is equal to `(1 - X)⁻¹ ^ (d + 1)`. +`PowerSeries.invOneSubPow S d` is equal to `(1 - X)⁻¹ ^ d`. -/ theorem invOneSubPow_eq_inv_one_sub_pow : - invOneSubPow d = (Units.mkOfMulEqOne (1 - X) (mk 1 : S⟦X⟧) - <| Eq.trans (mul_comm _ _) mk_one_mul_one_sub_eq_one)⁻¹ ^ (d + 1) := by - rw [inv_pow] - exact (DivisionMonoid.inv_eq_of_mul _ (invOneSubPow d) <| by - rw [← Units.val_eq_one, Units.val_mul, Units.val_pow_eq_pow_val] - exact (invOneSubPow d).inv_val).symm + invOneSubPow S d = + (Units.mkOfMulEqOne (1 - X) (mk 1 : S⟦X⟧) <| + Eq.trans (mul_comm _ _) (mk_one_mul_one_sub_eq_one S))⁻¹ ^ d := by + induction d with + | zero => exact Eq.symm <| pow_zero _ + | succ d _ => + rw [inv_pow] + exact (DivisionMonoid.inv_eq_of_mul _ (invOneSubPow S (d + 1)) <| by + rw [← Units.val_eq_one, Units.val_mul, Units.val_pow_eq_pow_val] + exact (invOneSubPow S (d + 1)).inv_val).symm theorem invOneSubPow_inv_eq_one_sub_pow : - (invOneSubPow d).inv = (1 - X : S⟦X⟧) ^ (d + 1) := rfl + (invOneSubPow S d).inv = (1 - X : S⟦X⟧) ^ d := by + induction d with + | zero => exact Eq.symm <| pow_zero _ + | succ d => rfl + +theorem invOneSubPow_inv_eq_one_of_eq_zero (h : d = 0) : + (invOneSubPow S d).inv = 1 := by + delta invOneSubPow + simp only [h, Units.inv_eq_val_inv, inv_one, Units.val_one] + +theorem mk_add_choose_mul_one_sub_pow_eq_one : + (mk fun n ↦ Nat.choose (d + n) d : S⟦X⟧) * ((1 - X) ^ (d + 1)) = 1 := + (invOneSubPow S (d + 1)).val_inv end invOneSubPow From e759b3b1e909c1ebe13a364745f063152a61a3aa Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 17 Oct 2024 19:04:09 +0000 Subject: [PATCH 028/131] feat: UniqueFactorizationMonoid.factors_rel_of_associated (#17260) --- Mathlib/RingTheory/Radical.lean | 12 ++++++---- .../RingTheory/UniqueFactorizationDomain.lean | 23 ++++++++++++++----- 2 files changed, 25 insertions(+), 10 deletions(-) diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean index c86cee4872a46..0c28c161cdf4f 100644 --- a/Mathlib/RingTheory/Radical.lean +++ b/Mathlib/RingTheory/Radical.lean @@ -46,6 +46,12 @@ variable {M : Type*} [CancelCommMonoidWithZero M] [NormalizationMonoid M] def primeFactors (a : M) : Finset M := (normalizedFactors a).toFinset +theorem _root_.Associated.primeFactors_eq {a b : M} (h : Associated a b) : + primeFactors a = primeFactors b := by + unfold primeFactors + rw [h.normalizedFactors_eq] + + /-- Radical of an element `a` in a unique factorization monoid is the product of the prime factors of `a`. @@ -62,10 +68,8 @@ theorem radical_one_eq : radical (1 : M) = 1 := by rw [radical, primeFactors, normalizedFactors_one, Multiset.toFinset_zero, Finset.prod_empty] theorem radical_eq_of_associated {a b : M} (h : Associated a b) : radical a = radical b := by - rcases iff_iff_and_or_not_and_not.mp h.eq_zero_iff with (⟨rfl, rfl⟩ | ⟨ha, hb⟩) - · rfl - · simp_rw [radical, primeFactors] - rw [(associated_iff_normalizedFactors_eq_normalizedFactors ha hb).mp h] + unfold radical + rw [h.primeFactors_eq] theorem radical_unit_eq_one {a : M} (h : IsUnit a) : radical a = 1 := (radical_eq_of_associated (associated_one_iff_isUnit.mpr h)).trans radical_one_eq diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index e144f86f1433b..dde98af42a008 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -554,6 +554,13 @@ theorem factors_pow_count_prod [DecidableEq α] {x : α} (hx : x ≠ 0) : _ = prod (factors x) := by rw [toFinset_sum_count_nsmul_eq (factors x)] _ ~ᵤ x := factors_prod hx +theorem factors_rel_of_associated {a b : α} (h : Associated a b) : + Multiset.Rel Associated (factors a) (factors b) := by + rcases iff_iff_and_or_not_and_not.mp h.eq_zero_iff with (⟨rfl, rfl⟩ | ⟨ha, hb⟩) + · simp + · refine factors_unique irreducible_of_factor irreducible_of_factor ?_ + exact ((factors_prod ha).trans h).trans (factors_prod hb).symm + end UniqueFactorizationMonoid namespace UniqueFactorizationMonoid @@ -721,13 +728,17 @@ theorem dvd_iff_normalizedFactors_le_normalizedFactors {x y : α} (hx : x ≠ 0) (normalizedFactors_prod hy).dvd_iff_dvd_right] apply Multiset.prod_dvd_prod_of_le +theorem _root_.Associated.normalizedFactors_eq {a b : α} (h : Associated a b) : + normalizedFactors a = normalizedFactors b := by + unfold normalizedFactors + have h' : ⇑(normalize (α := α)) = Associates.out ∘ Associates.mk := funext Associates.out_mk + rw [h', ← Multiset.map_map, ← Multiset.map_map, + Associates.rel_associated_iff_map_eq_map.mp (factors_rel_of_associated h)] + theorem associated_iff_normalizedFactors_eq_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - x ~ᵤ y ↔ normalizedFactors x = normalizedFactors y := by - refine - ⟨fun h => ?_, fun h => - (normalizedFactors_prod hx).symm.trans (_root_.trans (by rw [h]) (normalizedFactors_prod hy))⟩ - apply le_antisymm <;> rw [← dvd_iff_normalizedFactors_le_normalizedFactors] - all_goals simp [*, h.dvd, h.symm.dvd] + x ~ᵤ y ↔ normalizedFactors x = normalizedFactors y := + ⟨Associated.normalizedFactors_eq, fun h => + (normalizedFactors_prod hx).symm.trans (_root_.trans (by rw [h]) (normalizedFactors_prod hy))⟩ theorem normalizedFactors_of_irreducible_pow {p : α} (hp : Irreducible p) (k : ℕ) : normalizedFactors (p ^ k) = Multiset.replicate k (normalize p) := by From 4092287b49c452c0efcf2dee7d9d5e5db095768d Mon Sep 17 00:00:00 2001 From: Etienne Date: Thu, 17 Oct 2024 19:04:10 +0000 Subject: [PATCH 029/131] feat: sufficient condition for a set of linear forms to span the dual space (#17548) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Consider a reflexive module and a set `s` of linear forms. If for any `z ≠ 0` there exists `f ∈ s` such that `f z ≠ 0`, then `s` spans the whole dual space. --- Mathlib/LinearAlgebra/Dual.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 0dfa5a4d643f0..23b6f1f84347a 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -690,6 +690,20 @@ 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'⟩ +/-- Consider a reflexive module and a set `s` of linear forms. If for any `z ≠ 0` there exists +`f ∈ s` such that `f z ≠ 0`, then `s` spans the whole dual space. -/ +theorem span_eq_top_of_ne_zero [IsReflexive R M] + {s : Set (M →ₗ[R] R)} [Free R ((M →ₗ[R] R) ⧸ (span R s))] + (h : ∀ z ≠ 0, ∃ f ∈ s, f z ≠ 0) : span R s = ⊤ := by + by_contra! hn + obtain ⟨φ, φne, hφ⟩ := exists_dual_map_eq_bot_of_lt_top hn.lt_top inferInstance + let φs := (evalEquiv R M).symm φ + have this f (hf : f ∈ s) : f φs = 0 := by + rw [← mem_bot R, ← hφ, mem_map] + exact ⟨f, subset_span hf, (apply_evalEquiv_symm_apply R M f φ).symm⟩ + obtain ⟨x, xs, hx⟩ := h φs (by simp [φne, φs]) + exact hx <| this x xs + variable {ι 𝕜 E : Type*} [Field 𝕜] [AddCommGroup E] [Module 𝕜 E] open LinearMap Set FiniteDimensional From b2dddd4c7b46f8feff35592d99c9e4ce57a3832f Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 17 Oct 2024 19:56:26 +0000 Subject: [PATCH 030/131] feat: sum of iterated derivatives (#16886) This introduces `Polynomial.sumIDeriv`, the sum of the iterated derivatives of a polynomial, to be used in particular in the proof of the Lindemann-Weierstrass theorem (see #6718). Included are: * `Polynomial.sumIDeriv_apply`, `Polynomial.sumIDeriv_apply_of_lt`, `Polynomial.sumIDeriv_apply_of_le`: `Polynomial.sumIDeriv` expressed as a sum * `Polynomial.sumIDeriv_C`, `Polynomial.sumIDeriv_X`: `Polynomial.sumIDeriv` applied to simple polynomials * `Polynomial.sumIDeriv_map`: `Polynomial.sumIDeriv` commutes with `Polynomial.map` * `Polynomial.sumIDeriv_derivative`: `Polynomial.sumIDeriv` commutes with `Polynomial.derivative` * `Polynomial.sumIDeriv_eq_self_add`: `sumIDeriv p = p + sumIDeriv (derivative p)` * `Polynomial.exists_iterate_derivative_eq_factorial_smul`: the `k`'th iterated derivative of a polynomial has a common factor `k!` * `Polynomial.aeval_iterate_derivative_of_lt`, `Polynomial.aeval_iterate_derivative_self`, `Polynomial.aeval_iterate_derivative_of_ge`: applying `Polynomial.aeval` to iterated derivatives * `Polynomial.aeval_sumIDeriv`, `Polynomial.aeval_sumIDeriv_of_pos`: applying `Polynomial.aeval` to `Polynomial.sumIDeriv` Co-authored-by: FR-vdash-bot Co-authored-by: FR --- Mathlib.lean | 1 + Mathlib/Algebra/Polynomial/Derivative.lean | 64 +++++ .../Polynomial/SumIteratedDerivative.lean | 245 ++++++++++++++++++ 3 files changed, 310 insertions(+) create mode 100644 Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean diff --git a/Mathlib.lean b/Mathlib.lean index b43069a00a054..91da140ade63c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -744,6 +744,7 @@ import Mathlib.Algebra.Polynomial.Roots import Mathlib.Algebra.Polynomial.Smeval import Mathlib.Algebra.Polynomial.SpecificDegree import Mathlib.Algebra.Polynomial.Splits +import Mathlib.Algebra.Polynomial.SumIteratedDerivative import Mathlib.Algebra.Polynomial.Taylor import Mathlib.Algebra.Polynomial.UnitTrinomial import Mathlib.Algebra.QuadraticDiscriminant diff --git a/Mathlib/Algebra/Polynomial/Derivative.lean b/Mathlib/Algebra/Polynomial/Derivative.lean index 9f22f0fa03689..9108354da435c 100644 --- a/Mathlib/Algebra/Polynomial/Derivative.lean +++ b/Mathlib/Algebra/Polynomial/Derivative.lean @@ -12,6 +12,7 @@ import Mathlib.GroupTheory.GroupAction.Ring ## Main definitions * `Polynomial.derivative`: The formal derivative of polynomials, expressed as a linear map. + * `Polynomial.derivativeFinsupp`: Iterated derivatives as a finite support function. -/ @@ -22,6 +23,8 @@ open Finset open Polynomial +open scoped Nat + namespace Polynomial universe u v w y z @@ -330,6 +333,21 @@ theorem coeff_iterate_derivative {k} (p : R[X]) (m : ℕ) : _ = Nat.descFactorial (m + k.succ) k.succ • p.coeff (m + k.succ) := by rw [Nat.succ_add_eq_add_succ] +theorem iterate_derivative_eq_sum (p : R[X]) (k : ℕ) : + derivative^[k] p = + ∑ x ∈ (derivative^[k] p).support, C ((x + k).descFactorial k • p.coeff (x + k)) * X ^ x := by + conv_lhs => rw [(derivative^[k] p).as_sum_support_C_mul_X_pow] + refine sum_congr rfl fun i _ ↦ ?_ + rw [coeff_iterate_derivative, Nat.descFactorial_eq_factorial_mul_choose] + +theorem iterate_derivative_eq_factorial_smul_sum (p : R[X]) (k : ℕ) : + derivative^[k] p = k ! • + ∑ x ∈ (derivative^[k] p).support, C ((x + k).choose k • p.coeff (x + k)) * X ^ x := by + conv_lhs => rw [iterate_derivative_eq_sum] + rw [smul_sum] + refine sum_congr rfl fun i _ ↦ ?_ + rw [← smul_mul_assoc, smul_C, smul_smul, Nat.descFactorial_eq_factorial_mul_choose] + theorem iterate_derivative_mul {n} (p q : R[X]) : derivative^[n] (p * q) = ∑ k ∈ range n.succ, (n.choose k • (derivative^[n - k] p * derivative^[k] q)) := by @@ -374,6 +392,52 @@ theorem iterate_derivative_mul {n} (p q : R[X]) : omega · rw [Nat.choose_zero_right, tsub_zero] +/-- +Iterated derivatives as a finite support function. +-/ +@[simps! apply_toFun] +noncomputable def derivativeFinsupp : R[X] →ₗ[R] ℕ →₀ R[X] where + toFun p := .onFinset (range (p.natDegree + 1)) (derivative^[·] p) fun i ↦ by + contrapose; simp_all [iterate_derivative_eq_zero, Nat.succ_le] + map_add' _ _ := by ext; simp + map_smul' _ _ := by ext; simp + +@[simp] +theorem support_derivativeFinsupp_subset_range {p : R[X]} {n : ℕ} (h : p.natDegree < n) : + (derivativeFinsupp p).support ⊆ range n := by + dsimp [derivativeFinsupp] + exact Finsupp.support_onFinset_subset.trans (Finset.range_subset.mpr h) + +@[simp] +theorem derivativeFinsupp_C (r : R) : derivativeFinsupp (C r : R[X]) = .single 0 (C r) := by + ext i : 1 + match i with + | 0 => simp + | i + 1 => simp + +@[simp] +theorem derivativeFinsupp_one : derivativeFinsupp (1 : R[X]) = .single 0 1 := by + simpa using derivativeFinsupp_C (1 : R) + +@[simp] +theorem derivativeFinsupp_X : derivativeFinsupp (X : R[X]) = .single 0 X + .single 1 1 := by + ext i : 1 + match i with + | 0 => simp + | 1 => simp + | (n + 2) => simp + +theorem derivativeFinsupp_map [Semiring S] (p : R[X]) (f : R →+* S) : + derivativeFinsupp (p.map f) = (derivativeFinsupp p).mapRange (·.map f) (by simp) := by + ext i : 1 + simp + +theorem derivativeFinsupp_derivative (p : R[X]) : + derivativeFinsupp (derivative p) = + (derivativeFinsupp p).comapDomain Nat.succ Nat.succ_injective.injOn := by + ext i : 1 + simp + end Semiring section CommSemiring diff --git a/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean b/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean new file mode 100644 index 0000000000000..b8b64c16ec9bf --- /dev/null +++ b/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean @@ -0,0 +1,245 @@ +/- +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.Polynomial.AlgebraMap +import Mathlib.Algebra.Polynomial.BigOperators +import Mathlib.Algebra.Polynomial.Degree.Lemmas +import Mathlib.Algebra.Polynomial.Derivative +import Mathlib.Algebra.Polynomial.RingDivision + +/-! +# Sum of iterated derivatives + +This file introduces `Polynomial.sumIDeriv`, the sum of the iterated derivatives of a polynomial, +as a linear map. This is used in particular in the proof of the Lindemann-Weierstrass theorem +(see #6718). + +## Main results + +* `Polynomial.sumIDeriv`: Sum of iterated derivatives of a polynomial, as a linear map +* `Polynomial.sumIDeriv_apply`, `Polynomial.sumIDeriv_apply_of_lt`, + `Polynomial.sumIDeriv_apply_of_le`: `Polynomial.sumIDeriv` expressed as a sum +* `Polynomial.sumIDeriv_C`, `Polynomial.sumIDeriv_X`: `Polynomial.sumIDeriv` applied to simple + polynomials +* `Polynomial.sumIDeriv_map`: `Polynomial.sumIDeriv` commutes with `Polynomial.map` +* `Polynomial.sumIDeriv_derivative`: `Polynomial.sumIDeriv` commutes with `Polynomial.derivative` +* `Polynomial.sumIDeriv_eq_self_add`: `sumIDeriv p = p + sumIDeriv (derivative p)` +* `Polynomial.exists_iterate_derivative_eq_factorial_smul`: the `k`'th iterated derivative of a + polynomial has a common factor `k!` +* `Polynomial.aeval_iterate_derivative_of_lt`, `Polynomial.aeval_iterate_derivative_self`, + `Polynomial.aeval_iterate_derivative_of_ge`: applying `Polynomial.aeval` to iterated derivatives +* `Polynomial.aeval_sumIDeriv`, `Polynomial.aeval_sumIDeriv_of_pos`: applying `Polynomial.aeval` to + `Polynomial.sumIDeriv` + +-/ + +open Finset +open scoped Nat + +namespace Polynomial + +variable {R S : Type*} + +section Semiring + +variable [Semiring R] [Semiring S] + +/-- +Sum of iterated derivatives of a polynomial, as a linear map + +This definition does not allow different weights for the derivatives. It is likely that it could be +extended to allow them, but this was not needed for the initial use case (the integration by part +of the integral $I_i$ in the +[Lindemann-Weierstrass](https://en.wikipedia.org/wiki/Lindemann%E2%80%93Weierstrass_theorem) +theorem). +-/ +noncomputable def sumIDeriv : R[X] →ₗ[R] R[X] := + Finsupp.lsum ℕ (fun _ ↦ LinearMap.id) ∘ₗ derivativeFinsupp + +theorem sumIDeriv_apply (p : R[X]) : + sumIDeriv p = ∑ i ∈ range (p.natDegree + 1), derivative^[i] p := by + dsimp [sumIDeriv] + exact Finsupp.sum_of_support_subset _ (by simp) _ (by simp) + +theorem sumIDeriv_apply_of_lt {p : R[X]} {n : ℕ} (hn : p.natDegree < n) : + sumIDeriv p = ∑ i ∈ range n, derivative^[i] p := by + dsimp [sumIDeriv] + exact Finsupp.sum_of_support_subset _ (by simp [hn]) _ (by simp) + +theorem sumIDeriv_apply_of_le {p : R[X]} {n : ℕ} (hn : p.natDegree ≤ n) : + sumIDeriv p = ∑ i ∈ range (n + 1), derivative^[i] p := by + dsimp [sumIDeriv] + exact Finsupp.sum_of_support_subset _ (by simp [Nat.lt_succ, hn]) _ (by simp) + +@[simp] +theorem sumIDeriv_C (a : R) : sumIDeriv (C a) = C a := by + rw [sumIDeriv_apply, natDegree_C, zero_add, sum_range_one, Function.iterate_zero_apply] + +@[simp] +theorem sumIDeriv_X : sumIDeriv X = X + C 1 := by + rw [sumIDeriv_apply, natDegree_X, sum_range_succ, sum_range_one, Function.iterate_zero_apply, + Function.iterate_one, derivative_X, eq_natCast, Nat.cast_one] + +@[simp] +theorem sumIDeriv_map (p : R[X]) (f : R →+* S) : + sumIDeriv (p.map f) = (sumIDeriv p).map f := by + let n := max (p.map f).natDegree p.natDegree + rw [sumIDeriv_apply_of_le (le_max_left _ _ : _ ≤ n)] + rw [sumIDeriv_apply_of_le (le_max_right _ _ : _ ≤ n)] + simp_rw [Polynomial.map_sum, iterate_derivative_map p f] + +theorem sumIDeriv_derivative (p : R[X]) : sumIDeriv (derivative p) = derivative (sumIDeriv p) := by + rw [sumIDeriv_apply_of_le ((natDegree_derivative_le p).trans tsub_le_self), sumIDeriv_apply, + derivative_sum] + simp_rw [← Function.iterate_succ_apply, Function.iterate_succ_apply'] + +theorem sumIDeriv_eq_self_add (p : R[X]) : sumIDeriv p = p + sumIDeriv (derivative p) := by + rw [sumIDeriv_derivative, sumIDeriv_apply, derivative_sum, sum_range_succ', sum_range_succ, + add_comm, ← add_zero (Finset.sum _ _)] + simp_rw [← Function.iterate_succ_apply' derivative, Nat.succ_eq_add_one, + Function.iterate_zero_apply, iterate_derivative_eq_zero (Nat.lt_succ_self _)] + +theorem exists_iterate_derivative_eq_factorial_smul (p : R[X]) (k : ℕ) : + ∃ gp : R[X], gp.natDegree ≤ p.natDegree - k ∧ derivative^[k] p = k ! • gp := by + refine ⟨_, (natDegree_sum_le _ _).trans ?_, iterate_derivative_eq_factorial_smul_sum p k⟩ + rw [fold_max_le] + refine ⟨Nat.zero_le _, fun i hi => ?_⟩ + dsimp only [Function.comp] + exact (natDegree_C_mul_le _ _).trans <| (natDegree_X_pow_le _).trans <| + (le_natDegree_of_mem_supp _ hi).trans <| natDegree_iterate_derivative _ _ + +end Semiring + +section CommSemiring + +variable [CommSemiring R] {A : Type*} [CommRing A] [Algebra R A] + +theorem aeval_iterate_derivative_of_lt (p : R[X]) (q : ℕ) (r : A) {p' : A[X]} + (hp : p.map (algebraMap R A) = (X - C r) ^ q * p') {k : ℕ} (hk : k < q) : + aeval r (derivative^[k] p) = 0 := by + have h (x) : (X - C r) ^ (q - (k - x)) = (X - C r) ^ 1 * (X - C r) ^ (q - (k - x) - 1) := by + rw [← pow_add, add_tsub_cancel_of_le] + rw [Nat.lt_iff_add_one_le] at hk + exact (le_tsub_of_add_le_left hk).trans (tsub_le_tsub_left (tsub_le_self : _ ≤ k) _) + rw [aeval_def, eval₂_eq_eval_map, ← iterate_derivative_map] + simp_rw [hp, iterate_derivative_mul, iterate_derivative_X_sub_pow, ← smul_mul_assoc, smul_smul, + h, ← mul_smul_comm, mul_assoc, ← mul_sum, eval_mul, pow_one, eval_sub, eval_X, eval_C, sub_self, + zero_mul] + +theorem aeval_iterate_derivative_self (p : R[X]) (q : ℕ) (r : A) {p' : A[X]} + (hp : p.map (algebraMap R A) = (X - C r) ^ q * p') : + aeval r (derivative^[q] p) = q ! • p'.eval r := by + have h (x) (h : 1 ≤ x) (h' : x ≤ q) : + (X - C r) ^ (q - (q - x)) = (X - C r) ^ 1 * (X - C r) ^ (q - (q - x) - 1) := by + rw [← pow_add, add_tsub_cancel_of_le] + rwa [tsub_tsub_cancel_of_le h'] + rw [aeval_def, eval₂_eq_eval_map, ← iterate_derivative_map] + simp_rw [hp, iterate_derivative_mul, iterate_derivative_X_sub_pow, ← smul_mul_assoc, smul_smul] + rw [sum_range_succ', Nat.choose_zero_right, one_mul, tsub_zero, Nat.descFactorial_self, tsub_self, + pow_zero, smul_mul_assoc, one_mul, Function.iterate_zero_apply, eval_add, eval_smul] + convert zero_add _ + rw [eval_finset_sum] + apply sum_eq_zero + intro x hx + rw [h (x + 1) le_add_self (Nat.add_one_le_iff.mpr (mem_range.mp hx)), pow_one, + eval_mul, eval_smul, eval_mul, eval_sub, eval_X, eval_C, sub_self, zero_mul, + smul_zero, zero_mul] + +variable (A) + +theorem aeval_iterate_derivative_of_ge (p : R[X]) (q : ℕ) {k : ℕ} (hk : q ≤ k) : + ∃ gp : R[X], gp.natDegree ≤ p.natDegree - k ∧ + ∀ r : A, aeval r (derivative^[k] p) = q ! • aeval r gp := by + obtain ⟨p', p'_le, hp'⟩ := exists_iterate_derivative_eq_factorial_smul p k + obtain ⟨k, rfl⟩ := Nat.exists_eq_add_of_le hk + refine ⟨((q + k).descFactorial k : R[X]) * p', (natDegree_C_mul_le _ _).trans p'_le, fun r => ?_⟩ + simp_rw [hp', nsmul_eq_mul, map_mul, map_natCast, ← mul_assoc, ← Nat.cast_mul, + Nat.add_descFactorial_eq_ascFactorial, Nat.factorial_mul_ascFactorial] + +theorem aeval_sumIDeriv (p : R[X]) (q : ℕ) : + ∃ gp : R[X], gp.natDegree ≤ p.natDegree - q ∧ + ∀ (r : A) {p' : A[X]}, p.map (algebraMap R A) = (X - C r) ^ q * p' → + aeval r (sumIDeriv p) = q ! • aeval r gp := by + have h (k) : + ∃ gp : R[X], gp.natDegree ≤ p.natDegree - q ∧ + ∀ (r : A) {p' : A[X]}, p.map (algebraMap R A) = (X - C r) ^ q * p' → + aeval r (derivative^[k] p) = q ! • aeval r gp := by + cases lt_or_ge k q with + | inl hk => + use 0 + rw [natDegree_zero] + use Nat.zero_le _ + intro r p' hp + rw [map_zero, smul_zero, aeval_iterate_derivative_of_lt p q r hp hk] + | inr hk => + obtain ⟨gp, gp_le, h⟩ := aeval_iterate_derivative_of_ge A p q hk + exact ⟨gp, gp_le.trans (tsub_le_tsub_left hk _), fun r p' _ => h r⟩ + choose c h using h + choose c_le hc using h + refine ⟨(range (p.natDegree + 1)).sum c, ?_, ?_⟩ + · refine (natDegree_sum_le _ _).trans ?_ + rw [fold_max_le] + exact ⟨Nat.zero_le _, fun i _ => c_le i⟩ + intro r p' hp + rw [sumIDeriv_apply, map_sum]; simp_rw [hc _ r hp, map_sum, smul_sum] + +theorem aeval_sumIDeriv_of_pos [Nontrivial A] [NoZeroDivisors A] (p : R[X]) {q : ℕ} (hq : 0 < q) : + ∃ gp : R[X], gp.natDegree ≤ p.natDegree - q ∧ + ∀ (inj_amap : Function.Injective (algebraMap R A)) (r : A) {p' : A[X]}, + p.map (algebraMap R A) = (X - C r) ^ (q - 1) * p' → + aeval r (sumIDeriv p) = (q - 1)! • p'.eval r + q ! • aeval r gp := by + rcases eq_or_ne p 0 with (rfl | p0) + · use 0 + rw [natDegree_zero] + use Nat.zero_le _ + intro _ r p' hp + rw [map_zero, map_zero, smul_zero, add_zero] + rw [Polynomial.map_zero] at hp + replace hp := (mul_eq_zero.mp hp.symm).resolve_left ?_ + rw [hp, eval_zero, smul_zero] + exact fun h => X_sub_C_ne_zero r (pow_eq_zero h) + let c k := if hk : q ≤ k then (aeval_iterate_derivative_of_ge A p q hk).choose else 0 + have c_le (k) : (c k).natDegree ≤ p.natDegree - k := by + dsimp only [c] + split_ifs with h + · exact (aeval_iterate_derivative_of_ge A p q h).choose_spec.1 + · rw [natDegree_zero]; exact Nat.zero_le _ + have hc (k) (hk : q ≤ k) : ∀ (r : A), aeval r (derivative^[k] p) = q ! • aeval r (c k) := by + simp_rw [c, dif_pos hk] + exact (aeval_iterate_derivative_of_ge A p q hk).choose_spec.2 + refine ⟨∑ x ∈ Ico q (p.natDegree + 1), c x, ?_, ?_⟩ + · refine (natDegree_sum_le _ _).trans ?_ + rw [fold_max_le] + exact ⟨Nat.zero_le _, fun i hi => (c_le i).trans (tsub_le_tsub_left (mem_Ico.mp hi).1 _)⟩ + intro inj_amap r p' hp + have : range (p.natDegree + 1) = range q ∪ Ico q (p.natDegree + 1) := by + rw [range_eq_Ico, Ico_union_Ico_eq_Ico hq.le] + have h := natDegree_map_le (algebraMap R A) p + rw [congr_arg natDegree hp, natDegree_mul, natDegree_pow, natDegree_X_sub_C, mul_one, + ← Nat.sub_add_comm (Nat.one_le_of_lt hq), tsub_le_iff_right] at h + exact le_of_add_le_left h + · exact pow_ne_zero _ (X_sub_C_ne_zero r) + · rintro rfl + rw [mul_zero, Polynomial.map_eq_zero_iff inj_amap] at hp + exact p0 hp + rw [← zero_add ((q - 1)! • p'.eval r)] + rw [sumIDeriv_apply, map_sum, map_sum, this] + have : range q = range (q - 1 + 1) := by rw [tsub_add_cancel_of_le (Nat.one_le_of_lt hq)] + rw [sum_union, this, sum_range_succ] + congr 2 + · apply sum_eq_zero + exact fun x hx => aeval_iterate_derivative_of_lt p _ r hp (mem_range.mp hx) + · rw [← aeval_iterate_derivative_self _ _ _ hp] + · rw [smul_sum, sum_congr rfl] + intro k hk + exact hc k (mem_Ico.mp hk).1 r + · rw [range_eq_Ico, disjoint_iff_inter_eq_empty, eq_empty_iff_forall_not_mem] + intro x hx + rw [mem_inter, mem_Ico, mem_Ico] at hx + exact hx.1.2.not_le hx.2.1 + +end CommSemiring + +end Polynomial From bdecdcb60f733cb48871ea7bf782a46027f518d7 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 19:56:27 +0000 Subject: [PATCH 031/131] chore: fix outdated `ext` porting notes (#17810) I went through all porting notes mentioning `ext` and fixed / removed those that no longer apply. --- Mathlib/Algebra/Polynomial/Monomial.lean | 9 +++------ Mathlib/AlgebraicGeometry/Limits.lean | 1 - Mathlib/CategoryTheory/EpiMono.lean | 2 -- Mathlib/Data/Set/Image.lean | 2 +- Mathlib/LinearAlgebra/Finsupp.lean | 6 ++---- Mathlib/LinearAlgebra/Matrix/Adjugate.lean | 3 +-- Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean | 3 +-- Mathlib/NumberTheory/Zsqrtd/Basic.lean | 3 +-- Mathlib/RepresentationTheory/Invariants.lean | 2 +- 9 files changed, 10 insertions(+), 21 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Monomial.lean b/Mathlib/Algebra/Polynomial/Monomial.lean index 29b15cda6f45e..535103330d899 100644 --- a/Mathlib/Algebra/Polynomial/Monomial.lean +++ b/Mathlib/Algebra/Polynomial/Monomial.lean @@ -56,12 +56,9 @@ theorem ringHom_ext {S} [Semiring S] {f g : R[X] →+* S} (h₁ : ∀ a, f (C a) set f' := f.comp (toFinsuppIso R).symm.toRingHom with hf' set g' := g.comp (toFinsuppIso R).symm.toRingHom with hg' have A : f' = g' := by - -- Porting note: Was `ext; simp [..]; simpa [..] using h₂`. - ext : 1 - · ext - simp [f', g', h₁, RingEquiv.toRingHom_eq_coe] - · refine MonoidHom.ext_mnat ?_ - simpa [RingEquiv.toRingHom_eq_coe] using h₂ + ext + simp [f', g', h₁, RingEquiv.toRingHom_eq_coe] + simpa using h₂ have B : f = f'.comp (toFinsuppIso R) := by rw [hf', RingHom.comp_assoc] ext x diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index d5614273639d9..0622355eef447 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -57,7 +57,6 @@ def Scheme.emptyTo (X : Scheme.{u}) : ∅ ⟶ X := @[ext] theorem Scheme.empty_ext {X : Scheme.{u}} (f g : ∅ ⟶ X) : f = g := - -- Porting note (#11041): `ext` regression Scheme.Hom.ext' (Subsingleton.elim (α := ∅ ⟶ _) _ _) theorem Scheme.eq_emptyTo {X : Scheme.{u}} (f : ∅ ⟶ X) : f = Scheme.emptyTo X := diff --git a/Mathlib/CategoryTheory/EpiMono.lean b/Mathlib/CategoryTheory/EpiMono.lean index 13cfd3e4ad1b3..8c8f738cc3b9f 100644 --- a/Mathlib/CategoryTheory/EpiMono.lean +++ b/Mathlib/CategoryTheory/EpiMono.lean @@ -38,7 +38,6 @@ such that `f ≫ retraction f = 𝟙 X`. Every split monomorphism is a monomorphism. -/ /- Porting note(#5171): removed @[nolint has_nonempty_instance] -/ -/- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule -/ @[ext, aesop apply safe (rule_sets := [CategoryTheory])] structure SplitMono {X Y : C} (f : X ⟶ Y) where /-- The map splitting `f` -/ @@ -64,7 +63,6 @@ such that `section_ f ≫ f = 𝟙 Y`. Every split epimorphism is an epimorphism. -/ /- Porting note(#5171): removed @[nolint has_nonempty_instance] -/ -/- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule -/ @[ext, aesop apply safe (rule_sets := [CategoryTheory])] structure SplitEpi {X Y : C} (f : X ⟶ Y) where /-- The map splitting `f` -/ diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 071d9c15c3f6c..00387c99f2202 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -950,7 +950,7 @@ theorem range_inclusion (h : s ⊆ t) : range (inclusion h) = { x : t | (x : α) -- When `f` is injective, see also `Equiv.ofInjective`. theorem leftInverse_rangeSplitting (f : α → β) : LeftInverse (rangeFactorization f) (rangeSplitting f) := fun x => by - apply Subtype.ext -- Porting note: why doesn't `ext` find this lemma? + ext simp only [rangeFactorization_coe] apply apply_rangeSplitting diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index 91e9feb4155da..5371c45091225 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -1331,8 +1331,7 @@ def splittingOfFinsuppSurjective (f : M →ₗ[R] α →₀ R) (s : Surjective f theorem splittingOfFinsuppSurjective_splits (f : M →ₗ[R] α →₀ R) (s : Surjective f) : f.comp (splittingOfFinsuppSurjective f s) = LinearMap.id := by - -- Porting note: `ext` can't find appropriate theorems. - refine lhom_ext' fun x => ext_ring <| Finsupp.ext fun y => ?_ + ext x dsimp [splittingOfFinsuppSurjective] congr rw [sum_single_index, one_smul] @@ -1357,8 +1356,7 @@ def splittingOfFunOnFintypeSurjective [Finite α] (f : M →ₗ[R] α → R) (s theorem splittingOfFunOnFintypeSurjective_splits [Finite α] (f : M →ₗ[R] α → R) (s : Surjective f) : f.comp (splittingOfFunOnFintypeSurjective f s) = LinearMap.id := by classical - -- Porting note: `ext` can't find appropriate theorems. - refine pi_ext' fun x => ext_ring <| funext fun y => ?_ + ext x y dsimp [splittingOfFunOnFintypeSurjective] rw [linearEquivFunOnFinite_symm_single, Finsupp.sum_single_index, one_smul, (s (Finsupp.single x 1)).choose_spec, Finsupp.single_eq_pi_single] diff --git a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean index ef3a1b2ec6823..516ba73963c50 100644 --- a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean +++ b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean @@ -114,8 +114,7 @@ theorem cramer_row_self (i : n) (h : ∀ j, b j = A j i) : A.cramer b = Pi.singl @[simp] theorem cramer_one : cramer (1 : Matrix n n α) = 1 := by - -- Porting note: was `ext i j` - refine LinearMap.pi_ext' (fun (i : n) => LinearMap.ext_ring (funext (fun (j : n) => ?_))) + ext i j convert congr_fun (cramer_row_self (1 : Matrix n n α) (Pi.single i 1) i _) j · simp · intro j diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean index e7d285c6ad0d2..a6138962526ff 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean @@ -93,8 +93,7 @@ theorem Matrix.represents_iff' {A : Matrix ι ι R} {f : Module.End R M} : have := LinearMap.congr_fun h (Pi.single i 1) rwa [PiToModule.fromEnd_apply_single_one, PiToModule.fromMatrix_apply_single_one] at this · intro h - -- Porting note: was `ext` - refine LinearMap.pi_ext' (fun i => LinearMap.ext_ring ?_) + ext simp_rw [LinearMap.comp_apply, LinearMap.coe_single, PiToModule.fromEnd_apply_single_one, PiToModule.fromMatrix_apply_single_one] apply h diff --git a/Mathlib/NumberTheory/Zsqrtd/Basic.lean b/Mathlib/NumberTheory/Zsqrtd/Basic.lean index 58b19978fb5b9..914914db4f6cb 100644 --- a/Mathlib/NumberTheory/Zsqrtd/Basic.lean +++ b/Mathlib/NumberTheory/Zsqrtd/Basic.lean @@ -931,8 +931,7 @@ def lift {d : ℤ} : { r : R // r * r = ↑d } ≃ (ℤ√d →+* R) where ext simp right_inv f := by - -- Porting note: was `ext` - refine hom_ext _ _ ?_ + ext simp /-- `lift r` is injective if `d` is non-square, and R has characteristic zero (that is, the map from diff --git a/Mathlib/RepresentationTheory/Invariants.lean b/Mathlib/RepresentationTheory/Invariants.lean index fa62df0a43948..898d8d3427513 100644 --- a/Mathlib/RepresentationTheory/Invariants.lean +++ b/Mathlib/RepresentationTheory/Invariants.lean @@ -134,7 +134,7 @@ def invariantsEquivRepHom (X Y : Rep k G) : (linHom X.ρ Y.ρ).invariants ≃ₗ map_add' _ _ := rfl map_smul' _ _ := rfl invFun f := ⟨f.hom, fun g => (mem_invariants_iff_comm _ g).2 (f.comm g)⟩ - left_inv _ := by apply Subtype.ext; ext; rfl -- Porting note: Added `apply Subtype.ext` + left_inv _ := by ext; rfl right_inv _ := by ext; rfl end Rep From e355e86a77c6deb2205be0bc9a4983e13b25eed4 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 17 Oct 2024 19:56:28 +0000 Subject: [PATCH 032/131] chore: fix "proof was" porting notes (#17836) Go through all porting notes that contain the string "proof was" and try to fix them. The following steps were used: * Golf the new proof. * If the current proof is nicer, delete the note. * Try to restore the old proof. * Classify some of the old failing proofs. --- .../Algebra/Category/ModuleCat/Subobject.lean | 4 +-- Mathlib/Algebra/Order/Group/Units.lean | 6 +--- Mathlib/Algebra/Order/Interval/Basic.lean | 7 ++--- Mathlib/Algebra/Star/Subalgebra.lean | 10 +++--- .../Analysis/Calculus/FDeriv/Measurable.lean | 1 - Mathlib/Data/List/OfFn.lean | 10 ++---- Mathlib/Data/Tree/Basic.lean | 4 +-- .../IsAlgClosed/AlgebraicClosure.lean | 31 +++++++------------ Mathlib/Geometry/Manifold/ChartedSpace.lean | 3 +- .../LinearAlgebra/CliffordAlgebra/Equivs.lean | 3 +- .../CliffordAlgebra/EvenEquiv.lean | 9 ++---- Mathlib/LinearAlgebra/LinearIndependent.lean | 4 +-- Mathlib/LinearAlgebra/SymplecticGroup.lean | 1 - Mathlib/MeasureTheory/Measure/Hausdorff.lean | 14 ++------- Mathlib/Order/Hom/Lattice.lean | 6 ---- Mathlib/RepresentationTheory/Rep.lean | 6 +--- Mathlib/RingTheory/AdjoinRoot.lean | 3 -- Mathlib/RingTheory/WittVector/IsPoly.lean | 4 --- Mathlib/Topology/Algebra/Monoid.lean | 2 -- Mathlib/Topology/Constructions.lean | 2 -- .../Topology/FiberBundle/Constructions.lean | 7 ++--- 21 files changed, 33 insertions(+), 104 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean index 134f3702fde18..a1508f23b2e54 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean @@ -82,7 +82,7 @@ noncomputable def toKernelSubobject {M N : ModuleCat.{v} R} {f : M ⟶ N} : @[simp] theorem toKernelSubobject_arrow {M N : ModuleCat R} {f : M ⟶ N} (x : LinearMap.ker f) : (kernelSubobject f).arrow (toKernelSubobject x) = x.1 := by - -- Porting note: The whole proof was just `simp [toKernelSubobject]`. + -- Porting note (#10959): the whole proof was just `simp [toKernelSubobject]`. suffices ((arrow ((kernelSubobject f))) ∘ (kernelSubobjectIso f ≪≫ kernelIsoKer f).inv) x = x by convert this rw [Iso.trans_inv, ← coe_comp, Category.assoc] @@ -105,7 +105,7 @@ theorem cokernel_π_imageSubobject_ext {L M N : ModuleCat.{v} R} (f : L ⟶ M) [ (g : (imageSubobject f : ModuleCat.{v} R) ⟶ N) [HasCokernel g] {x y : N} (l : L) (w : x = y + g (factorThruImageSubobject f l)) : cokernel.π g x = cokernel.π g y := by subst w - -- Porting note: The proof from here used to just be `simp`. + -- Porting note (#10959): The proof from here used to just be `simp`. simp only [map_add, add_right_eq_self] change ((cokernel.π g) ∘ (g) ∘ (factorThruImageSubobject f)) l = 0 rw [← coe_comp, ← coe_comp, Category.assoc] diff --git a/Mathlib/Algebra/Order/Group/Units.lean b/Mathlib/Algebra/Order/Group/Units.lean index af674406ea62a..c32449fad0ca7 100644 --- a/Mathlib/Algebra/Order/Group/Units.lean +++ b/Mathlib/Algebra/Order/Group/Units.lean @@ -20,11 +20,7 @@ variable {α : Type*} additive group."] instance Units.orderedCommGroup [OrderedCommMonoid α] : OrderedCommGroup αˣ := { Units.instPartialOrderUnits, Units.instCommGroupUnits with - mul_le_mul_left := fun _ _ h _ => (@mul_le_mul_left' α _ _ _ _ _ h _) } - --- Porting note: the mathlib3 proof was --- mul_le_mul_left := fun a b h c => (mul_le_mul_left' (h : (a : α) ≤ b) _ : (c : α) * a ≤ c * b) } --- see https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/elaboration.20failure.20in.20algebra.2Eorder.2Egroup.2Eunits + mul_le_mul_left := fun _ _ h _ => (mul_le_mul_left' (α := α) h _) } /-- The units of a linearly ordered commutative monoid form a linearly ordered commutative group. -/ @[to_additive "The units of a linearly ordered commutative additive monoid form a diff --git a/Mathlib/Algebra/Order/Interval/Basic.lean b/Mathlib/Algebra/Order/Interval/Basic.lean index dcd8a5d261e09..5fc5f067acadf 100644 --- a/Mathlib/Algebra/Order/Interval/Basic.lean +++ b/Mathlib/Algebra/Order/Interval/Basic.lean @@ -611,11 +611,8 @@ theorem length_sub_le : (s - t).length ≤ s.length + t.length := by simpa [sub_eq_add_neg] using length_add_le s (-t) theorem length_sum_le (f : ι → Interval α) (s : Finset ι) : - (∑ i ∈ s, f i).length ≤ ∑ i ∈ s, (f i).length := by - -- Porting note: Old proof was `:= Finset.le_sum_of_subadditive _ length_zero length_add_le _ _` - apply Finset.le_sum_of_subadditive - · exact length_zero - · exact length_add_le + (∑ i ∈ s, f i).length ≤ ∑ i ∈ s, (f i).length := + Finset.le_sum_of_subadditive _ length_zero length_add_le _ _ end Interval diff --git a/Mathlib/Algebra/Star/Subalgebra.lean b/Mathlib/Algebra/Star/Subalgebra.lean index 24e84fa477d77..a88d2390accc9 100644 --- a/Mathlib/Algebra/Star/Subalgebra.lean +++ b/Mathlib/Algebra/Star/Subalgebra.lean @@ -107,8 +107,7 @@ theorem toSubalgebra_le_iff {S₁ S₂ : StarSubalgebra R A} : equalities. -/ protected def copy (S : StarSubalgebra R A) (s : Set A) (hs : s = ↑S) : StarSubalgebra R A where toSubalgebra := Subalgebra.copy S.toSubalgebra s hs - star_mem' := @fun a ha => hs ▸ (S.star_mem' (by simpa [hs] using ha) : star a ∈ (S : Set A)) - -- Porting note: the old proof kept crashing Lean + star_mem' {a} ha := hs ▸ S.star_mem' (by simpa [hs] using ha) @[simp] theorem coe_copy (S : StarSubalgebra R A) (s : Set A) (hs : s = ↑S) : (S.copy s hs : Set A) = s := @@ -700,10 +699,9 @@ theorem ext_adjoin_singleton {a : A} [FunLike F (adjoin R ({a} : Set A)) B] variable [FunLike F A B] [AlgHomClass F R A B] [StarHomClass F A B] (f g : F) /-- The equalizer of two star `R`-algebra homomorphisms. -/ -def equalizer : StarSubalgebra R A := - { toSubalgebra := AlgHom.equalizer (f : A →ₐ[R] B) g - star_mem' := @fun a (ha : f a = g a) => by simpa only [← map_star] using congrArg star ha } --- Porting note: much like `StarSubalgebra.copy` the old proof was broken and hard to fix +def equalizer : StarSubalgebra R A where + toSubalgebra := AlgHom.equalizer (f : A →ₐ[R] B) g + star_mem' {a} (ha : f a = g a) := by simpa only [← map_star] using congrArg star ha @[simp] theorem mem_equalizer (x : A) : x ∈ StarAlgHom.equalizer f g ↔ f x = g x := diff --git a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean index 28e1539f0e516..5801ab5775fa7 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean @@ -592,7 +592,6 @@ theorem D_subset_differentiable_set {K : Set F} (hK : IsComplete K) : _ ≤ ‖L e p q - L e p r‖ + ‖L e p r - L e' p' r‖ + ‖L e' p' r - L e' p' q'‖ := (le_trans (norm_add_le _ _) (add_le_add_right (norm_add_le _ _) _)) _ ≤ 4 * (1 / 2) ^ e + 4 * (1 / 2) ^ e + 4 * (1 / 2) ^ e := by gcongr - -- Porting note: proof was `by apply_rules [add_le_add]` _ = 12 * (1 / 2) ^ e := by ring /- For definiteness, use `L0 e = L e (n e) (n e)`, to have a single sequence. We claim that this diff --git a/Mathlib/Data/List/OfFn.lean b/Mathlib/Data/List/OfFn.lean index 9aa74790b948b..a3271be86ffd8 100644 --- a/Mathlib/Data/List/OfFn.lean +++ b/Mathlib/Data/List/OfFn.lean @@ -200,14 +200,8 @@ def ofFnRec {C : List α → Sort*} (h : ∀ (n) (f : Fin n → α), C (List.ofF @[simp] theorem ofFnRec_ofFn {C : List α → Sort*} (h : ∀ (n) (f : Fin n → α), C (List.ofFn f)) {n : ℕ} - (f : Fin n → α) : @ofFnRec _ C h (List.ofFn f) = h _ f := by - -- Porting note: Old proof was - -- equivSigmaTuple.rightInverse_symm.cast_eq (fun s => h s.1 s.2) ⟨n, f⟩ - have := (@equivSigmaTuple α).rightInverse_symm - dsimp [equivSigmaTuple] at this - have := this.cast_eq (fun s => h s.1 s.2) ⟨n, f⟩ - dsimp only at this - rw [ofFnRec, ← this] + (f : Fin n → α) : @ofFnRec _ C h (List.ofFn f) = h _ f := + equivSigmaTuple.rightInverse_symm.cast_eq (fun s => h s.1 s.2) ⟨n, f⟩ theorem exists_iff_exists_tuple {P : List α → Prop} : (∃ l : List α, P l) ↔ ∃ (n : _) (f : Fin n → α), P (List.ofFn f) := diff --git a/Mathlib/Data/Tree/Basic.lean b/Mathlib/Data/Tree/Basic.lean index 95ebf3246593c..df733700e6e4a 100644 --- a/Mathlib/Data/Tree/Basic.lean +++ b/Mathlib/Data/Tree/Basic.lean @@ -103,9 +103,7 @@ compile_inductive% Tree @[elab_as_elim] def unitRecOn {motive : Tree Unit → Sort*} (t : Tree Unit) (base : motive nil) (ind : ∀ x y, motive x → motive y → motive (x △ y)) : motive t := - -- Porting note: Old proof was `t.recOn base fun u => u.recOn ind` but - -- structure eta makes it unnecessary (https://github.com/leanprover/lean4/issues/777). - t.recOn base fun _u => ind + t.recOn base fun _u ↦ ind theorem left_node_right_eq_self : ∀ {x : Tree Unit} (_hx : x ≠ nil), x.left △ x.right = x | nil, h => by trivial diff --git a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index a00d2cba8705e..ea7fe25e09f58 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -164,12 +164,8 @@ instance Step.algebraSucc (n) : Algebra (Step k n) (Step k (n + 1)) := (toStepSucc k n).toAlgebra theorem toStepSucc.exists_root {n} {f : Polynomial (Step k n)} (hfm : f.Monic) - (hfi : Irreducible f) : ∃ x : Step k (n + 1), f.eval₂ (toStepSucc k n) x = 0 := by --- Porting note: original proof was `@AdjoinMonic.exists_root _ (Step.field k n) _ hfm hfi`, --- but it timeouts. - obtain ⟨x, hx⟩ := @AdjoinMonic.exists_root _ (Step.field k n) _ hfm hfi --- Porting note: using `hx` instead of `by apply hx` timeouts. - exact ⟨x, by apply hx⟩ + (hfi : Irreducible f) : ∃ x : Step k (n + 1), f.eval₂ (toStepSucc k n) x = 0 := + @AdjoinMonic.exists_root _ (Step.field k n) _ hfm hfi -- Porting note: the following two declarations were added during the port to be used in the -- definition of toStepOfLE @@ -186,29 +182,24 @@ private theorem toStepOfLE'.succ (m n : ℕ) (h : m ≤ n) : def toStepOfLE (m n : ℕ) (h : m ≤ n) : Step k m →+* Step k n where toFun := toStepOfLE' k m n h map_one' := by --- Porting note: original proof was `induction' h with n h ih; · exact Nat.leRecOn_self 1` --- `rw [Nat.leRecOn_succ h, ih, RingHom.map_one]` induction' h with a h ih · exact Nat.leRecOn_self 1 - · rw [toStepOfLE'.succ k m a h]; simp [ih] + · simp [toStepOfLE'.succ k m a h, ih] map_mul' x y := by --- Porting note: original proof was `induction' h with n h ih; · simp_rw [Nat.leRecOn_self]` --- `simp_rw [Nat.leRecOn_succ h, ih, RingHom.map_mul]` + simp only induction' h with a h ih - · dsimp [toStepOfLE']; simp_rw [Nat.leRecOn_self] - · simp_rw [toStepOfLE'.succ k m a h]; simp only at ih; simp [ih] --- Porting note: original proof was `induction' h with n h ih; · exact Nat.leRecOn_self 0` --- `rw [Nat.leRecOn_succ h, ih, RingHom.map_zero]` + · simp_rw [toStepOfLE', Nat.leRecOn_self] + · simp [toStepOfLE'.succ k m a h, ih] map_zero' := by + simp only induction' h with a h ih · exact Nat.leRecOn_self 0 - · simp_rw [toStepOfLE'.succ k m a h]; simp only at ih; simp [ih] + · simp [toStepOfLE'.succ k m a h, ih] map_add' x y := by --- Porting note: original proof was `induction' h with n h ih; · simp_rw [Nat.leRecOn_self]` --- `simp_rw [Nat.leRecOn_succ h, ih, RingHom.map_add]` + simp only induction' h with a h ih - · dsimp [toStepOfLE']; simp_rw [Nat.leRecOn_self] - · simp_rw [toStepOfLE'.succ k m a h]; simp only at ih; simp [ih] + · simp_rw [toStepOfLE', Nat.leRecOn_self] + · simp [toStepOfLE'.succ k m a h, ih] @[simp] theorem coe_toStepOfLE (m n : ℕ) (h : m ≤ n) : diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 6a69ac58da981..5acd6a8cea119 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -758,9 +758,8 @@ section variable {ι : Type*} {Hi : ι → Type*} --- Porting note: Old proof was `Pi.inhabited _`. instance modelPiInhabited [∀ i, Inhabited (Hi i)] : Inhabited (ModelPi Hi) := - ⟨fun _ ↦ default⟩ + Pi.instInhabited instance [∀ i, TopologicalSpace (Hi i)] : TopologicalSpace (ModelPi Hi) := Pi.topologicalSpace diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean index 156e71925668f..2285548a3e107 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean @@ -392,7 +392,6 @@ theorem equiv_ι (r : R) : CliffordAlgebraDualNumber.equiv (ι (R := R) _ r) = r @[simp] theorem equiv_symm_eps : CliffordAlgebraDualNumber.equiv.symm (eps : R[ε]) = ι (0 : QuadraticForm R R) 1 := - -- Porting note: Original proof was `DualNumber.lift_apply_eps _` - DualNumber.lift_apply_eps (R := R) (B := CliffordAlgebra (0 : QuadraticForm R R)) _ + DualNumber.lift_apply_eps _ end CliffordAlgebraDualNumber diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean index 119f96a21ae0d..e9ab3a38b7895 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean @@ -152,13 +152,8 @@ def ofEven : CliffordAlgebra.even (Q' Q) →ₐ[R] CliffordAlgebra Q := by theorem ofEven_ι (x y : M × R) : ofEven Q ((even.ι (Q' Q)).bilin x y) = - (ι Q x.1 + algebraMap R _ x.2) * (ι Q y.1 - algebraMap R _ y.2) := by - -- Porting note: entire proof was the term-mode `even.lift_ι (Q' Q) _ x y` - unfold ofEven - lift_lets - intro f - -- TODO: replacing `?_` with `_` takes way longer? - exact @even.lift_ι R (M × R) _ _ _ (Q' Q) _ _ _ ⟨f, ?_, ?_⟩ x y + (ι Q x.1 + algebraMap R _ x.2) * (ι Q y.1 - algebraMap R _ y.2) := + even.lift_ι (Q' Q) _ x y theorem toEven_comp_ofEven : (toEven Q).comp (ofEven Q) = AlgHom.id R _ := even.algHom_ext (Q' Q) <| diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index ccd018f0d3ddf..3f46f76b260ab 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -931,9 +931,7 @@ theorem LinearIndependent.independent_span_singleton (hv : LinearIndependent R v obtain ⟨⟨r, rfl⟩, hm⟩ := hm suffices r = 0 by simp [this] apply linearIndependent_iff_not_smul_mem_span.mp hv i - -- Porting note: The original proof was using `convert hm`. - suffices v '' (univ \ {i}) = range fun j : { j // j ≠ i } => v j by - rwa [this] + convert hm ext simp diff --git a/Mathlib/LinearAlgebra/SymplecticGroup.lean b/Mathlib/LinearAlgebra/SymplecticGroup.lean index c638f98ae8bd9..55b9036e4c5a4 100644 --- a/Mathlib/LinearAlgebra/SymplecticGroup.lean +++ b/Mathlib/LinearAlgebra/SymplecticGroup.lean @@ -86,7 +86,6 @@ open Matrix theorem mem_iff {A : Matrix (l ⊕ l) (l ⊕ l) R} : A ∈ symplecticGroup l R ↔ A * J l R * Aᵀ = J l R := by simp [symplecticGroup] --- Porting note: Previous proof was `by infer_instance` instance coeMatrix : Coe (symplecticGroup l R) (Matrix (l ⊕ l) (l ⊕ l) R) := ⟨Subtype.val⟩ diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index d4023bb58018d..ce69db687e782 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -1026,18 +1026,8 @@ theorem hausdorffMeasure_smul_right_image [NormedAddCommGroup E] [NormedSpace -- break lineMap into pieces suffices μH[1] ((‖v‖ • ·) '' (LinearMap.toSpanSingleton ℝ E (‖v‖⁻¹ • v) '' s)) = ‖v‖₊ • μH[1] s by - -- Porting note: proof was shorter, could need some golf - simp only [hausdorffMeasure_real, nnreal_smul_coe_apply] - convert this - · simp only [image_smul, LinearMap.toSpanSingleton_apply, Set.image_image] - ext e - simp only [mem_image] - refine ⟨fun ⟨x, h⟩ => ⟨x, ?_⟩, fun ⟨x, h⟩ => ⟨x, ?_⟩⟩ - · rw [smul_comm (norm _), smul_comm (norm _), inv_smul_smul₀ hn] - exact h - · rw [smul_comm (norm _), smul_comm (norm _), inv_smul_smul₀ hn] at h - exact h - · exact hausdorffMeasure_real.symm + simpa only [Set.image_image, smul_comm (norm _), inv_smul_smul₀ hn, + LinearMap.toSpanSingleton_apply] using this have iso_smul : Isometry (LinearMap.toSpanSingleton ℝ E (‖v‖⁻¹ • v)) := by refine AddMonoidHomClass.isometry_of_norm _ fun x => (norm_smul _ _).trans ?_ rw [norm_smul, norm_inv, norm_norm, inv_mul_cancel₀ hn, mul_one, LinearMap.id_apply] diff --git a/Mathlib/Order/Hom/Lattice.lean b/Mathlib/Order/Hom/Lattice.lean index 94af3f7566504..4179afdc3722b 100644 --- a/Mathlib/Order/Hom/Lattice.lean +++ b/Mathlib/Order/Hom/Lattice.lean @@ -1567,7 +1567,6 @@ theorem withTop_id : (SupHom.id α).withTop = SupHom.id _ := DFunLike.coe_inject @[simp] theorem withTop_comp (f : SupHom β γ) (g : SupHom α β) : (f.comp g).withTop = f.withTop.comp g.withTop := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊥` to the domain and codomain of a `SupHom`. -/ @@ -1588,7 +1587,6 @@ theorem withBot_id : (SupHom.id α).withBot = SupBotHom.id _ := DFunLike.coe_inj @[simp] theorem withBot_comp (f : SupHom β γ) (g : SupHom α β) : (f.comp g).withBot = f.withBot.comp g.withBot := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊤` to the codomain of a `SupHom`. -/ @@ -1638,7 +1636,6 @@ theorem withTop_id : (InfHom.id α).withTop = InfTopHom.id _ := DFunLike.coe_inj @[simp] theorem withTop_comp (f : InfHom β γ) (g : InfHom α β) : (f.comp g).withTop = f.withTop.comp g.withTop := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊥` to the domain and codomain of an `InfHom`. -/ @@ -1658,7 +1655,6 @@ theorem withBot_id : (InfHom.id α).withBot = InfHom.id _ := DFunLike.coe_inject @[simp] theorem withBot_comp (f : InfHom β γ) (g : InfHom α β) : (f.comp g).withBot = f.withBot.comp g.withBot := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊤` to the codomain of an `InfHom`. -/ @@ -1708,7 +1704,6 @@ theorem withTop_id : (LatticeHom.id α).withTop = LatticeHom.id _ := @[simp] theorem withTop_comp (f : LatticeHom β γ) (g : LatticeHom α β) : (f.comp g).withTop = f.withTop.comp g.withTop := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊥` to the domain and codomain of a `LatticeHom`. -/ @@ -1729,7 +1724,6 @@ theorem withBot_id : (LatticeHom.id α).withBot = LatticeHom.id _ := @[simp] theorem withBot_comp (f : LatticeHom β γ) (g : LatticeHom α β) : (f.comp g).withBot = f.withBot.comp g.withBot := --- Porting note: Proof was `DFunLike.coe_injective (Option.map_comp_map _ _).symm` DFunLike.coe_injective <| Eq.symm <| Option.map_comp_map _ _ /-- Adjoins a `⊤` and `⊥` to the domain and codomain of a `LatticeHom`. -/ diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 8ebf69973568b..e6d2b8e5cd51e 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -183,13 +183,9 @@ theorem linearization_μ_hom (X Y : Action (Type u) (MonCat.of G)) : @[simp] theorem linearization_μ_inv_hom (X Y : Action (Type u) (MonCat.of G)) : (inv ((linearization k G).μ X Y)).hom = (finsuppTensorFinsupp' k X.V Y.V).symm.toLinearMap := by --- Porting note (#11039): broken proof was -/- simp_rw [← Action.forget_map, Functor.map_inv, Action.forget_map, linearization_μ_hom] - apply IsIso.inv_eq_of_hom_inv_id _ - exact LinearMap.ext fun x => LinearEquiv.symm_apply_apply _ _-/ rw [← Action.forget_map, Functor.map_inv] apply IsIso.inv_eq_of_hom_inv_id - exact LinearMap.ext fun x => LinearEquiv.symm_apply_apply (finsuppTensorFinsupp' k X.V Y.V) x + exact LinearMap.ext fun x ↦ LinearEquiv.symm_apply_apply _ _ @[simp] theorem linearization_ε_hom : (linearization k G).ε.hom = Finsupp.lsingle PUnit.unit := diff --git a/Mathlib/RingTheory/AdjoinRoot.lean b/Mathlib/RingTheory/AdjoinRoot.lean index bba7f1e3e7457..eed173fefa7a8 100644 --- a/Mathlib/RingTheory/AdjoinRoot.lean +++ b/Mathlib/RingTheory/AdjoinRoot.lean @@ -205,7 +205,6 @@ theorem aeval_eq (p : R[X]) : aeval (root f) p = mk f p := rw [_root_.map_mul, aeval_C, map_pow, aeval_X, RingHom.map_mul, mk_C, RingHom.map_pow, mk_X] rfl --- Porting note: the following proof was partly in term-mode, but I was not able to fix it. theorem adjoinRoot_eq_top : Algebra.adjoin R ({root f} : Set (AdjoinRoot f)) = ⊤ := by refine Algebra.eq_top_iff.2 fun x => ?_ induction x using AdjoinRoot.induction_on with @@ -402,7 +401,6 @@ def modByMonicHom (hg : g.Monic) : AdjoinRoot g →ₗ[R] R[X] := theorem modByMonicHom_mk (hg : g.Monic) (f : R[X]) : modByMonicHom hg (mk g f) = f %ₘ g := rfl --- Porting note: the following proof was partly in term-mode, but I was not able to fix it. theorem mk_leftInverse (hg : g.Monic) : Function.LeftInverse (mk g) (modByMonicHom hg) := by intro f induction f using AdjoinRoot.induction_on @@ -716,7 +714,6 @@ def quotAdjoinRootEquivQuotPolynomialQuot : ((Ideal.quotEquivOfEq (by rw [map_span, Set.image_singleton])).trans (Polynomial.quotQuotEquivComm I f).symm)) --- Porting note: mathlib3 proof was a long `rw` that timeouts. @[simp] theorem quotAdjoinRootEquivQuotPolynomialQuot_mk_of (p : R[X]) : quotAdjoinRootEquivQuotPolynomialQuot I f (Ideal.Quotient.mk (I.map (of f)) (mk f p)) = diff --git a/Mathlib/RingTheory/WittVector/IsPoly.lean b/Mathlib/RingTheory/WittVector/IsPoly.lean index 69617badd6694..87bd44d9849c4 100644 --- a/Mathlib/RingTheory/WittVector/IsPoly.lean +++ b/Mathlib/RingTheory/WittVector/IsPoly.lean @@ -311,15 +311,11 @@ end ZeroOne /-- Addition of Witt vectors is a polynomial function. -/ -- Porting note: replaced `@[is_poly]` with `instance`. instance addIsPoly₂ [Fact p.Prime] : IsPoly₂ p fun _ _ => (· + ·) := - -- porting note: the proof was - -- `⟨⟨wittAdd p, by intros; dsimp only [WittVector.hasAdd]; simp [eval]⟩⟩` ⟨⟨wittAdd p, by intros; ext; exact add_coeff _ _ _⟩⟩ /-- Multiplication of Witt vectors is a polynomial function. -/ -- Porting note: replaced `@[is_poly]` with `instance`. instance mulIsPoly₂ [Fact p.Prime] : IsPoly₂ p fun _ _ => (· * ·) := - -- porting note: the proof was - -- `⟨⟨wittMul p, by intros; dsimp only [WittVector.hasMul]; simp [eval]⟩⟩` ⟨⟨wittMul p, by intros; ext; exact mul_coeff _ _ _⟩⟩ -- unfortunately this is not universe polymorphic, merely because `f` isn't diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index 126c6636dd313..a938067727496 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -302,8 +302,6 @@ theorem isClosed_setOf_map_mul [Mul M₁] [Mul M₂] [ContinuousMul M₂] : isClosed_iInter fun x => isClosed_iInter fun y => isClosed_eq (continuous_apply _) - -- Porting note: proof was: - -- `((continuous_apply _).mul (continuous_apply _))` (by continuity) -- Porting note: split variables command over two lines, can't change explicitness at the same time diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 646dbbbcf8c5a..e0bb583947c37 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -880,11 +880,9 @@ theorem continuous_isRight : Continuous (isRight : X ⊕ Y → Bool) := continuous_sum_dom.2 ⟨continuous_const, continuous_const⟩ @[continuity, fun_prop] --- Porting note: the proof was `continuous_sup_rng_left continuous_coinduced_rng` theorem continuous_inl : Continuous (@inl X Y) := ⟨fun _ => And.left⟩ @[continuity, fun_prop] --- Porting note: the proof was `continuous_sup_rng_right continuous_coinduced_rng` theorem continuous_inr : Continuous (@inr X Y) := ⟨fun _ => And.right⟩ @[fun_prop, continuity] diff --git a/Mathlib/Topology/FiberBundle/Constructions.lean b/Mathlib/Topology/FiberBundle/Constructions.lean index a9ac5e6eb5149..5ecf7487e2490 100644 --- a/Mathlib/Topology/FiberBundle/Constructions.lean +++ b/Mathlib/Topology/FiberBundle/Constructions.lean @@ -250,11 +250,8 @@ universe u v w₁ w₂ U variable {B : Type u} (F : Type v) (E : B → Type w₁) {B' : Type w₂} (f : B' → B) -instance [∀ x : B, TopologicalSpace (E x)] : ∀ x : B', TopologicalSpace ((f *ᵖ E) x) := by - -- Porting note: Original proof was `delta_instance Bundle.Pullback` - intro x - rw [Bundle.Pullback] - infer_instance +instance [∀ x : B, TopologicalSpace (E x)] : ∀ x : B', TopologicalSpace ((f *ᵖ E) x) := + inferInstanceAs (∀ x, TopologicalSpace (E (f x))) variable [TopologicalSpace B'] [TopologicalSpace (TotalSpace F E)] From a4d283b5f2f1bfb9f5c6d8c8c441d5e9d3b3373b Mon Sep 17 00:00:00 2001 From: Peter Nelson <71660771+apnelson1@users.noreply.github.com> Date: Thu, 17 Oct 2024 19:56:29 +0000 Subject: [PATCH 033/131] feat(Order/BooleanAlgebra, Data/Set/Basic): diff_le_diff_iff (#17876) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds a lemma stating that, for `x, y ≤ z` in a boolean algebra, we have `z \ x ≤ z \ y` iff `y ≤ x`. We also add the same lemma for sets/finsets. --- Mathlib/Data/Finset/Basic.lean | 4 ++++ Mathlib/Data/Set/Basic.lean | 4 ++++ Mathlib/Order/BooleanAlgebra.lean | 5 +++++ 3 files changed, 13 insertions(+) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index d8f077efcc2ea..f226586bced60 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -1828,6 +1828,10 @@ theorem sdiff_empty : s \ ∅ = s := theorem sdiff_subset_sdiff (hst : s ⊆ t) (hvu : v ⊆ u) : s \ u ⊆ t \ v := sdiff_le_sdiff hst hvu +theorem sdiff_subset_sdiff_iff_subset {r : Finset α} (hs : s ⊆ r) (ht : t ⊆ r) : + r \ s ⊆ r \ t ↔ t ⊆ s := + sdiff_le_sdiff_iff_le hs ht + @[simp, norm_cast] theorem coe_sdiff (s₁ s₂ : Finset α) : ↑(s₁ \ s₂) = (s₁ \ s₂ : Set α) := Set.ext fun _ => mem_sdiff diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 2cc7a256bcfdf..63e22e76e74a8 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -1497,6 +1497,10 @@ theorem diff_subset_diff_left {s₁ s₂ t : Set α} (h : s₁ ⊆ s₂) : s₁ theorem diff_subset_diff_right {s t u : Set α} (h : t ⊆ u) : s \ u ⊆ s \ t := sdiff_le_sdiff_left ‹t ≤ u› +theorem diff_subset_diff_iff_subset {r : Set α} (hs : s ⊆ r) (ht : t ⊆ r) : + r \ s ⊆ r \ t ↔ t ⊆ s := + sdiff_le_sdiff_iff_le hs ht + theorem compl_eq_univ_diff (s : Set α) : sᶜ = univ \ s := top_sdiff.symm diff --git a/Mathlib/Order/BooleanAlgebra.lean b/Mathlib/Order/BooleanAlgebra.lean index 60726797718c4..ba0cb4558f27b 100644 --- a/Mathlib/Order/BooleanAlgebra.lean +++ b/Mathlib/Order/BooleanAlgebra.lean @@ -359,6 +359,11 @@ theorem sdiff_eq_comm (hy : y ≤ x) (hz : z ≤ x) : x \ y = z ↔ x \ z = y := theorem eq_of_sdiff_eq_sdiff (hxz : x ≤ z) (hyz : y ≤ z) (h : z \ x = z \ y) : x = y := by rw [← sdiff_sdiff_eq_self hxz, h, sdiff_sdiff_eq_self hyz] +theorem sdiff_le_sdiff_iff_le (hx : x ≤ z) (hy : y ≤ z) : z \ x ≤ z \ y ↔ y ≤ x := by + refine ⟨fun h ↦ ?_, sdiff_le_sdiff_left⟩ + rw [← sdiff_sdiff_eq_self hx, ← sdiff_sdiff_eq_self hy] + exact sdiff_le_sdiff_left h + theorem sdiff_sdiff_left' : (x \ y) \ z = x \ y ⊓ x \ z := by rw [sdiff_sdiff_left, sdiff_sup] theorem sdiff_sdiff_sup_sdiff : z \ (x \ y ⊔ y \ x) = z ⊓ (z \ x ⊔ y) ⊓ (z \ y ⊔ x) := From 78d236b37033084670ab8e8b7b0514f719bc6696 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Thu, 17 Oct 2024 20:34:59 +0000 Subject: [PATCH 034/131] feat(GroupTheory/GroupAction/Blocks): lemmas to handle blocks on finite sets (#14029) This PR proves specific properties of blocks when the action is on a finite set Co-authored-by: Thomas Browning --- Mathlib/GroupTheory/GroupAction/Blocks.lean | 138 ++++++++++++++++++-- Mathlib/GroupTheory/Index.lean | 20 ++- 2 files changed, 147 insertions(+), 11 deletions(-) diff --git a/Mathlib/GroupTheory/GroupAction/Blocks.lean b/Mathlib/GroupTheory/GroupAction/Blocks.lean index dddd660c93991..e88c5972e23aa 100644 --- a/Mathlib/GroupTheory/GroupAction/Blocks.lean +++ b/Mathlib/GroupTheory/GroupAction/Blocks.lean @@ -4,11 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir -/ +import Mathlib.Algebra.Group.Subgroup.Actions +import Mathlib.Data.Set.Card import Mathlib.Data.Setoid.Partition import Mathlib.GroupTheory.GroupAction.Basic import Mathlib.GroupTheory.GroupAction.Pointwise import Mathlib.GroupTheory.GroupAction.SubMulAction -import Mathlib.Algebra.Group.Subgroup.Actions +import Mathlib.GroupTheory.Index +import Mathlib.Tactic.IntervalCases /-! # Blocks @@ -22,9 +25,21 @@ Given `SMul G X`, an action of a type `G` on a type `X`, we define The non-existence of nontrivial blocks is the definition of primitive actions. +## Results for actions on finite sets + +- `IsBlock.ncard_block_mul_ncard_orbit_eq` : The cardinality of a block +multiplied by the number of its translates is the cardinal of the ambient type + +- `IsBlock.is_univ_of_large_block` : a too large block is equal to `Set.univ` + +- `IsBlock.is_subsingleton` : a too small block is a subsingleton + +- `IsBlock.of_subset` : the intersections of the translates of a finite subset +that contain a given point is a block + ## References -We follow [wieland1964]. +We follow [wielandt1964]. -/ @@ -200,12 +215,13 @@ theorem isBlock_orbit (a : X) : IsBlock G (orbit G a) := variable (X) /-- The full set is a (trivial) block -/ -theorem isFixedBlock_top : IsFixedBlock G (⊤ : Set X) := - fun _ ↦ by simp only [Set.top_eq_univ, Set.smul_set_univ] +theorem isFixedBlock_univ : IsFixedBlock G (Set.univ : Set X) := + fun _ ↦ by simp only [Set.smul_set_univ] +@[deprecated (since := "2024-09-14")] alias isFixedBlock_top := isFixedBlock_univ /-- The full set is a (trivial) block -/ -theorem isBlock_top : IsBlock G (⊤ : Set X) := - (isFixedBlock_top _).isBlock +theorem isBlock_univ : IsBlock G (Set.univ : Set X) := + (isFixedBlock_univ _).isBlock variable {X} @@ -282,9 +298,9 @@ theorem IsBlock.inter {B₁ B₂ : Set X} (h₁ : IsBlock G B₁) (h₂ : IsBloc theorem IsBlock.iInter {ι : Type*} {B : ι → Set X} (hB : ∀ i : ι, IsBlock G (B i)) : IsBlock G (⋂ i, B i) := by by_cases hι : (IsEmpty ι) - · -- ι = ∅, block = ⊤ - suffices (⋂ i : ι, B i) = Set.univ by simpa only [this] using isBlock_top X - simpa only [Set.top_eq_univ, Set.iInter_eq_univ] using (hι.elim' ·) + · -- ι = ∅, block = univ + suffices (⋂ i : ι, B i) = Set.univ by simpa only [this] using isBlock_univ X + simpa only [Set.iInter_eq_univ] using (hι.elim' ·) rw [IsBlock.def_one] intro g rw [Set.smul_set_iInter] @@ -470,6 +486,110 @@ def block_stabilizerOrderIso [htGX : IsPretransitive G X] (a : X) : end Stabilizer +section Finite + +namespace IsBlock + +variable [IsPretransitive G X] {B : Set X} + +theorem ncard_block_eq_relindex (hB : IsBlock G B) {x : X} (hx : x ∈ B) : + B.ncard = (stabilizer G x).relindex (stabilizer G B) := by + have key : (stabilizer G x).subgroupOf (stabilizer G B) = stabilizer (stabilizer G B) x := by + ext; rfl + rw [Subgroup.relindex, key, index_stabilizer, hB.orbit_stabilizer_eq hx] + +/-- The cardinality of the ambient space is the product of the cardinality of a block + by the cardinality of the set of translates of that block -/ +theorem ncard_block_mul_ncard_orbit_eq (hB : IsBlock G B) (hB_ne : B.Nonempty) : + Set.ncard B * Set.ncard (orbit G B) = Nat.card X := by + obtain ⟨x, hx⟩ := hB_ne + rw [ncard_block_eq_relindex hB hx, ← index_stabilizer, + Subgroup.relindex_mul_index (hB.stabilizer_le hx), index_stabilizer_of_transitive] + +/-- The cardinality of a block divides the cardinality of the ambient type -/ +theorem ncard_dvd_card (hB : IsBlock G B) (hB_ne : B.Nonempty) : + Set.ncard B ∣ Nat.card X := + Dvd.intro _ (hB.ncard_block_mul_ncard_orbit_eq hB_ne) + +/-- A too large block is equal to `univ` -/ +theorem eq_univ_card_lt [hX : Finite X] (hB : IsBlock G B) (hB' : Nat.card X < Set.ncard B * 2) : + B = Set.univ := by + rcases Set.eq_empty_or_nonempty B with rfl | hB_ne + · simp only [Set.ncard_empty, zero_mul, not_lt_zero'] at hB' + have key := hB.ncard_block_mul_ncard_orbit_eq hB_ne + rw [← key, mul_lt_mul_iff_of_pos_left (by rwa [Set.ncard_pos])] at hB' + interval_cases (orbit G B).ncard + · rw [mul_zero, eq_comm, Nat.card_eq_zero, or_iff_left hX.not_infinite] at key + exact (IsEmpty.exists_iff.mp hB_ne).elim + · rw [mul_one, ← Set.ncard_univ] at key + rw [Set.eq_of_subset_of_ncard_le (Set.subset_univ B) key.ge] + +/-- If a block has too many translates, then it is a (sub)singleton -/ +theorem subsingleton_of_card_lt [Finite X] (hB : IsBlock G B) + (hB' : Nat.card X < 2 * Set.ncard (orbit G B)) : + B.Subsingleton := by + suffices Set.ncard B < 2 by + rw [Nat.lt_succ_iff, Set.ncard_le_one_iff_eq] at this + cases this with + | inl h => rw [h]; exact Set.subsingleton_empty + | inr h => + obtain ⟨a, ha⟩ := h; rw [ha]; exact Set.subsingleton_singleton + cases Set.eq_empty_or_nonempty B with + | inl h => rw [h, Set.ncard_empty]; norm_num + | inr h => + rw [← hB.ncard_block_mul_ncard_orbit_eq h, lt_iff_not_ge] at hB' + rw [← not_le] + exact fun hb ↦ hB' (Nat.mul_le_mul_right _ hb) + +/- The assumption `B.Finite` is necessary : + For G = ℤ acting on itself, a = 0 and B = ℕ, the translates `k • B` of the statement + are just `k + ℕ`, for `k ≤ 0`, and the corresponding intersection is `ℕ`, which is not a block. + (Remark by Thomas Browning) -/ +-- Note : add {B} because otherwise Lean includes `hB : IsBlock G B` +/-- The intersection of the translates of a *finite* subset which contain a given point +is a block (Wielandt, th. 7.3 )-/ +theorem of_subset {B : Set X} (a : X) (hfB : B.Finite) : + IsBlock G (⋂ (k : G) (_ : a ∈ k • B), k • B) := by + let B' := ⋂ (k : G) (_ : a ∈ k • B), k • B + cases' Set.eq_empty_or_nonempty B with hfB_e hfB_ne + · simp [hfB_e, isBlock_univ] + have hB'₀ : ∀ (k : G) (_ : a ∈ k • B), B' ≤ k • B := by + intro k hk + exact Set.biInter_subset_of_mem hk + have hfB' : B'.Finite := by + obtain ⟨b, hb : b ∈ B⟩ := hfB_ne + obtain ⟨k, hk : k • b = a⟩ := exists_smul_eq G b a + apply Set.Finite.subset (Set.Finite.map _ hfB) (hB'₀ k ⟨b, hb, hk⟩) + have hag : ∀ g : G, a ∈ g • B' → B' ≤ g • B' := by + intro g hg x hx + -- a = g • b; b ∈ B'; a ∈ k • B → b ∈ k • B + simp only [B', Set.mem_iInter, Set.mem_smul_set_iff_inv_smul_mem, + smul_smul, ← mul_inv_rev] at hg hx ⊢ + exact fun _ ↦ hx _ ∘ hg _ + have hag' (g : G) (hg : a ∈ g • B') : B' = g • B' := by + rw [eq_comm, ← mem_stabilizer_iff, mem_stabilizer_of_finite_iff_le_smul _ hfB'] + exact hag g hg + rw [mk_notempty_one] + intro g hg + rw [← Set.nonempty_iff_ne_empty] at hg + obtain ⟨b : X, hb' : b ∈ g • B', hb : b ∈ B'⟩ := Set.nonempty_def.mp hg + obtain ⟨k : G, hk : k • a = b⟩ := exists_smul_eq G a b + have hak : a ∈ k⁻¹ • B' := by + refine ⟨b, hb, ?_⟩ + simp only [← hk, inv_smul_smul] + have hagk : a ∈ (k⁻¹ * g) • B' := by + rw [mul_smul, Set.mem_smul_set_iff_inv_smul_mem, inv_inv, hk] + exact hb' + have hkB' : B' = k⁻¹ • B' := hag' k⁻¹ hak + have hgkB' : B' = (k⁻¹ * g) • B' := hag' (k⁻¹ * g) hagk + rw [mul_smul] at hgkB' + rw [← smul_eq_iff_eq_inv_smul] at hkB' hgkB' + rw [← hgkB', hkB'] + +end IsBlock + +end Finite + end Group end MulAction diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 99284c35f6952..977fdac89d177 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -3,8 +3,9 @@ Copyright (c) 2021 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ -import Mathlib.Data.Finite.Card import Mathlib.Algebra.BigOperators.GroupWithZero.Finset +import Mathlib.Data.Finite.Card +import Mathlib.Data.Set.Card import Mathlib.GroupTheory.Coset.Card import Mathlib.GroupTheory.Finiteness import Mathlib.GroupTheory.GroupAction.Quotient @@ -30,7 +31,7 @@ Several theorems proved in this file are known as Lagrange's theorem. - `relindex_mul_index` : If `H ≤ K`, then `H.relindex K * K.index = H.index` - `index_dvd_of_le` : If `H ≤ K`, then `K.index ∣ H.index` - `relindex_mul_relindex` : `relindex` is multiplicative in towers - +- `MulAction.index_stabilizer`: the index of the stabilizer is the cardinality of the orbit -/ @@ -558,6 +559,21 @@ end FiniteIndex end Subgroup +namespace MulAction + +variable (G : Type*) {X : Type*} [Group G] [MulAction G X] (x : X) + +theorem index_stabilizer : + (stabilizer G x).index = (orbit G x).ncard := + (Nat.card_congr (MulAction.orbitEquivQuotientStabilizer G x)).symm.trans + (Set.Nat.card_coe_set_eq (orbit G x)) + +theorem index_stabilizer_of_transitive [IsPretransitive G X] : + (stabilizer G x).index = Nat.card X := by + rw [index_stabilizer, orbit_eq_univ, Set.ncard_univ] + +end MulAction + namespace MonoidHom open Finset From 34de7267074d0f2532903e3e72770b04bfa15ae5 Mon Sep 17 00:00:00 2001 From: jouglasheen Date: Thu, 17 Oct 2024 20:35:01 +0000 Subject: [PATCH 035/131] feat: a nonarchimedean group is totally disconnected (#16687) Prove that a nonarchimedean group is a totally disconnected topological space. Co-authored-by: Kevin Buzzard Co-authored-by: jouglasheen <159258007+jouglasheen@users.noreply.github.com> --- Mathlib.lean | 1 + .../Nonarchimedean/TotallyDisconnected.lean | 60 +++++++++++++++++++ 2 files changed, 61 insertions(+) create mode 100644 Mathlib/Topology/Algebra/Nonarchimedean/TotallyDisconnected.lean diff --git a/Mathlib.lean b/Mathlib.lean index 91da140ade63c..d7f45c2607328 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4555,6 +4555,7 @@ import Mathlib.Topology.Algebra.Nonarchimedean.AdicTopology import Mathlib.Topology.Algebra.Nonarchimedean.Bases import Mathlib.Topology.Algebra.Nonarchimedean.Basic import Mathlib.Topology.Algebra.Nonarchimedean.Completion +import Mathlib.Topology.Algebra.Nonarchimedean.TotallyDisconnected import Mathlib.Topology.Algebra.OpenSubgroup import Mathlib.Topology.Algebra.Order.Archimedean import Mathlib.Topology.Algebra.Order.Compact diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/TotallyDisconnected.lean b/Mathlib/Topology/Algebra/Nonarchimedean/TotallyDisconnected.lean new file mode 100644 index 0000000000000..fd36afa2ef8c6 --- /dev/null +++ b/Mathlib/Topology/Algebra/Nonarchimedean/TotallyDisconnected.lean @@ -0,0 +1,60 @@ +/- +Copyright (c) 2024 Jou Glasheen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jou Glasheen, Kevin Buzzard, David Loeffler, Yongle Hu, Johan Commelin +-/ + +import Mathlib.Topology.Algebra.Nonarchimedean.Basic + +/-! +# Total separatedness of nonarchimedean groups + +In this file, we prove that a nonarchimedean group is a totally separated topological space. +The fact that a nonarchimedean group is a totally disconnected topological space +is implied by the fact that a nonarchimedean group is totally separated. + +## Main results + +- `NonarchimedeanGroup.instTotallySeparated`: + A nonarchimedean group is a totally separated topological space. + +## Notation + + - `G` : Is a nonarchimedean group. + - `V` : Is an open subgroup which is a neighbourhood of the identity in `G`. + +## References + +See Proposition 2.3.9 and Problem 63 in [F. Q. Gouvêa, *p-adic numbers*][gouvea1997]. +-/ + +open Pointwise TopologicalSpace + +variable {G : Type*} [TopologicalSpace G] [Group G] [NonarchimedeanGroup G] [T2Space G] + +namespace NonarchimedeanGroup + +@[to_additive] +lemma exists_openSubgroup_separating {a b : G} (h : a ≠ b) : + ∃ (V : OpenSubgroup G), Disjoint (a • (V : Set G)) (b • V) := by + obtain ⟨u, v, _, open_v, mem_u, mem_v, dis⟩ := t2_separation (h ∘ inv_mul_eq_one.mp) + obtain ⟨V, hV⟩ := is_nonarchimedean v (open_v.mem_nhds mem_v) + use V + simp only [Disjoint, Set.le_eq_subset, Set.bot_eq_empty, Set.subset_empty_iff] + intros x mem_aV mem_bV + by_contra! con + obtain ⟨s, hs⟩ := con + have hsa : s ∈ a • (V : Set G) := mem_aV hs + have hsb : s ∈ b • (V : Set G) := mem_bV hs + rw [mem_leftCoset_iff] at hsa hsb + refine dis.subset_compl_right mem_u (hV ?_) + simpa [mul_assoc] using mul_mem hsa (inv_mem hsb) + +@[to_additive] +instance (priority := 100) instTotallySeparated : TotallySeparatedSpace G where + isTotallySeparated_univ x _ y _ hxy := by + obtain ⟨V, dxy⟩ := exists_openSubgroup_separating hxy + exact ⟨_, _, V.isOpen.smul x, (V.isClosed.smul x).isOpen_compl, mem_own_leftCoset .., + dxy.subset_compl_left <| mem_own_leftCoset .., by simp, disjoint_compl_right⟩ + +end NonarchimedeanGroup From 389e09795accaebbea904dccedfed6322feb5b86 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Thu, 17 Oct 2024 20:59:58 +0000 Subject: [PATCH 036/131] chore(Algebra.Order.Monoid.TypeTags): split into unbundled and bundled ordered algebra (#16047) We move the unbundled ordered algebra results out of `Algebra.Order.Monoid.TypeTags` into `Algebra.Order.Monoid.Unbundled.TypeTags` to avoid importing bundled ordered algebra results unless necessary. --- Mathlib.lean | 1 + Mathlib/Algebra/Order/Monoid/ToMulBot.lean | 2 +- Mathlib/Algebra/Order/Monoid/TypeTags.lean | 104 +--------------- .../Order/Monoid/Unbundled/TypeTags.lean | 111 ++++++++++++++++++ Mathlib/Algebra/Order/SuccPred/TypeTags.lean | 2 +- 5 files changed, 116 insertions(+), 104 deletions(-) create mode 100644 Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean diff --git a/Mathlib.lean b/Mathlib.lean index d7f45c2607328..721ac852ea637 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -653,6 +653,7 @@ import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE import Mathlib.Algebra.Order.Monoid.Unbundled.MinMax import Mathlib.Algebra.Order.Monoid.Unbundled.OrderDual import Mathlib.Algebra.Order.Monoid.Unbundled.Pow +import Mathlib.Algebra.Order.Monoid.Unbundled.TypeTags import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop import Mathlib.Algebra.Order.Monoid.Units import Mathlib.Algebra.Order.Monoid.WithTop diff --git a/Mathlib/Algebra/Order/Monoid/ToMulBot.lean b/Mathlib/Algebra/Order/Monoid/ToMulBot.lean index 96f8542fa024c..ab578ee5f849e 100644 --- a/Mathlib/Algebra/Order/Monoid/ToMulBot.lean +++ b/Mathlib/Algebra/Order/Monoid/ToMulBot.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import Mathlib.Algebra.Order.GroupWithZero.Canonical -import Mathlib.Algebra.Order.Monoid.TypeTags +import Mathlib.Algebra.Order.Monoid.Unbundled.TypeTags import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop diff --git a/Mathlib/Algebra/Order/Monoid/TypeTags.lean b/Mathlib/Algebra/Order/Monoid/TypeTags.lean index c8a194531eb5b..31b5a417097b5 100644 --- a/Mathlib/Algebra/Order/Monoid/TypeTags.lean +++ b/Mathlib/Algebra/Order/Monoid/TypeTags.lean @@ -3,62 +3,13 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ -import Mathlib.Algebra.Group.TypeTags -import Mathlib.Algebra.Order.Monoid.Defs +import Mathlib.Algebra.Order.Monoid.Unbundled.TypeTags import Mathlib.Algebra.Order.Monoid.Canonical.Defs -/-! # Ordered monoid structures on `Multiplicative α` and `Additive α`. -/ +/-! # Bundled ordered monoid structures on `Multiplicative α` and `Additive α`. -/ variable {α : Type*} -instance : ∀ [LE α], LE (Multiplicative α) := - fun {inst} => inst - -instance : ∀ [LE α], LE (Additive α) := - fun {inst} => inst - -instance : ∀ [LT α], LT (Multiplicative α) := - fun {inst} => inst - -instance : ∀ [LT α], LT (Additive α) := - fun {inst} => inst - -instance Multiplicative.preorder : ∀ [Preorder α], Preorder (Multiplicative α) := - fun {inst} => inst - -instance Additive.preorder : ∀ [Preorder α], Preorder (Additive α) := - fun {inst} => inst - -instance Multiplicative.partialOrder : ∀ [PartialOrder α], PartialOrder (Multiplicative α) := - fun {inst} => inst - -instance Additive.partialOrder : ∀ [PartialOrder α], PartialOrder (Additive α) := - fun {inst} => inst - -instance Multiplicative.linearOrder : ∀ [LinearOrder α], LinearOrder (Multiplicative α) := - fun {inst} => inst - -instance Additive.linearOrder : ∀ [LinearOrder α], LinearOrder (Additive α) := - fun {inst} => inst - -instance Multiplicative.orderBot [LE α] : ∀ [OrderBot α], OrderBot (Multiplicative α) := - fun {inst} => inst - -instance Additive.orderBot [LE α] : ∀ [OrderBot α], OrderBot (Additive α) := - fun {inst} => inst - -instance Multiplicative.orderTop [LE α] : ∀ [OrderTop α], OrderTop (Multiplicative α) := - fun {inst} => inst - -instance Additive.orderTop [LE α] : ∀ [OrderTop α], OrderTop (Additive α) := - fun {inst} => inst - -instance Multiplicative.boundedOrder [LE α] : ∀ [BoundedOrder α], BoundedOrder (Multiplicative α) := - fun {inst} => inst - -instance Additive.boundedOrder [LE α] : ∀ [BoundedOrder α], BoundedOrder (Additive α) := - fun {inst} => inst - instance Multiplicative.orderedCommMonoid [OrderedAddCommMonoid α] : OrderedCommMonoid (Multiplicative α) := { Multiplicative.partialOrder, Multiplicative.commMonoid with @@ -87,13 +38,6 @@ instance Additive.linearOrderedAddCommMonoid [LinearOrderedCommMonoid α] : LinearOrderedAddCommMonoid (Additive α) := { Additive.linearOrder, Additive.orderedAddCommMonoid with } -instance Multiplicative.existsMulOfLe [Add α] [LE α] [ExistsAddOfLE α] : - ExistsMulOfLE (Multiplicative α) := - ⟨@exists_add_of_le α _ _ _⟩ - -instance Additive.existsAddOfLe [Mul α] [LE α] [ExistsMulOfLE α] : ExistsAddOfLE (Additive α) := - ⟨@exists_mul_of_le α _ _ _⟩ - instance Multiplicative.canonicallyOrderedCommMonoid [CanonicallyOrderedAddCommMonoid α] : CanonicallyOrderedCommMonoid (Multiplicative α) := { Multiplicative.orderedCommMonoid, Multiplicative.orderBot, @@ -112,47 +56,3 @@ instance Multiplicative.canonicallyLinearOrderedCommMonoid instance [CanonicallyLinearOrderedCommMonoid α] : CanonicallyLinearOrderedAddCommMonoid (Additive α) := { Additive.canonicallyOrderedAddCommMonoid, Additive.linearOrder with } - -namespace Additive - -variable [Preorder α] - -@[simp] -theorem ofMul_le {a b : α} : ofMul a ≤ ofMul b ↔ a ≤ b := - Iff.rfl - -@[simp] -theorem ofMul_lt {a b : α} : ofMul a < ofMul b ↔ a < b := - Iff.rfl - -@[simp] -theorem toMul_le {a b : Additive α} : toMul a ≤ toMul b ↔ a ≤ b := - Iff.rfl - -@[simp] -theorem toMul_lt {a b : Additive α} : toMul a < toMul b ↔ a < b := - Iff.rfl - -end Additive - -namespace Multiplicative - -variable [Preorder α] - -@[simp] -theorem ofAdd_le {a b : α} : ofAdd a ≤ ofAdd b ↔ a ≤ b := - Iff.rfl - -@[simp] -theorem ofAdd_lt {a b : α} : ofAdd a < ofAdd b ↔ a < b := - Iff.rfl - -@[simp] -theorem toAdd_le {a b : Multiplicative α} : toAdd a ≤ toAdd b ↔ a ≤ b := - Iff.rfl - -@[simp] -theorem toAdd_lt {a b : Multiplicative α} : toAdd a < toAdd b ↔ a < b := - Iff.rfl - -end Multiplicative diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean new file mode 100644 index 0000000000000..1647d43c6a434 --- /dev/null +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean @@ -0,0 +1,111 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE +import Mathlib.Order.BoundedOrder + +/-! # Ordered monoid structures on `Multiplicative α` and `Additive α`. -/ + +variable {α : Type*} + +instance : ∀ [LE α], LE (Multiplicative α) := + fun {inst} => inst + +instance : ∀ [LE α], LE (Additive α) := + fun {inst} => inst + +instance : ∀ [LT α], LT (Multiplicative α) := + fun {inst} => inst + +instance : ∀ [LT α], LT (Additive α) := + fun {inst} => inst + +instance Multiplicative.preorder : ∀ [Preorder α], Preorder (Multiplicative α) := + fun {inst} => inst + +instance Additive.preorder : ∀ [Preorder α], Preorder (Additive α) := + fun {inst} => inst + +instance Multiplicative.partialOrder : ∀ [PartialOrder α], PartialOrder (Multiplicative α) := + fun {inst} => inst + +instance Additive.partialOrder : ∀ [PartialOrder α], PartialOrder (Additive α) := + fun {inst} => inst + +instance Multiplicative.linearOrder : ∀ [LinearOrder α], LinearOrder (Multiplicative α) := + fun {inst} => inst + +instance Additive.linearOrder : ∀ [LinearOrder α], LinearOrder (Additive α) := + fun {inst} => inst + +instance Multiplicative.orderBot [LE α] : ∀ [OrderBot α], OrderBot (Multiplicative α) := + fun {inst} => inst + +instance Additive.orderBot [LE α] : ∀ [OrderBot α], OrderBot (Additive α) := + fun {inst} => inst + +instance Multiplicative.orderTop [LE α] : ∀ [OrderTop α], OrderTop (Multiplicative α) := + fun {inst} => inst + +instance Additive.orderTop [LE α] : ∀ [OrderTop α], OrderTop (Additive α) := + fun {inst} => inst + +instance Multiplicative.boundedOrder [LE α] : ∀ [BoundedOrder α], BoundedOrder (Multiplicative α) := + fun {inst} => inst + +instance Additive.boundedOrder [LE α] : ∀ [BoundedOrder α], BoundedOrder (Additive α) := + fun {inst} => inst + +instance Multiplicative.existsMulOfLe [Add α] [LE α] [ExistsAddOfLE α] : + ExistsMulOfLE (Multiplicative α) := + ⟨@exists_add_of_le α _ _ _⟩ + +instance Additive.existsAddOfLe [Mul α] [LE α] [ExistsMulOfLE α] : ExistsAddOfLE (Additive α) := + ⟨@exists_mul_of_le α _ _ _⟩ + +namespace Additive + +variable [Preorder α] + +@[simp] +theorem ofMul_le {a b : α} : ofMul a ≤ ofMul b ↔ a ≤ b := + Iff.rfl + +@[simp] +theorem ofMul_lt {a b : α} : ofMul a < ofMul b ↔ a < b := + Iff.rfl + +@[simp] +theorem toMul_le {a b : Additive α} : toMul a ≤ toMul b ↔ a ≤ b := + Iff.rfl + +@[simp] +theorem toMul_lt {a b : Additive α} : toMul a < toMul b ↔ a < b := + Iff.rfl + +end Additive + +namespace Multiplicative + +variable [Preorder α] + +@[simp] +theorem ofAdd_le {a b : α} : ofAdd a ≤ ofAdd b ↔ a ≤ b := + Iff.rfl + +@[simp] +theorem ofAdd_lt {a b : α} : ofAdd a < ofAdd b ↔ a < b := + Iff.rfl + +@[simp] +theorem toAdd_le {a b : Multiplicative α} : toAdd a ≤ toAdd b ↔ a ≤ b := + Iff.rfl + +@[simp] +theorem toAdd_lt {a b : Multiplicative α} : toAdd a < toAdd b ↔ a < b := + Iff.rfl + +end Multiplicative diff --git a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean index 4293cd851694a..7c69e6dc40162 100644 --- a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean +++ b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean @@ -3,8 +3,8 @@ 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.Order.Monoid.TypeTags import Mathlib.Order.SuccPred.Archimedean +import Mathlib.Algebra.Order.Monoid.Unbundled.TypeTags /-! # Successor and predecessor on type tags From 23a1a4bd2117e0e4a6391122bcf08d7c2edf6008 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Thu, 17 Oct 2024 22:32:03 +0000 Subject: [PATCH 037/131] feat(AlgebraicGeometry): proper morphisms of schemes (#17863) We define proper morphisms of schemes and show standard stability properties. Partly from the valuative criterion project. Co-authored by: Andrew Yang --- Mathlib.lean | 1 + .../Morphisms/ClosedImmersion.lean | 14 ++++ .../Morphisms/FiniteType.lean | 3 + .../AlgebraicGeometry/Morphisms/Proper.lean | 69 +++++++++++++++++++ .../Morphisms/UniversallyClosed.lean | 12 +++- 5 files changed, 98 insertions(+), 1 deletion(-) create mode 100644 Mathlib/AlgebraicGeometry/Morphisms/Proper.lean diff --git a/Mathlib.lean b/Mathlib.lean index 721ac852ea637..6871945d645ea 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -868,6 +868,7 @@ import Mathlib.AlgebraicGeometry.Morphisms.FiniteType import Mathlib.AlgebraicGeometry.Morphisms.IsIso import Mathlib.AlgebraicGeometry.Morphisms.OpenImmersion import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion +import Mathlib.AlgebraicGeometry.Morphisms.Proper import Mathlib.AlgebraicGeometry.Morphisms.QuasiCompact import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index c78e3b86d608d..932325058e426 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -6,6 +6,7 @@ Authors: Amelia Livingston, Christian Merten, Jonas van der Schaaf import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties +import Mathlib.AlgebraicGeometry.Morphisms.FiniteType import Mathlib.AlgebraicGeometry.ResidueField import Mathlib.RingTheory.RingHom.Surjective @@ -269,4 +270,17 @@ lemma IsClosedImmersion.stableUnderBaseChange : exact ⟨inferInstance, RingHom.surjective_stableUnderBaseChange.pullback_fst_app_top _ RingHom.surjective_respectsIso f _ hsurj⟩ +/-- Closed immersions are locally of finite type. -/ +instance (priority := 900) {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsClosedImmersion f] : + LocallyOfFiniteType f := by + wlog hY : IsAffine Y + · rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := @LocallyOfFiniteType) _ + (iSup_affineOpens_eq_top Y)] + intro U + have H : IsClosedImmersion (f ∣_ U) := IsLocalAtTarget.restrict h U + exact this _ U.2 + obtain ⟨_, hf⟩ := h.isAffine_surjective_of_isAffine + rw [HasRingHomProperty.iff_of_isAffine (P := @LocallyOfFiniteType)] + exact RingHom.FiniteType.of_surjective (Scheme.Hom.app f ⊤) hf + end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean b/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean index def22435c69c0..a5bbcc0abb5f2 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean @@ -59,6 +59,9 @@ theorem locallyOfFiniteType_of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [LocallyOfFiniteType (f ≫ g)] : LocallyOfFiniteType f := HasRingHomProperty.of_comp (fun _ _ ↦ RingHom.FiniteType.of_comp_finiteType) ‹_› +instance : MorphismProperty.IsMultiplicative @LocallyOfFiniteType where + id_mem _ := inferInstance + open scoped TensorProduct in lemma locallyOfFiniteType_stableUnderBaseChange : MorphismProperty.StableUnderBaseChange @LocallyOfFiniteType := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Proper.lean b/Mathlib/AlgebraicGeometry/Morphisms/Proper.lean new file mode 100644 index 0000000000000..b54565bfb3af8 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Morphisms/Proper.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2024 Christian Merten, Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten, Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.Morphisms.Separated +import Mathlib.AlgebraicGeometry.Morphisms.UniversallyClosed +import Mathlib.AlgebraicGeometry.Morphisms.FiniteType + +/-! + +# Proper morphisms + +A morphism of schemes is proper if it is separated, universally closed and (locally) of finite type. +Note that we don't require quasi-compact, since this is implied by universally closed (TODO). + +-/ + + +noncomputable section + +open CategoryTheory + +universe u + +namespace AlgebraicGeometry + +variable {X Y : Scheme.{u}} (f : X ⟶ Y) + +/-- A morphism is proper if it is separated, universally closed and locally of finite type. -/ +@[mk_iff] +class IsProper extends IsSeparated f, UniversallyClosed f, LocallyOfFiniteType f : Prop where + +lemma isProper_eq : @IsProper = + (@IsSeparated ⊓ @UniversallyClosed : MorphismProperty Scheme) ⊓ @LocallyOfFiniteType := by + ext X Y f + rw [isProper_iff, ← and_assoc] + rfl + +namespace IsProper + +instance : MorphismProperty.RespectsIso @IsProper := by + rw [isProper_eq] + infer_instance + +instance stableUnderComposition : MorphismProperty.IsStableUnderComposition @IsProper := by + rw [isProper_eq] + infer_instance + +instance : MorphismProperty.IsMultiplicative @IsProper := by + rw [isProper_eq] + infer_instance + +instance (priority := 900) [IsClosedImmersion f] : IsProper f where + +lemma stableUnderBaseChange : MorphismProperty.StableUnderBaseChange @IsProper := by + rw [isProper_eq] + exact MorphismProperty.StableUnderBaseChange.inf + (MorphismProperty.StableUnderBaseChange.inf + IsSeparated.stableUnderBaseChange universallyClosed_stableUnderBaseChange) + locallyOfFiniteType_stableUnderBaseChange + +instance : IsLocalAtTarget @IsProper := by + rw [isProper_eq] + infer_instance + +end IsProper + +end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean index 6a5489d992476..166651638bcaa 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.AlgebraicGeometry.Morphisms.Constructors +import Mathlib.AlgebraicGeometry.Morphisms.ClosedImmersion import Mathlib.Topology.LocalAtTarget /-! @@ -40,6 +40,13 @@ class UniversallyClosed (f : X ⟶ Y) : Prop where theorem universallyClosed_eq : @UniversallyClosed = universally (topologically @IsClosedMap) := by ext X Y f; rw [universallyClosed_iff] +instance (priority := 900) [IsClosedImmersion f] : UniversallyClosed f := by + rw [universallyClosed_eq] + intro X' Y' i₁ i₂ f' hf + have hf' : IsClosedImmersion f' := + IsClosedImmersion.stableUnderBaseChange hf.flip inferInstance + exact hf'.base_closed.isClosedMap + theorem universallyClosed_respectsIso : RespectsIso @UniversallyClosed := universallyClosed_eq.symm ▸ universally_respectsIso (topologically @IsClosedMap) @@ -59,6 +66,9 @@ instance universallyClosedTypeComp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [hf : UniversallyClosed f] [hg : UniversallyClosed g] : UniversallyClosed (f ≫ g) := comp_mem _ _ _ hf hg +instance : MorphismProperty.IsMultiplicative @UniversallyClosed where + id_mem _ := inferInstance + instance universallyClosed_fst {X Y Z : Scheme} (f : X ⟶ Z) (g : Y ⟶ Z) [hg : UniversallyClosed g] : UniversallyClosed (pullback.fst f g) := universallyClosed_stableUnderBaseChange.fst f g hg From f1ebf24bc990d6d810d4780353eb74d9ecb39803 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 18 Oct 2024 00:49:50 +0000 Subject: [PATCH 038/131] refactor: add `compactSpace_spectrum` and `spectrum_nonempty` fields to the `ContinuousFunctionalCalculus` (#17801) --- Mathlib/Algebra/Algebra/Quasispectrum.lean | 3 +++ .../Instances.lean | 1 + .../NonUnital.lean | 21 +++++++++++++------ .../ContinuousFunctionalCalculus/Order.lean | 12 +++++------ .../Restrict.lean | 10 +++++++++ .../ContinuousFunctionalCalculus/Unital.lean | 7 +++++++ .../ContinuousFunctionalCalculus/Rpow.lean | 8 ++++--- .../Matrix/HermitianFunctionalCalculus.lean | 5 +++++ 8 files changed, 51 insertions(+), 16 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Quasispectrum.lean b/Mathlib/Algebra/Algebra/Quasispectrum.lean index 33a5807e877b8..e94372dc43fdc 100644 --- a/Mathlib/Algebra/Algebra/Quasispectrum.lean +++ b/Mathlib/Algebra/Algebra/Quasispectrum.lean @@ -254,6 +254,9 @@ lemma quasispectrum.not_isUnit_mem (a : A) {r : R} (hr : ¬ IsUnit r) : r ∈ qu lemma quasispectrum.zero_mem [Nontrivial R] (a : A) : 0 ∈ quasispectrum R a := quasispectrum.not_isUnit_mem a <| by simp +theorem quasispectrum.nonempty [Nontrivial R] (a : A) : (quasispectrum R a).Nonempty := + Set.nonempty_of_mem <| quasispectrum.zero_mem R a + instance quasispectrum.instZero [Nontrivial R] (a : A) : Zero (quasispectrum R a) where zero := ⟨0, quasispectrum.zero_mem R a⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index 50abcebd079d7..ea67a1f84c970 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -155,6 +155,7 @@ instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [NormedRing A [CStarRing A] [CompleteSpace A] [NormedAlgebra ℂ A] [StarModule ℂ A] : ContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop) where predicate_zero := isStarNormal_zero + spectrum_nonempty a _ := spectrum.nonempty a exists_cfc_of_predicate a ha := by refine ⟨(elementalStarAlgebra ℂ a).subtype.comp <| continuousFunctionalCalculus a, ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index 361cdc8aefebc..d4a2d892c661e 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -73,11 +73,16 @@ class NonUnitalContinuousFunctionalCalculus (R : Type*) {A : Type*} (p : outPara [ContinuousStar R] [NonUnitalRing A] [StarRing A] [TopologicalSpace A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] : Prop where predicate_zero : p 0 + [compactSpace_quasispectrum : ∀ a : A, CompactSpace (σₙ R a)] exists_cfc_of_predicate : ∀ a, p a → ∃ φ : C(σₙ R a, R)₀ →⋆ₙₐ[R] A, ClosedEmbedding φ ∧ φ ⟨(ContinuousMap.id R).restrict <| σₙ R a, rfl⟩ = a ∧ (∀ f, σₙ R (φ f) = Set.range f) ∧ ∀ f, p (φ f) --- TODO: try to unify with the unital version. The `ℝ≥0` case makes it tricky. +-- this instance should not be activated everywhere but it is useful when developing generic API +-- for the continuous functional calculus +scoped[NonUnitalContinuousFunctionalCalculus] +attribute [instance] NonUnitalContinuousFunctionalCalculus.compactSpace_quasispectrum + /-- A class guaranteeing that the non-unital continuous functional calculus is uniquely determined by the properties that it is a continuous non-unital star algebra homomorphism mapping the (restriction of) the identity to `a`. This is the necessary tool used to establish `cfcₙHom_comp` @@ -685,7 +690,10 @@ lemma cfcₙHom_of_cfcHom_map_quasispectrum {a : A} (ha : p a) : simp only [Set.mem_singleton_iff] at hx ⊢ rw [show x = 0 from Subtype.val_injective hx, map_zero] -variable [CompleteSpace R] [h_cpct : ∀ a : A, CompactSpace (spectrum R a)] +variable [CompleteSpace R] + +-- gives access to the `ContinuousFunctionalCalculus.compactSpace_spectrum` instance +open scoped ContinuousFunctionalCalculus lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : ClosedEmbedding (cfcₙHom_of_cfcHom R ha) := by @@ -709,6 +717,10 @@ lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : instance ContinuousFunctionalCalculus.toNonUnital : NonUnitalContinuousFunctionalCalculus R p where predicate_zero := cfc_predicate_zero R + compactSpace_quasispectrum a := by + have h_cpct : CompactSpace (spectrum R a) := inferInstance + simp only [← isCompact_iff_compactSpace, quasispectrum_eq_spectrum_union_zero] at h_cpct ⊢ + exact h_cpct |>.union isCompact_singleton exists_cfc_of_predicate _ ha := ⟨cfcₙHom_of_cfcHom R ha, closedEmbedding_cfcₙHom_of_cfcHom ha, @@ -716,12 +728,9 @@ instance ContinuousFunctionalCalculus.toNonUnital : NonUnitalContinuousFunctiona cfcₙHom_of_cfcHom_map_quasispectrum ha, fun _ ↦ cfcHom_predicate ha _⟩ +open scoped NonUnitalContinuousFunctionalCalculus in lemma cfcₙHom_eq_cfcₙHom_of_cfcHom [UniqueNonUnitalContinuousFunctionalCalculus R A] {a : A} (ha : p a) : cfcₙHom (R := R) ha = cfcₙHom_of_cfcHom R ha := by - have h_cpct' : CompactSpace (σₙ R a) := by - specialize h_cpct a - simp_rw [← isCompact_iff_compactSpace, quasispectrum_eq_spectrum_union_zero] at h_cpct ⊢ - exact h_cpct.union isCompact_singleton refine UniqueNonUnitalContinuousFunctionalCalculus.eq_of_continuous_of_map_id (σₙ R a) ?_ _ _ ?_ ?_ ?_ · simp diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index 8536353d9475b..10e0978f35a5c 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -85,20 +85,20 @@ lemma cfc_nnreal_le_iff {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] [ rw [cfc_nnreal_eq_real, cfc_nnreal_eq_real, cfc_le_iff ..] simp [NNReal.coe_le_coe, ← ha_spec.image] +open ContinuousFunctionalCalculus in /-- In a unital `ℝ`-algebra `A` with a continuous functional calculus, an element `a : A` is larger than some `algebraMap ℝ A r` if and only if every element of the `ℝ`-spectrum is nonnegative. -/ lemma CFC.exists_pos_algebraMap_le_iff {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] - [PartialOrder A] [StarOrderedRing A] [Algebra ℝ A] [NonnegSpectrumClass ℝ A] + [PartialOrder A] [StarOrderedRing A] [Algebra ℝ A] [NonnegSpectrumClass ℝ A] [Nontrivial A] [ContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] - {a : A} [CompactSpace (spectrum ℝ a)] - (h_non : (spectrum ℝ a).Nonempty) (ha : IsSelfAdjoint a := by cfc_tac) : + {a : A} (ha : IsSelfAdjoint a := by cfc_tac) : (∃ r > 0, algebraMap ℝ A r ≤ a) ↔ (∀ x ∈ spectrum ℝ a, 0 < x) := by have h_cpct : IsCompact (spectrum ℝ a) := isCompact_iff_compactSpace.mpr inferInstance simp_rw [algebraMap_le_iff_le_spectrum (a := a)] refine ⟨?_, fun h ↦ ?_⟩ · rintro ⟨r, hr, hr_le⟩ exact (hr.trans_le <| hr_le · ·) - · obtain ⟨r, hr, hr_min⟩ := h_cpct.exists_isMinOn h_non continuousOn_id + · obtain ⟨r, hr, hr_min⟩ := h_cpct.exists_isMinOn (spectrum_nonempty ℝ a ha) continuousOn_id exact ⟨r, h _ hr, hr_min⟩ section CStar_unital @@ -218,10 +218,8 @@ lemma CStarAlgebra.isUnit_of_le {a b : A} (h₀ : IsUnit a) (ha : 0 ≤ a := by rw [← spectrum.zero_not_mem_iff ℝ≥0] at h₀ ⊢ nontriviality A have hb := (show 0 ≤ a from ha).trans hab - have ha' := IsSelfAdjoint.of_nonneg ha |>.spectrum_nonempty - have hb' := IsSelfAdjoint.of_nonneg hb |>.spectrum_nonempty rw [zero_not_mem_iff, SpectrumRestricts.nnreal_lt_iff (.nnreal_of_nonneg ‹_›), - NNReal.coe_zero, ← CFC.exists_pos_algebraMap_le_iff ‹_›] at h₀ ⊢ + NNReal.coe_zero, ← CFC.exists_pos_algebraMap_le_iff] at h₀ ⊢ peel h₀ with r hr _ exact this.trans hab diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean index cf36028a7922e..bb43228df3fdf 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean @@ -98,6 +98,12 @@ protected theorem cfc (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) (h : ∀ a, p a ↔ q a ∧ SpectrumRestricts a f) : ContinuousFunctionalCalculus R p where predicate_zero := h0 + spectrum_nonempty a ha := ((h a).mp ha).2.image ▸ + (ContinuousFunctionalCalculus.spectrum_nonempty a ((h a).mp ha).1 |>.image f) + compactSpace_spectrum a := by + have := ContinuousFunctionalCalculus.compactSpace_spectrum (R := S) a + rw [← isCompact_iff_compactSpace] at this ⊢ + simpa using halg.toClosedEmbedding.isCompact_preimage this exists_cfc_of_predicate a ha := by refine ⟨((h a).mp ha).2.starAlgHom (cfcHom ((h a).mp ha).1 (R := S)), ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ @@ -235,6 +241,10 @@ protected theorem cfc (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) (h : ∀ a, p a ↔ q a ∧ QuasispectrumRestricts a f) : NonUnitalContinuousFunctionalCalculus R p where predicate_zero := h0 + compactSpace_quasispectrum a := by + have := NonUnitalContinuousFunctionalCalculus.compactSpace_quasispectrum (R := S) a + rw [← isCompact_iff_compactSpace] at this ⊢ + simpa using halg.toClosedEmbedding.isCompact_preimage this exists_cfc_of_predicate a ha := by refine ⟨((h a).mp ha).2.nonUnitalStarAlgHom (cfcₙHom ((h a).mp ha).1 (R := S)), ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index 8b6473cba0de3..29fa35441c946 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -162,10 +162,17 @@ class ContinuousFunctionalCalculus (R : Type*) {A : Type*} (p : outParam (A → [CommSemiring R] [StarRing R] [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] [Ring A] [StarRing A] [TopologicalSpace A] [Algebra R A] : Prop where predicate_zero : p 0 + [compactSpace_spectrum (a : A) : CompactSpace (spectrum R a)] + spectrum_nonempty [Nontrivial A] (a : A) (ha : p a) : (spectrum R a).Nonempty exists_cfc_of_predicate : ∀ a, p a → ∃ φ : C(spectrum R a, R) →⋆ₐ[R] A, ClosedEmbedding φ ∧ φ ((ContinuousMap.id R).restrict <| spectrum R a) = a ∧ (∀ f, spectrum R (φ f) = Set.range f) ∧ ∀ f, p (φ f) +-- this instance should not be activated everywhere but it is useful when developing generic API +-- for the continuous functional calculus +scoped[ContinuousFunctionalCalculus] +attribute [instance] ContinuousFunctionalCalculus.compactSpace_spectrum + /-- A class guaranteeing that the continuous functional calculus is uniquely determined by the properties that it is a continuous star algebra homomorphism mapping the (restriction of) the identity to `a`. This is the necessary tool used to establish `cfcHom_comp` and the more common diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean index 9fcf8e77c60a6..fda3885c78a00 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean @@ -5,8 +5,8 @@ Authors: Frédéric Dupuis -/ import Mathlib.Analysis.SpecialFunctions.Pow.Real -import Mathlib.Analysis.Normed.Algebra.Spectrum import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital +import Mathlib.Analysis.SpecialFunctions.Pow.Continuity /-! # Real powers defined via the continuous functional calculus @@ -306,8 +306,10 @@ lemma rpow_intCast (a : Aˣ) (n : ℤ) (ha : (0 : A) ≤ a := by cfc_tac) : section unital_vs_nonunital -variable [∀ (a : A), CompactSpace (spectrum ℝ a)] - [UniqueNonUnitalContinuousFunctionalCalculus ℝ≥0 A] +variable [UniqueNonUnitalContinuousFunctionalCalculus ℝ≥0 A] + +-- provides instance `ContinuousFunctionalCalculus.compactSpace_spectrum` +open scoped ContinuousFunctionalCalculus lemma nnrpow_eq_rpow {a : A} {x : ℝ≥0} (hx : 0 < x) : a ^ x = a ^ (x : ℝ) := by rw [nnrpow_def (A := A), rpow_def, cfcₙ_eq_cfc] diff --git a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean index ff884d2e289ad..5c9c8e5757541 100644 --- a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean +++ b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean @@ -135,6 +135,11 @@ instance instContinuousFunctionalCalculus : rw [star_eq_conjTranspose, diagonal_conjTranspose] congr! simp [Pi.star_def, Function.comp_def] + spectrum_nonempty a ha := by + obtain (h | h) := isEmpty_or_nonempty n + · obtain ⟨x, y, hxy⟩ := exists_pair_ne (Matrix n n 𝕜) + exact False.elim <| Matrix.of.symm.injective.ne hxy <| Subsingleton.elim _ _ + · exact eigenvalues_eq_spectrum_real ha ▸ Set.range_nonempty _ predicate_zero := .zero _ instance instUniqueContinuousFunctionalCalculus : From 3f9cdcd1db12e4b662c93305015ce28d909918d8 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Fri, 18 Oct 2024 01:51:14 +0000 Subject: [PATCH 039/131] feat(Tactic/Linter/Lint): improve `dupNamespace` linter (#17631) Refactor the the way the ids are extracted for the `dupNamespace` linter. We also extract this to a new file in anticipation of re-using that function in another linter. The new version finds auto-generated names, such as `to_additive` names, but ignores "internal" names. Like the previous linter, it also supports `export` and `alias`. Co-authored-by: adomani Co-authored-by: damiano Co-authored-by: Moritz Firsching --- Mathlib.lean | 1 + Mathlib/Tactic.lean | 1 + Mathlib/Tactic/DeclarationNames.lean | 45 ++++++++++++++++++++++++++++ Mathlib/Tactic/Linter/Lint.lean | 32 ++++++++------------ test/Lint.lean | 28 +++++++++++++++-- 5 files changed, 85 insertions(+), 22 deletions(-) create mode 100644 Mathlib/Tactic/DeclarationNames.lean diff --git a/Mathlib.lean b/Mathlib.lean index 6871945d645ea..8963ee37e6ec3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4298,6 +4298,7 @@ import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Conv import Mathlib.Tactic.Convert import Mathlib.Tactic.Core +import Mathlib.Tactic.DeclarationNames import Mathlib.Tactic.DefEqTransformations import Mathlib.Tactic.DeprecateMe import Mathlib.Tactic.DeriveFintype diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index c8d7b7ae27494..cd494193a67b4 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -62,6 +62,7 @@ import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Conv import Mathlib.Tactic.Convert import Mathlib.Tactic.Core +import Mathlib.Tactic.DeclarationNames import Mathlib.Tactic.DefEqTransformations import Mathlib.Tactic.DeprecateMe import Mathlib.Tactic.DeriveFintype diff --git a/Mathlib/Tactic/DeclarationNames.lean b/Mathlib/Tactic/DeclarationNames.lean new file mode 100644 index 0000000000000..3003bab4db21d --- /dev/null +++ b/Mathlib/Tactic/DeclarationNames.lean @@ -0,0 +1,45 @@ +/- +Copyright (c) 2024 Moritz Firsching. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa, Moritz Firsching +-/ + +import Lean.DeclarationRange +import Lean.ResolveName + +/-! +This file contains functions that are used by multiple linters. +-/ + +open Lean Parser Elab Command Meta +namespace Mathlib.Linter + +/-- +If `pos` is a `String.Pos`, then `getNamesFrom pos` returns the array of identifiers +for the names of the declarations whose syntax begins in position at least `pos`. +-/ +def getNamesFrom {m} [Monad m] [MonadEnv m] [MonadFileMap m] (pos : String.Pos) : + m (Array Syntax) := do + let drs := declRangeExt.getState (← getEnv) + let fm ← getFileMap + let mut nms := #[] + for (nm, rgs) in drs do + if pos ≤ fm.ofPosition rgs.range.pos then + let ofPos1 := fm.ofPosition rgs.selectionRange.pos + let ofPos2 := fm.ofPosition rgs.selectionRange.endPos + nms := nms.push (mkIdentFrom (.ofRange ⟨ofPos1, ofPos2⟩) nm) + return nms + +/-- +If `stx` is a syntax node for an `export` statement, then `getAliasSyntax stx` returns the array of +identifiers with the "exported" names. +-/ +def getAliasSyntax {m} [Monad m] [MonadResolveName m] (stx : Syntax) : m (Array Syntax) := do + let mut aliases := #[] + if let `(export $_ ($ids*)) := stx then + let currNamespace ← getCurrNamespace + for idStx in ids do + let id := idStx.getId + aliases := aliases.push + (mkIdentFrom (.ofRange (idStx.raw.getRange?.getD default)) (currNamespace ++ id)) + return aliases diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index 4c0a8fb5df9a4..8641d3a70bb7d 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ import Batteries.Tactic.Lint +import Mathlib.Tactic.DeclarationNames /-! # Linters for Mathlib @@ -77,29 +78,20 @@ namespace DupNamespaceLinter open Lean Parser Elab Command Meta -/-- `getIds stx` extracts the `declId` nodes from the `Syntax` `stx`. -If `stx` is an `alias` or an `export`, then it extracts an `ident`, instead of a `declId`. -/ -partial -def getIds : Syntax → Array Syntax - | .node _ `Batteries.Tactic.Alias.alias args => args[2:3] - | .node _ ``Lean.Parser.Command.export args => (args[3:4] : Array Syntax).map (·[0]) - | stx@(.node _ _ args) => - ((args.attach.map fun ⟨a, _⟩ ↦ getIds a).foldl (· ++ ·) #[stx]).filter (·.getKind == ``declId) - | _ => default - @[inherit_doc linter.dupNamespace] def dupNamespace : Linter where run := withSetOptionIn fun stx ↦ do if Linter.getLinterValue linter.dupNamespace (← getOptions) then - match getIds stx with - | #[id] => - let ns := (← getScope).currNamespace - let declName := ns ++ (if id.getKind == ``declId then id[0].getId else id.getId) - let nm := declName.components - let some (dup, _) := nm.zip (nm.tailD []) |>.find? fun (x, y) ↦ x == y - | return - Linter.logLint linter.dupNamespace id - m!"The namespace '{dup}' is duplicated in the declaration '{declName}'" - | _ => return + let mut aliases := #[] + if let some exp := stx.find? (·.isOfKind `Lean.Parser.Command.export) then + aliases ← getAliasSyntax exp + for id in (← getNamesFrom (stx.getPos?.getD default)) ++ aliases do + let declName := id.getId + if declName.hasMacroScopes then continue + let nm := declName.components + let some (dup, _) := nm.zip (nm.tailD []) |>.find? fun (x, y) ↦ x == y + | continue + Linter.logLint linter.dupNamespace id + m!"The namespace '{dup}' is duplicated in the declaration '{declName}'" initialize addLinter dupNamespace diff --git a/test/Lint.lean b/test/Lint.lean index 25bf9c9147d13..7f6b2c917ddfb 100644 --- a/test/Lint.lean +++ b/test/Lint.lean @@ -7,6 +7,23 @@ import Mathlib.Order.SetNotation set_option linter.dupNamespace false +namespace termG +-- this creates a hygienic declaration starting with `termG.termG.«_@».test.Lint...` +-- and the linter ignores it +set_option linter.dupNamespace true in +local notation "G" => Unit + +/-- info: [termG, termG] -/ +#guard_msgs in +open Lean in +run_meta + let env ← getEnv + let consts := env.constants.toList.find? (·.1.getRoot == `termG) + let reps := (consts.map (·.1.components.take 2)).getD default + logInfo m!"{reps}" + guard (reps[0]! == reps[1]!) +end termG + /-- warning: The namespace 'add' is duplicated in the declaration 'add.add' note: this linter can be disabled with `set_option linter.dupNamespace false` @@ -25,8 +42,12 @@ note: this linter can be disabled with `set_option linter.dupNamespace false` set_option linter.dupNamespace true in def Foo.foo := True --- the `dupNamespace` linter does not notice that `to_additive` created `Foo.add.add`. +/-- +warning: The namespace 'add' is duplicated in the declaration 'Foo.add.add' +note: this linter can be disabled with `set_option linter.dupNamespace false` +-/ #guard_msgs in +set_option linter.dupNamespace true in @[to_additive] theorem add.mul : True := .intro -- However, the declaration `Foo.add.add` is present in the environment. @@ -52,10 +73,13 @@ namespace add /-- warning: The namespace 'add' is duplicated in the declaration 'add.add' note: this linter can be disabled with `set_option linter.dupNamespace false` +--- +warning: The namespace 'add' is duplicated in the declaration 'add.add' +note: this linter can be disabled with `set_option linter.dupNamespace false` -/ #guard_msgs in set_option linter.dupNamespace true in -export Nat (add) +export Nat (add add_comm add) end add From 1814dc4dcc7a88545743b89f61260b1db6d9fd02 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 02:46:57 +0000 Subject: [PATCH 040/131] refactor(Equiv/Set): use `Disjoint` (#17880) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ... instead of `s ∩ t ⊆ ∅` in `Equiv.Set.union`. Also add `simps` to `Equiv.Set.univ`. --- Mathlib/Data/Finset/Basic.lean | 2 +- Mathlib/Data/Set/Card.lean | 3 +-- Mathlib/GroupTheory/Perm/Finite.lean | 8 +++---- Mathlib/Logic/Equiv/Set.lean | 32 ++++++++++++--------------- Mathlib/SetTheory/Cardinal/Basic.lean | 2 +- 5 files changed, 21 insertions(+), 26 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index f226586bced60..a94497b850efa 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -3019,7 +3019,7 @@ open Finset /-- The disjoint union of finsets is a sum -/ def Finset.union (s t : Finset α) (h : Disjoint s t) : s ⊕ t ≃ (s ∪ t : Finset α) := - Equiv.Set.ofEq (coe_union _ _) |>.trans (Equiv.Set.union (disjoint_coe.mpr h).le_bot) |>.symm + Equiv.Set.ofEq (coe_union _ _) |>.trans (Equiv.Set.union (disjoint_coe.mpr h)) |>.symm @[simp] theorem Finset.union_symm_inl (h : Disjoint s t) (x : s) : diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index 57aba3430a761..fa03f476c8309 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -109,8 +109,7 @@ theorem encard_ne_zero : s.encard ≠ 0 ↔ s.Nonempty := by theorem encard_union_eq (h : Disjoint s t) : (s ∪ t).encard = s.encard + t.encard := by classical - have e := (Equiv.Set.union (by rwa [subset_empty_iff, ← disjoint_iff_inter_eq_empty])).symm - simp [encard, ← PartENat.card_congr e, PartENat.card_sum, PartENat.withTopEquiv] + simp [encard, PartENat.card_congr (Equiv.Set.union h), PartENat.card_sum, PartENat.withTopEquiv] theorem encard_insert_of_not_mem {a : α} (has : a ∉ s) : (insert a s).encard = s.encard + 1 := by rw [← union_singleton, encard_union_eq (by simpa), encard_singleton] diff --git a/Mathlib/GroupTheory/Perm/Finite.lean b/Mathlib/GroupTheory/Perm/Finite.lean index e7ead8ca2a46c..8c9cb242cb255 100644 --- a/Mathlib/GroupTheory/Perm/Finite.lean +++ b/Mathlib/GroupTheory/Perm/Finite.lean @@ -184,9 +184,9 @@ theorem Disjoint.isConj_mul [Finite α] {σ τ π ρ : Perm α} (hc1 : IsConj σ have hd2'' := disjoint_coe.2 (disjoint_iff_disjoint_support.1 hd2) refine isConj_of_support_equiv ?_ ?_ · refine - ((Equiv.Set.ofEq hd1').trans (Equiv.Set.union hd1''.le_bot)).trans + ((Equiv.Set.ofEq hd1').trans (Equiv.Set.union hd1'')).trans ((Equiv.sumCongr (subtypeEquiv f fun a => ?_) (subtypeEquiv g fun a => ?_)).trans - ((Equiv.Set.ofEq hd2').trans (Equiv.Set.union hd2''.le_bot)).symm) <;> + ((Equiv.Set.ofEq hd2').trans (Equiv.Set.union hd2'')).symm) <;> · simp only [Set.mem_image, toEmbedding_apply, exists_eq_right, support_conj, coe_map, apply_eq_iff_eq] · intro x hx @@ -195,7 +195,7 @@ theorem Disjoint.isConj_mul [Finite α] {σ τ π ρ : Perm α} (hc1 : IsConj σ rw [hd1', Set.mem_union] at hx cases' hx with hxσ hxτ · rw [mem_coe, mem_support] at hxσ - rw [Set.union_apply_left hd1''.le_bot _, Set.union_apply_left hd1''.le_bot _] + rw [Set.union_apply_left, Set.union_apply_left] · simp only [subtypeEquiv_apply, Perm.coe_mul, Sum.map_inl, comp_apply, Set.union_symm_apply_left, Subtype.coe_mk, apply_eq_iff_eq] have h := (hd2 (f x)).resolve_left ?_ @@ -206,7 +206,7 @@ theorem Disjoint.isConj_mul [Finite α] {σ τ π ρ : Perm α} (hc1 : IsConj σ · rwa [Subtype.coe_mk, Perm.mul_apply, (hd1 x).resolve_left hxσ, mem_coe, apply_mem_support, mem_support] · rw [mem_coe, ← apply_mem_support, mem_support] at hxτ - rw [Set.union_apply_right hd1''.le_bot _, Set.union_apply_right hd1''.le_bot _] + rw [Set.union_apply_right, Set.union_apply_right] · simp only [subtypeEquiv_apply, Perm.coe_mul, Sum.map_inr, comp_apply, Set.union_symm_apply_right, Subtype.coe_mk, apply_eq_iff_eq] have h := (hd2 (g (τ x))).resolve_right ?_ diff --git a/Mathlib/Logic/Equiv/Set.lean b/Mathlib/Logic/Equiv/Set.lean index fb24c904db068..a734cb9993db0 100644 --- a/Mathlib/Logic/Equiv/Set.lean +++ b/Mathlib/Logic/Equiv/Set.lean @@ -174,8 +174,8 @@ def image {α β : Type*} (e : α ≃ β) (s : Set α) : namespace Set --- Porting note: Removed attribute @[simps apply symm_apply] /-- `univ α` is equivalent to `α`. -/ +@[simps apply symm_apply] protected def univ (α) : @univ α ≃ α := ⟨Subtype.val, fun a => ⟨a, trivial⟩, fun ⟨_, _⟩ => rfl, fun _ => rfl⟩ @@ -203,25 +203,25 @@ protected def union' {α} {s t : Set α} (p : α → Prop) [DecidablePred p] (hs rcases o with (⟨x, h⟩ | ⟨x, h⟩) <;> [simp [hs _ h]; simp [ht _ h]] /-- If sets `s` and `t` are disjoint, then `s ∪ t` is equivalent to `s ⊕ t`. -/ -protected def union {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : s ∩ t ⊆ ∅) : +protected def union {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : Disjoint s t) : (s ∪ t : Set α) ≃ s ⊕ t := - Set.union' (fun x => x ∈ s) (fun _ => id) fun _ xt xs => H ⟨xs, xt⟩ + Set.union' (fun x => x ∈ s) (fun _ => id) fun _ xt xs => Set.disjoint_left.mp H xs xt -theorem union_apply_left {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : s ∩ t ⊆ ∅) +theorem union_apply_left {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : Disjoint s t) {a : (s ∪ t : Set α)} (ha : ↑a ∈ s) : Equiv.Set.union H a = Sum.inl ⟨a, ha⟩ := dif_pos ha -theorem union_apply_right {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : s ∩ t ⊆ ∅) +theorem union_apply_right {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : Disjoint s t) {a : (s ∪ t : Set α)} (ha : ↑a ∈ t) : Equiv.Set.union H a = Sum.inr ⟨a, ha⟩ := - dif_neg fun h => H ⟨h, ha⟩ + dif_neg fun h => Set.disjoint_left.mp H h ha @[simp] -theorem union_symm_apply_left {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : s ∩ t ⊆ ∅) +theorem union_symm_apply_left {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : Disjoint s t) (a : s) : (Equiv.Set.union H).symm (Sum.inl a) = ⟨a, by simp⟩ := rfl @[simp] -theorem union_symm_apply_right {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : s ∩ t ⊆ ∅) +theorem union_symm_apply_right {α} {s t : Set α} [DecidablePred fun x => x ∈ s] (H : Disjoint s t) (a : t) : (Equiv.Set.union H).symm (Sum.inr a) = ⟨a, by simp⟩ := rfl @@ -247,7 +247,7 @@ protected def insert {α} {s : Set.{u} α} [DecidablePred (· ∈ s)] {a : α} ( (insert a s : Set α) ≃ s ⊕ PUnit.{u + 1} := calc (insert a s : Set α) ≃ ↥(s ∪ {a}) := Equiv.Set.ofEq (by simp) - _ ≃ s ⊕ ({a} : Set α) := Equiv.Set.union fun x ⟨hx, _⟩ => by simp_all + _ ≃ s ⊕ ({a} : Set α) := Equiv.Set.union <| by simpa _ ≃ s ⊕ PUnit.{u + 1} := sumCongr (Equiv.refl _) (Equiv.Set.singleton _) @[simp] @@ -273,7 +273,7 @@ theorem insert_apply_right {α} {s : Set.{u} α} [DecidablePred (· ∈ s)] {a : /-- If `s : Set α` is a set with decidable membership, then `s ⊕ sᶜ` is equivalent to `α`. -/ protected def sumCompl {α} (s : Set α) [DecidablePred (· ∈ s)] : s ⊕ (sᶜ : Set α) ≃ α := calc - s ⊕ (sᶜ : Set α) ≃ ↥(s ∪ sᶜ) := (Equiv.Set.union (by simp [Set.ext_iff])).symm + s ⊕ (sᶜ : Set α) ≃ ↥(s ∪ sᶜ) := (Equiv.Set.union disjoint_compl_right).symm _ ≃ @univ α := Equiv.Set.ofEq (by simp) _ ≃ α := Equiv.Set.univ _ @@ -289,15 +289,11 @@ theorem sumCompl_apply_inr {α : Type u} (s : Set α) [DecidablePred (· ∈ s)] theorem sumCompl_symm_apply_of_mem {α : Type u} {s : Set α} [DecidablePred (· ∈ s)] {x : α} (hx : x ∈ s) : (Equiv.Set.sumCompl s).symm x = Sum.inl ⟨x, hx⟩ := by - have : ((⟨x, Or.inl hx⟩ : (s ∪ sᶜ : Set α)) : α) ∈ s := hx - rw [Equiv.Set.sumCompl] - simpa using Set.union_apply_left (by simp) this + simp [Equiv.Set.sumCompl, Equiv.Set.univ, union_apply_left, hx] theorem sumCompl_symm_apply_of_not_mem {α : Type u} {s : Set α} [DecidablePred (· ∈ s)] {x : α} (hx : x ∉ s) : (Equiv.Set.sumCompl s).symm x = Sum.inr ⟨x, hx⟩ := by - have : ((⟨x, Or.inr hx⟩ : (s ∪ sᶜ : Set α)) : α) ∈ sᶜ := hx - rw [Equiv.Set.sumCompl] - simpa using Set.union_apply_right (by simp) this + simp [Equiv.Set.sumCompl, Equiv.Set.univ, union_apply_right, hx] @[simp] theorem sumCompl_symm_apply {α : Type*} {s : Set α} [DecidablePred (· ∈ s)] {x : s} : @@ -315,7 +311,7 @@ protected def sumDiffSubset {α} {s t : Set α} (h : s ⊆ t) [DecidablePred (· s ⊕ (t \ s : Set α) ≃ t := calc s ⊕ (t \ s : Set α) ≃ (s ∪ t \ s : Set α) := - (Equiv.Set.union (by simp [inter_diff_self])).symm + (Equiv.Set.union disjoint_sdiff_self_right).symm _ ≃ t := Equiv.Set.ofEq (by simp [union_diff_self, union_eq_self_of_subset_left h]) @[simp] @@ -346,7 +342,7 @@ protected def unionSumInter {α : Type u} (s t : Set α) [DecidablePred (· ∈ (s ∪ t : Set α) ⊕ (s ∩ t : Set α) ≃ (s ∪ t \ s : Set α) ⊕ (s ∩ t : Set α) := by rw [union_diff_self] _ ≃ (s ⊕ (t \ s : Set α)) ⊕ (s ∩ t : Set α) := - sumCongr (Set.union <| subset_empty_iff.2 (inter_diff_self _ _)) (Equiv.refl _) + sumCongr (Set.union disjoint_sdiff_self_right) (Equiv.refl _) _ ≃ s ⊕ ((t \ s : Set α) ⊕ (s ∩ t : Set α)) := sumAssoc _ _ _ _ ≃ s ⊕ (t \ s ∪ s ∩ t : Set α) := sumCongr (Equiv.refl _) diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 8c37e772ce499..22ec5d425d8c1 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1954,7 +1954,7 @@ theorem mk_union_le {α : Type u} (S T : Set α) : #(S ∪ T : Set α) ≤ #S + theorem mk_union_of_disjoint {α : Type u} {S T : Set α} (H : Disjoint S T) : #(S ∪ T : Set α) = #S + #T := by classical - exact Quot.sound ⟨Equiv.Set.union H.le_bot⟩ + exact Quot.sound ⟨Equiv.Set.union H⟩ theorem mk_insert {α : Type u} {s : Set α} {a : α} (h : a ∉ s) : #(insert a s : Set α) = #s + 1 := by From fb09b71265c60a58ebf2c1d4fd6b9477d4506d14 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 18 Oct 2024 03:51:31 +0000 Subject: [PATCH 041/131] chore: cleanup of pairwise imports (#17888) --- Mathlib/Data/Fin/Basic.lean | 3 +++ Mathlib/Data/List/OfFn.lean | 6 +++--- 2 files changed, 6 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 715342aa56963..5dbafff567583 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -626,6 +626,9 @@ theorem leftInverse_cast (eq : n = m) : LeftInverse (cast eq.symm) (cast eq) := theorem rightInverse_cast (eq : n = m) : RightInverse (cast eq.symm) (cast eq) := fun _ => rfl +theorem cast_lt_cast (eq : n = m) {a b : Fin n} : cast eq a < cast eq b ↔ a < b := + Iff.rfl + theorem cast_le_cast (eq : n = m) {a b : Fin n} : cast eq a ≤ cast eq b ↔ a ≤ b := Iff.rfl diff --git a/Mathlib/Data/List/OfFn.lean b/Mathlib/Data/List/OfFn.lean index a3271be86ffd8..27001ba3ea4b0 100644 --- a/Mathlib/Data/List/OfFn.lean +++ b/Mathlib/Data/List/OfFn.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Batteries.Data.List.OfFn -import Batteries.Data.List.Pairwise import Mathlib.Data.Fin.Tuple.Basic /-! @@ -178,8 +177,9 @@ theorem ofFn_fin_repeat {m} (a : Fin m → α) (n : ℕ) : @[simp] theorem pairwise_ofFn {R : α → α → Prop} {n} {f : Fin n → α} : (ofFn f).Pairwise R ↔ ∀ ⦃i j⦄, i < j → R (f i) (f j) := by - simp only [pairwise_iff_get, (Fin.rightInverse_cast (length_ofFn f)).surjective.forall, get_ofFn, - ← Fin.not_le, Fin.cast_le_cast] + simp only [pairwise_iff_getElem, length_ofFn, List.getElem_ofFn, + (Fin.rightInverse_cast (length_ofFn f)).surjective.forall, Fin.forall_iff, Fin.cast_mk, + Fin.mk_lt_mk, forall_comm (α := (_ : Prop)) (β := ℕ)] /-- Lists are equivalent to the sigma type of tuples of a given length. -/ @[simps] From 077e0efd658ccaa95474fb130feeded70c602f6a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 18 Oct 2024 04:24:40 +0000 Subject: [PATCH 042/131] chore: move deprecated file Mathlib/Data/ByteArray into Mathlib/Deprecated (#17884) --- Mathlib.lean | 2 +- Mathlib/{Data => Deprecated}/ByteArray.lean | 0 scripts/noshake.json | 3 ++- 3 files changed, 3 insertions(+), 2 deletions(-) rename Mathlib/{Data => Deprecated}/ByteArray.lean (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 8963ee37e6ec3..9a66b6063fc59 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2145,7 +2145,6 @@ import Mathlib.Data.Bool.Count import Mathlib.Data.Bool.Set import Mathlib.Data.Bracket import Mathlib.Data.Bundle -import Mathlib.Data.ByteArray import Mathlib.Data.Char import Mathlib.Data.Complex.Abs import Mathlib.Data.Complex.Basic @@ -2670,6 +2669,7 @@ import Mathlib.Data.ZMod.Quotient import Mathlib.Data.ZMod.Units import Mathlib.Deprecated.AlgebraClasses import Mathlib.Deprecated.Aliases +import Mathlib.Deprecated.ByteArray import Mathlib.Deprecated.Combinator import Mathlib.Deprecated.Equiv import Mathlib.Deprecated.Group diff --git a/Mathlib/Data/ByteArray.lean b/Mathlib/Deprecated/ByteArray.lean similarity index 100% rename from Mathlib/Data/ByteArray.lean rename to Mathlib/Deprecated/ByteArray.lean diff --git a/scripts/noshake.json b/scripts/noshake.json index 681e7f87ee2b5..3e93d378ca0e2 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -344,14 +344,15 @@ "Mathlib.Lean.Meta": ["Batteries.Logic"], "Mathlib.Lean.Expr.ExtraRecognizers": ["Mathlib.Data.Set.Operations"], "Mathlib.Lean.Expr.Basic": ["Batteries.Logic"], - "Mathlib.Deprecated.NatLemmas": ["Batteries.Data.Nat.Lemmas", "Batteries.WF"], "Mathlib.GroupTheory.MonoidLocalization.Basic": ["Mathlib.Init.Data.Prod"], "Mathlib.Geometry.Manifold.Sheaf.Smooth": ["Mathlib.CategoryTheory.Sites.Whiskering"], "Mathlib.Geometry.Manifold.PoincareConjecture": ["Mathlib.AlgebraicTopology.FundamentalGroupoid.SimplyConnected", "Mathlib.Util.Superscript"], + "Mathlib.Deprecated.NatLemmas": ["Batteries.Data.Nat.Lemmas", "Batteries.WF"], "Mathlib.Deprecated.MinMax": ["Mathlib.Order.MinMax"], + "Mathlib.Deprecated.ByteArray": ["Batteries.Data.ByteSubarray"], "Mathlib.Data.Vector.Basic": ["Mathlib.Control.Applicative"], "Mathlib.Data.Set.Image": ["Batteries.Tactic.Congr"], "Mathlib.Data.Seq.Parallel": ["Mathlib.Init.Data.Prod"], From 1b9ced20cbaf344811df3cbe11b727d3989bf2e2 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 18 Oct 2024 04:24:41 +0000 Subject: [PATCH 043/131] chore: cleanup of imports involving List.count (#17885) --- Archive/MiuLanguage/DecisionNec.lean | 2 -- Mathlib/Combinatorics/Enumerative/DyckWord.lean | 1 - Mathlib/Data/List/Count.lean | 3 +-- 3 files changed, 1 insertion(+), 5 deletions(-) diff --git a/Archive/MiuLanguage/DecisionNec.lean b/Archive/MiuLanguage/DecisionNec.lean index 694b599ac899d..3bd06ca5463c6 100644 --- a/Archive/MiuLanguage/DecisionNec.lean +++ b/Archive/MiuLanguage/DecisionNec.lean @@ -5,9 +5,7 @@ Authors: Gihan Marasingha -/ import Archive.MiuLanguage.Basic import Mathlib.Data.List.Basic -import Mathlib.Data.List.Count import Mathlib.Data.Nat.ModEq -import Mathlib.Tactic.Ring /-! # Decision procedure: necessary condition diff --git a/Mathlib/Combinatorics/Enumerative/DyckWord.lean b/Mathlib/Combinatorics/Enumerative/DyckWord.lean index cd58430af7c76..5924d03e54bd1 100644 --- a/Mathlib/Combinatorics/Enumerative/DyckWord.lean +++ b/Mathlib/Combinatorics/Enumerative/DyckWord.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Jeremy Tan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Tan -/ -import Batteries.Data.List.Count import Mathlib.Combinatorics.Enumerative.Catalan import Mathlib.Tactic.Positivity diff --git a/Mathlib/Data/List/Count.lean b/Mathlib/Data/List/Count.lean index 89475cec9fe1e..13c127b563b2a 100644 --- a/Mathlib/Data/List/Count.lean +++ b/Mathlib/Data/List/Count.lean @@ -10,8 +10,7 @@ import Mathlib.Tactic.Common # Counting in lists This file proves basic properties of `List.countP` and `List.count`, which count the number of -elements of a list satisfying a predicate and equal to a given element respectively. Their -definitions can be found in `Batteries.Data.List.Basic`. +elements of a list satisfying a predicate and equal to a given element respectively. -/ assert_not_exists Set.range From 4358f93967f3a6a35e2e299ff9de9e7007336549 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Fri, 18 Oct 2024 05:30:33 +0000 Subject: [PATCH 044/131] feat(NumberTheory/FermatNumber): fermatNumber n + 2 (#17659) We provide a simple non-linear recurrence relation for the Fermat numbers. Since this involves subtraction, both Nat and Int versions are given. Co-authored-by: Moritz Firsching --- Mathlib/NumberTheory/Fermat.lean | 30 +++++++++++++++++++++++++++++- 1 file changed, 29 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index 5fe9b6b32bc73..848dc49be4810 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -8,6 +8,8 @@ import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Algebra.Order.Star.Basic import Mathlib.Data.Nat.Prime.Defs import Mathlib.Tactic.Ring.RingNF +import Mathlib.Tactic.Linarith + /-! # Fermat numbers @@ -62,6 +64,32 @@ theorem fermatNumber_succ (n : ℕ) : fermatNumber (n + 1) = (fermatNumber n - 1 rw [fermatNumber, pow_succ, mul_comm, Nat.pow_mul'] rfl +theorem two_mul_fermatNumber_sub_one_sq_le_fermatNumber_sq (n : ℕ) : + 2 * (fermatNumber n - 1) ^ 2 ≤ (fermatNumber (n + 1)) ^ 2 := by + simp only [fermatNumber, add_tsub_cancel_right] + have : 0 ≤ 1 + 2 ^ (2 ^ n * 4) := le_add_left 0 (Nat.add 1 _) + ring_nf + linarith + +theorem fermatNumber_eq_fermatNumber_sq_sub_two_mul_fermatNumber_sub_one_sq (n : ℕ) : + fermatNumber (n + 2) = (fermatNumber (n + 1)) ^ 2 - 2 * (fermatNumber n - 1) ^ 2 := by + simp only [fermatNumber, add_sub_self_right] + rw [← add_sub_self_right (2 ^ 2 ^ (n + 2) + 1) <| 2 * 2 ^ 2 ^ (n + 1)] + ring_nf + +end Nat + +open Nat + +theorem Int.fermatNumber_eq_fermatNumber_sq_sub_two_mul_fermatNumber_sub_one_sq (n : ℕ) : + (fermatNumber (n + 2) : ℤ) = (fermatNumber (n + 1)) ^ 2 - 2 * (fermatNumber n - 1) ^ 2 := by + rw [Nat.fermatNumber_eq_fermatNumber_sq_sub_two_mul_fermatNumber_sub_one_sq, + Nat.cast_sub <| two_mul_fermatNumber_sub_one_sq_le_fermatNumber_sq n] + simp only [fermatNumber, push_cast, add_tsub_cancel_right] + +namespace Nat + +open Finset /-- **Goldbach's theorem** : no two distinct Fermat numbers share a common factor greater than one. @@ -78,4 +106,4 @@ theorem coprime_fermatNumber_fermatNumber {k n : ℕ} (h : k ≠ n) : refine ((dvd_prime prime_two).mp h_m).elim id (fun h_two ↦ ?_) exact ((odd_fermatNumber _).not_two_dvd_nat (h_two ▸ h_n)).elim - end Nat +end Nat From 5d0c5274420fa8a2989d403a5058acb316ab06e4 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 05:56:05 +0000 Subject: [PATCH 045/131] feat(PathConnected): review API, add missing lemmas (#17855) --- .../Complex/UpperHalfPlane/Topology.lean | 3 +- Mathlib/Analysis/Convex/Normed.lean | 2 +- Mathlib/Topology/Connected/PathConnected.lean | 204 +++++++++--------- 3 files changed, 103 insertions(+), 106 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean index 3c32e2e099f23..caecc196aeb8d 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean @@ -55,8 +55,7 @@ instance : T4Space ℍ := inferInstance instance : ContractibleSpace ℍ := (convex_halfspace_im_gt 0).contractibleSpace ⟨I, one_pos.trans_eq I_im.symm⟩ -instance : LocPathConnectedSpace ℍ := - locPathConnected_of_isOpen <| isOpen_lt continuous_const Complex.continuous_im +instance : LocPathConnectedSpace ℍ := openEmbedding_coe.locPathConnectedSpace instance : NoncompactSpace ℍ := by refine ⟨fun h => ?_⟩ diff --git a/Mathlib/Analysis/Convex/Normed.lean b/Mathlib/Analysis/Convex/Normed.lean index 235a27716dee4..fae3e58b55f18 100644 --- a/Mathlib/Analysis/Convex/Normed.lean +++ b/Mathlib/Analysis/Convex/Normed.lean @@ -112,7 +112,7 @@ instance (priority := 100) NormedSpace.instPathConnectedSpace : PathConnectedSpa TopologicalAddGroup.pathConnectedSpace instance (priority := 100) NormedSpace.instLocPathConnectedSpace : LocPathConnectedSpace E := - locPathConnected_of_bases (fun _ => Metric.nhds_basis_ball) fun x r r_pos => + .of_bases (fun _ => Metric.nhds_basis_ball) fun x r r_pos => (convex_ball x r).isPathConnected <| by simp [r_pos] theorem Wbtw.dist_add_dist {x y z : P} (h : Wbtw ℝ x y z) : diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 644fcb694b3a8..e5798309d30d4 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -785,13 +785,35 @@ theorem Specializes.joinedIn (h : x ⤳ y) (hx : x ∈ F) (hy : y ∈ F) : Joine theorem Inseparable.joinedIn (h : Inseparable x y) (hx : x ∈ F) (hy : y ∈ F) : JoinedIn F x y := h.specializes.joinedIn hx hy -/-! ### Path component -/ +theorem JoinedIn.map_continuousOn (h : JoinedIn F x y) {f : X → Y} (hf : ContinuousOn f F) : + JoinedIn (f '' F) (f x) (f y) := + let ⟨γ, hγ⟩ := h + ⟨γ.map' <| hf.mono (range_subset_iff.mpr hγ), fun t ↦ mem_image_of_mem _ (hγ t)⟩ + +theorem JoinedIn.map (h : JoinedIn F x y) {f : X → Y} (hf : Continuous f) : + JoinedIn (f '' F) (f x) (f y) := + h.map_continuousOn hf.continuousOn + +theorem Inducing.joinedIn_image {f : X → Y} (hf : Inducing f) (hx : x ∈ F) (hy : y ∈ F) : + JoinedIn (f '' F) (f x) (f y) ↔ JoinedIn F x y := by + refine ⟨?_, (.map · hf.continuous)⟩ + rintro ⟨γ, hγ⟩ + choose γ' hγ'F hγ' using hγ + have h₀ : x ⤳ γ' 0 := by rw [← hf.specializes_iff, hγ', γ.source] + have h₁ : γ' 1 ⤳ y := by rw [← hf.specializes_iff, hγ', γ.target] + have h : JoinedIn F (γ' 0) (γ' 1) := by + refine ⟨⟨⟨γ', ?_⟩, rfl, rfl⟩, hγ'F⟩ + simpa only [hf.continuous_iff, comp_def, hγ'] using map_continuous γ + exact (h₀.joinedIn hx (hγ'F _)).trans <| h.trans <| h₁.joinedIn (hγ'F _) hy +/-! ### Path component -/ /-- The path component of `x` is the set of points that can be joined to `x`. -/ def pathComponent (x : X) := { y | Joined x y } +theorem mem_pathComponent_iff : x ∈ pathComponent y ↔ Joined y x := .rfl + @[simp] theorem mem_pathComponent_self (x : X) : x ∈ pathComponent x := Joined.refl x @@ -867,24 +889,16 @@ theorem IsPathConnected.image' (hF : IsPathConnected F) exact hf.mono (range_subset_iff.2 (hx y_in).somePath_mem) /-- If `f` is continuous and `F` is path-connected, so is `f(F)`. -/ -theorem IsPathConnected.image (hF : IsPathConnected F) {f : X → Y} - (hf : Continuous f) : IsPathConnected (f '' F) := hF.image' hf.continuousOn +theorem IsPathConnected.image (hF : IsPathConnected F) {f : X → Y} (hf : Continuous f) : + IsPathConnected (f '' F) := + hF.image' hf.continuousOn /-- If `f : X → Y` is a `Inducing`, `f(F)` is path-connected iff `F` is. -/ nonrec theorem Inducing.isPathConnected_iff {f : X → Y} (hf : Inducing f) : IsPathConnected F ↔ IsPathConnected (f '' F) := by - refine ⟨fun hF ↦ hF.image hf.continuous, fun hF ↦ ?_⟩ - simp? [isPathConnected_iff] at hF ⊢ says - simp only [isPathConnected_iff, image_nonempty, mem_image, forall_exists_index, - and_imp, forall_apply_eq_imp_iff₂] at hF ⊢ - refine ⟨hF.1, fun x hx y hy ↦ ?_⟩ - rcases hF.2 x hx y hy with ⟨γ, hγ⟩ - choose γ' hγ' hγγ' using hγ - have key₁ : Inseparable x (γ' 0) := by rw [← hf.inseparable_iff, hγγ' 0, γ.source] - have key₂ : Inseparable (γ' 1) y := by rw [← hf.inseparable_iff, hγγ' 1, γ.target] - refine key₁.joinedIn hx (hγ' 0) |>.trans ⟨⟨⟨γ', ?_⟩, rfl, rfl⟩, hγ'⟩ |>.trans - (key₂.joinedIn (hγ' 1) hy) - simpa [hf.continuous_iff] using γ.continuous.congr fun t ↦ (hγγ' t).symm + simp only [IsPathConnected, forall_mem_image, exists_mem_image] + refine exists_congr fun x ↦ and_congr_right fun hx ↦ forall₂_congr fun y hy ↦ ?_ + rw [hf.joinedIn_image hx hy] /-- If `h : X → Y` is a homeomorphism, `h(s)` is path-connected iff `s` is. -/ @[simp] @@ -922,10 +936,7 @@ theorem IsPathConnected.union {U V : Set X} (hU : IsPathConnected U) (hV : IsPat ambient type `U` (when `U` contains `W`). -/ theorem IsPathConnected.preimage_coe {U W : Set X} (hW : IsPathConnected W) (hWU : W ⊆ U) : IsPathConnected (((↑) : U → X) ⁻¹' W) := by - rcases hW with ⟨x, x_in, hx⟩ - use ⟨x, hWU x_in⟩, by simp [x_in] - rintro ⟨y, hyU⟩ hyW - exact ⟨(hx hyW).joined_subtype.somePath.map (continuous_inclusion hWU), by simp⟩ + rwa [inducing_subtype_val.isPathConnected_iff, Subtype.image_preimage_val, inter_eq_right.2 hWU] theorem IsPathConnected.exists_path_through_family {n : ℕ} {s : Set X} (h : IsPathConnected s) (p : Fin (n + 1) → X) (hp : ∀ i, p i ∈ s) : @@ -997,6 +1008,7 @@ theorem IsPathConnected.exists_path_through_family' {n : ℕ} /-- A topological space is path-connected if it is non-empty and every two points can be joined by a continuous path. -/ +@[mk_iff] class PathConnectedSpace (X : Type*) [TopologicalSpace X] : Prop where /-- A path-connected space must be nonempty. -/ nonempty : Nonempty X @@ -1025,31 +1037,12 @@ def somePath (x y : X) : Path x y := end PathConnectedSpace -theorem isPathConnected_iff_pathConnectedSpace : IsPathConnected F ↔ PathConnectedSpace F := by - rw [isPathConnected_iff] - constructor - · rintro ⟨⟨x, x_in⟩, h⟩ - refine ⟨⟨⟨x, x_in⟩⟩, ?_⟩ - rintro ⟨y, y_in⟩ ⟨z, z_in⟩ - have H := h y y_in z z_in - rwa [joinedIn_iff_joined y_in z_in] at H - · rintro ⟨⟨x, x_in⟩, H⟩ - refine ⟨⟨x, x_in⟩, fun y y_in z z_in => ?_⟩ - rw [joinedIn_iff_joined y_in z_in] - apply H - theorem pathConnectedSpace_iff_univ : PathConnectedSpace X ↔ IsPathConnected (univ : Set X) := by - constructor - · intro h - haveI := @PathConnectedSpace.nonempty X _ _ - inhabit X - refine ⟨default, mem_univ _, ?_⟩ - intros y _hy - simpa using PathConnectedSpace.joined default y - · intro h - have h' := h.joinedIn - cases' h with x h - exact ⟨⟨x⟩, by simpa using h'⟩ + simp [pathConnectedSpace_iff, isPathConnected_iff, nonempty_iff_univ_nonempty] + +theorem isPathConnected_iff_pathConnectedSpace : IsPathConnected F ↔ PathConnectedSpace F := by + rw [pathConnectedSpace_iff_univ, inducing_subtype_val.isPathConnected_iff, image_univ, + Subtype.range_val_subtype, setOf_mem_eq] theorem isPathConnected_univ [PathConnectedSpace X] : IsPathConnected (univ : Set X) := pathConnectedSpace_iff_univ.mp inferInstance @@ -1111,6 +1104,7 @@ end PathConnectedSpace /-! ### Locally path connected spaces -/ +section LocPathConnectedSpace /-- A topological space is locally path connected, at every point, path connected neighborhoods form a neighborhood basis. -/ @@ -1120,66 +1114,70 @@ class LocPathConnectedSpace (X : Type*) [TopologicalSpace X] : Prop where export LocPathConnectedSpace (path_connected_basis) -theorem locPathConnected_of_bases {p : ι → Prop} {s : X → ι → Set X} - (h : ∀ x, (𝓝 x).HasBasis p (s x)) (h' : ∀ x i, p i → IsPathConnected (s x i)) : - LocPathConnectedSpace X := by - constructor - intro x - apply (h x).to_hasBasis - · intro i pi - exact ⟨s x i, ⟨(h x).mem_of_mem pi, h' x i pi⟩, by rfl⟩ - · rintro U ⟨U_in, _hU⟩ - rcases (h x).mem_iff.mp U_in with ⟨i, pi, hi⟩ - tauto - -theorem pathConnectedSpace_iff_connectedSpace [LocPathConnectedSpace X] : - PathConnectedSpace X ↔ ConnectedSpace X := by - constructor - · intro h - infer_instance - · intro hX - rw [pathConnectedSpace_iff_eq] - use Classical.arbitrary X - refine IsClopen.eq_univ ⟨?_, ?_⟩ (by simp) - · rw [isClosed_iff_nhds] - intro y H - rcases (path_connected_basis y).ex_mem with ⟨U, ⟨U_in, hU⟩⟩ - rcases H U U_in with ⟨z, hz, hz'⟩ - exact (hU.joinedIn z hz y <| mem_of_mem_nhds U_in).joined.mem_pathComponent hz' - · rw [isOpen_iff_mem_nhds] - intro y y_in - rcases (path_connected_basis y).ex_mem with ⟨U, ⟨U_in, hU⟩⟩ - apply mem_of_superset U_in - rw [← pathComponent_congr y_in] - exact hU.subset_pathComponent (mem_of_mem_nhds U_in) - -theorem pathConnected_subset_basis [LocPathConnectedSpace X] {U : Set X} (h : IsOpen U) - (hx : x ∈ U) : (𝓝 x).HasBasis (fun s : Set X => s ∈ 𝓝 x ∧ IsPathConnected s ∧ s ⊆ U) id := +theorem LocPathConnectedSpace.of_bases {p : X → ι → Prop} {s : X → ι → Set X} + (h : ∀ x, (𝓝 x).HasBasis (p x) (s x)) (h' : ∀ x i, p x i → IsPathConnected (s x i)) : + LocPathConnectedSpace X where + path_connected_basis x := by + rw [hasBasis_self] + intro t ht + rcases (h x).mem_iff.mp ht with ⟨i, hpi, hi⟩ + exact ⟨s x i, (h x).mem_of_mem hpi, h' x i hpi, hi⟩ + +@[deprecated (since := "2024-10-16")] +alias locPathConnected_of_bases := LocPathConnectedSpace.of_bases + +variable [LocPathConnectedSpace X] + +/-- In a locally path connected space, each path component is an open set. -/ +protected theorem IsOpen.pathComponent (x : X) : IsOpen (pathComponent x) := by + rw [isOpen_iff_mem_nhds] + intro y hxy + rcases (path_connected_basis y).ex_mem with ⟨V, hVy, hVc⟩ + filter_upwards [hVy] with z hz + exact hxy.out.trans (hVc.joinedIn _ (mem_of_mem_nhds hVy) _ hz).joined + +/-- In a locally path connected space, each path component is a closed set. -/ +protected theorem IsClosed.pathComponent (x : X) : IsClosed (pathComponent x) := by + rw [← isOpen_compl_iff, isOpen_iff_mem_nhds] + intro y hxy + rcases (path_connected_basis y).ex_mem with ⟨V, hVy, hVc⟩ + filter_upwards [hVy] with z hz hxz + exact hxy <| hxz.trans (hVc.joinedIn _ hz _ (mem_of_mem_nhds hVy)).joined + +/-- In a locally path connected space, each path component is a clopen set. -/ +protected theorem IsClopen.pathComponent (x : X) : IsClopen (pathComponent x) := + ⟨.pathComponent x, .pathComponent x⟩ + +theorem pathConnectedSpace_iff_connectedSpace : PathConnectedSpace X ↔ ConnectedSpace X := by + refine ⟨fun _ ↦ inferInstance, fun h ↦ ⟨inferInstance, fun x y ↦ ?_⟩⟩ + rw [← mem_pathComponent_iff, (IsClopen.pathComponent _).eq_univ] <;> simp + +theorem pathComponent_eq_connectedComponent (x : X) : pathComponent x = connectedComponent x := + (pathComponent_subset_component x).antisymm <| + (IsClopen.pathComponent x).connectedComponent_subset (mem_pathComponent_self _) + +theorem pathConnected_subset_basis {U : Set X} (h : IsOpen U) (hx : x ∈ U) : + (𝓝 x).HasBasis (fun s : Set X => s ∈ 𝓝 x ∧ IsPathConnected s ∧ s ⊆ U) id := (path_connected_basis x).hasBasis_self_subset (IsOpen.mem_nhds h hx) -theorem locPathConnected_of_isOpen [LocPathConnectedSpace X] {U : Set X} (h : IsOpen U) : - LocPathConnectedSpace U := - ⟨by - rintro ⟨x, x_in⟩ - rw [nhds_subtype_eq_comap] - constructor - intro V - rw [(HasBasis.comap ((↑) : U → X) (pathConnected_subset_basis h x_in)).mem_iff] - constructor - · rintro ⟨W, ⟨W_in, hW, hWU⟩, hWV⟩ - exact ⟨Subtype.val ⁻¹' W, ⟨⟨preimage_mem_comap W_in, hW.preimage_coe hWU⟩, hWV⟩⟩ - · rintro ⟨W, ⟨W_in, hW⟩, hWV⟩ - refine - ⟨(↑) '' W, - ⟨Filter.image_coe_mem_of_mem_comap (IsOpen.mem_nhds h x_in) W_in, - hW.image continuous_subtype_val, Subtype.coe_image_subset U W⟩, - ?_⟩ - rintro x ⟨y, ⟨y_in, hy⟩⟩ - rw [← Subtype.coe_injective hy] - tauto⟩ - -theorem IsOpen.isConnected_iff_isPathConnected [LocPathConnectedSpace X] {U : Set X} - (U_op : IsOpen U) : IsPathConnected U ↔ IsConnected U := by +theorem OpenEmbedding.locPathConnectedSpace {e : Y → X} (he : OpenEmbedding e) : + LocPathConnectedSpace Y := + have (y : Y) : + (𝓝 y).HasBasis (fun s ↦ s ∈ 𝓝 (e y) ∧ IsPathConnected s ∧ s ⊆ range e) (e ⁻¹' ·) := + he.basis_nhds <| pathConnected_subset_basis he.isOpen_range (mem_range_self _) + .of_bases this fun x s ⟨_, hs, hse⟩ ↦ by + rwa [he.isPathConnected_iff, image_preimage_eq_of_subset hse] + +theorem IsOpen.locPathConnectedSpace {U : Set X} (h : IsOpen U) : LocPathConnectedSpace U := + (openEmbedding_subtype_val h).locPathConnectedSpace + +@[deprecated (since := "2024-10-17")] +alias locPathConnected_of_isOpen := IsOpen.locPathConnectedSpace + +theorem IsOpen.isConnected_iff_isPathConnected {U : Set X} (U_op : IsOpen U) : + IsConnected U ↔ IsPathConnected U := by rw [isConnected_iff_connectedSpace, isPathConnected_iff_pathConnectedSpace] - haveI := locPathConnected_of_isOpen U_op - exact pathConnectedSpace_iff_connectedSpace + haveI := U_op.locPathConnectedSpace + exact pathConnectedSpace_iff_connectedSpace.symm + +end LocPathConnectedSpace From 319d7a37f0e813c9603b8df05901330e49aaf0eb Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 18 Oct 2024 06:38:39 +0000 Subject: [PATCH 046/131] chore: move `NNRat` APIs earlier (#17269) No need to import bundled ordered algebraic typeclasses in `ring` tactic now. --- Archive/Imo/Imo1998Q2.lean | 2 +- Archive/Imo/Imo2013Q1.lean | 7 ++- .../ContinuedFractions/ConvergentsEquiv.lean | 1 + Mathlib/Algebra/Field/Rat.lean | 52 +++++++++++++++++-- Mathlib/Algebra/Order/Field/Rat.lean | 22 -------- Mathlib/Algebra/Order/Module/Rat.lean | 1 + Mathlib/Algebra/Order/Star/Basic.lean | 4 +- Mathlib/Data/FP/Basic.lean | 2 +- Mathlib/Data/Finset/Density.lean | 1 + Mathlib/Data/NNRat/BigOperators.lean | 1 + Mathlib/Data/NNRat/Defs.lean | 11 +++- Mathlib/Data/NNRat/Lemmas.lean | 7 --- Mathlib/Data/Nat/BitIndices.lean | 6 +-- Mathlib/Data/Rat/Cast/Defs.lean | 7 +-- Mathlib/Data/Rat/Cast/Lemmas.lean | 2 +- Mathlib/Data/Rat/Cast/Order.lean | 2 +- Mathlib/Data/Rat/Star.lean | 4 +- Mathlib/Data/Real/Basic.lean | 1 + Mathlib/NumberTheory/Fermat.lean | 1 + Mathlib/Tactic/Polyrith.lean | 1 - Mathlib/Tactic/Ring/Basic.lean | 3 +- test/cancel_denoms.lean | 1 + test/linear_combination'.lean | 1 + test/linear_combination.lean | 1 + test/module.lean | 1 + test/ring.lean | 1 + 26 files changed, 88 insertions(+), 55 deletions(-) diff --git a/Archive/Imo/Imo1998Q2.lean b/Archive/Imo/Imo1998Q2.lean index 367aec35fd5fc..db7ae35345496 100644 --- a/Archive/Imo/Imo1998Q2.lean +++ b/Archive/Imo/Imo1998Q2.lean @@ -5,7 +5,7 @@ Authors: Oliver Nash -/ import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.Field.Basic -import Mathlib.Algebra.Ring.Int +import Mathlib.Algebra.Order.Field.Rat import Mathlib.GroupTheory.GroupAction.Ring import Mathlib.Tactic.NoncommRing import Mathlib.Tactic.Ring diff --git a/Archive/Imo/Imo2013Q1.lean b/Archive/Imo/Imo2013Q1.lean index 42b43f937926c..295886d7e18ea 100644 --- a/Archive/Imo/Imo2013Q1.lean +++ b/Archive/Imo/Imo2013Q1.lean @@ -3,12 +3,11 @@ Copyright (c) 2021 David Renshaw. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Renshaw -/ -import Mathlib.Algebra.BigOperators.Pi -import Mathlib.Algebra.Order.Ring.Abs -import Mathlib.Data.PNat.Basic -import Mathlib.Tactic.Ring +import Mathlib.Algebra.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Field.Rat import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.Positivity.Basic +import Mathlib.Tactic.Ring /-! # IMO 2013 Q1 diff --git a/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean b/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean index 99e76fa64a5f5..fa5e59f20d0ad 100644 --- a/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean +++ b/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Kevin Kappelmann. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Kappelmann -/ +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Algebra.ContinuedFractions.ContinuantsRecurrence import Mathlib.Algebra.ContinuedFractions.TerminatedStable import Mathlib.Tactic.FieldSimp diff --git a/Mathlib/Algebra/Field/Rat.lean b/Mathlib/Algebra/Field/Rat.lean index ee2d2e48680d1..ad60ac2ab4bb0 100644 --- a/Mathlib/Algebra/Field/Rat.lean +++ b/Mathlib/Algebra/Field/Rat.lean @@ -13,11 +13,6 @@ This file contains the field instance on the rational numbers. See note [foundational algebra order theory]. -## TODO - -Move the `Semifield ℚ≥0` instance here. This will involve proving it by hand rather than relying on -the `Nonneg` machinery. - ## Tags rat, rationals, field, ℚ, numerator, denominator, num, denom @@ -44,4 +39,51 @@ These also prevent non-computable instances being used to construct these instan instance instDivisionRing : DivisionRing ℚ := inferInstance +protected lemma inv_nonneg {a : ℚ} (ha : 0 ≤ a) : 0 ≤ a⁻¹ := by + rw [inv_def'] + exact divInt_nonneg (Int.ofNat_nonneg a.den) (num_nonneg.mpr ha) + +protected lemma div_nonneg {a b : ℚ} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a / b := + mul_nonneg ha (Rat.inv_nonneg hb) + end Rat + +namespace NNRat + +instance instInv : Inv ℚ≥0 where + inv x := ⟨x⁻¹, Rat.inv_nonneg x.2⟩ + +instance instDiv : Div ℚ≥0 where + div x y := ⟨x / y, Rat.div_nonneg x.2 y.2⟩ + +@[simp, norm_cast] lemma coe_inv (q : ℚ≥0) : ((q⁻¹ : ℚ≥0) : ℚ) = (q : ℚ)⁻¹ := rfl +@[simp, norm_cast] lemma coe_div (p q : ℚ≥0) : ((p / q : ℚ≥0) : ℚ) = p / q := rfl + +lemma inv_def (q : ℚ≥0) : q⁻¹ = divNat q.den q.num := by ext; simp [Rat.inv_def', num_coe, den_coe] +lemma div_def (p q : ℚ≥0) : p / q = divNat (p.num * q.den) (p.den * q.num) := by + ext; simp [Rat.div_def', num_coe, den_coe] + +lemma num_inv_of_ne_zero {q : ℚ≥0} (hq : q ≠ 0) : q⁻¹.num = q.den := by + rw [inv_def, divNat, num, coe_mk, Rat.divInt_ofNat, ← Rat.mk_eq_mkRat _ _ (num_ne_zero.mpr hq), + Int.natAbs_ofNat] + simpa using q.coprime_num_den.symm + +lemma den_inv_of_ne_zero {q : ℚ≥0} (hq : q ≠ 0) : q⁻¹.den = q.num := by + rw [inv_def, divNat, den, coe_mk, Rat.divInt_ofNat, ← Rat.mk_eq_mkRat _ _ (num_ne_zero.mpr hq)] + simpa using q.coprime_num_den.symm + +@[simp] +lemma num_div_den (q : ℚ≥0) : (q.num : ℚ≥0) / q.den = q := by + ext1 + rw [coe_div, coe_natCast, coe_natCast, num, ← Int.cast_natCast] + exact (cast_def _).symm + +instance instSemifield : Semifield ℚ≥0 where + __ := instNNRatCommSemiring + inv_zero := by ext; simp + mul_inv_cancel q h := by ext; simp [h] + nnratCast_def q := q.num_div_den.symm + nnqsmul q a := q * a + nnqsmul_def q a := rfl + +end NNRat diff --git a/Mathlib/Algebra/Order/Field/Rat.lean b/Mathlib/Algebra/Order/Field/Rat.lean index 228df54d19272..7b6ba3bfd819a 100644 --- a/Mathlib/Algebra/Order/Field/Rat.lean +++ b/Mathlib/Algebra/Order/Field/Rat.lean @@ -32,25 +32,3 @@ end Rat -- instances for performance deriving instance CanonicallyLinearOrderedSemifield, LinearOrderedSemifield, LinearOrderedCommGroupWithZero for NNRat - -/-! ### Miscellaneous lemmas -/ - -namespace NNRat - -@[simp, norm_cast] lemma coe_inv (q : ℚ≥0) : ((q⁻¹ : ℚ≥0) : ℚ) = (q : ℚ)⁻¹ := rfl -@[simp, norm_cast] lemma coe_div (p q : ℚ≥0) : ((p / q : ℚ≥0) : ℚ) = p / q := rfl - -lemma inv_def (q : ℚ≥0) : q⁻¹ = divNat q.den q.num := by ext; simp [Rat.inv_def', num_coe, den_coe] -lemma div_def (p q : ℚ≥0) : p / q = divNat (p.num * q.den) (p.den * q.num) := by - ext; simp [Rat.div_def', num_coe, den_coe] - -lemma num_inv_of_ne_zero {q : ℚ≥0} (hq : q ≠ 0) : q⁻¹.num = q.den := by - rw [inv_def, divNat, num, coe_mk, Rat.divInt_ofNat, ← Rat.mk_eq_mkRat _ _ (num_ne_zero.mpr hq), - Int.natAbs_ofNat] - simpa using q.coprime_num_den.symm - -lemma den_inv_of_ne_zero {q : ℚ≥0} (hq : q ≠ 0) : q⁻¹.den = q.num := by - rw [inv_def, divNat, den, coe_mk, Rat.divInt_ofNat, ← Rat.mk_eq_mkRat _ _ (num_ne_zero.mpr hq)] - simpa using q.coprime_num_den.symm - -end NNRat diff --git a/Mathlib/Algebra/Order/Module/Rat.lean b/Mathlib/Algebra/Order/Module/Rat.lean index 288226c006963..21b376b2c7771 100644 --- a/Mathlib/Algebra/Order/Module/Rat.lean +++ b/Mathlib/Algebra/Order/Module/Rat.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.Order.Module.Defs +import Mathlib.Data.NNRat.Lemmas import Mathlib.Data.Rat.Cast.Order /-! diff --git a/Mathlib/Algebra/Order/Star/Basic.lean b/Mathlib/Algebra/Order/Star/Basic.lean index bcda0ea6851ab..6370c2c91944e 100644 --- a/Mathlib/Algebra/Order/Star/Basic.lean +++ b/Mathlib/Algebra/Order/Star/Basic.lean @@ -3,10 +3,10 @@ Copyright (c) 2023 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.Algebra.Group.Submonoid.Operations +import Mathlib.Algebra.Order.Group.Defs +import Mathlib.Algebra.Order.Group.Nat import Mathlib.Algebra.Star.SelfAdjoint import Mathlib.Algebra.Star.StarRingHom -import Mathlib.Algebra.Regular.Basic import Mathlib.Tactic.ContinuousFunctionalCalculus /-! # Star ordered rings diff --git a/Mathlib/Data/FP/Basic.lean b/Mathlib/Data/FP/Basic.lean index c854edb12daef..703af9e82aadb 100644 --- a/Mathlib/Data/FP/Basic.lean +++ b/Mathlib/Data/FP/Basic.lean @@ -85,7 +85,7 @@ theorem Float.Zero.valid : ValidFinite emin 0 := ring_nf rw [mul_comm] assumption - le_trans C.precMax (Nat.le_mul_of_pos_left _ two_pos), + le_trans C.precMax (Nat.le_mul_of_pos_left _ Nat.zero_lt_two), by (rw [max_eq_right]; simp [sub_eq_add_neg, Int.ofNat_zero_le])⟩ @[nolint docBlame] diff --git a/Mathlib/Data/Finset/Density.lean b/Mathlib/Data/Finset/Density.lean index 002513c40e808..08b5292155a45 100644 --- a/Mathlib/Data/Finset/Density.lean +++ b/Mathlib/Data/Finset/Density.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.Order.Field.Basic +import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Fintype.Card import Mathlib.Data.NNRat.Order import Mathlib.Data.Rat.Cast.CharZero diff --git a/Mathlib/Data/NNRat/BigOperators.lean b/Mathlib/Data/NNRat/BigOperators.lean index e49a23c8a08e6..755e9be98bc81 100644 --- a/Mathlib/Data/NNRat/BigOperators.lean +++ b/Mathlib/Data/NNRat/BigOperators.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Algebra.Order.BigOperators.Ring.Finset +import Mathlib.Algebra.Order.Ring.Rat import Mathlib.Data.NNRat.Defs /-! # Casting lemmas for non-negative rational numbers involving sums and products diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index 5bda9ba5b3826..3f411698b84cd 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -14,7 +14,8 @@ import Mathlib.Algebra.Ring.Rat This file defines the nonnegative rationals as a subtype of `Rat` and provides its basic algebraic order structure. -Note that `NNRat` is not declared as a `Field` here. See `Data.NNRat.Lemmas` for that instance. +Note that `NNRat` is not declared as a `Semifield` here. See `Mathlib.Algebra.Field.Rat` for that +instance. We also define an instance `CanLift ℚ ℚ≥0`. This instance can be used by the `lift` tactic to replace `x : ℚ` and `hx : 0 ≤ x` in the proof context with `x : ℚ≥0` while replacing all occurrences @@ -63,8 +64,16 @@ namespace NNRat variable {p q : ℚ≥0} +instance instNontrivial : Nontrivial ℚ≥0 where exists_pair_ne := ⟨1, 0, by decide⟩ +instance instOrderBot : OrderBot ℚ≥0 where + bot := 0 + bot_le q := q.2 + @[simp] lemma val_eq_cast (q : ℚ≥0) : q.1 = q := rfl +instance instCharZero : CharZero ℚ≥0 where + cast_injective a b hab := by simpa using congr_arg num hab + instance canLift : CanLift ℚ ℚ≥0 (↑) fun q ↦ 0 ≤ q where prf q hq := ⟨⟨q, hq⟩, rfl⟩ diff --git a/Mathlib/Data/NNRat/Lemmas.lean b/Mathlib/Data/NNRat/Lemmas.lean index 4e1ee467939eb..91aadd51015c4 100644 --- a/Mathlib/Data/NNRat/Lemmas.lean +++ b/Mathlib/Data/NNRat/Lemmas.lean @@ -61,13 +61,6 @@ namespace NNRat variable {p q : ℚ≥0} -@[simp] -lemma num_div_den (q : ℚ≥0) : (q.num : ℚ≥0) / q.den = q := by - ext : 1 - rw [coe_div, coe_natCast, coe_natCast, num, ← Int.cast_natCast, - Int.natAbs_of_nonneg (Rat.num_nonneg.2 q.cast_nonneg)] - exact Rat.num_div_den q - /-- A recursor for nonnegative rationals in terms of numerators and denominators. -/ protected def rec {α : ℚ≥0 → Sort*} (h : ∀ m n : ℕ, α (m / n)) (q : ℚ≥0) : α q := by rw [← num_div_den q]; apply h diff --git a/Mathlib/Data/Nat/BitIndices.lean b/Mathlib/Data/Nat/BitIndices.lean index 79bd2eb583ef4..bd77462f10032 100644 --- a/Mathlib/Data/Nat/BitIndices.lean +++ b/Mathlib/Data/Nat/BitIndices.lean @@ -3,12 +3,12 @@ Copyright (c) 2024 Peter Nelson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Peter Nelson -/ -import Mathlib.Data.List.Sort -import Mathlib.Data.Nat.Bitwise import Mathlib.Algebra.BigOperators.Ring.List import Mathlib.Algebra.Order.BigOperators.Group.List import Mathlib.Algebra.Order.Star.Basic -import Mathlib.Algebra.Order.Sub.Defs +import Mathlib.Algebra.Order.Sub.Basic +import Mathlib.Data.List.Sort +import Mathlib.Data.Nat.Bitwise /-! # Bit Indices diff --git a/Mathlib/Data/Rat/Cast/Defs.lean b/Mathlib/Data/Rat/Cast/Defs.lean index 739dbc17cb0ce..5458a333b2468 100644 --- a/Mathlib/Data/Rat/Cast/Defs.lean +++ b/Mathlib/Data/Rat/Cast/Defs.lean @@ -3,13 +3,12 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ +import Mathlib.Algebra.Field.Basic import Mathlib.Algebra.Field.Rat import Mathlib.Algebra.Group.Commute.Basic -import Mathlib.Algebra.GroupWithZero.Units.Lemmas -import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Int.Cast.Lemmas -import Mathlib.Data.NNRat.Lemmas import Mathlib.Data.Rat.Lemmas +import Mathlib.Order.Nat /-! # Casts for Rational Numbers @@ -24,6 +23,8 @@ casting lemmas showing the well-behavedness of this injection. rat, rationals, field, ℚ, numerator, denominator, num, denom, cast, coercion, casting -/ +assert_not_exists OrderedAddCommMonoid + variable {F ι α β : Type*} namespace NNRat diff --git a/Mathlib/Data/Rat/Cast/Lemmas.lean b/Mathlib/Data/Rat/Cast/Lemmas.lean index 4089d6b56d2e4..f61ef4b456824 100644 --- a/Mathlib/Data/Rat/Cast/Lemmas.lean +++ b/Mathlib/Data/Rat/Cast/Lemmas.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Algebra.Field.Basic +import Mathlib.Algebra.Order.Nonneg.Field import Mathlib.Data.Rat.Cast.Defs import Mathlib.Tactic.Positivity.Basic diff --git a/Mathlib/Data/Rat/Cast/Order.lean b/Mathlib/Data/Rat/Cast/Order.lean index 2a1d627e9d6b0..c85aa1ca681f0 100644 --- a/Mathlib/Data/Rat/Cast/Order.lean +++ b/Mathlib/Data/Rat/Cast/Order.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Algebra.Order.Ring.Rat +import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Rat.Cast.CharZero import Mathlib.Tactic.Positivity.Core import Mathlib.Algebra.Order.Field.Basic diff --git a/Mathlib/Data/Rat/Star.lean b/Mathlib/Data/Rat/Star.lean index 8f4b38afcab0e..1e69ef6c77136 100644 --- a/Mathlib/Data/Rat/Star.lean +++ b/Mathlib/Data/Rat/Star.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux, Yaël Dillies -/ import Mathlib.Algebra.GroupWithZero.Commute +import Mathlib.Algebra.Order.Field.Rat +import Mathlib.Algebra.Order.Monoid.Submonoid import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Algebra.Order.Star.Basic -import Mathlib.Data.NNRat.Lemmas -import Mathlib.Algebra.Order.Monoid.Submonoid import Mathlib.Tactic.FieldSimp /-! diff --git a/Mathlib/Data/Real/Basic.lean b/Mathlib/Data/Real/Basic.lean index 9efe0fb5d9ae3..3465de02f2a9a 100644 --- a/Mathlib/Data/Real/Basic.lean +++ b/Mathlib/Data/Real/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.Order.CauSeq.Completion +import Mathlib.Algebra.Order.Field.Rat /-! # Real numbers from Cauchy sequences diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index 848dc49be4810..c90f62582cae8 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -5,6 +5,7 @@ Authors: Moritz Firsching -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Order.Ring.Basic +import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Algebra.Order.Star.Basic import Mathlib.Data.Nat.Prime.Defs import Mathlib.Tactic.Ring.RingNF diff --git a/Mathlib/Tactic/Polyrith.lean b/Mathlib/Tactic/Polyrith.lean index 1b3ef56ab08c8..b6da0ce70e38b 100644 --- a/Mathlib/Tactic/Polyrith.lean +++ b/Mathlib/Tactic/Polyrith.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 Dhruv Bhatia. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dhruv Bhatia, Eric Wieser, Mario Carneiro -/ -import Mathlib.Algebra.Order.Field.Rat import Mathlib.Tactic.LinearCombination /-! diff --git a/Mathlib/Tactic/Ring/Basic.lean b/Mathlib/Tactic/Ring/Basic.lean index c5ac7fd426cd8..b25109521ba02 100644 --- a/Mathlib/Tactic/Ring/Basic.lean +++ b/Mathlib/Tactic/Ring/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Aurélien Saue, Anne Baanen -/ -import Mathlib.Algebra.Order.Ring.Rat import Mathlib.Tactic.NormNum.Inv import Mathlib.Tactic.NormNum.Pow import Mathlib.Util.AtomM @@ -75,6 +74,8 @@ This feature wasn't needed yet, so it's not implemented yet. ring, semiring, exponent, power -/ +assert_not_exists OrderedAddCommMonoid + namespace Mathlib.Tactic namespace Ring diff --git a/test/cancel_denoms.lean b/test/cancel_denoms.lean index b717b67f56ee2..4d1001ecbcd6e 100644 --- a/test/cancel_denoms.lean +++ b/test/cancel_denoms.lean @@ -1,3 +1,4 @@ +import Mathlib.Algebra.Order.Field.Rat import Mathlib.Tactic.CancelDenoms import Mathlib.Tactic.Ring diff --git a/test/linear_combination'.lean b/test/linear_combination'.lean index 14516504da879..36056748d622d 100644 --- a/test/linear_combination'.lean +++ b/test/linear_combination'.lean @@ -1,3 +1,4 @@ +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Tactic.LinearCombination' import Mathlib.Tactic.Linarith diff --git a/test/linear_combination.lean b/test/linear_combination.lean index 9f652a8186a63..6e8b23a79630f 100644 --- a/test/linear_combination.lean +++ b/test/linear_combination.lean @@ -1,3 +1,4 @@ +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Tactic.LinearCombination import Mathlib.Tactic.Linarith diff --git a/test/module.lean b/test/module.lean index ee59dd02d0d9f..b3ec8f57b5789 100644 --- a/test/module.lean +++ b/test/module.lean @@ -3,6 +3,7 @@ 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.Algebra.Order.Field.Defs import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.LinearCombination import Mathlib.Tactic.Module diff --git a/test/ring.lean b/test/ring.lean index 5abd490bfa500..ffdfe755e4dec 100644 --- a/test/ring.lean +++ b/test/ring.lean @@ -1,3 +1,4 @@ +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.Ring From 9b16c5d7fb54a2f624ce50e9b4f9de1e2a26b963 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 07:23:26 +0000 Subject: [PATCH 047/131] chore(Topology/Algebra/Order): move 2 files (#17879) These two files don't deal with algebra (groups, monoids etc); move them to `Topology/Order` instead. --- Mathlib.lean | 4 ++-- Mathlib/Analysis/Calculus/LocalExtr/Rolle.lean | 2 +- Mathlib/Topology/MetricSpace/Bounded.lean | 2 +- Mathlib/Topology/MetricSpace/ProperSpace/Lemmas.lean | 2 +- Mathlib/Topology/{Algebra => }/Order/Compact.lean | 0 Mathlib/Topology/{Algebra => }/Order/Rolle.lean | 2 +- test/LibrarySearch/IsCompact.lean | 2 +- 7 files changed, 7 insertions(+), 7 deletions(-) rename Mathlib/Topology/{Algebra => }/Order/Compact.lean (100%) rename Mathlib/Topology/{Algebra => }/Order/Rolle.lean (98%) diff --git a/Mathlib.lean b/Mathlib.lean index 9a66b6063fc59..6b0d74489e5e7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4561,12 +4561,10 @@ import Mathlib.Topology.Algebra.Nonarchimedean.Completion import Mathlib.Topology.Algebra.Nonarchimedean.TotallyDisconnected import Mathlib.Topology.Algebra.OpenSubgroup import Mathlib.Topology.Algebra.Order.Archimedean -import Mathlib.Topology.Algebra.Order.Compact import Mathlib.Topology.Algebra.Order.Field import Mathlib.Topology.Algebra.Order.Floor import Mathlib.Topology.Algebra.Order.Group import Mathlib.Topology.Algebra.Order.LiminfLimsup -import Mathlib.Topology.Algebra.Order.Rolle import Mathlib.Topology.Algebra.Order.UpperLower import Mathlib.Topology.Algebra.Polynomial import Mathlib.Topology.Algebra.PontryaginDual @@ -4823,6 +4821,7 @@ import Mathlib.Topology.Order.Bornology import Mathlib.Topology.Order.Bounded import Mathlib.Topology.Order.Category.AlexDisc import Mathlib.Topology.Order.Category.FrameAdjunction +import Mathlib.Topology.Order.Compact import Mathlib.Topology.Order.DenselyOrdered import Mathlib.Topology.Order.ExtendFrom import Mathlib.Topology.Order.ExtrClosure @@ -4846,6 +4845,7 @@ import Mathlib.Topology.Order.OrderClosed import Mathlib.Topology.Order.PartialSups import Mathlib.Topology.Order.Priestley import Mathlib.Topology.Order.ProjIcc +import Mathlib.Topology.Order.Rolle import Mathlib.Topology.Order.ScottTopology import Mathlib.Topology.Order.T5 import Mathlib.Topology.Order.UpperLowerSetTopology diff --git a/Mathlib/Analysis/Calculus/LocalExtr/Rolle.lean b/Mathlib/Analysis/Calculus/LocalExtr/Rolle.lean index bd4dace2f8094..ef3e8f191d380 100644 --- a/Mathlib/Analysis/Calculus/LocalExtr/Rolle.lean +++ b/Mathlib/Analysis/Calculus/LocalExtr/Rolle.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Anatole Dedecker -/ import Mathlib.Analysis.Calculus.LocalExtr.Basic -import Mathlib.Topology.Algebra.Order.Rolle +import Mathlib.Topology.Order.Rolle /-! # Rolle's Theorem diff --git a/Mathlib/Topology/MetricSpace/Bounded.lean b/Mathlib/Topology/MetricSpace/Bounded.lean index 5ce1a3fd60a5e..7bd2d18acad75 100644 --- a/Mathlib/Topology/MetricSpace/Bounded.lean +++ b/Mathlib/Topology/MetricSpace/Bounded.lean @@ -3,7 +3,7 @@ Copyright (c) 2015, 2017 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel -/ -import Mathlib.Topology.Algebra.Order.Compact +import Mathlib.Topology.Order.Compact import Mathlib.Topology.MetricSpace.ProperSpace import Mathlib.Topology.MetricSpace.Cauchy import Mathlib.Topology.EMetricSpace.Diam diff --git a/Mathlib/Topology/MetricSpace/ProperSpace/Lemmas.lean b/Mathlib/Topology/MetricSpace/ProperSpace/Lemmas.lean index d582d83e1a6e6..f75b63205c5b6 100644 --- a/Mathlib/Topology/MetricSpace/ProperSpace/Lemmas.lean +++ b/Mathlib/Topology/MetricSpace/ProperSpace/Lemmas.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Topology.Algebra.Order.Compact +import Mathlib.Topology.Order.Compact import Mathlib.Topology.MetricSpace.ProperSpace import Mathlib.Topology.Order.IntermediateValue import Mathlib.Topology.Order.LocalExtr diff --git a/Mathlib/Topology/Algebra/Order/Compact.lean b/Mathlib/Topology/Order/Compact.lean similarity index 100% rename from Mathlib/Topology/Algebra/Order/Compact.lean rename to Mathlib/Topology/Order/Compact.lean diff --git a/Mathlib/Topology/Algebra/Order/Rolle.lean b/Mathlib/Topology/Order/Rolle.lean similarity index 98% rename from Mathlib/Topology/Algebra/Order/Rolle.lean rename to Mathlib/Topology/Order/Rolle.lean index 57a3118c4e740..31fa140f80c8a 100644 --- a/Mathlib/Topology/Algebra/Order/Rolle.lean +++ b/Mathlib/Topology/Order/Rolle.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Topology.Order.ExtendFrom -import Mathlib.Topology.Algebra.Order.Compact +import Mathlib.Topology.Order.Compact import Mathlib.Topology.Order.LocalExtr import Mathlib.Topology.Order.T5 diff --git a/test/LibrarySearch/IsCompact.lean b/test/LibrarySearch/IsCompact.lean index 0428a0cc91ba1..0ac0af98ca6a4 100644 --- a/test/LibrarySearch/IsCompact.lean +++ b/test/LibrarySearch/IsCompact.lean @@ -1,5 +1,5 @@ import Mathlib.Topology.Instances.Real -import Mathlib.Topology.Algebra.Order.Compact +import Mathlib.Topology.Order.Compact -- TODO: uses sorry, but is hidden behind the `apply?` /-- warning: declaration uses 'sorry' -/ From 672f75b15914a0b4f9b9f6aa3108cb0950e1cf65 Mon Sep 17 00:00:00 2001 From: Jiang Jiedong Date: Fri, 18 Oct 2024 07:51:31 +0000 Subject: [PATCH 048/131] refactor: rename `IsLocalRingHom` to `IsLocalHom` (#17403) Replace the `IsLocalRingHom` with `IsLocalHom`. Add instances for `IsLocalHom` in `CommRingCat` to resolve the problem caused by #6045. For a more detailed explanation of the problem, see https://github.com/leanprover-community/mathlib4/pull/17403#issuecomment-2418322525. Further golf some proof related to `IsLocalHom` instance search. Co-authored-by: Jiang Jiedong <107380768+jjdishere@users.noreply.github.com> --- Mathlib/Algebra/Associated/Basic.lean | 2 +- .../Algebra/Category/Ring/Constructions.lean | 14 +++-- Mathlib/Algebra/Category/Ring/Instances.lean | 52 +++++++++++++--- Mathlib/Algebra/Group/Units/Equiv.lean | 8 ++- Mathlib/Algebra/Group/Units/Hom.lean | 59 ++++++++++++------- .../Algebra/GroupWithZero/Units/Lemmas.lean | 11 ++-- Mathlib/Algebra/Polynomial/Expand.lean | 7 ++- Mathlib/AlgebraicGeometry/AffineScheme.lean | 1 - .../GammaSpecAdjunction.lean | 2 - .../PrimeSpectrum/Basic.lean | 11 ++-- .../ProjectiveSpectrum/Scheme.lean | 2 +- Mathlib/AlgebraicGeometry/ResidueField.lean | 9 --- Mathlib/AlgebraicGeometry/Scheme.lean | 5 +- Mathlib/AlgebraicGeometry/Spec.lean | 6 +- Mathlib/AlgebraicGeometry/Stalk.lean | 14 ++--- Mathlib/AlgebraicGeometry/StructureSheaf.lean | 8 +-- Mathlib/FieldTheory/Separable.lean | 2 +- .../RingedSpace/LocallyRingedSpace.lean | 26 ++++---- .../LocallyRingedSpace/HasColimits.lean | 40 +++++++------ .../LocallyRingedSpace/ResidueField.lean | 11 ---- .../Geometry/RingedSpace/OpenImmersion.lean | 7 +-- .../LinearAlgebra/ExteriorAlgebra/Basic.lean | 8 ++- Mathlib/RingTheory/Henselian.lean | 9 ++- Mathlib/RingTheory/Ideal/Basic.lean | 2 +- .../LocalRing/ResidueField/Basic.lean | 27 +++++---- .../RingTheory/LocalRing/RingHom/Basic.lean | 47 ++++++++++----- Mathlib/RingTheory/Localization/AtPrime.lean | 10 +++- Mathlib/RingTheory/MvPowerSeries/Inverse.lean | 10 +++- Mathlib/RingTheory/PowerSeries/Inverse.lean | 10 +++- Mathlib/RingTheory/SurjectiveOnStalks.lean | 5 +- .../RingTheory/Valuation/ValExtension.lean | 10 +++- .../Valuation/ValuationSubring.lean | 2 +- 32 files changed, 269 insertions(+), 168 deletions(-) diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index a5cad90ce7c77..503074287a0cc 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -249,7 +249,7 @@ theorem Irreducible.dvd_comm [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irr ⟨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 := + {f : F} [IsLocalHom 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 _)⟩ diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index c2cc69f71a187..d9bb67a3267cd 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -224,7 +224,7 @@ def equalizerForkIsLimit : IsLimit (equalizerFork f g) := by · intro m hm exact RingHom.ext fun x => Subtype.ext <| ConcreteCategory.congr_hom hm x -instance : IsLocalRingHom (equalizerFork f g).ι := by +instance : IsLocalHom (equalizerFork f g).ι := by constructor rintro ⟨a, h₁ : _ = _⟩ (⟨⟨x, y, h₃, h₄⟩, rfl : x = _⟩ : IsUnit a) have : y ∈ RingHom.eqLocus f g := by @@ -234,8 +234,9 @@ instance : IsLocalRingHom (equalizerFork f g).ι := by rw [isUnit_iff_exists_inv] exact ⟨⟨y, this⟩, Subtype.eq h₃⟩ -instance equalizer_ι_isLocalRingHom (F : WalkingParallelPair ⥤ CommRingCat.{u}) : - IsLocalRingHom (limit.π F WalkingParallelPair.zero) := by +@[instance] +theorem equalizer_ι_isLocalHom (F : WalkingParallelPair ⥤ CommRingCat.{u}) : + IsLocalHom (limit.π F WalkingParallelPair.zero) := by have := limMap_π (diagramIsoParallelPair F).hom WalkingParallelPair.zero rw [← IsIso.comp_inv_eq] at this rw [← this] @@ -244,15 +245,18 @@ instance equalizer_ι_isLocalRingHom (F : WalkingParallelPair ⥤ CommRingCat.{u equalizerForkIsLimit (F.map WalkingParallelPairHom.left) (F.map WalkingParallelPairHom.right)⟩ WalkingParallelPair.zero] - change IsLocalRingHom ((lim.map _ ≫ _ ≫ (equalizerFork _ _).ι) ≫ _) + change IsLocalHom ((lim.map _ ≫ _ ≫ (equalizerFork _ _).ι) ≫ _) infer_instance +@[deprecated (since := "2024-10-10")] +alias equalizer_ι_isLocalRingHom := equalizer_ι_isLocalHom + open CategoryTheory.Limits.WalkingParallelPair Opposite open CategoryTheory.Limits.WalkingParallelPairHom instance equalizer_ι_is_local_ring_hom' (F : WalkingParallelPairᵒᵖ ⥤ CommRingCat.{u}) : - IsLocalRingHom (limit.π F (Opposite.op WalkingParallelPair.one)) := by + IsLocalHom (limit.π F (Opposite.op WalkingParallelPair.one)) := by have : _ = limit.π F (walkingParallelPairOpEquiv.functor.obj _) := (limit.isoLimitCone_inv_π ⟨_, IsLimit.whiskerEquivalence (limit.isLimit F) walkingParallelPairOpEquiv⟩ diff --git a/Mathlib/Algebra/Category/Ring/Instances.lean b/Mathlib/Algebra/Category/Ring/Instances.lean index fc588d27c8248..def5d4f5edf11 100644 --- a/Mathlib/Algebra/Category/Ring/Instances.lean +++ b/Mathlib/Algebra/Category/Ring/Instances.lean @@ -11,6 +11,7 @@ import Mathlib.RingTheory.LocalRing.RingHom.Basic # Ring-theoretic results in terms of categorical languages -/ +universe u open CategoryTheory @@ -36,16 +37,53 @@ instance Localization.epi' {R : CommRingCat} (M : Submonoid R) : rcases R with ⟨α, str⟩ exact IsLocalization.epi M _ -instance CommRingCat.isLocalRingHom_comp {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) - [IsLocalRingHom g] [IsLocalRingHom f] : IsLocalRingHom (f ≫ g) := - RingHom.isLocalRingHom_comp _ _ +/- +These three instances solve the issue where the `FunLike` instances provided by +`CommRingCat.instFunLike'`, `CommRingCat.instFunLike''`, and `CommRingCat.instFunLike'''` +are not syntactically equal to `CommRingCat.instFunLike` when applied to +objects of the form `CommRingCat.of R`. +To prevent infinite loops, the priority of these three instances must be set lower +than that of other instances. +-/ +instance (priority := 50) {R : Type*} [CommRing R] {S : CommRingCat} (f : CommRingCat.of R ⟶ S) + [IsLocalHom (R := CommRingCat.of R) f] : IsLocalHom f := + inferInstance + +instance (priority := 50) {R : CommRingCat} {S : Type*} [CommRing S] (f : R ⟶ CommRingCat.of S) + [IsLocalHom (S := CommRingCat.of S) f] : IsLocalHom f := + inferInstance + +instance (priority := 50) {R S : Type u} [CommRing R] [CommRing S] + (f : CommRingCat.of R ⟶ CommRingCat.of S) + [IsLocalHom (R := CommRingCat.of R) (S := CommRingCat.of S) f] : IsLocalHom f := + inferInstance + +-- This instance handles the coercion of a morphism into a real `RingHom`. +instance {R S : CommRingCat} (f : R ⟶ S) [IsLocalHom f] : + IsLocalHom (F := R →+* S) f := + inferInstance -theorem isLocalRingHom_of_iso {R S : CommRingCat} (f : R ≅ S) : IsLocalRingHom f.hom := +@[instance] +theorem CommRingCat.isLocalHom_comp {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) + [IsLocalHom g] [IsLocalHom f] : IsLocalHom (f ≫ g) := + RingHom.isLocalHom_comp _ _ + +@[deprecated (since := "2024-10-10")] +alias CommRingCat.isLocalRingHom_comp := CommRingCat.isLocalHom_comp + +theorem isLocalHom_of_iso {R S : CommRingCat} (f : R ≅ S) : IsLocalHom f.hom := { map_nonunit := fun a ha => by convert f.inv.isUnit_map ha exact (RingHom.congr_fun f.hom_inv_id _).symm } +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_iso := isLocalHom_of_iso + -- see Note [lower instance priority] -instance (priority := 100) isLocalRingHom_of_isIso {R S : CommRingCat} (f : R ⟶ S) [IsIso f] : - IsLocalRingHom f := - isLocalRingHom_of_iso (asIso f) +@[instance 100] +theorem isLocalHom_of_isIso {R S : CommRingCat} (f : R ⟶ S) [IsIso f] : + IsLocalHom f := + isLocalHom_of_iso (asIso f) + +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_isIso := isLocalHom_of_isIso diff --git a/Mathlib/Algebra/Group/Units/Equiv.lean b/Mathlib/Algebra/Group/Units/Equiv.lean index 12d6fb5b2285c..27f2a366dd06d 100644 --- a/Mathlib/Algebra/Group/Units/Equiv.lean +++ b/Mathlib/Algebra/Group/Units/Equiv.lean @@ -193,9 +193,13 @@ theorem MulEquiv.inv_symm (G : Type*) [DivisionCommMonoid G] : (MulEquiv.inv G).symm = MulEquiv.inv G := rfl -instance isLocalRingHom_equiv [Monoid M] [Monoid N] [EquivLike F M N] - [MulEquivClass F M N] (f : F) : IsLocalRingHom f where +@[instance] +theorem isLocalHom_equiv [Monoid M] [Monoid N] [EquivLike F M N] + [MulEquivClass F M N] (f : F) : IsLocalHom 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` + +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_equiv := isLocalHom_equiv diff --git a/Mathlib/Algebra/Group/Units/Hom.lean b/Mathlib/Algebra/Group/Units/Hom.lean index 844235b25ef68..5a210213562a4 100644 --- a/Mathlib/Algebra/Group/Units/Hom.lean +++ b/Mathlib/Algebra/Group/Units/Hom.lean @@ -17,7 +17,7 @@ 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 +* `IsLocalHom`: 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 @@ -28,7 +28,7 @@ also contains unrelated results about `Units` that depend on `MonoidHom`. 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`. +Add a `@[to_additive]` version of `IsLocalHom`. -/ assert_not_exists MonoidWithZero @@ -179,7 +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`. -/ +/-- Prefer `IsLocalHom.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) : @@ -208,7 +208,7 @@ theorem liftRight_inv_mul (f : M →* N) (h : ∀ x, IsUnit (f x)) (x) : end Monoid end IsUnit -section IsLocalRingHom +section IsLocalHom variable {G R S T F : Type*} @@ -218,13 +218,16 @@ variable [Monoid R] [Monoid S] [Monoid T] [FunLike F R S] 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`. -/ +class IsLocalHom (f : F) : Prop where + /-- A local homomorphism `f : R ⟶ S` will send nonunits of `R` to nonunits of `S`. -/ map_nonunit : ∀ a, IsUnit (f a) → IsUnit a +@[deprecated (since := "2024-10-10")] +alias IsLocalRingHom := IsLocalHom + @[simp] -theorem IsUnit.of_map (f : F) [IsLocalRingHom f] (a : R) (h : IsUnit (f a)) : IsUnit a := - IsLocalRingHom.map_nonunit a h +theorem IsUnit.of_map (f : F) [IsLocalHom f] (a : R) (h : IsUnit (f a)) : IsUnit a := + IsLocalHom.map_nonunit a h -- TODO : remove alias, change the parenthesis of `f` and `a` alias isUnit_of_map_unit := IsUnit.of_map @@ -232,24 +235,38 @@ 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 isUnit_map_iff (f : F) [IsLocalHom f] (a : R) : IsUnit (f a) ↔ IsUnit a := + ⟨IsLocalHom.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 +theorem isLocalHom_of_leftInverse [FunLike G S R] [MonoidHomClass G S R] + {f : F} (g : G) (hfg : Function.LeftInverse g f) : IsLocalHom 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) +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_leftInverse := isLocalHom_of_leftInverse + +@[instance] +theorem MonoidHom.isLocalHom_comp (g : S →* T) (f : R →* S) [IsLocalHom g] + [IsLocalHom f] : IsLocalHom (g.comp f) where + map_nonunit a := IsLocalHom.map_nonunit a ∘ IsLocalHom.map_nonunit (f := g) (f a) + +@[deprecated (since := "2024-10-10")] +alias MonoidHom.isLocalRingHom_comp := MonoidHom.isLocalHom_comp -- see note [lower instance priority] -instance (priority := 100) isLocalRingHom_toMonoidHom (f : F) [IsLocalRingHom f] : - IsLocalRingHom (f : R →* S) := - ⟨IsLocalRingHom.map_nonunit (f := f)⟩ +@[instance 100] +theorem isLocalHom_toMonoidHom (f : F) [IsLocalHom f] : + IsLocalHom (f : R →* S) := + ⟨IsLocalHom.map_nonunit (f := f)⟩ -theorem MonoidHom.isLocalRingHom_of_comp (f : R →* S) (g : S →* T) [IsLocalRingHom (g.comp f)] : - IsLocalRingHom f := +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_toMonoidHom := isLocalHom_toMonoidHom + +theorem MonoidHom.isLocalHom_of_comp (f : R →* S) (g : S →* T) [IsLocalHom (g.comp f)] : + IsLocalHom f := ⟨fun _ ha => (isUnit_map_iff (g.comp f) _).mp (ha.map g)⟩ -end IsLocalRingHom +@[deprecated (since := "2024-10-10")] +alias MonoidHom.isLocalRingHom_of_comp := MonoidHom.isLocalHom_of_comp + +end IsLocalHom diff --git a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean index 407a0fb1b6c75..2d5790637992e 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean @@ -22,8 +22,8 @@ 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 +lemma isLocalHom_of_exists_map_ne_one [FunLike F G₀ M] [MonoidHomClass F G₀ M] {f : F} + (hf : ∃ x : G₀, f x ≠ 1) : IsLocalHom f where map_nonunit a h := by rcases eq_or_ne a 0 with (rfl | h) · obtain ⟨t, ht⟩ := hf @@ -33,9 +33,12 @@ lemma isLocalRingHom_of_exists_map_ne_one [FunLike F G₀ M] [MonoidHomClass F G exact (h.mul_right_cancel this).symm · exact ⟨⟨a, a⁻¹, mul_inv_cancel₀ h, inv_mul_cancel₀ h⟩, rfl⟩ +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_exists_map_ne_one := isLocalHom_of_exists_map_ne_one + 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⟩ + (f : F) : IsLocalHom f := + isLocalHom_of_exists_map_ne_one ⟨0, by simp⟩ end Monoid diff --git a/Mathlib/Algebra/Polynomial/Expand.lean b/Mathlib/Algebra/Polynomial/Expand.lean index ae212d4a08496..33ed390954f87 100644 --- a/Mathlib/Algebra/Polynomial/Expand.lean +++ b/Mathlib/Algebra/Polynomial/Expand.lean @@ -268,17 +268,20 @@ section IsDomain variable (R : Type u) [CommRing R] [IsDomain R] -theorem isLocalRingHom_expand {p : ℕ} (hp : 0 < p) : IsLocalRingHom (expand R p) := by +theorem isLocalHom_expand {p : ℕ} (hp : 0 < p) : IsLocalHom (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] +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_expand := isLocalHom_expand + 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 + let _ := isLocalHom_expand R hp.bot_lt hf.of_map theorem of_irreducible_expand_pow {p : ℕ} (hp : p ≠ 0) {f : R[X]} {n : ℕ} : diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 0859e4051faef..74ea6759b202f 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -288,7 +288,6 @@ lemma isoSpec_hom_base_apply (x : U) : (Iso.eq_comp_inv _).mpr (Scheme.Opens.germ_stalkIso_hom U (V := ⊤) x trivial), X.presheaf.germ_res_assoc, Spec.map_comp, Scheme.comp_base_apply] congr 1 - have := isLocalRingHom_of_isIso (U.stalkIso x).inv exact LocalRing.comap_closedPoint (U.stalkIso x).inv lemma isoSpec_inv_app_top : diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index bd775fd8d0f99..577af21ca59d1 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -271,8 +271,6 @@ def identityToΓSpec : 𝟭 LocallyRingedSpace.{u} ⟶ Γ.rightOp ⋙ Spec.toLoc dsimp show PrimeSpectrum.comap (f.c.app (op ⊤)) (X.toΓSpecFun x) = Y.toΓSpecFun (f.base x) dsimp [toΓSpecFun] - -- TODO: this instance was found automatically before #6045 - have := @AlgebraicGeometry.LocallyRingedSpace.isLocalRingHomStalkMap X Y rw [← LocalRing.comap_closedPoint (f.stalkMap x), ← PrimeSpectrum.comap_comp_apply, ← PrimeSpectrum.comap_comp_apply] congr 2 diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 26c8303a68816..05c884eb52c60 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -624,17 +624,20 @@ def closedPoint : PrimeSpectrum R := variable {R} -theorem isLocalRingHom_iff_comap_closedPoint {S : Type v} [CommSemiring S] [LocalRing S] - (f : R →+* S) : IsLocalRingHom f ↔ PrimeSpectrum.comap f (closedPoint S) = closedPoint R := by +theorem isLocalHom_iff_comap_closedPoint {S : Type v} [CommSemiring S] [LocalRing S] + (f : R →+* S) : IsLocalHom f ↔ PrimeSpectrum.comap f (closedPoint S) = closedPoint R := by -- Porting note: inline `this` does **not** work have := (local_hom_TFAE f).out 0 4 rw [this, PrimeSpectrum.ext_iff] rfl +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_iff_comap_closedPoint := isLocalHom_iff_comap_closedPoint + @[simp] theorem comap_closedPoint {S : Type v} [CommSemiring S] [LocalRing S] (f : R →+* S) - [IsLocalRingHom f] : PrimeSpectrum.comap f (closedPoint S) = closedPoint R := - (isLocalRingHom_iff_comap_closedPoint f).mp inferInstance + [IsLocalHom f] : PrimeSpectrum.comap f (closedPoint S) = closedPoint R := + (isLocalHom_iff_comap_closedPoint f).mp inferInstance theorem specializes_closedPoint (x : PrimeSpectrum R) : x ⤳ closedPoint R := (PrimeSpectrum.le_iff_specializes _ _).mp (LocalRing.le_maximalIdeal x.2.1) diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 106326628a9e4..77da59dbcba24 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -642,7 +642,7 @@ lemma toSpec_base_apply_eq_comap {f} (x : Proj| pbo f) : rw [awayToΓ_ΓToStalk, CommRingCat.comp_eq_ring_hom_comp, PrimeSpectrum.comap_comp] exact congr(PrimeSpectrum.comap _ $(@LocalRing.comap_closedPoint (HomogeneousLocalization.AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal) _ _ - ((Proj| pbo f).presheaf.stalk x) _ _ _ (isLocalRingHom_of_isIso _))) + ((Proj| pbo f).presheaf.stalk x) _ _ _ (isLocalHom_of_isIso _))) lemma toSpec_base_apply_eq {f} (x : Proj| pbo f) : (toSpec 𝒜 f).base x = ProjIsoSpecTopComponent.toSpec 𝒜 f x := diff --git a/Mathlib/AlgebraicGeometry/ResidueField.lean b/Mathlib/AlgebraicGeometry/ResidueField.lean index a429de3ad3174..ac1d5d1e8dcc2 100644 --- a/Mathlib/AlgebraicGeometry/ResidueField.lean +++ b/Mathlib/AlgebraicGeometry/ResidueField.lean @@ -83,15 +83,6 @@ lemma basicOpen_eq_bot_iff_forall_evaluation_eq_zero (f : X.presheaf.obj (op U)) variable {X Y : Scheme.{u}} (f : X ⟶ Y) - --- 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.base x) →+* X.presheaf.stalk x) - (f.stalkMap x) := - f.1.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. -/ def Hom.residueFieldMap (f : X.Hom Y) (x : X) : diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index b958a149bf9ec..d77a37bef9383 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -640,6 +640,9 @@ namespace Scheme variable {X Y : Scheme.{u}} (f : X ⟶ Y) +instance (x) : IsLocalHom (f.stalkMap x) := + f.prop x + @[simp] lemma stalkMap_id (X : Scheme.{u}) (x : X) : (𝟙 X : X ⟶ X).stalkMap x = 𝟙 (X.presheaf.stalk x) := @@ -723,7 +726,7 @@ open LocalRing @[simp] lemma Spec_closedPoint {R S : CommRingCat} [LocalRing R] [LocalRing S] - {f : R ⟶ S} [IsLocalRingHom f] : (Spec.map f).base (closedPoint S) = closedPoint R := + {f : R ⟶ S} [IsLocalHom f] : (Spec.map f).base (closedPoint S) = closedPoint R := LocalRing.comap_closedPoint f end LocalRing diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index 25719eb7d693f..4d4595fade0bd 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -230,7 +230,7 @@ The induced map of a ring homomorphism on the prime spectra, as a morphism of lo def Spec.locallyRingedSpaceMap {R S : CommRingCat.{u}} (f : R ⟶ S) : Spec.locallyRingedSpaceObj S ⟶ Spec.locallyRingedSpaceObj R := LocallyRingedSpace.Hom.mk (Spec.sheafedSpaceMap f) fun p => - IsLocalRingHom.mk fun a ha => by + IsLocalHom.mk fun a ha => by -- Here, we are showing that the map on prime spectra induced by `f` is really a morphism of -- *locally* ringed spaces, i.e. that the induced map on the stalks is a local ring -- homomorphism. @@ -238,11 +238,9 @@ 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 + replace ha := IsLocalHom.map_nonunit (f := (Localization.localRingHom (PrimeSpectrum.comap f p).asIdeal p.asIdeal f _)) _ ha convert RingHom.isUnit_map (stalkIso R (PrimeSpectrum.comap f p)).inv ha erw [← comp_apply, show stalkToFiberRingHom R _ = (stalkIso _ _).hom from rfl, diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index c5d01c4906552..251d379d23a14 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -198,11 +198,9 @@ def stalkClosedPointTo : X.presheaf.stalk (f.base (closedPoint R)) ⟶ R := f.stalkMap (closedPoint R) ≫ (stalkClosedPointIso R).hom -instance isLocalRingHom_stalkClosedPointTo : - IsLocalRingHom (stalkClosedPointTo f) := by - apply (config := { allowSynthFailures := true }) RingHom.isLocalRingHom_comp - · apply isLocalRingHom_of_iso - · apply f.prop +instance isLocalHom_stalkClosedPointTo : + IsLocalHom (stalkClosedPointTo f) := + inferInstanceAs <| IsLocalHom (f.stalkMap (closedPoint R) ≫ (stalkClosedPointIso R).hom) lemma preimage_eq_top_of_closedPoint_mem {U : Opens X} (hU : f.base (closedPoint R) ∈ U) : f ⁻¹ᵁ U = ⊤ := @@ -234,7 +232,7 @@ lemma germ_stalkClosedPointTo (U : Opens X) (hU : f.base (closedPoint R) ∈ U) @[reassoc] lemma germ_stalkClosedPointTo_Spec_fromSpecStalk - {x : X} (f : X.presheaf.stalk x ⟶ R) [IsLocalRingHom f] (U : Opens X) (hU) : + {x : X} (f : X.presheaf.stalk x ⟶ R) [IsLocalHom f] (U : Opens X) (hU) : X.presheaf.germ U _ hU ≫ stalkClosedPointTo (Spec.map f ≫ X.fromSpecStalk x) = X.presheaf.germ U x (by simpa using hU) ≫ f := by have : (Spec.map f ≫ X.fromSpecStalk x).base (closedPoint R) = x := by @@ -278,7 +276,7 @@ variable {R} omit [LocalRing R] in /-- useful lemma for applications of `SpecToEquivOfLocalRing` -/ lemma SpecToEquivOfLocalRing_eq_iff - {f₁ f₂ : Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalRingHom f }} : + {f₁ f₂ : Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalHom f }} : f₁ = f₂ ↔ ∃ h₁ : f₁.1 = f₂.1, f₁.2.1 = (X.presheaf.stalkCongr (by rw [h₁]; rfl)).hom ≫ f₂.2.1 := by constructor @@ -297,7 +295,7 @@ Given a local ring `R` and scheme `X`, morphisms `Spec R ⟶ X` corresponds to p @[simps] noncomputable def SpecToEquivOfLocalRing : - (Spec R ⟶ X) ≃ Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalRingHom f } where + (Spec R ⟶ X) ≃ Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalHom f } where toFun f := ⟨f.base (closedPoint R), Scheme.stalkClosedPointTo f, inferInstance⟩ invFun xf := Spec.map xf.2.1 ≫ X.fromSpecStalk xf.1 left_inv := Scheme.Spec_stalkClosedPointTo_fromSpecStalk diff --git a/Mathlib/AlgebraicGeometry/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/StructureSheaf.lean index 7541f38a34b7c..6d0abf3ae594a 100644 --- a/Mathlib/AlgebraicGeometry/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/StructureSheaf.lean @@ -542,14 +542,14 @@ 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) : IsLocalHom (stalkToFiberRingHom R x) := + isLocalHom_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 _ +instance (x : PrimeSpectrum R) : IsLocalHom (localizationToStalk R x) := + isLocalHom_of_isIso _ @[simp, reassoc] theorem stalkToFiberRingHom_localizationToStalk (x : PrimeSpectrum.Top R) : diff --git a/Mathlib/FieldTheory/Separable.lean b/Mathlib/FieldTheory/Separable.lean index 63d3ad68e42f2..b8cfde792d439 100644 --- a/Mathlib/FieldTheory/Separable.lean +++ b/Mathlib/FieldTheory/Separable.lean @@ -350,7 +350,7 @@ theorem separable_or {f : F[X]} (hf : Irreducible f) : have := natDegree_eq_zero_of_derivative_eq_zero H have := (natDegree_pos_iff_degree_pos.mpr <| degree_pos_of_irreducible hf).ne' contradiction - haveI := isLocalRingHom_expand F hp + haveI := isLocalHom_expand F hp exact Or.inr ⟨by rw [separable_iff_derivative_ne_zero hf, Classical.not_not, H], contract p f, diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean index 72d5eb095aa96..f0aedb017be00 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean @@ -11,7 +11,7 @@ import Mathlib.Geometry.RingedSpace.Stalks We define (bundled) locally ringed spaces (as `SheafedSpace CommRing` along with the fact that the stalks are local rings), and morphisms between these (morphisms in `SheafedSpace` with -`IsLocalRingHom` on the stalk maps). +`IsLocalHom` on the stalk maps). -/ -- Explicit universe annotations were used in this file to improve performance #12737 @@ -73,7 +73,7 @@ def 𝒪 : Sheaf CommRingCat X.toTopCat := structure Hom (X Y : LocallyRingedSpace.{u}) extends X.toPresheafedSpace.Hom Y.toPresheafedSpace : Type _ where /-- the underlying morphism induces a local ring homomorphism on stalks -/ - prop : ∀ x, IsLocalRingHom (toHom.stalkMap x) + prop : ∀ x, IsLocalHom (toHom.stalkMap x) /-- A morphism of locally ringed spaces as a morphism of sheafed spaces. -/ abbrev Hom.toShHom {X Y : LocallyRingedSpace.{u}} (f : X.Hom Y) : @@ -104,18 +104,23 @@ 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.toShHom.stalkMap x -instance isLocalRingHomStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : - IsLocalRingHom (f.stalkMap x) := +@[instance] +theorem isLocalHomStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : + IsLocalHom (f.stalkMap x) := f.2 x -instance isLocalRingHomValStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : - IsLocalRingHom (f.toShHom.stalkMap x) := +@[instance] +theorem isLocalHomValStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : + IsLocalHom (f.toShHom.stalkMap x) := f.2 x +@[deprecated (since := "2024-10-10")] +alias isLocalRingHomValStalkMap := isLocalHomValStalkMap + /-- The identity morphism on a locally ringed space. -/ @[simps! toShHom] def id (X : LocallyRingedSpace.{u}) : Hom X X := - ⟨𝟙 X.toSheafedSpace, fun x => by erw [PresheafedSpace.stalkMap.id]; apply isLocalRingHom_id⟩ + ⟨𝟙 X.toSheafedSpace, fun x => by erw [PresheafedSpace.stalkMap.id]; infer_instance⟩ instance (X : LocallyRingedSpace.{u}) : Inhabited (Hom X X) := ⟨id X⟩ @@ -124,7 +129,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.toShHom ≫ g.toShHom, fun x => by erw [PresheafedSpace.stalkMap.comp] - exact @RingHom.isLocalRingHom_comp _ _ _ _ _ _ _ _ (f.2 _) (g.2 _)⟩ + infer_instance⟩ /-- The category of locally ringed spaces. -/ instance : Category LocallyRingedSpace.{u} where @@ -180,12 +185,11 @@ See also `isoOfSheafedSpaceIso`. @[simps! toShHom] def homOfSheafedSpaceHomOfIsIso {X Y : LocallyRingedSpace.{u}} (f : X.toSheafedSpace ⟶ Y.toSheafedSpace) [IsIso f] : X ⟶ Y := - Hom.mk f fun x => + Hom.mk f fun _ => -- Here we need to see that the stalk maps are really local ring homomorphisms. -- This can be solved by type class inference, because stalk maps of isomorphisms -- are isomorphisms and isomorphisms are local ring homomorphisms. - show IsLocalRingHom ((SheafedSpace.forgetToPresheafedSpace.map f).stalkMap x) by - infer_instance + inferInstance /-- Given two locally ringed spaces `X` and `Y`, an isomorphism between `X` and `Y` as _sheafed_ spaces can be lifted to an isomorphism `X ⟶ Y` as locally ringed spaces. diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index a45f070a1bf4f..01ff44c393aaa 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -102,7 +102,7 @@ noncomputable def coproductCofanIsColimit : IsColimit (coproductCofan F) where (colimit.ι_desc (C := SheafedSpace.{u+1, u, u} CommRingCatMax.{u, u}) (forgetToSheafedSpace.mapCocone s) i : _)] haveI : - IsLocalRingHom + IsLocalHom (((forgetToSheafedSpace.mapCocone s).ι.app i).stalkMap y) := (s.ι.app i).2 y infer_instance⟩ @@ -131,9 +131,10 @@ variable {X Y : LocallyRingedSpace.{v}} (f g : X ⟶ Y) namespace HasCoequalizer -instance coequalizer_π_app_isLocalRingHom +@[instance] +theorem coequalizer_π_app_isLocalHom (U : TopologicalSpace.Opens (coequalizer f.toShHom g.toShHom).carrier) : - IsLocalRingHom ((coequalizer.π f.toShHom g.toShHom : _).c.app (op U)) := by + IsLocalHom ((coequalizer.π f.toShHom g.toShHom : _).c.app (op U)) := by have := ι_comp_coequalizerComparison f.toShHom g.toShHom SheafedSpace.forgetToPresheafedSpace rw [← PreservesCoequalizer.iso_hom] at this erw [SheafedSpace.congr_app this.symm (op U)] @@ -144,6 +145,9 @@ instance coequalizer_π_app_isLocalRingHom PresheafedSpace.c_isIso_of_iso _ infer_instance +@[deprecated (since := "2024-10-10")] +alias coequalizer_π_app_isLocalRingHom := coequalizer_π_app_isLocalHom + /-! We roughly follow the construction given in [MR0302656]. Given a pair `f, g : X ⟶ Y` of morphisms of locally ringed spaces, we want to show that the stalk map of @@ -210,8 +214,9 @@ theorem imageBasicOpen_image_open : erw [imageBasicOpen_image_preimage] exact (imageBasicOpen f g U s).2 -instance coequalizer_π_stalk_isLocalRingHom (x : Y) : - IsLocalRingHom ((coequalizer.π f.toShHom g.toShHom : _).stalkMap x) := by +@[instance] +theorem coequalizer_π_stalk_isLocalHom (x : Y) : + IsLocalHom ((coequalizer.π f.toShHom g.toShHom : _).stalkMap x) := by constructor rintro a ha rcases TopCat.Presheaf.germ_exist _ _ a with ⟨U, hU, s, rfl⟩ @@ -240,6 +245,9 @@ instance coequalizer_π_stalk_isLocalRingHom (x : Y) : convert @RingedSpace.isUnit_res_basicOpen Y.toRingedSpace (unop _) (((coequalizer.π f.toShHom g.toShHom).c.app (op U)) s) +@[deprecated (since := "2024-10-10")] +alias coequalizer_π_stalk_isLocalRingHom := coequalizer_π_stalk_isLocalHom + end HasCoequalizer /-- The coequalizer of two locally ringed space in the category of sheafed spaces is a locally @@ -249,20 +257,18 @@ noncomputable def coequalizer : LocallyRingedSpace where localRing x := by obtain ⟨y, rfl⟩ := (TopCat.epi_iff_surjective (coequalizer.π f.toShHom g.toShHom).base).mp inferInstance x - -- TODO: this instance was found automatically before #6045 - have _ : IsLocalRingHom ((coequalizer.π f.toShHom g.toShHom).stalkMap y) := inferInstance exact ((coequalizer.π f.toShHom g.toShHom : _).stalkMap y).domain_localRing /-- The explicit coequalizer cofork of locally ringed spaces. -/ noncomputable def coequalizerCofork : Cofork f g := @Cofork.ofπ _ _ _ _ f g (coequalizer f g) ⟨coequalizer.π f.toShHom g.toShHom, -- Porting note: this used to be automatic - HasCoequalizer.coequalizer_π_stalk_isLocalRingHom _ _⟩ + HasCoequalizer.coequalizer_π_stalk_isLocalHom _ _⟩ (LocallyRingedSpace.Hom.ext' (coequalizer.condition f.toShHom g.toShHom)) -theorem isLocalRingHom_stalkMap_congr {X Y : RingedSpace} (f g : X ⟶ Y) (H : f = g) (x) - (h : IsLocalRingHom (f.stalkMap x)) : - IsLocalRingHom (g.stalkMap x) := by +theorem isLocalHom_stalkMap_congr {X Y : RingedSpace} (f g : X ⟶ Y) (H : f = g) (x) + (h : IsLocalHom (f.stalkMap x)) : + IsLocalHom (g.stalkMap x) := by rw [PresheafedSpace.stalkMap.congr_hom _ _ H.symm x]; infer_instance /-- The cofork constructed in `coequalizer_cofork` is indeed a colimit cocone. -/ @@ -274,19 +280,19 @@ noncomputable def coequalizerCoforkIsColimit : IsColimit (coequalizerCofork f g) · intro x rcases (TopCat.epi_iff_surjective (coequalizer.π f.toShHom g.toShHom).base).mp inferInstance x with ⟨y, rfl⟩ - -- Porting note: was `apply isLocalRingHom_of_comp _ (PresheafedSpace.stalkMap ...)`, this + -- Porting note: was `apply isLocalHom_of_comp _ (PresheafedSpace.stalkMap ...)`, this -- used to allow you to provide the proof that `... ≫ ...` is a local ring homomorphism later, -- but this is no longer possible set h := _ - change IsLocalRingHom h - suffices _ : IsLocalRingHom (((coequalizerCofork f g).π.1.stalkMap _).comp h) by - apply isLocalRingHom_of_comp _ ((coequalizerCofork f g).π.1.stalkMap _) + change IsLocalHom h + suffices _ : IsLocalHom (((coequalizerCofork f g).π.1.stalkMap _).comp h) by + apply isLocalHom_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).π.toShHom.stalkMap y) + change IsLocalHom (h ≫ (coequalizerCofork f g).π.toShHom.stalkMap y) erw [← PresheafedSpace.stalkMap.comp] - apply isLocalRingHom_stalkMap_congr _ _ (coequalizer.π_desc s.π.toShHom e).symm y + apply isLocalHom_stalkMap_congr _ _ (coequalizer.π_desc s.π.toShHom e).symm y infer_instance constructor · exact LocallyRingedSpace.Hom.ext' (coequalizer.π_desc _ _) diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean index 89243a71df33e..3a60a5f355de1 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean @@ -85,14 +85,6 @@ lemma Γevaluation_ne_zero_iff_mem_basicOpen (x : X) (f : X.presheaf.obj (op ⊤ 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.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. -/ def residueFieldMap (x : X) : Y.residueField (f.base x) ⟶ X.residueField x := @@ -114,9 +106,6 @@ lemma residueFieldMap_comp {Z : LocallyRingedSpace.{u}} (g : Y ⟶ Z) (x : X) : residueFieldMap (f ≫ g) x = residueFieldMap g (f.base x) ≫ residueFieldMap f x := by simp only [comp_toShHom, SheafedSpace.comp_base, Function.comp_apply, residueFieldMap] simp_rw [stalkMap_comp] - haveI : IsLocalRingHom (g.stalkMap (f.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 a29fd4205a61a..959214c257295 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -947,7 +947,7 @@ instance : SheafedSpace.IsOpenImmersion (LocallyRingedSpace.forgetToSheafedSpace H -- note to reviewers: is there a `count_heartbeats` for this? -set_option synthInstance.maxHeartbeats 30000 in +set_option synthInstance.maxHeartbeats 40000 in /-- An explicit pullback cone over `cospan f g` if `f` is an open immersion. -/ def pullbackConeOfLeft : PullbackCone f g := by refine PullbackCone.mk ?_ @@ -966,7 +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 +set_option synthInstance.maxHeartbeats 40000 in /-- The constructed `pullbackConeOfLeft` is indeed limiting. -/ def pullbackConeOfLeftIsLimit : IsLimit (pullbackConeOfLeft f g) := PullbackCone.isLimitAux' _ fun s => by @@ -986,8 +986,7 @@ 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] - -- TODO: This instance is found by `infer_instance` before #6045. - apply CommRingCat.isLocalRingHom_comp + infer_instance · intro m _ h₂ rw [← cancel_mono (pullbackConeOfLeft f g).snd] exact h₂.trans <| LocallyRingedSpace.Hom.ext' diff --git a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean index cd48d0d7611d4..258bd811ae3b5 100644 --- a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean @@ -164,8 +164,12 @@ 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) +@[instance] +theorem isLocalHom_algebraMap : IsLocalHom (algebraMap R (ExteriorAlgebra R M)) := + isLocalHom_of_leftInverse _ (algebraMap_leftInverse M) + +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_algebraMap := isLocalHom_algebraMap 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/RingTheory/Henselian.lean b/Mathlib/RingTheory/Henselian.lean index b1762c379f824..42918161d3de0 100644 --- a/Mathlib/RingTheory/Henselian.lean +++ b/Mathlib/RingTheory/Henselian.lean @@ -60,8 +60,8 @@ universe u v open Polynomial LocalRing Polynomial Function List -theorem isLocalRingHom_of_le_jacobson_bot {R : Type*} [CommRing R] (I : Ideal R) - (h : I ≤ Ideal.jacobson ⊥) : IsLocalRingHom (Ideal.Quotient.mk I) := by +theorem isLocalHom_of_le_jacobson_bot {R : Type*} [CommRing R] (I : Ideal R) + (h : I ≤ Ideal.jacobson ⊥) : IsLocalHom (Ideal.Quotient.mk I) := by constructor intro a h have : IsUnit (Ideal.Quotient.mk (Ideal.jacobson ⊥) a) := by @@ -79,6 +79,9 @@ theorem isLocalRingHom_of_le_jacobson_bot {R : Type*} [CommRing R] (I : Ideal R) simp? at h1 says simp only [mul_one, sub_add_cancel, IsUnit.mul_iff] at h1 exact h1.1 +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_le_jacobson_bot := isLocalHom_of_le_jacobson_bot + /-- A ring `R` is *Henselian* at an ideal `I` if the following condition holds: for every polynomial `f` over `R`, with a *simple* root `a₀` over the quotient ring `R/I`, there exists a lift `a : R` of `a₀` that is a root of `f`. @@ -193,7 +196,7 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] exact (ih.eval f).trans h₁ have hf'c : ∀ n, IsUnit (f'.eval (c n)) := by intro n - haveI := isLocalRingHom_of_le_jacobson_bot I (IsAdicComplete.le_jacobson_bot I) + haveI := isLocalHom_of_le_jacobson_bot I (IsAdicComplete.le_jacobson_bot I) apply IsUnit.of_map (Ideal.Quotient.mk I) convert h₂ using 1 exact SModEq.def.mp ((hc_mod n).eval _) diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 69a42212b475e..1db54f75b9e6e 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -808,7 +808,7 @@ theorem one_not_mem_nonunits [Monoid α] : (1 : α) ∉ nonunits α := -- 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 α := + [IsLocalHom 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 α := diff --git a/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean b/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean index ca82438849309..5dc19df864940 100644 --- a/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean @@ -41,7 +41,7 @@ instance ResidueField.algebra : Algebra R (ResidueField R) := theorem ResidueField.algebraMap_eq : algebraMap R (ResidueField R) = residue R := rfl -instance : IsLocalRingHom (LocalRing.residue R) := +instance : IsLocalHom (LocalRing.residue R) := ⟨fun _ ha => Classical.not_not.mp (Ideal.Quotient.eq_zero_iff_mem.not.mp (isUnit_iff_ne_zero.mp ha))⟩ @@ -50,22 +50,22 @@ variable {R} namespace ResidueField /-- A local ring homomorphism into a field can be descended onto the residue field. -/ -def lift {R S : Type*} [CommRing R] [LocalRing R] [Field S] (f : R →+* S) [IsLocalRingHom f] : +def lift {R S : Type*} [CommRing R] [LocalRing R] [Field S] (f : R →+* S) [IsLocalHom f] : LocalRing.ResidueField R →+* S := Ideal.Quotient.lift _ f fun a ha => by_contradiction fun h => ha (isUnit_of_map_unit f a (isUnit_iff_ne_zero.mpr h)) theorem lift_comp_residue {R S : Type*} [CommRing R] [LocalRing R] [Field S] (f : R →+* S) - [IsLocalRingHom f] : (lift f).comp (residue R) = f := + [IsLocalHom f] : (lift f).comp (residue R) = f := RingHom.ext fun _ => rfl @[simp] theorem lift_residue_apply {R S : Type*} [CommRing R] [LocalRing R] [Field S] (f : R →+* S) - [IsLocalRingHom f] (x) : lift f (residue R x) = f x := + [IsLocalHom f] (x) : lift f (residue R x) = f x := rfl /-- The map on residue fields induced by a local homomorphism between local rings -/ -def map (f : R →+* S) [IsLocalRingHom f] : ResidueField R →+* ResidueField S := +def map (f : R →+* S) [IsLocalHom f] : ResidueField R →+* ResidueField S := Ideal.Quotient.lift (maximalIdeal R) ((Ideal.Quotient.mk _).comp f) fun a ha => by erw [Ideal.Quotient.eq_zero_iff_mem] exact map_nonunit f a ha @@ -79,16 +79,16 @@ theorem map_id : /-- The composite of two `LocalRing.ResidueField.map`s is the `LocalRing.ResidueField.map` of the composite. -/ -theorem map_comp (f : T →+* R) (g : R →+* S) [IsLocalRingHom f] [IsLocalRingHom g] : +theorem map_comp (f : T →+* R) (g : R →+* S) [IsLocalHom f] [IsLocalHom g] : LocalRing.ResidueField.map (g.comp f) = (LocalRing.ResidueField.map g).comp (LocalRing.ResidueField.map f) := Ideal.Quotient.ringHom_ext <| RingHom.ext fun _ => rfl -theorem map_comp_residue (f : R →+* S) [IsLocalRingHom f] : +theorem map_comp_residue (f : R →+* S) [IsLocalHom f] : (ResidueField.map f).comp (residue R) = (residue S).comp f := rfl -theorem map_residue (f : R →+* S) [IsLocalRingHom f] (r : R) : +theorem map_residue (f : R →+* S) [IsLocalHom f] (r : R) : ResidueField.map f (residue R r) = residue S (f r) := rfl @@ -96,8 +96,8 @@ theorem map_id_apply (x : ResidueField R) : map (RingHom.id R) x = x := DFunLike.congr_fun map_id x @[simp] -theorem map_map (f : R →+* S) (g : S →+* T) (x : ResidueField R) [IsLocalRingHom f] - [IsLocalRingHom g] : map g (map f x) = map (g.comp f) x := +theorem map_map (f : R →+* S) (g : S →+* T) (x : ResidueField R) [IsLocalHom f] + [IsLocalHom g] : map g (map f x) = map (g.comp f) x := DFunLike.congr_fun (map_comp f g).symm x /-- A ring isomorphism defines an isomorphism of residue fields. -/ @@ -147,7 +147,7 @@ end MulSemiringAction section FiniteDimensional -variable [Algebra R S] [IsLocalRingHom (algebraMap R S)] +variable [Algebra R S] [IsLocalHom (algebraMap R S)] noncomputable instance : Algebra (ResidueField R) (ResidueField S) := (ResidueField.map (algebraMap R S)).toAlgebra @@ -175,13 +175,16 @@ end FiniteDimensional end ResidueField -theorem isLocalRingHom_residue : IsLocalRingHom (LocalRing.residue R) := by +theorem isLocalHom_residue : IsLocalHom (LocalRing.residue R) := by constructor intro a ha by_contra h erw [Ideal.Quotient.eq_zero_iff_mem.mpr ((LocalRing.mem_maximalIdeal _).mpr h)] at ha exact ha.ne_zero rfl +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_residue := isLocalHom_residue + end end LocalRing diff --git a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean index 4b44e62f3ee66..f14d328b4582a 100644 --- a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean @@ -21,25 +21,40 @@ section variable [Semiring R] [Semiring S] [Semiring T] -instance isLocalRingHom_id (R : Type*) [Semiring R] : IsLocalRingHom (RingHom.id R) where +@[instance] +theorem isLocalHom_id (R : Type*) [Semiring R] : IsLocalHom (RingHom.id R) where map_nonunit _ := id +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_id := isLocalHom_id + -- 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 100] +theorem isLocalHom_toRingHom {F : Type*} [FunLike F R S] + [RingHomClass F R S] (f : F) [IsLocalHom f] : IsLocalHom (f : R →+* S) := + ⟨IsLocalHom.map_nonunit (f := f)⟩ + +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_toRingHom := isLocalHom_toRingHom -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) +@[instance] +theorem RingHom.isLocalHom_comp (g : S →+* T) (f : R →+* S) [IsLocalHom g] + [IsLocalHom f] : IsLocalHom (g.comp f) where + map_nonunit a := IsLocalHom.map_nonunit a ∘ IsLocalHom.map_nonunit (f := g) (f a) -theorem isLocalRingHom_of_comp (f : R →+* S) (g : S →+* T) [IsLocalRingHom (g.comp f)] : - IsLocalRingHom f := +@[deprecated (since := "2024-10-10")] +alias RingHom.isLocalRingHom_comp := RingHom.isLocalHom_comp + +theorem isLocalHom_of_comp (f : R →+* S) (g : S →+* T) [IsLocalHom (g.comp f)] : + IsLocalHom f := ⟨fun _ ha => (isUnit_map_iff (g.comp f) _).mp (g.isUnit_map ha)⟩ +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_of_comp := isLocalHom_of_comp + /-- If `f : R →+* S` is a local ring hom, then `R` is a local ring if `S` is. -/ theorem RingHom.domain_localRing {R S : Type*} [CommSemiring R] [CommSemiring S] [H : LocalRing S] - (f : R →+* S) [IsLocalRingHom f] : LocalRing R := by + (f : R →+* S) [IsLocalHom f] : LocalRing R := by haveI : Nontrivial R := pullback_nonzero f f.map_zero f.map_one apply LocalRing.of_nonunits_add intro a b @@ -57,7 +72,7 @@ variable [CommSemiring R] [LocalRing R] [CommSemiring S] [LocalRing S] /-- The image of the maximal ideal of the source is contained within the maximal ideal of the target. -/ -theorem map_nonunit (f : R →+* S) [IsLocalRingHom f] (a : R) (h : a ∈ maximalIdeal R) : +theorem map_nonunit (f : R →+* S) [IsLocalHom f] (a : R) (h : a ∈ maximalIdeal R) : f a ∈ maximalIdeal S := fun H => h <| isUnit_of_map_unit f a H end @@ -73,7 +88,7 @@ i.e. any preimage of a unit is still a unit. https://stacks.math.columbia.edu/ta -/ theorem local_hom_TFAE (f : R →+* S) : List.TFAE - [IsLocalRingHom f, f '' (maximalIdeal R).1 ⊆ maximalIdeal S, + [IsLocalHom f, f '' (maximalIdeal R).1 ⊆ maximalIdeal S, (maximalIdeal R).map f ≤ maximalIdeal S, maximalIdeal R ≤ (maximalIdeal S).comap f, (maximalIdeal S).comap f = maximalIdeal R] := by tfae_have 1 → 2 @@ -89,19 +104,19 @@ theorem local_hom_TFAE (f : R →+* S) : end theorem of_surjective [CommSemiring R] [LocalRing R] [CommSemiring S] [Nontrivial S] (f : R →+* S) - [IsLocalRingHom f] (hf : Function.Surjective f) : LocalRing S := + [IsLocalHom f] (hf : Function.Surjective f) : LocalRing S := of_isUnit_or_isUnit_of_isUnit_add (by intro a b hab obtain ⟨a, rfl⟩ := hf a obtain ⟨b, rfl⟩ := hf b rw [← map_add] at hab exact - (isUnit_or_isUnit_of_isUnit_add <| IsLocalRingHom.map_nonunit _ hab).imp f.isUnit_map + (isUnit_or_isUnit_of_isUnit_add <| IsLocalHom.map_nonunit _ hab).imp f.isUnit_map f.isUnit_map) /-- If `f : R →+* S` is a surjective local ring hom, then the induced units map is surjective. -/ theorem surjective_units_map_of_local_ringHom [CommRing R] [CommRing S] (f : R →+* S) - (hf : Function.Surjective f) (h : IsLocalRingHom f) : + (hf : Function.Surjective f) (h : IsLocalHom f) : Function.Surjective (Units.map <| f.toMonoidHom) := by intro a obtain ⟨b, hb⟩ := hf (a : S) @@ -113,7 +128,7 @@ theorem surjective_units_map_of_local_ringHom [CommRing R] [CommRing S] (f : R /-- Every ring hom `f : K →+* R` from a division ring `K` to a nontrivial ring `R` is a local ring hom. -/ instance (priority := 100) {K R} [DivisionRing K] [CommRing R] [Nontrivial R] - (f : K →+* R) : IsLocalRingHom f where + (f : K →+* R) : IsLocalHom f where map_nonunit r hr := by simpa only [isUnit_iff_ne_zero, ne_eq, map_eq_zero] using hr.ne_zero end LocalRing diff --git a/Mathlib/RingTheory/Localization/AtPrime.lean b/Mathlib/RingTheory/Localization/AtPrime.lean index cec0198c34a92..90a63bac86990 100644 --- a/Mathlib/RingTheory/Localization/AtPrime.lean +++ b/Mathlib/RingTheory/Localization/AtPrime.lean @@ -218,14 +218,18 @@ theorem localRingHom_mk' (J : Ideal P) [J.IsPrime] (f : R →+* P) (hIJ : I = J. (⟨f y, le_comap_primeCompl_iff.mpr (ge_of_eq hIJ) y.2⟩ : J.primeCompl) := map_mk' _ _ _ -instance isLocalRingHom_localRingHom (J : Ideal P) [hJ : J.IsPrime] (f : R →+* P) - (hIJ : I = J.comap f) : IsLocalRingHom (localRingHom I J f hIJ) := - IsLocalRingHom.mk fun x hx => by +@[instance] +theorem isLocalHom_localRingHom (J : Ideal P) [hJ : J.IsPrime] (f : R →+* P) + (hIJ : I = J.comap f) : IsLocalHom (localRingHom I J f hIJ) := + IsLocalHom.mk fun x hx => by rcases IsLocalization.mk'_surjective I.primeCompl x with ⟨r, s, rfl⟩ rw [localRingHom_mk'] at hx rw [AtPrime.isUnit_mk'_iff] at hx ⊢ exact fun hr => hx ((SetLike.ext_iff.mp hIJ r).mp hr) +@[deprecated (since := "2024-10-10")] +alias isLocalRingHom_localRingHom := isLocalHom_localRingHom + theorem localRingHom_unique (J : Ideal P) [J.IsPrime] (f : R →+* P) (hIJ : I = J.comap f) {j : Localization.AtPrime I →+* Localization.AtPrime J} (hj : ∀ x : R, j (algebraMap _ _ x) = algebraMap _ _ (f x)) : localRingHom I J f hIJ = j := diff --git a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean index d5160e8067589..ce619c30ce099 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean @@ -31,7 +31,7 @@ Instances are defined: * Formal power series over a local ring form a local ring. * The morphism `MvPowerSeries.map σ f : MvPowerSeries σ A →* MvPowerSeries σ B` - induced by a local morphism `f : A →+* B` (`IsLocalRingHom f`) + induced by a local morphism `f : A →+* B` (`IsLocalHom f`) of commutative rings is a *local* morphism. -/ @@ -167,13 +167,14 @@ end CommRing section LocalRing -variable {S : Type*} [CommRing R] [CommRing S] (f : R →+* S) [IsLocalRingHom f] +variable {S : Type*} [CommRing R] [CommRing S] (f : R →+* S) [IsLocalHom f] -- Thanks to the linter for informing us that this instance does -- not actually need R and S to be local rings! /-- The map between multivariate formal power series over the same indexing set induced by a local ring hom `A → B` is local -/ -instance map.isLocalRingHom : IsLocalRingHom (map σ f) := +@[instance] +theorem map.isLocalHom : IsLocalHom (map σ f) := ⟨by rintro φ ⟨ψ, h⟩ replace h := congr_arg (constantCoeff σ S) h @@ -183,6 +184,9 @@ instance map.isLocalRingHom : IsLocalRingHom (map σ f) := rcases isUnit_of_map_unit f _ this with ⟨c, hc⟩ exact isUnit_of_mul_eq_one φ (invOfUnit φ c) (mul_invOfUnit φ c hc.symm)⟩ +@[deprecated (since := "2024-10-10")] +alias map.isLocalRingHom := map.isLocalHom + end LocalRing section Field diff --git a/Mathlib/RingTheory/PowerSeries/Inverse.lean b/Mathlib/RingTheory/PowerSeries/Inverse.lean index aec805253fd3d..4a9202d3df4a7 100644 --- a/Mathlib/RingTheory/PowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/PowerSeries/Inverse.lean @@ -266,10 +266,14 @@ end Field section LocalRing -variable {S : Type*} [CommRing R] [CommRing S] (f : R →+* S) [IsLocalRingHom f] +variable {S : Type*} [CommRing R] [CommRing S] (f : R →+* S) [IsLocalHom f] -instance map.isLocalRingHom : IsLocalRingHom (map f) := - MvPowerSeries.map.isLocalRingHom f +@[instance] +theorem map.isLocalHom : IsLocalHom (map f) := + MvPowerSeries.map.isLocalHom f + +@[deprecated (since := "2024-10-10")] +alias map.isLocalRingHom := map.isLocalHom variable [LocalRing R] [LocalRing S] diff --git a/Mathlib/RingTheory/SurjectiveOnStalks.lean b/Mathlib/RingTheory/SurjectiveOnStalks.lean index 94c5443047b71..c8c8f65ad3cba 100644 --- a/Mathlib/RingTheory/SurjectiveOnStalks.lean +++ b/Mathlib/RingTheory/SurjectiveOnStalks.lean @@ -174,7 +174,7 @@ lemma SurjectiveOnStalks.baseChange one_mul, mul_one, id_apply, ← e] rw [Algebra.algebraMap_eq_smul_one, ← smul_tmul', smul_mul_assoc] -lemma surjectiveOnStalks_iff_of_isLocalRingHom [LocalRing S] [IsLocalRingHom f] : +lemma surjectiveOnStalks_iff_of_isLocalHom [LocalRing S] [IsLocalHom f] : f.SurjectiveOnStalks ↔ Function.Surjective f := by refine ⟨fun H x ↦ ?_, fun h ↦ surjectiveOnStalks_of_surjective h⟩ obtain ⟨y, r, c, hc, hr, e⟩ := @@ -185,4 +185,7 @@ lemma surjectiveOnStalks_iff_of_isLocalRingHom [LocalRing S] [IsLocalRingHom f] apply hc.mul_right_injective simp only [← _root_.map_mul, ← mul_assoc, IsUnit.mul_val_inv, one_mul, e] +@[deprecated (since := "2024-10-10")] +alias surjectiveOnStalks_iff_of_isLocalRingHom := surjectiveOnStalks_iff_of_isLocalHom + end RingHom diff --git a/Mathlib/RingTheory/Valuation/ValExtension.lean b/Mathlib/RingTheory/Valuation/ValExtension.lean index 566aa0f741a77..11ee048743d9c 100644 --- a/Mathlib/RingTheory/Valuation/ValExtension.lean +++ b/Mathlib/RingTheory/Valuation/ValExtension.lean @@ -143,16 +143,20 @@ theorem algebraMap_injective [IsValExtension vK vA] [Nontrivial A] : ext apply RingHom.injective (algebraMap K A) h -instance instIsLocalRingHomValuationInteger {S ΓS: Type*} [CommRing S] +@[instance] +theorem instIsLocalHomValuationInteger {S ΓS: Type*} [CommRing S] [LinearOrderedCommGroupWithZero ΓS] - [Algebra R S] [IsLocalRingHom (algebraMap R S)] {vS : Valuation S ΓS} - [IsValExtension vR vS] : IsLocalRingHom (algebraMap vR.integer vS.integer) where + [Algebra R S] [IsLocalHom (algebraMap R S)] {vS : Valuation S ΓS} + [IsValExtension vR vS] : IsLocalHom (algebraMap vR.integer vS.integer) where map_nonunit r hr := by apply (Valuation.integer.integers (v := vR)).isUnit_of_one · exact (isUnit_map_iff (algebraMap R S) _).mp (hr.map (algebraMap _ S)) · apply (Valuation.integer.integers (v := vS)).one_of_isUnit at hr exact (val_map_eq_one_iff vR vS _).mp hr +@[deprecated (since := "2024-10-10")] +alias instIsLocalRingHomValuationInteger := instIsLocalHomValuationInteger + end integer end IsValExtension diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index 48d6df21160fa..d6d738435853c 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -646,7 +646,7 @@ theorem ker_unitGroupToResidueFieldUnits : theorem surjective_unitGroupToResidueFieldUnits : Function.Surjective A.unitGroupToResidueFieldUnits := (LocalRing.surjective_units_map_of_local_ringHom _ Ideal.Quotient.mk_surjective - LocalRing.isLocalRingHom_residue).comp + LocalRing.isLocalHom_residue).comp (MulEquiv.surjective _) /-- The quotient of the unit group of `A` by the principal unit group of `A` agrees with From 1f26636368232f82cf41fc3ab1793f6d61a9edc2 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Fri, 18 Oct 2024 08:51:11 +0000 Subject: [PATCH 049/131] chore(Algebra/Module/Equiv/Defs): remove `FunLike` instance (#17477) Co-authored-by: Moritz Firsching --- Mathlib/Algebra/Module/Equiv/Defs.lean | 9 --------- 1 file changed, 9 deletions(-) diff --git a/Mathlib/Algebra/Module/Equiv/Defs.lean b/Mathlib/Algebra/Module/Equiv/Defs.lean index 7d08836e2d79f..430a86911a3e0 100644 --- a/Mathlib/Algebra/Module/Equiv/Defs.lean +++ b/Mathlib/Algebra/Module/Equiv/Defs.lean @@ -168,15 +168,6 @@ instance : EquivLike (M ≃ₛₗ[σ] M₂) M M₂ where left_inv := LinearEquiv.left_inv right_inv := LinearEquiv.right_inv -/-- Helper instance for when inference gets stuck on following the normal chain -`EquivLike → FunLike`. - -TODO: this instance doesn't appear to be necessary: remove it (after benchmarking?) --/ -instance : FunLike (M ≃ₛₗ[σ] M₂) M M₂ where - coe := DFunLike.coe - coe_injective' := DFunLike.coe_injective - instance : SemilinearEquivClass (M ≃ₛₗ[σ] M₂) σ M M₂ where map_add := (·.map_add') --map_add' Porting note (#11215): TODO why did I need to change this? map_smulₛₗ := (·.map_smul') --map_smul' Porting note (#11215): TODO why did I need to change this? From 6dca5634f0e8d885d4dae292bfa910bb62a9e524 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 08:51:13 +0000 Subject: [PATCH 050/131] chore(MeasureTheory): use `Disjoint` (#17890) --- .../ConditionalExpectation/CondexpL1.lean | 8 ++++---- Mathlib/MeasureTheory/Function/LpSpace.lean | 4 ++-- Mathlib/MeasureTheory/Integral/Bochner.lean | 9 ++++----- Mathlib/MeasureTheory/Integral/SetToL1.lean | 16 ++++++++-------- 4 files changed, 18 insertions(+), 19 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean index 633ff287c5581..45d4dfd32c71c 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean @@ -135,7 +135,7 @@ theorem norm_condexpIndL1Fin_le (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x exact lintegral_nnnorm_condexpIndSMul_le hm hs hμs x theorem condexpIndL1Fin_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) - (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : + (hμt : μ t ≠ ∞) (hst : Disjoint s t) (x : G) : condexpIndL1Fin hm (hs.union ht) ((measure_union_le s t).trans_lt (lt_top_iff_ne_top.mpr (ENNReal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne x = condexpIndL1Fin hm hs hμs x + condexpIndL1Fin hm ht hμt x := by @@ -224,7 +224,7 @@ theorem continuous_condexpIndL1 : Continuous fun x : G => condexpIndL1 hm μ s x continuous_of_linear_of_bound condexpIndL1_add condexpIndL1_smul norm_condexpIndL1_le theorem condexpIndL1_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) - (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : + (hμt : μ t ≠ ∞) (hst : Disjoint s t) (x : G) : condexpIndL1 hm μ (s ∪ t) x = condexpIndL1 hm μ s x + condexpIndL1 hm μ t x := by have hμst : μ (s ∪ t) ≠ ∞ := ((measure_union_le s t).trans_lt (lt_top_iff_ne_top.mpr (ENNReal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne @@ -283,12 +283,12 @@ theorem norm_condexpInd_le : ‖(condexpInd G hm μ s : G →L[ℝ] α →₁[μ ContinuousLinearMap.opNorm_le_bound _ ENNReal.toReal_nonneg norm_condexpInd_apply_le theorem condexpInd_disjoint_union_apply (hs : MeasurableSet s) (ht : MeasurableSet t) - (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : + (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : Disjoint s t) (x : G) : condexpInd G hm μ (s ∪ t) x = condexpInd G hm μ s x + condexpInd G hm μ t x := condexpIndL1_disjoint_union hs ht hμs hμt hst x theorem condexpInd_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) - (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) : (condexpInd G hm μ (s ∪ t) : G →L[ℝ] α →₁[μ] G) = + (hμt : μ t ≠ ∞) (hst : Disjoint s t) : (condexpInd G hm μ (s ∪ t) : G →L[ℝ] α →₁[μ] G) = condexpInd G hm μ s + condexpInd G hm μ t := by ext1 x; push_cast; exact condexpInd_disjoint_union_apply hs ht hμs hμt hst x diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index a20a1c45824a7..7e5941d364518 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -857,7 +857,7 @@ theorem memℒp_add_of_disjoint {f g : α → E} (h : Disjoint (support f) (supp /-- The indicator of a disjoint union of two sets is the sum of the indicators of the sets. -/ theorem indicatorConstLp_disjoint_union {s t : Set α} (hs : MeasurableSet s) (ht : MeasurableSet t) - (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (c : E) : + (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : Disjoint s t) (c : E) : indicatorConstLp p (hs.union ht) (measure_union_ne_top hμs hμt) c = indicatorConstLp p hs hμs c + indicatorConstLp p ht hμt c := by ext1 @@ -865,7 +865,7 @@ theorem indicatorConstLp_disjoint_union {s t : Set α} (hs : MeasurableSet s) (h refine EventuallyEq.trans ?_ (EventuallyEq.add indicatorConstLp_coeFn.symm indicatorConstLp_coeFn.symm) - rw [Set.indicator_union_of_disjoint (Set.disjoint_iff_inter_eq_empty.mpr hst) _] + rw [Set.indicator_union_of_disjoint hst] end IndicatorConstLp diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index b4eab8d959db6..b497a13fc272b 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -201,18 +201,17 @@ theorem weightedSMul_null {s : Set α} (h_zero : μ s = 0) : (weightedSMul μ s ext1 x; rw [weightedSMul_apply, h_zero]; simp theorem weightedSMul_union' (s t : Set α) (ht : MeasurableSet t) (hs_finite : μ s ≠ ∞) - (ht_finite : μ t ≠ ∞) (h_inter : s ∩ t = ∅) : + (ht_finite : μ t ≠ ∞) (hdisj : Disjoint s t) : (weightedSMul μ (s ∪ t) : F →L[ℝ] F) = weightedSMul μ s + weightedSMul μ t := by ext1 x - simp_rw [add_apply, weightedSMul_apply, - measure_union (Set.disjoint_iff_inter_eq_empty.mpr h_inter) ht, + simp_rw [add_apply, weightedSMul_apply, measure_union hdisj ht, ENNReal.toReal_add hs_finite ht_finite, add_smul] @[nolint unusedArguments] theorem weightedSMul_union (s t : Set α) (_hs : MeasurableSet s) (ht : MeasurableSet t) - (hs_finite : μ s ≠ ∞) (ht_finite : μ t ≠ ∞) (h_inter : s ∩ t = ∅) : + (hs_finite : μ s ≠ ∞) (ht_finite : μ t ≠ ∞) (hdisj : Disjoint s t) : (weightedSMul μ (s ∪ t) : F →L[ℝ] F) = weightedSMul μ s + weightedSMul μ t := - weightedSMul_union' s t ht hs_finite ht_finite h_inter + weightedSMul_union' s t ht hs_finite ht_finite hdisj theorem weightedSMul_smul [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass ℝ 𝕜 F] (c : 𝕜) (s : Set α) (x : F) : weightedSMul μ s (c • x) = c • weightedSMul μ s x := by diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index e4fe8a61b787a..3d16c95fed2c1 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -9,7 +9,7 @@ import Mathlib.MeasureTheory.Function.SimpleFuncDenseLp # Extension of a linear function from indicators to L1 Let `T : Set α → E →L[ℝ] F` be additive for measurable sets with finite measure, in the sense that -for `s, t` two such sets, `s ∩ t = ∅ → T (s ∪ t) = T s + T t`. `T` is akin to a bilinear map on +for `s, t` two such sets, `Disjoint s t → T (s ∪ t) = T s + T t`. `T` is akin to a bilinear map on `Set α × E`, or a linear map on indicator functions. This file constructs an extension of `T` to integrable simple functions, which are finite sums of @@ -23,7 +23,7 @@ expectation of an integrable function in `MeasureTheory.Function.ConditionalExpe ## Main Definitions - `FinMeasAdditive μ T`: the property that `T` is additive on measurable sets with finite measure. - For two such sets, `s ∩ t = ∅ → T (s ∪ t) = T s + T t`. + For two such sets, `Disjoint s t → T (s ∪ t) = T s + T t`. - `DominatedFinMeasAdditive μ T C`: `FinMeasAdditive μ T ∧ ∀ s, ‖T s‖ ≤ C * (μ s).toReal`. This is the property needed to perform the extension from indicators to L1. - `setToL1 (hT : DominatedFinMeasAdditive μ T C) : (α →₁[μ] E) →L[ℝ] F`: the extension of `T` @@ -90,7 +90,8 @@ section FinMeasAdditive sets with finite measure is the sum of its values on each set. -/ def FinMeasAdditive {β} [AddMonoid β] {_ : MeasurableSpace α} (μ : Measure α) (T : Set α → β) : Prop := - ∀ s t, MeasurableSet s → MeasurableSet t → μ s ≠ ∞ → μ t ≠ ∞ → s ∩ t = ∅ → T (s ∪ t) = T s + T t + ∀ s t, MeasurableSet s → MeasurableSet t → μ s ≠ ∞ → μ t ≠ ∞ → Disjoint s t → + T (s ∪ t) = T s + T t namespace FinMeasAdditive @@ -133,7 +134,7 @@ theorem smul_measure_iff (c : ℝ≥0∞) (hc_ne_zero : c ≠ 0) (hc_ne_top : c theorem map_empty_eq_zero {β} [AddCancelMonoid β] {T : Set α → β} (hT : FinMeasAdditive μ T) : T ∅ = 0 := by have h_empty : μ ∅ ≠ ∞ := (measure_empty.le.trans_lt ENNReal.coe_lt_top).ne - specialize hT ∅ ∅ MeasurableSet.empty MeasurableSet.empty h_empty h_empty (Set.inter_empty ∅) + specialize hT ∅ ∅ MeasurableSet.empty MeasurableSet.empty h_empty h_empty (disjoint_empty _) rw [Set.union_empty] at hT nth_rw 1 [← add_zero (T ∅)] at hT exact (add_left_cancel hT).symm @@ -160,10 +161,9 @@ theorem map_iUnion_fin_meas_set_eq_sum (T : Set α → β) (T_empty : T ∅ = 0) · congr; convert Finset.iSup_insert a s S · exact (measure_biUnion_lt_top s.finite_toSet fun i hi ↦ (hps i <| Finset.mem_insert_of_mem hi).lt_top).ne - · simp_rw [Set.inter_iUnion] - refine iUnion_eq_empty.mpr fun i => iUnion_eq_empty.mpr fun hi => ?_ - rw [← Set.disjoint_iff_inter_eq_empty] - refine h_disj a (Finset.mem_insert_self a s) i (Finset.mem_insert_of_mem hi) fun hai => ?_ + · simp_rw [Set.disjoint_iUnion_right] + intro i hi + refine h_disj a (Finset.mem_insert_self a s) i (Finset.mem_insert_of_mem hi) fun hai ↦ ?_ rw [← hai] at hi exact has hi From ec349bacca2d1e67eae5142d16b30a202510578d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 09:18:36 +0000 Subject: [PATCH 051/131] chore(PosPart): Make more lemmas simp (#17053) --- Mathlib/Algebra/Order/Group/PosPart.lean | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Order/Group/PosPart.lean b/Mathlib/Algebra/Order/Group/PosPart.lean index 085e56c1b0be2..eb8835fa80134 100644 --- a/Mathlib/Algebra/Order/Group/PosPart.lean +++ b/Mathlib/Algebra/Order/Group/PosPart.lean @@ -65,7 +65,7 @@ instance instLeOnePart : LeOnePart α where @[to_additive] lemma oneLePart_mono : Monotone (·⁺ᵐ : α → α) := fun _a _b hab ↦ sup_le_sup_right hab _ -@[to_additive (attr := simp)] lemma oneLePart_one : (1 : α)⁺ᵐ = 1 := sup_idem _ +@[to_additive (attr := simp high)] lemma oneLePart_one : (1 : α)⁺ᵐ = 1 := sup_idem _ @[to_additive (attr := simp)] lemma leOnePart_one : (1 : α)⁻ᵐ = 1 := by simp [leOnePart] @@ -79,8 +79,10 @@ instance instLeOnePart : LeOnePart α where @[to_additive] lemma inv_le_leOnePart (a : α) : a⁻¹ ≤ a⁻ᵐ := le_sup_left @[to_additive (attr := simp)] lemma oneLePart_eq_self : a⁺ᵐ = a ↔ 1 ≤ a := sup_eq_left +@[to_additive (attr := simp)] lemma oneLePart_eq_one : a⁺ᵐ = 1 ↔ a ≤ 1 := sup_eq_right -@[to_additive] lemma oneLePart_eq_one : a⁺ᵐ = 1 ↔ a ≤ 1 := sup_eq_right +@[to_additive (attr := simp)] alias ⟨_, oneLePart_of_one_le⟩ := oneLePart_eq_self +@[to_additive (attr := simp)] alias ⟨_, oneLePart_of_le_one⟩ := oneLePart_eq_one /-- See also `leOnePart_eq_inv`. -/ @[to_additive "See also `negPart_eq_neg`."] @@ -114,6 +116,9 @@ variable [CovariantClass α α (· * ·) (· ≤ ·)] @[to_additive (attr := simp)] lemma leOnePart_eq_one : a⁻ᵐ = 1 ↔ 1 ≤ a := by simp [leOnePart_eq_one'] +@[to_additive (attr := simp)] alias ⟨_, leOnePart_of_le_one⟩ := leOnePart_eq_inv +@[to_additive (attr := simp)] alias ⟨_, leOnePart_of_one_le⟩ := leOnePart_eq_one + @[to_additive (attr := simp) negPart_pos] lemma one_lt_ltOnePart (ha : a < 1) : 1 < a⁻ᵐ := by rwa [leOnePart_eq_inv.2 ha.le, one_lt_inv'] From 3b5cedb7b2da9faae67402d7a66a75e5f44b666e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 09:18:37 +0000 Subject: [PATCH 052/131] feat(Topology): define `ContinuousEval{,Const}` classes (#17319) --- Mathlib.lean | 2 + .../NormedSpace/Multilinear/Basic.lean | 27 +++--- Mathlib/Analysis/ODE/PicardLindelof.lean | 4 +- .../Topology/Algebra/ContinuousMonoidHom.lean | 15 +++- .../Algebra/Module/Alternating/Topology.lean | 26 +++--- .../Algebra/Module/Multilinear/Topology.lean | 27 +++--- .../Algebra/Module/StrongTopology.lean | 30 ++++--- Mathlib/Topology/CompactOpen.lean | 41 +++++---- Mathlib/Topology/Connected/PathConnected.lean | 14 +-- Mathlib/Topology/ContinuousMap/Algebra.lean | 40 +++++---- .../ContinuousMap/ContinuousMapZero.lean | 12 ++- Mathlib/Topology/Hom/ContinuousEval.lean | 69 ++++++++++++++ Mathlib/Topology/Hom/ContinuousEvalConst.lean | 89 +++++++++++++++++++ Mathlib/Topology/Homotopy/HSpaces.lean | 2 +- Mathlib/Topology/Homotopy/HomotopyGroup.lean | 8 +- Mathlib/Topology/Homotopy/Path.lean | 10 +-- 16 files changed, 307 insertions(+), 109 deletions(-) create mode 100644 Mathlib/Topology/Hom/ContinuousEval.lean create mode 100644 Mathlib/Topology/Hom/ContinuousEvalConst.lean diff --git a/Mathlib.lean b/Mathlib.lean index 6b0d74489e5e7..1e5548b290acc 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4716,6 +4716,8 @@ import Mathlib.Topology.Filter import Mathlib.Topology.GDelta import Mathlib.Topology.Germ import Mathlib.Topology.Gluing +import Mathlib.Topology.Hom.ContinuousEval +import Mathlib.Topology.Hom.ContinuousEvalConst import Mathlib.Topology.Hom.Open import Mathlib.Topology.Homeomorph import Mathlib.Topology.Homotopy.Basic diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index ceb978b4e25e7..7f88338734d62 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -80,16 +80,19 @@ variable {𝕜 ι : Type*} {E : ι → Type*} {F : Type*} [NormedField 𝕜] [Finite ι] [∀ i, SeminormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] [TopologicalSpace F] [AddCommGroup F] [TopologicalAddGroup F] [Module 𝕜 F] -/-- Applying a multilinear map to a vector is continuous in both coordinates. -/ -theorem ContinuousMultilinearMap.continuous_eval : - Continuous fun p : ContinuousMultilinearMap 𝕜 E F × ∀ i, E i => p.1 p.2 := by - cases nonempty_fintype ι - let _ := TopologicalAddGroup.toUniformSpace F - have := comm_topologicalAddGroup_is_uniform (G := F) - refine (UniformOnFun.continuousOn_eval₂ fun m ↦ ?_).comp_continuous - (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⟩ +instance ContinuousMultilinearMap.instContinuousEval : + ContinuousEval (ContinuousMultilinearMap 𝕜 E F) (Π i, E i) F where + continuous_eval := by + cases nonempty_fintype ι + let _ := TopologicalAddGroup.toUniformSpace F + have := comm_topologicalAddGroup_is_uniform (G := F) + refine (UniformOnFun.continuousOn_eval₂ fun m ↦ ?_).comp_continuous + (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⟩ + +@[deprecated (since := "2024-10-05")] +protected alias ContinuousMultilinearMap.continuous_eval := continuous_eval namespace ContinuousLinearMap @@ -97,8 +100,8 @@ variable {G : Type*} [AddCommGroup G] [TopologicalSpace G] [Module 𝕜 G] [Cont (f : G →L[𝕜] ContinuousMultilinearMap 𝕜 E F) lemma continuous_uncurry_of_multilinear : - Continuous (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) := - ContinuousMultilinearMap.continuous_eval.comp <| .prodMap (map_continuous f) continuous_id + Continuous (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) := by + fun_prop lemma continuousOn_uncurry_of_multilinear {s} : ContinuousOn (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) s := diff --git a/Mathlib/Analysis/ODE/PicardLindelof.lean b/Mathlib/Analysis/ODE/PicardLindelof.lean index 1f6996fcbf20e..bf6d91625d9e7 100644 --- a/Mathlib/Analysis/ODE/PicardLindelof.lean +++ b/Mathlib/Analysis/ODE/PicardLindelof.lean @@ -223,10 +223,10 @@ instance [CompleteSpace E] : CompleteSpace v.FunSpace := by 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 ?_ + refine (isClosed_eq (continuous_eval_const _) continuous_const).inter ?_ have : IsClosed {f : Icc v.tMin v.tMax → E | LipschitzWith v.C f} := isClosed_setOf_lipschitzWith v.C - exact this.preimage ContinuousMap.continuous_coe + exact this.preimage continuous_coeFun theorem intervalIntegrable_vComp (t₁ t₂ : ℝ) : IntervalIntegrable f.vComp volume t₁ t₂ := f.continuous_vComp.intervalIntegrable _ _ diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index 4e852f064519b..a41b781ceaadc 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -240,6 +240,15 @@ theorem embedding_toContinuousMap : Embedding (toContinuousMap : ContinuousMonoidHom A B → C(A, B)) := ⟨inducing_toContinuousMap A B, toContinuousMap_injective⟩ +@[to_additive] +instance instContinuousEvalConst : ContinuousEvalConst (ContinuousMonoidHom A B) A B := + .of_continuous_forget (inducing_toContinuousMap A B).continuous + +@[to_additive] +instance instContinuousEval [LocallyCompactPair A B] : + ContinuousEval (ContinuousMonoidHom A B) A B := + .of_continuous_forget (inducing_toContinuousMap A B).continuous + @[to_additive] lemma range_toContinuousMap : Set.range (toContinuousMap : ContinuousMonoidHom A B → C(A, B)) = @@ -254,10 +263,10 @@ theorem closedEmbedding_toContinuousMap [ContinuousMul B] [T2Space B] : toEmbedding := embedding_toContinuousMap A B isClosed_range := by simp only [range_toContinuousMap, Set.setOf_and, Set.setOf_forall] - refine .inter (isClosed_singleton.preimage (ContinuousMap.continuous_eval_const 1)) <| + refine .inter (isClosed_singleton.preimage (continuous_eval_const 1)) <| isClosed_iInter fun x ↦ isClosed_iInter fun y ↦ ?_ - exact isClosed_eq (ContinuousMap.continuous_eval_const (x * y)) <| - .mul (ContinuousMap.continuous_eval_const x) (ContinuousMap.continuous_eval_const y) + exact isClosed_eq (continuous_eval_const (x * y)) <| + .mul (continuous_eval_const x) (continuous_eval_const y) variable {A B C D E} diff --git a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean index 32efbbd5a39c7..52bcb32bef151 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean @@ -36,7 +36,7 @@ lemma isClosed_range_toContinuousMultilinearMap [ContinuousSMul 𝕜 E] [T2Space ContinuousMultilinearMap 𝕜 (fun _ : ι ↦ E) F)) := by simp only [range_toContinuousMultilinearMap, setOf_forall] repeat refine isClosed_iInter fun _ ↦ ?_ - exact isClosed_singleton.preimage (ContinuousMultilinearMap.continuous_eval_const _) + exact isClosed_singleton.preimage (continuous_eval_const _) end IsClosedRange @@ -125,6 +125,10 @@ lemma embedding_toContinuousMultilinearMap : haveI := comm_topologicalAddGroup_is_uniform (G := F) isUniformEmbedding_toContinuousMultilinearMap.embedding +instance instTopologicalAddGroup : TopologicalAddGroup (E [⋀^ι]→L[𝕜] F) := + embedding_toContinuousMultilinearMap.topologicalAddGroup + (toContinuousMultilinearMapLinear (R := ℕ)) + @[continuity, fun_prop] lemma continuous_toContinuousMultilinearMap : Continuous (toContinuousMultilinearMap : (E [⋀^ι]→L[𝕜] F → _)) := @@ -159,21 +163,19 @@ lemma closedEmbedding_toContinuousMultilinearMap [T2Space F] : (E [⋀^ι]→L[𝕜] F) → ContinuousMultilinearMap 𝕜 (fun _ : ι ↦ E) F) := ⟨embedding_toContinuousMultilinearMap, isClosed_range_toContinuousMultilinearMap⟩ -@[continuity, fun_prop] -theorem continuous_eval_const (x : ι → E) : - Continuous fun p : E [⋀^ι]→L[𝕜] F ↦ p x := - (ContinuousMultilinearMap.continuous_eval_const x).comp continuous_toContinuousMultilinearMap +instance instContinuousEvalConst : ContinuousEvalConst (E [⋀^ι]→L[𝕜] F) (ι → E) F := + .of_continuous_forget continuous_toContinuousMultilinearMap -theorem continuous_coe_fun : - Continuous (DFunLike.coe : E [⋀^ι]→L[𝕜] F → (ι → E) → F) := - continuous_pi continuous_eval_const +@[deprecated (since := "2024-10-05")] +protected alias continuous_eval_const := continuous_eval_const + +@[deprecated (since := "2024-10-05")] +protected alias continuous_coe_fun := continuous_coeFun instance instT2Space [T2Space F] : T2Space (E [⋀^ι]→L[𝕜] F) := - .of_injective_continuous DFunLike.coe_injective continuous_coe_fun + .of_injective_continuous DFunLike.coe_injective continuous_coeFun -instance instT3Space [T2Space F] : T2Space (E [⋀^ι]→L[𝕜] F) := - letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F - haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform +instance instT3Space [T2Space F] : T3Space (E [⋀^ι]→L[𝕜] F) := inferInstance section RestrictScalars diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean index baf4b78201f9a..19af25db016c5 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Topology.Algebra.Module.Multilinear.Bounded import Mathlib.Topology.Algebra.Module.UniformConvergence +import Mathlib.Topology.Hom.ContinuousEvalConst /-! # Topology on continuous multilinear maps @@ -152,6 +153,11 @@ end UniformAddGroup variable [TopologicalSpace F] [TopologicalAddGroup F] +instance instTopologicalAddGroup : TopologicalAddGroup (ContinuousMultilinearMap 𝕜 E F) := + letI := TopologicalAddGroup.toUniformSpace F + haveI := comm_topologicalAddGroup_is_uniform (G := F) + inferInstance + instance instContinuousConstSMul {M : Type*} [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜 M F] [ContinuousConstSMul M F] : ContinuousConstSMul M (ContinuousMultilinearMap 𝕜 E F) := by @@ -188,20 +194,21 @@ theorem hasBasis_nhds_zero : variable [∀ i, ContinuousSMul 𝕜 (E i)] -theorem continuous_eval_const (x : Π i, E i) : - Continuous fun p : ContinuousMultilinearMap 𝕜 E F ↦ p x := by - letI := TopologicalAddGroup.toUniformSpace F - haveI := comm_topologicalAddGroup_is_uniform (G := F) - exact (uniformContinuous_eval_const x).continuous +instance : ContinuousEvalConst (ContinuousMultilinearMap 𝕜 E F) (Π i, E i) F where + continuous_eval_const x := + let _ := TopologicalAddGroup.toUniformSpace F + have _ := comm_topologicalAddGroup_is_uniform (G := F) + (uniformContinuous_eval_const x).continuous +@[deprecated (since := "2024-10-05")] protected alias continuous_eval_const := continuous_eval_const @[deprecated (since := "2024-04-10")] alias continuous_eval_left := continuous_eval_const - -theorem continuous_coe_fun : - Continuous (DFunLike.coe : ContinuousMultilinearMap 𝕜 E F → (Π i, E i) → F) := - continuous_pi continuous_eval_const +@[deprecated (since := "2024-10-05")] protected alias continuous_coe_fun := continuous_coeFun instance instT2Space [T2Space F] : T2Space (ContinuousMultilinearMap 𝕜 E F) := - .of_injective_continuous DFunLike.coe_injective continuous_coe_fun + .of_injective_continuous DFunLike.coe_injective continuous_coeFun + +instance instT3Space [T2Space F] : T3Space (ContinuousMultilinearMap 𝕜 E F) := + inferInstance section RestrictScalars diff --git a/Mathlib/Topology/Algebra/Module/StrongTopology.lean b/Mathlib/Topology/Algebra/Module/StrongTopology.lean index 8e988265fd745..3132c5f3d16b8 100644 --- a/Mathlib/Topology/Algebra/Module/StrongTopology.lean +++ b/Mathlib/Topology/Algebra/Module/StrongTopology.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ import Mathlib.Topology.Algebra.Module.UniformConvergence +import Mathlib.Topology.Hom.ContinuousEvalConst /-! # Strong topologies on the space of continuous linear maps @@ -151,12 +152,19 @@ instance instTopologicalAddGroup [TopologicalSpace F] [TopologicalAddGroup F] haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform infer_instance +theorem continuousEvalConst [TopologicalSpace F] [TopologicalAddGroup F] + (𝔖 : Set (Set E)) (h𝔖 : ⋃₀ 𝔖 = Set.univ) : + ContinuousEvalConst (UniformConvergenceCLM σ F 𝔖) E F where + continuous_eval_const x := by + letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F + haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform + exact (UniformOnFun.uniformContinuous_eval h𝔖 x).continuous.comp + (embedding_coeFn σ F 𝔖).continuous + theorem t2Space [TopologicalSpace F] [TopologicalAddGroup F] [T2Space F] - (𝔖 : 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𝔖 - exact (embedding_coeFn σ F 𝔖).t2Space + (𝔖 : Set (Set E)) (h𝔖 : ⋃₀ 𝔖 = univ) : T2Space (UniformConvergenceCLM σ F 𝔖) := + have := continuousEvalConst σ F 𝔖 h𝔖 + .of_injective_continuous DFunLike.coe_injective continuous_coeFun instance instDistribMulAction (M : Type*) [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜₂ M F] [TopologicalSpace F] [TopologicalAddGroup F] [ContinuousConstSMul M F] (𝔖 : Set (Set E)) : @@ -321,11 +329,13 @@ instance uniformSpace [UniformSpace F] [UniformAddGroup F] : UniformSpace (E → instance uniformAddGroup [UniformSpace F] [UniformAddGroup F] : UniformAddGroup (E →SL[σ] F) := UniformConvergenceCLM.instUniformAddGroup σ F _ -instance [TopologicalSpace F] [TopologicalAddGroup F] [ContinuousSMul 𝕜₁ E] [T2Space F] : - T2Space (E →SL[σ] F) := - UniformConvergenceCLM.t2Space σ F _ - (Set.eq_univ_of_forall fun x => - Set.mem_sUnion_of_mem (Set.mem_singleton x) (isVonNBounded_singleton x)) +instance instContinuousEvalConst [TopologicalSpace F] [TopologicalAddGroup F] + [ContinuousSMul 𝕜₁ E] : ContinuousEvalConst (E →SL[σ] F) E F := + UniformConvergenceCLM.continuousEvalConst σ F _ Bornology.isVonNBounded_covers + +instance instT2Space [TopologicalSpace F] [TopologicalAddGroup F] [ContinuousSMul 𝕜₁ E] + [T2Space F] : T2Space (E →SL[σ] F) := + UniformConvergenceCLM.t2Space σ F _ Bornology.isVonNBounded_covers protected theorem hasBasis_nhds_zero_of_basis [TopologicalSpace F] [TopologicalAddGroup F] {ι : Type*} {p : ι → Prop} {b : ι → Set F} (h : (𝓝 0 : Filter F).HasBasis p b) : diff --git a/Mathlib/Topology/CompactOpen.lean b/Mathlib/Topology/CompactOpen.lean index c4e0af021885c..9c199df740de2 100644 --- a/Mathlib/Topology/CompactOpen.lean +++ b/Mathlib/Topology/CompactOpen.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Reid Barton. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Reid Barton -/ +import Mathlib.Topology.Hom.ContinuousEval import Mathlib.Topology.ContinuousMap.Basic /-! @@ -160,27 +161,24 @@ section Ev /-- The evaluation map `C(X, Y) × X → Y` is continuous if `X, Y` is a locally compact pair of spaces. -/ -@[continuity] -theorem continuous_eval [LocallyCompactPair X Y] : Continuous fun p : C(X, Y) × X => p.1 p.2 := by - simp_rw [continuous_iff_continuousAt, ContinuousAt, (nhds_basis_opens _).tendsto_right_iff] - rintro ⟨f, x⟩ U ⟨hx : f x ∈ U, hU : IsOpen U⟩ - rcases exists_mem_nhds_isCompact_mapsTo f.continuous (hU.mem_nhds hx) with ⟨K, hxK, hK, hKU⟩ - filter_upwards [prod_mem_nhds (eventually_mapsTo hK hU hKU) hxK] using fun _ h ↦ h.1 h.2 +instance [LocallyCompactPair X Y] : ContinuousEval C(X, Y) X Y where + continuous_eval := by + simp_rw [continuous_iff_continuousAt, ContinuousAt, (nhds_basis_opens _).tendsto_right_iff] + rintro ⟨f, x⟩ U ⟨hx : f x ∈ U, hU : IsOpen U⟩ + rcases exists_mem_nhds_isCompact_mapsTo f.continuous (hU.mem_nhds hx) with ⟨K, hxK, hK, hKU⟩ + filter_upwards [prod_mem_nhds (eventually_mapsTo hK hU hKU) hxK] using fun _ h ↦ h.1 h.2 -/-- Evaluation of a continuous map `f` at a point `x` is continuous in `f`. +@[deprecated (since := "2024-10-01")] protected alias continuous_eval := continuous_eval -Porting note: merged `continuous_eval_const` with `continuous_eval_const'` removing unneeded -assumptions. -/ -@[continuity] -theorem continuous_eval_const (a : X) : Continuous fun f : C(X, Y) => f a := - continuous_def.2 fun U hU ↦ by simpa using isOpen_setOf_mapsTo (isCompact_singleton (x := a)) hU +instance : ContinuousEvalConst C(X, Y) X Y where + continuous_eval_const x := + continuous_def.2 fun U hU ↦ by simpa using isOpen_setOf_mapsTo isCompact_singleton hU -/-- Coercion from `C(X, Y)` with compact-open topology to `X → Y` with pointwise convergence -topology is a continuous map. +@[deprecated (since := "2024-10-01")] protected alias continuous_eval_const := continuous_eval_const -Porting note: merged `continuous_coe` with `continuous_coe'` removing unneeded assumptions. -/ +@[deprecated continuous_coeFun (since := "2024-10-01")] theorem continuous_coe : Continuous ((⇑) : C(X, Y) → (X → Y)) := - continuous_pi continuous_eval_const + continuous_coeFun lemma isClosed_setOf_mapsTo {t : Set Y} (ht : IsClosed t) (s : Set X) : IsClosed {f : C(X, Y) | MapsTo f s t} := @@ -192,7 +190,7 @@ lemma isClopen_setOf_mapsTo (hK : IsCompact K) (hU : IsClopen U) : @[norm_cast] lemma specializes_coe {f g : C(X, Y)} : ⇑f ⤳ ⇑g ↔ f ⤳ g := by - refine ⟨fun h ↦ ?_, fun h ↦ h.map continuous_coe⟩ + refine ⟨fun h ↦ ?_, fun h ↦ h.map continuous_coeFun⟩ suffices ∀ K, IsCompact K → ∀ U, IsOpen U → MapsTo g K U → MapsTo f K U by simpa [specializes_iff_pure, nhds_compactOpen] exact fun K _ U hU hg x hx ↦ (h.map (continuous_apply x)).mem_open hU (hg hx) @@ -202,7 +200,7 @@ lemma inseparable_coe {f g : C(X, Y)} : Inseparable (f : X → Y) g ↔ Insepara simp only [inseparable_iff_specializes_and, specializes_coe] instance [T0Space Y] : T0Space C(X, Y) := - t0Space_of_injective_of_continuous DFunLike.coe_injective continuous_coe + t0Space_of_injective_of_continuous DFunLike.coe_injective continuous_coeFun instance [R0Space Y] : R0Space C(X, Y) where specializes_symmetric f g h := by @@ -210,10 +208,10 @@ instance [R0Space Y] : R0Space C(X, Y) where exact h.symm instance [T1Space Y] : T1Space C(X, Y) := - t1Space_of_injective_of_continuous DFunLike.coe_injective continuous_coe + t1Space_of_injective_of_continuous DFunLike.coe_injective continuous_coeFun instance [R1Space Y] : R1Space C(X, Y) := - .of_continuous_specializes_imp continuous_coe fun _ _ ↦ specializes_coe.1 + .of_continuous_specializes_imp continuous_coeFun fun _ _ ↦ specializes_coe.1 instance [T2Space Y] : T2Space C(X, Y) := inferInstance @@ -388,7 +386,8 @@ 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.prodMap continuous_id) + dsimp [Function.comp_def] + exact (continuous_fst.fst.eval continuous_fst.snd).eval continuous_snd /-- 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/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index e5798309d30d4..4cc61d6513e73 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -198,14 +198,14 @@ compact-open topology on the space `C(I,X)` of continuous maps from `I` to `X`. instance topologicalSpace : TopologicalSpace (Path x y) := TopologicalSpace.induced ((↑) : _ → C(I, X)) ContinuousMap.compactOpen -theorem continuous_eval : Continuous fun p : Path x y × I => p.1 p.2 := - ContinuousMap.continuous_eval.comp <| - (continuous_induced_dom (α := Path x y)).prodMap continuous_id +instance : ContinuousEval (Path x y) I X := .of_continuous_forget continuous_induced_dom -@[continuity] +@[deprecated (since := "2024-10-04")] protected alias continuous_eval := continuous_eval + +@[deprecated Continuous.eval (since := "2024-10-04")] theorem _root_.Continuous.path_eval {Y} [TopologicalSpace Y] {f : Y → Path x y} {g : Y → I} - (hf : Continuous f) (hg : Continuous g) : Continuous fun y => f y (g y) := - Continuous.comp continuous_eval (hf.prod_mk hg) + (hf : Continuous f) (hg : Continuous g) : Continuous fun y => f y (g y) := by + continuity theorem continuous_uncurry_iff {Y} [TopologicalSpace Y] {g : Y → Path x y} : Continuous ↿g ↔ Continuous g := @@ -427,7 +427,7 @@ theorem symm_continuous_family {ι : Type*} [TopologicalSpace ι] @[continuity] theorem continuous_symm : Continuous (symm : Path x y → Path y x) := - continuous_uncurry_iff.mp <| symm_continuous_family _ (continuous_fst.path_eval continuous_snd) + continuous_uncurry_iff.mp <| symm_continuous_family _ (continuous_fst.eval continuous_snd) @[continuity] theorem continuous_uncurry_extend_of_continuous_family {ι : Type*} [TopologicalSpace ι] diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index 4e62f17e049be..b6ff6659b4878 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -372,25 +372,27 @@ instance [CommGroup β] [TopologicalGroup β] : TopologicalGroup C(α, β) where uniformContinuous_inv.comp_tendstoUniformlyOn (tendsto_iff_forall_compact_tendstoUniformlyOn.mp Filter.tendsto_id K hK) --- TODO: rewrite the next three lemmas for products and deduce sum case via `to_additive`, once --- definition of `tprod` is in place -/-- If `α` is locally compact, and an infinite sum of functions in `C(α, β)` -converges to `g` (for the compact-open topology), then the pointwise sum converges to `g x` for -all `x ∈ α`. -/ -theorem hasSum_apply {γ : Type*} [AddCommMonoid β] [ContinuousAdd β] - {f : γ → C(α, β)} {g : C(α, β)} (hf : HasSum f g) (x : α) : - HasSum (fun i : γ => f i x) (g x) := by - let ev : C(α, β) →+ β := (Pi.evalAddMonoidHom _ x).comp coeFnAddMonoidHom - exact hf.map ev (ContinuousMap.continuous_eval_const x) - -theorem summable_apply [AddCommMonoid β] [ContinuousAdd β] {γ : Type*} {f : γ → C(α, β)} - (hf : Summable f) (x : α) : Summable fun i : γ => f i x := - (hasSum_apply hf.hasSum x).summable - -theorem tsum_apply [T2Space β] [AddCommMonoid β] [ContinuousAdd β] {γ : Type*} {f : γ → C(α, β)} - (hf : Summable f) (x : α) : - ∑' i : γ, f i x = (∑' i : γ, f i) x := - (hasSum_apply hf.hasSum x).tsum_eq +/-- If an infinite product of functions in `C(α, β)` converges to `g` +(for the compact-open topology), then the pointwise product converges to `g x` for all `x ∈ α`. -/ +@[to_additive + "If an infinite sum of functions in `C(α, β)` converges to `g` (for the compact-open topology), +then the pointwise sum converges to `g x` for all `x ∈ α`."] +theorem hasProd_apply {γ : Type*} [CommMonoid β] [ContinuousMul β] + {f : γ → C(α, β)} {g : C(α, β)} (hf : HasProd f g) (x : α) : + HasProd (fun i : γ => f i x) (g x) := by + let ev : C(α, β) →* β := (Pi.evalMonoidHom _ x).comp coeFnMonoidHom + exact hf.map ev (continuous_eval_const x) + +@[to_additive] +theorem multipliable_apply [CommMonoid β] [ContinuousMul β] {γ : Type*} {f : γ → C(α, β)} + (hf : Multipliable f) (x : α) : Multipliable fun i : γ => f i x := + (hasProd_apply hf.hasProd x).multipliable + +@[to_additive] +theorem tprod_apply [T2Space β] [CommMonoid β] [ContinuousMul β] {γ : Type*} {f : γ → C(α, β)} + (hf : Multipliable f) (x : α) : + ∏' i : γ, f i x = (∏' i : γ, f i) x := + (hasProd_apply hf.hasProd x).tprod_eq end ContinuousMap diff --git a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean index 5fa65e8b7c342..575b4e26ad07c 100644 --- a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean @@ -85,15 +85,25 @@ lemma embedding_toContinuousMap : Embedding ((↑) : C(X, R)₀ → C(X, R)) whe inj _ _ h := ext fun x ↦ congr($(h) x) instance [T0Space R] : T0Space C(X, R)₀ := embedding_toContinuousMap.t0Space +instance [R0Space R] : R0Space C(X, R)₀ := embedding_toContinuousMap.r0Space instance [T1Space R] : T1Space C(X, R)₀ := embedding_toContinuousMap.t1Space +instance [R1Space R] : R1Space C(X, R)₀ := embedding_toContinuousMap.r1Space instance [T2Space R] : T2Space C(X, R)₀ := embedding_toContinuousMap.t2Space +instance [RegularSpace R] : RegularSpace C(X, R)₀ := embedding_toContinuousMap.regularSpace +instance [T3Space R] : T3Space C(X, R)₀ := embedding_toContinuousMap.t3Space + +instance instContinuousEvalConst : ContinuousEvalConst C(X, R)₀ X R := + .of_continuous_forget embedding_toContinuousMap.continuous + +instance instContinuousEval [LocallyCompactPair X R] : ContinuousEval C(X, R)₀ X R := + .of_continuous_forget embedding_toContinuousMap.continuous lemma closedEmbedding_toContinuousMap [T1Space R] : ClosedEmbedding ((↑) : C(X, R)₀ → C(X, R)) where toEmbedding := embedding_toContinuousMap isClosed_range := by rw [range_toContinuousMap] - exact isClosed_singleton.preimage <| ContinuousMap.continuous_eval_const 0 + exact isClosed_singleton.preimage <| continuous_eval_const 0 @[fun_prop] lemma continuous_comp_left {X Y Z : Type*} [TopologicalSpace X] diff --git a/Mathlib/Topology/Hom/ContinuousEval.lean b/Mathlib/Topology/Hom/ContinuousEval.lean new file mode 100644 index 0000000000000..7667c58b0d186 --- /dev/null +++ b/Mathlib/Topology/Hom/ContinuousEval.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2024 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Topology.Hom.ContinuousEvalConst +import Mathlib.Topology.ContinuousMap.Defs + +/-! +# Bundled maps with evaluation continuous in both variables + +In this file we define a class `ContinuousEval F X Y` +saying that `F` is a bundled morphism class (in the sense of `FunLike`) +with a topology such that `fun (f, x) : F × X ↦ f x` is a continuous function. +-/ + +open scoped Topology +open Filter + +/-- A typeclass saying that `F` is a bundled morphism class (in the sense of `FunLike`) +with a topology such that `fun (f, x) : F × X ↦ f x` is a continuous function. -/ +class ContinuousEval (F : Type*) (X Y : outParam Type*) [FunLike F X Y] + [TopologicalSpace F] [TopologicalSpace X] [TopologicalSpace Y] : Prop where + /-- Evaluation of a bundled morphism at a point is continuous in both variables. -/ + continuous_eval : Continuous fun fx : F × X ↦ fx.1 fx.2 + +export ContinuousEval (continuous_eval) + +variable {F X Y Z : Type*} [FunLike F X Y] + [TopologicalSpace F] [TopologicalSpace X] [TopologicalSpace Y] [ContinuousEval F X Y] + [TopologicalSpace Z] {f : Z → F} {g : Z → X} {s : Set Z} {z : Z} + +@[continuity, fun_prop] +protected theorem Continuous.eval (hf : Continuous f) (hg : Continuous g) : + Continuous fun z ↦ f z (g z) := + continuous_eval.comp (hf.prod_mk hg) + +/-- If a type `F'` of bundled morphisms admits a continuous projection +to a type satisfying `ContinuousEval`, +then `F'` satisfies this predicate too. + +The word "forget" in the name is motivated by the term "forgetful functor". -/ +theorem ContinuousEval.of_continuous_forget {F' : Type*} [FunLike F' X Y] [TopologicalSpace F'] + {f : F' → F} (hc : Continuous f) (hf : ∀ g, ⇑(f g) = g := by intro; rfl) : + ContinuousEval F' X Y where + continuous_eval := by simpa only [← hf] using hc.fst'.eval continuous_snd + +instance (priority := 100) ContinuousEval.toContinuousMapClass : ContinuousMapClass F X Y where + map_continuous _ := continuous_const.eval continuous_id + +instance (priority := 100) ContinuousEval.toContinuousEvalConst : ContinuousEvalConst F X Y where + continuous_eval_const _ := continuous_id.eval continuous_const + +protected theorem Filter.Tendsto.eval {α : Type*} {l : Filter α} {f : α → F} {f₀ : F} + {g : α → X} {x₀ : X} (hf : Tendsto f l (𝓝 f₀)) (hg : Tendsto g l (𝓝 x₀)) : + Tendsto (fun a ↦ f a (g a)) l (𝓝 (f₀ x₀)) := + (ContinuousEval.continuous_eval.tendsto _).comp (hf.prod_mk_nhds hg) + +protected nonrec theorem ContinuousAt.eval (hf : ContinuousAt f z) (hg : ContinuousAt g z) : + ContinuousAt (fun z ↦ f z (g z)) z := + hf.eval hg + +protected nonrec theorem ContinuousWithinAt.eval (hf : ContinuousWithinAt f s z) + (hg : ContinuousWithinAt g s z) : ContinuousWithinAt (fun z ↦ f z (g z)) s z := + hf.eval hg + +protected theorem ContinuousOn.eval (hf : ContinuousOn f s) (hg : ContinuousOn g s) : + ContinuousOn (fun z ↦ f z (g z)) s := + fun z hz ↦ (hf z hz).eval (hg z hz) diff --git a/Mathlib/Topology/Hom/ContinuousEvalConst.lean b/Mathlib/Topology/Hom/ContinuousEvalConst.lean new file mode 100644 index 0000000000000..6793028a59d88 --- /dev/null +++ b/Mathlib/Topology/Hom/ContinuousEvalConst.lean @@ -0,0 +1,89 @@ +/- +Copyright (c) 2024 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Topology.Constructions + +/-! +# Bundled morphisms with continuous evaluation at a point + +In this file we define a typeclass +saying that `F` is a type of bundled morphisms (in the sense of `DFunLike`) +with a topology on `F` such that evaluation at a point is continuous in `f : F`. + +## Implementation Notes + +For now, we define the typeclass for non-dependent bundled functions only. +Whenever we add a type of bundled dependent functions with a topology having this property, +we may decide to generalize from `FunLike` to `DFunLike`. +-/ + +open scoped Topology +open Filter + +/-- A typeclass saying that `F` is a type of bundled morphisms (in the sense of `DFunLike`) +with a topology on `F` such that evaluation at a point is continuous in `f : F`. -/ +class ContinuousEvalConst (F : Type*) (α X : outParam Type*) [FunLike F α X] + [TopologicalSpace F] [TopologicalSpace X] : Prop where + continuous_eval_const (x : α) : Continuous fun f : F ↦ f x + +export ContinuousEvalConst (continuous_eval_const) + +section ContinuousEvalConst + +variable {F α X Z : Type*} [FunLike F α X] [TopologicalSpace F] [TopologicalSpace X] + [ContinuousEvalConst F α X] [TopologicalSpace Z] {f : Z → F} {s : Set Z} {z : Z} + +/-- If a type `F'` of bundled morphisms admits a continuous projection +to a type satisfying `ContinuousEvalConst`, +then `F'` satisfies this predicate too. + +The word "forget" in the name is motivated by the term "forgetful functor". -/ +theorem ContinuousEvalConst.of_continuous_forget {F' : Type*} [FunLike F' α X] [TopologicalSpace F'] + {f : F' → F} (hc : Continuous f) (hf : ∀ g, ⇑(f g) = g := by intro; rfl) : + ContinuousEvalConst F' α X where + continuous_eval_const x := by simpa only [← hf] using (continuous_eval_const x).comp hc + +@[continuity, fun_prop] +protected theorem Continuous.eval_const (hf : Continuous f) (x : α) : Continuous (f · x) := + (continuous_eval_const x).comp hf + +theorem continuous_coeFun : Continuous (DFunLike.coe : F → α → X) := + continuous_pi continuous_eval_const + +protected theorem Continuous.coeFun (hf : Continuous f) : Continuous fun z ↦ ⇑(f z) := + continuous_pi hf.eval_const + +protected theorem Filter.Tendsto.eval_const {ι : Type*} {l : Filter ι} {f : ι → F} {g : F} + (hf : Tendsto f l (𝓝 g)) (a : α) : Tendsto (f · a) l (𝓝 (g a)) := + ((continuous_id.eval_const a).tendsto _).comp hf + +protected theorem Filter.Tendsto.coeFun {ι : Type*} {l : Filter ι} {f : ι → F} {g : F} + (hf : Tendsto f l (𝓝 g)) : Tendsto (fun i ↦ ⇑(f i)) l (𝓝 ⇑g) := + (continuous_id.coeFun.tendsto _).comp hf + +protected nonrec theorem ContinuousAt.eval_const (hf : ContinuousAt f z) (x : α) : + ContinuousAt (f · x) z := + hf.eval_const x + +protected nonrec theorem ContinuousAt.coeFun (hf : ContinuousAt f z) : + ContinuousAt (fun z ↦ ⇑(f z)) z := + hf.coeFun + +protected nonrec theorem ContinuousWithinAt.eval_const (hf : ContinuousWithinAt f s z) (x : α) : + ContinuousWithinAt (f · x) s z := + hf.eval_const x + +protected nonrec theorem ContinuousWithinAt.coeFun (hf : ContinuousWithinAt f s z) : + ContinuousWithinAt (fun z ↦ ⇑(f z)) s z := + hf.coeFun + +protected theorem ContinuousOn.eval_const (hf : ContinuousOn f s) (x : α) : + ContinuousOn (f · x) s := + fun z hz ↦ (hf z hz).eval_const x + +protected theorem ContinuousOn.coeFun (hf : ContinuousOn f s) (x : α) : ContinuousOn (f · x) s := + fun z hz ↦ (hf z hz).eval_const x + +end ContinuousEvalConst diff --git a/Mathlib/Topology/Homotopy/HSpaces.lean b/Mathlib/Topology/Homotopy/HSpaces.lean index cf728834e55ae..fba0212ad88f5 100644 --- a/Mathlib/Topology/Homotopy/HSpaces.lean +++ b/Mathlib/Topology/Homotopy/HSpaces.lean @@ -212,7 +212,7 @@ def delayReflRight (θ : I) (γ : Path x y) : Path x y where theorem continuous_delayReflRight : Continuous fun p : I × Path x y => delayReflRight p.1 p.2 := continuous_uncurry_iff.mp <| - (continuous_snd.comp continuous_fst).path_eval <| + (continuous_snd.comp continuous_fst).eval <| continuous_qRight.comp <| continuous_snd.prod_mk <| continuous_fst.comp continuous_fst theorem delayReflRight_zero (γ : Path x y) : delayReflRight 0 γ = γ.trans (Path.refl y) := by diff --git a/Mathlib/Topology/Homotopy/HomotopyGroup.lean b/Mathlib/Topology/Homotopy/HomotopyGroup.lean index fac80ea72c606..a9ec3d6dc3be4 100644 --- a/Mathlib/Topology/Homotopy/HomotopyGroup.lean +++ b/Mathlib/Topology/Homotopy/HomotopyGroup.lean @@ -111,6 +111,11 @@ theorem ext (f g : Ω^ N X x) (H : ∀ y, f y = g y) : f = g := theorem mk_apply (f : C(I^N, X)) (H y) : (⟨f, H⟩ : Ω^ N X x) y = f y := rfl +instance instContinuousEval : ContinuousEval (Ω^ N X x) (I^N) X := + .of_continuous_forget continuous_subtype_val + +instance instContinuousEvalConst : ContinuousEvalConst (Ω^ N X x) (I^N) X := inferInstance + /-- Copy of a `GenLoop` with a new map from the unit cube equal to the old one. Useful to fix definitional equalities. -/ def copy (f : Ω^ N X x) (g : (I^N) → X) (h : g = f) : Ω^ N X x := @@ -184,7 +189,7 @@ def toLoop (i : N) (p : Ω^ N X x) : Ω (Ω^ { j // j ≠ i } X x) const where theorem continuous_toLoop (i : N) : Continuous (@toLoop N X _ x _ i) := Path.continuous_uncurry_iff.1 <| Continuous.subtype_mk - (ContinuousMap.continuous_eval.comp <| + (continuous_eval.comp <| Continuous.prodMap (ContinuousMap.continuous_curry.comp <| (ContinuousMap.continuous_comp_left _).comp continuous_subtype_val) @@ -372,7 +377,6 @@ def genLoopHomeoOfIsEmpty (N x) [IsEmpty N] : Ω^ N X x ≃ₜ X where invFun y := ⟨ContinuousMap.const _ y, fun _ ⟨i, _⟩ => isEmptyElim i⟩ left_inv f := by ext; exact congr_arg f (Subsingleton.elim _ _) right_inv _ := rfl - continuous_toFun := (ContinuousMap.continuous_eval_const (0 : N → I)).comp continuous_induced_dom continuous_invFun := ContinuousMap.const'.2.subtype_mk _ /-- The homotopy "group" indexed by an empty type is in bijection with diff --git a/Mathlib/Topology/Homotopy/Path.lean b/Mathlib/Topology/Homotopy/Path.lean index 254ed2b7f70dc..a5ad060e6071e 100644 --- a/Mathlib/Topology/Homotopy/Path.lean +++ b/Mathlib/Topology/Homotopy/Path.lean @@ -188,15 +188,7 @@ def reparam (p : Path x₀ x₁) (f : I → I) (hf : Continuous f) (hf₀ : f 0 · rw [Set.mem_singleton_iff] at hx rw [hx] simp [hf₁] - continuous_toFun := by - -- Porting note: was `continuity` in auto-param - refine continuous_const.path_eval ?_ - apply Continuous.subtype_mk - apply Continuous.add <;> apply Continuous.mul - · exact continuous_induced_dom.comp (unitInterval.continuous_symm.comp continuous_fst) - · continuity - · continuity - · continuity + continuous_toFun := by fun_prop /-- Suppose `F : Homotopy p q`. Then we have a `Homotopy p.symm q.symm` by reversing the second argument. From 8e0ebee297fc69a4ac143794e1e26ad52d25a587 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 09:18:38 +0000 Subject: [PATCH 053/131] chore(*): fix some TODOs (#17891) Also drop an unneeded assumption in a lemma. --- .../Order/Monoid/Unbundled/ExistsOfLE.lean | 7 ++++--- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 5 ++--- Mathlib/Analysis/Oscillation.lean | 2 +- Mathlib/Data/Finset/Card.lean | 3 +-- Mathlib/Data/NNReal/Basic.lean | 20 ++++++++----------- .../Integral/RieszMarkovKakutani.lean | 2 +- .../MetricSpace/HausdorffDistance.lean | 2 +- 7 files changed, 18 insertions(+), 23 deletions(-) diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean index da6e069dc2d22..bfca4ed38d9a9 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean @@ -65,20 +65,21 @@ end MulOneClass section ExistsMulOfLE variable [LinearOrder α] [DenselyOrdered α] [Monoid α] [ExistsMulOfLE α] - [CovariantClass α α (· * ·) (· < ·)] [ContravariantClass α α (· * ·) (· < ·)] {a b : α} + [ContravariantClass α α (· * ·) (· < ·)] {a b : α} @[to_additive] theorem le_of_forall_one_lt_le_mul (h : ∀ ε : α, 1 < ε → a ≤ b * ε) : a ≤ b := le_of_forall_le_of_dense fun x hxb => by obtain ⟨ε, rfl⟩ := exists_mul_of_le hxb.le - exact h _ ((lt_mul_iff_one_lt_right' b).1 hxb) + exact h _ (one_lt_of_lt_mul_right hxb) @[to_additive] theorem le_of_forall_one_lt_lt_mul' (h : ∀ ε : α, 1 < ε → a < b * ε) : a ≤ b := le_of_forall_one_lt_le_mul fun ε hε => (h ε hε).le @[to_additive] -theorem le_iff_forall_one_lt_lt_mul' : a ≤ b ↔ ∀ ε, 1 < ε → a < b * ε := +theorem le_iff_forall_one_lt_lt_mul' [CovariantClass α α (· * ·) (· < ·)] : + a ≤ b ↔ ∀ ε, 1 < ε → a < b * ε := ⟨fun h _ => lt_mul_of_le_of_one_lt h, le_of_forall_one_lt_lt_mul'⟩ end ExistsMulOfLE diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index f4d4b2b3d92cc..23b46aad1701c 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -1562,9 +1562,8 @@ end prodMap section AlgebraInverse -variable (𝕜) {R : Type*} [NormedRing R] --- Porting note: this couldn't be on the same line as the binder type update of `𝕜` -variable [NormedAlgebra 𝕜 R] +variable (𝕜) +variable {R : Type*} [NormedRing R] [NormedAlgebra 𝕜 R] open NormedRing ContinuousLinearMap Ring diff --git a/Mathlib/Analysis/Oscillation.lean b/Mathlib/Analysis/Oscillation.lean index c6327e3b9a059..cfc6cf9116dd9 100644 --- a/Mathlib/Analysis/Oscillation.lean +++ b/Mathlib/Analysis/Oscillation.lean @@ -52,7 +52,7 @@ namespace ContinuousWithinAt theorem oscillationWithin_eq_zero [TopologicalSpace E] {f : E → F} {D : Set E} {x : E} (hf : ContinuousWithinAt f D x) : oscillationWithin f D x = 0 := by - refine le_antisymm (le_of_forall_pos_le_add fun ε hε _ ↦ ?_) (zero_le _) + refine le_antisymm (_root_.le_of_forall_pos_le_add fun ε hε ↦ ?_) (zero_le _) rw [zero_add] have : ball (f x) (ε / 2) ∈ (𝓝[D] x).map f := hf <| ball_mem_nhds _ (by simp [ne_of_gt hε]) refine (biInf_le diam this).trans (le_of_le_of_eq diam_ball ?_) diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 376e7fe70747f..67ac01fae4261 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -25,8 +25,7 @@ This defines the cardinality of a `Finset` and provides induction principles for -/ assert_not_exists MonoidWithZero --- TODO: After a lot more work, --- assert_not_exists OrderedCommMonoid +assert_not_exists OrderedCommMonoid open Function Multiset Nat diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index f0ef943d2ea06..0db61f23290b1 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -68,8 +68,8 @@ scoped notation "ℝ≥0" => NNReal noncomputable instance : FloorSemiring ℝ≥0 := Nonneg.floorSemiring instance instDenselyOrdered : DenselyOrdered ℝ≥0 := Nonneg.instDenselyOrdered instance : OrderBot ℝ≥0 := inferInstance -instance : Archimedean ℝ≥0 := Nonneg.instArchimedean -instance : MulArchimedean ℝ≥0 := Nonneg.instMulArchimedean +instance instArchimedean : Archimedean ℝ≥0 := Nonneg.instArchimedean +instance instMulArchimedean : MulArchimedean ℝ≥0 := Nonneg.instMulArchimedean noncomputable instance : Sub ℝ≥0 := Nonneg.sub noncomputable instance : OrderedSub ℝ≥0 := Nonneg.orderedSub @@ -493,17 +493,13 @@ theorem le_iInf_add_iInf {ι ι' : Sort*} [Nonempty ι] [Nonempty ι'] {f : ι rw [← NNReal.coe_le_coe, NNReal.coe_add, coe_iInf, coe_iInf] exact le_ciInf_add_ciInf h -example : Archimedean ℝ≥0 := by infer_instance +-- Short-circuit instance search +instance instCovariantClassAddLE : CovariantClass ℝ≥0 ℝ≥0 (· + ·) (· ≤ ·) := inferInstance +instance instContravariantClassAddLT : ContravariantClass ℝ≥0 ℝ≥0 (· + ·) (· < ·) := inferInstance +instance instCovariantClassMulLE : CovariantClass ℝ≥0 ℝ≥0 (· * ·) (· ≤ ·) := inferInstance --- Porting note (#11215): TODO: remove? -instance covariant_add : CovariantClass ℝ≥0 ℝ≥0 (· + ·) (· ≤ ·) := inferInstance - -instance contravariant_add : ContravariantClass ℝ≥0 ℝ≥0 (· + ·) (· < ·) := inferInstance - -instance covariant_mul : CovariantClass ℝ≥0 ℝ≥0 (· * ·) (· ≤ ·) := inferInstance - --- Porting note (#11215): TODO: delete? -nonrec theorem le_of_forall_pos_le_add {a b : ℝ≥0} (h : ∀ ε, 0 < ε → a ≤ b + ε) : a ≤ b := +@[deprecated le_of_forall_pos_le_add (since := "2024-10-17")] +protected theorem le_of_forall_pos_le_add {a b : ℝ≥0} (h : ∀ ε, 0 < ε → a ≤ b + ε) : a ≤ b := le_of_forall_pos_le_add h theorem lt_iff_exists_rat_btwn (a b : ℝ≥0) : diff --git a/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean b/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean index b6a9c18db7857..da3bb92699829 100644 --- a/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean +++ b/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean @@ -86,7 +86,7 @@ theorem exists_lt_rieszContentAux_add_pos (K : Compacts X) {ε : ℝ≥0} (εpos finitely subadditive: `λ(K₁ ∪ K₂) ≤ λ(K₁) + λ(K₂)` for any compact subsets `K₁, K₂ ⊆ X`. -/ theorem rieszContentAux_sup_le (K1 K2 : Compacts X) : rieszContentAux Λ (K1 ⊔ K2) ≤ rieszContentAux Λ K1 + rieszContentAux Λ K2 := by - apply NNReal.le_of_forall_pos_le_add + apply _root_.le_of_forall_pos_le_add intro ε εpos --get test functions s.t. `λ(Ki) ≤ Λfi ≤ λ(Ki) + ε/2, i=1,2` obtain ⟨f1, f_test_function_K1⟩ := exists_lt_rieszContentAux_add_pos Λ K1 (half_pos εpos) diff --git a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean index b898627b74b4e..06799e111037e 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean @@ -451,7 +451,7 @@ theorem infEdist_ne_top (h : s.Nonempty) : infEdist x s ≠ ⊤ := by rcases h with ⟨y, hy⟩ exact ne_top_of_le_ne_top (edist_ne_top _ _) (infEdist_le_edist_of_mem hy) --- Porting note (#11215): TODO: make it a `simp` lemma +@[simp] theorem infEdist_eq_top_iff : infEdist x s = ∞ ↔ s = ∅ := by rcases s.eq_empty_or_nonempty with rfl | hs <;> simp [*, Nonempty.ne_empty, infEdist_ne_top] From 34f08fb6b3608c4b798f493e995f4eb615a9b01d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 18 Oct 2024 09:52:32 +0000 Subject: [PATCH 054/131] feat: generalize results from EuclideanSpace to PiLp (#17663) The EuclideanSpace results are left in place for convenience. This generalizes leanprover-community/mathlib3#15363 Co-authored-by: Eric Wieser --- Mathlib.lean | 2 + .../Analysis/Calculus/ContDiff/WithLp.lean | 41 ++++++++++++++ Mathlib/Analysis/Calculus/FDeriv/WithLp.lean | 56 +++++++++++++++++++ .../Analysis/InnerProductSpace/Calculus.lean | 54 ++++++++---------- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 13 +---- Mathlib/Analysis/Normed/Lp/PiLp.lean | 16 +++++- 6 files changed, 141 insertions(+), 41 deletions(-) create mode 100644 Mathlib/Analysis/Calculus/ContDiff/WithLp.lean create mode 100644 Mathlib/Analysis/Calculus/FDeriv/WithLp.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1e5548b290acc..3c455bcceb7e9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1004,6 +1004,7 @@ import Mathlib.Analysis.Calculus.ContDiff.Defs import Mathlib.Analysis.Calculus.ContDiff.FTaylorSeries import Mathlib.Analysis.Calculus.ContDiff.FiniteDimension import Mathlib.Analysis.Calculus.ContDiff.RCLike +import Mathlib.Analysis.Calculus.ContDiff.WithLp import Mathlib.Analysis.Calculus.Darboux import Mathlib.Analysis.Calculus.Deriv.Abs import Mathlib.Analysis.Calculus.Deriv.Add @@ -1041,6 +1042,7 @@ import Mathlib.Analysis.Calculus.FDeriv.Prod import Mathlib.Analysis.Calculus.FDeriv.RestrictScalars import Mathlib.Analysis.Calculus.FDeriv.Star import Mathlib.Analysis.Calculus.FDeriv.Symmetric +import Mathlib.Analysis.Calculus.FDeriv.WithLp import Mathlib.Analysis.Calculus.FormalMultilinearSeries import Mathlib.Analysis.Calculus.Gradient.Basic import Mathlib.Analysis.Calculus.Implicit diff --git a/Mathlib/Analysis/Calculus/ContDiff/WithLp.lean b/Mathlib/Analysis/Calculus/ContDiff/WithLp.lean new file mode 100644 index 0000000000000..4c99b51e455ba --- /dev/null +++ b/Mathlib/Analysis/Calculus/ContDiff/WithLp.lean @@ -0,0 +1,41 @@ +/- +Copyright (c) 2022 Anatole Dedecker. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anatole Dedecker, Eric Wieser +-/ +import Mathlib.Analysis.Calculus.ContDiff.Basic +import Mathlib.Analysis.Normed.Lp.PiLp + +/-! +# Derivatives on `WithLp` +-/ + +section PiLp + +open ContinuousLinearMap + +variable {𝕜 ι : Type*} {E : ι → Type*} {H : Type*} +variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup H] [∀ i, NormedAddCommGroup (E i)] + [∀ i, NormedSpace 𝕜 (E i)] [NormedSpace 𝕜 H] [Fintype ι] (p) [Fact (1 ≤ p)] + {f : H → PiLp p E} {f' : H →L[𝕜] PiLp p E} {t : Set H} {y : H} + +theorem contDiffWithinAt_piLp {n : ℕ∞} : + ContDiffWithinAt 𝕜 n f t y ↔ ∀ i, ContDiffWithinAt 𝕜 n (fun x => f x i) t y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiffWithinAt_iff, contDiffWithinAt_pi] + rfl + +theorem contDiffAt_piLp {n : ℕ∞} : + ContDiffAt 𝕜 n f y ↔ ∀ i, ContDiffAt 𝕜 n (fun x => f x i) y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiffAt_iff, contDiffAt_pi] + rfl + +theorem contDiffOn_piLp {n : ℕ∞} : + ContDiffOn 𝕜 n f t ↔ ∀ i, ContDiffOn 𝕜 n (fun x => f x i) t := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiffOn_iff, contDiffOn_pi] + rfl + +theorem contDiff_piLp {n : ℕ∞} : ContDiff 𝕜 n f ↔ ∀ i, ContDiff 𝕜 n fun x => f x i := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiff_iff, contDiff_pi] + rfl + +end PiLp diff --git a/Mathlib/Analysis/Calculus/FDeriv/WithLp.lean b/Mathlib/Analysis/Calculus/FDeriv/WithLp.lean new file mode 100644 index 0000000000000..1953b606870a9 --- /dev/null +++ b/Mathlib/Analysis/Calculus/FDeriv/WithLp.lean @@ -0,0 +1,56 @@ +/- +Copyright (c) 2022 Anatole Dedecker. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anatole Dedecker, Eric Wieser +-/ +import Mathlib.Analysis.Calculus.FDeriv.Prod +import Mathlib.Analysis.Calculus.FDeriv.Equiv +import Mathlib.Analysis.Normed.Lp.PiLp + + +/-! +# Derivatives on `WithLp` +-/ + +section PiLp + +open ContinuousLinearMap + +variable {𝕜 ι : Type*} {E : ι → Type*} {H : Type*} +variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup H] [∀ i, NormedAddCommGroup (E i)] + [∀ i, NormedSpace 𝕜 (E i)] [NormedSpace 𝕜 H] [Fintype ι] (p) [Fact (1 ≤ p)] + {f : H → PiLp p E} {f' : H →L[𝕜] PiLp p E} {t : Set H} {y : H} + +theorem differentiableWithinAt_piLp : + DifferentiableWithinAt 𝕜 f t y ↔ ∀ i, DifferentiableWithinAt 𝕜 (fun x => f x i) t y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_differentiableWithinAt_iff, + differentiableWithinAt_pi] + rfl + +theorem differentiableAt_piLp : + DifferentiableAt 𝕜 f y ↔ ∀ i, DifferentiableAt 𝕜 (fun x => f x i) y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_differentiableAt_iff, differentiableAt_pi] + rfl + +theorem differentiableOn_piLp : + DifferentiableOn 𝕜 f t ↔ ∀ i, DifferentiableOn 𝕜 (fun x => f x i) t := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_differentiableOn_iff, differentiableOn_pi] + rfl + +theorem differentiable_piLp : Differentiable 𝕜 f ↔ ∀ i, Differentiable 𝕜 fun x => f x i := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_differentiable_iff, differentiable_pi] + rfl + +theorem hasStrictFDerivAt_piLp : + HasStrictFDerivAt f f' y ↔ + ∀ i, HasStrictFDerivAt (fun x => f x i) (PiLp.proj _ _ i ∘L f') y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_hasStrictFDerivAt_iff, hasStrictFDerivAt_pi'] + rfl + +theorem hasFDerivWithinAt_piLp : + HasFDerivWithinAt f f' t y ↔ + ∀ i, HasFDerivWithinAt (fun x => f x i) (PiLp.proj _ _ i ∘L f') t y := by + rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_hasFDerivWithinAt_iff, hasFDerivWithinAt_pi'] + rfl + +end PiLp diff --git a/Mathlib/Analysis/InnerProductSpace/Calculus.lean b/Mathlib/Analysis/InnerProductSpace/Calculus.lean index 44808e4d72423..018aa26547ec5 100644 --- a/Mathlib/Analysis/InnerProductSpace/Calculus.lean +++ b/Mathlib/Analysis/InnerProductSpace/Calculus.lean @@ -6,6 +6,8 @@ Authors: Yury Kudryashov import Mathlib.Analysis.InnerProductSpace.PiL2 import Mathlib.Analysis.SpecialFunctions.Sqrt import Mathlib.Analysis.NormedSpace.HomeomorphBall +import Mathlib.Analysis.Calculus.ContDiff.WithLp +import Mathlib.Analysis.Calculus.FDeriv.WithLp /-! # Calculus in inner product spaces @@ -264,60 +266,52 @@ end DerivInner section PiLike +/-! ### Convenience aliases of `PiLp` lemmas for `EuclideanSpace` -/ + open ContinuousLinearMap variable {𝕜 ι H : Type*} [RCLike 𝕜] [NormedAddCommGroup H] [NormedSpace 𝕜 H] [Fintype ι] {f : H → EuclideanSpace 𝕜 ι} {f' : H →L[𝕜] EuclideanSpace 𝕜 ι} {t : Set H} {y : H} theorem differentiableWithinAt_euclidean : - DifferentiableWithinAt 𝕜 f t y ↔ ∀ i, DifferentiableWithinAt 𝕜 (fun x => f x i) t y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_differentiableWithinAt_iff, differentiableWithinAt_pi] - rfl + DifferentiableWithinAt 𝕜 f t y ↔ ∀ i, DifferentiableWithinAt 𝕜 (fun x => f x i) t y := + differentiableWithinAt_piLp _ theorem differentiableAt_euclidean : - DifferentiableAt 𝕜 f y ↔ ∀ i, DifferentiableAt 𝕜 (fun x => f x i) y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_differentiableAt_iff, differentiableAt_pi] - rfl + DifferentiableAt 𝕜 f y ↔ ∀ i, DifferentiableAt 𝕜 (fun x => f x i) y := + differentiableAt_piLp _ theorem differentiableOn_euclidean : - DifferentiableOn 𝕜 f t ↔ ∀ i, DifferentiableOn 𝕜 (fun x => f x i) t := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_differentiableOn_iff, differentiableOn_pi] - rfl + DifferentiableOn 𝕜 f t ↔ ∀ i, DifferentiableOn 𝕜 (fun x => f x i) t := + differentiableOn_piLp _ -theorem differentiable_euclidean : Differentiable 𝕜 f ↔ ∀ i, Differentiable 𝕜 fun x => f x i := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_differentiable_iff, differentiable_pi] - rfl +theorem differentiable_euclidean : Differentiable 𝕜 f ↔ ∀ i, Differentiable 𝕜 fun x => f x i := + differentiable_piLp _ theorem hasStrictFDerivAt_euclidean : HasStrictFDerivAt f f' y ↔ - ∀ i, HasStrictFDerivAt (fun x => f x i) (EuclideanSpace.proj i ∘L f') y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_hasStrictFDerivAt_iff, hasStrictFDerivAt_pi'] - rfl + ∀ i, HasStrictFDerivAt (fun x => f x i) (PiLp.proj _ _ i ∘L f') y := + hasStrictFDerivAt_piLp _ theorem hasFDerivWithinAt_euclidean : HasFDerivWithinAt f f' t y ↔ - ∀ i, HasFDerivWithinAt (fun x => f x i) (EuclideanSpace.proj i ∘L f') t y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_hasFDerivWithinAt_iff, hasFDerivWithinAt_pi'] - rfl + ∀ i, HasFDerivWithinAt (fun x => f x i) (PiLp.proj _ _ i ∘L f') t y := + hasFDerivWithinAt_piLp _ theorem contDiffWithinAt_euclidean {n : ℕ∞} : - ContDiffWithinAt 𝕜 n f t y ↔ ∀ i, ContDiffWithinAt 𝕜 n (fun x => f x i) t y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_contDiffWithinAt_iff, contDiffWithinAt_pi] - rfl + ContDiffWithinAt 𝕜 n f t y ↔ ∀ i, ContDiffWithinAt 𝕜 n (fun x => f x i) t y := + contDiffWithinAt_piLp _ theorem contDiffAt_euclidean {n : ℕ∞} : - ContDiffAt 𝕜 n f y ↔ ∀ i, ContDiffAt 𝕜 n (fun x => f x i) y := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_contDiffAt_iff, contDiffAt_pi] - rfl + ContDiffAt 𝕜 n f y ↔ ∀ i, ContDiffAt 𝕜 n (fun x => f x i) y := + contDiffAt_piLp _ theorem contDiffOn_euclidean {n : ℕ∞} : - ContDiffOn 𝕜 n f t ↔ ∀ i, ContDiffOn 𝕜 n (fun x => f x i) t := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_contDiffOn_iff, contDiffOn_pi] - rfl + ContDiffOn 𝕜 n f t ↔ ∀ i, ContDiffOn 𝕜 n (fun x => f x i) t := + contDiffOn_piLp _ -theorem contDiff_euclidean {n : ℕ∞} : ContDiff 𝕜 n f ↔ ∀ i, ContDiff 𝕜 n fun x => f x i := by - rw [← (EuclideanSpace.equiv ι 𝕜).comp_contDiff_iff, contDiff_pi] - rfl +theorem contDiff_euclidean {n : ℕ∞} : ContDiff 𝕜 n f ↔ ∀ i, ContDiff 𝕜 n fun x => f x i := + contDiff_piLp _ end PiLike diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 2c14b91ede9f1..3a77d1f7c5328 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -200,18 +200,11 @@ abbrev EuclideanSpace.equiv : EuclideanSpace 𝕜 ι ≃L[𝕜] ι → 𝕜 := variable {ι 𝕜} --- TODO : This should be generalized to `PiLp`. /-- The projection on the `i`-th coordinate of `EuclideanSpace 𝕜 ι`, as a linear map. -/ -@[simps!] -def EuclideanSpace.projₗ (i : ι) : EuclideanSpace 𝕜 ι →ₗ[𝕜] 𝕜 := - (LinearMap.proj i).comp (WithLp.linearEquiv 2 𝕜 (ι → 𝕜) : EuclideanSpace 𝕜 ι →ₗ[𝕜] ι → 𝕜) +abbrev EuclideanSpace.projₗ (i : ι) : EuclideanSpace 𝕜 ι →ₗ[𝕜] 𝕜 := PiLp.projₗ _ _ i --- TODO : This should be generalized to `PiLp`. -/-- The projection on the `i`-th coordinate of `EuclideanSpace 𝕜 ι`, -as a continuous linear map. -/ -@[simps! apply coe] -def EuclideanSpace.proj (i : ι) : EuclideanSpace 𝕜 ι →L[𝕜] 𝕜 := - ⟨EuclideanSpace.projₗ i, continuous_apply i⟩ +/-- The projection on the `i`-th coordinate of `EuclideanSpace 𝕜 ι`, as a continuous linear map. -/ +abbrev EuclideanSpace.proj (i : ι) : EuclideanSpace 𝕜 ι →L[𝕜] 𝕜 := PiLp.proj _ _ i section DecEq diff --git a/Mathlib/Analysis/Normed/Lp/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean index 3902519c5b489..3b9cc58963b91 100644 --- a/Mathlib/Analysis/Normed/Lp/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -91,7 +91,7 @@ section /- Register simplification lemmas for the applications of `PiLp` elements, as the usual lemmas for Pi types will not trigger. -/ variable {𝕜 p α} -variable [SeminormedRing 𝕜] [∀ i, SeminormedAddCommGroup (β i)] +variable [Semiring 𝕜] [∀ i, SeminormedAddCommGroup (β i)] variable [∀ i, Module 𝕜 (β i)] (c : 𝕜) variable (x y : PiLp p β) (i : ι) @@ -121,6 +121,13 @@ theorem smul_apply : (c • x) i = c • x i := @[simp] theorem neg_apply : (-x) i = -x i := rfl + +variable (p) in +/-- The projection on the `i`-th coordinate of `WithLp p (∀ i, α i)`, as a linear map. -/ +@[simps!] +def projₗ (i : ι) : PiLp p β →ₗ[𝕜] β i := + (LinearMap.proj i : (∀ i, β i) →ₗ[𝕜] β i) ∘ₗ (WithLp.linearEquiv p 𝕜 (∀ i, β i)).toLinearMap + end /-! Note that the unapplied versions of these lemmas are deliberately omitted, as they break @@ -891,6 +898,13 @@ protected def continuousLinearEquiv : PiLp p β ≃L[𝕜] ∀ i, β i where continuous_toFun := continuous_equiv _ _ continuous_invFun := continuous_equiv_symm _ _ +variable {𝕜} in +/-- The projection on the `i`-th coordinate of `PiLp p β`, as a continuous linear map. -/ +@[simps!] +def proj (i : ι) : PiLp p β →L[𝕜] β i where + __ := projₗ p β i + cont := continuous_apply i + end Fintype section Basis From 6338c7313383353c208346451dbdf0cc037ac8a2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 18 Oct 2024 09:52:33 +0000 Subject: [PATCH 055/131] fix: improve line-wrapping of (D)Finsupp reprs (#17882) These were previously wrapping very poorly. Also update tests to use `#guard_msgs` more effectively, and fix a typo in the delaborators for DFinsupp. --- Mathlib/Data/DFinsupp/Notation.lean | 10 +++---- Mathlib/Data/Finsupp/Notation.lean | 6 ++-- test/dfinsupp_notation.lean | 41 ++++++++++++++++++++------- test/finsupp_notation.lean | 44 +++++++++++++++++++++++++---- 4 files changed, 78 insertions(+), 23 deletions(-) diff --git a/Mathlib/Data/DFinsupp/Notation.lean b/Mathlib/Data/DFinsupp/Notation.lean index f4782385250a8..bdb9cdad1a53f 100644 --- a/Mathlib/Data/DFinsupp/Notation.lean +++ b/Mathlib/Data/DFinsupp/Notation.lean @@ -44,13 +44,13 @@ def elabUpdate₀ : Elab.Term.TermElab | _ => fun _ => Elab.throwUnsupportedSyntax /-- Unexpander for the `fun₀ | i => x` notation. -/ -@[app_unexpander Finsupp.single] +@[app_unexpander DFinsupp.single] def singleUnexpander : Lean.PrettyPrinter.Unexpander | `($_ $pat $val) => `(fun₀ | $pat => $val) | _ => throw () /-- Unexpander for the `fun₀ | i => x` notation. -/ -@[app_unexpander Finsupp.update] +@[app_unexpander DFinsupp.update] def updateUnexpander : Lean.PrettyPrinter.Unexpander | `($_ $f $pat $val) => match f with | `(fun₀ $xs:matchAlt*) => `(fun₀ $xs:matchAlt* | $pat => $val) @@ -68,9 +68,9 @@ unsafe instance {α : Type*} {β : α → Type*} [Repr α] [∀ i, Repr (β i)] if vals.length = 0 then "0" else - let ret := "fun₀" ++ - Std.Format.join (vals_dedup.map <| - fun a => f!" | " ++ a.1 ++ f!" => " ++ a.2) + let ret : Std.Format := f!"fun₀" ++ .nest 2 ( + .group (.join <| vals_dedup.map fun a => + .line ++ .group (f!"| {a.1} =>" ++ .line ++ a.2))) if p ≥ leadPrec then Format.paren ret else ret end DFinsupp diff --git a/Mathlib/Data/Finsupp/Notation.lean b/Mathlib/Data/Finsupp/Notation.lean index 2091f9b3cb8d1..12c11f9e628e5 100644 --- a/Mathlib/Data/Finsupp/Notation.lean +++ b/Mathlib/Data/Finsupp/Notation.lean @@ -83,9 +83,9 @@ unsafe instance instRepr {α β} [Repr α] [Repr β] [Zero β] : Repr (α →₀ if f.support.card = 0 then "0" else - let ret := "fun₀" ++ - Std.Format.join (f.support.val.unquot.map <| - fun a => " | " ++ repr a ++ " => " ++ repr (f a)) + let ret : Std.Format := f!"fun₀" ++ .nest 2 ( + .group (.join <| f.support.val.unquot.map fun a => + .line ++ .group (f!"| {repr a} =>" ++ .line ++ repr (f a)))) if p ≥ leadPrec then Format.paren ret else ret -- This cannot be put in `Mathlib.Data.DFinsupp.Notation` where it belongs, since doc-strings diff --git a/test/dfinsupp_notation.lean b/test/dfinsupp_notation.lean index ed6511532b5f9..ea4d3815079ae 100644 --- a/test/dfinsupp_notation.lean +++ b/test/dfinsupp_notation.lean @@ -7,20 +7,41 @@ example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4 : Π₀ i, Fin (i + 10)) 1 = 3 := by example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4 : Π₀ i, Fin (i + 10)) 2 = 3 := by simp example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4 : Π₀ i, Fin (i + 10)) 3 = 4 := by simp +section Repr + +/-- info: fun₀ | 1 => 3 | 2 => 3 : Π₀ (i : ℕ), Fin (i + 10) -/ +#guard_msgs in +#check (fun₀ | 1 => 3 | 2 => 3 : Π₀ i, Fin (i + 10)) + +/-- info: fun₀ | 2 => 7 -/ +#guard_msgs in +#eval ((fun₀ | 1 => 3 | 2 => 3) + (fun₀ | 1 => -3 | 2 => 4) : Π₀ _, ℤ) + /-- -info: +info: fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 -/ #guard_msgs in -#eval show Lean.MetaM Unit from - guard <| - reprStr (fun₀ | 1 => 3 | 2 => 3 : Π₀ i, Fin (i + 10)) - = "fun₀ | 1 => 3 | 2 => 3" +#eval (fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 : Π₀ _ : List String, ℕ) + +end Repr +section PrettyPrinter /-- -info: +info: fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 : Π₀ (i : List String), ℕ -/ #guard_msgs in -#eval show Lean.MetaM Unit from - guard <| - reprStr ((fun₀ | 1 => 3 | 2 => 3) + (fun₀ | 1 => -3 | 2 => 4) : Π₀ _, ℤ) - = "fun₀ | 2 => 7" +#check (fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 : Π₀ _ : List String, ℕ) + +end PrettyPrinter diff --git a/test/finsupp_notation.lean b/test/finsupp_notation.lean index e042b67a8903f..08e871dde99cb 100644 --- a/test/finsupp_notation.lean +++ b/test/finsupp_notation.lean @@ -12,14 +12,48 @@ example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4) 2 = 3 := by example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4) 3 = 4 := by simp +section repr + +/-- info: fun₀ | 1 => 3 | 2 => 3 -/ +#guard_msgs in +#eval (Finsupp.mk {1, 2} (fun | 1 | 2 => 3 | _ => 0) (fun x => by aesop)) + /-- -info: +info: fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 -/ #guard_msgs in -#eval show Lean.MetaM Unit from - guard <| - reprStr (Finsupp.mk {1, 2} (fun | 1 | 2 => 3 | _ => 0) (fun x => by aesop)) - = "fun₀ | 1 => 3 | 2 => 3" +#eval Finsupp.mk + {["there are five words here", "and five more words here"], + ["there are seven words but only here"], + ["just two"]} + (fun + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 + | _ => 0) + (fun x => by aesop) + +end repr + +section PrettyPrinter + +/-- +info: fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 : List String →₀ ℕ +-/ +#guard_msgs in +#check fun₀ + | ["there are five words here", "and five more words here"] => 5 + | ["there are seven words but only here"] => 7 + | ["just two"] => 2 + +end PrettyPrinter + /-! ## (computable) number theory examples -/ From fd77f31972aa69d64ea2889b8d154417042e4aba Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Fri, 18 Oct 2024 09:52:34 +0000 Subject: [PATCH 056/131] chore(CategoryTheory): Make terminology for final functors consistent (#17895) Removes the remaining mentions of "cofinal functors". --- Mathlib/CategoryTheory/Limits/Final.lean | 32 ++++++++++----------- Mathlib/CategoryTheory/Limits/Presheaf.lean | 2 +- 2 files changed, 17 insertions(+), 17 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 2341721a15f55..965ee9e0e669e 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -36,7 +36,7 @@ This readily implies 2., as `comp_hasColimit`, `hasColimit_of_comp`, and `colimi From 2. we can specialize to `G = coyoneda.obj (op d)` to obtain 3., as `colimitCompCoyonedaIso`. -From 3., we prove 1. directly in `cofinal_of_colimit_comp_coyoneda_iso_pUnit`. +From 3., we prove 1. directly in `final_of_colimit_comp_coyoneda_iso_pUnit`. Dually, we prove that if a functor `F : C ⥤ D` is initial, then any functor `G : D ⥤ E` has a limit if and only if `F ⋙ G` does, and these limits are isomorphic via `limit.pre G F`. @@ -170,13 +170,13 @@ instance (d : D) : Nonempty (StructuredArrow d F) := variable {E : Type u₃} [Category.{v₃} E] (G : D ⥤ E) /-- -When `F : C ⥤ D` is cofinal, we denote by `lift F d` an arbitrary choice of object in `C` such that +When `F : C ⥤ D` is final, we denote by `lift F d` an arbitrary choice of object in `C` such that there exists a morphism `d ⟶ F.obj (lift F d)`. -/ def lift (d : D) : C := (Classical.arbitrary (StructuredArrow d F)).right -/-- When `F : C ⥤ D` is cofinal, we denote by `homToLift` an arbitrary choice of morphism +/-- When `F : C ⥤ D` is final, we denote by `homToLift` an arbitrary choice of morphism `d ⟶ F.obj (lift F d)`. -/ def homToLift (d : D) : d ⟶ F.obj (lift F d) := @@ -267,7 +267,7 @@ theorem colimit_cocone_comp_aux (s : Cocone (F ⋙ G)) (j : C) : variable (F G) -/-- If `F` is cofinal, +/-- If `F` is final, the category of cocones on `F ⋙ G` is equivalent to the category of cocones on `G`, for any `G : D ⥤ E`. -/ @@ -280,13 +280,13 @@ def coconesEquiv : Cocone (F ⋙ G) ≌ Cocone G where variable {G} -/-- When `F : C ⥤ D` is cofinal, and `t : Cocone G` for some `G : D ⥤ E`, +/-- When `F : C ⥤ D` is final, and `t : Cocone G` for some `G : D ⥤ E`, `t.whisker F` is a colimit cocone exactly when `t` is. -/ def isColimitWhiskerEquiv (t : Cocone G) : IsColimit (t.whisker F) ≃ IsColimit t := IsColimit.ofCoconeEquiv (coconesEquiv F G).symm -/-- When `F` is cofinal, and `t : Cocone (F ⋙ G)`, +/-- When `F` is final, and `t : Cocone (F ⋙ G)`, `extendCocone.obj t` is a colimit cocone exactly when `t` is. -/ def isColimitExtendCoconeEquiv (t : Cocone (F ⋙ G)) : @@ -312,7 +312,7 @@ section variable (G) -/-- When `F : C ⥤ D` is cofinal, and `G : D ⥤ E` has a colimit, then `F ⋙ G` has a colimit also and +/-- When `F : C ⥤ D` is final, and `G : D ⥤ E` has a colimit, then `F ⋙ G` has a colimit also and `colimit (F ⋙ G) ≅ colimit G` https://stacks.math.columbia.edu/tag/04E7 @@ -328,7 +328,7 @@ def colimitCoconeOfComp (t : ColimitCocone (F ⋙ G)) : ColimitCocone G where cocone := extendCocone.obj t.cocone isColimit := (isColimitExtendCoconeEquiv F _).symm t.isColimit -/-- When `F` is cofinal, and `F ⋙ G` has a colimit, then `G` has a colimit also. +/-- When `F` is final, and `F ⋙ G` has a colimit, then `G` has a colimit also. We can't make this an instance, because `F` is not determined by the goal. (Even if this weren't a problem, it would cause a loop with `comp_hasColimit`.) @@ -345,7 +345,7 @@ section -- Porting note: this instance does not seem to be found automatically --attribute [local instance] hasColimit_of_comp -/-- When `F` is cofinal, and `F ⋙ G` has a colimit, then `G` has a colimit also and +/-- When `F` is final, and `F ⋙ G` has a colimit, then `G` has a colimit also and `colimit (F ⋙ G) ≅ colimit G` https://stacks.math.columbia.edu/tag/04E7 @@ -390,9 +390,9 @@ theorem zigzag_of_eqvGen_quot_rel {F : C ⥤ D} {d : D} {f₁ f₂ : ΣX, d ⟶ end Final -/-- If `colimit (F ⋙ coyoneda.obj (op d)) ≅ PUnit` for all `d : D`, then `F` is cofinal. +/-- If `colimit (F ⋙ coyoneda.obj (op d)) ≅ PUnit` for all `d : D`, then `F` is final. -/ -theorem cofinal_of_colimit_comp_coyoneda_iso_pUnit +theorem final_of_colimit_comp_coyoneda_iso_pUnit (I : ∀ d, colimit (F ⋙ coyoneda.obj (op d)) ≅ PUnit) : Final F := ⟨fun d => by have : Nonempty (StructuredArrow d F) := by @@ -410,18 +410,18 @@ theorem cofinal_of_colimit_comp_coyoneda_iso_pUnit clear e y₁ y₂ exact Final.zigzag_of_eqvGen_quot_rel t⟩ -/-- A variant of `cofinal_of_colimit_comp_coyoneda_iso_pUnit` where we bind the various claims +/-- A variant of `final_of_colimit_comp_coyoneda_iso_pUnit` where we bind the various claims about `colimit (F ⋙ coyoneda.obj (Opposite.op d))` for each `d : D` into a single claim about the presheaf `colimit (F ⋙ yoneda)`. -/ -theorem cofinal_of_isTerminal_colimit_comp_yoneda +theorem final_of_isTerminal_colimit_comp_yoneda (h : IsTerminal (colimit (F ⋙ yoneda))) : Final F := by - refine cofinal_of_colimit_comp_coyoneda_iso_pUnit _ (fun d => ?_) + refine final_of_colimit_comp_coyoneda_iso_pUnit _ (fun d => ?_) refine Types.isTerminalEquivIsoPUnit _ ?_ let b := IsTerminal.isTerminalObj ((evaluation _ _).obj (Opposite.op d)) _ h exact b.ofIso <| preservesColimitIso ((evaluation _ _).obj (Opposite.op d)) (F ⋙ yoneda) /-- If the universal morphism `colimit (F ⋙ coyoneda.obj (op d)) ⟶ colimit (coyoneda.obj (op d))` -is an isomorphism (as it always is when `F` is cofinal), +is an isomorphism (as it always is when `F` is final), then `colimit (F ⋙ coyoneda.obj (op d)) ≅ PUnit` (simply because `colimit (coyoneda.obj (op d)) ≅ PUnit`). -/ @@ -437,7 +437,7 @@ variable {C : Type v} [Category.{v} C] {D : Type v} [Category.{v} D] (F : C ⥤ theorem final_iff_isIso_colimit_pre : Final F ↔ ∀ G : D ⥤ Type v, IsIso (colimit.pre G F) := ⟨fun _ => inferInstance, - fun _ => cofinal_of_colimit_comp_coyoneda_iso_pUnit _ fun _ => Final.colimitCompCoyonedaIso _ _⟩ + fun _ => final_of_colimit_comp_coyoneda_iso_pUnit _ fun _ => Final.colimitCompCoyonedaIso _ _⟩ end SmallCategory diff --git a/Mathlib/CategoryTheory/Limits/Presheaf.lean b/Mathlib/CategoryTheory/Limits/Presheaf.lean index 5dc588dc02b38..79b189b11e50c 100644 --- a/Mathlib/CategoryTheory/Limits/Presheaf.lean +++ b/Mathlib/CategoryTheory/Limits/Presheaf.lean @@ -559,7 +559,7 @@ variable {I : Type v₁} [SmallCategory I] (F : I ⥤ C) Proposition 2.6.3(ii) in [Kashiwara2006] -/ theorem final_toCostructuredArrow_comp_pre {c : Cocone (F ⋙ yoneda)} (hc : IsColimit c) : Functor.Final (c.toCostructuredArrow ⋙ CostructuredArrow.pre F yoneda c.pt) := by - apply Functor.cofinal_of_isTerminal_colimit_comp_yoneda + apply Functor.final_of_isTerminal_colimit_comp_yoneda suffices IsTerminal (colimit ((c.toCostructuredArrow ⋙ CostructuredArrow.pre F yoneda c.pt) ⋙ CostructuredArrow.toOver yoneda c.pt)) by From 23752ee05683cc476abb6f9eb98b63c0eb5d3da1 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 18 Oct 2024 10:20:47 +0000 Subject: [PATCH 057/131] feat: the algebraic closure of an algebraic k-algebra is an algebraic closure of k (#17894) Co-authored-by: FR --- Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index ea7fe25e09f58..d58f8dcbc128b 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -403,4 +403,9 @@ instance [CharZero k] : CharZero (AlgebraicClosure k) := instance {p : ℕ} [CharP k p] : CharP (AlgebraicClosure k) p := charP_of_injective_algebraMap (RingHom.injective (algebraMap k (AlgebraicClosure k))) p +instance {L : Type*} [Field k] [Field L] [Algebra k L] [Algebra.IsAlgebraic k L] : + IsAlgClosure k (AlgebraicClosure L) where + isAlgebraic := .trans (L := L) + isAlgClosed := inferInstance + end AlgebraicClosure From 1cc3f60d578ae8b32e7e308018e9f409a5dbaedc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 18 Oct 2024 10:35:29 +0000 Subject: [PATCH 058/131] chore(Order/RelIso/Set): golf `wellFounded_iff_wellFounded_subrel` (#17606) We also move it out of the initial segment file - it doesn't need that machinery to be proven. --- Mathlib/Order/InitialSeg.lean | 14 -------------- Mathlib/Order/RelIso/Set.lean | 11 +++++++++++ 2 files changed, 11 insertions(+), 14 deletions(-) diff --git a/Mathlib/Order/InitialSeg.lean b/Mathlib/Order/InitialSeg.lean index 513501c25072e..e53d78efa0277 100644 --- a/Mathlib/Order/InitialSeg.lean +++ b/Mathlib/Order/InitialSeg.lean @@ -461,20 +461,6 @@ protected theorem acc [IsTrans β s] (f : r ≺i s) (a : α) : Acc r a ↔ Acc s end PrincipalSeg -/-- A relation is well-founded iff every principal segment of it is well-founded. - -In this lemma we use `Subrel` to indicate its principal segments because it's usually more -convenient to use. --/ -theorem wellFounded_iff_wellFounded_subrel {β : Type*} {s : β → β → Prop} [IsTrans β s] : - WellFounded s ↔ ∀ b, WellFounded (Subrel s { b' | s b' b }) := by - refine - ⟨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.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] : WellFounded s ↔ ∀ (α : Type u) (r : α → α → Prop) (_ : r ≺i s), WellFounded r := ⟨fun wf _ _ f => RelHomClass.wellFounded f.toRelEmbedding wf, fun h => diff --git a/Mathlib/Order/RelIso/Set.lean b/Mathlib/Order/RelIso/Set.lean index 961b9773f5912..ec7b62b18b16f 100644 --- a/Mathlib/Order/RelIso/Set.lean +++ b/Mathlib/Order/RelIso/Set.lean @@ -98,3 +98,14 @@ theorem RelIso.preimage_eq_image_symm (e : r ≃r s) (t : Set β) : e ⁻¹' t = rw [e.symm.image_eq_preimage_symm]; rfl end image + +theorem Acc.of_subrel {r : α → α → Prop} [IsTrans α r] {b : α} (a : { a // r a b }) + (h : Acc (Subrel r { a | r a b }) a) : Acc r a.1 := + h.recOn fun a _ IH ↦ ⟨_, fun _ hb ↦ IH ⟨_, _root_.trans hb a.2⟩ hb⟩ + +/-- A relation `r` is well-founded iff every downward-interval `{ a | r a b }` of it is +well-founded. -/ +theorem wellFounded_iff_wellFounded_subrel {r : α → α → Prop} [IsTrans α r] : + WellFounded r ↔ ∀ b, WellFounded (Subrel r { a | r a b }) where + mp h _ := InvImage.wf Subtype.val h + mpr h := ⟨fun a ↦ ⟨_, fun b hr ↦ ((h a).apply _).of_subrel ⟨b, hr⟩⟩⟩ From cbaeaf5ebf33c9f777f35f6cf3a44029e7d35c0e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 18 Oct 2024 10:35:30 +0000 Subject: [PATCH 059/131] chore(Topology): deprecate `Homeomorph.toContinuousMap` (#17714) Use `_root_.toContinuousMap` instead. --- Mathlib/MeasureTheory/Measure/Content.lean | 6 +++--- Mathlib/Topology/Category/TopCat/Basic.lean | 4 ++-- Mathlib/Topology/ContinuousMap/Algebra.lean | 4 ++-- Mathlib/Topology/ContinuousMap/Basic.lean | 16 ++++++---------- Mathlib/Topology/ContinuousMap/Compact.lean | 4 ++-- .../ContinuousMap/ContinuousMapZero.lean | 8 ++++---- Mathlib/Topology/ContinuousMap/Defs.lean | 4 ++++ .../Topology/ContinuousMap/Polynomial.lean | 2 +- .../Topology/ContinuousMap/Weierstrass.lean | 2 +- Mathlib/Topology/Homotopy/HomotopyGroup.lean | 19 +++++++++---------- Mathlib/Topology/Sets/Opens.lean | 8 ++++---- 11 files changed, 38 insertions(+), 39 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Content.lean b/Mathlib/MeasureTheory/Measure/Content.lean index b07a09b9ad57e..77169173b08ad 100644 --- a/Mathlib/MeasureTheory/Measure/Content.lean +++ b/Mathlib/MeasureTheory/Measure/Content.lean @@ -190,7 +190,7 @@ theorem innerContent_iUnion_nat [R1Space G] ⦃U : ℕ → Set G⦄ rwa [Opens.iSup_def] at this theorem innerContent_comap (f : G ≃ₜ G) (h : ∀ ⦃K : Compacts G⦄, μ (K.map f f.continuous) = μ K) - (U : Opens G) : μ.innerContent (Opens.comap f.toContinuousMap U) = μ.innerContent U := by + (U : Opens G) : μ.innerContent (Opens.comap f U) = μ.innerContent U := by refine (Compacts.equiv f).surjective.iSup_congr _ fun K => iSup_congr_Prop image_subset_iff ?_ intro hK simp only [Equiv.coe_fn_mk, Subtype.mk_eq_mk, Compacts.equiv] @@ -200,7 +200,7 @@ theorem innerContent_comap (f : G ≃ₜ G) (h : ∀ ⦃K : Compacts G⦄, μ (K theorem is_mul_left_invariant_innerContent [Group G] [TopologicalGroup G] (h : ∀ (g : G) {K : Compacts G}, μ (K.map _ <| continuous_mul_left g) = μ K) (g : G) (U : Opens G) : - μ.innerContent (Opens.comap (Homeomorph.mulLeft g).toContinuousMap U) = μ.innerContent U := by + μ.innerContent (Opens.comap (Homeomorph.mulLeft g) U) = μ.innerContent U := by convert μ.innerContent_comap (Homeomorph.mulLeft g) (fun K => h g) U @[to_additive] @@ -211,7 +211,7 @@ theorem innerContent_pos_of_is_mul_left_invariant [Group G] [TopologicalGroup G] rcases compact_covered_by_mul_left_translates K.2 this with ⟨s, hs⟩ suffices μ K ≤ s.card * μ.innerContent U by exact (ENNReal.mul_pos_iff.mp <| hK.bot_lt.trans_le this).2 - have : (K : Set G) ⊆ ↑(⨆ g ∈ s, Opens.comap (Homeomorph.mulLeft g).toContinuousMap U) := by + have : (K : Set G) ⊆ ↑(⨆ g ∈ s, Opens.comap (Homeomorph.mulLeft g : C(G, G)) U) := by simpa only [Opens.iSup_def, Opens.coe_comap, Subtype.coe_mk] refine (μ.le_innerContent _ _ this).trans ?_ refine diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 8cf13a2f355f4..f1a2371e9013c 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -126,8 +126,8 @@ def trivial : Type u ⥤ TopCat.{u} where @[simps] def isoOfHomeo {X Y : TopCat.{u}} (f : X ≃ₜ Y) : X ≅ Y where -- Porting note: previously ⟨f⟩ for hom (inv) and tidy closed proofs - hom := f.toContinuousMap - inv := f.symm.toContinuousMap + hom := (f : C(X, Y)) + inv := (f.symm : C(Y, X)) hom_inv_id := by ext; exact f.symm_apply_apply _ inv_hom_id := by ext; exact f.apply_symm_apply _ diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index b6ff6659b4878..59525a680230d 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -1010,7 +1010,7 @@ variable [ContinuousStar A] [Algebra 𝕜 A] actually a homeomorphism. -/ @[simps] def compStarAlgEquiv' (f : X ≃ₜ Y) : C(Y, A) ≃⋆ₐ[𝕜] C(X, A) := - { f.toContinuousMap.compStarAlgHom' 𝕜 A with + { (f : C(X, Y)).compStarAlgHom' 𝕜 A with toFun := (f : C(X, Y)).compStarAlgHom' 𝕜 A invFun := (f.symm : C(Y, X)).compStarAlgHom' 𝕜 A left_inv := fun g => by @@ -1019,7 +1019,7 @@ def compStarAlgEquiv' (f : X ≃ₜ Y) : C(Y, A) ≃⋆ₐ[𝕜] C(X, A) := right_inv := fun g => by simp only [ContinuousMap.compStarAlgHom'_apply, ContinuousMap.comp_assoc, symm_comp_toContinuousMap, ContinuousMap.comp_id] - map_smul' := fun k a => map_smul (f.toContinuousMap.compStarAlgHom' 𝕜 A) k a } + map_smul' := fun k a => map_smul ((f : C(X, Y)).compStarAlgHom' 𝕜 A) k a } end Homeomorph diff --git a/Mathlib/Topology/ContinuousMap/Basic.lean b/Mathlib/Topology/ContinuousMap/Basic.lean index bd8aca0de948b..69e4f1e0ee255 100644 --- a/Mathlib/Topology/ContinuousMap/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Basic.lean @@ -429,19 +429,15 @@ namespace Homeomorph variable {α β γ : Type*} [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] variable (f : α ≃ₜ β) (g : β ≃ₜ γ) +instance instContinuousMapClass : ContinuousMapClass (α ≃ₜ β) α β where + map_continuous f := f.continuous_toFun + /-- The forward direction of a homeomorphism, as a bundled continuous map. -/ -@[simps] -def toContinuousMap (e : α ≃ₜ β) : C(α, β) := +@[simps, deprecated _root_.toContinuousMap (since := "2024-10-12")] +protected def toContinuousMap (e : α ≃ₜ β) : C(α, β) := ⟨e, e.continuous_toFun⟩ -/-- `Homeomorph.toContinuousMap` as a coercion. -/ -instance : Coe (α ≃ₜ β) C(α, β) := - ⟨Homeomorph.toContinuousMap⟩ - --- Porting note: Syntactic tautology -/- theorem toContinuousMap_as_coe : f.toContinuousMap = f := - rfl --/ +attribute [deprecated ContinuousMap.coe_apply (since := "2024-10-12")] toContinuousMap_apply @[simp] theorem coe_refl : (Homeomorph.refl α : C(α, α)) = ContinuousMap.id α := diff --git a/Mathlib/Topology/ContinuousMap/Compact.lean b/Mathlib/Topology/ContinuousMap/Compact.lean index 552c3e592bff1..e5b573631c838 100644 --- a/Mathlib/Topology/ContinuousMap/Compact.lean +++ b/Mathlib/Topology/ContinuousMap/Compact.lean @@ -455,8 +455,8 @@ theorem compRightContinuousMap_apply {X Y : Type*} (T : Type*) [TopologicalSpace def compRightHomeomorph {X Y : Type*} (T : Type*) [TopologicalSpace X] [CompactSpace X] [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 + toFun := compRightContinuousMap T f + invFun := compRightContinuousMap T f.symm left_inv g := ext fun _ => congr_arg g (f.apply_symm_apply _) right_inv g := ext fun _ => congr_arg g (f.symm_apply_apply _) diff --git a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean index 575b4e26ad07c..6d699f5ba3042 100644 --- a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean @@ -303,15 +303,15 @@ alias uniformEmbedding_comp := isUniformEmbedding_comp sending `0 : X` to `0 : Y`. -/ def _root_.UniformEquiv.arrowCongrLeft₀ {Y : Type*} [TopologicalSpace Y] [Zero Y] (f : X ≃ₜ Y) (hf : f 0 = 0) : C(X, R)₀ ≃ᵤ C(Y, R)₀ where - toFun g := g.comp ⟨f.symm.toContinuousMap, (f.toEquiv.apply_eq_iff_eq_symm_apply.eq ▸ hf).symm⟩ - invFun g := g.comp ⟨f.toContinuousMap, hf⟩ + toFun g := g.comp ⟨f.symm, (f.toEquiv.apply_eq_iff_eq_symm_apply.eq ▸ hf).symm⟩ + invFun g := g.comp ⟨f, hf⟩ left_inv g := ext fun _ ↦ congrArg g <| f.left_inv _ right_inv g := ext fun _ ↦ congrArg g <| f.right_inv _ uniformContinuous_toFun := isUniformEmbedding_toContinuousMap.uniformContinuous_iff.mpr <| - ContinuousMap.uniformContinuous_comp_left f.symm.toContinuousMap |>.comp + ContinuousMap.uniformContinuous_comp_left (f.symm : C(Y, X)) |>.comp isUniformEmbedding_toContinuousMap.uniformContinuous uniformContinuous_invFun := isUniformEmbedding_toContinuousMap.uniformContinuous_iff.mpr <| - ContinuousMap.uniformContinuous_comp_left f.toContinuousMap |>.comp + ContinuousMap.uniformContinuous_comp_left (f : C(X, Y)) |>.comp isUniformEmbedding_toContinuousMap.uniformContinuous end UniformSpace diff --git a/Mathlib/Topology/ContinuousMap/Defs.lean b/Mathlib/Topology/ContinuousMap/Defs.lean index f576214a27758..e1ee866070f68 100644 --- a/Mathlib/Topology/ContinuousMap/Defs.lean +++ b/Mathlib/Topology/ContinuousMap/Defs.lean @@ -93,6 +93,10 @@ protected theorem coe_coe {F : Type*} [FunLike F X Y] [ContinuousMapClass F X Y] ⇑(f : C(X, Y)) = f := rfl +protected theorem coe_apply {F : Type*} [FunLike F X Y] [ContinuousMapClass F X Y] (f : F) (x : X) : + (f : C(X, Y)) x = f x := + rfl + @[ext] theorem ext {f g : C(X, Y)} (h : ∀ a, f a = g a) : f = g := DFunLike.ext _ _ h diff --git a/Mathlib/Topology/ContinuousMap/Polynomial.lean b/Mathlib/Topology/ContinuousMap/Polynomial.lean index 83a3dd5e6a0ed..237d38cd410f8 100644 --- a/Mathlib/Topology/ContinuousMap/Polynomial.lean +++ b/Mathlib/Topology/ContinuousMap/Polynomial.lean @@ -164,7 +164,7 @@ open ContinuousMap /-- The preimage of polynomials on `[0,1]` under the pullback map by `x ↦ (b-a) * x + a` is the polynomials on `[a,b]`. -/ theorem polynomialFunctions.comap_compRightAlgHom_iccHomeoI (a b : ℝ) (h : a < b) : - (polynomialFunctions I).comap (compRightAlgHom ℝ ℝ (iccHomeoI a b h).symm.toContinuousMap) = + (polynomialFunctions I).comap (compRightAlgHom ℝ ℝ (iccHomeoI a b h).symm) = polynomialFunctions (Set.Icc a b) := by ext f fconstructor diff --git a/Mathlib/Topology/ContinuousMap/Weierstrass.lean b/Mathlib/Topology/ContinuousMap/Weierstrass.lean index fbfdd66fb2023..8e903e39101e1 100644 --- a/Mathlib/Topology/ContinuousMap/Weierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/Weierstrass.lean @@ -55,7 +55,7 @@ theorem polynomialFunctions_closure_eq_top (a b : ℝ) : · -- We can pullback continuous functions on `[a,b]` to continuous functions on `[0,1]`, -- by precomposing with an affine map. let W : C(Set.Icc a b, ℝ) →ₐ[ℝ] C(I, ℝ) := - compRightAlgHom ℝ ℝ (iccHomeoI a b h).symm.toContinuousMap + compRightAlgHom ℝ ℝ (iccHomeoI a b h).symm -- This operation is itself a homeomorphism -- (with respect to the norm topologies on continuous functions). let W' : C(Set.Icc a b, ℝ) ≃ₜ C(I, ℝ) := compRightHomeomorph ℝ (iccHomeoI a b h).symm diff --git a/Mathlib/Topology/Homotopy/HomotopyGroup.lean b/Mathlib/Topology/Homotopy/HomotopyGroup.lean index a9ec3d6dc3be4..21ad316e36339 100644 --- a/Mathlib/Topology/Homotopy/HomotopyGroup.lean +++ b/Mathlib/Topology/Homotopy/HomotopyGroup.lean @@ -180,7 +180,7 @@ variable [DecidableEq N] @[simps] def toLoop (i : N) (p : Ω^ N X x) : Ω (Ω^ { j // j ≠ i } X x) const where toFun t := - ⟨(p.val.comp (Cube.insertAt i).toContinuousMap).curry t, fun y yH => + ⟨(p.val.comp (Cube.insertAt i)).curry t, fun y yH => p.property (Cube.insertAt i (t, y)) (Cube.insertAt_boundary i <| Or.inr yH)⟩ source' := by ext t; refine p.property (Cube.insertAt i (0, t)) ⟨i, Or.inl ?_⟩; simp target' := by ext t; refine p.property (Cube.insertAt i (1, t)) ⟨i, Or.inr ?_⟩; simp @@ -200,10 +200,10 @@ theorem continuous_toLoop (i : N) : Continuous (@toLoop N X _ x _ i) := @[simps] def fromLoop (i : N) (p : Ω (Ω^ { j // j ≠ i } X x) const) : Ω^ N X x := ⟨(ContinuousMap.comp ⟨Subtype.val, by fun_prop⟩ p.toContinuousMap).uncurry.comp - (Cube.splitAt i).toContinuousMap, + (Cube.splitAt i), by rintro y ⟨j, Hj⟩ - simp only [ContinuousMap.comp_apply, toContinuousMap_apply, + simp only [ContinuousMap.comp_apply, ContinuousMap.coe_coe, funSplitAt_apply, ContinuousMap.uncurry_apply, ContinuousMap.coe_mk, Function.uncurry_apply_pair] obtain rfl | Hne := eq_or_ne j i @@ -243,8 +243,8 @@ theorem fromLoop_apply (i : N) {p : Ω (Ω^ { j // j ≠ i } X x) const} {t : I^ /-- Composition with `Cube.insertAt` as a continuous map. -/ abbrev cCompInsert (i : N) : C(C(I^N, X), C(I × I^{ j // j ≠ i }, X)) := - ⟨fun f => f.comp (Cube.insertAt i).toContinuousMap, - (Cube.insertAt i).toContinuousMap.continuous_comp_left⟩ + ⟨fun f => f.comp (Cube.insertAt i), + (toContinuousMap <| Cube.insertAt i).continuous_comp_left⟩ /-- A homotopy between `n+1`-dimensional loops `p` and `q` constant on the boundary seen as a homotopy between two paths in the space of `n`-dimensional paths. -/ @@ -280,7 +280,7 @@ theorem homotopicTo (i : N) {p q : Ω^ N X x} : C(I × I^N, X) := (ContinuousMap.comp ⟨_, ContinuousMap.continuous_uncurry⟩ (ContinuousMap.comp ⟨Subtype.val, by continuity⟩ H.toContinuousMap).curry).uncurry.comp <| - (ContinuousMap.id I).prodMap (Cube.splitAt i).toContinuousMap + (ContinuousMap.id I).prodMap (Cube.splitAt i) theorem homotopicFrom (i : N) {p q : Ω^ N X x} : (toLoop i p).Homotopic (toLoop i q) → Homotopic p q := by @@ -289,15 +289,14 @@ theorem homotopicFrom (i : N) {p q : Ω^ N X x} : · rintro t y ⟨j, jH⟩ erw [homotopyFrom_apply] obtain rfl | h := eq_or_ne j i - · simp only [Prod.map_apply, id_eq, toContinuousMap_apply, funSplitAt_apply, - Function.uncurry_apply_pair] + · simp only [Prod.map_apply, id_eq, funSplitAt_apply, Function.uncurry_apply_pair] rw [H.eq_fst] exacts [congr_arg p ((Cube.splitAt j).left_inv _), jH] · rw [p.2 _ ⟨j, jH⟩]; apply boundary; exact ⟨⟨j, h⟩, jH⟩ all_goals intro apply (homotopyFrom_apply _ _ _).trans - simp only [Prod.map_apply, id_eq, toContinuousMap_apply, funSplitAt_apply, + simp only [Prod.map_apply, id_eq, funSplitAt_apply, Function.uncurry_apply_pair, ContinuousMap.HomotopyWith.apply_zero, ContinuousMap.HomotopyWith.apply_one, ne_eq, Path.coe_toContinuousMap, toLoop_apply_coe, ContinuousMap.curry_apply, ContinuousMap.comp_apply] @@ -315,7 +314,7 @@ def transAt (i : N) (f g : Ω^ N X x) : Ω^ N X x := (by ext1; symm dsimp only [Path.trans, fromLoop, Path.coe_mk_mk, Function.comp_apply, mk_apply, - ContinuousMap.comp_apply, toContinuousMap_apply, funSplitAt_apply, + ContinuousMap.comp_apply, ContinuousMap.coe_coe, funSplitAt_apply, ContinuousMap.uncurry_apply, ContinuousMap.coe_mk, Function.uncurry_apply_pair] split_ifs · show f _ = _; congr 1 diff --git a/Mathlib/Topology/Sets/Opens.lean b/Mathlib/Topology/Sets/Opens.lean index 1b2ffd4204a29..bd142710c0990 100644 --- a/Mathlib/Topology/Sets/Opens.lean +++ b/Mathlib/Topology/Sets/Opens.lean @@ -379,10 +379,10 @@ theorem comap_injective [T0Space β] : Injective (comap : C(α, β) → FrameHom /-- A homeomorphism induces an order-preserving equivalence on open sets, by taking comaps. -/ @[simps (config := .asFn) apply] def _root_.Homeomorph.opensCongr (f : α ≃ₜ β) : Opens α ≃o Opens β where - toFun := Opens.comap f.symm.toContinuousMap - invFun := Opens.comap f.toContinuousMap - left_inv := fun _ => ext <| f.toEquiv.preimage_symm_preimage _ - right_inv := fun _ => ext <| f.toEquiv.symm_preimage_preimage _ + toFun := Opens.comap (f.symm : C(β, α)) + invFun := Opens.comap (f : C(α, β)) + left_inv _ := ext <| f.toEquiv.preimage_symm_preimage _ + right_inv _ := ext <| f.toEquiv.symm_preimage_preimage _ map_rel_iff' := by simp only [← SetLike.coe_subset_coe]; exact f.symm.surjective.preimage_subset_preimage_iff From a6bc4c5ee86f2ccb63f6d63406cd3db02d3de29a Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 18 Oct 2024 11:08:02 +0000 Subject: [PATCH 060/131] chore: reorganize the `ContinuousOn` file (#17901) Before adding some API for `ContinuousWithinAt` to have it match the one for `DifferentiableWithinAt`, and `ContDiffWithinAt`, and manifold versions, I'd rather start from a clean file. The file was a mess of random things added at random places, I tried to make it coherent. Nothing has been added or removed, no name has been changed, pure shuffling around. --- Mathlib/Topology/Constructions.lean | 7 + Mathlib/Topology/ContinuousOn.lean | 732 +++++++++++++++------------- 2 files changed, 398 insertions(+), 341 deletions(-) diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index e0bb583947c37..82bce23e4c0c9 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -1134,6 +1134,13 @@ lemma isClosed_preimage_val {s t : Set X} : IsClosed (s ↓∩ t) ↔ s ∩ clos Subtype.image_preimage_coe, subset_antisymm_iff, and_iff_left, Set.subset_inter_iff, and_iff_right] exacts [Set.inter_subset_left, Set.subset_inter Set.inter_subset_left subset_closure] + +theorem frontier_inter_open_inter {s t : Set X} (ht : IsOpen t) : + frontier (s ∩ t) ∩ t = frontier s ∩ t := by + simp only [Set.inter_comm _ t, ← Subtype.preimage_coe_eq_preimage_coe_iff, + ht.isOpenMap_subtype_val.preimage_frontier_eq_frontier_preimage continuous_subtype_val, + Subtype.preimage_coe_self_inter] + end Subtype section Quotient diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index a09c8362c391f..d69f7aec79123 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -9,13 +9,14 @@ import Mathlib.Topology.Constructions /-! # Neighborhoods and continuity relative to a subset -This file defines relative versions +This file develops API on the relative versions * `nhdsWithin` of `nhds` * `ContinuousOn` of `Continuous` * `ContinuousWithinAt` of `ContinuousAt` -and proves their basic properties, including the relationships between +related to continuity, which are defined in previous definition files. +Their basic properties studied in this file include the relationships between these restricted notions and the corresponding notions for the subtype equipped with the subspace topology. @@ -32,6 +33,10 @@ open Set Filter Function Topology Filter variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} variable [TopologicalSpace α] +/-! +## Properties of the neighborhood-within filter +-/ + @[simp] theorem nhds_bind_nhdsWithin {a : α} {s : Set α} : ((𝓝 a).bind fun x => 𝓝[s] x) = 𝓝[s] a := bind_inf_principal.trans <| congr_arg₂ _ nhds_bind_nhds rfl @@ -299,6 +304,8 @@ instance Pi.instNeBotNhdsWithinIoi [Nonempty ι] [∀ i, Preorder (π i)] {x : [∀ i, (𝓝[>] x i).NeBot] : (𝓝[>] x).NeBot := Pi.instNeBotNhdsWithinIio (π := fun i ↦ (π i)ᵒᵈ) (x := fun i ↦ OrderDual.toDual (x i)) +end Pi + theorem Filter.Tendsto.piecewise_nhdsWithin {f g : α → β} {t : Set α} [∀ x, Decidable (x ∈ t)] {a : α} {s : Set α} {l : Filter β} (h₀ : Tendsto f (𝓝[s ∩ t] a) l) (h₁ : Tendsto g (𝓝[s ∩ tᶜ] a) l) : Tendsto (piecewise t f g) (𝓝[s] a) l := by @@ -438,8 +445,16 @@ theorem tendsto_nhdsWithin_iff_subtype {s : Set α} {a : α} (h : a ∈ s) (f : Tendsto f (𝓝[s] a) l ↔ Tendsto (s.restrict f) (𝓝 ⟨a, h⟩) l := by rw [nhdsWithin_eq_map_subtype_coe h, tendsto_map'_iff]; rfl +/-! +## Local continuity properties of functions +-/ + variable [TopologicalSpace β] [TopologicalSpace γ] [TopologicalSpace δ] +/-! +### `ContinuousWithinAt` +-/ + /-- If a function is continuous within `s` at `x`, then it tends to `f x` within `s` by definition. We register this fact for use with the dot notation, especially to use `Filter.Tendsto.comp` as `ContinuousWithinAt.comp` will have a different meaning. -/ @@ -447,10 +462,6 @@ theorem ContinuousWithinAt.tendsto {f : α → β} {s : Set α} {x : α} (h : Co Tendsto f (𝓝[s] x) (𝓝 (f x)) := h -theorem ContinuousOn.continuousWithinAt {f : α → β} {s : Set α} {x : α} (hf : ContinuousOn f s) - (hx : x ∈ s) : ContinuousWithinAt f s x := - hf x hx - theorem continuousWithinAt_univ (f : α → β) (x : α) : ContinuousWithinAt f Set.univ x ↔ ContinuousAt f x := by rw [ContinuousAt, ContinuousWithinAt, nhdsWithin_univ] @@ -471,93 +482,44 @@ theorem ContinuousWithinAt.tendsto_nhdsWithin_image {f : α → β} {x : α} {s (h : ContinuousWithinAt f s x) : Tendsto f (𝓝[s] x) (𝓝[f '' s] f x) := h.tendsto_nhdsWithin (mapsTo_image _ _) -theorem ContinuousWithinAt.prod_map {f : α → γ} {g : β → δ} {s : Set α} {t : Set β} {x : α} {y : β} - (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g t y) : - ContinuousWithinAt (Prod.map f g) (s ×ˢ t) (x, y) := by - unfold ContinuousWithinAt at * - rw [nhdsWithin_prod_eq, Prod.map, nhds_prod_eq] - exact hf.prod_map hg - -theorem continuousWithinAt_prod_of_discrete_left [DiscreteTopology α] - {f : α × β → γ} {s : Set (α × β)} {x : α × β} : - ContinuousWithinAt f s x ↔ ContinuousWithinAt (f ⟨x.1, ·⟩) {b | (x.1, b) ∈ s} x.2 := by - rw [← x.eta]; simp_rw [ContinuousWithinAt, nhdsWithin, nhds_prod_eq, nhds_discrete, pure_prod, - ← map_inf_principal_preimage]; rfl - -theorem continuousWithinAt_prod_of_discrete_right [DiscreteTopology β] - {f : α × β → γ} {s : Set (α × β)} {x : α × β} : - ContinuousWithinAt f s x ↔ ContinuousWithinAt (f ⟨·, x.2⟩) {a | (a, x.2) ∈ s} x.1 := by - rw [← x.eta]; simp_rw [ContinuousWithinAt, nhdsWithin, nhds_prod_eq, nhds_discrete, prod_pure, - ← map_inf_principal_preimage]; rfl - -theorem continuousAt_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} {x : α × β} : - ContinuousAt f x ↔ ContinuousAt (f ⟨x.1, ·⟩) x.2 := by - simp_rw [← continuousWithinAt_univ]; exact continuousWithinAt_prod_of_discrete_left - -theorem continuousAt_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} {x : α × β} : - ContinuousAt f x ↔ ContinuousAt (f ⟨·, x.2⟩) x.1 := by - simp_rw [← continuousWithinAt_univ]; exact continuousWithinAt_prod_of_discrete_right - -theorem continuousOn_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} {s : Set (α × β)} : - ContinuousOn f s ↔ ∀ a, ContinuousOn (f ⟨a, ·⟩) {b | (a, b) ∈ s} := by - simp_rw [ContinuousOn, Prod.forall, continuousWithinAt_prod_of_discrete_left]; rfl - -theorem continuousOn_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} {s : Set (α × β)} : - ContinuousOn f s ↔ ∀ b, ContinuousOn (f ⟨·, b⟩) {a | (a, b) ∈ s} := by - simp_rw [ContinuousOn, Prod.forall, continuousWithinAt_prod_of_discrete_right]; apply forall_swap - -/-- If a function `f a b` is such that `y ↦ f a b` is continuous for all `a`, and `a` lives in a -discrete space, then `f` is continuous, and vice versa. -/ -theorem continuous_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} : - Continuous f ↔ ∀ a, Continuous (f ⟨a, ·⟩) := by - simp_rw [continuous_iff_continuousOn_univ]; exact continuousOn_prod_of_discrete_left - -theorem continuous_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} : - Continuous f ↔ ∀ b, Continuous (f ⟨·, b⟩) := by - simp_rw [continuous_iff_continuousOn_univ]; exact continuousOn_prod_of_discrete_right - -theorem isOpenMap_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} : - IsOpenMap f ↔ ∀ a, IsOpenMap (f ⟨a, ·⟩) := by - simp_rw [isOpenMap_iff_nhds_le, Prod.forall, nhds_prod_eq, nhds_discrete, pure_prod, map_map] - rfl - -theorem isOpenMap_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} : - IsOpenMap f ↔ ∀ b, IsOpenMap (f ⟨·, b⟩) := by - simp_rw [isOpenMap_iff_nhds_le, Prod.forall, forall_swap (α := α) (β := β), nhds_prod_eq, - nhds_discrete, prod_pure, map_map]; rfl +theorem nhdsWithin_le_comap {x : α} {s : Set α} {f : α → β} (ctsf : ContinuousWithinAt f s x) : + 𝓝[s] x ≤ comap f (𝓝[f '' s] f x) := + ctsf.tendsto_nhdsWithin_image.le_comap -theorem continuousWithinAt_pi {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] - {f : α → ∀ i, π i} {s : Set α} {x : α} : - ContinuousWithinAt f s x ↔ ∀ i, ContinuousWithinAt (fun y => f y i) s x := - tendsto_pi_nhds +theorem ContinuousWithinAt.preimage_mem_nhdsWithin {f : α → β} {x : α} {s : Set α} {t : Set β} + (h : ContinuousWithinAt f s x) (ht : t ∈ 𝓝 (f x)) : f ⁻¹' t ∈ 𝓝[s] x := + h ht -theorem continuousOn_pi {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] - {f : α → ∀ i, π i} {s : Set α} : ContinuousOn f s ↔ ∀ i, ContinuousOn (fun y => f y i) s := - ⟨fun h i x hx => tendsto_pi_nhds.1 (h x hx) i, fun h x hx => tendsto_pi_nhds.2 fun i => h i x hx⟩ +theorem ContinuousWithinAt.preimage_mem_nhdsWithin' {f : α → β} {x : α} {s : Set α} {t : Set β} + (h : ContinuousWithinAt f s x) (ht : t ∈ 𝓝[f '' s] f x) : f ⁻¹' t ∈ 𝓝[s] x := + h.tendsto_nhdsWithin (mapsTo_image _ _) ht -@[fun_prop] -theorem continuousOn_pi' {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] - {f : α → ∀ i, π i} {s : Set α} (hf : ∀ i, ContinuousOn (fun y => f y i) s) : - ContinuousOn f s := - continuousOn_pi.2 hf +theorem ContinuousWithinAt.preimage_mem_nhdsWithin'' + {f : α → β} {x : α} {y : β} {s t : Set β} + (h : ContinuousWithinAt f (f ⁻¹' s) x) (ht : t ∈ 𝓝[s] y) (hxy : y = f x) : + f ⁻¹' t ∈ 𝓝[f ⁻¹' s] x := by + rw [hxy] at ht + exact h.preimage_mem_nhdsWithin' (nhdsWithin_mono _ (image_preimage_subset f s) ht) -theorem ContinuousWithinAt.fin_insertNth {n} {π : Fin (n + 1) → Type*} - [∀ i, TopologicalSpace (π i)] (i : Fin (n + 1)) {f : α → π i} {a : α} {s : Set α} - (hf : ContinuousWithinAt f s a) {g : α → ∀ j : Fin n, π (i.succAbove j)} - (hg : ContinuousWithinAt g s a) : ContinuousWithinAt (fun a => i.insertNth (f a) (g a)) s a := - hf.tendsto.fin_insertNth i hg +theorem continuousWithinAt_of_not_mem_closure {f : α → β} {s : Set α} {x : α} (hx : x ∉ closure s) : + ContinuousWithinAt f s x := by + rw [mem_closure_iff_nhdsWithin_neBot, not_neBot] at hx + rw [ContinuousWithinAt, hx] + exact tendsto_bot -nonrec theorem ContinuousOn.fin_insertNth {n} {π : Fin (n + 1) → Type*} - [∀ i, TopologicalSpace (π i)] (i : Fin (n + 1)) {f : α → π i} {s : Set α} - (hf : ContinuousOn f s) {g : α → ∀ j : Fin n, π (i.succAbove j)} (hg : ContinuousOn g s) : - ContinuousOn (fun a => i.insertNth (f a) (g a)) s := fun a ha => - (hf a ha).fin_insertNth i (hg a ha) +/-! +### `ContinuousOn` +-/ theorem continuousOn_iff {f : α → β} {s : Set α} : ContinuousOn f s ↔ ∀ x ∈ s, ∀ t : Set β, IsOpen t → f x ∈ t → ∃ u, IsOpen u ∧ x ∈ u ∧ u ∩ s ⊆ f ⁻¹' t := by simp only [ContinuousOn, ContinuousWithinAt, tendsto_nhds, mem_nhdsWithin] +theorem ContinuousOn.continuousWithinAt {f : α → β} {s : Set α} {x : α} (hf : ContinuousOn f s) + (hx : x ∈ s) : ContinuousWithinAt f s x := + hf x hx + theorem continuousOn_iff_continuous_restrict {f : α → β} {s : Set α} : ContinuousOn f s ↔ Continuous (s.restrict f) := by rw [ContinuousOn, continuous_iff_continuousAt]; constructor @@ -606,10 +568,6 @@ theorem continuousOn_iff_isClosed {f : α → β} {s : Set α} : simp only [Subtype.preimage_coe_eq_preimage_coe_iff, eq_comm, Set.inter_comm s] rw [continuousOn_iff_continuous_restrict, continuous_iff_isClosed]; simp only [this] -theorem ContinuousOn.prod_map {f : α → γ} {g : β → δ} {s : Set α} {t : Set β} - (hf : ContinuousOn f s) (hg : ContinuousOn g t) : ContinuousOn (Prod.map f g) (s ×ˢ t) := - fun ⟨x, y⟩ ⟨hx, hy⟩ => ContinuousWithinAt.prod_map (hf x hx) (hg y hy) - theorem continuous_of_cover_nhds {ι : Sort*} {f : α → β} {s : ι → Set α} (hs : ∀ x : α, ∃ i, s i ∈ 𝓝 x) (hf : ∀ i, ContinuousOn f (s i)) : Continuous f := @@ -629,9 +587,66 @@ theorem Set.Subsingleton.continuousOn {s : Set α} (hs : s.Subsingleton) (f : α ContinuousOn f s := hs.induction_on (continuousOn_empty f) (continuousOn_singleton f) -theorem nhdsWithin_le_comap {x : α} {s : Set α} {f : α → β} (ctsf : ContinuousWithinAt f s x) : - 𝓝[s] x ≤ comap f (𝓝[f '' s] f x) := - ctsf.tendsto_nhdsWithin_image.le_comap +theorem continuousOn_open_iff {f : α → β} {s : Set α} (hs : IsOpen s) : + ContinuousOn f s ↔ ∀ t, IsOpen t → IsOpen (s ∩ f ⁻¹' t) := by + rw [continuousOn_iff'] + constructor + · intro h t ht + rcases h t ht with ⟨u, u_open, hu⟩ + rw [inter_comm, hu] + apply IsOpen.inter u_open hs + · intro h t ht + refine ⟨s ∩ f ⁻¹' t, h t ht, ?_⟩ + rw [@inter_comm _ s (f ⁻¹' t), inter_assoc, inter_self] + +theorem ContinuousOn.isOpen_inter_preimage {f : α → β} {s : Set α} {t : Set β} + (hf : ContinuousOn f s) (hs : IsOpen s) (ht : IsOpen t) : IsOpen (s ∩ f ⁻¹' t) := + (continuousOn_open_iff hs).1 hf t ht + +theorem ContinuousOn.isOpen_preimage {f : α → β} {s : Set α} {t : Set β} (h : ContinuousOn f s) + (hs : IsOpen s) (hp : f ⁻¹' t ⊆ s) (ht : IsOpen t) : IsOpen (f ⁻¹' t) := by + convert (continuousOn_open_iff hs).mp h t ht + rw [inter_comm, inter_eq_self_of_subset_left hp] + +theorem ContinuousOn.preimage_isClosed_of_isClosed {f : α → β} {s : Set α} {t : Set β} + (hf : ContinuousOn f s) (hs : IsClosed s) (ht : IsClosed t) : IsClosed (s ∩ f ⁻¹' t) := by + rcases continuousOn_iff_isClosed.1 hf t ht with ⟨u, hu⟩ + rw [inter_comm, hu.2] + apply IsClosed.inter hu.1 hs + +theorem ContinuousOn.preimage_interior_subset_interior_preimage {f : α → β} {s : Set α} {t : Set β} + (hf : ContinuousOn f s) (hs : IsOpen s) : s ∩ f ⁻¹' interior t ⊆ s ∩ interior (f ⁻¹' t) := + calc + s ∩ f ⁻¹' interior t ⊆ interior (s ∩ f ⁻¹' t) := + interior_maximal (inter_subset_inter (Subset.refl _) (preimage_mono interior_subset)) + (hf.isOpen_inter_preimage hs isOpen_interior) + _ = s ∩ interior (f ⁻¹' t) := by rw [interior_inter, hs.interior_eq] + +theorem continuousOn_of_locally_continuousOn {f : α → β} {s : Set α} + (h : ∀ x ∈ s, ∃ t, IsOpen t ∧ x ∈ t ∧ ContinuousOn f (s ∩ t)) : ContinuousOn f s := by + intro x xs + rcases h x xs with ⟨t, open_t, xt, ct⟩ + have := ct x ⟨xs, xt⟩ + rwa [ContinuousWithinAt, ← nhdsWithin_restrict _ xt open_t] at this + +theorem continuousOn_to_generateFrom_iff {β} {s : Set α} {T : Set (Set β)} {f : α → β} : + @ContinuousOn α β _ (.generateFrom T) f s ↔ ∀ x ∈ s, ∀ t ∈ T, f x ∈ t → f ⁻¹' t ∈ 𝓝[s] x := + forall₂_congr fun x _ => by + delta ContinuousWithinAt + simp only [TopologicalSpace.nhds_generateFrom, tendsto_iInf, tendsto_principal, mem_setOf_eq, + and_imp] + exact forall_congr' fun t => forall_swap + +-- Porting note: dropped an unneeded assumption +theorem continuousOn_isOpen_of_generateFrom {β : Type*} {s : Set α} {T : Set (Set β)} {f : α → β} + (h : ∀ t ∈ T, IsOpen (s ∩ f ⁻¹' t)) : + @ContinuousOn α β _ (.generateFrom T) f s := + continuousOn_to_generateFrom_iff.2 fun _x hx t ht hxt => mem_nhdsWithin.2 + ⟨_, h t ht, ⟨hx, hxt⟩, fun _y hy => hy.1.2⟩ + +/-! +### Congruence and monotonicity properties with respect to sets +-/ theorem ContinuousWithinAt.mono {f : α → β} {s t : Set α} {x : α} (h : ContinuousWithinAt f t x) (hs : s ⊆ t) : ContinuousWithinAt f s x := @@ -661,31 +676,6 @@ theorem ContinuousWithinAt.union {f : α → β} {s t : Set α} {x : α} (hs : C (ht : ContinuousWithinAt f t x) : ContinuousWithinAt f (s ∪ t) x := continuousWithinAt_union.2 ⟨hs, ht⟩ -theorem ContinuousWithinAt.mem_closure_image {f : α → β} {s : Set α} {x : α} - (h : ContinuousWithinAt f s x) (hx : x ∈ closure s) : f x ∈ closure (f '' s) := - haveI := mem_closure_iff_nhdsWithin_neBot.1 hx - mem_closure_of_tendsto h <| mem_of_superset self_mem_nhdsWithin (subset_preimage_image f s) - -theorem ContinuousWithinAt.mem_closure {f : α → β} {s : Set α} {x : α} {A : Set β} - (h : ContinuousWithinAt f s x) (hx : x ∈ closure s) (hA : MapsTo f s A) : f x ∈ closure A := - closure_mono (image_subset_iff.2 hA) (h.mem_closure_image hx) - -theorem Set.MapsTo.closure_of_continuousWithinAt {f : α → β} {s : Set α} {t : Set β} - (h : MapsTo f s t) (hc : ∀ x ∈ closure s, ContinuousWithinAt f s x) : - MapsTo f (closure s) (closure t) := fun x hx => (hc x hx).mem_closure hx h - -theorem Set.MapsTo.closure_of_continuousOn {f : α → β} {s : Set α} {t : Set β} (h : MapsTo f s t) - (hc : ContinuousOn f (closure s)) : MapsTo f (closure s) (closure t) := - h.closure_of_continuousWithinAt fun x hx => (hc x hx).mono subset_closure - -theorem ContinuousWithinAt.image_closure {f : α → β} {s : Set α} - (hf : ∀ x ∈ closure s, ContinuousWithinAt f s x) : f '' closure s ⊆ closure (f '' s) := - ((mapsTo_image f s).closure_of_continuousWithinAt hf).image_subset - -theorem ContinuousOn.image_closure {f : α → β} {s : Set α} (hf : ContinuousOn f (closure s)) : - f '' closure s ⊆ closure (f '' s) := - ContinuousWithinAt.image_closure fun x hx => (hf x hx).mono subset_closure - @[simp] theorem continuousWithinAt_singleton {f : α → β} {x : α} : ContinuousWithinAt f {x} x := by simp only [ContinuousWithinAt, nhdsWithin_singleton, tendsto_pure_nhds] @@ -712,50 +702,16 @@ theorem continuousWithinAt_compl_self {f : α → β} {a : α} : ContinuousWithinAt f {a}ᶜ a ↔ ContinuousAt f a := by rw [compl_eq_univ_diff, continuousWithinAt_diff_self, continuousWithinAt_univ] -@[simp] -theorem continuousWithinAt_update_same [DecidableEq α] {f : α → β} {s : Set α} {x : α} {y : β} : - ContinuousWithinAt (update f x y) s x ↔ Tendsto f (𝓝[s \ {x}] x) (𝓝 y) := - calc - ContinuousWithinAt (update f x y) s x ↔ Tendsto (update f x y) (𝓝[s \ {x}] x) (𝓝 y) := by - { rw [← continuousWithinAt_diff_self, ContinuousWithinAt, update_same] } - _ ↔ Tendsto f (𝓝[s \ {x}] x) (𝓝 y) := - tendsto_congr' <| eventually_nhdsWithin_iff.2 <| Eventually.of_forall - fun _ hz => update_noteq hz.2 _ _ - -@[simp] -theorem continuousAt_update_same [DecidableEq α] {f : α → β} {x : α} {y : β} : - ContinuousAt (Function.update f x y) x ↔ Tendsto f (𝓝[≠] x) (𝓝 y) := by - rw [← continuousWithinAt_univ, continuousWithinAt_update_same, compl_eq_univ_diff] - -theorem IsOpenMap.continuousOn_image_of_leftInvOn {f : α → β} {s : Set α} - (h : IsOpenMap (s.restrict f)) {finv : β → α} (hleft : LeftInvOn finv f s) : - ContinuousOn finv (f '' s) := by - refine continuousOn_iff'.2 fun t ht => ⟨f '' (t ∩ s), ?_, ?_⟩ - · rw [← image_restrict] - exact h _ (ht.preimage continuous_subtype_val) - · rw [inter_eq_self_of_subset_left (image_subset f inter_subset_right), hleft.image_inter'] - -theorem IsOpenMap.continuousOn_range_of_leftInverse {f : α → β} (hf : IsOpenMap f) {finv : β → α} - (hleft : Function.LeftInverse finv f) : ContinuousOn finv (range f) := by - rw [← image_univ] - exact (hf.restrict isOpen_univ).continuousOn_image_of_leftInvOn fun x _ => hleft x +theorem ContinuousOn.mono {f : α → β} {s t : Set α} (hf : ContinuousOn f s) (h : t ⊆ s) : + ContinuousOn f t := fun x hx => (hf x (h hx)).mono_left (nhdsWithin_mono _ h) -theorem ContinuousOn.congr_mono {f g : α → β} {s s₁ : Set α} (h : ContinuousOn f s) - (h' : EqOn g f s₁) (h₁ : s₁ ⊆ s) : ContinuousOn g s₁ := by - intro x hx - unfold ContinuousWithinAt - have A := (h x (h₁ hx)).mono h₁ - unfold ContinuousWithinAt at A - rw [← h' hx] at A - exact A.congr' h'.eventuallyEq_nhdsWithin.symm +theorem antitone_continuousOn {f : α → β} : Antitone (ContinuousOn f) := fun _s _t hst hf => + hf.mono hst -theorem ContinuousOn.congr {f g : α → β} {s : Set α} (h : ContinuousOn f s) (h' : EqOn g f s) : - ContinuousOn g s := - h.congr_mono h' (Subset.refl _) -theorem continuousOn_congr {f g : α → β} {s : Set α} (h' : EqOn g f s) : - ContinuousOn g s ↔ ContinuousOn f s := - ⟨fun h => ContinuousOn.congr h h'.symm, fun h => h.congr h'⟩ +/-! +### Relation between `ContinuousAt` and `ContinuousWithinAt` +-/ theorem ContinuousAt.continuousWithinAt {f : α → β} {s : Set α} {x : α} (h : ContinuousAt f x) : ContinuousWithinAt f s x := @@ -780,7 +736,61 @@ theorem ContinuousOn.continuousAt {f : α → β} {s : Set α} {x : α} (h : Con theorem ContinuousAt.continuousOn {f : α → β} {s : Set α} (hcont : ∀ x ∈ s, ContinuousAt f x) : ContinuousOn f s := fun x hx => (hcont x hx).continuousWithinAt -theorem ContinuousWithinAt.comp {g : β → γ} {f : α → β} {s : Set α} {t : Set β} {x : α} +@[fun_prop] +theorem Continuous.continuousOn {f : α → β} {s : Set α} (h : Continuous f) : ContinuousOn f s := by + rw [continuous_iff_continuousOn_univ] at h + exact h.mono (subset_univ _) + +theorem Continuous.continuousWithinAt {f : α → β} {s : Set α} {x : α} (h : Continuous f) : + ContinuousWithinAt f s x := + h.continuousAt.continuousWithinAt + + +/-! +### Congruence properties with respect to functions +-/ + +theorem ContinuousOn.congr_mono {f g : α → β} {s s₁ : Set α} (h : ContinuousOn f s) + (h' : EqOn g f s₁) (h₁ : s₁ ⊆ s) : ContinuousOn g s₁ := by + intro x hx + unfold ContinuousWithinAt + have A := (h x (h₁ hx)).mono h₁ + unfold ContinuousWithinAt at A + rw [← h' hx] at A + exact A.congr' h'.eventuallyEq_nhdsWithin.symm + +theorem ContinuousOn.congr {f g : α → β} {s : Set α} (h : ContinuousOn f s) (h' : EqOn g f s) : + ContinuousOn g s := + h.congr_mono h' (Subset.refl _) + +theorem continuousOn_congr {f g : α → β} {s : Set α} (h' : EqOn g f s) : + ContinuousOn g s ↔ ContinuousOn f s := + ⟨fun h => ContinuousOn.congr h h'.symm, fun h => h.congr h'⟩ + +theorem Filter.EventuallyEq.congr_continuousWithinAt {f g : α → β} {s : Set α} {x : α} + (h : f =ᶠ[𝓝[s] x] g) (hx : f x = g x) : + ContinuousWithinAt f s x ↔ ContinuousWithinAt g s x := by + rw [ContinuousWithinAt, hx, tendsto_congr' h, ContinuousWithinAt] + +theorem ContinuousWithinAt.congr_of_eventuallyEq {f f₁ : α → β} {s : Set α} {x : α} + (h : ContinuousWithinAt f s x) (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : + ContinuousWithinAt f₁ s x := + (h₁.congr_continuousWithinAt hx).2 h + +theorem ContinuousWithinAt.congr {f f₁ : α → β} {s : Set α} {x : α} (h : ContinuousWithinAt f s x) + (h₁ : ∀ y ∈ s, f₁ y = f y) (hx : f₁ x = f x) : ContinuousWithinAt f₁ s x := + h.congr_of_eventuallyEq (mem_of_superset self_mem_nhdsWithin h₁) hx + +theorem ContinuousWithinAt.congr_mono {f g : α → β} {s s₁ : Set α} {x : α} + (h : ContinuousWithinAt f s x) (h' : EqOn g f s₁) (h₁ : s₁ ⊆ s) (hx : g x = f x) : + ContinuousWithinAt g s₁ x := + (h.mono h₁).congr h' hx + +/-! +### Composition +-/ + +theorem ContinuousWithinAt.comp {g : β → γ} {f : α → β} {s : Set α} {t : Set β} {x : α} (hg : ContinuousWithinAt g t (f x)) (hf : ContinuousWithinAt f s x) (h : MapsTo f s t) : ContinuousWithinAt (g ∘ f) s x := hg.tendsto.comp (hf.tendsto_nhdsWithin h) @@ -803,26 +813,11 @@ theorem ContinuousOn.comp'' {g : β → γ} {f : α → β} {s : Set α} {t : Se (hf : ContinuousOn f s) (h : Set.MapsTo f s t) : ContinuousOn (fun x => g (f x)) s := ContinuousOn.comp hg hf h -theorem ContinuousOn.mono {f : α → β} {s t : Set α} (hf : ContinuousOn f s) (h : t ⊆ s) : - ContinuousOn f t := fun x hx => (hf x (h hx)).mono_left (nhdsWithin_mono _ h) - -theorem antitone_continuousOn {f : α → β} : Antitone (ContinuousOn f) := fun _s _t hst hf => - hf.mono hst - @[fun_prop] theorem ContinuousOn.comp' {g : β → γ} {f : α → β} {s : Set α} {t : Set β} (hg : ContinuousOn g t) (hf : ContinuousOn f s) : ContinuousOn (g ∘ f) (s ∩ f ⁻¹' t) := hg.comp (hf.mono inter_subset_left) inter_subset_right -@[fun_prop] -theorem Continuous.continuousOn {f : α → β} {s : Set α} (h : Continuous f) : ContinuousOn f s := by - rw [continuous_iff_continuousOn_univ] at h - exact h.mono (subset_univ _) - -theorem Continuous.continuousWithinAt {f : α → β} {s : Set α} {x : α} (h : Continuous f) : - ContinuousWithinAt f s x := - h.continuousAt.continuousWithinAt - theorem Continuous.comp_continuousOn {g : β → γ} {f : α → β} {s : Set α} (hg : Continuous g) (hf : ContinuousOn f s) : ContinuousOn (g ∘ f) s := hg.continuousOn.comp hf (mapsTo_univ _ _) @@ -839,60 +834,185 @@ theorem ContinuousOn.comp_continuous {g : β → γ} {f : α → β} {s : Set β rw [continuous_iff_continuousOn_univ] at * exact hg.comp hf fun x _ => hs x +theorem ContinuousAt.comp₂_continuousWithinAt {f : β × γ → δ} {g : α → β} {h : α → γ} {x : α} + {s : Set α} (hf : ContinuousAt f (g x, h x)) (hg : ContinuousWithinAt g s x) + (hh : ContinuousWithinAt h s x) : + ContinuousWithinAt (fun x ↦ f (g x, h x)) s x := + ContinuousAt.comp_continuousWithinAt hf (hg.prod_mk_nhds hh) + +theorem ContinuousAt.comp₂_continuousWithinAt_of_eq {f : β × γ → δ} {g : α → β} + {h : α → γ} {x : α} {s : Set α} {y : β × γ} (hf : ContinuousAt f y) + (hg : ContinuousWithinAt g s x) (hh : ContinuousWithinAt h s x) (e : (g x, h x) = y) : + ContinuousWithinAt (fun x ↦ f (g x, h x)) s x := by + rw [← e] at hf + exact hf.comp₂_continuousWithinAt hg hh + + +/-! +### Image +-/ + +theorem ContinuousWithinAt.mem_closure_image {f : α → β} {s : Set α} {x : α} + (h : ContinuousWithinAt f s x) (hx : x ∈ closure s) : f x ∈ closure (f '' s) := + haveI := mem_closure_iff_nhdsWithin_neBot.1 hx + mem_closure_of_tendsto h <| mem_of_superset self_mem_nhdsWithin (subset_preimage_image f s) + +theorem ContinuousWithinAt.mem_closure {f : α → β} {s : Set α} {x : α} {A : Set β} + (h : ContinuousWithinAt f s x) (hx : x ∈ closure s) (hA : MapsTo f s A) : f x ∈ closure A := + closure_mono (image_subset_iff.2 hA) (h.mem_closure_image hx) + +theorem Set.MapsTo.closure_of_continuousWithinAt {f : α → β} {s : Set α} {t : Set β} + (h : MapsTo f s t) (hc : ∀ x ∈ closure s, ContinuousWithinAt f s x) : + MapsTo f (closure s) (closure t) := fun x hx => (hc x hx).mem_closure hx h + +theorem Set.MapsTo.closure_of_continuousOn {f : α → β} {s : Set α} {t : Set β} (h : MapsTo f s t) + (hc : ContinuousOn f (closure s)) : MapsTo f (closure s) (closure t) := + h.closure_of_continuousWithinAt fun x hx => (hc x hx).mono subset_closure + +theorem ContinuousWithinAt.image_closure {f : α → β} {s : Set α} + (hf : ∀ x ∈ closure s, ContinuousWithinAt f s x) : f '' closure s ⊆ closure (f '' s) := + ((mapsTo_image f s).closure_of_continuousWithinAt hf).image_subset + +theorem ContinuousOn.image_closure {f : α → β} {s : Set α} (hf : ContinuousOn f (closure s)) : + f '' closure s ⊆ closure (f '' s) := + ContinuousWithinAt.image_closure fun x hx => (hf x hx).mono subset_closure + +/-! +### Product +-/ + +theorem ContinuousWithinAt.prod_map {f : α → γ} {g : β → δ} {s : Set α} {t : Set β} {x : α} {y : β} + (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g t y) : + ContinuousWithinAt (Prod.map f g) (s ×ˢ t) (x, y) := by + unfold ContinuousWithinAt at * + rw [nhdsWithin_prod_eq, Prod.map, nhds_prod_eq] + exact hf.prod_map hg + +theorem continuousWithinAt_prod_of_discrete_left [DiscreteTopology α] + {f : α × β → γ} {s : Set (α × β)} {x : α × β} : + ContinuousWithinAt f s x ↔ ContinuousWithinAt (f ⟨x.1, ·⟩) {b | (x.1, b) ∈ s} x.2 := by + rw [← x.eta]; simp_rw [ContinuousWithinAt, nhdsWithin, nhds_prod_eq, nhds_discrete, pure_prod, + ← map_inf_principal_preimage]; rfl + +theorem continuousWithinAt_prod_of_discrete_right [DiscreteTopology β] + {f : α × β → γ} {s : Set (α × β)} {x : α × β} : + ContinuousWithinAt f s x ↔ ContinuousWithinAt (f ⟨·, x.2⟩) {a | (a, x.2) ∈ s} x.1 := by + rw [← x.eta]; simp_rw [ContinuousWithinAt, nhdsWithin, nhds_prod_eq, nhds_discrete, prod_pure, + ← map_inf_principal_preimage]; rfl + +theorem continuousAt_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} {x : α × β} : + ContinuousAt f x ↔ ContinuousAt (f ⟨x.1, ·⟩) x.2 := by + simp_rw [← continuousWithinAt_univ]; exact continuousWithinAt_prod_of_discrete_left + +theorem continuousAt_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} {x : α × β} : + ContinuousAt f x ↔ ContinuousAt (f ⟨·, x.2⟩) x.1 := by + simp_rw [← continuousWithinAt_univ]; exact continuousWithinAt_prod_of_discrete_right + +theorem continuousOn_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} {s : Set (α × β)} : + ContinuousOn f s ↔ ∀ a, ContinuousOn (f ⟨a, ·⟩) {b | (a, b) ∈ s} := by + simp_rw [ContinuousOn, Prod.forall, continuousWithinAt_prod_of_discrete_left]; rfl + +theorem continuousOn_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} {s : Set (α × β)} : + ContinuousOn f s ↔ ∀ b, ContinuousOn (f ⟨·, b⟩) {a | (a, b) ∈ s} := by + simp_rw [ContinuousOn, Prod.forall, continuousWithinAt_prod_of_discrete_right]; apply forall_swap + +/-- If a function `f a b` is such that `y ↦ f a b` is continuous for all `a`, and `a` lives in a +discrete space, then `f` is continuous, and vice versa. -/ +theorem continuous_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} : + Continuous f ↔ ∀ a, Continuous (f ⟨a, ·⟩) := by + simp_rw [continuous_iff_continuousOn_univ]; exact continuousOn_prod_of_discrete_left + +theorem continuous_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} : + Continuous f ↔ ∀ b, Continuous (f ⟨·, b⟩) := by + simp_rw [continuous_iff_continuousOn_univ]; exact continuousOn_prod_of_discrete_right + +theorem isOpenMap_prod_of_discrete_left [DiscreteTopology α] {f : α × β → γ} : + IsOpenMap f ↔ ∀ a, IsOpenMap (f ⟨a, ·⟩) := by + simp_rw [isOpenMap_iff_nhds_le, Prod.forall, nhds_prod_eq, nhds_discrete, pure_prod, map_map] + rfl + +theorem isOpenMap_prod_of_discrete_right [DiscreteTopology β] {f : α × β → γ} : + IsOpenMap f ↔ ∀ b, IsOpenMap (f ⟨·, b⟩) := by + simp_rw [isOpenMap_iff_nhds_le, Prod.forall, forall_swap (α := α) (β := β), nhds_prod_eq, + nhds_discrete, prod_pure, map_map]; rfl + +theorem ContinuousOn.prod_map {f : α → γ} {g : β → δ} {s : Set α} {t : Set β} + (hf : ContinuousOn f s) (hg : ContinuousOn g t) : ContinuousOn (Prod.map f g) (s ×ˢ t) := + fun ⟨x, y⟩ ⟨hx, hy⟩ => ContinuousWithinAt.prod_map (hf x hx) (hg y hy) + +theorem ContinuousWithinAt.prod {f : α → β} {g : α → γ} {s : Set α} {x : α} + (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g s x) : + ContinuousWithinAt (fun x => (f x, g x)) s x := + hf.prod_mk_nhds hg + @[fun_prop] -theorem continuousOn_apply {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] - (i : ι) (s) : ContinuousOn (fun p : ∀ i, π i => p i) s := - Continuous.continuousOn (continuous_apply i) +theorem ContinuousOn.prod {f : α → β} {g : α → γ} {s : Set α} (hf : ContinuousOn f s) + (hg : ContinuousOn g s) : ContinuousOn (fun x => (f x, g x)) s := fun x hx => + ContinuousWithinAt.prod (hf x hx) (hg x hx) -theorem ContinuousWithinAt.preimage_mem_nhdsWithin {f : α → β} {x : α} {s : Set α} {t : Set β} - (h : ContinuousWithinAt f s x) (ht : t ∈ 𝓝 (f x)) : f ⁻¹' t ∈ 𝓝[s] x := - h ht +theorem continuousOn_fst {s : Set (α × β)} : ContinuousOn Prod.fst s := + continuous_fst.continuousOn -theorem Set.LeftInvOn.map_nhdsWithin_eq {f : α → β} {g : β → α} {x : β} {s : Set β} - (h : LeftInvOn f g s) (hx : f (g x) = x) (hf : ContinuousWithinAt f (g '' s) (g x)) - (hg : ContinuousWithinAt g s x) : map g (𝓝[s] x) = 𝓝[g '' s] g x := by - apply le_antisymm - · exact hg.tendsto_nhdsWithin (mapsTo_image _ _) - · have A : g ∘ f =ᶠ[𝓝[g '' s] g x] id := - h.rightInvOn_image.eqOn.eventuallyEq_of_mem self_mem_nhdsWithin - refine le_map_of_right_inverse A ?_ - simpa only [hx] using hf.tendsto_nhdsWithin (h.mapsTo (surjOn_image _ _)) +theorem continuousWithinAt_fst {s : Set (α × β)} {p : α × β} : ContinuousWithinAt Prod.fst s p := + continuous_fst.continuousWithinAt -theorem Function.LeftInverse.map_nhds_eq {f : α → β} {g : β → α} {x : β} - (h : Function.LeftInverse f g) (hf : ContinuousWithinAt f (range g) (g x)) - (hg : ContinuousAt g x) : map g (𝓝 x) = 𝓝[range g] g x := by - simpa only [nhdsWithin_univ, image_univ] using - (h.leftInvOn univ).map_nhdsWithin_eq (h x) (by rwa [image_univ]) hg.continuousWithinAt +@[fun_prop] +theorem ContinuousOn.fst {f : α → β × γ} {s : Set α} (hf : ContinuousOn f s) : + ContinuousOn (fun x => (f x).1) s := + continuous_fst.comp_continuousOn hf -theorem ContinuousWithinAt.preimage_mem_nhdsWithin' {f : α → β} {x : α} {s : Set α} {t : Set β} - (h : ContinuousWithinAt f s x) (ht : t ∈ 𝓝[f '' s] f x) : f ⁻¹' t ∈ 𝓝[s] x := - h.tendsto_nhdsWithin (mapsTo_image _ _) ht +theorem ContinuousWithinAt.fst {f : α → β × γ} {s : Set α} {a : α} (h : ContinuousWithinAt f s a) : + ContinuousWithinAt (fun x => (f x).fst) s a := + continuousAt_fst.comp_continuousWithinAt h -theorem ContinuousWithinAt.preimage_mem_nhdsWithin'' - {f : α → β} {x : α} {y : β} {s t : Set β} - (h : ContinuousWithinAt f (f ⁻¹' s) x) (ht : t ∈ 𝓝[s] y) (hxy : y = f x) : - f ⁻¹' t ∈ 𝓝[f ⁻¹' s] x := by - rw [hxy] at ht - exact h.preimage_mem_nhdsWithin' (nhdsWithin_mono _ (image_preimage_subset f s) ht) +theorem continuousOn_snd {s : Set (α × β)} : ContinuousOn Prod.snd s := + continuous_snd.continuousOn -theorem Filter.EventuallyEq.congr_continuousWithinAt {f g : α → β} {s : Set α} {x : α} - (h : f =ᶠ[𝓝[s] x] g) (hx : f x = g x) : - ContinuousWithinAt f s x ↔ ContinuousWithinAt g s x := by - rw [ContinuousWithinAt, hx, tendsto_congr' h, ContinuousWithinAt] +theorem continuousWithinAt_snd {s : Set (α × β)} {p : α × β} : ContinuousWithinAt Prod.snd s p := + continuous_snd.continuousWithinAt -theorem ContinuousWithinAt.congr_of_eventuallyEq {f f₁ : α → β} {s : Set α} {x : α} - (h : ContinuousWithinAt f s x) (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : - ContinuousWithinAt f₁ s x := - (h₁.congr_continuousWithinAt hx).2 h +@[fun_prop] +theorem ContinuousOn.snd {f : α → β × γ} {s : Set α} (hf : ContinuousOn f s) : + ContinuousOn (fun x => (f x).2) s := + continuous_snd.comp_continuousOn hf -theorem ContinuousWithinAt.congr {f f₁ : α → β} {s : Set α} {x : α} (h : ContinuousWithinAt f s x) - (h₁ : ∀ y ∈ s, f₁ y = f y) (hx : f₁ x = f x) : ContinuousWithinAt f₁ s x := - h.congr_of_eventuallyEq (mem_of_superset self_mem_nhdsWithin h₁) hx +theorem ContinuousWithinAt.snd {f : α → β × γ} {s : Set α} {a : α} (h : ContinuousWithinAt f s a) : + ContinuousWithinAt (fun x => (f x).snd) s a := + continuousAt_snd.comp_continuousWithinAt h -theorem ContinuousWithinAt.congr_mono {f g : α → β} {s s₁ : Set α} {x : α} - (h : ContinuousWithinAt f s x) (h' : EqOn g f s₁) (h₁ : s₁ ⊆ s) (hx : g x = f x) : - ContinuousWithinAt g s₁ x := - (h.mono h₁).congr h' hx +theorem continuousWithinAt_prod_iff {f : α → β × γ} {s : Set α} {x : α} : + ContinuousWithinAt f s x ↔ + ContinuousWithinAt (Prod.fst ∘ f) s x ∧ ContinuousWithinAt (Prod.snd ∘ f) s x := + ⟨fun h => ⟨h.fst, h.snd⟩, fun ⟨h1, h2⟩ => h1.prod h2⟩ + +/-! +### Pi +-/ + +theorem continuousWithinAt_pi {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] + {f : α → ∀ i, π i} {s : Set α} {x : α} : + ContinuousWithinAt f s x ↔ ∀ i, ContinuousWithinAt (fun y => f y i) s x := + tendsto_pi_nhds + +theorem continuousOn_pi {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] + {f : α → ∀ i, π i} {s : Set α} : ContinuousOn f s ↔ ∀ i, ContinuousOn (fun y => f y i) s := + ⟨fun h i x hx => tendsto_pi_nhds.1 (h x hx) i, fun h x hx => tendsto_pi_nhds.2 fun i => h i x hx⟩ + +@[fun_prop] +theorem continuousOn_pi' {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] + {f : α → ∀ i, π i} {s : Set α} (hf : ∀ i, ContinuousOn (fun y => f y i) s) : + ContinuousOn f s := + continuousOn_pi.2 hf + +@[fun_prop] +theorem continuousOn_apply {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] + (i : ι) (s) : ContinuousOn (fun p : ∀ i, π i => p i) s := + Continuous.continuousOn (continuous_apply i) + + +/-! +### Specific functions +-/ @[fun_prop] theorem continuousOn_const {s : Set α} {c : β} : ContinuousOn (fun _ => c) s := @@ -916,85 +1036,33 @@ protected theorem ContinuousOn.iterate {f : α → α} {s : Set α} (hcont : Con | 0 => continuousOn_id | (n + 1) => (hcont.iterate hmaps n).comp hcont hmaps -theorem continuousOn_open_iff {f : α → β} {s : Set α} (hs : IsOpen s) : - ContinuousOn f s ↔ ∀ t, IsOpen t → IsOpen (s ∩ f ⁻¹' t) := by - rw [continuousOn_iff'] - constructor - · intro h t ht - rcases h t ht with ⟨u, u_open, hu⟩ - rw [inter_comm, hu] - apply IsOpen.inter u_open hs - · intro h t ht - refine ⟨s ∩ f ⁻¹' t, h t ht, ?_⟩ - rw [@inter_comm _ s (f ⁻¹' t), inter_assoc, inter_self] - -theorem ContinuousOn.isOpen_inter_preimage {f : α → β} {s : Set α} {t : Set β} - (hf : ContinuousOn f s) (hs : IsOpen s) (ht : IsOpen t) : IsOpen (s ∩ f ⁻¹' t) := - (continuousOn_open_iff hs).1 hf t ht - -theorem ContinuousOn.isOpen_preimage {f : α → β} {s : Set α} {t : Set β} (h : ContinuousOn f s) - (hs : IsOpen s) (hp : f ⁻¹' t ⊆ s) (ht : IsOpen t) : IsOpen (f ⁻¹' t) := by - convert (continuousOn_open_iff hs).mp h t ht - rw [inter_comm, inter_eq_self_of_subset_left hp] - -theorem ContinuousOn.preimage_isClosed_of_isClosed {f : α → β} {s : Set α} {t : Set β} - (hf : ContinuousOn f s) (hs : IsClosed s) (ht : IsClosed t) : IsClosed (s ∩ f ⁻¹' t) := by - rcases continuousOn_iff_isClosed.1 hf t ht with ⟨u, hu⟩ - rw [inter_comm, hu.2] - apply IsClosed.inter hu.1 hs - -theorem ContinuousOn.preimage_interior_subset_interior_preimage {f : α → β} {s : Set α} {t : Set β} - (hf : ContinuousOn f s) (hs : IsOpen s) : s ∩ f ⁻¹' interior t ⊆ s ∩ interior (f ⁻¹' t) := - calc - s ∩ f ⁻¹' interior t ⊆ interior (s ∩ f ⁻¹' t) := - interior_maximal (inter_subset_inter (Subset.refl _) (preimage_mono interior_subset)) - (hf.isOpen_inter_preimage hs isOpen_interior) - _ = s ∩ interior (f ⁻¹' t) := by rw [interior_inter, hs.interior_eq] - -theorem continuousOn_of_locally_continuousOn {f : α → β} {s : Set α} - (h : ∀ x ∈ s, ∃ t, IsOpen t ∧ x ∈ t ∧ ContinuousOn f (s ∩ t)) : ContinuousOn f s := by - intro x xs - rcases h x xs with ⟨t, open_t, xt, ct⟩ - have := ct x ⟨xs, xt⟩ - rwa [ContinuousWithinAt, ← nhdsWithin_restrict _ xt open_t] at this - -theorem continuousOn_to_generateFrom_iff {β} {s : Set α} {T : Set (Set β)} {f : α → β} : - @ContinuousOn α β _ (.generateFrom T) f s ↔ ∀ x ∈ s, ∀ t ∈ T, f x ∈ t → f ⁻¹' t ∈ 𝓝[s] x := - forall₂_congr fun x _ => by - delta ContinuousWithinAt - simp only [TopologicalSpace.nhds_generateFrom, tendsto_iInf, tendsto_principal, mem_setOf_eq, - and_imp] - exact forall_congr' fun t => forall_swap - --- Porting note: dropped an unneeded assumption -theorem continuousOn_isOpen_of_generateFrom {β : Type*} {s : Set α} {T : Set (Set β)} {f : α → β} - (h : ∀ t ∈ T, IsOpen (s ∩ f ⁻¹' t)) : - @ContinuousOn α β _ (.generateFrom T) f s := - continuousOn_to_generateFrom_iff.2 fun _x hx t ht hxt => mem_nhdsWithin.2 - ⟨_, h t ht, ⟨hx, hxt⟩, fun _y hy => hy.1.2⟩ - -theorem ContinuousWithinAt.prod {f : α → β} {g : α → γ} {s : Set α} {x : α} - (hf : ContinuousWithinAt f s x) (hg : ContinuousWithinAt g s x) : - ContinuousWithinAt (fun x => (f x, g x)) s x := - hf.prod_mk_nhds hg +theorem ContinuousWithinAt.fin_insertNth {n} {π : Fin (n + 1) → Type*} + [∀ i, TopologicalSpace (π i)] (i : Fin (n + 1)) {f : α → π i} {a : α} {s : Set α} + (hf : ContinuousWithinAt f s a) {g : α → ∀ j : Fin n, π (i.succAbove j)} + (hg : ContinuousWithinAt g s a) : ContinuousWithinAt (fun a => i.insertNth (f a) (g a)) s a := + hf.tendsto.fin_insertNth i hg -@[fun_prop] -theorem ContinuousOn.prod {f : α → β} {g : α → γ} {s : Set α} (hf : ContinuousOn f s) - (hg : ContinuousOn g s) : ContinuousOn (fun x => (f x, g x)) s := fun x hx => - ContinuousWithinAt.prod (hf x hx) (hg x hx) +nonrec theorem ContinuousOn.fin_insertNth {n} {π : Fin (n + 1) → Type*} + [∀ i, TopologicalSpace (π i)] (i : Fin (n + 1)) {f : α → π i} {s : Set α} + (hf : ContinuousOn f s) {g : α → ∀ j : Fin n, π (i.succAbove j)} (hg : ContinuousOn g s) : + ContinuousOn (fun a => i.insertNth (f a) (g a)) s := fun a ha => + (hf a ha).fin_insertNth i (hg a ha) -theorem ContinuousAt.comp₂_continuousWithinAt {f : β × γ → δ} {g : α → β} {h : α → γ} {x : α} - {s : Set α} (hf : ContinuousAt f (g x, h x)) (hg : ContinuousWithinAt g s x) - (hh : ContinuousWithinAt h s x) : - ContinuousWithinAt (fun x ↦ f (g x, h x)) s x := - ContinuousAt.comp_continuousWithinAt hf (hg.prod hh) +theorem Set.LeftInvOn.map_nhdsWithin_eq {f : α → β} {g : β → α} {x : β} {s : Set β} + (h : LeftInvOn f g s) (hx : f (g x) = x) (hf : ContinuousWithinAt f (g '' s) (g x)) + (hg : ContinuousWithinAt g s x) : map g (𝓝[s] x) = 𝓝[g '' s] g x := by + apply le_antisymm + · exact hg.tendsto_nhdsWithin (mapsTo_image _ _) + · have A : g ∘ f =ᶠ[𝓝[g '' s] g x] id := + h.rightInvOn_image.eqOn.eventuallyEq_of_mem self_mem_nhdsWithin + refine le_map_of_right_inverse A ?_ + simpa only [hx] using hf.tendsto_nhdsWithin (h.mapsTo (surjOn_image _ _)) -theorem ContinuousAt.comp₂_continuousWithinAt_of_eq {f : β × γ → δ} {g : α → β} - {h : α → γ} {x : α} {s : Set α} {y : β × γ} (hf : ContinuousAt f y) - (hg : ContinuousWithinAt g s x) (hh : ContinuousWithinAt h s x) (e : (g x, h x) = y) : - ContinuousWithinAt (fun x ↦ f (g x, h x)) s x := by - rw [← e] at hf - exact hf.comp₂_continuousWithinAt hg hh +theorem Function.LeftInverse.map_nhds_eq {f : α → β} {g : β → α} {x : β} + (h : Function.LeftInverse f g) (hf : ContinuousWithinAt f (range g) (g x)) + (hg : ContinuousAt g x) : map g (𝓝 x) = 𝓝[range g] g x := by + simpa only [nhdsWithin_univ, image_univ] using + (h.leftInvOn univ).map_nhdsWithin_eq (h x) (by rwa [image_univ]) hg.continuousWithinAt theorem Inducing.continuousWithinAt_iff {f : α → β} {g : β → γ} (hg : Inducing g) {s : Set α} {x : α} : ContinuousWithinAt f s x ↔ ContinuousWithinAt (g ∘ f) s x := by @@ -1024,11 +1092,37 @@ theorem QuotientMap.continuousOn_isOpen_iff {f : α → β} {g : β → γ} (h : simp only [continuousOn_iff_continuous_restrict, (h.restrictPreimage_isOpen hs).continuous_iff] rfl -theorem continuousWithinAt_of_not_mem_closure {f : α → β} {s : Set α} {x : α} (hx : x ∉ closure s) : - ContinuousWithinAt f s x := by - rw [mem_closure_iff_nhdsWithin_neBot, not_neBot] at hx - rw [ContinuousWithinAt, hx] - exact tendsto_bot +theorem IsOpenMap.continuousOn_image_of_leftInvOn {f : α → β} {s : Set α} + (h : IsOpenMap (s.restrict f)) {finv : β → α} (hleft : LeftInvOn finv f s) : + ContinuousOn finv (f '' s) := by + refine continuousOn_iff'.2 fun t ht => ⟨f '' (t ∩ s), ?_, ?_⟩ + · rw [← image_restrict] + exact h _ (ht.preimage continuous_subtype_val) + · rw [inter_eq_self_of_subset_left (image_subset f inter_subset_right), hleft.image_inter'] + +theorem IsOpenMap.continuousOn_range_of_leftInverse {f : α → β} (hf : IsOpenMap f) {finv : β → α} + (hleft : Function.LeftInverse finv f) : ContinuousOn finv (range f) := by + rw [← image_univ] + exact (hf.restrict isOpen_univ).continuousOn_image_of_leftInvOn fun x _ => hleft x + +/-! +### Continuity of piecewise defined functions +-/ + +@[simp] +theorem continuousWithinAt_update_same [DecidableEq α] {f : α → β} {s : Set α} {x : α} {y : β} : + ContinuousWithinAt (update f x y) s x ↔ Tendsto f (𝓝[s \ {x}] x) (𝓝 y) := + calc + ContinuousWithinAt (update f x y) s x ↔ Tendsto (update f x y) (𝓝[s \ {x}] x) (𝓝 y) := by + { rw [← continuousWithinAt_diff_self, ContinuousWithinAt, update_same] } + _ ↔ Tendsto f (𝓝[s \ {x}] x) (𝓝 y) := + tendsto_congr' <| eventually_nhdsWithin_iff.2 <| Eventually.of_forall + fun _ hz => update_noteq hz.2 _ _ + +@[simp] +theorem continuousAt_update_same [DecidableEq α] {f : α → β} {x : α} {y : β} : + ContinuousAt (Function.update f x y) x ↔ Tendsto f (𝓝[≠] x) (𝓝 y) := by + rw [← continuousWithinAt_univ, continuousWithinAt_update_same, compl_eq_univ_diff] theorem ContinuousOn.if' {s : Set α} {p : α → Prop} {f g : α → β} [∀ a, Decidable (p a)] (hpf : ∀ a ∈ s ∩ frontier { a | p a }, @@ -1177,48 +1271,4 @@ theorem continuousOn_piecewise_ite' {s s' t : Set α} {f f' : α → β} [∀ x, theorem continuousOn_piecewise_ite {s s' t : Set α} {f f' : α → β} [∀ x, Decidable (x ∈ t)] (h : ContinuousOn f s) (h' : ContinuousOn f' s') (H : s ∩ frontier t = s' ∩ frontier t) (Heq : EqOn f f' (s ∩ frontier t)) : ContinuousOn (t.piecewise f f') (t.ite s s') := - continuousOn_piecewise_ite' (h.mono inter_subset_left) (h'.mono inter_subset_left) H - Heq - -theorem frontier_inter_open_inter {s t : Set α} (ht : IsOpen t) : - frontier (s ∩ t) ∩ t = frontier s ∩ t := by - simp only [Set.inter_comm _ t, ← Subtype.preimage_coe_eq_preimage_coe_iff, - ht.isOpenMap_subtype_val.preimage_frontier_eq_frontier_preimage continuous_subtype_val, - Subtype.preimage_coe_self_inter] - -theorem continuousOn_fst {s : Set (α × β)} : ContinuousOn Prod.fst s := - continuous_fst.continuousOn - -theorem continuousWithinAt_fst {s : Set (α × β)} {p : α × β} : ContinuousWithinAt Prod.fst s p := - continuous_fst.continuousWithinAt - -@[fun_prop] -theorem ContinuousOn.fst {f : α → β × γ} {s : Set α} (hf : ContinuousOn f s) : - ContinuousOn (fun x => (f x).1) s := - continuous_fst.comp_continuousOn hf - -theorem ContinuousWithinAt.fst {f : α → β × γ} {s : Set α} {a : α} (h : ContinuousWithinAt f s a) : - ContinuousWithinAt (fun x => (f x).fst) s a := - continuousAt_fst.comp_continuousWithinAt h - -theorem continuousOn_snd {s : Set (α × β)} : ContinuousOn Prod.snd s := - continuous_snd.continuousOn - -theorem continuousWithinAt_snd {s : Set (α × β)} {p : α × β} : ContinuousWithinAt Prod.snd s p := - continuous_snd.continuousWithinAt - -@[fun_prop] -theorem ContinuousOn.snd {f : α → β × γ} {s : Set α} (hf : ContinuousOn f s) : - ContinuousOn (fun x => (f x).2) s := - continuous_snd.comp_continuousOn hf - -theorem ContinuousWithinAt.snd {f : α → β × γ} {s : Set α} {a : α} (h : ContinuousWithinAt f s a) : - ContinuousWithinAt (fun x => (f x).snd) s a := - continuousAt_snd.comp_continuousWithinAt h - -theorem continuousWithinAt_prod_iff {f : α → β × γ} {s : Set α} {x : α} : - ContinuousWithinAt f s x ↔ - ContinuousWithinAt (Prod.fst ∘ f) s x ∧ ContinuousWithinAt (Prod.snd ∘ f) s x := - ⟨fun h => ⟨h.fst, h.snd⟩, fun ⟨h1, h2⟩ => h1.prod h2⟩ - -end Pi + continuousOn_piecewise_ite' (h.mono inter_subset_left) (h'.mono inter_subset_left) H Heq From 6c837c78f211be4560cf2b63adb5347e9854b808 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 18 Oct 2024 12:42:54 +0000 Subject: [PATCH 061/131] chore: delete >1year old deprecations (#17903) --- Mathlib/Algebra/BigOperators/Group/List.lean | 1 - Mathlib/Topology/Algebra/Module/Basic.lean | 5 ----- Mathlib/Topology/Separation.lean | 5 ----- Mathlib/Topology/UniformSpace/CompactConvergence.lean | 5 ----- 4 files changed, 16 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 42ca241b7d835..c61a25c867599 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -626,7 +626,6 @@ namespace MonoidHom protected theorem map_list_prod (f : M →* N) (l : List M) : f l.prod = (l.map f).prod := map_list_prod f l -attribute [deprecated map_list_prod (since := "2023-01-10")] MonoidHom.map_list_prod attribute [deprecated map_list_sum (since := "2024-05-02")] AddMonoidHom.map_list_sum end MonoidHom diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 637ef9b981a96..4296c9b917999 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -450,11 +450,6 @@ theorem map_smul_of_tower {R S : Type*} [Semiring S] [SMul R M₁] [Module S M f (c • x) = c • f x := LinearMap.CompatibleSMul.map_smul (f : M₁ →ₗ[S] M₂) c x -@[deprecated _root_.map_sum (since := "2023-09-16")] -protected theorem map_sum {ι : Type*} (f : M₁ →SL[σ₁₂] M₂) (s : Finset ι) (g : ι → M₁) : - f (∑ i ∈ s, g i) = ∑ i ∈ s, f (g i) := - map_sum .. - @[simp, norm_cast] theorem coe_coe (f : M₁ →SL[σ₁₂] M₂) : ⇑(f : M₁ →ₛₗ[σ₁₂] M₂) = f := rfl diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 4d39f3cd21001..792e2e67b2544 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -2122,11 +2122,6 @@ theorem exists_open_between_and_isCompact_closure [LocallyCompactSpace X] [Regul refine ⟨interior L, isOpen_interior, KL, A.trans LU, ?_⟩ exact L_compact.closure_of_subset interior_subset -@[deprecated WeaklyLocallyCompactSpace.locallyCompactSpace (since := "2023-09-03")] -theorem locally_compact_of_compact [T2Space X] [CompactSpace X] : - LocallyCompactSpace X := - inferInstance - end LocallyCompactRegularSpace section T25 diff --git a/Mathlib/Topology/UniformSpace/CompactConvergence.lean b/Mathlib/Topology/UniformSpace/CompactConvergence.lean index f13aaeff613af..c70587a4d1137 100644 --- a/Mathlib/Topology/UniformSpace/CompactConvergence.lean +++ b/Mathlib/Topology/UniformSpace/CompactConvergence.lean @@ -252,11 +252,6 @@ theorem tendsto_iff_tendstoLocallyUniformly [WeaklyLocallyCompactSpace α] : obtain ⟨n, hn₁, hn₂⟩ := exists_compact_mem_nhds x exact ⟨n, hn₂, h n hn₁ V hV⟩ -@[deprecated tendsto_iff_tendstoLocallyUniformly (since := "2023-09-03")] -theorem tendstoLocallyUniformly_of_tendsto [WeaklyLocallyCompactSpace α] (h : Tendsto F p (𝓝 f)) : - TendstoLocallyUniformly (fun i a => F i a) f p := - tendsto_iff_tendstoLocallyUniformly.1 h - section Functorial variable {γ δ : Type*} [TopologicalSpace γ] [UniformSpace δ] From 8f520637f08a283ae0e192c0a218a1b4e67ec8ea Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 18 Oct 2024 13:37:51 +0000 Subject: [PATCH 062/131] chore: redefine `Nat.bitCasesOn` `Nat.binaryRec` (#15571) The new definitions of `Nat.bitCasesOn` and `Nat.binaryRec*` have better runtime performance. Also they have lighter weight dependencies now and may be upstreamed to `Batteries` later. --- Mathlib/Data/Nat/Bits.lean | 133 +++++++++++++++++++--------------- Mathlib/Data/Nat/Bitwise.lean | 10 ++- Mathlib/Data/Nat/Size.lean | 1 - 3 files changed, 79 insertions(+), 65 deletions(-) diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index c6e41071fc214..e580011668c94 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -123,10 +123,24 @@ lemma bit_val (b n) : bit b n = 2 * n + cond b 1 0 := by lemma bit_decomp (n : Nat) : bit (bodd n) (div2 n) = n := (bit_val _ _).trans <| (Nat.add_comm _ _).trans <| bodd_add_div2 _ -/-- For a predicate `C : Nat → Sort*`, if instances can be +theorem shiftRight_one (n) : n >>> 1 = n / 2 := rfl + +@[simp] +theorem bit_decide_mod_two_eq_one_shiftRight_one (n : Nat) : bit (n % 2 = 1) (n >>> 1) = n := by + simp only [bit, shiftRight_one] + cases mod_two_eq_zero_or_one n with | _ h => simpa [h] using Nat.div_add_mod n 2 + +theorem bit_testBit_zero_shiftRight_one (n : Nat) : bit (n.testBit 0) (n >>> 1) = n := by + simp + +/-- For a predicate `motive : Nat → Sort*`, if instances can be constructed for natural numbers of the form `bit b n`, they can be constructed for any given natural number. -/ -def bitCasesOn {C : Nat → Sort u} (n) (h : ∀ b n, C (bit b n)) : C n := bit_decomp n ▸ h _ _ +def bitCasesOn {motive : Nat → Sort u} (n) (h : ∀ b n, motive (bit b n)) : motive n := + -- `1 &&& n != 0` is faster than `n.testBit 0`. This may change when we have faster `testBit`. + let x := h (1 &&& n != 0) (n >>> 1) + -- `congrArg motive _ ▸ x` is defeq to `x` in non-dependent case + congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x lemma bit_zero : bit false 0 = 0 := rfl @@ -158,21 +172,16 @@ lemma binaryRec_decreasing (h : n ≠ 0) : div2 n < n := by rwa [Nat.mul_one] at this /-- A recursion principle for `bit` representations of natural numbers. - For a predicate `C : Nat → Sort*`, if instances can be + For a predicate `motive : Nat → Sort*`, if instances can be constructed for natural numbers of the form `bit b n`, they can be constructed for all natural numbers. -/ -def binaryRec {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) : ∀ n, C n := - fun n => - if n0 : n = 0 then by - simp only [n0] - exact z - else by - let n' := div2 n - have _x : bit (bodd n) n' = n := by - apply bit_decomp n - rw [← _x] - exact f (bodd n) n' (binaryRec z f n') - decreasing_by exact binaryRec_decreasing n0 +def binaryRec {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) + (n : Nat) : motive n := + if n0 : n = 0 then congrArg motive n0 ▸ z + else + let x := f (1 &&& n != 0) (n >>> 1) (binaryRec z f (n >>> 1)) + congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x +decreasing_by exact bitwise_rec_lemma n0 /-- `size n` : Returns the size of a natural number in bits i.e. the length of its binary representation -/ @@ -191,7 +200,8 @@ def ldiff : ℕ → ℕ → ℕ := bitwise fun a b => a && not b @[simp] -lemma binaryRec_zero {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) : +lemma binaryRec_zero {motive : Nat → Sort u} + (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) : binaryRec z f 0 = z := by rw [binaryRec] rfl @@ -200,7 +210,8 @@ lemma binaryRec_zero {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit lemma binaryRec_one {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) : binaryRec z f 1 = f true 0 z := by rw [binaryRec] - simp + simp only [succ_ne_self, ↓reduceDIte, reduceShiftRight, binaryRec_zero] + rfl /-! bitwise ops -/ @@ -249,25 +260,6 @@ lemma testBit_bit_succ (m b n) : testBit (bit b n) (succ m) = testBit n m := by simp only [bodd_eq_one_and_ne_zero] at this exact this -lemma binaryRec_eq {C : Nat → Sort u} {z : C 0} {f : ∀ b n, C n → C (bit b n)} - (h : f false 0 z = z) (b n) : binaryRec z f (bit b n) = f b n (binaryRec z f n) := by - rw [binaryRec] - split_ifs with h' - · generalize binaryRec z f (bit b n) = e - revert e - have bf := bodd_bit b n - have n0 := div2_bit b n - rw [h'] at bf n0 - simp only [bodd_zero, div2_zero] at bf n0 - subst bf n0 - rw [binaryRec_zero] - intros - rw [h, eq_mpr_eq_cast, cast_eq] - · simp only; generalize_proofs h - revert h - rw [bodd_bit, div2_bit] - intros; simp only [eq_mpr_eq_cast, cast_eq] - /-! ### `boddDiv2_eq` and `bodd` -/ @@ -296,28 +288,41 @@ theorem bit_ne_zero (b) {n} (h : n ≠ 0) : bit b n ≠ 0 := by cases b <;> dsimp [bit] <;> omega @[simp] -theorem bitCasesOn_bit {C : ℕ → Sort u} (H : ∀ b n, C (bit b n)) (b : Bool) (n : ℕ) : - bitCasesOn (bit b n) H = H b n := - eq_of_heq <| (eqRec_heq _ _).trans <| by rw [bodd_bit, div2_bit] +theorem bit_div_two (b n) : bit b n / 2 = n := by + rw [bit_val, Nat.add_comm, add_mul_div_left, div_eq_of_lt, Nat.zero_add] + · cases b <;> decide + · decide + +@[simp] +theorem bit_shiftRight_one (b n) : bit b n >>> 1 = n := + bit_div_two b n @[simp] -theorem bitCasesOn_bit0 {C : ℕ → Sort u} (H : ∀ b n, C (bit b n)) (n : ℕ) : +theorem bitCasesOn_bit {motive : ℕ → Sort u} (h : ∀ b n, motive (bit b n)) (b : Bool) (n : ℕ) : + bitCasesOn (bit b n) h = h b n := by + change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ h _ _ = h b n + generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e + rw [testBit_bit_zero, bit_shiftRight_one] + intros; rfl + +@[simp] +theorem bitCasesOn_bit0 {motive : ℕ → Sort u} (H : ∀ b n, motive (bit b n)) (n : ℕ) : bitCasesOn (2 * n) H = H false n := bitCasesOn_bit H false n @[simp] -theorem bitCasesOn_bit1 {C : ℕ → Sort u} (H : ∀ b n, C (bit b n)) (n : ℕ) : +theorem bitCasesOn_bit1 {motive : ℕ → Sort u} (H : ∀ b n, motive (bit b n)) (n : ℕ) : bitCasesOn (2 * n + 1) H = H true n := bitCasesOn_bit H true n -theorem bit_cases_on_injective {C : ℕ → Sort u} : - Function.Injective fun H : ∀ b n, C (bit b n) => fun n => bitCasesOn n H := by +theorem bit_cases_on_injective {motive : ℕ → Sort u} : + Function.Injective fun H : ∀ b n, motive (bit b n) => fun n => bitCasesOn n H := by intro H₁ H₂ h ext b n simpa only [bitCasesOn_bit] using congr_fun h (bit b n) @[simp] -theorem bit_cases_on_inj {C : ℕ → Sort u} (H₁ H₂ : ∀ b n, C (bit b n)) : +theorem bit_cases_on_inj {motive : ℕ → Sort u} (H₁ H₂ : ∀ b n, motive (bit b n)) : ((fun n => bitCasesOn n H₁) = fun n => bitCasesOn n H₂) ↔ H₁ = H₂ := bit_cases_on_injective.eq_iff @@ -340,27 +345,34 @@ The same as `binaryRec_eq`, but that one unfortunately requires `f` to be the identity when appending `false` to `0`. Here, we allow you to explicitly say that that case is not happening, i.e. supplying `n = 0 → b = true`. -/ -theorem binaryRec_eq' {C : ℕ → Sort*} {z : C 0} {f : ∀ b n, C n → C (bit b n)} (b n) - (h : f false 0 z = z ∨ (n = 0 → b = true)) : +theorem binaryRec_eq' {motive : ℕ → Sort*} {z : motive 0} {f : ∀ b n, motive n → motive (bit b n)} + (b n) (h : f false 0 z = z ∨ (n = 0 → b = true)) : binaryRec z f (bit b n) = f b n (binaryRec z f n) := by - rw [binaryRec] - split_ifs with h' - · rcases bit_eq_zero_iff.mp h' with ⟨rfl, rfl⟩ - rw [binaryRec_zero] + by_cases h' : bit b n = 0 + case pos => + obtain ⟨rfl, rfl⟩ := bit_eq_zero_iff.mp h' simp only [imp_false, or_false, eq_self_iff_true, not_true, reduceCtorEq] at h + unfold binaryRec exact h.symm - · dsimp only [] - generalize_proofs e - revert e - rw [bodd_bit, div2_bit] - intros - rfl + case neg => + rw [binaryRec, dif_neg h'] + change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ f _ _ _ = _ + generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e + rw [testBit_bit_zero, bit_shiftRight_one] + intros; rfl + +theorem binaryRec_eq {motive : Nat → Sort u} {z : motive 0} + {f : ∀ b n, motive n → motive (bit b n)} + (h : f false 0 z = z) (b n) : + binaryRec z f (bit b n) = f b n (binaryRec z f n) := + binaryRec_eq' b n (.inl h) /-- The same as `binaryRec`, but the induction step can assume that if `n=0`, the bit being appended is `true`-/ @[elab_as_elim] -def binaryRec' {C : ℕ → Sort*} (z : C 0) (f : ∀ b n, (n = 0 → b = true) → C n → C (bit b n)) : - ∀ n, C n := +def binaryRec' {motive : ℕ → Sort*} (z : motive 0) + (f : ∀ b n, (n = 0 → b = true) → motive n → motive (bit b n)) : + ∀ n, motive n := binaryRec z fun b n ih => if h : n = 0 → b = true then f b n h ih else by @@ -370,8 +382,9 @@ def binaryRec' {C : ℕ → Sort*} (z : C 0) (f : ∀ b n, (n = 0 → b = true) /-- The same as `binaryRec`, but special casing both 0 and 1 as base cases -/ @[elab_as_elim] -def binaryRecFromOne {C : ℕ → Sort*} (z₀ : C 0) (z₁ : C 1) (f : ∀ b n, n ≠ 0 → C n → C (bit b n)) : - ∀ n, C n := +def binaryRecFromOne {motive : ℕ → Sort*} (z₀ : motive 0) (z₁ : motive 1) + (f : ∀ b n, n ≠ 0 → motive n → motive (bit b n)) : + ∀ n, motive n := binaryRec' z₀ fun b n h ih => if h' : n = 0 then by rw [h', h h'] diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 98bc466c54c8d..2a45dc564a4b8 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -69,10 +69,12 @@ lemma bitwise_of_ne_zero {n m : Nat} (hn : n ≠ 0) (hm : m ≠ 0) : theorem binaryRec_of_ne_zero {C : Nat → Sort*} (z : C 0) (f : ∀ b n, C n → C (bit b n)) {n} (h : n ≠ 0) : binaryRec z f n = bit_decomp n ▸ f (bodd n) (div2 n) (binaryRec z f (div2 n)) := by - rw [Eq.rec_eq_cast] - rw [binaryRec] - dsimp only - rw [dif_neg h, eq_mpr_eq_cast] + cases n using bitCasesOn with + | h b n => + rw [binaryRec_eq' _ _ (by right; simpa [bit_eq_zero_iff] using h)] + generalize_proofs h; revert h + rw [bodd_bit, div2_bit] + simp @[simp] lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by rfl) (a m b n) : diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index f3f51b3aa4f02..5c7eb08714c3c 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -44,7 +44,6 @@ theorem size_bit {b n} (h : bit b n ≠ 0) : size (bit b n) = succ (size n) := b lhs rw [binaryRec] simp [h] - rw [div2_bit] section From e84469bdcf9629d3ca28d01dd2889a5ee3fb5742 Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 18 Oct 2024 13:37:52 +0000 Subject: [PATCH 063/131] =?UTF-8?q?chore:=20do=20not=20use=20`=C2=B7=20?= =?UTF-8?q?=E2=89=88=20=C2=B7`=20in=20`Quotient.eq`=20(#17594)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit When `Setoid` is not an instance, we obviously want to avoid using `· ≈ ·`. Currently mathlib replicates the API around `Quotient.mk` for `Quotient.mk''`, but this is not needed because `Quotient.mk` doesn't use instance parameters now. After adjusting the API we can stop using `Quotient.mk''`. --- Mathlib/Algebra/BigOperators/Group/Finset.lean | 2 +- Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean | 4 ++-- Mathlib/Data/Multiset/Basic.lean | 2 +- Mathlib/Data/Quot.lean | 5 ++++- Mathlib/Data/Setoid/Basic.lean | 7 +++++-- Mathlib/ModelTheory/DirectLimit.lean | 2 +- 6 files changed, 14 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index cda549919c99c..ed79af0319d28 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1618,7 +1618,7 @@ theorem prod_partition (R : Setoid α) [DecidableRel R.r] : /-- If we can partition a product into subsets that cancel out, then the whole product cancels. -/ @[to_additive "If we can partition a sum into subsets that cancel out, then the whole sum cancels."] theorem prod_cancels_of_partition_cancels (R : Setoid α) [DecidableRel R.r] - (h : ∀ x ∈ s, ∏ a ∈ s.filter fun y => y ≈ x, f a = 1) : ∏ x ∈ s, f x = 1 := by + (h : ∀ x ∈ s, ∏ a ∈ s.filter fun y => R y x, f a = 1) : ∏ x ∈ s, f x = 1 := by rw [prod_partition R, ← Finset.prod_eq_one] intro xbar xbar_in_s obtain ⟨x, x_in_s, rfl⟩ := mem_image.mp xbar_in_s diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean index d0976e216dd81..e4838de2f5dd0 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -1169,7 +1169,7 @@ lemma add_of_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W'.add P Q = W'.dblXYZ P lemma add_smul_of_equiv {P Q : Fin 3 → R} (h : P ≈ Q) {u v : R} (hu : IsUnit u) (hv : IsUnit v) : W'.add (u • P) (v • Q) = u ^ 4 • W'.add P Q := by have smul : P ≈ Q ↔ u • P ≈ v • Q := by - erw [← Quotient.eq, ← Quotient.eq, smul_eq P hu, smul_eq Q hv] + erw [← Quotient.eq_iff_equiv, ← Quotient.eq_iff_equiv, smul_eq P hu, smul_eq Q hv] rfl rw [add_of_equiv <| smul.mp h, dblXYZ_smul, add_of_equiv h] @@ -1185,7 +1185,7 @@ lemma add_of_not_equiv {P Q : Fin 3 → R} (h : ¬P ≈ Q) : W'.add P Q = W'.add lemma add_smul_of_not_equiv {P Q : Fin 3 → R} (h : ¬P ≈ Q) {u v : R} (hu : IsUnit u) (hv : IsUnit v) : W'.add (u • P) (v • Q) = (u * v) ^ 2 • W'.add P Q := by have smul : P ≈ Q ↔ u • P ≈ v • Q := by - erw [← Quotient.eq, ← Quotient.eq, smul_eq P hu, smul_eq Q hv] + erw [← Quotient.eq_iff_equiv, ← Quotient.eq_iff_equiv, smul_eq P hu, smul_eq Q hv] rfl rw [add_of_not_equiv <| h.comp smul.mpr, addXYZ_smul, add_of_not_equiv h] diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 8f15a338625d1..e04ef7fa7a1dc 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -69,7 +69,7 @@ instance [DecidableEq α] (l₁ l₂ : List α) : Decidable (l₁ ≈ l₂) := -- Porting note: `Quotient.recOnSubsingleton₂ s₁ s₂` was in parens which broke elaboration instance decidableEq [DecidableEq α] : DecidableEq (Multiset α) - | s₁, s₂ => Quotient.recOnSubsingleton₂ s₁ s₂ fun _ _ => decidable_of_iff' _ Quotient.eq + | s₁, s₂ => Quotient.recOnSubsingleton₂ s₁ s₂ fun _ _ => decidable_of_iff' _ Quotient.eq_iff_equiv /-- defines a size for a multiset by referring to the size of the underlying list -/ protected diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 8f524efb7730b..a11fa73582072 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -274,9 +274,12 @@ 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⟧ ↔ r x y := ⟨Quotient.exact, Quotient.sound⟩ +theorem Quotient.eq_iff_equiv {r : Setoid α} {x y : α} : Quotient.mk r x = ⟦y⟧ ↔ x ≈ y := + Quotient.eq + theorem Quotient.forall {α : Sort*} {s : Setoid α} {p : Quotient s → Prop} : (∀ a, p a) ↔ ∀ a : α, p ⟦a⟧ := ⟨fun h _ ↦ h _, fun h a ↦ a.ind h⟩ diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index 63241cbb3689d..c34439b36e1e1 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -221,9 +221,12 @@ lemma sInf_equiv {S : Set (Setoid α)} {x y : α} : letI := sInf S x ≈ y ↔ ∀ s ∈ S, s.Rel x y := Iff.rfl +lemma sInf_iff {S : Set (Setoid α)} {x y : α} : + sInf S x y ↔ ∀ s ∈ S, s x y := Iff.rfl + lemma quotient_mk_sInf_eq {S : Set (Setoid α)} {x y : α} : - Quotient.mk (sInf S) x = Quotient.mk (sInf S) y ↔ ∀ s ∈ S, s.Rel x y := by - simp [sInf_equiv] + Quotient.mk (sInf S) x = Quotient.mk (sInf S) y ↔ ∀ s ∈ S, s x y := by + simp [sInf_iff] /-- The map induced between quotients by a setoid inequality. -/ def map_of_le {s t : Setoid α} (h : s ≤ t) : Quotient s → Quotient t := diff --git a/Mathlib/ModelTheory/DirectLimit.lean b/Mathlib/ModelTheory/DirectLimit.lean index 295d49fb6ebf9..bc5cfde3f62f1 100644 --- a/Mathlib/ModelTheory/DirectLimit.lean +++ b/Mathlib/ModelTheory/DirectLimit.lean @@ -349,7 +349,7 @@ def lift (g : ∀ i, G i ↪[L] P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i rw [← Quotient.out_eq x, ← Quotient.out_eq y, Quotient.lift_mk, Quotient.lift_mk] at xy obtain ⟨i, hx, hy⟩ := directed_of (· ≤ ·) x.out.1 y.out.1 rw [← Hg x.out.1 i hx, ← Hg y.out.1 i hy] at xy - rw [← Quotient.out_eq x, ← Quotient.out_eq y, Quotient.eq, equiv_iff G f hx hy] + rw [← Quotient.out_eq x, ← Quotient.out_eq y, Quotient.eq_iff_equiv, equiv_iff G f hx hy] exact (g i).injective xy map_fun' F x := by obtain ⟨i, y, rfl⟩ := exists_quotient_mk'_sigma_mk'_eq G f x From 344c0bb0e582a25235a93bb531130cb0433c6392 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 14:18:26 +0000 Subject: [PATCH 064/131] feat: a subgroup is `1` iff it's trivial (#17906) From LeanCamCombi --- Mathlib/Algebra/Group/Subgroup/Pointwise.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index 95c62df222f42..f42601fa4cc09 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -41,6 +41,10 @@ lemma smul_coe_set [Group G] [SetLike S G] [SubgroupClass S G] {s : S} {a : G} ( a • (s : Set G) = s := by ext; simp [Set.mem_smul_set_iff_inv_smul_mem, mul_mem_cancel_left, ha] +@[norm_cast, to_additive] +lemma coe_set_eq_one [Group G] {s : Subgroup G} : (s : Set G) = 1 ↔ s = ⊥ := + (SetLike.ext'_iff.trans (by rfl)).symm + @[to_additive (attr := simp)] lemma op_smul_coe_set [Group G] [SetLike S G] [SubgroupClass S G] {s : S} {a : G} (ha : a ∈ s) : MulOpposite.op a • (s : Set G) = s := by From 2a3ae36b046fab48ee07482c6b66d7d5f1e6eff3 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Fri, 18 Oct 2024 14:51:42 +0000 Subject: [PATCH 065/131] feat (RingTheory/HahnSeries/Summable) : transport summable families along Equiv (#16872) This PR defines a new summable family given a summable family of Hahn series and an equivalence of parametrizing types. We show that the new family has the same formal sum as the old one. --- Mathlib/RingTheory/HahnSeries/Summable.lean | 25 ++++++++++++++++----- 1 file changed, 20 insertions(+), 5 deletions(-) diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index 4405c005d34e1..45a21761c48c3 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -42,7 +42,7 @@ open Pointwise noncomputable section -variable {Γ : Type*} {R : Type*} +variable {Γ Γ' R V α β : Type*} namespace HahnSeries @@ -61,13 +61,11 @@ theorem isPWO_iUnion_support_powers [LinearOrderedCancelAddCommMonoid Γ] [Ring section -variable (Γ) (R) [PartialOrder Γ] [AddCommMonoid R] - /-- A family of Hahn series whose formal coefficient-wise sum is a Hahn series. For each coefficient of the sum to be well-defined, we require that only finitely many series are nonzero at any given coefficient. For the formal sum to be a Hahn series, we require that the union of the supports of the constituent series is partially well-ordered. -/ -structure SummableFamily (α : Type*) where +structure SummableFamily (Γ) (R) [PartialOrder Γ] [AddCommMonoid R] (α : Type*) where /-- A parametrized family of Hahn series. -/ toFun : α → HahnSeries Γ R isPWO_iUnion_support' : Set.IsPWO (⋃ a : α, (toFun a).support) @@ -79,7 +77,7 @@ namespace SummableFamily section AddCommMonoid -variable [PartialOrder Γ] [AddCommMonoid R] {α : Type*} +variable [PartialOrder Γ] [AddCommMonoid R] instance : FunLike (SummableFamily Γ R α) α (HahnSeries Γ R) where coe := toFun @@ -180,6 +178,23 @@ theorem hsum_add {s t : SummableFamily Γ R α} : (s + t).hsum = s.hsum + t.hsum simp only [hsum_coeff, add_coeff, add_apply] exact finsum_add_distrib (s.finite_co_support _) (t.finite_co_support _) +/-- A summable family induced by an equivalence of the parametrizing type. -/ +@[simps] +def Equiv (e : α ≃ β) (s : SummableFamily Γ R α) : SummableFamily Γ R β where + toFun b := s (e.symm b) + isPWO_iUnion_support' := by + refine Set.IsPWO.mono s.isPWO_iUnion_support fun g => ?_ + simp only [Set.mem_iUnion, mem_support, ne_eq, forall_exists_index] + exact fun b hg => Exists.intro (e.symm b) hg + finite_co_support' g := + (Equiv.set_finite_iff e.subtypeEquivOfSubtype').mp <| s.finite_co_support' g + +@[simp] +theorem hsum_equiv (e : α ≃ β) (s : SummableFamily Γ R α) : (Equiv e s).hsum = s.hsum := by + ext g + simp only [hsum_coeff, Equiv_toFun] + exact finsum_eq_of_bijective e.symm (Equiv.bijective e.symm) fun x => rfl + end AddCommMonoid section AddCommGroup From c289ac29f16b5d6733857a95139dc57fa8ebc646 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Fri, 18 Oct 2024 14:51:43 +0000 Subject: [PATCH 066/131] feat(RingTheory/Algebraic): add some results on `Transcendental` (#17867) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit - `transcendental_iff`, `transcendental_iff_ker_eq_bot`: similar to the API of `AlgebraicIndependent` - `IsAlgebraic.of_aeval`, `Transcendental.aeval`: behavior of transcendental over `aeval` - `Polynomial.transcendental`, `Polynomial.transcendental_X`: some specific polynomials are transcendental over its base ring - `Transcendental.linearIndependent_sub_inv`: if `E / F` is a field extension, `x` is an element of `E` transcendental over `F`, then `{(x - a)⁻¹ | a : F}` is linearly independent over `F` - `transcendental_algebraMap_iff`: negation of `isAlgebraic_algebraMap_iff` --- Mathlib/RingTheory/Algebraic.lean | 69 +++++++++++++++++++++++++++++++ 1 file changed, 69 insertions(+) diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index 8685279579238..9942e6c347371 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -40,6 +40,63 @@ theorem is_transcendental_of_subsingleton [Subsingleton R] (x : A) : Transcenden variable {R} +/-- An element `x` is transcendental over `R` if and only if for any polynomial `p`, +`Polynomial.aeval x p = 0` implies `p = 0`. This is similar to `algebraicIndependent_iff`. -/ +theorem transcendental_iff {x : A} : + Transcendental R x ↔ ∀ p : R[X], aeval x p = 0 → p = 0 := by + rw [Transcendental, IsAlgebraic, not_exists] + congr! 1; tauto + +variable (R) in +theorem Polynomial.transcendental_X : Transcendental R (X (R := R)) := by + simp [transcendental_iff] + +theorem IsAlgebraic.of_aeval {r : A} (f : R[X]) (hf : f.natDegree ≠ 0) + (hf' : f.leadingCoeff ∈ nonZeroDivisors R) (H : IsAlgebraic R (aeval r f)) : + IsAlgebraic R r := by + obtain ⟨p, h1, h2⟩ := H + have : (p.comp f).coeff (p.natDegree * f.natDegree) ≠ 0 := fun h ↦ h1 <| by + rwa [coeff_comp_degree_mul_degree hf, + mul_right_mem_nonZeroDivisors_eq_zero_iff (pow_mem hf' _), + leadingCoeff_eq_zero] at h + exact ⟨p.comp f, fun h ↦ this (by simp [h]), by rwa [aeval_comp]⟩ + +theorem Transcendental.aeval {r : A} (H : Transcendental R r) (f : R[X]) (hf : f.natDegree ≠ 0) + (hf' : f.leadingCoeff ∈ nonZeroDivisors R) : + Transcendental R (aeval r f) := fun h ↦ H (h.of_aeval f hf hf') + +theorem Polynomial.transcendental (f : R[X]) (hf : f.natDegree ≠ 0) + (hf' : f.leadingCoeff ∈ nonZeroDivisors R) : + Transcendental R f := by + simpa using (transcendental_X R).aeval f hf hf' + +/-- If `E / F` is a field extension, `x` is an element of `E` transcendental over `F`, +then `{(x - a)⁻¹ | a : F}` is linearly independent over `F`. -/ +theorem Transcendental.linearIndependent_sub_inv + {F E : Type*} [Field F] [Field E] [Algebra F E] {x : E} (H : Transcendental F x) : + LinearIndependent F fun a ↦ (x - algebraMap F E a)⁻¹ := by + rw [transcendental_iff] at H + refine linearIndependent_iff'.2 fun s m hm i hi ↦ ?_ + have hnz (a : F) : x - algebraMap F E a ≠ 0 := fun h ↦ + X_sub_C_ne_zero a <| H (.X - .C a) (by simp [h]) + let b := s.prod fun j ↦ x - algebraMap F E j + have h1 : ∀ i ∈ s, m i • (b * (x - algebraMap F E i)⁻¹) = + m i • (s.erase i).prod fun j ↦ x - algebraMap F E j := fun i hi ↦ by + simp_rw [b, ← s.prod_erase_mul _ hi, mul_inv_cancel_right₀ (hnz i)] + replace hm := congr(b * $(hm)) + simp_rw [mul_zero, Finset.mul_sum, mul_smul_comm, Finset.sum_congr rfl h1] at hm + let p : Polynomial F := s.sum fun i ↦ .C (m i) * (s.erase i).prod fun j ↦ .X - .C j + replace hm := congr(Polynomial.aeval i $(H p (by simp_rw [← hm, p, map_sum, map_mul, map_prod, + map_sub, aeval_X, aeval_C, Algebra.smul_def]))) + have h2 : ∀ j ∈ s.erase i, m j * ((s.erase j).prod fun x ↦ i - x) = 0 := fun j hj ↦ by + have := Finset.mem_erase_of_ne_of_mem (Finset.ne_of_mem_erase hj).symm hi + simp_rw [← (s.erase j).prod_erase_mul _ this, sub_self, mul_zero] + simp_rw [map_zero, p, map_sum, map_mul, map_prod, map_sub, aeval_X, + aeval_C, Algebra.id.map_eq_self, ← s.sum_erase_add _ hi, + Finset.sum_eq_zero h2, zero_add] at hm + exact eq_zero_of_ne_zero_of_mul_right_eq_zero (Finset.prod_ne_zero_iff.2 fun j hj ↦ + sub_ne_zero.2 (Finset.ne_of_mem_erase hj).symm) hm + /-- A subalgebra is algebraic if all its elements are algebraic. -/ nonrec def Subalgebra.IsAlgebraic (S : Subalgebra R A) : Prop := @@ -86,10 +143,18 @@ theorem isAlgebraic_iff_not_injective {x : A} : IsAlgebraic R x ↔ ¬Function.Injective (Polynomial.aeval x : R[X] →ₐ[R] A) := by simp only [IsAlgebraic, injective_iff_map_eq_zero, not_forall, and_comm, exists_prop] +/-- An element `x` is transcendental over `R` if and only if the map `Polynomial.aeval x` +is injective. This is similar to `algebraicIndependent_iff_injective_aeval`. -/ theorem transcendental_iff_injective {x : A} : Transcendental R x ↔ Function.Injective (Polynomial.aeval x : R[X] →ₐ[R] A) := isAlgebraic_iff_not_injective.not_left +/-- An element `x` is transcendental over `R` if and only if the kernel of the ring homomorphism +`Polynomial.aeval x` is the zero ideal. This is similar to `algebraicIndependent_iff_ker_eq_bot`. -/ +theorem transcendental_iff_ker_eq_bot {x : A} : + Transcendental R x ↔ RingHom.ker (aeval (R := R) x) = ⊥ := by + rw [transcendental_iff_injective, RingHom.injective_iff_ker_eq_bot] + end section zero_ne_one @@ -174,6 +239,10 @@ theorem isAlgebraic_algebraMap_iff {a : S} (h : Function.Injective (algebraMap S IsAlgebraic R (algebraMap S A a) ↔ IsAlgebraic R a := isAlgebraic_algHom_iff (IsScalarTower.toAlgHom R S A) h +theorem transcendental_algebraMap_iff {a : S} (h : Function.Injective (algebraMap S A)) : + Transcendental R (algebraMap S A a) ↔ Transcendental R a := by + simp_rw [Transcendental, isAlgebraic_algebraMap_iff h] + theorem IsAlgebraic.of_pow {r : A} {n : ℕ} (hn : 0 < n) (ht : IsAlgebraic R (r ^ n)) : IsAlgebraic R r := by obtain ⟨p, p_nonzero, hp⟩ := ht From 66fe0b7279d0edf044d905b5a4068aa929b78e7b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 14:51:44 +0000 Subject: [PATCH 067/131] doc(Simps): deduplicate clause (#17916) Both clauses were different before #8296 --- Mathlib/Tactic/Simps/Basic.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/Tactic/Simps/Basic.lean b/Mathlib/Tactic/Simps/Basic.lean index 1caba0e4825ab..3b28a506c4d61 100644 --- a/Mathlib/Tactic/Simps/Basic.lean +++ b/Mathlib/Tactic/Simps/Basic.lean @@ -317,8 +317,7 @@ Some common uses: This will generate `foo_apply` lemmas for each declaration `foo`. * If you prefer `coe_foo` lemmas that state equalities between functions, use `initialize_simps_projections MulHom (toFun → coe, as_prefix coe)` - In this case you have to use `@[simps (config := .asFn)]` or equivalently - `@[simps (config := .asFn)]` whenever you call `@[simps]`. + In this case you have to use `@[simps (config := .asFn)]` whenever you call `@[simps]`. * You can also initialize to use both, in which case you have to choose which one to use by default, by using either of the following ``` From 9e3fa1e4f5106b00d9b078c99479dae53c6adbb7 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Fri, 18 Oct 2024 15:39:36 +0000 Subject: [PATCH 068/131] feat(FieldTheory/Adjoin): add some results of `extendScalars` (#15238) - `extendScalars_{self|top|inf|sup}` for intermediate fields and subfields - also add `restrictScalars_{inf|sup}` and `bot_toSubfield` - add a missing helper lemma `toSubfield_injective` --- Mathlib/FieldTheory/Adjoin.lean | 82 ++++++++++++++++++- .../FieldTheory/IntermediateField/Basic.lean | 15 ++-- 2 files changed, 89 insertions(+), 8 deletions(-) diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index aa48dc241889b..f5504ce434121 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -114,6 +114,9 @@ theorem mem_bot {x : E} : x ∈ (⊥ : IntermediateField F E) ↔ x ∈ Set.rang @[simp] theorem bot_toSubalgebra : (⊥ : IntermediateField F E).toSubalgebra = ⊥ := rfl +theorem bot_toSubfield : (⊥ : IntermediateField F E).toSubfield = (algebraMap F E).fieldRange := + rfl + @[simp] theorem coe_top : ↑(⊤ : IntermediateField F E) = (Set.univ : Set E) := rfl @@ -267,16 +270,31 @@ def topEquiv : (⊤ : IntermediateField F E) ≃ₐ[F] E := -- Porting note: this theorem is now generated by the `@[simps!]` above. +section RestrictScalars + @[simp] theorem restrictScalars_bot_eq_self (K : IntermediateField F E) : (⊥ : IntermediateField K E).restrictScalars _ = K := SetLike.coe_injective Subtype.range_coe +variable {K : Type*} [Field K] [Algebra K E] [Algebra K F] [IsScalarTower K F E] + @[simp] -theorem restrictScalars_top {K : Type*} [Field K] [Algebra K E] [Algebra K F] - [IsScalarTower K F E] : (⊤ : IntermediateField F E).restrictScalars K = ⊤ := +theorem restrictScalars_top : (⊤ : IntermediateField F E).restrictScalars K = ⊤ := rfl +variable (K) +variable (L L' : IntermediateField F E) + +theorem restrictScalars_sup : + L.restrictScalars K ⊔ L'.restrictScalars K = (L ⊔ L').restrictScalars K := + toSubfield_injective (by simp) + +theorem restrictScalars_inf : + L.restrictScalars K ⊓ L'.restrictScalars K = (L ⊓ L').restrictScalars K := rfl + +end RestrictScalars + variable {K : Type*} [Field K] [Algebra F K] @[simp] @@ -1466,3 +1484,63 @@ theorem comap_map (f : L →ₐ[K] L') (S : IntermediateField K L) : (S.map f).c SetLike.coe_injective (Set.preimage_image_eq _ f.injective) end IntermediateField + +section ExtendScalars + +variable {K : Type*} [Field K] {L : Type*} [Field L] [Algebra K L] + +namespace Subfield + +variable (F : Subfield L) + +@[simp] +theorem extendScalars_self : extendScalars (le_refl F) = ⊥ := by + ext x + rw [mem_extendScalars, IntermediateField.mem_bot] + refine ⟨fun h ↦ ⟨⟨x, h⟩, rfl⟩, ?_⟩ + rintro ⟨y, rfl⟩ + exact y.2 + +@[simp] +theorem extendScalars_top : extendScalars (le_top : F ≤ ⊤) = ⊤ := + IntermediateField.toSubfield_injective (by simp) + +variable {F} +variable {E E' : Subfield L} (h : F ≤ E) (h' : F ≤ E') + +theorem extendScalars_sup : + extendScalars h ⊔ extendScalars h' = extendScalars (le_sup_of_le_left h : F ≤ E ⊔ E') := + ((extendScalars.orderIso F).map_sup ⟨_, h⟩ ⟨_, h'⟩).symm + +theorem extendScalars_inf : extendScalars h ⊓ extendScalars h' = extendScalars (le_inf h h') := + ((extendScalars.orderIso F).map_inf ⟨_, h⟩ ⟨_, h'⟩).symm + +end Subfield + +namespace IntermediateField + +variable (F : IntermediateField K L) + +@[simp] +theorem extendScalars_self : extendScalars (le_refl F) = ⊥ := + restrictScalars_injective K (by simp) + +@[simp] +theorem extendScalars_top : extendScalars (le_top : F ≤ ⊤) = ⊤ := + restrictScalars_injective K (by simp) + +variable {F} +variable {E E' : IntermediateField K L} (h : F ≤ E) (h' : F ≤ E') + +theorem extendScalars_sup : + extendScalars h ⊔ extendScalars h' = extendScalars (le_sup_of_le_left h : F ≤ E ⊔ E') := + ((extendScalars.orderIso F).map_sup ⟨_, h⟩ ⟨_, h'⟩).symm + +theorem extendScalars_inf : extendScalars h ⊓ extendScalars h' = extendScalars (le_inf h h') := + ((extendScalars.orderIso F).map_inf ⟨_, h⟩ ⟨_, h'⟩).symm + +end IntermediateField + +end ExtendScalars + +set_option linter.style.longFile 1700 diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index 1c807c5364473..54dbf4d6b374d 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -494,14 +494,17 @@ theorem coe_inclusion {E F : IntermediateField K L} (hEF : E ≤ F) (e : E) : variable {S} -theorem toSubalgebra_injective : - Function.Injective (IntermediateField.toSubalgebra : IntermediateField K L → _) := by - intro S S' h +theorem toSubalgebra_injective : Function.Injective (toSubalgebra : IntermediateField K L → _) := by + intro _ _ h + ext + simp_rw [← mem_toSubalgebra, h] + +theorem toSubfield_injective : Function.Injective (toSubfield : IntermediateField K L → _) := by + intro _ _ h ext - rw [← mem_toSubalgebra, ← mem_toSubalgebra, h] + simp_rw [← mem_toSubfield, h] -theorem map_injective (f : L →ₐ[K] L') : - Function.Injective (map f) := by +theorem map_injective (f : L →ₐ[K] L') : Function.Injective (map f) := by intro _ _ h rwa [← toSubalgebra_injective.eq_iff, toSubalgebra_map, toSubalgebra_map, (Subalgebra.map_injective f.injective).eq_iff, toSubalgebra_injective.eq_iff] at h From c5b97e2f7eecede13acd99994d4be2239dbd2ea3 Mon Sep 17 00:00:00 2001 From: Yongle Hu Date: Fri, 18 Oct 2024 15:39:37 +0000 Subject: [PATCH 069/131] feat(NumberTheory/NumberField/Basic): lemmas for extension of ring of integers (#17784) Add some lemmas and `instance`s for extension of number fields or ring of integers. Co-authored-by: Yongle Hu @mbkybky Co-authored-by: Jiedong Jiang @jjdishere Co-authored-by: Hu Yongle <2065545849@qq.com> --- Mathlib/NumberTheory/NumberField/Basic.lean | 64 ++++++++++++++++++++- 1 file changed, 63 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index cd142c881c370..64c8f95606e20 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -59,6 +59,23 @@ protected theorem isAlgebraic [NumberField K] : Algebra.IsAlgebraic ℚ K := instance [NumberField K] [NumberField L] [Algebra K L] : FiniteDimensional K L := Module.Finite.of_restrictScalars_finite ℚ K L +/-- A finite extension of a number field is a number field. -/ +theorem of_module_finite [NumberField K] [Algebra K L] [Module.Finite K L] : NumberField L where + to_charZero := charZero_of_injective_algebraMap (algebraMap K L).injective + to_finiteDimensional := + letI := charZero_of_injective_algebraMap (algebraMap K L).injective + Module.Finite.trans K L + +variable {K} {L} in +instance of_intermediateField [NumberField K] [NumberField L] [Algebra K L] + (E : IntermediateField K L) : NumberField E := + of_module_finite K E + +theorem of_tower [NumberField K] [NumberField L] [Algebra K L] (E : Type*) [Field E] + [Algebra K E] [Algebra E L] [IsScalarTower K E L] : NumberField E := + letI := Module.Finite.left K E L + of_module_finite K E + /-- The ring of integers (or number ring) corresponding to a number field is the integral closure of ℤ in the number field. @@ -89,7 +106,7 @@ instance : Nontrivial (𝓞 K) := inferInstanceAs (Nontrivial (integralClosure _ _)) instance {L : Type*} [Ring L] [Algebra K L] : Algebra (𝓞 K) L := inferInstanceAs (Algebra (integralClosure _ _) L) -instance {L : Type*} [Ring L] [Algebra K L] : IsScalarTower (𝓞 K) K L := +instance {L : Type*} [Ring L] [Algebra K L] : IsScalarTower (𝓞 K) K L := inferInstanceAs (IsScalarTower (integralClosure _ _) K L) variable {K} @@ -222,6 +239,9 @@ instance [NumberField K] : IsFractionRing (𝓞 K) K := instance : IsIntegralClosure (𝓞 K) ℤ K := integralClosure.isIntegralClosure _ _ +instance : Algebra.IsIntegral ℤ (𝓞 K) := + IsIntegralClosure.isIntegral_algebra ℤ K + instance [NumberField K] : IsIntegrallyClosed (𝓞 K) := integralClosure.isIntegrallyClosedOfFiniteExtension ℚ @@ -283,6 +303,48 @@ def restrict_monoidHom [MulOneClass M] (f : M →* K) (h : ∀ x, IsIntegral ℤ map_one' := by simp only [restrict, map_one, mk_one] map_mul' x y := by simp only [restrict, map_mul, mk_mul_mk _] +section extension + +variable (K L : Type*) [Field K] [Field L] [Algebra K L] + +instance : IsScalarTower (𝓞 K) (𝓞 L) L := + IsScalarTower.of_algebraMap_eq' rfl + +instance : IsIntegralClosure (𝓞 L) (𝓞 K) L := + IsIntegralClosure.tower_top (R := ℤ) + +/-- The ring of integers of `L` is isomorphic to any integral closure of `𝓞 K` in `L` -/ +protected noncomputable def algEquiv (R : Type*) [CommRing R] [Algebra (𝓞 K) R] [Algebra R L] + [IsScalarTower (𝓞 K) R L] [IsIntegralClosure R (𝓞 K) L] : 𝓞 L ≃ₐ[𝓞 K] R := + (IsIntegralClosure.equiv (𝓞 K) R L _).symm + +/-- Any extension between ring of integers is integral. -/ +instance extension_algebra_isIntegral : Algebra.IsIntegral (𝓞 K) (𝓞 L) := + IsIntegralClosure.isIntegral_algebra (𝓞 K) L + +/-- Any extension between ring of integers of number fields is noetherian. -/ +instance extension_isNoetherian [NumberField K] [NumberField L] : IsNoetherian (𝓞 K) (𝓞 L) := + IsIntegralClosure.isNoetherian (𝓞 K) K L (𝓞 L) + +/-- The kernel of the algebraMap between ring of integers is `⊥`. -/ +theorem ker_algebraMap_eq_bot : RingHom.ker (algebraMap (𝓞 K) (𝓞 L)) = ⊥ := + (RingHom.ker_eq_bot_iff_eq_zero (algebraMap (𝓞 K) (𝓞 L))).mpr <| fun x hx => by + have h : (algebraMap K L) x = (algebraMap (𝓞 K) (𝓞 L)) x := rfl + simp only [hx, map_zero, map_eq_zero, RingOfIntegers.coe_eq_zero_iff] at h + exact h + +/-- The algebraMap between ring of integers is injective. -/ +theorem algebraMap.injective : Function.Injective (algebraMap (𝓞 K) (𝓞 L)) := + (RingHom.injective_iff_ker_eq_bot (algebraMap (𝓞 K) (𝓞 L))).mpr (ker_algebraMap_eq_bot K L) + +instance : NoZeroSMulDivisors (𝓞 K) (𝓞 L) := + NoZeroSMulDivisors.of_algebraMap_injective (algebraMap.injective K L) + +instance : NoZeroSMulDivisors (𝓞 K) L := + NoZeroSMulDivisors.trans (𝓞 K) (𝓞 L) L + +end extension + end RingOfIntegers variable [NumberField K] From 8b31f1fbde480a5c530264eeea00d051d4c586c5 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Fri, 18 Oct 2024 15:48:17 +0000 Subject: [PATCH 070/131] feat(GroupTheory/ArchimedeanDensely): discrete iff not densely ordered (#16618) Stronger statements than previously proven `Or`s --- Mathlib/GroupTheory/ArchimedeanDensely.lean | 108 ++++++++++++++++---- 1 file changed, 86 insertions(+), 22 deletions(-) diff --git a/Mathlib/GroupTheory/ArchimedeanDensely.lean b/Mathlib/GroupTheory/ArchimedeanDensely.lean index 8b00a52cd33de..cf344500028b4 100644 --- a/Mathlib/GroupTheory/ArchimedeanDensely.lean +++ b/Mathlib/GroupTheory/ArchimedeanDensely.lean @@ -191,6 +191,21 @@ lemma LinearOrderedAddCommGroup.discrete_or_denselyOrdered (G : Type*) · simp [hz.left] · simpa [lt_sub_iff_add_lt'] using hz.right +/-- Any linearly ordered archimedean additive group is either isomorphic (and order-isomorphic) +to the integers, or is densely ordered, exclusively. -/ +lemma LinearOrderedAddCommGroup.discrete_iff_not_denselyOrdered (G : Type*) + [LinearOrderedAddCommGroup G] [Archimedean G] : + Nonempty (G ≃+o ℤ) ↔ ¬ DenselyOrdered G := by + suffices ∀ (_ : G ≃+o ℤ), ¬ DenselyOrdered G by + rcases LinearOrderedAddCommGroup.discrete_or_denselyOrdered G with ⟨⟨h⟩⟩|h + · simpa [this h] using ⟨h⟩ + · simp only [h, not_true_eq_false, iff_false, not_nonempty_iff] + exact ⟨fun H ↦ (this H) h⟩ + intro e H + rw [denselyOrdered_iff_of_orderIsoClass e] at H + obtain ⟨_, _⟩ := exists_between (one_pos (α := ℤ)) + linarith + variable (G) in /-- Any linearly ordered mul-archimedean group is either isomorphic (and order-isomorphic) to the multiplicative integers, or is densely ordered. -/ @@ -201,31 +216,28 @@ lemma LinearOrderedCommGroup.discrete_or_denselyOrdered : rintro ⟨f, hf⟩ exact ⟨AddEquiv.toMultiplicative' f, hf⟩ -/-- Any nontrivial (has other than 0 and 1) linearly ordered mul-archimedean group with zero is -either isomorphic (and order-isomorphic) to `ℤₘ₀`, or is densely ordered. -/ -lemma LinearOrderedCommGroupWithZero.discrete_or_denselyOrdered (G : Type*) - [LinearOrderedCommGroupWithZero G] [Nontrivial Gˣ] [MulArchimedean G] : - Nonempty (G ≃*o ℤₘ₀) ∨ DenselyOrdered G := by - classical - refine (LinearOrderedCommGroup.discrete_or_denselyOrdered Gˣ).imp ?_ ?_ - · intro ⟨f⟩ - refine ⟨OrderMonoidIso.trans - ⟨WithZero.withZeroUnitsEquiv.symm, ?_⟩ ⟨f.withZero, ?_⟩⟩ - · intro - simp only [WithZero.withZeroUnitsEquiv, MulEquiv.symm_mk, - MulEquiv.toEquiv_eq_coe, Equiv.toFun_as_coe, EquivLike.coe_coe, MulEquiv.coe_mk, - Equiv.coe_fn_symm_mk ] - split_ifs <;> - simp_all [← Units.val_le_val] - · intro a b - induction a <;> induction b <;> - simp [MulEquiv.withZero] +variable (G) in +/-- Any linearly ordered mul-archimedean group is either isomorphic (and order-isomorphic) +to the multiplicative integers, or is densely ordered, exclusively. -/ +@[to_additive existing] +lemma LinearOrderedCommGroup.discrete_iff_not_denselyOrdered : + Nonempty (G ≃*o Multiplicative ℤ) ↔ ¬ DenselyOrdered G := by + let e : G ≃o Additive G := OrderIso.refl G + rw [denselyOrdered_iff_of_orderIsoClass e, + ← LinearOrderedAddCommGroup.discrete_iff_not_denselyOrdered (Additive G)] + refine Nonempty.congr ?_ ?_ <;> intro f + · exact ⟨MulEquiv.toAdditive' f, by simp⟩ + · exact ⟨MulEquiv.toAdditive'.symm f, by simp⟩ + +lemma denselyOrdered_units_iff {G₀ : Type*} [LinearOrderedCommGroupWithZero G₀] [Nontrivial G₀ˣ] : + DenselyOrdered G₀ˣ ↔ DenselyOrdered G₀ := by + constructor · intro H refine ⟨fun x y h ↦ ?_⟩ rcases (zero_le' (a := x)).eq_or_lt with rfl|hx - · lift y to Gˣ using h.ne'.isUnit - obtain ⟨z, hz⟩ := exists_ne (1 : Gˣ) - refine ⟨(y * |z|ₘ⁻¹ : Gˣ), ?_, ?_⟩ + · lift y to G₀ˣ using h.ne'.isUnit + obtain ⟨z, hz⟩ := exists_ne (1 : G₀ˣ) + refine ⟨(y * |z|ₘ⁻¹ : G₀ˣ), ?_, ?_⟩ · simp [zero_lt_iff] · rw [Units.val_lt_val] simp [hz] @@ -233,3 +245,55 @@ lemma LinearOrderedCommGroupWithZero.discrete_or_denselyOrdered (G : Type*) (by simp [← Units.val_lt_val, h]) refine ⟨z, ?_, ?_⟩ <;> simpa [← Units.val_lt_val] + · intro H + refine ⟨fun x y h ↦ ?_⟩ + obtain ⟨z, hz⟩ := exists_between (Units.val_lt_val.mpr h) + rcases (zero_le' (a := z)).eq_or_lt with rfl|hz' + · simp at hz + refine ⟨Units.mk0 z hz'.ne', ?_⟩ + simp [← Units.val_lt_val, hz] + +/-- Any nontrivial (has other than 0 and 1) linearly ordered mul-archimedean group with zero is +either isomorphic (and order-isomorphic) to `ℤₘ₀`, or is densely ordered. -/ +lemma LinearOrderedCommGroupWithZero.discrete_or_denselyOrdered (G : Type*) + [LinearOrderedCommGroupWithZero G] [Nontrivial Gˣ] [MulArchimedean G] : + Nonempty (G ≃*o WithZero (Multiplicative ℤ)) ∨ DenselyOrdered G := by + classical + rw [← denselyOrdered_units_iff] + refine (LinearOrderedCommGroup.discrete_or_denselyOrdered Gˣ).imp_left ?_ + intro ⟨f⟩ + refine ⟨OrderMonoidIso.trans + ⟨WithZero.withZeroUnitsEquiv.symm, ?_⟩ ⟨f.withZero, ?_⟩⟩ + · intro + simp only [WithZero.withZeroUnitsEquiv, MulEquiv.symm_mk, + MulEquiv.toEquiv_eq_coe, Equiv.toFun_as_coe, EquivLike.coe_coe, MulEquiv.coe_mk, + Equiv.coe_fn_symm_mk ] + split_ifs <;> + simp_all [← Units.val_le_val] + · intro a b + induction a <;> induction b <;> + simp [MulEquiv.withZero] + +open WithZero in +/-- Any nontrivial (has other than 0 and 1) linearly ordered mul-archimedean group with zero is +either isomorphic (and order-isomorphic) to `ℤₘ₀`, or is densely ordered, exclusively -/ +lemma LinearOrderedCommGroupWithZero.discrete_iff_not_denselyOrdered (G : Type*) + [LinearOrderedCommGroupWithZero G] [Nontrivial Gˣ] [MulArchimedean G] : + Nonempty (G ≃*o WithZero (Multiplicative ℤ)) ↔ ¬ DenselyOrdered G := by + rw [← denselyOrdered_units_iff, + ← LinearOrderedCommGroup.discrete_iff_not_denselyOrdered] + refine Nonempty.congr ?_ ?_ <;> intro f + · refine ⟨MulEquiv.unzero (withZeroUnitsEquiv.trans f), ?_⟩ + intros + simp only [MulEquiv.unzero, withZeroUnitsEquiv, MulEquiv.trans_apply, + MulEquiv.coe_mk, Equiv.coe_fn_mk, recZeroCoe_coe, OrderMonoidIso.coe_mulEquiv, + MulEquiv.symm_trans_apply, MulEquiv.symm_mk, Equiv.coe_fn_symm_mk, map_eq_zero, coe_ne_zero, + ↓reduceDIte, unzero_coe, MulEquiv.toEquiv_eq_coe, Equiv.toFun_as_coe, EquivLike.coe_coe] + rw [← Units.val_le_val, ← map_le_map_iff f, ← coe_le_coe, coe_unzero, coe_unzero] + · refine ⟨withZeroUnitsEquiv.symm.trans (MulEquiv.withZero f), ?_⟩ + intros + simp only [withZeroUnitsEquiv, MulEquiv.symm_mk, MulEquiv.withZero, + MulEquiv.toMonoidHom_eq_coe, MulEquiv.toEquiv_eq_coe, Equiv.toFun_as_coe, EquivLike.coe_coe, + MulEquiv.trans_apply, MulEquiv.coe_mk, Equiv.coe_fn_symm_mk, Equiv.coe_fn_mk] + split_ifs <;> + simp_all [← Units.val_le_val] From 94bb3a8e89c75b20472ec5a37df0f225cb08e003 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Fri, 18 Oct 2024 15:58:51 +0000 Subject: [PATCH 071/131] feat(CategoryTheory/Monoidal): add `Mon_Class` type class (#17185) [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Algebras.20in.20monoidal.20categories.2C.20bundled.20or.20semi-bundled.3F) --- Mathlib/CategoryTheory/Monoidal/Bimon_.lean | 41 +++++++++++- Mathlib/CategoryTheory/Monoidal/Comon_.lean | 67 ++++++++++++++++++- Mathlib/CategoryTheory/Monoidal/Hopf_.lean | 35 +++++++++- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 72 ++++++++++++++++++++- 4 files changed, 208 insertions(+), 7 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean index 109f528c80123..e195b65504704 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean @@ -29,7 +29,46 @@ universe v₁ v₂ u₁ u₂ u open CategoryTheory MonoidalCategory -variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory.{v₁} C] [BraidedCategory C] +variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] [BraidedCategory C] + +open scoped Mon_Class Comon_Class + +/-- +A bimonoid object in a braided category `C` is a object that is simultaneously monoid and comonoid +objects, and structure morphisms of them satisfy appropriate consistency conditions. +-/ +class Bimon_Class (M : C) extends Mon_Class M, Comon_Class M where + /- For the names of the conditions below, the unprimed names are reserved for the version where + the argument `M` is explicit. -/ + mul_comul' : μ[M] ≫ Δ[M] = (Δ[M] ⊗ Δ[M]) ≫ tensor_μ M M M M ≫ (μ[M] ⊗ μ[M]) := by aesop_cat + one_comul' : η[M] ≫ Δ[M] = η[M ⊗ M] := by aesop_cat + mul_counit' : μ[M] ≫ ε[M] = ε[M ⊗ M] := by aesop_cat + one_counit' : η[M] ≫ ε[M] = 𝟙 (𝟙_ C) := by aesop_cat + +namespace Bimon_Class + +/- The simp attribute is reserved for the unprimed versions. -/ +attribute [reassoc] mul_comul' one_comul' mul_counit' one_counit' + +variable (M : C) [Bimon_Class M] + +@[reassoc (attr := simp)] +theorem mul_comul (M : C) [Bimon_Class M] : + μ[M] ≫ Δ[M] = (Δ[M] ⊗ Δ[M]) ≫ tensor_μ M M M M ≫ (μ[M] ⊗ μ[M]) := + mul_comul' + +@[reassoc (attr := simp)] +theorem one_comul (M : C) [Bimon_Class M] : η[M] ≫ Δ[M] = η[M ⊗ M] := one_comul' + +@[reassoc (attr := simp)] +theorem mul_counit (M : C) [Bimon_Class M] : μ[M] ≫ ε[M] = ε[M ⊗ M] := mul_counit' + +@[reassoc (attr := simp)] +theorem one_counit (M : C) [Bimon_Class M] : η[M] ≫ ε[M] = 𝟙 (𝟙_ C) := one_counit' + +end Bimon_Class + +variable (C) /-- A bimonoid object in a braided category `C` is a comonoid object in the (monoidal) diff --git a/Mathlib/CategoryTheory/Monoidal/Comon_.lean b/Mathlib/CategoryTheory/Monoidal/Comon_.lean index 498a20b561c79..ff63ce4bbb0eb 100644 --- a/Mathlib/CategoryTheory/Monoidal/Comon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Comon_.lean @@ -29,7 +29,49 @@ universe v₁ v₂ u₁ u₂ u open CategoryTheory MonoidalCategory -variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory.{v₁} C] +variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] + +/-- A comonoid object internal to a monoidal category. + +When the monoidal category is preadditive, this is also sometimes called a "coalgebra object". +-/ +class Comon_Class (X : C) where + /-- The counit morphism of a comonoid object. -/ + counit : X ⟶ 𝟙_ C + /-- The comultiplication morphism of a comonoid object. -/ + comul : X ⟶ X ⊗ X + /- For the names of the conditions below, the unprimed names are reserved for the version where + the argument `X` is explicit. -/ + counit_comul' : comul ≫ counit ▷ X = (λ_ X).inv := by aesop_cat + comul_counit' : comul ≫ X ◁ counit = (ρ_ X).inv := by aesop_cat + comul_assoc' : comul ≫ X ◁ comul = comul ≫ (comul ▷ X) ≫ (α_ X X X).hom := by aesop_cat + +namespace Comon_Class + +@[inherit_doc] scoped notation "Δ" => Comon_Class.comul +@[inherit_doc] scoped notation "Δ["M"]" => Comon_Class.comul (X := M) +@[inherit_doc] scoped notation "ε" => Comon_Class.counit +@[inherit_doc] scoped notation "ε["M"]" => Comon_Class.counit (X := M) + +/- The simp attribute is reserved for the unprimed versions. -/ +attribute [reassoc] counit_comul' comul_counit' comul_assoc' + +@[reassoc (attr := simp)] +theorem counit_comul (X : C) [Comon_Class X] : Δ ≫ ε ▷ X = (λ_ X).inv := counit_comul' + +@[reassoc (attr := simp)] +theorem comul_counit (X : C) [Comon_Class X] : Δ ≫ X ◁ ε = (ρ_ X).inv := comul_counit' + +@[reassoc (attr := simp)] +theorem comul_assoc (X : C) [Comon_Class X] : + Δ ≫ X ◁ Δ = Δ ≫ Δ ▷ X ≫ (α_ X X X).hom := + comul_assoc' + +end Comon_Class + +open scoped Comon_Class + +variable (C) /-- A comonoid object internal to a monoidal category. @@ -52,6 +94,24 @@ attribute [reassoc (attr := simp)] Comon_.comul_assoc namespace Comon_ +variable {C} + +/-- Construct an object of `Comon_ C` from an object `X : C` and `Comon_Class X` instance. -/ +@[simps] +def mk' (X : C) [Comon_Class X] : Comon_ C where + X := X + counit := ε + comul := Δ + +instance {M : Comon_ C} : Comon_Class M.X where + counit := M.counit + comul := M.comul + counit_comul' := M.counit_comul + comul_counit' := M.comul_counit + comul_assoc' := M.comul_assoc + +variable (C) + /-- The trivial comonoid object. We later show this is terminal in `Comon_ C`. -/ @[simps] @@ -251,6 +311,9 @@ variable [BraidedCategory C] theorem tensorObj_X (A B : Comon_ C) : (A ⊗ B).X = A.X ⊗ B.X := rfl +instance (A B : C) [Comon_Class A] [Comon_Class B] : Comon_Class (A ⊗ B) := + inferInstanceAs <| Comon_Class (Comon_.mk' A ⊗ Comon_.mk' B).X + theorem tensorObj_counit (A B : Comon_ C) : (A ⊗ B).counit = (A.counit ⊗ B.counit) ≫ (λ_ _).hom := rfl @@ -281,7 +344,7 @@ theorem tensorObj_comul (A B : Comon_ C) : /-- The forgetful functor from `Comon_ C` to `C` is monoidal when `C` is braided monoidal. -/ def forgetMonoidal : MonoidalFunctor (Comon_ C) C := { forget C with - ε := 𝟙 _ + «ε» := 𝟙 _ μ := fun _ _ => 𝟙 _ } @[simp] theorem forgetMonoidal_toFunctor : (forgetMonoidal C).toFunctor = forget C := rfl diff --git a/Mathlib/CategoryTheory/Monoidal/Hopf_.lean b/Mathlib/CategoryTheory/Monoidal/Hopf_.lean index 4bf50784e7cc7..e9e5220ef1ddc 100644 --- a/Mathlib/CategoryTheory/Monoidal/Hopf_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Hopf_.lean @@ -22,7 +22,40 @@ universe v₁ v₂ u₁ u₂ u open CategoryTheory MonoidalCategory -variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory.{v₁} C] [BraidedCategory C] +variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] [BraidedCategory C] + +open scoped Mon_Class Comon_Class + +/-- +A Hopf monoid in a braided category `C` is a bimonoid object in `C` equipped with an antipode. +-/ +class Hopf_Class (X : C) extends Bimon_Class X where + /-- The antipode is an endomorphism of the underlying object of the Hopf monoid. -/ + antipode : X ⟶ X + /- For the names of the conditions below, the unprimed names are reserved for the version where + the argument `X` is explicit. -/ + antipode_left' : Δ ≫ antipode ▷ X ≫ μ = ε ≫ η := by aesop_cat + antipode_right' : Δ ≫ X ◁ antipode ≫ μ = ε ≫ η := by aesop_cat + +namespace Hopf_Class + +@[inherit_doc] scoped notation "𝒮" => Hopf_Class.antipode +@[inherit_doc] scoped notation "𝒮["X"]" => Hopf_Class.antipode (X := X) + +/- The simp attribute is reserved for the unprimed versions. -/ +attribute [reassoc] antipode_left' antipode_right' + +/-- The object is provided as an explicit argument. -/ +@[reassoc (attr := simp)] +theorem antipode_left (X : C) [Hopf_Class X] : Δ ≫ 𝒮 ▷ X ≫ μ = ε ≫ η := antipode_left' + +/-- The object is provided as an explicit argument. -/ +@[reassoc (attr := simp)] +theorem antipode_right (X : C) [Hopf_Class X] : Δ ≫ X ◁ 𝒮 ≫ μ = ε ≫ η := antipode_right' + +end Hopf_Class + +variable (C) /-- A Hopf monoid in a braided category `C` is a bimonoid object in `C` equipped with an antipode. diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index fc0cec7c2faf7..c04c4f4d55c07 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -23,7 +23,51 @@ universe v₁ v₂ u₁ u₂ u open CategoryTheory MonoidalCategory -variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory.{v₁} C] +variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] + +/-- A monoid object internal to a monoidal category. + +When the monoidal category is preadditive, this is also sometimes called an "algebra object". +-/ +class Mon_Class (X : C) where + /-- The unit morphism of a monoid object. -/ + one : 𝟙_ C ⟶ X + /-- The multiplication morphism of a monoid object. -/ + mul : X ⊗ X ⟶ X + /- For the names of the conditions below, the unprimed names are reserved for the version where + the argument `X` is explicit. -/ + one_mul' : one ▷ X ≫ mul = (λ_ X).hom := by aesop_cat + mul_one' : X ◁ one ≫ mul = (ρ_ X).hom := by aesop_cat + -- Obviously there is some flexibility stating this axiom. + -- This one has left- and right-hand sides matching the statement of `Monoid.mul_assoc`, + -- and chooses to place the associator on the right-hand side. + -- The heuristic is that unitors and associators "don't have much weight". + mul_assoc' : (mul ▷ X) ≫ mul = (α_ X X X).hom ≫ (X ◁ mul) ≫ mul := by aesop_cat + +namespace Mon_Class + +@[inherit_doc] scoped notation "μ" => Mon_Class.mul +@[inherit_doc] scoped notation "μ["M"]" => Mon_Class.mul (X := M) +@[inherit_doc] scoped notation "η" => Mon_Class.one +@[inherit_doc] scoped notation "η["M"]" => Mon_Class.one (X := M) + +/- The simp attribute is reserved for the unprimed versions. -/ +attribute [reassoc] one_mul' mul_one' mul_assoc' + +@[reassoc (attr := simp)] +theorem one_mul (X : C) [Mon_Class X] : η ▷ X ≫ μ = (λ_ X).hom := one_mul' + +@[reassoc (attr := simp)] +theorem mul_one (X : C) [Mon_Class X] : X ◁ η ≫ μ = (ρ_ X).hom := mul_one' + +@[reassoc (attr := simp)] +theorem mul_assoc (X : C) [Mon_Class X] : μ ▷ X ≫ μ = (α_ X X X).hom ≫ X ◁ μ ≫ μ := mul_assoc' + +end Mon_Class + +open scoped Mon_Class + +variable (C) /-- A monoid object internal to a monoidal category. @@ -50,6 +94,24 @@ attribute [reassoc (attr := simp)] Mon_.mul_assoc namespace Mon_ +variable {C} + +/-- Construct an object of `Mon_ C` from an object `X : C` and `Mon_Class X` instance. -/ +@[simps] +def mk' (X : C) [Mon_Class X] : Mon_ C where + X := X + one := η + mul := μ + +instance {M : Mon_ C} : Mon_Class M.X where + one := M.one + mul := M.mul + one_mul' := M.one_mul + mul_one' := M.mul_one + mul_assoc' := M.mul_assoc + +variable (C) + /-- The trivial monoid object. We later show this is initial in `Mon_ C`. -/ @[simps] @@ -235,7 +297,7 @@ def monToLaxMonoidal : Mon_ C ⥤ LaxMonoidalFunctor (Discrete PUnit.{u + 1}) C { obj := fun _ => A.X map := fun _ => 𝟙 _ ε := A.one - μ := fun _ _ => A.mul } + «μ» := fun _ _ => A.mul } map f := { app := fun _ => f.hom } @@ -496,13 +558,17 @@ theorem tensor_mul (M N : Mon_ C) : (M ⊗ N).mul = instance monMonoidal : MonoidalCategory (Mon_ C) where tensorHom_def := by intros; ext; simp [tensorHom_def] +@[simps!] +instance {M N : C} [Mon_Class M] [Mon_Class N] : Mon_Class (M ⊗ N) := + inferInstanceAs <| Mon_Class (Mon_.mk' M ⊗ Mon_.mk' N).X + variable (C) /-- The forgetful functor from `Mon_ C` to `C` is monoidal when `C` is braided monoidal. -/ def forgetMonoidal : MonoidalFunctor (Mon_ C) C := { forget C with ε := 𝟙 _ - μ := fun _ _ => 𝟙 _ } + «μ» := fun _ _ => 𝟙 _ } @[simp] theorem forgetMonoidal_toFunctor : (forgetMonoidal C).toFunctor = forget C := rfl @[simp] theorem forgetMonoidal_ε : (forgetMonoidal C).ε = 𝟙 (𝟙_ C) := rfl From c6a02ffa70f93156c098b09f6b8d612ed1eb8a52 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 18 Oct 2024 16:20:05 +0000 Subject: [PATCH 072/131] chore: more `erw`s that can become `rw`s (#17923) Automatically found by the `erw` linter from #17638. --- Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean | 2 +- Mathlib/CategoryTheory/Shift/Basic.lean | 2 +- Mathlib/Condensed/Discrete/Colimit.lean | 4 ++-- Mathlib/Geometry/RingedSpace/OpenImmersion.lean | 4 ++-- Mathlib/Topology/Order/ScottTopology.lean | 4 ++-- 5 files changed, 8 insertions(+), 8 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean b/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean index cb15f1505053c..abcc9d116cda0 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean @@ -514,7 +514,7 @@ lemma mk_π {X : C} (π : F.obj zero ⟶ X) (h : F.map left ≫ π = F.map right (mk π h).π = π := rfl lemma condition (G : ReflexiveCofork F) : F.map left ≫ G.π = F.map right ≫ G.π := by - erw [Cocone.w G left, Cocone.w G right] + rw [Cocone.w G left, Cocone.w G right] @[simp] lemma app_one_eq_π (G : ReflexiveCofork F) : G.ι.app zero = G.π := rfl diff --git a/Mathlib/CategoryTheory/Shift/Basic.lean b/Mathlib/CategoryTheory/Shift/Basic.lean index 7acb86cbee29d..889c3bf069f82 100644 --- a/Mathlib/CategoryTheory/Shift/Basic.lean +++ b/Mathlib/CategoryTheory/Shift/Basic.lean @@ -519,7 +519,7 @@ lemma shiftFunctorCompIsoId_add'_inv_app : dsimp [shiftFunctorCompIsoId] simp only [Functor.map_comp, Category.assoc] congr 1 - erw [← NatTrans.naturality] + rw [← NatTrans.naturality] dsimp rw [← cancel_mono ((shiftFunctorAdd' C p' p 0 hp).inv.app X), Iso.hom_inv_id_app, Category.assoc, Category.assoc, Category.assoc, Category.assoc, diff --git a/Mathlib/Condensed/Discrete/Colimit.lean b/Mathlib/Condensed/Discrete/Colimit.lean index 75f3c6ee7dfeb..9e6930f3aa849 100644 --- a/Mathlib/Condensed/Discrete/Colimit.lean +++ b/Mathlib/Condensed/Discrete/Colimit.lean @@ -276,7 +276,7 @@ lemma isoLocallyConstantOfIsColimit_inv (X : Profinite.{u}ᵒᵖ ⥤ Type (u+1)) simp only [types_comp_apply, isoFinYoneda_inv_app, counitApp_app] apply presheaf_ext.{u, u+1} (X := X) (Y := X) (f := f) intro x - erw [incl_of_counitAppApp] + rw [incl_of_counitAppApp] simp only [counitAppAppImage, CompHausLike.coe_of] letI : Fintype (fiber.{u, u+1} f x) := Fintype.ofInjective (sigmaIncl.{u, u+1} f x).1 Subtype.val_injective @@ -551,7 +551,7 @@ lemma isoLocallyConstantOfIsColimit_inv (X : LightProfinite.{u}ᵒᵖ ⥤ Type u simp only [types_comp_apply, isoFinYoneda_inv_app, counitApp_app] apply presheaf_ext.{u, u} (X := X) (Y := X) (f := f) intro x - erw [incl_of_counitAppApp] + rw [incl_of_counitAppApp] simp only [counitAppAppImage, CompHausLike.coe_of] letI : Fintype (fiber.{u, u} f x) := Fintype.ofInjective (sigmaIncl.{u, u} f x).1 Subtype.val_injective diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 959214c257295..11dbae9b63d30 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -1120,8 +1120,8 @@ theorem lift_range (H' : Set.range g.base ⊆ Set.range f.base) : have : _ = (pullback.fst f g).base := PreservesPullback.iso_hom_fst (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forget _) f g - rw [LocallyRingedSpace.comp_base, ← this, ← Category.assoc, coe_comp] - rw [Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ] + rw [LocallyRingedSpace.comp_base, ← this, ← Category.assoc, coe_comp, Set.range_comp, + Set.range_iff_surjective.mpr, Set.image_univ] -- Porting note (#11224): change `rw` to `erw` on this lemma · erw [TopCat.pullback_fst_range] ext diff --git a/Mathlib/Topology/Order/ScottTopology.lean b/Mathlib/Topology/Order/ScottTopology.lean index 0a203e27ad789..2a58f0caeeb47 100644 --- a/Mathlib/Topology/Order/ScottTopology.lean +++ b/Mathlib/Topology/Order/ScottTopology.lean @@ -177,7 +177,7 @@ variable {α D} lemma isOpen_iff [IsScottHausdorff α D] : IsOpen s ↔ ∀ ⦃d : Set α⦄, d ∈ D → d.Nonempty → DirectedOn (· ≤ ·) d → ∀ ⦃a : α⦄, IsLUB d a → - a ∈ s → ∃ b ∈ d, Ici b ∩ d ⊆ s := by erw [topology_eq_scottHausdorff (α := α) (D := D)]; rfl + a ∈ s → ∃ b ∈ d, Ici b ∩ d ⊆ s := by rw [topology_eq_scottHausdorff (α := α) (D := D)]; rfl lemma dirSupInaccOn_of_isOpen [IsScottHausdorff α D] (h : IsOpen s) : DirSupInaccOn D s := fun d hd₀ hd₁ hd₂ a hda hd₃ ↦ by @@ -241,7 +241,7 @@ lemma topology_eq [IsScott α D] : ‹_› = scott α D := topology_eq_scott variable {α} {D} {s : Set α} {a : α} lemma isOpen_iff_isUpperSet_and_scottHausdorff_open [IsScott α D] : - IsOpen s ↔ IsUpperSet s ∧ IsOpen[scottHausdorff α D] s := by erw [topology_eq α D]; rfl + IsOpen s ↔ IsUpperSet s ∧ IsOpen[scottHausdorff α D] s := by rw [topology_eq α D]; rfl lemma isOpen_iff_isUpperSet_and_dirSupInaccOn [IsScott α D] : IsOpen s ↔ IsUpperSet s ∧ DirSupInaccOn D s := by From 7b3afdd0e7374470ac2af311f49a06197ede6db1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 16:28:16 +0000 Subject: [PATCH 073/131] feat: more lemmas about the size of pointwise sets (#17231) Complete the API and follow the convention that `Cardinal.mk` is called `card` and `Nat.card` is called `natCard` in lemma names. From LeanCamCombi --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 82 ++++++++++++------- .../GroupWithZero/Pointwise/Set/Card.lean | 27 ++++++ Mathlib/SetTheory/Cardinal/Basic.lean | 8 ++ 4 files changed, 89 insertions(+), 29 deletions(-) create mode 100644 Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3c455bcceb7e9..b0b4fbb4d3724 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -330,6 +330,7 @@ import Mathlib.Algebra.GroupWithZero.NonZeroDivisors import Mathlib.Algebra.GroupWithZero.Opposite import Mathlib.Algebra.GroupWithZero.Pi import Mathlib.Algebra.GroupWithZero.Pointwise.Set.Basic +import Mathlib.Algebra.GroupWithZero.Pointwise.Set.Card import Mathlib.Algebra.GroupWithZero.Prod import Mathlib.Algebra.GroupWithZero.Semiconj import Mathlib.Algebra.GroupWithZero.ULift diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 341ed441df932..80a073b792385 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -3,60 +3,84 @@ Copyright (c) 2024 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.Data.Set.Pointwise.Finite import Mathlib.SetTheory.Cardinal.Finite /-! -# Cardinalities of pointwise operations on sets. +# Cardinalities of pointwise operations on sets -/ -namespace Set - -open Pointwise - -variable {α β : Type*} +open scoped Cardinal Pointwise -section MulAction -variable [Group α] [MulAction α β] +namespace Set +variable {G M α : Type*} -@[to_additive (attr := simp)] -lemma card_smul_set (a : α) (s : Set β) : Nat.card ↥(a • s) = Nat.card s := - Nat.card_image_of_injective (MulAction.injective a) _ +section Mul +variable [Mul M] {s t : Set M} -end MulAction +@[to_additive] +lemma _root_.Cardinal.mk_mul_le : #(s * t) ≤ #s * #t := by + rw [← image2_mul]; exact Cardinal.mk_image2_le -section IsCancelMul -variable [Mul α] [IsCancelMul α] {s t : Set α} +variable [IsCancelMul M] @[to_additive] -lemma card_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by - classical +lemma natCard_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by obtain h | h := (s * t).infinite_or_finite · simp [Set.Infinite.card_eq_zero h] - obtain ⟨hs, ht⟩ | rfl | rfl := finite_mul.1 h - · lift s to Finset α using hs - lift t to Finset α using ht - rw [← Finset.coe_mul] - simpa [-Finset.coe_mul] using Finset.card_mul_le - all_goals simp + simp only [Nat.card, ← Cardinal.toNat_mul] + refine Cardinal.toNat_le_toNat Cardinal.mk_mul_le ?_ + aesop (add simp [Cardinal.mul_lt_aleph0_iff, finite_mul]) + +@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_mul_le := natCard_mul_le -end IsCancelMul +end Mul section InvolutiveInv -variable [InvolutiveInv α] {s t : Set α} +variable [InvolutiveInv G] {s t : Set G} + +@[to_additive (attr := simp)] +lemma _root_.Cardinal.mk_inv (s : Set G) : #↥(s⁻¹) = #s := by + rw [← image_inv, Cardinal.mk_image_eq_of_injOn _ _ inv_injective.injOn] @[to_additive (attr := simp)] -lemma card_inv (s : Set α) : Nat.card ↥(s⁻¹) = Nat.card s := by +lemma natCard_inv (s : Set G) : Nat.card ↥(s⁻¹) = Nat.card s := by rw [← image_inv, Nat.card_image_of_injective inv_injective] +@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_inv := natCard_inv + end InvolutiveInv +section DivInvMonoid +variable [DivInvMonoid M] {s t : Set M} + +@[to_additive] +lemma _root_.Cardinal.mk_div_le : #(s / t) ≤ #s * #t := by + rw [← image2_div]; exact Cardinal.mk_image2_le + +end DivInvMonoid + section Group -variable [Group α] {s t : Set α} +variable [Group G] {s t : Set G} @[to_additive] -lemma card_div_le : Nat.card (s / t) ≤ Nat.card s * Nat.card t := by - rw [div_eq_mul_inv, ← card_inv t]; exact card_mul_le +lemma natCard_div_le : Nat.card (s / t) ≤ Nat.card s * Nat.card t := by + rw [div_eq_mul_inv, ← natCard_inv t]; exact natCard_mul_le + +@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_div_le := natCard_div_le + +variable [MulAction G α] + +@[to_additive (attr := simp)] +lemma _root_.Cardinal.mk_smul_set (a : G) (s : Set α) : #↥(a • s) = #s := + Cardinal.mk_image_eq_of_injOn _ _ (MulAction.injective a).injOn + +@[to_additive (attr := simp)] +lemma natCard_smul_set (a : G) (s : Set α) : Nat.card ↥(a • s) = Nat.card s := + Nat.card_image_of_injective (MulAction.injective a) _ + +@[to_additive (attr := deprecated (since := "2024-09-30"))] +alias card_smul_set := Cardinal.mk_smul_set end Group end Set diff --git a/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean b/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean new file mode 100644 index 0000000000000..8466c8876f1c1 --- /dev/null +++ b/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean @@ -0,0 +1,27 @@ +/- +Copyright (c) 2024 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.Group.Pointwise.Set.Basic +import Mathlib.Algebra.GroupWithZero.Action.Basic +import Mathlib.SetTheory.Cardinal.Finite + +/-! +# Cardinality of sets under pointwise group with zero operations +-/ + +open scoped Cardinal Pointwise + +variable {G G₀ M M₀ : Type*} + +namespace Set +variable [GroupWithZero G₀] [Zero M₀] [MulActionWithZero G₀ M₀] {a : G₀} + +lemma _root_.Cardinal.mk_smul_set₀ (ha : a ≠ 0) (s : Set M₀) : #↥(a • s) = #s := + Cardinal.mk_image_eq_of_injOn _ _ (MulAction.injective₀ ha).injOn + +lemma natCard_smul_set₀ (ha : a ≠ 0) (s : Set M₀) : Nat.card ↥(a • s) = Nat.card s := + Nat.card_image_of_injective (MulAction.injective₀ ha) _ + +end Set diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 22ec5d425d8c1..fcc9e9174c5c2 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1820,9 +1820,17 @@ theorem mk_emptyCollection_iff {α : Type u} {s : Set α} : #s = 0 ↔ s = ∅ : theorem mk_univ {α : Type u} : #(@univ α) = #α := mk_congr (Equiv.Set.univ α) +@[simp] lemma mk_setProd {α β : Type u} (s : Set α) (t : Set β) : #(s ×ˢ t) = #s * #t := by + rw [mul_def, mk_congr (Equiv.Set.prod ..)] + theorem mk_image_le {α β : Type u} {f : α → β} {s : Set α} : #(f '' s) ≤ #s := mk_le_of_surjective surjective_onto_image +lemma mk_image2_le {α β γ : Type u} {f : α → β → γ} {s : Set α} {t : Set β} : + #(image2 f s t) ≤ #s * #t := by + rw [← image_uncurry_prod, ← mk_setProd] + exact mk_image_le + theorem mk_image_le_lift {α : Type u} {β : Type v} {f : α → β} {s : Set α} : lift.{u} #(f '' s) ≤ lift.{v} #s := lift_mk_le.{0}.mpr ⟨Embedding.ofSurjective _ surjective_onto_image⟩ From 7e96f0c587bdf06977946b7227123bfd4ae81dc4 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 18 Oct 2024 16:28:18 +0000 Subject: [PATCH 074/131] chore: remove unused open statements (#17899) --- Mathlib/Algebra/Group/EvenFunction.lean | 2 -- Mathlib/Algebra/Polynomial/Module/Basic.lean | 2 +- Mathlib/Analysis/Calculus/LogDeriv.lean | 2 +- Mathlib/Analysis/Normed/Group/Ultra.lean | 2 -- .../Analysis/SpecialFunctions/Trigonometric/Cotangent.lean | 4 ++-- Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean | 1 - Mathlib/Combinatorics/SetFamily/KruskalKatona.lean | 2 +- Mathlib/Data/Nat/Choose/Lucas.lean | 2 +- Mathlib/GroupTheory/CosetCover.lean | 2 +- Mathlib/GroupTheory/GroupAction/Blocks.lean | 2 +- Mathlib/MeasureTheory/Measure/SeparableMeasure.lean | 2 +- Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean | 2 +- Mathlib/NumberTheory/Fermat.lean | 1 - Mathlib/NumberTheory/JacobiSum/Basic.lean | 2 +- Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean | 2 +- Mathlib/NumberTheory/NumberField/Units/Regulator.lean | 2 +- Mathlib/Order/LiminfLimsup.lean | 2 +- Mathlib/Probability/Martingale/BorelCantelli.lean | 2 +- Mathlib/RingTheory/RootsOfUnity/Lemmas.lean | 2 +- 19 files changed, 16 insertions(+), 22 deletions(-) diff --git a/Mathlib/Algebra/Group/EvenFunction.lean b/Mathlib/Algebra/Group/EvenFunction.lean index 3896969ce6795..4713a0549c61f 100644 --- a/Mathlib/Algebra/Group/EvenFunction.lean +++ b/Mathlib/Algebra/Group/EvenFunction.lean @@ -18,8 +18,6 @@ conflicting with the root-level definitions `Even` and `Odd` (which, for functio function takes even resp. odd _values_, a wholly different concept). -/ -open scoped BigOperators - namespace Function variable {α β : Type*} [Neg α] diff --git a/Mathlib/Algebra/Polynomial/Module/Basic.lean b/Mathlib/Algebra/Polynomial/Module/Basic.lean index 9a18a3951bc2e..c22c4aea59f64 100644 --- a/Mathlib/Algebra/Polynomial/Module/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Module/Basic.lean @@ -14,7 +14,7 @@ This is defined as a type alias `PolynomialModule R M := ℕ →₀ M`, since th module structures on `ℕ →₀ M` of interest. See the docstring of `PolynomialModule` for details. -/ universe u v -open Polynomial BigOperators +open Polynomial /-- The `R[X]`-module `M[X]` for an `R`-module `M`. This is isomorphic (as an `R`-module) to `M[X]` when `M` is a ring. diff --git a/Mathlib/Analysis/Calculus/LogDeriv.lean b/Mathlib/Analysis/Calculus/LogDeriv.lean index 495f3de5c2775..7b8a8970a06f3 100644 --- a/Mathlib/Analysis/Calculus/LogDeriv.lean +++ b/Mathlib/Analysis/Calculus/LogDeriv.lean @@ -17,7 +17,7 @@ noncomputable section open Filter Function -open scoped Topology BigOperators Classical +open scoped Topology Classical variable {𝕜 𝕜': Type*} [NontriviallyNormedField 𝕜] [NontriviallyNormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] diff --git a/Mathlib/Analysis/Normed/Group/Ultra.lean b/Mathlib/Analysis/Normed/Group/Ultra.lean index 2b8c5f341691f..8881f035276fb 100644 --- a/Mathlib/Analysis/Normed/Group/Ultra.lean +++ b/Mathlib/Analysis/Normed/Group/Ultra.lean @@ -33,8 +33,6 @@ ultrametric, nonarchimedean -/ open Metric NNReal -open scoped BigOperators - namespace IsUltrametricDist section Group diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Cotangent.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Cotangent.lean index d9d5e56fb97ef..1db8bfb1259ff 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Cotangent.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Cotangent.lean @@ -11,9 +11,9 @@ import Mathlib.Analysis.Complex.UpperHalfPlane.Exp This file contains lemmas about the cotangent function, including useful series expansions. -/ -open Real Complex BigOperators Filter +open Real Complex -open scoped UpperHalfPlane Topology +open scoped UpperHalfPlane lemma Complex.cot_eq_exp_ratio (z : ℂ) : cot z = (Complex.exp (2 * I * z) + 1) / (I * (1 - Complex.exp (2 * I * z))) := by diff --git a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean index a3bad041f207e..164ae582fd759 100644 --- a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean +++ b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean @@ -22,7 +22,6 @@ elements of sum zero. -/ open Finset MvPolynomial -open scoped BigOperators variable {ι : Type*} diff --git a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean index 8b5ce0135fec7..23daec75ae3d2 100644 --- a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean +++ b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean @@ -119,7 +119,7 @@ protected lemma IsInitSeg.shadow [Finite α] (h₁ : IsInitSeg 𝒜 r) : IsInitS end Colex open Finset Colex Nat UV -open scoped BigOperators FinsetFamily +open scoped FinsetFamily variable {α : Type*} [LinearOrder α] {s U V : Finset α} {n : ℕ} diff --git a/Mathlib/Data/Nat/Choose/Lucas.lean b/Mathlib/Data/Nat/Choose/Lucas.lean index 94e97c1f452af..0df7d593456f8 100644 --- a/Mathlib/Data/Nat/Choose/Lucas.lean +++ b/Mathlib/Data/Nat/Choose/Lucas.lean @@ -22,7 +22,7 @@ k_i` modulo `p`, where `n_i` and `k_i` are the base-`p` digits of `n` and `k`, r open Finset hiding choose -open Nat BigOperators Polynomial +open Nat Polynomial namespace Choose diff --git a/Mathlib/GroupTheory/CosetCover.lean b/Mathlib/GroupTheory/CosetCover.lean index a8ba591868ffa..155e9fd467601 100644 --- a/Mathlib/GroupTheory/CosetCover.lean +++ b/Mathlib/GroupTheory/CosetCover.lean @@ -42,7 +42,7 @@ set of all minimal polynomials (not proved here). -/ -open scoped Pointwise BigOperators +open scoped Pointwise namespace Subgroup diff --git a/Mathlib/GroupTheory/GroupAction/Blocks.lean b/Mathlib/GroupTheory/GroupAction/Blocks.lean index e88c5972e23aa..cc63c90b25d16 100644 --- a/Mathlib/GroupTheory/GroupAction/Blocks.lean +++ b/Mathlib/GroupTheory/GroupAction/Blocks.lean @@ -43,7 +43,7 @@ We follow [wielandt1964]. -/ -open scoped BigOperators Pointwise +open scoped Pointwise namespace MulAction diff --git a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean index 167c0d5bdfb53..cb0e5e2a0c352 100644 --- a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean @@ -63,7 +63,7 @@ written `≠ ∞` rather than `< ∞`. See `Ne.lt_top` and `ne_of_lt` to switch separable measure, measure-dense, Lp space, second-countable -/ -open MeasurableSpace Set ENNReal TopologicalSpace BigOperators symmDiff Filter Real +open MeasurableSpace Set ENNReal TopologicalSpace symmDiff Real namespace MeasureTheory diff --git a/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean b/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean index 2d7efbda8ea35..54708abe18b61 100644 --- a/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean +++ b/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean @@ -49,7 +49,7 @@ namespace FirstOrder namespace Field -open Ring FreeCommRing BigOperators Polynomial Language +open Ring FreeCommRing Polynomial Language /-- A generic monic polynomial of degree `n` as an element of the free commutative ring in `n+1` variables, with a variable for each diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index c90f62582cae8..0e3ab436b2414 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -26,7 +26,6 @@ for all natural numbers `n`. namespace Nat open Finset -open scoped BigOperators /-- Fermat numbers: the `n`-th Fermat number is defined as `2^(2^n) + 1`. -/ def fermatNumber (n : ℕ) : ℕ := 2 ^ (2 ^ n) + 1 diff --git a/Mathlib/NumberTheory/JacobiSum/Basic.lean b/Mathlib/NumberTheory/JacobiSum/Basic.lean index a936f799462ca..7e6fd3e597e49 100644 --- a/Mathlib/NumberTheory/JacobiSum/Basic.lean +++ b/Mathlib/NumberTheory/JacobiSum/Basic.lean @@ -28,7 +28,7 @@ but generalize where appropriate. This is based on Lean code written as part of the bachelor's thesis of Alexander Spahl. -/ -open BigOperators Finset +open Finset /-! ### Jacobi sums: definition and first properties diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 4970ba0001a5a..cc34b948b40e0 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -40,7 +40,7 @@ open scoped NumberField noncomputable section -open NumberField NumberField.InfinitePlace NumberField.Units BigOperators +open NumberField NumberField.InfinitePlace NumberField.Units variable (K : Type*) [Field K] diff --git a/Mathlib/NumberTheory/NumberField/Units/Regulator.lean b/Mathlib/NumberTheory/NumberField/Units/Regulator.lean index 34d7ce2339491..cd7cb76b5bf07 100644 --- a/Mathlib/NumberTheory/NumberField/Units/Regulator.lean +++ b/Mathlib/NumberTheory/NumberField/Units/Regulator.lean @@ -32,7 +32,7 @@ namespace NumberField.Units variable (K : Type*) [Field K] -open MeasureTheory Classical BigOperators NumberField.InfinitePlace +open MeasureTheory Classical NumberField.InfinitePlace NumberField NumberField.Units.dirichletUnitTheorem variable [NumberField K] diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 0d6faea4d46f6..300e1d0160937 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -285,7 +285,7 @@ end Relation section add_and_sum -open Filter BigOperators Set +open Filter Set variable {α : Type*} {f : Filter α} variable {R : Type*} diff --git a/Mathlib/Probability/Martingale/BorelCantelli.lean b/Mathlib/Probability/Martingale/BorelCantelli.lean index 4897ac3d82a8e..0f483a8596a0b 100644 --- a/Mathlib/Probability/Martingale/BorelCantelli.lean +++ b/Mathlib/Probability/Martingale/BorelCantelli.lean @@ -34,7 +34,7 @@ and `ProbabilityTheory.measure_limsup_eq_one` for the second (which does). open Filter -open scoped NNReal ENNReal MeasureTheory ProbabilityTheory BigOperators Topology +open scoped NNReal ENNReal MeasureTheory ProbabilityTheory Topology namespace MeasureTheory diff --git a/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean b/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean index e6fded4815f54..33d3afe29785a 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean @@ -23,7 +23,7 @@ variable {R : Type*} [CommRing R] [IsDomain R] namespace IsPrimitiveRoot -open Finset Polynomial BigOperators +open Finset Polynomial /-- If `μ` is a primitive `n`th root of unity in `R`, then `∏(1≤k Date: Fri, 18 Oct 2024 16:28:19 +0000 Subject: [PATCH 075/131] chore: remove `CoeFun` instances where `FunLike` is available (#17900) During the port we found that `FunLike` is robust enough not to need an extra `CoeFun` shortcut. Let's see if that rule can be consistently applied to the whole of the library. There is still duplication between `FunLike` and `CoeFun` for `Grp`, `Mon`, `CommGrp` and `CommMon`, which will need a more thorough fix. See also #17866. This is the second half of #17870, and doesn't affect the benchmarks. --- Mathlib/Algebra/Algebra/Hom.lean | 2 -- Mathlib/Algebra/Algebra/NonUnitalHom.lean | 5 ----- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 4 ++-- Mathlib/Algebra/Polynomial/Module/Basic.lean | 2 +- Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean | 3 --- Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean | 5 ----- Mathlib/LinearAlgebra/QuadraticForm/Basic.lean | 5 ----- Mathlib/MeasureTheory/OuterMeasure/Defs.lean | 3 --- Mathlib/ModelTheory/Basic.lean | 6 ------ Mathlib/ModelTheory/ElementaryMaps.lean | 3 --- Mathlib/Order/OmegaCompletePartialOrder.lean | 1 - Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean | 3 --- Mathlib/Topology/Algebra/Module/WeakDual.lean | 5 ----- Mathlib/Topology/Connected/PathConnected.lean | 6 ------ Mathlib/Topology/Homeomorph.lean | 2 -- Mathlib/Topology/MetricSpace/Dilation.lean | 3 --- Mathlib/Topology/MetricSpace/DilationEquiv.lean | 3 --- 17 files changed, 3 insertions(+), 58 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Hom.lean b/Mathlib/Algebra/Algebra/Hom.lean index 55dcdc81a91cb..5c5eac38616b5 100644 --- a/Mathlib/Algebra/Algebra/Hom.lean +++ b/Mathlib/Algebra/Algebra/Hom.lean @@ -85,8 +85,6 @@ section Semiring variable [CommSemiring R] [Semiring A] [Semiring B] [Semiring C] [Semiring D] variable [Algebra R A] [Algebra R B] [Algebra R C] [Algebra R D] --- Porting note: we don't port specialized `CoeFun` instances if there is `DFunLike` instead - instance funLike : FunLike (A →ₐ[R] B) A B where coe f := f.toFun coe_injective' f g h := by diff --git a/Mathlib/Algebra/Algebra/NonUnitalHom.lean b/Mathlib/Algebra/Algebra/NonUnitalHom.lean index 34a5518dc4de3..7eb27b49ee74a 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalHom.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalHom.lean @@ -156,11 +156,6 @@ variable [NonUnitalNonAssocSemiring A] [DistribMulAction R A] variable [NonUnitalNonAssocSemiring B] [DistribMulAction S B] variable [NonUnitalNonAssocSemiring C] [DistribMulAction T C] --- Porting note: Replaced with DFunLike instance --- /-- see Note [function coercion] -/ --- instance : CoeFun (A →ₙₐ[R] B) fun _ => A → B := --- ⟨toFun⟩ - instance : DFunLike (A →ₛₙₐ[φ] B) A fun _ => B where coe f := f.toFun coe_injective' := by rintro ⟨⟨⟨f, _⟩, _⟩, _⟩ ⟨⟨⟨g, _⟩, _⟩, _⟩ h; congr diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 7125b7d636d97..767b5198c52d8 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -86,7 +86,7 @@ instance MonoidAlgebra.instIsCancelAdd [IsCancelAdd k] : IsCancelAdd (MonoidAlge inferInstanceAs (IsCancelAdd (G →₀ k)) instance MonoidAlgebra.coeFun : CoeFun (MonoidAlgebra k G) fun _ => G → k := - Finsupp.instCoeFun + inferInstanceAs (CoeFun (G →₀ k) _) end @@ -823,7 +823,7 @@ instance instIsCancelAdd [IsCancelAdd k] : IsCancelAdd (AddMonoidAlgebra k G) := inferInstanceAs (IsCancelAdd (G →₀ k)) instance coeFun : CoeFun k[G] fun _ => G → k := - Finsupp.instCoeFun + inferInstanceAs (CoeFun (G →₀ k) _) end AddMonoidAlgebra diff --git a/Mathlib/Algebra/Polynomial/Module/Basic.lean b/Mathlib/Algebra/Polynomial/Module/Basic.lean index c22c4aea59f64..047d5c7f86864 100644 --- a/Mathlib/Algebra/Polynomial/Module/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Module/Basic.lean @@ -57,7 +57,7 @@ instance instFunLike : FunLike (PolynomialModule R M) ℕ M := Finsupp.instFunLike instance : CoeFun (PolynomialModule R M) fun _ => ℕ → M := - Finsupp.instCoeFun + inferInstanceAs <| CoeFun (_ →₀ _) _ theorem zero_apply (i : ℕ) : (0 : PolynomialModule R M) i = 0 := Finsupp.zero_apply diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean index 0849ae203d5a8..3dcf47e22249d 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean @@ -89,9 +89,6 @@ instance equivLike : EquivLike (P₁ ≃ᵃ[k] P₂) P₁ P₂ where right_inv f := f.right_inv coe_injective' _ _ h _ := toAffineMap_injective (DFunLike.coe_injective h) -instance : CoeFun (P₁ ≃ᵃ[k] P₂) fun _ => P₁ → P₂ := - DFunLike.hasCoeToFun - instance : CoeOut (P₁ ≃ᵃ[k] P₂) (P₁ ≃ P₂) := ⟨AffineEquiv.toEquiv⟩ diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean index edeab9e15af30..71f99e5c85dde 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean @@ -67,11 +67,6 @@ instance AffineMap.instFunLike (k : Type*) {V1 : Type*} (P1 : Type*) {V2 : Type* apply vadd_right_cancel (f p) rw [← f_add, h, ← g_add] -instance AffineMap.hasCoeToFun (k : Type*) {V1 : Type*} (P1 : Type*) {V2 : Type*} (P2 : Type*) - [Ring k] [AddCommGroup V1] [Module k V1] [AffineSpace V1 P1] [AddCommGroup V2] [Module k V2] - [AffineSpace V2 P2] : CoeFun (P1 →ᵃ[k] P2) fun _ => P1 → P2 := - DFunLike.hasCoeToFun - namespace LinearMap variable {k : Type*} {V₁ : Type*} {V₂ : Type*} [Ring k] [AddCommGroup V₁] [Module k V₁] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index fe2fe2b66dfe2..39bc41327d511 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -161,11 +161,6 @@ instance instFunLike : FunLike (QuadraticMap R M N) M N where coe := toFun coe_injective' x y h := by cases x; cases y; congr -/-- Helper instance for when there's too many metavariables to apply -`DFunLike.hasCoeToFun` directly. -/ -instance : CoeFun (QuadraticMap R M N) fun _ => M → N := - ⟨DFunLike.coe⟩ - variable (Q) /-- The `simp` normal form for a quadratic map is `DFunLike.coe`, not `toFun`. -/ diff --git a/Mathlib/MeasureTheory/OuterMeasure/Defs.lean b/Mathlib/MeasureTheory/OuterMeasure/Defs.lean index e1169e732155e..a776ff90e2c9c 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Defs.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Defs.lean @@ -62,9 +62,6 @@ instance : FunLike (OuterMeasure α) (Set α) ℝ≥0∞ where coe m := m.measureOf coe_injective' | ⟨_, _, _, _⟩, ⟨_, _, _, _⟩, rfl => rfl -instance instCoeFun : CoeFun (OuterMeasure α) (fun _ => Set α → ℝ≥0∞) := - inferInstance - @[simp] theorem measureOf_eq_coe (m : OuterMeasure α) : m.measureOf = m := rfl instance : OuterMeasureClass (OuterMeasure α) α where diff --git a/Mathlib/ModelTheory/Basic.lean b/Mathlib/ModelTheory/Basic.lean index 6846aee6eb65c..08419dda7c6f0 100644 --- a/Mathlib/ModelTheory/Basic.lean +++ b/Mathlib/ModelTheory/Basic.lean @@ -296,9 +296,6 @@ instance homClass : HomClass L (M →[L] N) M N where instance [L.IsAlgebraic] : StrongHomClass L (M →[L] N) M N := HomClass.strongHomClassOfIsAlgebraic -instance hasCoeToFun : CoeFun (M →[L] N) fun _ => M → N := - DFunLike.hasCoeToFun - @[simp] theorem toFun_eq_coe {f : M →[L] N} : f.toFun = (f : M → N) := rfl @@ -557,9 +554,6 @@ def symm (f : M ≃[L] N) : N ≃[L] M := refine (f.map_rel' r (f.toEquiv.symm ∘ x)).symm.trans ?_ rw [← Function.comp_assoc, Equiv.toFun_as_coe, Equiv.self_comp_symm, Function.id_comp] } -instance hasCoeToFun : CoeFun (M ≃[L] N) fun _ => M → N := - DFunLike.hasCoeToFun - @[simp] theorem symm_symm (f : M ≃[L] N) : f.symm.symm = f := diff --git a/Mathlib/ModelTheory/ElementaryMaps.lean b/Mathlib/ModelTheory/ElementaryMaps.lean index f774228f18d50..a8c9fc2716718 100644 --- a/Mathlib/ModelTheory/ElementaryMaps.lean +++ b/Mathlib/ModelTheory/ElementaryMaps.lean @@ -67,9 +67,6 @@ instance instFunLike : FunLike (M ↪ₑ[L] N) M N where ext x exact funext_iff.1 h x -instance : CoeFun (M ↪ₑ[L] N) fun _ => M → N := - DFunLike.hasCoeToFun - @[simp] theorem map_boundedFormula (f : M ↪ₑ[L] N) {α : Type*} {n : ℕ} (φ : L.BoundedFormula α n) (v : α → M) (xs : Fin n → M) : φ.Realize (f ∘ v) (f ∘ xs) ↔ φ.Realize v xs := by diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 0d7b0cf6a3804..179530a5796f4 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -72,7 +72,6 @@ variable [Preorder α] [Preorder β] [Preorder γ] instance : FunLike (Chain α) ℕ α := inferInstanceAs <| FunLike (ℕ →o α) ℕ α instance : OrderHomClass (Chain α) ℕ α := inferInstanceAs <| OrderHomClass (ℕ →o α) ℕ α -instance : CoeFun (Chain α) fun _ => ℕ → α := ⟨DFunLike.coe⟩ instance [Inhabited α] : Inhabited (Chain α) := ⟨⟨default, fun _ _ _ => le_rfl⟩⟩ diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean index ad413d2ba21ae..1a0a5179e3380 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean @@ -78,9 +78,6 @@ instance continuousMapClass : ContinuousMapClass (ContinuousMultilinearMap R M₁ M₂) (∀ i, M₁ i) M₂ where map_continuous := ContinuousMultilinearMap.cont -instance : CoeFun (ContinuousMultilinearMap R M₁ M₂) fun _ => (∀ i, M₁ i) → M₂ := - ⟨fun f => f⟩ - /-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections. -/ def Simps.apply (L₁ : ContinuousMultilinearMap R M₁ M₂) (v : ∀ i, M₁ i) : M₂ := diff --git a/Mathlib/Topology/Algebra/Module/WeakDual.lean b/Mathlib/Topology/Algebra/Module/WeakDual.lean index 1a71efa833d83..45e6b09628962 100644 --- a/Mathlib/Topology/Algebra/Module/WeakDual.lean +++ b/Mathlib/Topology/Algebra/Module/WeakDual.lean @@ -102,11 +102,6 @@ instance instFunLike : FunLike (WeakDual 𝕜 E) E 𝕜 := instance instContinuousLinearMapClass : ContinuousLinearMapClass (WeakDual 𝕜 E) 𝕜 E 𝕜 := ContinuousLinearMap.continuousSemilinearMapClass -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun` -directly. -/ -instance : CoeFun (WeakDual 𝕜 E) fun _ => E → 𝕜 := - DFunLike.hasCoeToFun - /-- If a monoid `M` distributively continuously acts on `𝕜` and this action commutes with multiplication on `𝕜`, then it acts on `WeakDual 𝕜 E`. -/ instance instMulAction (M) [Monoid M] [DistribMulAction M 𝕜] [SMulCommClass 𝕜 M 𝕜] diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 4cc61d6513e73..93ef11a0f5e30 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -85,12 +85,6 @@ instance Path.funLike : FunLike (Path x y) I X where instance Path.continuousMapClass : ContinuousMapClass (Path x y) I X where map_continuous γ := show Continuous γ.toContinuousMap by fun_prop --- Porting note: not necessary in light of the instance above -/- -instance : CoeFun (Path x y) fun _ => I → X := - ⟨fun p => p.toFun⟩ --/ - @[ext] protected theorem Path.ext : ∀ {γ₁ γ₂ : Path x y}, (γ₁ : I → X) = γ₂ → γ₁ = γ₂ := by rintro ⟨⟨x, h11⟩, h12, h13⟩ ⟨⟨x, h21⟩, h22, h23⟩ rfl diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index 09be530f7a37a..c406c1e326c58 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -60,8 +60,6 @@ instance : EquivLike (X ≃ₜ Y) X Y where right_inv h := h.right_inv coe_injective' _ _ H _ := toEquiv_injective <| DFunLike.ext' H -instance : CoeFun (X ≃ₜ Y) fun _ ↦ X → Y := ⟨DFunLike.coe⟩ - @[simp] theorem homeomorph_mk_coe (a : X ≃ Y) (b c) : (Homeomorph.mk a b c : X → Y) = a := rfl diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index 2622d97a69454..0ece9c35abd15 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -87,9 +87,6 @@ instance funLike : FunLike (α →ᵈ β) α β where instance toDilationClass : DilationClass (α →ᵈ β) α β where edist_eq' f := edist_eq' f -instance : CoeFun (α →ᵈ β) fun _ => α → β := - DFunLike.hasCoeToFun - @[simp] theorem toFun_eq_coe {f : α →ᵈ β} : f.toFun = (f : α → β) := rfl diff --git a/Mathlib/Topology/MetricSpace/DilationEquiv.lean b/Mathlib/Topology/MetricSpace/DilationEquiv.lean index d510cd58fb9e7..3fffba7e08305 100644 --- a/Mathlib/Topology/MetricSpace/DilationEquiv.lean +++ b/Mathlib/Topology/MetricSpace/DilationEquiv.lean @@ -60,9 +60,6 @@ instance : EquivLike (X ≃ᵈ Y) X Y where instance : DilationEquivClass (X ≃ᵈ Y) X Y where edist_eq' f := f.edist_eq' -instance : CoeFun (X ≃ᵈ Y) fun _ ↦ (X → Y) where - coe f := f - @[simp] theorem coe_toEquiv (e : X ≃ᵈ Y) : ⇑e.toEquiv = e := rfl @[ext] From 2178cceeb3452f1a474caa729442569a4646056d Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Fri, 18 Oct 2024 16:57:13 +0000 Subject: [PATCH 076/131] chore: update Mathlib dependencies 2024-10-18 (#17922) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index b86cb2dfc481e..c93ba44db832e 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -15,7 +15,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "fa3d73a2cf077f4b14c7840352ac7b08aeb6eb41", + "rev": "1357f4f49450abb9dfd4783e38219f4ce84f9785", "name": "Qq", "manifestFile": "lake-manifest.json", "inputRev": "master", From 1ef8dd30084da1bb42462baaf2dbfcafc07701c9 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Fri, 18 Oct 2024 17:35:48 +0000 Subject: [PATCH 077/131] feat (RingTheory/HahnSeries/Summable) : singleton summable family (#16871) This PR defines a summable family made of a single Hahn series, and shows that the formal sum of the family is equal to the original Hahn series. --- Mathlib/RingTheory/HahnSeries/Summable.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index 45a21761c48c3..a4a57ee6c1624 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -178,6 +178,19 @@ theorem hsum_add {s t : SummableFamily Γ R α} : (s + t).hsum = s.hsum + t.hsum simp only [hsum_coeff, add_coeff, add_apply] exact finsum_add_distrib (s.finite_co_support _) (t.finite_co_support _) +/-- The summable family made of a single Hahn series. -/ +@[simps] +def single (x : HahnSeries Γ R) : SummableFamily Γ R Unit where + toFun _ := x + isPWO_iUnion_support' := + Eq.mpr (congrArg (fun s ↦ s.IsPWO) (Set.iUnion_const x.support)) x.isPWO_support + finite_co_support' g := Set.toFinite {a | ((fun _ ↦ x) a).coeff g ≠ 0} + +@[simp] +theorem hsum_single (x : HahnSeries Γ R) : (single x).hsum = x := by + ext g + simp only [hsum_coeff, single_toFun, finsum_unique] + /-- A summable family induced by an equivalence of the parametrizing type. -/ @[simps] def Equiv (e : α ≃ β) (s : SummableFamily Γ R α) : SummableFamily Γ R β where From 75204348c7d435ac1d22d07c32f5d2836d9548ed Mon Sep 17 00:00:00 2001 From: Thmoas-Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Fri, 18 Oct 2024 17:54:26 +0000 Subject: [PATCH 078/131] feat(FieldTheory/Galois): Finite galois subextensions (#16978) Define the type of finite Galois subextensions and prove basic properties about it. Co-authored-by: Yuyang Zhao Nailin Guan <3571819596@qq.com> Jujian Zhang Moves: - Mathlib.FieldTheory.Galois -> Mathlib.FieldTheory.Galois.Basic --- Mathlib.lean | 1 + Mathlib/FieldTheory/Galois/Infinite.lean | 214 +++++++++++++++++++++++ Mathlib/FieldTheory/KrullTopology.lean | 11 ++ Mathlib/FieldTheory/Normal.lean | 45 +++++ Mathlib/FieldTheory/NormalClosure.lean | 17 ++ 5 files changed, 288 insertions(+) create mode 100644 Mathlib/FieldTheory/Galois/Infinite.lean diff --git a/Mathlib.lean b/Mathlib.lean index b0b4fbb4d3724..c66a3419a9aeb 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2724,6 +2724,7 @@ import Mathlib.FieldTheory.Finite.Trace import Mathlib.FieldTheory.Finiteness import Mathlib.FieldTheory.Fixed import Mathlib.FieldTheory.Galois.Basic +import Mathlib.FieldTheory.Galois.Infinite import Mathlib.FieldTheory.IntermediateField.Algebraic import Mathlib.FieldTheory.IntermediateField.Basic import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure diff --git a/Mathlib/FieldTheory/Galois/Infinite.lean b/Mathlib/FieldTheory/Galois/Infinite.lean new file mode 100644 index 0000000000000..d3c2345887e0d --- /dev/null +++ b/Mathlib/FieldTheory/Galois/Infinite.lean @@ -0,0 +1,214 @@ +/- +Copyright (c) 2024 Nailin Guan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nailin Guan, Yuyang Zhao, Jujian Zhang +-/ + +import Mathlib.Algebra.Category.Grp.FiniteGrp +import Mathlib.CategoryTheory.Category.Preorder +import Mathlib.FieldTheory.NormalClosure +import Mathlib.FieldTheory.SeparableClosure + +/-! + +# Main definitions and results + +In a field extension `K/k` + +* `FiniteGaloisIntermediateField` : The type of intermediate fields of `K/k` + that are finite and Galois over `k` + +* `adjoin` : The finite Galois intermediate field obtained from the normal closure of adjoining a + finite `s : Set K` to `k`. + +* `finGaloisGroup L` : The (finite) Galois group `Gal(L/k)` associated to a + `L : FiniteGaloisIntermediateField k K` `L`. + +* `finGaloisGroupMap` : For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` + giving the restriction of `Gal(L₁/k)` to `Gal(L₂/k)` + +* `finGaloisGroupFunctor` : Mapping `FiniteGaloisIntermediateField` ordered by inverse inclusion to + its corresponding Galois Group as `FiniteGrp`. + +# TODO + +* `FiniteGaloisIntermediateField` should be a `ConditionallyCompleteLattice` but isn't proved yet. + +-/ + +open CategoryTheory Opposite +open scoped IntermediateField + +variable (k K : Type*) [Field k] [Field K] [Algebra k K] + +/-- The type of intermediate fields of `K/k` that are finite and Galois over `k` -/ +structure FiniteGaloisIntermediateField extends IntermediateField k K where + [finiteDimensional : FiniteDimensional k toIntermediateField] + [isGalois : IsGalois k toIntermediateField] + +namespace FiniteGaloisIntermediateField + +instance : Coe (FiniteGaloisIntermediateField k K) (IntermediateField k K) where + coe := toIntermediateField + +instance : CoeSort (FiniteGaloisIntermediateField k K) (Type _) where + coe L := L.toIntermediateField + +instance (L : FiniteGaloisIntermediateField k K) : FiniteDimensional k L := + L.finiteDimensional + +instance (L : FiniteGaloisIntermediateField k K) : IsGalois k L := L.isGalois + +variable {k K} + +lemma val_injective : Function.Injective (toIntermediateField (k := k) (K := K)) := by + rintro ⟨⟩ ⟨⟩ eq + simpa only [mk.injEq] using eq + +/-- Turns the collection of finite Galois IntermediateFields of `K/k` into a lattice. -/ + +instance (L₁ L₂ : IntermediateField k K) [IsGalois k L₁] [IsGalois k L₂] : + IsGalois k ↑(L₁ ⊔ L₂) where + +instance (L₁ L₂ : IntermediateField k K) [FiniteDimensional k L₁] : + FiniteDimensional k ↑(L₁ ⊓ L₂) := + .of_injective (IntermediateField.inclusion inf_le_left).toLinearMap + (IntermediateField.inclusion inf_le_left).injective + +instance (L₁ L₂ : IntermediateField k K) [FiniteDimensional k L₂] : + FiniteDimensional k ↑(L₁ ⊓ L₂) := + .of_injective (IntermediateField.inclusion inf_le_right).toLinearMap + (IntermediateField.inclusion inf_le_right).injective + +instance (L₁ L₂ : IntermediateField k K) [Algebra.IsSeparable k L₁] : + Algebra.IsSeparable k ↑(L₁ ⊓ L₂) := + .of_algHom _ _ (IntermediateField.inclusion inf_le_left) + +instance (L₁ L₂ : IntermediateField k K) [Algebra.IsSeparable k L₂] : + Algebra.IsSeparable k ↑(L₁ ⊓ L₂) := + .of_algHom _ _ (IntermediateField.inclusion inf_le_right) + +instance (L₁ L₂ : IntermediateField k K) [IsGalois k L₁] [IsGalois k L₂] : + IsGalois k ↑(L₁ ⊓ L₂) where + +instance : Sup (FiniteGaloisIntermediateField k K) where + sup L₁ L₂ := .mk <| L₁ ⊔ L₂ + +instance : Inf (FiniteGaloisIntermediateField k K) where + inf L₁ L₂ := .mk <| L₁ ⊓ L₂ + +instance : Lattice (FiniteGaloisIntermediateField k K) := + val_injective.lattice _ (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) + +instance : OrderBot (FiniteGaloisIntermediateField k K) where + bot := .mk ⊥ + bot_le _ := bot_le (α := IntermediateField _ _) + +@[simp] +lemma le_iff (L₁ L₂ : FiniteGaloisIntermediateField k K) : + L₁ ≤ L₂ ↔ L₁.toIntermediateField ≤ L₂.toIntermediateField := + Iff.rfl + +variable (k) in +/-- The minimal (finite) Galois intermediate field containing a finite set `s : Set K` in a +Galois extension `K/k` defined as the the normal closure of the field obtained by adjoining +the set `s : Set K` to `k`. -/ +noncomputable def adjoin [IsGalois k K] (s : Set K) [Finite s] : + FiniteGaloisIntermediateField k K := { + normalClosure k (IntermediateField.adjoin k (s : Set K)) K with + finiteDimensional := + letI : FiniteDimensional k (IntermediateField.adjoin k (s : Set K)) := + IntermediateField.finiteDimensional_adjoin <| fun z _ => + IsAlgebraic.isIntegral (Algebra.IsAlgebraic.isAlgebraic z) + normalClosure.is_finiteDimensional k (IntermediateField.adjoin k (s : Set K)) K + isGalois := IsGalois.normalClosure k (IntermediateField.adjoin k (s : Set K)) K } + +@[simp] +lemma adjoin_val [IsGalois k K] (s : Set K) [Finite s] : + (FiniteGaloisIntermediateField.adjoin k s) = + normalClosure k (IntermediateField.adjoin k s) K := + rfl + +variable (k) in +lemma subset_adjoin [IsGalois k K] (s : Set K) [Finite s] : + s ⊆ (adjoin k s).toIntermediateField := + (IntermediateField.subset_adjoin k s).trans (IntermediateField.le_normalClosure _) + +theorem adjoin_simple_le_iff [IsGalois k K] {x : K} {L : FiniteGaloisIntermediateField k K} : + adjoin k {x} ≤ L ↔ x ∈ L.toIntermediateField := by + simp only [le_iff, adjoin_val, IntermediateField.normalClosure_le_iff_of_normal, + IntermediateField.adjoin_le_iff, Set.le_eq_subset, Set.singleton_subset_iff, SetLike.mem_coe] + +@[simp] +theorem adjoin_map [IsGalois k K] (f : K →ₐ[k] K) (s : Set K) [Finite s] : + adjoin k (f '' s) = adjoin k s := by + apply val_injective; dsimp [adjoin_val] + rw [← IntermediateField.adjoin_map, IntermediateField.normalClosure_map_eq] + +@[simp] +theorem adjoin_simple_map_algHom [IsGalois k K] (f : K →ₐ[k] K) (x : K) : + adjoin k {f x} = adjoin k {x} := by + simpa only [Set.image_singleton] using adjoin_map f { x } + +@[simp] +theorem adjoin_simple_map_algEquiv [IsGalois k K] (f : K ≃ₐ[k] K) (x : K) : + adjoin k {f x} = adjoin k {x} := + adjoin_simple_map_algHom (f : K →ₐ[k] K) x + +end FiniteGaloisIntermediateField + +section Profinite + +variable {k K} + +/-- The (finite) Galois group `Gal(L / k)` associated to a +`L : FiniteGaloisIntermediateField k K` `L`. -/ +def FiniteGaloisIntermediateField.finGaloisGroup (L : FiniteGaloisIntermediateField k K) : + FiniteGrp := + letI := AlgEquiv.fintype k L + FiniteGrp.of <| L ≃ₐ[k] L + +/-- For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` + the restriction homomorphism from `Gal(L₁/k)` to `Gal(L₂/k)` -/ +noncomputable def finGaloisGroupMap {L₁ L₂ : (FiniteGaloisIntermediateField k K)ᵒᵖ} + (le : L₁ ⟶ L₂) : L₁.unop.finGaloisGroup ⟶ L₂.unop.finGaloisGroup := + haveI : Normal k L₂.unop := IsGalois.to_normal + letI : Algebra L₂.unop L₁.unop := RingHom.toAlgebra (Subsemiring.inclusion <| leOfHom le.1) + haveI : IsScalarTower k L₂.unop L₁.unop := IsScalarTower.of_algebraMap_eq (congrFun rfl) + FiniteGrp.ofHom (AlgEquiv.restrictNormalHom L₂.unop) + +namespace finGaloisGroupMap + +@[simp] +lemma map_id (L : (FiniteGaloisIntermediateField k K)ᵒᵖ) : + (finGaloisGroupMap (𝟙 L)) = 𝟙 L.unop.finGaloisGroup := + AlgEquiv.restrictNormalHom_id _ _ + +@[simp] +lemma map_comp {L₁ L₂ L₃ : (FiniteGaloisIntermediateField k K)ᵒᵖ} (f : L₁ ⟶ L₂) (g : L₂ ⟶ L₃) : + finGaloisGroupMap (f ≫ g) = finGaloisGroupMap f ≫ finGaloisGroupMap g := by + iterate 2 + induction L₁ with | _ L₁ => ?_ + induction L₂ with | _ L₂ => ?_ + induction L₃ with | _ L₃ => ?_ + letI : Algebra L₃ L₂ := RingHom.toAlgebra (Subsemiring.inclusion g.unop.le) + letI : Algebra L₂ L₁ := RingHom.toAlgebra (Subsemiring.inclusion f.unop.le) + letI : Algebra L₃ L₁ := RingHom.toAlgebra (Subsemiring.inclusion (g.unop.le.trans f.unop.le)) + haveI : IsScalarTower k L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower k L₃ L₁ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower k L₃ L₂ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower L₃ L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl + apply IsScalarTower.AlgEquiv.restrictNormalHom_comp k L₃ L₂ L₁ + +end finGaloisGroupMap + +variable (k K) in +/-- The functor from `FiniteGaloisIntermediateField` (ordered by reverse inclusion) to `FiniteGrp`, +mapping each intermediate field `K/L/k` to `Gal (L/k)`.-/ +noncomputable def finGaloisGroupFunctor : (FiniteGaloisIntermediateField k K)ᵒᵖ ⥤ FiniteGrp where + obj L := L.unop.finGaloisGroup + map := finGaloisGroupMap + map_id := finGaloisGroupMap.map_id + map_comp := finGaloisGroupMap.map_comp + +end Profinite diff --git a/Mathlib/FieldTheory/KrullTopology.lean b/Mathlib/FieldTheory/KrullTopology.lean index 7f7f53c3109a7..c5751d36c7ff4 100644 --- a/Mathlib/FieldTheory/KrullTopology.lean +++ b/Mathlib/FieldTheory/KrullTopology.lean @@ -178,6 +178,17 @@ instance krullTopology (K L : Type*) [Field K] [Field L] [Algebra K L] : instance (K L : Type*) [Field K] [Field L] [Algebra K L] : TopologicalGroup (L ≃ₐ[K] L) := GroupFilterBasis.isTopologicalGroup (galGroupBasis K L) +open scoped Topology in +lemma krullTopology_mem_nhds_one (K L : Type*) [Field K] [Field L] [Algebra K L] + (s : Set (L ≃ₐ[K] L)) : s ∈ 𝓝 1 ↔ ∃ E : IntermediateField K L, + FiniteDimensional K E ∧ (E.fixingSubgroup : Set (L ≃ₐ[K] L)) ⊆ s := by + rw [GroupFilterBasis.nhds_one_eq] + constructor + · rintro ⟨-, ⟨-, ⟨E, fin, rfl⟩, rfl⟩, hE⟩ + exact ⟨E, fin, hE⟩ + · rintro ⟨E, fin, hE⟩ + exact ⟨E.fixingSubgroup, ⟨E.fixingSubgroup, ⟨E, fin, rfl⟩, rfl⟩, hE⟩ + section KrullT2 open scoped Topology Filter diff --git a/Mathlib/FieldTheory/Normal.lean b/Mathlib/FieldTheory/Normal.lean index 98ae051a5fd57..242f0ef840a7d 100644 --- a/Mathlib/FieldTheory/Normal.lean +++ b/Mathlib/FieldTheory/Normal.lean @@ -355,10 +355,55 @@ theorem AlgEquiv.restrict_liftNormal (χ : K₁ ≃ₐ[F] K₁) [Normal F K₁] (algebraMap K₁ E).injective (Eq.trans (AlgEquiv.restrictNormal_commutes _ K₁ x) (χ.liftNormal_commutes E x)) +/-- The group homomorphism given by restricting an algebra isomorphism to a normal subfield +is surjective. -/ theorem AlgEquiv.restrictNormalHom_surjective [Normal F K₁] [Normal F E] : Function.Surjective (AlgEquiv.restrictNormalHom K₁ : (E ≃ₐ[F] E) → K₁ ≃ₐ[F] K₁) := fun χ => ⟨χ.liftNormal E, χ.restrict_liftNormal E⟩ +/-- The group homomorphism given by restricting an algebra isomorphism to itself +is the identity map. -/ +@[simp] +theorem AlgEquiv.restrictNormalHom_id (F K : Type*) + [Field F] [Field K] [Algebra F K] [Normal F K] : + AlgEquiv.restrictNormalHom K = MonoidHom.id (K ≃ₐ[F] K) := by + ext f x + dsimp only [restrictNormalHom, MonoidHom.mk'_apply, MonoidHom.id_apply] + apply (algebraMap K K).injective + rw [AlgEquiv.restrictNormal_commutes] + simp only [Algebra.id.map_eq_id, RingHom.id_apply] + +namespace IsScalarTower + +/-- In a scalar tower `K₃/K₂/K₁/F` with `K₁` and `K₂` are normal over `F`, the group homomorphism +given by the restriction of algebra isomorphisms of `K₃` to `K₁` is equal to the composition of +the group homomorphism given by the restricting an algebra isomorphism of `K₃` to `K₂` and +the group homomorphism given by the restricting an algebra isomorphism of `K₂` to `K₁` -/ +theorem AlgEquiv.restrictNormalHom_comp (F K₁ K₂ K₃ : Type*) + [Field F] [Field K₁] [Field K₂] [Field K₃] + [Algebra F K₁] [Algebra F K₂] [Algebra F K₃] [Algebra K₁ K₂] [Algebra K₁ K₃] [Algebra K₂ K₃] + [IsScalarTower F K₁ K₃] [IsScalarTower F K₁ K₂] [IsScalarTower F K₂ K₃] [IsScalarTower K₁ K₂ K₃] + [Normal F K₁] [Normal F K₂] : + AlgEquiv.restrictNormalHom K₁ = + (AlgEquiv.restrictNormalHom K₁).comp + (AlgEquiv.restrictNormalHom (F := F) (K₁ := K₃) K₂) := by + ext f x + apply (algebraMap K₁ K₃).injective + rw [IsScalarTower.algebraMap_eq K₁ K₂ K₃] + simp only [AlgEquiv.restrictNormalHom, MonoidHom.mk'_apply, RingHom.coe_comp, Function.comp_apply, + ← algebraMap_apply, AlgEquiv.restrictNormal_commutes, MonoidHom.coe_comp] + +theorem AlgEquiv.restrictNormalHom_comp_apply (K₁ K₂ : Type*) {F K₃ : Type*} + [Field F] [Field K₁] [Field K₂] [Field K₃] + [Algebra F K₁] [Algebra F K₂] [Algebra F K₃] [Algebra K₁ K₂] [Algebra K₁ K₃] [Algebra K₂ K₃] + [IsScalarTower F K₁ K₃] [IsScalarTower F K₁ K₂] [IsScalarTower F K₂ K₃] [IsScalarTower K₁ K₂ K₃] + [Normal F K₁] [Normal F K₂] (f : K₃ ≃ₐ[F] K₃) : + AlgEquiv.restrictNormalHom K₁ f = + (AlgEquiv.restrictNormalHom K₁) (AlgEquiv.restrictNormalHom K₂ f) := by + rw [IsScalarTower.AlgEquiv.restrictNormalHom_comp F K₁ K₂ K₃, MonoidHom.comp_apply] + +end IsScalarTower + open IntermediateField in theorem Normal.minpoly_eq_iff_mem_orbit [h : Normal F E] {x y : E} : minpoly F x = minpoly F y ↔ x ∈ MulAction.orbit (E ≃ₐ[F] E) y := by diff --git a/Mathlib/FieldTheory/NormalClosure.lean b/Mathlib/FieldTheory/NormalClosure.lean index 680dc94bf8709..fac28822a66a5 100644 --- a/Mathlib/FieldTheory/NormalClosure.lean +++ b/Mathlib/FieldTheory/NormalClosure.lean @@ -282,4 +282,21 @@ lemma normal_iff_forall_map_eq : Normal F K ↔ ∀ σ : L →ₐ[F] L, K.map σ lemma normal_iff_forall_map_eq' : Normal F K ↔ ∀ σ : L ≃ₐ[F] L, K.map ↑σ = K := ⟨fun h σ ↦ normal_iff_forall_map_eq.1 h σ, fun h ↦ normal_iff_forall_map_le'.2 (fun σ ↦ (h σ).le)⟩ +@[simp] +lemma normalClosure_map_eq (K : IntermediateField F L) (σ : L →ₐ[F] L) : + normalClosure F (K.map σ) L = normalClosure F K L := by + have (σ : L ≃ₐ[F] L) : normalClosure F (K.map (σ : L →ₐ[F] L)) L = normalClosure F K L := by + simp_rw [normalClosure_def'', map_map] + exact (Equiv.mulRight σ).iSup_congr fun _ ↦ rfl + exact this ((Algebra.IsAlgebraic.algEquivEquivAlgHom _ _).symm σ) + +@[simp] +theorem normalClosure_le_iff_of_normal {K₁ K₂ : IntermediateField F L} [Normal F K₂] : + normalClosure F K₁ L ≤ K₂ ↔ K₁ ≤ K₂ := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rw [normalClosure_le_iff] at h + simpa only [fieldRange_val] using h K₁.val + · rw [← normalClosure_of_normal K₂] + exact normalClosure_mono K₁ K₂ h + end IntermediateField From 8610319f04e6b637e3c597905fa5cb405e5549fe Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 18 Oct 2024 18:06:40 +0000 Subject: [PATCH 079/131] =?UTF-8?q?doc(SetTheory/Game/PGame):=20expand=20d?= =?UTF-8?q?ocstrings=20for=20`=E2=89=A4`=20and=20`=E2=A7=8F`=20(#16809)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/SetTheory/Game/PGame.lean | 11 ++++++++--- 1 file changed, 8 insertions(+), 3 deletions(-) diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index afa7e0b220917..83f9d1f8d2559 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -23,6 +23,9 @@ takes two types (thought of as indexing the possible moves for the players Left pair of functions out of these types to `SetTheory.PGame` (thought of as describing the resulting game after making a move). +We may denote a game as $\{L | R\}$, where $L$ and $R$ stand for the collections of left and right +moves. This notation is not currently used in Mathlib. + Combinatorial games themselves, as a quotient of pregames, are constructed in `Game.lean`. ## Conway induction @@ -349,15 +352,17 @@ instance isEmpty_one_rightMoves : IsEmpty (RightMoves 1) := /-- The less or equal relation on pre-games. -If `0 ≤ x`, then Left can win `x` as the second player. -/ +If `0 ≤ x`, then Left can win `x` as the second player. `x ≤ y` means that `0 ≤ y - x`. +See `PGame.le_iff_sub_nonneg`. -/ instance le : LE PGame := ⟨Sym2.GameAdd.fix wf_isOption fun x y le => (∀ i, ¬le y (x.moveLeft i) (Sym2.GameAdd.snd_fst <| IsOption.moveLeft i)) ∧ ∀ j, ¬le (y.moveRight j) x (Sym2.GameAdd.fst_snd <| IsOption.moveRight j)⟩ -/-- The less or fuzzy relation on pre-games. +/-- The less or fuzzy relation on pre-games. `x ⧏ y` is defined as `¬ y ≤ x`. -If `0 ⧏ x`, then Left can win `x` as the first player. -/ +If `0 ⧏ x`, then Left can win `x` as the first player. `x ⧏ y` means that `0 ⧏ y - x`. +See `PGame.lf_iff_sub_zero_lf`. -/ def LF (x y : PGame) : Prop := ¬y ≤ x From df601e35b5257a5c391ee5ca5648201c4f46cc78 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Fri, 18 Oct 2024 18:31:29 +0000 Subject: [PATCH 080/131] feat(AlgebraicGeometry/EllipticCurve/Projective): implement group operations on point representatives for projective coordinates (#9418) Define the analogous secant-and-tangent negation `neg` and addition `add` on `PointRep` over `F`, and lift them to `PointClass`. Define `Point` as a `PointClass` that is nonsingular. Prove in `neg_equiv` and `add_equiv` that these operations preserve the equivalence. Prove in `nonsingular_neg` and `nonsingular_add` that these operations preserve nonsingularity by reducing it to the affine case, and lift these proofs to `PointClass`. This is the third in a series of four PRs leading to #8485 --- .../EllipticCurve/Projective.lean | 282 +++++++++++++++++- scripts/no_lints_prime_decls.txt | 1 + 2 files changed, 282 insertions(+), 1 deletion(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean index 3d4ea8337acc9..00a1a91f99235 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean @@ -14,7 +14,8 @@ import Mathlib.Tactic.LinearCombination' This file defines the type of points on a Weierstrass curve as a tuple, consisting of an equivalence class of triples up to scaling by a unit, satisfying a Weierstrass equation with a nonsingular -condition. +condition. This file also defines the negation and addition operations of the group law for this +type, and proves that they respect the Weierstrass equation and the nonsingular condition. ## Mathematical background @@ -32,6 +33,7 @@ given by a tuple consisting of $[x:y:z]$ and the nonsingular condition on any re As in `Mathlib.AlgebraicGeometry.EllipticCurve.Affine`, the set of nonsingular rational points forms an abelian group under the same secant-and-tangent process, but the polynomials involved are homogeneous, and any instances of division become multiplication in the $Z$-coordinate. +Note that most computational proofs follow from their analogous proofs for affine coordinates. ## Main definitions @@ -39,10 +41,16 @@ homogeneous, and any instances of division become multiplication in the $Z$-coor * `WeierstrassCurve.Projective.toAffine`: the Weierstrass curve in affine coordinates. * `WeierstrassCurve.Projective.Nonsingular`: the nonsingular condition on a point representative. * `WeierstrassCurve.Projective.NonsingularLift`: the nonsingular condition on a point class. + * `WeierstrassCurve.Projective.neg`: the negation operation on a point representative. + * `WeierstrassCurve.Projective.negMap`: the negation operation on a point class. + * `WeierstrassCurve.Projective.add`: the addition operation on a point representative. + * `WeierstrassCurve.Projective.addMap`: the addition operation on a point class. ## Main statements * `WeierstrassCurve.Projective.polynomial_relation`: Euler's homogeneous function theorem. + * `WeierstrassCurve.Projective.nonsingular_neg`: negation preserves the nonsingular condition. + * `WeierstrassCurve.Projective.nonsingular_add`: addition preserves the nonsingular condition. ## Implementation notes @@ -138,6 +146,11 @@ lemma smul_equiv (P : Fin 3 → R) {u : R} (hu : IsUnit u) : u • P ≈ P := lemma smul_eq (P : Fin 3 → R) {u : R} (hu : IsUnit u) : (⟦u • P⟧ : PointClass R) = ⟦P⟧ := Quotient.eq.mpr <| smul_equiv P hu +lemma smul_equiv_smul (P Q : Fin 3 → R) {u v : R} (hu : IsUnit u) (hv : IsUnit v) : + u • P ≈ v • Q ↔ P ≈ Q := by + erw [← Quotient.eq_iff_equiv, ← Quotient.eq_iff_equiv, smul_eq P hu, smul_eq Q hv] + rfl + variable (W') in /-- The coercion to a Weierstrass curve in affine coordinates. -/ abbrev toAffine : Affine R := @@ -1163,4 +1176,271 @@ lemma addXYZ_of_Z_ne_zero {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equati end Addition +section Negation + +/-! ### Negation on point representatives -/ + +variable (W') in +/-- The negation of a point representative. -/ +def neg (P : Fin 3 → R) : Fin 3 → R := + ![P x, W'.negY P, P z] + +lemma neg_X (P : Fin 3 → R) : W'.neg P x = P x := + rfl + +lemma neg_Y (P : Fin 3 → R) : W'.neg P y = W'.negY P := + rfl + +lemma neg_Z (P : Fin 3 → R) : W'.neg P z = P z := + rfl + +lemma neg_smul (P : Fin 3 → R) (u : R) : W'.neg (u • P) = u • W'.neg P := by + simpa only [neg, negY_smul] using (smul_fin3 (W'.neg P) u).symm + +lemma neg_smul_equiv (P : Fin 3 → R) {u : R} (hu : IsUnit u) : W'.neg (u • P) ≈ W'.neg P := + ⟨hu.unit, (neg_smul ..).symm⟩ + +lemma neg_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W'.neg P ≈ W'.neg Q := by + rcases h with ⟨u, rfl⟩ + exact neg_smul_equiv Q u.isUnit + +lemma neg_of_Z_eq_zero [NoZeroDivisors R] {P : Fin 3 → R} (hP : W'.Equation P) (hPz : P z = 0) : + W'.neg P = -P y • ![0, 1, 0] := by + erw [neg, X_eq_zero_of_Z_eq_zero hP hPz, negY_of_Z_eq_zero hP hPz, hPz, smul_fin3, mul_zero, + mul_one] + +lemma neg_of_Z_ne_zero {P : Fin 3 → F} (hPz : P z ≠ 0) : + W.neg P = P z • ![P x / P z, W.toAffine.negY (P x / P z) (P y / P z), 1] := by + erw [neg, smul_fin3, mul_div_cancel₀ _ hPz, ← negY_of_Z_ne_zero hPz, mul_div_cancel₀ _ hPz, + mul_one] + +private lemma nonsingular_neg_of_Z_ne_zero {P : Fin 3 → F} (hP : W.Nonsingular P) (hPz : P z ≠ 0) : + W.Nonsingular ![P x / P z, W.toAffine.negY (P x / P z) (P y / P z), 1] := + (nonsingular_some ..).mpr <| Affine.nonsingular_neg <| (nonsingular_of_Z_ne_zero hPz).mp hP + +lemma nonsingular_neg {P : Fin 3 → F} (hP : W.Nonsingular P) : W.Nonsingular <| W.neg P := by + by_cases hPz : P z = 0 + · simp only [neg_of_Z_eq_zero hP.left hPz, nonsingular_smul _ (isUnit_Y_of_Z_eq_zero hP hPz).neg, + nonsingular_zero] + · simp only [neg_of_Z_ne_zero hPz, nonsingular_smul _ <| Ne.isUnit hPz, + nonsingular_neg_of_Z_ne_zero hP hPz] + +lemma addZ_neg (P : Fin 3 → R) : W'.addZ P (W'.neg P) = 0 := by + rw [addZ, neg_X, neg_Y, neg_Z, negY] + ring1 + +lemma addX_neg (P : Fin 3 → R) : W'.addX P (W'.neg P) = 0 := by + rw [addX, neg_X, neg_Y, neg_Z, negY] + ring1 + +lemma negAddY_neg {P : Fin 3 → R} (hP : W'.Equation P) : W'.negAddY P (W'.neg P) = W'.dblZ P := by + linear_combination (norm := (rw [negAddY, neg_X, neg_Y, neg_Z, dblZ, negY]; ring1)) + -3 * (P y - W'.negY P) * (equation_iff _).mp hP + +lemma addY_neg {P : Fin 3 → R} (hP : W'.Equation P) : W'.addY P (W'.neg P) = -W'.dblZ P := by + rw [addY, negY_eq, addX_neg, negAddY_neg hP, addZ_neg, mul_zero, sub_zero, mul_zero, sub_zero] + +lemma addXYZ_neg {P : Fin 3 → R} (hP : W'.Equation P) : + W'.addXYZ P (W'.neg P) = -W'.dblZ P • ![0, 1, 0] := by + erw [addXYZ, addX_neg, addY_neg hP, addZ_neg, smul_fin3, mul_zero, mul_one] + +variable (W') in +/-- The negation of a point class. If `P` is a point representative, +then `W'.negMap ⟦P⟧` is definitionally equivalent to `W'.neg P`. -/ +def negMap (P : PointClass R) : PointClass R := + P.map W'.neg fun _ _ => neg_equiv + +lemma negMap_eq (P : Fin 3 → R) : W'.negMap ⟦P⟧ = ⟦W'.neg P⟧ := + rfl + +lemma negMap_of_Z_eq_zero {P : Fin 3 → F} (hP : W.Nonsingular P) (hPz : P z = 0) : + W.negMap ⟦P⟧ = ⟦![0, 1, 0]⟧ := by + rw [negMap_eq, neg_of_Z_eq_zero hP.left hPz, smul_eq _ (isUnit_Y_of_Z_eq_zero hP hPz).neg] + +lemma negMap_of_Z_ne_zero {P : Fin 3 → F} (hPz : P z ≠ 0) : + W.negMap ⟦P⟧ = ⟦![P x / P z, W.toAffine.negY (P x / P z) (P y / P z), 1]⟧ := by + rw [negMap_eq, neg_of_Z_ne_zero hPz, smul_eq _ <| Ne.isUnit hPz] + +lemma nonsingularLift_negMap {P : PointClass F} (hP : W.NonsingularLift P) : + W.NonsingularLift <| W.negMap P := by + rcases P with ⟨_⟩ + exact nonsingular_neg hP + +end Negation + +section Addition + +/-! ### Addition on point representatives -/ + +open Classical in +variable (W') in +/-- The addition of two point representatives. -/ +noncomputable def add (P Q : Fin 3 → R) : Fin 3 → R := + if P ≈ Q then W'.dblXYZ P else W'.addXYZ P Q + +lemma add_of_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W'.add P Q = W'.dblXYZ P := + if_pos h + +lemma add_smul_of_equiv {P Q : Fin 3 → R} (h : P ≈ Q) {u v : R} (hu : IsUnit u) (hv : IsUnit v) : + W'.add (u • P) (v • Q) = u ^ 4 • W'.add P Q := by + rw [add_of_equiv <| (smul_equiv_smul P Q hu hv).mpr h, dblXYZ_smul, add_of_equiv h] + +lemma add_self (P : Fin 3 → R) : W'.add P P = W'.dblXYZ P := + add_of_equiv <| Setoid.refl _ + +lemma add_of_eq {P Q : Fin 3 → R} (h : P = Q) : W'.add P Q = W'.dblXYZ P := + h ▸ add_self P + +lemma add_of_not_equiv {P Q : Fin 3 → R} (h : ¬P ≈ Q) : W'.add P Q = W'.addXYZ P Q := + if_neg h + +lemma add_smul_of_not_equiv {P Q : Fin 3 → R} (h : ¬P ≈ Q) {u v : R} (hu : IsUnit u) + (hv : IsUnit v) : W'.add (u • P) (v • Q) = (u * v) ^ 2 • W'.add P Q := by + rw [add_of_not_equiv <| h.comp (smul_equiv_smul P Q hu hv).mp, addXYZ_smul, add_of_not_equiv h] + +lemma add_smul_equiv (P Q : Fin 3 → R) {u v : R} (hu : IsUnit u) (hv : IsUnit v) : + W'.add (u • P) (v • Q) ≈ W'.add P Q := by + by_cases h : P ≈ Q + · exact ⟨hu.unit ^ 4, by convert (add_smul_of_equiv h hu hv).symm⟩ + · exact ⟨(hu.unit * hv.unit) ^ 2, by convert (add_smul_of_not_equiv h hu hv).symm⟩ + +lemma add_equiv {P P' Q Q' : Fin 3 → R} (hP : P ≈ P') (hQ : Q ≈ Q') : + W'.add P Q ≈ W'.add P' Q' := by + rcases hP, hQ with ⟨⟨u, rfl⟩, ⟨v, rfl⟩⟩ + exact add_smul_equiv P' Q' u.isUnit v.isUnit + +lemma add_of_Z_eq_zero {P Q : Fin 3 → F} (hP : W.Nonsingular P) (hQ : W.Nonsingular Q) + (hPz : P z = 0) (hQz : Q z = 0) : W.add P Q = P y ^ 4 • ![0, 1, 0] := by + rw [add, if_pos <| equiv_of_Z_eq_zero hP hQ hPz hQz, dblXYZ_of_Z_eq_zero hP.left hPz] + +lemma add_of_Z_eq_zero_left [NoZeroDivisors R] {P Q : Fin 3 → R} (hP : W'.Equation P) + (hPz : P z = 0) (hQz : Q z ≠ 0) : W'.add P Q = (P y ^ 2 * Q z) • Q := by + rw [add, if_neg <| not_equiv_of_Z_eq_zero_left hPz hQz, addXYZ_of_Z_eq_zero_left hP hPz] + +lemma add_of_Z_eq_zero_right [NoZeroDivisors R] {P Q : Fin 3 → R} (hQ : W'.Equation Q) + (hPz : P z ≠ 0) (hQz : Q z = 0) : W'.add P Q = -(Q y ^ 2 * P z) • P := by + rw [add, if_neg <| not_equiv_of_Z_eq_zero_right hPz hQz, addXYZ_of_Z_eq_zero_right hQ hQz] + +lemma add_of_Y_eq {P Q : Fin 3 → F} (hP : W.Equation P) (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z = Q x * P z) (hy : P y * Q z = Q y * P z) (hy' : P y * Q z = W.negY Q * P z) : + W.add P Q = W.dblU P • ![0, 1, 0] := by + rw [add, if_pos <| equiv_of_X_eq_of_Y_eq hPz hQz hx hy, dblXYZ_of_Y_eq hP hPz hQz hx hy hy'] + +lemma add_of_Y_ne {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equation Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hx : P x * Q z = Q x * P z) (hy : P y * Q z ≠ Q y * P z) : + W.add P Q = addU P Q • ![0, 1, 0] := by + rw [add, if_neg <| not_equiv_of_Y_ne hy, addXYZ_of_X_eq hP hQ hPz hQz hx] + +lemma add_of_Y_ne' {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equation Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hx : P x * Q z = Q x * P z) (hy : P y * Q z ≠ W.negY Q * P z) : + W.add P Q = W.dblZ P • + ![W.toAffine.addX (P x / P z) (Q x / Q z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), + W.toAffine.addY (P x / P z) (Q x / Q z) (P y / P z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), 1] := by + rw [add, if_pos <| equiv_of_X_eq_of_Y_eq hPz hQz hx <| Y_eq_of_Y_ne' hP hQ hPz hQz hx hy, + dblXYZ_of_Z_ne_zero hP hQ hPz hQz hx hy] + +lemma add_of_X_ne {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equation Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hx : P x * Q z ≠ Q x * P z) : W.add P Q = W.addZ P Q • + ![W.toAffine.addX (P x / P z) (Q x / Q z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), + W.toAffine.addY (P x / P z) (Q x / Q z) (P y / P z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), 1] := by + rw [add, if_neg <| not_equiv_of_X_ne hx, addXYZ_of_Z_ne_zero hP hQ hPz hQz hx] + +private lemma nonsingular_add_of_Z_ne_zero {P Q : Fin 3 → F} (hP : W.Nonsingular P) + (hQ : W.Nonsingular Q) (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hxy : P x * Q z = Q x * P z → P y * Q z ≠ W.negY Q * P z) : W.Nonsingular + ![W.toAffine.addX (P x / P z) (Q x / Q z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), + W.toAffine.addY (P x / P z) (Q x / Q z) (P y / P z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), 1] := + (nonsingular_some ..).mpr <| Affine.nonsingular_add ((nonsingular_of_Z_ne_zero hPz).mp hP) + ((nonsingular_of_Z_ne_zero hQz).mp hQ) (by rwa [← X_eq_iff hPz hQz, ne_eq, ← Y_eq_iff' hPz hQz]) + +lemma nonsingular_add {P Q : Fin 3 → F} (hP : W.Nonsingular P) (hQ : W.Nonsingular Q) : + W.Nonsingular <| W.add P Q := by + by_cases hPz : P z = 0 + · by_cases hQz : Q z = 0 + · simp only [add_of_Z_eq_zero hP hQ hPz hQz, + nonsingular_smul _ <| (isUnit_Y_of_Z_eq_zero hP hPz).pow 4, nonsingular_zero] + · simpa only [add_of_Z_eq_zero_left hP.left hPz hQz, + nonsingular_smul _ <| ((isUnit_Y_of_Z_eq_zero hP hPz).pow 2).mul <| Ne.isUnit hQz] + · by_cases hQz : Q z = 0 + · simpa only [add_of_Z_eq_zero_right hQ.left hPz hQz, + nonsingular_smul _ (((isUnit_Y_of_Z_eq_zero hQ hQz).pow 2).mul <| Ne.isUnit hPz).neg] + · by_cases hxy : P x * Q z = Q x * P z → P y * Q z ≠ W.negY Q * P z + · by_cases hx : P x * Q z = Q x * P z + · simp only [add_of_Y_ne' hP.left hQ.left hPz hQz hx <| hxy hx, + nonsingular_smul _ <| isUnit_dblZ_of_Y_ne' hP.left hQ.left hPz hQz hx <| hxy hx, + nonsingular_add_of_Z_ne_zero hP hQ hPz hQz hxy] + · simp only [add_of_X_ne hP.left hQ.left hPz hQz hx, + nonsingular_smul _ <| isUnit_addZ_of_X_ne hP.left hQ.left hx, + nonsingular_add_of_Z_ne_zero hP hQ hPz hQz hxy] + · rw [_root_.not_imp, not_ne_iff] at hxy + by_cases hy : P y * Q z = Q y * P z + · simp only [add_of_Y_eq hP.left hPz hQz hxy.left hy hxy.right, nonsingular_smul _ <| + isUnit_dblU_of_Y_eq hP hPz hQz hxy.left hy hxy.right, nonsingular_zero] + · simp only [add_of_Y_ne hP.left hQ.left hPz hQz hxy.left hy, + nonsingular_smul _ <| isUnit_addU_of_Y_ne hPz hQz hy, nonsingular_zero] + +variable (W') in +/-- The addition of two point classes. If `P` is a point representative, +then `W.addMap ⟦P⟧ ⟦Q⟧` is definitionally equivalent to `W.add P Q`. -/ +noncomputable def addMap (P Q : PointClass R) : PointClass R := + Quotient.map₂ W'.add (fun _ _ hP _ _ hQ => add_equiv hP hQ) P Q + +lemma addMap_eq (P Q : Fin 3 → R) : W'.addMap ⟦P⟧ ⟦Q⟧ = ⟦W'.add P Q⟧ := + rfl + +lemma addMap_of_Z_eq_zero_left {P : Fin 3 → F} {Q : PointClass F} (hP : W.Nonsingular P) + (hQ : W.NonsingularLift Q) (hPz : P z = 0) : W.addMap ⟦P⟧ Q = Q := by + rcases Q with ⟨Q⟩ + by_cases hQz : Q z = 0 + · erw [addMap_eq, add_of_Z_eq_zero hP hQ hPz hQz, + smul_eq _ <| (isUnit_Y_of_Z_eq_zero hP hPz).pow 4, Quotient.eq] + exact Setoid.symm <| equiv_zero_of_Z_eq_zero hQ hQz + · erw [addMap_eq, add_of_Z_eq_zero_left hP.left hPz hQz, + smul_eq _ <| ((isUnit_Y_of_Z_eq_zero hP hPz).pow 2).mul <| Ne.isUnit hQz] + rfl + +lemma addMap_of_Z_eq_zero_right {P : PointClass F} {Q : Fin 3 → F} (hP : W.NonsingularLift P) + (hQ : W.Nonsingular Q) (hQz : Q z = 0) : W.addMap P ⟦Q⟧ = P := by + rcases P with ⟨P⟩ + by_cases hPz : P z = 0 + · erw [addMap_eq, add_of_Z_eq_zero hP hQ hPz hQz, + smul_eq _ <| (isUnit_Y_of_Z_eq_zero hP hPz).pow 4, Quotient.eq] + exact Setoid.symm <| equiv_zero_of_Z_eq_zero hP hPz + · erw [addMap_eq, add_of_Z_eq_zero_right hQ.left hPz hQz, + smul_eq _ (((isUnit_Y_of_Z_eq_zero hQ hQz).pow 2).mul <| Ne.isUnit hPz).neg] + rfl + +lemma addMap_of_Y_eq {P Q : Fin 3 → F} (hP : W.Nonsingular P) (hQ : W.Equation Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hx : P x * Q z = Q x * P z) (hy' : P y * Q z = W.negY Q * P z) : + W.addMap ⟦P⟧ ⟦Q⟧ = ⟦![0, 1, 0]⟧ := by + by_cases hy : P y * Q z = Q y * P z + · rw [addMap_eq, add_of_Y_eq hP.left hPz hQz hx hy hy', + smul_eq _ <| isUnit_dblU_of_Y_eq hP hPz hQz hx hy hy'] + · rw [addMap_eq, add_of_Y_ne hP.left hQ hPz hQz hx hy, + smul_eq _ <| isUnit_addU_of_Y_ne hPz hQz hy] + +lemma addMap_of_Z_ne_zero {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equation Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hxy : P x * Q z = Q x * P z → P y * Q z ≠ W.negY Q * P z) : W.addMap ⟦P⟧ ⟦Q⟧ = + ⟦![W.toAffine.addX (P x / P z) (Q x / Q z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), + W.toAffine.addY (P x / P z) (Q x / Q z) (P y / P z) + (W.toAffine.slope (P x / P z) (Q x / Q z) (P y / P z) (Q y / Q z)), 1]⟧ := by + by_cases hx : P x * Q z = Q x * P z + · rw [addMap_eq, add_of_Y_ne' hP hQ hPz hQz hx <| hxy hx, + smul_eq _ <| isUnit_dblZ_of_Y_ne' hP hQ hPz hQz hx <| hxy hx] + · rw [addMap_eq, add_of_X_ne hP hQ hPz hQz hx, smul_eq _ <| isUnit_addZ_of_X_ne hP hQ hx] + +lemma nonsingularLift_addMap {P Q : PointClass F} (hP : W.NonsingularLift P) + (hQ : W.NonsingularLift Q) : W.NonsingularLift <| W.addMap P Q := by + rcases P; rcases Q + exact nonsingular_add hP hQ + +end Addition + end WeierstrassCurve.Projective diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index 8affabbf01733..671479377b629 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -4796,6 +4796,7 @@ WeierstrassCurve.leadingCoeff_preΨ' WeierstrassCurve.map_preΨ' WeierstrassCurve.natDegree_coeff_preΨ' WeierstrassCurve.natDegree_preΨ' +WeierstrassCurve.Projective.add_of_Y_ne' WeierstrassCurve.Projective.addX_eq' WeierstrassCurve.Projective.addY_of_X_eq' WeierstrassCurve.Projective.addZ_eq' From abad5bb64ad342b9f7f122f073225ba5c551664f Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 18 Oct 2024 18:38:50 +0000 Subject: [PATCH 081/131] chore(Data/Nat/Bits): do not use tactic to make definitions (#17926) --- Mathlib/Data/Nat/Bits.lean | 16 +++++++++------- 1 file changed, 9 insertions(+), 7 deletions(-) diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index e580011668c94..63469d742f212 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -375,10 +375,11 @@ def binaryRec' {motive : ℕ → Sort*} (z : motive 0) ∀ n, motive n := binaryRec z fun b n ih => if h : n = 0 → b = true then f b n h ih - else by - convert z - rw [bit_eq_zero_iff] - simpa using h + else + have : bit b n = 0 := by + rw [bit_eq_zero_iff] + cases n <;> cases b <;> simp at h ⊢ + congrArg motive this ▸ z /-- The same as `binaryRec`, but special casing both 0 and 1 as base cases -/ @[elab_as_elim] @@ -386,9 +387,10 @@ def binaryRecFromOne {motive : ℕ → Sort*} (z₀ : motive 0) (z₁ : motive 1 (f : ∀ b n, n ≠ 0 → motive n → motive (bit b n)) : ∀ n, motive n := binaryRec' z₀ fun b n h ih => - if h' : n = 0 then by - rw [h', h h'] - exact z₁ + if h' : n = 0 then + have : bit b n = bit true 0 := by + rw [h', h h'] + congrArg motive this ▸ z₁ else f b n h' ih @[simp] From 7d20facb45c8983b0c9e7d2db53146a6642b8ef8 Mon Sep 17 00:00:00 2001 From: FordUniver Date: Fri, 18 Oct 2024 19:09:03 +0000 Subject: [PATCH 082/131] feat: variant of the binary AM-GM inequality (#17877) added variant of binary AM-GM inequality Co-authored-by: FordUniver <61389961+FordUniver@users.noreply.github.com> --- Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean | 17 +++++++++++++++-- Mathlib/Data/Nat/Cast/Defs.lean | 4 ++++ 2 files changed, 19 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean index 2aa45b63b34c9..9c0bbff7ea773 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean @@ -751,8 +751,8 @@ lemma max_mul_mul_le_max_mul_max [PosMulMono α] [MulPosMono α] (b c : α) (ha mul_le_mul (le_max_right a c) (le_max_right b d) hd (le_trans ha (le_max_left a c)) max_le (by simpa [mul_comm, max_comm] using ba) (by simpa [mul_comm, max_comm] using cd) -/-- Binary **arithmetic mean-geometric mean inequality** (aka AM-GM inequality) for linearly ordered -commutative semirings. -/ +/-- Binary, squared, and division-free **arithmetic mean-geometric mean inequality** +(aka AM-GM inequality) for linearly ordered commutative semirings. -/ lemma two_mul_le_add_sq [ExistsAddOfLE α] [MulPosStrictMono α] [ContravariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· ≤ ·)] (a b : α) : 2 * a * b ≤ a ^ 2 + b ^ 2 := by @@ -761,6 +761,19 @@ lemma two_mul_le_add_sq [ExistsAddOfLE α] [MulPosStrictMono α] alias two_mul_le_add_pow_two := two_mul_le_add_sq +/-- Binary, squared, and division-free **arithmetic mean-geometric mean inequality** +(aka AM-GM inequality) for linearly ordered commutative semirings. -/ +lemma four_mul_le_sq_add [ExistsAddOfLE α] [MulPosStrictMono α] + [ContravariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· ≤ ·)] + (a b : α) : 4 * a * b ≤ (a + b) ^ 2 := by + calc 4 * a * b + _ = 2 * a * b + 2 * a * b := by rw [mul_assoc, two_add_two_eq_four.symm, add_mul, mul_assoc] + _ ≤ a ^ 2 + b ^ 2 + 2 * a * b := by gcongr; exact two_mul_le_add_sq _ _ + _ = a ^ 2 + 2 * a * b + b ^ 2 := by rw [add_right_comm] + _ = (a + b) ^ 2 := (add_sq a b).symm + +alias four_mul_le_pow_two_add := four_mul_le_sq_add + end LinearOrderedCommSemiring section LinearOrderedRing diff --git a/Mathlib/Data/Nat/Cast/Defs.lean b/Mathlib/Data/Nat/Cast/Defs.lean index 937ea2f0f088c..eb1e5e1690360 100644 --- a/Mathlib/Data/Nat/Cast/Defs.lean +++ b/Mathlib/Data/Nat/Cast/Defs.lean @@ -199,3 +199,7 @@ theorem three_add_one_eq_four [AddMonoidWithOne R] : 3 + 1 = (4 : R) := by ← Nat.cast_add, ← Nat.cast_add, ← Nat.cast_add] apply congrArg decide + +theorem two_add_two_eq_four [AddMonoidWithOne R] : 2 + 2 = (4 : R) := by + simp [← one_add_one_eq_two, ← Nat.cast_one, ← three_add_one_eq_four, + ← two_add_one_eq_three, add_assoc] From 63c47e1d61626a0e8d3209728b4a86fc90b87cd7 Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 18 Oct 2024 19:09:04 +0000 Subject: [PATCH 083/131] chore(Data/Nat/Bits): attribute `elab_as_elim` `inline` `specialize` (#17897) Add attribute `elab_as_elim` to binary recursors. Also add `inline` and `specialize` for performance. Without `specialize`, the `popcountTR` in the test is not tail-recursive, and so crashes with ``` error: libc++abi: terminating due to uncaught exception of type lean::throwable: deep recursion was detected at 'interpreter' ``` Co-authored-by: Eric Wieser --- Mathlib/Data/Nat/Bits.lean | 10 ++++++---- Mathlib/Data/Nat/Fib/Basic.lean | 8 ++++---- Mathlib/Data/Nat/Size.lean | 14 +++++++------- test/BinaryRec.lean | 8 ++++++++ 4 files changed, 25 insertions(+), 15 deletions(-) create mode 100644 test/BinaryRec.lean diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 63469d742f212..371661d05995b 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -136,6 +136,7 @@ theorem bit_testBit_zero_shiftRight_one (n : Nat) : bit (n.testBit 0) (n >>> 1) /-- For a predicate `motive : Nat → Sort*`, if instances can be constructed for natural numbers of the form `bit b n`, they can be constructed for any given natural number. -/ +@[inline] def bitCasesOn {motive : Nat → Sort u} (n) (h : ∀ b n, motive (bit b n)) : motive n := -- `1 &&& n != 0` is faster than `n.testBit 0`. This may change when we have faster `testBit`. let x := h (1 &&& n != 0) (n >>> 1) @@ -175,6 +176,7 @@ lemma binaryRec_decreasing (h : n ≠ 0) : div2 n < n := by For a predicate `motive : Nat → Sort*`, if instances can be constructed for natural numbers of the form `bit b n`, they can be constructed for all natural numbers. -/ +@[elab_as_elim, specialize] def binaryRec {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) (n : Nat) : motive n := if n0 : n = 0 then congrArg motive n0 ▸ z @@ -208,7 +210,7 @@ lemma binaryRec_zero {motive : Nat → Sort u} @[simp] lemma binaryRec_one {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) : - binaryRec z f 1 = f true 0 z := by + binaryRec (motive := C) z f 1 = f true 0 z := by rw [binaryRec] simp only [succ_ne_self, ↓reduceDIte, reduceShiftRight, binaryRec_zero] rfl @@ -369,7 +371,7 @@ theorem binaryRec_eq {motive : Nat → Sort u} {z : motive 0} /-- The same as `binaryRec`, but the induction step can assume that if `n=0`, the bit being appended is `true`-/ -@[elab_as_elim] +@[elab_as_elim, specialize] def binaryRec' {motive : ℕ → Sort*} (z : motive 0) (f : ∀ b n, (n = 0 → b = true) → motive n → motive (bit b n)) : ∀ n, motive n := @@ -382,7 +384,7 @@ def binaryRec' {motive : ℕ → Sort*} (z : motive 0) congrArg motive this ▸ z /-- The same as `binaryRec`, but special casing both 0 and 1 as base cases -/ -@[elab_as_elim] +@[elab_as_elim, specialize] def binaryRecFromOne {motive : ℕ → Sort*} (z₀ : motive 0) (z₁ : motive 1) (f : ∀ b n, n ≠ 0 → motive n → motive (bit b n)) : ∀ n, motive n := @@ -399,7 +401,7 @@ theorem zero_bits : bits 0 = [] := by simp [Nat.bits] @[simp] theorem bits_append_bit (n : ℕ) (b : Bool) (hn : n = 0 → b = true) : (bit b n).bits = b :: n.bits := by - rw [Nat.bits, binaryRec_eq'] + rw [Nat.bits, Nat.bits, binaryRec_eq'] simpa @[simp] diff --git a/Mathlib/Data/Nat/Fib/Basic.lean b/Mathlib/Data/Nat/Fib/Basic.lean index 4d5698a009a29..587e609b5ab7b 100644 --- a/Mathlib/Data/Nat/Fib/Basic.lean +++ b/Mathlib/Data/Nat/Fib/Basic.lean @@ -200,12 +200,12 @@ theorem fast_fib_aux_bit_tt (n : ℕ) : · simp theorem fast_fib_aux_eq (n : ℕ) : fastFibAux n = (fib n, fib (n + 1)) := by - apply Nat.binaryRec _ (fun b n' ih => _) n + refine Nat.binaryRec ?_ ?_ n · simp [fastFibAux] · rintro (_|_) n' ih <;> - simp only [fast_fib_aux_bit_ff, fast_fib_aux_bit_tt, congr_arg Prod.fst ih, - congr_arg Prod.snd ih, Prod.mk.inj_iff] <;> - simp [bit, fib_two_mul, fib_two_mul_add_one, fib_two_mul_add_two] + simp only [fast_fib_aux_bit_ff, fast_fib_aux_bit_tt, congr_arg Prod.fst ih, + congr_arg Prod.snd ih, Prod.mk.inj_iff] <;> + simp [bit, fib_two_mul, fib_two_mul_add_one, fib_two_mul_add_two] theorem fast_fib_eq (n : ℕ) : fastFib n = fib n := by rw [fastFib, fast_fib_aux_eq] diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index 5c7eb08714c3c..1fdafff0b8fd0 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -39,7 +39,7 @@ theorem size_zero : size 0 = 0 := by simp [size] @[simp] theorem size_bit {b n} (h : bit b n ≠ 0) : size (bit b n) = succ (size n) := by - rw [size] + unfold size conv => lhs rw [binaryRec] @@ -81,7 +81,7 @@ theorem size_shiftLeft {m} (h : m ≠ 0) (n) : size (m <<< n) = size m + n := by theorem lt_size_self (n : ℕ) : n < 2 ^ size n := by rw [← one_shiftLeft] have : ∀ {n}, n = 0 → n < 1 <<< (size n) := by simp - apply binaryRec _ _ n + refine binaryRec ?_ ?_ n · apply this rfl intro b n IH by_cases h : bit b n = 0 @@ -91,11 +91,11 @@ theorem lt_size_self (n : ℕ) : n < 2 ^ size n := by theorem size_le {m n : ℕ} : size m ≤ n ↔ m < 2 ^ n := ⟨fun h => lt_of_lt_of_le (lt_size_self _) (pow_le_pow_of_le_right (by decide) h), by - rw [← one_shiftLeft]; revert n - apply binaryRec _ _ m - · intro n - simp - · intro b m IH n h + rw [← one_shiftLeft] + induction m using binaryRec generalizing n with + | z => simp + | f b m IH => + intro h by_cases e : bit b m = 0 · simp [e] rw [size_bit e] diff --git a/test/BinaryRec.lean b/test/BinaryRec.lean new file mode 100644 index 0000000000000..1112d600d4af8 --- /dev/null +++ b/test/BinaryRec.lean @@ -0,0 +1,8 @@ +import Mathlib.Data.Nat.Bits + +def Nat.popcountTR (n : Nat) : Nat := + n.binaryRec (·) (fun b _ f x ↦ f (x + b.toNat)) 0 + +/-- info: 1 -/ +#guard_msgs in +#eval Nat.popcountTR (2 ^ 20240) From 547e878793a20c0a59b55ba0cca411f38d0e1bfa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 21:00:38 +0000 Subject: [PATCH 084/131] =?UTF-8?q?feat:=20`Disjoint=20(a=20=E2=80=A2=20s)?= =?UTF-8?q?=20t=20=E2=86=94=20Disjoint=20s=20(a=E2=81=BB=C2=B9=20=E2=80=A2?= =?UTF-8?q?=20t)`=20(#17907)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Data/Set/Pointwise/SMul.lean | 17 ++++++++++++++--- 1 file changed, 14 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 00fe29ebfea63..9e6917fe0172c 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -330,7 +330,7 @@ theorem smul_set_inter : a • (s ∩ t) = a • s ∩ a • t := image_inter <| MulAction.injective a @[to_additive] -theorem smul_set_iInter {ι : Type*} +theorem smul_set_iInter {ι : Sort*} (a : α) (t : ι → Set β) : (a • ⋂ i, t i) = ⋂ i, a • t i := image_iInter (MulAction.bijective a) t @@ -401,8 +401,19 @@ lemma inv_op_smul_set_distrib (a : α) (s : Set α) : (op a • s)⁻¹ = a⁻¹ ext; simp [mem_smul_set_iff_inv_smul_mem] @[to_additive (attr := simp)] -lemma smul_set_disjoint_iff : Disjoint (a • s) (a • t) ↔ Disjoint s t := by - simp [disjoint_iff, ← smul_set_inter] +lemma disjoint_smul_set : Disjoint (a • s) (a • t) ↔ Disjoint s t := + disjoint_image_iff <| MulAction.injective _ + +@[to_additive] +lemma disjoint_smul_set_left : Disjoint (a • s) t ↔ Disjoint s (a⁻¹ • t) := by + simpa using disjoint_smul_set (a := a) (t := a⁻¹ • t) + +@[to_additive] +lemma disjoint_smul_set_right : Disjoint s (a • t) ↔ Disjoint (a⁻¹ • s) t := by + simpa using disjoint_smul_set (a := a) (s := a⁻¹ • s) + +@[to_additive (attr := deprecated (since := "2024-10-18"))] +alias smul_set_disjoint_iff := disjoint_smul_set end Group From 8959a60fa36fe31f229b95fa4bb3612b36b29160 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Fri, 18 Oct 2024 21:29:10 +0000 Subject: [PATCH 085/131] feat(CategoryTheory/Monoidal): add `IsMon_Hom` type class (#17186) --- Mathlib/CategoryTheory/Monoidal/Bimon_.lean | 4 ++++ Mathlib/CategoryTheory/Monoidal/Comon_.lean | 9 +++++++++ Mathlib/CategoryTheory/Monoidal/Mon_.lean | 9 +++++++++ 3 files changed, 22 insertions(+) diff --git a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean index e195b65504704..d522837730846 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean @@ -68,6 +68,10 @@ theorem one_counit (M : C) [Bimon_Class M] : η[M] ≫ ε[M] = 𝟙 (𝟙_ C) := end Bimon_Class +/-- The property that a morphism between bimonoid objects is a bimonoid morphism. -/ +class IsBimon_Hom {M N : C} [Bimon_Class M] [Bimon_Class N] (f : M ⟶ N) extends + IsMon_Hom f, IsComon_Hom f : Prop + variable (C) /-- diff --git a/Mathlib/CategoryTheory/Monoidal/Comon_.lean b/Mathlib/CategoryTheory/Monoidal/Comon_.lean index ff63ce4bbb0eb..ac2283e8e7402 100644 --- a/Mathlib/CategoryTheory/Monoidal/Comon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Comon_.lean @@ -71,6 +71,15 @@ end Comon_Class open scoped Comon_Class +variable {M N : C} [Comon_Class M] [Comon_Class N] + +/-- The property that a morphism between comonoid objects is a comonoid morphism. -/ +class IsComon_Hom (f : M ⟶ N) : Prop where + hom_counit : f ≫ ε = ε := by aesop_cat + hom_comul : f ≫ Δ = Δ ≫ (f ⊗ f) := by aesop_cat + +attribute [reassoc (attr := simp)] IsComon_Hom.hom_counit IsComon_Hom.hom_comul + variable (C) /-- A comonoid object internal to a monoidal category. diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index c04c4f4d55c07..52b4de6677a39 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -67,6 +67,15 @@ end Mon_Class open scoped Mon_Class +variable {M N : C} [Mon_Class M] [Mon_Class N] + +/-- The property that a morphism between monoid objects is a monoid morphism. -/ +class IsMon_Hom (f : M ⟶ N) : Prop where + one_hom : η ≫ f = η := by aesop_cat + mul_hom : μ ≫ f = (f ⊗ f) ≫ μ := by aesop_cat + +attribute [reassoc (attr := simp)] IsMon_Hom.one_hom IsMon_Hom.mul_hom + variable (C) /-- A monoid object internal to a monoidal category. From de5d061f84a1fdff844f983aea367c3bdbf676b8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 18 Oct 2024 21:29:11 +0000 Subject: [PATCH 086/131] chore: Rename `OpenEmbedding` to `IsOpenEmbedding` (#17898) `Function.Embedding` is a type while `Embedding` is a proposition, and there are many other kinds of embeddings than topological embeddings. Hence this PR is a step towards 1. renaming `Embedding` to `IsEmbedding` and similarly for neighborhing declarations (which `OpenEmbedding` is) 2. namespacing it inside `Topology` [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/rename.20.60Inducing.60.20and.20.60Embedding.60.3F). See #15993 for context. --- Mathlib/AlgebraicGeometry/AffineScheme.lean | 14 +- Mathlib/AlgebraicGeometry/Cover/Open.lean | 2 +- Mathlib/AlgebraicGeometry/FunctionField.lean | 4 +- Mathlib/AlgebraicGeometry/Gluing.lean | 9 +- Mathlib/AlgebraicGeometry/Limits.lean | 4 +- .../AlgebraicGeometry/Morphisms/Affine.lean | 2 +- .../AlgebraicGeometry/Morphisms/IsIso.lean | 2 +- .../Morphisms/OpenImmersion.lean | 4 +- .../Morphisms/Preimmersion.lean | 2 +- .../Morphisms/UnderlyingMap.lean | 14 +- Mathlib/AlgebraicGeometry/Noetherian.lean | 4 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 25 ++-- .../PrimeSpectrum/Basic.lean | 7 +- .../ProjectiveSpectrum/Scheme.lean | 13 +- Mathlib/AlgebraicGeometry/Properties.lean | 4 +- Mathlib/AlgebraicGeometry/Restrict.lean | 30 ++--- Mathlib/AlgebraicGeometry/Scheme.lean | 2 +- .../Complex/UpperHalfPlane/Manifold.lean | 4 +- .../Complex/UpperHalfPlane/Topology.lean | 15 ++- Mathlib/Analysis/Normed/Ring/Units.lean | 7 +- Mathlib/Analysis/SpecialFunctions/Exp.lean | 11 +- Mathlib/Geometry/Manifold/ChartedSpace.lean | 12 +- .../Geometry/Manifold/ContMDiff/Basic.lean | 27 ++-- .../Instances/UnitsOfNormedAlgebra.lean | 10 +- .../Manifold/LocalInvariantProperties.lean | 4 +- .../Manifold/SmoothManifoldWithCorners.lean | 8 +- .../RingedSpace/LocallyRingedSpace.lean | 10 +- .../Geometry/RingedSpace/OpenImmersion.lean | 53 ++++---- .../Geometry/RingedSpace/PresheafedSpace.lean | 16 +-- .../RingedSpace/PresheafedSpace/Gluing.lean | 19 +-- .../Geometry/RingedSpace/SheafedSpace.lean | 8 +- Mathlib/Geometry/RingedSpace/Stalks.lean | 10 +- .../Constructions/BorelSpace/Basic.lean | 5 +- .../EisensteinSeries/UniformConvergence.lean | 2 +- Mathlib/Topology/Algebra/ConstMulAction.lean | 2 +- .../Algebra/Constructions/DomMulAct.lean | 12 +- .../Algebra/Nonarchimedean/Basic.lean | 2 +- Mathlib/Topology/Bases.lean | 4 +- .../Category/CompHausLike/Limits.lean | 28 ++-- Mathlib/Topology/Category/Stonean/Limits.lean | 4 +- Mathlib/Topology/Category/TopCat/Basic.lean | 38 ++++-- .../Category/TopCat/Limits/Products.lean | 7 +- .../Category/TopCat/Limits/Pullbacks.lean | 48 ++++--- .../Topology/Category/TopCat/OpenNhds.lean | 7 +- Mathlib/Topology/Category/TopCat/Opens.lean | 33 +++-- Mathlib/Topology/Clopen.lean | 2 +- .../Topology/Compactification/OnePoint.lean | 23 ++-- .../Topology/Compactness/LocallyCompact.lean | 7 +- .../Topology/Connected/LocallyConnected.lean | 9 +- Mathlib/Topology/Connected/PathConnected.lean | 7 +- Mathlib/Topology/Constructions.lean | 72 +++++++---- Mathlib/Topology/ContinuousOn.lean | 5 +- Mathlib/Topology/Defs/Induced.lean | 7 +- Mathlib/Topology/Gluing.lean | 24 ++-- Mathlib/Topology/Homeomorph.lean | 16 ++- Mathlib/Topology/Inseparable.lean | 5 +- Mathlib/Topology/Instances/ENNReal.lean | 11 +- Mathlib/Topology/Instances/ENat.lean | 7 +- Mathlib/Topology/Instances/EReal.lean | 9 +- Mathlib/Topology/Irreducible.lean | 2 +- Mathlib/Topology/IsLocalHomeomorph.lean | 54 +++++--- Mathlib/Topology/LocalAtTarget.lean | 30 +++-- Mathlib/Topology/LocallyClosed.lean | 5 +- Mathlib/Topology/Maps/Basic.lean | 121 ++++++++++++------ Mathlib/Topology/PartialHomeomorph.lean | 29 +++-- Mathlib/Topology/QuasiSeparated.lean | 14 +- Mathlib/Topology/SeparatedMap.lean | 11 +- Mathlib/Topology/Separation.lean | 17 ++- Mathlib/Topology/Sets/Opens.lean | 14 +- Mathlib/Topology/Sheaves/LocalPredicate.lean | 4 +- .../Sheaves/SheafCondition/Sites.lean | 19 ++- Mathlib/Topology/Sheaves/Stalks.lean | 7 +- Mathlib/Topology/Sober.lean | 7 +- Mathlib/Topology/Support.lean | 2 +- scripts/no_lints_prime_decls.txt | 8 +- 75 files changed, 677 insertions(+), 419 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 74ea6759b202f..5a69b75178c41 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -278,7 +278,7 @@ def isoSpec : ↑U ≅ Spec Γ(X, U) := haveI : IsAffine U := hU U.toScheme.isoSpec ≪≫ Scheme.Spec.mapIso - (X.presheaf.mapIso (eqToIso U.openEmbedding_obj_top).op).op + (X.presheaf.mapIso (eqToIso U.isOpenEmbedding_obj_top).op).op open LocalRing in lemma isoSpec_hom_base_apply (x : U) : @@ -395,8 +395,8 @@ theorem _root_.AlgebraicGeometry.Scheme.Hom.isAffineOpen_iff_of_isOpenImmersion (f : AlgebraicGeometry.Scheme.Hom X Y) [H : IsOpenImmersion f] {U : X.Opens} : IsAffineOpen (f ''ᵁ U) ↔ IsAffineOpen U := by refine ⟨fun hU => @isAffine_of_isIso _ _ - (IsOpenImmersion.isoOfRangeEq (X.ofRestrict U.openEmbedding ≫ f) (Y.ofRestrict _) ?_).hom ?_ hU, - fun hU => hU.image_of_isOpenImmersion f⟩ + (IsOpenImmersion.isoOfRangeEq (X.ofRestrict U.isOpenEmbedding ≫ f) (Y.ofRestrict _) ?_).hom + ?_ hU, fun hU => hU.image_of_isOpenImmersion f⟩ · rw [Scheme.comp_base, coe_comp, Set.range_comp] dsimp [Opens.coe_inclusion', Scheme.restrict] erw [Subtype.range_coe, Subtype.range_coe] -- now `erw` after #13170 @@ -491,7 +491,7 @@ theorem ι_basicOpen_preimage (r : Γ(X, ⊤)) : IsAffineOpen ((X.basicOpen r).ι ⁻¹ᵁ U) := by apply (X.basicOpen r).ι.isAffineOpen_iff_of_isOpenImmersion.mp dsimp [Scheme.Hom.opensFunctor, LocallyRingedSpace.IsOpenImmersion.opensFunctor] - rw [Opens.functor_obj_map_obj, Opens.openEmbedding_obj_top, inf_comm, + rw [Opens.functor_obj_map_obj, Opens.isOpenEmbedding_obj_top, inf_comm, ← Scheme.basicOpen_res _ _ (homOfLE le_top).op] exact hU.basicOpen _ @@ -504,12 +504,12 @@ theorem exists_basicOpen_le {V : X.Opens} (x : V) (h : ↑x ∈ U) : ((Opens.map U.inclusion').obj V).isOpen have : U.ι ''ᵁ (U.toScheme.basicOpen r) = - X.basicOpen (X.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op r) := by + X.basicOpen (X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op r) := by refine (Scheme.image_basicOpen U.ι r).trans ?_ rw [Scheme.basicOpen_res_eq] simp only [Scheme.Opens.toScheme_presheaf_obj, Scheme.Opens.ι_appIso, Iso.refl_inv, CommRingCat.id_apply] - use X.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op r + use X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op r rw [← this] exact ⟨Set.image_subset_iff.mpr h₂, ⟨_, h⟩, h₁, rfl⟩ @@ -573,7 +573,7 @@ theorem isLocalization_of_eq_basicOpen {V : X.Opens} (i : V ⟶ U) (e : V = X.ba instance _root_.AlgebraicGeometry.Γ_restrict_isLocalization (X : Scheme.{u}) [IsAffine X] (r : Γ(X, ⊤)) : IsLocalization.Away r Γ(X.basicOpen r, ⊤) := - (isAffineOpen_top X).isLocalization_of_eq_basicOpen r _ (Opens.openEmbedding_obj_top _) + (isAffineOpen_top X).isLocalization_of_eq_basicOpen r _ (Opens.isOpenEmbedding_obj_top _) include hU in theorem basicOpen_basicOpen_is_basicOpen (g : Γ(X, X.basicOpen f)) : diff --git a/Mathlib/AlgebraicGeometry/Cover/Open.lean b/Mathlib/AlgebraicGeometry/Cover/Open.lean index 836bfdf366a33..32b601db856c3 100644 --- a/Mathlib/AlgebraicGeometry/Cover/Open.lean +++ b/Mathlib/AlgebraicGeometry/Cover/Open.lean @@ -248,7 +248,7 @@ theorem OpenCover.compactSpace {X : Scheme.{u}} (𝒰 : X.OpenCover) [Finite (TopCat.homeoOfIso (asIso (IsOpenImmersion.isoOfRangeEq (𝒰.map i) - (X.ofRestrict (Opens.openEmbedding ⟨_, (𝒰.IsOpen i).base_open.isOpen_range⟩)) + (X.ofRestrict (Opens.isOpenEmbedding ⟨_, (𝒰.IsOpen i).base_open.isOpen_range⟩)) Subtype.range_coe.symm).hom.base)) /-- Given open covers `{ Uᵢ }` and `{ Uⱼ }`, we may form the open cover `{ Uᵢ ∩ Uⱼ }`. -/ diff --git a/Mathlib/AlgebraicGeometry/FunctionField.lean b/Mathlib/AlgebraicGeometry/FunctionField.lean index 242c360e3de84..6833058adc17f 100644 --- a/Mathlib/AlgebraicGeometry/FunctionField.lean +++ b/Mathlib/AlgebraicGeometry/FunctionField.lean @@ -135,7 +135,7 @@ theorem IsAffineOpen.primeIdealOf_genericPoint {X : Scheme} [IsIntegral X] {U : delta IsAffineOpen.primeIdealOf convert genericPoint_eq_of_isOpenImmersion - (U.toScheme.isoSpec.hom ≫ Spec.map (X.presheaf.map (eqToHom U.openEmbedding_obj_top).op)) + (U.toScheme.isoSpec.hom ≫ Spec.map (X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top).op)) -- Porting note: this was `ext1` apply Subtype.ext exact (genericPoint_eq_of_isOpenImmersion U.ι).symm @@ -146,7 +146,7 @@ theorem functionField_isFractionRing_of_isAffineOpen [IsIntegral X] (U : X.Opens haveI : IsAffine _ := hU haveI : IsIntegral U := @isIntegral_of_isAffine_of_isDomain _ _ _ - (by rw [Scheme.Opens.toScheme_presheaf_obj, Opens.openEmbedding_obj_top]; infer_instance) + (by rw [Scheme.Opens.toScheme_presheaf_obj, Opens.isOpenEmbedding_obj_top]; infer_instance) delta IsFractionRing Scheme.functionField convert hU.isLocalization_stalk ⟨genericPoint X, (((genericPoint_spec X).mem_open_set_iff U.isOpen).mpr (by simpa using ‹Nonempty U›))⟩ using 1 diff --git a/Mathlib/AlgebraicGeometry/Gluing.lean b/Mathlib/AlgebraicGeometry/Gluing.lean index 9eeee3d9e71e3..c38450ee48155 100644 --- a/Mathlib/AlgebraicGeometry/Gluing.lean +++ b/Mathlib/AlgebraicGeometry/Gluing.lean @@ -379,10 +379,13 @@ theorem fromGlued_open_map : IsOpenMap 𝒰.fromGlued.base := by exact Set.preimage_image_eq _ 𝒰.fromGlued_injective · exact ⟨hx, 𝒰.covers x⟩ -theorem fromGlued_openEmbedding : OpenEmbedding 𝒰.fromGlued.base := - openEmbedding_of_continuous_injective_open +theorem fromGlued_isOpenEmbedding : IsOpenEmbedding 𝒰.fromGlued.base := + isOpenEmbedding_of_continuous_injective_open (by fun_prop) 𝒰.fromGlued_injective 𝒰.fromGlued_open_map +@[deprecated (since := "2024-10-18")] +alias fromGlued_openEmbedding := fromGlued_isOpenEmbedding + instance : Epi 𝒰.fromGlued.base := by rw [TopCat.epi_iff_surjective] intro x @@ -393,7 +396,7 @@ instance : Epi 𝒰.fromGlued.base := by exact h instance fromGlued_open_immersion : IsOpenImmersion 𝒰.fromGlued := - IsOpenImmersion.of_stalk_iso _ 𝒰.fromGlued_openEmbedding + IsOpenImmersion.of_stalk_iso _ 𝒰.fromGlued_isOpenEmbedding instance : IsIso 𝒰.fromGlued := let F := Scheme.forgetToLocallyRingedSpace ⋙ LocallyRingedSpace.forgetToSheafedSpace ⋙ diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index 0622355eef447..1480c9d5d315f 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -250,7 +250,7 @@ lemma sigmaι_eq_iff (i j : ι) (x y) : by_cases h : i = j · subst h simp only [Sigma.mk.inj_iff, heq_eq_eq, true_and] - exact ((disjointGlueData f).ι i).openEmbedding.inj H + exact ((disjointGlueData f).ι i).isOpenEmbedding.inj H · obtain (e | ⟨z, _⟩) := (Scheme.GlueData.ι_eq_iff _ _ _ _ _).mp H · exact (h (Sigma.mk.inj_iff.mp e).1).elim · simp only [disjointGlueData_J, disjointGlueData_V, h, ↓reduceIte] at z @@ -269,7 +269,7 @@ lemma disjoint_opensRange_sigmaι (i j : ι) (h : i ≠ j) : lemma exists_sigmaι_eq (x : (∐ f : _)) : ∃ i y, (Sigma.ι f i).base y = x := by obtain ⟨i, y, e⟩ := (disjointGlueData f).ι_jointly_surjective ((sigmaIsoGlued f).hom.base x) - refine ⟨i, y, (sigmaIsoGlued f).hom.openEmbedding.inj ?_⟩ + refine ⟨i, y, (sigmaIsoGlued f).hom.isOpenEmbedding.inj ?_⟩ rwa [← Scheme.comp_base_apply, ι_sigmaIsoGlued_hom] lemma iSup_opensRange_sigmaι : ⨆ i, (Sigma.ι f i).opensRange = ⊤ := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean index ecd431a4735b4..bcb0509ca655d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean @@ -129,7 +129,7 @@ lemma isAffine_of_isAffineOpen_basicOpen (s : Set Γ(X, ⊤)) refine IsIso.comp_isIso' ?_ inferInstance convert isIso_ΓSpec_adjunction_unit_app_basicOpen i.1 using 0 refine congr(IsIso ((ΓSpec.adjunction.unit.app X).app $(?_))) - rw [Opens.openEmbedding_obj_top] + rw [Opens.isOpenEmbedding_obj_top] /-- If `s` is a spanning set of `Γ(X, U)`, such that each `X.basicOpen i` is affine, then `U` is also diff --git a/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean b/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean index 384f3493f9a1e..41f31c85bebb9 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean @@ -31,7 +31,7 @@ lemma isomorphisms_eq_stalkwise : ext X Y f exact ⟨fun H ↦ inferInstanceAs (IsIso (TopCat.isoOfHomeo (H.1.1.toHomeomeomorph_of_surjective H.2)).hom), fun (_ : IsIso f.base) ↦ - let e := (TopCat.homeoOfIso <| asIso f.base); ⟨e.openEmbedding, e.surjective⟩⟩ + let e := (TopCat.homeoOfIso <| asIso f.base); ⟨e.isOpenEmbedding, e.surjective⟩⟩ instance : IsLocalAtTarget (isomorphisms Scheme) := isomorphisms_eq_isOpenImmersion_inf_surjective ▸ inferInstance diff --git a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean index 43a94b8d2d1a0..1a4868bc3510b 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean @@ -29,13 +29,13 @@ namespace AlgebraicGeometry variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) theorem isOpenImmersion_iff_stalk {f : X ⟶ Y} : IsOpenImmersion f ↔ - OpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := by + IsOpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := by constructor · intro h; exact ⟨h.1, inferInstance⟩ · rintro ⟨h₁, h₂⟩; exact IsOpenImmersion.of_stalk_iso f h₁ theorem isOpenImmersion_eq_inf : - @IsOpenImmersion = (topologically OpenEmbedding) ⊓ + @IsOpenImmersion = (topologically IsOpenEmbedding) ⊓ stalkwise (fun f ↦ Function.Bijective f) := by ext exact isOpenImmersion_iff_stalk.trans diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean index a5fd6903836cb..9336dbb858cde 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean @@ -60,7 +60,7 @@ instance : IsLocalAtTarget @IsPreimmersion := isPreimmersion_eq_inf ▸ inferInstance instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsOpenImmersion f] : IsPreimmersion f where - base_embedding := f.openEmbedding.toEmbedding + base_embedding := f.isOpenEmbedding.toEmbedding surj_on_stalks _ := (ConcreteCategory.bijective_of_isIso _).2 instance : MorphismProperty.IsMultiplicative @IsPreimmersion where diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean index ff349f67ab6a8..00d8b44dc43f0 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean @@ -18,7 +18,7 @@ of the underlying map of topological spaces, including - `IsOpenMap` - `IsClosedMap` - `Embedding` -- `OpenEmbedding` +- `IsOpenEmbedding` - `ClosedEmbedding` -/ @@ -119,17 +119,17 @@ instance embedding_isLocalAtTarget : IsLocalAtTarget (topologically Embedding) : end Embedding -section OpenEmbedding +section IsOpenEmbedding -instance : (topologically OpenEmbedding).RespectsIso := - topologically_respectsIso _ (fun e ↦ e.openEmbedding) (fun _ _ hf hg ↦ hg.comp hf) +instance : (topologically IsOpenEmbedding).RespectsIso := + topologically_respectsIso _ (fun e ↦ e.isOpenEmbedding) (fun _ _ hf hg ↦ hg.comp hf) -instance openEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically OpenEmbedding) := +instance isOpenEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically IsOpenEmbedding) := topologically_isLocalAtTarget _ (fun _ s hf ↦ hf.restrictPreimage s) - (fun _ _ _ hU hfcont hf ↦ (openEmbedding_iff_openEmbedding_of_iSup_eq_top hU hfcont).mpr hf) + (fun _ _ _ hU hfcont hf ↦ (isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top hU hfcont).mpr hf) -end OpenEmbedding +end IsOpenEmbedding section ClosedEmbedding diff --git a/Mathlib/AlgebraicGeometry/Noetherian.lean b/Mathlib/AlgebraicGeometry/Noetherian.lean index 45c741d358319..121bca1ca0d2b 100644 --- a/Mathlib/AlgebraicGeometry/Noetherian.lean +++ b/Mathlib/AlgebraicGeometry/Noetherian.lean @@ -199,7 +199,7 @@ instance (priority := 100) {Z : Scheme} [IsLocallyNoetherian X] apply (quasiCompact_iff_forall_affine f).mpr intro U hU rw [Opens.map_coe, ← Set.preimage_inter_range] - apply f.openEmbedding.toInducing.isCompact_preimage' + apply f.isOpenEmbedding.toInducing.isCompact_preimage' · apply (noetherianSpace_set_iff _).mp · convert noetherianSpace_of_isAffineOpen U hU apply IsLocallyNoetherian.component_noetherian ⟨U, hU⟩ @@ -213,7 +213,7 @@ instance (priority := 100) IsLocallyNoetherian.quasiSeparatedSpace [IsLocallyNoe QuasiSeparatedSpace X := by apply (quasiSeparatedSpace_iff_affine X).mpr intro U V - have hInd := U.2.fromSpec.openEmbedding.toInducing + have hInd := U.2.fromSpec.isOpenEmbedding.toInducing apply (hInd.isCompact_preimage_iff ?_).mp · rw [← Set.preimage_inter_range, IsAffineOpen.range_fromSpec, Set.inter_comm] apply hInd.isCompact_preimage' diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index 7b3769c508404..bc36b58d339d1 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -56,7 +56,7 @@ protected def scheme (X : LocallyRingedSpace.{u}) refine SheafedSpace.forgetToPresheafedSpace.preimageIso ?_ apply PresheafedSpace.IsOpenImmersion.isoOfRangeEq (PresheafedSpace.ofRestrict _ _) f.1 · exact Subtype.range_coe_subtype - · exact Opens.openEmbedding _ -- Porting note (#11187): was `infer_instance` + · exact Opens.isOpenEmbedding _ -- Porting note (#11187): was `infer_instance` end LocallyRingedSpace.IsOpenImmersion @@ -71,13 +71,16 @@ namespace Scheme.Hom variable {X Y : Scheme.{u}} (f : Scheme.Hom X Y) [H : IsOpenImmersion f] -theorem openEmbedding : OpenEmbedding f.base := +theorem isOpenEmbedding : IsOpenEmbedding f.base := H.base_open +@[deprecated (since := "2024-10-18")] +alias openEmbedding := isOpenEmbedding + /-- The image of an open immersion as an open set. -/ @[simps] def opensRange : Y.Opens := - ⟨_, f.openEmbedding.isOpen_range⟩ + ⟨_, f.isOpenEmbedding.isOpen_range⟩ /-- The functor `opens X ⥤ opens Y` associated with an open immersion `f : X ⟶ Y`. -/ abbrev opensFunctor : X.Opens ⥤ Y.Opens := @@ -103,7 +106,7 @@ lemma image_top_eq_opensRange : f ''ᵁ ⊤ = f.opensRange := by @[simp] lemma preimage_image_eq (U : X.Opens) : f ⁻¹ᵁ f ''ᵁ U = U := by apply Opens.ext - simp [Set.preimage_image_eq _ f.openEmbedding.inj] + simp [Set.preimage_image_eq _ f.isOpenEmbedding.inj] lemma image_le_image_iff (f : X ⟶ Y) [IsOpenImmersion f] (U U' : X.Opens) : f ''ᵁ U ≤ f ''ᵁ U' ↔ U ≤ U' := by @@ -194,7 +197,7 @@ def IsOpenImmersion.opensEquiv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion X.Opens ≃ { U : Y.Opens // U ≤ f.opensRange } where toFun U := ⟨f ''ᵁ U, Set.image_subset_range _ _⟩ invFun U := f ⁻¹ᵁ U - left_inv _ := Opens.ext (Set.preimage_image_eq _ f.openEmbedding.inj) + left_inv _ := Opens.ext (Set.preimage_image_eq _ f.isOpenEmbedding.inj) right_inv U := Subtype.ext (Opens.ext (Set.image_preimage_eq_of_subset U.2)) namespace Scheme @@ -202,7 +205,7 @@ namespace Scheme instance basic_open_isOpenImmersion {R : CommRingCat.{u}} (f : R) : IsOpenImmersion (Spec.map (CommRingCat.ofHom (algebraMap R (Localization.Away f)))) := by apply SheafedSpace.IsOpenImmersion.of_stalk_iso (H := ?_) - · exact (PrimeSpectrum.localization_away_openEmbedding (Localization.Away f) f : _) + · exact (PrimeSpectrum.localization_away_isOpenEmbedding (Localization.Away f) f : _) · intro x exact Spec_map_localization_isIso R (Submonoid.powers f) x @@ -292,7 +295,7 @@ end PresheafedSpace.IsOpenImmersion section Restrict -variable {U : TopCat.{u}} (X : Scheme.{u}) {f : U ⟶ TopCat.of X} (h : OpenEmbedding f) +variable {U : TopCat.{u}} (X : Scheme.{u}) {f : U ⟶ TopCat.of X} (h : IsOpenEmbedding f) /-- The restriction of a Scheme along an open embedding. -/ @[simps! (config := .lemmasOnly) carrier, simps! presheaf_obj] @@ -351,7 +354,7 @@ theorem to_iso {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsOpenImmersion f] [Epi f.b LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) (@PresheafedSpace.IsOpenImmersion.to_iso _ _ _ _ f.toPshHom h _) _ -theorem of_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) (hf : OpenEmbedding f.base) +theorem of_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) (hf : IsOpenEmbedding f.base) [∀ x, IsIso (f.stalkMap x)] : IsOpenImmersion f := haveI (x : X) : IsIso (f.toShHom.stalkMap x) := inferInstanceAs <| IsIso (f.stalkMap x) SheafedSpace.IsOpenImmersion.of_stalk_iso f.toShHom hf @@ -368,10 +371,10 @@ lemma of_comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) [IsOpenImmersion infer_instance IsIso.of_isIso_comp_left (f := g.stalkMap (f.base x)) _ IsOpenImmersion.of_stalk_iso _ <| - OpenEmbedding.of_comp _ (Scheme.Hom.openEmbedding g) (Scheme.Hom.openEmbedding (f ≫ g)) + IsOpenEmbedding.of_comp _ (Scheme.Hom.isOpenEmbedding g) (Scheme.Hom.isOpenEmbedding (f ≫ g)) theorem iff_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) : - IsOpenImmersion f ↔ OpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := + IsOpenImmersion f ↔ IsOpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := ⟨fun H => ⟨H.1, fun x ↦ inferInstanceAs <| IsIso (f.toPshHom.stalkMap x)⟩, fun ⟨h₁, h₂⟩ => @IsOpenImmersion.of_stalk_iso _ _ f h₁ h₂⟩ @@ -391,7 +394,7 @@ theorem _root_.AlgebraicGeometry.isIso_iff_stalk_iso {X Y : Scheme.{u}} (f : X (Equiv.ofBijective _ ⟨h₂.inj, (TopCat.epi_iff_surjective _).mp h₁⟩) h₂.continuous h₂.isOpenMap)).hom infer_instance - · intro H; exact ⟨inferInstance, (TopCat.homeoOfIso (asIso f.base)).openEmbedding⟩ + · intro H; exact ⟨inferInstance, (TopCat.homeoOfIso (asIso f.base)).isOpenEmbedding⟩ /-- An open immersion induces an isomorphism from the domain onto the image -/ def isoRestrict : X ≅ (Z.restrict H.base_open : _) := diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 05c884eb52c60..f918a39487b28 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -457,13 +457,16 @@ theorem localization_away_comap_range (S : Type v) [CommSemiring S] [Algebra R S · rintro h₁ _ ⟨⟨n, rfl⟩, h₃⟩ exact h₁ (x.2.mem_of_pow_mem _ h₃) -theorem localization_away_openEmbedding (S : Type v) [CommSemiring S] [Algebra R S] (r : R) - [IsLocalization.Away r S] : OpenEmbedding (comap (algebraMap R S)) := +theorem localization_away_isOpenEmbedding (S : Type v) [CommSemiring S] [Algebra R S] (r : R) + [IsLocalization.Away r S] : IsOpenEmbedding (comap (algebraMap R S)) := { toEmbedding := localization_comap_embedding S (Submonoid.powers r) isOpen_range := by rw [localization_away_comap_range S r] exact isOpen_basicOpen } +@[deprecated (since := "2024-10-18")] +alias localization_away_openEmbedding := localization_away_isOpenEmbedding + theorem isCompact_basicOpen (f : R) : IsCompact (basicOpen f : Set (PrimeSpectrum R)) := by rw [← localization_away_comap_range (Localization (Submonoid.powers f))] exact isCompact_range (map_continuous _) diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 77da59dbcba24..dae30030d3507 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -122,12 +122,13 @@ local notation3 "Proj.T" => PresheafedSpace.carrier <| SheafedSpace.toPresheafed /-- `Proj` restrict to some open set -/ macro "Proj| " U:term : term => - `((Proj.toLocallyRingedSpace 𝒜).restrict (Opens.openEmbedding (X := Proj.T) ($U : Opens Proj.T))) + `((Proj.toLocallyRingedSpace 𝒜).restrict + (Opens.isOpenEmbedding (X := Proj.T) ($U : Opens Proj.T))) /-- the underlying topological space of `Proj` restricted to some open set -/ local notation "Proj.T| " U => PresheafedSpace.carrier <| SheafedSpace.toPresheafedSpace <| LocallyRingedSpace.toSheafedSpace - <| (LocallyRingedSpace.restrict Proj (Opens.openEmbedding (X := Proj.T) (U : Opens Proj.T))) + <| (LocallyRingedSpace.restrict Proj (Opens.isOpenEmbedding (X := Proj.T) (U : Opens Proj.T))) /-- basic open sets in `Proj` -/ local notation "pbo " x => ProjectiveSpectrum.basicOpen 𝒜 x @@ -610,13 +611,13 @@ Mathematically, the map is the same as `awayToSection`. -/ def awayToΓ (f) : CommRingCat.of (A⁰_ f) ⟶ LocallyRingedSpace.Γ.obj (op <| Proj| pbo f) := awayToSection 𝒜 f ≫ (ProjectiveSpectrum.Proj.structureSheaf 𝒜).1.map - (homOfLE (Opens.openEmbedding_obj_top _).le).op + (homOfLE (Opens.isOpenEmbedding_obj_top _).le).op lemma awayToΓ_ΓToStalk (f) (x) : awayToΓ 𝒜 f ≫ (Proj| pbo f).presheaf.Γgerm x = HomogeneousLocalization.mapId 𝒜 (Submonoid.powers_le.mpr x.2) ≫ (Proj.stalkIso' 𝒜 x.1).toCommRingCatIso.inv ≫ - ((Proj.toLocallyRingedSpace 𝒜).restrictStalkIso (Opens.openEmbedding _) x).inv := by + ((Proj.toLocallyRingedSpace 𝒜).restrictStalkIso (Opens.isOpenEmbedding _) x).inv := by rw [awayToΓ, Category.assoc, ← Category.assoc _ (Iso.inv _), Iso.eq_comp_inv, Category.assoc, Category.assoc, Presheaf.Γgerm] rw [LocallyRingedSpace.restrictStalkIso_hom_eq_germ] @@ -781,7 +782,7 @@ lemma toStalk_specStalkEquiv (f) (x : pbo f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 lemma stalkMap_toSpec (f) (x : pbo f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : (toSpec 𝒜 f).stalkMap x = (specStalkEquiv 𝒜 f x f_deg hm).hom ≫ (Proj.stalkIso' 𝒜 x.1).toCommRingCatIso.inv ≫ - ((Proj.toLocallyRingedSpace 𝒜).restrictStalkIso (Opens.openEmbedding _) x).inv := + ((Proj.toLocallyRingedSpace 𝒜).restrictStalkIso (Opens.isOpenEmbedding _) x).inv := IsLocalization.ringHom_ext (R := A⁰_ f) ((toSpec 𝒜 f).base x).asIdeal.primeCompl (S := (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).base x)) <| (toStalk_stalkMap_toSpec _ _ _).trans <| by @@ -794,7 +795,7 @@ lemma isIso_toSpec (f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : rw [stalkMap_toSpec 𝒜 f x f_deg hm]; infer_instance haveI : LocallyRingedSpace.IsOpenImmersion (toSpec 𝒜 f) := LocallyRingedSpace.IsOpenImmersion.of_stalk_iso (toSpec 𝒜 f) - (TopCat.homeoOfIso (asIso <| (toSpec 𝒜 f).base)).openEmbedding + (TopCat.homeoOfIso (asIso <| (toSpec 𝒜 f).base)).isOpenEmbedding exact LocallyRingedSpace.IsOpenImmersion.to_iso _ end ProjectiveSpectrum.Proj diff --git a/Mathlib/AlgebraicGeometry/Properties.lean b/Mathlib/AlgebraicGeometry/Properties.lean index a5bbb08d90c15..84f113a767c02 100644 --- a/Mathlib/AlgebraicGeometry/Properties.lean +++ b/Mathlib/AlgebraicGeometry/Properties.lean @@ -39,8 +39,8 @@ instance : QuasiSober X := by quasiSober_of_open_cover (Set.range fun x => Set.range <| (X.affineCover.map x).base) · rintro ⟨_, i, rfl⟩; exact (X.affineCover.IsOpen i).base_open.isOpen_range · rintro ⟨_, i, rfl⟩ - exact @OpenEmbedding.quasiSober _ _ _ _ _ (Homeomorph.ofEmbedding _ - (X.affineCover.IsOpen i).base_open.toEmbedding).symm.openEmbedding PrimeSpectrum.quasiSober + exact @IsOpenEmbedding.quasiSober _ _ _ _ _ (Homeomorph.ofEmbedding _ + (X.affineCover.IsOpen i).base_open.toEmbedding).symm.isOpenEmbedding PrimeSpectrum.quasiSober · rw [Set.top_eq_univ, Set.sUnion_range, Set.eq_univ_iff_forall] intro x; exact ⟨_, ⟨_, rfl⟩, X.affineCover.covers x⟩ diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index 3eadf0265e3fc..2e8d4801e0e78 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -41,7 +41,7 @@ namespace Scheme.Opens /-- Open subset of a scheme as a scheme. -/ @[coe] def toScheme {X : Scheme.{u}} (U : X.Opens) : Scheme.{u} := - X.restrict U.openEmbedding + X.restrict U.isOpenEmbedding instance : CoeOut X.Opens Scheme := ⟨toScheme⟩ @@ -85,7 +85,7 @@ lemma range_ι : Set.range U.ι.base = U := Subtype.range_val lemma ι_image_top : U.ι ''ᵁ ⊤ = U := - U.openEmbedding_obj_top + U.isOpenEmbedding_obj_top lemma ι_image_le (W : U.toScheme.Opens) : U.ι ''ᵁ W ≤ U := by simp_rw [← U.ι_image_top] @@ -105,7 +105,7 @@ lemma ι_app_self : U.ι.app U = X.presheaf.map (eqToHom (X := U.ι ''ᵁ _) (by lemma eq_presheaf_map_eqToHom {V W : Opens U} (e : U.ι ''ᵁ V = U.ι ''ᵁ W) : X.presheaf.map (eqToHom e).op = - U.toScheme.presheaf.map (eqToHom <| U.openEmbedding.functor_obj_injective e).op := rfl + U.toScheme.presheaf.map (eqToHom <| U.isOpenEmbedding.functor_obj_injective e).op := rfl @[simp] lemma nonempty_iff : Nonempty U.toScheme ↔ (U : Set X).Nonempty := by @@ -121,20 +121,20 @@ def topIso : Γ(U, ⊤) ≅ Γ(X, U) := /-- The stalks of an open subscheme are isomorphic to the stalks of the original scheme. -/ def stalkIso {X : Scheme.{u}} (U : X.Opens) (x : U) : U.toScheme.presheaf.stalk x ≅ X.presheaf.stalk x.1 := - X.restrictStalkIso (Opens.openEmbedding _) _ + X.restrictStalkIso (Opens.isOpenEmbedding _) _ @[reassoc (attr := simp)] lemma germ_stalkIso_hom {X : Scheme.{u}} (U : X.Opens) {V : U.toScheme.Opens} (x : U) (hx : x ∈ V) : U.toScheme.presheaf.germ V x hx ≫ (U.stalkIso x).hom = X.presheaf.germ (U.ι ''ᵁ V) x.1 ⟨x, hx, rfl⟩ := - PresheafedSpace.restrictStalkIso_hom_eq_germ _ U.openEmbedding _ _ _ + PresheafedSpace.restrictStalkIso_hom_eq_germ _ U.isOpenEmbedding _ _ _ @[reassoc] lemma germ_stalkIso_inv {X : Scheme.{u}} (U : X.Opens) (V : U.toScheme.Opens) (x : U) (hx : x ∈ V) : X.presheaf.germ (U.ι ''ᵁ V) x ⟨x, hx, rfl⟩ ≫ (U.stalkIso x).inv = U.toScheme.presheaf.germ V x hx := - PresheafedSpace.restrictStalkIso_inv_eq_germ X.toPresheafedSpace U.openEmbedding V x hx + PresheafedSpace.restrictStalkIso_inv_eq_germ X.toPresheafedSpace U.isOpenEmbedding V x hx end Scheme.Opens @@ -168,9 +168,9 @@ instance ΓRestrictAlgebra {X : Scheme.{u}} (U : X.Opens) : lemma Scheme.map_basicOpen' (r : Γ(U, ⊤)) : U.ι ''ᵁ (U.toScheme.basicOpen r) = X.basicOpen - (X.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op r) := by - refine (Scheme.image_basicOpen (X.ofRestrict U.openEmbedding) r).trans ?_ - rw [← Scheme.basicOpen_res_eq _ _ (eqToHom U.openEmbedding_obj_top).op] + (X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op r) := by + refine (Scheme.image_basicOpen (X.ofRestrict U.isOpenEmbedding) r).trans ?_ + rw [← Scheme.basicOpen_res_eq _ _ (eqToHom U.isOpenEmbedding_obj_top).op] rw [← comp_apply, ← CategoryTheory.Functor.map_comp, ← op_comp, eqToHom_trans, eqToHom_refl, op_id, CategoryTheory.Functor.map_id] congr @@ -231,7 +231,7 @@ theorem Scheme.restrictFunctor_map_ofRestrict {U V : X.Opens} (i : U ⟶ V) : theorem Scheme.restrictFunctor_map_base {U V : X.Opens} (i : U ⟶ V) : (X.restrictFunctor.map i).left.base = (Opens.toTopCat _).map i := by ext a; refine Subtype.ext ?_ -- Porting note: `ext` did not pick up `Subtype.ext` - exact (congr_arg (fun f : X.restrict U.openEmbedding ⟶ X => f.base a) + exact (congr_arg (fun f : X.restrict U.isOpenEmbedding ⟶ X => f.base a) (X.restrictFunctor_map_ofRestrict i)) theorem Scheme.restrictFunctor_map_app_aux {U V : X.Opens} (i : U ⟶ V) (W : Opens V) : @@ -258,7 +258,7 @@ isomorphic to the structure sheaf. -/ @[simps!] def Scheme.restrictFunctorΓ : X.restrictFunctor.op ⋙ (Over.forget X).op ⋙ Scheme.Γ ≅ X.presheaf := NatIso.ofComponents - (fun U => X.presheaf.mapIso ((eqToIso (unop U).openEmbedding_obj_top).symm.op : _)) + (fun U => X.presheaf.mapIso ((eqToIso (unop U).isOpenEmbedding_obj_top).symm.op : _)) (by intro U V i dsimp @@ -348,7 +348,7 @@ theorem isPullback_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Open rw [← Category.id_comp f] refine (IsPullback.of_horiz_isIso ⟨?_⟩).paste_horiz - (IsPullback.of_hasPullback f (Y.ofRestrict U.openEmbedding)).flip + (IsPullback.of_hasPullback f (Y.ofRestrict U.isOpenEmbedding)).flip -- Porting note: changed `rw` to `erw` erw [pullbackRestrictIsoRestrict_inv_fst]; rw [Category.comp_id] @@ -434,8 +434,8 @@ theorem morphismRestrict_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V theorem Γ_map_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : Scheme.Γ.map (f ∣_ U).op = - Y.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op ≫ - f.app U ≫ X.presheaf.map (eqToHom (f ⁻¹ᵁ U).openEmbedding_obj_top).op := by + Y.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op ≫ + f.app U ≫ X.presheaf.map (eqToHom (f ⁻¹ᵁ U).isOpenEmbedding_obj_top).op := by rw [Scheme.Γ_map_op, morphismRestrict_app f U ⊤, f.naturality_assoc, ← X.presheaf.map_comp] rfl @@ -486,7 +486,7 @@ def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : /-- Restricting a morphism twice onto a basic open set is isomorphic to one restriction. -/ def morphismRestrictRestrictBasicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (r : Γ(Y, U)) : Arrow.mk (f ∣_ U ∣_ - U.toScheme.basicOpen (Y.presheaf.map (eqToHom U.openEmbedding_obj_top).op r)) ≅ + U.toScheme.basicOpen (Y.presheaf.map (eqToHom U.isOpenEmbedding_obj_top).op r)) ≅ Arrow.mk (f ∣_ Y.basicOpen r) := by refine morphismRestrictRestrict _ _ _ ≪≫ morphismRestrictEq _ ?_ have e := Scheme.preimage_basicOpen U.ι r diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index d77a37bef9383..18cc563e90561 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -46,7 +46,7 @@ structure Scheme extends LocallyRingedSpace where ∀ x : toLocallyRingedSpace, ∃ (U : OpenNhds x) (R : CommRingCat), Nonempty - (toLocallyRingedSpace.restrict U.openEmbedding ≅ Spec.toLocallyRingedSpace.obj (op R)) + (toLocallyRingedSpace.restrict U.isOpenEmbedding ≅ Spec.toLocallyRingedSpace.obj (op R)) namespace Scheme diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean index 2bc3767cffe15..45772d8d62ff6 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean @@ -19,10 +19,10 @@ open scoped UpperHalfPlane Manifold namespace UpperHalfPlane noncomputable instance : ChartedSpace ℂ ℍ := - UpperHalfPlane.openEmbedding_coe.singletonChartedSpace + UpperHalfPlane.isOpenEmbedding_coe.singletonChartedSpace instance : SmoothManifoldWithCorners 𝓘(ℂ) ℍ := - UpperHalfPlane.openEmbedding_coe.singleton_smoothManifoldWithCorners 𝓘(ℂ) + UpperHalfPlane.isOpenEmbedding_coe.singleton_smoothManifoldWithCorners 𝓘(ℂ) /-- The inclusion map `ℍ → ℂ` is a smooth map of manifolds. -/ theorem smooth_coe : Smooth 𝓘(ℂ) 𝓘(ℂ) ((↑) : ℍ → ℂ) := fun _ => contMDiffAt_extChartAt diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean index caecc196aeb8d..8aec80774131b 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean @@ -30,8 +30,11 @@ namespace UpperHalfPlane instance : TopologicalSpace ℍ := instTopologicalSpaceSubtype -theorem openEmbedding_coe : OpenEmbedding ((↑) : ℍ → ℂ) := - IsOpen.openEmbedding_subtype_val <| isOpen_lt continuous_const Complex.continuous_im +theorem isOpenEmbedding_coe : IsOpenEmbedding ((↑) : ℍ → ℂ) := + IsOpen.isOpenEmbedding_subtypeVal <| isOpen_lt continuous_const Complex.continuous_im + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_coe := isOpenEmbedding_coe theorem embedding_coe : Embedding ((↑) : ℍ → ℂ) := embedding_subtype_val @@ -55,7 +58,7 @@ instance : T4Space ℍ := inferInstance instance : ContractibleSpace ℍ := (convex_halfspace_im_gt 0).contractibleSpace ⟨I, one_pos.trans_eq I_im.symm⟩ -instance : LocPathConnectedSpace ℍ := openEmbedding_coe.locPathConnectedSpace +instance : LocPathConnectedSpace ℍ := isOpenEmbedding_coe.locPathConnectedSpace instance : NoncompactSpace ℍ := by refine ⟨fun h => ?_⟩ @@ -65,7 +68,7 @@ instance : NoncompactSpace ℍ := by exact absurd ((this 0).1 (@left_mem_Ici ℝ _ 0)) (@lt_irrefl ℝ _ 0) instance : LocallyCompactSpace ℍ := - openEmbedding_coe.locallyCompactSpace + isOpenEmbedding_coe.locallyCompactSpace section strips @@ -118,13 +121,13 @@ theorem ModularGroup_T_zpow_mem_verticalStrip (z : ℍ) {N : ℕ} (hn : 0 < N) : end strips /-- A continuous section `ℂ → ℍ` of the natural inclusion map, bundled as a `PartialHomeomorph`. -/ -def ofComplex : PartialHomeomorph ℂ ℍ := (openEmbedding_coe.toPartialHomeomorph _).symm +def ofComplex : PartialHomeomorph ℂ ℍ := (isOpenEmbedding_coe.toPartialHomeomorph _).symm /-- Extend a function on `ℍ` arbitrarily to a function on all of `ℂ`. -/ scoped notation "↑ₕ" f => f ∘ ofComplex lemma ofComplex_apply (z : ℍ) : ofComplex (z : ℂ) = z := - OpenEmbedding.toPartialHomeomorph_left_inv .. + IsOpenEmbedding.toPartialHomeomorph_left_inv .. lemma ofComplex_apply_eq_ite (w : ℂ) : ofComplex w = if hw : 0 < w.im then ⟨w, hw⟩ else Classical.choice inferInstance := by diff --git a/Mathlib/Analysis/Normed/Ring/Units.lean b/Mathlib/Analysis/Normed/Ring/Units.lean index 86883d4c5d017..eed32209f935a 100644 --- a/Mathlib/Analysis/Normed/Ring/Units.lean +++ b/Mathlib/Analysis/Normed/Ring/Units.lean @@ -200,15 +200,18 @@ open MulOpposite Filter NormedRing /-- In a normed ring with summable geometric series, the coercion from `Rˣ` (equipped with the induced topology from the embedding in `R × R`) to `R` is an open embedding. -/ -theorem openEmbedding_val : OpenEmbedding (val : Rˣ → R) where +theorem isOpenEmbedding_val : IsOpenEmbedding (val : Rˣ → R) where toEmbedding := embedding_val_mk' (fun _ ⟨u, hu⟩ ↦ hu ▸ (inverse_continuousAt u).continuousWithinAt) Ring.inverse_unit isOpen_range := Units.isOpen +@[deprecated (since := "2024-10-18")] +alias openEmbedding_val := isOpenEmbedding_val + /-- In a normed ring with summable geometric series, the coercion from `Rˣ` (equipped with the induced topology from the embedding in `R × R`) to `R` is an open map. -/ theorem isOpenMap_val : IsOpenMap (val : Rˣ → R) := - openEmbedding_val.isOpenMap + isOpenEmbedding_val.isOpenMap end Units diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index 5d7dc418a5c49..9a02b4d3dc803 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -369,16 +369,19 @@ theorem tendsto_exp_comp_nhds_zero {f : α → ℝ} : Tendsto (fun x => exp (f x)) l (𝓝 0) ↔ Tendsto f l atBot := by simp_rw [← comp_apply (f := exp), ← tendsto_comap_iff, comap_exp_nhds_zero] -theorem openEmbedding_exp : OpenEmbedding exp := - isOpen_Ioi.openEmbedding_subtype_val.comp expOrderIso.toHomeomorph.openEmbedding +theorem isOpenEmbedding_exp : IsOpenEmbedding exp := + isOpen_Ioi.isOpenEmbedding_subtypeVal.comp expOrderIso.toHomeomorph.isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_exp := isOpenEmbedding_exp @[simp] theorem map_exp_nhds (x : ℝ) : map exp (𝓝 x) = 𝓝 (exp x) := - openEmbedding_exp.map_nhds_eq x + isOpenEmbedding_exp.map_nhds_eq x @[simp] theorem comap_exp_nhds_exp (x : ℝ) : comap exp (𝓝 (exp x)) = 𝓝 x := - (openEmbedding_exp.nhds_eq_comap x).symm + (isOpenEmbedding_exp.nhds_eq_comap x).symm theorem isLittleO_pow_exp_atTop {n : ℕ} : (fun x : ℝ => x ^ n) =o[atTop] Real.exp := by simpa [isLittleO_iff_tendsto fun x hx => ((exp_pos x).ne' hx).elim] using diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 5acd6a8cea119..006514be87487 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -1049,7 +1049,7 @@ variable (e : PartialHomeomorph α H) /-- If a single partial homeomorphism `e` from a space `α` into `H` has source covering the whole space `α`, then that partial homeomorphism induces an `H`-charted space structure on `α`. (This condition is equivalent to `e` being an open embedding of `α` into `H`; see -`OpenEmbedding.singletonChartedSpace`.) -/ +`IsOpenEmbedding.singletonChartedSpace`.) -/ def singletonChartedSpace (h : e.source = Set.univ) : ChartedSpace H α where atlas := {e} chartAt _ := e @@ -1085,24 +1085,24 @@ theorem singleton_hasGroupoid (h : e.source = Set.univ) (G : StructureGroupoid H end PartialHomeomorph -namespace OpenEmbedding +namespace IsOpenEmbedding variable [Nonempty α] /-- An open embedding of `α` into `H` induces an `H`-charted space structure on `α`. See `PartialHomeomorph.singletonChartedSpace`. -/ -def singletonChartedSpace {f : α → H} (h : OpenEmbedding f) : ChartedSpace H α := +def singletonChartedSpace {f : α → H} (h : IsOpenEmbedding f) : ChartedSpace H α := (h.toPartialHomeomorph f).singletonChartedSpace (toPartialHomeomorph_source _ _) -theorem singletonChartedSpace_chartAt_eq {f : α → H} (h : OpenEmbedding f) {x : α} : +theorem singletonChartedSpace_chartAt_eq {f : α → H} (h : IsOpenEmbedding f) {x : α} : ⇑(@chartAt H _ α _ h.singletonChartedSpace x) = f := rfl -theorem singleton_hasGroupoid {f : α → H} (h : OpenEmbedding f) (G : StructureGroupoid H) +theorem singleton_hasGroupoid {f : α → H} (h : IsOpenEmbedding f) (G : StructureGroupoid H) [ClosedUnderRestriction G] : @HasGroupoid _ _ _ _ h.singletonChartedSpace G := (h.toPartialHomeomorph f).singleton_hasGroupoid (toPartialHomeomorph_source _ _) G -end OpenEmbedding +end IsOpenEmbedding end Singleton diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index cd4f4365ba286..b79ff42cc1ad9 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -13,7 +13,7 @@ In this file, we show that standard operations on smooth maps between smooth man * `contMDiff_id` gives the smoothness of the identity * `contMDiff_const` gives the smoothness of constant functions * `contMDiff_inclusion` shows that the inclusion between open sets of a topological space is smooth -* `contMDiff_openEmbedding` shows that if `M` has a `ChartedSpace` structure induced by an open +* `contMDiff_isOpenEmbedding` shows that if `M` has a `ChartedSpace` structure induced by an open embedding `e : M → H`, then `e` is smooth. ## Tags @@ -358,11 +358,11 @@ end ChartedSpace section -variable {e : M → H} (h : OpenEmbedding e) {n : WithTop ℕ} +variable {e : M → H} (h : IsOpenEmbedding e) {n : WithTop ℕ} /-- If the `ChartedSpace` structure on a manifold `M` is given by an open embedding `e : M → H`, then `e` is smooth. -/ -lemma contMDiff_openEmbedding [Nonempty M] : +lemma contMDiff_isOpenEmbedding [Nonempty M] : haveI := h.singletonChartedSpace; ContMDiff I I n e := by haveI := h.singleton_smoothManifoldWithCorners I rw [@contMDiff_iff _ _ _ _ _ _ _ _ _ _ h.singletonChartedSpace] @@ -385,12 +385,15 @@ lemma contMDiff_openEmbedding [Nonempty M] : h.toPartialHomeomorph_target] at this exact this +@[deprecated (since := "2024-10-18")] +alias contMDiff_openEmbedding := contMDiff_isOpenEmbedding + variable {I I'} /-- If the `ChartedSpace` structure on a manifold `M` is given by an open embedding `e : M → H`, then the inverse of `e` is smooth. -/ -lemma contMDiffOn_openEmbedding_symm [Nonempty M] : +lemma contMDiffOn_isOpenEmbedding_symm [Nonempty M] : haveI := h.singletonChartedSpace; ContMDiffOn I I - n (OpenEmbedding.toPartialHomeomorph e h).symm (range e) := by + n (IsOpenEmbedding.toPartialHomeomorph e h).symm (range e) := by haveI := h.singleton_smoothManifoldWithCorners I rw [@contMDiffOn_iff] constructor @@ -409,21 +412,27 @@ lemma contMDiffOn_openEmbedding_symm [Nonempty M] : apply I.right_inv exact mem_of_subset_of_mem (extChartAt_target_subset_range _ _) hz.1 +@[deprecated (since := "2024-10-18")] +alias contMDiffOn_openEmbedding_symm := contMDiffOn_isOpenEmbedding_symm + variable [ChartedSpace H M] -variable [Nonempty M'] {e' : M' → H'} (h' : OpenEmbedding e') +variable [Nonempty M'] {e' : M' → H'} (h' : IsOpenEmbedding e') /-- Let `M'` be a manifold whose chart structure is given by an open embedding `e'` into its model space `H'`. Then the smoothness of `e' ∘ f : M → H'` implies the smoothness of `f`. This is useful, for example, when `e' ∘ f = g ∘ e` for smooth maps `e : M → X` and `g : X → H'`. -/ -lemma ContMDiff.of_comp_openEmbedding {f : M → M'} (hf : ContMDiff I I' n (e' ∘ f)) : +lemma ContMDiff.of_comp_isOpenEmbedding {f : M → M'} (hf : ContMDiff I I' n (e' ∘ f)) : haveI := h'.singletonChartedSpace; ContMDiff I I' n f := by have : f = (h'.toPartialHomeomorph e').symm ∘ e' ∘ f := by ext - rw [Function.comp_apply, Function.comp_apply, OpenEmbedding.toPartialHomeomorph_left_inv] + rw [Function.comp_apply, Function.comp_apply, IsOpenEmbedding.toPartialHomeomorph_left_inv] rw [this] apply @ContMDiffOn.comp_contMDiff _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ - h'.singletonChartedSpace _ _ (range e') _ (contMDiffOn_openEmbedding_symm h') hf + h'.singletonChartedSpace _ _ (range e') _ (contMDiffOn_isOpenEmbedding_symm h') hf simp +@[deprecated (since := "2024-10-18")] +alias ContMDiff.of_comp_openEmbedding := ContMDiff.of_comp_isOpenEmbedding + end diff --git a/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean b/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean index e7690b95e11f0..afe2f7f2112cb 100644 --- a/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean +++ b/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean @@ -33,7 +33,7 @@ namespace Units variable {R : Type*} [NormedRing R] [CompleteSpace R] instance : ChartedSpace R Rˣ := - openEmbedding_val.singletonChartedSpace + isOpenEmbedding_val.singletonChartedSpace theorem chartAt_apply {a : Rˣ} {b : Rˣ} : chartAt R a b = b := rfl @@ -44,17 +44,17 @@ theorem chartAt_source {a : Rˣ} : (chartAt R a).source = Set.univ := variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [NormedAlgebra 𝕜 R] instance : SmoothManifoldWithCorners 𝓘(𝕜, R) Rˣ := - openEmbedding_val.singleton_smoothManifoldWithCorners 𝓘(𝕜, R) + isOpenEmbedding_val.singleton_smoothManifoldWithCorners 𝓘(𝕜, R) /-- For a complete normed ring `R`, the embedding of the units `Rˣ` into `R` is a smooth map between manifolds. -/ lemma contMDiff_val {m : ℕ∞} : ContMDiff 𝓘(𝕜, R) 𝓘(𝕜, R) m (val : Rˣ → R) := - contMDiff_openEmbedding 𝓘(𝕜, R) Units.openEmbedding_val + contMDiff_isOpenEmbedding 𝓘(𝕜, R) Units.isOpenEmbedding_val /-- The units of a complete normed ring form a Lie group. -/ instance : LieGroup 𝓘(𝕜, R) Rˣ where smooth_mul := by - apply ContMDiff.of_comp_openEmbedding Units.openEmbedding_val + apply ContMDiff.of_comp_isOpenEmbedding Units.isOpenEmbedding_val have : (val : Rˣ → R) ∘ (fun x : Rˣ × Rˣ => x.1 * x.2) = (fun x : R × R => x.1 * x.2) ∘ (fun x : Rˣ × Rˣ => (x.1, x.2)) := by ext; simp rw [this] @@ -65,7 +65,7 @@ instance : LieGroup 𝓘(𝕜, R) Rˣ where rw [contMDiff_iff_contDiff] exact contDiff_mul smooth_inv := by - apply ContMDiff.of_comp_openEmbedding Units.openEmbedding_val + apply ContMDiff.of_comp_isOpenEmbedding Units.isOpenEmbedding_val have : (val : Rˣ → R) ∘ (fun x : Rˣ => x⁻¹) = Ring.inverse ∘ val := by ext; simp rw [this, ContMDiff] refine fun x => ContMDiffAt.comp x ?_ (contMDiff_val x) diff --git a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean index 5ea4b4a42e6b8..6b89039b66f57 100644 --- a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean +++ b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean @@ -490,7 +490,7 @@ theorem liftPropAt_iff_comp_subtype_val (hG : LocalInvariantProp G G' P) {U : Op LiftPropAt P f x ↔ LiftPropAt P (f ∘ Subtype.val) x := by simp only [LiftPropAt, liftPropWithinAt_iff'] congrm ?_ ∧ ?_ - · simp_rw [continuousWithinAt_univ, U.openEmbedding'.continuousAt_iff] + · simp_rw [continuousWithinAt_univ, U.isOpenEmbedding'.continuousAt_iff] · apply hG.congr_iff exact (U.chartAt_subtype_val_symm_eventuallyEq).fun_comp (chartAt H' (f x) ∘ f) @@ -500,7 +500,7 @@ theorem liftPropAt_iff_comp_inclusion (hG : LocalInvariantProp G G' P) {U V : Op simp only [LiftPropAt, liftPropWithinAt_iff'] congrm ?_ ∧ ?_ · simp_rw [continuousWithinAt_univ, - (TopologicalSpace.Opens.openEmbedding_of_le hUV).continuousAt_iff] + (TopologicalSpace.Opens.isOpenEmbedding_of_le hUV).continuousAt_iff] · apply hG.congr_iff exact (TopologicalSpace.Opens.chartAt_inclusion_symm_eventuallyEq hUV).fun_comp (chartAt H' (f (Set.inclusion hUV x)) ∘ f) diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index d6cebfb34829f..ac7279ed0be9c 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -724,13 +724,17 @@ theorem PartialHomeomorph.singleton_smoothManifoldWithCorners @SmoothManifoldWithCorners.mk' _ _ _ _ _ _ _ _ _ _ (id _) <| e.singleton_hasGroupoid h (contDiffGroupoid ∞ I) -theorem OpenEmbedding.singleton_smoothManifoldWithCorners {𝕜 : Type*} [NontriviallyNormedField 𝕜] +theorem IsOpenEmbedding.singleton_smoothManifoldWithCorners {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [Nonempty M] {f : M → H} - (h : OpenEmbedding f) : + (h : IsOpenEmbedding f) : @SmoothManifoldWithCorners 𝕜 _ E _ _ H _ I M _ h.singletonChartedSpace := (h.toPartialHomeomorph f).singleton_smoothManifoldWithCorners I (by simp) +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.singleton_smoothManifoldWithCorners := + IsOpenEmbedding.singleton_smoothManifoldWithCorners + namespace TopologicalSpace.Opens open TopologicalSpace diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean index f0aedb017be00..2fc0a8b8893ef 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean @@ -217,8 +217,8 @@ instance is_sheafedSpace_iso {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) [IsIso /-- The restriction of a locally ringed space along an open embedding. -/ @[simps!] -def restrict {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) : - LocallyRingedSpace where +def restrict {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} + (h : IsOpenEmbedding f) : LocallyRingedSpace where localRing := by intro x -- We show that the stalk of the restriction is isomorphic to the original stalk, @@ -228,13 +228,13 @@ def restrict {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h /-- The canonical map from the restriction to the subspace. -/ def ofRestrict {U : TopCat} (X : LocallyRingedSpace.{u}) - {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) : X.restrict h ⟶ X := + {f : U ⟶ X.toTopCat} (h : IsOpenEmbedding f) : X.restrict h ⟶ X := ⟨X.toPresheafedSpace.ofRestrict h, fun _ => inferInstance⟩ /-- The restriction of a locally ringed space `X` to the top subspace is isomorphic to `X` itself. -/ def restrictTopIso (X : LocallyRingedSpace.{u}) : - X.restrict (Opens.openEmbedding ⊤) ≅ X := + X.restrict (Opens.isOpenEmbedding ⊤) ≅ X := isoOfSheafedSpaceIso X.toSheafedSpace.restrictTopIso /-- The global sections, notated Gamma. @@ -436,7 +436,7 @@ theorem preimage_basicOpen {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) {U : Ope rw [← stalkMap_germ_apply] at hx exact (isUnit_map_iff (f.toShHom.stalkMap _) _).mp hx -variable {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) +variable {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h : IsOpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) /-- For an open embedding `f : U ⟶ X` and a point `x : U`, we get an isomorphism between the stalk diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 11dbae9b63d30..f549a8e9ee841 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -68,7 +68,7 @@ spaces, such that the sheaf map `Y(V) ⟶ f _* X(V)` is an iso for each `V ⊆ U class PresheafedSpace.IsOpenImmersion {X Y : PresheafedSpace C} (f : X ⟶ Y) : Prop where /-- the underlying continuous map of underlying spaces from the source to an open subset of the target. -/ - base_open : OpenEmbedding f.base + base_open : IsOpenEmbedding f.base /-- the underlying sheaf morphism is an isomorphism on each open subset -/ c_iso : ∀ U : Opens X, IsIso (f.c.app (op (base_open.isOpenMap.functor.obj U))) @@ -214,7 +214,7 @@ theorem app_inv_app' (U : Opens Y) (hU : (U : Set Y) ⊆ Set.range f.base) : /-- An isomorphism is an open immersion. -/ instance ofIso {X Y : PresheafedSpace C} (H : X ≅ Y) : IsOpenImmersion H.hom where - base_open := (TopCat.homeoOfIso ((forget C).mapIso H)).openEmbedding + base_open := (TopCat.homeoOfIso ((forget C).mapIso H)).isOpenEmbedding -- Porting note: `inferInstance` will fail if Lean is not told that `H.hom.c` is iso c_iso _ := letI : IsIso H.hom.c := c_isIso_of_iso H.hom; inferInstance @@ -223,7 +223,7 @@ instance (priority := 100) ofIsIso {X Y : PresheafedSpace C} (f : X ⟶ Y) [IsIs AlgebraicGeometry.PresheafedSpace.IsOpenImmersion.ofIso (asIso f) instance ofRestrict {X : TopCat} (Y : PresheafedSpace C) {f : X ⟶ Y.carrier} - (hf : OpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) where + (hf : IsOpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) where base_open := hf c_iso U := by dsimp @@ -243,7 +243,7 @@ instance ofRestrict {X : TopCat} (Y : PresheafedSpace C) {f : X ⟶ Y.carrier} @[elementwise, simp] theorem ofRestrict_invApp {C : Type*} [Category C] (X : PresheafedSpace C) {Y : TopCat} - {f : Y ⟶ TopCat.of X.carrier} (h : OpenEmbedding f) (U : Opens (X.restrict h).carrier) : + {f : Y ⟶ TopCat.of X.carrier} (h : IsOpenEmbedding f) (U : Opens (X.restrict h).carrier) : (PresheafedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := by delta invApp rw [IsIso.comp_inv_eq, Category.id_comp] @@ -286,7 +286,7 @@ variable {X Y Z : PresheafedSpace C} (f : X ⟶ Z) [hf : IsOpenImmersion f] (g : /-- (Implementation.) The projection map when constructing the pullback along an open immersion. -/ def pullbackConeOfLeftFst : - Y.restrict (TopCat.snd_openEmbedding_of_left_openEmbedding hf.base_open g.base) ⟶ X where + Y.restrict (TopCat.snd_isOpenEmbedding_of_left hf.base_open g.base) ⟶ X where base := pullback.fst _ _ c := { app := fun U => @@ -377,7 +377,7 @@ def pullbackConeOfLeftLift : s.pt ⟶ (pullbackConeOfLeft f g).pt where dsimp [s'] rw [Function.comp_def, ← Set.preimage_preimage] rw [Set.preimage_image_eq _ - (TopCat.snd_openEmbedding_of_left_openEmbedding hf.base_open g.base).inj] + (TopCat.snd_isOpenEmbedding_of_left hf.base_open g.base).inj] rfl)) naturality := fun U V i => by erw [s.snd.c.naturality_assoc] @@ -727,7 +727,7 @@ whose forgetful functor reflects isomorphisms, preserves limits and filtered col Then a morphism `X ⟶ Y` that is a topological open embedding is an open immersion iff every stalk map is an iso. -/ -theorem of_stalk_iso {X Y : SheafedSpace C} (f : X ⟶ Y) (hf : OpenEmbedding f.base) +theorem of_stalk_iso {X Y : SheafedSpace C} (f : X ⟶ Y) (hf : IsOpenEmbedding f.base) [H : ∀ x : X.1, IsIso (f.stalkMap x)] : SheafedSpace.IsOpenImmersion f := { base_open := hf c_iso := fun U => by @@ -738,7 +738,7 @@ theorem of_stalk_iso {X Y : SheafedSpace C} (f : X ⟶ Y) (hf : OpenEmbedding f. specialize H y delta PresheafedSpace.Hom.stalkMap at H haveI H' := - TopCat.Presheaf.stalkPushforward.stalkPushforward_iso_of_openEmbedding C hf X.presheaf y + TopCat.Presheaf.stalkPushforward.stalkPushforward_iso_of_isOpenEmbedding C hf X.presheaf y have := IsIso.comp_isIso' H (@IsIso.inv_isIso _ _ _ _ _ H') rwa [Category.assoc, IsIso.hom_inv_id, Category.comp_id] at this } @@ -812,12 +812,12 @@ theorem app_inv_app' (U : Opens Y) (hU : (U : Set Y) ⊆ Set.range f.base) : PresheafedSpace.IsOpenImmersion.app_invApp f U instance ofRestrict {X : TopCat} (Y : SheafedSpace C) {f : X ⟶ Y.carrier} - (hf : OpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) := + (hf : IsOpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) := PresheafedSpace.IsOpenImmersion.ofRestrict _ hf @[elementwise, simp] theorem ofRestrict_invApp {C : Type*} [Category C] (X : SheafedSpace C) {Y : TopCat} - {f : Y ⟶ TopCat.of X.carrier} (h : OpenEmbedding f) (U : Opens (X.restrict h).carrier) : + {f : Y ⟶ TopCat.of X.carrier} (h : IsOpenEmbedding f) (U : Opens (X.restrict h).carrier) : (SheafedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ h U @@ -839,7 +839,7 @@ section Prod variable [HasLimits C] {ι : Type v} (F : Discrete ι ⥤ SheafedSpace.{_, v, v} C) [HasColimit F] (i : Discrete ι) -theorem sigma_ι_openEmbedding : OpenEmbedding (colimit.ι F i).base := by +theorem sigma_ι_isOpenEmbedding : IsOpenEmbedding (colimit.ι F i).base := by rw [← show _ = (colimit.ι F i).base from ι_preservesColimitsIso_inv (SheafedSpace.forget C) F i] have : _ = _ ≫ colimit.ι (Discrete.functor ((F ⋙ SheafedSpace.forget C).obj ∘ Discrete.mk)) i := HasColimit.isoOfNatIso_ι_hom Discrete.natIsoFunctor i @@ -850,17 +850,20 @@ theorem sigma_ι_openEmbedding : OpenEmbedding (colimit.ι F i).base := by rw [← Iso.eq_comp_inv] at this cases i rw [this, ← Category.assoc] - -- Porting note: `simp_rw` can't use `TopCat.openEmbedding_iff_comp_isIso` and - -- `TopCat.openEmbedding_iff_isIso_comp`. + -- Porting note: `simp_rw` can't use `TopCat.isOpenEmbedding_iff_comp_isIso` and + -- `TopCat.isOpenEmbedding_iff_isIso_comp`. -- See https://github.com/leanprover-community/mathlib4/issues/5026 - erw [TopCat.openEmbedding_iff_comp_isIso, TopCat.openEmbedding_iff_comp_isIso, - TopCat.openEmbedding_iff_comp_isIso, TopCat.openEmbedding_iff_isIso_comp] - exact openEmbedding_sigmaMk + erw [TopCat.isOpenEmbedding_iff_comp_isIso, TopCat.isOpenEmbedding_iff_comp_isIso, + TopCat.isOpenEmbedding_iff_comp_isIso, TopCat.isOpenEmbedding_iff_isIso_comp] + exact isOpenEmbedding_sigmaMk + +@[deprecated (since := "2024-10-18")] +alias sigma_ι_openEmbedding := sigma_ι_isOpenEmbedding theorem image_preimage_is_empty (j : Discrete ι) (h : i ≠ j) (U : Opens (F.obj i)) : (Opens.map (colimit.ι (F ⋙ SheafedSpace.forgetToPresheafedSpace) j).base).obj ((Opens.map (preservesColimitIso SheafedSpace.forgetToPresheafedSpace F).inv.base).obj - ((sigma_ι_openEmbedding F i).isOpenMap.functor.obj U)) = + ((sigma_ι_isOpenEmbedding F i).isOpenMap.functor.obj U)) = ⊥ := by ext x apply iff_false_intro @@ -880,15 +883,15 @@ theorem image_preimage_is_empty (j : Discrete ι) (h : i ≠ j) (U : Opens (F.ob instance sigma_ι_isOpenImmersion [HasStrictTerminalObjects C] : SheafedSpace.IsOpenImmersion (colimit.ι F i) where - base_open := sigma_ι_openEmbedding F i + base_open := sigma_ι_isOpenEmbedding F i c_iso U := by have e : colimit.ι F i = _ := (ι_preservesColimitsIso_inv SheafedSpace.forgetToPresheafedSpace F i).symm have H : - OpenEmbedding + IsOpenEmbedding (colimit.ι (F ⋙ SheafedSpace.forgetToPresheafedSpace) i ≫ (preservesColimitIso SheafedSpace.forgetToPresheafedSpace F).inv).base := - e ▸ sigma_ι_openEmbedding F i + e ▸ sigma_ι_isOpenEmbedding F i suffices IsIso <| (colimit.ι (F ⋙ SheafedSpace.forgetToPresheafedSpace) i ≫ (preservesColimitIso SheafedSpace.forgetToPresheafedSpace F).inv).c.app <| op (H.isOpenMap.functor.obj U) by @@ -921,7 +924,7 @@ end SheafedSpace.IsOpenImmersion namespace LocallyRingedSpace.IsOpenImmersion -instance (X : LocallyRingedSpace) {U : TopCat} (f : U ⟶ X.toTopCat) (hf : OpenEmbedding f) : +instance (X : LocallyRingedSpace) {U : TopCat} (f : U ⟶ X.toTopCat) (hf : IsOpenEmbedding f) : LocallyRingedSpace.IsOpenImmersion (X.ofRestrict hf) := PresheafedSpace.IsOpenImmersion.ofRestrict X.toPresheafedSpace hf @@ -951,7 +954,7 @@ set_option synthInstance.maxHeartbeats 40000 in /-- An explicit pullback cone over `cospan f g` if `f` is an open immersion. -/ def pullbackConeOfLeft : PullbackCone f g := by refine PullbackCone.mk ?_ - (Y.ofRestrict (TopCat.snd_openEmbedding_of_left_openEmbedding H.base_open g.base)) ?_ + (Y.ofRestrict (TopCat.snd_isOpenEmbedding_of_left H.base_open g.base)) ?_ · use PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftFst f.1 g.1 intro x have := PresheafedSpace.stalkMap.congr_hom _ _ @@ -1154,7 +1157,7 @@ whose forgetful functor reflects isomorphisms, preserves limits and filtered col Then a morphism `X ⟶ Y` that is a topological open embedding is an open immersion iff every stalk map is an iso. -/ -theorem of_stalk_iso {X Y : LocallyRingedSpace} (f : X ⟶ Y) (hf : OpenEmbedding f.base) +theorem of_stalk_iso {X Y : LocallyRingedSpace} (f : X ⟶ Y) (hf : IsOpenEmbedding f.base) [stalk_iso : ∀ x : X.1, IsIso (f.stalkMap x)] : LocallyRingedSpace.IsOpenImmersion f := SheafedSpace.IsOpenImmersion.of_stalk_iso _ hf (H := stalk_iso) @@ -1223,12 +1226,12 @@ theorem app_inv_app' (U : Opens Y) (hU : (U : Set Y) ⊆ Set.range f.base) : PresheafedSpace.IsOpenImmersion.app_invApp f.1 U instance ofRestrict {X : TopCat} (Y : LocallyRingedSpace) {f : X ⟶ Y.carrier} - (hf : OpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) := + (hf : IsOpenEmbedding f) : IsOpenImmersion (Y.ofRestrict hf) := PresheafedSpace.IsOpenImmersion.ofRestrict _ hf @[elementwise, simp] theorem ofRestrict_invApp (X : LocallyRingedSpace) {Y : TopCat} - {f : Y ⟶ TopCat.of X.carrier} (h : OpenEmbedding f) (U : Opens (X.restrict h).carrier) : + {f : Y ⟶ TopCat.of X.carrier} (h : IsOpenEmbedding f) (U : Opens (X.restrict h).carrier) : (LocallyRingedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ h U diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean index 12c2d277fbd66..366aeeda7b4b0 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean @@ -293,7 +293,7 @@ section Restrict -/ @[simps] def restrict {U : TopCat} (X : PresheafedSpace C) {f : U ⟶ (X : TopCat)} - (h : OpenEmbedding f) : PresheafedSpace C where + (h : IsOpenEmbedding f) : PresheafedSpace C where carrier := U presheaf := h.isOpenMap.functor.op ⋙ X.presheaf @@ -301,7 +301,7 @@ def restrict {U : TopCat} (X : PresheafedSpace C) {f : U ⟶ (X : TopCat)} -/ @[simps] def ofRestrict {U : TopCat} (X : PresheafedSpace C) {f : U ⟶ (X : TopCat)} - (h : OpenEmbedding f) : X.restrict h ⟶ X where + (h : IsOpenEmbedding f) : X.restrict h ⟶ X where base := f c := { app := fun V => X.presheaf.map (h.isOpenMap.adjunction.counit.app V.unop).op @@ -310,8 +310,8 @@ def ofRestrict {U : TopCat} (X : PresheafedSpace C) {f : U ⟶ (X : TopCat)} rw [← map_comp, ← map_comp] rfl } -instance ofRestrict_mono {U : TopCat} (X : PresheafedSpace C) (f : U ⟶ X.1) (hf : OpenEmbedding f) : - Mono (X.ofRestrict hf) := by +instance ofRestrict_mono {U : TopCat} (X : PresheafedSpace C) (f : U ⟶ X.1) + (hf : IsOpenEmbedding f) : Mono (X.ofRestrict hf) := by haveI : Mono f := (TopCat.mono_iff_injective _).mpr hf.inj constructor intro Z g₁ g₂ eq @@ -339,14 +339,14 @@ instance ofRestrict_mono {U : TopCat} (X : PresheafedSpace C) (f : U ⟶ X.1) (h simpa using h theorem restrict_top_presheaf (X : PresheafedSpace C) : - (X.restrict (Opens.openEmbedding ⊤)).presheaf = + (X.restrict (Opens.isOpenEmbedding ⊤)).presheaf = (Opens.inclusionTopIso X.carrier).inv _* X.presheaf := by dsimp rw [Opens.inclusion'_top_functor X.carrier] rfl theorem ofRestrict_top_c (X : PresheafedSpace C) : - (X.ofRestrict (Opens.openEmbedding ⊤)).c = + (X.ofRestrict (Opens.isOpenEmbedding ⊤)).c = eqToHom (by rw [restrict_top_presheaf, ← Presheaf.Pushforward.comp_eq] @@ -365,14 +365,14 @@ theorem ofRestrict_top_c (X : PresheafedSpace C) : subspace. -/ @[simps] -def toRestrictTop (X : PresheafedSpace C) : X ⟶ X.restrict (Opens.openEmbedding ⊤) where +def toRestrictTop (X : PresheafedSpace C) : X ⟶ X.restrict (Opens.isOpenEmbedding ⊤) where base := (Opens.inclusionTopIso X.carrier).inv c := eqToHom (restrict_top_presheaf X) /-- The isomorphism from the restriction to the top subspace. -/ @[simps] -def restrictTopIso (X : PresheafedSpace C) : X.restrict (Opens.openEmbedding ⊤) ≅ X where +def restrictTopIso (X : PresheafedSpace C) : X.restrict (Opens.isOpenEmbedding ⊤) ≅ X where hom := X.ofRestrict _ inv := X.toRestrictTop hom_inv_id := by diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index bcfd0917ee844..16c8a58289001 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -121,14 +121,17 @@ abbrev toTopGlueData : TopCat.GlueData := { f_open := fun i j => (D.f_open i j).base_open toGlueData := 𝖣.mapGlueData (forget C) } -theorem ι_openEmbedding [HasLimits C] (i : D.J) : OpenEmbedding (𝖣.ι i).base := by +theorem ι_isOpenEmbedding [HasLimits C] (i : D.J) : IsOpenEmbedding (𝖣.ι i).base := by rw [← show _ = (𝖣.ι i).base from 𝖣.ι_gluedIso_inv (PresheafedSpace.forget _) _] -- Porting note: added this erewrite erw [coe_comp] refine - OpenEmbedding.comp - (TopCat.homeoOfIso (𝖣.gluedIso (PresheafedSpace.forget _)).symm).openEmbedding - (D.toTopGlueData.ι_openEmbedding i) + IsOpenEmbedding.comp + (TopCat.homeoOfIso (𝖣.gluedIso (PresheafedSpace.forget _)).symm).isOpenEmbedding + (D.toTopGlueData.ι_isOpenEmbedding i) + +@[deprecated (since := "2024-10-18")] +alias ι_openEmbedding := ι_isOpenEmbedding theorem pullback_base (i j k : D.J) (S : Set (D.V (i, j)).carrier) : (π₂ i, j, k) '' ((π₁ i, j, k) ⁻¹' S) = D.f i k ⁻¹' (D.f i j '' S) := by @@ -237,7 +240,7 @@ theorem snd_invApp_t_app (i j k : D.J) (U : Opens (pullback (D.f i j) (D.f i k)) variable [HasLimits C] theorem ι_image_preimage_eq (i j : D.J) (U : Opens (D.U i).carrier) : - (Opens.map (𝖣.ι j).base).obj ((D.ι_openEmbedding i).isOpenMap.functor.obj U) = + (Opens.map (𝖣.ι j).base).obj ((D.ι_isOpenEmbedding i).isOpenMap.functor.obj U) = (opensFunctor (D.f j i)).obj ((Opens.map (𝖣.t j i).base).obj ((Opens.map (𝖣.f i j).base).obj U)) := by ext1 @@ -265,7 +268,7 @@ theorem ι_image_preimage_eq (i j : D.J) (U : Opens (D.U i).carrier) : def opensImagePreimageMap (i j : D.J) (U : Opens (D.U i).carrier) : (D.U i).presheaf.obj (op U) ⟶ (D.U j).presheaf.obj (op <| - (Opens.map (𝖣.ι j).base).obj ((D.ι_openEmbedding i).isOpenMap.functor.obj U)) := + (Opens.map (𝖣.ι j).base).obj ((D.ι_isOpenEmbedding i).isOpenMap.functor.obj U)) := (D.f i j).c.app (op U) ≫ (D.t j i).c.app _ ≫ (D.f_open j i).invApp _ (unop _) ≫ @@ -311,7 +314,7 @@ the image `ι '' U` in the glued space is the limit of this diagram. -/ abbrev diagramOverOpen {i : D.J} (U : Opens (D.U i).carrier) : -- Porting note : ↓ these need to be explicit (WalkingMultispan D.diagram.fstFrom D.diagram.sndFrom)ᵒᵖ ⥤ C := - componentwiseDiagram 𝖣.diagram.multispan ((D.ι_openEmbedding i).isOpenMap.functor.obj U) + componentwiseDiagram 𝖣.diagram.multispan ((D.ι_isOpenEmbedding i).isOpenMap.functor.obj U) /-- (Implementation) The projection from the limit of `diagram_over_open` to a component of `D.U j`. -/ @@ -490,7 +493,7 @@ instance componentwise_diagram_π_isIso (i : D.J) (U : Opens (D.U i).carrier) : exact Iso.inv_hom_id ((D.U i).presheaf.mapIso (eqToIso _)) instance ιIsOpenImmersion (i : D.J) : IsOpenImmersion (𝖣.ι i) where - base_open := D.ι_openEmbedding i + base_open := D.ι_isOpenEmbedding i c_iso U := by erw [← colimitPresheafObjIsoComponentwiseLimit_hom_π]; infer_instance /-- The following diagram is a pullback, i.e. `Vᵢⱼ` is the intersection of `Uᵢ` and `Uⱼ` in `X`. diff --git a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean index 3af5f2d7e91d0..4a94f1819e128 100644 --- a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean @@ -156,20 +156,20 @@ open TopCat.Presheaf /-- The restriction of a sheafed space along an open embedding into the space. -/ -def restrict {U : TopCat} (X : SheafedSpace C) {f : U ⟶ (X : TopCat)} (h : OpenEmbedding f) : +def restrict {U : TopCat} (X : SheafedSpace C) {f : U ⟶ (X : TopCat)} (h : IsOpenEmbedding f) : SheafedSpace C := - { X.toPresheafedSpace.restrict h with IsSheaf := isSheaf_of_openEmbedding h X.IsSheaf } + { X.toPresheafedSpace.restrict h with IsSheaf := isSheaf_of_isOpenEmbedding h X.IsSheaf } /-- The map from the restriction of a presheafed space. -/ @[simps!] def ofRestrict {U : TopCat} (X : SheafedSpace C) {f : U ⟶ (X : TopCat)} - (h : OpenEmbedding f) : X.restrict h ⟶ X := X.toPresheafedSpace.ofRestrict h + (h : IsOpenEmbedding f) : X.restrict h ⟶ X := X.toPresheafedSpace.ofRestrict h /-- The restriction of a sheafed space `X` to the top subspace is isomorphic to `X` itself. -/ @[simps! hom inv] -def restrictTopIso (X : SheafedSpace C) : X.restrict (Opens.openEmbedding ⊤) ≅ X := +def restrictTopIso (X : SheafedSpace C) : X.restrict (Opens.isOpenEmbedding ⊤) ≅ X := isoMk (X.toPresheafedSpace.restrictTopIso) /-- The global sections, notated Gamma. diff --git a/Mathlib/Geometry/RingedSpace/Stalks.lean b/Mathlib/Geometry/RingedSpace/Stalks.lean index 12c20e0348392..8f6f8f792d874 100644 --- a/Mathlib/Geometry/RingedSpace/Stalks.lean +++ b/Mathlib/Geometry/RingedSpace/Stalks.lean @@ -57,7 +57,7 @@ section Restrict of `X` at `f x` and the stalk of the restriction of `X` along `f` at t `x`. -/ def restrictStalkIso {U : TopCat} (X : PresheafedSpace.{_, _, v} C) {f : U ⟶ (X : TopCat.{v})} - (h : OpenEmbedding f) (x : U) : (X.restrict h).presheaf.stalk x ≅ X.presheaf.stalk (f x) := + (h : IsOpenEmbedding f) (x : U) : (X.restrict h).presheaf.stalk x ≅ X.presheaf.stalk (f x) := haveI := initial_of_adjunction (h.isOpenMap.adjunctionNhds x) Final.colimitIso (h.isOpenMap.functorNhds x).op ((OpenNhds.inclusion (f x)).op ⋙ X.presheaf) -- As a left adjoint, the functor `h.is_open_map.functor_nhds x` is initial. @@ -67,7 +67,7 @@ def restrictStalkIso {U : TopCat} (X : PresheafedSpace.{_, _, v} C) {f : U ⟶ ( -- Porting note (#11119): removed `simp` attribute, for left hand side is not in simple normal form. @[elementwise, reassoc] theorem restrictStalkIso_hom_eq_germ {U : TopCat} (X : PresheafedSpace.{_, _, v} C) - {f : U ⟶ (X : TopCat.{v})} (h : OpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) : + {f : U ⟶ (X : TopCat.{v})} (h : IsOpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) : (X.restrict h).presheaf.germ _ x hx ≫ (restrictStalkIso X h x).hom = X.presheaf.germ (h.isOpenMap.functor.obj V) (f x) ⟨x, hx, rfl⟩ := colimit.ι_pre ((OpenNhds.inclusion (f x)).op ⋙ X.presheaf) (h.isOpenMap.functorNhds x).op @@ -77,14 +77,14 @@ theorem restrictStalkIso_hom_eq_germ {U : TopCat} (X : PresheafedSpace.{_, _, v} -- as the simpNF linter claims they never apply. @[simp, elementwise, reassoc] theorem restrictStalkIso_inv_eq_germ {U : TopCat} (X : PresheafedSpace.{_, _, v} C) - {f : U ⟶ (X : TopCat.{v})} (h : OpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) : + {f : U ⟶ (X : TopCat.{v})} (h : IsOpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) : X.presheaf.germ (h.isOpenMap.functor.obj V) (f x) ⟨x, hx, rfl⟩ ≫ (restrictStalkIso X h x).inv = (X.restrict h).presheaf.germ _ x hx := by rw [← restrictStalkIso_hom_eq_germ, Category.assoc, Iso.hom_inv_id, Category.comp_id] theorem restrictStalkIso_inv_eq_ofRestrict {U : TopCat} (X : PresheafedSpace.{_, _, v} C) - {f : U ⟶ (X : TopCat.{v})} (h : OpenEmbedding f) (x : U) : + {f : U ⟶ (X : TopCat.{v})} (h : IsOpenEmbedding f) (x : U) : (X.restrictStalkIso h x).inv = (X.ofRestrict h).stalkMap x := by -- We can't use `ext` here due to https://github.com/leanprover/std4/pull/159 refine colimit.hom_ext fun V => ?_ @@ -99,7 +99,7 @@ theorem restrictStalkIso_inv_eq_ofRestrict {U : TopCat} (X : PresheafedSpace.{_, exact (colimit.w ((OpenNhds.inclusion (f x)).op ⋙ X.presheaf) i.op).symm instance ofRestrict_stalkMap_isIso {U : TopCat} (X : PresheafedSpace.{_, _, v} C) - {f : U ⟶ (X : TopCat.{v})} (h : OpenEmbedding f) (x : U) : + {f : U ⟶ (X : TopCat.{v})} (h : IsOpenEmbedding f) (x : U) : IsIso ((X.ofRestrict h).stalkMap x) := by rw [← restrictStalkIso_inv_eq_ofRestrict]; infer_instance diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index a09f3bc5eb9bd..601ec698d26dd 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -636,10 +636,13 @@ protected theorem ClosedEmbedding.measurableEmbedding {f : α → β} (h : Close MeasurableEmbedding f := h.toEmbedding.measurableEmbedding h.isClosed_range.measurableSet -protected theorem OpenEmbedding.measurableEmbedding {f : α → β} (h : OpenEmbedding f) : +protected theorem IsOpenEmbedding.measurableEmbedding {f : α → β} (h : IsOpenEmbedding f) : MeasurableEmbedding f := h.toEmbedding.measurableEmbedding h.isOpen_range.measurableSet +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.measurableEmbedding := IsOpenEmbedding.measurableEmbedding + instance Empty.borelSpace : BorelSpace Empty := ⟨borel_eq_top_of_discrete.symm⟩ diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 7268d68dd4275..408d7863a320a 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -190,7 +190,7 @@ lemma eisensteinSeries_tendstoLocallyUniformlyOn {k : ℤ} {N : ℕ} (hk : 3 ≤ apply TendstoLocallyUniformlyOn.comp (s := ⊤) _ _ _ (PartialHomeomorph.continuousOn_symm _) · simp only [SlashInvariantForm.toFun_eq_coe, Set.top_eq_univ, tendstoLocallyUniformlyOn_univ] apply eisensteinSeries_tendstoLocallyUniformly hk - · simp only [OpenEmbedding.toPartialHomeomorph_target, Set.top_eq_univ, mapsTo_range_iff, + · simp only [IsOpenEmbedding.toPartialHomeomorph_target, Set.top_eq_univ, mapsTo_range_iff, Set.mem_univ, forall_const] end summability diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index 7f721b2b8dbdd..e8e5e5ff0b50e 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -258,7 +258,7 @@ theorem subset_interior_smul_right {s : Set G} {t : Set α} : s • interior t @[to_additive (attr := simp)] theorem smul_mem_nhds_smul_iff {t : Set α} (g : G) {a : α} : g • t ∈ 𝓝 (g • a) ↔ t ∈ 𝓝 a := - (Homeomorph.smul g).openEmbedding.image_mem_nhds + (Homeomorph.smul g).isOpenEmbedding.image_mem_nhds @[to_additive] alias ⟨_, smul_mem_nhds_smul⟩ := smul_mem_nhds_smul_iff diff --git a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean index 9c399af02e789..a907bb2550a2a 100644 --- a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean +++ b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean @@ -49,15 +49,21 @@ theorem coe_mkHomeomorph_symm : ⇑(mkHomeomorph : M ≃ₜ Mᵈᵐᵃ).symm = m @[to_additive] theorem inducing_mk : Inducing (@mk M) := mkHomeomorph.inducing @[to_additive] theorem embedding_mk : Embedding (@mk M) := mkHomeomorph.embedding -@[to_additive] theorem openEmbedding_mk : OpenEmbedding (@mk M) := mkHomeomorph.openEmbedding +@[to_additive] theorem isOpenEmbedding_mk : IsOpenEmbedding (@mk M) := mkHomeomorph.isOpenEmbedding @[to_additive] theorem closedEmbedding_mk : ClosedEmbedding (@mk M) := mkHomeomorph.closedEmbedding @[to_additive] theorem quotientMap_mk : QuotientMap (@mk M) := mkHomeomorph.quotientMap +@[deprecated (since := "2024-10-18")] +alias openEmbedding_mk := isOpenEmbedding_mk + @[to_additive] theorem inducing_mk_symm : Inducing (@mk M).symm := mkHomeomorph.symm.inducing @[to_additive] theorem embedding_mk_symm : Embedding (@mk M).symm := mkHomeomorph.symm.embedding @[to_additive] -theorem openEmbedding_mk_symm : OpenEmbedding (@mk M).symm := mkHomeomorph.symm.openEmbedding +theorem isOpenEmbedding_mk_symm : IsOpenEmbedding (@mk M).symm := mkHomeomorph.symm.isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_mk_symm := isOpenEmbedding_mk_symm @[to_additive] theorem closedEmbedding_mk_symm : ClosedEmbedding (@mk M).symm := mkHomeomorph.symm.closedEmbedding @@ -108,7 +114,7 @@ instance instCompactSpace [CompactSpace M] : CompactSpace Mᵈᵐᵃ := @[to_additive] instance instLocallyCompactSpace [LocallyCompactSpace M] : LocallyCompactSpace Mᵈᵐᵃ := - openEmbedding_mk_symm.locallyCompactSpace + isOpenEmbedding_mk_symm.locallyCompactSpace @[to_additive] instance instWeaklyLocallyCompactSpace [WeaklyLocallyCompactSpace M] : diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean b/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean index 56b953c25c271..06b301dde6cb3 100644 --- a/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean +++ b/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean @@ -60,7 +60,7 @@ variable {K : Type*} [Group K] [TopologicalSpace K] [NonarchimedeanGroup K] /-- If a topological group embeds into a nonarchimedean group, then it is nonarchimedean. -/ @[to_additive] -theorem nonarchimedean_of_emb (f : G →* H) (emb : OpenEmbedding f) : NonarchimedeanGroup H := +theorem nonarchimedean_of_emb (f : G →* H) (emb : IsOpenEmbedding f) : NonarchimedeanGroup H := { is_nonarchimedean := fun U hU => have h₁ : f ⁻¹' U ∈ 𝓝 (1 : G) := by apply emb.continuous.tendsto diff --git a/Mathlib/Topology/Bases.lean b/Mathlib/Topology/Bases.lean index 46d7a2dd5d1df..e953974c7c1dd 100644 --- a/Mathlib/Topology/Bases.lean +++ b/Mathlib/Topology/Bases.lean @@ -887,8 +887,8 @@ theorem IsTopologicalBasis.sum {s : Set (Set α)} (hs : IsTopologicalBasis s) {t IsTopologicalBasis ((fun u => Sum.inl '' u) '' s ∪ (fun u => Sum.inr '' u) '' t) := by apply isTopologicalBasis_of_isOpen_of_nhds · rintro u (⟨w, hw, rfl⟩ | ⟨w, hw, rfl⟩) - · exact openEmbedding_inl.isOpenMap w (hs.isOpen hw) - · exact openEmbedding_inr.isOpenMap w (ht.isOpen hw) + · exact isOpenEmbedding_inl.isOpenMap w (hs.isOpen hw) + · exact isOpenEmbedding_inr.isOpenMap w (ht.isOpen hw) · rintro (x | x) u hxu u_open · obtain ⟨v, vs, xv, vu⟩ : ∃ v ∈ s, x ∈ v ∧ v ⊆ Sum.inl ⁻¹' u := hs.exists_subset_of_mem_open hxu (isOpen_sum_iff.1 u_open).1 diff --git a/Mathlib/Topology/Category/CompHausLike/Limits.lean b/Mathlib/Topology/Category/CompHausLike/Limits.lean index 845abc60f56e7..3e5a965f2dce5 100644 --- a/Mathlib/Topology/Category/CompHausLike/Limits.lean +++ b/Mathlib/Topology/Category/CompHausLike/Limits.lean @@ -150,20 +150,26 @@ variable {P : TopCat.{u} → Prop} [HasExplicitFiniteCoproducts.{0} P] example : HasFiniteCoproducts (CompHausLike.{u} P) := inferInstance /-- The inclusion maps into the explicit finite coproduct are open embeddings. -/ -lemma finiteCoproduct.openEmbedding_ι (a : α) : - OpenEmbedding (finiteCoproduct.ι X a) := - openEmbedding_sigmaMk (σ := fun a ↦ (X a)) +lemma finiteCoproduct.isOpenEmbedding_ι (a : α) : + IsOpenEmbedding (finiteCoproduct.ι X a) := + isOpenEmbedding_sigmaMk (σ := fun a ↦ (X a)) + +@[deprecated (since := "2024-10-18")] +alias finiteCoproduct.openEmbedding_ι := finiteCoproduct.isOpenEmbedding_ι /-- The inclusion maps into the abstract finite coproduct are open embeddings. -/ -lemma Sigma.openEmbedding_ι (a : α) : - OpenEmbedding (Sigma.ι X a) := by - refine OpenEmbedding.of_comp _ (homeoOfIso ((colimit.isColimit _).coconePointUniqueUpToIso - (finiteCoproduct.isColimit X))).openEmbedding ?_ - convert finiteCoproduct.openEmbedding_ι X a +lemma Sigma.isOpenEmbedding_ι (a : α) : + IsOpenEmbedding (Sigma.ι X a) := by + refine IsOpenEmbedding.of_comp _ (homeoOfIso ((colimit.isColimit _).coconePointUniqueUpToIso + (finiteCoproduct.isColimit X))).isOpenEmbedding ?_ + convert finiteCoproduct.isOpenEmbedding_ι X a ext x change (Sigma.ι X a ≫ _) x = _ simp +@[deprecated (since := "2024-10-18")] +alias Sigma.openEmbedding_ι := Sigma.isOpenEmbedding_ι + /-- The functor to `TopCat` preserves finite coproducts if they exist. -/ instance (P) [HasExplicitFiniteCoproducts.{0} P] : PreservesFiniteCoproducts (compHausLikeToTop P) := by @@ -332,12 +338,12 @@ instance [HasExplicitPullbacksOfInclusions P] : HasPullbacksOfInclusions (CompHa theorem hasPullbacksOfInclusions (hP' : ∀ ⦃X Y B : CompHausLike.{u} P⦄ (f : X ⟶ B) (g : Y ⟶ B) - (_ : OpenEmbedding f), HasExplicitPullback f g) : + (_ : IsOpenEmbedding f), HasExplicitPullback f g) : HasExplicitPullbacksOfInclusions P := { hasProp := by intro _ _ _ f apply hP' - exact Sigma.openEmbedding_ι _ _ } + exact Sigma.isOpenEmbedding_ι _ _ } /-- The functor to `TopCat` preserves pullbacks of inclusions if they exist. -/ noncomputable instance [HasExplicitPullbacksOfInclusions P] : @@ -350,7 +356,7 @@ instance [HasExplicitPullbacksOfInclusions P] : FinitaryExtensive (CompHausLike finitaryExtensive_of_preserves_and_reflects (compHausLikeToTop P) theorem finitaryExtensive (hP' : ∀ ⦃X Y B : CompHausLike.{u} P⦄ (f : X ⟶ B) (g : Y ⟶ B) - (_ : OpenEmbedding f), HasExplicitPullback f g) : + (_ : IsOpenEmbedding f), HasExplicitPullback f g) : FinitaryExtensive (CompHausLike P) := have := hasPullbacksOfInclusions hP' finitaryExtensive_of_preserves_and_reflects (compHausLikeToTop P) diff --git a/Mathlib/Topology/Category/Stonean/Limits.lean b/Mathlib/Topology/Category/Stonean/Limits.lean index 9a344de319c22..ed7151753cb73 100644 --- a/Mathlib/Topology/Category/Stonean/Limits.lean +++ b/Mathlib/Topology/Category/Stonean/Limits.lean @@ -24,7 +24,7 @@ namespace Stonean instance : HasExplicitFiniteCoproducts.{w, u} (fun Y ↦ ExtremallyDisconnected Y) where hasProp _ := { hasProp := show ExtremallyDisconnected (Σ (_a : _), _) from inferInstance} -variable {X Y Z : Stonean} {f : X ⟶ Z} (i : Y ⟶ Z) (hi : OpenEmbedding f) +variable {X Y Z : Stonean} {f : X ⟶ Z} (i : Y ⟶ Z) (hi : IsOpenEmbedding f) include hi lemma extremallyDisconnected_preimage : ExtremallyDisconnected (i ⁻¹' (Set.range f)) where @@ -35,7 +35,7 @@ lemma extremallyDisconnected_preimage : ExtremallyDisconnected (i ⁻¹' (Set.ra rw [← (closure U).preimage_image_eq Subtype.coe_injective, ← h.1.closedEmbedding_subtype_val.closure_image_eq U] exact isOpen_induced (ExtremallyDisconnected.open_closure _ - (h.2.openEmbedding_subtype_val.isOpenMap U hU)) + (h.2.isOpenEmbedding_subtypeVal.isOpenMap U hU)) lemma extremallyDisconnected_pullback : ExtremallyDisconnected {xy : X × Y | f xy.1 = i xy.2} := have := extremallyDisconnected_preimage i hi diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index f1a2371e9013c..57a2905802c66 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -168,29 +168,41 @@ lemma isIso_of_bijective_of_isClosedMap {X Y : TopCat.{u}} (f : X ⟶ Y) inferInstanceAs <| IsIso (TopCat.isoOfHomeo e).hom -- Porting note: simpNF requested partially simped version below -theorem openEmbedding_iff_comp_isIso {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : - OpenEmbedding (f ≫ g) ↔ OpenEmbedding f := - (TopCat.homeoOfIso (asIso g)).openEmbedding.of_comp_iff f +theorem isOpenEmbedding_iff_comp_isIso {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : + IsOpenEmbedding (f ≫ g) ↔ IsOpenEmbedding f := + (TopCat.homeoOfIso (asIso g)).isOpenEmbedding.of_comp_iff f + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_comp_isIso := isOpenEmbedding_iff_comp_isIso @[simp] -theorem openEmbedding_iff_comp_isIso' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : - OpenEmbedding ((forget TopCat).map f ≫ (forget TopCat).map g) ↔ OpenEmbedding f := by +theorem isOpenEmbedding_iff_comp_isIso' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : + IsOpenEmbedding ((forget TopCat).map f ≫ (forget TopCat).map g) ↔ IsOpenEmbedding f := by simp only [← Functor.map_comp] - exact openEmbedding_iff_comp_isIso f g + exact isOpenEmbedding_iff_comp_isIso f g + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_comp_isIso' := isOpenEmbedding_iff_comp_isIso' -- Porting note: simpNF requested partially simped version below -theorem openEmbedding_iff_isIso_comp {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] : - OpenEmbedding (f ≫ g) ↔ OpenEmbedding g := by +theorem isOpenEmbedding_iff_isIso_comp {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] : + IsOpenEmbedding (f ≫ g) ↔ IsOpenEmbedding g := by constructor · intro h - convert h.comp (TopCat.homeoOfIso (asIso f).symm).openEmbedding + convert h.comp (TopCat.homeoOfIso (asIso f).symm).isOpenEmbedding exact congrArg _ (IsIso.inv_hom_id_assoc f g).symm - · exact fun h => h.comp (TopCat.homeoOfIso (asIso f)).openEmbedding + · exact fun h => h.comp (TopCat.homeoOfIso (asIso f)).isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_isIso_comp := isOpenEmbedding_iff_isIso_comp @[simp] -theorem openEmbedding_iff_isIso_comp' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] : - OpenEmbedding ((forget TopCat).map f ≫ (forget TopCat).map g) ↔ OpenEmbedding g := by +theorem isOpenEmbedding_iff_isIso_comp' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] : + IsOpenEmbedding ((forget TopCat).map f ≫ (forget TopCat).map g) ↔ IsOpenEmbedding g := by simp only [← Functor.map_comp] - exact openEmbedding_iff_isIso_comp f g + exact isOpenEmbedding_iff_isIso_comp f g + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_isIso_comp' := isOpenEmbedding_iff_isIso_comp' end TopCat diff --git a/Mathlib/Topology/Category/TopCat/Limits/Products.lean b/Mathlib/Topology/Category/TopCat/Limits/Products.lean index cd45197d42a40..132e2144e2090 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Products.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Products.lean @@ -283,7 +283,8 @@ def binaryCofanIsColimit (X Y : TopCat.{u}) : IsColimit (TopCat.binaryCofan X Y) theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : Nonempty (IsColimit c) ↔ - OpenEmbedding c.inl ∧ OpenEmbedding c.inr ∧ IsCompl (Set.range c.inl) (Set.range c.inr) := by + IsOpenEmbedding c.inl ∧ IsOpenEmbedding c.inr ∧ + IsCompl (Set.range c.inl) (Set.range c.inr) := by classical constructor · rintro ⟨h⟩ @@ -293,9 +294,9 @@ theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : h.comp_coconePointUniqueUpToIso_inv (binaryCofanIsColimit X Y) ⟨WalkingPair.right⟩] dsimp refine ⟨(homeoOfIso <| h.coconePointUniqueUpToIso - (binaryCofanIsColimit X Y)).symm.openEmbedding.comp openEmbedding_inl, + (binaryCofanIsColimit X Y)).symm.isOpenEmbedding.comp isOpenEmbedding_inl, (homeoOfIso <| h.coconePointUniqueUpToIso - (binaryCofanIsColimit X Y)).symm.openEmbedding.comp openEmbedding_inr, ?_⟩ + (binaryCofanIsColimit X Y)).symm.isOpenEmbedding.comp isOpenEmbedding_inr, ?_⟩ erw [Set.range_comp, ← eq_compl_iff_isCompl, Set.range_comp _ Sum.inr, ← Set.image_compl_eq (homeoOfIso <| h.coconePointUniqueUpToIso (binaryCofanIsColimit X Y)).symm.bijective, Set.compl_range_inr, Set.image_comp] diff --git a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean index 7427588e292b7..e51d68906803f 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean @@ -284,10 +284,10 @@ W ⟶ Y X ⟶ Z ``` -/ -theorem pullback_map_openEmbedding_of_open_embeddings {W X Y Z S T : TopCat.{u}} (f₁ : W ⟶ S) - (f₂ : X ⟶ S) (g₁ : Y ⟶ T) (g₂ : Z ⟶ T) {i₁ : W ⟶ Y} {i₂ : X ⟶ Z} (H₁ : OpenEmbedding i₁) - (H₂ : OpenEmbedding i₂) (i₃ : S ⟶ T) [H₃ : Mono i₃] (eq₁ : f₁ ≫ i₃ = i₁ ≫ g₁) - (eq₂ : f₂ ≫ i₃ = i₂ ≫ g₂) : OpenEmbedding (pullback.map f₁ f₂ g₁ g₂ i₁ i₂ i₃ eq₁ eq₂) := by +theorem pullback_map_isOpenEmbedding {W X Y Z S T : TopCat.{u}} (f₁ : W ⟶ S) + (f₂ : X ⟶ S) (g₁ : Y ⟶ T) (g₂ : Z ⟶ T) {i₁ : W ⟶ Y} {i₂ : X ⟶ Z} (H₁ : IsOpenEmbedding i₁) + (H₂ : IsOpenEmbedding i₂) (i₃ : S ⟶ T) [H₃ : Mono i₃] (eq₁ : f₁ ≫ i₃ = i₁ ≫ g₁) + (eq₂ : f₂ ≫ i₃ = i₂ ≫ g₂) : IsOpenEmbedding (pullback.map f₁ f₂ g₁ g₂ i₁ i₂ i₃ eq₁ eq₂) := by constructor · apply pullback_map_embedding_of_embeddings f₁ f₂ g₁ g₂ H₁.toEmbedding H₂.toEmbedding i₃ eq₁ eq₂ @@ -298,6 +298,9 @@ theorem pullback_map_openEmbedding_of_open_embeddings {W X Y Z S T : TopCat.{u}} · apply ContinuousMap.continuous_toFun · exact H₂.isOpen_range +@[deprecated (since := "2024-10-18")] +alias pullback_map_openEmbedding_of_open_embeddings := pullback_map_isOpenEmbedding + theorem snd_embedding_of_left_embedding {X Y S : TopCat} {f : X ⟶ S} (H : Embedding f) (g : Y ⟶ S) : Embedding <| ⇑(pullback.snd f g) := by convert (homeoOfIso (asIso (pullback.snd (𝟙 S) g))).embedding.comp @@ -320,30 +323,39 @@ theorem embedding_of_pullback_embeddings {X Y S : TopCat} {f : X ⟶ S} {g : Y rw [← coe_comp, ← limit.w _ WalkingCospan.Hom.inr] rfl -theorem snd_openEmbedding_of_left_openEmbedding {X Y S : TopCat} {f : X ⟶ S} (H : OpenEmbedding f) - (g : Y ⟶ S) : OpenEmbedding <| ⇑(pullback.snd f g) := by - convert (homeoOfIso (asIso (pullback.snd (𝟙 S) g))).openEmbedding.comp - (pullback_map_openEmbedding_of_open_embeddings (i₂ := 𝟙 Y) f g (𝟙 _) g H - (homeoOfIso (Iso.refl _)).openEmbedding (𝟙 _) rfl (by simp)) +theorem snd_isOpenEmbedding_of_left {X Y S : TopCat} {f : X ⟶ S} (H : IsOpenEmbedding f) + (g : Y ⟶ S) : IsOpenEmbedding <| ⇑(pullback.snd f g) := by + convert (homeoOfIso (asIso (pullback.snd (𝟙 S) g))).isOpenEmbedding.comp + (pullback_map_isOpenEmbedding (i₂ := 𝟙 Y) f g (𝟙 _) g H + (homeoOfIso (Iso.refl _)).isOpenEmbedding (𝟙 _) rfl (by simp)) erw [← coe_comp] simp -theorem fst_openEmbedding_of_right_openEmbedding {X Y S : TopCat} (f : X ⟶ S) {g : Y ⟶ S} - (H : OpenEmbedding g) : OpenEmbedding <| ⇑(pullback.fst f g) := by - convert (homeoOfIso (asIso (pullback.fst f (𝟙 S)))).openEmbedding.comp - (pullback_map_openEmbedding_of_open_embeddings (i₁ := 𝟙 X) f g f (𝟙 _) - (homeoOfIso (Iso.refl _)).openEmbedding H (𝟙 _) rfl (by simp)) +@[deprecated (since := "2024-10-18")] +alias snd_openEmbedding_of_left_openEmbedding := snd_isOpenEmbedding_of_left + +theorem fst_isOpenEmbedding_of_right {X Y S : TopCat} (f : X ⟶ S) {g : Y ⟶ S} + (H : IsOpenEmbedding g) : IsOpenEmbedding <| ⇑(pullback.fst f g) := by + convert (homeoOfIso (asIso (pullback.fst f (𝟙 S)))).isOpenEmbedding.comp + (pullback_map_isOpenEmbedding (i₁ := 𝟙 X) f g f (𝟙 _) + (homeoOfIso (Iso.refl _)).isOpenEmbedding H (𝟙 _) rfl (by simp)) erw [← coe_comp] simp +@[deprecated (since := "2024-10-18")] +alias fst_openEmbedding_of_right_openEmbedding := fst_isOpenEmbedding_of_right + /-- If `X ⟶ S`, `Y ⟶ S` are open embeddings, then so is `X ×ₛ Y ⟶ S`. -/ -theorem openEmbedding_of_pullback_open_embeddings {X Y S : TopCat} {f : X ⟶ S} {g : Y ⟶ S} - (H₁ : OpenEmbedding f) (H₂ : OpenEmbedding g) : - OpenEmbedding (limit.π (cospan f g) WalkingCospan.one) := by - convert H₂.comp (snd_openEmbedding_of_left_openEmbedding H₁ g) +theorem isOpenEmbedding_of_pullback_open_embeddings {X Y S : TopCat} {f : X ⟶ S} {g : Y ⟶ S} + (H₁ : IsOpenEmbedding f) (H₂ : IsOpenEmbedding g) : + IsOpenEmbedding (limit.π (cospan f g) WalkingCospan.one) := by + convert H₂.comp (snd_isOpenEmbedding_of_left H₁ g) rw [← coe_comp, ← limit.w _ WalkingCospan.Hom.inr] rfl +@[deprecated (since := "2024-10-18")] +alias openEmbedding_of_pullback_open_embeddings := isOpenEmbedding_of_pullback_open_embeddings + theorem fst_iso_of_right_embedding_range_subset {X Y S : TopCat} (f : X ⟶ S) {g : Y ⟶ S} (hg : Embedding g) (H : Set.range f ⊆ Set.range g) : IsIso (pullback.fst f g) := by diff --git a/Mathlib/Topology/Category/TopCat/OpenNhds.lean b/Mathlib/Topology/Category/TopCat/OpenNhds.lean index 863d3a8c4d612..fdfe88f737c64 100644 --- a/Mathlib/Topology/Category/TopCat/OpenNhds.lean +++ b/Mathlib/Topology/Category/TopCat/OpenNhds.lean @@ -86,8 +86,11 @@ def inclusion (x : X) : OpenNhds x ⥤ Opens X := theorem inclusion_obj (x : X) (U) (p) : (inclusion x).obj ⟨U, p⟩ = U := rfl -theorem openEmbedding {x : X} (U : OpenNhds x) : OpenEmbedding U.1.inclusion' := - U.1.openEmbedding +theorem isOpenEmbedding {x : X} (U : OpenNhds x) : IsOpenEmbedding U.1.inclusion' := + U.1.isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias openEmbedding := isOpenEmbedding /-- The preimage functor from neighborhoods of `f x` to neighborhoods of `x`. -/ def map (x : X) : OpenNhds (f x) ⥤ OpenNhds x where diff --git a/Mathlib/Topology/Category/TopCat/Opens.lean b/Mathlib/Topology/Category/TopCat/Opens.lean index eb8eb48fb88be..de97a8d22ea44 100644 --- a/Mathlib/Topology/Category/TopCat/Opens.lean +++ b/Mathlib/Topology/Category/TopCat/Opens.lean @@ -119,8 +119,11 @@ def inclusion' {X : TopCat.{u}} (U : Opens X) : (toTopCat X).obj U ⟶ X where theorem coe_inclusion' {X : TopCat} {U : Opens X} : (inclusion' U : U → X) = Subtype.val := rfl -theorem openEmbedding {X : TopCat.{u}} (U : Opens X) : OpenEmbedding (inclusion' U) := - IsOpen.openEmbedding_subtype_val U.2 +theorem isOpenEmbedding {X : TopCat.{u}} (U : Opens X) : IsOpenEmbedding (inclusion' U) := + IsOpen.isOpenEmbedding_subtypeVal U.2 + +@[deprecated (since := "2024-10-18")] +alias openEmbedding := isOpenEmbedding /-- The inclusion of the top open subset (i.e. the whole space) is an isomorphism. -/ @@ -294,20 +297,26 @@ instance IsOpenMap.functorFullOfMono {X Y : TopCat} {f : X ⟶ Y} (hf : IsOpenMa instance IsOpenMap.functor_faithful {X Y : TopCat} {f : X ⟶ Y} (hf : IsOpenMap f) : hf.functor.Faithful where -lemma OpenEmbedding.functor_obj_injective {X Y : TopCat} {f : X ⟶ Y} (hf : OpenEmbedding f) : +lemma IsOpenEmbedding.functor_obj_injective {X Y : TopCat} {f : X ⟶ Y} (hf : IsOpenEmbedding f) : Function.Injective hf.isOpenMap.functor.obj := fun _ _ e ↦ Opens.ext (Set.image_injective.mpr hf.inj (congr_arg (↑· : Opens Y → Set Y) e)) +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.functor_obj_injective := IsOpenEmbedding.functor_obj_injective + namespace TopologicalSpace.Opens open TopologicalSpace @[simp] -theorem openEmbedding_obj_top {X : TopCat} (U : Opens X) : - U.openEmbedding.isOpenMap.functor.obj ⊤ = U := by +theorem isOpenEmbedding_obj_top {X : TopCat} (U : Opens X) : + U.isOpenEmbedding.isOpenMap.functor.obj ⊤ = U := by ext1 exact Set.image_univ.trans Subtype.range_coe +@[deprecated (since := "2024-10-18")] +alias openEmbedding_obj_top := isOpenEmbedding_obj_top + @[simp] theorem inclusion'_map_eq_top {X : TopCat} (U : Opens X) : (Opens.map U.inclusion').obj U = ⊤ := by ext1 @@ -315,10 +324,10 @@ theorem inclusion'_map_eq_top {X : TopCat} (U : Opens X) : (Opens.map U.inclusio @[simp] theorem adjunction_counit_app_self {X : TopCat} (U : Opens X) : - U.openEmbedding.isOpenMap.adjunction.counit.app U = eqToHom (by simp) := Subsingleton.elim _ _ + U.isOpenEmbedding.isOpenMap.adjunction.counit.app U = eqToHom (by simp) := Subsingleton.elim _ _ theorem inclusion'_top_functor (X : TopCat) : - (@Opens.openEmbedding X ⊤).isOpenMap.functor = map (inclusionTopIso X).inv := by + (@Opens.isOpenEmbedding X ⊤).isOpenMap.functor = map (inclusionTopIso X).inv := by refine CategoryTheory.Functor.ext ?_ ?_ · intro U ext x @@ -347,23 +356,23 @@ lemma set_range_inclusion' {X : TopCat} (U : Opens X) : @[simp] theorem functor_map_eq_inf {X : TopCat} (U V : Opens X) : - U.openEmbedding.isOpenMap.functor.obj ((Opens.map U.inclusion').obj V) = V ⊓ U := by + U.isOpenEmbedding.isOpenMap.functor.obj ((Opens.map U.inclusion').obj V) = V ⊓ U := by ext1 simp only [IsOpenMap.functor_obj_coe, map_coe, coe_inf, Set.image_preimage_eq_inter_range, set_range_inclusion' U] -theorem map_functor_eq' {X U : TopCat} (f : U ⟶ X) (hf : OpenEmbedding f) (V) : +theorem map_functor_eq' {X U : TopCat} (f : U ⟶ X) (hf : IsOpenEmbedding f) (V) : ((Opens.map f).obj <| hf.isOpenMap.functor.obj V) = V := Opens.ext <| Set.preimage_image_eq _ hf.inj @[simp] theorem map_functor_eq {X : TopCat} {U : Opens X} (V : Opens U) : - ((Opens.map U.inclusion').obj <| U.openEmbedding.isOpenMap.functor.obj V) = V := - TopologicalSpace.Opens.map_functor_eq' _ U.openEmbedding V + ((Opens.map U.inclusion').obj <| U.isOpenEmbedding.isOpenMap.functor.obj V) = V := + TopologicalSpace.Opens.map_functor_eq' _ U.isOpenEmbedding V @[simp] theorem adjunction_counit_map_functor {X : TopCat} {U : Opens X} (V : Opens U) : - U.openEmbedding.isOpenMap.adjunction.counit.app (U.openEmbedding.isOpenMap.functor.obj V) = + U.isOpenEmbedding.isOpenMap.adjunction.counit.app (U.isOpenEmbedding.isOpenMap.functor.obj V) = eqToHom (by dsimp; rw [map_functor_eq V]) := by subsingleton diff --git a/Mathlib/Topology/Clopen.lean b/Mathlib/Topology/Clopen.lean index ff10a04698b33..511935a8e3b98 100644 --- a/Mathlib/Topology/Clopen.lean +++ b/Mathlib/Topology/Clopen.lean @@ -118,7 +118,7 @@ theorem isClopen_range_inr : IsClopen (range (Sum.inr : Y → X ⊕ Y)) := theorem isClopen_range_sigmaMk {X : ι → Type*} [∀ i, TopologicalSpace (X i)] {i : ι} : IsClopen (Set.range (@Sigma.mk ι X i)) := - ⟨closedEmbedding_sigmaMk.isClosed_range, openEmbedding_sigmaMk.isOpen_range⟩ + ⟨closedEmbedding_sigmaMk.isClosed_range, isOpenEmbedding_sigmaMk.isOpen_range⟩ protected theorem QuotientMap.isClopen_preimage {f : X → Y} (hf : QuotientMap f) {s : Set Y} : IsClopen (f ⁻¹' s) ↔ IsClopen s := diff --git a/Mathlib/Topology/Compactification/OnePoint.lean b/Mathlib/Topology/Compactification/OnePoint.lean index 622f86227bd24..e870661f60d4f 100644 --- a/Mathlib/Topology/Compactification/OnePoint.lean +++ b/Mathlib/Topology/Compactification/OnePoint.lean @@ -241,28 +241,31 @@ theorem continuous_coe : Continuous ((↑) : X → OnePoint X) := theorem isOpenMap_coe : IsOpenMap ((↑) : X → OnePoint X) := fun _ => isOpen_image_coe.2 -theorem openEmbedding_coe : OpenEmbedding ((↑) : X → OnePoint X) := - openEmbedding_of_continuous_injective_open continuous_coe coe_injective isOpenMap_coe +theorem isOpenEmbedding_coe : IsOpenEmbedding ((↑) : X → OnePoint X) := + isOpenEmbedding_of_continuous_injective_open continuous_coe coe_injective isOpenMap_coe + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_coe := isOpenEmbedding_coe theorem isOpen_range_coe : IsOpen (range ((↑) : X → OnePoint X)) := - openEmbedding_coe.isOpen_range + isOpenEmbedding_coe.isOpen_range theorem isClosed_infty : IsClosed ({∞} : Set (OnePoint X)) := by rw [← compl_range_coe, isClosed_compl_iff] exact isOpen_range_coe theorem nhds_coe_eq (x : X) : 𝓝 ↑x = map ((↑) : X → OnePoint X) (𝓝 x) := - (openEmbedding_coe.map_nhds_eq x).symm + (isOpenEmbedding_coe.map_nhds_eq x).symm theorem nhdsWithin_coe_image (s : Set X) (x : X) : 𝓝[(↑) '' s] (x : OnePoint X) = map (↑) (𝓝[s] x) := - (openEmbedding_coe.toEmbedding.map_nhdsWithin_eq _ _).symm + (isOpenEmbedding_coe.toEmbedding.map_nhdsWithin_eq _ _).symm theorem nhdsWithin_coe (s : Set (OnePoint X)) (x : X) : 𝓝[s] ↑x = map (↑) (𝓝[(↑) ⁻¹' s] x) := - (openEmbedding_coe.map_nhdsWithin_preimage_eq _ _).symm + (isOpenEmbedding_coe.map_nhdsWithin_preimage_eq _ _).symm theorem comap_coe_nhds (x : X) : comap ((↑) : X → OnePoint X) (𝓝 x) = 𝓝 x := - (openEmbedding_coe.toInducing.nhds_eq_comap x).symm + (isOpenEmbedding_coe.toInducing.nhds_eq_comap x).symm /-- If `x` is not an isolated point of `X`, then `x : OnePoint X` is not an isolated point of `OnePoint X`. -/ @@ -427,18 +430,18 @@ theorem denseRange_coe [NoncompactSpace X] : DenseRange ((↑) : X → OnePoint exact dense_compl_singleton _ theorem isDenseEmbedding_coe [NoncompactSpace X] : IsDenseEmbedding ((↑) : X → OnePoint X) := - { openEmbedding_coe with dense := denseRange_coe } + { isOpenEmbedding_coe with dense := denseRange_coe } @[deprecated (since := "2024-09-30")] alias denseEmbedding_coe := isDenseEmbedding_coe @[simp, norm_cast] theorem specializes_coe {x y : X} : (x : OnePoint X) ⤳ y ↔ x ⤳ y := - openEmbedding_coe.toInducing.specializes_iff + isOpenEmbedding_coe.toInducing.specializes_iff @[simp, norm_cast] theorem inseparable_coe {x y : X} : Inseparable (x : OnePoint X) y ↔ Inseparable x y := - openEmbedding_coe.toInducing.inseparable_iff + isOpenEmbedding_coe.toInducing.inseparable_iff theorem not_specializes_infty_coe {x : X} : ¬Specializes ∞ (x : OnePoint X) := isClosed_infty.not_specializes rfl (coe_ne_infty x) diff --git a/Mathlib/Topology/Compactness/LocallyCompact.lean b/Mathlib/Topology/Compactness/LocallyCompact.lean index 8d10e1ce3da2e..5bba2a34cb127 100644 --- a/Mathlib/Topology/Compactness/LocallyCompact.lean +++ b/Mathlib/Topology/Compactness/LocallyCompact.lean @@ -199,10 +199,13 @@ protected theorem ClosedEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f (hf : ClosedEmbedding f) : LocallyCompactSpace X := hf.toInducing.locallyCompactSpace hf.isClosed_range.isLocallyClosed -protected theorem OpenEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} - (hf : OpenEmbedding f) : LocallyCompactSpace X := +protected theorem IsOpenEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} + (hf : IsOpenEmbedding f) : LocallyCompactSpace X := hf.toInducing.locallyCompactSpace hf.isOpen_range.isLocallyClosed +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.locallyCompactSpace := IsOpenEmbedding.locallyCompactSpace + protected theorem IsLocallyClosed.locallyCompactSpace [LocallyCompactSpace X] {s : Set X} (hs : IsLocallyClosed s) : LocallyCompactSpace s := embedding_subtype_val.locallyCompactSpace <| by rwa [Subtype.range_val] diff --git a/Mathlib/Topology/Connected/LocallyConnected.lean b/Mathlib/Topology/Connected/LocallyConnected.lean index 52d6a10719be2..065d40db50471 100644 --- a/Mathlib/Topology/Connected/LocallyConnected.lean +++ b/Mathlib/Topology/Connected/LocallyConnected.lean @@ -120,8 +120,8 @@ theorem locallyConnectedSpace_of_connected_bases {ι : Type*} (b : α → ι → (fun i hi => ⟨b x i, ⟨(hbasis x).mem_of_mem hi, hconnected x i hi⟩, subset_rfl⟩) fun s hs => ⟨(hbasis x).index s hs.1, ⟨(hbasis x).property_index hs.1, (hbasis x).set_index_subset hs.1⟩⟩ -theorem OpenEmbedding.locallyConnectedSpace [LocallyConnectedSpace α] [TopologicalSpace β] - {f : β → α} (h : OpenEmbedding f) : LocallyConnectedSpace β := by +theorem IsOpenEmbedding.locallyConnectedSpace [LocallyConnectedSpace α] [TopologicalSpace β] + {f : β → α} (h : IsOpenEmbedding f) : LocallyConnectedSpace β := by refine locallyConnectedSpace_of_connected_bases (fun _ s ↦ f ⁻¹' s) (fun x s ↦ (IsOpen s ∧ f x ∈ s ∧ IsConnected s) ∧ s ⊆ range f) (fun x ↦ ?_) (fun x s hxs ↦ hxs.1.2.2.isPreconnected.preimage_of_isOpenMap h.inj h.isOpenMap hxs.2) @@ -129,8 +129,11 @@ theorem OpenEmbedding.locallyConnectedSpace [LocallyConnectedSpace α] [Topologi exact LocallyConnectedSpace.open_connected_basis (f x) |>.restrict_subset (h.isOpen_range.mem_nhds <| mem_range_self _) |>.comap _ +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.locallyConnectedSpace := IsOpenEmbedding.locallyConnectedSpace + theorem IsOpen.locallyConnectedSpace [LocallyConnectedSpace α] {U : Set α} (hU : IsOpen U) : LocallyConnectedSpace U := - hU.openEmbedding_subtype_val.locallyConnectedSpace + hU.isOpenEmbedding_subtypeVal.locallyConnectedSpace end LocallyConnectedSpace diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 93ef11a0f5e30..3a8377f938164 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -1154,7 +1154,7 @@ theorem pathConnected_subset_basis {U : Set X} (h : IsOpen U) (hx : x ∈ U) : (𝓝 x).HasBasis (fun s : Set X => s ∈ 𝓝 x ∧ IsPathConnected s ∧ s ⊆ U) id := (path_connected_basis x).hasBasis_self_subset (IsOpen.mem_nhds h hx) -theorem OpenEmbedding.locPathConnectedSpace {e : Y → X} (he : OpenEmbedding e) : +theorem IsOpenEmbedding.locPathConnectedSpace {e : Y → X} (he : IsOpenEmbedding e) : LocPathConnectedSpace Y := have (y : Y) : (𝓝 y).HasBasis (fun s ↦ s ∈ 𝓝 (e y) ∧ IsPathConnected s ∧ s ⊆ range e) (e ⁻¹' ·) := @@ -1162,8 +1162,11 @@ theorem OpenEmbedding.locPathConnectedSpace {e : Y → X} (he : OpenEmbedding e) .of_bases this fun x s ⟨_, hs, hse⟩ ↦ by rwa [he.isPathConnected_iff, image_preimage_eq_of_subset hse] +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.locPathConnectedSpace := IsOpenEmbedding.locPathConnectedSpace + theorem IsOpen.locPathConnectedSpace {U : Set X} (h : IsOpen U) : LocPathConnectedSpace U := - (openEmbedding_subtype_val h).locPathConnectedSpace + (isOpenEmbedding_subtypeVal h).locPathConnectedSpace @[deprecated (since := "2024-10-17")] alias locPathConnected_of_isOpen := IsOpen.locPathConnectedSpace diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 82bce23e4c0c9..85297dd644e16 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -823,11 +823,14 @@ protected theorem IsOpenMap.prodMap {f : X → Y} {g : Z → W} (hf : 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) +protected theorem IsOpenEmbedding.prodMap {f : X → Y} {g : Z → W} (hf : IsOpenEmbedding f) + (hg : IsOpenEmbedding g) : IsOpenEmbedding (Prod.map f g) := + isOpenEmbedding_of_embedding_open (hf.1.prodMap hg.1) (hf.isOpenMap.prodMap hg.isOpenMap) -@[deprecated (since := "2024-10-05")] alias OpenEmbedding.prod := OpenEmbedding.prodMap +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.prodMap := IsOpenEmbedding.prodMap + +@[deprecated (since := "2024-10-05")] alias IsOpenEmbedding.prod := IsOpenEmbedding.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 @@ -902,23 +905,29 @@ theorem isOpenMap_inl : IsOpenMap (@inl X Y) := fun u hu => by theorem isOpenMap_inr : IsOpenMap (@inr X Y) := fun u hu => by simpa [isOpen_sum_iff, preimage_image_eq u Sum.inr_injective] -theorem openEmbedding_inl : OpenEmbedding (@inl X Y) := - openEmbedding_of_continuous_injective_open continuous_inl inl_injective isOpenMap_inl +theorem isOpenEmbedding_inl : IsOpenEmbedding (@inl X Y) := + isOpenEmbedding_of_continuous_injective_open continuous_inl inl_injective isOpenMap_inl + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_inl := isOpenEmbedding_inl + +theorem isOpenEmbedding_inr : IsOpenEmbedding (@inr X Y) := + isOpenEmbedding_of_continuous_injective_open continuous_inr inr_injective isOpenMap_inr -theorem openEmbedding_inr : OpenEmbedding (@inr X Y) := - openEmbedding_of_continuous_injective_open continuous_inr inr_injective isOpenMap_inr +@[deprecated (since := "2024-10-18")] +alias openEmbedding_inr := isOpenEmbedding_inr theorem embedding_inl : Embedding (@inl X Y) := - openEmbedding_inl.1 + isOpenEmbedding_inl.1 theorem embedding_inr : Embedding (@inr X Y) := - openEmbedding_inr.1 + isOpenEmbedding_inr.1 theorem isOpen_range_inl : IsOpen (range (inl : X → X ⊕ Y)) := - openEmbedding_inl.2 + isOpenEmbedding_inl.2 theorem isOpen_range_inr : IsOpen (range (inr : Y → X ⊕ Y)) := - openEmbedding_inr.2 + isOpenEmbedding_inr.2 theorem isClosed_range_inl : IsClosed (range (inl : X → X ⊕ Y)) := by rw [← isOpen_compl_iff, compl_range_inl] @@ -935,10 +944,10 @@ theorem closedEmbedding_inr : ClosedEmbedding (inr : Y → X ⊕ Y) := ⟨embedding_inr, isClosed_range_inr⟩ theorem nhds_inl (x : X) : 𝓝 (inl x : X ⊕ Y) = map inl (𝓝 x) := - (openEmbedding_inl.map_nhds_eq _).symm + (isOpenEmbedding_inl.map_nhds_eq _).symm theorem nhds_inr (y : Y) : 𝓝 (inr y : X ⊕ Y) = map inr (𝓝 y) := - (openEmbedding_inr.map_nhds_eq _).symm + (isOpenEmbedding_inr.map_nhds_eq _).symm @[simp] theorem continuous_sum_map {f : X → Y} {g : Z → W} : @@ -1006,12 +1015,15 @@ theorem Continuous.subtype_val {f : Y → Subtype p} (hf : Continuous f) : Continuous fun x => (f x : X) := continuous_subtype_val.comp hf -theorem IsOpen.openEmbedding_subtype_val {s : Set X} (hs : IsOpen s) : - OpenEmbedding ((↑) : s → X) := +theorem IsOpen.isOpenEmbedding_subtypeVal {s : Set X} (hs : IsOpen s) : + IsOpenEmbedding ((↑) : s → X) := ⟨embedding_subtype_val, (@Subtype.range_coe _ s).symm ▸ hs⟩ +@[deprecated (since := "2024-10-18")] +alias IsOpen.openEmbedding_subtype_val := IsOpen.isOpenEmbedding_subtypeVal + theorem IsOpen.isOpenMap_subtype_val {s : Set X} (hs : IsOpen s) : IsOpenMap ((↑) : s → X) := - hs.openEmbedding_subtype_val.isOpenMap + hs.isOpenEmbedding_subtypeVal.isOpenMap theorem IsOpenMap.restrict {f : X → Y} (hf : IsOpenMap f) {s : Set X} (hs : IsOpen s) : IsOpenMap (s.restrict f) := @@ -1123,8 +1135,8 @@ then its restriction to the preimage of an open set is a quotient map too. -/ theorem QuotientMap.restrictPreimage_isOpen {f : X → Y} (hf : QuotientMap f) {s : Set Y} (hs : IsOpen s) : QuotientMap (s.restrictPreimage f) := by refine quotientMap_iff.2 ⟨hf.surjective.restrictPreimage _, fun U ↦ ?_⟩ - rw [hs.openEmbedding_subtype_val.open_iff_image_open, ← hf.isOpen_preimage, - (hs.preimage hf.continuous).openEmbedding_subtype_val.open_iff_image_open, + rw [hs.isOpenEmbedding_subtypeVal.open_iff_image_open, ← hf.isOpen_preimage, + (hs.preimage hf.continuous).isOpenEmbedding_subtypeVal.open_iff_image_open, image_val_preimage_restrictPreimage] open scoped Set.Notation in @@ -1528,10 +1540,13 @@ theorem isClosedMap_sigmaMk {i : ι} : IsClosedMap (@Sigma.mk ι σ i) := by theorem isClosed_range_sigmaMk {i : ι} : IsClosed (range (@Sigma.mk ι σ i)) := isClosedMap_sigmaMk.isClosed_range -theorem openEmbedding_sigmaMk {i : ι} : OpenEmbedding (@Sigma.mk ι σ i) := - openEmbedding_of_continuous_injective_open continuous_sigmaMk sigma_mk_injective +theorem isOpenEmbedding_sigmaMk {i : ι} : IsOpenEmbedding (@Sigma.mk ι σ i) := + isOpenEmbedding_of_continuous_injective_open continuous_sigmaMk sigma_mk_injective isOpenMap_sigmaMk +@[deprecated (since := "2024-10-18")] +alias openEmbedding_sigmaMk := isOpenEmbedding_sigmaMk + theorem closedEmbedding_sigmaMk {i : ι} : ClosedEmbedding (@Sigma.mk ι σ i) := closedEmbedding_of_continuous_injective_closed continuous_sigmaMk sigma_mk_injective isClosedMap_sigmaMk @@ -1540,7 +1555,7 @@ theorem embedding_sigmaMk {i : ι} : Embedding (@Sigma.mk ι σ i) := closedEmbedding_sigmaMk.1 theorem Sigma.nhds_mk (i : ι) (x : σ i) : 𝓝 (⟨i, x⟩ : Sigma σ) = Filter.map (Sigma.mk i) (𝓝 x) := - (openEmbedding_sigmaMk.map_nhds_eq x).symm + (isOpenEmbedding_sigmaMk.map_nhds_eq x).symm theorem Sigma.nhds_eq (x : Sigma σ) : 𝓝 x = Filter.map (Sigma.mk x.1) (𝓝 x.2) := by cases x @@ -1600,7 +1615,7 @@ theorem isOpenMap_sigma {f : Sigma σ → X} : IsOpenMap f ↔ ∀ i, IsOpenMap theorem isOpenMap_sigma_map {f₁ : ι → κ} {f₂ : ∀ i, σ i → τ (f₁ i)} : IsOpenMap (Sigma.map f₁ f₂) ↔ ∀ i, IsOpenMap (f₂ i) := isOpenMap_sigma.trans <| - forall_congr' fun i => (@openEmbedding_sigmaMk _ _ _ (f₁ i)).isOpenMap_iff.symm + forall_congr' fun i => (@isOpenEmbedding_sigmaMk _ _ _ (f₁ i)).isOpenMap_iff.symm theorem inducing_sigma_map {f₁ : ι → κ} {f₂ : ∀ i, σ i → τ (f₁ i)} (h₁ : Injective f₁) : Inducing (Sigma.map f₁ f₂) ↔ ∀ i, Inducing (f₂ i) := by @@ -1611,11 +1626,14 @@ theorem embedding_sigma_map {f₁ : ι → κ} {f₂ : ∀ i, σ i → τ (f₁ Embedding (Sigma.map f₁ f₂) ↔ ∀ i, Embedding (f₂ i) := by simp only [embedding_iff, Injective.sigma_map, inducing_sigma_map h, forall_and, h.sigma_map_iff] -theorem openEmbedding_sigma_map {f₁ : ι → κ} {f₂ : ∀ i, σ i → τ (f₁ i)} (h : Injective f₁) : - OpenEmbedding (Sigma.map f₁ f₂) ↔ ∀ i, OpenEmbedding (f₂ i) := by - simp only [openEmbedding_iff_embedding_open, isOpenMap_sigma_map, embedding_sigma_map h, +theorem isOpenEmbedding_sigma_map {f₁ : ι → κ} {f₂ : ∀ i, σ i → τ (f₁ i)} (h : Injective f₁) : + IsOpenEmbedding (Sigma.map f₁ f₂) ↔ ∀ i, IsOpenEmbedding (f₂ i) := by + simp only [isOpenEmbedding_iff_embedding_open, isOpenMap_sigma_map, embedding_sigma_map h, forall_and] +@[deprecated (since := "2024-10-18")] +alias openEmbedding_sigma_map := isOpenEmbedding_sigma_map + end Sigma section ULift @@ -1688,4 +1706,4 @@ theorem Filter.Eventually.prod_nhdsSet {p : X × Y → Prop} {px : X → Prop} { end NhdsSet -set_option linter.style.longFile 1700 +set_option linter.style.longFile 1900 diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index d69f7aec79123..ef53ffeb6a397 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -1081,12 +1081,15 @@ theorem Embedding.map_nhdsWithin_eq {f : α → β} (hf : Embedding f) (s : Set rw [nhdsWithin, Filter.map_inf hf.inj, hf.map_nhds_eq, map_principal, ← nhdsWithin_inter', inter_eq_self_of_subset_right (image_subset_range _ _)] -theorem OpenEmbedding.map_nhdsWithin_preimage_eq {f : α → β} (hf : OpenEmbedding f) (s : Set β) +theorem IsOpenEmbedding.map_nhdsWithin_preimage_eq {f : α → β} (hf : IsOpenEmbedding f) (s : Set β) (x : α) : map f (𝓝[f ⁻¹' s] x) = 𝓝[s] f x := by rw [hf.toEmbedding.map_nhdsWithin_eq, image_preimage_eq_inter_range] apply nhdsWithin_eq_nhdsWithin (mem_range_self _) hf.isOpen_range rw [inter_assoc, inter_self] +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.map_nhdsWithin_preimage_eq := IsOpenEmbedding.map_nhdsWithin_preimage_eq + theorem QuotientMap.continuousOn_isOpen_iff {f : α → β} {g : β → γ} (h : QuotientMap f) {s : Set β} (hs : IsOpen s) : ContinuousOn g s ↔ ContinuousOn (g ∘ f) (f ⁻¹' s) := by simp only [continuousOn_iff_continuous_restrict, (h.restrictPreimage_isOpen hs).continuous_iff] diff --git a/Mathlib/Topology/Defs/Induced.lean b/Mathlib/Topology/Defs/Induced.lean index c64958ba70443..0aaa1d46b8de9 100644 --- a/Mathlib/Topology/Defs/Induced.lean +++ b/Mathlib/Topology/Defs/Induced.lean @@ -28,7 +28,7 @@ as well as topology inducing maps, topological embeddings, and quotient maps. * `Embedding`: a map `f : X → Y` is an *embedding*, if it is a topology inducing map and it is injective. -* `OpenEmbedding`: a map `f : X → Y` is an *open embedding*, +* `IsOpenEmbedding`: a map `f : X → Y` is an *open embedding*, if it is an embedding and its range is open. An open embedding is an open map. @@ -112,10 +112,13 @@ structure Embedding [TopologicalSpace X] [TopologicalSpace Y] (f : X → Y) exte /-- An open embedding is an embedding with open range. -/ @[mk_iff] -structure OpenEmbedding (f : X → Y) extends Embedding f : Prop where +structure IsOpenEmbedding (f : X → Y) extends Embedding f : Prop where /-- The range of an open embedding is an open set. -/ isOpen_range : IsOpen <| range f +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding := IsOpenEmbedding + /-- A closed embedding is an embedding with closed image. -/ @[mk_iff] structure ClosedEmbedding (f : X → Y) extends Embedding f : Prop where diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index 753550992301e..87c5a52e0a494 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -46,7 +46,7 @@ provided. * `TopCat.GlueData.preimage_range`: The preimage of the image of `U i` in `U j` is `V i j`. * `TopCat.GlueData.preimage_image_eq_image`: The preimage of the image of some `U ⊆ U i` is given by XXX. -* `TopCat.GlueData.ι_openEmbedding`: Each of the `ι i`s are open embeddings. +* `TopCat.GlueData.ι_isOpenEmbedding`: Each of the `ι i`s are open embeddings. -/ @@ -84,7 +84,7 @@ conditions are stated in a less categorical way. -/ -- porting note (#5171): removed @[nolint has_nonempty_instance] structure GlueData extends GlueData TopCat where - f_open : ∀ i j, OpenEmbedding (f i j) + f_open : ∀ i j, IsOpenEmbedding (f i j) f_mono := fun i j => (TopCat.mono_iff_injective _).mpr (f_open i j).toEmbedding.inj namespace GlueData @@ -292,10 +292,13 @@ theorem open_image_open (i : D.J) (U : Opens (𝖣.U i)) : IsOpen (𝖣.ι i '' apply (D.t j i ≫ D.f i j).continuous_toFun.isOpen_preimage exact U.isOpen -theorem ι_openEmbedding (i : D.J) : OpenEmbedding (𝖣.ι i) := - openEmbedding_of_continuous_injective_open (𝖣.ι i).continuous_toFun (D.ι_injective i) fun U h => +theorem ι_isOpenEmbedding (i : D.J) : IsOpenEmbedding (𝖣.ι i) := + isOpenEmbedding_of_continuous_injective_open (𝖣.ι i).continuous_toFun (D.ι_injective i) fun U h => D.open_image_open i ⟨U, h⟩ +@[deprecated (since := "2024-10-18")] +alias ι_openEmbedding := ι_isOpenEmbedding + /-- A family of gluing data consists of 1. An index type `J` 2. A bundled topological space `U i` for each `i : J`. @@ -356,7 +359,7 @@ def mk' (h : MkCore.{u}) : TopCat.GlueData where -- Porting note (#12129): additional beta reduction needed beta_reduce exact (h.V_id i).symm ▸ (Opens.inclusionTopIso (h.U i)).isIso_hom - f_open := fun i j : h.J => (h.V i j).openEmbedding + f_open := fun i j : h.J => (h.V i j).isOpenEmbedding t := h.t t_id i := by ext; rw [h.t_id]; rfl t' := h.t' @@ -404,7 +407,7 @@ def ofOpenSubsets : TopCat.GlueData.{u} := cocycle := fun _ _ _ _ _ => rfl } /-- The canonical map from the glue of a family of open subsets `α` into `α`. -This map is an open embedding (`fromOpenSubsetsGlue_openEmbedding`), +This map is an open embedding (`fromOpenSubsetsGlue_isOpenEmbedding`), and its range is `⋃ i, (U i : Set α)` (`range_fromOpenSubsetsGlue`). -/ def fromOpenSubsetsGlue : (ofOpenSubsets U).toGlueData.glued ⟶ TopCat.of α := @@ -439,7 +442,7 @@ theorem fromOpenSubsetsGlue_isOpenMap : IsOpenMap (fromOpenSubsetsGlue U) := by use Set.inter_subset_left constructor · rw [← Set.image_preimage_eq_inter_range] - apply (Opens.openEmbedding (X := TopCat.of α) (U i)).isOpenMap + apply (Opens.isOpenEmbedding (X := TopCat.of α) (U i)).isOpenMap convert hs i using 1 erw [← ι_fromOpenSubsetsGlue, coe_comp, Set.preimage_comp] -- porting note: `congr 1` did nothing, so I replaced it with `apply congr_arg` @@ -449,10 +452,13 @@ theorem fromOpenSubsetsGlue_isOpenMap : IsOpenMap (fromOpenSubsetsGlue U) := by rw [ι_fromOpenSubsetsGlue_apply] exact Set.mem_range_self _ -theorem fromOpenSubsetsGlue_openEmbedding : OpenEmbedding (fromOpenSubsetsGlue U) := - openEmbedding_of_continuous_injective_open (ContinuousMap.continuous_toFun _) +theorem fromOpenSubsetsGlue_isOpenEmbedding : IsOpenEmbedding (fromOpenSubsetsGlue U) := + isOpenEmbedding_of_continuous_injective_open (ContinuousMap.continuous_toFun _) (fromOpenSubsetsGlue_injective U) (fromOpenSubsetsGlue_isOpenMap U) +@[deprecated (since := "2024-10-18")] +alias fromOpenSubsetsGlue_openEmbedding := fromOpenSubsetsGlue_isOpenEmbedding + theorem range_fromOpenSubsetsGlue : Set.range (fromOpenSubsetsGlue U) = ⋃ i, (U i : Set α) := by ext constructor diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index c406c1e326c58..ce885a2ca891d 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -332,8 +332,11 @@ theorem isClosed_image (h : X ≃ₜ Y) {s : Set X} : IsClosed (h '' s) ↔ IsCl protected theorem isClosedMap (h : X ≃ₜ Y) : IsClosedMap h := fun _ => h.isClosed_image.2 -protected theorem openEmbedding (h : X ≃ₜ Y) : OpenEmbedding h := - openEmbedding_of_embedding_open h.embedding h.isOpenMap +theorem isOpenEmbedding (h : X ≃ₜ Y) : IsOpenEmbedding h := + isOpenEmbedding_of_embedding_open h.embedding h.isOpenMap + +@[deprecated (since := "2024-10-18")] +alias openEmbedding := isOpenEmbedding protected theorem closedEmbedding (h : X ≃ₜ Y) : ClosedEmbedding h := closedEmbedding_of_embedding_closed h.embedding h.isClosedMap @@ -401,7 +404,7 @@ theorem locallyConnectedSpace [i : LocallyConnectedSpace Y] (h : X ≃ₜ Y) : the domain is a locally compact space. -/ theorem locallyCompactSpace_iff (h : X ≃ₜ Y) : LocallyCompactSpace X ↔ LocallyCompactSpace Y := by - exact ⟨fun _ => h.symm.openEmbedding.locallyCompactSpace, + exact ⟨fun _ => h.symm.isOpenEmbedding.locallyCompactSpace, fun _ => h.closedEmbedding.locallyCompactSpace⟩ /-- If a bijective map `e : X ≃ Y` is continuous and open, then it is a homeomorphism. -/ @@ -908,10 +911,13 @@ protected lemma isClosedMap : IsClosedMap f := (hf.homeomorph f).isClosedMap protected lemma inducing : Inducing f := (hf.homeomorph f).inducing protected lemma quotientMap : QuotientMap f := (hf.homeomorph f).quotientMap protected lemma embedding : Embedding f := (hf.homeomorph f).embedding -protected lemma openEmbedding : OpenEmbedding f := (hf.homeomorph f).openEmbedding +lemma isOpenEmbedding : IsOpenEmbedding f := (hf.homeomorph f).isOpenEmbedding protected lemma closedEmbedding : ClosedEmbedding f := (hf.homeomorph f).closedEmbedding lemma isDenseEmbedding : IsDenseEmbedding f := (hf.homeomorph f).isDenseEmbedding +@[deprecated (since := "2024-10-18")] +alias openEmbedding := isOpenEmbedding + @[deprecated (since := "2024-09-30")] alias denseEmbedding := isDenseEmbedding @@ -932,7 +938,7 @@ lemma isHomeomorph_iff_exists_inverse : IsHomeomorph f ↔ Continuous f ∧ ∃ /-- A map is a homeomorphism iff it is a surjective embedding. -/ lemma isHomeomorph_iff_embedding_surjective : IsHomeomorph f ↔ Embedding f ∧ Surjective f where mp hf := ⟨hf.embedding, hf.surjective⟩ - mpr h := ⟨h.1.continuous, ((openEmbedding_iff f).2 ⟨h.1, h.2.range_eq ▸ isOpen_univ⟩).isOpenMap, + mpr h := ⟨h.1.continuous, ((isOpenEmbedding_iff f).2 ⟨h.1, h.2.range_eq ▸ isOpen_univ⟩).isOpenMap, h.1.inj, h.2⟩ /-- A map is a homeomorphism iff it is continuous, closed and bijective. -/ diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index a179072f0d072..85b288a5061a8 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -369,9 +369,12 @@ lemma Inducing.generalizingMap (hf : Inducing f) (h : StableUnderGeneralization obtain ⟨y, rfl⟩ := h e ⟨x, rfl⟩ exact ⟨_, hf.specializes_iff.mp e, rfl⟩ -lemma OpenEmbedding.generalizingMap (hf : OpenEmbedding f) : GeneralizingMap f := +lemma IsOpenEmbedding.generalizingMap (hf : IsOpenEmbedding f) : GeneralizingMap f := hf.toInducing.generalizingMap hf.isOpen_range.stableUnderGeneralization +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.generalizingMap := IsOpenEmbedding.generalizingMap + lemma SpecializingMap.stableUnderSpecialization_range (h : SpecializingMap f) : StableUnderSpecialization (range f) := @image_univ _ _ f ▸ stableUnderSpecialization_univ.image h diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 40974db3f9717..a6719468d25d3 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -59,11 +59,14 @@ theorem isOpen_Ico_zero : IsOpen (Ico 0 b) := by rw [ENNReal.Ico_eq_Iio] exact isOpen_Iio -theorem openEmbedding_coe : OpenEmbedding ((↑) : ℝ≥0 → ℝ≥0∞) := +theorem isOpenEmbedding_coe : IsOpenEmbedding ((↑) : ℝ≥0 → ℝ≥0∞) := ⟨embedding_coe, by rw [range_coe']; exact isOpen_Iio⟩ +@[deprecated (since := "2024-10-18")] +alias openEmbedding_coe := isOpenEmbedding_coe + theorem coe_range_mem_nhds : range ((↑) : ℝ≥0 → ℝ≥0∞) ∈ 𝓝 (r : ℝ≥0∞) := - IsOpen.mem_nhds openEmbedding_coe.isOpen_range <| mem_range_self _ + IsOpen.mem_nhds isOpenEmbedding_coe.isOpen_range <| mem_range_self _ @[norm_cast] theorem tendsto_coe {f : Filter α} {m : α → ℝ≥0} {a : ℝ≥0} : @@ -79,7 +82,7 @@ theorem continuous_coe_iff {α} [TopologicalSpace α] {f : α → ℝ≥0} : embedding_coe.continuous_iff.symm theorem nhds_coe {r : ℝ≥0} : 𝓝 (r : ℝ≥0∞) = (𝓝 r).map (↑) := - (openEmbedding_coe.map_nhds_eq r).symm + (isOpenEmbedding_coe.map_nhds_eq r).symm theorem tendsto_nhds_coe_iff {α : Type*} {l : Filter α} {x : ℝ≥0} {f : ℝ≥0∞ → α} : Tendsto f (𝓝 ↑x) l ↔ Tendsto (f ∘ (↑) : ℝ≥0 → α) (𝓝 x) l := by @@ -91,7 +94,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.prodMap openEmbedding_coe).map_nhds_eq (r, p)).symm + ((isOpenEmbedding_coe.prodMap isOpenEmbedding_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/ENat.lean b/Mathlib/Topology/Instances/ENat.lean index ee7773df5b64b..0217f6abc247c 100644 --- a/Mathlib/Topology/Instances/ENat.lean +++ b/Mathlib/Topology/Instances/ENat.lean @@ -33,11 +33,14 @@ instance : OrderTopology ℕ∞ := ⟨rfl⟩ theorem embedding_natCast : Embedding ((↑) : ℕ → ℕ∞) := Nat.strictMono_cast.embedding_of_ordConnected <| range_natCast ▸ ordConnected_Iio -theorem openEmbedding_natCast : OpenEmbedding ((↑) : ℕ → ℕ∞) := +theorem isOpenEmbedding_natCast : IsOpenEmbedding ((↑) : ℕ → ℕ∞) := ⟨embedding_natCast, range_natCast ▸ isOpen_Iio⟩ +@[deprecated (since := "2024-10-18")] +alias openEmbedding_natCast := isOpenEmbedding_natCast + theorem nhds_natCast (n : ℕ) : 𝓝 (n : ℕ∞) = pure (n : ℕ∞) := by - simp [← openEmbedding_natCast.map_nhds_eq] + simp [← isOpenEmbedding_natCast.map_nhds_eq] @[simp] protected theorem nhds_eq_pure {n : ℕ∞} (h : n ≠ ⊤) : 𝓝 n = pure n := by diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index f6f81f163a7bd..1b66243074048 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -51,9 +51,12 @@ instance : SecondCountableTopology EReal := theorem embedding_coe : Embedding ((↑) : ℝ → EReal) := coe_strictMono.embedding_of_ordConnected <| by rw [range_coe_eq_Ioo]; exact ordConnected_Ioo -theorem openEmbedding_coe : OpenEmbedding ((↑) : ℝ → EReal) := +theorem isOpenEmbedding_coe : IsOpenEmbedding ((↑) : ℝ → EReal) := ⟨embedding_coe, by simp only [range_coe_eq_Ioo, isOpen_Ioo]⟩ +@[deprecated (since := "2024-10-18")] +alias openEmbedding_coe := isOpenEmbedding_coe + @[norm_cast] theorem tendsto_coe {α : Type*} {f : Filter α} {m : α → ℝ} {a : ℝ} : Tendsto (fun a => (m a : EReal)) f (𝓝 ↑a) ↔ Tendsto m f (𝓝 a) := @@ -66,11 +69,11 @@ theorem continuous_coe_iff {f : α → ℝ} : (Continuous fun a => (f a : EReal) embedding_coe.continuous_iff.symm theorem nhds_coe {r : ℝ} : 𝓝 (r : EReal) = (𝓝 r).map (↑) := - (openEmbedding_coe.map_nhds_eq r).symm + (isOpenEmbedding_coe.map_nhds_eq r).symm theorem nhds_coe_coe {r p : ℝ} : 𝓝 ((r : EReal), (p : EReal)) = (𝓝 (r, p)).map fun p : ℝ × ℝ => (↑p.1, ↑p.2) := - ((openEmbedding_coe.prodMap openEmbedding_coe).map_nhds_eq (r, p)).symm + ((isOpenEmbedding_coe.prodMap isOpenEmbedding_coe).map_nhds_eq (r, p)).symm theorem tendsto_toReal {a : EReal} (ha : a ≠ ⊤) (h'a : a ≠ ⊥) : Tendsto EReal.toReal (𝓝 a) (𝓝 a.toReal) := by diff --git a/Mathlib/Topology/Irreducible.lean b/Mathlib/Topology/Irreducible.lean index 17b268468df10..4855d9e0f6e57 100644 --- a/Mathlib/Topology/Irreducible.lean +++ b/Mathlib/Topology/Irreducible.lean @@ -297,7 +297,7 @@ theorem IsPreirreducible.interior (ht : IsPreirreducible t) : IsPreirreducible ( ht.open_subset isOpen_interior interior_subset theorem IsPreirreducible.preimage (ht : IsPreirreducible t) {f : Y → X} - (hf : OpenEmbedding f) : IsPreirreducible (f ⁻¹' t) := by + (hf : IsOpenEmbedding f) : IsPreirreducible (f ⁻¹' t) := by rintro U V hU hV ⟨x, hx, hx'⟩ ⟨y, hy, hy'⟩ obtain ⟨_, h₁, ⟨y, h₂, rfl⟩, ⟨y', h₃, h₄⟩⟩ := ht _ _ (hf.isOpenMap _ hU) (hf.isOpenMap _ hV) ⟨f x, hx, Set.mem_image_of_mem f hx'⟩ diff --git a/Mathlib/Topology/IsLocalHomeomorph.lean b/Mathlib/Topology/IsLocalHomeomorph.lean index 6351032d8967e..d75d32d051089 100644 --- a/Mathlib/Topology/IsLocalHomeomorph.lean +++ b/Mathlib/Topology/IsLocalHomeomorph.lean @@ -39,22 +39,26 @@ the source of some `e : PartialHomeomorph X Y` with `f = e`. -/ def IsLocalHomeomorphOn := ∀ x ∈ s, ∃ e : PartialHomeomorph X Y, x ∈ e.source ∧ f = e -theorem isLocalHomeomorphOn_iff_openEmbedding_restrict {f : X → Y} : - IsLocalHomeomorphOn f s ↔ ∀ x ∈ s, ∃ U ∈ 𝓝 x, OpenEmbedding (U.restrict f) := by +theorem isLocalHomeomorphOn_iff_isOpenEmbedding_restrict {f : X → Y} : + IsLocalHomeomorphOn f s ↔ ∀ x ∈ s, ∃ U ∈ 𝓝 x, IsOpenEmbedding (U.restrict f) := by refine ⟨fun h x hx ↦ ?_, fun h x hx ↦ ?_⟩ · obtain ⟨e, hxe, rfl⟩ := h x hx - exact ⟨e.source, e.open_source.mem_nhds hxe, e.openEmbedding_restrict⟩ + exact ⟨e.source, e.open_source.mem_nhds hxe, e.isOpenEmbedding_restrict⟩ · obtain ⟨U, hU, emb⟩ := h x hx - have : OpenEmbedding ((interior U).restrict f) := by + have : IsOpenEmbedding ((interior U).restrict f) := by refine emb.comp ⟨embedding_inclusion interior_subset, ?_⟩ rw [Set.range_inclusion]; exact isOpen_induced isOpen_interior - obtain ⟨cont, inj, openMap⟩ := openEmbedding_iff_continuous_injective_open.mp this + obtain ⟨cont, inj, openMap⟩ := isOpenEmbedding_iff_continuous_injective_open.mp this haveI : Nonempty X := ⟨x⟩ exact ⟨PartialHomeomorph.ofContinuousOpenRestrict (Set.injOn_iff_injective.mpr inj).toPartialEquiv (continuousOn_iff_continuous_restrict.mpr cont) openMap isOpen_interior, mem_interior_iff_mem_nhds.mpr hU, rfl⟩ +@[deprecated (since := "2024-10-18")] +alias isLocalHomeomorphOn_iff_openEmbedding_restrict := + isLocalHomeomorphOn_iff_isOpenEmbedding_restrict + namespace IsLocalHomeomorphOn /-- Proves that `f` satisfies `IsLocalHomeomorphOn f s`. The condition `h` is weaker than the @@ -141,14 +145,20 @@ theorem isLocalHomeomorph_iff_isLocalHomeomorphOn_univ : protected theorem IsLocalHomeomorph.isLocalHomeomorphOn (hf : IsLocalHomeomorph f) : IsLocalHomeomorphOn f s := fun x _ ↦ hf x -theorem isLocalHomeomorph_iff_openEmbedding_restrict {f : X → Y} : - IsLocalHomeomorph f ↔ ∀ x : X, ∃ U ∈ 𝓝 x, OpenEmbedding (U.restrict f) := by +theorem isLocalHomeomorph_iff_isOpenEmbedding_restrict {f : X → Y} : + IsLocalHomeomorph f ↔ ∀ x : X, ∃ U ∈ 𝓝 x, IsOpenEmbedding (U.restrict f) := by simp_rw [isLocalHomeomorph_iff_isLocalHomeomorphOn_univ, - isLocalHomeomorphOn_iff_openEmbedding_restrict, imp_iff_right (Set.mem_univ _)] + isLocalHomeomorphOn_iff_isOpenEmbedding_restrict, imp_iff_right (Set.mem_univ _)] + +@[deprecated (since := "2024-10-18")] +alias isLocalHomeomorph_iff_openEmbedding_restrict := isLocalHomeomorph_iff_isOpenEmbedding_restrict + +theorem IsOpenEmbedding.isLocalHomeomorph (hf : IsOpenEmbedding f) : IsLocalHomeomorph f := + isLocalHomeomorph_iff_isOpenEmbedding_restrict.mpr fun _ ↦ + ⟨_, Filter.univ_mem, hf.comp (Homeomorph.Set.univ X).isOpenEmbedding⟩ -theorem OpenEmbedding.isLocalHomeomorph (hf : OpenEmbedding f) : IsLocalHomeomorph f := - isLocalHomeomorph_iff_openEmbedding_restrict.mpr fun _ ↦ - ⟨_, Filter.univ_mem, hf.comp (Homeomorph.Set.univ X).openEmbedding⟩ +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.isLocalHomeomorph := IsOpenEmbedding.isLocalHomeomorph variable (f) @@ -194,15 +204,18 @@ protected theorem comp (hg : IsLocalHomeomorph g) (hf : IsLocalHomeomorph f) : (hg.isLocalHomeomorphOn.comp hf.isLocalHomeomorphOn (Set.univ.mapsTo_univ f)) /-- An injective local homeomorphism is an open embedding. -/ -theorem openEmbedding_of_injective (hf : IsLocalHomeomorph f) (hi : f.Injective) : - OpenEmbedding f := - openEmbedding_of_continuous_injective_open hf.continuous hi hf.isOpenMap +theorem isOpenEmbedding_of_injective (hf : IsLocalHomeomorph f) (hi : f.Injective) : + IsOpenEmbedding f := + isOpenEmbedding_of_continuous_injective_open hf.continuous hi hf.isOpenMap + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_of_injective := isOpenEmbedding_of_injective /-- A surjective embedding is a homeomorphism. -/ noncomputable def _root_.Embedding.toHomeomeomorph_of_surjective (hf : Embedding f) (hsurj : Function.Surjective f) : X ≃ₜ Y := Homeomorph.homeomorphOfContinuousOpen (Equiv.ofBijective f ⟨hf.inj, hsurj⟩) - hf.continuous (hf.toOpenEmbedding_of_surjective hsurj).isOpenMap + hf.continuous (hf.toIsOpenEmbedding_of_surjective hsurj).isOpenMap /-- A bijective local homeomorphism is a homeomorphism. -/ noncomputable def toHomeomorph_of_bijective (hf : IsLocalHomeomorph f) (hb : f.Bijective) : @@ -210,9 +223,12 @@ noncomputable def toHomeomorph_of_bijective (hf : IsLocalHomeomorph f) (hb : f.B Homeomorph.homeomorphOfContinuousOpen (Equiv.ofBijective f hb) hf.continuous hf.isOpenMap /-- Continuous local sections of a local homeomorphism are open embeddings. -/ -theorem openEmbedding_of_comp (hf : IsLocalHomeomorph g) (hgf : OpenEmbedding (g ∘ f)) - (cont : Continuous f) : OpenEmbedding f := - (hgf.isLocalHomeomorph.of_comp hf cont).openEmbedding_of_injective hgf.inj.of_comp +theorem isOpenEmbedding_of_comp (hf : IsLocalHomeomorph g) (hgf : IsOpenEmbedding (g ∘ f)) + (cont : Continuous f) : IsOpenEmbedding f := + (hgf.isLocalHomeomorph.of_comp hf cont).isOpenEmbedding_of_injective hgf.inj.of_comp + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_of_comp := isOpenEmbedding_of_comp open TopologicalSpace in /-- Ranges of continuous local sections of a local homeomorphism @@ -221,7 +237,7 @@ theorem isTopologicalBasis (hf : IsLocalHomeomorph f) : IsTopologicalBasis {U : Set X | ∃ V : Set Y, IsOpen V ∧ ∃ s : C(V,X), f ∘ s = (↑) ∧ Set.range s = U} := by refine isTopologicalBasis_of_isOpen_of_nhds ?_ fun x U hx hU ↦ ?_ · rintro _ ⟨U, hU, s, hs, rfl⟩ - refine (openEmbedding_of_comp hf (hs ▸ ⟨embedding_subtype_val, ?_⟩) s.continuous).isOpen_range + refine (isOpenEmbedding_of_comp hf (hs ▸ ⟨embedding_subtype_val, ?_⟩) s.continuous).isOpen_range rwa [Subtype.range_val] · obtain ⟨f, hxf, rfl⟩ := hf x refine ⟨f.source ∩ U, ⟨f.target ∩ f.symm ⁻¹' U, f.symm.isOpen_inter_preimage hU, diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 184b90a9c30e4..4e1a6f06f340a 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -12,7 +12,7 @@ import Mathlib.Topology.LocallyClosed We show that the following properties of continuous maps are local at the target : - `Inducing` - `Embedding` -- `OpenEmbedding` +- `IsOpenEmbedding` - `ClosedEmbedding` -/ @@ -40,12 +40,18 @@ theorem Set.restrictPreimage_embedding (s : Set β) (h : Embedding f) : alias Embedding.restrictPreimage := Set.restrictPreimage_embedding -theorem Set.restrictPreimage_openEmbedding (s : Set β) (h : OpenEmbedding f) : - OpenEmbedding (s.restrictPreimage f) := +theorem Set.restrictPreimage_isOpenEmbedding (s : Set β) (h : IsOpenEmbedding f) : + IsOpenEmbedding (s.restrictPreimage f) := ⟨h.1.restrictPreimage s, (s.range_restrictPreimage f).symm ▸ continuous_subtype_val.isOpen_preimage _ h.isOpen_range⟩ -alias OpenEmbedding.restrictPreimage := Set.restrictPreimage_openEmbedding +@[deprecated (since := "2024-10-18")] +alias Set.restrictPreimage_openEmbedding := Set.restrictPreimage_isOpenEmbedding + +alias IsOpenEmbedding.restrictPreimage := Set.restrictPreimage_isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.restrictPreimage := IsOpenEmbedding.restrictPreimage theorem Set.restrictPreimage_closedEmbedding (s : Set β) (h : ClosedEmbedding f) : ClosedEmbedding (s.restrictPreimage f) := @@ -96,7 +102,7 @@ theorem isOpen_iff_coe_preimage_of_iSup_eq_top (s : Set β) : -- Porting note: rewrote to avoid ´simp´ issues rw [isOpen_iff_inter_of_iSup_eq_top hU s] refine forall_congr' fun i => ?_ - rw [(U _).2.openEmbedding_subtype_val.open_iff_image_open] + rw [(U _).2.isOpenEmbedding_subtypeVal.open_iff_image_open] erw [Set.image_preimage_eq_inter_range] rw [Subtype.range_coe, Opens.carrier_eq_coe] @@ -110,7 +116,7 @@ theorem isLocallyClosed_iff_coe_preimage_of_iSup_eq_top (s : Set β) : rw [isOpen_iff_coe_preimage_of_iSup_eq_top hU] exact forall_congr' fun i ↦ by have : coborder ((↑) ⁻¹' s : Set (U i)) = Subtype.val ⁻¹' coborder s := by - exact (U i).isOpen.openEmbedding_subtype_val.coborder_preimage _ + exact (U i).isOpen.isOpenEmbedding_subtypeVal.coborder_preimage _ rw [this] theorem isOpenMap_iff_isOpenMap_of_iSup_eq_top : @@ -150,7 +156,7 @@ theorem inducing_iff_inducing_of_iSup_eq_top (h : Continuous f) : (show f x ∈ iSup U by rw [hU] trivial) - rw [← OpenEmbedding.map_nhds_eq (h.1 _ (U i).2).openEmbedding_subtype_val ⟨x, hi⟩, + rw [← IsOpenEmbedding.map_nhds_eq (h.1 _ (U i).2).isOpenEmbedding_subtypeVal ⟨x, hi⟩, (H i) ⟨x, hi⟩, Filter.subtype_coe_map_comap, Function.comp_apply, Subtype.coe_mk, inf_eq_left, Filter.le_principal_iff] exact Filter.preimage_mem_comap ((U i).2.mem_nhds hi) @@ -165,15 +171,19 @@ theorem embedding_iff_embedding_of_iSup_eq_top (h : Continuous f) : convert congr_arg SetLike.coe hU simp -theorem openEmbedding_iff_openEmbedding_of_iSup_eq_top (h : Continuous f) : - OpenEmbedding f ↔ ∀ i, OpenEmbedding ((U i).1.restrictPreimage f) := by - simp_rw [openEmbedding_iff] +theorem isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top (h : Continuous f) : + IsOpenEmbedding f ↔ ∀ i, IsOpenEmbedding ((U i).1.restrictPreimage f) := by + simp_rw [isOpenEmbedding_iff] rw [forall_and] apply and_congr · apply embedding_iff_embedding_of_iSup_eq_top <;> assumption · simp_rw [Set.range_restrictPreimage] apply isOpen_iff_coe_preimage_of_iSup_eq_top hU +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_openEmbedding_of_iSup_eq_top := + isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top + theorem closedEmbedding_iff_closedEmbedding_of_iSup_eq_top (h : Continuous f) : ClosedEmbedding f ↔ ∀ i, ClosedEmbedding ((U i).1.restrictPreimage f) := by simp_rw [closedEmbedding_iff] diff --git a/Mathlib/Topology/LocallyClosed.lean b/Mathlib/Topology/LocallyClosed.lean index fd260f0e11c03..6aa387864372e 100644 --- a/Mathlib/Topology/LocallyClosed.lean +++ b/Mathlib/Topology/LocallyClosed.lean @@ -91,10 +91,13 @@ lemma coborder_preimage (hf : IsOpenMap f) (hf' : Continuous f) (s : Set Y) : (hf.coborder_preimage_subset s).antisymm (hf'.preimage_coborder_subset s) protected -lemma OpenEmbedding.coborder_preimage (hf : OpenEmbedding f) (s : Set Y) : +lemma IsOpenEmbedding.coborder_preimage (hf : IsOpenEmbedding f) (s : Set Y) : coborder (f ⁻¹' s) = f ⁻¹' (coborder s) := coborder_preimage hf.isOpenMap hf.continuous s +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.coborder_preimage := IsOpenEmbedding.coborder_preimage + lemma isClosed_preimage_val_coborder : IsClosed (coborder s ↓∩ s) := by rw [isClosed_preimage_val, inter_eq_right.mpr subset_coborder, coborder_inter_closure] diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index de7a4ffef4499..c4aaea387693f 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -20,7 +20,7 @@ This file introduces the following properties of a map `f : X → Y` between top are identified by `f` are also inseparable in the topology on `X`. * `Embedding f` means `f` is inducing and also injective. Equivalently, `f` identifies `X` with a subspace of `Y`. -* `OpenEmbedding f` means `f` is an embedding with open image, so it identifies `X` with an +* `IsOpenEmbedding f` means `f` is an embedding with open image, so it identifies `X` with an open subspace of `Y`. Equivalently, `f` is an embedding and an open map. * `ClosedEmbedding f` similarly means `f` is an embedding with closed image, so it identifies `X` with a closed subspace of `Y`. Equivalently, `f` is an embedding and a closed map. @@ -470,99 +470,138 @@ theorem IsClosedMap.mapClusterPt_iff_lift'_closure end IsClosedMap -section OpenEmbedding +section IsOpenEmbedding variable [TopologicalSpace X] [TopologicalSpace Y] -theorem OpenEmbedding.isOpenMap (hf : OpenEmbedding f) : IsOpenMap f := +theorem IsOpenEmbedding.isOpenMap (hf : IsOpenEmbedding f) : IsOpenMap f := hf.toEmbedding.toInducing.isOpenMap hf.isOpen_range -theorem OpenEmbedding.map_nhds_eq (hf : OpenEmbedding f) (x : X) : +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.isOpenMap := IsOpenEmbedding.isOpenMap + +theorem IsOpenEmbedding.map_nhds_eq (hf : IsOpenEmbedding f) (x : X) : map f (𝓝 x) = 𝓝 (f x) := hf.toEmbedding.map_nhds_of_mem _ <| hf.isOpen_range.mem_nhds <| mem_range_self _ -theorem OpenEmbedding.open_iff_image_open (hf : OpenEmbedding f) {s : Set X} : +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.map_nhds_eq := IsOpenEmbedding.map_nhds_eq + +theorem IsOpenEmbedding.open_iff_image_open (hf : IsOpenEmbedding f) {s : Set X} : IsOpen s ↔ IsOpen (f '' s) := ⟨hf.isOpenMap s, fun h => by convert ← h.preimage hf.toEmbedding.continuous apply preimage_image_eq _ hf.inj⟩ -theorem OpenEmbedding.tendsto_nhds_iff [TopologicalSpace Z] {f : ι → Y} {l : Filter ι} {y : Y} - (hg : OpenEmbedding g) : Tendsto f l (𝓝 y) ↔ Tendsto (g ∘ f) l (𝓝 (g y)) := +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.open_iff_image_open := IsOpenEmbedding.open_iff_image_open + +theorem IsOpenEmbedding.tendsto_nhds_iff [TopologicalSpace Z] {f : ι → Y} {l : Filter ι} {y : Y} + (hg : IsOpenEmbedding g) : Tendsto f l (𝓝 y) ↔ Tendsto (g ∘ f) l (𝓝 (g y)) := hg.toEmbedding.tendsto_nhds_iff -theorem OpenEmbedding.tendsto_nhds_iff' (hf : OpenEmbedding f) {l : Filter Z} {x : X} : +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.tendsto_nhds_iff := IsOpenEmbedding.tendsto_nhds_iff + +theorem IsOpenEmbedding.tendsto_nhds_iff' (hf : IsOpenEmbedding f) {l : Filter Z} {x : X} : Tendsto (g ∘ f) (𝓝 x) l ↔ Tendsto g (𝓝 (f x)) l := by rw [Tendsto, ← map_map, hf.map_nhds_eq]; rfl -theorem OpenEmbedding.continuousAt_iff [TopologicalSpace Z] (hf : OpenEmbedding f) {x : X} : +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.tendsto_nhds_iff' := IsOpenEmbedding.tendsto_nhds_iff' + +theorem IsOpenEmbedding.continuousAt_iff [TopologicalSpace Z] (hf : IsOpenEmbedding f) {x : X} : ContinuousAt (g ∘ f) x ↔ ContinuousAt g (f x) := hf.tendsto_nhds_iff' -theorem OpenEmbedding.continuous (hf : OpenEmbedding f) : Continuous f := +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.continuousAt_iff := IsOpenEmbedding.continuousAt_iff + +theorem IsOpenEmbedding.continuous (hf : IsOpenEmbedding f) : Continuous f := hf.toEmbedding.continuous -theorem OpenEmbedding.open_iff_preimage_open (hf : OpenEmbedding f) {s : Set Y} +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.continuous := IsOpenEmbedding.continuous + +theorem IsOpenEmbedding.open_iff_preimage_open (hf : IsOpenEmbedding f) {s : Set Y} (hs : s ⊆ range f) : IsOpen s ↔ IsOpen (f ⁻¹' s) := by rw [hf.open_iff_image_open, image_preimage_eq_inter_range, inter_eq_self_of_subset_left hs] -theorem openEmbedding_of_embedding_open (h₁ : Embedding f) (h₂ : IsOpenMap f) : - OpenEmbedding f := +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.open_iff_preimage_open := IsOpenEmbedding.open_iff_preimage_open + +theorem isOpenEmbedding_of_embedding_open (h₁ : Embedding f) (h₂ : IsOpenMap f) : + IsOpenEmbedding f := ⟨h₁, h₂.isOpen_range⟩ -/-- A surjective embedding is an `OpenEmbedding`. -/ -theorem _root_.Embedding.toOpenEmbedding_of_surjective (hf : Embedding f) (hsurj : f.Surjective) : - OpenEmbedding f := +@[deprecated (since := "2024-10-18")] +alias openEmbedding_of_embedding_open := isOpenEmbedding_of_embedding_open + +/-- A surjective embedding is an `IsOpenEmbedding`. -/ +theorem _root_.Embedding.toIsOpenEmbedding_of_surjective (hf : Embedding f) (hsurj : f.Surjective) : + IsOpenEmbedding f := ⟨hf, hsurj.range_eq ▸ isOpen_univ⟩ -theorem openEmbedding_iff_embedding_open : - OpenEmbedding f ↔ Embedding f ∧ IsOpenMap f := - ⟨fun h => ⟨h.1, h.isOpenMap⟩, fun h => openEmbedding_of_embedding_open h.1 h.2⟩ +@[deprecated (since := "2024-10-18")] +alias _root_.Embedding.toOpenEmbedding_of_surjective := Embedding.toIsOpenEmbedding_of_surjective + +theorem isOpenEmbedding_iff_embedding_open : + IsOpenEmbedding f ↔ Embedding f ∧ IsOpenMap f := + ⟨fun h => ⟨h.1, h.isOpenMap⟩, fun h => isOpenEmbedding_of_embedding_open h.1 h.2⟩ -theorem openEmbedding_of_continuous_injective_open - (h₁ : Continuous f) (h₂ : Injective f) (h₃ : IsOpenMap f) : OpenEmbedding f := by - simp only [openEmbedding_iff_embedding_open, embedding_iff, inducing_iff_nhds, *, and_true] +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_embedding_open := isOpenEmbedding_iff_embedding_open + +theorem isOpenEmbedding_of_continuous_injective_open + (h₁ : Continuous f) (h₂ : Injective f) (h₃ : IsOpenMap f) : IsOpenEmbedding f := by + simp only [isOpenEmbedding_iff_embedding_open, embedding_iff, inducing_iff_nhds, *, and_true] exact fun x => le_antisymm (h₁.tendsto _).le_comap (@comap_map _ _ (𝓝 x) _ h₂ ▸ comap_mono (h₃.nhds_le _)) -theorem openEmbedding_iff_continuous_injective_open : - OpenEmbedding f ↔ Continuous f ∧ Injective f ∧ IsOpenMap f := +theorem isOpenEmbedding_iff_continuous_injective_open : + IsOpenEmbedding f ↔ Continuous f ∧ Injective f ∧ IsOpenMap f := ⟨fun h => ⟨h.continuous, h.inj, h.isOpenMap⟩, fun h => - openEmbedding_of_continuous_injective_open h.1 h.2.1 h.2.2⟩ + isOpenEmbedding_of_continuous_injective_open h.1 h.2.1 h.2.2⟩ + +@[deprecated (since := "2024-10-18")] +alias openEmbedding_iff_continuous_injective_open := isOpenEmbedding_iff_continuous_injective_open -theorem openEmbedding_id : OpenEmbedding (@id X) := +theorem isOpenEmbedding_id : IsOpenEmbedding (@id X) := ⟨embedding_id, IsOpenMap.id.isOpen_range⟩ -namespace OpenEmbedding +@[deprecated (since := "2024-10-18")] +alias openEmbedding_id := isOpenEmbedding_id + +namespace IsOpenEmbedding variable [TopologicalSpace Z] -protected theorem comp (hg : OpenEmbedding g) - (hf : OpenEmbedding f) : OpenEmbedding (g ∘ f) := +protected theorem comp (hg : IsOpenEmbedding g) + (hf : IsOpenEmbedding f) : IsOpenEmbedding (g ∘ f) := ⟨hg.1.comp hf.1, (hg.isOpenMap.comp hf.isOpenMap).isOpen_range⟩ -theorem isOpenMap_iff (hg : OpenEmbedding g) : +theorem isOpenMap_iff (hg : IsOpenEmbedding g) : IsOpenMap f ↔ IsOpenMap (g ∘ f) := by simp_rw [isOpenMap_iff_nhds_le, ← map_map, comp, ← hg.map_nhds_eq, Filter.map_le_map_iff hg.inj] -theorem of_comp_iff (f : X → Y) (hg : OpenEmbedding g) : - OpenEmbedding (g ∘ f) ↔ OpenEmbedding f := by - simp only [openEmbedding_iff_continuous_injective_open, ← hg.isOpenMap_iff, ← +theorem of_comp_iff (f : X → Y) (hg : IsOpenEmbedding g) : + IsOpenEmbedding (g ∘ f) ↔ IsOpenEmbedding f := by + simp only [isOpenEmbedding_iff_continuous_injective_open, ← hg.isOpenMap_iff, ← hg.1.continuous_iff, hg.inj.of_comp_iff] -theorem of_comp (f : X → Y) (hg : OpenEmbedding g) - (h : OpenEmbedding (g ∘ f)) : OpenEmbedding f := - (OpenEmbedding.of_comp_iff f hg).1 h +theorem of_comp (f : X → Y) (hg : IsOpenEmbedding g) + (h : IsOpenEmbedding (g ∘ f)) : IsOpenEmbedding f := + (IsOpenEmbedding.of_comp_iff f hg).1 h -theorem of_isEmpty [IsEmpty X] (f : X → Y) : OpenEmbedding f := - openEmbedding_of_embedding_open (.of_subsingleton f) (IsOpenMap.of_isEmpty f) +theorem of_isEmpty [IsEmpty X] (f : X → Y) : IsOpenEmbedding f := + isOpenEmbedding_of_embedding_open (.of_subsingleton f) (IsOpenMap.of_isEmpty f) -theorem image_mem_nhds {f : X → Y} (hf : OpenEmbedding f) {s : Set X} {x : X} : +theorem image_mem_nhds {f : X → Y} (hf : IsOpenEmbedding f) {s : Set X} {x : X} : f '' s ∈ 𝓝 (f x) ↔ s ∈ 𝓝 x := by rw [← hf.map_nhds_eq, mem_map, preimage_image_eq _ hf.inj] -end OpenEmbedding +end IsOpenEmbedding -end OpenEmbedding +end IsOpenEmbedding section ClosedEmbedding diff --git a/Mathlib/Topology/PartialHomeomorph.lean b/Mathlib/Topology/PartialHomeomorph.lean index 3cecb43406b4e..4d2b050b3eb3c 100644 --- a/Mathlib/Topology/PartialHomeomorph.lean +++ b/Mathlib/Topology/PartialHomeomorph.lean @@ -1092,18 +1092,24 @@ def toHomeomorphOfSourceEqUnivTargetEqUniv (h : e.source = (univ : Set X)) (h' : continuous_invFun := by simpa only [continuous_iff_continuousOn_univ, h'] using e.continuousOn_symm -theorem openEmbedding_restrict : OpenEmbedding (e.source.restrict e) := by - refine openEmbedding_of_continuous_injective_open (e.continuousOn.comp_continuous +theorem isOpenEmbedding_restrict : IsOpenEmbedding (e.source.restrict e) := by + refine isOpenEmbedding_of_continuous_injective_open (e.continuousOn.comp_continuous continuous_subtype_val Subtype.prop) e.injOn.injective fun V hV ↦ ?_ rw [Set.restrict_eq, Set.image_comp] exact e.isOpen_image_of_subset_source (e.open_source.isOpenMap_subtype_val V hV) fun _ ⟨x, _, h⟩ ↦ h ▸ x.2 +@[deprecated (since := "2024-10-18")] +alias openEmbedding_restrict := isOpenEmbedding_restrict + /-- A partial homeomorphism whose source is all of `X` defines an open embedding of `X` into `Y`. -The converse is also true; see `OpenEmbedding.toPartialHomeomorph`. -/ -theorem to_openEmbedding (h : e.source = Set.univ) : OpenEmbedding e := - e.openEmbedding_restrict.comp - ((Homeomorph.setCongr h).trans <| Homeomorph.Set.univ X).symm.openEmbedding +The converse is also true; see `IsOpenEmbedding.toPartialHomeomorph`. -/ +theorem to_isOpenEmbedding (h : e.source = Set.univ) : IsOpenEmbedding e := + e.isOpenEmbedding_restrict.comp + ((Homeomorph.setCongr h).trans <| Homeomorph.Set.univ X).symm.isOpenEmbedding + +@[deprecated (since := "2024-10-18")] +alias to_openEmbedding := to_isOpenEmbedding end PartialHomeomorph @@ -1156,12 +1162,13 @@ theorem trans_transPartialHomeomorph (e : X ≃ₜ Y) (e' : Y ≃ₜ Z) (f'' : P end Homeomorph -namespace OpenEmbedding +namespace IsOpenEmbedding -variable (f : X → Y) (h : OpenEmbedding f) +variable (f : X → Y) (h : IsOpenEmbedding f) /-- An open embedding of `X` into `Y`, with `X` nonempty, defines a partial homeomorphism -whose source is all of `X`. The converse is also true; see `PartialHomeomorph.to_openEmbedding`. -/ +whose source is all of `X`. The converse is also true; see +`PartialHomeomorph.to_isOpenEmbedding`. -/ @[simps! (config := mfld_cfg) apply source target] noncomputable def toPartialHomeomorph [Nonempty X] : PartialHomeomorph X Y := PartialHomeomorph.ofContinuousOpen (h.toEmbedding.inj.injOn.toPartialEquiv f univ) @@ -1178,7 +1185,7 @@ lemma toPartialHomeomorph_right_inv {x : Y} (hx : x ∈ Set.range f) : rw [← congr_fun (h.toPartialHomeomorph_apply f), PartialHomeomorph.right_inv] rwa [toPartialHomeomorph_target] -end OpenEmbedding +end IsOpenEmbedding /-! inclusion of an open set in a topological space -/ namespace TopologicalSpace.Opens @@ -1191,7 +1198,7 @@ variable (s : Opens X) (hs : Nonempty s) /-- The inclusion of an open subset `s` of a space `X` into `X` is a partial homeomorphism from the subtype `s` to `X`. -/ noncomputable def partialHomeomorphSubtypeCoe : PartialHomeomorph s X := - OpenEmbedding.toPartialHomeomorph _ s.2.openEmbedding_subtype_val + IsOpenEmbedding.toPartialHomeomorph _ s.2.isOpenEmbedding_subtypeVal @[simp, mfld_simps] theorem partialHomeomorphSubtypeCoe_coe : (s.partialHomeomorphSubtypeCoe hs : s → X) = (↑) := diff --git a/Mathlib/Topology/QuasiSeparated.lean b/Mathlib/Topology/QuasiSeparated.lean index 869b89e96ebba..73240b9ea2c7e 100644 --- a/Mathlib/Topology/QuasiSeparated.lean +++ b/Mathlib/Topology/QuasiSeparated.lean @@ -22,7 +22,7 @@ open subsets, but their intersection `(0, 1]` is not. of any pairs of compact open subsets of `s` are still compact. - `QuasiSeparatedSpace`: A topological space is quasi-separated if the intersections of any pairs of compact open subsets are still compact. -- `QuasiSeparatedSpace.of_openEmbedding`: If `f : α → β` is an open embedding, and `β` is +- `QuasiSeparatedSpace.of_isOpenEmbedding`: If `f : α → β` is an open embedding, and `β` is a quasi-separated space, then so is `α`. -/ @@ -79,7 +79,7 @@ theorem IsQuasiSeparated.image_of_embedding {s : Set α} (H : IsQuasiSeparated s rw [Set.image_preimage_eq_inter_range, Set.inter_eq_left] exact hV.trans (Set.image_subset_range _ _) -theorem OpenEmbedding.isQuasiSeparated_iff (h : OpenEmbedding f) {s : Set α} : +theorem IsOpenEmbedding.isQuasiSeparated_iff (h : IsOpenEmbedding f) {s : Set α} : IsQuasiSeparated s ↔ IsQuasiSeparated (f '' s) := by refine ⟨fun hs => hs.image_of_embedding h.toEmbedding, ?_⟩ intro H U V hU hU' hU'' hV hV' hV'' @@ -88,10 +88,13 @@ theorem OpenEmbedding.isQuasiSeparated_iff (h : OpenEmbedding f) {s : Set α} : H (f '' U) (f '' V) (Set.image_subset _ hU) (h.isOpenMap _ hU') (hU''.image h.continuous) (Set.image_subset _ hV) (h.isOpenMap _ hV') (hV''.image h.continuous) +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.isQuasiSeparated_iff := IsOpenEmbedding.isQuasiSeparated_iff + theorem isQuasiSeparated_iff_quasiSeparatedSpace (s : Set α) (hs : IsOpen s) : IsQuasiSeparated s ↔ QuasiSeparatedSpace s := by rw [← isQuasiSeparated_univ_iff] - convert (hs.openEmbedding_subtype_val.isQuasiSeparated_iff (s := Set.univ)).symm + convert (hs.isOpenEmbedding_subtypeVal.isQuasiSeparated_iff (s := Set.univ)).symm simp theorem IsQuasiSeparated.of_subset {s t : Set α} (ht : IsQuasiSeparated t) (h : s ⊆ t) : @@ -110,7 +113,10 @@ theorem IsQuasiSeparated.of_quasiSeparatedSpace (s : Set α) [QuasiSeparatedSpac IsQuasiSeparated s := isQuasiSeparated_univ.of_subset (Set.subset_univ _) -theorem QuasiSeparatedSpace.of_openEmbedding (h : OpenEmbedding f) [QuasiSeparatedSpace β] : +theorem QuasiSeparatedSpace.of_isOpenEmbedding (h : IsOpenEmbedding f) [QuasiSeparatedSpace β] : QuasiSeparatedSpace α := isQuasiSeparated_univ_iff.mp (h.isQuasiSeparated_iff.mpr <| IsQuasiSeparated.of_quasiSeparatedSpace _) + +@[deprecated (since := "2024-10-18")] +alias QuasiSeparatedSpace.of_openEmbedding := QuasiSeparatedSpace.of_isOpenEmbedding diff --git a/Mathlib/Topology/SeparatedMap.lean b/Mathlib/Topology/SeparatedMap.lean index 471b35d8086e1..ccb0ea0f688df 100644 --- a/Mathlib/Topology/SeparatedMap.lean +++ b/Mathlib/Topology/SeparatedMap.lean @@ -141,15 +141,18 @@ theorem isLocallyInjective_iff_isOpen_diagonal {f : X → Y} : exact ⟨t₁ ∩ t₂, Filter.inter_mem h₁ h₂, fun x₁ h₁ x₂ h₂ he ↦ @t_sub ⟨(x₁, x₂), he⟩ (prod_sub ⟨h₁.1, h₂.2⟩)⟩ -theorem IsLocallyInjective_iff_openEmbedding {f : X → Y} : - IsLocallyInjective f ↔ OpenEmbedding (toPullbackDiag f) := by +theorem IsLocallyInjective_iff_isOpenEmbedding {f : X → Y} : + IsLocallyInjective f ↔ IsOpenEmbedding (toPullbackDiag f) := by rw [isLocallyInjective_iff_isOpen_diagonal, ← range_toPullbackDiag] exact ⟨fun h ↦ ⟨embedding_toPullbackDiag f, h⟩, fun h ↦ h.isOpen_range⟩ +@[deprecated (since := "2024-10-18")] +alias IsLocallyInjective_iff_openEmbedding := IsLocallyInjective_iff_isOpenEmbedding + theorem isLocallyInjective_iff_isOpenMap {f : X → Y} : IsLocallyInjective f ↔ IsOpenMap (toPullbackDiag f) := - IsLocallyInjective_iff_openEmbedding.trans - ⟨OpenEmbedding.isOpenMap, openEmbedding_of_continuous_injective_open + IsLocallyInjective_iff_isOpenEmbedding.trans + ⟨IsOpenEmbedding.isOpenMap, isOpenEmbedding_of_continuous_injective_open (embedding_toPullbackDiag f).continuous (injective_toPullbackDiag f)⟩ theorem discreteTopology_iff_locallyInjective (y : Y) : diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 792e2e67b2544..1ba06927812eb 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -1544,7 +1544,7 @@ Hausdorff spaces: `f x ≠ f y`. We use this lemma to prove that topological spaces defined using `induced` are Hausdorff spaces. -* `separated_by_openEmbedding` says that for an open embedding `f : X → Y` of a Hausdorff space +* `separated_by_isOpenEmbedding` says that for an open embedding `f : X → Y` of a Hausdorff space `X`, the images of two distinct points `x y : X`, `x ≠ y` can be separated by open neighborhoods. We use this lemma to prove that topological spaces defined using `coinduced` are Hausdorff spaces. -/ @@ -1560,13 +1560,16 @@ theorem separated_by_continuous [TopologicalSpace Y] [T2Space Y] let ⟨u, v, uo, vo, xu, yv, uv⟩ := t2_separation h ⟨f ⁻¹' u, f ⁻¹' v, uo.preimage hf, vo.preimage hf, xu, yv, uv.preimage _⟩ -theorem separated_by_openEmbedding [TopologicalSpace Y] [T2Space X] - {f : X → Y} (hf : OpenEmbedding f) {x y : X} (h : x ≠ y) : +theorem separated_by_isOpenEmbedding [TopologicalSpace Y] [T2Space X] + {f : X → Y} (hf : IsOpenEmbedding f) {x y : X} (h : x ≠ y) : ∃ u v : Set Y, IsOpen u ∧ IsOpen v ∧ f x ∈ u ∧ f y ∈ v ∧ Disjoint u v := let ⟨u, v, uo, vo, xu, yv, uv⟩ := t2_separation h ⟨f '' u, f '' v, hf.isOpenMap _ uo, hf.isOpenMap _ vo, mem_image_of_mem _ xu, mem_image_of_mem _ yv, disjoint_image_of_injective hf.inj uv⟩ +@[deprecated (since := "2024-10-18")] +alias separated_by_openEmbedding := separated_by_isOpenEmbedding + instance {p : X → Prop} [T2Space X] : T2Space (Subtype p) := inferInstance instance Prod.t2Space [T2Space X] [TopologicalSpace Y] [T2Space Y] : T2Space (X × Y) := @@ -1591,10 +1594,10 @@ instance [T2Space X] [TopologicalSpace Y] [T2Space Y] : T2Space (X ⊕ Y) := by constructor rintro (x | x) (y | y) h - · exact separated_by_openEmbedding openEmbedding_inl <| ne_of_apply_ne _ h + · exact separated_by_isOpenEmbedding isOpenEmbedding_inl <| ne_of_apply_ne _ h · exact separated_by_continuous continuous_isLeft <| by simp · exact separated_by_continuous continuous_isLeft <| by simp - · exact separated_by_openEmbedding openEmbedding_inr <| ne_of_apply_ne _ h + · exact separated_by_isOpenEmbedding isOpenEmbedding_inr <| ne_of_apply_ne _ h instance Pi.t2Space {Y : X → Type v} [∀ a, TopologicalSpace (Y a)] [∀ a, T2Space (Y a)] : T2Space (∀ a, Y a) := @@ -1606,7 +1609,7 @@ instance Sigma.t2Space {ι} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [ rintro ⟨i, x⟩ ⟨j, y⟩ neq rcases eq_or_ne i j with (rfl | h) · replace neq : x ≠ y := ne_of_apply_ne _ neq - exact separated_by_openEmbedding openEmbedding_sigmaMk neq + exact separated_by_isOpenEmbedding isOpenEmbedding_sigmaMk neq · let _ := (⊥ : TopologicalSpace ι); have : DiscreteTopology ι := ⟨rfl⟩ exact separated_by_continuous (continuous_def.2 fun u _ => isOpen_sigma_fst_preimage u) h @@ -2597,7 +2600,7 @@ theorem loc_compact_Haus_tot_disc_of_zero_dim [TotallyDisconnectedSpace H] : let v : Set u := ((↑) : u → s) ⁻¹' V have : ((↑) : u → H) = ((↑) : s → H) ∘ ((↑) : u → s) := rfl have f0 : Embedding ((↑) : u → H) := embedding_subtype_val.comp embedding_subtype_val - have f1 : OpenEmbedding ((↑) : u → H) := by + have f1 : IsOpenEmbedding ((↑) : u → H) := by refine ⟨f0, ?_⟩ · have : Set.range ((↑) : u → H) = interior s := by rw [this, Set.range_comp, Subtype.range_coe, Subtype.image_preimage_coe] diff --git a/Mathlib/Topology/Sets/Opens.lean b/Mathlib/Topology/Sets/Opens.lean index bd142710c0990..1252c062af227 100644 --- a/Mathlib/Topology/Sets/Opens.lean +++ b/Mathlib/Topology/Sets/Opens.lean @@ -243,16 +243,22 @@ def frameMinimalAxioms : Frame.MinimalAxioms (Opens α) where instance instFrame : Frame (Opens α) := .ofMinimalAxioms frameMinimalAxioms -theorem openEmbedding' (U : Opens α) : OpenEmbedding (Subtype.val : U → α) := - U.isOpen.openEmbedding_subtype_val +theorem isOpenEmbedding' (U : Opens α) : IsOpenEmbedding (Subtype.val : U → α) := + U.isOpen.isOpenEmbedding_subtypeVal -theorem openEmbedding_of_le {U V : Opens α} (i : U ≤ V) : - OpenEmbedding (Set.inclusion <| SetLike.coe_subset_coe.2 i) := +@[deprecated (since := "2024-10-18")] +alias openEmbedding' := isOpenEmbedding' + +theorem isOpenEmbedding_of_le {U V : Opens α} (i : U ≤ V) : + IsOpenEmbedding (Set.inclusion <| SetLike.coe_subset_coe.2 i) := { toEmbedding := embedding_inclusion i isOpen_range := by rw [Set.range_inclusion i] exact U.isOpen.preimage continuous_subtype_val } +@[deprecated (since := "2024-10-18")] +alias openEmbedding_of_le := isOpenEmbedding_of_le + theorem not_nonempty_iff_eq_bot (U : Opens α) : ¬Set.Nonempty (U : Set α) ↔ U = ⊥ := by rw [← coe_inj, coe_bot, ← Set.not_nonempty_iff_eq_empty] diff --git a/Mathlib/Topology/Sheaves/LocalPredicate.lean b/Mathlib/Topology/Sheaves/LocalPredicate.lean index 8d32b5f87a342..bf6e5eb7f9b01 100644 --- a/Mathlib/Topology/Sheaves/LocalPredicate.lean +++ b/Mathlib/Topology/Sheaves/LocalPredicate.lean @@ -73,7 +73,7 @@ variable (X) @[simps!] def continuousPrelocal (T : TopCat.{v}) : PrelocalPredicate fun _ : X => T where pred {_} f := Continuous f - res {_ _} i _ h := Continuous.comp h (Opens.openEmbedding_of_le i.le).continuous + res {_ _} i _ h := Continuous.comp h (Opens.isOpenEmbedding_of_le i.le).continuous /-- Satisfying the inhabited linter. -/ instance inhabitedPrelocalPredicate (T : TopCat.{v}) : @@ -114,7 +114,7 @@ def continuousLocal (T : TopCat.{v}) : LocalPredicate fun _ : X => T := dsimp at w rw [continuous_iff_continuousAt] at w specialize w ⟨x, m⟩ - simpa using (Opens.openEmbedding_of_le i.le).continuousAt_iff.1 w } + simpa using (Opens.isOpenEmbedding_of_le i.le).continuousAt_iff.1 w } /-- Satisfying the inhabited linter. -/ instance inhabitedLocalPredicate (T : TopCat.{v}) : Inhabited (LocalPredicate fun _ : X => T) := diff --git a/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean b/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean index 07a4029efeb34..88541de424f6b 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean @@ -136,14 +136,14 @@ theorem coverDense_inducedFunctor {B : ι → Opens X} (h : Opens.IsBasis (Set.r end TopCat.Opens -section OpenEmbedding +section IsOpenEmbedding open TopCat.Presheaf Opposite variable {C : Type u} [Category.{v} C] variable {X Y : TopCat.{w}} {f : X ⟶ Y} {F : Y.Presheaf C} -theorem OpenEmbedding.compatiblePreserving (hf : OpenEmbedding f) : +theorem IsOpenEmbedding.compatiblePreserving (hf : IsOpenEmbedding f) : CompatiblePreserving (Opens.grothendieckTopology Y) hf.isOpenMap.functor := by haveI : Mono f := (TopCat.mono_iff_injective f).mpr hf.inj apply compatiblePreservingOfDownwardsClosed @@ -152,6 +152,9 @@ theorem OpenEmbedding.compatiblePreserving (hf : OpenEmbedding f) : obtain ⟨_, _, rfl⟩ := i.le h exact ⟨_, rfl⟩ +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.compatiblePreserving := IsOpenEmbedding.compatiblePreserving + theorem IsOpenMap.coverPreserving (hf : IsOpenMap f) : CoverPreserving (Opens.grothendieckTopology X) (Opens.grothendieckTopology Y) hf.functor := by constructor @@ -160,18 +163,24 @@ theorem IsOpenMap.coverPreserving (hf : IsOpenMap f) : exact ⟨_, hf.functor.map i, ⟨_, i, 𝟙 _, hV, rfl⟩, Set.mem_image_of_mem f hxV⟩ -lemma OpenEmbedding.functor_isContinuous (h : OpenEmbedding f) : +lemma IsOpenEmbedding.functor_isContinuous (h : IsOpenEmbedding f) : h.isOpenMap.functor.IsContinuous (Opens.grothendieckTopology X) (Opens.grothendieckTopology Y) := by apply Functor.isContinuous_of_coverPreserving · exact h.compatiblePreserving · exact h.isOpenMap.coverPreserving -theorem TopCat.Presheaf.isSheaf_of_openEmbedding (h : OpenEmbedding f) (hF : F.IsSheaf) : +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.functor_isContinuous := IsOpenEmbedding.functor_isContinuous + +theorem TopCat.Presheaf.isSheaf_of_isOpenEmbedding (h : IsOpenEmbedding f) (hF : F.IsSheaf) : IsSheaf (h.isOpenMap.functor.op ⋙ F) := by have := h.functor_isContinuous exact Functor.op_comp_isSheaf _ _ _ ⟨_, hF⟩ +@[deprecated (since := "2024-10-18")] +alias TopCat.Presheaf.isSheaf_of_openEmbedding := TopCat.Presheaf.isSheaf_of_isOpenEmbedding + variable (f) instance : RepresentablyFlat (Opens.map f) := by @@ -203,7 +212,7 @@ instance : (Opens.map f).IsContinuous (Opens.grothendieckTopology Y) · exact compatiblePreserving_opens_map f · exact coverPreserving_opens_map f -end OpenEmbedding +end IsOpenEmbedding namespace TopCat.Sheaf diff --git a/Mathlib/Topology/Sheaves/Stalks.lean b/Mathlib/Topology/Sheaves/Stalks.lean index 74df85eb39c03..125a7e70f0fc4 100644 --- a/Mathlib/Topology/Sheaves/Stalks.lean +++ b/Mathlib/Topology/Sheaves/Stalks.lean @@ -208,8 +208,8 @@ theorem comp (ℱ : X.Presheaf C) (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : ext simp [germ, stalkPushforward] -theorem stalkPushforward_iso_of_openEmbedding {f : X ⟶ Y} (hf : OpenEmbedding f) (F : X.Presheaf C) - (x : X) : IsIso (F.stalkPushforward _ f x) := by +theorem stalkPushforward_iso_of_isOpenEmbedding {f : X ⟶ Y} (hf : IsOpenEmbedding f) + (F : X.Presheaf C) (x : X) : IsIso (F.stalkPushforward _ f x) := by haveI := Functor.initial_of_adjunction (hf.isOpenMap.adjunctionNhds x) convert ((Functor.Final.colimitIso (hf.isOpenMap.functorNhds x).op @@ -234,6 +234,9 @@ theorem stalkPushforward_iso_of_openEmbedding {f : X ⟶ Y} (hf : OpenEmbedding refine ((homOfLE ?_).op : op (unop U) ⟶ _) exact Set.image_preimage_subset _ _ +@[deprecated (since := "2024-10-18")] +alias stalkPushforward_iso_of_openEmbedding := stalkPushforward_iso_of_isOpenEmbedding + end stalkPushforward section stalkPullback diff --git a/Mathlib/Topology/Sober.lean b/Mathlib/Topology/Sober.lean index 82c12d98222d6..06f9286a3e5d3 100644 --- a/Mathlib/Topology/Sober.lean +++ b/Mathlib/Topology/Sober.lean @@ -182,7 +182,7 @@ theorem ClosedEmbedding.quasiSober {f : α → β} (hf : ClosedEmbedding f) [Qua apply image_injective.mpr hf.inj rw [← hx.def, ← hf.closure_image_eq, image_singleton] -theorem OpenEmbedding.quasiSober {f : α → β} (hf : OpenEmbedding f) [QuasiSober β] : +theorem IsOpenEmbedding.quasiSober {f : α → β} (hf : IsOpenEmbedding f) [QuasiSober β] : QuasiSober α where sober hS hS' := by have hS'' := hS.image f hf.continuous.continuousOn @@ -205,6 +205,9 @@ theorem OpenEmbedding.quasiSober {f : α → β} (hf : OpenEmbedding f) [QuasiSo exact fun hy => ⟨fun h => hT.closure_eq ▸ closure_mono inter_subset_left h, fun h => subset_closure ⟨h, hy⟩⟩ +@[deprecated (since := "2024-10-18")] +alias OpenEmbedding.quasiSober := IsOpenEmbedding.quasiSober + /-- A space is quasi sober if it can be covered by open quasi sober subsets. -/ theorem quasiSober_of_open_cover (S : Set (Set α)) (hS : ∀ s : S, IsOpen (s : Set α)) [hS' : ∀ s : S, QuasiSober s] (hS'' : ⋃₀ S = ⊤) : QuasiSober α := by @@ -216,7 +219,7 @@ theorem quasiSober_of_open_cover (S : Set (Set α)) (hS : ∀ s : S, IsOpen (s : trivial haveI : QuasiSober U := hS' ⟨U, hU⟩ have H : IsPreirreducible ((↑) ⁻¹' t : Set U) := - h.2.preimage (hS ⟨U, hU⟩).openEmbedding_subtype_val + h.2.preimage (hS ⟨U, hU⟩).isOpenEmbedding_subtypeVal replace H : IsIrreducible ((↑) ⁻¹' t : Set U) := ⟨⟨⟨x, hU'⟩, by simpa using hx⟩, H⟩ use H.genericPoint have := continuous_subtype_val.closure_preimage_subset _ H.isGenericPoint_genericPoint_closure.mem diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index c0f19ee3e63a7..11a71da0311a5 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -257,7 +257,7 @@ theorem continuous_extend_one [TopologicalSpace β] {U : Set α'} (hU : IsOpen U continuous_of_mulTSupport fun x h ↦ by rw [show x = ↑(⟨x, Subtype.coe_image_subset _ _ (supp.mulTSupport_extend_one_subset continuous_subtype_val h)⟩ : U) by rfl, - ← (hU.openEmbedding_subtype_val).continuousAt_iff, extend_comp Subtype.val_injective] + ← (hU.isOpenEmbedding_subtypeVal).continuousAt_iff, extend_comp Subtype.val_injective] exact cont.continuousAt /-- If `f` has compact multiplicative support, then `f` tends to 1 at infinity. -/ diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index 671479377b629..e1817a2092239 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -3521,7 +3521,7 @@ ONote.NF.below_of_lt' ONote.nf_repr_split' ONote.NF.snd' ONote.split_eq_scale_split' -OpenEmbedding.tendsto_nhds_iff' +IsOpenEmbedding.tendsto_nhds_iff' openSegment_eq_image' openSegment_eq_Ioo' Option.bind_congr' @@ -4639,8 +4639,8 @@ toIocMod_sub_zsmul' toIocMod_zsmul_add' toIxxMod_total' TopCat.GlueData.preimage_image_eq_image' -TopCat.openEmbedding_iff_comp_isIso' -TopCat.openEmbedding_iff_isIso_comp' +TopCat.isOpenEmbedding_iff_comp_isIso' +TopCat.isOpenEmbedding_iff_isIso_comp' TopCat.Presheaf.germ_stalkSpecializes' TopCat.Presheaf.pushforward_eq' TopCat.Presheaf.pushforward_map_app' @@ -4650,7 +4650,7 @@ TopologicalSpace.Opens.coe_inclusion' TopologicalSpace.Opens.map_comp_obj' TopologicalSpace.Opens.map_functor_eq' TopologicalSpace.Opens.map_id_obj' -TopologicalSpace.Opens.openEmbedding' +TopologicalSpace.Opens.isOpenEmbedding' TopologicalSpace.Opens.set_range_forget_map_inclusion' TopologicalSpace.Opens.set_range_inclusion' TopologicalSpace.SecondCountableTopology.mk' From b9efa9d8b132e603ec517380352c5b5a29a25e88 Mon Sep 17 00:00:00 2001 From: peabrainiac Date: Fri, 18 Oct 2024 21:51:41 +0000 Subject: [PATCH 087/131] feat(Topology/Connected/PathConnected): some lemmas for `pathComponentIn` (#16983) Adds some basic lemmas about `pathComponentIn`, analogous to existing lemmas about `pathComponent` and `connectedComponentIn`. Co-authored-by: peabrainiac <43812953+peabrainiac@users.noreply.github.com> --- Mathlib/Topology/Connected/PathConnected.lean | 33 +++++++++++++++++++ 1 file changed, 33 insertions(+) diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 3a8377f938164..b9a49ba4fcd65 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -848,6 +848,24 @@ theorem Joined.mem_pathComponent (hyz : Joined y z) (hxy : y ∈ pathComponent x z ∈ pathComponent x := hxy.trans hyz +theorem mem_pathComponentIn_self (h : x ∈ F) : x ∈ pathComponentIn x F := + JoinedIn.refl h + +theorem pathComponentIn_subset : pathComponentIn x F ⊆ F := + fun _ hy ↦ hy.target_mem + +theorem pathComponentIn_nonempty_iff : (pathComponentIn x F).Nonempty ↔ x ∈ F := + ⟨fun ⟨_, ⟨γ, hγ⟩⟩ ↦ γ.source ▸ hγ 0, fun hx ↦ ⟨x, mem_pathComponentIn_self hx⟩⟩ + +theorem pathComponentIn_congr (h : x ∈ pathComponentIn y F) : + pathComponentIn x F = pathComponentIn y F := by + ext; exact ⟨h.trans, h.symm.trans⟩ + +@[gcongr] +theorem pathComponentIn_mono {G : Set X} (h : F ⊆ G) : + pathComponentIn x F ⊆ pathComponentIn x G := + fun _ ⟨γ, hγ⟩ ↦ ⟨γ, fun t ↦ h (hγ t)⟩ + /-! ### Path connected sets -/ @@ -913,11 +931,26 @@ theorem IsPathConnected.mem_pathComponent (h : IsPathConnected F) (x_in : x ∈ theorem IsPathConnected.subset_pathComponent (h : IsPathConnected F) (x_in : x ∈ F) : F ⊆ pathComponent x := fun _y y_in => h.mem_pathComponent x_in y_in +theorem IsPathConnected.subset_pathComponentIn {s : Set X} (hs : IsPathConnected s) + (hxs : x ∈ s) (hsF : s ⊆ F) : s ⊆ pathComponentIn x F := + fun y hys ↦ (hs.joinedIn x hxs y hys).mono hsF + theorem isPathConnected_singleton (x : X) : IsPathConnected ({x} : Set X) := by refine ⟨x, rfl, ?_⟩ rintro y rfl exact JoinedIn.refl rfl +theorem isPathConnected_pathComponentIn (h : x ∈ F) : IsPathConnected (pathComponentIn x F) := + ⟨x, mem_pathComponentIn_self h, fun ⟨γ, hγ⟩ ↦ by + refine ⟨γ, fun t ↦ + ⟨(γ.truncateOfLE t.2.1).cast (γ.extend_zero.symm) (γ.extend_extends' t).symm, fun t' ↦ ?_⟩⟩ + dsimp [Path.truncateOfLE, Path.truncate] + exact γ.extend_extends' ⟨min (max t'.1 0) t.1, by simp [t.2.1, t.2.2]⟩ ▸ hγ _⟩ + +theorem isPathConnected_pathComponent : IsPathConnected (pathComponent x) := by + rw [← pathComponentIn_univ] + exact isPathConnected_pathComponentIn (mem_univ x) + theorem IsPathConnected.union {U V : Set X} (hU : IsPathConnected U) (hV : IsPathConnected V) (hUV : (U ∩ V).Nonempty) : IsPathConnected (U ∪ V) := by rcases hUV with ⟨x, xU, xV⟩ From fa8e2105aa86b708d0b03e08ba31c88c1a1f982c Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Fri, 18 Oct 2024 21:51:43 +0000 Subject: [PATCH 088/131] chore: update Mathlib dependencies 2024-10-18 (#17929) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index c93ba44db832e..f8536ac19fc8c 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "c521f0185f4dd42b6aa4898010d5ba5357c57c9f", + "rev": "9e3d0d810e9021767c360756f066574f72e8fcce", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 992ec73d5d2540c7d7bd628b704ef49909a13835 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 19 Oct 2024 00:06:23 +0000 Subject: [PATCH 089/131] feat: generalize `LinearMap.exists_rightInverse_of_surjective` to projective modules (#17560) Previously it only held in vector spaces. --- Mathlib/Algebra/Module/Projective.lean | 6 +++++- Mathlib/LinearAlgebra/Basis/VectorSpace.lean | 9 --------- Mathlib/LinearAlgebra/Dimension/LinearMap.lean | 1 + Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean | 1 + Mathlib/Topology/Algebra/SeparationQuotient.lean | 1 + 5 files changed, 8 insertions(+), 10 deletions(-) diff --git a/Mathlib/Algebra/Module/Projective.lean b/Mathlib/Algebra/Module/Projective.lean index b4387f01f63c9..7a68117b7f456 100644 --- a/Mathlib/Algebra/Module/Projective.lean +++ b/Mathlib/Algebra/Module/Projective.lean @@ -91,7 +91,7 @@ theorem projective_def' : /-- A projective R-module has the property that maps from it lift along surjections. -/ theorem projective_lifting_property [h : Projective R P] (f : M →ₗ[R] N) (g : P →ₗ[R] N) - (hf : Function.Surjective f) : ∃ h : P →ₗ[R] M, f.comp h = g := by + (hf : Function.Surjective f) : ∃ h : P →ₗ[R] M, f ∘ₗ h = g := by /- Here's the first step of the proof. Recall that `X →₀ R` is Lean's way of talking about the free `R`-module @@ -110,6 +110,10 @@ theorem projective_lifting_property [h : Projective R P] (f : M →ₗ[R] N) (g conv_rhs => rw [← hs p] simp [φ, Finsupp.linearCombination_apply, Function.surjInv_eq hf, map_finsupp_sum] +theorem _root_.LinearMap.exists_rightInverse_of_surjective [Projective R P] + (f : M →ₗ[R] P) (hf_surj : range f = ⊤) : ∃ g : P →ₗ[R] M, f ∘ₗ g = LinearMap.id := + projective_lifting_property f (.id : P →ₗ[R] P) (LinearMap.range_eq_top.1 hf_surj) + /-- A module which satisfies the universal property is projective: If all surjections of `R`-modules `(P →₀ R) →ₗ[R] P` have `R`-linear left inverse maps, then `P` is projective. -/ diff --git a/Mathlib/LinearAlgebra/Basis/VectorSpace.lean b/Mathlib/LinearAlgebra/Basis/VectorSpace.lean index e7a219e569dab..c16a5b57e71ea 100644 --- a/Mathlib/LinearAlgebra/Basis/VectorSpace.lean +++ b/Mathlib/LinearAlgebra/Basis/VectorSpace.lean @@ -266,15 +266,6 @@ theorem Submodule.exists_isCompl (p : Submodule K V) : ∃ q : Submodule K V, Is instance Submodule.complementedLattice : ComplementedLattice (Submodule K V) := ⟨Submodule.exists_isCompl⟩ -theorem LinearMap.exists_rightInverse_of_surjective (f : V →ₗ[K] V') (hf_surj : range f = ⊤) : - ∃ g : V' →ₗ[K] V, f.comp g = LinearMap.id := by - let C := Basis.ofVectorSpaceIndex K V' - let hC := Basis.ofVectorSpace K V' - haveI : Inhabited V := ⟨0⟩ - refine ⟨(hC.constr ℕ : _ → _) (C.restrict (invFun f)), hC.ext fun c => ?_⟩ - rw [LinearMap.comp_apply, hC.constr_basis] - simp [hC, rightInverse_invFun (LinearMap.range_eq_top.1 hf_surj) c] - /-- Any linear map `f : p →ₗ[K] V'` defined on a subspace `p` can be extended to the whole space. -/ theorem LinearMap.exists_extend {p : Submodule K V} (f : p →ₗ[K] V') : diff --git a/Mathlib/LinearAlgebra/Dimension/LinearMap.lean b/Mathlib/LinearAlgebra/Dimension/LinearMap.lean index ca2cb869a8ec5..d9f144980e339 100644 --- a/Mathlib/LinearAlgebra/Dimension/LinearMap.lean +++ b/Mathlib/LinearAlgebra/Dimension/LinearMap.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ +import Mathlib.Algebra.Module.Projective import Mathlib.LinearAlgebra.Dimension.DivisionRing import Mathlib.LinearAlgebra.Dimension.FreeAndStrongRankCondition diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 172f5f1c0d87a..372c3c0e74552 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ +import Mathlib.Algebra.Module.Projective import Mathlib.FieldTheory.Finiteness /-! diff --git a/Mathlib/Topology/Algebra/SeparationQuotient.lean b/Mathlib/Topology/Algebra/SeparationQuotient.lean index 83775bba6cbe3..8a99f4dc7ef34 100644 --- a/Mathlib/Topology/Algebra/SeparationQuotient.lean +++ b/Mathlib/Topology/Algebra/SeparationQuotient.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.LinearAlgebra.Basis.VectorSpace +import Mathlib.Algebra.Module.Projective import Mathlib.Topology.Algebra.Module.Basic import Mathlib.Topology.Maps.OpenQuotient From 6e38c6649191d8bf77deb739de42ae8df4581c0d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 19 Oct 2024 00:45:04 +0000 Subject: [PATCH 090/131] feat: algebraic properties of `LinearMap.compMultilinear` (#17932) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This also demotes `LinearMap.compAlternatingMap` to a plain function, for consistency. To make up for this loss, it introduces `LinearMap.compAlternatingMapₗ` as the stronger linear version. --- Mathlib/LinearAlgebra/Alternating/Basic.lean | 56 ++++++++++++++++---- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 47 ++++++++++++++-- 2 files changed, 88 insertions(+), 15 deletions(-) diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index 994b3b3d983db..f3a7aafc8973a 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -407,19 +407,12 @@ end AlternatingMap namespace LinearMap -variable {N₂ : Type*} [AddCommMonoid N₂] [Module R N₂] +variable {S : Type*} {N₂ : Type*} [AddCommMonoid N₂] [Module R N₂] /-- Composing an alternating map with a linear map on the left gives again an alternating map. -/ -def compAlternatingMap (g : N →ₗ[R] N₂) : (M [⋀^ι]→ₗ[R] N) →+ (M [⋀^ι]→ₗ[R] N₂) where - toFun f := - { g.compMultilinearMap (f : MultilinearMap R (fun _ : ι => M) N) with - map_eq_zero_of_eq' := fun v i j h hij => by simp [f.map_eq_zero_of_eq v h hij] } - map_zero' := by - ext - simp - map_add' a b := by - ext - simp +def compAlternatingMap (g : N →ₗ[R] N₂) (f : M [⋀^ι]→ₗ[R] N) : M [⋀^ι]→ₗ[R] N₂ where + __ := g.compMultilinearMap (f : MultilinearMap R (fun _ : ι => M) N) + map_eq_zero_of_eq' v i j h hij := by simp [f.map_eq_zero_of_eq v h hij] @[simp] theorem coe_compAlternatingMap (g : N →ₗ[R] N₂) (f : M [⋀^ι]→ₗ[R] N) : @@ -431,6 +424,47 @@ theorem compAlternatingMap_apply (g : N →ₗ[R] N₂) (f : M [⋀^ι]→ₗ[R] g.compAlternatingMap f m = g (f m) := rfl +@[simp] +theorem compAlternatingMap_zero (g : N →ₗ[R] N₂) : + g.compAlternatingMap (0 : M [⋀^ι]→ₗ[R] N) = 0 := + AlternatingMap.ext fun _ => map_zero g + +@[simp] +theorem zero_compAlternatingMap (f: M [⋀^ι]→ₗ[R] N) : + (0 : N →ₗ[R] N₂).compAlternatingMap f = 0 := rfl + +@[simp] +theorem compAlternatingMap_add (g : N →ₗ[R] N₂) (f₁ f₂ : M [⋀^ι]→ₗ[R] N) : + g.compAlternatingMap (f₁ + f₂) = g.compAlternatingMap f₁ + g.compAlternatingMap f₂ := + AlternatingMap.ext fun _ => map_add g _ _ + +@[simp] +theorem add_compAlternatingMap (g₁ g₂ : N →ₗ[R] N₂) (f: M [⋀^ι]→ₗ[R] N) : + (g₁ + g₂).compAlternatingMap f = g₁.compAlternatingMap f + g₂.compAlternatingMap f := rfl + +@[simp] +theorem compAlternatingMap_smul [Monoid S] [DistribMulAction S N] [DistribMulAction S N₂] + [SMulCommClass R S N] [SMulCommClass R S N₂] [CompatibleSMul N N₂ S R] + (g : N →ₗ[R] N₂) (s : S) (f : M [⋀^ι]→ₗ[R] N) : + g.compAlternatingMap (s • f) = s • g.compAlternatingMap f := + AlternatingMap.ext fun _ => g.map_smul_of_tower _ _ + +@[simp] +theorem smul_compAlternatingMap [Monoid S] [DistribMulAction S N₂] [SMulCommClass R S N₂] + (g : N →ₗ[R] N₂) (s : S) (f : M [⋀^ι]→ₗ[R] N) : + (s • g).compAlternatingMap f = s • g.compAlternatingMap f := rfl + +variable (S) in +/-- `LinearMap.compAlternatingMap` as an `S`-linear map. -/ +@[simps] +def compAlternatingMapₗ [Semiring S] [Module S N] [Module S N₂] + [SMulCommClass R S N] [SMulCommClass R S N₂] [LinearMap.CompatibleSMul N N₂ S R] + (g : N →ₗ[R] N₂) : + (M [⋀^ι]→ₗ[R] N) →ₗ[S] (M [⋀^ι]→ₗ[R] N₂) where + toFun := g.compAlternatingMap + map_add' := g.compAlternatingMap_add + map_smul' := g.compAlternatingMap_smul + theorem smulRight_eq_comp {R M₁ M₂ ι : Type*} [CommSemiring R] [AddCommMonoid M₁] [AddCommMonoid M₂] [Module R M₁] [Module R M₂] (f : M₁ [⋀^ι]→ₗ[R] R) (z : M₂) : f.smulRight z = (LinearMap.id.smulRight z).compAlternatingMap f := diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 5fa2ebd1c7ff4..8db7f46936b45 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -787,6 +787,36 @@ theorem compMultilinearMap_apply (g : M₂ →ₗ[R] M₃) (f : MultilinearMap R g.compMultilinearMap f m = g (f m) := rfl +@[simp] +theorem compMultilinearMap_zero (g : M₂ →ₗ[R] M₃) : + g.compMultilinearMap (0 : MultilinearMap R M₁ M₂) = 0 := + MultilinearMap.ext fun _ => map_zero g + +@[simp] +theorem zero_compMultilinearMap (f: MultilinearMap R M₁ M₂) : + (0 : M₂ →ₗ[R] M₃).compMultilinearMap f = 0 := rfl + +@[simp] +theorem compMultilinearMap_add (g : M₂ →ₗ[R] M₃) (f₁ f₂ : MultilinearMap R M₁ M₂) : + g.compMultilinearMap (f₁ + f₂) = g.compMultilinearMap f₁ + g.compMultilinearMap f₂ := + MultilinearMap.ext fun _ => map_add g _ _ + +@[simp] +theorem add_compMultilinearMap (g₁ g₂ : M₂ →ₗ[R] M₃) (f: MultilinearMap R M₁ M₂) : + (g₁ + g₂).compMultilinearMap f = g₁.compMultilinearMap f + g₂.compMultilinearMap f := rfl + +@[simp] +theorem compMultilinearMap_smul [Monoid S] [DistribMulAction S M₂] [DistribMulAction S M₃] + [SMulCommClass R S M₂] [SMulCommClass R S M₃] [CompatibleSMul M₂ M₃ S R] + (g : M₂ →ₗ[R] M₃) (s : S) (f : MultilinearMap R M₁ M₂) : + g.compMultilinearMap (s • f) = s • g.compMultilinearMap f := + MultilinearMap.ext fun _ => g.map_smul_of_tower _ _ + +@[simp] +theorem smul_compMultilinearMap [Monoid S] [DistribMulAction S M₃] [SMulCommClass R S M₃] + (g : M₂ →ₗ[R] M₃) (s : S) (f : MultilinearMap R M₁ M₂) : + (s • g).compMultilinearMap f = s • g.compMultilinearMap f := rfl + /-- The multilinear version of `LinearMap.subtype_comp_codRestrict` -/ @[simp] theorem subtype_compMultilinearMap_codRestrict (f : MultilinearMap R M₁ M₂) (p : Submodule R M₂) @@ -835,12 +865,23 @@ instance : Module S (MultilinearMap R M₁ M₂) := instance [NoZeroSMulDivisors S M₂] : NoZeroSMulDivisors S (MultilinearMap R M₁ M₂) := coe_injective.noZeroSMulDivisors _ rfl coe_smul +variable [AddCommMonoid M₃] [Module S M₃] [Module R M₃] [SMulCommClass R S M₃] + +variable (S) in +/-- `LinearMap.compMultilinearMap` as an `S`-linear map. -/ +@[simps] +def _root_.LinearMap.compMultilinearMapₗ [Semiring S] [Module S M₂] [Module S M₃] + [SMulCommClass R S M₂] [SMulCommClass R S M₃] [LinearMap.CompatibleSMul M₂ M₃ S R] + (g : M₂ →ₗ[R] M₃) : + MultilinearMap R M₁ M₂ →ₗ[S] MultilinearMap R M₁ M₃ where + toFun := g.compMultilinearMap + map_add' := g.compMultilinearMap_add + map_smul' := g.compMultilinearMap_smul + variable (R S M₁ M₂ M₃) section OfSubsingleton -variable [AddCommMonoid M₃] [Module S M₃] [Module R M₃] [SMulCommClass R S M₃] - /-- Linear equivalence between linear maps `M₂ →ₗ[R] M₃` and one-multilinear maps `MultilinearMap R (fun _ : ι ↦ M₂) M₃`. -/ @[simps (config := { simpRhs := true })] @@ -904,8 +945,6 @@ def constLinearEquivOfIsEmpty [IsEmpty ι] : M₂ ≃ₗ[S] MultilinearMap R M left_inv _ := rfl right_inv f := ext fun _ => MultilinearMap.congr_arg f <| Subsingleton.elim _ _ -variable [AddCommMonoid M₃] [Module R M₃] [Module S M₃] [SMulCommClass R S M₃] - /-- `MultilinearMap.domDomCongr` as a `LinearEquiv`. -/ @[simps apply symm_apply] def domDomCongrLinearEquiv {ι₁ ι₂} (σ : ι₁ ≃ ι₂) : From f6b96180cc040c8fdf7eb10746000229ef5eba25 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sat, 19 Oct 2024 01:19:58 +0000 Subject: [PATCH 091/131] chore: update Mathlib dependencies 2024-10-19 (#17933) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index f8536ac19fc8c..6f116ba1524dd 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "9e3d0d810e9021767c360756f066574f72e8fcce", + "rev": "1e0bf50b357069e1d658512a579a5faac6587c40", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From bd5ba6b8fd201d0165ca6068e0a9437a56c4b58f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 19 Oct 2024 02:20:12 +0000 Subject: [PATCH 092/131] chore(SetTheory/Ordinal/Basic): deprecate duplicate `Nat.cast` lemmas (#17857) --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 37 +++++++++++------------ Mathlib/SetTheory/Ordinal/Notation.lean | 18 +++++------ 2 files changed, 27 insertions(+), 28 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 90d58666720d4..b43f781637c4a 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -2099,6 +2099,10 @@ theorem Ordinal.not_bddAbove_compl_of_small (s : Set Ordinal.{u}) [hs : Small.{u namespace Ordinal +instance instCharZero : CharZero Ordinal := by + refine ⟨fun a b h ↦ ?_⟩ + rwa [← Cardinal.ord_nat, ← Cardinal.ord_nat, Cardinal.ord_inj, Nat.cast_inj] at h + @[simp] theorem one_add_natCast (m : ℕ) : 1 + (m : Ordinal) = succ m := by rw [← Nat.cast_one, ← Nat.cast_add, add_comm] @@ -2121,42 +2125,37 @@ theorem natCast_mul (m : ℕ) : ∀ n : ℕ, ((m * n : ℕ) : Ordinal) = m * n @[deprecated (since := "2024-04-17")] alias nat_cast_mul := natCast_mul -/-- Alias of `Nat.cast_le`, specialized to `Ordinal` --/ -theorem natCast_le {m n : ℕ} : (m : Ordinal) ≤ n ↔ m ≤ n := by - rw [← Cardinal.ord_nat, ← Cardinal.ord_nat, Cardinal.ord_le_ord, Cardinal.natCast_le] +@[deprecated Nat.cast_le (since := "2024-10-17")] +theorem natCast_le {m n : ℕ} : (m : Ordinal) ≤ n ↔ m ≤ n := Nat.cast_le @[deprecated (since := "2024-04-17")] alias nat_cast_le := natCast_le -/-- Alias of `Nat.cast_inj`, specialized to `Ordinal` --/ -theorem natCast_inj {m n : ℕ} : (m : Ordinal) = n ↔ m = n := by - simp only [le_antisymm_iff, natCast_le] +@[deprecated Nat.cast_inj (since := "2024-10-17")] +theorem natCast_inj {m n : ℕ} : (m : Ordinal) = n ↔ m = n := Nat.cast_inj @[deprecated (since := "2024-04-17")] alias nat_cast_inj := natCast_inj -instance charZero : CharZero Ordinal where - cast_injective _ _ := natCast_inj.mp - -/-- Alias of `Nat.cast_lt`, specialized to `Ordinal` --/ +@[deprecated Nat.cast_lt (since := "2024-10-17")] theorem natCast_lt {m n : ℕ} : (m : Ordinal) < n ↔ m < n := Nat.cast_lt @[deprecated (since := "2024-04-17")] alias nat_cast_lt := natCast_lt -/-- Alias of `Nat.cast_eq_zero`, specialized to `Ordinal` --/ +@[deprecated Nat.cast_eq_zero (since := "2024-10-17")] theorem natCast_eq_zero {n : ℕ} : (n : Ordinal) = 0 ↔ n = 0 := Nat.cast_eq_zero @[deprecated (since := "2024-04-17")] alias nat_cast_eq_zero := natCast_eq_zero -/-- Alias of `Nat.cast_eq_zero`, specialized to `Ordinal` --/ +@[deprecated Nat.cast_ne_zero (since := "2024-10-17")] theorem natCast_ne_zero {n : ℕ} : (n : Ordinal) ≠ 0 ↔ n ≠ 0 := Nat.cast_ne_zero @[deprecated (since := "2024-04-17")] alias nat_cast_ne_zero := natCast_ne_zero -/-- Alias of `Nat.cast_pos'`, specialized to `Ordinal` --/ +@[deprecated Nat.cast_pos' (since := "2024-10-17")] theorem natCast_pos {n : ℕ} : (0 : Ordinal) < n ↔ 0 < n := Nat.cast_pos' @[deprecated (since := "2024-04-17")] @@ -2165,10 +2164,10 @@ alias nat_cast_pos := natCast_pos @[simp, norm_cast] theorem natCast_sub (m n : ℕ) : ((m - n : ℕ) : Ordinal) = m - n := by rcases le_total m n with h | h - · rw [tsub_eq_zero_iff_le.2 h, Ordinal.sub_eq_zero_iff_le.2 (natCast_le.2 h)] + · rw [tsub_eq_zero_iff_le.2 h, Ordinal.sub_eq_zero_iff_le.2 (Nat.cast_le.2 h)] rfl · apply (add_left_cancel n).1 - rw [← Nat.cast_add, add_tsub_cancel_of_le h, Ordinal.add_sub_cancel_of_le (natCast_le.2 h)] + rw [← Nat.cast_add, add_tsub_cancel_of_le h, Ordinal.add_sub_cancel_of_le (Nat.cast_le.2 h)] @[deprecated (since := "2024-04-17")] alias nat_cast_sub := natCast_sub @@ -2177,12 +2176,12 @@ alias nat_cast_sub := natCast_sub theorem natCast_div (m n : ℕ) : ((m / n : ℕ) : Ordinal) = m / n := by rcases eq_or_ne n 0 with (rfl | hn) · simp - · have hn' := natCast_ne_zero.2 hn + · have hn' : (n : Ordinal) ≠ 0 := Nat.cast_ne_zero.2 hn apply le_antisymm - · rw [le_div hn', ← natCast_mul, natCast_le, mul_comm] + · rw [le_div hn', ← natCast_mul, Nat.cast_le, mul_comm] apply Nat.div_mul_le_self - · rw [div_le hn', ← add_one_eq_succ, ← Nat.cast_succ, ← natCast_mul, natCast_lt, mul_comm, ← - Nat.div_lt_iff_lt_mul (Nat.pos_of_ne_zero hn)] + · rw [div_le hn', ← add_one_eq_succ, ← Nat.cast_succ, ← natCast_mul, Nat.cast_lt, mul_comm, + ← Nat.div_lt_iff_lt_mul (Nat.pos_of_ne_zero hn)] apply Nat.lt_succ_self @[deprecated (since := "2024-04-17")] diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 050c6b4f8f653..bab9ce4002c72 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -139,7 +139,7 @@ theorem repr_one : repr (ofNat 1) = (1 : ℕ) := repr_ofNat 1 theorem omega0_le_oadd (e n a) : ω ^ repr e ≤ repr (oadd e n a) := by refine le_trans ?_ (le_add_right _ _) - simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos (repr e) omega0_pos).2 (natCast_le.2 n.2) + simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos (repr e) omega0_pos).2 (Nat.cast_le.2 n.2) @[deprecated (since := "2024-09-30")] alias omega_le_oadd := omega0_le_oadd @@ -275,7 +275,7 @@ theorem oadd_lt_oadd_2 {e o₁ o₂ : ONote} {n₁ n₂ : ℕ+} (h₁ : NF (oadd oadd e n₁ o₁ < oadd e n₂ o₂ := by simp only [lt_def, repr] refine lt_of_lt_of_le ((add_lt_add_iff_left _).2 h₁.snd'.repr_lt) (le_trans ?_ (le_add_right _ _)) - rwa [← mul_succ,Ordinal.mul_le_mul_iff_left (opow_pos _ omega0_pos), succ_le_iff, natCast_lt] + rwa [← mul_succ,Ordinal.mul_le_mul_iff_left (opow_pos _ omega0_pos), succ_le_iff, Nat.cast_lt] theorem oadd_lt_oadd_3 {e n a₁ a₂} (h : a₁ < a₂) : oadd e n a₁ < oadd e n a₂ := by rw [lt_def]; unfold repr @@ -451,7 +451,7 @@ theorem repr_add : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ + o₂) = rep cases he' : e' <;> simp only [he', zero_def, opow_zero, repr, gt_iff_lt] at this ⊢ <;> exact lt_of_le_of_lt (le_add_right _ _) this · simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos (repr e') omega0_pos).2 - (natCast_le.2 n'.pos) + (Nat.cast_le.2 n'.pos) · rw [ee, ← add_assoc, ← mul_add] theorem sub_nfBelow : ∀ {o₁ o₂ b}, NFBelow o₁ b → NF o₂ → NFBelow (o₁ - o₂) b @@ -507,7 +507,7 @@ theorem repr_sub : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ - o₂) = rep refine (Ordinal.sub_eq_of_add_eq <| add_absorp h₂.snd'.repr_lt <| le_trans ?_ (le_add_right _ _)).symm - simpa using mul_le_mul_left' (natCast_le.2 <| Nat.succ_pos _) _ + exact Ordinal.le_mul_left _ (Nat.cast_lt.2 <| Nat.succ_pos _) · exact (Ordinal.sub_eq_of_add_eq <| add_absorp (h₂.below_of_lt ee).repr_lt <| omega0_le_oadd _ _ _).symm @@ -564,7 +564,7 @@ theorem repr_mul : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ * o₂) = rep simp [(· * ·)] have ao : repr a₁ + ω ^ repr e₁ * (n₁ : ℕ) = ω ^ repr e₁ * (n₁ : ℕ) := by apply add_absorp h₁.snd'.repr_lt - simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos _ omega0_pos).2 (natCast_le.2 n₁.2) + simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos _ omega0_pos).2 (Nat.cast_le.2 n₁.2) by_cases e0 : e₂ = 0 · cases' Nat.exists_eq_succ_of_ne_zero n₂.ne_zero with x xe simp only [e0, repr, PNat.mul_coe, natCast_mul, opow_zero, one_mul] @@ -578,7 +578,7 @@ theorem repr_mul : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ * o₂) = rep congr 2 have := mt repr_inj.1 e0 rw [add_mul_limit ao (isLimit_opow_left isLimit_omega0 this), mul_assoc, - mul_omega0_dvd (natCast_pos.2 n₁.pos) (nat_lt_omega0 _)] + mul_omega0_dvd (Nat.cast_pos'.2 n₁.pos) (nat_lt_omega0 _)] simpa using opow_dvd_opow ω (one_le_iff_ne_zero.2 this) /-- Calculate division and remainder of `o` mod `ω`: @@ -857,7 +857,7 @@ theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω · exact lt_of_lt_of_le Rl (opow_le_opow_right omega0_pos <| - mul_le_mul_left' (succ_le_succ_iff.2 (natCast_le.2 (le_of_lt k.lt_succ_self))) _) + mul_le_mul_left' (succ_le_succ_iff.2 (Nat.cast_le.2 (le_of_lt k.lt_succ_self))) _) calc (ω0 ^ (k.succ : Ordinal)) * α' + R' _ = (ω0 ^ succ (k : Ordinal)) * α' + ((ω0 ^ (k : Ordinal)) * α' * m + R) := by @@ -869,12 +869,12 @@ theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω dvd_add (dvd_mul_of_dvd_left (by simpa using opow_dvd_opow ω (one_le_iff_ne_zero.2 e0)) _) d rw [mul_add (ω0 ^ (k : Ordinal)), add_assoc, ← mul_assoc, ← opow_succ, add_mul_limit _ (isLimit_iff_omega0_dvd.2 ⟨ne_of_gt α0, αd⟩), mul_assoc, - @mul_omega0_dvd n (natCast_pos.2 n.pos) (nat_lt_omega0 _) _ αd] + @mul_omega0_dvd n (Nat.cast_pos'.2 n.pos) (nat_lt_omega0 _) _ αd] apply @add_absorp _ (repr a0 * succ ↑k) · refine principal_add_omega0_opow _ ?_ Rl rw [opow_mul, opow_succ, Ordinal.mul_lt_mul_iff_left ω00] exact No.snd'.repr_lt - · have := mul_le_mul_left' (one_le_iff_pos.2 <| natCast_pos.2 n.pos) (ω0 ^ succ (k : Ordinal)) + · have := mul_le_mul_left' (one_le_iff_pos.2 <| Nat.cast_pos'.2 n.pos) (ω0 ^ succ (k : Ordinal)) rw [opow_mul] simpa [-opow_succ] · cases m From 81e65433ee271e679931bc2674586fce8fd5db8e Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 05:19:36 +0000 Subject: [PATCH 093/131] chore(Data/Nat/{Bits, Bitwise}): use `Bool.toNat` (#17925) --- Mathlib/Data/Nat/Bits.lean | 6 +++--- Mathlib/Data/Nat/Bitwise.lean | 12 ++++-------- Mathlib/Data/Nat/Size.lean | 4 ++-- 3 files changed, 9 insertions(+), 13 deletions(-) diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 371661d05995b..71f01660fc34f 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -73,7 +73,7 @@ lemma bodd_mul (m n : ℕ) : bodd (m * n) = (bodd m && bodd n) := by simp only [mul_succ, bodd_add, IH, bodd_succ] cases bodd m <;> cases bodd n <;> rfl -lemma mod_two_of_bodd (n : ℕ) : n % 2 = cond (bodd n) 1 0 := by +lemma mod_two_of_bodd (n : ℕ) : n % 2 = (bodd n).toNat := by have := congr_arg bodd (mod_add_div n 2) simp? [not] at this says simp only [bodd_add, bodd_mul, bodd_succ, not, bodd_zero, Bool.false_and, Bool.bne_false] @@ -100,7 +100,7 @@ lemma div2_succ (n : ℕ) : div2 (n + 1) = cond (bodd n) (succ (div2 n)) (div2 n attribute [local simp] Nat.add_comm Nat.add_assoc Nat.add_left_comm Nat.mul_comm Nat.mul_assoc -lemma bodd_add_div2 : ∀ n, cond (bodd n) 1 0 + 2 * div2 n = n +lemma bodd_add_div2 : ∀ n, (bodd n).toNat + 2 * div2 n = n | 0 => rfl | succ n => by simp only [bodd_succ, Bool.cond_not, div2_succ, Nat.mul_comm] @@ -117,7 +117,7 @@ lemma div2_val (n) : div2 n = n / 2 := by /-- `bit b` appends the digit `b` to the binary representation of its natural number input. -/ def bit (b : Bool) : ℕ → ℕ := cond b (2 * · + 1) (2 * ·) -lemma bit_val (b n) : bit b n = 2 * n + cond b 1 0 := by +lemma bit_val (b n) : bit b n = 2 * n + b.toNat := by cases b <;> rfl lemma bit_decomp (n : Nat) : bit (bodd n) (div2 n) = n := diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 2a45dc564a4b8..cce15c0d55eee 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -89,22 +89,18 @@ lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by <;> simp_all (config := {decide := true}) [two_mul] lemma bit_mod_two (a : Bool) (x : ℕ) : - bit a x % 2 = if a then 1 else 0 := by - #adaptation_note /-- nightly-2024-03-16: simp was - -- simp (config := { unfoldPartialApp := true }) only [bit, bit1, bit0, ← mul_two, - -- Bool.cond_eq_ite] -/ - simp only [bit, ite_apply, ← mul_two, Bool.cond_eq_ite] - split_ifs <;> simp [Nat.add_mod] + bit a x % 2 = a.toNat := by + cases a <;> simp [bit_val, mul_add_mod] @[simp] lemma bit_mod_two_eq_zero_iff (a x) : bit a x % 2 = 0 ↔ !a := by - rw [bit_mod_two]; split_ifs <;> simp_all + simp [bit_mod_two] @[simp] lemma bit_mod_two_eq_one_iff (a x) : bit a x % 2 = 1 ↔ a := by - rw [bit_mod_two]; split_ifs <;> simp_all + simp [bit_mod_two] @[simp] theorem lor_bit : ∀ a m b n, bit a m ||| bit b n = bit (a || b) (m ||| n) := diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index 1fdafff0b8fd0..fec3892450d06 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -19,8 +19,8 @@ theorem shiftLeft_eq_mul_pow (m) : ∀ n, m <<< n = m * 2 ^ n := shiftLeft_eq _ theorem shiftLeft'_tt_eq_mul_pow (m) : ∀ n, shiftLeft' true m n + 1 = (m + 1) * 2 ^ n | 0 => by simp [shiftLeft', pow_zero, Nat.one_mul] | k + 1 => by - rw [shiftLeft', bit_val, cond_true, add_assoc, ← Nat.mul_add_one, shiftLeft'_tt_eq_mul_pow m k, - mul_left_comm, mul_comm 2, pow_succ] + rw [shiftLeft', bit_val, Bool.toNat_true, add_assoc, ← Nat.mul_add_one, + shiftLeft'_tt_eq_mul_pow m k, mul_left_comm, mul_comm 2, pow_succ] end From 74600feef71163f17723cd491710cbe1cec142d7 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 19 Oct 2024 06:27:29 +0000 Subject: [PATCH 094/131] refactor(Order/Filter): don't use `generate` for `CompleteLattice` (#17799) --- Mathlib/Analysis/Complex/Basic.lean | 4 +- .../RotationNumber/TranslationNumber.lean | 2 +- Mathlib/MeasureTheory/Group/Measure.lean | 2 +- Mathlib/Order/Filter/Basic.lean | 44 ++++++-------- Mathlib/Order/Filter/Defs.lean | 59 +++++++++++-------- 5 files changed, 56 insertions(+), 55 deletions(-) diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index e4257988bc45b..55726f263e5bf 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -598,10 +598,10 @@ theorem ofReal_tsum (f : α → ℝ) : (↑(∑' a, f a) : ℂ) = ∑' a, ↑(f RCLike.ofReal_tsum _ _ theorem hasSum_re {f : α → ℂ} {x : ℂ} (h : HasSum f x) : HasSum (fun x => (f x).re) x.re := - RCLike.hasSum_re _ h + RCLike.hasSum_re ℂ h theorem hasSum_im {f : α → ℂ} {x : ℂ} (h : HasSum f x) : HasSum (fun x => (f x).im) x.im := - RCLike.hasSum_im _ h + RCLike.hasSum_im ℂ h theorem re_tsum {f : α → ℂ} (h : Summable f) : (∑' a, f a).re = ∑' a, (f a).re := RCLike.re_tsum _ h diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index 745facc0985ce..65ee8350a4451 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -607,7 +607,7 @@ theorem tendsto_translationNumber_of_dist_bounded_aux (x : ℕ → ℝ) (C : ℝ theorem translationNumber_eq_of_dist_bounded {f g : CircleDeg1Lift} (C : ℝ) (H : ∀ n : ℕ, dist ((f ^ n) 0) ((g ^ n) 0) ≤ C) : τ f = τ g := Eq.symm <| g.translationNumber_eq_of_tendsto_aux <| - f.tendsto_translationNumber_of_dist_bounded_aux _ C H + f.tendsto_translationNumber_of_dist_bounded_aux (fun n ↦ (g ^ n) 0) C H @[simp] theorem translationNumber_one : τ 1 = 0 := diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 3e36cb114970b..e058bb5e8e9af 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -775,7 +775,7 @@ nonrec theorem _root_.MulEquiv.isHaarMeasure_map [BorelSpace G] [TopologicalGrou [TopologicalGroup H] (e : G ≃* H) (he : Continuous e) (hesymm : Continuous e.symm) : IsHaarMeasure (Measure.map e μ) := let f : G ≃ₜ H := .mk e - isHaarMeasure_map μ e he e.surjective f.closedEmbedding.tendsto_cocompact + isHaarMeasure_map μ (e : G →* H) he e.surjective f.closedEmbedding.tendsto_cocompact /-- A convenience wrapper for MeasureTheory.Measure.isAddHaarMeasure_map`. -/ instance _root_.ContinuousLinearEquiv.isAddHaarMeasure_map diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 84c4397e07ac6..8b546db30b150 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -83,6 +83,9 @@ instance inhabitedMem : Inhabited { s : Set α // s ∈ f } := theorem filter_eq_iff : f = g ↔ f.sets = g.sets := ⟨congr_arg _, filter_eq⟩ +@[simp] theorem sets_subset_sets : f.sets ⊆ g.sets ↔ g ≤ f := .rfl +@[simp] theorem sets_ssubset_sets : f.sets ⊂ g.sets ↔ g < f := .rfl + /-- 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 := @@ -253,21 +256,20 @@ theorem mem_inf_iff_superset {f g : Filter α} {s : Set α} : section CompleteLattice -/- We lift the complete lattice along the Galois connection `generate` / `sets`. Unfortunately, - we want to have different definitional equalities for some lattice operations. So we define them - upfront and change the lattice operations for the complete lattice instance. -/ -instance instCompleteLatticeFilter : CompleteLattice (Filter α) := - { @OrderDual.instCompleteLattice _ (giGenerate α).liftCompleteLattice with - le := (· ≤ ·) - top := ⊤ - le_top := fun _ _s hs => (mem_top.1 hs).symm ▸ univ_mem - inf := (· ⊓ ·) - inf_le_left := fun _ _ _ => mem_inf_of_left - inf_le_right := fun _ _ _ => mem_inf_of_right - le_inf := fun _ _ _ h₁ h₂ _s ⟨_a, ha, _b, hb, hs⟩ => hs.symm ▸ inter_mem (h₁ ha) (h₂ hb) - sSup := join ∘ 𝓟 - le_sSup := fun _ _f hf _s hs => hs hf - sSup_le := fun _ _f hf _s hs _g hg => hf _ hg hs } +/- Complete lattice structure on `Filter α`. -/ +instance instCompleteLatticeFilter : CompleteLattice (Filter α) where + le_sup_left _ _ _ h := h.1 + le_sup_right _ _ _ h := h.2 + sup_le _ _ _ h₁ h₂ _ h := ⟨h₁ h, h₂ h⟩ + inf_le_left _ _ _ := mem_inf_of_left + inf_le_right _ _ _ := mem_inf_of_right + le_inf := fun _ _ _ h₁ h₂ _s ⟨_a, ha, _b, hb, hs⟩ => hs.symm ▸ inter_mem (h₁ ha) (h₂ hb) + le_sSup _ _ h₁ _ h₂ := h₂ h₁ + sSup_le _ _ h₁ _ h₂ _ h₃ := h₁ _ h₃ h₂ + sInf_le _ _ h₁ _ h₂ := by rw [← Filter.sSup_lowerBounds]; exact fun _ h₃ ↦ h₃ h₁ h₂ + le_sInf _ _ h₁ _ h₂ := by rw [← Filter.sSup_lowerBounds] at h₂; exact h₂ h₁ + le_top _ _ := univ_mem' + bot_le _ _ _ := trivial instance : Inhabited (Filter α) := ⟨⊥⟩ @@ -324,10 +326,6 @@ theorem mem_sup {f g : Filter α} {s : Set α} : s ∈ f ⊔ g ↔ s ∈ f ∧ s theorem union_mem_sup {f g : Filter α} {s t : Set α} (hs : s ∈ f) (ht : t ∈ g) : s ∪ t ∈ f ⊔ g := ⟨mem_of_superset hs subset_union_left, mem_of_superset ht subset_union_right⟩ -@[simp] -theorem mem_sSup {x : Set α} {s : Set (Filter α)} : x ∈ sSup s ↔ ∀ f ∈ s, x ∈ (f : Filter α) := - Iff.rfl - @[simp] theorem mem_iSup {x : Set α} {f : ι → Filter α} : x ∈ iSup f ↔ ∀ i, x ∈ f i := by simp only [← Filter.mem_sets, iSup_sets_eq, mem_iInter] @@ -337,7 +335,7 @@ theorem iSup_neBot {f : ι → Filter α} : (⨆ i, f i).NeBot ↔ ∃ i, (f i). simp [neBot_iff] theorem iInf_eq_generate (s : ι → Filter α) : iInf s = generate (⋃ i, (s i).sets) := - show generate _ = generate _ from congr_arg _ <| congr_arg sSup <| (range_comp _ _).symm + eq_of_forall_le_iff fun _ ↦ by simp [le_generate_iff] theorem mem_iInf_of_mem {f : ι → Filter α} (i : ι) {s} (hs : s ∈ f i) : s ∈ ⨅ i, f i := iInf_le f i hs @@ -598,7 +596,7 @@ abbrev coframeMinimalAxioms : Coframe.MinimalAxioms (Filter α) := iInf_sup_le_sup_sInf := fun f s t ⟨h₁, h₂⟩ => by classical rw [iInf_subtype'] - rw [sInf_eq_iInf', iInf_sets_eq_finite, mem_iUnion] at h₂ + rw [sInf_eq_iInf', ← Filter.mem_sets, iInf_sets_eq_finite, mem_iUnion] at h₂ obtain ⟨u, hu⟩ := h₂ rw [← Finset.inf_eq_iInf] at hu suffices ⨅ i : s, f ⊔ ↑i ≤ f ⊔ u.inf fun i => ↑i from this ⟨h₁, hu⟩ @@ -1572,10 +1570,6 @@ instance : LawfulFunctor (Filter : Type u → Type u) where theorem pure_sets (a : α) : (pure a : Filter α).sets = { s | a ∈ s } := rfl -@[simp] -theorem mem_pure {a : α} {s : Set α} : s ∈ (pure a : Filter α) ↔ a ∈ s := - Iff.rfl - @[simp] theorem eventually_pure {a : α} {p : α → Prop} : (∀ᶠ x in pure a, p x) ↔ p a := Iff.rfl diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean index cd23371bb3f2f..06e6369f2f301 100644 --- a/Mathlib/Order/Filter/Defs.lean +++ b/Mathlib/Order/Filter/Defs.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Jeremy Avigad -/ import Mathlib.Data.Set.Basic import Mathlib.Order.SetNotation +import Mathlib.Order.Bounds.Defs /-! # Definitions about filters @@ -141,6 +142,15 @@ scoped notation "𝓟" => Filter.principal @[simp] theorem mem_principal : s ∈ 𝓟 t ↔ t ⊆ s := Iff.rfl +/-- `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 where + pure x := .copy (𝓟 {x}) {s | x ∈ s} fun _ ↦ by simp + +@[simp] +theorem mem_pure {a : α} {s : Set α} : s ∈ (pure a : Filter α) ↔ a ∈ s := + Iff.rfl + /-- The *kernel* of a filter is the intersection of all its sets. -/ def ker (f : Filter α) : Set α := ⋂₀ f.sets @@ -164,11 +174,25 @@ instance : PartialOrder (Filter α) where 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 _) }⟩ +instance instSupSet : SupSet (Filter α) where + sSup S := join (𝓟 S) + +@[simp] theorem mem_sSup {S : Set (Filter α)} : s ∈ sSup S ↔ ∀ f ∈ S, s ∈ f := .rfl + +/-- Infimum of a set of filters. +This definition is marked as irreducible +so that Lean doesn't try to unfold it when unifying expressions. -/ +@[irreducible] +protected def sInf (s : Set (Filter α)) : Filter α := sSup (lowerBounds s) + +instance instInfSet : InfSet (Filter α) where + sInf := Filter.sInf + +protected theorem sSup_lowerBounds (s : Set (Filter α)) : sSup (lowerBounds s) = sInf s := by + simp [sInf, Filter.sInf] + +instance : Top (Filter α) where + top := .copy (sSup (Set.range pure)) {s | ∀ x, x ∈ s} <| by simp theorem mem_top_iff_forall {s : Set α} : s ∈ (⊤ : Filter α) ↔ ∀ x, x ∈ s := Iff.rfl @@ -178,11 +202,7 @@ 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 } + bot := .copy (sSup ∅) univ <| by simp @[simp] theorem mem_bot {s : Set α} : s ∈ (⊥ : Filter α) := @@ -190,7 +210,7 @@ theorem mem_bot {s : Set α} : s ∈ (⊥ : Filter α) := /-- The infimum of filters is the filter generated by intersections of elements of the two filters. -/ -instance : Inf (Filter α) := +instance instInf : Inf (Filter α) := ⟨fun f g : Filter α => { sets := { s | ∃ a ∈ f, ∃ b ∈ g, s = a ∩ b } univ_sets := ⟨_, univ_mem, _, univ_mem, by simp⟩ @@ -205,12 +225,8 @@ instance : Inf (Filter α) := 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⟩ } +instance instSup : Sup (Filter α) where + sup f g := .copy (sSup {f, g}) {s | s ∈ f ∧ s ∈ g} <| by simp /-- 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 @@ -319,15 +335,6 @@ in adding quantifiers to the middle of `Tendsto`s. See 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⟩ From a0521c9b7e636372cca49d817075d7da0fd4a681 Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 06:53:15 +0000 Subject: [PATCH 095/131] chore: deprecate `Setoid.Rel` (#16260) --- .../Limits/Shapes/SingleObj.lean | 2 +- Mathlib/Data/Quot.lean | 2 + Mathlib/Data/Setoid/Basic.lean | 94 ++++++++++--------- Mathlib/Data/Setoid/Partition.lean | 38 ++++---- Mathlib/GroupTheory/Congruence/Defs.lean | 8 +- Mathlib/GroupTheory/DoubleCoset.lean | 10 +- Mathlib/GroupTheory/GroupAction/Basic.lean | 17 ++-- Mathlib/GroupTheory/GroupAction/ConjAct.lean | 2 +- Mathlib/GroupTheory/GroupAction/Quotient.lean | 9 +- Mathlib/Order/Partition/Finpartition.lean | 4 +- .../RingTheory/AdicCompletion/Algebra.lean | 2 +- Mathlib/RingTheory/Congruence/Basic.lean | 4 +- .../RingTheory/Valuation/ValuationRing.lean | 14 +-- .../Topology/Algebra/ProperAction/Basic.lean | 2 +- Mathlib/Topology/DiscreteQuotient.lean | 16 ++-- Mathlib/Topology/MetricSpace/Contracting.lean | 4 +- Mathlib/Topology/Separation.lean | 2 +- 17 files changed, 118 insertions(+), 112 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean b/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean index e501871d0f757..a4e8274b8176e 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean @@ -75,7 +75,7 @@ variable {G : Type v} [Group G] (J : SingleObj G ⥤ Type u) equivalent to the `MulAction.orbitRel` equivalence relation on `J.obj (SingleObj.star G)`. -/ lemma Types.Quot.Rel.iff_orbitRel (x y : J.obj (SingleObj.star G)) : Types.Quot.Rel J ⟨SingleObj.star G, x⟩ ⟨SingleObj.star G, y⟩ - ↔ Setoid.Rel (MulAction.orbitRel G (J.obj (SingleObj.star G))) x y := by + ↔ MulAction.orbitRel G (J.obj (SingleObj.star G)) x y := by have h (g : G) : y = g • x ↔ g • x = y := ⟨symm, symm⟩ conv => rhs; rw [Setoid.comm'] show (∃ g : G, y = g • x) ↔ (∃ g : G, g • x = y) diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index a11fa73582072..2bec860ba95b9 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -26,6 +26,8 @@ run_cmd Lean.Elab.Command.liftTermElabM do Lean.Meta.registerCoercion ``Setoid.r (some { numArgs := 2, coercee := 1, type := .coeFun }) +/-- When writing a lemma about `someSetoid x y` (which uses this instance), +call it `someSetoid_apply` not `someSetoid_r`. -/ instance : CoeFun (Setoid α) (fun _ ↦ α → α → Prop) where coe := @Setoid.r _ diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index c34439b36e1e1..e4b60325f1b21 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -15,9 +15,6 @@ theorems for quotients of arbitrary types. ## Implementation notes -The function `Rel` and lemmas ending in ' make it easier to talk about different -equivalence relations on the same type. - The complete lattice instance for equivalence relations could have been defined by lifting the Galois insertion of equivalence relations on α into binary relations on α, and then using `CompleteLattice.copy` to define a complete lattice instance with more appropriate @@ -39,47 +36,57 @@ attribute [trans] Setoid.trans variable {α : Type*} {β : Type*} /-- A version of `Setoid.r` that takes the equivalence relation as an explicit argument. -/ +@[deprecated (since := "2024-08-29")] def Setoid.Rel (r : Setoid α) : α → α → Prop := @Setoid.r _ r +set_option linter.deprecated false in +@[deprecated (since := "2024-10-09")] instance Setoid.decidableRel (r : Setoid α) [h : DecidableRel r.r] : DecidableRel r.Rel := h +set_option linter.deprecated false in /-- A version of `Quotient.eq'` compatible with `Setoid.Rel`, to make rewriting possible. -/ +@[deprecated Quotient.eq' (since := "2024-10-09")] theorem Quotient.eq_rel {r : Setoid α} {x y} : (Quotient.mk' x : Quotient r) = Quotient.mk' y ↔ r.Rel x y := Quotient.eq namespace Setoid -@[ext] +attribute [ext] ext + +set_option linter.deprecated false in +@[deprecated Setoid.ext (since := "2024-10-09")] theorem ext' {r s : Setoid α} (H : ∀ a b, r.Rel a b ↔ s.Rel a b) : r = s := ext H -theorem ext_iff {r s : Setoid α} : r = s ↔ ∀ a b, r.Rel a b ↔ s.Rel a b := +set_option linter.deprecated false in +@[deprecated Setoid.ext_iff (since := "2024-10-09")] +theorem ext'_iff {r s : Setoid α} : r = s ↔ ∀ a b, r.Rel a b ↔ s.Rel a b := ⟨fun h _ _ => h ▸ Iff.rfl, ext'⟩ /-- Two equivalence relations are equal iff their underlying binary operations are equal. -/ -theorem eq_iff_rel_eq {r₁ r₂ : Setoid α} : r₁ = r₂ ↔ r₁.Rel = r₂.Rel := - ⟨fun h => h ▸ rfl, fun h => Setoid.ext' fun _ _ => h ▸ Iff.rfl⟩ +theorem eq_iff_rel_eq {r₁ r₂ : Setoid α} : r₁ = r₂ ↔ ⇑r₁ = ⇑r₂ := + ⟨fun h => h ▸ rfl, fun h => Setoid.ext fun _ _ => h ▸ Iff.rfl⟩ /-- Defining `≤` for equivalence relations. -/ instance : LE (Setoid α) := - ⟨fun r s => ∀ ⦃x y⦄, r.Rel x y → s.Rel x y⟩ + ⟨fun r s => ∀ ⦃x y⦄, r x y → s x y⟩ -theorem le_def {r s : Setoid α} : r ≤ s ↔ ∀ {x y}, r.Rel x y → s.Rel x y := +theorem le_def {r s : Setoid α} : r ≤ s ↔ ∀ {x y}, r x y → s x y := Iff.rfl @[refl] -theorem refl' (r : Setoid α) (x) : r.Rel x x := r.iseqv.refl x +theorem refl' (r : Setoid α) (x) : r x x := r.iseqv.refl x @[symm] -theorem symm' (r : Setoid α) : ∀ {x y}, r.Rel x y → r.Rel y x := r.iseqv.symm +theorem symm' (r : Setoid α) : ∀ {x y}, r x y → r y x := r.iseqv.symm @[trans] -theorem trans' (r : Setoid α) : ∀ {x y z}, r.Rel x y → r.Rel y z → r.Rel x z := r.iseqv.trans +theorem trans' (r : Setoid α) : ∀ {x y z}, r x y → r y z → r x z := r.iseqv.trans -theorem comm' (s : Setoid α) {x y} : s.Rel x y ↔ s.Rel y x := +theorem comm' (s : Setoid α) {x y} : s x y ↔ s y x := ⟨s.symm', s.symm'⟩ /-- The kernel of a function is an equivalence relation. -/ @@ -89,7 +96,7 @@ def ker (f : α → β) : Setoid α := /-- The kernel of the quotient map induced by an equivalence relation r equals r. -/ @[simp] theorem ker_mk_eq (r : Setoid α) : ker (@Quotient.mk'' _ r) = r := - ext' fun _ _ => Quotient.eq + ext fun _ _ => Quotient.eq theorem ker_apply_mk_out {f : α → β} (a : α) : f (⟦a⟧ : Quotient (Setoid.ker f)).out = f a := @Quotient.mk_out _ (Setoid.ker f) a @@ -98,7 +105,7 @@ theorem ker_apply_mk_out' {f : α → β} (a : α) : f (Quotient.mk _ a : Quotient <| Setoid.ker f).out' = f a := @Quotient.mk_out' _ (Setoid.ker f) a -theorem ker_def {f : α → β} {x y : α} : (ker f).Rel x y ↔ f x = f y := +theorem ker_def {f : α → β} {x y : α} : ker f x y ↔ f x = f y := Iff.rfl /-- Given types `α`, `β`, the product of two equivalence relations `r` on `α` and `s` on `β`: @@ -106,7 +113,7 @@ theorem ker_def {f : α → β} {x y : α} : (ker f).Rel x y ↔ f x = f y := by `r` and `x₂` is related to `y₂` by `s`. -/ protected def prod (r : Setoid α) (s : Setoid β) : Setoid (α × β) where - r x y := r.Rel x.1 y.1 ∧ s.Rel x.2 y.2 + r x y := r x.1 y.1 ∧ s x.2 y.2 iseqv := ⟨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⟩⟩ @@ -157,28 +164,28 @@ noncomputable def piQuotientEquiv {ι : Sort*} {α : ι → Sort*} (r : ∀ i, S /-- The infimum of two equivalence relations. -/ instance : Inf (Setoid α) := ⟨fun r s => - ⟨fun x y => r.Rel x y ∧ s.Rel x y, + ⟨fun x y => r x y ∧ s x y, ⟨fun x => ⟨r.refl' x, s.refl' x⟩, fun h => ⟨r.symm' h.1, s.symm' h.2⟩, fun h1 h2 => ⟨r.trans' h1.1 h2.1, s.trans' h1.2 h2.2⟩⟩⟩⟩ /-- The infimum of 2 equivalence relations r and s is the same relation as the infimum of the underlying binary operations. -/ -theorem inf_def {r s : Setoid α} : (r ⊓ s).Rel = r.Rel ⊓ s.Rel := +theorem inf_def {r s : Setoid α} : ⇑(r ⊓ s) = ⇑r ⊓ ⇑s := rfl -theorem inf_iff_and {r s : Setoid α} {x y} : (r ⊓ s).Rel x y ↔ r.Rel x y ∧ s.Rel x y := +theorem inf_iff_and {r s : Setoid α} {x y} : (r ⊓ s) x y ↔ r x y ∧ s x y := Iff.rfl /-- The infimum of a set of equivalence relations. -/ instance : InfSet (Setoid α) := ⟨fun S => - { r := fun x y => ∀ r ∈ S, r.Rel x y + { r := fun x y => ∀ r ∈ S, r x y iseqv := ⟨fun x r _ => r.refl' x, fun h r hr => r.symm' <| h r hr, fun h1 h2 r hr => r.trans' (h1 r hr) <| h2 r hr⟩ }⟩ /-- The underlying binary operation of the infimum of a set of equivalence relations is the infimum of the set's image under the map to the underlying binary operation. -/ -theorem sInf_def {s : Set (Setoid α)} : (sInf s).Rel = sInf (Rel '' s) := by +theorem sInf_def {s : Set (Setoid α)} : ⇑(sInf s) = sInf ((⇑) '' s) := by ext simp only [sInf_image, iInf_apply, iInf_Prop_eq] rfl @@ -189,7 +196,7 @@ instance : PartialOrder (Setoid α) where le_refl _ _ _ := id le_trans _ _ _ hr hs _ _ h := hs <| hr h lt_iff_le_not_le _ _ := Iff.rfl - le_antisymm _ _ h1 h2 := Setoid.ext' fun _ _ => ⟨fun h => h1 h, fun h => h2 h⟩ + le_antisymm _ _ h1 h2 := Setoid.ext fun _ _ => ⟨fun h => h1 h, fun h => h2 h⟩ /-- The complete lattice of equivalence relations on a type, with bottom element `=` and top element the trivial equivalence relation. -/ @@ -206,20 +213,20 @@ instance completeLattice : CompleteLattice (Setoid α) := bot_le := fun r x _ h => h ▸ r.2.1 x } @[simp] -theorem top_def : (⊤ : Setoid α).Rel = ⊤ := +theorem top_def : ⇑(⊤ : Setoid α) = ⊤ := rfl @[simp] -theorem bot_def : (⊥ : Setoid α).Rel = (· = ·) := +theorem bot_def : ⇑(⊥ : Setoid α) = (· = ·) := rfl -theorem eq_top_iff {s : Setoid α} : s = (⊤ : Setoid α) ↔ ∀ x y : α, s.Rel x y := by +theorem eq_top_iff {s : Setoid α} : s = (⊤ : Setoid α) ↔ ∀ x y : α, s x y := by rw [_root_.eq_top_iff, Setoid.le_def, Setoid.top_def] simp only [Pi.top_apply, Prop.top_eq_true, forall_true_left] lemma sInf_equiv {S : Set (Setoid α)} {x y : α} : letI := sInf S - x ≈ y ↔ ∀ s ∈ S, s.Rel x y := Iff.rfl + x ≈ y ↔ ∀ s ∈ S, s x y := Iff.rfl lemma sInf_iff {S : Set (Setoid α)} {x y : α} : sInf S x y ↔ ∀ s ∈ S, s x y := Iff.rfl @@ -245,7 +252,7 @@ open Relation /-- The inductively defined equivalence closure of a binary relation r is the infimum of the set of all equivalence relations containing r. -/ theorem eqvGen_eq (r : α → α → Prop) : - EqvGen.setoid r = sInf { s : Setoid α | ∀ ⦃x y⦄, r x y → s.Rel x y } := + EqvGen.setoid r = sInf { s : Setoid α | ∀ ⦃x y⦄, r x y → s x y } := le_antisymm (fun _ _ H => EqvGen.rec (fun _ _ h _ hs => hs h) (refl' _) (fun _ _ _ => symm' _) @@ -255,20 +262,20 @@ theorem eqvGen_eq (r : α → α → Prop) : /-- The supremum of two equivalence relations r and s is the equivalence closure of the binary relation `x is related to y by r or s`. -/ theorem sup_eq_eqvGen (r s : Setoid α) : - r ⊔ s = EqvGen.setoid fun x y => r.Rel x y ∨ s.Rel x y := by + r ⊔ s = EqvGen.setoid fun x y => r x y ∨ s x y := by rw [eqvGen_eq] apply congr_arg sInf simp only [le_def, or_imp, ← forall_and] /-- The supremum of 2 equivalence relations r and s is the equivalence closure of the supremum of the underlying binary operations. -/ -theorem sup_def {r s : Setoid α} : r ⊔ s = EqvGen.setoid (r.Rel ⊔ s.Rel) := by +theorem sup_def {r s : Setoid α} : r ⊔ s = EqvGen.setoid (⇑r ⊔ ⇑s) := by rw [sup_eq_eqvGen]; rfl /-- The supremum of a set S of equivalence relations is the equivalence closure of the binary relation `there exists r ∈ S relating x and y`. -/ theorem sSup_eq_eqvGen (S : Set (Setoid α)) : - sSup S = EqvGen.setoid fun x y => ∃ r : Setoid α, r ∈ S ∧ r.Rel x y := by + sSup S = EqvGen.setoid fun x y => ∃ r : Setoid α, r ∈ S ∧ r x y := by rw [eqvGen_eq] apply congr_arg sInf simp only [upperBounds, le_def, and_imp, exists_imp] @@ -277,7 +284,7 @@ theorem sSup_eq_eqvGen (S : Set (Setoid α)) : /-- The supremum of a set of equivalence relations is the equivalence closure of the supremum of the set's image under the map to the underlying binary operation. -/ -theorem sSup_def {s : Set (Setoid α)} : sSup s = EqvGen.setoid (sSup (Rel '' s)) := by +theorem sSup_def {s : Set (Setoid α)} : sSup s = EqvGen.setoid (sSup ((⇑) '' s)) := by rw [sSup_eq_eqvGen, sSup_image] congr with (x y) simp only [iSup_apply, iSup_Prop_eq, exists_prop] @@ -288,13 +295,12 @@ theorem eqvGen_of_setoid (r : Setoid α) : EqvGen.setoid r.r = r := le_antisymm (by rw [eqvGen_eq]; exact sInf_le fun _ _ => id) EqvGen.rel /-- Equivalence closure is idempotent. -/ -@[simp] -theorem eqvGen_idem (r : α → α → Prop) : EqvGen.setoid (EqvGen.setoid r).Rel = EqvGen.setoid r := +theorem eqvGen_idem (r : α → α → Prop) : EqvGen.setoid (EqvGen.setoid r) = EqvGen.setoid r := eqvGen_of_setoid _ /-- The equivalence closure of a binary relation r is contained in any equivalence relation containing r. -/ -theorem eqvGen_le {r : α → α → Prop} {s : Setoid α} (h : ∀ x y, r x y → s.Rel x y) : +theorem eqvGen_le {r : α → α → Prop} {s : Setoid α} (h : ∀ x y, r x y → s x y) : EqvGen.setoid r ≤ s := by rw [eqvGen_eq]; exact sInf_le h /-- Equivalence closure of binary relations is monotone. -/ @@ -304,7 +310,7 @@ theorem eqvGen_mono {r s : α → α → Prop} (h : ∀ x y, r x y → s x y) : /-- There is a Galois insertion of equivalence relations on α into binary relations on α, with equivalence closure the lower adjoint. -/ -def gi : @GaloisInsertion (α → α → Prop) (Setoid α) _ _ EqvGen.setoid Rel where +def gi : @GaloisInsertion (α → α → Prop) (Setoid α) _ _ EqvGen.setoid (⇑) where choice r _ := EqvGen.setoid r gc _ s := ⟨fun H _ _ h => H <| EqvGen.rel _ _ h, fun H => eqvGen_of_setoid s ▸ eqvGen_mono H⟩ le_l_u x := (eqvGen_of_setoid x).symm ▸ le_refl x @@ -320,7 +326,7 @@ theorem injective_iff_ker_bot (f : α → β) : Injective f ↔ ker f = ⊥ := (@eq_bot_iff (Setoid α) _ _ (ker f)).symm /-- The elements related to x ∈ α by the kernel of f are those in the preimage of f(x) under f. -/ -theorem ker_iff_mem_preimage {f : α → β} {x y} : (ker f).Rel x y ↔ x ∈ f ⁻¹' {f y} := +theorem ker_iff_mem_preimage {f : α → β} {x y} : ker f x y ↔ x ∈ f ⁻¹' {f y} := Iff.rfl /-- Equivalence between functions `α → β` such that `r x y → f x = f y` and functions @@ -345,7 +351,7 @@ theorem ker_lift_injective (f : α → β) : Injective (@Quotient.lift _ _ (ker /-- Given a map f from α to β, the kernel of f is the unique equivalence relation on α whose induced map from the quotient of α to β is injective. -/ -theorem ker_eq_lift_of_injective {r : Setoid α} (f : α → β) (H : ∀ x y, r.Rel x y → f x = f y) +theorem ker_eq_lift_of_injective {r : Setoid α} (f : α → β) (H : ∀ x y, r x y → f x = f y) (h : Injective (Quotient.lift f H)) : ker f = r := le_antisymm (fun x y hk => @@ -387,13 +393,13 @@ variable {r f} closure of the relation on `f`'s image defined by '`x ≈ y` iff the elements of `f⁻¹(x)` are related to the elements of `f⁻¹(y)` by `r`.' -/ def map (r : Setoid α) (f : α → β) : Setoid β := - Relation.EqvGen.setoid fun x y => ∃ a b, f a = x ∧ f b = y ∧ r.Rel a b + Relation.EqvGen.setoid fun x y => ∃ a b, f a = x ∧ f b = y ∧ r a b /-- Given a surjective function f whose kernel is contained in an equivalence relation r, the equivalence relation on f's codomain defined by x ≈ y ↔ the elements of f⁻¹(x) are related to the elements of f⁻¹(y) by r. -/ def mapOfSurjective (r) (f : α → β) (h : ker f ≤ r) (hf : Surjective f) : Setoid β := - ⟨fun x y => ∃ a b, f a = x ∧ f b = y ∧ r.Rel a b, + ⟨fun x y => ∃ a b, f a = x ∧ f b = y ∧ r a b, ⟨fun x => let ⟨y, hy⟩ := hf x ⟨y, y, hy, hy, r.refl' y⟩, @@ -411,9 +417,9 @@ relation on `α` defined by '`x ≈ y` iff `f(x)` is related to `f(y)` by `r`'. See note [reducible non-instances]. -/ abbrev comap (f : α → β) (r : Setoid β) : Setoid α := - ⟨r.Rel on f, r.iseqv.comap _⟩ + ⟨r on f, r.iseqv.comap _⟩ -theorem comap_rel (f : α → β) (r : Setoid β) (x y : α) : (comap f r).Rel x y ↔ r.Rel (f x) (f y) := +theorem comap_rel (f : α → β) (r : Setoid β) (x y : α) : comap f r x y ↔ r (f x) (f y) := Iff.rfl /-- Given a map `f : N → M` and an equivalence relation `r` on `β`, the equivalence relation @@ -424,7 +430,8 @@ theorem comap_eq {f : α → β} {r : Setoid β} : comap f r = ker (@Quotient.mk /-- The second isomorphism theorem for sets. -/ noncomputable def comapQuotientEquiv (f : α → β) (r : Setoid β) : Quotient (comap f r) ≃ Set.range (@Quotient.mk'' _ r ∘ f) := - (Quotient.congrRight <| ext_iff.1 comap_eq).trans <| quotientKerEquivRange <| Quotient.mk'' ∘ f + (Quotient.congrRight <| Setoid.ext_iff.1 comap_eq).trans <| quotientKerEquivRange <| + Quotient.mk'' ∘ f variable (r f) @@ -454,7 +461,7 @@ def correspondence (r : Setoid α) : { s // r ≤ s } ≃o Setoid (Quotient r) w (fun h ↦ s.1.trans' (s.1.trans' (s.2 h₁) h) (s.1.symm' (s.2 h₂))), ⟨Quotient.ind s.1.2.1, @fun x y ↦ Quotient.inductionOn₂ x y fun _ _ ↦ s.1.2.2, @fun x y z ↦ Quotient.inductionOn₃ x y z fun _ _ _ ↦ s.1.2.3⟩⟩ - invFun s := ⟨comap Quotient.mk' s, fun x y h => by rw [comap_rel, eq_rel.2 h]⟩ + invFun s := ⟨comap Quotient.mk' s, fun x y h => by rw [comap_rel, Quotient.eq'.2 h]⟩ left_inv _ := rfl right_inv _ := ext fun x y ↦ Quotient.inductionOn₂ x y fun _ _ ↦ Iff.rfl map_rel_iff' := @@ -478,7 +485,6 @@ theorem Quotient.subsingleton_iff {s : Setoid α} : Subsingleton (Quotient s) refine (surjective_quotient_mk' _).forall.trans (forall_congr' fun a => ?_) refine (surjective_quotient_mk' _).forall.trans (forall_congr' fun b => ?_) simp_rw [Prop.top_eq_true, true_implies, Quotient.eq'] - rfl theorem Quot.subsingleton_iff (r : α → α → Prop) : Subsingleton (Quot r) ↔ Relation.EqvGen r = ⊤ := by diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index b44853cfa153b..27e668bcb2ece 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -53,9 +53,9 @@ def mkClasses (c : Set (Set α)) (H : ∀ a, ∃! b ∈ c, a ∈ b) : Setoid α /-- Makes the equivalence classes of an equivalence relation. -/ def classes (r : Setoid α) : Set (Set α) := - { s | ∃ y, s = { x | r.Rel x y } } + { s | ∃ y, s = { x | r x y } } -theorem mem_classes (r : Setoid α) (y) : { x | r.Rel x y } ∈ r.classes := +theorem mem_classes (r : Setoid α) (y) : { x | r x y } ∈ r.classes := ⟨y, rfl⟩ theorem classes_ker_subset_fiber_set {β : Type*} (f : α → β) : @@ -74,17 +74,17 @@ theorem card_classes_ker_le {α β : Type*} [Fintype β] (f : α → β) /-- Two equivalence relations are equal iff all their equivalence classes are equal. -/ theorem eq_iff_classes_eq {r₁ r₂ : Setoid α} : - r₁ = r₂ ↔ ∀ x, { y | r₁.Rel x y } = { y | r₂.Rel x y } := - ⟨fun h _x => h ▸ rfl, fun h => ext' fun x => Set.ext_iff.1 <| h x⟩ + r₁ = r₂ ↔ ∀ x, { y | r₁ x y } = { y | r₂ x y } := + ⟨fun h _x => h ▸ rfl, fun h => ext fun x => Set.ext_iff.1 <| h x⟩ -theorem rel_iff_exists_classes (r : Setoid α) {x y} : r.Rel x y ↔ ∃ c ∈ r.classes, x ∈ c ∧ y ∈ c := +theorem rel_iff_exists_classes (r : Setoid α) {x y} : r x y ↔ ∃ c ∈ r.classes, x ∈ c ∧ y ∈ c := ⟨fun h => ⟨_, r.mem_classes y, h, r.refl' y⟩, fun ⟨c, ⟨z, hz⟩, hx, hy⟩ => by subst c exact r.trans' hx (r.symm' hy)⟩ /-- Two equivalence relations are equal iff their equivalence classes are equal. -/ theorem classes_inj {r₁ r₂ : Setoid α} : r₁ = r₂ ↔ r₁.classes = r₂.classes := - ⟨fun h => h ▸ rfl, fun h => ext' fun a b => by simp only [rel_iff_exists_classes, exists_prop, h]⟩ + ⟨fun h => h ▸ rfl, fun h => ext fun a b => by simp only [rel_iff_exists_classes, exists_prop, h]⟩ /-- The empty set is not an equivalence class. -/ theorem empty_not_mem_classes {r : Setoid α} : ∅ ∉ r.classes := fun ⟨y, hy⟩ => @@ -92,7 +92,7 @@ theorem empty_not_mem_classes {r : Setoid α} : ∅ ∉ r.classes := fun ⟨y, h /-- Equivalence classes partition the type. -/ theorem classes_eqv_classes {r : Setoid α} (a) : ∃! b ∈ r.classes, a ∈ b := - ExistsUnique.intro { x | r.Rel x a } ⟨r.mem_classes a, r.refl' _⟩ <| by + ExistsUnique.intro { x | r x a } ⟨r.mem_classes a, r.refl' _⟩ <| by rintro y ⟨⟨_, rfl⟩, ha⟩ ext x exact ⟨fun hx => r.trans' hx (r.symm' ha), fun hx => r.trans' hx ha⟩ @@ -105,7 +105,7 @@ theorem eq_of_mem_classes {r : Setoid α} {x b} (hc : b ∈ r.classes) (hb : x /-- The elements of a set of sets partitioning α are the equivalence classes of the equivalence relation defined by the set of sets. -/ theorem eq_eqv_class_of_mem {c : Set (Set α)} (H : ∀ a, ∃! b ∈ c, a ∈ b) {s y} - (hs : s ∈ c) (hy : y ∈ s) : s = { x | (mkClasses c H).Rel x y } := by + (hs : s ∈ c) (hy : y ∈ s) : s = { x | mkClasses c H x y } := by ext x constructor · intro hx _s' hs' hx' @@ -117,11 +117,11 @@ theorem eq_eqv_class_of_mem {c : Set (Set α)} (H : ∀ a, ∃! b ∈ c, a ∈ b /-- The equivalence classes of the equivalence relation defined by a set of sets partitioning α are elements of the set of sets. -/ theorem eqv_class_mem {c : Set (Set α)} (H : ∀ a, ∃! b ∈ c, a ∈ b) {y} : - { x | (mkClasses c H).Rel x y } ∈ c := + { x | mkClasses c H x y } ∈ c := (H y).elim fun _ hc _ => eq_eqv_class_of_mem H hc.1 hc.2 ▸ hc.1 theorem eqv_class_mem' {c : Set (Set α)} (H : ∀ a, ∃! b ∈ c, a ∈ b) {x} : - { y : α | (mkClasses c H).Rel x y } ∈ c := by + { y : α | mkClasses c H x y } ∈ c := by convert @Setoid.eqv_class_mem _ _ H x using 3 rw [Setoid.comm'] @@ -145,13 +145,13 @@ def setoidOfDisjointUnion {c : Set (Set α)} (hu : Set.sUnion c = @Set.univ α) /-- The equivalence relation made from the equivalence classes of an equivalence relation r equals r. -/ theorem mkClasses_classes (r : Setoid α) : mkClasses r.classes classes_eqv_classes = r := - ext' fun x _y => - ⟨fun h => r.symm' (h { z | r.Rel z x } (r.mem_classes x) <| r.refl' x), fun h _b hb hx => + ext fun x _y => + ⟨fun h => r.symm' (h { z | r z x } (r.mem_classes x) <| r.refl' x), fun h _b hb hx => eq_of_mem_classes (r.mem_classes x) (r.refl' x) hb hx ▸ r.symm' h⟩ @[simp] theorem sUnion_classes (r : Setoid α) : ⋃₀ r.classes = Set.univ := - Set.eq_univ_of_forall fun x => Set.mem_sUnion.2 ⟨{ y | r.Rel y x }, ⟨x, rfl⟩, Setoid.refl _⟩ + Set.eq_univ_of_forall fun x => Set.mem_sUnion.2 ⟨{ y | r y x }, ⟨x, rfl⟩, Setoid.refl _⟩ /-- The equivalence between the quotient by an equivalence relation and its type of equivalence classes. -/ @@ -177,8 +177,8 @@ noncomputable def quotientEquivClasses (r : Setoid α) : Quotient r ≃ Setoid.c @[simp] lemma quotientEquivClasses_mk_eq (r : Setoid α) (a : α) : - (quotientEquivClasses r (Quotient.mk r a) : Set α) = { x | r.Rel x a } := - (@Subtype.ext_iff_val _ _ _ ⟨{ x | r.Rel x a }, Setoid.mem_classes r a⟩).mp rfl + (quotientEquivClasses r (Quotient.mk r a) : Set α) = { x | r x a } := + (@Subtype.ext_iff_val _ _ _ ⟨{ x | r x a }, Setoid.mem_classes r a⟩).mp rfl section Partition @@ -217,7 +217,7 @@ theorem IsPartition.sUnion_eq_univ {c : Set (Set α)} (hc : IsPartition c) : ⋃ /-- All elements of a partition of α are the equivalence class of some y ∈ α. -/ theorem exists_of_mem_partition {c : Set (Set α)} (hc : IsPartition c) {s} (hs : s ∈ c) : - ∃ y, s = { x | (mkClasses c hc.2).Rel x y } := + ∃ y, s = { x | mkClasses c hc.2 x y } := let ⟨y, hy⟩ := nonempty_of_mem_partition hc hs ⟨y, eq_eqv_class_of_mem hc.2 hs hy⟩ @@ -367,7 +367,7 @@ protected abbrev setoid (hs : IndexedPartition s) : Setoid α := theorem index_some (i : ι) : hs.index (hs.some i) = i := (mem_iff_index_eq _).1 <| hs.some_mem i -theorem some_index (x : α) : hs.setoid.Rel (hs.some (hs.index x)) x := +theorem some_index (x : α) : hs.setoid (hs.some (hs.index x)) x := hs.index_some (hs.index x) /-- The quotient associated to an indexed partition. -/ @@ -382,7 +382,7 @@ instance [Inhabited α] : Inhabited hs.Quotient := ⟨hs.proj default⟩ theorem proj_eq_iff {x y : α} : hs.proj x = hs.proj y ↔ hs.index x = hs.index y := - Quotient.eq_rel + Quotient.eq'' @[simp] theorem proj_some_index (x : α) : hs.proj (hs.some (hs.index x)) = hs.proj x := @@ -423,7 +423,7 @@ theorem index_out' (x : hs.Quotient) : hs.index x.out' = hs.index (hs.out x) := theorem proj_out (x : hs.Quotient) : hs.proj (hs.out x) = x := Quotient.inductionOn' x fun x => Quotient.sound' <| hs.some_index x -theorem class_of {x : α} : setOf (hs.setoid.Rel x) = s (hs.index x) := +theorem class_of {x : α} : setOf (hs.setoid x) = s (hs.index x) := Set.ext fun _y => eq_comm.trans hs.mem_iff_index_eq.symm theorem proj_fiber (x : hs.Quotient) : hs.proj ⁻¹' {x} = s (hs.equivQuotient.symm x) := diff --git a/Mathlib/GroupTheory/Congruence/Defs.lean b/Mathlib/GroupTheory/Congruence/Defs.lean index ced9c0f1660d1..118baa87520bd 100644 --- a/Mathlib/GroupTheory/Congruence/Defs.lean +++ b/Mathlib/GroupTheory/Congruence/Defs.lean @@ -166,7 +166,7 @@ theorem ext {c d : Con M} (H : ∀ x y, c x y ↔ d x y) : c = d := @[to_additive "The map sending an additive congruence relation to its underlying equivalence relation is injective."] theorem toSetoid_inj {c d : Con M} (H : c.toSetoid = d.toSetoid) : c = d := - ext <| ext_iff.1 H + ext <| Setoid.ext_iff.1 H /-- Two congruence relations are equal iff their underlying binary relations are equal. -/ @[to_additive "Two additive congruence relations are equal iff their underlying binary relations @@ -334,7 +334,7 @@ instance : InfSet (Con M) where @[to_additive "The infimum of a set of additive congruence relations is the same as the infimum of the set's image under the map to the underlying equivalence relation."] theorem sInf_toSetoid (S : Set (Con M)) : (sInf S).toSetoid = sInf (toSetoid '' S) := - Setoid.ext' fun x y => + Setoid.ext fun x y => ⟨fun h r ⟨c, hS, hr⟩ => by rw [← hr]; exact h c hS, fun h c hS => h c.toSetoid ⟨c, hS, rfl⟩⟩ /-- The infimum of a set of congruence relations is the same as the infimum of the set's image @@ -572,8 +572,8 @@ def correspondence : { d // c ≤ d } ≃o Con c.Quotient where constructor · intros h x y hs rcases h ⟨x, y, rfl, rfl, hs⟩ with ⟨a, b, hx, hy, ht⟩ - exact t.1.trans (t.1.symm <| t.2 <| Quotient.eq_rel.1 hx) - (t.1.trans ht (t.2 <| Quotient.eq_rel.1 hy)) + exact t.1.trans (t.1.symm <| t.2 <| Quotient.eq'.1 hx) + (t.1.trans ht (t.2 <| Quotient.eq'.1 hy)) · intros h _ _ hs rcases hs with ⟨a, b, hx, hy, Hs⟩ exact ⟨a, b, hx, hy, h Hs⟩ diff --git a/Mathlib/GroupTheory/DoubleCoset.lean b/Mathlib/GroupTheory/DoubleCoset.lean index a7d7e29638510..d80688cff9e27 100644 --- a/Mathlib/GroupTheory/DoubleCoset.lean +++ b/Mathlib/GroupTheory/DoubleCoset.lean @@ -72,15 +72,15 @@ def Quotient (H K : Set G) : Type _ := _root_.Quotient (setoid H K) theorem rel_iff {H K : Subgroup G} {x y : G} : - (setoid ↑H ↑K).Rel x y ↔ ∃ a ∈ H, ∃ b ∈ K, y = a * x * b := + setoid ↑H ↑K x y ↔ ∃ a ∈ H, ∃ b ∈ K, y = a * x * b := Iff.trans ⟨fun (hxy : doset x H K = doset y H K) => hxy ▸ mem_doset_self H K y, fun hxy => (doset_eq_of_mem hxy).symm⟩ mem_doset theorem bot_rel_eq_leftRel (H : Subgroup G) : - (setoid ↑(⊥ : Subgroup G) ↑H).Rel = (QuotientGroup.leftRel H).Rel := by + ⇑(setoid ↑(⊥ : Subgroup G) ↑H) = ⇑(QuotientGroup.leftRel H) := by ext a b - rw [rel_iff, Setoid.Rel, QuotientGroup.leftRel_apply] + rw [rel_iff, QuotientGroup.leftRel_apply] constructor · rintro ⟨a, rfl : a = 1, b, hb, rfl⟩ change a⁻¹ * (1 * a * b) ∈ H @@ -89,9 +89,9 @@ theorem bot_rel_eq_leftRel (H : Subgroup G) : exact ⟨1, rfl, a⁻¹ * b, h, by rw [one_mul, mul_inv_cancel_left]⟩ theorem rel_bot_eq_right_group_rel (H : Subgroup G) : - (setoid ↑H ↑(⊥ : Subgroup G)).Rel = (QuotientGroup.rightRel H).Rel := by + ⇑(setoid ↑H ↑(⊥ : Subgroup G)) = ⇑(QuotientGroup.rightRel H) := by ext a b - rw [rel_iff, Setoid.Rel, QuotientGroup.rightRel_apply] + rw [rel_iff, QuotientGroup.rightRel_apply] constructor · rintro ⟨b, hb, a, rfl : a = 1, rfl⟩ change b * a * 1 * a⁻¹ ∈ H diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index 9f99144a7769a..547d1fc8f93c9 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -373,12 +373,11 @@ def orbitRel : Setoid α where variable {G α} @[to_additive] -theorem orbitRel_apply {a b : α} : (orbitRel G α).Rel a b ↔ a ∈ orbit G b := +theorem orbitRel_apply {a b : α} : orbitRel G α a b ↔ a ∈ orbit G b := Iff.rfl -@[to_additive] -lemma orbitRel_r_apply {a b : α} : (orbitRel G _).r a b ↔ a ∈ orbit G b := - Iff.rfl +@[to_additive (attr := deprecated (since := "2024-10-18"))] +alias orbitRel_r_apply := orbitRel_apply @[to_additive] lemma orbitRel_subgroup_le (H : Subgroup G) : orbitRel H α ≤ orbitRel G α := @@ -468,7 +467,7 @@ theorem pretransitive_iff_subsingleton_quotient : · refine Quot.inductionOn a (fun x ↦ ?_) exact Quot.inductionOn b (fun y ↦ Quot.sound <| exists_smul_eq G y x) · have h : Quotient.mk (orbitRel G α) b = ⟦a⟧ := Subsingleton.elim _ _ - exact Quotient.eq_rel.mp h + exact Quotient.eq''.mp h /-- If `α` is non-empty, an action is pretransitive if and only if the quotient has exactly one element. -/ @@ -511,7 +510,7 @@ lemma orbitRel.Quotient.orbit_injective : Injective (orbitRel.Quotient.orbit : orbitRel.Quotient G α → Set α) := by intro x y h simp_rw [orbitRel.Quotient.orbit_eq_orbit_out _ Quotient.out_eq', orbit_eq_iff, - ← orbitRel_r_apply] at h + ← orbitRel_apply] at h simpa [← Quotient.eq''] using h @[to_additive (attr := simp)] @@ -596,7 +595,7 @@ lemma orbitRel.Quotient.mem_subgroup_orbit_iff' {H : Subgroup G} {x : orbitRel.Q at hb rw [orbitRel.Quotient.mem_subgroup_orbit_iff] convert hb using 1 - rw [orbit_eq_iff, ← orbitRel_r_apply, ← Quotient.eq'', Quotient.out_eq', @Quotient.mk''_eq_mk] + rw [orbit_eq_iff, ← orbitRel_apply, ← Quotient.eq'', Quotient.out_eq', @Quotient.mk''_eq_mk] rw [orbitRel.Quotient.mem_orbit, h, @Quotient.mk''_eq_mk] variable (G) (α) @@ -725,7 +724,7 @@ theorem stabilizer_smul_eq_stabilizer_map_conj (g : G) (a : α) : inv_mul_cancel, one_smul, ← mem_stabilizer_iff, Subgroup.mem_map_equiv, MulAut.conj_symm_apply] /-- A bijection between the stabilizers of two elements in the same orbit. -/ -noncomputable def stabilizerEquivStabilizerOfOrbitRel {a b : α} (h : (orbitRel G α).Rel a b) : +noncomputable def stabilizerEquivStabilizerOfOrbitRel {a b : α} (h : orbitRel G α a b) : stabilizer G a ≃* stabilizer G b := let g : G := Classical.choose h have hg : g • b = a := Classical.choose_spec h @@ -749,7 +748,7 @@ theorem stabilizer_vadd_eq_stabilizer_map_conj (g : G) (a : α) : AddAut.conj_symm_apply] /-- A bijection between the stabilizers of two elements in the same orbit. -/ -noncomputable def stabilizerEquivStabilizerOfOrbitRel {a b : α} (h : (orbitRel G α).Rel a b) : +noncomputable def stabilizerEquivStabilizerOfOrbitRel {a b : α} (h : orbitRel G α a b) : stabilizer G a ≃+ stabilizer G b := let g : G := Classical.choose h have hg : g +ᵥ b = a := Classical.choose_spec h diff --git a/Mathlib/GroupTheory/GroupAction/ConjAct.lean b/Mathlib/GroupTheory/GroupAction/ConjAct.lean index 1f5c561c568ef..866e1f3b36304 100644 --- a/Mathlib/GroupTheory/GroupAction/ConjAct.lean +++ b/Mathlib/GroupTheory/GroupAction/ConjAct.lean @@ -246,7 +246,7 @@ theorem fixedPoints_eq_center : fixedPoints (ConjAct G) G = center G := by theorem mem_orbit_conjAct {g h : G} : g ∈ orbit (ConjAct G) h ↔ IsConj g h := by rw [isConj_comm, isConj_iff, mem_orbit_iff]; rfl -theorem orbitRel_conjAct : (orbitRel (ConjAct G) G).Rel = IsConj := +theorem orbitRel_conjAct : ⇑(orbitRel (ConjAct G) G) = IsConj := funext₂ fun g h => by rw [orbitRel_apply, mem_orbit_conjAct] theorem orbit_eq_carrier_conjClasses (g : G) : diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 3babc32ec4071..f58416df124ac 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -329,8 +329,7 @@ noncomputable def equivSubgroupOrbitsSetoidComap (H : Subgroup α) (ω : Ω) : have hx := x.property simp only [Set.mem_preimage, Set.mem_singleton_iff] at hx rwa [orbitRel.Quotient.mem_orbit, @Quotient.mk''_eq_mk]⟩⟧) fun a b h ↦ by - change Setoid.Rel _ _ _ at h - rw [Setoid.comap_rel, Setoid.Rel, ← Quotient.eq'', @Quotient.mk''_eq_mk] at h + rw [Setoid.comap_rel, ← Quotient.eq'', @Quotient.mk''_eq_mk] at h simp only [orbitRel.Quotient.subgroup_quotient_eq_iff] exact h left_inv := by @@ -373,7 +372,7 @@ noncomputable def equivSubgroupOrbitsQuotientGroup [IsPretransitive α β] orbitRel.Quotient H β ≃ α ⧸ H where toFun := fun q ↦ q.liftOn' (fun y ↦ (exists_smul_eq α y x).choose) (by intro y₁ y₂ h - rw [orbitRel_r_apply] at h + rw [orbitRel_apply] at h rw [Quotient.eq'', leftRel_eq] dsimp only rcases h with ⟨g, rfl⟩ @@ -389,12 +388,12 @@ noncomputable def equivSubgroupOrbitsQuotientGroup [IsPretransitive α β] intro g₁ g₂ h rw [leftRel_eq] at h simp only - rw [← @Quotient.mk''_eq_mk, Quotient.eq'', orbitRel_r_apply] + rw [← @Quotient.mk''_eq_mk, Quotient.eq'', orbitRel_apply] exact ⟨⟨_, h⟩, by simp [mul_smul]⟩) left_inv := fun y ↦ by induction' y using Quotient.inductionOn' with y simp only [Quotient.liftOn'_mk''] - rw [← @Quotient.mk''_eq_mk, Quotient.eq'', orbitRel_r_apply] + rw [← @Quotient.mk''_eq_mk, Quotient.eq'', orbitRel_apply] convert mem_orbit_self _ rw [inv_smul_eq_iff, (exists_smul_eq α y x).choose_spec] right_inv := fun g ↦ by diff --git a/Mathlib/Order/Partition/Finpartition.lean b/Mathlib/Order/Partition/Finpartition.lean index 2392533c815c4..0c1bbdebaec82 100644 --- a/Mathlib/Order/Partition/Finpartition.lean +++ b/Mathlib/Order/Partition/Finpartition.lean @@ -581,12 +581,12 @@ def ofSetoid (s : Setoid α) [DecidableRel s.r] : Finpartition (univ : Finset α sup_parts := by ext a simp only [sup_image, Function.id_comp, mem_univ, mem_sup, mem_filter, true_and, iff_true] - use a; exact s.refl a + use a not_bot_mem := by rw [bot_eq_empty, mem_image, not_exists] intro a simp only [filter_eq_empty_iff, not_forall, mem_univ, forall_true_left, true_and, not_not] - use a; exact s.refl a + use a theorem mem_part_ofSetoid_iff_rel {s : Setoid α} [DecidableRel s.r] {b : α} : b ∈ (ofSetoid s).part a ↔ s.r a b := by diff --git a/Mathlib/RingTheory/AdicCompletion/Algebra.lean b/Mathlib/RingTheory/AdicCompletion/Algebra.lean index f9ab3648fed20..6f7c6168e2c6c 100644 --- a/Mathlib/RingTheory/AdicCompletion/Algebra.lean +++ b/Mathlib/RingTheory/AdicCompletion/Algebra.lean @@ -197,7 +197,7 @@ theorem smul_mk {m n : ℕ} (hmn : m ≤ n) (r : AdicCauchySequence I R) good definitional behaviour for the module instance on adic completions -/ instance : SMul (R ⧸ (I • ⊤ : Ideal R)) (M ⧸ (I • ⊤ : Submodule R M)) where smul r x := - Quotient.liftOn r (· • x) fun b₁ b₂ (h : Setoid.Rel _ b₁ b₂) ↦ by + Quotient.liftOn r (· • x) fun b₁ b₂ h ↦ by refine Quotient.inductionOn' x (fun x ↦ ?_) have h : b₁ - b₂ ∈ (I : Submodule R R) := by rwa [show I = I • ⊤ by simp, ← Submodule.quotientRel_def] diff --git a/Mathlib/RingTheory/Congruence/Basic.lean b/Mathlib/RingTheory/Congruence/Basic.lean index 131575077c37b..90eea0cd93488 100644 --- a/Mathlib/RingTheory/Congruence/Basic.lean +++ b/Mathlib/RingTheory/Congruence/Basic.lean @@ -74,7 +74,7 @@ instance : FunLike (RingCon R) R (R → Prop) where rcases x with ⟨⟨x, _⟩, _⟩ rcases y with ⟨⟨y, _⟩, _⟩ congr! - rw [Setoid.ext_iff, (show x.Rel = y.Rel from h)] + rw [Setoid.ext_iff, (show ⇑x = ⇑y from h)] simp theorem rel_eq_coe : c.r = c := @@ -445,7 +445,7 @@ instance : InfSet (RingCon R) where /-- The infimum of a set of congruence relations is the same as the infimum of the set's image under the map to the underlying equivalence relation. -/ theorem sInf_toSetoid (S : Set (RingCon R)) : (sInf S).toSetoid = sInf ((·.toSetoid) '' S) := - Setoid.ext' fun x y => + Setoid.ext fun x y => ⟨fun h r ⟨c, hS, hr⟩ => by rw [← hr]; exact h c hS, fun h c hS => h c.toSetoid ⟨c, hS, rfl⟩⟩ /-- The infimum of a set of congruence relations is the same as the infimum of the set's image diff --git a/Mathlib/RingTheory/Valuation/ValuationRing.lean b/Mathlib/RingTheory/Valuation/ValuationRing.lean index e95d7d322ff0d..266cd8e9d5884 100644 --- a/Mathlib/RingTheory/Valuation/ValuationRing.lean +++ b/Mathlib/RingTheory/Valuation/ValuationRing.lean @@ -142,22 +142,22 @@ noncomputable instance linearOrder : LinearOrder (ValueGroup A K) where noncomputable instance linearOrderedCommGroupWithZero : LinearOrderedCommGroupWithZero (ValueGroup A K) := { linearOrder .. with - mul_assoc := by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩; apply Quotient.sound'; rw [mul_assoc]; apply Setoid.refl' - one_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [one_mul]; apply Setoid.refl' - mul_one := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_one]; apply Setoid.refl' - mul_comm := by rintro ⟨a⟩ ⟨b⟩; apply Quotient.sound'; rw [mul_comm]; apply Setoid.refl' + mul_assoc := by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩; apply Quotient.sound'; rw [mul_assoc] + one_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [one_mul] + mul_one := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_one] + mul_comm := by rintro ⟨a⟩ ⟨b⟩; apply Quotient.sound'; rw [mul_comm] mul_le_mul_left := by rintro ⟨a⟩ ⟨b⟩ ⟨c, rfl⟩ ⟨d⟩ use c; simp only [Algebra.smul_def]; ring - zero_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [zero_mul]; apply Setoid.refl' - mul_zero := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_zero]; apply Setoid.refl' + zero_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [zero_mul] + mul_zero := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_zero] zero_le_one := ⟨0, by rw [zero_smul]⟩ exists_pair_ne := by use 0, 1 intro c; obtain ⟨d, hd⟩ := Quotient.exact' c apply_fun fun t => d⁻¹ • t at hd simp only [inv_smul_smul, smul_zero, one_ne_zero] at hd - inv_zero := by apply Quotient.sound'; rw [inv_zero]; apply Setoid.refl' + inv_zero := by apply Quotient.sound'; rw [inv_zero] mul_inv_cancel := by rintro ⟨a⟩ ha apply Quotient.sound' diff --git a/Mathlib/Topology/Algebra/ProperAction/Basic.lean b/Mathlib/Topology/Algebra/ProperAction/Basic.lean index fa7a99510f216..92fc69d56d9f4 100644 --- a/Mathlib/Topology/Algebra/ProperAction/Basic.lean +++ b/Mathlib/Topology/Algebra/ProperAction/Basic.lean @@ -118,7 +118,7 @@ theorem t2Space_quotient_mulAction_of_properSMul [ProperSMul G X] : · ext ⟨x₁, x₂⟩ simp only [mem_preimage, map_apply, mem_diagonal_iff, mem_range, Prod.mk.injEq, Prod.exists, exists_eq_right] - rw [Quotient.eq_rel, MulAction.orbitRel_apply, MulAction.mem_orbit_iff] + rw [Quotient.eq', MulAction.orbitRel_apply, MulAction.mem_orbit_iff] all_goals infer_instance /-- If a T2 group acts properly on a topological space, then this topological space is T2. -/ diff --git a/Mathlib/Topology/DiscreteQuotient.lean b/Mathlib/Topology/DiscreteQuotient.lean index ed3253cad0993..bad892ec52aeb 100644 --- a/Mathlib/Topology/DiscreteQuotient.lean +++ b/Mathlib/Topology/DiscreteQuotient.lean @@ -69,7 +69,7 @@ variable {α X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [Topologic @[ext] -- Porting note: in Lean 4, uses projection to `r` instead of `Setoid`. structure DiscreteQuotient (X : Type*) [TopologicalSpace X] extends Setoid X where /-- For every point `x`, the set `{ y | Rel x y }` is an open set. -/ - protected isOpen_setOf_rel : ∀ x, IsOpen (setOf (toSetoid.Rel x)) + protected isOpen_setOf_rel : ∀ x, IsOpen (setOf (toSetoid x)) namespace DiscreteQuotient @@ -81,13 +81,13 @@ lemma toSetoid_injective : Function.Injective (@toSetoid X _) /-- Construct a discrete quotient from a clopen set. -/ def ofIsClopen {A : Set X} (h : IsClopen A) : DiscreteQuotient X where toSetoid := ⟨fun x y => x ∈ A ↔ y ∈ A, fun _ => Iff.rfl, Iff.symm, Iff.trans⟩ - isOpen_setOf_rel x := by by_cases hx : x ∈ A <;> simp [Setoid.Rel, hx, h.1, h.2, ← compl_setOf] + isOpen_setOf_rel x := by by_cases hx : x ∈ A <;> simp [hx, h.1, h.2, ← compl_setOf] -theorem refl : ∀ x, S.Rel x x := S.refl' +theorem refl : ∀ x, S.toSetoid x x := S.refl' -theorem symm (x y : X) : S.Rel x y → S.Rel y x := S.symm' +theorem symm (x y : X) : S.toSetoid x y → S.toSetoid y x := S.symm' -theorem trans (x y z : X) : S.Rel x y → S.Rel y z → S.Rel x z := S.trans' +theorem trans (x y z : X) : S.toSetoid x y → S.toSetoid y z → S.toSetoid x z := S.trans' /-- The setoid whose quotient yields the discrete quotient. -/ add_decl_doc toSetoid @@ -101,7 +101,7 @@ instance : TopologicalSpace S := /-- The projection from `X` to the given discrete quotient. -/ def proj : X → S := Quotient.mk'' -theorem fiber_eq (x : X) : S.proj ⁻¹' {S.proj x} = setOf (S.Rel x) := +theorem fiber_eq (x : X) : S.proj ⁻¹' {S.proj x} = setOf (S.toSetoid x) := Set.ext fun _ => eq_comm.trans Quotient.eq'' theorem proj_surjective : Function.Surjective S.proj := @@ -130,7 +130,7 @@ theorem isOpen_preimage (A : Set S) : IsOpen (S.proj ⁻¹' A) := theorem isClosed_preimage (A : Set S) : IsClosed (S.proj ⁻¹' A) := (S.isClopen_preimage A).1 -theorem isClopen_setOf_rel (x : X) : IsClopen (setOf (S.Rel x)) := by +theorem isClopen_setOf_rel (x : X) : IsClopen (setOf (S.toSetoid x)) := by rw [← fiber_eq] apply isClopen_preimage @@ -359,7 +359,7 @@ lemma comp_finsetClopens [CompactSpace X] : (Set.image (fun (t : Clopens X) ↦ t.carrier) ∘ Finset.toSet) ∘ finsetClopens X = fun ⟨f, _⟩ ↦ f.classes := by ext d - simp only [Setoid.classes, Setoid.Rel, Set.mem_setOf_eq, Function.comp_apply, + simp only [Setoid.classes, Set.mem_setOf_eq, Function.comp_apply, finsetClopens, Set.coe_toFinset, Set.mem_image, Set.mem_range, exists_exists_eq_and] constructor diff --git a/Mathlib/Topology/MetricSpace/Contracting.lean b/Mathlib/Topology/MetricSpace/Contracting.lean index 2fcec8e182a86..ac3f24a9253fa 100644 --- a/Mathlib/Topology/MetricSpace/Contracting.lean +++ b/Mathlib/Topology/MetricSpace/Contracting.lean @@ -137,7 +137,7 @@ theorem efixedPoint_eq_of_edist_lt_top (hf : ContractingWith K f) {x : α} (hx : efixedPoint f hf x hx = efixedPoint f hf y hy := by refine (hf.eq_or_edist_eq_top_of_fixedPoints ?_ ?_).elim id fun h' ↦ False.elim (ne_of_lt ?_ h') <;> try apply efixedPoint_isFixedPt - change edistLtTopSetoid.Rel _ _ + change edistLtTopSetoid _ _ trans x · apply Setoid.symm' -- Porting note: Originally `symm` exact hf.edist_efixedPoint_lt_top hx @@ -221,7 +221,7 @@ theorem efixedPoint_eq_of_edist_lt_top' (hf : ContractingWith K f) {s : Set α} efixedPoint' f hsc hsf hfs x hxs hx = efixedPoint' f htc htf hft y hyt hy := by refine (hf.eq_or_edist_eq_top_of_fixedPoints ?_ ?_).elim id fun h' ↦ False.elim (ne_of_lt ?_ h') <;> try apply efixedPoint_isFixedPt' - change edistLtTopSetoid.Rel _ _ + change edistLtTopSetoid _ _ trans x · apply Setoid.symm' -- Porting note: Originally `symm` apply edist_efixedPoint_lt_top' diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 1ba06927812eb..d2e132acfed38 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -1632,7 +1632,7 @@ instance : TopologicalSpace (t2Quotient X) := /-- The map from a topological space to its largest T2 quotient. -/ def mk : X → t2Quotient X := Quotient.mk (t2Setoid X) -lemma mk_eq {x y : X} : mk x = mk y ↔ ∀ s : Setoid X, T2Space (Quotient s) → s.Rel x y := +lemma mk_eq {x y : X} : mk x = mk y ↔ ∀ s : Setoid X, T2Space (Quotient s) → s x y := Setoid.quotient_mk_sInf_eq variable (X) From 90c6be64b3540bf42afd057554434af4597f1367 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 19 Oct 2024 06:53:17 +0000 Subject: [PATCH 096/131] =?UTF-8?q?feat(MeasureTheory):=20`=CE=BC.comap=20?= =?UTF-8?q?Prod.swap=20=3D=20=CE=BC.map=20Prod.swap`=20(#17918)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From PFR --- Mathlib/MeasureTheory/Measure/MeasureSpace.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 7a7ab5f8c32c6..3c03fb33601d8 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -1975,6 +1975,13 @@ theorem quasiMeasurePreserving_symm (μ : Measure α) (e : α ≃ᵐ β) : end MeasurableEquiv +namespace MeasureTheory.Measure +variable {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + +lemma comap_swap (μ : Measure (α × β)) : μ.comap Prod.swap = μ.map Prod.swap := + (MeasurableEquiv.prodComm ..).comap_symm + +end MeasureTheory.Measure end set_option linter.style.longFile 2000 From 61509b3d7569a6b81e6686f22dd8db1b56156e09 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 19 Oct 2024 07:30:32 +0000 Subject: [PATCH 097/131] chore(SetTheory/Ordinal/Basic): make `Ordinal.typein` a `PrincipalSeg` (#17599) This allows us to golf most of its API down to one-liners. --- Mathlib/Order/InitialSeg.lean | 7 ++ Mathlib/SetTheory/Ordinal/Basic.lean | 113 ++++++++++++++------------- 2 files changed, 65 insertions(+), 55 deletions(-) diff --git a/Mathlib/Order/InitialSeg.lean b/Mathlib/Order/InitialSeg.lean index e53d78efa0277..b1e211077acf2 100644 --- a/Mathlib/Order/InitialSeg.lean +++ b/Mathlib/Order/InitialSeg.lean @@ -289,6 +289,10 @@ theorem mem_range_of_rel [IsTrans β s] (f : r ≺i s) {a : α} {b : β} (h : s @[deprecated mem_range_of_rel (since := "2024-09-21")] alias init := mem_range_of_rel +theorem surjOn (f : r ≺i s) : Set.SurjOn f Set.univ { b | s b f.top } := by + intro b h + simpa using mem_range_of_rel_top _ h + /-- A principal segment is in particular an initial segment. -/ instance hasCoeInitialSeg [IsTrans β s] : Coe (r ≺i s) (r ≼i s) := ⟨fun f => ⟨f.toRelEmbedding, fun _ _ => f.mem_range_of_rel⟩⟩ @@ -383,6 +387,9 @@ instance [IsWellOrder β s] : Subsingleton (r ≺i s) := cases g have := RelEmbedding.coe_fn_injective ef; congr ⟩ +protected theorem eq [IsWellOrder β s] (f g : r ≺i s) (a) : f a = g a := by + rw [Subsingleton.elim f g] + theorem top_eq [IsWellOrder γ t] (e : r ≃r s) (f : r ≺i t) (g : s ≺i t) : f.top = g.top := by rw [Subsingleton.elim f (PrincipalSeg.equivLT e g)]; rfl diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 2de607c6ac406..bc2a81d9ed4ff 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -139,11 +139,6 @@ instance inhabited : Inhabited Ordinal := instance one : One Ordinal := ⟨type <| @EmptyRelation PUnit⟩ -/-- The order type of an element inside a well order. For the embedding as a principal segment, see -`typein.principalSeg`. -/ -def typein (r : α → α → Prop) [IsWellOrder α r] (a : α) : Ordinal := - type (Subrel r { b | r b a }) - @[simp] theorem type_def' (w : WellOrder) : ⟦w⟧ = type w.r := by cases w @@ -375,69 +370,73 @@ def principalSegToType {α β : Ordinal} (h : α < β) : α.toType by - 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] - (f : r ≼i s) (a : α) : Ordinal.typein s (f a) = Ordinal.typein r a := - Eq.symm <| - Quotient.sound - ⟨RelIso.ofSurjective - (RelEmbedding.codRestrict _ ((Subrel.relEmbedding _ _).trans f) fun ⟨x, h⟩ => by - rw [RelEmbedding.trans_apply]; exact f.toRelEmbedding.map_rel_iff.2 h) - fun ⟨y, h⟩ => by - rcases f.mem_range_of_rel h with ⟨a, rfl⟩ - exact ⟨⟨a, f.toRelEmbedding.map_rel_iff.1 h⟩, - Subtype.eq <| RelEmbedding.trans_apply _ _ _⟩⟩ +theorem typein_top {α β} {r : α → α → Prop} {s : β → β → Prop} + [IsWellOrder α r] [IsWellOrder β s] (f : r ≺i s) : typein s f.top = type r := + f.subrelIso.ordinal_type_eq @[simp] theorem typein_lt_typein (r : α → α → Prop) [IsWellOrder α r] {a b : α} : typein r a < typein r b ↔ r a b := - ⟨fun ⟨f⟩ => by - have : f.top.1 = a := by - let f' := PrincipalSeg.ofElement r a - let g' := f.trans (PrincipalSeg.ofElement r b) - have : g'.top = f'.top := by rw [Subsingleton.elim f' g'] - exact this - rw [← this] - exact f.top.2, fun h => - ⟨PrincipalSeg.codRestrict _ (PrincipalSeg.ofElement r a) (fun x => @trans _ r _ _ _ _ x.2 h) h⟩⟩ + (typein r).map_rel_iff + +theorem mem_range_typein_iff (r : α → α → Prop) [IsWellOrder α r] {o} : + o ∈ Set.range (typein r) ↔ o < type r := + (typein r).mem_range_iff_rel theorem typein_surj (r : α → α → Prop) [IsWellOrder α r] {o} (h : o < type r) : - ∃ a, typein r a = o := - inductionOn o (fun _ _ _ ⟨f⟩ => ⟨f.top, typein_top _⟩) h + o ∈ Set.range (typein r) := + (typein r).mem_range_of_rel_top h + +theorem typein_surjOn (r : α → α → Prop) [IsWellOrder α r] : + Set.SurjOn (typein r) Set.univ (Set.Iio (type r)) := + (typein r).surjOn theorem typein_injective (r : α → α → Prop) [IsWellOrder α r] : Injective (typein r) := - injective_of_increasing r (· < ·) (typein r) (typein_lt_typein r).2 + (typein r).injective -@[simp] theorem typein_inj (r : α → α → Prop) [IsWellOrder α r] {a b} : typein r a = typein r b ↔ a = b := - (typein_injective r).eq_iff - -/-- Principal segment version of the `typein` function, embedding a well order into ordinals as a -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 _ ↦ ⟨fun ⟨a, h⟩ ↦ h ▸ typein_lt_type r a, typein_surj r⟩⟩ - -@[simp] -theorem typein.principalSeg_coe (r : α → α → Prop) [IsWellOrder α r] : - (typein.principalSeg r : α → Ordinal) = typein r := - rfl + (typein r).inj /-! ### Enumerating elements in a well-order with ordinals. -/ @@ -450,16 +449,16 @@ the elements of `α`. -/ @[simps! symm_apply_coe] def enum (r : α → α → Prop) [IsWellOrder α r] : @RelIso { o // o < type r } α (Subrel (· < ·) { o | o < type r }) r := - (typein.principalSeg r).subrelIso + (typein r).subrelIso @[simp] theorem typein_enum (r : α → α → Prop) [IsWellOrder α r] {o} (h : o < type r) : typein r (enum r ⟨o, h⟩) = o := - (typein.principalSeg r).apply_subrelIso _ + (typein r).apply_subrelIso _ theorem enum_type {α β} {r : α → α → Prop} {s : β → β → Prop} [IsWellOrder α r] [IsWellOrder β s] (f : s ≺i r) {h : type s < type r} : enum r ⟨type s, h⟩ = f.top := - (typein.principalSeg r).injective <| (typein_enum _ _).trans (typein_top _).symm + (typein r).injective <| (typein_enum _ _).trans (typein_top _).symm @[simp] theorem enum_typein (r : α → α → Prop) [IsWellOrder α r] (a : α) : @@ -502,6 +501,10 @@ theorem induction {p : Ordinal.{u} → Prop} (i : Ordinal.{u}) (h : ∀ j, (∀ p i := lt_wf.induction i h +theorem typein_apply {α β} {r : α → α → Prop} {s : β → β → Prop} [IsWellOrder α r] [IsWellOrder β s] + (f : r ≼i s) (a : α) : typein s (f a) = typein r a := by + rw [← f.leLT_apply _ a, (f.leLT _).eq] + /-! ### Cardinality of ordinals -/ @@ -1031,7 +1034,7 @@ def liftPrincipalSeg : Ordinal.{u} Date: Sat, 19 Oct 2024 07:30:33 +0000 Subject: [PATCH 098/131] chore(Algebra/Module): split off `NoZeroSMulDivisors` (#17909) The definition of `NoZeroSMulDivisors` doesn't require us to know about modules. By splitting off this definition, we can streamline the downstream import graph a bit. In the process, generalize a few instances from `Module` to `SMulWithZero`. New files: * `NoZeroSMulDivisors/Defs.lean`: definition and basic results that don't require `Module` * `NoZeroSMulDivisors/Basic.lean`: basic results that require `Module` * `NoZeroSMulDivisors/Pi.lean`: results on `NoZeroSMulDivisors _ (_ -> _)` that don't require `Module` * `NoZeroSMulDivisors/Prod.lean`: results on `NoZeroSMulDivisors _ (_ x _)` that don't require `Module` --- Mathlib.lean | 4 + Mathlib/Algebra/Algebra/Basic.lean | 1 + Mathlib/Algebra/BigOperators/Finprod.lean | 2 +- Mathlib/Algebra/Group/EvenFunction.lean | 4 +- Mathlib/Algebra/ModEq.lean | 1 + Mathlib/Algebra/Module/Basic.lean | 25 +-- Mathlib/Algebra/Module/Card.lean | 2 +- Mathlib/Algebra/Module/Defs.lean | 176 ------------------ Mathlib/Algebra/Module/LinearMap/Basic.lean | 2 +- Mathlib/Algebra/Module/LinearMap/End.lean | 1 + Mathlib/Algebra/Module/Pi.lean | 13 -- Mathlib/Algebra/Module/Prod.lean | 13 -- Mathlib/Algebra/Module/Rat.lean | 1 + Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean | 155 +++++++++++++++ Mathlib/Algebra/NoZeroSMulDivisors/Defs.lean | 78 ++++++++ Mathlib/Algebra/NoZeroSMulDivisors/Pi.lean | 34 ++++ Mathlib/Algebra/NoZeroSMulDivisors/Prod.lean | 32 ++++ Mathlib/Algebra/Order/Module/Defs.lean | 2 +- Mathlib/Algebra/Order/Star/Basic.lean | 1 + Mathlib/Data/ENNReal/Operations.lean | 1 - Mathlib/Data/Set/Pointwise/SMul.lean | 1 + Mathlib/GroupTheory/Subgroup/Saturated.lean | 2 +- Mathlib/LinearAlgebra/BilinearMap.lean | 1 + Mathlib/LinearAlgebra/LinearPMap.lean | 1 - Mathlib/LinearAlgebra/Multilinear/Basic.lean | 1 + Mathlib/LinearAlgebra/Span.lean | 1 + Mathlib/RingTheory/Nilpotent/Basic.lean | 2 +- 27 files changed, 321 insertions(+), 236 deletions(-) create mode 100644 Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean create mode 100644 Mathlib/Algebra/NoZeroSMulDivisors/Defs.lean create mode 100644 Mathlib/Algebra/NoZeroSMulDivisors/Pi.lean create mode 100644 Mathlib/Algebra/NoZeroSMulDivisors/Prod.lean diff --git a/Mathlib.lean b/Mathlib.lean index c66a3419a9aeb..66e48c38b8ae0 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -544,6 +544,10 @@ import Mathlib.Algebra.MvPolynomial.Rename import Mathlib.Algebra.MvPolynomial.Supported import Mathlib.Algebra.MvPolynomial.Variables import Mathlib.Algebra.NeZero +import Mathlib.Algebra.NoZeroSMulDivisors.Basic +import Mathlib.Algebra.NoZeroSMulDivisors.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Pi +import Mathlib.Algebra.NoZeroSMulDivisors.Prod import Mathlib.Algebra.Notation import Mathlib.Algebra.Opposites import Mathlib.Algebra.Order.AbsoluteValue diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 66932daf4008d..e9e129ad5c0fd 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -9,6 +9,7 @@ import Mathlib.Algebra.Module.Equiv.Basic import Mathlib.Algebra.Module.Submodule.Ker import Mathlib.Algebra.Module.Submodule.RestrictScalars import Mathlib.Algebra.Module.ULift +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Ring.Subring.Basic import Mathlib.Data.Nat.Cast.Order.Basic import Mathlib.Data.Int.CharZero diff --git a/Mathlib/Algebra/BigOperators/Finprod.lean b/Mathlib/Algebra/BigOperators/Finprod.lean index fbfabaaebfca3..b6b2d96428073 100644 --- a/Mathlib/Algebra/BigOperators/Finprod.lean +++ b/Mathlib/Algebra/BigOperators/Finprod.lean @@ -5,7 +5,7 @@ Authors: Kexing Ying, Kevin Buzzard, Yury Kudryashov -/ import Mathlib.Algebra.BigOperators.GroupWithZero.Finset import Mathlib.Algebra.Group.FiniteSupport -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Set.Subsingleton diff --git a/Mathlib/Algebra/Group/EvenFunction.lean b/Mathlib/Algebra/Group/EvenFunction.lean index 4713a0549c61f..81ddf35dcd57c 100644 --- a/Mathlib/Algebra/Group/EvenFunction.lean +++ b/Mathlib/Algebra/Group/EvenFunction.lean @@ -3,9 +3,9 @@ Copyright (c) 2024 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import Mathlib.Algebra.Group.Action.Pi -import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.BigOperators.Group.Finset +import Mathlib.Algebra.Group.Action.Pi +import Mathlib.Algebra.NoZeroSMulDivisors.Basic /-! # Even and odd functions diff --git a/Mathlib/Algebra/ModEq.lean b/Mathlib/Algebra/ModEq.lean index a275aad2c7254..2744ba932b93f 100644 --- a/Mathlib/Algebra/ModEq.lean +++ b/Mathlib/Algebra/ModEq.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Data.Int.ModEq import Mathlib.Algebra.Field.Basic +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Order.Ring.Int import Mathlib.GroupTheory.QuotientGroup.Basic diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index 8656e47667f24..18deb70317a22 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Field.Basic import Mathlib.Algebra.Group.Action.Pi import Mathlib.Algebra.Group.Indicator import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs /-! # Further basic results about modules. @@ -95,30 +96,6 @@ theorem inv_intCast_smul_comm {α E : Type*} (R : Type*) [AddCommGroup E] [Divis @[deprecated (since := "2024-04-17")] alias inv_int_cast_smul_comm := inv_intCast_smul_comm - - -section NoZeroSMulDivisors - -section Module - -instance [AddCommGroup M] [NoZeroSMulDivisors ℤ M] : NoZeroSMulDivisors ℕ M := - ⟨fun {c x} hcx ↦ by rwa [← Nat.cast_smul_eq_nsmul ℤ c x, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ - -end Module - -section GroupWithZero - -variable [GroupWithZero R] [AddMonoid M] [DistribMulAction R M] - --- see note [lower instance priority] -/-- This instance applies to `DivisionSemiring`s, in particular `NNReal` and `NNRat`. -/ -instance (priority := 100) GroupWithZero.toNoZeroSMulDivisors : NoZeroSMulDivisors R M := - ⟨fun {a _} h ↦ or_iff_not_imp_left.2 fun ha ↦ (smul_eq_zero_iff_eq <| Units.mk0 a ha).1 h⟩ - -end GroupWithZero - -end NoZeroSMulDivisors - namespace Function lemma support_smul_subset_left [Zero R] [Zero M] [SMulWithZero R M] (f : α → R) (g : α → M) : diff --git a/Mathlib/Algebra/Module/Card.lean b/Mathlib/Algebra/Module/Card.lean index e9c8b1d5446ec..3941ffd972d1d 100644 --- a/Mathlib/Algebra/Module/Card.lean +++ b/Mathlib/Algebra/Module/Card.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.SetTheory.Cardinal.Basic /-! diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index a024554de986a..c51d03d323fe8 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -408,182 +408,6 @@ instance AddCommGroup.intIsScalarTower {R : Type u} {M : Type v} [Ring R] [AddCo [Module R M] : IsScalarTower ℤ R M where smul_assoc n x y := ((smulAddHom R M).flip y).map_zsmul x n -section NoZeroSMulDivisors - -/-! ### `NoZeroSMulDivisors` - -This section defines the `NoZeroSMulDivisors` class, and includes some tests -for the vanishing of elements (especially in modules over division rings). --/ - - -/-- `NoZeroSMulDivisors R M` states that a scalar multiple is `0` only if either argument is `0`. -This is a version of saying that `M` is torsion free, without assuming `R` is zero-divisor free. - -The main application of `NoZeroSMulDivisors R M`, when `M` is a module, -is the result `smul_eq_zero`: a scalar multiple is `0` iff either argument is `0`. - -It is a generalization of the `NoZeroDivisors` class to heterogeneous multiplication. --/ -@[mk_iff] -class NoZeroSMulDivisors (R M : Type*) [Zero R] [Zero M] [SMul R M] : Prop where - /-- If scalar multiplication yields zero, either the scalar or the vector was zero. -/ - eq_zero_or_eq_zero_of_smul_eq_zero : ∀ {c : R} {x : M}, c • x = 0 → c = 0 ∨ x = 0 - -export NoZeroSMulDivisors (eq_zero_or_eq_zero_of_smul_eq_zero) - -/-- Pullback a `NoZeroSMulDivisors` instance along an injective function. -/ -theorem Function.Injective.noZeroSMulDivisors {R M N : Type*} [Zero R] [Zero M] [Zero N] - [SMul R M] [SMul R N] [NoZeroSMulDivisors R N] (f : M → N) (hf : Function.Injective f) - (h0 : f 0 = 0) (hs : ∀ (c : R) (x : M), f (c • x) = c • f x) : NoZeroSMulDivisors R M := - ⟨fun {_ _} h => - Or.imp_right (@hf _ _) <| h0.symm ▸ eq_zero_or_eq_zero_of_smul_eq_zero (by rw [← hs, h, h0])⟩ - --- See note [lower instance priority] -instance (priority := 100) NoZeroDivisors.toNoZeroSMulDivisors [Zero R] [Mul R] - [NoZeroDivisors R] : NoZeroSMulDivisors R R := - ⟨fun {_ _} => eq_zero_or_eq_zero_of_mul_eq_zero⟩ - -theorem smul_ne_zero [Zero R] [Zero M] [SMul R M] [NoZeroSMulDivisors R M] {c : R} {x : M} - (hc : c ≠ 0) (hx : x ≠ 0) : c • x ≠ 0 := fun h => - (eq_zero_or_eq_zero_of_smul_eq_zero h).elim hc hx - -section SMulWithZero - -variable [Zero R] [Zero M] [SMulWithZero R M] [NoZeroSMulDivisors R M] {c : R} {x : M} - -@[simp] -theorem smul_eq_zero : c • x = 0 ↔ c = 0 ∨ x = 0 := - ⟨eq_zero_or_eq_zero_of_smul_eq_zero, fun h => - h.elim (fun h => h.symm ▸ zero_smul R x) fun h => h.symm ▸ smul_zero c⟩ - -theorem smul_ne_zero_iff : c • x ≠ 0 ↔ c ≠ 0 ∧ x ≠ 0 := by rw [Ne, smul_eq_zero, not_or] - -lemma smul_eq_zero_iff_left (hx : x ≠ 0) : c • x = 0 ↔ c = 0 := by simp [hx] -lemma smul_eq_zero_iff_right (hc : c ≠ 0) : c • x = 0 ↔ x = 0 := by simp [hc] -lemma smul_ne_zero_iff_left (hx : x ≠ 0) : c • x ≠ 0 ↔ c ≠ 0 := by simp [hx] -lemma smul_ne_zero_iff_right (hc : c ≠ 0) : c • x ≠ 0 ↔ x ≠ 0 := by simp [hc] - -end SMulWithZero - -section Module - -section Nat - -theorem Nat.noZeroSMulDivisors - (R) (M) [Semiring R] [CharZero R] [AddCommMonoid M] [Module R M] [NoZeroSMulDivisors R M] : - NoZeroSMulDivisors ℕ M where - eq_zero_or_eq_zero_of_smul_eq_zero {c x} := by rw [← Nat.cast_smul_eq_nsmul R, smul_eq_zero]; simp - -theorem two_nsmul_eq_zero - (R) (M) [Semiring R] [CharZero R] [AddCommMonoid M] [Module R M] [NoZeroSMulDivisors R M] - {v : M} : 2 • v = 0 ↔ v = 0 := by - haveI := Nat.noZeroSMulDivisors R M - simp [smul_eq_zero] - -end Nat - -variable [Semiring R] -variable (R M) - -/-- If `M` is an `R`-module with one and `M` has characteristic zero, then `R` has characteristic -zero as well. Usually `M` is an `R`-algebra. -/ -theorem CharZero.of_module (M) [AddCommMonoidWithOne M] [CharZero M] [Module R M] : CharZero R := by - refine ⟨fun m n h => @Nat.cast_injective M _ _ _ _ ?_⟩ - rw [← nsmul_one, ← nsmul_one, ← Nat.cast_smul_eq_nsmul R, ← Nat.cast_smul_eq_nsmul R, h] - -end Module - -section AddCommGroup - --- `R` can still be a semiring here -variable [Semiring R] [AddCommGroup M] [Module R M] - -section SMulInjective - -variable (M) - -theorem smul_right_injective [NoZeroSMulDivisors R M] {c : R} (hc : c ≠ 0) : - Function.Injective (c • · : M → M) := - (injective_iff_map_eq_zero (smulAddHom R M c)).2 fun _ ha => (smul_eq_zero.mp ha).resolve_left hc - -variable {M} - -theorem smul_right_inj [NoZeroSMulDivisors R M] {c : R} (hc : c ≠ 0) {x y : M} : - c • x = c • y ↔ x = y := - (smul_right_injective M hc).eq_iff - -end SMulInjective - -section Nat - -theorem self_eq_neg - (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] - {v : M} : v = -v ↔ v = 0 := by - rw [← two_nsmul_eq_zero R M, two_smul, add_eq_zero_iff_eq_neg] - -theorem neg_eq_self - (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] - {v : M} : -v = v ↔ v = 0 := by - rw [eq_comm, self_eq_neg R M] - -theorem self_ne_neg - (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] - {v : M} : v ≠ -v ↔ v ≠ 0 := - (self_eq_neg R M).not - -theorem neg_ne_self - (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] - {v : M} : -v ≠ v ↔ v ≠ 0 := - (neg_eq_self R M).not - -end Nat - -end AddCommGroup - -section Module - -variable [Ring R] [AddCommGroup M] [Module R M] - -section SMulInjective - -variable (R) -variable [NoZeroSMulDivisors R M] - -theorem smul_left_injective {x : M} (hx : x ≠ 0) : Function.Injective fun c : R => c • x := - fun c d h => - sub_eq_zero.mp - ((smul_eq_zero.mp - (calc - (c - d) • x = c • x - d • x := sub_smul c d x - _ = 0 := sub_eq_zero.mpr h - )).resolve_right - hx) - -end SMulInjective - -instance [NoZeroSMulDivisors ℤ M] : NoZeroSMulDivisors ℕ M := - ⟨fun {c x} hcx ↦ by rwa [← Nat.cast_smul_eq_nsmul ℤ, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ - -variable (R M) - -theorem NoZeroSMulDivisors.int_of_charZero - (R) (M) [Ring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] [CharZero R] : - NoZeroSMulDivisors ℤ M := - ⟨fun {z x} h ↦ by simpa [← smul_one_smul R z x] using h⟩ - -/-- Only a ring of characteristic zero can have a non-trivial module without additive or -scalar torsion. -/ -theorem CharZero.of_noZeroSMulDivisors [Nontrivial M] [NoZeroSMulDivisors ℤ M] : CharZero R := by - refine ⟨fun {n m h} ↦ ?_⟩ - obtain ⟨x, hx⟩ := exists_ne (0 : M) - replace h : (n : ℤ) • x = (m : ℤ) • x := by simp [← Nat.cast_smul_eq_nsmul R, h] - simpa using smul_left_injective ℤ hx h - -end Module - -end NoZeroSMulDivisors - -- Porting note (#10618): simp can prove this --@[simp] theorem Nat.smul_one_eq_cast {R : Type*} [NonAssocSemiring R] (m : ℕ) : m • (1 : R) = ↑m := by diff --git a/Mathlib/Algebra/Module/LinearMap/Basic.lean b/Mathlib/Algebra/Module/LinearMap/Basic.lean index 1a68ecbd89506..baff9dbf891d4 100644 --- a/Mathlib/Algebra/Module/LinearMap/Basic.lean +++ b/Mathlib/Algebra/Module/LinearMap/Basic.lean @@ -5,7 +5,7 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Frédéric Dupuis, Heather Macbeth -/ import Mathlib.Algebra.Module.LinearMap.Defs -import Mathlib.Algebra.Module.Pi +import Mathlib.Algebra.NoZeroSMulDivisors.Pi import Mathlib.GroupTheory.GroupAction.DomAct.Basic /-! diff --git a/Mathlib/Algebra/Module/LinearMap/End.lean b/Mathlib/Algebra/Module/LinearMap/End.lean index f75595188a8c1..a861127390ad8 100644 --- a/Mathlib/Algebra/Module/LinearMap/End.lean +++ b/Mathlib/Algebra/Module/LinearMap/End.lean @@ -6,6 +6,7 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne -/ import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.Module.LinearMap.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs /-! # Endomorphisms of a module diff --git a/Mathlib/Algebra/Module/Pi.lean b/Mathlib/Algebra/Module/Pi.lean index 1d6a967b8db91..23c86123bb947 100644 --- a/Mathlib/Algebra/Module/Pi.lean +++ b/Mathlib/Algebra/Module/Pi.lean @@ -83,17 +83,4 @@ instance module' {g : I → Type*} {r : ∀ i, Semiring (f i)} {m : ∀ i, AddCo -- Porting note: not sure why `apply zero_smul` fails here. rw [zero_smul] -instance noZeroSMulDivisors (α) [Semiring α] [∀ i, AddCommMonoid <| f i] - [∀ i, Module α <| f i] [∀ i, NoZeroSMulDivisors α <| f i] : - NoZeroSMulDivisors α (∀ i : I, f i) := - ⟨fun {_ _} h => - or_iff_not_imp_left.mpr fun hc => - funext fun i => (smul_eq_zero.mp (congr_fun h i)).resolve_left hc⟩ - -/-- A special case of `Pi.noZeroSMulDivisors` for non-dependent types. Lean struggles to -synthesize this instance by itself elsewhere in the library. -/ -instance _root_.Function.noZeroSMulDivisors {ι α β : Type*} [Semiring α] [AddCommMonoid β] - [Module α β] [NoZeroSMulDivisors α β] : NoZeroSMulDivisors α (ι → β) := - Pi.noZeroSMulDivisors _ - end Pi diff --git a/Mathlib/Algebra/Module/Prod.lean b/Mathlib/Algebra/Module/Prod.lean index 11b934777a06a..de41fcccad536 100644 --- a/Mathlib/Algebra/Module/Prod.lean +++ b/Mathlib/Algebra/Module/Prod.lean @@ -35,17 +35,4 @@ instance instModule [Semiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M add_smul := fun _ _ _ => mk.inj_iff.mpr ⟨add_smul _ _ _, add_smul _ _ _⟩ zero_smul := fun _ => mk.inj_iff.mpr ⟨zero_smul _ _, zero_smul _ _⟩ } -instance noZeroSMulDivisors {r : Semiring R} [AddCommMonoid M] [AddCommMonoid N] - [Module R M] [Module R N] [NoZeroSMulDivisors R M] [NoZeroSMulDivisors R N] : - NoZeroSMulDivisors R (M × N) := - { eq_zero_or_eq_zero_of_smul_eq_zero := by -- Porting note: in mathlib3 there is no need for `by`/ - -- `intro`/`exact`, i.e. the following works: - -- ⟨fun c ⟨x, y⟩ h => - -- or_iff_not_imp_left.mpr fun hc => - intro c ⟨x, y⟩ h - exact or_iff_not_imp_left.mpr fun hc => - mk.inj_iff.mpr - ⟨(smul_eq_zero.mp (congr_arg fst h)).resolve_left hc, - (smul_eq_zero.mp (congr_arg snd h)).resolve_left hc⟩ } - end Prod diff --git a/Mathlib/Algebra/Module/Rat.lean b/Mathlib/Algebra/Module/Rat.lean index c0ed84a4366ed..3ce5e5ac011a8 100644 --- a/Mathlib/Algebra/Module/Rat.lean +++ b/Mathlib/Algebra/Module/Rat.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Module.Basic +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Field.Rat import Mathlib.Algebra.Order.Field.Rat diff --git a/Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean b/Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean new file mode 100644 index 0000000000000..f33114e9583a3 --- /dev/null +++ b/Mathlib/Algebra/NoZeroSMulDivisors/Basic.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2015 Nathaniel Thomas. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen, Yury Kudryashov, Joseph Myers, Heather Macbeth, Kim Morrison, Yaël Dillies +-/ +import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs + +/-! +# `NoZeroSMulDivisors` + +This file defines the `NoZeroSMulDivisors` class, and includes some tests +for the vanishing of elements (especially in modules over division rings). +-/ + +assert_not_exists Multiset +assert_not_exists Set.indicator +assert_not_exists Pi.single_smul₀ +assert_not_exists Field + +section NoZeroSMulDivisors + +variable {R M : Type*} + +section Module + +section Nat + +theorem Nat.noZeroSMulDivisors + (R) (M) [Semiring R] [CharZero R] [AddCommMonoid M] [Module R M] [NoZeroSMulDivisors R M] : + NoZeroSMulDivisors ℕ M where + eq_zero_or_eq_zero_of_smul_eq_zero {c x} := by rw [← Nat.cast_smul_eq_nsmul R, smul_eq_zero]; simp + +theorem two_nsmul_eq_zero + (R) (M) [Semiring R] [CharZero R] [AddCommMonoid M] [Module R M] [NoZeroSMulDivisors R M] + {v : M} : 2 • v = 0 ↔ v = 0 := by + haveI := Nat.noZeroSMulDivisors R M + simp [smul_eq_zero] + +end Nat + +variable [Semiring R] +variable (R M) + +/-- If `M` is an `R`-module with one and `M` has characteristic zero, then `R` has characteristic +zero as well. Usually `M` is an `R`-algebra. -/ +theorem CharZero.of_module (M) [AddCommMonoidWithOne M] [CharZero M] [Module R M] : CharZero R := by + refine ⟨fun m n h => @Nat.cast_injective M _ _ _ _ ?_⟩ + rw [← nsmul_one, ← nsmul_one, ← Nat.cast_smul_eq_nsmul R, ← Nat.cast_smul_eq_nsmul R, h] + +end Module + +section AddCommGroup + +-- `R` can still be a semiring here +variable [Semiring R] [AddCommGroup M] [Module R M] + +section SMulInjective + +variable (M) + +theorem smul_right_injective [NoZeroSMulDivisors R M] {c : R} (hc : c ≠ 0) : + Function.Injective (c • · : M → M) := + (injective_iff_map_eq_zero (smulAddHom R M c)).2 fun _ ha => (smul_eq_zero.mp ha).resolve_left hc + +variable {M} + +theorem smul_right_inj [NoZeroSMulDivisors R M] {c : R} (hc : c ≠ 0) {x y : M} : + c • x = c • y ↔ x = y := + (smul_right_injective M hc).eq_iff + +end SMulInjective + +section Nat + +theorem self_eq_neg + (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] + {v : M} : v = -v ↔ v = 0 := by + rw [← two_nsmul_eq_zero R M, two_smul, add_eq_zero_iff_eq_neg] + +theorem neg_eq_self + (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] + {v : M} : -v = v ↔ v = 0 := by + rw [eq_comm, self_eq_neg R M] + +theorem self_ne_neg + (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] + {v : M} : v ≠ -v ↔ v ≠ 0 := + (self_eq_neg R M).not + +theorem neg_ne_self + (R) (M) [Semiring R] [CharZero R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] + {v : M} : -v ≠ v ↔ v ≠ 0 := + (neg_eq_self R M).not + +end Nat + +end AddCommGroup + +section Module + +variable [Ring R] [AddCommGroup M] [Module R M] + +section SMulInjective + +variable (R) +variable [NoZeroSMulDivisors R M] + +theorem smul_left_injective {x : M} (hx : x ≠ 0) : Function.Injective fun c : R => c • x := + fun c d h => + sub_eq_zero.mp + ((smul_eq_zero.mp + (calc + (c - d) • x = c • x - d • x := sub_smul c d x + _ = 0 := sub_eq_zero.mpr h + )).resolve_right + hx) + +end SMulInjective + +instance [NoZeroSMulDivisors ℤ M] : NoZeroSMulDivisors ℕ M := + ⟨fun {c x} hcx ↦ by rwa [← Nat.cast_smul_eq_nsmul ℤ, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ + +variable (R M) + +theorem NoZeroSMulDivisors.int_of_charZero + (R) (M) [Ring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] [CharZero R] : + NoZeroSMulDivisors ℤ M := + ⟨fun {z x} h ↦ by simpa [← smul_one_smul R z x] using h⟩ + +/-- Only a ring of characteristic zero can have a non-trivial module without additive or +scalar torsion. -/ +theorem CharZero.of_noZeroSMulDivisors [Nontrivial M] [NoZeroSMulDivisors ℤ M] : CharZero R := by + refine ⟨fun {n m h} ↦ ?_⟩ + obtain ⟨x, hx⟩ := exists_ne (0 : M) + replace h : (n : ℤ) • x = (m : ℤ) • x := by simp [← Nat.cast_smul_eq_nsmul R, h] + simpa using smul_left_injective ℤ hx h + +instance [AddCommGroup M] [NoZeroSMulDivisors ℤ M] : NoZeroSMulDivisors ℕ M := + ⟨fun {c x} hcx ↦ by rwa [← Nat.cast_smul_eq_nsmul ℤ c x, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ + +end Module + +section GroupWithZero + +variable [GroupWithZero R] [AddMonoid M] [DistribMulAction R M] + +-- see note [lower instance priority] +/-- This instance applies to `DivisionSemiring`s, in particular `NNReal` and `NNRat`. -/ +instance (priority := 100) GroupWithZero.toNoZeroSMulDivisors : NoZeroSMulDivisors R M := + ⟨fun {a _} h ↦ or_iff_not_imp_left.2 fun ha ↦ (smul_eq_zero_iff_eq <| Units.mk0 a ha).1 h⟩ + +end GroupWithZero + +end NoZeroSMulDivisors diff --git a/Mathlib/Algebra/NoZeroSMulDivisors/Defs.lean b/Mathlib/Algebra/NoZeroSMulDivisors/Defs.lean new file mode 100644 index 0000000000000..caa70e20f08e3 --- /dev/null +++ b/Mathlib/Algebra/NoZeroSMulDivisors/Defs.lean @@ -0,0 +1,78 @@ +/- +Copyright (c) 2015 Nathaniel Thomas. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen, Yury Kudryashov, Joseph Myers, Heather Macbeth, Kim Morrison, Yaël Dillies +-/ +import Mathlib.Algebra.SMulWithZero + +/-! +# `NoZeroSMulDivisors` + +This file defines the `NoZeroSMulDivisors` class, and includes some tests +for the vanishing of elements (especially in modules over division rings). +-/ + +assert_not_exists Multiset +assert_not_exists Set.indicator +assert_not_exists Pi.single_smul₀ +assert_not_exists Field +assert_not_exists Module + +section NoZeroSMulDivisors + +/-! ### `NoZeroSMulDivisors` + +-/ + +variable {R M : Type*} + +/-- `NoZeroSMulDivisors R M` states that a scalar multiple is `0` only if either argument is `0`. +This is a version of saying that `M` is torsion free, without assuming `R` is zero-divisor free. + +The main application of `NoZeroSMulDivisors R M`, when `M` is a module, +is the result `smul_eq_zero`: a scalar multiple is `0` iff either argument is `0`. + +It is a generalization of the `NoZeroDivisors` class to heterogeneous multiplication. +-/ +@[mk_iff] +class NoZeroSMulDivisors (R M : Type*) [Zero R] [Zero M] [SMul R M] : Prop where + /-- If scalar multiplication yields zero, either the scalar or the vector was zero. -/ + eq_zero_or_eq_zero_of_smul_eq_zero : ∀ {c : R} {x : M}, c • x = 0 → c = 0 ∨ x = 0 + +export NoZeroSMulDivisors (eq_zero_or_eq_zero_of_smul_eq_zero) + +/-- Pullback a `NoZeroSMulDivisors` instance along an injective function. -/ +theorem Function.Injective.noZeroSMulDivisors {R M N : Type*} [Zero R] [Zero M] [Zero N] + [SMul R M] [SMul R N] [NoZeroSMulDivisors R N] (f : M → N) (hf : Function.Injective f) + (h0 : f 0 = 0) (hs : ∀ (c : R) (x : M), f (c • x) = c • f x) : NoZeroSMulDivisors R M := + ⟨fun {_ _} h => + Or.imp_right (@hf _ _) <| h0.symm ▸ eq_zero_or_eq_zero_of_smul_eq_zero (by rw [← hs, h, h0])⟩ + +-- See note [lower instance priority] +instance (priority := 100) NoZeroDivisors.toNoZeroSMulDivisors [Zero R] [Mul R] + [NoZeroDivisors R] : NoZeroSMulDivisors R R := + ⟨fun {_ _} => eq_zero_or_eq_zero_of_mul_eq_zero⟩ + +theorem smul_ne_zero [Zero R] [Zero M] [SMul R M] [NoZeroSMulDivisors R M] {c : R} {x : M} + (hc : c ≠ 0) (hx : x ≠ 0) : c • x ≠ 0 := fun h => + (eq_zero_or_eq_zero_of_smul_eq_zero h).elim hc hx + +section SMulWithZero + +variable [Zero R] [Zero M] [SMulWithZero R M] [NoZeroSMulDivisors R M] {c : R} {x : M} + +@[simp] +theorem smul_eq_zero : c • x = 0 ↔ c = 0 ∨ x = 0 := + ⟨eq_zero_or_eq_zero_of_smul_eq_zero, fun h => + h.elim (fun h => h.symm ▸ zero_smul R x) fun h => h.symm ▸ smul_zero c⟩ + +theorem smul_ne_zero_iff : c • x ≠ 0 ↔ c ≠ 0 ∧ x ≠ 0 := by rw [Ne, smul_eq_zero, not_or] + +lemma smul_eq_zero_iff_left (hx : x ≠ 0) : c • x = 0 ↔ c = 0 := by simp [hx] +lemma smul_eq_zero_iff_right (hc : c ≠ 0) : c • x = 0 ↔ x = 0 := by simp [hc] +lemma smul_ne_zero_iff_left (hx : x ≠ 0) : c • x ≠ 0 ↔ c ≠ 0 := by simp [hx] +lemma smul_ne_zero_iff_right (hc : c ≠ 0) : c • x ≠ 0 ↔ x ≠ 0 := by simp [hc] + +end SMulWithZero + +end NoZeroSMulDivisors diff --git a/Mathlib/Algebra/NoZeroSMulDivisors/Pi.lean b/Mathlib/Algebra/NoZeroSMulDivisors/Pi.lean new file mode 100644 index 0000000000000..ac7b42a019040 --- /dev/null +++ b/Mathlib/Algebra/NoZeroSMulDivisors/Pi.lean @@ -0,0 +1,34 @@ +/- +Copyright (c) 2018 Simon Hudon. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen, Yaël Dillies +-/ +import Mathlib.Algebra.NoZeroSMulDivisors.Defs +import Mathlib.Algebra.Group.Action.Pi + +/-! +# Pi instances for NoZeroSMulDivisors + +This file defines instances for NoZeroSMulDivisors on Pi types. +-/ + + +universe u v + +variable {I : Type u} + +-- The indexing type +variable {f : I → Type v} + +instance Pi.noZeroSMulDivisors (α) [Zero α] [∀ i, Zero <| f i] + [∀ i, SMulWithZero α <| f i] [∀ i, NoZeroSMulDivisors α <| f i] : + NoZeroSMulDivisors α (∀ i : I, f i) := + ⟨fun {_ _} h => + or_iff_not_imp_left.mpr fun hc => + funext fun i => (smul_eq_zero.mp (congr_fun h i)).resolve_left hc⟩ + +/-- A special case of `Pi.noZeroSMulDivisors` for non-dependent types. Lean struggles to +synthesize this instance by itself elsewhere in the library. -/ +instance _root_.Function.noZeroSMulDivisors {ι α β : Type*} [Zero α] [Zero β] + [SMulWithZero α β] [NoZeroSMulDivisors α β] : NoZeroSMulDivisors α (ι → β) := + Pi.noZeroSMulDivisors _ diff --git a/Mathlib/Algebra/NoZeroSMulDivisors/Prod.lean b/Mathlib/Algebra/NoZeroSMulDivisors/Prod.lean new file mode 100644 index 0000000000000..e836979c0377a --- /dev/null +++ b/Mathlib/Algebra/NoZeroSMulDivisors/Prod.lean @@ -0,0 +1,32 @@ +/- +Copyright (c) 2018 Simon Hudon. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen +-/ +import Mathlib.Algebra.NoZeroSMulDivisors.Defs +import Mathlib.Algebra.Group.Action.Prod + +/-! +# Prod instances for NoZeroSMulDivisors + +This file defines a NoZeroSMulDivisors instance for the binary product of actions. +-/ + +variable {R M N : Type*} + +namespace Prod + +instance noZeroSMulDivisors [Zero R] [Zero M] [Zero N] + [SMulWithZero R M] [SMulWithZero R N] [NoZeroSMulDivisors R M] [NoZeroSMulDivisors R N] : + NoZeroSMulDivisors R (M × N) := + { eq_zero_or_eq_zero_of_smul_eq_zero := by -- Porting note: in mathlib3 there is no need for `by`/ + -- `intro`/`exact`, i.e. the following works: + -- ⟨fun c ⟨x, y⟩ h => + -- or_iff_not_imp_left.mpr fun hc => + intro c ⟨x, y⟩ h + exact or_iff_not_imp_left.mpr fun hc => + mk.inj_iff.mpr + ⟨(smul_eq_zero.mp (congr_arg fst h)).resolve_left hc, + (smul_eq_zero.mp (congr_arg snd h)).resolve_left hc⟩ } + +end Prod diff --git a/Mathlib/Algebra/Order/Module/Defs.lean b/Mathlib/Algebra/Order/Module/Defs.lean index 8323669b29434..a1172e5a8524d 100644 --- a/Mathlib/Algebra/Order/Module/Defs.lean +++ b/Mathlib/Algebra/Order/Module/Defs.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Order.Field.Defs import Mathlib.Algebra.Order.GroupWithZero.Action.Synonym import Mathlib.Tactic.GCongr diff --git a/Mathlib/Algebra/Order/Star/Basic.lean b/Mathlib/Algebra/Order/Star/Basic.lean index 6370c2c91944e..6df890e8638a3 100644 --- a/Mathlib/Algebra/Order/Star/Basic.lean +++ b/Mathlib/Algebra/Order/Star/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ +import Mathlib.Algebra.NoZeroSMulDivisors.Defs import Mathlib.Algebra.Order.Group.Defs import Mathlib.Algebra.Order.Group.Nat import Mathlib.Algebra.Star.SelfAdjoint diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 473c3d6c65d78..26c974025ccf6 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -5,7 +5,6 @@ Authors: Johannes Hölzl, Yury Kudryashov -/ import Mathlib.Algebra.BigOperators.WithTop import Mathlib.Algebra.GroupWithZero.Divisibility -import Mathlib.Algebra.Module.Basic import Mathlib.Data.ENNReal.Basic /-! diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 9e6917fe0172c..bd562b3fb02ad 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Group.Pi.Basic import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs import Mathlib.Data.Set.Pairwise.Basic /-! diff --git a/Mathlib/GroupTheory/Subgroup/Saturated.lean b/Mathlib/GroupTheory/Subgroup/Saturated.lean index 1d7d366b914ad..6cbc7b2c45d71 100644 --- a/Mathlib/GroupTheory/Subgroup/Saturated.lean +++ b/Mathlib/GroupTheory/Subgroup/Saturated.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import Mathlib.Algebra.Group.Subgroup.Basic -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs /-! # Saturated subgroups diff --git a/Mathlib/LinearAlgebra/BilinearMap.lean b/Mathlib/LinearAlgebra/BilinearMap.lean index 0430d8a20645b..c720aff5ff9b6 100644 --- a/Mathlib/LinearAlgebra/BilinearMap.lean +++ b/Mathlib/LinearAlgebra/BilinearMap.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Mario Carneiro -/ import Mathlib.Algebra.Module.Submodule.Ker +import Mathlib.Algebra.NoZeroSMulDivisors.Basic /-! # Basics on bilinear maps diff --git a/Mathlib/LinearAlgebra/LinearPMap.lean b/Mathlib/LinearAlgebra/LinearPMap.lean index ff0ddc27cc274..6cc05feeabb39 100644 --- a/Mathlib/LinearAlgebra/LinearPMap.lean +++ b/Mathlib/LinearAlgebra/LinearPMap.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Moritz Doll -/ import Mathlib.LinearAlgebra.Prod -import Mathlib.Algebra.Module.Basic /-! # Partially defined linear maps diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 8db7f46936b45..ea1242276149d 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Algebra.Algebra.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Pi import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Fintype.BigOperators import Mathlib.Data.Fintype.Sort diff --git a/Mathlib/LinearAlgebra/Span.lean b/Mathlib/LinearAlgebra/Span.lean index 8e873cfc1c533..478735c139125 100644 --- a/Mathlib/LinearAlgebra/Span.lean +++ b/Mathlib/LinearAlgebra/Span.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Module.Prod import Mathlib.Algebra.Module.Submodule.EqLocus import Mathlib.Algebra.Module.Submodule.Equiv import Mathlib.Algebra.Module.Submodule.RestrictScalars +import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Ring.Idempotents import Mathlib.Data.Set.Pointwise.SMul import Mathlib.Order.CompactlyGenerated.Basic diff --git a/Mathlib/RingTheory/Nilpotent/Basic.lean b/Mathlib/RingTheory/Nilpotent/Basic.lean index 71e196c394db9..9eddb6417010a 100644 --- a/Mathlib/RingTheory/Nilpotent/Basic.lean +++ b/Mathlib/RingTheory/Nilpotent/Basic.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.GeomSum import Mathlib.Algebra.Group.Action.Prod import Mathlib.Algebra.GroupWithZero.NonZeroDivisors -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.NoZeroSMulDivisors.Defs import Mathlib.Algebra.SMulWithZero import Mathlib.Data.Nat.Choose.Sum import Mathlib.Data.Nat.Lattice From 418a5eb7aec3fb639097cb13f74fc031ac4057f2 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 19 Oct 2024 07:30:34 +0000 Subject: [PATCH 099/131] chore: remove CoeFun instances where FunLike is available (#17911) During the port we found that `FunLike` is robust enough not to need an extra `CoeFun` shortcut. Let's see if that rule can be consistently applied to the whole of the library. There is still duplication between `FunLike` and `CoeFun` for `Grp`, `Mon`, `CommGrp` and `CommMon`, which will need a more thorough fix. See also #17866. I am currently bisecting to figure out exactly why the benchmarks are indicating a slowdown with no obvious cause. --- Mathlib/Algebra/Order/AbsoluteValue.lean | 5 ----- Mathlib/Algebra/Ring/CentroidHom.lean | 7 ------- .../Analysis/Distribution/SchwartzSpace.lean | 4 ---- Mathlib/Analysis/Normed/Algebra/Norm.lean | 8 ------- Mathlib/Analysis/Normed/Group/Seminorm.lean | 21 ------------------- 5 files changed, 45 deletions(-) diff --git a/Mathlib/Algebra/Order/AbsoluteValue.lean b/Mathlib/Algebra/Order/AbsoluteValue.lean index 0dcc6acdcfbb4..fae16b0ffcd6d 100644 --- a/Mathlib/Algebra/Order/AbsoluteValue.lean +++ b/Mathlib/Algebra/Order/AbsoluteValue.lean @@ -77,11 +77,6 @@ def Simps.apply (f : AbsoluteValue R S) : R → S := initialize_simps_projections AbsoluteValue (toMulHom_toFun → apply) -/-- Helper instance for when there's too many metavariables to apply `DFunLike.has_coe_to_fun` -directly. -/ -instance : CoeFun (AbsoluteValue R S) fun _ => R → S := - DFunLike.hasCoeToFun - @[simp] theorem coe_toMulHom : ⇑abv.toMulHom = abv := rfl diff --git a/Mathlib/Algebra/Ring/CentroidHom.lean b/Mathlib/Algebra/Ring/CentroidHom.lean index 37a70431efb16..b1b90bb34b9d1 100644 --- a/Mathlib/Algebra/Ring/CentroidHom.lean +++ b/Mathlib/Algebra/Ring/CentroidHom.lean @@ -102,13 +102,6 @@ instance : CentroidHomClass (CentroidHom α) α where map_mul_right f := f.map_mul_right' -/-- Helper instance for when there's too many metavariables to apply `DFunLike.CoeFun` -directly. -/ -/- Porting note: Lean gave me `unknown constant 'DFunLike.CoeFun'` and says `CoeFun` is a type -mismatch, so I used `library_search`. -/ -instance : CoeFun (CentroidHom α) fun _ ↦ α → α := - inferInstanceAs (CoeFun (CentroidHom α) fun _ ↦ α → α) - -- Porting note: removed @[simp]; not in normal form. (`toAddMonoidHom_eq_coe` below ensures that -- the LHS simplifies to the RHS anyway.) theorem toFun_eq_coe {f : CentroidHom α} : f.toFun = f := rfl diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index ba8aae073aff9..a9cd6e02b3fbe 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -90,10 +90,6 @@ instance instFunLike : FunLike 𝓢(E, F) E F where coe f := f.toFun coe_injective' f g h := by cases f; cases g; congr -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun`. -/ -instance instCoeFun : CoeFun 𝓢(E, F) fun _ => E → F := - DFunLike.hasCoeToFun - /-- All derivatives of a Schwartz function are rapidly decaying. -/ theorem decay (f : 𝓢(E, F)) (k n : ℕ) : ∃ C : ℝ, 0 < C ∧ ∀ x, ‖x‖ ^ k * ‖iteratedFDeriv ℝ n f x‖ ≤ C := by diff --git a/Mathlib/Analysis/Normed/Algebra/Norm.lean b/Mathlib/Analysis/Normed/Algebra/Norm.lean index eb815ac91298e..715ad8694cf6f 100644 --- a/Mathlib/Analysis/Normed/Algebra/Norm.lean +++ b/Mathlib/Analysis/Normed/Algebra/Norm.lean @@ -70,10 +70,6 @@ instance algebraNormClass : AlgebraNormClass (AlgebraNorm R S) R S where eq_zero_of_map_eq_zero f := f.eq_zero_of_map_eq_zero' _ map_smul_eq_mul f := f.smul' -/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ -instance : CoeFun (AlgebraNorm R S) fun _ => S → ℝ := - DFunLike.hasCoeToFun - theorem toFun_eq_coe (p : AlgebraNorm R S) : p.toFun = p := rfl @[ext] @@ -160,10 +156,6 @@ instance mulAlgebraNormClass : MulAlgebraNormClass (MulAlgebraNorm R S) R S wher eq_zero_of_map_eq_zero f := f.eq_zero_of_map_eq_zero' _ map_smul_eq_mul f := f.smul' -/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ -instance : CoeFun (MulAlgebraNorm R S) fun _ => S → ℝ := - DFunLike.hasCoeToFun - theorem toFun_eq_coe (p : MulAlgebraNorm R S) : p.toFun = p := rfl @[ext] diff --git a/Mathlib/Analysis/Normed/Group/Seminorm.lean b/Mathlib/Analysis/Normed/Group/Seminorm.lean index eede911fb046d..a8d9220ba89cc 100644 --- a/Mathlib/Analysis/Normed/Group/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Group/Seminorm.lean @@ -181,12 +181,6 @@ instance groupSeminormClass : GroupSeminormClass (GroupSeminorm E) E ℝ where map_mul_le_add f := f.mul_le' map_inv_eq_map f := f.inv' -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun`. -/ -@[to_additive "Helper instance for when there's too many metavariables to apply -`DFunLike.hasCoeToFun`. "] -instance : CoeFun (GroupSeminorm E) fun _ => E → ℝ := - ⟨DFunLike.coe⟩ - @[to_additive (attr := simp)] theorem toFun_eq_coe : p.toFun = p := rfl @@ -447,10 +441,6 @@ instance nonarchAddGroupSeminormClass : map_zero f := f.map_zero' map_neg_eq_map' f := f.neg' -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun`. -/ -instance : CoeFun (NonarchAddGroupSeminorm E) fun _ => E → ℝ := - ⟨DFunLike.coe⟩ - -- Porting note: `simpNF` said the left hand side simplified to this @[simp] theorem toZeroHom_eq_coe : ⇑p.toZeroHom = p := by @@ -667,13 +657,6 @@ instance groupNormClass : GroupNormClass (GroupNorm E) E ℝ where map_inv_eq_map f := f.inv' eq_one_of_map_eq_zero f := f.eq_one_of_map_eq_zero' _ -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun` -directly. -/ -@[to_additive "Helper instance for when there's too many metavariables to apply -`DFunLike.hasCoeToFun` directly. "] -instance : CoeFun (GroupNorm E) fun _ => E → ℝ := - DFunLike.hasCoeToFun - -- Porting note: `simpNF` told me the left-hand side simplified to this @[to_additive (attr := simp)] theorem toGroupSeminorm_eq_coe : ⇑p.toGroupSeminorm = p := @@ -799,10 +782,6 @@ instance nonarchAddGroupNormClass : NonarchAddGroupNormClass (NonarchAddGroupNor map_neg_eq_map' f := f.neg' eq_zero_of_map_eq_zero f := f.eq_zero_of_map_eq_zero' _ -/-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun`. -/ -noncomputable instance : CoeFun (NonarchAddGroupNorm E) fun _ => E → ℝ := - DFunLike.hasCoeToFun - -- Porting note: `simpNF` told me the left-hand side simplified to this @[simp] theorem toNonarchAddGroupSeminorm_eq_coe : ⇑p.toNonarchAddGroupSeminorm = p := From 66f63643834286e06540e60e459b22c7bbd25475 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sat, 19 Oct 2024 08:15:11 +0000 Subject: [PATCH 100/131] chore: move trans tactic to Batteries (batteries#1001) (#17931) Co-authored-by: F. G. Dorais --- Mathlib.lean | 1 - Mathlib/Algebra/Notation.lean | 1 + Mathlib/Order/Defs.lean | 2 +- Mathlib/Tactic.lean | 1 - Mathlib/Tactic/Common.lean | 1 - Mathlib/Tactic/MoveAdd.lean | 1 + Mathlib/Tactic/Relation/Trans.lean | 222 ------------------------ Mathlib/Tactic/ToAdditive/Frontend.lean | 4 +- Mathlib/Tactic/Widget/CongrM.lean | 1 + lake-manifest.json | 2 +- scripts/noshake.json | 2 +- test/trans.lean | 109 ------------ 12 files changed, 8 insertions(+), 339 deletions(-) delete mode 100644 Mathlib/Tactic/Relation/Trans.lean delete mode 100644 test/trans.lean diff --git a/Mathlib.lean b/Mathlib.lean index 66e48c38b8ae0..e24762ba7699b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4447,7 +4447,6 @@ import Mathlib.Tactic.ReduceModChar import Mathlib.Tactic.ReduceModChar.Ext import Mathlib.Tactic.Relation.Rfl import Mathlib.Tactic.Relation.Symm -import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.Rename import Mathlib.Tactic.RenameBVar import Mathlib.Tactic.Replace diff --git a/Mathlib/Algebra/Notation.lean b/Mathlib/Algebra/Notation.lean index 0ea472b668eb9..ffb3097c7ba73 100644 --- a/Mathlib/Algebra/Notation.lean +++ b/Mathlib/Algebra/Notation.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ +import Mathlib.Tactic.TypeStar import Mathlib.Tactic.ToAdditive /-! diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index 7a46f0ce76cf7..75d603ee2c080 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ import Batteries.Classes.Order +import Batteries.Tactic.Trans import Mathlib.Data.Ordering.Basic import Mathlib.Tactic.Lemma -import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.SplitIfs import Mathlib.Tactic.TypeStar diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index cd494193a67b4..efac05b1b7da1 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -203,7 +203,6 @@ import Mathlib.Tactic.ReduceModChar import Mathlib.Tactic.ReduceModChar.Ext import Mathlib.Tactic.Relation.Rfl import Mathlib.Tactic.Relation.Symm -import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.Rename import Mathlib.Tactic.RenameBVar import Mathlib.Tactic.Replace diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index f9c7a25eb782c..c864e4fc60f26 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -79,7 +79,6 @@ import Mathlib.Tactic.PushNeg import Mathlib.Tactic.RSuffices import Mathlib.Tactic.Recover import Mathlib.Tactic.Relation.Rfl -import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.Rename import Mathlib.Tactic.RenameBVar import Mathlib.Tactic.Says diff --git a/Mathlib/Tactic/MoveAdd.lean b/Mathlib/Tactic/MoveAdd.lean index 094ce46b60265..87cf3bdb6add4 100644 --- a/Mathlib/Tactic/MoveAdd.lean +++ b/Mathlib/Tactic/MoveAdd.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Damiano Testa -/ import Mathlib.Algebra.Group.Basic +import Mathlib.Lean.Meta /-! diff --git a/Mathlib/Tactic/Relation/Trans.lean b/Mathlib/Tactic/Relation/Trans.lean deleted file mode 100644 index 290d23e9aeb70..0000000000000 --- a/Mathlib/Tactic/Relation/Trans.lean +++ /dev/null @@ -1,222 +0,0 @@ -/- -Copyright (c) 2022 Siddhartha Gadgil. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Siddhartha Gadgil, Mario Carneiro --/ -import Mathlib.Lean.Meta -import Mathlib.Lean.Elab.Tactic.Basic -import Lean.Elab.Tactic.ElabTerm -import Mathlib.Tactic.TypeStar - -/-! -# `trans` tactic - -This implements the `trans` tactic, which can apply transitivity theorems with an optional middle -variable argument. --/ - -namespace Mathlib.Tactic -open Lean Meta Elab - -initialize registerTraceClass `Tactic.trans - -/-- Discrimation tree settings for the `trans` extension. -/ -def transExt.config : WhnfCoreConfig := {} - -/-- Environment extension storing transitivity lemmas -/ -initialize transExt : - SimpleScopedEnvExtension (Name × Array DiscrTree.Key) (DiscrTree Name) ← - registerSimpleScopedEnvExtension { - addEntry := fun dt (n, ks) ↦ dt.insertCore ks n - initial := {} - } - -initialize registerBuiltinAttribute { - name := `trans - descr := "transitive relation" - add := fun decl _ kind ↦ MetaM.run' do - let declTy := (← getConstInfo decl).type - let (xs, _, targetTy) ← withReducible <| forallMetaTelescopeReducing declTy - let fail := throwError - "@[trans] attribute only applies to lemmas proving - x ∼ y → y ∼ z → x ∼ z, got {indentExpr declTy} with target {indentExpr targetTy}" - let .app (.app rel _) _ := targetTy | fail - let some yzHyp := xs.back? | fail - let some xyHyp := xs.pop.back? | fail - let .app (.app _ _) _ ← inferType yzHyp | fail - let .app (.app _ _) _ ← inferType xyHyp | fail - let key ← withReducible <| DiscrTree.mkPath rel transExt.config - transExt.add (decl, key) kind -} - -universe u v in -/-- Composition using the `Trans` class in the homogeneous case. -/ -def _root_.Trans.simple {α : Sort u} {r : α → α → Sort v} {a b c : α} [Trans r r r] : - r a b → r b c → r a c := trans - -universe u v w in -/-- Composition using the `Trans` class in the general case. -/ -def _root_.Trans.het {α β γ : Sort*} {a : α} {b : β} {c : γ} - {r : α → β → Sort u} {s : β → γ → Sort v} {t : outParam (α → γ → Sort w)} [Trans r s t] : - r a b → s b c → t a c := trans - -open Lean.Elab.Tactic - -/-- solving `e ← mkAppM' f #[x]` -/ -def getExplicitFuncArg? (e : Expr) : MetaM (Option <| Expr × Expr) := do - match e with - | Expr.app f a => do - if ← isDefEq (← mkAppM' f #[a]) e then - return some (f, a) - else - getExplicitFuncArg? f - | _ => return none - -/-- solving `tgt ← mkAppM' rel #[x, z]` given `tgt = f z` -/ -def getExplicitRelArg? (tgt f z : Expr) : MetaM (Option <| Expr × Expr) := do - match f with - | Expr.app rel x => do - let check: Bool ← do - try - let folded ← mkAppM' rel #[x, z] - isDefEq folded tgt - catch _ => - pure false - if check then - return some (rel, x) - else - getExplicitRelArg? tgt rel z - | _ => return none - -/-- refining `tgt ← mkAppM' rel #[x, z]` dropping more arguments if possible -/ -def getExplicitRelArgCore (tgt rel x z : Expr) : MetaM (Expr × Expr) := do - match rel with - | Expr.app rel' _ => do - let check: Bool ← do - try - let folded ← mkAppM' rel' #[x, z] - isDefEq folded tgt - catch _ => - pure false - if !check then - return (rel, x) - else - getExplicitRelArgCore tgt rel' x z - | _ => return (rel ,x) - -/-- Internal definition for `trans` tactic. Either a binary relation or a non-dependent -arrow. -/ -inductive TransRelation - | app (rel : Expr) - | implies (name : Name) (bi : BinderInfo) - -/-- Finds an explicit binary relation in the argument, if possible. -/ -def getRel (tgt : Expr) : MetaM (Option (TransRelation × Expr × Expr)) := do - match tgt with - | .forallE name binderType body info => return .some (.implies name info, binderType, body) - | .app f z => - match (← getExplicitRelArg? tgt f z) with - | some (rel, x) => - let (rel, x) ← getExplicitRelArgCore tgt rel x z - return some (.app rel, x, z) - | none => - return none - | _ => return none - -/-- -`trans` applies to a goal whose target has the form `t ~ u` where `~` is a transitive relation, -that is, a relation which has a transitivity lemma tagged with the attribute [trans]. - -* `trans s` replaces the goal with the two subgoals `t ~ s` and `s ~ u`. -* If `s` is omitted, then a metavariable is used instead. - -Additionally, `trans` also applies to a goal whose target has the form `t → u`, -in which case it replaces the goal with `t → s` and `s → u`. --/ -elab "trans" t?:(ppSpace colGt term)? : tactic => withMainContext do - let tgt ← getMainTarget'' - let .some (rel, x, z) ← getRel tgt | - throwError (m!"transitivity lemmas only apply to binary relations and " ++ - m!"non-dependent arrows, not {indentExpr tgt}") - match rel with - | .implies name info => - -- only consider non-dependent arrows - if z.hasLooseBVars then - throwError "`trans` is not implemented for dependent arrows{indentExpr tgt}" - -- parse the intermeditate term - let middleType ← mkFreshExprMVar none - let t'? ← t?.mapM (elabTermWithHoles · middleType (← getMainTag)) - let middle ← (t'?.map (pure ·.1)).getD (mkFreshExprMVar middleType) - liftMetaTactic fun goal => do - -- create two new goals - let g₁ ← mkFreshExprMVar (some <| .forallE name x middle info) .synthetic - let g₂ ← mkFreshExprMVar (some <| .forallE name middle z info) .synthetic - -- close the original goal with `fun x => g₂ (g₁ x)` - goal.assign (.lam name x (.app g₂ (.app g₁ (.bvar 0))) .default) - pure <| [g₁.mvarId!, g₂.mvarId!] ++ if let some (_, gs') := t'? then gs' else [middle.mvarId!] - return - | .app rel => - trace[Tactic.trans]"goal decomposed" - trace[Tactic.trans]"rel: {indentExpr rel}" - trace[Tactic.trans]"x: {indentExpr x}" - trace[Tactic.trans]"z: {indentExpr z}" - -- first trying the homogeneous case - try - let ty ← inferType x - let t'? ← t?.mapM (elabTermWithHoles · ty (← getMainTag)) - let s ← saveState - trace[Tactic.trans]"trying homogeneous case" - let lemmas := - (← (transExt.getState (← getEnv)).getUnify rel transExt.config).push ``Trans.simple - for lem in lemmas do - trace[Tactic.trans]"trying lemma {lem}" - try - liftMetaTactic fun g ↦ do - let lemTy ← inferType (← mkConstWithLevelParams lem) - let arity ← withReducible <| forallTelescopeReducing lemTy fun es _ ↦ pure es.size - let y ← (t'?.map (pure ·.1)).getD (mkFreshExprMVar ty) - let g₁ ← mkFreshExprMVar (some <| ← mkAppM' rel #[x, y]) .synthetic - let g₂ ← mkFreshExprMVar (some <| ← mkAppM' rel #[y, z]) .synthetic - g.assign (← mkAppOptM lem (mkArray (arity - 2) none ++ #[some g₁, some g₂])) - pure <| [g₁.mvarId!, g₂.mvarId!] ++ - if let some (_, gs') := t'? then gs' else [y.mvarId!] - return - catch _ => s.restore - pure () - catch _ => - trace[Tactic.trans]"trying heterogeneous case" - let t'? ← t?.mapM (elabTermWithHoles · none (← getMainTag)) - let s ← saveState - for lem in (← (transExt.getState (← getEnv)).getUnify rel transExt.config).push - ``HEq.trans |>.push ``HEq.trans do - try - liftMetaTactic fun g ↦ do - trace[Tactic.trans]"trying lemma {lem}" - let lemTy ← inferType (← mkConstWithLevelParams lem) - let arity ← withReducible <| forallTelescopeReducing lemTy fun es _ ↦ pure es.size - trace[Tactic.trans]"arity: {arity}" - trace[Tactic.trans]"lemma-type: {lemTy}" - let y ← (t'?.map (pure ·.1)).getD (mkFreshExprMVar none) - trace[Tactic.trans]"obtained y: {y}" - trace[Tactic.trans]"rel: {indentExpr rel}" - trace[Tactic.trans]"x:{indentExpr x}" - trace[Tactic.trans]"z: {indentExpr z}" - let g₂ ← mkFreshExprMVar (some <| ← mkAppM' rel #[y, z]) .synthetic - trace[Tactic.trans]"obtained g₂: {g₂}" - let g₁ ← mkFreshExprMVar (some <| ← mkAppM' rel #[x, y]) .synthetic - trace[Tactic.trans]"obtained g₁: {g₁}" - g.assign (← mkAppOptM lem (mkArray (arity - 2) none ++ #[some g₁, some g₂])) - pure <| [g₁.mvarId!, g₂.mvarId!] ++ if let some (_, gs') := t'? then gs' else [y.mvarId!] - return - catch e => - trace[Tactic.trans]"failed: {e.toMessageData}" - s.restore - throwError m!"no applicable transitivity lemma found for {indentExpr tgt}" - -syntax "transitivity" (ppSpace colGt term)? : tactic -set_option hygiene false in -macro_rules - | `(tactic| transitivity) => `(tactic| trans) - | `(tactic| transitivity $e) => `(tactic| trans $e) - -end Mathlib.Tactic diff --git a/Mathlib/Tactic/ToAdditive/Frontend.lean b/Mathlib/Tactic/ToAdditive/Frontend.lean index 277e01a289f6b..c0d0ef7afdcfa 100644 --- a/Mathlib/Tactic/ToAdditive/Frontend.lean +++ b/Mathlib/Tactic/ToAdditive/Frontend.lean @@ -16,7 +16,7 @@ import Lean.Meta.Tactic.Rfl import Lean.Meta.Match.MatcherInfo import Batteries.Lean.NameMapAttribute import Batteries.Tactic.Lint -- useful to lint this file and for DiscrTree.elements -import Mathlib.Tactic.Relation.Trans -- just to copy the attribute +import Batteries.Tactic.Trans import Mathlib.Tactic.Eqns -- just to copy the attribute import Mathlib.Tactic.Simps.Basic @@ -1167,7 +1167,7 @@ partial def applyAttributes (stx : Syntax) (rawAttrs : Array Syntax) (thisAttr s (fun b n => (b.tree.values.any fun t => t.declName = n)) thisAttr `ext src tgt warnAttr stx Lean.Meta.Rfl.reflExt (·.values.contains ·) thisAttr `refl src tgt warnAttr stx Lean.Meta.Symm.symmExt (·.values.contains ·) thisAttr `symm src tgt - warnAttr stx Mathlib.Tactic.transExt (·.values.contains ·) thisAttr `trans src tgt + warnAttr stx Batteries.Tactic.transExt (·.values.contains ·) thisAttr `trans src tgt warnAttr stx Lean.Meta.coeExt (·.contains ·) thisAttr `coe src tgt warnParametricAttr stx Lean.Linter.deprecatedAttr thisAttr `deprecated src tgt -- the next line also warns for `@[to_additive, simps]`, because of the application times diff --git a/Mathlib/Tactic/Widget/CongrM.lean b/Mathlib/Tactic/Widget/CongrM.lean index 9c259c3daaf30..dad4520d652cc 100644 --- a/Mathlib/Tactic/Widget/CongrM.lean +++ b/Mathlib/Tactic/Widget/CongrM.lean @@ -6,6 +6,7 @@ Authors: Patrick Massot import Mathlib.Tactic.Widget.SelectPanelUtils import Mathlib.Tactic.CongrM +import Batteries.Lean.Position /-! # CongrM widget diff --git a/lake-manifest.json b/lake-manifest.json index 6f116ba1524dd..8d8fe8356466a 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "1e0bf50b357069e1d658512a579a5faac6587c40", + "rev": "10130798199d306703dee5ab2567961444ebbd04", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/scripts/noshake.json b/scripts/noshake.json index 3e93d378ca0e2..a98fca062e875 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -148,7 +148,6 @@ "Mathlib.Tactic.Recover", "Mathlib.Tactic.ReduceModChar.Ext", "Mathlib.Tactic.Relation.Symm", - "Mathlib.Tactic.Relation.Trans", "Mathlib.Tactic.Rename", "Mathlib.Tactic.RenameBVar", "Mathlib.Tactic.Replace", @@ -325,6 +324,7 @@ "Mathlib.Order.RelClasses": ["Batteries.WF"], "Mathlib.Order.Filter.ListTraverse": ["Mathlib.Control.Traversable.Instances"], + "Mathlib.Order.Defs": ["Batteries.Tactic.Trans"], "Mathlib.Order.Basic": ["Batteries.Tactic.Classical"], "Mathlib.NumberTheory.Harmonic.Defs": ["Mathlib.Algebra.Order.BigOperators.Ring.Finset", diff --git a/test/trans.lean b/test/trans.lean deleted file mode 100644 index 50b5bc01e7b20..0000000000000 --- a/test/trans.lean +++ /dev/null @@ -1,109 +0,0 @@ -import Mathlib.Tactic.Relation.Trans -import Mathlib.Tactic.GuardGoalNums - -set_option autoImplicit true --- testing that the attribute is recognized and used -def nleq (a b : Nat) : Prop := a ≤ b - -@[trans] def nleq_trans : nleq a b → nleq b c → nleq a c := Nat.le_trans - -example (a b c : Nat) : nleq a b → nleq b c → nleq a c := by - intro h₁ h₂ - trans b - assumption - assumption - -example (a b c : Nat) : nleq a b → nleq b c → nleq a c := by intros; trans <;> assumption - --- using `Trans` typeclass -@[trans] def eq_trans {a b c : α} : a = b → b = c → a = c := by - intro h₁ h₂ - apply Eq.trans h₁ h₂ - -example (a b c : Nat) : a = b → b = c → a = c := by intros; trans <;> assumption - -example (a b c : Nat) : a = b → b = c → a = c := by - intro h₁ h₂ - trans b - assumption - assumption - -example : @Trans Nat Nat Nat (· ≤ ·) (· ≤ ·) (· ≤ ·) := inferInstance - -example (a b c : Nat) : a ≤ b → b ≤ c → a ≤ c := by - intros h₁ h₂ - trans ?b - case b => exact b - exact h₁ - exact h₂ - -example (a b c : α) (R : α → α → Prop) [Trans R R R] : R a b → R b c → R a c := by - intros h₁ h₂ - trans ?b - case b => exact b - exact h₁ - exact h₂ - -example (a b c : Nat) : a ≤ b → b ≤ c → a ≤ c := by - intros h₁ h₂ - trans - exact h₁ - exact h₂ - -example (a b c : Nat) : a ≤ b → b ≤ c → a ≤ c := by intros; trans <;> assumption - -example (a b c : Nat) : a < b → b < c → a < c := by - intro h₁ h₂ - trans b - assumption - assumption - -example (a b c : Nat) : a < b → b < c → a < c := by intros; trans <;> assumption - -example (x n p : Nat) (h₁ : n * Nat.succ p ≤ x) : n * p ≤ x := by - trans - · apply Nat.mul_le_mul_left; apply Nat.le_succ - · apply h₁ - -example (a : α) (c : γ) : ∀ b : β, HEq a b → HEq b c → HEq a c := by - intro b h₁ h₂ - trans b - assumption - assumption - -def MyLE (n m : Nat) := ∃ k, n + k = m - -@[trans] theorem MyLE.trans {n m k : Nat} (h1 : MyLE n m) (h2 : MyLE m k) : MyLE n k := by - cases h1 - cases h2 - subst_vars - exact ⟨_, Eq.symm <| Nat.add_assoc _ _ _⟩ - -example {n m k : Nat} (h1 : MyLE n m) (h2 : MyLE m k) : MyLE n k := by - trans <;> assumption - -/-- `trans` for implications. -/ -example {A B C : Prop} (h : A → B) (g : B → C) : A → C := by - trans B - · guard_target =ₛ A → B -- ensure we have `B` and not a free metavariable. - exact h - · guard_target =ₛ B → C - exact g - -set_option linter.unusedTactic false in -/-- `trans` for arrows between types. -/ -example {A B C : Type} (h : A → B) (g : B → C) : A → C := by - trans - guard_goal_nums 3 -- 3rd goal is the middle term - · exact h - · exact g - -universe u v w - -set_option linter.unusedTactic false in -/-- `trans` for arrows between types. -/ -example {A : Type u} {B : Type v} {C : Type w} (h : A → B) (g : B → C) : A → C := by - trans - guard_goal_nums 3 -- 3rd goal is the middle term - · exact h - · exact g From 9e3db2c4fc34b715fd1876983c6b12759a0ccb0d Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 12:38:54 +0000 Subject: [PATCH 101/131] perf: de-simp `tsub_eq_zero_of_le` (#17512) As mentioned [on Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Infrastructure.20for.20tracking.20frequently.20applied.20simp.20theorems/near/475331955): > `tsub_eq_zero_of_le` got worse in [the unmerged] #17444. It seems that Lean is searching `CanonicallyOrderedAdd` for types that are not ordered at all, then doing some slow unification of instances with metavariables that is impossible to succeed. --- Archive/Imo/Imo1986Q5.lean | 2 +- Mathlib/Algebra/MvPolynomial/Derivation.lean | 2 +- Mathlib/Algebra/Order/Sub/Basic.lean | 2 -- Mathlib/Data/ENNReal/Inv.lean | 2 +- Mathlib/Data/ENNReal/Operations.lean | 2 +- Mathlib/Data/Nat/Factorization/Basic.lean | 2 +- Mathlib/Data/Real/ConjExponents.lean | 4 ++-- Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean | 3 ++- Mathlib/MeasureTheory/Measure/DiracProba.lean | 2 +- Mathlib/Topology/Instances/ENNReal.lean | 2 +- Mathlib/Topology/Instances/ENat.lean | 2 +- Mathlib/Topology/MetricSpace/ThickenedIndicator.lean | 2 +- 12 files changed, 13 insertions(+), 14 deletions(-) diff --git a/Archive/Imo/Imo1986Q5.lean b/Archive/Imo/Imo1986Q5.lean index 82fe5961c1647..e8ca0c898f1d5 100644 --- a/Archive/Imo/Imo1986Q5.lean +++ b/Archive/Imo/Imo1986Q5.lean @@ -71,7 +71,7 @@ theorem isGood_iff {f : ℝ≥0 → ℝ≥0} : IsGood f ↔ f = fun x ↦ 2 / (2 refine ⟨fun hf ↦ funext hf.map_eq, ?_⟩ rintro rfl constructor - case map_two => simp + case map_two => simp [tsub_self] case map_ne_zero => intro x hx; simpa [tsub_eq_zero_iff_le] case map_add_rev => intro x y diff --git a/Mathlib/Algebra/MvPolynomial/Derivation.lean b/Mathlib/Algebra/MvPolynomial/Derivation.lean index 49d6860a55c77..1d88e3c8c50e2 100644 --- a/Mathlib/Algebra/MvPolynomial/Derivation.lean +++ b/Mathlib/Algebra/MvPolynomial/Derivation.lean @@ -45,7 +45,7 @@ theorem mkDerivationₗ_C (f : σ → A) (r : R) : mkDerivationₗ R f (C r) = 0 (mkDerivationₗ_monomial f _ _).trans (smul_zero _) theorem mkDerivationₗ_X (f : σ → A) (i : σ) : mkDerivationₗ R f (X i) = f i := - (mkDerivationₗ_monomial f _ _).trans <| by simp + (mkDerivationₗ_monomial f _ _).trans <| by simp [tsub_self] @[simp] theorem derivation_C (D : Derivation R (MvPolynomial σ R) A) (a : R) : D (C a) = 0 := diff --git a/Mathlib/Algebra/Order/Sub/Basic.lean b/Mathlib/Algebra/Order/Sub/Basic.lean index b2d6abd8dd435..e1b98ddf85486 100644 --- a/Mathlib/Algebra/Order/Sub/Basic.lean +++ b/Mathlib/Algebra/Order/Sub/Basic.lean @@ -32,8 +32,6 @@ theorem tsub_eq_zero_iff_le : a - b = 0 ↔ a ≤ b := by alias ⟨_, tsub_eq_zero_of_le⟩ := tsub_eq_zero_iff_le -attribute [simp] tsub_eq_zero_of_le - theorem tsub_self (a : α) : a - a = 0 := tsub_eq_zero_of_le le_rfl diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 0f9c29a783d26..c0a3bde921b26 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -831,7 +831,7 @@ lemma smul_sSup {R} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] lemma sub_iSup [Nonempty ι] (ha : a ≠ ∞) : a - ⨆ i, f i = ⨅ i, a - f i := by obtain ⟨i, hi⟩ | h := em (∃ i, a < f i) · rw [tsub_eq_zero_iff_le.2 <| le_iSup_of_le _ hi.le, (iInf_eq_bot _).2, bot_eq_zero] - exact fun x hx ↦ ⟨i, by simpa [hi.le]⟩ + exact fun x hx ↦ ⟨i, by simpa [hi.le, tsub_eq_zero_of_le]⟩ simp_rw [not_exists, not_lt] at h refine le_antisymm (le_iInf fun i ↦ tsub_le_tsub_left (le_iSup ..) _) <| ENNReal.le_sub_of_add_le_left (ne_top_of_le_ne_top ha <| iSup_le h) <| diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 26c974025ccf6..8318373b6a1bf 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -405,7 +405,7 @@ theorem sub_right_inj {a b c : ℝ≥0∞} (ha : a ≠ ∞) (hb : b ≤ a) (hc : (cancel_of_ne <| ne_top_of_le_ne_top ha hc) hb hc theorem sub_mul (h : 0 < b → b < a → c ≠ ∞) : (a - b) * c = a * c - b * c := by - rcases le_or_lt a b with hab | hab; · simp [hab, mul_right_mono hab] + rcases le_or_lt a b with hab | hab; · simp [hab, mul_right_mono hab, tsub_eq_zero_of_le] rcases eq_or_lt_of_le (zero_le b) with (rfl | hb); · simp exact (cancel_of_ne <| mul_ne_top hab.ne_top (h hb hab)).tsub_mul diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 17af0fffcfa82..c8c06e616e2d7 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -187,7 +187,7 @@ theorem exists_factorization_lt_of_lt {a b : ℕ} (ha : a ≠ 0) (hab : a < b) : theorem factorization_div {d n : ℕ} (h : d ∣ n) : (n / d).factorization = n.factorization - d.factorization := by rcases eq_or_ne d 0 with (rfl | hd); · simp [zero_dvd_iff.mp h] - rcases eq_or_ne n 0 with (rfl | hn); · simp + rcases eq_or_ne n 0 with (rfl | hn); · simp [tsub_eq_zero_of_le] apply add_left_injective d.factorization simp only rw [tsub_add_cancel_of_le <| (Nat.factorization_le_iff_dvd hd hn).mpr h, ← diff --git a/Mathlib/Data/Real/ConjExponents.lean b/Mathlib/Data/Real/ConjExponents.lean index 04bfc20da0ec1..581ac0eb8bed3 100644 --- a/Mathlib/Data/Real/ConjExponents.lean +++ b/Mathlib/Data/Real/ConjExponents.lean @@ -272,7 +272,7 @@ protected lemma conjExponent (hp : 1 ≤ p) : p.IsConjExponent (conjExponent p) refine (AddLECancellable.eq_tsub_iff_add_eq_of_le (α := ℝ≥0∞) (by simpa) (by simpa)).1 ?_ rw [inv_eq_iff_eq_inv] obtain rfl | hp₁ := hp.eq_or_lt - · simp + · simp [tsub_eq_zero_of_le] obtain rfl | hp := eq_or_ne p ∞ · simp calc @@ -317,7 +317,7 @@ lemma mul_eq_add : p * q = p + q := by lemma div_conj_eq_sub_one : p / q = p - 1 := by obtain rfl | hq := eq_or_ne q ∞ - · simp [h.symm.conj_eq] + · simp [h.symm.conj_eq, tsub_eq_zero_of_le] refine ENNReal.eq_sub_of_add_eq one_ne_top ?_ rw [← ENNReal.div_self h.symm.ne_zero hq, ← ENNReal.add_div, ← h.mul_eq_add, mul_div_assoc, ENNReal.div_self h.symm.ne_zero hq, mul_one] diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index 9dfc57bb16291..ed5e038f6f5c4 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -255,7 +255,8 @@ instance instMeasurableMul₂ : MeasurableMul₂ ℝ≥0∞ := by instance instMeasurableSub₂ : MeasurableSub₂ ℝ≥0∞ := ⟨by apply measurable_of_measurable_nnreal_nnreal <;> - simp [← WithTop.coe_sub]; exact continuous_sub.measurable.coe_nnreal_ennreal⟩ + simp [← WithTop.coe_sub, tsub_eq_zero_of_le]; + exact continuous_sub.measurable.coe_nnreal_ennreal⟩ instance instMeasurableInv : MeasurableInv ℝ≥0∞ := ⟨continuous_inv.measurable⟩ diff --git a/Mathlib/MeasureTheory/Measure/DiracProba.lean b/Mathlib/MeasureTheory/Measure/DiracProba.lean index 18bc9a7f40a5e..51377c0135958 100644 --- a/Mathlib/MeasureTheory/Measure/DiracProba.lean +++ b/Mathlib/MeasureTheory/Measure/DiracProba.lean @@ -36,7 +36,7 @@ lemma CompletelyRegularSpace.exists_BCNN {X : Type*} [TopologicalSpace X] [Compl set g' := BoundedContinuousFunction.mkOfBound ⟨fun x ↦ Real.toNNReal (g x), continuous_real_toNNReal.comp g_cont.subtype_val⟩ 1 g_bdd set f := 1 - g' - refine ⟨f, by simp [f, g', gx_zero], fun y y_in_K ↦ by simp [f, g', g_one_on_K y_in_K]⟩ + refine ⟨f, by simp [f, g', gx_zero], fun y y_in_K ↦ by simp [f, g', g_one_on_K y_in_K, tsub_self]⟩ namespace MeasureTheory diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index a6719468d25d3..f5c72e34f691d 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -422,7 +422,7 @@ theorem continuousOn_sub_left (a : ℝ≥0∞) : ContinuousOn (a - ·) { x : ℝ theorem continuous_sub_right (a : ℝ≥0∞) : Continuous fun x : ℝ≥0∞ => x - a := by by_cases a_infty : a = ∞ - · simp [a_infty, continuous_const] + · simp [a_infty, continuous_const, tsub_eq_zero_of_le] · rw [show (fun x => x - a) = (fun p : ℝ≥0∞ × ℝ≥0∞ => p.fst - p.snd) ∘ fun x => ⟨x, a⟩ by rfl] apply ContinuousOn.comp_continuous continuousOn_sub (continuous_id'.prod_mk continuous_const) intro x diff --git a/Mathlib/Topology/Instances/ENat.lean b/Mathlib/Topology/Instances/ENat.lean index 0217f6abc247c..c69882b280e04 100644 --- a/Mathlib/Topology/Instances/ENat.lean +++ b/Mathlib/Topology/Instances/ENat.lean @@ -92,7 +92,7 @@ protected theorem continuousAt_sub {a b : ℕ∞} (h : a ≠ ⊤ ∨ b ≠ ⊤) simpa [ContinuousAt, nhds_prod_eq] using tendsto_pure_nhds _ _ | (a : ℕ), ⊤, _ => suffices ∀ᶠ b in 𝓝 ⊤, (a - b : ℕ∞) = 0 by - simpa [ContinuousAt, nhds_prod_eq] + simpa [ContinuousAt, nhds_prod_eq, tsub_eq_zero_of_le] filter_upwards [le_mem_nhds (WithTop.coe_lt_top a)] with b using tsub_eq_zero_of_le | ⊤, (b : ℕ), _ => suffices ∀ n : ℕ, ∀ᶠ a : ℕ∞ in 𝓝 ⊤, b + n < a by diff --git a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean index 50b3c349f536c..1df15588c66b4 100644 --- a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean +++ b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean @@ -86,7 +86,7 @@ theorem thickenedIndicatorAux_zero {δ : ℝ} (δ_pos : 0 < δ) (E : Set α) {x have key := tsub_le_tsub (@rfl _ (1 : ℝ≥0∞)).le (ENNReal.div_le_div x_out (@rfl _ (ENNReal.ofReal δ : ℝ≥0∞)).le) rw [ENNReal.div_self (ne_of_gt (ENNReal.ofReal_pos.mpr δ_pos)) ofReal_ne_top] at key - simpa using key + simpa [tsub_self] using key theorem thickenedIndicatorAux_mono {δ₁ δ₂ : ℝ} (hle : δ₁ ≤ δ₂) (E : Set α) : thickenedIndicatorAux δ₁ E ≤ thickenedIndicatorAux δ₂ E := From 2b96a2f7f0879e2bd4f2517079e0752c24206277 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Calle=20S=C3=B6nne?= Date: Sat, 19 Oct 2024 13:30:30 +0000 Subject: [PATCH 102/131] feat(CategoryTheory/Bicategory): Grothendieck construction for a pseudofunctor (#15419) In this PR we add the Grothendieck construction for pseudofunctors. --- Mathlib.lean | 1 + .../Bicategory/Grothendieck.lean | 122 ++++++++++++++++++ docs/references.bib | 10 ++ 3 files changed, 133 insertions(+) create mode 100644 Mathlib/CategoryTheory/Bicategory/Grothendieck.lean diff --git a/Mathlib.lean b/Mathlib.lean index e24762ba7699b..f51a308d8a11e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1462,6 +1462,7 @@ import Mathlib.CategoryTheory.Bicategory.Functor.Oplax import Mathlib.CategoryTheory.Bicategory.Functor.Prelax import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor import Mathlib.CategoryTheory.Bicategory.FunctorBicategory +import Mathlib.CategoryTheory.Bicategory.Grothendieck import Mathlib.CategoryTheory.Bicategory.Kan.Adjunction import Mathlib.CategoryTheory.Bicategory.Kan.HasKan import Mathlib.CategoryTheory.Bicategory.Kan.IsKan diff --git a/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean b/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean new file mode 100644 index 0000000000000..d7eb112f1f6ba --- /dev/null +++ b/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean @@ -0,0 +1,122 @@ +/- +Copyright (c) 2024 Calle Sönne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Calle Sönne +-/ + +import Mathlib.CategoryTheory.Bicategory.LocallyDiscrete + +/-! +# The Grothendieck construction + +Given a category `𝒮` and any pseudofunctor `F` from `𝒮ᵒᵖ` to `Cat`, we associate to it a category +`∫ F`, equipped with a functor `∫ F ⥤ 𝒮`. + +The category `∫ F` is defined as follows: +* Objects: pairs `(S, a)` where `S` is an object of the base category and `a` is an object of the + category `F(S)`. +* Morphisms: morphisms `(R, b) ⟶ (S, a)` are defined as pairs `(f, h)` where `f : R ⟶ S` is a + morphism in `𝒮` and `h : b ⟶ F(f)(a)` + +The projection functor `∫ F ⥤ 𝒮` is then given by projecting to the first factors, i.e. +* On objects, it sends `(S, a)` to `S` +* On morphisms, it sends `(f, h)` to `f` + +## References +[Vistoli2008] "Notes on Grothendieck Topologies, Fibered Categories and Descent Theory" by +Angelo Vistoli +-/ + +namespace CategoryTheory + +universe w v₁ v₂ v₃ u₁ u₂ u₃ + +open CategoryTheory Functor Category Opposite Discrete Bicategory + +variable {𝒮 : Type u₁} [Category.{v₁} 𝒮] {F : Pseudofunctor (LocallyDiscrete 𝒮ᵒᵖ) Cat.{v₂, u₂}} + +/-- The type of objects in the fibered category associated to a presheaf valued in types. -/ +@[ext] +structure Pseudofunctor.Grothendieck (F : Pseudofunctor (LocallyDiscrete 𝒮ᵒᵖ) Cat.{v₂, u₂}) where + /-- The underlying object in the base category. -/ + base : 𝒮 + /-- The object in the fiber of the base object. -/ + fiber : F.obj ⟨op base⟩ + +namespace Pseudofunctor.Grothendieck + +/-- Notation for the Grothendieck category associated to a pseudofunctor `F`. -/ +scoped prefix:75 "∫ " => Pseudofunctor.Grothendieck + +/-- A morphism in the Grothendieck category `F : C ⥤ Cat` consists of +`base : X.base ⟶ Y.base` and `f.fiber : (F.map base).obj X.fiber ⟶ Y.fiber`. +-/ +structure Hom (X Y : ∫ F) where + /-- The morphism between base objects. -/ + base : X.base ⟶ Y.base + /-- The morphism in the fiber over the domain. -/ + fiber : X.fiber ⟶ (F.map base.op.toLoc).obj Y.fiber + +@[simps!] +instance categoryStruct : CategoryStruct (∫ F) where + Hom X Y := Hom X Y + id X := { + base := 𝟙 X.base + fiber := (F.mapId ⟨op X.base⟩).inv.app X.fiber } + comp {_ _ Z} f g := { + base := f.base ≫ g.base + fiber := f.fiber ≫ (F.map f.base.op.toLoc).map g.fiber ≫ + (F.mapComp g.base.op.toLoc f.base.op.toLoc).inv.app Z.fiber } + +section + +variable {a b : ∫ F} + +@[ext (iff := false)] +lemma Hom.ext (f g : a ⟶ b) (hfg₁ : f.base = g.base) + (hfg₂ : f.fiber = g.fiber ≫ eqToHom (hfg₁ ▸ rfl)) : f = g := by + cases f; cases g + congr + dsimp at hfg₁ + rw [← conj_eqToHom_iff_heq _ _ rfl (hfg₁ ▸ rfl)] + simpa only [eqToHom_refl, id_comp] using hfg₂ + +lemma Hom.ext_iff (f g : a ⟶ b) : + f = g ↔ ∃ (hfg : f.base = g.base), f.fiber = g.fiber ≫ eqToHom (hfg ▸ rfl) where + mp hfg := ⟨by rw [hfg], by simp [hfg]⟩ + mpr := fun ⟨hfg₁, hfg₂⟩ => Hom.ext f g hfg₁ hfg₂ + +lemma Hom.congr {a b : ∫ F} {f g : a ⟶ b} (h : f = g) : + f.fiber = g.fiber ≫ eqToHom (h ▸ rfl) := by + simp [h] + +end + +/-- The category structure on `∫ F`. -/ +instance category : Category (∫ F) where + toCategoryStruct := Pseudofunctor.Grothendieck.categoryStruct + id_comp {a b} f := by + ext + · simp + · simp [F.mapComp_id_right_inv_app, Strict.rightUnitor_eqToIso, ← NatTrans.naturality_assoc] + comp_id {a b} f := by + ext + · simp + · simp [F.mapComp_id_left_inv_app, ← Functor.map_comp_assoc, Strict.leftUnitor_eqToIso] + assoc f g h := by + ext + · simp + · simp [← NatTrans.naturality_assoc, F.mapComp_assoc_right_inv_app, Strict.associator_eqToIso] + +variable (F) + +/-- The projection `∫ F ⥤ 𝒮` given by projecting both objects and homs to the first +factor. -/ +@[simps] +def forget : ∫ F ⥤ 𝒮 where + obj X := X.base + map f := f.base + +end Pseudofunctor.Grothendieck + +end CategoryTheory diff --git a/docs/references.bib b/docs/references.bib index 39e500681a649..e94cd696fd232 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -3287,6 +3287,16 @@ @Book{ verdier1996 mrnumber = {1453167} } +@Misc{ vistoli2004, + author = {Vistoli, Angelo}, + title = {Notes on {Grothendieck} topologies, fibered categories and + descent theory}, + year = {2004}, + howpublished = {Preprint, {arXiv}:math/0412512 [math.{AG}] (2004)}, + url = {https://arxiv.org/abs/math/0412512}, + arxiv = {arXiv:math/0412512} +} + @Book{ wall2018analytic, title = {Analytic Theory of Continued Fractions}, author = {Wall, H.S.}, From 92412db77694c08a00efa5b1028264e8c7079f7b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Calle=20S=C3=B6nne?= Date: Sat, 19 Oct 2024 13:51:18 +0000 Subject: [PATCH 103/131] feat(FiberedCategory/Fiber): define the fibers of a functor (#13603) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define the fibers of a functor `p : 𝒳 ⥤ 𝒮` over objects `S : 𝒮`, and develop some basic API. Relatively little API is developed as in an upcoming PR we introduce a class `HasFibers` which should be preferred. Co-authored-by: Paul Lezeau --- Mathlib.lean | 1 + .../CategoryTheory/FiberedCategory/Fiber.lean | 127 ++++++++++++++++++ .../FiberedCategory/HomLift.lean | 10 +- 3 files changed, 135 insertions(+), 3 deletions(-) create mode 100644 Mathlib/CategoryTheory/FiberedCategory/Fiber.lean diff --git a/Mathlib.lean b/Mathlib.lean index f51a308d8a11e..fbd2d2cec89b5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1547,6 +1547,7 @@ import Mathlib.CategoryTheory.Extensive import Mathlib.CategoryTheory.FiberedCategory.BasedCategory import Mathlib.CategoryTheory.FiberedCategory.Cartesian import Mathlib.CategoryTheory.FiberedCategory.Cocartesian +import Mathlib.CategoryTheory.FiberedCategory.Fiber import Mathlib.CategoryTheory.FiberedCategory.Fibered import Mathlib.CategoryTheory.FiberedCategory.HomLift import Mathlib.CategoryTheory.Filtered.Basic diff --git a/Mathlib/CategoryTheory/FiberedCategory/Fiber.lean b/Mathlib/CategoryTheory/FiberedCategory/Fiber.lean new file mode 100644 index 0000000000000..94cfb1727d195 --- /dev/null +++ b/Mathlib/CategoryTheory/FiberedCategory/Fiber.lean @@ -0,0 +1,127 @@ +/- +Copyright (c) 2024 Calle Sönne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Calle Sönne, Paul Lezeau +-/ + +import Mathlib.CategoryTheory.FiberedCategory.HomLift +import Mathlib.CategoryTheory.Functor.Const + +/-! + +# Fibers of a functors + +In this file we define, for a functor `p : 𝒳 ⥤ 𝒴`, the fiber categories `Fiber p S` for every +`S : 𝒮` as follows +- An object in `Fiber p S` is a pair `(a, ha)` where `a : 𝒳` and `ha : p.obj a = S`. +- A morphism in `Fiber p S` is a morphism `φ : a ⟶ b` in 𝒳 such that `p.map φ = 𝟙 S`. + +For any category `C` equipped with a functor `F : C ⥤ 𝒳` such that `F ⋙ p` is constant at `S`, +we define a functor `inducedFunctor : C ⥤ Fiber p S` that `F` factors through. +-/ + +universe v₁ u₁ v₂ u₂ v₃ u₃ + +namespace CategoryTheory + +open IsHomLift + +namespace Functor + +variable {𝒮 : Type u₁} {𝒳 : Type u₂} [Category.{v₁} 𝒮] [Category.{v₂} 𝒳] + +/-- `Fiber p S` is the type of elements of `𝒳` mapping to `S` via `p`. -/ +def Fiber (p : 𝒳 ⥤ 𝒮) (S : 𝒮) := { a : 𝒳 // p.obj a = S } + +namespace Fiber + +variable {p : 𝒳 ⥤ 𝒮} {S : 𝒮} + +/-- `Fiber p S` has the structure of a category with morphisms being those lying over `𝟙 S`. -/ +instance fiberCategory : Category (Fiber p S) where + Hom a b := {φ : a.1 ⟶ b.1 // IsHomLift p (𝟙 S) φ} + id a := ⟨𝟙 a.1, IsHomLift.id a.2⟩ + comp φ ψ := ⟨φ.val ≫ ψ.val, by have := φ.2; have := ψ.2; infer_instance⟩ + +/-- The functor including `Fiber p S` into `𝒳`. -/ +def fiberInclusion : Fiber p S ⥤ 𝒳 where + obj a := a.1 + map φ := φ.1 + +instance {a b : Fiber p S} (φ : a ⟶ b) : IsHomLift p (𝟙 S) (fiberInclusion.map φ) := φ.2 + +@[ext] +lemma hom_ext {a b : Fiber p S} {φ ψ : a ⟶ b} + (h : fiberInclusion.map φ = fiberInclusion.map ψ) : φ = ψ := + Subtype.ext h + +instance : (fiberInclusion : Fiber p S ⥤ _).Faithful where + +/-- For fixed `S : 𝒮` this is the natural isomorphism between `fiberInclusion ⋙ p` and the constant +function valued at `S`. -/ +@[simps!] +def fiberInclusionCompIsoConst : fiberInclusion ⋙ p ≅ (const (Fiber p S)).obj S := + NatIso.ofComponents (fun X ↦ eqToIso X.2) + (fun φ ↦ by simp [IsHomLift.fac' p (𝟙 S) (fiberInclusion.map φ)]) + +lemma fiberInclusion_comp_eq_const : fiberInclusion ⋙ p = (const (Fiber p S)).obj S := + Functor.ext (fun x ↦ x.2) (fun _ _ φ ↦ IsHomLift.fac' p (𝟙 S) (fiberInclusion.map φ)) + +/-- The object of the fiber over `S` corresponding to a `a : 𝒳` such that `p(a) = S`. -/ +def mk {p : 𝒳 ⥤ 𝒮} {S : 𝒮} {a : 𝒳} (ha : p.obj a = S) : Fiber p S := ⟨a, ha⟩ + +@[simp] +lemma fiberInclusion_mk {p : 𝒳 ⥤ 𝒮} {S : 𝒮} {a : 𝒳} (ha : p.obj a = S) : + fiberInclusion.obj (mk ha) = a := + rfl + +/-- The morphism in the fiber over `S` corresponding to a morphism in `𝒳` lifting `𝟙 S`. -/ +def homMk (p : 𝒳 ⥤ 𝒮) (S : 𝒮) {a b : 𝒳} (φ : a ⟶ b) [IsHomLift p (𝟙 S) φ] : + mk (domain_eq p (𝟙 S) φ) ⟶ mk (codomain_eq p (𝟙 S) φ) := + ⟨φ, inferInstance⟩ + +@[simp] +lemma fiberInclusion_homMk (p : 𝒳 ⥤ 𝒮) (S : 𝒮) {a b : 𝒳} (φ : a ⟶ b) [IsHomLift p (𝟙 S) φ] : + fiberInclusion.map (homMk p S φ) = φ := + rfl + +@[simp] +lemma homMk_id (p : 𝒳 ⥤ 𝒮) (S : 𝒮) (a : 𝒳) [IsHomLift p (𝟙 S) (𝟙 a)] : + homMk p S (𝟙 a) = 𝟙 (mk (domain_eq p (𝟙 S) (𝟙 a))) := + rfl + +@[simp] +lemma homMk_comp {a b c : 𝒳} (φ : a ⟶ b) (ψ : b ⟶ c) [IsHomLift p (𝟙 S) φ] + [IsHomLift p (𝟙 S) ψ] : homMk p S φ ≫ homMk p S ψ = homMk p S (φ ≫ ψ) := + rfl + +section + +variable {p : 𝒳 ⥤ 𝒮} {S : 𝒮} {C : Type u₃} [Category.{v₃} C] {F : C ⥤ 𝒳} + (hF : F ⋙ p = (const C).obj S) + +/-- Given a functor `F : C ⥤ 𝒳` such that `F ⋙ p` is constant at some `S : 𝒮`, then +we get an induced functor `C ⥤ Fiber p S` that `F` factors through. -/ +@[simps] +def inducedFunctor : C ⥤ Fiber p S where + obj x := ⟨F.obj x, by simp only [← comp_obj, hF, const_obj_obj]⟩ + map φ := ⟨F.map φ, of_commsq _ _ _ _ _ <| by simpa using (eqToIso hF).hom.naturality φ⟩ + +@[simp] +lemma inducedFunctor_map {X Y : C} (f : X ⟶ Y) : + fiberInclusion.map ((inducedFunctor hF).map f) = F.map f := rfl + +/-- Given a functor `F : C ⥤ 𝒳` such that `F ⋙ p` is constant at some `S : 𝒮`, then +we get a natural isomorphism between `inducedFunctor _ ⋙ fiberInclusion` and `F`. -/ +@[simps!] +def inducedFunctorCompIsoSelf : (inducedFunctor hF) ⋙ fiberInclusion ≅ F := Iso.refl _ + +lemma inducedFunctor_comp : (inducedFunctor hF) ⋙ fiberInclusion = F := rfl + +end + +end Fiber + +end Functor + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean b/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean index b9bc045f14acb..5d64e6e16c258 100644 --- a/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean +++ b/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean @@ -106,12 +106,16 @@ lemma of_fac' {R S : 𝒮} {a b : 𝒳} (f : R ⟶ S) (φ : a ⟶ b) (ha : p.obj obtain rfl : f = p.map φ := by simpa using h.symm infer_instance -lemma of_commSq {R S : 𝒮} {a b : 𝒳} (f : R ⟶ S) (φ : a ⟶ b) (ha : p.obj a = R) (hb : p.obj b = S) - (h : CommSq (p.map φ) (eqToHom ha) (eqToHom hb) f) : p.IsHomLift f φ := by +lemma of_commsq {R S : 𝒮} {a b : 𝒳} (f : R ⟶ S) (φ : a ⟶ b) (ha : p.obj a = R) (hb : p.obj b = S) + (h : p.map φ ≫ eqToHom hb = (eqToHom ha) ≫ f) : p.IsHomLift f φ := by subst ha hb - obtain rfl : f = p.map φ := by simpa using h.1.symm + obtain rfl : f = p.map φ := by simpa using h.symm infer_instance +lemma of_commSq {R S : 𝒮} {a b : 𝒳} (f : R ⟶ S) (φ : a ⟶ b) (ha : p.obj a = R) (hb : p.obj b = S) + (h : CommSq (p.map φ) (eqToHom ha) (eqToHom hb) f) : p.IsHomLift f φ := + of_commsq p f φ ha hb h.1 + instance comp {R S T : 𝒮} {a b c : 𝒳} (f : R ⟶ S) (g : S ⟶ T) (φ : a ⟶ b) (ψ : b ⟶ c) [p.IsHomLift f φ] [p.IsHomLift g ψ] : p.IsHomLift (f ≫ g) (φ ≫ ψ) := by apply of_commSq From 0d932abadf1d7aad0fd9efbae434ff5e6fa24667 Mon Sep 17 00:00:00 2001 From: Daniel Weber Date: Sat, 19 Oct 2024 14:07:56 +0000 Subject: [PATCH 104/131] feat: define rooted trees (#16272) This is preparation for proving Kruskal's tree theorem. Co-authored-by: Daniel Weber <55664973+Command-Master@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Order/SuccPred/Tree.lean | 228 +++++++++++++++++++++++++++++++ 2 files changed, 229 insertions(+) create mode 100644 Mathlib/Order/SuccPred/Tree.lean diff --git a/Mathlib.lean b/Mathlib.lean index fbd2d2cec89b5..4c5d00a1fa732 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3797,6 +3797,7 @@ import Mathlib.Order.SuccPred.IntervalSucc import Mathlib.Order.SuccPred.Limit import Mathlib.Order.SuccPred.LinearLocallyFinite import Mathlib.Order.SuccPred.Relation +import Mathlib.Order.SuccPred.Tree import Mathlib.Order.SupClosed import Mathlib.Order.SupIndep import Mathlib.Order.SymmDiff diff --git a/Mathlib/Order/SuccPred/Tree.lean b/Mathlib/Order/SuccPred/Tree.lean new file mode 100644 index 0000000000000..ea7ab54e7fb12 --- /dev/null +++ b/Mathlib/Order/SuccPred/Tree.lean @@ -0,0 +1,228 @@ +/- +Copyright (c) 2024 Daniel Weber. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Daniel Weber +-/ + +import Mathlib.Order.SuccPred.Archimedean +import Mathlib.Data.Nat.Find +import Mathlib.Order.Atoms +import Mathlib.Data.SetLike.Basic + +/-! +# Rooted trees + +This file proves basic results about rooted trees, represented using the ancestorship order. +This is a `PartialOrder`, with `PredOrder` with the immediate parent as a predecessor, and an +`OrderBot` which is the root. We also have an `IsPredArchimedean` assumption to prevent infinite +dangling chains. + +--/ + +variable {α : Type*} [PartialOrder α] [PredOrder α] [IsPredArchimedean α] + +namespace IsPredArchimedean + +variable [OrderBot α] + +section DecidableEq + +variable [DecidableEq α] + +/-- +The unique atom less than an element in an `OrderBot` with archimedean predecessor. +-/ +def findAtom (r : α) : α := + Order.pred^[Nat.find (bot_le (a := r)).exists_pred_iterate - 1] r + +@[simp] +lemma findAtom_le (r : α) : findAtom r ≤ r := + Order.pred_iterate_le _ _ + +@[simp] +lemma findAtom_bot : findAtom (⊥ : α) = ⊥ := by + apply Function.iterate_fixed + simp + +@[simp] +lemma pred_findAtom (r : α) : Order.pred (findAtom r) = ⊥ := by + unfold findAtom + generalize h : Nat.find (bot_le (a := r)).exists_pred_iterate = n + cases n + · have : Order.pred^[0] r = ⊥ := by + rw [← h] + apply Nat.find_spec (bot_le (a := r)).exists_pred_iterate + simp only [Function.iterate_zero, id_eq] at this + simp [this] + · simp only [Nat.add_sub_cancel_right, ← Function.iterate_succ_apply', Nat.succ_eq_add_one] + rw [← h] + apply Nat.find_spec (bot_le (a := r)).exists_pred_iterate + +@[simp] +lemma findAtom_eq_bot {r : α} : + findAtom r = ⊥ ↔ r = ⊥ where + mp h := by + unfold findAtom at h + have := Nat.find_min' (bot_le (a := r)).exists_pred_iterate h + replace : Nat.find (bot_le (a := r)).exists_pred_iterate = 0 := by omega + simpa [this] using h + mpr h := by simp [h] + +lemma findAtom_ne_bot {r : α} : + findAtom r ≠ ⊥ ↔ r ≠ ⊥ := findAtom_eq_bot.not + +lemma isAtom_findAtom {r : α} (hr : r ≠ ⊥) : + IsAtom (findAtom r) := by + constructor + · simp [hr] + · intro b hb + apply Order.le_pred_of_lt at hb + simpa using hb + +@[simp] +lemma isAtom_findAtom_iff {r : α} : + IsAtom (findAtom r) ↔ r ≠ ⊥ where + mpr := isAtom_findAtom + mp h nh := by simp only [nh, findAtom_bot] at h; exact h.1 rfl + +end DecidableEq + +instance instIsAtomic : IsAtomic α where + eq_bot_or_exists_atom_le b := by classical + rw [or_iff_not_imp_left] + intro hb + use findAtom b, isAtom_findAtom hb, findAtom_le b + +end IsPredArchimedean + +/-- +The type of rooted trees. +-/ +structure RootedTree where + /-- The type representing the elements in the tree. -/ + α : Type* + /-- The type should be a `SemilatticeInf`, + where `inf` is the least common ancestor in the tree. -/ + [semilatticeInf : SemilatticeInf α] + /-- The type should have a bottom, the root. -/ + [orderBot : OrderBot α] + /-- The type should have a predecessor for every element, its parent. -/ + [predOrder : PredOrder α] + /-- The predecessor relationship should be archimedean. -/ + [isPredArchimedean : IsPredArchimedean α] + +attribute [coe] RootedTree.α + +instance : CoeSort RootedTree Type* := ⟨RootedTree.α⟩ + +attribute [instance] RootedTree.semilatticeInf RootedTree.predOrder + RootedTree.orderBot RootedTree.isPredArchimedean + +/-- +A subtree is represented by its root, therefore this is a type synonym. +-/ +def SubRootedTree (t : RootedTree) : Type* := t + +/-- +The root of a `SubRootedTree`. +-/ +def SubRootedTree.root {t : RootedTree} (v : SubRootedTree t) : t := v + +/-- +The `SubRootedTree` rooted at a given node. +-/ +def RootedTree.subtree (t : RootedTree) (r : t) : SubRootedTree t := r + +@[simp] +lemma RootedTree.root_subtree (t : RootedTree) (r : t) : (t.subtree r).root = r := rfl + +@[simp] +lemma RootedTree.subtree_root (t : RootedTree) (v : SubRootedTree t) : t.subtree v.root = v := rfl + +@[ext] +lemma SubRootedTree.ext {t : RootedTree} {v₁ v₂ : SubRootedTree t} + (h : v₁.root = v₂.root) : v₁ = v₂ := h + +instance (t : RootedTree) : SetLike (SubRootedTree t) t where + coe v := Set.Ici v.root + coe_injective' a₁ a₂ h := by + simpa only [Set.Ici_inj, ← SubRootedTree.ext_iff] using h + +lemma SubRootedTree.mem_iff {t : RootedTree} {r : SubRootedTree t} {v : t} : + v ∈ r ↔ r.root ≤ v := Iff.rfl + +/-- +The coercion from a `SubRootedTree` to a `RootedTree`. +-/ +@[coe, reducible] +noncomputable def SubRootedTree.coeTree {t : RootedTree} (r : SubRootedTree t) : RootedTree where + α := Set.Ici r.root + +noncomputable instance (t : RootedTree) : CoeOut (SubRootedTree t) RootedTree := + ⟨SubRootedTree.coeTree⟩ + +@[simp] +lemma SubRootedTree.bot_mem_iff {t : RootedTree} (r : SubRootedTree t) : + ⊥ ∈ r ↔ r.root = ⊥ := by + simp [mem_iff] + +/-- +All of the immediate subtrees of a given rooted tree, that is subtrees which are rooted at a direct +child of the root (or, order theoretically, at an atom). +-/ +def RootedTree.subtrees (t : RootedTree) : Set (SubRootedTree t) := + {x | IsAtom x.root} + +variable {t : RootedTree} + +lemma SubRootedTree.root_ne_bot_of_mem_subtrees (r : SubRootedTree t) (hr : r ∈ t.subtrees) : + r.root ≠ ⊥ := by + simp only [RootedTree.subtrees, Set.mem_setOf_eq] at hr + exact hr.1 + +lemma RootedTree.mem_subtrees_disjoint_iff {t₁ t₂ : SubRootedTree t} + (ht₁ : t₁ ∈ t.subtrees) (ht₂ : t₂ ∈ t.subtrees) (v₁ v₂ : t) (h₁ : v₁ ∈ t₁) + (h₂ : v₂ ∈ t₂) : + Disjoint v₁ v₂ ↔ t₁ ≠ t₂ where + mp h := by + intro nh + have : t₁.root ≤ (v₁ : t) ⊓ (v₂ : t) := by + simp only [le_inf_iff] + exact ⟨h₁, nh ▸ h₂⟩ + rw [h.eq_bot] at this + simp only [le_bot_iff] at this + exact t₁.root_ne_bot_of_mem_subtrees ht₁ this + mpr h := by + rw [SubRootedTree.mem_iff] at h₁ h₂ + contrapose! h + rw [disjoint_iff, ← ne_eq, ← bot_lt_iff_ne_bot] at h + rcases lt_or_le_of_directed (by simp : v₁ ⊓ v₂ ≤ v₁) h₁ with oh | oh + · simp_all [RootedTree.subtrees, IsAtom.lt_iff] + rw [le_inf_iff] at oh + ext + simpa only [ht₂.le_iff_eq ht₁.1, ht₁.le_iff_eq ht₂.1, eq_comm, or_self] using + le_total_of_directed oh.2 h₂ + +lemma RootedTree.subtrees_disjoint : t.subtrees.PairwiseDisjoint ((↑) : _ → Set t) := by + intro t₁ ht₁ t₂ ht₂ h + rw [Function.onFun_apply, Set.disjoint_left] + intro a ha hb + rw [← mem_subtrees_disjoint_iff ht₁ ht₂ a a ha hb, disjoint_self] at h + subst h + simp only [SetLike.mem_coe, SubRootedTree.bot_mem_iff] at ha + exact t₁.root_ne_bot_of_mem_subtrees ht₁ ha + +/-- +The immediate subtree of `t` containing `v`, or all of `t` if `v` is the root. +-/ +def RootedTree.subtreeOf (t : RootedTree) [DecidableEq t] (v : t) : SubRootedTree t := + t.subtree (IsPredArchimedean.findAtom v) + +@[simp] +lemma RootedTree.mem_subtreeOf [DecidableEq t] {v : t} : + v ∈ t.subtreeOf v := by + simp [SubRootedTree.mem_iff, RootedTree.subtreeOf] + +lemma RootedTree.subtreeOf_mem_subtrees [DecidableEq t] {v : t} (hr : v ≠ ⊥) : + t.subtreeOf v ∈ t.subtrees := by + simpa [RootedTree.subtrees, RootedTree.subtreeOf] From c8922d33629788528cbfa02f9c56afb909d1c72d Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 15:25:40 +0000 Subject: [PATCH 105/131] chore: make `Quotient.mk''` an abbrev of `Quotient.mk _` (#16264) The plan is to deprecate `Quotient.mk''` in future PRs. --- Mathlib/Algebra/BigOperators/Group/Finset.lean | 4 ++-- Mathlib/Algebra/CharZero/Quotient.lean | 6 +++--- Mathlib/Data/List/Cycle.lean | 6 +++--- Mathlib/Data/Multiset/Basic.lean | 3 +++ Mathlib/Data/Quot.lean | 6 ++---- Mathlib/Data/Set/Image.lean | 2 +- Mathlib/GroupTheory/Coset/Basic.lean | 4 ++-- Mathlib/GroupTheory/GroupAction/Quotient.lean | 6 +++--- Mathlib/GroupTheory/PGroup.lean | 2 +- 9 files changed, 20 insertions(+), 19 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index ed79af0319d28..1dbe92862bdb6 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1611,13 +1611,13 @@ theorem dvd_prod_of_mem (f : α → β) {a : α} {s : Finset α} (ha : a ∈ s) /-- A product can be partitioned into a product of products, each equivalent under a setoid. -/ @[to_additive "A sum can be partitioned into a sum of sums, each equivalent under a setoid."] theorem prod_partition (R : Setoid α) [DecidableRel R.r] : - ∏ x ∈ s, f x = ∏ xbar ∈ s.image Quotient.mk'', ∏ y ∈ s.filter (⟦·⟧ = xbar), f y := by + ∏ x ∈ s, f x = ∏ xbar ∈ s.image (Quotient.mk _), ∏ y ∈ s.filter (⟦·⟧ = xbar), f y := by refine (Finset.prod_image' f fun x _hx => ?_).symm rfl /-- If we can partition a product into subsets that cancel out, then the whole product cancels. -/ @[to_additive "If we can partition a sum into subsets that cancel out, then the whole sum cancels."] -theorem prod_cancels_of_partition_cancels (R : Setoid α) [DecidableRel R.r] +theorem prod_cancels_of_partition_cancels (R : Setoid α) [DecidableRel R] (h : ∀ x ∈ s, ∏ a ∈ s.filter fun y => R y x, f a = 1) : ∏ x ∈ s, f x = 1 := by rw [prod_partition R, ← Finset.prod_eq_one] intro xbar xbar_in_s diff --git a/Mathlib/Algebra/CharZero/Quotient.lean b/Mathlib/Algebra/CharZero/Quotient.lean index 8a44865763c7e..79b566ed37ce1 100644 --- a/Mathlib/Algebra/CharZero/Quotient.lean +++ b/Mathlib/Algebra/CharZero/Quotient.lean @@ -51,11 +51,11 @@ namespace QuotientAddGroup theorem zmultiples_zsmul_eq_zsmul_iff {ψ θ : R ⧸ AddSubgroup.zmultiples p} {z : ℤ} (hz : z ≠ 0) : z • ψ = z • θ ↔ ∃ k : Fin z.natAbs, ψ = θ + ((k : ℕ) • (p / z) : R) := by - induction ψ using Quotient.inductionOn' - induction θ using Quotient.inductionOn' + induction ψ using Quotient.inductionOn + induction θ using Quotient.inductionOn -- Porting note: Introduced Zp notation to shorten lines let Zp := AddSubgroup.zmultiples p - have : (Quotient.mk'' : R → R ⧸ Zp) = ((↑) : R → R ⧸ Zp) := rfl + have : (Quotient.mk _ : R → R ⧸ Zp) = ((↑) : R → R ⧸ Zp) := rfl simp only [this] simp_rw [← QuotientAddGroup.mk_zsmul, ← QuotientAddGroup.mk_add, QuotientAddGroup.eq_iff_sub_mem, ← smul_sub, ← sub_sub] diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 9c8cbaa143209..18607fde71b44 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -824,11 +824,11 @@ theorem chain_ne_nil (r : α → α → Prop) {l : List α} : theorem chain_map {β : Type*} {r : α → α → Prop} (f : β → α) {s : Cycle β} : Chain r (s.map f) ↔ Chain (fun a b => r (f a) (f b)) s := - Quotient.inductionOn' s fun l => by + Quotient.inductionOn s fun l => by cases' l with a l · rfl - dsimp only [Chain, ← mk''_eq_coe, Quotient.liftOn'_mk'', Cycle.map, Quotient.map', Quot.map, - Quotient.mk'', Quotient.liftOn', Quotient.liftOn, Quot.liftOn_mk, List.map] + dsimp only [Chain, Quotient.liftOn_mk, Cycle.map, Quotient.map', Quot.map, + Quotient.liftOn', Quotient.liftOn, Quot.liftOn_mk, List.map] rw [← concat_eq_append, ← List.map_concat, List.chain_map f] simp diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index e04ef7fa7a1dc..97366784110f6 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -67,6 +67,9 @@ theorem coe_eq_coe {l₁ l₂ : List α} : (l₁ : Multiset α) = l₂ ↔ l₁ instance [DecidableEq α] (l₁ l₂ : List α) : Decidable (l₁ ≈ l₂) := inferInstanceAs (Decidable (l₁ ~ l₂)) +instance [DecidableEq α] (l₁ l₂ : List α) : Decidable (isSetoid α l₁ l₂) := + inferInstanceAs (Decidable (l₁ ~ l₂)) + -- Porting note: `Quotient.recOnSubsingleton₂ s₁ s₂` was in parens which broke elaboration instance decidableEq [DecidableEq α] : DecidableEq (Multiset α) | s₁, s₂ => Quotient.recOnSubsingleton₂ s₁ s₂ fun _ _ => decidable_of_iff' _ Quotient.eq_iff_equiv diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 2bec860ba95b9..8dfbafe4dbf99 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -546,8 +546,8 @@ several different quotient relations on a type, for example quotient groups, rin -- Porting note: Quotient.mk' is the equivalent of Lean 3's `Quotient.mk` /-- A version of `Quotient.mk` taking `{s : Setoid α}` as an implicit argument instead of an instance argument. -/ -protected def mk'' (a : α) : Quotient s₁ := - Quot.mk s₁.1 a +protected abbrev mk'' (a : α) : Quotient s₁ := + ⟦a⟧ /-- `Quotient.mk''` is a surjective function. -/ theorem surjective_Quotient_mk'' : Function.Surjective (Quotient.mk'' : α → Quotient s₁) := @@ -693,7 +693,6 @@ protected theorem eq' {s₁ : Setoid α} {a b : α} : @Quotient.mk' α s₁ a = @Quotient.mk' α s₁ b ↔ s₁ a b := Quotient.eq -@[simp] protected theorem eq'' {a b : α} : @Quotient.mk'' α s₁ a = Quotient.mk'' b ↔ s₁ a b := Quotient.eq @@ -725,7 +724,6 @@ protected theorem liftOn₂'_mk {t : Setoid β} (f : α → β → γ) (h) (a : 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 : α) : (Quotient.mk s x).map' f h = (Quotient.mk t (f x)) := rfl diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 00387c99f2202..1da452c8964d3 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -836,7 +836,7 @@ theorem range_quotient_lift [s : Setoid ι] (hf) : theorem range_quotient_mk' {s : Setoid α} : range (Quotient.mk' : α → Quotient s) = univ := range_quot_mk _ -@[simp] lemma Quotient.range_mk'' {sa : Setoid α} : range (Quotient.mk'' (s₁ := sa)) = univ := +lemma Quotient.range_mk'' {sa : Setoid α} : range (Quotient.mk'' (s₁ := sa)) = univ := range_quotient_mk @[simp] diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index e7aa809a354ae..c5df256757ef5 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -491,7 +491,7 @@ open MulAction in lemma orbit_mk_eq_smul (x : α) : MulAction.orbitRel.Quotient.orbit (x : α ⧸ s) = x • s := by ext rw [orbitRel.Quotient.mem_orbit] - simpa [mem_smul_set_iff_inv_smul_mem, ← leftRel_apply] using Setoid.comm' _ + simpa [mem_smul_set_iff_inv_smul_mem, ← leftRel_apply, Quotient.eq''] using Setoid.comm' _ @[to_additive] lemma orbit_eq_out'_smul (x : α ⧸ s) : MulAction.orbitRel.Quotient.orbit x = x.out' • s := by @@ -530,7 +530,7 @@ noncomputable def groupEquivQuotientProdSubgroup : α ≃ (α ⧸ s) × s := show (_root_.Subtype fun x : α => Quotient.mk'' x = L) ≃ _root_.Subtype fun x : α => Quotient.mk'' x = Quotient.mk'' _ - simp [-Quotient.eq''] + simp rfl _ ≃ Σ _L : α ⧸ s, s := Equiv.sigmaCongrRight fun _ => leftCosetEquivSubgroup _ _ ≃ (α ⧸ s) × s := Equiv.sigmaEquivProd _ _ diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index f58416df124ac..e532945e65310 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -320,10 +320,10 @@ noncomputable def equivSubgroupOrbitsSetoidComap (H : Subgroup α) (ω : Ω) : toFun := fun q ↦ q.liftOn' (fun x ↦ ⟦⟨↑x, by simp only [Set.mem_preimage, Set.mem_singleton_iff] have hx := x.property - rwa [orbitRel.Quotient.mem_orbit, @Quotient.mk''_eq_mk] at hx⟩⟧) fun a b h ↦ by - simp only [← Quotient.eq'', Quotient.mk''_eq_mk, + rwa [orbitRel.Quotient.mem_orbit] at hx⟩⟧) fun a b h ↦ by + simp only [← Quotient.eq, orbitRel.Quotient.subgroup_quotient_eq_iff] at h - simp only [← Quotient.mk''_eq_mk, Quotient.eq''] at h ⊢ + simp only [Quotient.eq] at h ⊢ exact h invFun := fun q ↦ q.liftOn' (fun x ↦ ⟦⟨↑x, by have hx := x.property diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index d882d7e22ca40..eb9cdfbf182a9 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -179,7 +179,7 @@ theorem card_modEq_card_fixedPoints : Nat.card α ≡ Nat.card (fixedPoints G α rw [Nat.card_eq_fintype_card] at hk have : k = 0 := by contrapose! hb - simp [-Quotient.eq'', key, hk, hb] + simp [-Quotient.eq, key, hk, hb] exact ⟨⟨b, mem_fixedPoints_iff_card_orbit_eq_one.2 <| by rw [hk, this, pow_zero]⟩, Finset.mem_univ _, ne_of_eq_of_ne Nat.cast_one one_ne_zero, rfl⟩ From 64ac27ea68f1e00fc29e2ce69a1c73c348dcdfb4 Mon Sep 17 00:00:00 2001 From: Paul Lezeau Date: Sat, 19 Oct 2024 16:03:51 +0000 Subject: [PATCH 106/131] chore(Order/Bounds/Basic): Split long file into two shorter files (#16539) This PR transfers some of the content from `Order/Bounds/Basic` into a new file, `Order/Bounds/Image`. --- Mathlib.lean | 1 + Mathlib/Order/Bounds/Basic.lean | 482 --------------------------- Mathlib/Order/Bounds/Image.lean | 498 ++++++++++++++++++++++++++++ Mathlib/Order/Bounds/OrderIso.lean | 2 +- Mathlib/Order/CompleteLattice.lean | 1 + Mathlib/Order/GaloisConnection.lean | 2 +- 6 files changed, 502 insertions(+), 484 deletions(-) create mode 100644 Mathlib/Order/Bounds/Image.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4c5d00a1fa732..5fcbb05388323 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3621,6 +3621,7 @@ import Mathlib.Order.Bounded import Mathlib.Order.BoundedOrder import Mathlib.Order.Bounds.Basic import Mathlib.Order.Bounds.Defs +import Mathlib.Order.Bounds.Image import Mathlib.Order.Bounds.OrderIso import Mathlib.Order.Category.BddDistLat import Mathlib.Order.Category.BddLat diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index fb6e31a954dff..66d0f19bb54b8 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -3,7 +3,6 @@ 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, Yury Kudryashov -/ -import Mathlib.Data.Set.NAry import Mathlib.Order.Bounds.Defs import Mathlib.Order.Directed import Mathlib.Order.Interval.Set.Basic @@ -905,484 +904,3 @@ theorem IsGLB.exists_between' (h : IsGLB s a) (h' : a ∉ s) (hb : a < b) : ∃ ⟨c, hcs, hac.lt_of_ne fun hac => h' <| hac.symm ▸ hcs, hcb⟩ end LinearOrder - -/-! -### Images of upper/lower bounds under monotone functions --/ - - -namespace MonotoneOn - -variable [Preorder α] [Preorder β] {f : α → β} {s t : Set α} {a : α} - -theorem mem_upperBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) (Has : a ∈ upperBounds s) - (Hat : a ∈ t) : f a ∈ upperBounds (f '' s) := - forall_mem_image.2 fun _ H => Hf (Hst H) Hat (Has H) - -theorem mem_upperBounds_image_self (Hf : MonotoneOn f t) : - a ∈ upperBounds t → a ∈ t → f a ∈ upperBounds (f '' t) := - Hf.mem_upperBounds_image subset_rfl - -theorem mem_lowerBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) (Has : a ∈ lowerBounds s) - (Hat : a ∈ t) : f a ∈ lowerBounds (f '' s) := - forall_mem_image.2 fun _ H => Hf Hat (Hst H) (Has H) - -theorem mem_lowerBounds_image_self (Hf : MonotoneOn f t) : - a ∈ lowerBounds t → a ∈ t → f a ∈ lowerBounds (f '' t) := - Hf.mem_lowerBounds_image subset_rfl - -theorem image_upperBounds_subset_upperBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) : - f '' (upperBounds s ∩ t) ⊆ upperBounds (f '' s) := by - rintro _ ⟨a, ha, rfl⟩ - exact Hf.mem_upperBounds_image Hst ha.1 ha.2 - -theorem image_lowerBounds_subset_lowerBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) : - f '' (lowerBounds s ∩ t) ⊆ lowerBounds (f '' s) := - Hf.dual.image_upperBounds_subset_upperBounds_image Hst - -/-- The image under a monotone function on a set `t` of a subset which has an upper bound in `t` - is bounded above. -/ -theorem map_bddAbove (Hf : MonotoneOn f t) (Hst : s ⊆ t) : - (upperBounds s ∩ t).Nonempty → BddAbove (f '' s) := fun ⟨C, hs, ht⟩ => - ⟨f C, Hf.mem_upperBounds_image Hst hs ht⟩ - -/-- The image under a monotone function on a set `t` of a subset which has a lower bound in `t` - is bounded below. -/ -theorem map_bddBelow (Hf : MonotoneOn f t) (Hst : s ⊆ t) : - (lowerBounds s ∩ t).Nonempty → BddBelow (f '' s) := fun ⟨C, hs, ht⟩ => - ⟨f C, Hf.mem_lowerBounds_image Hst hs ht⟩ - -/-- A monotone map sends a least element of a set to a least element of its image. -/ -theorem map_isLeast (Hf : MonotoneOn f t) (Ha : IsLeast t a) : IsLeast (f '' t) (f a) := - ⟨mem_image_of_mem _ Ha.1, Hf.mem_lowerBounds_image_self Ha.2 Ha.1⟩ - -/-- A monotone map sends a greatest element of a set to a greatest element of its image. -/ -theorem map_isGreatest (Hf : MonotoneOn f t) (Ha : IsGreatest t a) : IsGreatest (f '' t) (f a) := - ⟨mem_image_of_mem _ Ha.1, Hf.mem_upperBounds_image_self Ha.2 Ha.1⟩ - -end MonotoneOn - -namespace AntitoneOn - -variable [Preorder α] [Preorder β] {f : α → β} {s t : Set α} {a : α} - -theorem mem_upperBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) (Has : a ∈ lowerBounds s) : - a ∈ t → f a ∈ upperBounds (f '' s) := - Hf.dual_right.mem_lowerBounds_image Hst Has - -theorem mem_upperBounds_image_self (Hf : AntitoneOn f t) : - a ∈ lowerBounds t → a ∈ t → f a ∈ upperBounds (f '' t) := - Hf.dual_right.mem_lowerBounds_image_self - -theorem mem_lowerBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : - a ∈ upperBounds s → a ∈ t → f a ∈ lowerBounds (f '' s) := - Hf.dual_right.mem_upperBounds_image Hst - -theorem mem_lowerBounds_image_self (Hf : AntitoneOn f t) : - a ∈ upperBounds t → a ∈ t → f a ∈ lowerBounds (f '' t) := - Hf.dual_right.mem_upperBounds_image_self - -theorem image_lowerBounds_subset_upperBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : - f '' (lowerBounds s ∩ t) ⊆ upperBounds (f '' s) := - Hf.dual_right.image_lowerBounds_subset_lowerBounds_image Hst - -theorem image_upperBounds_subset_lowerBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : - f '' (upperBounds s ∩ t) ⊆ lowerBounds (f '' s) := - Hf.dual_right.image_upperBounds_subset_upperBounds_image Hst - -/-- The image under an antitone function of a set which is bounded above is bounded below. -/ -theorem map_bddAbove (Hf : AntitoneOn f t) (Hst : s ⊆ t) : - (upperBounds s ∩ t).Nonempty → BddBelow (f '' s) := - Hf.dual_right.map_bddAbove Hst - -/-- The image under an antitone function of a set which is bounded below is bounded above. -/ -theorem map_bddBelow (Hf : AntitoneOn f t) (Hst : s ⊆ t) : - (lowerBounds s ∩ t).Nonempty → BddAbove (f '' s) := - Hf.dual_right.map_bddBelow Hst - -/-- An antitone map sends a greatest element of a set to a least element of its image. -/ -theorem map_isGreatest (Hf : AntitoneOn f t) : IsGreatest t a → IsLeast (f '' t) (f a) := - Hf.dual_right.map_isGreatest - -/-- An antitone map sends a least element of a set to a greatest element of its image. -/ -theorem map_isLeast (Hf : AntitoneOn f t) : IsLeast t a → IsGreatest (f '' t) (f a) := - Hf.dual_right.map_isLeast - -end AntitoneOn - -namespace Monotone - -variable [Preorder α] [Preorder β] {f : α → β} (Hf : Monotone f) {a : α} {s : Set α} - -include Hf - -theorem mem_upperBounds_image (Ha : a ∈ upperBounds s) : f a ∈ upperBounds (f '' s) := - forall_mem_image.2 fun _ H => Hf (Ha H) - -theorem mem_lowerBounds_image (Ha : a ∈ lowerBounds s) : f a ∈ lowerBounds (f '' s) := - forall_mem_image.2 fun _ H => Hf (Ha H) - -theorem image_upperBounds_subset_upperBounds_image : - f '' upperBounds s ⊆ upperBounds (f '' s) := by - rintro _ ⟨a, ha, rfl⟩ - exact Hf.mem_upperBounds_image ha - -theorem image_lowerBounds_subset_lowerBounds_image : f '' lowerBounds s ⊆ lowerBounds (f '' s) := - Hf.dual.image_upperBounds_subset_upperBounds_image - -/-- The image under a monotone function of a set which is bounded above is bounded above. See also -`BddAbove.image2`. -/ -theorem map_bddAbove : BddAbove s → BddAbove (f '' s) - | ⟨C, hC⟩ => ⟨f C, Hf.mem_upperBounds_image hC⟩ - -/-- The image under a monotone function of a set which is bounded below is bounded below. See also -`BddBelow.image2`. -/ -theorem map_bddBelow : BddBelow s → BddBelow (f '' s) - | ⟨C, hC⟩ => ⟨f C, Hf.mem_lowerBounds_image hC⟩ - -/-- A monotone map sends a least element of a set to a least element of its image. -/ -theorem map_isLeast (Ha : IsLeast s a) : IsLeast (f '' s) (f a) := - ⟨mem_image_of_mem _ Ha.1, Hf.mem_lowerBounds_image Ha.2⟩ - -/-- A monotone map sends a greatest element of a set to a greatest element of its image. -/ -theorem map_isGreatest (Ha : IsGreatest s a) : IsGreatest (f '' s) (f a) := - ⟨mem_image_of_mem _ Ha.1, Hf.mem_upperBounds_image Ha.2⟩ - -end Monotone - -namespace Antitone - -variable [Preorder α] [Preorder β] {f : α → β} (hf : Antitone f) {a : α} {s : Set α} - -include hf - -theorem mem_upperBounds_image : a ∈ lowerBounds s → f a ∈ upperBounds (f '' s) := - hf.dual_right.mem_lowerBounds_image - -theorem mem_lowerBounds_image : a ∈ upperBounds s → f a ∈ lowerBounds (f '' s) := - hf.dual_right.mem_upperBounds_image - -theorem image_lowerBounds_subset_upperBounds_image : f '' lowerBounds s ⊆ upperBounds (f '' s) := - hf.dual_right.image_lowerBounds_subset_lowerBounds_image - -theorem image_upperBounds_subset_lowerBounds_image : f '' upperBounds s ⊆ lowerBounds (f '' s) := - hf.dual_right.image_upperBounds_subset_upperBounds_image - -/-- The image under an antitone function of a set which is bounded above is bounded below. -/ -theorem map_bddAbove : BddAbove s → BddBelow (f '' s) := - hf.dual_right.map_bddAbove - -/-- The image under an antitone function of a set which is bounded below is bounded above. -/ -theorem map_bddBelow : BddBelow s → BddAbove (f '' s) := - hf.dual_right.map_bddBelow - -/-- An antitone map sends a greatest element of a set to a least element of its image. -/ -theorem map_isGreatest : IsGreatest s a → IsLeast (f '' s) (f a) := - hf.dual_right.map_isGreatest - -/-- An antitone map sends a least element of a set to a greatest element of its image. -/ -theorem map_isLeast : IsLeast s a → IsGreatest (f '' s) (f a) := - hf.dual_right.map_isLeast - -end Antitone - -section Image2 - -variable [Preorder α] [Preorder β] [Preorder γ] {f : α → β → γ} {s : Set α} {t : Set β} {a : α} - {b : β} - -section MonotoneMonotone - -variable (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Monotone (f a)) - -include h₀ h₁ - -theorem mem_upperBounds_image2 (ha : a ∈ upperBounds s) (hb : b ∈ upperBounds t) : - f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem mem_lowerBounds_image2 (ha : a ∈ lowerBounds s) (hb : b ∈ lowerBounds t) : - f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem image2_upperBounds_upperBounds_subset : - image2 f (upperBounds s) (upperBounds t) ⊆ upperBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ mem_upperBounds_image2 h₀ h₁ ha hb - -theorem image2_lowerBounds_lowerBounds_subset : - image2 f (lowerBounds s) (lowerBounds t) ⊆ lowerBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ mem_lowerBounds_image2 h₀ h₁ ha hb - -/-- See also `Monotone.map_bddAbove`. -/ -protected theorem BddAbove.image2 : - BddAbove s → BddAbove t → BddAbove (image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_upperBounds_image2 h₀ h₁ ha hb⟩ - -/-- See also `Monotone.map_bddBelow`. -/ -protected theorem BddBelow.image2 : - BddBelow s → BddBelow t → BddBelow (image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_lowerBounds_image2 h₀ h₁ ha hb⟩ - -protected theorem IsGreatest.image2 (ha : IsGreatest s a) (hb : IsGreatest t b) : - IsGreatest (image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, mem_upperBounds_image2 h₀ h₁ ha.2 hb.2⟩ - -protected theorem IsLeast.image2 (ha : IsLeast s a) (hb : IsLeast t b) : - IsLeast (image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, mem_lowerBounds_image2 h₀ h₁ ha.2 hb.2⟩ - -end MonotoneMonotone - -section MonotoneAntitone - -variable (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Antitone (f a)) - -include h₀ h₁ - -theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds (ha : a ∈ upperBounds s) - (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds (ha : a ∈ lowerBounds s) - (hb : b ∈ upperBounds t) : f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem image2_upperBounds_lowerBounds_subset_upperBounds_image2 : - image2 f (upperBounds s) (lowerBounds t) ⊆ upperBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha hb - -theorem image2_lowerBounds_upperBounds_subset_lowerBounds_image2 : - image2 f (lowerBounds s) (upperBounds t) ⊆ lowerBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha hb - -theorem BddAbove.bddAbove_image2_of_bddBelow : - BddAbove s → BddBelow t → BddAbove (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha hb⟩ - -theorem BddBelow.bddBelow_image2_of_bddAbove : - BddBelow s → BddAbove t → BddBelow (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha hb⟩ - -theorem IsGreatest.isGreatest_image2_of_isLeast (ha : IsGreatest s a) (hb : IsLeast t b) : - IsGreatest (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, - mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ - -theorem IsLeast.isLeast_image2_of_isGreatest (ha : IsLeast s a) (hb : IsGreatest t b) : - IsLeast (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, - mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ - -end MonotoneAntitone - -section AntitoneAntitone - -variable (h₀ : ∀ b, Antitone (swap f b)) (h₁ : ∀ a, Antitone (f a)) - -include h₀ h₁ - -theorem mem_upperBounds_image2_of_mem_lowerBounds (ha : a ∈ lowerBounds s) - (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem mem_lowerBounds_image2_of_mem_upperBounds (ha : a ∈ upperBounds s) - (hb : b ∈ upperBounds t) : f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem image2_upperBounds_upperBounds_subset_upperBounds_image2 : - image2 f (lowerBounds s) (lowerBounds t) ⊆ upperBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha hb - -theorem image2_lowerBounds_lowerBounds_subset_lowerBounds_image2 : - image2 f (upperBounds s) (upperBounds t) ⊆ lowerBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha hb - -theorem BddBelow.image2_bddAbove : BddBelow s → BddBelow t → BddAbove (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha hb⟩ - -theorem BddAbove.image2_bddBelow : BddAbove s → BddAbove t → BddBelow (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha hb⟩ - -theorem IsLeast.isGreatest_image2 (ha : IsLeast s a) (hb : IsLeast t b) : - IsGreatest (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ - -theorem IsGreatest.isLeast_image2 (ha : IsGreatest s a) (hb : IsGreatest t b) : - IsLeast (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ - -end AntitoneAntitone - -section AntitoneMonotone - -variable (h₀ : ∀ b, Antitone (swap f b)) (h₁ : ∀ a, Monotone (f a)) - -include h₀ h₁ - -theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds (ha : a ∈ lowerBounds s) - (hb : b ∈ upperBounds t) : f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds (ha : a ∈ upperBounds s) - (hb : b ∈ lowerBounds t) : f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy - -theorem image2_lowerBounds_upperBounds_subset_upperBounds_image2 : - image2 f (lowerBounds s) (upperBounds t) ⊆ upperBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha hb - -theorem image2_upperBounds_lowerBounds_subset_lowerBounds_image2 : - image2 f (upperBounds s) (lowerBounds t) ⊆ lowerBounds (image2 f s t) := - image2_subset_iff.2 fun _ ha _ hb ↦ - mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha hb - -theorem BddBelow.bddAbove_image2_of_bddAbove : - BddBelow s → BddAbove t → BddAbove (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha hb⟩ - -theorem BddAbove.bddBelow_image2_of_bddAbove : - BddAbove s → BddBelow t → BddBelow (Set.image2 f s t) := by - rintro ⟨a, ha⟩ ⟨b, hb⟩ - exact ⟨f a b, mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha hb⟩ - -theorem IsLeast.isGreatest_image2_of_isGreatest (ha : IsLeast s a) (hb : IsGreatest t b) : - IsGreatest (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, - mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ - -theorem IsGreatest.isLeast_image2_of_isLeast (ha : IsGreatest s a) (hb : IsLeast t b) : - IsLeast (Set.image2 f s t) (f a b) := - ⟨mem_image2_of_mem ha.1 hb.1, - mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ - -end AntitoneMonotone - -end Image2 - -section Prod - -variable {α β : Type*} [Preorder α] [Preorder β] - -lemma bddAbove_prod {s : Set (α × β)} : - BddAbove s ↔ BddAbove (Prod.fst '' s) ∧ BddAbove (Prod.snd '' s) := - ⟨fun ⟨p, hp⟩ ↦ ⟨⟨p.1, forall_mem_image.2 fun _q hq ↦ (hp hq).1⟩, - ⟨p.2, forall_mem_image.2 fun _q hq ↦ (hp hq).2⟩⟩, - fun ⟨⟨x, hx⟩, ⟨y, hy⟩⟩ ↦ ⟨⟨x, y⟩, fun _p hp ↦ - ⟨hx <| mem_image_of_mem _ hp, hy <| mem_image_of_mem _ hp⟩⟩⟩ - -lemma bddBelow_prod {s : Set (α × β)} : - BddBelow s ↔ BddBelow (Prod.fst '' s) ∧ BddBelow (Prod.snd '' s) := - bddAbove_prod (α := αᵒᵈ) (β := βᵒᵈ) - -lemma bddAbove_range_prod {F : ι → α × β} : - BddAbove (range F) ↔ BddAbove (range <| Prod.fst ∘ F) ∧ BddAbove (range <| Prod.snd ∘ F) := by - simp only [bddAbove_prod, ← range_comp] - -lemma bddBelow_range_prod {F : ι → α × β} : - BddBelow (range F) ↔ BddBelow (range <| Prod.fst ∘ F) ∧ BddBelow (range <| Prod.snd ∘ F) := - bddAbove_range_prod (α := αᵒᵈ) (β := βᵒᵈ) - -theorem isLUB_prod {s : Set (α × β)} (p : α × β) : - IsLUB s p ↔ IsLUB (Prod.fst '' s) p.1 ∧ IsLUB (Prod.snd '' s) p.2 := by - refine - ⟨fun H => - ⟨⟨monotone_fst.mem_upperBounds_image H.1, fun a ha => ?_⟩, - ⟨monotone_snd.mem_upperBounds_image H.1, fun a ha => ?_⟩⟩, - fun H => ⟨?_, ?_⟩⟩ - · suffices h : (a, p.2) ∈ upperBounds s from (H.2 h).1 - exact fun q hq => ⟨ha <| mem_image_of_mem _ hq, (H.1 hq).2⟩ - · suffices h : (p.1, a) ∈ upperBounds s from (H.2 h).2 - exact fun q hq => ⟨(H.1 hq).1, ha <| mem_image_of_mem _ hq⟩ - · exact fun q hq => ⟨H.1.1 <| mem_image_of_mem _ hq, H.2.1 <| mem_image_of_mem _ hq⟩ - · exact fun q hq => - ⟨H.1.2 <| monotone_fst.mem_upperBounds_image hq, - H.2.2 <| monotone_snd.mem_upperBounds_image hq⟩ - -theorem isGLB_prod {s : Set (α × β)} (p : α × β) : - IsGLB s p ↔ IsGLB (Prod.fst '' s) p.1 ∧ IsGLB (Prod.snd '' s) p.2 := - @isLUB_prod αᵒᵈ βᵒᵈ _ _ _ _ - -end Prod - - -section Pi - -variable {π : α → Type*} [∀ a, Preorder (π a)] - -lemma bddAbove_pi {s : Set (∀ a, π a)} : - BddAbove s ↔ ∀ a, BddAbove (Function.eval a '' s) := - ⟨fun ⟨f, hf⟩ a ↦ ⟨f a, forall_mem_image.2 fun _ hg ↦ hf hg a⟩, - fun h ↦ ⟨fun a ↦ (h a).some, fun _ hg a ↦ (h a).some_mem <| mem_image_of_mem _ hg⟩⟩ - -lemma bddBelow_pi {s : Set (∀ a, π a)} : - BddBelow s ↔ ∀ a, BddBelow (Function.eval a '' s) := - bddAbove_pi (π := fun a ↦ (π a)ᵒᵈ) - -lemma bddAbove_range_pi {F : ι → ∀ a, π a} : - BddAbove (range F) ↔ ∀ a, BddAbove (range fun i ↦ F i a) := by - simp only [bddAbove_pi, ← range_comp] - rfl - -lemma bddBelow_range_pi {F : ι → ∀ a, π a} : - BddBelow (range F) ↔ ∀ a, BddBelow (range fun i ↦ F i a) := - bddAbove_range_pi (π := fun a ↦ (π a)ᵒᵈ) - -theorem isLUB_pi {s : Set (∀ a, π a)} {f : ∀ a, π a} : - IsLUB s f ↔ ∀ a, IsLUB (Function.eval a '' s) (f a) := by - classical - refine - ⟨fun H a => ⟨(Function.monotone_eval a).mem_upperBounds_image H.1, fun b hb => ?_⟩, fun H => - ⟨?_, ?_⟩⟩ - · suffices h : Function.update f a b ∈ upperBounds s from Function.update_same a b f ▸ H.2 h a - exact fun g hg => le_update_iff.2 ⟨hb <| mem_image_of_mem _ hg, fun i _ => H.1 hg i⟩ - · exact fun g hg a => (H a).1 (mem_image_of_mem _ hg) - · exact fun g hg a => (H a).2 ((Function.monotone_eval a).mem_upperBounds_image hg) - -theorem isGLB_pi {s : Set (∀ a, π a)} {f : ∀ a, π a} : - IsGLB s f ↔ ∀ a, IsGLB (Function.eval a '' s) (f a) := - @isLUB_pi α (fun a => (π a)ᵒᵈ) _ s f - -end Pi - -theorem IsGLB.of_image [Preorder α] [Preorder β] {f : α → β} (hf : ∀ {x y}, f x ≤ f y ↔ x ≤ y) - {s : Set α} {x : α} (hx : IsGLB (f '' s) (f x)) : IsGLB s x := - ⟨fun _ hy => hf.1 <| hx.1 <| mem_image_of_mem _ hy, fun _ hy => - hf.1 <| hx.2 <| Monotone.mem_lowerBounds_image (fun _ _ => hf.2) hy⟩ - -theorem IsLUB.of_image [Preorder α] [Preorder β] {f : α → β} (hf : ∀ {x y}, f x ≤ f y ↔ x ≤ y) - {s : Set α} {x : α} (hx : IsLUB (f '' s) (f x)) : IsLUB s x := - ⟨fun _ hy => hf.1 <| hx.1 <| mem_image_of_mem _ hy, fun _ hy => - hf.1 <| hx.2 <| Monotone.mem_upperBounds_image (fun _ _ => hf.2) hy⟩ - -lemma BddAbove.range_mono [Preorder β] {f : α → β} (g : α → β) (h : ∀ a, f a ≤ g a) - (hbdd : BddAbove (range g)) : BddAbove (range f) := by - obtain ⟨C, hC⟩ := hbdd - use C - rintro - ⟨x, rfl⟩ - exact (h x).trans (hC <| mem_range_self x) - -lemma BddBelow.range_mono [Preorder β] (f : α → β) {g : α → β} (h : ∀ a, f a ≤ g a) - (hbdd : BddBelow (range f)) : BddBelow (range g) := - BddAbove.range_mono (β := βᵒᵈ) f h hbdd - -lemma BddAbove.range_comp {γ : Type*} [Preorder β] [Preorder γ] {f : α → β} {g : β → γ} - (hf : BddAbove (range f)) (hg : Monotone g) : BddAbove (range (fun x => g (f x))) := by - change BddAbove (range (g ∘ f)) - simpa only [Set.range_comp] using hg.map_bddAbove hf - -lemma BddBelow.range_comp {γ : Type*} [Preorder β] [Preorder γ] {f : α → β} {g : β → γ} - (hf : BddBelow (range f)) (hg : Monotone g) : BddBelow (range (fun x => g (f x))) := by - change BddBelow (range (g ∘ f)) - simpa only [Set.range_comp] using hg.map_bddBelow hf diff --git a/Mathlib/Order/Bounds/Image.lean b/Mathlib/Order/Bounds/Image.lean new file mode 100644 index 0000000000000..910c466332520 --- /dev/null +++ b/Mathlib/Order/Bounds/Image.lean @@ -0,0 +1,498 @@ +/- +Copyright (c) 2017 Paul Lezeau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Yury Kudryashov, Paul Lezeau +-/ +import Mathlib.Data.Set.NAry +import Mathlib.Order.Bounds.Defs + +/-! + +# Images of upper/lower bounds under monotone functions + +In this file we prove various results about the behaviour of bounds under monotone/antitone maps. +-/ + +open Function Set + +open OrderDual (toDual ofDual) + +universe u v w x + +variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} + +namespace MonotoneOn + +variable [Preorder α] [Preorder β] {f : α → β} {s t : Set α} {a : α} + +theorem mem_upperBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) (Has : a ∈ upperBounds s) + (Hat : a ∈ t) : f a ∈ upperBounds (f '' s) := + forall_mem_image.2 fun _ H => Hf (Hst H) Hat (Has H) + +theorem mem_upperBounds_image_self (Hf : MonotoneOn f t) : + a ∈ upperBounds t → a ∈ t → f a ∈ upperBounds (f '' t) := + Hf.mem_upperBounds_image subset_rfl + +theorem mem_lowerBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) (Has : a ∈ lowerBounds s) + (Hat : a ∈ t) : f a ∈ lowerBounds (f '' s) := + forall_mem_image.2 fun _ H => Hf Hat (Hst H) (Has H) + +theorem mem_lowerBounds_image_self (Hf : MonotoneOn f t) : + a ∈ lowerBounds t → a ∈ t → f a ∈ lowerBounds (f '' t) := + Hf.mem_lowerBounds_image subset_rfl + +theorem image_upperBounds_subset_upperBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) : + f '' (upperBounds s ∩ t) ⊆ upperBounds (f '' s) := by + rintro _ ⟨a, ha, rfl⟩ + exact Hf.mem_upperBounds_image Hst ha.1 ha.2 + +theorem image_lowerBounds_subset_lowerBounds_image (Hf : MonotoneOn f t) (Hst : s ⊆ t) : + f '' (lowerBounds s ∩ t) ⊆ lowerBounds (f '' s) := + Hf.dual.image_upperBounds_subset_upperBounds_image Hst + +/-- The image under a monotone function on a set `t` of a subset which has an upper bound in `t` + is bounded above. -/ +theorem map_bddAbove (Hf : MonotoneOn f t) (Hst : s ⊆ t) : + (upperBounds s ∩ t).Nonempty → BddAbove (f '' s) := fun ⟨C, hs, ht⟩ => + ⟨f C, Hf.mem_upperBounds_image Hst hs ht⟩ + +/-- The image under a monotone function on a set `t` of a subset which has a lower bound in `t` + is bounded below. -/ +theorem map_bddBelow (Hf : MonotoneOn f t) (Hst : s ⊆ t) : + (lowerBounds s ∩ t).Nonempty → BddBelow (f '' s) := fun ⟨C, hs, ht⟩ => + ⟨f C, Hf.mem_lowerBounds_image Hst hs ht⟩ + +/-- A monotone map sends a least element of a set to a least element of its image. -/ +theorem map_isLeast (Hf : MonotoneOn f t) (Ha : IsLeast t a) : IsLeast (f '' t) (f a) := + ⟨mem_image_of_mem _ Ha.1, Hf.mem_lowerBounds_image_self Ha.2 Ha.1⟩ + +/-- A monotone map sends a greatest element of a set to a greatest element of its image. -/ +theorem map_isGreatest (Hf : MonotoneOn f t) (Ha : IsGreatest t a) : IsGreatest (f '' t) (f a) := + ⟨mem_image_of_mem _ Ha.1, Hf.mem_upperBounds_image_self Ha.2 Ha.1⟩ + +end MonotoneOn + +namespace AntitoneOn + +variable [Preorder α] [Preorder β] {f : α → β} {s t : Set α} {a : α} + +theorem mem_upperBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) (Has : a ∈ lowerBounds s) : + a ∈ t → f a ∈ upperBounds (f '' s) := + Hf.dual_right.mem_lowerBounds_image Hst Has + +theorem mem_upperBounds_image_self (Hf : AntitoneOn f t) : + a ∈ lowerBounds t → a ∈ t → f a ∈ upperBounds (f '' t) := + Hf.dual_right.mem_lowerBounds_image_self + +theorem mem_lowerBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : + a ∈ upperBounds s → a ∈ t → f a ∈ lowerBounds (f '' s) := + Hf.dual_right.mem_upperBounds_image Hst + +theorem mem_lowerBounds_image_self (Hf : AntitoneOn f t) : + a ∈ upperBounds t → a ∈ t → f a ∈ lowerBounds (f '' t) := + Hf.dual_right.mem_upperBounds_image_self + +theorem image_lowerBounds_subset_upperBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : + f '' (lowerBounds s ∩ t) ⊆ upperBounds (f '' s) := + Hf.dual_right.image_lowerBounds_subset_lowerBounds_image Hst + +theorem image_upperBounds_subset_lowerBounds_image (Hf : AntitoneOn f t) (Hst : s ⊆ t) : + f '' (upperBounds s ∩ t) ⊆ lowerBounds (f '' s) := + Hf.dual_right.image_upperBounds_subset_upperBounds_image Hst + +/-- The image under an antitone function of a set which is bounded above is bounded below. -/ +theorem map_bddAbove (Hf : AntitoneOn f t) (Hst : s ⊆ t) : + (upperBounds s ∩ t).Nonempty → BddBelow (f '' s) := + Hf.dual_right.map_bddAbove Hst + +/-- The image under an antitone function of a set which is bounded below is bounded above. -/ +theorem map_bddBelow (Hf : AntitoneOn f t) (Hst : s ⊆ t) : + (lowerBounds s ∩ t).Nonempty → BddAbove (f '' s) := + Hf.dual_right.map_bddBelow Hst + +/-- An antitone map sends a greatest element of a set to a least element of its image. -/ +theorem map_isGreatest (Hf : AntitoneOn f t) : IsGreatest t a → IsLeast (f '' t) (f a) := + Hf.dual_right.map_isGreatest + +/-- An antitone map sends a least element of a set to a greatest element of its image. -/ +theorem map_isLeast (Hf : AntitoneOn f t) : IsLeast t a → IsGreatest (f '' t) (f a) := + Hf.dual_right.map_isLeast + +end AntitoneOn + +namespace Monotone + +variable [Preorder α] [Preorder β] {f : α → β} (Hf : Monotone f) {a : α} {s : Set α} + +include Hf + +theorem mem_upperBounds_image (Ha : a ∈ upperBounds s) : f a ∈ upperBounds (f '' s) := + forall_mem_image.2 fun _ H => Hf (Ha H) + +theorem mem_lowerBounds_image (Ha : a ∈ lowerBounds s) : f a ∈ lowerBounds (f '' s) := + forall_mem_image.2 fun _ H => Hf (Ha H) + +theorem image_upperBounds_subset_upperBounds_image : + f '' upperBounds s ⊆ upperBounds (f '' s) := by + rintro _ ⟨a, ha, rfl⟩ + exact Hf.mem_upperBounds_image ha + +theorem image_lowerBounds_subset_lowerBounds_image : f '' lowerBounds s ⊆ lowerBounds (f '' s) := + Hf.dual.image_upperBounds_subset_upperBounds_image + +/-- The image under a monotone function of a set which is bounded above is bounded above. See also +`BddAbove.image2`. -/ +theorem map_bddAbove : BddAbove s → BddAbove (f '' s) + | ⟨C, hC⟩ => ⟨f C, Hf.mem_upperBounds_image hC⟩ + +/-- The image under a monotone function of a set which is bounded below is bounded below. See also +`BddBelow.image2`. -/ +theorem map_bddBelow : BddBelow s → BddBelow (f '' s) + | ⟨C, hC⟩ => ⟨f C, Hf.mem_lowerBounds_image hC⟩ + +/-- A monotone map sends a least element of a set to a least element of its image. -/ +theorem map_isLeast (Ha : IsLeast s a) : IsLeast (f '' s) (f a) := + ⟨mem_image_of_mem _ Ha.1, Hf.mem_lowerBounds_image Ha.2⟩ + +/-- A monotone map sends a greatest element of a set to a greatest element of its image. -/ +theorem map_isGreatest (Ha : IsGreatest s a) : IsGreatest (f '' s) (f a) := + ⟨mem_image_of_mem _ Ha.1, Hf.mem_upperBounds_image Ha.2⟩ + +end Monotone + +namespace Antitone + +variable [Preorder α] [Preorder β] {f : α → β} (hf : Antitone f) {a : α} {s : Set α} + +include hf + +theorem mem_upperBounds_image : a ∈ lowerBounds s → f a ∈ upperBounds (f '' s) := + hf.dual_right.mem_lowerBounds_image + +theorem mem_lowerBounds_image : a ∈ upperBounds s → f a ∈ lowerBounds (f '' s) := + hf.dual_right.mem_upperBounds_image + +theorem image_lowerBounds_subset_upperBounds_image : f '' lowerBounds s ⊆ upperBounds (f '' s) := + hf.dual_right.image_lowerBounds_subset_lowerBounds_image + +theorem image_upperBounds_subset_lowerBounds_image : f '' upperBounds s ⊆ lowerBounds (f '' s) := + hf.dual_right.image_upperBounds_subset_upperBounds_image + +/-- The image under an antitone function of a set which is bounded above is bounded below. -/ +theorem map_bddAbove : BddAbove s → BddBelow (f '' s) := + hf.dual_right.map_bddAbove + +/-- The image under an antitone function of a set which is bounded below is bounded above. -/ +theorem map_bddBelow : BddBelow s → BddAbove (f '' s) := + hf.dual_right.map_bddBelow + +/-- An antitone map sends a greatest element of a set to a least element of its image. -/ +theorem map_isGreatest : IsGreatest s a → IsLeast (f '' s) (f a) := + hf.dual_right.map_isGreatest + +/-- An antitone map sends a least element of a set to a greatest element of its image. -/ +theorem map_isLeast : IsLeast s a → IsGreatest (f '' s) (f a) := + hf.dual_right.map_isLeast + +end Antitone + +section Image2 + +variable [Preorder α] [Preorder β] [Preorder γ] {f : α → β → γ} {s : Set α} {t : Set β} {a : α} + {b : β} + +section MonotoneMonotone + +variable (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Monotone (f a)) + +include h₀ h₁ + +theorem mem_upperBounds_image2 (ha : a ∈ upperBounds s) (hb : b ∈ upperBounds t) : + f a b ∈ upperBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem mem_lowerBounds_image2 (ha : a ∈ lowerBounds s) (hb : b ∈ lowerBounds t) : + f a b ∈ lowerBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem image2_upperBounds_upperBounds_subset : + image2 f (upperBounds s) (upperBounds t) ⊆ upperBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ mem_upperBounds_image2 h₀ h₁ ha hb + +theorem image2_lowerBounds_lowerBounds_subset : + image2 f (lowerBounds s) (lowerBounds t) ⊆ lowerBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ mem_lowerBounds_image2 h₀ h₁ ha hb + +/-- See also `Monotone.map_bddAbove`. -/ +protected theorem BddAbove.image2 : + BddAbove s → BddAbove t → BddAbove (image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_upperBounds_image2 h₀ h₁ ha hb⟩ + +/-- See also `Monotone.map_bddBelow`. -/ +protected theorem BddBelow.image2 : + BddBelow s → BddBelow t → BddBelow (image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_lowerBounds_image2 h₀ h₁ ha hb⟩ + +protected theorem IsGreatest.image2 (ha : IsGreatest s a) (hb : IsGreatest t b) : + IsGreatest (image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, mem_upperBounds_image2 h₀ h₁ ha.2 hb.2⟩ + +protected theorem IsLeast.image2 (ha : IsLeast s a) (hb : IsLeast t b) : + IsLeast (image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, mem_lowerBounds_image2 h₀ h₁ ha.2 hb.2⟩ + +end MonotoneMonotone + +section MonotoneAntitone + +variable (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Antitone (f a)) + +include h₀ h₁ + +theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds (ha : a ∈ upperBounds s) + (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds (ha : a ∈ lowerBounds s) + (hb : b ∈ upperBounds t) : f a b ∈ lowerBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem image2_upperBounds_lowerBounds_subset_upperBounds_image2 : + image2 f (upperBounds s) (lowerBounds t) ⊆ upperBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha hb + +theorem image2_lowerBounds_upperBounds_subset_lowerBounds_image2 : + image2 f (lowerBounds s) (upperBounds t) ⊆ lowerBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha hb + +theorem BddAbove.bddAbove_image2_of_bddBelow : + BddAbove s → BddBelow t → BddAbove (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha hb⟩ + +theorem BddBelow.bddBelow_image2_of_bddAbove : + BddBelow s → BddAbove t → BddBelow (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha hb⟩ + +theorem IsGreatest.isGreatest_image2_of_isLeast (ha : IsGreatest s a) (hb : IsLeast t b) : + IsGreatest (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, + mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ + +theorem IsLeast.isLeast_image2_of_isGreatest (ha : IsLeast s a) (hb : IsGreatest t b) : + IsLeast (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, + mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ + +end MonotoneAntitone + +section AntitoneAntitone + +variable (h₀ : ∀ b, Antitone (swap f b)) (h₁ : ∀ a, Antitone (f a)) + +include h₀ h₁ + +theorem mem_upperBounds_image2_of_mem_lowerBounds (ha : a ∈ lowerBounds s) + (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem mem_lowerBounds_image2_of_mem_upperBounds (ha : a ∈ upperBounds s) + (hb : b ∈ upperBounds t) : f a b ∈ lowerBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem image2_upperBounds_upperBounds_subset_upperBounds_image2 : + image2 f (lowerBounds s) (lowerBounds t) ⊆ upperBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha hb + +theorem image2_lowerBounds_lowerBounds_subset_lowerBounds_image2 : + image2 f (upperBounds s) (upperBounds t) ⊆ lowerBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha hb + +theorem BddBelow.image2_bddAbove : BddBelow s → BddBelow t → BddAbove (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha hb⟩ + +theorem BddAbove.image2_bddBelow : BddAbove s → BddAbove t → BddBelow (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha hb⟩ + +theorem IsLeast.isGreatest_image2 (ha : IsLeast s a) (hb : IsLeast t b) : + IsGreatest (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, mem_upperBounds_image2_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ + +theorem IsGreatest.isLeast_image2 (ha : IsGreatest s a) (hb : IsGreatest t b) : + IsLeast (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, mem_lowerBounds_image2_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ + +end AntitoneAntitone + +section AntitoneMonotone + +variable (h₀ : ∀ b, Antitone (swap f b)) (h₁ : ∀ a, Monotone (f a)) + +include h₀ h₁ + +theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds (ha : a ∈ lowerBounds s) + (hb : b ∈ upperBounds t) : f a b ∈ upperBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds (ha : a ∈ upperBounds s) + (hb : b ∈ lowerBounds t) : f a b ∈ lowerBounds (image2 f s t) := + forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + +theorem image2_lowerBounds_upperBounds_subset_upperBounds_image2 : + image2 f (lowerBounds s) (upperBounds t) ⊆ upperBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha hb + +theorem image2_upperBounds_lowerBounds_subset_lowerBounds_image2 : + image2 f (upperBounds s) (lowerBounds t) ⊆ lowerBounds (image2 f s t) := + image2_subset_iff.2 fun _ ha _ hb ↦ + mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha hb + +theorem BddBelow.bddAbove_image2_of_bddAbove : + BddBelow s → BddAbove t → BddAbove (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha hb⟩ + +theorem BddAbove.bddBelow_image2_of_bddAbove : + BddAbove s → BddBelow t → BddBelow (Set.image2 f s t) := by + rintro ⟨a, ha⟩ ⟨b, hb⟩ + exact ⟨f a b, mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha hb⟩ + +theorem IsLeast.isGreatest_image2_of_isGreatest (ha : IsLeast s a) (hb : IsGreatest t b) : + IsGreatest (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, + mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds h₀ h₁ ha.2 hb.2⟩ + +theorem IsGreatest.isLeast_image2_of_isLeast (ha : IsGreatest s a) (hb : IsLeast t b) : + IsLeast (Set.image2 f s t) (f a b) := + ⟨mem_image2_of_mem ha.1 hb.1, + mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds h₀ h₁ ha.2 hb.2⟩ + +end AntitoneMonotone + +end Image2 + +section Prod + +variable {α β : Type*} [Preorder α] [Preorder β] + +lemma bddAbove_prod {s : Set (α × β)} : + BddAbove s ↔ BddAbove (Prod.fst '' s) ∧ BddAbove (Prod.snd '' s) := + ⟨fun ⟨p, hp⟩ ↦ ⟨⟨p.1, forall_mem_image.2 fun _q hq ↦ (hp hq).1⟩, + ⟨p.2, forall_mem_image.2 fun _q hq ↦ (hp hq).2⟩⟩, + fun ⟨⟨x, hx⟩, ⟨y, hy⟩⟩ ↦ ⟨⟨x, y⟩, fun _p hp ↦ + ⟨hx <| mem_image_of_mem _ hp, hy <| mem_image_of_mem _ hp⟩⟩⟩ + +lemma bddBelow_prod {s : Set (α × β)} : + BddBelow s ↔ BddBelow (Prod.fst '' s) ∧ BddBelow (Prod.snd '' s) := + bddAbove_prod (α := αᵒᵈ) (β := βᵒᵈ) + +lemma bddAbove_range_prod {F : ι → α × β} : + BddAbove (range F) ↔ BddAbove (range <| Prod.fst ∘ F) ∧ BddAbove (range <| Prod.snd ∘ F) := by + simp only [bddAbove_prod, ← range_comp] + +lemma bddBelow_range_prod {F : ι → α × β} : + BddBelow (range F) ↔ BddBelow (range <| Prod.fst ∘ F) ∧ BddBelow (range <| Prod.snd ∘ F) := + bddAbove_range_prod (α := αᵒᵈ) (β := βᵒᵈ) + +theorem isLUB_prod {s : Set (α × β)} (p : α × β) : + IsLUB s p ↔ IsLUB (Prod.fst '' s) p.1 ∧ IsLUB (Prod.snd '' s) p.2 := by + refine + ⟨fun H => + ⟨⟨monotone_fst.mem_upperBounds_image H.1, fun a ha => ?_⟩, + ⟨monotone_snd.mem_upperBounds_image H.1, fun a ha => ?_⟩⟩, + fun H => ⟨?_, ?_⟩⟩ + · suffices h : (a, p.2) ∈ upperBounds s from (H.2 h).1 + exact fun q hq => ⟨ha <| mem_image_of_mem _ hq, (H.1 hq).2⟩ + · suffices h : (p.1, a) ∈ upperBounds s from (H.2 h).2 + exact fun q hq => ⟨(H.1 hq).1, ha <| mem_image_of_mem _ hq⟩ + · exact fun q hq => ⟨H.1.1 <| mem_image_of_mem _ hq, H.2.1 <| mem_image_of_mem _ hq⟩ + · exact fun q hq => + ⟨H.1.2 <| monotone_fst.mem_upperBounds_image hq, + H.2.2 <| monotone_snd.mem_upperBounds_image hq⟩ + +theorem isGLB_prod {s : Set (α × β)} (p : α × β) : + IsGLB s p ↔ IsGLB (Prod.fst '' s) p.1 ∧ IsGLB (Prod.snd '' s) p.2 := + @isLUB_prod αᵒᵈ βᵒᵈ _ _ _ _ + +end Prod + + +section Pi + +variable {π : α → Type*} [∀ a, Preorder (π a)] + +lemma bddAbove_pi {s : Set (∀ a, π a)} : + BddAbove s ↔ ∀ a, BddAbove (Function.eval a '' s) := + ⟨fun ⟨f, hf⟩ a ↦ ⟨f a, forall_mem_image.2 fun _ hg ↦ hf hg a⟩, + fun h ↦ ⟨fun a ↦ (h a).some, fun _ hg a ↦ (h a).some_mem <| mem_image_of_mem _ hg⟩⟩ + +lemma bddBelow_pi {s : Set (∀ a, π a)} : + BddBelow s ↔ ∀ a, BddBelow (Function.eval a '' s) := + bddAbove_pi (π := fun a ↦ (π a)ᵒᵈ) + +lemma bddAbove_range_pi {F : ι → ∀ a, π a} : + BddAbove (range F) ↔ ∀ a, BddAbove (range fun i ↦ F i a) := by + simp only [bddAbove_pi, ← range_comp] + rfl + +lemma bddBelow_range_pi {F : ι → ∀ a, π a} : + BddBelow (range F) ↔ ∀ a, BddBelow (range fun i ↦ F i a) := + bddAbove_range_pi (π := fun a ↦ (π a)ᵒᵈ) + +theorem isLUB_pi {s : Set (∀ a, π a)} {f : ∀ a, π a} : + IsLUB s f ↔ ∀ a, IsLUB (Function.eval a '' s) (f a) := by + classical + refine + ⟨fun H a => ⟨(Function.monotone_eval a).mem_upperBounds_image H.1, fun b hb => ?_⟩, fun H => + ⟨?_, ?_⟩⟩ + · suffices h : Function.update f a b ∈ upperBounds s from Function.update_same a b f ▸ H.2 h a + exact fun g hg => le_update_iff.2 ⟨hb <| mem_image_of_mem _ hg, fun i _ => H.1 hg i⟩ + · exact fun g hg a => (H a).1 (mem_image_of_mem _ hg) + · exact fun g hg a => (H a).2 ((Function.monotone_eval a).mem_upperBounds_image hg) + +theorem isGLB_pi {s : Set (∀ a, π a)} {f : ∀ a, π a} : + IsGLB s f ↔ ∀ a, IsGLB (Function.eval a '' s) (f a) := + @isLUB_pi α (fun a => (π a)ᵒᵈ) _ s f + +end Pi + +theorem IsGLB.of_image [Preorder α] [Preorder β] {f : α → β} (hf : ∀ {x y}, f x ≤ f y ↔ x ≤ y) + {s : Set α} {x : α} (hx : IsGLB (f '' s) (f x)) : IsGLB s x := + ⟨fun _ hy => hf.1 <| hx.1 <| mem_image_of_mem _ hy, fun _ hy => + hf.1 <| hx.2 <| Monotone.mem_lowerBounds_image (fun _ _ => hf.2) hy⟩ + +theorem IsLUB.of_image [Preorder α] [Preorder β] {f : α → β} (hf : ∀ {x y}, f x ≤ f y ↔ x ≤ y) + {s : Set α} {x : α} (hx : IsLUB (f '' s) (f x)) : IsLUB s x := + ⟨fun _ hy => hf.1 <| hx.1 <| mem_image_of_mem _ hy, fun _ hy => + hf.1 <| hx.2 <| Monotone.mem_upperBounds_image (fun _ _ => hf.2) hy⟩ + +lemma BddAbove.range_mono [Preorder β] {f : α → β} (g : α → β) (h : ∀ a, f a ≤ g a) + (hbdd : BddAbove (range g)) : BddAbove (range f) := by + obtain ⟨C, hC⟩ := hbdd + use C + rintro - ⟨x, rfl⟩ + exact (h x).trans (hC <| mem_range_self x) + +lemma BddBelow.range_mono [Preorder β] (f : α → β) {g : α → β} (h : ∀ a, f a ≤ g a) + (hbdd : BddBelow (range f)) : BddBelow (range g) := + BddAbove.range_mono (β := βᵒᵈ) f h hbdd + +lemma BddAbove.range_comp {γ : Type*} [Preorder β] [Preorder γ] {f : α → β} {g : β → γ} + (hf : BddAbove (range f)) (hg : Monotone g) : BddAbove (range (fun x => g (f x))) := by + change BddAbove (range (g ∘ f)) + simpa only [Set.range_comp] using hg.map_bddAbove hf + +lemma BddBelow.range_comp {γ : Type*} [Preorder β] [Preorder γ] {f : α → β} {g : β → γ} + (hf : BddBelow (range f)) (hg : Monotone g) : BddBelow (range (fun x => g (f x))) := by + change BddBelow (range (g ∘ f)) + simpa only [Set.range_comp] using hg.map_bddBelow hf diff --git a/Mathlib/Order/Bounds/OrderIso.lean b/Mathlib/Order/Bounds/OrderIso.lean index 913266a83c0cc..d3fda94252a64 100644 --- a/Mathlib/Order/Bounds/OrderIso.lean +++ b/Mathlib/Order/Bounds/OrderIso.lean @@ -3,7 +3,7 @@ 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, Yury Kudryashov -/ -import Mathlib.Order.Bounds.Basic +import Mathlib.Order.Bounds.Image import Mathlib.Order.Hom.Set /-! diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index 2d47e7b97ad07..b2c06f1f0a3d0 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl -/ import Mathlib.Data.Bool.Set import Mathlib.Data.Nat.Set +import Mathlib.Data.Set.NAry import Mathlib.Data.Set.Prod import Mathlib.Data.ULift import Mathlib.Order.Bounds.Basic diff --git a/Mathlib/Order/GaloisConnection.lean b/Mathlib/Order/GaloisConnection.lean index 7420bcaf290d1..5995f271ba962 100644 --- a/Mathlib/Order/GaloisConnection.lean +++ b/Mathlib/Order/GaloisConnection.lean @@ -6,7 +6,7 @@ Authors: Johannes Hölzl import Mathlib.Order.CompleteLattice import Mathlib.Order.Synonym import Mathlib.Order.Hom.Set -import Mathlib.Order.Bounds.Basic +import Mathlib.Order.Bounds.Image /-! # Galois connections, insertions and coinsertions From e06fe3e853532f908b0e52e0e65db18e4dd65a76 Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 16:03:52 +0000 Subject: [PATCH 107/131] =?UTF-8?q?chore:=20do=20not=20use=20`=C2=B7=20?= =?UTF-8?q?=E2=89=88=20=C2=B7`=20in=20`Quotient.mk=5Fout`=20(#17940)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Data/Quot.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 8dfbafe4dbf99..b9ff59544232d 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -355,7 +355,7 @@ noncomputable def Quotient.out {s : Setoid α} : Quotient s → α := theorem Quotient.out_eq {s : Setoid α} (q : Quotient s) : ⟦q.out⟧ = q := Quot.out_eq q -theorem Quotient.mk_out {s : Setoid α} (a : α) : (⟦a⟧ : Quotient s).out ≈ a := +theorem Quotient.mk_out {s : Setoid α} (a : α) : s (⟦a⟧ : Quotient s).out a := Quotient.exact (Quotient.out_eq _) theorem Quotient.mk_eq_iff_out {s : Setoid α} {x : α} {y : Quotient s} : From 2c5ee0d7efdb723afdf718540ddfabaa477015fc Mon Sep 17 00:00:00 2001 From: Yongle Hu Date: Sat, 19 Oct 2024 18:37:47 +0000 Subject: [PATCH 108/131] feat(RingTheory/Ideal/Over): define a class for an ideal lying over an ideal (#17415) Define `class` of an ideal lies over an ideal. Co-authored-by: Yongle Hu @mbkybky Co-authored-by: Jiedong Jiang @jjdishere Co-authored-by: Hu Yongle <2065545849@qq.com> Co-authored-by: Yongle Hu <2065545849@qq.com> --- Mathlib/RingTheory/FiniteType.lean | 5 + Mathlib/RingTheory/Ideal/Maps.lean | 16 +++ Mathlib/RingTheory/Ideal/Over.lean | 126 +++++++++++++++++- .../IsIntegralClosure/Basic.lean | 25 ++++ Mathlib/RingTheory/Noetherian.lean | 8 +- 5 files changed, 176 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/FiniteType.lean b/Mathlib/RingTheory/FiniteType.lean index 8eb4df96ddf80..97c79c644e56a 100644 --- a/Mathlib/RingTheory/FiniteType.lean +++ b/Mathlib/RingTheory/FiniteType.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import Mathlib.RingTheory.Adjoin.Tower +import Mathlib.RingTheory.Ideal.QuotientOperations /-! # Finiteness conditions in commutative algebra @@ -111,6 +112,10 @@ theorem trans [Algebra S A] [IsScalarTower R S A] (hRS : FiniteType R S) (hSA : FiniteType R A := ⟨fg_trans' hRS.1 hSA.1⟩ +instance quotient (R : Type*) {S : Type*} [CommSemiring R] [CommRing S] [Algebra R S] (I : Ideal S) + [h : Algebra.FiniteType R S] : Algebra.FiniteType R (S ⧸ I) := + Algebra.FiniteType.trans h inferInstance + /-- An algebra is finitely generated if and only if it is a quotient of a free algebra whose variables are indexed by a finset. -/ theorem iff_quotient_freeAlgebra : diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index faa5fc206c934..80ddaa2813840 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -846,3 +846,19 @@ def idealMap (I : Ideal R) : I →ₗ[R] I.map (algebraMap R S) := (fun _ ↦ Ideal.mem_map_of_mem _) end Algebra + +namespace NoZeroSMulDivisors + +theorem of_ker_algebraMap_eq_bot (R A : Type*) [CommRing R] [Semiring A] [Algebra R A] + [NoZeroDivisors A] (h : RingHom.ker (algebraMap R A) = ⊥) : NoZeroSMulDivisors R A := + of_algebraMap_injective ((RingHom.injective_iff_ker_eq_bot _).mpr h) + +theorem ker_algebraMap_eq_bot (R A : Type*) [CommRing R] [Ring A] [Nontrivial A] [Algebra R A] + [NoZeroSMulDivisors R A] : RingHom.ker (algebraMap R A) = ⊥ := + (RingHom.injective_iff_ker_eq_bot _).mp (algebraMap_injective R A) + +theorem iff_ker_algebraMap_eq_bot {R A : Type*} [CommRing R] [Ring A] [IsDomain A] [Algebra R A] : + NoZeroSMulDivisors R A ↔ RingHom.ker (algebraMap R A) = ⊥ := + iff_algebraMap_injective.trans (RingHom.injective_iff_ker_eq_bot (algebraMap R A)) + +end NoZeroSMulDivisors diff --git a/Mathlib/RingTheory/Ideal/Over.lean b/Mathlib/RingTheory/Ideal/Over.lean index f8586e0f2559a..bbc4a002eeb87 100644 --- a/Mathlib/RingTheory/Ideal/Over.lean +++ b/Mathlib/RingTheory/Ideal/Over.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ -import Mathlib.RingTheory.Algebraic import Mathlib.RingTheory.Localization.AtPrime import Mathlib.RingTheory.Localization.Integral @@ -429,4 +428,129 @@ lemma map_eq_top_iff {R S} [CommRing R] [CommRing S] end IsDomain +section ideal_liesOver + +section Semiring + +variable (A : Type*) [CommSemiring A] {B : Type*} [Semiring B] [Algebra A B] + (P : Ideal B) (p : Ideal A) + +/-- The ideal obtained by pulling back the ideal `P` from `B` to `A`. -/ +abbrev under : Ideal A := Ideal.comap (algebraMap A B) P + +theorem under_def : P.under A = Ideal.comap (algebraMap A B) P := rfl + +instance IsPrime.under [hP : P.IsPrime] : (P.under A).IsPrime := + hP.comap (algebraMap A B) + +variable {A} + +/-- `P` lies over `p` if `p` is the preimage of `P` of the `algebraMap`. -/ +class LiesOver : Prop where + over : p = P.under A + +instance over_under : P.LiesOver (P.under A) where over := rfl + +theorem over_def [P.LiesOver p] : p = P.under A := LiesOver.over + +theorem mem_of_liesOver [P.LiesOver p] (x : A) : x ∈ p ↔ algebraMap A B x ∈ P := by + rw [P.over_def p] + rfl + +end Semiring + +section CommSemiring + +variable {A : Type*} [CommSemiring A] {B : Type*} [CommSemiring B] {C : Type*} [Semiring C] + [Algebra A B] [Algebra B C] [Algebra A C] [IsScalarTower A B C] + (𝔓 : Ideal C) (P : Ideal B) (p : Ideal A) + +@[simp] +theorem under_under : (𝔓.under B).under A = 𝔓.under A := by + simp_rw [comap_comap, ← IsScalarTower.algebraMap_eq] + +theorem LiesOver.trans [𝔓.LiesOver P] [P.LiesOver p] : 𝔓.LiesOver p where + over := by rw [P.over_def p, 𝔓.over_def P, under_under] + +theorem LiesOver.tower_bot [hp : 𝔓.LiesOver p] [hP : 𝔓.LiesOver P] : P.LiesOver p where + over := by rw [𝔓.over_def p, 𝔓.over_def P, under_under] + +variable (B) + +instance under_liesOver_of_liesOver [𝔓.LiesOver p] : (𝔓.under B).LiesOver p := + LiesOver.tower_bot 𝔓 (𝔓.under B) p + +end CommSemiring + +section CommRing + +namespace Quotient + +variable (R : Type*) [CommSemiring R] {A B : Type*} [CommRing A] [CommRing B] [Algebra A B] + [Algebra R A] [Algebra R B] [IsScalarTower R A B] (P : Ideal B) (p : Ideal A) [P.LiesOver p] + +/-- If `P` lies over `p`, then canonically `B ⧸ P` is a `A ⧸ p`-algebra. -/ +instance algebraOfLiesOver : Algebra (A ⧸ p) (B ⧸ P) := + algebraQuotientOfLEComap (le_of_eq (P.over_def p)) + +instance isScalarTower_of_liesOver : IsScalarTower R (A ⧸ p) (B ⧸ P) := + IsScalarTower.of_algebraMap_eq' <| + congrArg (algebraMap B (B ⧸ P)).comp (IsScalarTower.algebraMap_eq R A B) + +/-- `B ⧸ P` is a finite `A ⧸ p`-module if `B` is a finite `A`-module. -/ +instance module_finite_of_liesOver [Module.Finite A B] : Module.Finite (A ⧸ p) (B ⧸ P) := + Module.Finite.of_restrictScalars_finite A (A ⧸ p) (B ⧸ P) + +example [Module.Finite A B] : Module.Finite (A ⧸ P.under A) (B ⧸ P) := inferInstance + +/-- `B ⧸ P` is a finitely generated `A ⧸ p`-algebra if `B` is a finitely generated `A`-algebra. -/ +instance algebra_finiteType_of_liesOver [Algebra.FiniteType A B] : + Algebra.FiniteType (A ⧸ p) (B ⧸ P) := + Algebra.FiniteType.of_restrictScalars_finiteType A (A ⧸ p) (B ⧸ P) + +/-- `B ⧸ P` is a Noetherian `A ⧸ p`-module if `B` is a Noetherian `A`-module. -/ +instance isNoetherian_of_liesOver [IsNoetherian A B] : IsNoetherian (A ⧸ p) (B ⧸ P) := + isNoetherian_of_tower A inferInstance + +theorem algebraMap_injective_of_liesOver : + Function.Injective (algebraMap (A ⧸ p) (B ⧸ P)) := by + rintro ⟨a⟩ ⟨b⟩ hab + apply Quotient.eq.mpr ((mem_of_liesOver P p (a - b)).mpr _) + rw [RingHom.map_sub] + exact Quotient.eq.mp hab + +instance [P.IsPrime] : NoZeroSMulDivisors (A ⧸ p) (B ⧸ P) := + NoZeroSMulDivisors.of_algebraMap_injective (Quotient.algebraMap_injective_of_liesOver P p) + +end Quotient + +end CommRing + +section IsIntegral + +variable {A : Type*} [CommRing A] {B : Type*} [CommRing B] [Algebra A B] [Algebra.IsIntegral A B] + (P : Ideal B) (p : Ideal A) [P.LiesOver p] + +variable (A) in +/-- If `B` is an integral `A`-algebra, `P` is a maximal ideal of `B`, then the pull back of + `P` is also a maximal ideal of `A`. -/ +instance IsMaximal.under [P.IsMaximal] : (P.under A).IsMaximal := + isMaximal_comap_of_isIntegral_of_isMaximal P + +theorem IsMaximal.of_liesOver_isMaximal [hpm : p.IsMaximal] [P.IsPrime] : P.IsMaximal := by + rw [P.over_def p] at hpm + exact isMaximal_of_isIntegral_of_isMaximal_comap P hpm + +theorem IsMaximal.of_isMaximal_liesOver [P.IsMaximal] : p.IsMaximal := by + rw [P.over_def p] + exact isMaximal_comap_of_isIntegral_of_isMaximal P + +/-- `B ⧸ P` is an integral `A ⧸ p`-algebra if `B` is a integral `A`-algebra. -/ +instance Quotient.algebra_isIntegral_of_liesOver : Algebra.IsIntegral (A ⧸ p) (B ⧸ P) := + Algebra.IsIntegral.tower_top A + +end IsIntegral + +end ideal_liesOver + end Ideal diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean index 59406a95b8ec6..eabc18718c598 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean @@ -558,6 +558,18 @@ nonrec theorem RingHom.IsIntegral.tower_bot (hg : Function.Injective g) haveI : IsScalarTower R S T := IsScalarTower.of_algebraMap_eq fun _ ↦ rfl fun x ↦ IsIntegral.tower_bot hg (hfg (g x)) +variable (T) in +/-- Let `T / S / R` be a tower of algebras, `T` is non-trivial and is a torsion free `S`-module, + then if `T` is an integral `R`-algebra, then `S` is an integral `R`-algebra. -/ +theorem Algebra.IsIntegral.tower_bot [Algebra R S] [Algebra R T] [Algebra S T] + [NoZeroSMulDivisors S T] [Nontrivial T] [IsScalarTower R S T] + [h : Algebra.IsIntegral R T] : Algebra.IsIntegral R S where + isIntegral := by + apply RingHom.IsIntegral.tower_bot (algebraMap R S) (algebraMap S T) + (NoZeroSMulDivisors.algebraMap_injective S T) + rw [← IsScalarTower.algebraMap_eq R S T] + exact h.isIntegral + theorem IsIntegral.tower_bot_of_field {R A B : Type*} [CommRing R] [Field A] [CommRing B] [Nontrivial B] [Algebra R A] [Algebra A B] [Algebra R B] [IsScalarTower R A B] {x : A} (h : IsIntegral R (algebraMap A B x)) : IsIntegral R x := @@ -571,6 +583,16 @@ theorem RingHom.isIntegralElem.of_comp {x : T} (h : (g.comp f).IsIntegralElem x) theorem RingHom.IsIntegral.tower_top (h : (g.comp f).IsIntegral) : g.IsIntegral := fun x ↦ RingHom.isIntegralElem.of_comp f g (h x) +variable (R) in +/-- Let `T / S / R` be a tower of algebras, `T` is an integral `R`-algebra, then it is integral + as an `S`-algebra. -/ +theorem Algebra.IsIntegral.tower_top [Algebra R S] [Algebra R T] [Algebra S T] [IsScalarTower R S T] + [h : Algebra.IsIntegral R T] : Algebra.IsIntegral S T where + isIntegral := by + apply RingHom.IsIntegral.tower_top (algebraMap R S) (algebraMap S T) + rw [← IsScalarTower.algebraMap_eq R S T] + exact h.isIntegral + theorem RingHom.IsIntegral.quotient {I : Ideal S} (hf : f.IsIntegral) : (Ideal.quotientMap I f le_rfl).IsIntegral := by rintro ⟨x⟩ @@ -578,6 +600,9 @@ theorem RingHom.IsIntegral.quotient {I : Ideal S} (hf : f.IsIntegral) : refine ⟨p.map (Ideal.Quotient.mk _), p_monic.map _, ?_⟩ simpa only [hom_eval₂, eval₂_map] using congr_arg (Ideal.Quotient.mk I) hpx +instance {I : Ideal A} [Algebra.IsIntegral R A] : Algebra.IsIntegral R (A ⧸ I) := + Algebra.IsIntegral.trans A + instance Algebra.IsIntegral.quotient {I : Ideal A} [Algebra.IsIntegral R A] : Algebra.IsIntegral (R ⧸ I.comap (algebraMap R A)) (A ⧸ I) := ⟨RingHom.IsIntegral.quotient (algebraMap R A) Algebra.IsIntegral.isIntegral⟩ diff --git a/Mathlib/RingTheory/Noetherian.lean b/Mathlib/RingTheory/Noetherian.lean index 903c8862e2b78..411b9fdf6c5a4 100644 --- a/Mathlib/RingTheory/Noetherian.lean +++ b/Mathlib/RingTheory/Noetherian.lean @@ -112,9 +112,11 @@ theorem isNoetherian_of_surjective (f : M →ₗ[R] P) (hf : LinearMap.range f = variable {M} -instance isNoetherian_quotient {R} [Ring R] {M} [AddCommGroup M] [Module R M] - (N : Submodule R M) [IsNoetherian R M] : IsNoetherian R (M ⧸ N) := - isNoetherian_of_surjective _ _ (LinearMap.range_eq_top.mpr N.mkQ_surjective) +instance isNoetherian_quotient {A M : Type*} [Ring A] [AddCommGroup M] [SMul R A] [Module R M] + [Module A M] [IsScalarTower R A M] (N : Submodule A M) [IsNoetherian R M] : + IsNoetherian R (M ⧸ N) := + isNoetherian_of_surjective M ((Submodule.mkQ N).restrictScalars R) <| + LinearMap.range_eq_top.mpr N.mkQ_surjective @[deprecated (since := "2024-04-27"), nolint defLemma] alias Submodule.Quotient.isNoetherian := isNoetherian_quotient From 37de4d77648a7d69b80dfc65ad9ef35e65696d2f Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Sat, 19 Oct 2024 19:08:11 +0000 Subject: [PATCH 109/131] feat(CategoryTheory/KanExtensions): Evaluating `lan` twice amounts to taking a colimit (#17531) --- .../Functor/KanExtension/Adjunction.lean | 44 +++++++++++++++++ .../Functor/KanExtension/Pointwise.lean | 48 +++++++++++++++++++ 2 files changed, 92 insertions(+) diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean index fe8433c905fb7..274273b8794c3 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean @@ -56,6 +56,28 @@ noncomputable def isPointwiseLeftKanExtensionLanUnit (LeftExtension.mk _ (L.lanUnit.app F)).IsPointwiseLeftKanExtension := isPointwiseLeftKanExtensionOfIsLeftKanExtension (F := F) _ (L.lanUnit.app F) +/-- If a left Kan extension is pointwise, then evaluating it at an object is isomorphic to +taking a colimit. -/ +noncomputable def lanObjObjIsoColimit (F : C ⥤ H) [HasPointwiseLeftKanExtension L F] (X : D) : + (L.lan.obj F).obj X ≅ Limits.colimit (CostructuredArrow.proj L X ⋙ F) := + LeftExtension.IsPointwiseLeftKanExtensionAt.isoColimit (F := F) + (isPointwiseLeftKanExtensionLanUnit L F X) + +@[reassoc (attr := simp)] +lemma ι_lanObjObjIsoColimit_inv + (F : C ⥤ H) [HasPointwiseLeftKanExtension L F] (X : D) (f : CostructuredArrow L X) : + Limits.colimit.ι _ f ≫ (L.lanObjObjIsoColimit F X).inv = + (L.lanUnit.app F).app f.left ≫ (L.lan.obj F).map f.hom := by + simp [lanObjObjIsoColimit, lanUnit] + +@[reassoc (attr := simp)] +lemma ι_lanObjObjIsoColimit_hom + (F : C ⥤ H) [HasPointwiseLeftKanExtension L F] (X : D) (f : CostructuredArrow L X) : + (L.lanUnit.app F).app f.left ≫ (L.lan.obj F).map f.hom ≫ (L.lanObjObjIsoColimit F X).hom = + Limits.colimit.ι (CostructuredArrow.proj L X ⋙ F) f := + LeftExtension.IsPointwiseLeftKanExtensionAt.ι_isoColimit_hom (F := F) + (isPointwiseLeftKanExtensionLanUnit L F X) f + variable (H) in /-- The left Kan extension functor `L.Lan` is left adjoint to the precomposition by `L`. -/ @@ -162,6 +184,28 @@ noncomputable def isPointwiseRightKanExtensionRanCounit (RightExtension.mk _ (L.ranCounit.app F)).IsPointwiseRightKanExtension := isPointwiseRightKanExtensionOfIsRightKanExtension (F := F) _ (L.ranCounit.app F) +/-- If a right Kan extension is pointwise, then evaluating it at an object is isomorphic to +taking a limit. -/ +noncomputable def ranObjObjIsoLimit (F : C ⥤ H) [HasPointwiseRightKanExtension L F] (X : D) : + (L.ran.obj F).obj X ≅ Limits.limit (StructuredArrow.proj X L ⋙ F) := + RightExtension.IsPointwiseRightKanExtensionAt.isoLimit (F := F) + (isPointwiseRightKanExtensionRanCounit L F X) + +@[reassoc (attr := simp)] +lemma ranObjObjIsoLimit_hom_π + (F : C ⥤ H) [HasPointwiseRightKanExtension L F] (X : D) (f : StructuredArrow X L) : + (L.ranObjObjIsoLimit F X).hom ≫ Limits.limit.π _ f = + (L.ran.obj F).map f.hom ≫ (L.ranCounit.app F).app f.right := by + simp [ranObjObjIsoLimit, ran, ranCounit] + +@[reassoc (attr := simp)] +lemma ranObjObjIsoLimit_inv_π + (F : C ⥤ H) [HasPointwiseRightKanExtension L F] (X : D) (f : StructuredArrow X L) : + (L.ranObjObjIsoLimit F X).inv ≫ (L.ran.obj F).map f.hom ≫ (L.ranCounit.app F).app f.right = + Limits.limit.π _ f := + RightExtension.IsPointwiseRightKanExtensionAt.isoLimit_inv_π (F := F) + (isPointwiseRightKanExtensionRanCounit L F X) f + variable (H) in /-- The right Kan extension functor `L.ran` is right adjoint to the precomposition by `L`. -/ diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean index 9d4cb7a6c2ea1..a1a044d1a3846 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean @@ -98,6 +98,30 @@ lemma IsPointwiseLeftKanExtensionAt.isIso_hom_app IsIso (E.hom.app X) := by simpa using h.isIso_ι_app_of_isTerminal _ CostructuredArrow.mkIdTerminal +namespace IsPointwiseLeftKanExtensionAt + +variable {E} {Y : D} (h : E.IsPointwiseLeftKanExtensionAt Y) + [HasColimit (CostructuredArrow.proj L Y ⋙ F)] + +/-- A pointwise left Kan extension of `F` along `L` applied to an object `Y` is isomorphic to +`colimit (CostructuredArrow.proj L Y ⋙ F)`. -/ +noncomputable def isoColimit : + E.right.obj Y ≅ colimit (CostructuredArrow.proj L Y ⋙ F) := + h.coconePointUniqueUpToIso (colimit.isColimit _) + +@[reassoc (attr := simp)] +lemma ι_isoColimit_inv (g : CostructuredArrow L Y) : + colimit.ι _ g ≫ h.isoColimit.inv = E.hom.app g.left ≫ E.right.map g.hom := + IsColimit.comp_coconePointUniqueUpToIso_inv _ _ _ + +@[reassoc (attr := simp)] +lemma ι_isoColimit_hom (g : CostructuredArrow L Y) : + E.hom.app g.left ≫ E.right.map g.hom ≫ h.isoColimit.hom = + colimit.ι (CostructuredArrow.proj L Y ⋙ F) g := by + simpa using h.comp_coconePointUniqueUpToIso_hom (colimit.isColimit _) g + +end IsPointwiseLeftKanExtensionAt + /-- A left extension `E : LeftExtension L F` is a pointwise left Kan extension when it is a pointwise left Kan extension at any object. -/ abbrev IsPointwiseLeftKanExtension := ∀ (Y : D), E.IsPointwiseLeftKanExtensionAt Y @@ -211,6 +235,30 @@ lemma IsPointwiseRightKanExtensionAt.isIso_hom_app IsIso (E.hom.app X) := by simpa using h.isIso_π_app_of_isInitial _ StructuredArrow.mkIdInitial +namespace IsPointwiseRightKanExtensionAt + +variable {E} {Y : D} (h : E.IsPointwiseRightKanExtensionAt Y) + [HasLimit (StructuredArrow.proj Y L ⋙ F)] + +/-- A pointwise right Kan extension of `F` along `L` applied to an object `Y` is isomorphic to +`limit (StructuredArrow.proj Y L ⋙ F)`. -/ +noncomputable def isoLimit : + E.left.obj Y ≅ limit (StructuredArrow.proj Y L ⋙ F) := + h.conePointUniqueUpToIso (limit.isLimit _) + +@[reassoc (attr := simp)] +lemma isoLimit_hom_π (g : StructuredArrow Y L) : + h.isoLimit.hom ≫ limit.π _ g = E.left.map g.hom ≫ E.hom.app g.right := + IsLimit.conePointUniqueUpToIso_hom_comp _ _ _ + +@[reassoc (attr := simp)] +lemma isoLimit_inv_π (g : StructuredArrow Y L) : + h.isoLimit.inv ≫ E.left.map g.hom ≫ E.hom.app g.right = + limit.π (StructuredArrow.proj Y L ⋙ F) g := by + simpa using h.conePointUniqueUpToIso_inv_comp (limit.isLimit _) g + +end IsPointwiseRightKanExtensionAt + /-- A right extension `E : RightExtension L F` is a pointwise right Kan extension when it is a pointwise right Kan extension at any object. -/ abbrev IsPointwiseRightKanExtension := ∀ (Y : D), E.IsPointwiseRightKanExtensionAt Y From 19e2abfdd236a1feb36b33a6e219be17097d587f Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sat, 19 Oct 2024 19:20:12 +0000 Subject: [PATCH 110/131] feat (AlgebraicGeometry/Modules): stalks of the sheaf `M^~` (#14247) Construct the isomorphism of $R$-modules from the stalk of the sheaf $M^{\~}$ on $Spec(R)$ to the localization $M_x$ as an $R$-module, where $R$ is a commutative ring and $M$ is an $R$-module. ## Main definition * `ModuleCat.tilde.stalkIso`: The isomorphism of `R`-modules from the stalk of `M^~` at `x` to the localization of `M` at the prime ideal corresponding to `x`. ## Technical note To get the `R`-module structure on the stalks on `M^~`, we had to define `ModuleCat.tildeInModuleCat`, which is `M^~` seen as sheaf of `R`-modules. We get it by applying a forgetful functor to `ModuleCat.tilde M`. Co-authored-by: Weihong Xu Co-authored-by: Sophie Morel Co-authored-by: Amelia Livingston This contribution was created as part of the AIM workshop "Formalizing algebraic geometry" in June 2024. Co-authored-by: morel Co-authored-by: Weihong Xu --- Mathlib/AlgebraicGeometry/Modules/Tilde.lean | 140 +++++++++++++++++++ 1 file changed, 140 insertions(+) diff --git a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean index 7fd34c2070d7a..d14b580db3d1a 100644 --- a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean +++ b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean @@ -22,6 +22,8 @@ such that `M^~(U)` is the set of dependent functions that are locally fractions. * `ModuleCat.tildeInType` : `M^~` as a sheaf of types groups. * `ModuleCat.tilde` : `M^~` as a sheaf of `𝒪_{Spec R}`-modules. +* `ModuleCat.tilde.stalkIso`: The isomorphism of `R`-modules from the stalk of `M^~` at `x` to +the localization of `M` at the prime ideal corresponding to `x`. ## Technical note @@ -252,6 +254,144 @@ noncomputable def localizationToStalk (x : PrimeSpectrum.Top R) : (TopCat.Presheaf.stalk (tildeInModuleCat M) x) := LocalizedModule.lift _ (toStalk M x) <| isUnit_toStalk M x + +/-- The ring homomorphism that takes a section of the structure sheaf of `R` on the open set `U`, +implemented as a subtype of dependent functions to localizations at prime ideals, and evaluates +the section on the point corresponding to a given prime ideal. -/ +def openToLocalization (U : Opens (PrimeSpectrum R)) (x : PrimeSpectrum R) (hx : x ∈ U) : + (tildeInModuleCat M).obj (op U) ⟶ + ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) where + toFun s := (s.1 ⟨x, hx⟩ : _) + map_add' _ _ := rfl + map_smul' _ _ := rfl + +/-- +The morphism of `R`-modules from the stalk of `M^~` at `x` to the localization of `M` at the +prime ideal of `R` corresponding to `x`. +-/ +noncomputable def stalkToFiberLinearMap (x : PrimeSpectrum.Top R) : + (tildeInModuleCat M).stalk x ⟶ + ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) := + Limits.colimit.desc ((OpenNhds.inclusion x).op ⋙ (tildeInModuleCat M)) + { pt := _ + ι := + { app := fun U => openToLocalization M ((OpenNhds.inclusion _).obj U.unop) x U.unop.2 } } + +@[simp] +theorem germ_comp_stalkToFiberLinearMap (U : Opens (PrimeSpectrum.Top R)) (x) (hx : x ∈ U) : + (tildeInModuleCat M).germ U x hx ≫ stalkToFiberLinearMap M x = + openToLocalization M U x hx := + Limits.colimit.ι_desc _ _ + +@[simp] +theorem stalkToFiberLinearMap_germ (U : Opens (PrimeSpectrum.Top R)) (x : PrimeSpectrum.Top R) + (hx : x ∈ U) (s : (tildeInModuleCat M).1.obj (op U)) : + stalkToFiberLinearMap M x + (TopCat.Presheaf.germ (tildeInModuleCat M) U x hx s) = (s.1 ⟨x, hx⟩ : _) := + DFunLike.ext_iff.1 (germ_comp_stalkToFiberLinearMap M U x hx) s + +@[reassoc (attr := simp), elementwise (attr := simp)] +theorem toOpen_germ (U : Opens (PrimeSpectrum.Top R)) (x) (hx : x ∈ U) : + toOpen M U ≫ M.tildeInModuleCat.germ U x hx = toStalk M x := by + rw [← toOpen_res M ⊤ U (homOfLE le_top : U ⟶ ⊤), Category.assoc, Presheaf.germ_res]; rfl + +@[reassoc (attr := simp)] +theorem toStalk_comp_stalkToFiberLinearMap (x : PrimeSpectrum.Top R) : + toStalk M x ≫ stalkToFiberLinearMap M x = + LocalizedModule.mkLinearMap x.asIdeal.primeCompl M := by + rw [toStalk, Category.assoc, germ_comp_stalkToFiberLinearMap]; rfl + +theorem stalkToFiberLinearMap_toStalk (x : PrimeSpectrum.Top R) (m : M) : + (stalkToFiberLinearMap M x) (toStalk M x m) = + LocalizedModule.mk m 1 := + LinearMap.ext_iff.1 (toStalk_comp_stalkToFiberLinearMap M x) _ + +/-- +If `U` is an open subset of `Spec R`, `m` is an element of `M` and `r` is an element of `R` +that is invertible on `U` (i.e. does not belong to any prime ideal corresponding to a point +in `U`), this is `m / r` seen as a section of `M^~` over `U`. +-/ +def const (m : M) (r : R) (U : Opens (PrimeSpectrum.Top R)) + (hu : ∀ x ∈ U, r ∈ (x : PrimeSpectrum.Top R).asIdeal.primeCompl) : + (tildeInModuleCat M).obj (op U) := + ⟨fun x => LocalizedModule.mk m ⟨r, hu x x.2⟩, fun x => + ⟨U, x.2, 𝟙 _, m, r, fun y => ⟨hu _ y.2, by + simpa only [LocalizedModule.mkLinearMap_apply, LocalizedModule.smul'_mk, + LocalizedModule.mk_eq] using ⟨1, by simp⟩⟩⟩⟩ + +@[simp] +theorem const_apply (m : M) (r : R) (U : Opens (PrimeSpectrum.Top R)) + (hu : ∀ x ∈ U, r ∈ (x : PrimeSpectrum.Top R).asIdeal.primeCompl) (x : U) : + (const M m r U hu).1 x = LocalizedModule.mk m ⟨r, hu x x.2⟩ := + rfl + +theorem exists_const (U) (s : (tildeInModuleCat M).obj (op U)) (x : PrimeSpectrum.Top R) + (hx : x ∈ U) : + ∃ (V : Opens (PrimeSpectrum.Top R)) (_ : x ∈ V) (i : V ⟶ U) (f : M) (g : R) (hg : _), + const M f g V hg = (tildeInModuleCat M).map i.op s := + let ⟨V, hxV, iVU, f, g, hfg⟩ := s.2 ⟨x, hx⟩ + ⟨V, hxV, iVU, f, g, fun y hyV => (hfg ⟨y, hyV⟩).1, Subtype.eq <| funext fun y => by + obtain ⟨h1, (h2 : g • s.1 ⟨y, _⟩ = LocalizedModule.mk f 1)⟩ := hfg y + exact show LocalizedModule.mk f ⟨g, by exact h1⟩ = s.1 (iVU y) by + set x := s.1 (iVU y); change g • x = _ at h2; clear_value x + induction x using LocalizedModule.induction_on with + | h a b => + rw [LocalizedModule.smul'_mk, LocalizedModule.mk_eq] at h2 + obtain ⟨c, hc⟩ := h2 + exact LocalizedModule.mk_eq.mpr ⟨c, by simpa using hc.symm⟩⟩ + +@[simp] +theorem res_const (f : M) (g : R) (U hu V hv i) : + (tildeInModuleCat M).map i (const M f g U hu) = const M f g V hv := + rfl + +@[simp] +theorem localizationToStalk_mk (x : PrimeSpectrum.Top R) (f : M) (s : x.asIdeal.primeCompl) : + localizationToStalk M x (LocalizedModule.mk f s) = + (tildeInModuleCat M).germ (PrimeSpectrum.basicOpen (s : R)) x s.2 + (const M f s (PrimeSpectrum.basicOpen s) fun _ => id) := + (Module.End_isUnit_iff _ |>.1 (isUnit_toStalk M x s)).injective <| by + erw [← LinearMap.mul_apply] + simp only [IsUnit.mul_val_inv, LinearMap.one_apply, Module.algebraMap_end_apply] + show (M.tildeInModuleCat.germ ⊤ x ⟨⟩) ((toOpen M ⊤) f) = _ + rw [← map_smul] + fapply TopCat.Presheaf.germ_ext (W := PrimeSpectrum.basicOpen s.1) (hxW := s.2) + · exact homOfLE le_top + · exact 𝟙 _ + refine Subtype.eq <| funext fun y => show LocalizedModule.mk f 1 = _ from ?_ + show LocalizedModule.mk f 1 = s.1 • LocalizedModule.mk f _ + rw [LocalizedModule.smul'_mk, LocalizedModule.mk_eq] + exact ⟨1, by simp⟩ + +/-- +The isomorphism of `R`-modules from the stalk of `M^~` at `x` to the localization of `M` at the +prime ideal corresponding to `x`. +-/ +@[simps] +noncomputable def stalkIso (x : PrimeSpectrum.Top R) : + TopCat.Presheaf.stalk (tildeInModuleCat M) x ≅ + ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) where + hom := stalkToFiberLinearMap M x + inv := localizationToStalk M x + hom_inv_id := TopCat.Presheaf.stalk_hom_ext _ fun U hxU ↦ ext _ fun s ↦ by + show localizationToStalk M x (stalkToFiberLinearMap M x (M.tildeInModuleCat.germ U x hxU s)) = + M.tildeInModuleCat.germ U x hxU s + rw [stalkToFiberLinearMap_germ] + obtain ⟨V, hxV, iVU, f, g, (hg : V ≤ PrimeSpectrum.basicOpen _), hs⟩ := + exists_const _ _ s x hxU + rw [← res_apply M U V iVU s ⟨x, hxV⟩, ← hs, const_apply, localizationToStalk_mk] + exact (tildeInModuleCat M).germ_ext V hxV (homOfLE hg) iVU <| hs ▸ rfl + inv_hom_id := by ext x; exact x.induction_on (fun _ _ => by simp) + +instance (x : PrimeSpectrum.Top R) : + IsLocalizedModule x.asIdeal.primeCompl (toStalk M x) := by + convert IsLocalizedModule.of_linearEquiv + (hf := localizedModuleIsLocalizedModule (M := M) x.asIdeal.primeCompl) + (e := (stalkIso M x).symm.toLinearEquiv) + simp only [of_coe, show (stalkIso M x).symm.toLinearEquiv.toLinearMap = (stalkIso M x).inv by rfl, + stalkIso_inv] + erw [LocalizedModule.lift_comp] + end Tilde end ModuleCat From 721c01e4de26b2bb91e0c32ddcdcdf842e97c405 Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 19 Oct 2024 19:27:05 +0000 Subject: [PATCH 111/131] chore(Data/Nat/{Bit, Bitwise}): make `Nat.bit_mod_two` `Nat.bit_eq_zero_iff` be `simp` (#17924) Also deprecate `Nat.bit_eq_zero`. It is identical to `Nat.bit_eq_zero_iff`. --- Mathlib/Data/Nat/Bits.lean | 1 + Mathlib/Data/Nat/Bitwise.lean | 13 +++++-------- Mathlib/Data/Nat/Multiplicity.lean | 7 +++---- 3 files changed, 9 insertions(+), 12 deletions(-) diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 71f01660fc34f..6ea51b387201e 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -328,6 +328,7 @@ theorem bit_cases_on_inj {motive : ℕ → Sort u} (H₁ H₂ : ∀ b n, motive ((fun n => bitCasesOn n H₁) = fun n => bitCasesOn n H₂) ↔ H₁ = H₂ := bit_cases_on_injective.eq_iff +@[simp] theorem bit_eq_zero_iff {n : ℕ} {b : Bool} : bit b n = 0 ↔ n = 0 ∧ b = false := by constructor · cases b <;> simp [Nat.bit]; omega diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index cce15c0d55eee..07873ec6fd976 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -88,19 +88,18 @@ lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by cases a <;> cases b <;> simp [h2, h4] <;> split_ifs <;> simp_all (config := {decide := true}) [two_mul] +@[simp] lemma bit_mod_two (a : Bool) (x : ℕ) : bit a x % 2 = a.toNat := by cases a <;> simp [bit_val, mul_add_mod] -@[simp] lemma bit_mod_two_eq_zero_iff (a x) : bit a x % 2 = 0 ↔ !a := by - simp [bit_mod_two] + simp -@[simp] lemma bit_mod_two_eq_one_iff (a x) : bit a x % 2 = 1 ↔ a := by - simp [bit_mod_two] + simp @[simp] theorem lor_bit : ∀ a m b n, bit a m ||| bit b n = bit (a || b) (m ||| n) := @@ -142,12 +141,10 @@ theorem bit_false : bit false = (2 * ·) := theorem bit_true : bit true = (2 * · + 1) := rfl -@[simp] -theorem bit_eq_zero {n : ℕ} {b : Bool} : n.bit b = 0 ↔ n = 0 ∧ b = false := by - cases b <;> simp [bit, Nat.mul_eq_zero] +@[deprecated (since := "2024-10-19")] alias bit_eq_zero := bit_eq_zero_iff theorem bit_ne_zero_iff {n : ℕ} {b : Bool} : n.bit b ≠ 0 ↔ n = 0 → b = true := by - simpa only [not_and, Bool.not_eq_false] using (@bit_eq_zero n b).not + simp /-- An alternative for `bitwise_bit` which replaces the `f false false = false` assumption with assumptions that neither `bit a m` nor `bit b n` are `0` diff --git a/Mathlib/Data/Nat/Multiplicity.lean b/Mathlib/Data/Nat/Multiplicity.lean index e86d693f824f5..61b86c32c88a8 100644 --- a/Mathlib/Data/Nat/Multiplicity.lean +++ b/Mathlib/Data/Nat/Multiplicity.lean @@ -6,7 +6,6 @@ Authors: Chris Hughes import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.GeomSum import Mathlib.Algebra.Order.Ring.Abs -import Mathlib.Data.Nat.Bitwise import Mathlib.Data.Nat.Log import Mathlib.Data.Nat.Prime.Defs import Mathlib.Data.Nat.Digits @@ -282,10 +281,10 @@ theorem emultiplicity_two_factorial_lt : ∀ {n : ℕ} (_ : n ≠ 0), emultiplic · intro b n ih h by_cases hn : n = 0 · subst hn - simp only [ne_eq, bit_eq_zero, true_and, Bool.not_eq_false] at h - simp only [h, bit_true, factorial, mul_one, Nat.isUnit_iff, cast_one] + simp only [ne_eq, bit_eq_zero_iff, true_and, Bool.not_eq_false] at h + simp only [h, factorial, mul_one, Nat.isUnit_iff, cast_one] rw [Prime.emultiplicity_one] - · simp [zero_lt_one] + · exact zero_lt_one · decide have : emultiplicity 2 (2 * n)! < (2 * n : ℕ) := by rw [prime_two.emultiplicity_factorial_mul] From fee48cc34cea5902fd163549acb01e413d7e7243 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Sat, 19 Oct 2024 19:57:03 +0000 Subject: [PATCH 112/131] =?UTF-8?q?feat(AlgebraicGeometry/ResidueField):?= =?UTF-8?q?=20classification=20of=20`Spec=20K=20=E2=9F=B6=20X`=20with=20`K?= =?UTF-8?q?`=20a=20field=20(#17768)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds more API for residue fields and in particular the classification of morphisms `Spec K ⟶ X` with `K` a field. From the valuative criterion project. Co-authored-by: Qi Ge Co-authored-by: Andrew Yang Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> Co-authored-by: Christian Merten --- .../Morphisms/ClosedImmersion.lean | 6 +- .../Morphisms/Preimmersion.lean | 27 +++++ Mathlib/AlgebraicGeometry/ResidueField.lean | 108 +++++++++++++++++- Mathlib/AlgebraicGeometry/Stalk.lean | 15 ++- 4 files changed, 150 insertions(+), 6 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index 932325058e426..03d557fa0c469 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -3,12 +3,10 @@ Copyright (c) 2023 Jonas van der Schaaf. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Christian Merten, Jonas van der Schaaf -/ -import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties import Mathlib.AlgebraicGeometry.Morphisms.FiniteType import Mathlib.AlgebraicGeometry.ResidueField -import Mathlib.RingTheory.RingHom.Surjective /-! @@ -134,6 +132,10 @@ theorem of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsClosedImmersion simp_rw [Scheme.stalkMap_comp] at h exact Function.Surjective.of_comp h +instance Spec_map_residue {X : Scheme.{u}} (x) : IsClosedImmersion (Spec.map (X.residue x)) := + IsClosedImmersion.spec_of_surjective (X.residue x) + Ideal.Quotient.mk_surjective + instance {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : QuasiCompact f where isCompact_preimage _ _ hU' := base_closed.isCompact_preimage hU' diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean index 9336dbb858cde..f504cc230196f 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang -/ import Mathlib.AlgebraicGeometry.Morphisms.UnderlyingMap import Mathlib.RingTheory.RingHom.Surjective +import Mathlib.RingTheory.SurjectiveOnStalks /-! @@ -95,6 +96,32 @@ theorem comp_iff {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsPreimmersion g] IsPreimmersion (f ≫ g) ↔ IsPreimmersion f := ⟨fun _ ↦ of_comp f g, fun _ ↦ inferInstance⟩ +lemma Spec_map_iff {R S : CommRingCat.{u}} (f : R ⟶ S) : + IsPreimmersion (Spec.map f) ↔ Embedding (PrimeSpectrum.comap f) ∧ f.SurjectiveOnStalks := by + haveI : (RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).RespectsIso := by + rw [← RingHom.toMorphismProperty_respectsIso_iff] + exact RingHom.surjective_respectsIso + refine ⟨fun ⟨h₁, h₂⟩ ↦ ⟨h₁, ?_⟩, fun ⟨h₁, h₂⟩ ↦ ⟨h₁, fun (x : PrimeSpectrum S) ↦ ?_⟩⟩ + · intro p hp + let e := Scheme.arrowStalkMapSpecIso f ⟨p, hp⟩ + apply ((RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).arrow_mk_iso_iff e).mp + exact h₂ ⟨p, hp⟩ + · let e := Scheme.arrowStalkMapSpecIso f x + apply ((RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).arrow_mk_iso_iff e).mpr + exact h₂ x.asIdeal x.isPrime + +lemma mk_Spec_map {R S : CommRingCat.{u}} {f : R ⟶ S} + (h₁ : Embedding (PrimeSpectrum.comap f)) (h₂ : f.SurjectiveOnStalks) : + IsPreimmersion (Spec.map f) := + (Spec_map_iff f).mpr ⟨h₁, h₂⟩ + +lemma of_isLocalization {R S : Type u} [CommRing R] (M : Submonoid R) [CommRing S] + [Algebra R S] [IsLocalization M S] : + IsPreimmersion (Spec.map (CommRingCat.ofHom <| algebraMap R S)) := + IsPreimmersion.mk_Spec_map + (PrimeSpectrum.localization_comap_embedding (R := R) S M) + (RingHom.surjectiveOnStalks_of_isLocalization (M := M) S) + end IsPreimmersion end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/ResidueField.lean b/Mathlib/AlgebraicGeometry/ResidueField.lean index ac1d5d1e8dcc2..c7fcb54274fd0 100644 --- a/Mathlib/AlgebraicGeometry/ResidueField.lean +++ b/Mathlib/AlgebraicGeometry/ResidueField.lean @@ -3,8 +3,8 @@ Copyright (c) 2024 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.Geometry.RingedSpace.LocallyRingedSpace.ResidueField import Mathlib.AlgebraicGeometry.Stalk +import Mathlib.Geometry.RingedSpace.LocallyRingedSpace.ResidueField /-! @@ -20,12 +20,15 @@ The following are in the `AlgebraicGeometry.Scheme` namespace: - `AlgebraicGeometry.Scheme.Hom.residueFieldMap`: A morphism of schemes induce a homomorphism of residue fields. - `AlgebraicGeometry.Scheme.fromSpecResidueField`: The canonical map `Spec κ(x) ⟶ X`. +- `AlgebraicGeometry.Scheme.SpecToEquivOfField`: morphisms `Spec K ⟶ X` for a field `K` correspond + to pairs of `x : X` with embedding `κ(x) ⟶ K`. + -/ universe u -open CategoryTheory TopologicalSpace Opposite +open CategoryTheory TopologicalSpace Opposite LocalRing noncomputable section @@ -45,12 +48,37 @@ instance (x : X) : Field (X.residueField x) := def residue (X : Scheme.{u}) (x) : X.presheaf.stalk x ⟶ X.residueField x := LocalRing.residue _ +/-- See `AlgebraicGeometry.IsClosedImmersion.Spec_map_residue` for the stronger result that +`Spec.map (X.residue x)` is a closed immersion. -/ +instance {X : Scheme.{u}} (x) : IsPreimmersion (Spec.map (X.residue x)) := + IsPreimmersion.mk_Spec_map + (PrimeSpectrum.closedEmbedding_comap_of_surjective _ _ (Ideal.Quotient.mk_surjective)).embedding + (RingHom.surjectiveOnStalks_of_surjective (Ideal.Quotient.mk_surjective)) + +@[simp] +lemma Spec_map_residue_apply {X : Scheme.{u}} (x : X) (s : Spec (X.residueField x)) : + (Spec.map (X.residue x)).base s = closedPoint (X.presheaf.stalk x) := + LocalRing.PrimeSpectrum.comap_residue _ s + lemma residue_surjective (X : Scheme.{u}) (x) : Function.Surjective (X.residue x) := Ideal.Quotient.mk_surjective instance (X : Scheme.{u}) (x) : Epi (X.residue x) := ConcreteCategory.epi_of_surjective _ (X.residue_surjective x) +/-- If `K` is a field and `f : 𝒪_{X, x} ⟶ K` is a ring map, then this is the induced +map `κ(x) ⟶ K`. -/ +def descResidueField {K : Type u} [Field K] {X : Scheme.{u}} {x : X} + (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f] : + X.residueField x ⟶ .of K := + LocalRing.ResidueField.lift (S := K) f + +@[reassoc (attr := simp)] +lemma residue_descResidueField {K : Type u} [Field K] {X : Scheme.{u}} {x} + (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f] : + X.residue x ≫ X.descResidueField f = f := + RingHom.ext fun _ ↦ rfl + /-- If `U` is an open of `X` containing `x`, we have a canonical ring map from the sections over `U` to the residue field of `x`. @@ -171,6 +199,10 @@ lemma residue_residueFieldCongr (X : Scheme) {x y : X} (h : x = y) : subst h simp +lemma Hom.residueFieldMap_congr {f g : X ⟶ Y} (e : f = g) (x : X) : + f.residueFieldMap x = (Y.residueFieldCongr (by subst e; rfl)).hom ≫ g.residueFieldMap x := by + subst e; simp + end congr section fromResidueField @@ -178,7 +210,12 @@ section fromResidueField /-- The canonical map `Spec κ(x) ⟶ X`. -/ def fromSpecResidueField (X : Scheme) (x : X) : Spec (X.residueField x) ⟶ X := - Spec.map (CommRingCat.ofHom (X.residue x)) ≫ X.fromSpecStalk x + Spec.map (X.residue x) ≫ X.fromSpecStalk x + +instance {X : Scheme.{u}} (x : X) : IsPreimmersion (X.fromSpecResidueField x) := by + dsimp only [Scheme.fromSpecResidueField] + rw [IsPreimmersion.comp_iff] + infer_instance @[reassoc (attr := simp)] lemma residueFieldCongr_fromSpecResidueField {x y : X} (h : x = y) : @@ -186,8 +223,73 @@ lemma residueFieldCongr_fromSpecResidueField {x y : X} (h : x = y) : X.fromSpecResidueField _ := by subst h; simp +@[reassoc] +lemma Hom.Spec_map_residueFieldMap_fromSpecResidueField (x : X) : + Spec.map (f.residueFieldMap x) ≫ Y.fromSpecResidueField _ = + X.fromSpecResidueField x ≫ f := by + dsimp only [fromSpecResidueField] + rw [Category.assoc, ← Spec_map_stalkMap_fromSpecStalk, ← Spec.map_comp_assoc, + ← Spec.map_comp_assoc] + rfl + +@[simp] +lemma fromSpecResidueField_apply (x : X.carrier) (s : Spec (X.residueField x)) : + (X.fromSpecResidueField x).base s = x := by + simp [fromSpecResidueField] + +lemma range_fromSpecResidueField (x : X.carrier) : + Set.range (X.fromSpecResidueField x).base = {x} := by + ext s + simp only [Set.mem_range, fromSpecResidueField_apply, Set.mem_singleton_iff, eq_comm (a := s)] + constructor + · rintro ⟨-, h⟩ + exact h + · rintro rfl + exact ⟨closedPoint (X.residueField x), rfl⟩ + +lemma descResidueField_fromSpecResidueField {K : Type*} [Field K] (X : Scheme) {x} + (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f] : + Spec.map (X.descResidueField f) ≫ + X.fromSpecResidueField x = Spec.map f ≫ X.fromSpecStalk x := by + simp [fromSpecResidueField, ← Spec.map_comp_assoc] + +lemma descResidueField_stalkClosedPointTo_fromSpecResidueField + (K : Type u) [Field K] (X : Scheme.{u}) (f : Spec (.of K) ⟶ X) : + Spec.map (X.descResidueField (Scheme.stalkClosedPointTo f)) ≫ + X.fromSpecResidueField (f.base (closedPoint K)) = f := by + erw [X.descResidueField_fromSpecResidueField] + rw [Scheme.Spec_stalkClosedPointTo_fromSpecStalk] + end fromResidueField +/-- A helper lemma to work with `AlgebraicGeometry.Scheme.SpecToEquivOfField`. -/ +lemma SpecToEquivOfField_eq_iff {K : Type*} [Field K] {X : Scheme} + {f₁ f₂ : Σ x : X.carrier, X.residueField x ⟶ .of K} : + f₁ = f₂ ↔ ∃ e : f₁.1 = f₂.1, f₁.2 = (X.residueFieldCongr e).hom ≫ f₂.2 := by + constructor + · rintro rfl + simp + · obtain ⟨f, _⟩ := f₁ + obtain ⟨g, _⟩ := f₂ + rintro ⟨(rfl : f = g), h⟩ + simpa + +/-- For a field `K` and a scheme `X`, the morphisms `Spec K ⟶ X` bijectively correspond +to pairs of points `x` of `X` and embeddings `κ(x) ⟶ K`. -/ +def SpecToEquivOfField (K : Type u) [Field K] (X : Scheme.{u}) : + (Spec (.of K) ⟶ X) ≃ Σ x, X.residueField x ⟶ .of K where + toFun f := + ⟨_, X.descResidueField (Scheme.stalkClosedPointTo f)⟩ + invFun xf := Spec.map xf.2 ≫ X.fromSpecResidueField xf.1 + left_inv := Scheme.descResidueField_stalkClosedPointTo_fromSpecResidueField K X + right_inv f := by + rw [SpecToEquivOfField_eq_iff] + simp only [CommRingCat.coe_of, Scheme.comp_coeBase, TopCat.coe_comp, Function.comp_apply, + Scheme.fromSpecResidueField_apply, exists_true_left] + rw [← Spec.map_inj, Spec.map_comp, ← cancel_mono (X.fromSpecResidueField _)] + erw [Scheme.descResidueField_stalkClosedPointTo_fromSpecResidueField] + simp + end Scheme end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index 251d379d23a14..92b56c3849cc3 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang, Fangming Li -/ import Mathlib.AlgebraicGeometry.AffineScheme +import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion /-! # Stalks of a Scheme @@ -68,6 +69,19 @@ noncomputable def Scheme.fromSpecStalk (X : Scheme) (x : X) : theorem IsAffineOpen.fromSpecStalk_eq_fromSpecStalk {x : X} (hxU : x ∈ U) : hU.fromSpecStalk hxU = X.fromSpecStalk x := fromSpecStalk_eq .. +instance IsAffineOpen.fromSpecStalk_isPreimmersion {X : Scheme.{u}} {U : Opens X} + (hU : IsAffineOpen U) (x : X) (hx : x ∈ U) : IsPreimmersion (hU.fromSpecStalk hx) := by + dsimp [IsAffineOpen.fromSpecStalk] + haveI : IsPreimmersion (Spec.map (X.presheaf.germ U x hx)) := + letI : Algebra Γ(X, U) (X.presheaf.stalk x) := (X.presheaf.germ U x hx).toAlgebra + haveI := hU.isLocalization_stalk ⟨x, hx⟩ + IsPreimmersion.of_isLocalization (R := Γ(X, U)) (S := X.presheaf.stalk x) + (hU.primeIdealOf ⟨x, hx⟩).asIdeal.primeCompl + apply IsPreimmersion.comp + +instance {X : Scheme.{u}} (x : X) : IsPreimmersion (X.fromSpecStalk x) := + IsAffineOpen.fromSpecStalk_isPreimmersion _ _ _ + lemma IsAffineOpen.fromSpecStalk_closedPoint {U : Opens X} (hU : IsAffineOpen U) {x : X} (hxU : x ∈ U) : (hU.fromSpecStalk hxU).base (closedPoint (X.presheaf.stalk x)) = x := by @@ -184,7 +198,6 @@ section stalkClosedPointTo variable {R} (f : Spec R ⟶ X) - namespace Scheme /-- From 227950d6190a304c6d697f37d74425bebe12917d Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Sat, 19 Oct 2024 21:32:09 +0000 Subject: [PATCH 113/131] feat: topology on module over topological ring (#16895) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a module over a topological ring, we define the module topology on this module to be the finest module making it into a topological module (i.e. the finest topology making addition and scalar multiplication continuous). NB this PR gave rise to a discussion about whether `⨆ a ∈ s, f a` or `sSup (f '' s)` should be the simp normal form, which was discussed on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/normal.20form.20for.20pullback.20of.20Inf.20of.20topologies/near/472676441) without any clear conclusion. Co-authored-by: Eric Wieser --- Mathlib.lean | 1 + Mathlib/Topology/Algebra/Module/Basic.lean | 13 + .../Algebra/Module/ModuleTopology.lean | 265 ++++++++++++++++++ Mathlib/Topology/Order.lean | 10 + 4 files changed, 289 insertions(+) create mode 100644 Mathlib/Topology/Algebra/Module/ModuleTopology.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5fcbb05388323..c75bbf3dd05cd 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4553,6 +4553,7 @@ import Mathlib.Topology.Algebra.Module.Determinant import Mathlib.Topology.Algebra.Module.FiniteDimension import Mathlib.Topology.Algebra.Module.LinearPMap import Mathlib.Topology.Algebra.Module.LocallyConvex +import Mathlib.Topology.Algebra.Module.ModuleTopology import Mathlib.Topology.Algebra.Module.Multilinear.Basic import Mathlib.Topology.Algebra.Module.Multilinear.Bounded import Mathlib.Topology.Algebra.Module.Multilinear.Topology diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 4296c9b917999..943df043c2c09 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -12,6 +12,7 @@ import Mathlib.Algebra.Algebra.Defs import Mathlib.LinearAlgebra.Projection import Mathlib.LinearAlgebra.Pi import Mathlib.LinearAlgebra.Finsupp +import Mathlib.Algebra.Module.Opposites /-! # Theory of topological modules and continuous linear maps. @@ -2379,4 +2380,16 @@ end Submodule end Quotient +namespace MulOpposite + +variable (R : Type*) [Semiring R] [τR : TopologicalSpace R] [TopologicalSemiring R] + {M : Type*} [AddCommMonoid M] [Module R M] [TopologicalSpace M] [ContinuousSMul R M] + +/-- The function `op` is a continuous linear equivalence. -/ +@[simps!] +def opContinuousLinearEquiv : M ≃L[R] Mᵐᵒᵖ where + __ := MulOpposite.opLinearEquiv R + +end MulOpposite + set_option linter.style.longFile 2500 diff --git a/Mathlib/Topology/Algebra/Module/ModuleTopology.lean b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean new file mode 100644 index 0000000000000..6fdfe9edc0b90 --- /dev/null +++ b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean @@ -0,0 +1,265 @@ +/- +Copyright (c) 2024 Kevin Buzzard. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kevin Buzzard, Will Sawin +-/ +import Mathlib.Topology.Algebra.Module.Basic + +/-! +# A "module topology" for modules over a topological ring + +If `R` is a topological ring acting on an additive abelian group `A`, we define the +*module topology* to be the finest topology on `A` making both the maps +`• : R × A → A` and `+ : A × A → A` continuous (with all the products having the +product topology). Note that `- : A → A` is also automatically continuous as it is `a ↦ (-1) • a`. + +This topology was suggested by Will Sawin [here](https://mathoverflow.net/a/477763/1384). + + +## Mathematical details + +I (buzzard) don't know of any reference for this other than Sawin's mathoverflow answer, +so I expand some of the details here. + +First note that the definition makes sense in far more generality (for example `R` just needs to +be a topological space acting on an additive monoid). + +Next note that there *is* a finest topology with this property! Indeed, topologies on a fixed +type form a complete lattice (infinite infs and sups exist). So if `τ` is the Inf of all +the topologies on `A` which make `+` and `•` continuous, then the claim is that `+` and `•` +are still continuous for `τ` (note that topologies are ordered so that finer topologies +are smaller). To show `+ : A × A → A` is continuous we equivalently need to show +that the pushforward of the product topology `τ × τ` along `+` is `≤ τ`, and because `τ` is +the greatest lower bound of the topologies making `•` and `+` continuous, +it suffices to show that it's `≤ σ` for any topology `σ` on `A` which makes `+` and `•` continuous. +However pushforward and products are monotone, so `τ × τ ≤ σ × σ`, and the pushforward of +`σ × σ` is `≤ σ` because that's precisely the statement that `+` is continuous for `σ`. +The proof for `•` follows mutatis mutandis. + +A *topological module* for a topological ring `R` is an `R`-module `A` with a topology +making `+` and `•` continuous. The discussion so far has shown that the module topology makes +an `R`-module `A` into a topological module, and moreover is the finest topology with this property. + +A crucial observation is that if `M` is a topological `R`-module, if `A` is an `R`-module with no +topology, and if `φ : A → M` is linear, then the pullback of `M`'s topology to `A` is a topology +making `A` into a topological module. Let's for example check that `•` is continuous. +If `U ⊆ A` is open then by definition of the pullback topology, `U = φ⁻¹(V)` for some open `V ⊆ M`, +and now the pullback of `U` under `•` is just the pullback along the continuous map +`id × φ : R × A → R × M` of the preimage of `V` under the continuous map `• : R × M → M`, +so it's open. The proof for `+` is similar. + +As a consequence of this, we see that if `φ : A → M` is a linear map between topological `R`-modules +modules and if `A` has the module topology, then `φ` is automatically continuous. +Indeed the argument above shows that if `A → M` is linear then the module +topology on `A` is `≤` the pullback of the module topology on `M` (because it's the inf of a set +containing this topology) which is the definition of continuity. + +We also deduce that the module topology is a functor from the category of `R`-modules +(`R` a topological ring) to the category of topological `R`-modules, and it is perhaps +unsurprising that this is an adjoint to the forgetful functor. Indeed, if `A` is an `R`-module +and `M` is a topological `R`-module, then the previous paragraph shows that +the linear maps `A → M` are precisely the continuous linear maps +from (`A` with its module topology) to `M`, so the module topology is a left adjoint +to the forgetful functor. + +This file develops the theory of the module topology. We prove that the module topology on +`R` as a module over itself is `R`'s original topology, and that the module topology +is isomorphism-invariant. + +## Main theorems + +* `TopologicalSemiring.toIsModuleTopology : IsModuleTopology R R`. The module + topology on `R` is `R`'s topology. +* `IsModuleTopology.iso [IsModuleTopology R A] (e : A ≃L[R] B) : IsModuleTopology R B`. If `A` and + `B` are `R`-modules with topologies, if `e` is a topological isomorphism between them, + and if `A` has the module topology, then `B` has the module topology. + +## TODO + +Forthcoming PRs from the FLT repo will show that the module topology on a (binary or finite) product +of modules is the product of the module topologies, and that the module topology on the quotient +of a module `M` is the quotient topology when `M` is equipped with the module topology. + +We will also show the slightly more subtle result that if `M`, `N` and `P` are `R`-modules +equipped with the module topology and if furthermore `M` is finite as an `R`-module, +then any bilinear map `M × N → P` is continuous. + +As a consequence of this, we deduce that if `R` is a commutative topological ring +and `A` is an `R`-algebra of finite type as `R`-module, then `A` with its module +topology becomes a topological ring (i.e., multiplication is continuous). + +Other TODOs (not done in the FLT repo): + +* The module topology is a functor from the category of `R`-modules + to the category of topological `R`-modules, and it's an adjoint to the forgetful functor. + +-/ + +section basics + +/- +This section is just boilerplate, defining the module topology and making +some infrastructure. Note that we don't even need that `R` is a ring +in the main definitions, just that it acts on `A`. +-/ +variable (R : Type*) [TopologicalSpace R] (A : Type*) [Add A] [SMul R A] + +/-- The module topology, for a module `A` over a topological ring `R`. It's the finest topology +making addition and the `R`-action continuous, or equivalently the finest topology making `A` +into a topological `R`-module. More precisely it's the Inf of the set of +topologies with these properties; theorems `continuousSMul` and `continuousAdd` show +that the module topology also has these properties. -/ +abbrev moduleTopology : TopologicalSpace A := + sInf {t | @ContinuousSMul R A _ _ t ∧ @ContinuousAdd A t _} + +/-- A class asserting that the topology on a module over a topological ring `R` is +the module topology. See `moduleTopology` for more discussion of the module topology. -/ +class IsModuleTopology [τA : TopologicalSpace A] : Prop where + /-- Note that this should not be used directly, and `eq_moduleTopology`, which takes `R` and `A` + explicitly, should be used instead. -/ + eq_moduleTopology' : τA = moduleTopology R A + +theorem eq_moduleTopology [τA : TopologicalSpace A] [IsModuleTopology R A] : + τA = moduleTopology R A := + IsModuleTopology.eq_moduleTopology' (R := R) (A := A) + +/-- Scalar multiplication `• : R × A → A` is continuous if `R` is a topological +ring, and `A` is an `R` module with the module topology. -/ +theorem ModuleTopology.continuousSMul : @ContinuousSMul R A _ _ (moduleTopology R A) := + /- Proof: We need to prove that the product topology is finer than the pullback + of the module topology. But the module topology is an Inf and thus a limit, + and pullback is a right adjoint, so it preserves limits. + We must thus show that the product topology is finer than an Inf, so it suffices + to show it's a lower bound, which is not hard. All this is wrapped into + `continuousSMul_sInf`. + -/ + continuousSMul_sInf fun _ h ↦ h.1 + +/-- Addition `+ : A × A → A` is continuous if `R` is a topological +ring, and `A` is an `R` module with the module topology. -/ +theorem ModuleTopology.continuousAdd : @ContinuousAdd A (moduleTopology R A) _ := + continuousAdd_sInf fun _ h ↦ h.2 + +instance IsModuleTopology.toContinuousSMul [TopologicalSpace A] [IsModuleTopology R A] : + ContinuousSMul R A := eq_moduleTopology R A ▸ ModuleTopology.continuousSMul R A + +-- this can't be an instance because typclass inference can't be expected to find `R`. +theorem IsModuleTopology.toContinuousAdd [TopologicalSpace A] [IsModuleTopology R A] : + ContinuousAdd A := eq_moduleTopology R A ▸ ModuleTopology.continuousAdd R A + +/-- The module topology is `≤` any topology making `A` into a topological module. -/ +theorem moduleTopology_le [τA : TopologicalSpace A] [ContinuousSMul R A] [ContinuousAdd A] : + moduleTopology R A ≤ τA := sInf_le ⟨inferInstance, inferInstance⟩ + +end basics + +namespace IsModuleTopology + +section basics + +variable {R : Type*} [TopologicalSpace R] + {A : Type*} [Add A] [SMul R A] [τA : TopologicalSpace A] + +/-- If `A` is a topological `R`-module and the identity map from (`A` with its given +topology) to (`A` with the module topology) is continuous, then the topology on `A` is +the module topology. -/ +theorem of_continuous_id [ContinuousAdd A] [ContinuousSMul R A] + (h : @Continuous A A τA (moduleTopology R A) id): + IsModuleTopology R A where + -- The topologies are equal because each is finer than the other. One inclusion + -- follows from the continuity hypothesis; the other is because the module topology + -- is the inf of all the topologies making `A` a topological module. + eq_moduleTopology' := le_antisymm (continuous_id_iff_le.1 h) (moduleTopology_le _ _) + +/-- The zero module has the module topology. -/ +instance instSubsingleton [Subsingleton A] : IsModuleTopology R A where + eq_moduleTopology' := Subsingleton.elim _ _ + +end basics + +section iso + +variable {R : Type*} [τR : TopologicalSpace R] [Semiring R] +variable {A : Type*} [AddCommMonoid A] [Module R A] [τA : TopologicalSpace A] [IsModuleTopology R A] +variable {B : Type*} [AddCommMonoid B] [Module R B] [τB : TopologicalSpace B] + +/-- If `A` and `B` are `R`-modules, homeomorphic via an `R`-linear homeomorphism, and if +`A` has the module topology, then so does `B`. -/ +theorem iso (e : A ≃L[R] B) : IsModuleTopology R B where + eq_moduleTopology' := by + -- get these in before I start putting new topologies on A and B and have to use `@` + let g : A →ₗ[R] B := e + let g' : B →ₗ[R] A := e.symm + let h : A →+ B := e + let h' : B →+ A := e.symm + simp_rw [e.toHomeomorph.symm.inducing.1, eq_moduleTopology R A, moduleTopology, induced_sInf] + apply congr_arg + ext τ -- from this point on the definitions of `g`, `g'` etc above don't work without `@`. + rw [Set.mem_image] + constructor + · rintro ⟨σ, ⟨hσ1, hσ2⟩, rfl⟩ + exact ⟨continuousSMul_induced g', continuousAdd_induced h'⟩ + · rintro ⟨h1, h2⟩ + use τ.induced e + rw [induced_compose] + refine ⟨⟨continuousSMul_induced g, continuousAdd_induced h⟩, ?_⟩ + nth_rw 2 [← induced_id (t := τ)] + simp + +end iso + +section self + +variable (R : Type*) [Semiring R] [τR : TopologicalSpace R] [TopologicalSemiring R] + +/-! +We now fix once and for all a topological semiring `R`. + +We first prove that the module topology on `R` considered as a module over itself, +is `R`'s topology. +-/ + +/-- The topology on a topological semiring `R` agrees with the module topology when considering +`R` as an `R`-module in the obvious way (i.e., via `Semiring.toModule`). -/ +instance _root_.TopologicalSemiring.toIsModuleTopology : IsModuleTopology R R := by + /- By a previous lemma it suffices to show that the identity from (R,usual) to + (R, module topology) is continuous. -/ + apply of_continuous_id + /- + The idea needed here is to rewrite the identity function as the composite of `r ↦ (r,1)` + from `R` to `R × R`, and multiplication `R × R → R`. + -/ + rw [show (id : R → R) = (fun rs ↦ rs.1 • rs.2) ∘ (fun r ↦ (r, 1)) by ext; simp] + /- + It thus suffices to show that each of these maps are continuous. For this claim to even make + sense, we need to topologise `R × R`. The trick is to do this by giving the first `R` the usual + topology `τR` and the second `R` the module topology. To do this we have to "fight mathlib" + a bit with `@`, because there is more than one topology on `R` here. + -/ + apply @Continuous.comp R (R × R) R τR (@instTopologicalSpaceProd R R τR (moduleTopology R R)) + (moduleTopology R R) + · /- + The map R × R → R is `•`, so by a fundamental property of the module topology, + this is continuous. -/ + exact @continuous_smul _ _ _ _ (moduleTopology R R) <| ModuleTopology.continuousSMul .. + · /- + The map `R → R × R` sending `r` to `(r,1)` is a map into a product, so it suffices to show + that each of the two factors is continuous. But the first is the identity function + on `(R, usual topology)` and the second is a constant function. -/ + exact @Continuous.prod_mk _ _ _ _ (moduleTopology R R) _ _ _ continuous_id <| + @continuous_const _ _ _ (moduleTopology R R) _ + +end self + +section MulOpposite + +variable (R : Type*) [Semiring R] [τR : TopologicalSpace R] [TopologicalSemiring R] + +/-- The module topology coming from the action of the topological ring `Rᵐᵒᵖ` on `R` + (via `Semiring.toOppositeModule`, i.e. via `(op r) • m = m * r`) is `R`'s topology. -/ +instance _root_.TopologicalSemiring.toOppositeIsModuleTopology : IsModuleTopology Rᵐᵒᵖ R := + .iso (MulOpposite.opContinuousLinearEquiv Rᵐᵒᵖ).symm + +end MulOpposite + +end IsModuleTopology diff --git a/Mathlib/Topology/Order.lean b/Mathlib/Topology/Order.lean index 05a207e173994..173ade1245b1d 100644 --- a/Mathlib/Topology/Order.lean +++ b/Mathlib/Topology/Order.lean @@ -393,6 +393,11 @@ theorem induced_iInf {ι : Sort w} {t : ι → TopologicalSpace α} : (⨅ i, t i).induced g = ⨅ i, (t i).induced g := (gc_coinduced_induced g).u_iInf +@[simp] +theorem induced_sInf {s : Set (TopologicalSpace α)} : + TopologicalSpace.induced g (sInf s) = sInf (TopologicalSpace.induced g '' s) := by + rw [sInf_eq_iInf', sInf_image', induced_iInf] + @[simp] theorem coinduced_bot : (⊥ : TopologicalSpace α).coinduced f = ⊥ := (gc_coinduced_induced f).l_bot @@ -406,6 +411,11 @@ theorem coinduced_iSup {ι : Sort w} {t : ι → TopologicalSpace α} : (⨆ i, t i).coinduced f = ⨆ i, (t i).coinduced f := (gc_coinduced_induced f).l_iSup +@[simp] +theorem coinduced_sSup {s : Set (TopologicalSpace α)} : + TopologicalSpace.coinduced f (sSup s) = sSup ((TopologicalSpace.coinduced f) '' s) := by + rw [sSup_eq_iSup', sSup_image', coinduced_iSup] + theorem induced_id [t : TopologicalSpace α] : t.induced id = t := TopologicalSpace.ext <| funext fun s => propext <| ⟨fun ⟨_, hs, h⟩ => h ▸ hs, fun hs => ⟨s, hs, rfl⟩⟩ From e9f0a88e5333f1edc2843a9164fb035a96406f5e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 19 Oct 2024 21:53:49 +0000 Subject: [PATCH 114/131] feat(Normed/Group): add `nndist_one_right` etc (#17954) Add - `nndist_one_right`, `nndist_zero_right`, `nndist_one_left`, `nndist_zero_left`; - `edist_one_right`, `edist_zero_right`, `edist_one_left`, `edist_zero_left`. --- Mathlib/Analysis/Normed/Group/Basic.lean | 14 ++++++++++++++ Mathlib/MeasureTheory/Function/L1Space.lean | 6 ++---- 2 files changed, 16 insertions(+), 4 deletions(-) diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index e379a64547598..b3742ef73f777 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -635,6 +635,13 @@ theorem nndist_eq_nnnorm_div (a b : E) : nndist a b = ‖a / b‖₊ := alias nndist_eq_nnnorm := nndist_eq_nnnorm_sub +@[to_additive (attr := simp)] +theorem nndist_one_right (a : E) : nndist a 1 = ‖a‖₊ := by simp [nndist_eq_nnnorm_div] + +@[to_additive (attr := simp)] +theorem edist_one_right (a : E) : edist a 1 = ‖a‖₊ := by + rw [edist_nndist, nndist_one_right] + @[to_additive (attr := simp) nnnorm_zero] theorem nnnorm_one' : ‖(1 : E)‖₊ = 0 := NNReal.eq norm_one' @@ -653,6 +660,13 @@ theorem nnnorm_mul_le' (a b : E) : ‖a * b‖₊ ≤ ‖a‖₊ + ‖b‖₊ := theorem nnnorm_inv' (a : E) : ‖a⁻¹‖₊ = ‖a‖₊ := NNReal.eq <| norm_inv' a +@[to_additive (attr := simp)] +theorem nndist_one_left (a : E) : nndist 1 a = ‖a‖₊ := by simp [nndist_eq_nnnorm_div] + +@[to_additive (attr := simp)] +theorem edist_one_left (a : E) : edist 1 a = ‖a‖₊ := by + rw [edist_nndist, nndist_one_left] + open scoped symmDiff in @[to_additive] theorem nndist_mulIndicator (s t : Set α) (f : α → E) (x : α) : diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 9994ae4d6e8b1..87c029742d760 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -1375,12 +1375,10 @@ theorem edist_toL1_toL1 (f g : α → β) (hf : Integrable f μ) (hg : Integrabl ENNReal.rpow_one, ne_eq, not_false_eq_true, div_self, ite_false] simp [edist_eq_coe_nnnorm_sub] -@[simp] theorem edist_toL1_zero (f : α → β) (hf : Integrable f μ) : edist (hf.toL1 f) 0 = ∫⁻ a, edist (f a) 0 ∂μ := by - simp only [toL1, Lp.edist_toLp_zero, eLpNorm, one_ne_zero, eLpNorm', one_toReal, ENNReal.rpow_one, - ne_eq, not_false_eq_true, div_self, ite_false] - simp [edist_eq_coe_nnnorm] + simp only [edist_zero_right, Lp.nnnorm_coe_ennreal, toL1_eq_mk, eLpNorm_aeeqFun] + apply eLpNorm_one_eq_lintegral_nnnorm variable {𝕜 : Type*} [NormedRing 𝕜] [Module 𝕜 β] [BoundedSMul 𝕜 β] From e0dccaab3f60adf16923fcc30de87ce08f296833 Mon Sep 17 00:00:00 2001 From: Thmoas-Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Sat, 19 Oct 2024 22:21:59 +0000 Subject: [PATCH 115/131] chore(FieldTheory/Galois): Move infinite Galois (#17945) Split the original file `FieldTheory.Galois.Infinite` into `GaloisClosure` and `Profinite`, for spliting the basic constructions, profinite (stated categorically), and the fundamental theorem completely. Moves: - Mathlib.FieldTheory.Galois.Infinite -> Mathlib.FieldTheory.Galois.GaloisClosure - Mathlib.FieldTheory.Galois.Infinite -> Mathlib.FieldTheory.Galois.Profinite --- Mathlib.lean | 3 +- .../{Infinite.lean => GaloisClosure.lean} | 72 +--------------- Mathlib/FieldTheory/Galois/Profinite.lean | 84 +++++++++++++++++++ 3 files changed, 87 insertions(+), 72 deletions(-) rename Mathlib/FieldTheory/Galois/{Infinite.lean => GaloisClosure.lean} (63%) create mode 100644 Mathlib/FieldTheory/Galois/Profinite.lean diff --git a/Mathlib.lean b/Mathlib.lean index c75bbf3dd05cd..3ace00fab33be 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2730,7 +2730,8 @@ import Mathlib.FieldTheory.Finite.Trace import Mathlib.FieldTheory.Finiteness import Mathlib.FieldTheory.Fixed import Mathlib.FieldTheory.Galois.Basic -import Mathlib.FieldTheory.Galois.Infinite +import Mathlib.FieldTheory.Galois.GaloisClosure +import Mathlib.FieldTheory.Galois.Profinite import Mathlib.FieldTheory.IntermediateField.Algebraic import Mathlib.FieldTheory.IntermediateField.Basic import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure diff --git a/Mathlib/FieldTheory/Galois/Infinite.lean b/Mathlib/FieldTheory/Galois/GaloisClosure.lean similarity index 63% rename from Mathlib/FieldTheory/Galois/Infinite.lean rename to Mathlib/FieldTheory/Galois/GaloisClosure.lean index d3c2345887e0d..c1f93236da4c0 100644 --- a/Mathlib/FieldTheory/Galois/Infinite.lean +++ b/Mathlib/FieldTheory/Galois/GaloisClosure.lean @@ -1,11 +1,9 @@ /- Copyright (c) 2024 Nailin Guan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Nailin Guan, Yuyang Zhao, Jujian Zhang +Authors: Nailin Guan, Yuyang Zhao -/ -import Mathlib.Algebra.Category.Grp.FiniteGrp -import Mathlib.CategoryTheory.Category.Preorder import Mathlib.FieldTheory.NormalClosure import Mathlib.FieldTheory.SeparableClosure @@ -21,24 +19,12 @@ In a field extension `K/k` * `adjoin` : The finite Galois intermediate field obtained from the normal closure of adjoining a finite `s : Set K` to `k`. -* `finGaloisGroup L` : The (finite) Galois group `Gal(L/k)` associated to a - `L : FiniteGaloisIntermediateField k K` `L`. - -* `finGaloisGroupMap` : For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` - giving the restriction of `Gal(L₁/k)` to `Gal(L₂/k)` - -* `finGaloisGroupFunctor` : Mapping `FiniteGaloisIntermediateField` ordered by inverse inclusion to - its corresponding Galois Group as `FiniteGrp`. - # TODO * `FiniteGaloisIntermediateField` should be a `ConditionallyCompleteLattice` but isn't proved yet. -/ -open CategoryTheory Opposite -open scoped IntermediateField - variable (k K : Type*) [Field k] [Field K] [Algebra k K] /-- The type of intermediate fields of `K/k` that are finite and Galois over `k` -/ @@ -156,59 +142,3 @@ theorem adjoin_simple_map_algEquiv [IsGalois k K] (f : K ≃ₐ[k] K) (x : K) : adjoin_simple_map_algHom (f : K →ₐ[k] K) x end FiniteGaloisIntermediateField - -section Profinite - -variable {k K} - -/-- The (finite) Galois group `Gal(L / k)` associated to a -`L : FiniteGaloisIntermediateField k K` `L`. -/ -def FiniteGaloisIntermediateField.finGaloisGroup (L : FiniteGaloisIntermediateField k K) : - FiniteGrp := - letI := AlgEquiv.fintype k L - FiniteGrp.of <| L ≃ₐ[k] L - -/-- For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` - the restriction homomorphism from `Gal(L₁/k)` to `Gal(L₂/k)` -/ -noncomputable def finGaloisGroupMap {L₁ L₂ : (FiniteGaloisIntermediateField k K)ᵒᵖ} - (le : L₁ ⟶ L₂) : L₁.unop.finGaloisGroup ⟶ L₂.unop.finGaloisGroup := - haveI : Normal k L₂.unop := IsGalois.to_normal - letI : Algebra L₂.unop L₁.unop := RingHom.toAlgebra (Subsemiring.inclusion <| leOfHom le.1) - haveI : IsScalarTower k L₂.unop L₁.unop := IsScalarTower.of_algebraMap_eq (congrFun rfl) - FiniteGrp.ofHom (AlgEquiv.restrictNormalHom L₂.unop) - -namespace finGaloisGroupMap - -@[simp] -lemma map_id (L : (FiniteGaloisIntermediateField k K)ᵒᵖ) : - (finGaloisGroupMap (𝟙 L)) = 𝟙 L.unop.finGaloisGroup := - AlgEquiv.restrictNormalHom_id _ _ - -@[simp] -lemma map_comp {L₁ L₂ L₃ : (FiniteGaloisIntermediateField k K)ᵒᵖ} (f : L₁ ⟶ L₂) (g : L₂ ⟶ L₃) : - finGaloisGroupMap (f ≫ g) = finGaloisGroupMap f ≫ finGaloisGroupMap g := by - iterate 2 - induction L₁ with | _ L₁ => ?_ - induction L₂ with | _ L₂ => ?_ - induction L₃ with | _ L₃ => ?_ - letI : Algebra L₃ L₂ := RingHom.toAlgebra (Subsemiring.inclusion g.unop.le) - letI : Algebra L₂ L₁ := RingHom.toAlgebra (Subsemiring.inclusion f.unop.le) - letI : Algebra L₃ L₁ := RingHom.toAlgebra (Subsemiring.inclusion (g.unop.le.trans f.unop.le)) - haveI : IsScalarTower k L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl - haveI : IsScalarTower k L₃ L₁ := IsScalarTower.of_algebraMap_eq' rfl - haveI : IsScalarTower k L₃ L₂ := IsScalarTower.of_algebraMap_eq' rfl - haveI : IsScalarTower L₃ L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl - apply IsScalarTower.AlgEquiv.restrictNormalHom_comp k L₃ L₂ L₁ - -end finGaloisGroupMap - -variable (k K) in -/-- The functor from `FiniteGaloisIntermediateField` (ordered by reverse inclusion) to `FiniteGrp`, -mapping each intermediate field `K/L/k` to `Gal (L/k)`.-/ -noncomputable def finGaloisGroupFunctor : (FiniteGaloisIntermediateField k K)ᵒᵖ ⥤ FiniteGrp where - obj L := L.unop.finGaloisGroup - map := finGaloisGroupMap - map_id := finGaloisGroupMap.map_id - map_comp := finGaloisGroupMap.map_comp - -end Profinite diff --git a/Mathlib/FieldTheory/Galois/Profinite.lean b/Mathlib/FieldTheory/Galois/Profinite.lean new file mode 100644 index 0000000000000..300f56169316b --- /dev/null +++ b/Mathlib/FieldTheory/Galois/Profinite.lean @@ -0,0 +1,84 @@ +/- +Copyright (c) 2024 Nailin Guan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nailin Guan, Yuyang Zhao, Jujian Zhang +-/ + +import Mathlib.Algebra.Category.Grp.FiniteGrp +import Mathlib.CategoryTheory.Category.Preorder +import Mathlib.FieldTheory.Galois.GaloisClosure + +/-! + +# Main definitions and results + +In a field extension `K/k` + +* `finGaloisGroup L` : The (finite) Galois group `Gal(L/k)` associated to a + `L : FiniteGaloisIntermediateField k K` `L`. + +* `finGaloisGroupMap` : For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` + giving the restriction of `Gal(L₁/k)` to `Gal(L₂/k)` + +* `finGaloisGroupFunctor` : Mapping `FiniteGaloisIntermediateField` ordered by inverse inclusion to + its corresponding Galois Group as `FiniteGrp`. + +-/ + +open CategoryTheory Opposite + +variable {k K : Type*} [Field k] [Field K] [Algebra k K] + +section Profinite + +/-- The (finite) Galois group `Gal(L / k)` associated to a +`L : FiniteGaloisIntermediateField k K` `L`. -/ +def FiniteGaloisIntermediateField.finGaloisGroup (L : FiniteGaloisIntermediateField k K) : + FiniteGrp := + letI := AlgEquiv.fintype k L + FiniteGrp.of <| L ≃ₐ[k] L + +/-- For `FiniteGaloisIntermediateField` s `L₁` and `L₂` with `L₂ ≤ L₁` + the restriction homomorphism from `Gal(L₁/k)` to `Gal(L₂/k)` -/ +noncomputable def finGaloisGroupMap {L₁ L₂ : (FiniteGaloisIntermediateField k K)ᵒᵖ} + (le : L₁ ⟶ L₂) : L₁.unop.finGaloisGroup ⟶ L₂.unop.finGaloisGroup := + haveI : Normal k L₂.unop := IsGalois.to_normal + letI : Algebra L₂.unop L₁.unop := RingHom.toAlgebra (Subsemiring.inclusion <| leOfHom le.1) + haveI : IsScalarTower k L₂.unop L₁.unop := IsScalarTower.of_algebraMap_eq (congrFun rfl) + FiniteGrp.ofHom (AlgEquiv.restrictNormalHom L₂.unop) + +namespace finGaloisGroupMap + +@[simp] +lemma map_id (L : (FiniteGaloisIntermediateField k K)ᵒᵖ) : + (finGaloisGroupMap (𝟙 L)) = 𝟙 L.unop.finGaloisGroup := + AlgEquiv.restrictNormalHom_id _ _ + +@[simp] +lemma map_comp {L₁ L₂ L₃ : (FiniteGaloisIntermediateField k K)ᵒᵖ} (f : L₁ ⟶ L₂) (g : L₂ ⟶ L₃) : + finGaloisGroupMap (f ≫ g) = finGaloisGroupMap f ≫ finGaloisGroupMap g := by + iterate 2 + induction L₁ with | _ L₁ => ?_ + induction L₂ with | _ L₂ => ?_ + induction L₃ with | _ L₃ => ?_ + letI : Algebra L₃ L₂ := RingHom.toAlgebra (Subsemiring.inclusion g.unop.le) + letI : Algebra L₂ L₁ := RingHom.toAlgebra (Subsemiring.inclusion f.unop.le) + letI : Algebra L₃ L₁ := RingHom.toAlgebra (Subsemiring.inclusion (g.unop.le.trans f.unop.le)) + haveI : IsScalarTower k L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower k L₃ L₁ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower k L₃ L₂ := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsScalarTower L₃ L₂ L₁ := IsScalarTower.of_algebraMap_eq' rfl + apply IsScalarTower.AlgEquiv.restrictNormalHom_comp k L₃ L₂ L₁ + +end finGaloisGroupMap + +variable (k K) in +/-- The functor from `FiniteGaloisIntermediateField` (ordered by reverse inclusion) to `FiniteGrp`, +mapping each intermediate field `K/L/k` to `Gal (L/k)`.-/ +noncomputable def finGaloisGroupFunctor : (FiniteGaloisIntermediateField k K)ᵒᵖ ⥤ FiniteGrp where + obj L := L.unop.finGaloisGroup + map := finGaloisGroupMap + map_id := finGaloisGroupMap.map_id + map_comp := finGaloisGroupMap.map_comp + +end Profinite From 7ca6dfcd941c73f5993541abab8addd016f5208c Mon Sep 17 00:00:00 2001 From: leanprover-community-bot-assistant Date: Sun, 20 Oct 2024 00:15:00 +0000 Subject: [PATCH 116/131] chore(scripts): update nolints.json (#17958) I am happy to remove some nolints for you! --- scripts/nolints.json | 1 - 1 file changed, 1 deletion(-) diff --git a/scripts/nolints.json b/scripts/nolints.json index 8b556db384b5e..438084797d9f2 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -502,7 +502,6 @@ ["docBlame", "Mathlib.Tactic.tacticMatch_target_"], ["docBlame", "Mathlib.Tactic.tacticSet!_"], ["docBlame", "Mathlib.Tactic.tacticSuffices_"], - ["docBlame", "Mathlib.Tactic.tacticTransitivity___"], ["docBlame", "Mathlib.Tactic.usingArg"], ["docBlame", "Mathlib.Tactic.withArgs"], ["docBlame", "Mathlib.WhatsNew.diffExtension"], From 17ca754012e0011ae092c49c91174c1cbe11a007 Mon Sep 17 00:00:00 2001 From: FR Date: Sun, 20 Oct 2024 07:05:32 +0000 Subject: [PATCH 117/131] =?UTF-8?q?chore:=20avoid=20some=20`(0=20<=20((?= =?UTF-8?q?=E2=86=91)=20:=20=E2=84=9D=E2=89=A50=20=E2=86=92=20=E2=84=9D)?= =?UTF-8?q?=20x)=20=3D=20(0=20<=20x)`=20defeq=20abuses=20(#17946)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean | 6 +++--- Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean | 4 ++-- Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean | 2 +- 3 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 8164bbd406fee..31f90aefe8f74 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -535,9 +535,9 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq_inner {u : E → F'} by_cases h3u : ∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ = 0 · rw [eLpNorm_nnreal_eq_lintegral h0p', h3u, ENNReal.zero_rpow_of_pos] <;> positivity have h4u : ∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ ≠ ∞ := by - refine lintegral_rpow_nnnorm_lt_top_of_eLpNorm'_lt_top (pos_iff_ne_zero.mpr h0p') ?_ |>.ne - dsimp only - rw [NNReal.val_eq_coe, ← eLpNorm_nnreal_eq_eLpNorm' h0p'] + refine lintegral_rpow_nnnorm_lt_top_of_eLpNorm'_lt_top + ((NNReal.coe_pos.trans pos_iff_ne_zero).mpr h0p') ?_ |>.ne + rw [← eLpNorm_nnreal_eq_eLpNorm' h0p'] exact hu.continuous.memℒp_of_hasCompactSupport (μ := μ) h2u |>.eLpNorm_lt_top have h5u : (∫⁻ x, ‖u x‖₊ ^ (p' : ℝ) ∂μ) ^ (1 / q) ≠ 0 := ENNReal.rpow_pos (pos_iff_ne_zero.mpr h3u) h4u |>.ne' diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index d2306464d57ff..acdcc41b5be97 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -79,7 +79,7 @@ theorem one_rpow (x : ℝ) : (1 : ℝ≥0) ^ x = 1 := NNReal.eq <| Real.one_rpow _ theorem rpow_add {x : ℝ≥0} (hx : x ≠ 0) (y z : ℝ) : x ^ (y + z) = x ^ y * x ^ z := - NNReal.eq <| Real.rpow_add (pos_iff_ne_zero.2 hx) _ _ + NNReal.eq <| Real.rpow_add ((NNReal.coe_pos.trans pos_iff_ne_zero).mpr hx) _ _ theorem rpow_add' (h : y + z ≠ 0) (x : ℝ≥0) : x ^ (y + z) = x ^ y * x ^ z := NNReal.eq <| Real.rpow_add' x.2 h @@ -146,7 +146,7 @@ lemma rpow_mul_intCast (x : ℝ≥0) (y : ℝ) (n : ℤ) : x ^ (y * n) = (x ^ y) theorem rpow_neg_one (x : ℝ≥0) : x ^ (-1 : ℝ) = x⁻¹ := by simp [rpow_neg] theorem rpow_sub {x : ℝ≥0} (hx : x ≠ 0) (y z : ℝ) : x ^ (y - z) = x ^ y / x ^ z := - NNReal.eq <| Real.rpow_sub (pos_iff_ne_zero.2 hx) y z + NNReal.eq <| Real.rpow_sub ((NNReal.coe_pos.trans pos_iff_ne_zero).mpr hx) y z theorem rpow_sub' (h : y - z ≠ 0) (x : ℝ≥0) : x ^ (y - z) = x ^ y / x ^ z := NNReal.eq <| Real.rpow_sub' x.2 h diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 5bda6afb8f9e1..ee75398764b6e 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -139,7 +139,7 @@ alias lintegral_rpow_nnnorm_eq_rpow_snorm' := lintegral_rpow_nnnorm_eq_rpow_eLpN lemma eLpNorm_nnreal_pow_eq_lintegral {f : α → F} {p : ℝ≥0} (hp : p ≠ 0) : eLpNorm f p μ ^ (p : ℝ) = ∫⁻ x, ‖f x‖₊ ^ (p : ℝ) ∂μ := by simp [eLpNorm_eq_eLpNorm' (by exact_mod_cast hp) ENNReal.coe_ne_top, - lintegral_rpow_nnnorm_eq_rpow_eLpNorm' (show 0 < (p : ℝ) from pos_iff_ne_zero.mpr hp)] + lintegral_rpow_nnnorm_eq_rpow_eLpNorm' ((NNReal.coe_pos.trans pos_iff_ne_zero).mpr hp)] @[deprecated (since := "2024-07-27")] alias snorm_nnreal_pow_eq_lintegral := eLpNorm_nnreal_pow_eq_lintegral From 2f60fa67d7c9457a0e591806fc3201269c401030 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 20 Oct 2024 07:37:59 +0000 Subject: [PATCH 118/131] =?UTF-8?q?feat(Kernel):=20`map=20(=CE=BA=20=C3=97?= =?UTF-8?q?=E2=82=96=20=CE=B7)=20Prod.swap=20=3D=20=CE=B7=20=C3=97?= =?UTF-8?q?=E2=82=96=20=CE=BA`=20(#17920)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and similar lemmas. From PFR Co-authored-by: Rémy Degenne --- Mathlib/Probability/Kernel/Composition.lean | 92 ++++++++++++++----- .../Kernel/Disintegration/Basic.lean | 2 +- .../Kernel/Disintegration/CDFToKernel.lean | 2 +- .../Kernel/Disintegration/Integral.lean | 4 +- .../Kernel/Disintegration/StandardBorel.lean | 2 +- .../Kernel/Disintegration/Unique.lean | 2 +- .../Probability/Kernel/IntegralCompProd.lean | 2 +- .../Probability/Kernel/MeasureCompProd.lean | 2 +- 8 files changed, 79 insertions(+), 29 deletions(-) diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index fd42f53736b76..9a2db5e2be4bc 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -208,8 +208,8 @@ theorem compProd_of_not_isSFiniteKernel_right (κ : Kernel α β) (η : Kernel ( rw [compProd, dif_neg] simp [h] -theorem compProd_apply (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel (α × β) γ) - [IsSFiniteKernel η] (a : α) (hs : MeasurableSet s) : +theorem compProd_apply (hs : MeasurableSet s) (κ : Kernel α β) [IsSFiniteKernel κ] + (η : Kernel (α × β) γ) [IsSFiniteKernel η] (a : α) : (κ ⊗ₖ η) a s = ∫⁻ b, η (a, b) {c | (b, c) ∈ s} ∂κ a := compProd_apply_eq_compProdFun κ η a hs @@ -229,7 +229,7 @@ lemma compProd_zero_left (κ : Kernel (α × β) γ) : (0 : Kernel α β) ⊗ₖ κ = 0 := by by_cases h : IsSFiniteKernel κ · ext a s hs - rw [Kernel.compProd_apply _ _ _ hs] + rw [Kernel.compProd_apply hs] simp · rw [Kernel.compProd_of_not_isSFiniteKernel_right _ _ h] @@ -238,10 +238,42 @@ lemma compProd_zero_right (κ : Kernel α β) (γ : Type*) {mγ : MeasurableSpac κ ⊗ₖ (0 : Kernel (α × β) γ) = 0 := by by_cases h : IsSFiniteKernel κ · ext a s hs - rw [Kernel.compProd_apply _ _ _ hs] + rw [Kernel.compProd_apply hs] simp · rw [Kernel.compProd_of_not_isSFiniteKernel_left _ _ h] +lemma compProd_preimage_fst {s : Set β} (hs : MeasurableSet s) (κ : Kernel α β) + (η : Kernel (α × β) γ) [IsSFiniteKernel κ] [IsMarkovKernel η] (x : α) : + (κ ⊗ₖ η) x (Prod.fst ⁻¹' s) = κ x s := by + rw [compProd_apply (measurable_fst hs)] + simp only [Set.mem_preimage] + classical + have : ∀ b : β, η (x, b) {_c | b ∈ s} = s.indicator (fun _ ↦ 1) b := by + intro b + by_cases hb : b ∈ s <;> simp [hb] + simp_rw [this] + rw [lintegral_indicator_const hs, one_mul] + +lemma compProd_deterministic_apply [MeasurableSingletonClass γ] {f : α × β → γ} (hf : Measurable f) + {s : Set (β × γ)} (hs : MeasurableSet s) (κ : Kernel α β) [IsSFiniteKernel κ] (x : α) : + (κ ⊗ₖ deterministic f hf) x s = κ x {b | (b, f (x, b)) ∈ s} := by + simp only [deterministic_apply, measurableSet_setOf, Set.mem_setOf_eq, Measure.dirac_apply, + Set.mem_setOf_eq, Set.indicator_apply, Pi.one_apply, compProd_apply hs] + let t := {b | (b, f (x, b)) ∈ s} + have ht : MeasurableSet t := (measurable_id.prod_mk (hf.comp measurable_prod_mk_left)) hs + rw [← lintegral_add_compl _ ht] + convert add_zero _ + · suffices ∀ b ∈ tᶜ, (if (b, f (x, b)) ∈ s then (1 : ℝ≥0∞) else 0) = 0 by + rw [setLIntegral_congr_fun ht.compl (ae_of_all _ this), lintegral_zero] + intro b hb + simp only [t, Set.mem_compl_iff, Set.mem_setOf_eq] at hb + simp [hb] + · suffices ∀ b ∈ t, (if (b, f (x, b)) ∈ s then (1 : ℝ≥0∞) else 0) = 1 by + rw [setLIntegral_congr_fun ht (ae_of_all _ this), setLIntegral_one] + intro b hb + simp only [t, Set.mem_setOf_eq] at hb + simp [hb] + section Ae /-! ### `ae` filter of the composition-product -/ @@ -257,14 +289,14 @@ theorem ae_kernel_lt_top (a : α) (h2s : (κ ⊗ₖ η) a s ≠ ∞) : have ht : MeasurableSet t := measurableSet_toMeasurable _ _ have h2t : (κ ⊗ₖ η) a t ≠ ∞ := by rwa [measure_toMeasurable] have ht_lt_top : ∀ᵐ b ∂κ a, η (a, b) (Prod.mk b ⁻¹' t) < ∞ := by - rw [Kernel.compProd_apply _ _ _ ht] at h2t + rw [Kernel.compProd_apply ht] at h2t exact ae_lt_top (Kernel.measurable_kernel_prod_mk_left' ht a) h2t filter_upwards [ht_lt_top] with b hb exact (this b).trans_lt hb theorem compProd_null (a : α) (hs : MeasurableSet s) : (κ ⊗ₖ η) a s = 0 ↔ (fun b => η (a, b) (Prod.mk b ⁻¹' s)) =ᵐ[κ a] 0 := by - rw [Kernel.compProd_apply _ _ _ hs, lintegral_eq_zero_iff] + rw [Kernel.compProd_apply hs, lintegral_eq_zero_iff] · rfl · exact Kernel.measurable_kernel_prod_mk_left' hs a @@ -308,8 +340,8 @@ variable {κ : Kernel α β} [IsSFiniteKernel κ] {η : Kernel (α × β) γ} [I theorem compProd_restrict {s : Set β} {t : Set γ} (hs : MeasurableSet s) (ht : MeasurableSet t) : Kernel.restrict κ hs ⊗ₖ Kernel.restrict η ht = Kernel.restrict (κ ⊗ₖ η) (hs.prod ht) := by ext a u hu - rw [compProd_apply _ _ _ hu, restrict_apply' _ _ _ hu, - compProd_apply _ _ _ (hu.inter (hs.prod ht))] + rw [compProd_apply hu, restrict_apply' _ _ _ hu, + compProd_apply (hu.inter (hs.prod ht))] simp only [Kernel.restrict_apply, Measure.restrict_apply' ht, Set.mem_inter_iff, Set.prod_mk_mem_set_prod_eq] have : @@ -383,7 +415,7 @@ theorem lintegral_compProd' (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kerne simp (config := { unfoldPartialApp := true }) only [SimpleFunc.const_zero, SimpleFunc.coe_piecewise, SimpleFunc.coe_const, SimpleFunc.coe_zero, Set.piecewise_eq_indicator, Function.const, lintegral_indicator_const hs] - rw [compProd_apply κ η _ hs, ← lintegral_const_mul c _] + rw [compProd_apply hs, ← lintegral_const_mul c _] swap · exact (measurable_kernel_prod_mk_left ((measurable_fst.snd.prod_mk measurable_snd) hs)).comp measurable_prod_mk_left @@ -485,11 +517,8 @@ theorem compProd_eq_sum_compProd_right (κ : Kernel α β) (η : Kernel (α × rw [Kernel.sum_comm] instance IsMarkovKernel.compProd (κ : Kernel α β) [IsMarkovKernel κ] (η : Kernel (α × β) γ) - [IsMarkovKernel η] : IsMarkovKernel (κ ⊗ₖ η) := - ⟨fun a => - ⟨by - rw [compProd_apply κ η a MeasurableSet.univ] - simp only [Set.mem_univ, Set.setOf_true, measure_univ, lintegral_one]⟩⟩ + [IsMarkovKernel η] : IsMarkovKernel (κ ⊗ₖ η) where + isProbabilityMeasure a := ⟨by simp [compProd_apply]⟩ theorem compProd_apply_univ_le (κ : Kernel α β) (η : Kernel (α × β) γ) [IsFiniteKernel η] (a : α) : (κ ⊗ₖ η) a Set.univ ≤ κ a Set.univ * IsFiniteKernel.bound η := by @@ -497,7 +526,7 @@ theorem compProd_apply_univ_le (κ : Kernel α β) (η : Kernel (α × β) γ) [ swap · rw [compProd_of_not_isSFiniteKernel_left _ _ hκ] simp - rw [compProd_apply κ η a MeasurableSet.univ] + rw [compProd_apply .univ] simp only [Set.mem_univ, Set.setOf_true] let Cη := IsFiniteKernel.bound η calc @@ -530,13 +559,13 @@ instance IsSFiniteKernel.compProd (κ : Kernel α β) (η : Kernel (α × β) γ lemma compProd_add_left (μ κ : Kernel α β) (η : Kernel (α × β) γ) [IsSFiniteKernel μ] [IsSFiniteKernel κ] [IsSFiniteKernel η] : - (μ + κ) ⊗ₖ η = μ ⊗ₖ η + κ ⊗ₖ η := by ext _ _ hs; simp [compProd_apply _ _ _ hs] + (μ + κ) ⊗ₖ η = μ ⊗ₖ η + κ ⊗ₖ η := by ext _ _ hs; simp [compProd_apply hs] lemma compProd_add_right (μ : Kernel α β) (κ η : Kernel (α × β) γ) [IsSFiniteKernel μ] [IsSFiniteKernel κ] [IsSFiniteKernel η] : μ ⊗ₖ (κ + η) = μ ⊗ₖ κ + μ ⊗ₖ η := by ext a s hs - simp only [compProd_apply _ _ _ hs, coe_add, Pi.add_apply, Measure.coe_add] + simp only [compProd_apply hs, coe_add, Pi.add_apply, Measure.coe_add] rw [lintegral_add_left] exact measurable_kernel_prod_mk_left' hs a @@ -545,7 +574,7 @@ lemma comapRight_compProd_id_prod {δ : Type*} {mδ : MeasurableSpace δ} {f : δ → γ} (hf : MeasurableEmbedding f) : comapRight (κ ⊗ₖ η) (MeasurableEmbedding.id.prod_mk hf) = κ ⊗ₖ (comapRight η hf) := by ext a t ht - rw [comapRight_apply' _ _ _ ht, compProd_apply, compProd_apply _ _ _ ht] + rw [comapRight_apply' _ _ _ ht, compProd_apply, compProd_apply ht] swap; · exact (MeasurableEmbedding.id.prod_mk hf).measurableSet_image.mpr ht refine lintegral_congr (fun b ↦ ?_) simp only [id_eq, Set.mem_image, Prod.mk.injEq, Prod.exists] @@ -1136,8 +1165,7 @@ section Prod /-! ### Product of two kernels -/ - -variable {γ : Type*} {mγ : MeasurableSpace γ} +variable {γ δ : Type*} {mγ : MeasurableSpace γ} {mδ : MeasurableSpace δ} /-- Product of two kernels. This is meaningful only when the kernels are s-finite. -/ noncomputable def prod (κ : Kernel α β) (η : Kernel α γ) : Kernel α (β × γ) := @@ -1148,7 +1176,7 @@ scoped[ProbabilityTheory] infixl:100 " ×ₖ " => ProbabilityTheory.Kernel.prod theorem prod_apply' (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel α γ) [IsSFiniteKernel η] (a : α) {s : Set (β × γ)} (hs : MeasurableSet s) : (κ ×ₖ η) a s = ∫⁻ b : β, (η a) {c : γ | (b, c) ∈ s} ∂κ a := by - simp_rw [prod, compProd_apply _ _ _ hs, swapLeft_apply _ _, prodMkLeft_apply, Prod.swap_prod_mk] + simp_rw [prod, compProd_apply hs, swapLeft_apply _ _, prodMkLeft_apply, Prod.swap_prod_mk] lemma prod_apply (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel α γ) [IsSFiniteKernel η] (a : α) : @@ -1192,6 +1220,28 @@ instance IsSFiniteKernel.prod (κ : Kernel α β) (η : Kernel α γ) : snd (κ ×ₖ η) = η := by ext x; simp [snd_apply, prod_apply] +lemma comap_prod_swap (κ : Kernel α β) (η : Kernel γ δ) [IsSFiniteKernel κ] [IsSFiniteKernel η] : + comap (prodMkRight α η ×ₖ prodMkLeft γ κ) Prod.swap measurable_swap + = map (prodMkRight γ κ ×ₖ prodMkLeft α η) Prod.swap := by + rw [ext_fun_iff] + intro x f hf + rw [lintegral_comap, lintegral_map _ measurable_swap _ hf, lintegral_prod _ _ _ hf, + lintegral_prod] + swap; · exact hf.comp measurable_swap + simp only [prodMkRight_apply, Prod.fst_swap, Prod.swap_prod_mk, lintegral_prodMkLeft, + Prod.snd_swap] + refine (lintegral_lintegral_swap ?_).symm + exact (hf.comp measurable_swap).aemeasurable + +lemma map_prod_swap (κ : Kernel α β) (η : Kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel η] : + map (κ ×ₖ η) Prod.swap = η ×ₖ κ := by + rw [ext_fun_iff] + intro x f hf + rw [lintegral_map _ measurable_swap _ hf, lintegral_prod, lintegral_prod _ _ _ hf] + swap; · exact hf.comp measurable_swap + refine (lintegral_lintegral_swap ?_).symm + exact hf.aemeasurable + lemma deterministic_prod_deterministic {f : α → β} {g : α → γ} (hf : Measurable f) (hg : Measurable g) : deterministic f hf ×ₖ deterministic g hg diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 85c938f894937..22934beefaf08 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -163,7 +163,7 @@ instance condKernelCountable.instIsCondKernel [∀ a, IsMarkovKernel (κCond a)] constructor ext a s hs conv_rhs => rw [← (κ a).disintegrate (κCond a)] - simp_rw [compProd_apply _ _ _ hs, condKernelCountable_apply, Measure.compProd_apply hs] + simp_rw [compProd_apply hs, condKernelCountable_apply, Measure.compProd_apply hs] congr end Countable diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index 3655fb84066ea..cf2267887837b 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -684,7 +684,7 @@ lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) lemma compProd_toKernel [IsFiniteKernel κ] [IsSFiniteKernel ν] (hf : IsCondKernelCDF f κ ν) : ν ⊗ₖ hf.toKernel f = κ := by ext a s hs - rw [Kernel.compProd_apply _ _ _ hs, lintegral_toKernel_mem hf a hs] + rw [Kernel.compProd_apply hs, lintegral_toKernel_mem hf a hs] end diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index 6f9b34c32006f..b3b312597fb26 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -39,14 +39,14 @@ variable [CountableOrCountablyGenerated α β] {κ : Kernel α (β × Ω)} [IsFi lemma lintegral_condKernel_mem (a : α) {s : Set (β × Ω)} (hs : MeasurableSet s) : ∫⁻ x, Kernel.condKernel κ (a, x) {y | (x, y) ∈ s} ∂(Kernel.fst κ a) = κ a s := by conv_rhs => rw [← κ.disintegrate κ.condKernel] - simp_rw [Kernel.compProd_apply _ _ _ hs] + simp_rw [Kernel.compProd_apply hs] lemma setLIntegral_condKernel_eq_measure_prod (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : ∫⁻ b in s, Kernel.condKernel κ (a, b) t ∂(Kernel.fst κ a) = κ a (s ×ˢ t) := by have : κ a (s ×ˢ t) = (Kernel.fst κ ⊗ₖ Kernel.condKernel κ) a (s ×ˢ t) := by congr; exact (κ.disintegrate _).symm - rw [this, Kernel.compProd_apply _ _ _ (hs.prod ht)] + rw [this, Kernel.compProd_apply (hs.prod ht)] classical have : ∀ b, Kernel.condKernel κ (a, b) {c | (b, c) ∈ s ×ˢ t} = s.indicator (fun b ↦ Kernel.condKernel κ (a, b) t) b := by diff --git a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean index db4a4a141b457..b6934d43877ad 100644 --- a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean @@ -259,7 +259,7 @@ lemma compProd_fst_borelMarkovFromReal_eq_comapRight_compProd congr rw [h_fst] ext a t ht : 2 - simp_rw [compProd_apply _ _ _ ht] + simp_rw [compProd_apply ht] refine lintegral_congr_ae ?_ have h_ae : ∀ᵐ t ∂(fst κ a), (a, t) ∈ {p : α × β | η p (range e)ᶜ = 0} := by rw [← h_fst] diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index 3a2d73da4ca97..786669c638b32 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -134,7 +134,7 @@ lemma Kernel.apply_eq_measure_condKernel_of_compProd_eq have : ρ a = (ρ a).fst ⊗ₘ Kernel.comap κ (fun b ↦ (a, b)) measurable_prod_mk_left := by ext s hs conv_lhs => rw [← hκ] - rw [Measure.compProd_apply hs, Kernel.compProd_apply _ _ _ hs] + rw [Measure.compProd_apply hs, Kernel.compProd_apply hs] rfl have h := eq_condKernel_of_measure_eq_compProd _ this rw [Kernel.fst_apply] diff --git a/Mathlib/Probability/Kernel/IntegralCompProd.lean b/Mathlib/Probability/Kernel/IntegralCompProd.lean index a8ad7e64ece6c..e88e7a0f5f916 100644 --- a/Mathlib/Probability/Kernel/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/IntegralCompProd.lean @@ -234,7 +234,7 @@ theorem integral_compProd : rotate_left · exact (Kernel.measurable_kernel_prod_mk_left' hs _).aemeasurable · exact ae_kernel_lt_top a h2s.ne - rw [Kernel.compProd_apply _ _ _ hs] + rw [Kernel.compProd_apply hs] rfl · intro f g _ i_f i_g hf hg simp_rw [integral_add' i_f i_g, Kernel.integral_integral_add' i_f i_g, hf, hg] diff --git a/Mathlib/Probability/Kernel/MeasureCompProd.lean b/Mathlib/Probability/Kernel/MeasureCompProd.lean index 9e4b40f0b1773..b2e9841c037ba 100644 --- a/Mathlib/Probability/Kernel/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/MeasureCompProd.lean @@ -55,7 +55,7 @@ lemma compProd_of_not_isSFiniteKernel (μ : Measure α) (κ : Kernel α β) (h : lemma compProd_apply [SFinite μ] [IsSFiniteKernel κ] {s : Set (α × β)} (hs : MeasurableSet s) : (μ ⊗ₘ κ) s = ∫⁻ a, κ a (Prod.mk a ⁻¹' s) ∂μ := by - simp_rw [compProd, Kernel.compProd_apply _ _ _ hs, Kernel.const_apply, Kernel.prodMkLeft_apply'] + simp_rw [compProd, Kernel.compProd_apply hs, Kernel.const_apply, Kernel.prodMkLeft_apply'] rfl lemma compProd_apply_prod [SFinite μ] [IsSFiniteKernel κ] From a555db73ee24eb1fe6ced83622d96b66e54d1d36 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 20 Oct 2024 10:03:18 +0000 Subject: [PATCH 119/131] chore(Filter): move more defs to `Defs` (#17949) --- Mathlib/Order/Filter/Defs.lean | 10 ++++++++++ Mathlib/Order/Filter/Lift.lean | 10 ---------- 2 files changed, 10 insertions(+), 10 deletions(-) diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean index 06e6369f2f301..4b9a6cb4f9a6b 100644 --- a/Mathlib/Order/Filter/Defs.lean +++ b/Mathlib/Order/Filter/Defs.lean @@ -340,6 +340,16 @@ instance : Bind Filter := instance : Functor Filter where map := @Filter.map +/-- A variant on `bind` using a function `g` taking a set instead of a member of `α`. +This is essentially a push-forward along a function mapping each set to a filter. -/ +protected def lift (f : Filter α) (g : Set α → Filter β) := + ⨅ s ∈ f, g s + +/-- Specialize `lift` to functions `Set α → Set β`. This can be viewed as a generalization of `map`. +This is essentially a push-forward along a function mapping each set to a set. -/ +protected def lift' (f : Filter α) (h : Set α → Set β) := + f.lift (𝓟 ∘ h) + end Filter namespace Mathlib.Tactic diff --git a/Mathlib/Order/Filter/Lift.lean b/Mathlib/Order/Filter/Lift.lean index 95e8d4708a1be..f485af84c1df0 100644 --- a/Mathlib/Order/Filter/Lift.lean +++ b/Mathlib/Order/Filter/Lift.lean @@ -19,11 +19,6 @@ variable {α β γ : Type*} {ι : Sort*} section lift -/-- A variant on `bind` using a function `g` taking a set instead of a member of `α`. -This is essentially a push-forward along a function mapping each set to a filter. -/ -protected def lift (f : Filter α) (g : Set α → Filter β) := - ⨅ s ∈ f, g s - variable {f f₁ f₂ : Filter α} {g g₁ g₂ : Set α → Filter β} @[simp] @@ -200,11 +195,6 @@ end lift section Lift' -/-- Specialize `lift` to functions `Set α → Set β`. This can be viewed as a generalization of `map`. -This is essentially a push-forward along a function mapping each set to a set. -/ -protected def lift' (f : Filter α) (h : Set α → Set β) := - f.lift (𝓟 ∘ h) - variable {f f₁ f₂ : Filter α} {h h₁ h₂ : Set α → Set β} @[simp] From 1ae9a84cb90c97164a22a64da2604a7324d0c9f6 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sun, 20 Oct 2024 11:19:44 +0000 Subject: [PATCH 120/131] feat(AlgebraicGeometry): affine morphisms are separated (#17962) --- .../Morphisms/ClosedImmersion.lean | 42 ++++++++++++++++++- .../Morphisms/Separated.lean | 23 ++++++++-- Mathlib/AlgebraicGeometry/Pullbacks.lean | 13 ++++++ 3 files changed, 72 insertions(+), 6 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index 03d557fa0c469..0e01ff0984171 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -3,10 +3,12 @@ Copyright (c) 2023 Jonas van der Schaaf. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Christian Merten, Jonas van der Schaaf -/ -import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated +import Mathlib.AlgebraicGeometry.Morphisms.Affine import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties import Mathlib.AlgebraicGeometry.Morphisms.FiniteType +import Mathlib.AlgebraicGeometry.Morphisms.IsIso import Mathlib.AlgebraicGeometry.ResidueField +import Mathlib.AlgebraicGeometry.Properties /-! @@ -110,7 +112,7 @@ instance spec_of_quotient_mk {R : CommRingCat.{u}} (I : Ideal R) : /-- Any morphism between affine schemes that is surjective on global sections is a closed immersion. -/ lemma of_surjective_of_isAffine {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X ⟶ Y) - (h : Function.Surjective (Scheme.Γ.map f.op)) : IsClosedImmersion f := by + (h : Function.Surjective (f.app ⊤)) : IsClosedImmersion f := by rw [MorphismProperty.arrow_mk_iso_iff @IsClosedImmersion (arrowIsoSpecΓOfIsAffine f)] apply spec_of_surjective exact h @@ -262,6 +264,17 @@ instance IsClosedImmersion.hasAffineProperty : HasAffineProperty @IsClosedImmers convert HasAffineProperty.of_isLocalAtTarget @IsClosedImmersion refine ⟨fun ⟨h₁, h₂⟩ ↦ of_surjective_of_isAffine _ h₂, by apply isAffine_surjective_of_isAffine⟩ +instance (priority := 900) {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsClosedImmersion f] : + IsAffineHom f := by + wlog hY : IsAffine Y + · rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := @IsAffineHom) _ + (iSup_affineOpens_eq_top Y)] + intro U + have H : IsClosedImmersion (f ∣_ U) := IsLocalAtTarget.restrict h U + exact this _ U.2 + rw [HasAffineProperty.iff_of_isAffine (P := @IsAffineHom)] + exact (IsClosedImmersion.isAffine_surjective_of_isAffine f).1 + /-- Being a closed immersion is stable under base change. -/ lemma IsClosedImmersion.stableUnderBaseChange : MorphismProperty.StableUnderBaseChange @IsClosedImmersion := by @@ -285,4 +298,29 @@ instance (priority := 900) {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsClosedImmersi rw [HasRingHomProperty.iff_of_isAffine (P := @LocallyOfFiniteType)] exact RingHom.FiniteType.of_surjective (Scheme.Hom.app f ⊤) hf +/-- A surjective closed immersion is an isomorphism when the target is reduced. -/ +lemma isIso_of_isClosedImmersion_of_surjective {X Y : Scheme.{u}} (f : X ⟶ Y) + [IsClosedImmersion f] [Surjective f] [IsReduced Y] : + IsIso f := by + wlog hY : IsAffine Y + · refine (IsLocalAtTarget.iff_of_openCover (P := .isomorphisms Scheme) Y.affineCover).mpr ?_ + intro i + apply (config := { allowSynthFailures := true }) this + · exact IsClosedImmersion.stableUnderBaseChange.snd _ _ inferInstance + · exact IsLocalAtTarget.of_isPullback (.of_hasPullback f (Y.affineCover.map i)) ‹_› + · exact isReduced_of_isOpenImmersion (Y.affineCover.map i) + · infer_instance + apply IsClosedImmersion.isIso_of_injective_of_isAffine + obtain ⟨hX, hf⟩ := HasAffineProperty.iff_of_isAffine.mp ‹IsClosedImmersion f› + let φ := f.app ⊤ + suffices RingHom.ker φ ≤ nilradical _ by + rwa [nilradical_eq_zero, Submodule.zero_eq_bot, le_bot_iff, + ← RingHom.injective_iff_ker_eq_bot] at this + refine (PrimeSpectrum.zeroLocus_eq_top_iff _).mp ?_ + rw [← range_specComap_of_surjective _ _ hf, Set.top_eq_univ, Set.range_iff_surjective] + have : Surjective (Spec.map (f.app ⊤)) := + (MorphismProperty.arrow_mk_iso_iff @Surjective (arrowIsoSpecΓOfIsAffine f)).mp + (inferInstanceAs (Surjective f)) + exact this.1 + end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean index 369855b7c262d..a1d81cfa22536 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean @@ -14,10 +14,6 @@ import Mathlib.CategoryTheory.MorphismProperty.Limits A morphism of schemes is separated if its diagonal morphism is a closed immmersion. -## TODO - -- Show affine morphisms are separated. - -/ @@ -75,6 +71,25 @@ instance : IsLocalAtTarget @IsSeparated := by rw [isSeparated_eq_diagonal_isClosedImmersion] infer_instance +instance (R S : CommRingCat.{u}) (f : R ⟶ S) : IsSeparated (Spec.map f) := by + constructor + letI := f.toAlgebra + show IsClosedImmersion (Limits.pullback.diagonal (Spec.map (CommRingCat.ofHom (algebraMap R S)))) + rw [diagonal_Spec_map, MorphismProperty.cancel_right_of_respectsIso @IsClosedImmersion] + exact .spec_of_surjective _ fun x ↦ ⟨.tmul R 1 x, + (Algebra.TensorProduct.lmul'_apply_tmul (R := R) (S := S) 1 x).trans (one_mul x)⟩ + +instance (priority := 100) [h : IsAffineHom f] : IsSeparated f := by + wlog hY : IsAffine Y + · rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := @IsSeparated) _ + (iSup_affineOpens_eq_top Y)] + intro U + have H : IsAffineHom (f ∣_ U) := IsLocalAtTarget.restrict h U + exact this _ U.2 + have : IsAffine X := HasAffineProperty.iff_of_isAffine.mp h + rw [MorphismProperty.arrow_mk_iso_iff @IsSeparated (arrowIsoSpecΓOfIsAffine f)] + infer_instance + end IsSeparated end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Pullbacks.lean b/Mathlib/AlgebraicGeometry/Pullbacks.lean index ca37134839d14..712933f44210e 100644 --- a/Mathlib/AlgebraicGeometry/Pullbacks.lean +++ b/Mathlib/AlgebraicGeometry/Pullbacks.lean @@ -565,6 +565,7 @@ def pullbackSpecIso : letI H := IsLimit.equivIsoLimit (PullbackCone.eta _) (PushoutCocone.isColimitEquivIsLimitOp _ (CommRingCat.pushoutCoconeIsColimit R S T)) limit.isoLimitCone ⟨_, isLimitPullbackConeMapOfIsLimit Scheme.Spec _ H⟩ + /-- The composition of the inverse of the isomorphism `pullbackSepcIso R S T` (from the pullback of `Spec S ⟶ Spec R` and `Spec T ⟶ Spec R` to `Spec (S ⊗[R] T)`) with the first projection is @@ -575,6 +576,7 @@ the morphism `Spec (S ⊗[R] T) ⟶ Spec S` obtained by applying `Spec.map` to t lemma pullbackSpecIso_inv_fst : (pullbackSpecIso R S T).inv ≫ pullback.fst _ _ = Spec.map (ofHom includeLeftRingHom) := limit.isoLimitCone_inv_π _ _ + /-- The composition of the inverse of the isomorphism `pullbackSepcIso R S T` (from the pullback of `Spec S ⟶ Spec R` and `Spec T ⟶ Spec R` to `Spec (S ⊗[R] T)`) with the second projection is @@ -586,6 +588,7 @@ lemma pullbackSpecIso_inv_snd : (pullbackSpecIso R S T).inv ≫ pullback.snd _ _ = Spec.map (ofHom (R := T) (S := S ⊗[R] T) (toRingHom includeRight)) := limit.isoLimitCone_inv_π _ _ + /-- The composition of the isomorphism `pullbackSepcIso R S T` (from the pullback of `Spec S ⟶ Spec R` and `Spec T ⟶ Spec R` to `Spec (S ⊗[R] T)`) with the morphism @@ -596,6 +599,7 @@ is the first projection. lemma pullbackSpecIso_hom_fst : (pullbackSpecIso R S T).hom ≫ Spec.map (ofHom includeLeftRingHom) = pullback.fst _ _ := by rw [← pullbackSpecIso_inv_fst, Iso.hom_inv_id_assoc] + /-- The composition of the isomorphism `pullbackSepcIso R S T` (from the pullback of `Spec S ⟶ Spec R` and `Spec T ⟶ Spec R` to `Spec (S ⊗[R] T)`) with the morphism @@ -607,6 +611,15 @@ lemma pullbackSpecIso_hom_snd : (pullbackSpecIso R S T).hom ≫ Spec.map (ofHom (toRingHom includeRight)) = pullback.snd _ _ := by rw [← pullbackSpecIso_inv_snd, Iso.hom_inv_id_assoc] +lemma diagonal_Spec_map : + pullback.diagonal (Spec.map (CommRingCat.ofHom (algebraMap R S))) = + Spec.map (CommRingCat.ofHom (Algebra.TensorProduct.lmul' R : S ⊗[R] S →ₐ[R] S).toRingHom) ≫ + (pullbackSpecIso R S S).inv := by + ext1 <;> simp only [pullback.diagonal_fst, pullback.diagonal_snd, ← Spec.map_comp, ← Spec.map_id, + AlgHom.toRingHom_eq_coe, Category.assoc, pullbackSpecIso_inv_fst, pullbackSpecIso_inv_snd] + · congr 1; ext x; show x = Algebra.TensorProduct.lmul' R (S := S) (x ⊗ₜ[R] 1); simp + · congr 1; ext x; show x = Algebra.TensorProduct.lmul' R (S := S) (1 ⊗ₜ[R] x); simp + end Spec From 6a88b5460a66aeed2fb64195c0925fd4d07334f1 Mon Sep 17 00:00:00 2001 From: FR Date: Sun, 20 Oct 2024 12:27:18 +0000 Subject: [PATCH 121/131] chore: move binary recursion on `Nat` into a new file (#15567) This is #13649 with fewer definition changes. --- Mathlib.lean | 1 + Mathlib/Data/Nat/BinaryRec.lean | 157 +++++++++++++++++++++++++++++++ Mathlib/Data/Nat/Bits.lean | 140 +-------------------------- Mathlib/Data/Nat/Bitwise.lean | 10 +- Mathlib/Data/Nat/EvenOddRec.lean | 3 +- Mathlib/Data/Nat/Fib/Basic.lean | 3 +- Mathlib/Data/Num/Basic.lean | 6 +- Mathlib/Data/Tree/Get.lean | 1 + scripts/noshake.json | 4 +- 9 files changed, 172 insertions(+), 153 deletions(-) create mode 100644 Mathlib/Data/Nat/BinaryRec.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3ace00fab33be..0eb264c536463 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2434,6 +2434,7 @@ import Mathlib.Data.NNRat.Lemmas import Mathlib.Data.NNRat.Order import Mathlib.Data.NNReal.Basic import Mathlib.Data.NNReal.Star +import Mathlib.Data.Nat.BinaryRec import Mathlib.Data.Nat.BitIndices import Mathlib.Data.Nat.Bits import Mathlib.Data.Nat.Bitwise diff --git a/Mathlib/Data/Nat/BinaryRec.lean b/Mathlib/Data/Nat/BinaryRec.lean new file mode 100644 index 0000000000000..0edc733fa68b0 --- /dev/null +++ b/Mathlib/Data/Nat/BinaryRec.lean @@ -0,0 +1,157 @@ +/- +Copyright (c) 2017 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Praneeth Kolichala, Yuyang Zhao +-/ + +/-! +# Binary recursion on `Nat` + +This file defines binary recursion on `Nat`. + +## Main results +* `Nat.binaryRec`: A recursion principle for `bit` representations of natural numbers. +* `Nat.binaryRec'`: The same as `binaryRec`, but the induction step can assume that if `n=0`, + the bit being appended is `true`. +* `Nat.binaryRecFromOne`: The same as `binaryRec`, but special casing both 0 and 1 as base cases. +-/ + +universe u + +namespace Nat + +/-- `bit b` appends the digit `b` to the binary representation of its natural number input. -/ +def bit (b : Bool) : Nat → Nat := cond b (2 * · + 1) (2 * ·) + +theorem shiftRight_one (n) : n >>> 1 = n / 2 := rfl + +@[simp] +theorem bit_decide_mod_two_eq_one_shiftRight_one (n : Nat) : bit (n % 2 = 1) (n >>> 1) = n := by + simp only [bit, shiftRight_one] + cases mod_two_eq_zero_or_one n with | _ h => simpa [h] using Nat.div_add_mod n 2 + +theorem bit_testBit_zero_shiftRight_one (n : Nat) : bit (n.testBit 0) (n >>> 1) = n := by + simp + +@[simp] +theorem bit_eq_zero_iff {n : Nat} {b : Bool} : bit b n = 0 ↔ n = 0 ∧ b = false := by + cases n <;> cases b <;> simp [bit, Nat.shiftLeft_succ, Nat.two_mul, ← Nat.add_assoc] + +/-- For a predicate `motive : Nat → Sort u`, if instances can be + constructed for natural numbers of the form `bit b n`, + they can be constructed for any given natural number. -/ +@[inline] +def bitCasesOn {motive : Nat → Sort u} (n) (h : ∀ b n, motive (bit b n)) : motive n := + -- `1 &&& n != 0` is faster than `n.testBit 0`. This may change when we have faster `testBit`. + let x := h (1 &&& n != 0) (n >>> 1) + -- `congrArg motive _ ▸ x` is defeq to `x` in non-dependent case + congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x + +/-- A recursion principle for `bit` representations of natural numbers. + For a predicate `motive : Nat → Sort u`, if instances can be + constructed for natural numbers of the form `bit b n`, + they can be constructed for all natural numbers. -/ +@[elab_as_elim, specialize] +def binaryRec {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) + (n : Nat) : motive n := + if n0 : n = 0 then congrArg motive n0 ▸ z + else + let x := f (1 &&& n != 0) (n >>> 1) (binaryRec z f (n >>> 1)) + congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x +decreasing_by exact bitwise_rec_lemma n0 + +/-- The same as `binaryRec`, but the induction step can assume that if `n=0`, + the bit being appended is `true`-/ +@[elab_as_elim, specialize] +def binaryRec' {motive : Nat → Sort u} (z : motive 0) + (f : ∀ b n, (n = 0 → b = true) → motive n → motive (bit b n)) : + ∀ n, motive n := + binaryRec z fun b n ih => + if h : n = 0 → b = true then f b n h ih + else + have : bit b n = 0 := by + rw [bit_eq_zero_iff] + cases n <;> cases b <;> simp at h ⊢ + congrArg motive this ▸ z + +/-- The same as `binaryRec`, but special casing both 0 and 1 as base cases -/ +@[elab_as_elim, specialize] +def binaryRecFromOne {motive : Nat → Sort u} (z₀ : motive 0) (z₁ : motive 1) + (f : ∀ b n, n ≠ 0 → motive n → motive (bit b n)) : + ∀ n, motive n := + binaryRec' z₀ fun b n h ih => + if h' : n = 0 then + have : bit b n = bit true 0 := by + rw [h', h h'] + congrArg motive this ▸ z₁ + else f b n h' ih + +theorem bit_val (b n) : bit b n = 2 * n + b.toNat := by + cases b <;> rfl + +@[simp] +theorem bit_div_two (b n) : bit b n / 2 = n := by + rw [bit_val, Nat.add_comm, add_mul_div_left, div_eq_of_lt, Nat.zero_add] + · cases b <;> decide + · decide + +@[simp] +theorem bit_mod_two (b n) : bit b n % 2 = b.toNat := by + cases b <;> simp [bit_val, mul_add_mod] + +@[simp] +theorem bit_shiftRight_one (b n) : bit b n >>> 1 = n := + bit_div_two b n + +theorem testBit_bit_zero (b n) : (bit b n).testBit 0 = b := by + simp + +@[simp] +theorem bitCasesOn_bit {motive : Nat → Sort u} (h : ∀ b n, motive (bit b n)) (b : Bool) (n : Nat) : + bitCasesOn (bit b n) h = h b n := by + change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ h _ _ = h b n + generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e + rw [testBit_bit_zero, bit_shiftRight_one] + intros; rfl + +unseal binaryRec in +@[simp] +theorem binaryRec_zero {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) : + binaryRec z f 0 = z := + rfl + +@[simp] +theorem binaryRec_one {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) : + binaryRec (motive := motive) z f 1 = f true 0 z := by + rw [binaryRec] + simp only [add_one_ne_zero, ↓reduceDIte, Nat.reduceShiftRight, binaryRec_zero] + rfl + +/-- +The same as `binaryRec_eq`, +but that one unfortunately requires `f` to be the identity when appending `false` to `0`. +Here, we allow you to explicitly say that that case is not happening, +i.e. supplying `n = 0 → b = true`. -/ +theorem binaryRec_eq' {motive : Nat → Sort u} {z : motive 0} + {f : ∀ b n, motive n → motive (bit b n)} (b n) + (h : f false 0 z = z ∨ (n = 0 → b = true)) : + binaryRec z f (bit b n) = f b n (binaryRec z f n) := by + by_cases h' : bit b n = 0 + case pos => + obtain ⟨rfl, rfl⟩ := bit_eq_zero_iff.mp h' + simp only [Bool.false_eq_true, imp_false, not_true_eq_false, or_false] at h + unfold binaryRec + exact h.symm + case neg => + rw [binaryRec, dif_neg h'] + change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ f _ _ _ = _ + generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e + rw [testBit_bit_zero, bit_shiftRight_one] + intros; rfl + +theorem binaryRec_eq {motive : Nat → Sort u} {z : motive 0} {f : ∀ b n, motive n → motive (bit b n)} + (h : f false 0 z = z) (b n) : + binaryRec z f (bit b n) = f b n (binaryRec z f n) := + binaryRec_eq' b n (.inl h) + +end Nat diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 6ea51b387201e..bdcc8d87a6a7c 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -5,8 +5,8 @@ Authors: Praneeth Kolichala -/ import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.Group.Nat -import Mathlib.Algebra.Group.Units.Basic import Mathlib.Data.Nat.Defs +import Mathlib.Data.Nat.BinaryRec import Mathlib.Data.List.Defs import Mathlib.Tactic.Convert import Mathlib.Tactic.GeneralizeProofs @@ -114,35 +114,9 @@ lemma div2_val (n) : div2 n = n / 2 := by (Nat.add_left_cancel (Eq.trans ?_ (Nat.mod_add_div n 2).symm)) rw [mod_two_of_bodd, bodd_add_div2] -/-- `bit b` appends the digit `b` to the binary representation of its natural number input. -/ -def bit (b : Bool) : ℕ → ℕ := cond b (2 * · + 1) (2 * ·) - -lemma bit_val (b n) : bit b n = 2 * n + b.toNat := by - cases b <;> rfl - lemma bit_decomp (n : Nat) : bit (bodd n) (div2 n) = n := (bit_val _ _).trans <| (Nat.add_comm _ _).trans <| bodd_add_div2 _ -theorem shiftRight_one (n) : n >>> 1 = n / 2 := rfl - -@[simp] -theorem bit_decide_mod_two_eq_one_shiftRight_one (n : Nat) : bit (n % 2 = 1) (n >>> 1) = n := by - simp only [bit, shiftRight_one] - cases mod_two_eq_zero_or_one n with | _ h => simpa [h] using Nat.div_add_mod n 2 - -theorem bit_testBit_zero_shiftRight_one (n : Nat) : bit (n.testBit 0) (n >>> 1) = n := by - simp - -/-- For a predicate `motive : Nat → Sort*`, if instances can be - constructed for natural numbers of the form `bit b n`, - they can be constructed for any given natural number. -/ -@[inline] -def bitCasesOn {motive : Nat → Sort u} (n) (h : ∀ b n, motive (bit b n)) : motive n := - -- `1 &&& n != 0` is faster than `n.testBit 0`. This may change when we have faster `testBit`. - let x := h (1 &&& n != 0) (n >>> 1) - -- `congrArg motive _ ▸ x` is defeq to `x` in non-dependent case - congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x - lemma bit_zero : bit false 0 = 0 := rfl @@ -172,19 +146,6 @@ lemma binaryRec_decreasing (h : n ≠ 0) : div2 n < n := by (lt_of_le_of_ne n.zero_le h.symm) rwa [Nat.mul_one] at this -/-- A recursion principle for `bit` representations of natural numbers. - For a predicate `motive : Nat → Sort*`, if instances can be - constructed for natural numbers of the form `bit b n`, - they can be constructed for all natural numbers. -/ -@[elab_as_elim, specialize] -def binaryRec {motive : Nat → Sort u} (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) - (n : Nat) : motive n := - if n0 : n = 0 then congrArg motive n0 ▸ z - else - let x := f (1 &&& n != 0) (n >>> 1) (binaryRec z f (n >>> 1)) - congrArg motive n.bit_testBit_zero_shiftRight_one ▸ x -decreasing_by exact bitwise_rec_lemma n0 - /-- `size n` : Returns the size of a natural number in bits i.e. the length of its binary representation -/ def size : ℕ → ℕ := @@ -201,20 +162,6 @@ def bits : ℕ → List Bool := def ldiff : ℕ → ℕ → ℕ := bitwise fun a b => a && not b -@[simp] -lemma binaryRec_zero {motive : Nat → Sort u} - (z : motive 0) (f : ∀ b n, motive n → motive (bit b n)) : - binaryRec z f 0 = z := by - rw [binaryRec] - rfl - -@[simp] -lemma binaryRec_one {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) : - binaryRec (motive := C) z f 1 = f true 0 z := by - rw [binaryRec] - simp only [succ_ne_self, ↓reduceDIte, reduceShiftRight, binaryRec_zero] - rfl - /-! bitwise ops -/ lemma bodd_bit (b n) : bodd (bit b n) = b := by @@ -242,13 +189,6 @@ lemma shiftLeft'_sub (b m) : ∀ {n k}, k ≤ n → shiftLeft' b m (n - k) = (sh lemma shiftLeft_sub : ∀ (m : Nat) {n k}, k ≤ n → m <<< (n - k) = (m <<< n) >>> k := fun _ _ _ hk => by simp only [← shiftLeft'_false, shiftLeft'_sub false _ hk] --- Not a `simp` lemma, as later `simp` will be able to prove this. -lemma testBit_bit_zero (b n) : testBit (bit b n) 0 = b := by - rw [testBit, bit] - cases b - · simp [← Nat.mul_two] - · simp [← Nat.mul_two] - lemma bodd_eq_one_and_ne_zero : ∀ n, bodd n = (1 &&& n != 0) | 0 => rfl | 1 => rfl @@ -289,24 +229,6 @@ theorem bit_add' : ∀ (b : Bool) (n m : ℕ), bit b (n + m) = bit b n + bit fal theorem bit_ne_zero (b) {n} (h : n ≠ 0) : bit b n ≠ 0 := by cases b <;> dsimp [bit] <;> omega -@[simp] -theorem bit_div_two (b n) : bit b n / 2 = n := by - rw [bit_val, Nat.add_comm, add_mul_div_left, div_eq_of_lt, Nat.zero_add] - · cases b <;> decide - · decide - -@[simp] -theorem bit_shiftRight_one (b n) : bit b n >>> 1 = n := - bit_div_two b n - -@[simp] -theorem bitCasesOn_bit {motive : ℕ → Sort u} (h : ∀ b n, motive (bit b n)) (b : Bool) (n : ℕ) : - bitCasesOn (bit b n) h = h b n := by - change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ h _ _ = h b n - generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e - rw [testBit_bit_zero, bit_shiftRight_one] - intros; rfl - @[simp] theorem bitCasesOn_bit0 {motive : ℕ → Sort u} (H : ∀ b n, motive (bit b n)) (n : ℕ) : bitCasesOn (2 * n) H = H false n := @@ -328,13 +250,6 @@ theorem bit_cases_on_inj {motive : ℕ → Sort u} (H₁ H₂ : ∀ b n, motive ((fun n => bitCasesOn n H₁) = fun n => bitCasesOn n H₂) ↔ H₁ = H₂ := bit_cases_on_injective.eq_iff -@[simp] -theorem bit_eq_zero_iff {n : ℕ} {b : Bool} : bit b n = 0 ↔ n = 0 ∧ b = false := by - constructor - · cases b <;> simp [Nat.bit]; omega - · rintro ⟨rfl, rfl⟩ - rfl - lemma bit_le : ∀ (b : Bool) {m n : ℕ}, m ≤ n → bit b m ≤ bit b n | true, _, _, h => by dsimp [bit]; omega | false, _, _, h => by dsimp [bit]; omega @@ -343,59 +258,6 @@ lemma bit_lt_bit (a b) (h : m < n) : bit a m < bit b n := calc bit a m < 2 * n := by cases a <;> dsimp [bit] <;> omega _ ≤ bit b n := by cases b <;> dsimp [bit] <;> omega -/-- -The same as `binaryRec_eq`, -but that one unfortunately requires `f` to be the identity when appending `false` to `0`. -Here, we allow you to explicitly say that that case is not happening, -i.e. supplying `n = 0 → b = true`. -/ -theorem binaryRec_eq' {motive : ℕ → Sort*} {z : motive 0} {f : ∀ b n, motive n → motive (bit b n)} - (b n) (h : f false 0 z = z ∨ (n = 0 → b = true)) : - binaryRec z f (bit b n) = f b n (binaryRec z f n) := by - by_cases h' : bit b n = 0 - case pos => - obtain ⟨rfl, rfl⟩ := bit_eq_zero_iff.mp h' - simp only [imp_false, or_false, eq_self_iff_true, not_true, reduceCtorEq] at h - unfold binaryRec - exact h.symm - case neg => - rw [binaryRec, dif_neg h'] - change congrArg motive (bit b n).bit_testBit_zero_shiftRight_one ▸ f _ _ _ = _ - generalize congrArg motive (bit b n).bit_testBit_zero_shiftRight_one = e; revert e - rw [testBit_bit_zero, bit_shiftRight_one] - intros; rfl - -theorem binaryRec_eq {motive : Nat → Sort u} {z : motive 0} - {f : ∀ b n, motive n → motive (bit b n)} - (h : f false 0 z = z) (b n) : - binaryRec z f (bit b n) = f b n (binaryRec z f n) := - binaryRec_eq' b n (.inl h) - -/-- The same as `binaryRec`, but the induction step can assume that if `n=0`, - the bit being appended is `true`-/ -@[elab_as_elim, specialize] -def binaryRec' {motive : ℕ → Sort*} (z : motive 0) - (f : ∀ b n, (n = 0 → b = true) → motive n → motive (bit b n)) : - ∀ n, motive n := - binaryRec z fun b n ih => - if h : n = 0 → b = true then f b n h ih - else - have : bit b n = 0 := by - rw [bit_eq_zero_iff] - cases n <;> cases b <;> simp at h ⊢ - congrArg motive this ▸ z - -/-- The same as `binaryRec`, but special casing both 0 and 1 as base cases -/ -@[elab_as_elim, specialize] -def binaryRecFromOne {motive : ℕ → Sort*} (z₀ : motive 0) (z₁ : motive 1) - (f : ∀ b n, n ≠ 0 → motive n → motive (bit b n)) : - ∀ n, motive n := - binaryRec' z₀ fun b n h ih => - if h' : n = 0 then - have : bit b n = bit true 0 := by - rw [h', h h'] - congrArg motive this ▸ z₁ - else f b n h' ih - @[simp] theorem zero_bits : bits 0 = [] := by simp [Nat.bits] diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 07873ec6fd976..c0cbec11309d7 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -3,13 +3,14 @@ Copyright (c) 2020 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel, Alex Keizer -/ +import Mathlib.Algebra.Group.Units.Basic +import Mathlib.Algebra.NeZero +import Mathlib.Algebra.Ring.Nat import Mathlib.Data.List.GetD import Mathlib.Data.Nat.Bits -import Mathlib.Algebra.Ring.Nat import Mathlib.Order.Basic import Mathlib.Tactic.AdaptationNote import Mathlib.Tactic.Common -import Mathlib.Algebra.NeZero /-! # Bitwise operations on natural numbers @@ -88,11 +89,6 @@ lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by cases a <;> cases b <;> simp [h2, h4] <;> split_ifs <;> simp_all (config := {decide := true}) [two_mul] -@[simp] -lemma bit_mod_two (a : Bool) (x : ℕ) : - bit a x % 2 = a.toNat := by - cases a <;> simp [bit_val, mul_add_mod] - lemma bit_mod_two_eq_zero_iff (a x) : bit a x % 2 = 0 ↔ !a := by simp diff --git a/Mathlib/Data/Nat/EvenOddRec.lean b/Mathlib/Data/Nat/EvenOddRec.lean index a5c452c173847..ce6c89abcd14f 100644 --- a/Mathlib/Data/Nat/EvenOddRec.lean +++ b/Mathlib/Data/Nat/EvenOddRec.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Stuart Presnell -/ import Mathlib.Algebra.Ring.Parity -import Mathlib.Data.Nat.Bits +import Mathlib.Data.Nat.BinaryRec + /-! # A recursion principle based on even and odd numbers. -/ namespace Nat diff --git a/Mathlib/Data/Nat/Fib/Basic.lean b/Mathlib/Data/Nat/Fib/Basic.lean index 587e609b5ab7b..399cabbe6107f 100644 --- a/Mathlib/Data/Nat/Fib/Basic.lean +++ b/Mathlib/Data/Nat/Fib/Basic.lean @@ -5,8 +5,9 @@ Authors: Kevin Kappelmann, Kyle Miller, Mario Carneiro -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Data.Finset.NatAntidiagonal -import Mathlib.Data.Nat.Bits import Mathlib.Data.Nat.GCD.Basic +import Mathlib.Data.Nat.BinaryRec +import Mathlib.Logic.Function.Iterate import Mathlib.Tactic.Ring import Mathlib.Tactic.Zify diff --git a/Mathlib/Data/Num/Basic.lean b/Mathlib/Data/Num/Basic.lean index 5e35ca5550ad8..07f592d41ee94 100644 --- a/Mathlib/Data/Num/Basic.lean +++ b/Mathlib/Data/Num/Basic.lean @@ -4,9 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Mario Carneiro -/ import Lean.Linter.Deprecated -import Mathlib.Data.Int.Notation import Mathlib.Algebra.Group.ZeroOne -import Mathlib.Data.Nat.Bits +import Mathlib.Data.Int.Notation +import Mathlib.Data.Nat.BinaryRec +import Mathlib.Tactic.TypeStar + /-! # Binary representation of integers using inductive types diff --git a/Mathlib/Data/Tree/Get.lean b/Mathlib/Data/Tree/Get.lean index e3adeba9eea04..b374b675b07ba 100644 --- a/Mathlib/Data/Tree/Get.lean +++ b/Mathlib/Data/Tree/Get.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Wojciech Nawrocki -/ import Mathlib.Data.Num.Basic +import Mathlib.Data.Ordering.Basic import Mathlib.Data.Tree.Basic /-! diff --git a/scripts/noshake.json b/scripts/noshake.json index a98fca062e875..1fa2973a3a767 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -41,6 +41,7 @@ "Mathlib.Control.Traversable.Lemmas", "Mathlib.Data.Finsupp.Notation", "Mathlib.Data.Int.Defs", + "Mathlib.Data.Int.Notation", "Mathlib.Data.Matrix.Notation", "Mathlib.Data.Matroid.Init", "Mathlib.Data.Nat.Notation", @@ -336,7 +337,6 @@ "Mathlib.Logic.Function.Defs": ["Mathlib.Tactic.AdaptationNote"], "Mathlib.Logic.Function.Basic": ["Batteries.Tactic.Init"], "Mathlib.Logic.Equiv.Defs": ["Mathlib.Data.Bool.Basic"], - "Mathlib.Logic.Basic": ["Mathlib.Data.Int.Notation"], "Mathlib.LinearAlgebra.Matrix.Transvection": ["Mathlib.Data.Matrix.DMatrix"], "Mathlib.LinearAlgebra.Basic": ["Mathlib.Algebra.BigOperators.Group.Finset"], "Mathlib.LinearAlgebra.AffineSpace.FiniteDimensional": ["Mathlib.Init.Core"], @@ -370,7 +370,6 @@ "Mathlib.Data.List.Basic": ["Mathlib.Control.Basic", "Mathlib.Data.Option.Basic"], "Mathlib.Data.LazyList.Basic": ["Mathlib.Lean.Thunk"], - "Mathlib.Data.Int.Order.Basic": ["Mathlib.Data.Int.Notation"], "Mathlib.Data.Int.Defs": ["Batteries.Data.Int.Order"], "Mathlib.Data.FunLike.Basic": ["Mathlib.Logic.Function.Basic"], "Mathlib.Data.Finset.Basic": ["Mathlib.Data.Finset.Attr"], @@ -423,7 +422,6 @@ "Mathlib.Algebra.Group.Units.Basic": ["Mathlib.Tactic.Subsingleton"], "Mathlib.Algebra.Group.Units": ["Mathlib.Tactic.Subsingleton"], "Mathlib.Algebra.Group.Pi.Basic": ["Batteries.Tactic.Classical"], - "Mathlib.Algebra.Group.Defs": ["Mathlib.Data.Int.Notation"], "Mathlib.Algebra.GeomSum": ["Mathlib.Algebra.Order.BigOperators.Ring.Finset"], "Mathlib.Algebra.Category.Ring.Basic": ["Mathlib.CategoryTheory.ConcreteCategory.ReflectsIso"], From d172902852dabd02b61ebafc692e751d29ecac1e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 20 Oct 2024 12:53:16 +0000 Subject: [PATCH 122/131] chore(Algebra/Order/BigOperators/Group/Finset): don't import `Ring` (#17939) By moving the positivity extension, one can shave off a bunch of imports --- Mathlib/Algebra/BigOperators/Finprod.lean | 1 + Mathlib/Algebra/Order/Antidiag/Finsupp.lean | 1 - .../Order/BigOperators/Group/Finset.lean | 45 ++----------------- Mathlib/Algebra/Order/Chebyshev.lean | 1 - Mathlib/Algebra/Order/Interval/Basic.lean | 1 + Mathlib/Analysis/Analytic/Inverse.lean | 1 + Mathlib/Analysis/Complex/AbelLimit.lean | 1 + .../Enumerative/Composition.lean | 3 +- .../Enumerative/DoubleCounting.lean | 1 + Mathlib/Combinatorics/Pigeonhole.lean | 4 +- .../SetFamily/AhlswedeZhang.lean | 2 - .../SetFamily/FourFunctions.lean | 1 - .../SimpleGraph/Regularity/Energy.lean | 1 - Mathlib/Data/Finsupp/Weight.lean | 4 -- Mathlib/Data/Nat/Choose/Sum.lean | 1 - Mathlib/Data/Set/Equitable.lean | 4 +- Mathlib/Data/Sign.lean | 2 + Mathlib/LinearAlgebra/Matrix/DotProduct.lean | 3 -- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 2 - Mathlib/NumberTheory/Divisors.lean | 2 +- Mathlib/NumberTheory/Harmonic/Defs.lean | 1 + Mathlib/Order/Partition/Equipartition.lean | 1 + Mathlib/Tactic/Positivity/Finset.lean | 37 +++++++++++++++ Mathlib/Testing/SlimCheck/Functions.lean | 2 +- Mathlib/Topology/Algebra/GroupWithZero.lean | 1 + .../Topology/Algebra/WithZeroTopology.lean | 1 + scripts/noshake.json | 3 +- test/positivity.lean | 1 + 28 files changed, 58 insertions(+), 70 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Finprod.lean b/Mathlib/Algebra/BigOperators/Finprod.lean index b6b2d96428073..33329a6e81269 100644 --- a/Mathlib/Algebra/BigOperators/Finprod.lean +++ b/Mathlib/Algebra/BigOperators/Finprod.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.BigOperators.GroupWithZero.Finset import Mathlib.Algebra.Group.FiniteSupport import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Data.Set.Subsingleton /-! diff --git a/Mathlib/Algebra/Order/Antidiag/Finsupp.lean b/Mathlib/Algebra/Order/Antidiag/Finsupp.lean index 3564d4ffb4076..150ff45ab21d7 100644 --- a/Mathlib/Algebra/Order/Antidiag/Finsupp.lean +++ b/Mathlib/Algebra/Order/Antidiag/Finsupp.lean @@ -5,7 +5,6 @@ Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández, Eric Wieser, Yaël Dillies -/ import Mathlib.Algebra.Order.Antidiag.Pi -import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Finsupp.Basic /-! diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index 8593406fe5e67..f799eabddb523 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -5,10 +5,9 @@ Authors: Johannes Hölzl -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Order.BigOperators.Group.Multiset +import Mathlib.Algebra.Order.Group.Nat import Mathlib.Data.Multiset.OrderedMonoid import Mathlib.Tactic.Bound.Attribute -import Mathlib.Tactic.NormNum.Basic -import Mathlib.Tactic.Positivity.Core /-! # Big operators on a finset in ordered groups @@ -17,6 +16,8 @@ This file contains the results concerning the interaction of multiset big operat groups/monoids. -/ +assert_not_exists Ring + open Function variable {ι α β M N G k R : Type*} @@ -581,43 +582,3 @@ theorem sup_powerset_len {α : Type*} [DecidableEq α] (x : Multiset α) : Eq.symm (finset_sum_eq_sup_iff_disjoint.mpr fun _ _ _ _ h => pairwise_disjoint_powersetCard x h) end Multiset - -namespace Mathlib.Meta.Positivity -open Qq Lean Meta Finset - -/-- The `positivity` extension which proves that `∑ i ∈ s, f i` is nonnegative if `f` is, and -positive if each `f i` is and `s` is nonempty. - -TODO: The following example does not work -``` -example (s : Finset ℕ) (f : ℕ → ℤ) (hf : ∀ n, 0 ≤ f n) : 0 ≤ s.sum f := by positivity -``` -because `compareHyp` can't look for assumptions behind binders. --/ -@[positivity Finset.sum _ _] -def evalFinsetSum : PositivityExt where eval {u α} zα pα e := do - match e with - | ~q(@Finset.sum $ι _ $instα $s $f) => - let i : Q($ι) ← mkFreshExprMVarQ q($ι) .syntheticOpaque - have body : Q($α) := .betaRev f #[i] - let rbody ← core zα pα body - let p_pos : Option Q(0 < $e) := ← (do - let .positive pbody := rbody | pure none -- Fail if the body is not provably positive - let .some ps ← proveFinsetNonempty s | pure none - let .some pα' ← trySynthInstanceQ q(OrderedCancelAddCommMonoid $α) | pure none - assertInstancesCommute - let pr : Q(∀ i, 0 < $f i) ← mkLambdaFVars #[i] pbody - return some q(@sum_pos $ι $α $pα' $f $s (fun i _ ↦ $pr i) $ps)) - -- Try to show that the sum is positive - if let some p_pos := p_pos then - return .positive p_pos - -- Fall back to showing that the sum is nonnegative - else - let pbody ← rbody.toNonneg - let pr : Q(∀ i, 0 ≤ $f i) ← mkLambdaFVars #[i] pbody - let pα' ← synthInstanceQ q(OrderedAddCommMonoid $α) - assertInstancesCommute - return .nonnegative q(@sum_nonneg $ι $α $pα' $f $s fun i _ ↦ $pr i) - | _ => throwError "not Finset.sum" - -end Mathlib.Meta.Positivity diff --git a/Mathlib/Algebra/Order/Chebyshev.lean b/Mathlib/Algebra/Order/Chebyshev.lean index 439d79b0c4e2b..7987886b72b00 100644 --- a/Mathlib/Algebra/Order/Chebyshev.lean +++ b/Mathlib/Algebra/Order/Chebyshev.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 Mantas Bakšys, Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mantas Bakšys, Yaël Dillies -/ -import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.Monovary import Mathlib.Algebra.Order.Rearrangement import Mathlib.GroupTheory.Perm.Cycle.Basic diff --git a/Mathlib/Algebra/Order/Interval/Basic.lean b/Mathlib/Algebra/Order/Interval/Basic.lean index 5fc5f067acadf..bcfec00e2cd85 100644 --- a/Mathlib/Algebra/Order/Interval/Basic.lean +++ b/Mathlib/Algebra/Order/Interval/Basic.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Order.Interval.Basic +import Mathlib.Tactic.Positivity.Core /-! # Interval arithmetic diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index 34ad1130c225c..4b94111adff37 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.Analytic.Composition import Mathlib.Analysis.Analytic.Linear +import Mathlib.Tactic.Positivity.Finset /-! diff --git a/Mathlib/Analysis/Complex/AbelLimit.lean b/Mathlib/Analysis/Complex/AbelLimit.lean index 20c8a9420d6f3..25d7f649514c6 100644 --- a/Mathlib/Analysis/Complex/AbelLimit.lean +++ b/Mathlib/Analysis/Complex/AbelLimit.lean @@ -6,6 +6,7 @@ Authors: Jeremy Tan import Mathlib.Analysis.Complex.Basic import Mathlib.Analysis.SpecificLimits.Normed import Mathlib.Tactic.Peel +import Mathlib.Tactic.Positivity.Finset /-! # Abel's limit theorem diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index 917e4a7e41dbc..c825e38016f80 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -6,7 +6,6 @@ Authors: Sébastien Gouëzel import Mathlib.Algebra.BigOperators.Fin import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Finset.Sort -import Mathlib.Data.Set.Subsingleton /-! # Compositions @@ -807,7 +806,7 @@ theorem card_boundaries_eq_succ_length : c.boundaries.card = c.length + 1 := theorem length_lt_card_boundaries : c.length < c.boundaries.card := by rw [c.card_boundaries_eq_succ_length] - exact lt_add_one _ + exact Nat.lt_add_one _ theorem lt_length (i : Fin c.length) : (i : ℕ) + 1 < c.boundaries.card := lt_tsub_iff_right.mp i.2 diff --git a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean index b929f72e6e626..3061124939269 100644 --- a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean +++ b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Ring.Nat /-! # Double countings diff --git a/Mathlib/Combinatorics/Pigeonhole.lean b/Mathlib/Combinatorics/Pigeonhole.lean index 3811b73e8b204..b91d6c658fd57 100644 --- a/Mathlib/Combinatorics/Pigeonhole.lean +++ b/Mathlib/Combinatorics/Pigeonhole.lean @@ -4,10 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller, Yury Kudryashov -/ import Mathlib.Algebra.Module.BigOperators -import Mathlib.Algebra.Module.Defs -import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.Nat.ModEq -import Mathlib.Data.Set.Finite /-! # Pigeonhole principles diff --git a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean index ea08e4e2baeef..73740d7fb77d3 100644 --- a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean +++ b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean @@ -6,10 +6,8 @@ Authors: Yaël Dillies, Vladimir Ivanov import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Algebra.Order.Field.Basic import Mathlib.Data.Finset.Sups import Mathlib.Tactic.FieldSimp -import Mathlib.Tactic.Positivity.Basic import Mathlib.Tactic.Ring /-! diff --git a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean index d3bd37943cd19..df03a09a30705 100644 --- a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean +++ b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.Pi import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Data.Finset.Sups -import Mathlib.Data.Set.Subsingleton import Mathlib.Order.Birkhoff import Mathlib.Order.Booleanisation import Mathlib.Order.Sublattice diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean index 54e5c27b5c906..ad2bb21be97a3 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean @@ -5,7 +5,6 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Combinatorics.SimpleGraph.Density import Mathlib.Data.Rat.BigOperators diff --git a/Mathlib/Data/Finsupp/Weight.lean b/Mathlib/Data/Finsupp/Weight.lean index ed767ee615ac8..075821985671b 100644 --- a/Mathlib/Data/Finsupp/Weight.lean +++ b/Mathlib/Data/Finsupp/Weight.lean @@ -3,11 +3,7 @@ Copyright (c) 2024 Antoine Chambert-Loir, María Inés de Frutos-Fernández. All Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández -/ - -import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.Module.Defs -import Mathlib.Algebra.Order.Ring.Defs -import Mathlib.Algebra.Order.Monoid.Canonical.Defs import Mathlib.LinearAlgebra.Finsupp /-! # weights of Finsupp functions diff --git a/Mathlib/Data/Nat/Choose/Sum.lean b/Mathlib/Data/Nat/Choose/Sum.lean index 04851939e6336..ab6812bc6fcb7 100644 --- a/Mathlib/Data/Nat/Choose/Sum.lean +++ b/Mathlib/Data/Nat/Choose/Sum.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.BigOperators.NatAntidiagonal import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Data.Nat.Choose.Basic import Mathlib.Tactic.Linarith import Mathlib.Tactic.Ring diff --git a/Mathlib/Data/Set/Equitable.lean b/Mathlib/Data/Set/Equitable.lean index 8eb08216d19ac..b2e7669c813a3 100644 --- a/Mathlib/Data/Set/Equitable.lean +++ b/Mathlib/Data/Set/Equitable.lean @@ -3,10 +3,8 @@ Copyright (c) 2021 Yaël Dillies, Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ -import Mathlib.Data.Set.Subsingleton import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Algebra.Group.Nat -import Mathlib.Data.Set.Basic +import Mathlib.Algebra.Order.Ring.Defs /-! # Equitable functions diff --git a/Mathlib/Data/Sign.lean b/Mathlib/Data/Sign.lean index 7e0033c7d8a2a..a492a43728d98 100644 --- a/Mathlib/Data/Sign.lean +++ b/Mathlib/Data/Sign.lean @@ -5,7 +5,9 @@ Authors: Eric Rodriguez -/ import Mathlib.Algebra.GroupWithZero.Units.Lemmas import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Ring.Cast import Mathlib.Data.Fintype.BigOperators + /-! # Sign function diff --git a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean index bb06aeea3eee8..bf76ac67e8efa 100644 --- a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean +++ b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean @@ -3,11 +3,8 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ -import Mathlib.Algebra.Ring.Regular import Mathlib.Algebra.Order.Star.Basic -import Mathlib.Data.Matrix.Basic import Mathlib.LinearAlgebra.StdBasis -import Mathlib.Algebra.Order.BigOperators.Group.Finset /-! # Dot product of two vectors diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index ea1242276149d..f37192686c399 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -5,10 +5,8 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Algebra.Algebra.Defs import Mathlib.Algebra.NoZeroSMulDivisors.Pi -import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Fintype.BigOperators import Mathlib.Data.Fintype.Sort -import Mathlib.Data.List.FinRange import Mathlib.LinearAlgebra.Pi import Mathlib.Logic.Equiv.Fintype import Mathlib.Tactic.Abel diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 3fb893b14cd82..d0f8d0cc61130 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.Nat.PrimeFin import Mathlib.Order.Interval.Finset.Nat diff --git a/Mathlib/NumberTheory/Harmonic/Defs.lean b/Mathlib/NumberTheory/Harmonic/Defs.lean index a7887fff3ff84..1169d86a0e470 100644 --- a/Mathlib/NumberTheory/Harmonic/Defs.lean +++ b/Mathlib/NumberTheory/Harmonic/Defs.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Order.Field.Basic import Mathlib.Tactic.Linarith +import Mathlib.Tactic.Positivity.Finset /-! diff --git a/Mathlib/Order/Partition/Equipartition.lean b/Mathlib/Order/Partition/Equipartition.lean index 7833e98aa3fe8..3969b84ca5eb0 100644 --- a/Mathlib/Order/Partition/Equipartition.lean +++ b/Mathlib/Order/Partition/Equipartition.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Yaël Dillies, Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ +import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.Set.Equitable import Mathlib.Logic.Equiv.Fin import Mathlib.Order.Partition.Finpartition diff --git a/Mathlib/Tactic/Positivity/Finset.lean b/Mathlib/Tactic/Positivity/Finset.lean index 6326f5d927658..bef478ebd38c6 100644 --- a/Mathlib/Tactic/Positivity/Finset.lean +++ b/Mathlib/Tactic/Positivity/Finset.lean @@ -3,7 +3,9 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Data.Finset.Density +import Mathlib.Tactic.NormNum.Basic import Mathlib.Tactic.Positivity.Core /-! @@ -52,6 +54,41 @@ def evalFinsetDens : PositivityExt where eval {u 𝕜} _ _ e := do return .positive q(@Nonempty.dens_pos $α $instα $s $ps) | _, _, _ => throwError "not Finset.dens" +/-- The `positivity` extension which proves that `∑ i ∈ s, f i` is nonnegative if `f` is, and +positive if each `f i` is and `s` is nonempty. + +TODO: The following example does not work +``` +example (s : Finset ℕ) (f : ℕ → ℤ) (hf : ∀ n, 0 ≤ f n) : 0 ≤ s.sum f := by positivity +``` +because `compareHyp` can't look for assumptions behind binders. +-/ +@[positivity Finset.sum _ _] +def evalFinsetSum : PositivityExt where eval {u α} zα pα e := do + match e with + | ~q(@Finset.sum $ι _ $instα $s $f) => + let i : Q($ι) ← mkFreshExprMVarQ q($ι) .syntheticOpaque + have body : Q($α) := .betaRev f #[i] + let rbody ← core zα pα body + let p_pos : Option Q(0 < $e) := ← (do + let .positive pbody := rbody | pure none -- Fail if the body is not provably positive + let .some ps ← proveFinsetNonempty s | pure none + let .some pα' ← trySynthInstanceQ q(OrderedCancelAddCommMonoid $α) | pure none + assertInstancesCommute + let pr : Q(∀ i, 0 < $f i) ← mkLambdaFVars #[i] pbody + return some q(@sum_pos $ι $α $pα' $f $s (fun i _ ↦ $pr i) $ps)) + -- Try to show that the sum is positive + if let some p_pos := p_pos then + return .positive p_pos + -- Fall back to showing that the sum is nonnegative + else + let pbody ← rbody.toNonneg + let pr : Q(∀ i, 0 ≤ $f i) ← mkLambdaFVars #[i] pbody + let pα' ← synthInstanceQ q(OrderedAddCommMonoid $α) + assertInstancesCommute + return .nonnegative q(@sum_nonneg $ι $α $pα' $f $s fun i _ ↦ $pr i) + | _ => throwError "not Finset.sum" + variable {α : Type*} {s : Finset α} example : 0 ≤ s.card := by positivity diff --git a/Mathlib/Testing/SlimCheck/Functions.lean b/Mathlib/Testing/SlimCheck/Functions.lean index 51eadae43cdec..7b73daf4bdbcb 100644 --- a/Mathlib/Testing/SlimCheck/Functions.lean +++ b/Mathlib/Testing/SlimCheck/Functions.lean @@ -289,7 +289,7 @@ theorem List.applyId_zip_eq [DecidableEq α] {xs ys : List α} (h₀ : List.Nodu | nil => cases h₂ | cons x' xs xs_ih => cases i - · simp only [length_cons, lt_add_iff_pos_left, add_pos_iff, zero_lt_one, or_true, + · simp only [length_cons, lt_add_iff_pos_left, add_pos_iff, Nat.lt_add_one, or_true, getElem?_eq_getElem, getElem_cons_zero, Option.some.injEq] at h₂ subst h₂ cases ys diff --git a/Mathlib/Topology/Algebra/GroupWithZero.lean b/Mathlib/Topology/Algebra/GroupWithZero.lean index b601051b87dc5..8874a996043c3 100644 --- a/Mathlib/Topology/Algebra/GroupWithZero.lean +++ b/Mathlib/Topology/Algebra/GroupWithZero.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Pi.Lemmas +import Mathlib.Algebra.GroupWithZero.Units.Equiv import Mathlib.Topology.Algebra.Monoid import Mathlib.Topology.Homeomorph diff --git a/Mathlib/Topology/Algebra/WithZeroTopology.lean b/Mathlib/Topology/Algebra/WithZeroTopology.lean index 8b290b9995546..59382fdd867a4 100644 --- a/Mathlib/Topology/Algebra/WithZeroTopology.lean +++ b/Mathlib/Topology/Algebra/WithZeroTopology.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot -/ +import Mathlib.Algebra.Order.GroupWithZero.Canonical import Mathlib.Topology.Algebra.GroupWithZero import Mathlib.Topology.Order.OrderClosed diff --git a/scripts/noshake.json b/scripts/noshake.json index 1fa2973a3a767..5ee9b18083934 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -219,7 +219,7 @@ "Mathlib.Tactic.ProxyType": ["Mathlib.Logic.Equiv.Defs"], "Mathlib.Tactic.ProdAssoc": ["Mathlib.Logic.Equiv.Defs"], "Mathlib.Tactic.Positivity.Finset": - ["Mathlib.Data.Finset.Density", "Mathlib.Data.Fintype.Card"], + ["Mathlib.Algebra.Order.BigOperators.Group.Finset", "Mathlib.Data.Finset.Density"], "Mathlib.Tactic.Positivity.Basic": ["Mathlib.Algebra.GroupPower.Order", "Mathlib.Algebra.Order.Group.PosPart", @@ -295,7 +295,6 @@ ["Mathlib.Algebra.Group.ZeroOne", "Mathlib.Tactic.Bound.Init"], "Mathlib.Tactic.Bound": ["Mathlib.Algebra.Order.AbsoluteValue", - "Mathlib.Algebra.Order.BigOperators.Group.Finset", "Mathlib.Algebra.Order.Floor", "Mathlib.Analysis.Analytic.Basic", "Mathlib.Tactic.Bound.Init"], diff --git a/test/positivity.lean b/test/positivity.lean index fc773c8abd962..07295214c3bbe 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -4,6 +4,7 @@ import Mathlib.Analysis.Normed.Group.Basic import Mathlib.Analysis.SpecialFunctions.Pow.Real import Mathlib.Analysis.SpecialFunctions.Log.Basic import Mathlib.MeasureTheory.Integral.Bochner +import Mathlib.Tactic.Positivity.Finset /-! # Tests for the `positivity` tactic From 6cd5a568ec89051613dd22e5053bbb64b11b575b Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sun, 20 Oct 2024 13:19:34 +0000 Subject: [PATCH 123/131] feat: induction principles for ContinuousMap via Stone-Weierstrass (#17831) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR provides induction lemmas for `C(s, 𝕜)` and `C(s, 𝕜)₀` with `RCLike 𝕜`. The idea is that when `s` is compact, the `closure` hypothesis will follow from the Stone-Weierstrass theorem and a limit property of the motive `p`. --- .../ContinuousMap/StoneWeierstrass.lean | 65 ++++++++++++++++++- 1 file changed, 64 insertions(+), 1 deletion(-) diff --git a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean index 8097bb16a5cf8..ecfa6bfa4e8f4 100644 --- a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean @@ -421,6 +421,37 @@ theorem polynomialFunctions.starClosure_topologicalClosure {𝕜 : Type*} [RCLik ContinuousMap.starSubalgebra_topologicalClosure_eq_top_of_separatesPoints _ (Subalgebra.separatesPoints_monotone le_sup_left (polynomialFunctions_separatesPoints s)) +/-- An induction principle for `C(s, 𝕜)`. -/ +@[elab_as_elim] +theorem ContinuousMap.induction_on {𝕜 : Type*} [RCLike 𝕜] {s : Set 𝕜} + {p : C(s, 𝕜) → Prop} (const : ∀ r, p (.const s r)) (id : p (.restrict s <| .id 𝕜)) + (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) + (star : ∀ f, p f → p (star f)) + (closure : (∀ f ∈ (polynomialFunctions s).starClosure, p f) → ∀ f, p f) (f : C(s, 𝕜)) : + p f := by + refine closure (fun f hf => ?_) f + rw [polynomialFunctions.starClosure_eq_adjoin_X] at hf + induction hf using StarAlgebra.adjoin_induction with + | mem f hf => + simp only [toContinuousMapOnAlgHom_apply, Set.mem_singleton_iff] at hf + rwa [hf, toContinuousMapOn_X_eq_restrict_id] + | algebraMap r => exact const r + | add _ _ _ _ hf hg => exact add _ _ hf hg + | mul _ _ _ _ hf hg => exact mul _ _ hf hg + | star _ _ hf => exact star _ hf + +open Topology in +@[elab_as_elim] +theorem ContinuousMap.induction_on_of_compact {𝕜 : Type*} [RCLike 𝕜] {s : Set 𝕜} [CompactSpace s] + {p : C(s, 𝕜) → Prop} (const : ∀ r, p (.const s r)) (id : p (.restrict s <| .id 𝕜)) + (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) + (star : ∀ f, p f → p (star f)) (frequently : ∀ f, (∃ᶠ g in 𝓝 f, p g) → p f) (f : C(s, 𝕜)) : + p f := by + refine f.induction_on const id add mul star fun h f ↦ frequently f ?_ + have := polynomialFunctions.starClosure_topologicalClosure s ▸ mem_top (x := f) + rw [← SetLike.mem_coe, topologicalClosure_coe, mem_closure_iff_frequently] at this + exact this.mp <| .of_forall h + /-- Continuous algebra homomorphisms from `C(s, ℝ)` into an `ℝ`-algebra `A` which agree at `X : 𝕜[X]` (interpreted as a continuous map) are, in fact, equal. -/ @[ext (iff := false)] @@ -548,7 +579,8 @@ lemma ker_evalStarAlgHom_eq_closure_adjoin_id (s : Set 𝕜) (h0 : 0 ∈ s) [Com end ContinuousMap -open ContinuousMapZero in +open scoped ContinuousMapZero + /-- If `s : Set 𝕜` with `RCLike 𝕜` is compact and contains `0`, then the non-unital star subalgebra generated by the identity function in `C(s, 𝕜)₀` is dense. This can be seen as a version of the Weierstrass approximation theorem. -/ @@ -566,4 +598,35 @@ lemma ContinuousMapZero.adjoin_id_dense {s : Set 𝕜} [Zero s] (h0 : ((0 : s) : ContinuousMap.evalStarAlgHom_apply, ContinuousMap.coe_coe] rw [show ⟨0, h0'⟩ = (0 : s) by ext; exact h0.symm, _root_.map_zero f] +/-- An induction principle for `C(s, 𝕜)₀`. -/ +@[elab_as_elim] +lemma ContinuousMapZero.induction_on {s : Set 𝕜} [Zero s] (h0 : ((0 : s) : 𝕜) = 0) + {p : C(s, 𝕜)₀ → Prop} (zero : p 0) (id : p (.id h0)) + (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) + (smul : ∀ (r : 𝕜) f, p f → p (r • f)) (star : ∀ f, p f → p (star f)) + (closure : (∀ f ∈ adjoin 𝕜 {(.id h0 : C(s, 𝕜)₀)}, p f) → ∀ f, p f) (f : C(s, 𝕜)₀) : + p f := by + refine closure (fun f hf => ?_) f + induction hf using NonUnitalStarAlgebra.adjoin_induction with + | mem f hf => + simp only [Set.mem_singleton_iff] at hf + rwa [hf] + | zero => exact zero + | add _ _ _ _ hf hg => exact add _ _ hf hg + | mul _ _ _ _ hf hg => exact mul _ _ hf hg + | smul _ _ _ hf => exact smul _ _ hf + | star _ _ hf => exact star _ hf + +open Topology in +@[elab_as_elim] +theorem ContinuousMapZero.induction_on_of_compact {s : Set 𝕜} [Zero s] (h0 : ((0 : s) : 𝕜) = 0) + [CompactSpace s] {p : C(s, 𝕜)₀ → Prop} (zero : p 0) (id : p (.id h0)) + (add : ∀ f g, p f → p g → p (f + g)) (mul : ∀ f g, p f → p g → p (f * g)) + (smul : ∀ (r : 𝕜) f, p f → p (r • f)) (star : ∀ f, p f → p (star f)) + (frequently : ∀ f, (∃ᶠ g in 𝓝 f, p g) → p f) (f : C(s, 𝕜)₀) : + p f := by + refine f.induction_on h0 zero id add mul smul star fun h f ↦ frequently f ?_ + have := (ContinuousMapZero.adjoin_id_dense h0).closure_eq ▸ Set.mem_univ (x := f) + exact mem_closure_iff_frequently.mp this |>.mp <| .of_forall h + end ContinuousMapZero From e7b15d2cf73d42c900766fa0e24e0762f13a8106 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sun, 20 Oct 2024 13:19:35 +0000 Subject: [PATCH 124/131] feat(Analysis/Complex/Positivity): new file (#17862) This adds results of the kind "if `f` is an entire function with all iterated derivatives at a point `c` nonnegative real, then `f z` is nonnegative real for `z = c + r` with `r` nonnegative real". From [EulerProducts](https://github.com/MichaelStollBayreuth/EulerProducts). This will be needed for the proof that $L(\chi, 1) \ne 0$ for nontrivial quadratic Dirichlet characters $\chi$ (which in turn is necessary for Dirichlet's Theorem on primes in arithmetic progressions; see [PNT+](https://github.com/AlexKontorovich/PrimeNumberTheoremAnd)). Co-authored-by: Michael Stoll <99838730+MichaelStollBayreuth@users.noreply.github.com> --- Mathlib.lean | 1 + .../Calculus/IteratedDeriv/Lemmas.lean | 42 ++++++++++ Mathlib/Analysis/Complex/Positivity.lean | 81 +++++++++++++++++++ 3 files changed, 124 insertions(+) create mode 100644 Mathlib/Analysis/Complex/Positivity.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0eb264c536463..de1b28985d7b6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1099,6 +1099,7 @@ import Mathlib.Analysis.Complex.OperatorNorm import Mathlib.Analysis.Complex.PhragmenLindelof import Mathlib.Analysis.Complex.Polynomial.Basic import Mathlib.Analysis.Complex.Polynomial.UnitTrinomial +import Mathlib.Analysis.Complex.Positivity import Mathlib.Analysis.Complex.ReImTopology import Mathlib.Analysis.Complex.RealDeriv import Mathlib.Analysis.Complex.RemovableSingularity diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean index 9b8588e207bce..0a4e4120acb15 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean @@ -15,6 +15,8 @@ This file contains a number of further results on `iteratedDerivWithin` that nee than are available in `Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean`. -/ +section one_dimensional + variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] @@ -123,3 +125,43 @@ lemma iteratedDeriv_comp_neg (n : ℕ) (f : 𝕜 → F) (a : 𝕜) : rw [iteratedDeriv_succ, iteratedDeriv_succ, ih', pow_succ', neg_mul, one_mul, deriv_comp_neg (f := fun x ↦ (-1 : 𝕜) ^ n • iteratedDeriv n f x), deriv_const_smul', neg_smul] + +open Topology in +lemma Filter.EventuallyEq.iteratedDeriv_eq (n : ℕ) {f g : 𝕜 → F} {x : 𝕜} (hfg : f =ᶠ[𝓝 x] g) : + iteratedDeriv n f x = iteratedDeriv n g x := by + simp only [← iteratedDerivWithin_univ, iteratedDerivWithin_eq_iteratedFDerivWithin] + rw [(hfg.filter_mono nhdsWithin_le_nhds).iteratedFDerivWithin_eq hfg.eq_of_nhds n] + +lemma Set.EqOn.iteratedDeriv_of_isOpen (hfg : Set.EqOn f g s) (hs : IsOpen s) (n : ℕ) : + Set.EqOn (iteratedDeriv n f) (iteratedDeriv n g) s := by + refine fun x hx ↦ Filter.EventuallyEq.iteratedDeriv_eq n ?_ + filter_upwards [IsOpen.mem_nhds hs hx] with a ha + exact hfg ha + +end one_dimensional + +/-! +### Invariance of iterated derivatives under translation +-/ + +section shift_invariance + +variable {𝕜 F} [NontriviallyNormedField 𝕜] [NormedAddCommGroup F] [NormedSpace 𝕜 F] + +/-- The iterated derivative commutes with shifting the function by a constant on the left. -/ +lemma iteratedDeriv_comp_const_add (n : ℕ) (f : 𝕜 → F) (s : 𝕜) : + iteratedDeriv n (fun z ↦ f (s + z)) = fun t ↦ iteratedDeriv n f (s + t) := by + induction n with + | zero => simp only [iteratedDeriv_zero] + | succ n IH => + simpa only [iteratedDeriv_succ, IH] using funext <| deriv_comp_const_add _ s + +/-- The iterated derivative commutes with shifting the function by a constant on the right. -/ +lemma iteratedDeriv_comp_add_const (n : ℕ) (f : 𝕜 → F) (s : 𝕜) : + iteratedDeriv n (fun z ↦ f (z + s)) = fun t ↦ iteratedDeriv n f (t + s) := by + induction n with + | zero => simp only [iteratedDeriv_zero] + | succ n IH => + simpa only [iteratedDeriv_succ, IH] using funext <| deriv_comp_add_const _ s + +end shift_invariance diff --git a/Mathlib/Analysis/Complex/Positivity.lean b/Mathlib/Analysis/Complex/Positivity.lean new file mode 100644 index 0000000000000..140e3a84386ff --- /dev/null +++ b/Mathlib/Analysis/Complex/Positivity.lean @@ -0,0 +1,81 @@ +/- +Copyright (c) 2024 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import Mathlib.Analysis.Complex.TaylorSeries + +/-! +# Nonnegativity of values of holomorphic functions + +We show that if `f` is holomorphic on an open disk `B(c,r)` and all iterated derivatives of `f` +at `c` are nonnegative real, then `f z ≥ 0` for all `z ≥ c` in the disk; see +`DifferentiableOn.nonneg_of_iteratedDeriv_nonneg`. We also provide a +variant `Differentiable.nonneg_of_iteratedDeriv_nonneg` for entire functions and versions +showing `f z ≥ f c` when all iterated derivatives except `f` itseld are nonnegative. +-/ + +open Complex + +open scoped ComplexOrder + +namespace DifferentiableOn + +/-- A function that is holomorphic on the open disk around `c` with radius `r` and whose iterated +derivatives at `c` are all nonnegative real has nonnegative real values on `c + [0,r)`. -/ +theorem nonneg_of_iteratedDeriv_nonneg {f : ℂ → ℂ} {c : ℂ} {r : ℝ} + (hf : DifferentiableOn ℂ f (Metric.ball c r)) (h : ∀ n, 0 ≤ iteratedDeriv n f c) ⦃z : ℂ⦄ + (hz₁ : c ≤ z) (hz₂ : z ∈ Metric.ball c r): + 0 ≤ f z := by + have H := taylorSeries_eq_on_ball' hz₂ hf + rw [← sub_nonneg] at hz₁ + have hz' := eq_re_of_ofReal_le hz₁ + rw [hz'] at hz₁ H + refine H ▸ tsum_nonneg fun n ↦ ?_ + rw [← ofReal_natCast, ← ofReal_pow, ← ofReal_inv, eq_re_of_ofReal_le (h n), ← ofReal_mul, + ← ofReal_mul] + norm_cast at hz₁ ⊢ + have := zero_re ▸ (Complex.le_def.mp (h n)).1 + positivity + +end DifferentiableOn + +namespace Differentiable + +/-- An entire function whose iterated derivatives at `c` are all nonnegative real has nonnegative +real values on `c + ℝ≥0`. -/ +theorem nonneg_of_iteratedDeriv_nonneg {f : ℂ → ℂ} (hf : Differentiable ℂ f) {c : ℂ} + (h : ∀ n, 0 ≤ iteratedDeriv n f c) ⦃z : ℂ⦄ (hz : c ≤ z) : + 0 ≤ f z := by + refine hf.differentiableOn.nonneg_of_iteratedDeriv_nonneg (r := (z - c).re + 1) h hz ?_ + rw [← sub_nonneg] at hz + rw [Metric.mem_ball, dist_eq, eq_re_of_ofReal_le hz] + simpa only [Complex.abs_of_nonneg (nonneg_iff.mp hz).1] using lt_add_one _ + +/-- An entire function whose iterated derivatives at `c` are all nonnegative real (except +possibly the value itself) has values of the form `f c + nonneg. real` on the set `c + ℝ≥0`. -/ +theorem apply_le_of_iteratedDeriv_nonneg {f : ℂ → ℂ} {c : ℂ} (hf : Differentiable ℂ f) + (h : ∀ n ≠ 0, 0 ≤ iteratedDeriv n f c) ⦃z : ℂ⦄ (hz : c ≤ z) : + f c ≤ f z := by + have h' (n : ℕ) : 0 ≤ iteratedDeriv n (f · - f c) c := by + cases n with + | zero => simp only [iteratedDeriv_zero, sub_self, le_refl] + | succ n => + specialize h (n + 1) n.succ_ne_zero + rw [iteratedDeriv_succ'] at h ⊢ + rwa [funext fun x ↦ deriv_sub_const (f := f) (x := x) (f c)] + exact sub_nonneg.mp <| nonneg_of_iteratedDeriv_nonneg (hf.sub_const _) h' hz + +/-- An entire function whose iterated derivatives at `c` are all real with alternating signs +(except possibly the value itself) has values of the form `f c + nonneg. real` along the +set `c - ℝ≥0`. -/ +theorem apply_le_of_iteratedDeriv_alternating {f : ℂ → ℂ} {c : ℂ} (hf : Differentiable ℂ f) + (h : ∀ n ≠ 0, 0 ≤ (-1) ^ n * iteratedDeriv n f c) ⦃z : ℂ⦄ (hz : z ≤ c) : + f c ≤ f z := by + convert apply_le_of_iteratedDeriv_nonneg (f := fun z ↦ f (-z)) + (hf.comp <| differentiable_neg) (fun n hn ↦ ?_) (neg_le_neg_iff.mpr hz) using 1 + · simp only [neg_neg] + · simp only [neg_neg] + · simpa only [iteratedDeriv_comp_neg, neg_neg, smul_eq_mul] using h n hn + +end Differentiable From 0dedd92ed7e30cdcb2c8ff536bed18ea97b79a4f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 20 Oct 2024 13:29:54 +0000 Subject: [PATCH 125/131] chore: Rename `ClosedEmbedding` to `IsClosedEmbedding` (#17937) `Function.Embedding` is a type while `Embedding` is a proposition, and there are many other kinds of embeddings than topological embeddings. Hence this PR is a step towards 1. renaming `Embedding` to `IsEmbedding` and similarly for neighborhing declarations (which `ClosedEmbedding` is) 2. namespacing it inside `Topology` [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/rename.20.60Inducing.60.20and.20.60Embedding.60.3F). See #15993 for context. --- .../Morphisms/ClosedImmersion.lean | 42 ++++----- .../Morphisms/UnderlyingMap.lean | 14 +-- .../PrimeSpectrum/Basic.lean | 12 ++- Mathlib/AlgebraicGeometry/ResidueField.lean | 2 +- .../Instances.lean | 27 +++--- .../NonUnital.lean | 36 ++++---- .../ContinuousFunctionalCalculus/Order.lean | 4 +- .../Restrict.lean | 46 +++++----- .../ContinuousFunctionalCalculus/Unique.lean | 4 +- .../ContinuousFunctionalCalculus/Unital.lean | 19 ++-- .../Analysis/CStarAlgebra/GelfandDuality.lean | 2 +- .../LineDeriv/IntegrationByParts.lean | 2 +- Mathlib/Analysis/Convex/Intrinsic.lean | 4 +- Mathlib/Analysis/Convolution.lean | 6 +- .../FunctionalSpaces/SobolevInequality.lean | 2 +- Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 4 +- Mathlib/Analysis/Normed/Module/Basic.lean | 4 +- .../Normed/Module/FiniteDimension.lean | 4 +- Mathlib/Analysis/Normed/Operator/Compact.lean | 2 +- .../ContinuousFunctionalCalculus/ExpLog.lean | 2 +- .../SpecialFunctions/Log/ENNRealLogExp.lean | 2 +- Mathlib/CategoryTheory/Galois/Topology.lean | 13 +-- .../Manifold/SmoothManifoldWithCorners.lean | 15 ++-- .../Geometry/Manifold/WhitneyEmbedding.lean | 4 +- .../Matrix/HermitianFunctionalCalculus.lean | 11 ++- .../Constructions/BorelSpace/Basic.lean | 10 ++- .../BorelSpace/ContinuousLinearMap.lean | 4 +- .../Constructions/Polish/EmbeddingReal.lean | 2 +- .../Function/StronglyMeasurable/Basic.lean | 4 +- .../Function/StronglyMeasurable/Lemmas.lean | 2 +- Mathlib/MeasureTheory/Group/Measure.lean | 2 +- Mathlib/MeasureTheory/Integral/Bochner.lean | 7 +- .../Integral/IntervalIntegral.lean | 8 +- .../MeasureTheory/Integral/SetIntegral.lean | 12 ++- .../Measure/Lebesgue/Integral.lean | 2 +- Mathlib/NumberTheory/Modular.lean | 10 +-- .../Algebra/Constructions/DomMulAct.lean | 22 +++-- .../Topology/Algebra/ContinuousMonoidHom.lean | 7 +- .../Algebra/Module/Alternating/Topology.lean | 7 +- .../Algebra/Module/FiniteDimension.lean | 18 ++-- .../Topology/Algebra/ProperAction/Basic.lean | 16 ++-- Mathlib/Topology/Algebra/StarSubalgebra.lean | 18 ++-- .../Category/LightProfinite/Sequence.lean | 9 +- .../Topology/Category/Profinite/Nobeling.lean | 6 +- Mathlib/Topology/Category/Stonean/Limits.lean | 2 +- Mathlib/Topology/Clopen.lean | 2 +- Mathlib/Topology/Compactness/Compact.lean | 28 ++++-- Mathlib/Topology/Compactness/Lindelof.lean | 27 ++++-- .../Topology/Compactness/LocallyCompact.lean | 16 ++-- Mathlib/Topology/Compactness/Paracompact.lean | 9 +- .../Topology/Compactness/SigmaCompact.lean | 9 +- Mathlib/Topology/Constructions.lean | 45 +++++++--- .../ContinuousMap/ContinuousMapZero.lean | 9 +- .../ContinuousMap/StoneWeierstrass.lean | 4 +- Mathlib/Topology/Defs/Induced.lean | 7 +- Mathlib/Topology/DiscreteSubset.lean | 2 +- Mathlib/Topology/FiberBundle/Basic.lean | 7 +- Mathlib/Topology/Homeomorph.lean | 18 ++-- Mathlib/Topology/Instances/CantorSet.lean | 4 +- Mathlib/Topology/Instances/EReal.lean | 5 +- Mathlib/Topology/Instances/Int.lean | 7 +- Mathlib/Topology/Instances/Irrational.lean | 2 +- Mathlib/Topology/Instances/NNReal.lean | 2 +- Mathlib/Topology/Instances/Nat.lean | 7 +- Mathlib/Topology/Instances/Rat.lean | 16 ++-- Mathlib/Topology/Instances/Real.lean | 2 +- Mathlib/Topology/KrullDimension.lean | 13 +-- Mathlib/Topology/LocalAtTarget.lean | 24 ++++-- Mathlib/Topology/Maps/Basic.lean | 54 +++++++----- Mathlib/Topology/Maps/Proper/Basic.lean | 20 +++-- .../Topology/MetricSpace/Antilipschitz.lean | 7 +- Mathlib/Topology/MetricSpace/Basic.lean | 9 +- Mathlib/Topology/MetricSpace/Dilation.lean | 8 +- Mathlib/Topology/MetricSpace/Isometry.lean | 9 +- Mathlib/Topology/MetricSpace/Polish.lean | 19 ++-- Mathlib/Topology/SeparatedMap.lean | 11 ++- Mathlib/Topology/Separation.lean | 53 +++++++----- Mathlib/Topology/Sober.lean | 5 +- Mathlib/Topology/Support.lean | 7 +- Mathlib/Topology/TietzeExtension.lean | 86 ++++++++++++------- Mathlib/Topology/UniformSpace/Ascoli.lean | 21 +++-- .../UniformSpace/CompleteSeparated.lean | 12 ++- .../UniformSpace/UniformEmbedding.lean | 7 +- 83 files changed, 655 insertions(+), 389 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index 0e01ff0984171..741e9d4b1146b 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -40,16 +40,19 @@ namespace AlgebraicGeometry topological map is a closed embedding and the induced stalk maps are surjective. -/ @[mk_iff] class IsClosedImmersion {X Y : Scheme} (f : X ⟶ Y) : Prop where - base_closed : ClosedEmbedding f.base + base_closed : IsClosedEmbedding f.base surj_on_stalks : ∀ x, Function.Surjective (f.stalkMap x) namespace IsClosedImmersion -lemma closedEmbedding {X Y : Scheme} (f : X ⟶ Y) - [IsClosedImmersion f] : ClosedEmbedding f.base := +lemma isClosedEmbedding {X Y : Scheme} (f : X ⟶ Y) + [IsClosedImmersion f] : IsClosedEmbedding f.base := IsClosedImmersion.base_closed -lemma eq_inf : @IsClosedImmersion = (topologically ClosedEmbedding) ⊓ +@[deprecated (since := "2024-10-20")] +alias closedEmbedding := isClosedEmbedding + +lemma eq_inf : @IsClosedImmersion = (topologically IsClosedEmbedding) ⊓ stalkwise (fun f ↦ Function.Surjective f) := by ext X Y f rw [isClosedImmersion_iff] @@ -57,7 +60,7 @@ lemma eq_inf : @IsClosedImmersion = (topologically ClosedEmbedding) ⊓ lemma iff_isPreimmersion {X Y : Scheme} {f : X ⟶ Y} : IsClosedImmersion f ↔ IsPreimmersion f ∧ IsClosed (Set.range f.base) := by - rw [and_comm, isClosedImmersion_iff, isPreimmersion_iff, ← and_assoc, closedEmbedding_iff, + rw [and_comm, isClosedImmersion_iff, isPreimmersion_iff, ← and_assoc, isClosedEmbedding_iff, @and_comm (Embedding _)] lemma of_isPreimmersion {X Y : Scheme} (f : X ⟶ Y) [IsPreimmersion f] @@ -69,7 +72,7 @@ instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : /-- Isomorphisms are closed immersions. -/ instance {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : IsClosedImmersion f where - base_closed := Homeomorph.closedEmbedding <| TopCat.homeoOfIso (asIso f.base) + base_closed := Homeomorph.isClosedEmbedding <| TopCat.homeoOfIso (asIso f.base) surj_on_stalks := fun _ ↦ (ConcreteCategory.bijective_of_isIso _).2 instance : MorphismProperty.IsMultiplicative @IsClosedImmersion where @@ -93,7 +96,7 @@ instance respectsIso : MorphismProperty.RespectsIso @IsClosedImmersion := by closed immersion. -/ theorem spec_of_surjective {R S : CommRingCat} (f : R ⟶ S) (h : Function.Surjective f) : IsClosedImmersion (Spec.map f) where - base_closed := PrimeSpectrum.closedEmbedding_comap_of_surjective _ _ h + base_closed := PrimeSpectrum.isClosedEmbedding_comap_of_surjective _ _ h surj_on_stalks x := by haveI : (RingHom.toMorphismProperty (fun f ↦ Function.Surjective f)).RespectsIso := by rw [← RingHom.toMorphismProperty_respectsIso_iff] @@ -121,14 +124,13 @@ lemma of_surjective_of_isAffine {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X theorem of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsClosedImmersion g] [IsClosedImmersion (f ≫ g)] : IsClosedImmersion f where base_closed := by - have h := closedEmbedding (f ≫ g) - rw [Scheme.comp_base] at h - apply closedEmbedding_of_continuous_injective_closed (Scheme.Hom.continuous f) - · exact Function.Injective.of_comp h.inj - · intro Z hZ - rw [ClosedEmbedding.closed_iff_image_closed (closedEmbedding g), - ← Set.image_comp] - exact ClosedEmbedding.isClosedMap h _ hZ + have h := isClosedEmbedding (f ≫ g) + simp only [Scheme.comp_coeBase, TopCat.coe_comp] at h + refine .of_continuous_injective_isClosedMap (Scheme.Hom.continuous f) h.inj.of_comp ?_ + intro Z hZ + rw [IsClosedEmbedding.closed_iff_image_closed (isClosedEmbedding g), + ← Set.image_comp] + exact h.isClosedMap _ hZ surj_on_stalks x := by have h := (f ≫ g).stalkMap_surjective x simp_rw [Scheme.stalkMap_comp] at h @@ -222,14 +224,14 @@ namespace IsClosedImmersion sections is injective, `f` is an isomorphism. -/ theorem isIso_of_injective_of_isAffine [IsClosedImmersion f] (hf : Function.Injective (f.app ⊤)) : IsIso f := (isIso_iff_stalk_iso f).mpr <| - have : CompactSpace X := (closedEmbedding f).compactSpace + have : CompactSpace X := (isClosedEmbedding f).compactSpace have hiso : IsIso f.base := TopCat.isIso_of_bijective_of_isClosedMap _ - ⟨(closedEmbedding f).inj, - surjective_of_isClosed_range_of_injective ((closedEmbedding f).isClosed_range) hf⟩ - ((closedEmbedding f).isClosedMap) + ⟨(isClosedEmbedding f).inj, + surjective_of_isClosed_range_of_injective ((isClosedEmbedding f).isClosed_range) hf⟩ + ((isClosedEmbedding f).isClosedMap) ⟨hiso, fun x ↦ (ConcreteCategory.isIso_iff_bijective _).mpr ⟨stalkMap_injective_of_isOpenMap_of_injective ((TopCat.homeoOfIso (asIso f.base)).isOpenMap) - (closedEmbedding f).inj hf _, f.stalkMap_surjective x⟩⟩ + (isClosedEmbedding f).inj hf _, f.stalkMap_surjective x⟩⟩ variable (f) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean index 00d8b44dc43f0..ed2f3268b2efc 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean @@ -19,7 +19,7 @@ of the underlying map of topological spaces, including - `IsClosedMap` - `Embedding` - `IsOpenEmbedding` -- `ClosedEmbedding` +- `IsClosedEmbedding` -/ @@ -131,16 +131,16 @@ instance isOpenEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically IsOpen end IsOpenEmbedding -section ClosedEmbedding +section IsClosedEmbedding -instance : (topologically ClosedEmbedding).RespectsIso := - topologically_respectsIso _ (fun e ↦ e.closedEmbedding) (fun _ _ hf hg ↦ hg.comp hf) +instance : (topologically IsClosedEmbedding).RespectsIso := + topologically_respectsIso _ (fun e ↦ e.isClosedEmbedding) (fun _ _ hf hg ↦ hg.comp hf) -instance closedEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically ClosedEmbedding) := +instance isClosedEmbedding_isLocalAtTarget : IsLocalAtTarget (topologically IsClosedEmbedding) := topologically_isLocalAtTarget _ (fun _ s hf ↦ hf.restrictPreimage s) - (fun _ _ _ hU hfcont hf ↦ (closedEmbedding_iff_closedEmbedding_of_iSup_eq_top hU hfcont).mpr hf) + (fun _ _ _ hU hfcont ↦ (isClosedEmbedding_iff_isClosedEmbedding_of_iSup_eq_top hU hfcont).mpr) -end ClosedEmbedding +end IsClosedEmbedding end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index f918a39487b28..2ad92e90d3efa 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -326,11 +326,14 @@ theorem isClosed_range_comap_of_surjective (hf : Surjective f) : rw [range_comap_of_surjective _ f hf] exact isClosed_zeroLocus _ -theorem closedEmbedding_comap_of_surjective (hf : Surjective f) : ClosedEmbedding (comap f) := +theorem isClosedEmbedding_comap_of_surjective (hf : Surjective f) : IsClosedEmbedding (comap f) := { induced := (comap_inducing_of_surjective S f hf).induced inj := comap_injective_of_surjective f hf isClosed_range := isClosed_range_comap_of_surjective S f hf } +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_comap_of_surjective := isClosedEmbedding_comap_of_surjective + end SpecOfSurjective section SpecProd @@ -351,17 +354,18 @@ noncomputable def primeSpectrumProdHomeo : PrimeSpectrum (R × S) ≃ₜ PrimeSpectrum R ⊕ PrimeSpectrum S := by refine ((primeSpectrumProd R S).symm.toHomeomorphOfInducing ?_).symm - refine (closedEmbedding_of_continuous_injective_closed ?_ (Equiv.injective _) ?_).toInducing + refine (IsClosedEmbedding.of_continuous_injective_isClosedMap ?_ + (Equiv.injective _) ?_).toInducing · rw [continuous_sum_dom] simp only [Function.comp_def, primeSpectrumProd_symm_inl, primeSpectrumProd_symm_inr] exact ⟨(comap _).2, (comap _).2⟩ · rw [isClosedMap_sum] constructor · simp_rw [primeSpectrumProd_symm_inl] - refine (closedEmbedding_comap_of_surjective _ _ ?_).isClosedMap + refine (isClosedEmbedding_comap_of_surjective _ _ ?_).isClosedMap exact Prod.fst_surjective · simp_rw [primeSpectrumProd_symm_inr] - refine (closedEmbedding_comap_of_surjective _ _ ?_).isClosedMap + refine (isClosedEmbedding_comap_of_surjective _ _ ?_).isClosedMap exact Prod.snd_surjective end SpecProd diff --git a/Mathlib/AlgebraicGeometry/ResidueField.lean b/Mathlib/AlgebraicGeometry/ResidueField.lean index c7fcb54274fd0..fd0360bc3bc0c 100644 --- a/Mathlib/AlgebraicGeometry/ResidueField.lean +++ b/Mathlib/AlgebraicGeometry/ResidueField.lean @@ -52,7 +52,7 @@ def residue (X : Scheme.{u}) (x) : X.presheaf.stalk x ⟶ X.residueField x := `Spec.map (X.residue x)` is a closed immersion. -/ instance {X : Scheme.{u}} (x) : IsPreimmersion (Spec.map (X.residue x)) := IsPreimmersion.mk_Spec_map - (PrimeSpectrum.closedEmbedding_comap_of_surjective _ _ (Ideal.Quotient.mk_surjective)).embedding + (PrimeSpectrum.isClosedEmbedding_comap_of_surjective _ _ Ideal.Quotient.mk_surjective).embedding (RingHom.surjectiveOnStalks_of_surjective (Ideal.Quotient.mk_surjective)) @[simp] diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index ea67a1f84c970..5d998afdff365 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -78,16 +78,19 @@ noncomputable def cfcₙAux : C(σₙ 𝕜 a, 𝕜)₀ →⋆ₙₐ[𝕜] A⁺¹ lemma cfcₙAux_id : cfcₙAux hp₁ a ha (ContinuousMapZero.id rfl) = a := cfcHom_id (hp₁.mpr ha) open Unitization in -lemma closedEmbedding_cfcₙAux : ClosedEmbedding (cfcₙAux hp₁ a ha) := by +lemma isClosedEmbedding_cfcₙAux : IsClosedEmbedding (cfcₙAux hp₁ a ha) := by simp only [cfcₙAux, NonUnitalStarAlgHom.coe_comp] - refine ((cfcHom_closedEmbedding (hp₁.mpr ha)).comp ?_).comp - ContinuousMapZero.closedEmbedding_toContinuousMap + refine ((cfcHom_isClosedEmbedding (hp₁.mpr ha)).comp ?_).comp + ContinuousMapZero.isClosedEmbedding_toContinuousMap let e : C(σₙ 𝕜 a, 𝕜) ≃ₜ C(σ 𝕜 (a : A⁺¹), 𝕜) := { (Homeomorph.compStarAlgEquiv' 𝕜 𝕜 <| .setCongr <| (quasispectrum_eq_spectrum_inr' 𝕜 𝕜 a).symm) with continuous_toFun := ContinuousMap.continuous_comp_left _ continuous_invFun := ContinuousMap.continuous_comp_left _ } - exact e.closedEmbedding + exact e.isClosedEmbedding + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_cfcₙAux := isClosedEmbedding_cfcₙAux lemma spec_cfcₙAux (f : C(σₙ 𝕜 a, 𝕜)₀) : σ 𝕜 (cfcₙAux hp₁ a ha f) = Set.range f := by rw [cfcₙAux, NonUnitalStarAlgHom.comp_assoc, NonUnitalStarAlgHom.comp_apply] @@ -103,7 +106,7 @@ variable [CompleteSpace A] lemma cfcₙAux_mem_range_inr (f : C(σₙ 𝕜 a, 𝕜)₀) : cfcₙAux hp₁ a ha f ∈ NonUnitalStarAlgHom.range (Unitization.inrNonUnitalStarAlgHom 𝕜 A) := by - have h₁ := (closedEmbedding_cfcₙAux hp₁ a ha).continuous.range_subset_closure_image_dense + have h₁ := (isClosedEmbedding_cfcₙAux hp₁ a ha).continuous.range_subset_closure_image_dense (ContinuousMapZero.adjoin_id_dense (s := σₙ 𝕜 a) rfl) ⟨f, rfl⟩ rw [← SetLike.mem_coe] refine closure_minimal ?_ ?_ h₁ @@ -133,11 +136,11 @@ theorem RCLike.nonUnitalContinuousFunctionalCalculus : have coe_ψ (f : C(σₙ 𝕜 a, 𝕜)₀) : ψ f = cfcₙAux hp₁ a ha f := congr_arg Subtype.val <| (inrRangeEquiv 𝕜 A).apply_symm_apply ⟨cfcₙAux hp₁ a ha f, cfcₙAux_mem_range_inr hp₁ a ha f⟩ - refine ⟨ψ, ?closedEmbedding, ?map_id, fun f ↦ ?map_spec, fun f ↦ ?isStarNormal⟩ - case closedEmbedding => - apply isometry_inr (𝕜 := 𝕜) (A := A) |>.closedEmbedding |>.of_comp_iff.mp + refine ⟨ψ, ?isClosedEmbedding, ?map_id, fun f ↦ ?map_spec, fun f ↦ ?isStarNormal⟩ + case isClosedEmbedding => + apply isometry_inr (𝕜 := 𝕜) (A := A) |>.isClosedEmbedding |>.of_comp_iff.mp have : inr ∘ ψ = cfcₙAux hp₁ a ha := by ext1; rw [Function.comp_apply, coe_ψ] - exact this ▸ closedEmbedding_cfcₙAux hp₁ a ha + exact this ▸ isClosedEmbedding_cfcₙAux hp₁ a ha case map_id => exact inr_injective (R := 𝕜) <| coe_ψ _ ▸ cfcₙAux_id hp₁ a ha case map_spec => exact quasispectrum_eq_spectrum_inr' 𝕜 𝕜 (ψ f) ▸ coe_ψ _ ▸ spec_cfcₙAux hp₁ a ha f @@ -158,13 +161,13 @@ instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [NormedRing A spectrum_nonempty a _ := spectrum.nonempty a exists_cfc_of_predicate a ha := by refine ⟨(elementalStarAlgebra ℂ a).subtype.comp <| continuousFunctionalCalculus a, - ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ - case hom_closedEmbedding => + ?hom_isClosedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ + case hom_isClosedEmbedding => -- note: Lean should find these for `StarAlgEquiv.isometry`, but it doesn't and so we -- provide them manually. have : SMulCommClass ℂ C(σ ℂ a, ℂ) C(σ ℂ a, ℂ) := Algebra.to_smulCommClass (A := C(σ ℂ a, ℂ)) have : IsScalarTower ℂ C(σ ℂ a, ℂ) C(σ ℂ a, ℂ) := IsScalarTower.right (A := C(σ ℂ a, ℂ)) - exact Isometry.closedEmbedding <| + exact Isometry.isClosedEmbedding <| isometry_subtype_coe.comp <| StarAlgEquiv.isometry (continuousFunctionalCalculus a) case hom_id => exact congr_arg Subtype.val <| continuousFunctionalCalculus_map_id a case hom_map_spectrum => diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index d4a2d892c661e..e837b34df4812 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -75,7 +75,7 @@ class NonUnitalContinuousFunctionalCalculus (R : Type*) {A : Type*} (p : outPara predicate_zero : p 0 [compactSpace_quasispectrum : ∀ a : A, CompactSpace (σₙ R a)] exists_cfc_of_predicate : ∀ a, p a → ∃ φ : C(σₙ R a, R)₀ →⋆ₙₐ[R] A, - ClosedEmbedding φ ∧ φ ⟨(ContinuousMap.id R).restrict <| σₙ R a, rfl⟩ = a ∧ + IsClosedEmbedding φ ∧ φ ⟨(ContinuousMap.id R).restrict <| σₙ R a, rfl⟩ = a ∧ (∀ f, σₙ R (φ f) = Set.range f) ∧ ∀ f, p (φ f) -- this instance should not be activated everywhere but it is useful when developing generic API @@ -128,13 +128,16 @@ the user should instead prefer `cfcₙ` over `cfcₙHom`. noncomputable def cfcₙHom : C(σₙ R a, R)₀ →⋆ₙₐ[R] A := (NonUnitalContinuousFunctionalCalculus.exists_cfc_of_predicate a ha).choose -lemma cfcₙHom_closedEmbedding : - ClosedEmbedding <| (cfcₙHom ha : C(σₙ R a, R)₀ →⋆ₙₐ[R] A) := +lemma cfcₙHom_isClosedEmbedding : + IsClosedEmbedding <| (cfcₙHom ha : C(σₙ R a, R)₀ →⋆ₙₐ[R] A) := (NonUnitalContinuousFunctionalCalculus.exists_cfc_of_predicate a ha).choose_spec.1 +@[deprecated (since := "2024-10-20")] +alias cfcₙHom_closedEmbedding := cfcₙHom_isClosedEmbedding + @[fun_prop] lemma cfcₙHom_continuous : Continuous (cfcₙHom ha : C(σₙ R a, R)₀ →⋆ₙₐ[R] A) := - cfcₙHom_closedEmbedding ha |>.continuous + cfcₙHom_isClosedEmbedding ha |>.continuous lemma cfcₙHom_id : cfcₙHom ha ⟨(ContinuousMap.id R).restrict <| σₙ R a, rfl⟩ = a := @@ -152,7 +155,7 @@ lemma cfcₙHom_predicate (f : C(σₙ R a, R)₀) : lemma cfcₙHom_eq_of_continuous_of_map_id [UniqueNonUnitalContinuousFunctionalCalculus R A] (φ : C(σₙ R a, R)₀ →⋆ₙₐ[R] A) (hφ₁ : Continuous φ) (hφ₂ : φ ⟨.restrict (σₙ R a) <| .id R, rfl⟩ = a) : cfcₙHom ha = φ := - (cfcₙHom ha).ext_continuousMap a φ (cfcₙHom_closedEmbedding ha).continuous hφ₁ <| by + (cfcₙHom ha).ext_continuousMap a φ (cfcₙHom_isClosedEmbedding ha).continuous hφ₁ <| by rw [cfcₙHom_id ha, hφ₂] theorem cfcₙHom_comp [UniqueNonUnitalContinuousFunctionalCalculus R A] (f : C(σₙ R a, R)₀) @@ -169,7 +172,7 @@ theorem cfcₙHom_comp [UniqueNonUnitalContinuousFunctionalCalculus R A] (f : C( let φ : C(σₙ R (cfcₙHom ha f), R)₀ →⋆ₙₐ[R] A := (cfcₙHom ha).comp ψ suffices cfcₙHom (cfcₙHom_predicate ha f) = φ from DFunLike.congr_fun this.symm g refine cfcₙHom_eq_of_continuous_of_map_id (cfcₙHom_predicate ha f) φ ?_ ?_ - · refine (cfcₙHom_closedEmbedding ha).continuous.comp <| continuous_induced_rng.mpr ?_ + · refine (cfcₙHom_isClosedEmbedding ha).continuous.comp <| continuous_induced_rng.mpr ?_ exact f'.toContinuousMap.continuous_comp_left.comp continuous_induced_dom · simp only [φ, ψ, NonUnitalStarAlgHom.comp_apply, NonUnitalStarAlgHom.coe_mk', NonUnitalAlgHom.coe_mk] @@ -187,7 +190,7 @@ noncomputable def cfcₙL {a : A} (ha : p a) : C(σₙ R a, R)₀ →L[R] A := { cfcₙHom ha with toFun := cfcₙHom ha map_smul' := map_smul _ - cont := (cfcₙHom_closedEmbedding ha).continuous } + cont := (cfcₙHom_isClosedEmbedding ha).continuous } end cfcₙL @@ -306,7 +309,7 @@ lemma eqOn_of_cfcₙ_eq_cfcₙ {f g : R → R} {a : A} (h : cfcₙ f a = cfcₙ (hg : ContinuousOn g (σₙ R a) := by cfc_cont_tac) (hg0 : g 0 = 0 := by cfc_zero_tac) : (σₙ R a).EqOn f g := by rw [cfcₙ_apply f a, cfcₙ_apply g a] at h - have := (cfcₙHom_closedEmbedding (show p a from ha) (R := R)).inj h + have := (cfcₙHom_isClosedEmbedding (show p a from ha) (R := R)).inj h intro x hx congrm($(this) ⟨x, hx⟩) @@ -695,12 +698,12 @@ variable [CompleteSpace R] -- gives access to the `ContinuousFunctionalCalculus.compactSpace_spectrum` instance open scoped ContinuousFunctionalCalculus -lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : - ClosedEmbedding (cfcₙHom_of_cfcHom R ha) := by +lemma isClosedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : + IsClosedEmbedding (cfcₙHom_of_cfcHom R ha) := by let f : C(spectrum R a, σₙ R a) := ⟨_, continuous_inclusion <| spectrum_subset_quasispectrum R a⟩ - refine (cfcHom_closedEmbedding ha).comp <| - (IsUniformInducing.isUniformEmbedding ⟨?_⟩).toClosedEmbedding + refine (cfcHom_isClosedEmbedding ha).comp <| + (IsUniformInducing.isUniformEmbedding ⟨?_⟩).toIsClosedEmbedding 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, @@ -715,6 +718,9 @@ lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : convert Filter.comap_const_of_mem this with ⟨u, v⟩ <;> ext ⟨x, rfl⟩ <;> [exact map_zero u; exact map_zero v] +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_cfcₙHom_of_cfcHom := isClosedEmbedding_cfcₙHom_of_cfcHom + instance ContinuousFunctionalCalculus.toNonUnital : NonUnitalContinuousFunctionalCalculus R p where predicate_zero := cfc_predicate_zero R compactSpace_quasispectrum a := by @@ -723,7 +729,7 @@ instance ContinuousFunctionalCalculus.toNonUnital : NonUnitalContinuousFunctiona exact h_cpct |>.union isCompact_singleton exists_cfc_of_predicate _ ha := ⟨cfcₙHom_of_cfcHom R ha, - closedEmbedding_cfcₙHom_of_cfcHom ha, + isClosedEmbedding_cfcₙHom_of_cfcHom ha, cfcHom_id ha, cfcₙHom_of_cfcHom_map_quasispectrum ha, fun _ ↦ cfcHom_predicate ha _⟩ @@ -734,8 +740,8 @@ lemma cfcₙHom_eq_cfcₙHom_of_cfcHom [UniqueNonUnitalContinuousFunctionalCalcu refine UniqueNonUnitalContinuousFunctionalCalculus.eq_of_continuous_of_map_id (σₙ R a) ?_ _ _ ?_ ?_ ?_ · simp - · exact (cfcₙHom_closedEmbedding (R := R) ha).continuous - · exact (closedEmbedding_cfcₙHom_of_cfcHom ha).continuous + · exact (cfcₙHom_isClosedEmbedding (R := R) ha).continuous + · exact (isClosedEmbedding_cfcₙHom_of_cfcHom ha).continuous · simpa only [cfcₙHom_id (R := R) ha] using (cfcHom_id ha).symm /-- When `cfc` is applied to a function that maps zero to zero, it is equivalent to using diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index 10e0978f35a5c..08951f013e6d4 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -370,8 +370,8 @@ lemma conjugate_le_norm_smul' {a b : A} (hb : IsSelfAdjoint b := by cfc_tac) : /-- The set of nonnegative elements in a C⋆-algebra is closed. -/ lemma isClosed_nonneg : IsClosed {a : A | 0 ≤ a} := by suffices IsClosed {a : Unitization ℂ A | 0 ≤ a} by - rw [Unitization.isometry_inr (𝕜 := ℂ) |>.closedEmbedding.closed_iff_image_closed] - convert this.inter <| (Unitization.isometry_inr (𝕜 := ℂ)).closedEmbedding.isClosed_range + rw [Unitization.isometry_inr (𝕜 := ℂ) |>.isClosedEmbedding.closed_iff_image_closed] + convert this.inter <| (Unitization.isometry_inr (𝕜 := ℂ)).isClosedEmbedding.isClosed_range ext a simp only [Set.mem_image, Set.mem_setOf_eq, Set.mem_inter_iff, Set.mem_range, ← exists_and_left] congr! 2 with x diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean index bb43228df3fdf..ac2ac678b6bee 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean @@ -82,14 +82,17 @@ lemma starAlgHom_id {a : A} {φ : C(spectrum S a, S) →⋆ₐ[S] A} {f : C(S, R variable [TopologicalSpace A] [ContinuousFunctionalCalculus S q] variable [CompleteSpace R] -lemma closedEmbedding_starAlgHom {a : A} {φ : C(spectrum S a, S) →⋆ₐ[S] A} - (hφ : ClosedEmbedding φ) {f : C(S, R)} (h : SpectrumRestricts a f) +lemma isClosedEmbedding_starAlgHom {a : A} {φ : C(spectrum S a, S) →⋆ₐ[S] A} + (hφ : IsClosedEmbedding φ) {f : C(S, R)} (h : SpectrumRestricts a f) (halg : IsUniformEmbedding (algebraMap R S)) : - ClosedEmbedding (h.starAlgHom φ) := - hφ.comp <| IsUniformEmbedding.toClosedEmbedding <| .comp + IsClosedEmbedding (h.starAlgHom φ) := + hφ.comp <| IsUniformEmbedding.toIsClosedEmbedding <| .comp (ContinuousMap.isUniformEmbedding_comp _ halg) (UniformEquiv.arrowCongr h.homeomorph.symm (.refl _) |>.isUniformEmbedding) +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_starAlgHom := isClosedEmbedding_starAlgHom + /-- Given a `ContinuousFunctionalCalculus S q`. If we form the predicate `p` for `a : A` characterized by: `q a` and the spectrum of `a` restricts to the scalar subring `R` via `f : C(S, R)`, then we can get a restricted functional calculus @@ -103,13 +106,13 @@ protected theorem cfc (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) compactSpace_spectrum a := by have := ContinuousFunctionalCalculus.compactSpace_spectrum (R := S) a rw [← isCompact_iff_compactSpace] at this ⊢ - simpa using halg.toClosedEmbedding.isCompact_preimage this + simpa using halg.toIsClosedEmbedding.isCompact_preimage this exists_cfc_of_predicate a ha := by refine ⟨((h a).mp ha).2.starAlgHom (cfcHom ((h a).mp ha).1 (R := S)), - ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ - case hom_closedEmbedding => - exact ((h a).mp ha).2.closedEmbedding_starAlgHom - (cfcHom_closedEmbedding ((h a).mp ha).1) halg + ?hom_isClosedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ + case hom_isClosedEmbedding => + exact ((h a).mp ha).2.isClosedEmbedding_starAlgHom + (cfcHom_isClosedEmbedding ((h a).mp ha).1) halg case hom_id => exact ((h a).mp ha).2.starAlgHom_id <| cfcHom_id ((h a).mp ha).1 case hom_map_spectrum => intro g @@ -143,7 +146,7 @@ lemma cfcHom_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R {a : A} (hpa : p a) (hqa : q a) (h : SpectrumRestricts a f) : cfcHom hpa = h.starAlgHom (cfcHom hqa) := by apply cfcHom_eq_of_continuous_of_map_id - · exact h.closedEmbedding_starAlgHom (cfcHom_closedEmbedding hqa) halg |>.continuous + · exact h.isClosedEmbedding_starAlgHom (cfcHom_isClosedEmbedding hqa) halg |>.continuous · exact h.starAlgHom_id (cfcHom_id hqa) lemma cfc_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) {a : A} (hpa : p a) @@ -222,15 +225,18 @@ lemma nonUnitalStarAlgHom_id {a : A} {φ : C(σₙ S a, S)₀ →⋆ₙₐ[S] A} variable [TopologicalSpace A] [NonUnitalContinuousFunctionalCalculus S q] variable [CompleteSpace R] -lemma closedEmbedding_nonUnitalStarAlgHom {a : A} {φ : C(σₙ S a, S)₀ →⋆ₙₐ[S] A} - (hφ : ClosedEmbedding φ) {f : C(S, R)} (h : QuasispectrumRestricts a f) +lemma isClosedEmbedding_nonUnitalStarAlgHom {a : A} {φ : C(σₙ S a, S)₀ →⋆ₙₐ[S] A} + (hφ : IsClosedEmbedding φ) {f : C(S, R)} (h : QuasispectrumRestricts a f) (halg : IsUniformEmbedding (algebraMap R S)) : - ClosedEmbedding (h.nonUnitalStarAlgHom φ) := by + IsClosedEmbedding (h.nonUnitalStarAlgHom φ) := by have : h.homeomorph.symm 0 = 0 := Subtype.ext (map_zero <| algebraMap _ _) - refine hφ.comp <| IsUniformEmbedding.toClosedEmbedding <| .comp + refine hφ.comp <| IsUniformEmbedding.toIsClosedEmbedding <| .comp (ContinuousMapZero.isUniformEmbedding_comp _ halg) (UniformEquiv.arrowCongrLeft₀ h.homeomorph.symm this |>.isUniformEmbedding) +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_nonUnitalStarAlgHom := isClosedEmbedding_nonUnitalStarAlgHom + variable [IsScalarTower R A A] [SMulCommClass R A A] /-- Given a `NonUnitalContinuousFunctionalCalculus S q`. If we form the predicate `p` for `a : A` @@ -244,13 +250,13 @@ protected theorem cfc (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) compactSpace_quasispectrum a := by have := NonUnitalContinuousFunctionalCalculus.compactSpace_quasispectrum (R := S) a rw [← isCompact_iff_compactSpace] at this ⊢ - simpa using halg.toClosedEmbedding.isCompact_preimage this + simpa using halg.toIsClosedEmbedding.isCompact_preimage this exists_cfc_of_predicate a ha := by refine ⟨((h a).mp ha).2.nonUnitalStarAlgHom (cfcₙHom ((h a).mp ha).1 (R := S)), - ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ - case hom_closedEmbedding => - exact ((h a).mp ha).2.closedEmbedding_nonUnitalStarAlgHom - (cfcₙHom_closedEmbedding ((h a).mp ha).1) halg + ?hom_isClosedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ + case hom_isClosedEmbedding => + exact ((h a).mp ha).2.isClosedEmbedding_nonUnitalStarAlgHom + (cfcₙHom_isClosedEmbedding ((h a).mp ha).1) halg case hom_id => exact ((h a).mp ha).2.nonUnitalStarAlgHom_id <| cfcₙHom_id ((h a).mp ha).1 case hom_map_spectrum => intro g @@ -289,7 +295,7 @@ lemma cfcₙHom_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap (hpa : p a) (hqa : q a) (h : QuasispectrumRestricts a f) : cfcₙHom hpa = h.nonUnitalStarAlgHom (cfcₙHom hqa) := by apply cfcₙHom_eq_of_continuous_of_map_id - · exact h.closedEmbedding_nonUnitalStarAlgHom (cfcₙHom_closedEmbedding hqa) halg |>.continuous + · exact h.isClosedEmbedding_nonUnitalStarAlgHom (cfcₙHom_isClosedEmbedding hqa) halg |>.continuous · exact h.nonUnitalStarAlgHom_id (cfcₙHom_id hqa) lemma cfcₙ_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) {a : A} diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean index 00f8c9bf97f18..5186931925074 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean @@ -169,7 +169,7 @@ instance NNReal.instUniqueContinuousFunctionalCalculus [UniqueContinuousFunction have : CompactSpace (spectrum ℝ a) := UniqueContinuousFunctionalCalculus.compactSpace_spectrum a rw [← isCompact_iff_compactSpace] at * rw [← spectrum.preimage_algebraMap ℝ] - exact closedEmbedding_subtype_val isClosed_nonneg |>.isCompact_preimage <| by assumption + exact isClosed_nonneg.isClosedEmbedding_subtypeVal.isCompact_preimage <| by assumption eq_of_continuous_of_map_id s hs φ ψ hφ hψ h := by let s' : Set ℝ := (↑) '' s let e : s ≃ₜ s' := @@ -376,7 +376,7 @@ instance NNReal.instUniqueNonUnitalContinuousFunctionalCalculus UniqueNonUnitalContinuousFunctionalCalculus.compactSpace_quasispectrum a rw [← isCompact_iff_compactSpace] at * rw [← quasispectrum.preimage_algebraMap ℝ] - exact closedEmbedding_subtype_val isClosed_nonneg |>.isCompact_preimage <| by assumption + exact isClosed_nonneg.isClosedEmbedding_subtypeVal.isCompact_preimage <| by assumption eq_of_continuous_of_map_id s hs _inst h0 φ ψ hφ hψ h := by let s' : Set ℝ := (↑) '' s let e : s ≃ₜ s' := diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index 29fa35441c946..c6070f8408685 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -165,7 +165,7 @@ class ContinuousFunctionalCalculus (R : Type*) {A : Type*} (p : outParam (A → [compactSpace_spectrum (a : A) : CompactSpace (spectrum R a)] spectrum_nonempty [Nontrivial A] (a : A) (ha : p a) : (spectrum R a).Nonempty exists_cfc_of_predicate : ∀ a, p a → ∃ φ : C(spectrum R a, R) →⋆ₐ[R] A, - ClosedEmbedding φ ∧ φ ((ContinuousMap.id R).restrict <| spectrum R a) = a ∧ + IsClosedEmbedding φ ∧ φ ((ContinuousMap.id R).restrict <| spectrum R a) = a ∧ (∀ f, spectrum R (φ f) = Set.range f) ∧ ∀ f, p (φ f) -- this instance should not be activated everywhere but it is useful when developing generic API @@ -228,13 +228,16 @@ user should instead prefer `cfc` over `cfcHom`. noncomputable def cfcHom : C(spectrum R a, R) →⋆ₐ[R] A := (ContinuousFunctionalCalculus.exists_cfc_of_predicate a ha).choose -lemma cfcHom_closedEmbedding : - ClosedEmbedding <| (cfcHom ha : C(spectrum R a, R) →⋆ₐ[R] A) := +lemma cfcHom_isClosedEmbedding : + IsClosedEmbedding <| (cfcHom ha : C(spectrum R a, R) →⋆ₐ[R] A) := (ContinuousFunctionalCalculus.exists_cfc_of_predicate a ha).choose_spec.1 +@[deprecated (since := "2024-10-20")] +alias cfcHom_closedEmbedding := cfcHom_isClosedEmbedding + @[fun_prop] lemma cfcHom_continuous : Continuous (cfcHom ha : C(spectrum R a, R) →⋆ₐ[R] A) := - cfcHom_closedEmbedding ha |>.continuous + cfcHom_isClosedEmbedding ha |>.continuous lemma cfcHom_id : cfcHom ha ((ContinuousMap.id R).restrict <| spectrum R a) = a := @@ -252,7 +255,7 @@ lemma cfcHom_predicate (f : C(spectrum R a, R)) : lemma cfcHom_eq_of_continuous_of_map_id [UniqueContinuousFunctionalCalculus R A] (φ : C(spectrum R a, R) →⋆ₐ[R] A) (hφ₁ : Continuous φ) (hφ₂ : φ (.restrict (spectrum R a) <| .id R) = a) : cfcHom ha = φ := - (cfcHom ha).ext_continuousMap a φ (cfcHom_closedEmbedding ha).continuous hφ₁ <| by + (cfcHom ha).ext_continuousMap a φ (cfcHom_isClosedEmbedding ha).continuous hφ₁ <| by rw [cfcHom_id ha, hφ₂] theorem cfcHom_comp [UniqueContinuousFunctionalCalculus R A] (f : C(spectrum R a, R)) @@ -263,7 +266,7 @@ theorem cfcHom_comp [UniqueContinuousFunctionalCalculus R A] (f : C(spectrum R a (cfcHom ha).comp <| ContinuousMap.compStarAlgHom' R R f' suffices cfcHom (cfcHom_predicate ha f) = φ from DFunLike.congr_fun this.symm g refine cfcHom_eq_of_continuous_of_map_id (cfcHom_predicate ha f) φ ?_ ?_ - · exact (cfcHom_closedEmbedding ha).continuous.comp f'.continuous_comp_left + · exact (cfcHom_isClosedEmbedding ha).continuous.comp f'.continuous_comp_left · simp only [φ, StarAlgHom.comp_apply, ContinuousMap.compStarAlgHom'_apply] congr ext x @@ -279,7 +282,7 @@ noncomputable def cfcL {a : A} (ha : p a) : C(spectrum R a, R) →L[R] A := { cfcHom ha with toFun := cfcHom ha map_smul' := map_smul _ - cont := (cfcHom_closedEmbedding ha).continuous } + cont := (cfcHom_isClosedEmbedding ha).continuous } end cfcL @@ -398,7 +401,7 @@ lemma eqOn_of_cfc_eq_cfc {f g : R → R} {a : A} (h : cfc f a = cfc g a) (hg : ContinuousOn g (spectrum R a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : (spectrum R a).EqOn f g := by rw [cfc_apply f a, cfc_apply g a] at h - have := (cfcHom_closedEmbedding (show p a from ha) (R := R)).inj h + have := (cfcHom_isClosedEmbedding (show p a from ha) (R := R)).inj h intro x hx congrm($(this) ⟨x, hx⟩) diff --git a/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean b/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean index af17ea013e3fc..a701ede3a9749 100644 --- a/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean +++ b/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean @@ -170,7 +170,7 @@ theorem gelfandTransform_bijective : Function.Bijective (gelfandTransform ℂ A) points in `C(characterSpace ℂ A, ℂ)` and is closed under `star`. -/ have h : rng.topologicalClosure = rng := le_antisymm (StarSubalgebra.topologicalClosure_minimal le_rfl - (gelfandTransform_isometry A).closedEmbedding.isClosed_range) + (gelfandTransform_isometry A).isClosedEmbedding.isClosed_range) (StarSubalgebra.le_topologicalClosure _) refine h ▸ ContinuousMap.starSubalgebra_topologicalClosure_eq_top_of_separatesPoints _ (fun _ _ => ?_) diff --git a/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean b/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean index 81863eb235bbc..21a62a2741891 100644 --- a/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean +++ b/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean @@ -130,7 +130,7 @@ theorem integral_bilinear_hasLineDerivAt_right_eq_neg_left_of_integrable -∫ (x : E' × ℝ), (B (f' (L.symm x))) (g (L.symm x)) ∂ν by have : μ = Measure.map L.symm ν := by simp [Measure.map_map L.symm.continuous.measurable L.continuous.measurable] - have hL : ClosedEmbedding L.symm := L.symm.toHomeomorph.closedEmbedding + have hL : IsClosedEmbedding L.symm := L.symm.toHomeomorph.isClosedEmbedding simpa [this, hL.integral_map] using H have L_emb : MeasurableEmbedding L := L.toHomeomorph.measurableEmbedding apply integral_bilinear_hasLineDerivAt_right_eq_neg_left_of_integrable_aux2 diff --git a/Mathlib/Analysis/Convex/Intrinsic.lean b/Mathlib/Analysis/Convex/Intrinsic.lean index 82cc1b1b0da2d..1706893f1b382 100644 --- a/Mathlib/Analysis/Convex/Intrinsic.lean +++ b/Mathlib/Analysis/Convex/Intrinsic.lean @@ -176,11 +176,11 @@ theorem intrinsicFrontier_union_intrinsicInterior (s : Set P) : theorem isClosed_intrinsicClosure (hs : IsClosed (affineSpan 𝕜 s : Set P)) : IsClosed (intrinsicClosure 𝕜 s) := - (closedEmbedding_subtype_val hs).isClosedMap _ isClosed_closure + hs.isClosedEmbedding_subtypeVal.isClosedMap _ isClosed_closure theorem isClosed_intrinsicFrontier (hs : IsClosed (affineSpan 𝕜 s : Set P)) : IsClosed (intrinsicFrontier 𝕜 s) := - (closedEmbedding_subtype_val hs).isClosedMap _ isClosed_frontier + hs.isClosedEmbedding_subtypeVal.isClosedMap _ isClosed_frontier @[simp] theorem affineSpan_intrinsicClosure (s : Set P) : diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 55806cb2f24fa..00f566fb58ae2 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1239,7 +1239,7 @@ theorem contDiffOn_convolution_right_with_param {f : G → E} {n : ℕ∞} (L : ((ContinuousLinearEquiv.arrowCongr isoE' isoF).symm : (E' →L[𝕜] F) →L[𝕜] eE' →L[𝕜] eF) L let R := fun q : eP × eG => (ef ⋆[eL, eμ] eg q.1) q.2 have R_contdiff : ContDiffOn 𝕜 n R ((isoP ⁻¹' s) ×ˢ univ) := by - have hek : IsCompact (isoG ⁻¹' k) := isoG.toHomeomorph.closedEmbedding.isCompact_preimage hk + have hek : IsCompact (isoG ⁻¹' k) := isoG.toHomeomorph.isClosedEmbedding.isCompact_preimage hk have hes : IsOpen (isoP ⁻¹' s) := isoP.continuous.isOpen_preimage _ hs refine contDiffOn_convolution_right_with_param_aux eL hes hek ?_ ?_ ?_ · intro p x hp hx @@ -1264,9 +1264,9 @@ theorem contDiffOn_convolution_right_with_param {f : G → E} {n : ℕ∞} (L : simp only [LinearIsometryEquiv.coe_coe, (· ∘ ·), ContinuousLinearEquiv.prod_symm, ContinuousLinearEquiv.prod_apply] simp only [R, convolution, coe_comp', ContinuousLinearEquiv.coe_coe, (· ∘ ·)] - rw [ClosedEmbedding.integral_map, ← isoF.integral_comp_comm] + rw [IsClosedEmbedding.integral_map, ← isoF.integral_comp_comm] · rfl - · exact isoG.symm.toHomeomorph.closedEmbedding + · exact isoG.symm.toHomeomorph.isClosedEmbedding simp_rw [this] at A exact A diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 31f90aefe8f74..39402193d24a1 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -334,7 +334,7 @@ theorem lintegral_pow_le_pow_lintegral_fderiv_aux [Fintype ι] _ ≤ ∫⁻ xᵢ in Iic (x i), ‖deriv (u ∘ update x i) xᵢ‖₊ := by apply le_trans (by simp) (HasCompactSupport.ennnorm_le_lintegral_Ici_deriv _ _ _) · exact hu.comp (by convert contDiff_update 1 x i) - · exact h2u.comp_closedEmbedding (closedEmbedding_update x i) + · exact h2u.comp_isClosedEmbedding (isClosedEmbedding_update x i) _ ≤ ∫⁻ xᵢ, (‖fderiv ℝ u (update x i xᵢ)‖₊ : ℝ≥0∞) := ?_ gcongr with y; swap · exact Measure.restrict_le_self diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index 04202dd932598..942cd2eeae3a2 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -133,7 +133,7 @@ instance instCompactSpaceNNReal {A : Type*} [NormedRing A] [NormedAlgebra ℝ A] (a : A) [CompactSpace (spectrum ℝ a)] : CompactSpace (spectrum ℝ≥0 a) := by rw [← isCompact_iff_compactSpace] at * rw [← preimage_algebraMap ℝ] - exact closedEmbedding_subtype_val isClosed_nonneg |>.isCompact_preimage <| by assumption + exact isClosed_nonneg.isClosedEmbedding_subtypeVal.isCompact_preimage <| by assumption section QuasispectrumCompact @@ -154,7 +154,7 @@ instance _root_.quasispectrum.instCompactSpaceNNReal [NormedSpace ℝ B] [IsScal CompactSpace (quasispectrum ℝ≥0 a) := by rw [← isCompact_iff_compactSpace] at * rw [← quasispectrum.preimage_algebraMap ℝ] - exact closedEmbedding_subtype_val isClosed_nonneg |>.isCompact_preimage <| by assumption + exact isClosed_nonneg.isClosedEmbedding_subtypeVal.isCompact_preimage <| by assumption end QuasispectrumCompact diff --git a/Mathlib/Analysis/Normed/Module/Basic.lean b/Mathlib/Analysis/Normed/Module/Basic.lean index d1bb26f18316d..09973f5989c6d 100644 --- a/Mathlib/Analysis/Normed/Module/Basic.lean +++ b/Mathlib/Analysis/Normed/Module/Basic.lean @@ -228,8 +228,8 @@ protected theorem NormedSpace.noncompactSpace : NoncompactSpace E := by exact ⟨fun h ↦ NormedSpace.unbounded_univ 𝕜 E h.isBounded⟩ · push_neg at H rcases exists_ne (0 : E) with ⟨x, hx⟩ - suffices ClosedEmbedding (Infinite.natEmbedding 𝕜 · • x) from this.noncompactSpace - refine closedEmbedding_of_pairwise_le_dist (norm_pos_iff.2 hx) fun k n hne ↦ ?_ + suffices IsClosedEmbedding (Infinite.natEmbedding 𝕜 · • x) from this.noncompactSpace + refine isClosedEmbedding_of_pairwise_le_dist (norm_pos_iff.2 hx) fun k n hne ↦ ?_ simp only [dist_eq_norm, ← sub_smul, norm_smul] rw [H, one_mul] rwa [sub_ne_zero, (Embedding.injective _).ne_iff] diff --git a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean index 4735888f210dd..34b802ea77856 100644 --- a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean +++ b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean @@ -519,8 +519,8 @@ lemma ProperSpace.of_locallyCompact_module [Nontrivial E] [LocallyCompactSpace E have : LocallyCompactSpace 𝕜 := by obtain ⟨v, hv⟩ : ∃ v : E, v ≠ 0 := exists_ne 0 let L : 𝕜 → E := fun t ↦ t • v - have : ClosedEmbedding L := closedEmbedding_smul_left hv - apply ClosedEmbedding.locallyCompactSpace this + have : IsClosedEmbedding L := isClosedEmbedding_smul_left hv + apply IsClosedEmbedding.locallyCompactSpace this .of_locallyCompactSpace 𝕜 @[deprecated (since := "2024-01-31")] diff --git a/Mathlib/Analysis/Normed/Operator/Compact.lean b/Mathlib/Analysis/Normed/Operator/Compact.lean index 633508c11661e..da69e15c03410 100644 --- a/Mathlib/Analysis/Normed/Operator/Compact.lean +++ b/Mathlib/Analysis/Normed/Operator/Compact.lean @@ -255,7 +255,7 @@ theorem IsCompactOperator.codRestrict {f : M₁ → M₂} (hf : IsCompactOperato (hV : ∀ x, f x ∈ V) (h_closed : IsClosed (V : Set M₂)) : IsCompactOperator (Set.codRestrict f V hV) := let ⟨_, hK, hKf⟩ := hf - ⟨_, (closedEmbedding_subtype_val h_closed).isCompact_preimage hK, hKf⟩ + ⟨_, h_closed.isClosedEmbedding_subtypeVal.isCompact_preimage hK, hKf⟩ end CodRestrict diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean index 8c7bd3f605070..65c353f847076 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean @@ -64,7 +64,7 @@ variable {𝕜 : Type*} {A : Type*} [RCLike 𝕜] {p : A → Prop} [NormedRing A lemma exp_eq_normedSpace_exp {a : A} (ha : p a := by cfc_tac) : cfc (exp 𝕜 : 𝕜 → 𝕜) a = exp 𝕜 a := by conv_rhs => rw [← cfc_id 𝕜 a ha, cfc_apply id a ha] - have h := (cfcHom_closedEmbedding (R := 𝕜) (show p a from ha)).continuous + have h := (cfcHom_isClosedEmbedding (R := 𝕜) (show p a from ha)).continuous have _ : ContinuousOn (exp 𝕜) (spectrum 𝕜 a) := exp_continuous.continuousOn simp_rw [← map_exp 𝕜 _ h, cfc_apply (exp 𝕜) a ha] congr 1 diff --git a/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLogExp.lean b/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLogExp.lean index 27f458ecd81d8..e257ffe48d381 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLogExp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLogExp.lean @@ -135,5 +135,5 @@ end Measurability end ENNReal instance : PolishSpace EReal := - ClosedEmbedding.polishSpace ⟨ENNReal.logOrderIso.symm.toHomeomorph.embedding, + IsClosedEmbedding.polishSpace ⟨ENNReal.logOrderIso.symm.toHomeomorph.embedding, ENNReal.logOrderIso.symm.toHomeomorph.range_coe ▸ isClosed_univ⟩ diff --git a/Mathlib/CategoryTheory/Galois/Topology.lean b/Mathlib/CategoryTheory/Galois/Topology.lean index 06ad6569ca157..f0bb5bf4ec1bb 100644 --- a/Mathlib/CategoryTheory/Galois/Topology.lean +++ b/Mathlib/CategoryTheory/Galois/Topology.lean @@ -79,26 +79,29 @@ lemma autEmbedding_range_isClosed : IsClosed (Set.range (autEmbedding F)) := by · fun_prop · fun_prop -lemma autEmbedding_closedEmbedding : ClosedEmbedding (autEmbedding F) where +lemma autEmbedding_isClosedEmbedding : IsClosedEmbedding (autEmbedding F) where induced := rfl inj := autEmbedding_injective F isClosed_range := autEmbedding_range_isClosed F -instance : CompactSpace (Aut F) := ClosedEmbedding.compactSpace (autEmbedding_closedEmbedding F) +@[deprecated (since := "2024-10-20")] +alias autEmbedding_closedEmbedding := autEmbedding_isClosedEmbedding + +instance : CompactSpace (Aut F) := IsClosedEmbedding.compactSpace (autEmbedding_isClosedEmbedding F) instance : T2Space (Aut F) := T2Space.of_injective_continuous (autEmbedding_injective F) continuous_induced_dom instance : TotallyDisconnectedSpace (Aut F) := - (Embedding.isTotallyDisconnected_range (autEmbedding_closedEmbedding F).embedding).mp + (Embedding.isTotallyDisconnected_range (autEmbedding_isClosedEmbedding F).embedding).mp (isTotallyDisconnected_of_totallyDisconnectedSpace _) instance : ContinuousMul (Aut F) := Inducing.continuousMul (autEmbedding F) - (autEmbedding_closedEmbedding F).toInducing + (autEmbedding_isClosedEmbedding F).toInducing instance : ContinuousInv (Aut F) := - Inducing.continuousInv (autEmbedding_closedEmbedding F).toInducing (fun _ ↦ rfl) + Inducing.continuousInv (autEmbedding_isClosedEmbedding F).toInducing (fun _ ↦ rfl) instance : TopologicalGroup (Aut F) := ⟨⟩ diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index ac7279ed0be9c..85e1eada2814c 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -280,19 +280,22 @@ protected theorem image_eq (s : Set H) : I '' s = I.symm ⁻¹' s ∩ range I := · rw [I.source_eq]; exact subset_univ _ · rw [inter_comm, I.target_eq, I.toPartialEquiv_coe_symm] -protected theorem closedEmbedding : ClosedEmbedding I := - I.leftInverse.closedEmbedding I.continuous_symm I.continuous +theorem isClosedEmbedding : IsClosedEmbedding I := + I.leftInverse.isClosedEmbedding I.continuous_symm I.continuous + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding := isClosedEmbedding theorem isClosed_range : IsClosed (range I) := - I.closedEmbedding.isClosed_range + I.isClosedEmbedding.isClosed_range @[deprecated (since := "2024-03-17")] alias closed_range := isClosed_range theorem map_nhds_eq (x : H) : map I (𝓝 x) = 𝓝[range I] I x := - I.closedEmbedding.toEmbedding.map_nhds_eq x + I.isClosedEmbedding.toEmbedding.map_nhds_eq x theorem map_nhdsWithin_eq (s : Set H) (x : H) : map I (𝓝[s] x) = 𝓝[I '' s] I x := - I.closedEmbedding.toEmbedding.map_nhdsWithin_eq s x + I.isClosedEmbedding.toEmbedding.map_nhdsWithin_eq s x theorem image_mem_nhdsWithin {x : H} {s : Set H} (hs : s ∈ 𝓝 x) : I '' s ∈ 𝓝[range I] I x := I.map_nhds_eq x ▸ image_mem_map hs @@ -348,7 +351,7 @@ open TopologicalSpace protected theorem secondCountableTopology [SecondCountableTopology E] (I : ModelWithCorners 𝕜 E H) : SecondCountableTopology H := - I.closedEmbedding.toEmbedding.secondCountableTopology + I.isClosedEmbedding.toEmbedding.secondCountableTopology end ModelWithCorners diff --git a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean index 3e3f2e4bb92fd..4e56b5f717089 100644 --- a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean +++ b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean @@ -128,9 +128,9 @@ supports of bump functions, then for some `n` it can be embedded into the `n`-di Euclidean space. -/ theorem exists_embedding_euclidean_of_compact [T2Space M] [CompactSpace M] : ∃ (n : ℕ) (e : M → EuclideanSpace ℝ (Fin n)), - Smooth I (𝓡 n) e ∧ ClosedEmbedding e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by + Smooth I (𝓡 n) e ∧ IsClosedEmbedding e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by rcases SmoothBumpCovering.exists_isSubordinate I isClosed_univ fun (x : M) _ => univ_mem with ⟨ι, f, -⟩ haveI := f.fintype rcases f.exists_immersion_euclidean with ⟨n, e, hsmooth, hinj, hinj_mfderiv⟩ - exact ⟨n, e, hsmooth, hsmooth.continuous.closedEmbedding hinj, hinj_mfderiv⟩ + exact ⟨n, e, hsmooth, hsmooth.continuous.isClosedEmbedding hinj, hinj_mfderiv⟩ diff --git a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean index 5c9c8e5757541..a1b13157f47c4 100644 --- a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean +++ b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean @@ -87,10 +87,10 @@ noncomputable def cfcAux : C(spectrum ℝ A, ℝ) →⋆ₐ[ℝ] (Matrix n n ext simp -lemma closedEmbedding_cfcAux : ClosedEmbedding hA.cfcAux := by +lemma isClosedEmbedding_cfcAux : IsClosedEmbedding hA.cfcAux := by have h0 : FiniteDimensional ℝ C(spectrum ℝ A, ℝ) := FiniteDimensional.of_injective (ContinuousMap.coeFnLinearMap ℝ (M := ℝ)) DFunLike.coe_injective - refine LinearMap.closedEmbedding_of_injective (𝕜 := ℝ) (E := C(spectrum ℝ A, ℝ)) + refine LinearMap.isClosedEmbedding_of_injective (𝕜 := ℝ) (E := C(spectrum ℝ A, ℝ)) (F := Matrix n n 𝕜) (f := hA.cfcAux) <| LinearMap.ker_eq_bot'.mpr fun f hf ↦ ?_ have h2 : diagonal (RCLike.ofReal ∘ f ∘ fun i ↦ ⟨hA.eigenvalues i, hA.eigenvalues_mem_spectrum_real i⟩) @@ -107,6 +107,9 @@ lemma closedEmbedding_cfcAux : ClosedEmbedding hA.cfcAux := by have := (diagonal_eq_diagonal_iff).mp h2 refine RCLike.ofReal_eq_zero.mp (this i) +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_cfcAux := isClosedEmbedding_cfcAux + lemma cfcAux_id : hA.cfcAux (.restrict (spectrum ℝ A) (.id ℝ)) = A := by conv_rhs => rw [hA.spectral_theorem] congr! @@ -117,7 +120,7 @@ instance instContinuousFunctionalCalculus : ContinuousFunctionalCalculus ℝ (IsSelfAdjoint : Matrix n n 𝕜 → Prop) where exists_cfc_of_predicate a ha := by replace ha : IsHermitian a := ha - refine ⟨ha.cfcAux, ha.closedEmbedding_cfcAux, ha.cfcAux_id, fun f ↦ ?map_spec, + refine ⟨ha.cfcAux, ha.isClosedEmbedding_cfcAux, ha.cfcAux_id, fun f ↦ ?map_spec, fun f ↦ ?hermitian⟩ case map_spec => apply Set.eq_of_subset_of_subset @@ -159,7 +162,7 @@ protected noncomputable def cfc (f : ℝ → ℝ) : Matrix n n 𝕜 := lemma cfc_eq (f : ℝ → ℝ) : cfc f A = hA.cfc f := by have hA' : IsSelfAdjoint A := hA - have := cfcHom_eq_of_continuous_of_map_id hA' hA.cfcAux hA.closedEmbedding_cfcAux.continuous + have := cfcHom_eq_of_continuous_of_map_id hA' hA.cfcAux hA.isClosedEmbedding_cfcAux.continuous hA.cfcAux_id rw [cfc_apply f A hA' (by rw [continuousOn_iff_continuous_restrict]; fun_prop), this] simp only [cfcAux_apply, ContinuousMap.coe_mk, Function.comp_def, Set.restrict_apply, diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 601ec698d26dd..c5518f43dd87a 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -472,9 +472,12 @@ is ae-measurable. -/ theorem Continuous.aemeasurable {f : α → γ} (h : Continuous f) {μ : Measure α} : AEMeasurable f μ := h.measurable.aemeasurable -theorem ClosedEmbedding.measurable {f : α → γ} (hf : ClosedEmbedding f) : Measurable f := +theorem IsClosedEmbedding.measurable {f : α → γ} (hf : IsClosedEmbedding f) : Measurable f := hf.continuous.measurable +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.measurable := IsClosedEmbedding.measurable + /-- If a function is defined piecewise in terms of functions which are continuous on their respective pieces, then it is measurable. -/ theorem ContinuousOn.measurable_piecewise {f g : α → γ} {s : Set α} [∀ j : α, Decidable (j ∈ s)] @@ -632,10 +635,13 @@ protected theorem Embedding.measurableEmbedding {f : α → β} (h₁ : Embeddin (((↑) : range f → β) ∘ (Homeomorph.ofEmbedding f h₁).toMeasurableEquiv) from (MeasurableEmbedding.subtype_coe h₂).comp (MeasurableEquiv.measurableEmbedding _) -protected theorem ClosedEmbedding.measurableEmbedding {f : α → β} (h : ClosedEmbedding f) : +protected theorem IsClosedEmbedding.measurableEmbedding {f : α → β} (h : IsClosedEmbedding f) : MeasurableEmbedding f := h.toEmbedding.measurableEmbedding h.isClosed_range.measurableSet +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.measurableEmbedding := IsClosedEmbedding.measurableEmbedding + protected theorem IsOpenEmbedding.measurableEmbedding {f : α → β} (h : IsOpenEmbedding f) : MeasurableEmbedding f := h.toEmbedding.measurableEmbedding h.isOpen_range.measurableSet diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/ContinuousLinearMap.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/ContinuousLinearMap.lean index 32e8e945358ec..c32b095c05312 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/ContinuousLinearMap.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/ContinuousLinearMap.lean @@ -90,10 +90,10 @@ variable [BorelSpace 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 theorem measurable_smul_const {f : α → 𝕜} {c : E} (hc : c ≠ 0) : (Measurable fun x => f x • c) ↔ Measurable f := - (closedEmbedding_smul_left hc).measurableEmbedding.measurable_comp_iff + (isClosedEmbedding_smul_left hc).measurableEmbedding.measurable_comp_iff theorem aemeasurable_smul_const {f : α → 𝕜} {μ : Measure α} {c : E} (hc : c ≠ 0) : AEMeasurable (fun x => f x • c) μ ↔ AEMeasurable f μ := - (closedEmbedding_smul_left hc).measurableEmbedding.aemeasurable_comp_iff + (isClosedEmbedding_smul_left hc).measurableEmbedding.aemeasurable_comp_iff end NormedSpace diff --git a/Mathlib/MeasureTheory/Constructions/Polish/EmbeddingReal.lean b/Mathlib/MeasureTheory/Constructions/Polish/EmbeddingReal.lean index 67a330bbb5bb7..49a72519b60d3 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish/EmbeddingReal.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish/EmbeddingReal.lean @@ -24,7 +24,7 @@ theorem exists_nat_measurableEquiv_range_coe_fin_of_finite [Finite α] : theorem measurableEquiv_range_coe_nat_of_infinite_of_countable [Infinite α] [Countable α] : Nonempty (α ≃ᵐ range ((↑) : ℕ → ℝ)) := by have : PolishSpace (range ((↑) : ℕ → ℝ)) := - Nat.closedEmbedding_coe_real.isClosedMap.isClosed_range.polishSpace + Nat.isClosedEmbedding_coe_real.isClosedMap.isClosed_range.polishSpace refine ⟨PolishSpace.Equiv.measurableEquiv ?_⟩ refine (nonempty_equiv_of_countable.some : α ≃ ℕ).trans ?_ exact Equiv.ofInjective ((↑) : ℕ → ℝ) Nat.cast_injective diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 7d3943e021a7c..012d894c0ae6f 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -691,7 +691,7 @@ theorem _root_.Embedding.comp_stronglyMeasurable_iff {m : MeasurableSpace α} [T ⟨fun H => stronglyMeasurable_iff_measurable_separable.2 ⟨?_, ?_⟩, fun H => hg.continuous.comp_stronglyMeasurable H⟩ · let G : β → range g := rangeFactorization g - have hG : ClosedEmbedding G := + have hG : IsClosedEmbedding G := { hg.codRestrict _ _ with isClosed_range := by rw [surjective_onto_range.range_eq] @@ -1513,7 +1513,7 @@ theorem _root_.Embedding.aestronglyMeasurable_comp_iff [PseudoMetrizableSpace β ⟨fun H => aestronglyMeasurable_iff_aemeasurable_separable.2 ⟨?_, ?_⟩, fun H => hg.continuous.comp_aestronglyMeasurable H⟩ · let G : β → range g := rangeFactorization g - have hG : ClosedEmbedding G := + have hG : IsClosedEmbedding G := { hg.codRestrict _ _ with isClosed_range := by rw [surjective_onto_range.range_eq]; exact isClosed_univ } have : AEMeasurable (G ∘ f) μ := AEMeasurable.subtype_mk H.aemeasurable diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean index 4251891161db2..95fbcc33ea7cf 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean @@ -46,7 +46,7 @@ variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] theorem aestronglyMeasurable_smul_const_iff {f : α → 𝕜} {c : E} (hc : c ≠ 0) : AEStronglyMeasurable (fun x => f x • c) μ ↔ AEStronglyMeasurable f μ := - (closedEmbedding_smul_left hc).toEmbedding.aestronglyMeasurable_comp_iff + (isClosedEmbedding_smul_left hc).toEmbedding.aestronglyMeasurable_comp_iff end NormedSpace diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index e058bb5e8e9af..49764e92934c6 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -775,7 +775,7 @@ nonrec theorem _root_.MulEquiv.isHaarMeasure_map [BorelSpace G] [TopologicalGrou [TopologicalGroup H] (e : G ≃* H) (he : Continuous e) (hesymm : Continuous e.symm) : IsHaarMeasure (Measure.map e μ) := let f : G ≃ₜ H := .mk e - isHaarMeasure_map μ (e : G →* H) he e.surjective f.closedEmbedding.tendsto_cocompact + isHaarMeasure_map μ (e : G →* H) he e.surjective f.isClosedEmbedding.tendsto_cocompact /-- A convenience wrapper for MeasureTheory.Measure.isAddHaarMeasure_map`. -/ instance _root_.ContinuousLinearEquiv.isAddHaarMeasure_map diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index b497a13fc272b..3c5f56fce5d8b 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1574,11 +1574,14 @@ theorem _root_.MeasurableEmbedding.integral_map {β} {_ : MeasurableSpace β} {f · rw [integral_non_aestronglyMeasurable hgm, integral_non_aestronglyMeasurable] exact fun hgf => hgm (hf.aestronglyMeasurable_map_iff.2 hgf) -theorem _root_.ClosedEmbedding.integral_map {β} [TopologicalSpace α] [BorelSpace α] - [TopologicalSpace β] [MeasurableSpace β] [BorelSpace β] {φ : α → β} (hφ : ClosedEmbedding φ) +theorem _root_.IsClosedEmbedding.integral_map {β} [TopologicalSpace α] [BorelSpace α] + [TopologicalSpace β] [MeasurableSpace β] [BorelSpace β] {φ : α → β} (hφ : IsClosedEmbedding φ) (f : β → G) : ∫ y, f y ∂Measure.map φ μ = ∫ x, f (φ x) ∂μ := hφ.measurableEmbedding.integral_map _ +@[deprecated (since := "2024-10-20")] +alias _root_.ClosedEmbedding.integral_map := _root_.IsClosedEmbedding.integral_map + theorem integral_map_equiv {β} [MeasurableSpace β] (e : α ≃ᵐ β) (f : β → G) : ∫ y, f y ∂Measure.map e μ = ∫ x, f (e x) ∂μ := e.measurableEmbedding.integral_map f diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 0c443e0792ad6..6b904ee39dabd 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -271,7 +271,7 @@ theorem comp_mul_left (hf : IntervalIntegrable f volume a b) (c : ℝ) : rcases eq_or_ne c 0 with (hc | hc); · rw [hc]; simp rw [intervalIntegrable_iff'] at hf ⊢ have A : MeasurableEmbedding fun x => x * c⁻¹ := - (Homeomorph.mulRight₀ _ (inv_ne_zero hc)).closedEmbedding.measurableEmbedding + (Homeomorph.mulRight₀ _ (inv_ne_zero hc)).isClosedEmbedding.measurableEmbedding rw [← Real.smul_map_volume_mul_right (inv_ne_zero hc), IntegrableOn, Measure.restrict_smul, integrable_smul_measure (by simpa : ENNReal.ofReal |c⁻¹| ≠ 0) ENNReal.ofReal_ne_top, ← IntegrableOn, MeasurableEmbedding.integrableOn_map_iff A] @@ -294,7 +294,7 @@ theorem comp_add_right (hf : IntervalIntegrable f volume a b) (c : ℝ) : · exact IntervalIntegrable.symm (this hf.symm (le_of_not_le h)) rw [intervalIntegrable_iff'] at hf ⊢ have A : MeasurableEmbedding fun x => x + c := - (Homeomorph.addRight c).closedEmbedding.measurableEmbedding + (Homeomorph.addRight c).isClosedEmbedding.measurableEmbedding rw [← map_add_right_eq_self volume c] at hf convert (MeasurableEmbedding.integrableOn_map_iff A).mp hf using 1 rw [preimage_add_const_uIcc] @@ -624,7 +624,7 @@ variable {a b c d : ℝ} (f : ℝ → E) theorem integral_comp_mul_right (hc : c ≠ 0) : (∫ x in a..b, f (x * c)) = c⁻¹ • ∫ x in a * c..b * c, f x := by have A : MeasurableEmbedding fun x => x * c := - (Homeomorph.mulRight₀ c hc).closedEmbedding.measurableEmbedding + (Homeomorph.mulRight₀ c hc).isClosedEmbedding.measurableEmbedding conv_rhs => rw [← Real.smul_map_volume_mul_right hc] simp_rw [integral_smul_measure, intervalIntegral, A.setIntegral_map, ENNReal.toReal_ofReal (abs_nonneg c)] @@ -661,7 +661,7 @@ theorem inv_smul_integral_comp_div (c) : @[simp] theorem integral_comp_add_right (d) : (∫ x in a..b, f (x + d)) = ∫ x in a + d..b + d, f x := have A : MeasurableEmbedding fun x => x + d := - (Homeomorph.addRight d).closedEmbedding.measurableEmbedding + (Homeomorph.addRight d).isClosedEmbedding.measurableEmbedding calc (∫ x in a..b, f (x + d)) = ∫ x in a + d..b + d, f x ∂Measure.map (fun x => x + d) volume := by simp [intervalIntegral, A.setIntegral_map] diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 76b9b535d4c5c..6376b28c74713 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -558,13 +558,19 @@ theorem _root_.MeasurableEmbedding.setIntegral_map {Y} {_ : MeasurableSpace Y} { @[deprecated (since := "2024-04-17")] alias _root_.MeasurableEmbedding.set_integral_map := _root_.MeasurableEmbedding.setIntegral_map -theorem _root_.ClosedEmbedding.setIntegral_map [TopologicalSpace X] [BorelSpace X] {Y} +theorem _root_.IsClosedEmbedding.setIntegral_map [TopologicalSpace X] [BorelSpace X] {Y} [MeasurableSpace Y] [TopologicalSpace Y] [BorelSpace Y] {g : X → Y} {f : Y → E} (s : Set Y) - (hg : ClosedEmbedding g) : ∫ y in s, f y ∂Measure.map g μ = ∫ x in g ⁻¹' s, f (g x) ∂μ := + (hg : IsClosedEmbedding g) : ∫ y in s, f y ∂Measure.map g μ = ∫ x in g ⁻¹' s, f (g x) ∂μ := hg.measurableEmbedding.setIntegral_map _ _ +@[deprecated (since := "2024-10-20")] +alias _root_.ClosedEmbedding.setIntegral_map := IsClosedEmbedding.setIntegral_map + @[deprecated (since := "2024-04-17")] -alias _root_.ClosedEmbedding.set_integral_map := _root_.ClosedEmbedding.setIntegral_map +alias _root_.IsClosedEmbedding.set_integral_map := _root_.IsClosedEmbedding.setIntegral_map + +@[deprecated (since := "2024-10-20")] +alias _root_.ClosedEmbedding.set_integral_map := IsClosedEmbedding.set_integral_map theorem MeasurePreserving.setIntegral_preimage_emb {Y} {_ : MeasurableSpace Y} {f : X → Y} {ν} (h₁ : MeasurePreserving f μ ν) (h₂ : MeasurableEmbedding f) (g : Y → E) (s : Set Y) : diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean index 1b31285651c1c..e4968335130b1 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean @@ -80,7 +80,7 @@ itself, it does not apply when `f` is more complicated -/ theorem integral_comp_neg_Iic {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] (c : ℝ) (f : ℝ → E) : (∫ x in Iic c, f (-x)) = ∫ x in Ioi (-c), f x := by have A : MeasurableEmbedding fun x : ℝ => -x := - (Homeomorph.neg ℝ).closedEmbedding.measurableEmbedding + (Homeomorph.neg ℝ).isClosedEmbedding.measurableEmbedding have := MeasurableEmbedding.setIntegral_map (μ := volume) A f (Ici (-c)) rw [Measure.map_neg_eq_self (volume : Measure ℝ)] at this simp_rw [← integral_Ici_eq_integral_Ioi, this, neg_preimage, preimage_neg_Ici, neg_neg] diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index d27a687df5826..339c0c6fe0f94 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -113,7 +113,7 @@ attribute [local simp] ContinuousLinearMap.coe_smul theorem tendsto_normSq_coprime_pair : Filter.Tendsto (fun p : Fin 2 → ℤ => normSq ((p 0 : ℂ) * z + p 1)) cofinite atTop := by -- using this instance rather than the automatic `Function.module` makes unification issues in - -- `LinearEquiv.closedEmbedding_of_injective` less bad later in the proof. + -- `LinearEquiv.isClosedEmbedding_of_injective` less bad later in the proof. letI : Module ℝ (Fin 2 → ℝ) := NormedSpace.toModule let π₀ : (Fin 2 → ℝ) →ₗ[ℝ] ℝ := LinearMap.proj 0 let π₁ : (Fin 2 → ℝ) →ₗ[ℝ] ℝ := LinearMap.proj 1 @@ -149,7 +149,7 @@ theorem tendsto_normSq_coprime_pair : rw [f_def, RingHom.map_add, RingHom.map_mul, mul_add, mul_left_comm, mul_conj, conj_ofReal, conj_ofReal, ← ofReal_mul, add_im, ofReal_im, zero_add, inv_mul_eq_iff_eq_mul₀ hz] simp only [ofReal_im, ofReal_re, mul_im, zero_add, mul_zero] - have hf' : ClosedEmbedding f := f.closedEmbedding_of_injective hf + have hf' : IsClosedEmbedding f := f.isClosedEmbedding_of_injective hf have h₂ : Tendsto (fun p : Fin 2 → ℤ => ((↑) : ℤ → ℝ) ∘ p) cofinite (cocompact _) := by convert Tendsto.pi_map_coprodᵢ fun _ => Int.tendsto_coe_cofinite · rw [coprodᵢ_cofinite] @@ -207,8 +207,8 @@ theorem tendsto_lcRow0 {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : Tendsto.pi_map_coprodᵢ fun _ : Fin 2 => Int.tendsto_coe_cofinite have hf₁ : Tendsto f₁ cofinite (cocompact _) := cocompact_ℝ_to_cofinite_ℤ_matrix.comp Subtype.coe_injective.tendsto_cofinite - have hf₂ : ClosedEmbedding (lcRow0Extend hcd) := - (lcRow0Extend hcd).toContinuousLinearEquiv.toHomeomorph.closedEmbedding + have hf₂ : IsClosedEmbedding (lcRow0Extend hcd) := + (lcRow0Extend hcd).toContinuousLinearEquiv.toHomeomorph.isClosedEmbedding convert hf₂.tendsto_cocompact.comp (hf₁.comp Subtype.coe_injective.tendsto_cofinite) using 1 ext ⟨g, rfl⟩ i j : 3 fin_cases i <;> [fin_cases j; skip] @@ -258,7 +258,7 @@ theorem tendsto_abs_re_smul {p : Fin 2 → ℤ} (hp : IsCoprime (p 0) (p 1)) : let f := Homeomorph.mulRight₀ _ this let ff := Homeomorph.addRight (((p 1 : ℂ) * z - p 0) / (((p 0 : ℂ) ^ 2 + (p 1 : ℂ) ^ 2) * (p 0 * z + p 1))).re - convert (f.trans ff).closedEmbedding.tendsto_cocompact.comp (tendsto_lcRow0 hp) with _ _ g + convert (f.trans ff).isClosedEmbedding.tendsto_cocompact.comp (tendsto_lcRow0 hp) with _ _ g change ((g : SL(2, ℤ)) • z).re = lcRow0 p ↑(↑g : SL(2, ℝ)) / ((p 0 : ℝ) ^ 2 + (p 1 : ℝ) ^ 2) + diff --git a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean index a907bb2550a2a..575ad9565a29c 100644 --- a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean +++ b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean @@ -50,9 +50,13 @@ theorem coe_mkHomeomorph_symm : ⇑(mkHomeomorph : M ≃ₜ Mᵈᵐᵃ).symm = m @[to_additive] theorem inducing_mk : Inducing (@mk M) := mkHomeomorph.inducing @[to_additive] theorem embedding_mk : Embedding (@mk M) := mkHomeomorph.embedding @[to_additive] theorem isOpenEmbedding_mk : IsOpenEmbedding (@mk M) := mkHomeomorph.isOpenEmbedding -@[to_additive] theorem closedEmbedding_mk : ClosedEmbedding (@mk M) := mkHomeomorph.closedEmbedding +@[to_additive] theorem isClosedEmbedding_mk : IsClosedEmbedding (@mk M) := + mkHomeomorph.isClosedEmbedding @[to_additive] theorem quotientMap_mk : QuotientMap (@mk M) := mkHomeomorph.quotientMap +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_mk := isClosedEmbedding_mk + @[deprecated (since := "2024-10-18")] alias openEmbedding_mk := isOpenEmbedding_mk @@ -66,7 +70,11 @@ theorem isOpenEmbedding_mk_symm : IsOpenEmbedding (@mk M).symm := mkHomeomorph.s alias openEmbedding_mk_symm := isOpenEmbedding_mk_symm @[to_additive] -theorem closedEmbedding_mk_symm : ClosedEmbedding (@mk M).symm := mkHomeomorph.symm.closedEmbedding +theorem isClosedEmbedding_mk_symm : IsClosedEmbedding (@mk M).symm := + mkHomeomorph.symm.isClosedEmbedding + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_mk_symm := isClosedEmbedding_mk_symm @[to_additive] theorem quotientMap_mk_symm : QuotientMap (@mk M).symm := mkHomeomorph.symm.quotientMap @@ -76,8 +84,8 @@ theorem quotientMap_mk_symm : QuotientMap (@mk M).symm := mkHomeomorph.symm.quot @[to_additive] instance instT2Space [T2Space M] : T2Space Mᵈᵐᵃ := embedding_mk_symm.t2Space @[to_additive] instance instT25Space [T25Space M] : T25Space Mᵈᵐᵃ := embedding_mk_symm.t25Space @[to_additive] instance instT3Space [T3Space M] : T3Space Mᵈᵐᵃ := embedding_mk_symm.t3Space -@[to_additive] instance instT4Space [T4Space M] : T4Space Mᵈᵐᵃ := closedEmbedding_mk_symm.t4Space -@[to_additive] instance instT5Space [T5Space M] : T5Space Mᵈᵐᵃ := closedEmbedding_mk_symm.t5Space +@[to_additive] instance instT4Space [T4Space M] : T4Space Mᵈᵐᵃ := isClosedEmbedding_mk_symm.t4Space +@[to_additive] instance instT5Space [T5Space M] : T5Space Mᵈᵐᵃ := isClosedEmbedding_mk_symm.t5Space @[to_additive] instance instR0Space [R0Space M] : R0Space Mᵈᵐᵃ := embedding_mk_symm.r0Space @[to_additive] instance instR1Space [R1Space M] : R1Space Mᵈᵐᵃ := embedding_mk_symm.r1Space @@ -86,11 +94,11 @@ theorem quotientMap_mk_symm : QuotientMap (@mk M).symm := mkHomeomorph.symm.quot instance instRegularSpace [RegularSpace M] : RegularSpace Mᵈᵐᵃ := embedding_mk_symm.regularSpace @[to_additive] -instance instNormalSpace [NormalSpace M] : NormalSpace Mᵈᵐᵃ := closedEmbedding_mk_symm.normalSpace +instance instNormalSpace [NormalSpace M] : NormalSpace Mᵈᵐᵃ := isClosedEmbedding_mk_symm.normalSpace @[to_additive] instance instCompletelyNormalSpace [CompletelyNormalSpace M] : CompletelyNormalSpace Mᵈᵐᵃ := - closedEmbedding_mk_symm.completelyNormalSpace + isClosedEmbedding_mk_symm.completelyNormalSpace @[to_additive] instance instDiscreteTopology [DiscreteTopology M] : DiscreteTopology Mᵈᵐᵃ := @@ -119,7 +127,7 @@ instance instLocallyCompactSpace [LocallyCompactSpace M] : LocallyCompactSpace M @[to_additive] instance instWeaklyLocallyCompactSpace [WeaklyLocallyCompactSpace M] : WeaklyLocallyCompactSpace Mᵈᵐᵃ := - closedEmbedding_mk_symm.weaklyLocallyCompactSpace + isClosedEmbedding_mk_symm.weaklyLocallyCompactSpace @[to_additive (attr := simp)] theorem map_mk_nhds (x : M) : map (mk : M → Mᵈᵐᵃ) (𝓝 x) = 𝓝 (mk x) := diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index a41b781ceaadc..faf9ab016a5e2 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -258,8 +258,8 @@ lemma range_toContinuousMap : exact ⟨{ f with map_one' := h1, map_mul' := hmul }, rfl⟩ @[to_additive] -theorem closedEmbedding_toContinuousMap [ContinuousMul B] [T2Space B] : - ClosedEmbedding (toContinuousMap : ContinuousMonoidHom A B → C(A, B)) where +theorem isClosedEmbedding_toContinuousMap [ContinuousMul B] [T2Space B] : + IsClosedEmbedding (toContinuousMap : ContinuousMonoidHom A B → C(A, B)) where toEmbedding := embedding_toContinuousMap A B isClosed_range := by simp only [range_toContinuousMap, Set.setOf_and, Set.setOf_forall] @@ -268,6 +268,9 @@ theorem closedEmbedding_toContinuousMap [ContinuousMul B] [T2Space B] : exact isClosed_eq (continuous_eval_const (x * y)) <| .mul (continuous_eval_const x) (continuous_eval_const y) +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_toContinuousMap := isClosedEmbedding_toContinuousMap + variable {A B C D E} @[to_additive] diff --git a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean index 52bcb32bef151..61776f9e7c9c2 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean @@ -158,11 +158,14 @@ theorem hasBasis_nhds_zero : variable [ContinuousSMul 𝕜 E] -lemma closedEmbedding_toContinuousMultilinearMap [T2Space F] : - ClosedEmbedding (toContinuousMultilinearMap : +lemma isClosedEmbedding_toContinuousMultilinearMap [T2Space F] : + IsClosedEmbedding (toContinuousMultilinearMap : (E [⋀^ι]→L[𝕜] F) → ContinuousMultilinearMap 𝕜 (fun _ : ι ↦ E) F) := ⟨embedding_toContinuousMultilinearMap, isClosed_range_toContinuousMultilinearMap⟩ +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_toContinuousMultilinearMap := isClosedEmbedding_toContinuousMultilinearMap + instance instContinuousEvalConst : ContinuousEvalConst (E [⋀^ι]→L[𝕜] F) (ι → E) F := .of_continuous_forget continuous_toContinuousMultilinearMap diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index 2cf817161d306..ccbd57b6b210b 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -518,24 +518,30 @@ theorem Submodule.closed_of_finiteDimensional s.complete_of_finiteDimensional.isClosed /-- An injective linear map with finite-dimensional domain is a closed embedding. -/ -theorem LinearMap.closedEmbedding_of_injective [T2Space E] [FiniteDimensional 𝕜 E] {f : E →ₗ[𝕜] F} - (hf : LinearMap.ker f = ⊥) : ClosedEmbedding f := +theorem LinearMap.isClosedEmbedding_of_injective [T2Space E] [FiniteDimensional 𝕜 E] {f : E →ₗ[𝕜] F} + (hf : LinearMap.ker f = ⊥) : IsClosedEmbedding f := let g := LinearEquiv.ofInjective f (LinearMap.ker_eq_bot.mp hf) { embedding_subtype_val.comp g.toContinuousLinearEquiv.toHomeomorph.embedding with isClosed_range := by haveI := f.finiteDimensional_range simpa [LinearMap.range_coe f] using f.range.closed_of_finiteDimensional } -theorem closedEmbedding_smul_left [T2Space E] {c : E} (hc : c ≠ 0) : - ClosedEmbedding fun x : 𝕜 => x • c := - LinearMap.closedEmbedding_of_injective (LinearMap.ker_toSpanSingleton 𝕜 E hc) +@[deprecated (since := "2024-10-20")] +alias LinearMap.closedEmbedding_of_injective := LinearMap.isClosedEmbedding_of_injective + +theorem isClosedEmbedding_smul_left [T2Space E] {c : E} (hc : c ≠ 0) : + IsClosedEmbedding fun x : 𝕜 => x • c := + LinearMap.isClosedEmbedding_of_injective (LinearMap.ker_toSpanSingleton 𝕜 E hc) + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_smul_left := isClosedEmbedding_smul_left -- `smul` is a closed map in the first argument. theorem isClosedMap_smul_left [T2Space E] (c : E) : IsClosedMap fun x : 𝕜 => x • c := by by_cases hc : c = 0 · simp_rw [hc, smul_zero] exact isClosedMap_const - · exact (closedEmbedding_smul_left hc).isClosedMap + · exact (isClosedEmbedding_smul_left hc).isClosedMap theorem ContinuousLinearMap.exists_right_inverse_of_surjective [FiniteDimensional 𝕜 F] (f : E →L[𝕜] F) (hf : LinearMap.range f = ⊤) : diff --git a/Mathlib/Topology/Algebra/ProperAction/Basic.lean b/Mathlib/Topology/Algebra/ProperAction/Basic.lean index 92fc69d56d9f4..8b179ad23ab0f 100644 --- a/Mathlib/Topology/Algebra/ProperAction/Basic.lean +++ b/Mathlib/Topology/Algebra/ProperAction/Basic.lean @@ -127,9 +127,7 @@ then this topological space is T2."] theorem t2Space_of_properSMul_of_t2Group [h_proper : ProperSMul G X] [T2Space G] : T2Space X := by let f := fun x : X ↦ ((1 : G), x) have proper_f : IsProperMap f := by - apply isProperMap_of_closedEmbedding - rw [closedEmbedding_iff] - constructor + refine IsClosedEmbedding.isProperMap ⟨?_, ?_⟩ · let g := fun gx : G × X ↦ gx.2 have : Function.LeftInverse g f := fun x ↦ by simp exact this.embedding (by fun_prop) (by fun_prop) @@ -150,18 +148,20 @@ then `H` also acts properly on `X`. -/ @[to_additive "If two groups `H` and `G` act on a topological space `X` such that `G` acts properly and there exists a group homomorphims `H → G` which is a closed embedding compatible with the actions, then `H` also acts properly on `X`."] -theorem properSMul_of_closedEmbedding {H : Type*} [Group H] [MulAction H X] [TopologicalSpace H] - [ProperSMul G X] (f : H →* G) (f_clemb : ClosedEmbedding f) +theorem properSMul_of_isClosedEmbedding {H : Type*} [Group H] [MulAction H X] [TopologicalSpace H] + [ProperSMul G X] (f : H →* G) (f_clemb : IsClosedEmbedding f) (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)) := this.prodMap isProperMap_id + have h : IsProperMap (Prod.map f (fun x : X ↦ x)) := f_clemb.isProperMap.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] exact h.comp <| ProperSMul.isProperMap_smul_pair +@[deprecated (since := "2024-10-20")] +alias properSMul_of_closedEmbedding := properSMul_of_isClosedEmbedding + /-- If `H` is a closed subgroup of `G` and `G` acts properly on X then so does `H`. -/ @[to_additive "If `H` is a closed subgroup of `G` and `G` acts properly on X then so does `H`."] instance {H : Subgroup G} [ProperSMul G X] [H_closed : IsClosed (H : Set G)] : ProperSMul H X := - properSMul_of_closedEmbedding H.subtype H_closed.closedEmbedding_subtype_val fun _ _ ↦ rfl + properSMul_of_isClosedEmbedding H.subtype H_closed.isClosedEmbedding_subtypeVal fun _ _ ↦ rfl diff --git a/Mathlib/Topology/Algebra/StarSubalgebra.lean b/Mathlib/Topology/Algebra/StarSubalgebra.lean index a8dd137d03c23..295e2416e5169 100644 --- a/Mathlib/Topology/Algebra/StarSubalgebra.lean +++ b/Mathlib/Topology/Algebra/StarSubalgebra.lean @@ -37,9 +37,9 @@ theorem embedding_inclusion {S₁ S₂ : StarSubalgebra R A} (h : S₁ ≤ S₂) { induced := Eq.symm induced_compose inj := Subtype.map_injective h Function.injective_id } -/-- The `StarSubalgebra.inclusion` of a closed star subalgebra is a `ClosedEmbedding`. -/ -theorem closedEmbedding_inclusion {S₁ S₂ : StarSubalgebra R A} (h : S₁ ≤ S₂) - (hS₁ : IsClosed (S₁ : Set A)) : ClosedEmbedding (inclusion h) := +/-- The `StarSubalgebra.inclusion` of a closed star subalgebra is a `IsClosedEmbedding`. -/ +theorem isClosedEmbedding_inclusion {S₁ S₂ : StarSubalgebra R A} (h : S₁ ≤ S₂) + (hS₁ : IsClosed (S₁ : Set A)) : IsClosedEmbedding (inclusion h) := { embedding_inclusion h with isClosed_range := isClosed_induced_iff.2 ⟨S₁, hS₁, by @@ -48,6 +48,9 @@ theorem closedEmbedding_inclusion {S₁ S₂ : StarSubalgebra R A} (h : S₁ ≤ · intro _ h' apply h h' ⟩ } +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_inclusion := isClosedEmbedding_inclusion + variable [TopologicalSemiring A] [ContinuousStar A] variable [TopologicalSpace B] [Semiring B] [Algebra R B] [StarRing B] @@ -99,7 +102,7 @@ theorem map_topologicalClosure_le [StarModule R B] [TopologicalSemiring B] [Cont image_closure_subset_closure_image hφ theorem topologicalClosure_map [StarModule R B] [TopologicalSemiring B] [ContinuousStar B] - (s : StarSubalgebra R A) (φ : A →⋆ₐ[R] B) (hφ : ClosedEmbedding φ) : + (s : StarSubalgebra R A) (φ : A →⋆ₐ[R] B) (hφ : IsClosedEmbedding φ) : (map φ s).topologicalClosure = map φ s.topologicalClosure := SetLike.coe_injective <| hφ.closure_image_eq _ @@ -205,8 +208,8 @@ theorem le_of_isClosed_of_mem {S : StarSubalgebra R A} (hS : IsClosed (S : Set A (hx : x ∈ S) : elementalStarAlgebra R x ≤ S := topologicalClosure_minimal (adjoin_le <| Set.singleton_subset_iff.2 hx) hS -/-- The coercion from an elemental algebra to the full algebra as a `ClosedEmbedding`. -/ -theorem closedEmbedding_coe (x : A) : ClosedEmbedding ((↑) : elementalStarAlgebra R x → A) := +/-- The coercion from an elemental algebra to the full algebra as a `IsClosedEmbedding`. -/ +theorem isClosedEmbedding_coe (x : A) : IsClosedEmbedding ((↑) : elementalStarAlgebra R x → A) := { induced := rfl inj := Subtype.coe_injective isClosed_range := by @@ -217,6 +220,9 @@ theorem closedEmbedding_coe (x : A) : ClosedEmbedding ((↑) : elementalStarAlge rintro ⟨y, rfl⟩ exact y.prop, fun hy => ⟨⟨y, hy⟩, rfl⟩⟩ } +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_coe := isClosedEmbedding_coe + @[elab_as_elim] theorem induction_on {x y : A} (hy : y ∈ elementalStarAlgebra R x) {P : (u : A) → u ∈ elementalStarAlgebra R x → Prop} diff --git a/Mathlib/Topology/Category/LightProfinite/Sequence.lean b/Mathlib/Topology/Category/LightProfinite/Sequence.lean index ae729f14b1979..adba9b5392add 100644 --- a/Mathlib/Topology/Category/LightProfinite/Sequence.lean +++ b/Mathlib/Topology/Category/LightProfinite/Sequence.lean @@ -29,8 +29,8 @@ noncomputable def natUnionInftyEmbedding : C(OnePoint ℕ, ℝ) where The continuous map from `ℕ∪{∞}` to `ℝ` sending `n` to `1/(n+1)` and `∞` to `0` is a closed embedding. -/ -lemma closedEmbedding_natUnionInftyEmbedding : ClosedEmbedding natUnionInftyEmbedding := by - refine closedEmbedding_of_continuous_injective_closed +lemma isClosedEmbedding_natUnionInftyEmbedding : IsClosedEmbedding natUnionInftyEmbedding := by + refine .of_continuous_injective_isClosedMap natUnionInftyEmbedding.continuous ?_ ?_ · rintro (_|n) (_|m) h · rfl @@ -45,7 +45,10 @@ lemma closedEmbedding_natUnionInftyEmbedding : ClosedEmbedding natUnionInftyEmbe rw [h] · exact fun _ hC => (hC.isCompact.image natUnionInftyEmbedding.continuous).isClosed -instance : MetrizableSpace (OnePoint ℕ) := closedEmbedding_natUnionInftyEmbedding.metrizableSpace +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_natUnionInftyEmbedding := isClosedEmbedding_natUnionInftyEmbedding + +instance : MetrizableSpace (OnePoint ℕ) := isClosedEmbedding_natUnionInftyEmbedding.metrizableSpace /-- The one point compactification of the natural numbers as a light profinite set. -/ abbrev NatUnionInfty : LightProfinite := of (OnePoint ℕ) diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index 3807d5656a53a..275dd04efa2cf 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -1779,7 +1779,7 @@ def GoodProducts.Basis (hC : IsClosed C) : end Induction -variable {S : Profinite} {ι : S → I → Bool} (hι : ClosedEmbedding ι) +variable {S : Profinite} {ι : S → I → Bool} (hι : IsClosedEmbedding ι) include hι /-- @@ -1801,8 +1801,8 @@ def Nobeling.ι : S → ({C : Set S // IsClopen C} → Bool) := fun s C => decid open scoped Classical in /-- The map `Nobeling.ι` is a closed embedding. -/ -theorem Nobeling.embedding : ClosedEmbedding (Nobeling.ι S) := by - apply Continuous.closedEmbedding +theorem Nobeling.embedding : IsClosedEmbedding (Nobeling.ι S) := by + apply Continuous.isClosedEmbedding · dsimp (config := { unfoldPartialApp := true }) [ι] refine continuous_pi ?_ intro C diff --git a/Mathlib/Topology/Category/Stonean/Limits.lean b/Mathlib/Topology/Category/Stonean/Limits.lean index ed7151753cb73..2b9c500e98a2e 100644 --- a/Mathlib/Topology/Category/Stonean/Limits.lean +++ b/Mathlib/Topology/Category/Stonean/Limits.lean @@ -33,7 +33,7 @@ lemma extremallyDisconnected_preimage : ExtremallyDisconnected (i ⁻¹' (Set.ra ⟨IsClosed.preimage i.continuous (isCompact_range f.continuous).isClosed, IsOpen.preimage i.continuous hi.isOpen_range⟩ rw [← (closure U).preimage_image_eq Subtype.coe_injective, - ← h.1.closedEmbedding_subtype_val.closure_image_eq U] + ← h.1.isClosedEmbedding_subtypeVal.closure_image_eq U] exact isOpen_induced (ExtremallyDisconnected.open_closure _ (h.2.isOpenEmbedding_subtypeVal.isOpenMap U hU)) diff --git a/Mathlib/Topology/Clopen.lean b/Mathlib/Topology/Clopen.lean index 511935a8e3b98..8d4aa07d73434 100644 --- a/Mathlib/Topology/Clopen.lean +++ b/Mathlib/Topology/Clopen.lean @@ -118,7 +118,7 @@ theorem isClopen_range_inr : IsClopen (range (Sum.inr : Y → X ⊕ Y)) := theorem isClopen_range_sigmaMk {X : ι → Type*} [∀ i, TopologicalSpace (X i)] {i : ι} : IsClopen (Set.range (@Sigma.mk ι X i)) := - ⟨closedEmbedding_sigmaMk.isClosed_range, isOpenEmbedding_sigmaMk.isOpen_range⟩ + ⟨isClosedEmbedding_sigmaMk.isClosed_range, isOpenEmbedding_sigmaMk.isOpen_range⟩ protected theorem QuotientMap.isClopen_preimage {f : X → Y} (hf : QuotientMap f) {s : Set Y} : IsClopen (f ⁻¹' s) ↔ IsClopen s := diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 0eb3000c7fbcf..ac4cbe1c1f71a 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -927,17 +927,23 @@ lemma Inducing.isCompact_preimage' {f : X → Y} (hf : Inducing f) {K : Set Y} (hf.isCompact_preimage_iff Kf).2 hK /-- The preimage of a compact set under a closed embedding is a compact set. -/ -theorem ClosedEmbedding.isCompact_preimage {f : X → Y} (hf : ClosedEmbedding f) +theorem IsClosedEmbedding.isCompact_preimage {f : X → Y} (hf : IsClosedEmbedding f) {K : Set Y} (hK : IsCompact K) : IsCompact (f ⁻¹' K) := hf.toInducing.isCompact_preimage (hf.isClosed_range) hK +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.isCompact_preimage := IsClosedEmbedding.isCompact_preimage + /-- A closed embedding is proper, ie, inverse images of compact sets are contained in compacts. -Moreover, the preimage of a compact set is compact, see `ClosedEmbedding.isCompact_preimage`. -/ -theorem ClosedEmbedding.tendsto_cocompact {f : X → Y} (hf : ClosedEmbedding f) : +Moreover, the preimage of a compact set is compact, see `IsClosedEmbedding.isCompact_preimage`. -/ +theorem IsClosedEmbedding.tendsto_cocompact {f : X → Y} (hf : IsClosedEmbedding f) : Tendsto f (Filter.cocompact X) (Filter.cocompact Y) := Filter.hasBasis_cocompact.tendsto_right_iff.mpr fun _K hK => (hf.isCompact_preimage hK).compl_mem_cocompact +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.tendsto_cocompact := IsClosedEmbedding.tendsto_cocompact + /-- Sets of subtype are compact iff the image under a coercion is. -/ theorem Subtype.isCompact_iff {p : X → Prop} {s : Set { x // p x }} : IsCompact s ↔ IsCompact ((↑) '' s : Set X) := @@ -956,14 +962,20 @@ theorem exists_nhds_ne_inf_principal_neBot (hs : IsCompact s) (hs' : s.Infinite) ∃ z ∈ s, (𝓝[≠] z ⊓ 𝓟 s).NeBot := hs'.exists_accPt_of_subset_isCompact hs Subset.rfl -protected theorem ClosedEmbedding.noncompactSpace [NoncompactSpace X] {f : X → Y} - (hf : ClosedEmbedding f) : NoncompactSpace Y := +protected theorem IsClosedEmbedding.noncompactSpace [NoncompactSpace X] {f : X → Y} + (hf : IsClosedEmbedding f) : NoncompactSpace Y := noncompactSpace_of_neBot hf.tendsto_cocompact.neBot -protected theorem ClosedEmbedding.compactSpace [h : CompactSpace Y] {f : X → Y} - (hf : ClosedEmbedding f) : CompactSpace X := +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.noncompactSpace := IsClosedEmbedding.noncompactSpace + +protected theorem IsClosedEmbedding.compactSpace [h : CompactSpace Y] {f : X → Y} + (hf : IsClosedEmbedding f) : CompactSpace X := ⟨by rw [hf.toInducing.isCompact_iff, image_univ]; exact hf.isClosed_range.isCompact⟩ +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.compactSpace := IsClosedEmbedding.compactSpace + theorem IsCompact.prod {t : Set Y} (hs : IsCompact s) (ht : IsCompact t) : IsCompact (s ×ˢ t) := by rw [isCompact_iff_ultrafilter_le_nhds'] at hs ht ⊢ @@ -981,7 +993,7 @@ instance (priority := 100) Finite.compactSpace [Finite X] : CompactSpace X where isCompact_univ := finite_univ.isCompact instance ULift.compactSpace [CompactSpace X] : CompactSpace (ULift.{v} X) := - ULift.closedEmbedding_down.compactSpace + ULift.isClosedEmbedding_down.compactSpace /-- The product of two compact spaces is compact. -/ instance [CompactSpace X] [CompactSpace Y] : CompactSpace (X × Y) := diff --git a/Mathlib/Topology/Compactness/Lindelof.lean b/Mathlib/Topology/Compactness/Lindelof.lean index 64c83c6ca1c47..9a5383d41086c 100644 --- a/Mathlib/Topology/Compactness/Lindelof.lean +++ b/Mathlib/Topology/Compactness/Lindelof.lean @@ -620,17 +620,24 @@ theorem Inducing.isLindelof_preimage {f : X → Y} (hf : Inducing f) (hf' : IsCl rwa [hf.isLindelof_iff, image_preimage_eq_inter_range] /-- The preimage of a Lindelöf set under a closed embedding is a Lindelöf set. -/ -theorem ClosedEmbedding.isLindelof_preimage {f : X → Y} (hf : ClosedEmbedding f) +theorem IsClosedEmbedding.isLindelof_preimage {f : X → Y} (hf : IsClosedEmbedding f) {K : Set Y} (hK : IsLindelof K) : IsLindelof (f ⁻¹' K) := hf.toInducing.isLindelof_preimage (hf.isClosed_range) hK +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.isLindelof_preimage := IsClosedEmbedding.isLindelof_preimage + /-- A closed embedding is proper, ie, inverse images of Lindelöf sets are contained in Lindelöf. -Moreover, the preimage of a Lindelöf set is Lindelöf, see `ClosedEmbedding.isLindelof_preimage`. -/ -theorem ClosedEmbedding.tendsto_coLindelof {f : X → Y} (hf : ClosedEmbedding f) : +Moreover, the preimage of a Lindelöf set is Lindelöf, see +`IsClosedEmbedding.isLindelof_preimage`. -/ +theorem IsClosedEmbedding.tendsto_coLindelof {f : X → Y} (hf : IsClosedEmbedding f) : Tendsto f (Filter.coLindelof X) (Filter.coLindelof Y) := hasBasis_coLindelof.tendsto_right_iff.mpr fun _K hK => (hf.isLindelof_preimage hK).compl_mem_coLindelof +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.tendsto_coLindelof := IsClosedEmbedding.tendsto_coLindelof + /-- Sets of subtype are Lindelöf iff the image under a coercion is. -/ theorem Subtype.isLindelof_iff {p : X → Prop} {s : Set { x // p x }} : IsLindelof s ↔ IsLindelof ((↑) '' s : Set X) := @@ -648,14 +655,20 @@ theorem IsLindelof.countable (hs : IsLindelof s) (hs' : DiscreteTopology s) : s. countable_coe_iff.mp (@countable_of_Lindelof_of_discrete _ _ (isLindelof_iff_LindelofSpace.mp hs) hs') -protected theorem ClosedEmbedding.nonLindelofSpace [NonLindelofSpace X] {f : X → Y} - (hf : ClosedEmbedding f) : NonLindelofSpace Y := +protected theorem IsClosedEmbedding.nonLindelofSpace [NonLindelofSpace X] {f : X → Y} + (hf : IsClosedEmbedding f) : NonLindelofSpace Y := nonLindelofSpace_of_neBot hf.tendsto_coLindelof.neBot -protected theorem ClosedEmbedding.LindelofSpace [h : LindelofSpace Y] {f : X → Y} - (hf : ClosedEmbedding f) : LindelofSpace X := +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.nonLindelofSpace := IsClosedEmbedding.nonLindelofSpace + +protected theorem IsClosedEmbedding.LindelofSpace [h : LindelofSpace Y] {f : X → Y} + (hf : IsClosedEmbedding f) : LindelofSpace X := ⟨by rw [hf.toInducing.isLindelof_iff, image_univ]; exact hf.isClosed_range.isLindelof⟩ +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.LindelofSpace := IsClosedEmbedding.LindelofSpace + /-- Countable topological spaces are Lindelof. -/ instance (priority := 100) Countable.LindelofSpace [Countable X] : LindelofSpace X where isLindelof_univ := countable_univ.isLindelof diff --git a/Mathlib/Topology/Compactness/LocallyCompact.lean b/Mathlib/Topology/Compactness/LocallyCompact.lean index 5bba2a34cb127..f479bd3f6e859 100644 --- a/Mathlib/Topology/Compactness/LocallyCompact.lean +++ b/Mathlib/Topology/Compactness/LocallyCompact.lean @@ -35,15 +35,18 @@ instance {ι : Type*} [Finite ι] {X : ι → Type*} [(i : ι) → TopologicalSp instance (priority := 100) [CompactSpace X] : WeaklyLocallyCompactSpace X where exists_compact_mem_nhds _ := ⟨univ, isCompact_univ, univ_mem⟩ -protected theorem ClosedEmbedding.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace Y] - {f : X → Y} (hf : ClosedEmbedding f) : WeaklyLocallyCompactSpace X where +protected theorem IsClosedEmbedding.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace Y] + {f : X → Y} (hf : IsClosedEmbedding f) : WeaklyLocallyCompactSpace X where exists_compact_mem_nhds x := let ⟨K, hK, hKx⟩ := exists_compact_mem_nhds (f x) ⟨f ⁻¹' K, hf.isCompact_preimage hK, hf.continuous.continuousAt hKx⟩ +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.weaklyLocallyCompactSpace := IsClosedEmbedding.weaklyLocallyCompactSpace + protected theorem IsClosed.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace X] {s : Set X} (hs : IsClosed s) : WeaklyLocallyCompactSpace s := - (closedEmbedding_subtype_val hs).weaklyLocallyCompactSpace + hs.isClosedEmbedding_subtypeVal.weaklyLocallyCompactSpace theorem IsOpenQuotientMap.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace X] {f : X → Y} (hf : IsOpenQuotientMap f) : WeaklyLocallyCompactSpace Y where @@ -195,10 +198,13 @@ theorem Inducing.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} (hf : rw [hf.isCompact_preimage_iff] exacts [hs.inter_right hZ, hUZ ▸ by gcongr] -protected theorem ClosedEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} - (hf : ClosedEmbedding f) : LocallyCompactSpace X := +protected theorem IsClosedEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} + (hf : IsClosedEmbedding f) : LocallyCompactSpace X := hf.toInducing.locallyCompactSpace hf.isClosed_range.isLocallyClosed +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.locallyCompactSpace := IsClosedEmbedding.locallyCompactSpace + protected theorem IsOpenEmbedding.locallyCompactSpace [LocallyCompactSpace Y] {f : X → Y} (hf : IsOpenEmbedding f) : LocallyCompactSpace X := hf.toInducing.locallyCompactSpace hf.isOpen_range.isLocallyClosed diff --git a/Mathlib/Topology/Compactness/Paracompact.lean b/Mathlib/Topology/Compactness/Paracompact.lean index ea4ab6733b402..638f75142b792 100644 --- a/Mathlib/Topology/Compactness/Paracompact.lean +++ b/Mathlib/Topology/Compactness/Paracompact.lean @@ -107,8 +107,8 @@ theorem precise_refinement_set [ParacompactSpace X] {s : Set X} (hs : IsClosed s · simp only [iUnion_option, ← compl_subset_iff_union] at vc exact Subset.trans (subset_compl_comm.1 <| vu Option.none) vc -theorem ClosedEmbedding.paracompactSpace [ParacompactSpace Y] {e : X → Y} (he : ClosedEmbedding e) : - ParacompactSpace X where +theorem IsClosedEmbedding.paracompactSpace [ParacompactSpace Y] {e : X → Y} + (he : IsClosedEmbedding e) : ParacompactSpace X where locallyFinite_refinement α s ho hu := by choose U hUo hU using fun a ↦ he.isOpen_iff.1 (ho a) simp only [← hU] at hu ⊢ @@ -119,8 +119,11 @@ theorem ClosedEmbedding.paracompactSpace [ParacompactSpace Y] {e : X → Y} (he hVf.preimage_continuous he.continuous, fun a ↦ ⟨a, preimage_mono (hVU a)⟩⟩ simpa only [range_subset_iff, mem_iUnion, iUnion_eq_univ_iff] using heV +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.paracompactSpace := IsClosedEmbedding.paracompactSpace + theorem Homeomorph.paracompactSpace_iff (e : X ≃ₜ Y) : ParacompactSpace X ↔ ParacompactSpace Y := - ⟨fun _ ↦ e.symm.closedEmbedding.paracompactSpace, fun _ ↦ e.closedEmbedding.paracompactSpace⟩ + ⟨fun _ ↦ e.symm.isClosedEmbedding.paracompactSpace, fun _ ↦ e.isClosedEmbedding.paracompactSpace⟩ /-- The product of a compact space and a paracompact space is a paracompact space. The formalization is based on https://dantopology.wordpress.com/2009/10/24/compact-x-paracompact-is-paracompact/ diff --git a/Mathlib/Topology/Compactness/SigmaCompact.lean b/Mathlib/Topology/Compactness/SigmaCompact.lean index ec5a6b0e2279f..9718f4ea99c27 100644 --- a/Mathlib/Topology/Compactness/SigmaCompact.lean +++ b/Mathlib/Topology/Compactness/SigmaCompact.lean @@ -249,17 +249,20 @@ instance [Countable ι] {X : ι → Type*} [∀ i, TopologicalSpace (X i)] refine ⟨max k n, k, le_max_left _ _, mem_image_of_mem _ ?_⟩ exact compactCovering_subset _ (le_max_right _ _) hn -protected theorem ClosedEmbedding.sigmaCompactSpace {e : Y → X} (he : ClosedEmbedding e) : +protected theorem IsClosedEmbedding.sigmaCompactSpace {e : Y → X} (he : IsClosedEmbedding e) : SigmaCompactSpace Y := ⟨⟨fun n => e ⁻¹' compactCovering X n, fun _ => he.isCompact_preimage (isCompact_compactCovering _ _), by rw [← preimage_iUnion, iUnion_compactCovering, preimage_univ]⟩⟩ +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.sigmaCompactSpace := IsClosedEmbedding.sigmaCompactSpace + theorem IsClosed.sigmaCompactSpace {s : Set X} (hs : IsClosed s) : SigmaCompactSpace s := - (closedEmbedding_subtype_val hs).sigmaCompactSpace + hs.isClosedEmbedding_subtypeVal.sigmaCompactSpace instance [SigmaCompactSpace Y] : SigmaCompactSpace (ULift.{u} Y) := - ULift.closedEmbedding_down.sigmaCompactSpace + ULift.isClosedEmbedding_down.sigmaCompactSpace /-- If `X` is a `σ`-compact space, then a locally finite family of nonempty sets of `X` can have only countably many elements, `Set.Countable` version. -/ diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 85297dd644e16..4a48f65c105f5 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -937,12 +937,18 @@ theorem isClosed_range_inr : IsClosed (range (inr : Y → X ⊕ Y)) := by rw [← isOpen_compl_iff, compl_range_inr] exact isOpen_range_inl -theorem closedEmbedding_inl : ClosedEmbedding (inl : X → X ⊕ Y) := +theorem isClosedEmbedding_inl : IsClosedEmbedding (inl : X → X ⊕ Y) := ⟨embedding_inl, isClosed_range_inl⟩ -theorem closedEmbedding_inr : ClosedEmbedding (inr : Y → X ⊕ Y) := +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_inl := isClosedEmbedding_inl + +theorem isClosedEmbedding_inr : IsClosedEmbedding (inr : Y → X ⊕ Y) := ⟨embedding_inr, isClosed_range_inr⟩ +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_inr := isClosedEmbedding_inr + theorem nhds_inl (x : X) : 𝓝 (inl x : X ⊕ Y) = map inl (𝓝 x) := (isOpenEmbedding_inl.map_nhds_eq _).symm @@ -981,7 +987,7 @@ theorem isClosedMap_sum {f : X ⊕ Y → Z} : IsClosedMap f ↔ (IsClosedMap fun a => f (.inl a)) ∧ IsClosedMap fun b => f (.inr b) := by constructor · intro h - exact ⟨h.comp closedEmbedding_inl.isClosedMap, h.comp closedEmbedding_inr.isClosedMap⟩ + exact ⟨h.comp isClosedEmbedding_inl.isClosedMap, h.comp isClosedEmbedding_inr.isClosedMap⟩ · rintro h Z hZ rw [isClosed_sum_iff] at hZ convert (h.1 _ hZ.1).union (h.2 _ hZ.2) @@ -1003,10 +1009,13 @@ theorem Inducing.of_codRestrict {f : X → Y} {t : Set Y} (ht : ∀ x, f x ∈ t theorem embedding_subtype_val : Embedding ((↑) : Subtype p → X) := ⟨inducing_subtype_val, Subtype.coe_injective⟩ -theorem closedEmbedding_subtype_val (h : IsClosed { a | p a }) : - ClosedEmbedding ((↑) : Subtype p → X) := +theorem IsClosedEmbedding.subtypeVal (h : IsClosed {a | p a}) : + IsClosedEmbedding ((↑) : Subtype p → X) := ⟨embedding_subtype_val, by rwa [Subtype.range_coe_subtype]⟩ +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_subtype_val := IsClosedEmbedding.subtypeVal + @[continuity, fun_prop] theorem continuous_subtype_val : Continuous (@Subtype.val X p) := continuous_induced_dom @@ -1029,13 +1038,15 @@ theorem IsOpenMap.restrict {f : X → Y} (hf : IsOpenMap f) {s : Set X} (hs : Is IsOpenMap (s.restrict f) := hf.comp hs.isOpenMap_subtype_val -nonrec theorem IsClosed.closedEmbedding_subtype_val {s : Set X} (hs : IsClosed s) : - ClosedEmbedding ((↑) : s → X) := - closedEmbedding_subtype_val hs +lemma IsClosed.isClosedEmbedding_subtypeVal {s : Set X} (hs : IsClosed s) : + IsClosedEmbedding ((↑) : s → X) := .subtypeVal hs + +@[deprecated (since := "2024-10-20")] +alias IsClosed.closedEmbedding_subtype_val := IsClosed.isClosedEmbedding_subtypeVal theorem IsClosed.isClosedMap_subtype_val {s : Set X} (hs : IsClosed s) : IsClosedMap ((↑) : s → X) := - hs.closedEmbedding_subtype_val.isClosedMap + hs.isClosedEmbedding_subtypeVal.isClosedMap @[continuity, fun_prop] theorem Continuous.subtype_mk {f : Y → X} (h : Continuous f) (hp : ∀ x, p (f x)) : @@ -1547,12 +1558,15 @@ theorem isOpenEmbedding_sigmaMk {i : ι} : IsOpenEmbedding (@Sigma.mk ι σ i) : @[deprecated (since := "2024-10-18")] alias openEmbedding_sigmaMk := isOpenEmbedding_sigmaMk -theorem closedEmbedding_sigmaMk {i : ι} : ClosedEmbedding (@Sigma.mk ι σ i) := - closedEmbedding_of_continuous_injective_closed continuous_sigmaMk sigma_mk_injective +theorem isClosedEmbedding_sigmaMk {i : ι} : IsClosedEmbedding (@Sigma.mk ι σ i) := + .of_continuous_injective_isClosedMap continuous_sigmaMk sigma_mk_injective isClosedMap_sigmaMk +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_sigmaMk := isClosedEmbedding_sigmaMk + theorem embedding_sigmaMk {i : ι} : Embedding (@Sigma.mk ι σ i) := - closedEmbedding_sigmaMk.1 + isClosedEmbedding_sigmaMk.1 theorem Sigma.nhds_mk (i : ι) (x : σ i) : 𝓝 (⟨i, x⟩ : Sigma σ) = Filter.map (Sigma.mk i) (𝓝 x) := (isOpenEmbedding_sigmaMk.map_nhds_eq x).symm @@ -1657,10 +1671,13 @@ theorem continuous_uLift_up [TopologicalSpace X] : Continuous (ULift.up : X → theorem embedding_uLift_down [TopologicalSpace X] : Embedding (ULift.down : ULift.{v, u} X → X) := ⟨⟨rfl⟩, ULift.down_injective⟩ -theorem ULift.closedEmbedding_down [TopologicalSpace X] : - ClosedEmbedding (ULift.down : ULift.{v, u} X → X) := +theorem ULift.isClosedEmbedding_down [TopologicalSpace X] : + IsClosedEmbedding (ULift.down : ULift.{v, u} X → X) := ⟨embedding_uLift_down, by simp only [ULift.down_surjective.range_eq, isClosed_univ]⟩ +@[deprecated (since := "2024-10-20")] +alias ULift.closedEmbedding_down := ULift.isClosedEmbedding_down + instance [TopologicalSpace X] [DiscreteTopology X] : DiscreteTopology (ULift X) := embedding_uLift_down.discreteTopology diff --git a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean index 6d699f5ba3042..6bb261df3d76d 100644 --- a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean @@ -98,13 +98,16 @@ instance instContinuousEvalConst : ContinuousEvalConst C(X, R)₀ X R := instance instContinuousEval [LocallyCompactPair X R] : ContinuousEval C(X, R)₀ X R := .of_continuous_forget embedding_toContinuousMap.continuous -lemma closedEmbedding_toContinuousMap [T1Space R] : - ClosedEmbedding ((↑) : C(X, R)₀ → C(X, R)) where +lemma isClosedEmbedding_toContinuousMap [T1Space R] : + IsClosedEmbedding ((↑) : C(X, R)₀ → C(X, R)) where toEmbedding := embedding_toContinuousMap isClosed_range := by rw [range_toContinuousMap] exact isClosed_singleton.preimage <| continuous_eval_const 0 +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_toContinuousMap := isClosedEmbedding_toContinuousMap + @[fun_prop] lemma continuous_comp_left {X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] [Zero X] [Zero Y] [Zero Z] (f : C(X, Y)₀) : @@ -288,7 +291,7 @@ alias uniformEmbedding_toContinuousMap := isUniformEmbedding_toContinuousMap instance [T1Space R] [CompleteSpace C(X, R)] : CompleteSpace C(X, R)₀ := completeSpace_iff_isComplete_range isUniformEmbedding_toContinuousMap.isUniformInducing - |>.mpr closedEmbedding_toContinuousMap.isClosed_range.isComplete + |>.mpr isClosedEmbedding_toContinuousMap.isClosed_range.isComplete lemma isUniformEmbedding_comp {Y : Type*} [UniformSpace Y] [Zero Y] (g : C(Y, R)₀) (hg : IsUniformEmbedding g) : IsUniformEmbedding (g.comp · : C(X, Y)₀ → C(X, R)₀) := diff --git a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean index ecfa6bfa4e8f4..7d274f86f8aab 100644 --- a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean @@ -588,8 +588,8 @@ lemma ContinuousMapZero.adjoin_id_dense {s : Set 𝕜} [Zero s] (h0 : ((0 : s) : [CompactSpace s] : Dense (adjoin 𝕜 {(.id h0 : C(s, 𝕜)₀)} : Set C(s, 𝕜)₀) := by have h0' : 0 ∈ s := h0 ▸ (0 : s).property rw [dense_iff_closure_eq, - ← closedEmbedding_toContinuousMap.injective.preimage_image (closure _), - ← closedEmbedding_toContinuousMap.closure_image_eq, ← coe_toContinuousMapHom, + ← isClosedEmbedding_toContinuousMap.injective.preimage_image (closure _), + ← isClosedEmbedding_toContinuousMap.closure_image_eq, ← coe_toContinuousMapHom, ← NonUnitalStarSubalgebra.coe_map, NonUnitalStarAlgHom.map_adjoin_singleton, toContinuousMapHom_apply, toContinuousMap_id h0, ← ContinuousMap.ker_evalStarAlgHom_eq_closure_adjoin_id s h0'] diff --git a/Mathlib/Topology/Defs/Induced.lean b/Mathlib/Topology/Defs/Induced.lean index 0aaa1d46b8de9..0e1df6198a3bd 100644 --- a/Mathlib/Topology/Defs/Induced.lean +++ b/Mathlib/Topology/Defs/Induced.lean @@ -32,7 +32,7 @@ as well as topology inducing maps, topological embeddings, and quotient maps. if it is an embedding and its range is open. An open embedding is an open map. -* `ClosedEmbedding`: a map `f : X → Y` is an *open embedding*, +* `IsClosedEmbedding`: a map `f : X → Y` is an *open embedding*, if it is an embedding and its range is open. An open embedding is an open map. @@ -121,10 +121,13 @@ alias OpenEmbedding := IsOpenEmbedding /-- A closed embedding is an embedding with closed image. -/ @[mk_iff] -structure ClosedEmbedding (f : X → Y) extends Embedding f : Prop where +structure IsClosedEmbedding (f : X → Y) extends Embedding f : Prop where /-- The range of a closed embedding is a closed set. -/ isClosed_range : IsClosed <| range f +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding := IsClosedEmbedding + /-- A function between topological spaces is a quotient map if it is surjective, and for all `s : Set Y`, `s` is open iff its preimage is an open set. -/ @[mk_iff quotientMap_iff'] diff --git a/Mathlib/Topology/DiscreteSubset.lean b/Mathlib/Topology/DiscreteSubset.lean index 5238fbf4e8e1e..825509afa0b34 100644 --- a/Mathlib/Topology/DiscreteSubset.lean +++ b/Mathlib/Topology/DiscreteSubset.lean @@ -69,7 +69,7 @@ lemma tendsto_cofinite_cocompact_of_discrete [DiscreteTopology X] lemma IsClosed.tendsto_coe_cofinite_of_discreteTopology {s : Set X} (hs : IsClosed s) (_hs' : DiscreteTopology s) : Tendsto ((↑) : s → X) cofinite (cocompact _) := - tendsto_cofinite_cocompact_of_discrete hs.closedEmbedding_subtype_val.tendsto_cocompact + tendsto_cofinite_cocompact_of_discrete hs.isClosedEmbedding_subtypeVal.tendsto_cocompact lemma IsClosed.tendsto_coe_cofinite_iff [T1Space X] [WeaklyLocallyCompactSpace X] {s : Set X} (hs : IsClosed s) : diff --git a/Mathlib/Topology/FiberBundle/Basic.lean b/Mathlib/Topology/FiberBundle/Basic.lean index 39a1e8d2edc2b..3fb66e7803f27 100644 --- a/Mathlib/Topology/FiberBundle/Basic.lean +++ b/Mathlib/Topology/FiberBundle/Basic.lean @@ -261,12 +261,15 @@ theorem continuous_totalSpaceMk (x : B) : Continuous (@TotalSpace.mk B F E x) := theorem totalSpaceMk_embedding (x : B) : Embedding (@TotalSpace.mk B F E x) := ⟨totalSpaceMk_inducing F E x, TotalSpace.mk_injective x⟩ -theorem totalSpaceMk_closedEmbedding [T1Space B] (x : B) : - ClosedEmbedding (@TotalSpace.mk B F E x) := +theorem totalSpaceMk_isClosedEmbedding [T1Space B] (x : B) : + IsClosedEmbedding (@TotalSpace.mk B F E x) := ⟨totalSpaceMk_embedding F E x, by rw [TotalSpace.range_mk] exact isClosed_singleton.preimage <| continuous_proj F E⟩ +@[deprecated (since := "2024-10-20")] +alias totalSpaceMk_closedEmbedding := totalSpaceMk_isClosedEmbedding + variable {E F} @[simp, mfld_simps] diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index ce885a2ca891d..d1f3d9a2db354 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -338,14 +338,17 @@ theorem isOpenEmbedding (h : X ≃ₜ Y) : IsOpenEmbedding h := @[deprecated (since := "2024-10-18")] alias openEmbedding := isOpenEmbedding -protected theorem closedEmbedding (h : X ≃ₜ Y) : ClosedEmbedding h := - closedEmbedding_of_embedding_closed h.embedding h.isClosedMap +theorem isClosedEmbedding (h : X ≃ₜ Y) : IsClosedEmbedding h := + .of_embedding_closed h.embedding h.isClosedMap + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding := isClosedEmbedding protected theorem normalSpace [NormalSpace X] (h : X ≃ₜ Y) : NormalSpace Y := - h.symm.closedEmbedding.normalSpace + h.symm.isClosedEmbedding.normalSpace protected theorem t4Space [T4Space X] (h : X ≃ₜ Y) : T4Space Y := - h.symm.closedEmbedding.t4Space + h.symm.isClosedEmbedding.t4Space theorem preimage_closure (h : X ≃ₜ Y) (s : Set Y) : h ⁻¹' closure s = closure (h ⁻¹' s) := h.isOpenMap.preimage_closure_eq_closure_preimage h.continuous _ @@ -368,7 +371,7 @@ theorem image_frontier (h : X ≃ₜ Y) (s : Set X) : h '' frontier s = frontier @[to_additive] theorem _root_.HasCompactMulSupport.comp_homeomorph {M} [One M] {f : Y → M} (hf : HasCompactMulSupport f) (φ : X ≃ₜ Y) : HasCompactMulSupport (f ∘ φ) := - hf.comp_closedEmbedding φ.closedEmbedding + hf.comp_isClosedEmbedding φ.isClosedEmbedding @[simp] theorem map_nhds_eq (h : X ≃ₜ Y) (x : X) : map h (𝓝 x) = 𝓝 (h x) := @@ -405,7 +408,7 @@ the domain is a locally compact space. -/ theorem locallyCompactSpace_iff (h : X ≃ₜ Y) : LocallyCompactSpace X ↔ LocallyCompactSpace Y := by exact ⟨fun _ => h.symm.isOpenEmbedding.locallyCompactSpace, - fun _ => h.closedEmbedding.locallyCompactSpace⟩ + fun _ => h.isClosedEmbedding.locallyCompactSpace⟩ /-- If a bijective map `e : X ≃ Y` is continuous and open, then it is a homeomorphism. -/ @[simps toEquiv] @@ -912,9 +915,10 @@ protected lemma inducing : Inducing f := (hf.homeomorph f).inducing protected lemma quotientMap : QuotientMap f := (hf.homeomorph f).quotientMap protected lemma embedding : Embedding f := (hf.homeomorph f).embedding lemma isOpenEmbedding : IsOpenEmbedding f := (hf.homeomorph f).isOpenEmbedding -protected lemma closedEmbedding : ClosedEmbedding f := (hf.homeomorph f).closedEmbedding +lemma isClosedEmbedding : IsClosedEmbedding f := (hf.homeomorph f).isClosedEmbedding lemma isDenseEmbedding : IsDenseEmbedding f := (hf.homeomorph f).isDenseEmbedding +@[deprecated (since := "2024-10-20")] alias closedEmbedding := isClosedEmbedding @[deprecated (since := "2024-10-18")] alias openEmbedding := isOpenEmbedding diff --git a/Mathlib/Topology/Instances/CantorSet.lean b/Mathlib/Topology/Instances/CantorSet.lean index 22167d7ac0fbb..69820c6161f0a 100644 --- a/Mathlib/Topology/Instances/CantorSet.lean +++ b/Mathlib/Topology/Instances/CantorSet.lean @@ -86,8 +86,8 @@ lemma isClosed_preCantorSet (n : ℕ) : IsClosed (preCantorSet n) := by | zero => exact isClosed_Icc | succ n ih => refine IsClosed.union ?_ ?_ - · simpa [f, div_eq_inv_mul] using f.closedEmbedding.closed_iff_image_closed.mp ih - · simpa [g, f, div_eq_inv_mul] using g.closedEmbedding.closed_iff_image_closed.mp ih + · simpa [f, div_eq_inv_mul] using f.isClosedEmbedding.closed_iff_image_closed.mp ih + · simpa [g, f, div_eq_inv_mul] using g.isClosedEmbedding.closed_iff_image_closed.mp ih /-- The ternary Cantor set is closed. -/ lemma isClosed_cantorSet : IsClosed cantorSet := diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index 1b66243074048..89c410a5d2b73 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -96,9 +96,12 @@ theorem embedding_coe_ennreal : Embedding ((↑) : ℝ≥0∞ → EReal) := coe_ennreal_strictMono.embedding_of_ordConnected <| by rw [range_coe_ennreal]; exact ordConnected_Ici -theorem closedEmbedding_coe_ennreal : ClosedEmbedding ((↑) : ℝ≥0∞ → EReal) := +theorem isClosedEmbedding_coe_ennreal : IsClosedEmbedding ((↑) : ℝ≥0∞ → EReal) := ⟨embedding_coe_ennreal, by rw [range_coe_ennreal]; exact isClosed_Ici⟩ +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_coe_ennreal := isClosedEmbedding_coe_ennreal + @[norm_cast] theorem tendsto_coe_ennreal {α : Type*} {f : Filter α} {m : α → ℝ≥0∞} {a : ℝ≥0∞} : Tendsto (fun a => (m a : EReal)) f (𝓝 ↑a) ↔ Tendsto m f (𝓝 a) := diff --git a/Mathlib/Topology/Instances/Int.lean b/Mathlib/Topology/Instances/Int.lean index 5ee25d7b929f2..536b95e356e15 100644 --- a/Mathlib/Topology/Instances/Int.lean +++ b/Mathlib/Topology/Instances/Int.lean @@ -45,8 +45,11 @@ theorem isUniformEmbedding_coe_real : IsUniformEmbedding ((↑) : ℤ → ℝ) : @[deprecated (since := "2024-10-01")] alias uniformEmbedding_coe_real := isUniformEmbedding_coe_real -theorem closedEmbedding_coe_real : ClosedEmbedding ((↑) : ℤ → ℝ) := - closedEmbedding_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist +theorem isClosedEmbedding_coe_real : IsClosedEmbedding ((↑) : ℤ → ℝ) := + isClosedEmbedding_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_coe_real := isClosedEmbedding_coe_real instance : MetricSpace ℤ := Int.isUniformEmbedding_coe_real.comapMetricSpace _ diff --git a/Mathlib/Topology/Instances/Irrational.lean b/Mathlib/Topology/Instances/Irrational.lean index fc7c1f606e691..7d4b39145e504 100644 --- a/Mathlib/Topology/Instances/Irrational.lean +++ b/Mathlib/Topology/Instances/Irrational.lean @@ -71,7 +71,7 @@ instance : DenselyOrdered { x // Irrational x } := theorem eventually_forall_le_dist_cast_div (hx : Irrational x) (n : ℕ) : ∀ᶠ ε : ℝ in 𝓝 0, ∀ m : ℤ, ε ≤ dist x (m / n) := by have A : IsClosed (range (fun m => (n : ℝ)⁻¹ * m : ℤ → ℝ)) := - ((isClosedMap_smul₀ (n⁻¹ : ℝ)).comp Int.closedEmbedding_coe_real.isClosedMap).isClosed_range + ((isClosedMap_smul₀ (n⁻¹ : ℝ)).comp Int.isClosedEmbedding_coe_real.isClosedMap).isClosed_range have B : x ∉ range (fun m => (n : ℝ)⁻¹ * m : ℤ → ℝ) := by rintro ⟨m, rfl⟩ simp at hx diff --git a/Mathlib/Topology/Instances/NNReal.lean b/Mathlib/Topology/Instances/NNReal.lean index 6f0a140493505..16383d3381a75 100644 --- a/Mathlib/Topology/Instances/NNReal.lean +++ b/Mathlib/Topology/Instances/NNReal.lean @@ -289,7 +289,7 @@ end Monotone instance instProperSpace : ProperSpace ℝ≥0 where isCompact_closedBall x r := by - have emb : ClosedEmbedding ((↑) : ℝ≥0 → ℝ) := Isometry.closedEmbedding fun _ ↦ congrFun rfl + have emb : IsClosedEmbedding ((↑) : ℝ≥0 → ℝ) := Isometry.isClosedEmbedding fun _ ↦ congrFun rfl exact emb.isCompact_preimage (K := Metric.closedBall x r) (isCompact_closedBall _ _) end NNReal diff --git a/Mathlib/Topology/Instances/Nat.lean b/Mathlib/Topology/Instances/Nat.lean index e90ce1cbe65b7..c5719cb578fc2 100644 --- a/Mathlib/Topology/Instances/Nat.lean +++ b/Mathlib/Topology/Instances/Nat.lean @@ -37,8 +37,11 @@ theorem isUniformEmbedding_coe_real : IsUniformEmbedding ((↑) : ℕ → ℝ) : @[deprecated (since := "2024-10-01")] alias uniformEmbedding_coe_real := isUniformEmbedding_coe_real -theorem closedEmbedding_coe_real : ClosedEmbedding ((↑) : ℕ → ℝ) := - closedEmbedding_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist +theorem isClosedEmbedding_coe_real : IsClosedEmbedding ((↑) : ℕ → ℝ) := + isClosedEmbedding_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_coe_real := isClosedEmbedding_coe_real instance : MetricSpace ℕ := Nat.isUniformEmbedding_coe_real.comapMetricSpace _ diff --git a/Mathlib/Topology/Instances/Rat.lean b/Mathlib/Topology/Instances/Rat.lean index 715361b5de0d9..8847d81d5d088 100644 --- a/Mathlib/Topology/Instances/Rat.lean +++ b/Mathlib/Topology/Instances/Rat.lean @@ -60,8 +60,11 @@ theorem Nat.isUniformEmbedding_coe_rat : IsUniformEmbedding ((↑) : ℕ → ℚ @[deprecated (since := "2024-10-01")] alias Nat.uniformEmbedding_coe_rat := Nat.isUniformEmbedding_coe_rat -theorem Nat.closedEmbedding_coe_rat : ClosedEmbedding ((↑) : ℕ → ℚ) := - closedEmbedding_of_pairwise_le_dist zero_lt_one <| by simpa using Nat.pairwise_one_le_dist +theorem Nat.isClosedEmbedding_coe_rat : IsClosedEmbedding ((↑) : ℕ → ℚ) := + isClosedEmbedding_of_pairwise_le_dist zero_lt_one <| by simpa using Nat.pairwise_one_le_dist + +@[deprecated (since := "2024-10-20")] +alias Nat.closedEmbedding_coe_rat := Nat.isClosedEmbedding_coe_rat @[norm_cast, simp] theorem Int.dist_cast_rat (x y : ℤ) : dist (x : ℚ) y = dist x y := by @@ -73,12 +76,15 @@ theorem Int.isUniformEmbedding_coe_rat : IsUniformEmbedding ((↑) : ℤ → ℚ @[deprecated (since := "2024-10-01")] alias Int.uniformEmbedding_coe_rat := Int.isUniformEmbedding_coe_rat -theorem Int.closedEmbedding_coe_rat : ClosedEmbedding ((↑) : ℤ → ℚ) := - closedEmbedding_of_pairwise_le_dist zero_lt_one <| by simpa using Int.pairwise_one_le_dist +theorem Int.isClosedEmbedding_coe_rat : IsClosedEmbedding ((↑) : ℤ → ℚ) := + isClosedEmbedding_of_pairwise_le_dist zero_lt_one <| by simpa using Int.pairwise_one_le_dist + +@[deprecated (since := "2024-10-20")] +alias Int.closedEmbedding_coe_rat := Int.isClosedEmbedding_coe_rat namespace Rat -instance : NoncompactSpace ℚ := Int.closedEmbedding_coe_rat.noncompactSpace +instance : NoncompactSpace ℚ := Int.isClosedEmbedding_coe_rat.noncompactSpace theorem uniformContinuous_add : UniformContinuous fun p : ℚ × ℚ => p.1 + p.2 := Rat.isUniformEmbedding_coe_real.isUniformInducing.uniformContinuous_iff.2 <| by diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index 927f93d6bf83a..c41f014575333 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -28,7 +28,7 @@ universe u v w variable {α : Type u} {β : Type v} {γ : Type w} -instance : NoncompactSpace ℝ := Int.closedEmbedding_coe_real.noncompactSpace +instance : NoncompactSpace ℝ := Int.isClosedEmbedding_coe_real.noncompactSpace theorem Real.uniformContinuous_add : UniformContinuous fun p : ℝ × ℝ => p.1 + p.2 := Metric.uniformContinuous_iff.2 fun _ε ε0 => diff --git a/Mathlib/Topology/KrullDimension.lean b/Mathlib/Topology/KrullDimension.lean index b5eee2e982ae4..9a25326df9e83 100644 --- a/Mathlib/Topology/KrullDimension.lean +++ b/Mathlib/Topology/KrullDimension.lean @@ -40,7 +40,7 @@ def IrreducibleCloseds.map {f : X → Y} (hf1 : Continuous f) (hf2 : IsClosedMap /-- Taking images under a closed embedding is strictly monotone on the preorder of irreducible closeds. -/ -lemma IrreducibleCloseds.map_strictMono {f : X → Y} (hf : ClosedEmbedding f) : +lemma IrreducibleCloseds.map_strictMono {f : X → Y} (hf : IsClosedEmbedding f) : StrictMono (IrreducibleCloseds.map hf.continuous hf.isClosedMap) := fun ⦃_ _⦄ UltV ↦ hf.inj.image_strictMono UltV @@ -48,16 +48,19 @@ lemma IrreducibleCloseds.map_strictMono {f : X → Y} (hf : ClosedEmbedding f) : If `f : X → Y` is a closed embedding, then the Krull dimension of `X` is less than or equal to the Krull dimension of `Y`. -/ -theorem ClosedEmbedding.topologicalKrullDim_le (f : X → Y) (hf : ClosedEmbedding f) : +theorem IsClosedEmbedding.topologicalKrullDim_le (f : X → Y) (hf : IsClosedEmbedding f) : topologicalKrullDim X ≤ topologicalKrullDim Y := krullDim_le_of_strictMono _ (IrreducibleCloseds.map_strictMono hf) +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.topologicalKrullDim_le := IsClosedEmbedding.topologicalKrullDim_le + /-- The topological Krull dimension is invariant under homeomorphisms -/ theorem IsHomeomorph.topologicalKrullDim_eq (f : X → Y) (h : IsHomeomorph f) : topologicalKrullDim X = topologicalKrullDim Y := have fwd : topologicalKrullDim X ≤ topologicalKrullDim Y := - ClosedEmbedding.topologicalKrullDim_le f h.closedEmbedding + IsClosedEmbedding.topologicalKrullDim_le f h.isClosedEmbedding have bwd : topologicalKrullDim Y ≤ topologicalKrullDim X := - ClosedEmbedding.topologicalKrullDim_le (h.homeomorph f).symm - (h.homeomorph f).symm.closedEmbedding + IsClosedEmbedding.topologicalKrullDim_le (h.homeomorph f).symm + (h.homeomorph f).symm.isClosedEmbedding le_antisymm fwd bwd diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 4e1a6f06f340a..c1e9f1b8153f2 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -13,7 +13,7 @@ We show that the following properties of continuous maps are local at the target - `Inducing` - `Embedding` - `IsOpenEmbedding` -- `ClosedEmbedding` +- `IsClosedEmbedding` -/ @@ -53,12 +53,18 @@ alias IsOpenEmbedding.restrictPreimage := Set.restrictPreimage_isOpenEmbedding @[deprecated (since := "2024-10-18")] alias OpenEmbedding.restrictPreimage := IsOpenEmbedding.restrictPreimage -theorem Set.restrictPreimage_closedEmbedding (s : Set β) (h : ClosedEmbedding f) : - ClosedEmbedding (s.restrictPreimage f) := +theorem Set.restrictPreimage_isClosedEmbedding (s : Set β) (h : IsClosedEmbedding f) : + IsClosedEmbedding (s.restrictPreimage f) := ⟨h.1.restrictPreimage s, (s.range_restrictPreimage f).symm ▸ inducing_subtype_val.isClosed_preimage _ h.isClosed_range⟩ -alias ClosedEmbedding.restrictPreimage := Set.restrictPreimage_closedEmbedding +@[deprecated (since := "2024-10-20")] +alias Set.restrictPreimage_closedEmbedding := Set.restrictPreimage_isClosedEmbedding + +alias IsClosedEmbedding.restrictPreimage := Set.restrictPreimage_isClosedEmbedding + +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.restrictPreimage := IsClosedEmbedding.restrictPreimage theorem IsClosedMap.restrictPreimage (H : IsClosedMap f) (s : Set β) : IsClosedMap (s.restrictPreimage f) := by @@ -184,11 +190,15 @@ theorem isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top (h : Continuous f) : alias openEmbedding_iff_openEmbedding_of_iSup_eq_top := isOpenEmbedding_iff_isOpenEmbedding_of_iSup_eq_top -theorem closedEmbedding_iff_closedEmbedding_of_iSup_eq_top (h : Continuous f) : - ClosedEmbedding f ↔ ∀ i, ClosedEmbedding ((U i).1.restrictPreimage f) := by - simp_rw [closedEmbedding_iff] +theorem isClosedEmbedding_iff_isClosedEmbedding_of_iSup_eq_top (h : Continuous f) : + IsClosedEmbedding f ↔ ∀ i, IsClosedEmbedding ((U i).1.restrictPreimage f) := by + simp_rw [isClosedEmbedding_iff] rw [forall_and] apply and_congr · apply embedding_iff_embedding_of_iSup_eq_top <;> assumption · simp_rw [Set.range_restrictPreimage] apply isClosed_iff_coe_preimage_of_iSup_eq_top hU + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_iff_closedEmbedding_of_iSup_eq_top := + isClosedEmbedding_iff_isClosedEmbedding_of_iSup_eq_top diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index c4aaea387693f..7933f9ee54473 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -22,7 +22,7 @@ This file introduces the following properties of a map `f : X → Y` between top a subspace of `Y`. * `IsOpenEmbedding f` means `f` is an embedding with open image, so it identifies `X` with an open subspace of `Y`. Equivalently, `f` is an embedding and an open map. -* `ClosedEmbedding f` similarly means `f` is an embedding with closed image, so it identifies +* `IsClosedEmbedding f` similarly means `f` is an embedding with closed image, so it identifies `X` with a closed subspace of `Y`. Equivalently, `f` is an embedding and a closed map. * `QuotientMap f` is the dual condition to `Embedding f`: `f` is surjective and the topology @@ -603,59 +603,69 @@ end IsOpenEmbedding end IsOpenEmbedding -section ClosedEmbedding +section IsClosedEmbedding variable [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] -namespace ClosedEmbedding +namespace IsClosedEmbedding -theorem tendsto_nhds_iff {g : ι → X} {l : Filter ι} {x : X} (hf : ClosedEmbedding f) : +theorem tendsto_nhds_iff {g : ι → X} {l : Filter ι} {x : X} (hf : IsClosedEmbedding f) : Tendsto g l (𝓝 x) ↔ Tendsto (f ∘ g) l (𝓝 (f x)) := hf.toEmbedding.tendsto_nhds_iff -theorem continuous (hf : ClosedEmbedding f) : Continuous f := +theorem continuous (hf : IsClosedEmbedding f) : Continuous f := hf.toEmbedding.continuous -theorem isClosedMap (hf : ClosedEmbedding f) : IsClosedMap f := +theorem isClosedMap (hf : IsClosedEmbedding f) : IsClosedMap f := hf.toEmbedding.toInducing.isClosedMap hf.isClosed_range -theorem closed_iff_image_closed (hf : ClosedEmbedding f) {s : Set X} : +theorem closed_iff_image_closed (hf : IsClosedEmbedding f) {s : Set X} : IsClosed s ↔ IsClosed (f '' s) := ⟨hf.isClosedMap s, fun h => by rw [← preimage_image_eq s hf.inj] exact h.preimage hf.continuous⟩ -theorem closed_iff_preimage_closed (hf : ClosedEmbedding f) {s : Set Y} +theorem closed_iff_preimage_closed (hf : IsClosedEmbedding f) {s : Set Y} (hs : s ⊆ range f) : IsClosed s ↔ IsClosed (f ⁻¹' s) := by rw [hf.closed_iff_image_closed, image_preimage_eq_of_subset hs] -theorem _root_.closedEmbedding_of_embedding_closed (h₁ : Embedding f) (h₂ : IsClosedMap f) : - ClosedEmbedding f := +theorem _root_.IsClosedEmbedding.of_embedding_closed (h₁ : Embedding f) (h₂ : IsClosedMap f) : + IsClosedEmbedding f := ⟨h₁, image_univ (f := f) ▸ h₂ univ isClosed_univ⟩ -theorem _root_.closedEmbedding_of_continuous_injective_closed (h₁ : Continuous f) (h₂ : Injective f) - (h₃ : IsClosedMap f) : ClosedEmbedding f := by - refine closedEmbedding_of_embedding_closed ⟨⟨?_⟩, h₂⟩ h₃ +@[deprecated (since := "2024-10-20")] +alias _root_.closedEmbedding_of_embedding_closed := _root_.IsClosedEmbedding.of_embedding_closed + +lemma _root_.IsClosedEmbedding.of_continuous_injective_isClosedMap (h₁ : Continuous f) + (h₂ : Injective f) (h₃ : IsClosedMap f) : IsClosedEmbedding f := by + refine .of_embedding_closed ⟨⟨?_⟩, h₂⟩ h₃ refine h₁.le_induced.antisymm fun s hs => ?_ refine ⟨(f '' sᶜ)ᶜ, (h₃ _ hs.isClosed_compl).isOpen_compl, ?_⟩ rw [preimage_compl, preimage_image_eq _ h₂, compl_compl] -theorem _root_.closedEmbedding_id : ClosedEmbedding (@id X) := +@[deprecated (since := "2024-10-20")] +alias _root_.closedEmbedding_of_continuous_injective_closed := + IsClosedEmbedding.of_continuous_injective_isClosedMap + +theorem _root_.isClosedEmbedding_id : IsClosedEmbedding (@id X) := ⟨embedding_id, IsClosedMap.id.isClosed_range⟩ -theorem comp (hg : ClosedEmbedding g) (hf : ClosedEmbedding f) : - ClosedEmbedding (g ∘ f) := +@[deprecated (since := "2024-10-20")] +alias _root_.closedEmbedding_id := _root_.isClosedEmbedding_id + +theorem comp (hg : IsClosedEmbedding g) (hf : IsClosedEmbedding f) : + IsClosedEmbedding (g ∘ f) := ⟨hg.toEmbedding.comp hf.toEmbedding, (hg.isClosedMap.comp hf.isClosedMap).isClosed_range⟩ -theorem of_comp_iff (hg : ClosedEmbedding g) : - ClosedEmbedding (g ∘ f) ↔ ClosedEmbedding f := by - simp_rw [closedEmbedding_iff, hg.toEmbedding.of_comp_iff, Set.range_comp, +theorem of_comp_iff (hg : IsClosedEmbedding g) : + IsClosedEmbedding (g ∘ f) ↔ IsClosedEmbedding f := by + simp_rw [isClosedEmbedding_iff, hg.toEmbedding.of_comp_iff, Set.range_comp, ← hg.closed_iff_image_closed] -theorem closure_image_eq (hf : ClosedEmbedding f) (s : Set X) : +theorem closure_image_eq (hf : IsClosedEmbedding f) (s : Set X) : closure (f '' s) = f '' closure s := hf.isClosedMap.closure_image_eq_of_continuous hf.continuous s -end ClosedEmbedding +end IsClosedEmbedding -end ClosedEmbedding +end IsClosedEmbedding diff --git a/Mathlib/Topology/Maps/Proper/Basic.lean b/Mathlib/Topology/Maps/Proper/Basic.lean index aadee505b3cac..bd8b9f6139ec2 100644 --- a/Mathlib/Topology/Maps/Proper/Basic.lean +++ b/Mathlib/Topology/Maps/Proper/Basic.lean @@ -273,17 +273,25 @@ protected lemma IsHomeomorph.isProperMap (hf : IsHomeomorph f) : IsProperMap f : @[simp] lemma isProperMap_id : IsProperMap (id : X → X) := IsHomeomorph.id.isProperMap /-- A closed embedding is proper. -/ -lemma isProperMap_of_closedEmbedding (hf : ClosedEmbedding f) : IsProperMap f := +lemma IsClosedEmbedding.isProperMap (hf : IsClosedEmbedding f) : IsProperMap f := isProperMap_of_isClosedMap_of_inj hf.continuous hf.inj hf.isClosedMap +@[deprecated (since := "2024-10-20")] +alias isProperMap_of_closedEmbedding := IsClosedEmbedding.isProperMap + /-- The coercion from a closed subset is proper. -/ -lemma isProperMap_subtype_val_of_closed {U : Set X} (hU : IsClosed U) : IsProperMap ((↑) : U → X) := - isProperMap_of_closedEmbedding hU.closedEmbedding_subtype_val +lemma IsClosed.isProperMap_subtypeVal {C : Set X} (hC : IsClosed C) : IsProperMap ((↑) : C → X) := + hC.isClosedEmbedding_subtypeVal.isProperMap + +@[deprecated (since := "2024-10-20")] +alias isProperMap_subtype_val_of_closed := IsClosed.isProperMap_subtypeVal /-- The restriction of a proper map to a closed subset is proper. -/ -lemma isProperMap_restr_of_proper_of_closed {U : Set X} (hf : IsProperMap f) (hU : IsClosed U) : - IsProperMap (fun x : U ↦ f x) := - IsProperMap.comp (isProperMap_subtype_val_of_closed hU) hf +lemma IsProperMap.restrict {C : Set X} (hf : IsProperMap f) (hC : IsClosed C) : + IsProperMap fun x : C ↦ f x := hC.isProperMap_subtypeVal.comp hf + +@[deprecated (since := "2024-10-20")] +alias isProperMap_restr_of_proper_of_closed := IsProperMap.restrict /-- The range of a proper map is closed. -/ lemma IsProperMap.isClosed_range (hf : IsProperMap f) : IsClosed (range f) := diff --git a/Mathlib/Topology/MetricSpace/Antilipschitz.lean b/Mathlib/Topology/MetricSpace/Antilipschitz.lean index 67c88ebcb3216..3ee1290ac0f7c 100644 --- a/Mathlib/Topology/MetricSpace/Antilipschitz.lean +++ b/Mathlib/Topology/MetricSpace/Antilipschitz.lean @@ -165,11 +165,14 @@ theorem isClosed_range {α β : Type*} [PseudoEMetricSpace α] [EMetricSpace β] IsClosed (range f) := (hf.isComplete_range hfc).isClosed -theorem closedEmbedding {α : Type*} {β : Type*} [EMetricSpace α] [EMetricSpace β] {K : ℝ≥0} +theorem isClosedEmbedding {α : Type*} {β : Type*} [EMetricSpace α] [EMetricSpace β] {K : ℝ≥0} {f : α → β} [CompleteSpace α] (hf : AntilipschitzWith K f) (hfc : UniformContinuous f) : - ClosedEmbedding f := + IsClosedEmbedding f := { (hf.isUniformEmbedding hfc).embedding with isClosed_range := hf.isClosed_range hfc } +@[deprecated (since := "2024-10-20")] +alias closedEmbedding := isClosedEmbedding + theorem subtype_coe (s : Set α) : AntilipschitzWith 1 ((↑) : s → α) := AntilipschitzWith.id.restrict s diff --git a/Mathlib/Topology/MetricSpace/Basic.lean b/Mathlib/Topology/MetricSpace/Basic.lean index 185ec4912deba..884f0a045a2a0 100644 --- a/Mathlib/Topology/MetricSpace/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Basic.lean @@ -56,10 +56,13 @@ theorem isClosed_of_pairwise_le_dist {s : Set γ} {ε : ℝ} (hε : 0 < ε) (hs : s.Pairwise fun x y => ε ≤ dist x y) : IsClosed s := isClosed_of_spaced_out (dist_mem_uniformity hε) <| by simpa using hs -theorem closedEmbedding_of_pairwise_le_dist {α : Type*} [TopologicalSpace α] [DiscreteTopology α] +theorem isClosedEmbedding_of_pairwise_le_dist {α : Type*} [TopologicalSpace α] [DiscreteTopology α] {ε : ℝ} (hε : 0 < ε) {f : α → γ} (hf : Pairwise fun x y => ε ≤ dist (f x) (f y)) : - ClosedEmbedding f := - closedEmbedding_of_spaced_out (dist_mem_uniformity hε) <| by simpa using hf + IsClosedEmbedding f := + isClosedEmbedding_of_spaced_out (dist_mem_uniformity hε) <| by simpa using hf + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_of_pairwise_le_dist := isClosedEmbedding_of_pairwise_le_dist /-- If `f : β → α` sends any two distinct points to points at distance at least `ε > 0`, then `f` is a uniform embedding with respect to the discrete uniformity on `β`. -/ diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index 0ece9c35abd15..a0d5113e469b3 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -432,9 +432,11 @@ protected theorem embedding [PseudoEMetricSpace β] [DilationClass F α β] (f : (Dilation.isUniformEmbedding f).embedding /-- A dilation from a complete emetric space is a closed embedding -/ -protected theorem closedEmbedding [CompleteSpace α] [EMetricSpace β] [DilationClass F α β] (f : F) : - ClosedEmbedding f := - (antilipschitz f).closedEmbedding (lipschitz f).uniformContinuous +lemma isClosedEmbedding [CompleteSpace α] [EMetricSpace β] [DilationClass F α β] (f : F) : + IsClosedEmbedding f := + (antilipschitz f).isClosedEmbedding (lipschitz f).uniformContinuous + +@[deprecated (since := "2024-10-20")] alias closedEmbedding := isClosedEmbedding end EmetricDilation diff --git a/Mathlib/Topology/MetricSpace/Isometry.lean b/Mathlib/Topology/MetricSpace/Isometry.lean index 34b17ebabb3e4..21177706ee722 100644 --- a/Mathlib/Topology/MetricSpace/Isometry.lean +++ b/Mathlib/Topology/MetricSpace/Isometry.lean @@ -177,9 +177,12 @@ protected theorem embedding (hf : Isometry f) : Embedding f := hf.isUniformEmbedding.embedding /-- An isometry from a complete emetric space is a closed embedding -/ -theorem closedEmbedding [CompleteSpace α] [EMetricSpace γ] {f : α → γ} (hf : Isometry f) : - ClosedEmbedding f := - hf.antilipschitz.closedEmbedding hf.lipschitz.uniformContinuous +theorem isClosedEmbedding [CompleteSpace α] [EMetricSpace γ] {f : α → γ} (hf : Isometry f) : + IsClosedEmbedding f := + hf.antilipschitz.isClosedEmbedding hf.lipschitz.uniformContinuous + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding := isClosedEmbedding end EmetricIsometry diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index 6db429aedccbf..c1e6a80a0300c 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -140,8 +140,8 @@ theorem exists_nat_nat_continuous_surjective (α : Type*) [TopologicalSpace α] exists_nat_nat_continuous_surjective_of_completeSpace α /-- Given a closed embedding into a Polish space, the source space is also Polish. -/ -theorem _root_.ClosedEmbedding.polishSpace [TopologicalSpace α] [TopologicalSpace β] [PolishSpace β] - {f : α → β} (hf : ClosedEmbedding f) : PolishSpace α := by +lemma _root_.IsClosedEmbedding.polishSpace [TopologicalSpace α] [TopologicalSpace β] [PolishSpace β] + {f : α → β} (hf : IsClosedEmbedding f) : PolishSpace α := by letI := upgradePolishSpace β letI : MetricSpace α := hf.toEmbedding.comapMetricSpace f haveI : SecondCountableTopology α := hf.toEmbedding.secondCountableTopology @@ -150,25 +150,28 @@ theorem _root_.ClosedEmbedding.polishSpace [TopologicalSpace α] [TopologicalSpa exact hf.isClosed_range.isComplete infer_instance +@[deprecated (since := "2024-10-20")] +alias _root_.ClosedEmbedding.polishSpace := _root_.IsClosedEmbedding.polishSpace + /-- Any countable discrete space is Polish. -/ instance (priority := 50) polish_of_countable [TopologicalSpace α] [h : Countable α] [DiscreteTopology α] : PolishSpace α := by obtain ⟨f, hf⟩ := h.exists_injective_nat - have : ClosedEmbedding f := by - apply closedEmbedding_of_continuous_injective_closed continuous_of_discreteTopology hf - exact fun t _ => isClosed_discrete _ + have : IsClosedEmbedding f := + .of_continuous_injective_isClosedMap continuous_of_discreteTopology hf + fun t _ ↦ isClosed_discrete _ exact this.polishSpace /-- Pulling back a Polish topology under an equiv gives again a Polish topology. -/ theorem _root_.Equiv.polishSpace_induced [t : TopologicalSpace β] [PolishSpace β] (f : α ≃ β) : @PolishSpace α (t.induced f) := letI : TopologicalSpace α := t.induced f - (f.toHomeomorphOfInducing ⟨rfl⟩).closedEmbedding.polishSpace + (f.toHomeomorphOfInducing ⟨rfl⟩).isClosedEmbedding.polishSpace /-- A closed subset of a Polish space is also Polish. -/ theorem _root_.IsClosed.polishSpace [TopologicalSpace α] [PolishSpace α] {s : Set α} (hs : IsClosed s) : PolishSpace s := - (IsClosed.closedEmbedding_subtype_val hs).polishSpace + hs.isClosedEmbedding_subtypeVal.polishSpace instance instPolishSpaceUniv [TopologicalSpace α] [PolishSpace α] : PolishSpace (univ : Set α) := @@ -211,7 +214,7 @@ theorem exists_polishSpace_forall_le {ι : Type*} [Countable ι] [t : Topologica .iInf ⟨none, Option.forall.2 ⟨le_rfl, hm⟩⟩ <| Option.forall.2 ⟨p, h'm⟩⟩ instance : PolishSpace ENNReal := - ClosedEmbedding.polishSpace ⟨ENNReal.orderIsoUnitIntervalBirational.toHomeomorph.embedding, + IsClosedEmbedding.polishSpace ⟨ENNReal.orderIsoUnitIntervalBirational.toHomeomorph.embedding, ENNReal.orderIsoUnitIntervalBirational.range_eq ▸ isClosed_univ⟩ end PolishSpace diff --git a/Mathlib/Topology/SeparatedMap.lean b/Mathlib/Topology/SeparatedMap.lean index ccb0ea0f688df..dae896a7a7e4b 100644 --- a/Mathlib/Topology/SeparatedMap.lean +++ b/Mathlib/Topology/SeparatedMap.lean @@ -86,15 +86,18 @@ theorem isSeparatedMap_iff_isClosed_diagonal {f : X → Y} : · obtain ⟨s₁, h₁, s₂, h₂, s_sub⟩ := mem_prod_iff.mp ht exact ⟨s₁, h₁, s₂, h₂, disjoint_left.2 fun x h₁ h₂ ↦ @t_sub ⟨(x, x), rfl⟩ (s_sub ⟨h₁, h₂⟩) rfl⟩ -theorem isSeparatedMap_iff_closedEmbedding {f : X → Y} : - IsSeparatedMap f ↔ ClosedEmbedding (toPullbackDiag f) := by +theorem isSeparatedMap_iff_isClosedEmbedding {f : X → Y} : + IsSeparatedMap f ↔ IsClosedEmbedding (toPullbackDiag f) := by rw [isSeparatedMap_iff_isClosed_diagonal, ← range_toPullbackDiag] exact ⟨fun h ↦ ⟨embedding_toPullbackDiag f, h⟩, fun h ↦ h.isClosed_range⟩ +@[deprecated (since := "2024-10-20")] +alias isSeparatedMap_iff_closedEmbedding := isSeparatedMap_iff_isClosedEmbedding + theorem isSeparatedMap_iff_isClosedMap {f : X → Y} : IsSeparatedMap f ↔ IsClosedMap (toPullbackDiag f) := - isSeparatedMap_iff_closedEmbedding.trans - ⟨ClosedEmbedding.isClosedMap, closedEmbedding_of_continuous_injective_closed + isSeparatedMap_iff_isClosedEmbedding.trans + ⟨IsClosedEmbedding.isClosedMap, .of_continuous_injective_isClosedMap (embedding_toPullbackDiag f).continuous (injective_toPullbackDiag f)⟩ open Function.Pullback in diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index d2e132acfed38..386aaf8de79c8 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -999,17 +999,18 @@ theorem disjoint_nhdsWithin_of_mem_discrete {s : Set X} [DiscreteTopology s] {x ⟨{x}ᶜ ∩ V, inter_mem_nhdsWithin _ h, disjoint_iff_inter_eq_empty.mpr (by rw [inter_assoc, h', compl_inter_self])⟩ -theorem closedEmbedding_update {ι : Type*} {β : ι → Type*} +theorem isClosedEmbedding_update {ι : Type*} {β : ι → Type*} [DecidableEq ι] [(i : ι) → TopologicalSpace (β i)] (x : (i : ι) → β i) (i : ι) [(i : ι) → T1Space (β i)] : - ClosedEmbedding (update x i) := by - apply closedEmbedding_of_continuous_injective_closed - · exact continuous_const.update i continuous_id - · exact update_injective x i - · intro s hs - rw [update_image] - apply isClosed_set_pi - simp [forall_update_iff, hs, isClosed_singleton] + IsClosedEmbedding (update x i) := by + refine .of_continuous_injective_isClosedMap (continuous_const.update i continuous_id) + (update_injective x i) fun s hs ↦ ?_ + rw [update_image] + apply isClosed_set_pi + simp [forall_update_iff, hs, isClosed_singleton] + +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_update := isClosedEmbedding_update /-! ### R₁ (preregular) spaces -/ @@ -1754,10 +1755,13 @@ theorem Function.LeftInverse.isClosed_range [T2Space X] {f : X → Y} {g : Y → @[deprecated (since := "2024-03-17")] alias Function.LeftInverse.closed_range := Function.LeftInverse.isClosed_range -theorem Function.LeftInverse.closedEmbedding [T2Space X] {f : X → Y} {g : Y → X} - (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : ClosedEmbedding g := +theorem Function.LeftInverse.isClosedEmbedding [T2Space X] {f : X → Y} {g : Y → X} + (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : IsClosedEmbedding g := ⟨h.embedding hf hg, h.isClosed_range hf hg⟩ +@[deprecated (since := "2024-10-20")] +alias Function.LeftInverse.closedEmbedding := Function.LeftInverse.isClosedEmbedding + theorem SeparatedNhds.of_isCompact_isCompact [T2Space X] {s t : Set X} (hs : IsCompact s) (ht : IsCompact t) (hst : Disjoint s t) : SeparatedNhds s t := by simp only [SeparatedNhds, prod_subset_compl_diagonal_iff_disjoint.symm] at hst ⊢ @@ -1852,9 +1856,12 @@ protected theorem Continuous.isClosedMap [CompactSpace X] [T2Space Y] {f : X → (h : Continuous f) : IsClosedMap f := fun _s hs => (hs.isCompact.image h).isClosed /-- A continuous injective map from a compact space to a Hausdorff space is a closed embedding. -/ -theorem Continuous.closedEmbedding [CompactSpace X] [T2Space Y] {f : X → Y} (h : Continuous f) - (hf : Function.Injective f) : ClosedEmbedding f := - closedEmbedding_of_continuous_injective_closed h hf h.isClosedMap +theorem Continuous.isClosedEmbedding [CompactSpace X] [T2Space Y] {f : X → Y} (h : Continuous f) + (hf : Function.Injective f) : IsClosedEmbedding f := + .of_continuous_injective_isClosedMap h hf h.isClosedMap + +@[deprecated (since := "2024-10-20")] +alias Continuous.closedEmbedding := Continuous.isClosedEmbedding /-- A continuous surjective map from a compact space to a Hausdorff space is a quotient map. -/ theorem QuotientMap.of_surjective_continuous [CompactSpace X] [T2Space Y] {f : X → Y} @@ -2257,14 +2264,17 @@ theorem normal_exists_closure_subset [NormalSpace X] {s t : Set X} (hs : IsClose exact fun x hxs hxt => hs't'.le_bot ⟨hxs, hxt⟩ /-- If the codomain of a closed embedding is a normal space, then so is the domain. -/ -protected theorem ClosedEmbedding.normalSpace [TopologicalSpace Y] [NormalSpace Y] {f : X → Y} - (hf : ClosedEmbedding f) : NormalSpace X where +protected theorem IsClosedEmbedding.normalSpace [TopologicalSpace Y] [NormalSpace Y] {f : X → Y} + (hf : IsClosedEmbedding f) : NormalSpace X where normal s t hs ht hst := by have H : SeparatedNhds (f '' s) (f '' t) := NormalSpace.normal (f '' s) (f '' t) (hf.isClosedMap s hs) (hf.isClosedMap t ht) (disjoint_image_of_injective hf.inj hst) exact (H.preimage hf.continuous).mono (subset_preimage_image _ _) (subset_preimage_image _ _) +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.normalSpace := IsClosedEmbedding.normalSpace + instance (priority := 100) NormalSpace.of_compactSpace_r1Space [CompactSpace X] [R1Space X] : NormalSpace X where normal _s _t hs ht := .of_isCompact_isCompact_isClosed hs.isCompact ht.isCompact ht @@ -2302,13 +2312,16 @@ theorem T4Space.of_compactSpace_t2Space [CompactSpace X] [T2Space X] : T4Space X := inferInstance /-- If the codomain of a closed embedding is a T₄ space, then so is the domain. -/ -protected theorem ClosedEmbedding.t4Space [TopologicalSpace Y] [T4Space Y] {f : X → Y} - (hf : ClosedEmbedding f) : T4Space X where +protected theorem IsClosedEmbedding.t4Space [TopologicalSpace Y] [T4Space Y] {f : X → Y} + (hf : IsClosedEmbedding f) : T4Space X where toT1Space := hf.toEmbedding.t1Space toNormalSpace := hf.normalSpace +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.t4Space := IsClosedEmbedding.t4Space + instance ULift.instT4Space [T4Space X] : T4Space (ULift X) := - ULift.closedEmbedding_down.t4Space + ULift.isClosedEmbedding_down.t4Space namespace SeparationQuotient @@ -2596,7 +2609,7 @@ theorem loc_compact_Haus_tot_disc_of_zero_dim [TotallyDisconnectedSpace H] : haveI : CompactSpace s := isCompact_iff_compactSpace.1 comp obtain ⟨V : Set s, VisClopen, Vx, V_sub⟩ := compact_exists_isClopen_in_isOpen u_open_in_s xs have VisClopen' : IsClopen (((↑) : s → H) '' V) := by - refine ⟨comp.isClosed.closedEmbedding_subtype_val.closed_iff_image_closed.1 VisClopen.1, ?_⟩ + refine ⟨comp.isClosed.isClosedEmbedding_subtypeVal.closed_iff_image_closed.1 VisClopen.1, ?_⟩ let v : Set u := ((↑) : u → s) ⁻¹' V have : ((↑) : u → H) = ((↑) : s → H) ∘ ((↑) : u → s) := rfl have f0 : Embedding ((↑) : u → H) := embedding_subtype_val.comp embedding_subtype_val diff --git a/Mathlib/Topology/Sober.lean b/Mathlib/Topology/Sober.lean index 06f9286a3e5d3..6cf96933116ef 100644 --- a/Mathlib/Topology/Sober.lean +++ b/Mathlib/Topology/Sober.lean @@ -172,7 +172,7 @@ noncomputable def irreducibleSetEquivPoints [QuasiSober α] [T0Space α] : simp [hs'.closure_eq, ht'.closure_eq] rfl -theorem ClosedEmbedding.quasiSober {f : α → β} (hf : ClosedEmbedding f) [QuasiSober β] : +theorem IsClosedEmbedding.quasiSober {f : α → β} (hf : IsClosedEmbedding f) [QuasiSober β] : QuasiSober α where sober hS hS' := by have hS'' := hS.image f hf.continuous.continuousOn @@ -182,6 +182,9 @@ theorem ClosedEmbedding.quasiSober {f : α → β} (hf : ClosedEmbedding f) [Qua apply image_injective.mpr hf.inj rw [← hx.def, ← hf.closure_image_eq, image_singleton] +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.quasiSober := IsClosedEmbedding.quasiSober + theorem IsOpenEmbedding.quasiSober {f : α → β} (hf : IsOpenEmbedding f) [QuasiSober β] : QuasiSober α where sober hS hS' := by diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index 11a71da0311a5..f0d049776a387 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -201,13 +201,16 @@ theorem _root_.hasCompactMulSupport_comp_left (hg : ∀ {x}, g x = 1 ↔ x = 1) simp_rw [hasCompactMulSupport_def, mulSupport_comp_eq g (@hg) f] @[to_additive] -theorem comp_closedEmbedding (hf : HasCompactMulSupport f) {g : α' → α} - (hg : ClosedEmbedding g) : HasCompactMulSupport (f ∘ g) := by +theorem comp_isClosedEmbedding (hf : HasCompactMulSupport f) {g : α' → α} + (hg : IsClosedEmbedding g) : HasCompactMulSupport (f ∘ g) := by rw [hasCompactMulSupport_def, Function.mulSupport_comp_eq_preimage] refine IsCompact.of_isClosed_subset (hg.isCompact_preimage hf) isClosed_closure ?_ rw [hg.toEmbedding.closure_eq_preimage_closure_image] exact preimage_mono (closure_mono <| image_preimage_subset _ _) +@[deprecated (since := "2024-10-20")] +alias comp_closedEmbedding := comp_isClosedEmbedding + @[to_additive] theorem comp₂_left (hf : HasCompactMulSupport f) (hf₂ : HasCompactMulSupport f₂) (hm : m 1 1 = 1) : diff --git a/Mathlib/Topology/TietzeExtension.lean b/Mathlib/Topology/TietzeExtension.lean index 9a2271b7f1202..70735e4ee49a4 100644 --- a/Mathlib/Topology/TietzeExtension.lean +++ b/Mathlib/Topology/TietzeExtension.lean @@ -20,7 +20,7 @@ bounded function, then there exists a bounded extension of the same norm. The proof mostly follows . We patch a small gap in the proof for unbounded functions, see -`exists_extension_forall_exists_le_ge_of_closedEmbedding`. +`exists_extension_forall_exists_le_ge_of_isClosedEmbedding`. In addition we provide a class `TietzeExtension` encoding the idea that a topological space satisfies the Tietze extension theorem. This allows us to get a version of the Tietze extension @@ -68,7 +68,7 @@ theorem ContinuousMap.exists_restrict_eq (hs : IsClosed s) (f : C(s, Y)) : nonempty topological space `X₁` into a normal topological space `X`. Let `f` be a continuous function on `X₁` with values in a `TietzeExtension` space `Y`. Then there exists a continuous function `g : C(X, Y)` such that `g ∘ e = f`. -/ -theorem ContinuousMap.exists_extension (he : ClosedEmbedding e) (f : C(X₁, Y)) : +theorem ContinuousMap.exists_extension (he : IsClosedEmbedding e) (f : C(X₁, Y)) : ∃ (g : C(X, Y)), g.comp ⟨e, he.continuous⟩ = f := by let e' : X₁ ≃ₜ Set.range e := Homeomorph.ofEmbedding _ he.toEmbedding obtain ⟨g, hg⟩ := (f.comp e'.symm).exists_restrict_eq he.isClosed_range @@ -81,7 +81,7 @@ continuous function `g : C(X, Y)` such that `g ∘ e = f`. This version is provided for convenience and backwards compatibility. Here the composition is phrased in terms of bare functions. -/ -theorem ContinuousMap.exists_extension' (he : ClosedEmbedding e) (f : C(X₁, Y)) : +theorem ContinuousMap.exists_extension' (he : IsClosedEmbedding e) (f : C(X₁, Y)) : ∃ (g : C(X, Y)), g ∘ e = f := f.exists_extension he |>.imp fun g hg ↦ by ext x; congrm($(hg) x) @@ -104,7 +104,7 @@ the radius is strictly positive, so finding this as an instance will fail. Instead, it is intended to be used as a constructor for theorems about sets which *do* satisfy `[TietzeExtension t]` under some hypotheses. -/ -theorem ContinuousMap.exists_extension_forall_mem (he : ClosedEmbedding e) +theorem ContinuousMap.exists_extension_forall_mem (he : IsClosedEmbedding e) {Y : Type v} [TopologicalSpace Y] (f : C(X₁, Y)) {t : Set Y} (hf : ∀ x, f x ∈ t) [ht : TietzeExtension.{u, v} t] : ∃ (g : C(X, Y)), (∀ x, g x ∈ t) ∧ g.comp ⟨e, he.continuous⟩ = f := by @@ -166,7 +166,7 @@ namespace BoundedContinuousFunction of a topological space into a normal topological space and `f : X →ᵇ ℝ` is a bounded continuous function, then there exists a bounded continuous function `g : Y →ᵇ ℝ` of the norm `‖g‖ ≤ ‖f‖ / 3` such that the distance between `g ∘ e` and `f` is at most `(2 / 3) * ‖f‖`. -/ -theorem tietze_extension_step (f : X →ᵇ ℝ) (e : C(X, Y)) (he : ClosedEmbedding e) : +theorem tietze_extension_step (f : X →ᵇ ℝ) (e : C(X, Y)) (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, ‖g‖ ≤ ‖f‖ / 3 ∧ dist (g.compContinuous e) f ≤ 2 / 3 * ‖f‖ := by have h3 : (0 : ℝ) < 3 := by norm_num1 have h23 : 0 < (2 / 3 : ℝ) := by norm_num1 @@ -216,8 +216,8 @@ theorem tietze_extension_step (f : X →ᵇ ℝ) (e : C(X, Y)) (he : ClosedEmbed embedding and bundled composition. If `e : C(X, Y)` is a closed embedding of a topological space into a normal topological space and `f : X →ᵇ ℝ` is a bounded continuous function, then there exists a bounded continuous function `g : Y →ᵇ ℝ` of the same norm such that `g ∘ e = f`. -/ -theorem exists_extension_norm_eq_of_closedEmbedding' (f : X →ᵇ ℝ) (e : C(X, Y)) - (he : ClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g.compContinuous e = f := by +theorem exists_extension_norm_eq_of_isClosedEmbedding' (f : X →ᵇ ℝ) (e : C(X, Y)) + (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g.compContinuous e = f := by /- For the proof, we iterate `tietze_extension_step`. Each time we apply it to the difference between the previous approximation and `f`. -/ choose F hF_norm hF_dist using fun f : X →ᵇ ℝ => tietze_extension_step f e he @@ -261,36 +261,42 @@ theorem exists_extension_norm_eq_of_closedEmbedding' (f : X →ᵇ ℝ) (e : C(X · rw [← hge] exact norm_compContinuous_le _ _ +@[deprecated (since := "2024-10-20")] +alias exists_extension_norm_eq_of_closedEmbedding' := exists_extension_norm_eq_of_isClosedEmbedding' + /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version with a closed embedding and unbundled composition. If `e : C(X, Y)` is a closed embedding of a topological space into a normal topological space and `f : X →ᵇ ℝ` is a bounded continuous function, then there exists a bounded continuous function `g : Y →ᵇ ℝ` of the same norm such that `g ∘ e = f`. -/ -theorem exists_extension_norm_eq_of_closedEmbedding (f : X →ᵇ ℝ) {e : X → Y} - (he : ClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g ∘ e = f := by - rcases exists_extension_norm_eq_of_closedEmbedding' f ⟨e, he.continuous⟩ he with ⟨g, hg, rfl⟩ +theorem exists_extension_norm_eq_of_isClosedEmbedding (f : X →ᵇ ℝ) {e : X → Y} + (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g ∘ e = f := by + rcases exists_extension_norm_eq_of_isClosedEmbedding' f ⟨e, he.continuous⟩ he with ⟨g, hg, rfl⟩ exact ⟨g, hg, rfl⟩ +@[deprecated (since := "2024-10-20")] +alias exists_extension_norm_eq_of_closedEmbedding := exists_extension_norm_eq_of_isClosedEmbedding + /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version for a closed set. If `f` is a bounded continuous real-valued function defined on a closed set in a normal topological space, then it can be extended to a bounded continuous function of the same norm defined on the whole space. -/ theorem exists_norm_eq_restrict_eq_of_closed {s : Set Y} (f : s →ᵇ ℝ) (hs : IsClosed s) : ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g.restrict s = f := - exists_extension_norm_eq_of_closedEmbedding' f ((ContinuousMap.id _).restrict s) - (closedEmbedding_subtype_val hs) + exists_extension_norm_eq_of_isClosedEmbedding' f ((ContinuousMap.id _).restrict s) + hs.isClosedEmbedding_subtypeVal /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version for a closed embedding and a bounded continuous function that takes values in a non-trivial closed interval. -See also `exists_extension_forall_mem_of_closedEmbedding` for a more general statement that works +See also `exists_extension_forall_mem_of_isClosedEmbedding` for a more general statement that works for any interval (finite or infinite, open or closed). If `e : X → Y` is a closed embedding and `f : X →ᵇ ℝ` is a bounded continuous function such that `f x ∈ [a, b]` for all `x`, where `a ≤ b`, then there exists a bounded continuous function `g : Y →ᵇ ℝ` such that `g y ∈ [a, b]` for all `y` and `g ∘ e = f`. -/ -theorem exists_extension_forall_mem_Icc_of_closedEmbedding (f : X →ᵇ ℝ) {a b : ℝ} {e : X → Y} - (hf : ∀ x, f x ∈ Icc a b) (hle : a ≤ b) (he : ClosedEmbedding e) : +theorem exists_extension_forall_mem_Icc_of_isClosedEmbedding (f : X →ᵇ ℝ) {a b : ℝ} {e : X → Y} + (hf : ∀ x, f x ∈ Icc a b) (hle : a ≤ b) (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, (∀ y, g y ∈ Icc a b) ∧ g ∘ e = f := by - rcases exists_extension_norm_eq_of_closedEmbedding (f - const X ((a + b) / 2)) he with + rcases exists_extension_norm_eq_of_isClosedEmbedding (f - const X ((a + b) / 2)) he with ⟨g, hgf, hge⟩ refine ⟨const Y ((a + b) / 2) + g, fun y => ?_, ?_⟩ · suffices ‖f - const X ((a + b) / 2)‖ ≤ (b - a) / 2 by @@ -302,13 +308,17 @@ theorem exists_extension_forall_mem_Icc_of_closedEmbedding (f : X →ᵇ ℝ) {a have : g (e x) = f x - (a + b) / 2 := congr_fun hge x simp [this] +@[deprecated (since := "2024-10-20")] +alias exists_extension_forall_mem_Icc_of_closedEmbedding := + exists_extension_forall_mem_Icc_of_isClosedEmbedding + /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version for a closed embedding. Let `e` be a closed embedding of a nonempty topological space `X` into a normal topological space `Y`. Let `f` be a bounded continuous real-valued function on `X`. Then there exists a bounded continuous function `g : Y →ᵇ ℝ` such that `g ∘ e = f` and each value `g y` belongs to a closed interval `[f x₁, f x₂]` for some `x₁` and `x₂`. -/ -theorem exists_extension_forall_exists_le_ge_of_closedEmbedding [Nonempty X] (f : X →ᵇ ℝ) - {e : X → Y} (he : ClosedEmbedding e) : +theorem exists_extension_forall_exists_le_ge_of_isClosedEmbedding [Nonempty X] (f : X →ᵇ ℝ) + {e : X → Y} (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, (∀ y, ∃ x₁ x₂, g y ∈ Icc (f x₁) (f x₂)) ∧ g ∘ e = f := by inhabit X -- Put `a = ⨅ x, f x` and `b = ⨆ x, f x` @@ -329,12 +339,12 @@ theorem exists_extension_forall_exists_le_ge_of_closedEmbedding [Nonempty X] (f have hsub : c - a = b - c := by field_simp [c] ring - /- Due to `exists_extension_forall_mem_Icc_of_closedEmbedding`, there exists an extension `g` + /- Due to `exists_extension_forall_mem_Icc_of_isClosedEmbedding`, there exists an extension `g` such that `g y ∈ [a, b]` for all `y`. However, if `a` and/or `b` do not belong to the range of `f`, then we need to ensure that these points do not belong to the range of `g`. This is done in two almost identical steps. First we deal with the case `∀ x, f x ≠ a`. -/ obtain ⟨g, hg_mem, hgf⟩ : ∃ g : Y →ᵇ ℝ, (∀ y, ∃ x, g y ∈ Icc (f x) b) ∧ g ∘ e = f := by - rcases exists_extension_forall_mem_Icc_of_closedEmbedding f hmem hle he with ⟨g, hg_mem, hgf⟩ + rcases exists_extension_forall_mem_Icc_of_isClosedEmbedding f hmem hle he with ⟨g, hg_mem, hgf⟩ -- If `a ∈ range f`, then we are done. rcases em (∃ x, f x = a) with (⟨x, rfl⟩ | ha') · exact ⟨g, fun y => ⟨x, hg_mem _⟩, hgf⟩ @@ -412,6 +422,10 @@ theorem exists_extension_forall_exists_le_ge_of_closedEmbedding [Nonempty X] (f · refine ⟨xl y, xu, ?_, hyxu.le⟩ simp [dg0 (Or.inr hc), hxl] +@[deprecated (since := "2024-10-20")] +alias exists_extension_forall_exists_le_ge_of_closedEmbedding := + exists_extension_forall_exists_le_ge_of_isClosedEmbedding + /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version for a closed embedding. Let `e` be a closed embedding of a nonempty topological space `X` into a normal topological space `Y`. Let `f` be a bounded continuous real-valued function on `X`. Let `t` be @@ -419,17 +433,21 @@ a nonempty convex set of real numbers (we use `OrdConnected` instead of `Convex` deduce this argument by typeclass search) such that `f x ∈ t` for all `x`. Then there exists a bounded continuous real-valued function `g : Y →ᵇ ℝ` such that `g y ∈ t` for all `y` and `g ∘ e = f`. -/ -theorem exists_extension_forall_mem_of_closedEmbedding (f : X →ᵇ ℝ) {t : Set ℝ} {e : X → Y} - [hs : OrdConnected t] (hf : ∀ x, f x ∈ t) (hne : t.Nonempty) (he : ClosedEmbedding e) : +theorem exists_extension_forall_mem_of_isClosedEmbedding (f : X →ᵇ ℝ) {t : Set ℝ} {e : X → Y} + [hs : OrdConnected t] (hf : ∀ x, f x ∈ t) (hne : t.Nonempty) (he : IsClosedEmbedding e) : ∃ g : Y →ᵇ ℝ, (∀ y, g y ∈ t) ∧ g ∘ e = f := by cases isEmpty_or_nonempty X · rcases hne with ⟨c, hc⟩ exact ⟨const Y c, fun _ => hc, funext fun x => isEmptyElim x⟩ - rcases exists_extension_forall_exists_le_ge_of_closedEmbedding f he with ⟨g, hg, hgf⟩ + rcases exists_extension_forall_exists_le_ge_of_isClosedEmbedding f he with ⟨g, hg, hgf⟩ refine ⟨g, fun y => ?_, hgf⟩ rcases hg y with ⟨xl, xu, h⟩ exact hs.out (hf _) (hf _) h +@[deprecated (since := "2024-10-20")] +alias exists_extension_forall_mem_of_closedEmbedding := + exists_extension_forall_mem_of_isClosedEmbedding + /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version for a closed set. Let `s` be a closed set in a normal topological space `Y`. Let `f` be a bounded continuous real-valued function on `s`. Let `t` be a nonempty convex set of real numbers (we use @@ -439,9 +457,8 @@ that `f x ∈ t` for all `x : s`. Then there exists a bounded continuous real-va theorem exists_forall_mem_restrict_eq_of_closed {s : Set Y} (f : s →ᵇ ℝ) (hs : IsClosed s) {t : Set ℝ} [OrdConnected t] (hf : ∀ x, f x ∈ t) (hne : t.Nonempty) : ∃ g : Y →ᵇ ℝ, (∀ y, g y ∈ t) ∧ g.restrict s = f := by - rcases exists_extension_forall_mem_of_closedEmbedding f hf hne - (closedEmbedding_subtype_val hs) with - ⟨g, hg, hgf⟩ + obtain ⟨g, hg, hgf⟩ := + exists_extension_forall_mem_of_isClosedEmbedding f hf hne hs.isClosedEmbedding_subtypeVal exact ⟨g, hg, DFunLike.coe_injective hgf⟩ end BoundedContinuousFunction @@ -454,8 +471,8 @@ topological space `Y`. Let `f` be a continuous real-valued function on `X`. Let convex set of real numbers (we use `OrdConnected` instead of `Convex` to automatically deduce this argument by typeclass search) such that `f x ∈ t` for all `x`. Then there exists a continuous real-valued function `g : C(Y, ℝ)` such that `g y ∈ t` for all `y` and `g ∘ e = f`. -/ -theorem exists_extension_forall_mem_of_closedEmbedding (f : C(X, ℝ)) {t : Set ℝ} {e : X → Y} - [hs : OrdConnected t] (hf : ∀ x, f x ∈ t) (hne : t.Nonempty) (he : ClosedEmbedding e) : +theorem exists_extension_forall_mem_of_isClosedEmbedding (f : C(X, ℝ)) {t : Set ℝ} {e : X → Y} + [hs : OrdConnected t] (hf : ∀ x, f x ∈ t) (hne : t.Nonempty) (he : IsClosedEmbedding e) : ∃ g : C(Y, ℝ), (∀ y, g y ∈ t) ∧ g ∘ e = f := by have h : ℝ ≃o Ioo (-1 : ℝ) 1 := orderIsoIooNegOneOne ℝ let F : X →ᵇ ℝ := @@ -474,7 +491,7 @@ theorem exists_extension_forall_mem_of_closedEmbedding (f : C(X, ℝ)) {t : Set rcases hz with ⟨z, hz, rfl⟩ exact ⟨z, hs.out hx hy hz, rfl⟩ have hFt : ∀ x, F x ∈ t' := fun x => mem_image_of_mem _ (hf x) - rcases F.exists_extension_forall_mem_of_closedEmbedding hFt (hne.image _) he with ⟨G, hG, hGF⟩ + rcases F.exists_extension_forall_mem_of_isClosedEmbedding hFt (hne.image _) he with ⟨G, hG, hGF⟩ let g : C(Y, ℝ) := ⟨h.symm ∘ codRestrict G _ fun y => ht_sub (hG y), h.symm.continuous.comp <| G.continuous.subtype_mk _⟩ @@ -487,8 +504,15 @@ theorem exists_extension_forall_mem_of_closedEmbedding (f : C(X, ℝ)) {t : Set · ext x exact hgG.2 (congr_fun hGF _) +@[deprecated (since := "2024-10-20")] +alias exists_extension_forall_mem_of_closedEmbedding := + exists_extension_forall_mem_of_isClosedEmbedding + @[deprecated (since := "2024-01-16")] -alias exists_extension_of_closedEmbedding := exists_extension' +alias exists_extension_of_isClosedEmbedding := exists_extension' + +@[deprecated (since := "2024-10-20")] +alias exists_extension_of_closedEmbedding := exists_extension_of_isClosedEmbedding /-- **Tietze extension theorem** for real-valued continuous maps, a version for a closed set. Let `s` be a closed set in a normal topological space `Y`. Let `f` be a continuous real-valued function @@ -500,7 +524,7 @@ theorem exists_restrict_eq_forall_mem_of_closed {s : Set Y} (f : C(s, ℝ)) {t : [OrdConnected t] (ht : ∀ x, f x ∈ t) (hne : t.Nonempty) (hs : IsClosed s) : ∃ g : C(Y, ℝ), (∀ y, g y ∈ t) ∧ g.restrict s = f := let ⟨g, hgt, hgf⟩ := - exists_extension_forall_mem_of_closedEmbedding f ht hne (closedEmbedding_subtype_val hs) + exists_extension_forall_mem_of_isClosedEmbedding f ht hne hs.isClosedEmbedding_subtypeVal ⟨g, hgt, coe_injective hgf⟩ @[deprecated (since := "2024-01-16")] alias exists_restrict_eq_of_closed := exists_restrict_eq diff --git a/Mathlib/Topology/UniformSpace/Ascoli.lean b/Mathlib/Topology/UniformSpace/Ascoli.lean index 351606b63c56c..4de38296cb8f6 100644 --- a/Mathlib/Topology/UniformSpace/Ascoli.lean +++ b/Mathlib/Topology/UniformSpace/Ascoli.lean @@ -456,14 +456,17 @@ and `F : ι → (X → α)`. Assume that: * For all `x`, the range of `i ↦ F i x` is contained in some fixed compact subset. Then `ι` is compact. -/ -theorem ArzelaAscoli.compactSpace_of_closedEmbedding [TopologicalSpace ι] {𝔖 : Set (Set X)} - (𝔖_compact : ∀ K ∈ 𝔖, IsCompact K) (F_clemb : ClosedEmbedding (UniformOnFun.ofFun 𝔖 ∘ F)) +theorem ArzelaAscoli.compactSpace_of_isClosedEmbedding [TopologicalSpace ι] {𝔖 : Set (Set X)} + (𝔖_compact : ∀ K ∈ 𝔖, IsCompact K) (F_clemb : IsClosedEmbedding (UniformOnFun.ofFun 𝔖 ∘ F)) (F_eqcont : ∀ K ∈ 𝔖, EquicontinuousOn F K) (F_pointwiseCompact : ∀ K ∈ 𝔖, ∀ x ∈ K, ∃ Q, IsCompact Q ∧ ∀ i, F i x ∈ Q) : CompactSpace ι := compactSpace_of_closed_inducing' 𝔖_compact F_clemb.toInducing F_clemb.isClosed_range F_eqcont F_pointwiseCompact +@[deprecated (since := "2024-10-20")] +alias ArzelaAscoli.compactSpace_of_closedEmbedding := ArzelaAscoli.compactSpace_of_isClosedEmbedding + /-- A version of the **Arzela-Ascoli theorem**. Let `X, ι` be topological spaces, `𝔖` a covering of `X` by compact subsets, `α` a T2 uniform space, @@ -474,13 +477,13 @@ Let `X, ι` be topological spaces, `𝔖` a covering of `X` by compact subsets, * For all `x ∈ ⋃₀ 𝔖`, the image of `s` under `i ↦ F i x` is contained in some fixed compact subset. Then `s` has compact closure in `ι`. -/ -theorem ArzelaAscoli.isCompact_closure_of_closedEmbedding [TopologicalSpace ι] [T2Space α] +theorem ArzelaAscoli.isCompact_closure_of_isClosedEmbedding [TopologicalSpace ι] [T2Space α] {𝔖 : Set (Set X)} (𝔖_compact : ∀ K ∈ 𝔖, IsCompact K) - (F_clemb : ClosedEmbedding (UniformOnFun.ofFun 𝔖 ∘ F)) + (F_clemb : IsClosedEmbedding (UniformOnFun.ofFun 𝔖 ∘ F)) {s : Set ι} (s_eqcont : ∀ K ∈ 𝔖, EquicontinuousOn (F ∘ ((↑) : s → ι)) K) (s_pointwiseCompact : ∀ K ∈ 𝔖, ∀ x ∈ K, ∃ Q, IsCompact Q ∧ ∀ i ∈ s, F i x ∈ Q) : IsCompact (closure s) := by - -- We apply `ArzelaAscoli.compactSpace_of_closedEmbedding` to the map + -- We apply `ArzelaAscoli.compactSpace_of_isClosedEmbedding` to the map -- `F ∘ (↑) : closure s → (X → α)`, for which all the hypotheses are easily verified. rw [isCompact_iff_compactSpace] have : ∀ K ∈ 𝔖, ∀ x ∈ K, Continuous (eval x ∘ F) := fun K hK x hx ↦ @@ -491,10 +494,14 @@ theorem ArzelaAscoli.isCompact_closure_of_closedEmbedding [TopologicalSpace ι] have cls_pointwiseCompact : ∀ K ∈ 𝔖, ∀ x ∈ K, ∃ Q, IsCompact Q ∧ ∀ i ∈ closure s, F i x ∈ Q := fun K hK x hx ↦ (s_pointwiseCompact K hK x hx).imp fun Q hQ ↦ ⟨hQ.1, closure_minimal hQ.2 <| hQ.1.isClosed.preimage (this K hK x hx)⟩ - exact ArzelaAscoli.compactSpace_of_closedEmbedding 𝔖_compact - (F_clemb.comp isClosed_closure.closedEmbedding_subtype_val) cls_eqcont + exact ArzelaAscoli.compactSpace_of_isClosedEmbedding 𝔖_compact + (F_clemb.comp isClosed_closure.isClosedEmbedding_subtypeVal) cls_eqcont fun K hK x hx ↦ (cls_pointwiseCompact K hK x hx).imp fun Q hQ ↦ ⟨hQ.1, by simpa using hQ.2⟩ +@[deprecated (since := "2024-10-20")] +alias ArzelaAscoli.isCompact_closure_of_closedEmbedding := + ArzelaAscoli.isCompact_closure_of_isClosedEmbedding + /-- A version of the **Arzela-Ascoli theorem**. If an equicontinuous family of continuous functions is compact in the pointwise topology, then it diff --git a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean index 7444f5a467be3..cc34886c49d41 100644 --- a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean +++ b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean @@ -27,13 +27,19 @@ theorem IsComplete.isClosed [UniformSpace α] [T0Space α] {s : Set α} (h : IsC rcases h f this inf_le_right with ⟨y, ys, fy⟩ rwa [(tendsto_nhds_unique' ha inf_le_left fy : a = y)] -theorem IsUniformEmbedding.toClosedEmbedding [UniformSpace α] [UniformSpace β] [CompleteSpace α] +theorem IsUniformEmbedding.toIsClosedEmbedding [UniformSpace α] [UniformSpace β] [CompleteSpace α] [T0Space β] {f : α → β} (hf : IsUniformEmbedding f) : - ClosedEmbedding f := + IsClosedEmbedding f := ⟨hf.embedding, hf.isUniformInducing.isComplete_range.isClosed⟩ +@[deprecated (since := "2024-10-20")] +alias IsUniformEmbedding.toClosedEmbedding := IsUniformEmbedding.toIsClosedEmbedding + @[deprecated (since := "2024-10-01")] -alias UniformEmbedding.toClosedEmbedding := IsUniformEmbedding.toClosedEmbedding +alias UniformEmbedding.toIsClosedEmbedding := IsUniformEmbedding.toIsClosedEmbedding + +@[deprecated (since := "2024-10-20")] +alias UniformEmbedding.toClosedEmbedding := UniformEmbedding.toIsClosedEmbedding namespace IsDenseInducing diff --git a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean index 36dee5874506e..565e1e1a1e8be 100644 --- a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean +++ b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean @@ -347,14 +347,17 @@ alias UniformEmbedding.isDenseEmbedding := IsUniformEmbedding.isDenseEmbedding @[deprecated (since := "2024-09-30")] alias IsUniformEmbedding.denseEmbedding := IsUniformEmbedding.isDenseEmbedding -theorem closedEmbedding_of_spaced_out {α} [TopologicalSpace α] [DiscreteTopology α] +theorem isClosedEmbedding_of_spaced_out {α} [TopologicalSpace α] [DiscreteTopology α] [T0Space β] {f : α → β} {s : Set (β × β)} (hs : s ∈ 𝓤 β) - (hf : Pairwise fun x y => (f x, f y) ∉ s) : ClosedEmbedding f := by + (hf : Pairwise fun x y => (f x, f y) ∉ s) : IsClosedEmbedding f := by rcases @DiscreteTopology.eq_bot α _ _ with rfl; let _ : UniformSpace α := ⊥ exact { (isUniformEmbedding_of_spaced_out hs hf).embedding with isClosed_range := isClosed_range_of_spaced_out hs hf } +@[deprecated (since := "2024-10-20")] +alias closedEmbedding_of_spaced_out := isClosedEmbedding_of_spaced_out + 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 From 8d27fbf359865a12c7bc7ed41bfe6ede9fd372a3 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Sun, 20 Oct 2024 16:57:21 +0000 Subject: [PATCH 126/131] feat(CategoryTheory): pull colimits out of hom functors with Yoneda on the LHS (#17440) A technical result on the way towards showing that Ind-objects are closed under filtered colimits. Co-authored-by: Markus Himmel --- Mathlib.lean | 1 + .../Limits/Preserves/Yoneda.lean | 79 +++++++++++++++++++ 2 files changed, 80 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean diff --git a/Mathlib.lean b/Mathlib.lean index de1b28985d7b6..ebba90f308a9c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1682,6 +1682,7 @@ import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Square import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Terminal import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Zero import Mathlib.CategoryTheory.Limits.Preserves.Ulift +import Mathlib.CategoryTheory.Limits.Preserves.Yoneda import Mathlib.CategoryTheory.Limits.Presheaf import Mathlib.CategoryTheory.Limits.Shapes.BinaryProducts import Mathlib.CategoryTheory.Limits.Shapes.Biproducts diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean new file mode 100644 index 0000000000000..629c22852c6a5 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean @@ -0,0 +1,79 @@ +/- +Copyright (c) 2024 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Limits.Preserves.Ulift +import Mathlib.CategoryTheory.Limits.FunctorToTypes + +/-! +# Yoneda preserves certain colimits + +Given a bifunctor `F : J ⥤ Cᵒᵖ ⥤ Type v`, we prove the isomorphism +`Hom(YX, colim_j F(j, -)) ≅ colim_j Hom(YX, F(j, -))`, where `Y` is the Yoneda embedding. + +We state this in two ways. One is functorial in `X` and stated as a natural isomorphism of functors +`yoneda.op ⋙ yoneda.obj (colimit F) ≅ yoneda.op ⋙ colimit (F ⋙ yoneda)`, and from this we +deduce the more traditional preservation statement +`PreservesColimit F (coyoneda.obj (op (yoneda.obj X)))`. + +The proof combines the Yoneda lemma with the fact that colimits in functor categories are computed +pointwise. + +## See also + +There is also a relative version of this statement where `F : J ⥤ Over A` for some presheaf +`A`, see `CategoryTheory.Comma.Presheaf`. + +-/ + +universe v₁ v₂ v₃ u₁ u₂ u₃ + +namespace CategoryTheory + +open CategoryTheory.Limits Opposite + +variable {C : Type u₁} [Category.{v₁} C] + +variable {J : Type u₂} [Category.{v₂} J] [HasColimitsOfShape J (Type v₁)] + [HasColimitsOfShape J (Type (max u₁ v₁))] (F : J ⥤ Cᵒᵖ ⥤ Type v₁) + +/-- Naturally in `X`, we have `Hom(YX, colim_i Fi) ≅ colim_i Hom(YX, Fi)`. -/ +noncomputable def yonedaYonedaColimit : + yoneda.op ⋙ yoneda.obj (colimit F) ≅ yoneda.op ⋙ colimit (F ⋙ yoneda) := calc + yoneda.op ⋙ yoneda.obj (colimit F) + ≅ colimit F ⋙ uliftFunctor.{u₁} := yonedaOpCompYonedaObj (colimit F) + _ ≅ F.flip ⋙ colim ⋙ uliftFunctor.{u₁} := + isoWhiskerRight (colimitIsoFlipCompColim F) uliftFunctor.{u₁} + _ ≅ F.flip ⋙ (whiskeringRight _ _ _).obj uliftFunctor.{u₁} ⋙ colim := + isoWhiskerLeft F.flip (preservesColimitNatIso uliftFunctor.{u₁}) + _ ≅ (yoneda.op ⋙ coyoneda ⋙ (whiskeringLeft _ _ _).obj F) ⋙ colim := isoWhiskerRight + (isoWhiskerRight largeCurriedYonedaLemma.symm ((whiskeringLeft _ _ _).obj F)) colim + _ ≅ yoneda.op ⋙ colimit (F ⋙ yoneda) := + isoWhiskerLeft yoneda.op (colimitIsoFlipCompColim (F ⋙ yoneda)).symm + +theorem yonedaYonedaColimit_app_inv {X : C} : ((yonedaYonedaColimit F).app (op X)).inv = + (colimitObjIsoColimitCompEvaluation _ _).hom ≫ + (colimit.post F (coyoneda.obj (op (yoneda.obj X)))) := by + dsimp [yonedaYonedaColimit] + simp only [Category.id_comp, Iso.cancel_iso_hom_left] + apply colimit.hom_ext + intro j + rw [colimit.ι_post, ι_colimMap_assoc] + simp only [← CategoryTheory.Functor.assoc, comp_evaluation] + rw [ι_preservesColimitsIso_inv_assoc, ← Functor.map_comp_assoc] + simp only [← comp_evaluation] + rw [colimitObjIsoColimitCompEvaluation_ι_inv, whiskerLeft_app] + ext η Y f + simp [largeCurriedYonedaLemma, yonedaOpCompYonedaObj, FunctorToTypes.colimit.map_ι_apply, + map_yonedaEquiv] + +noncomputable instance {X : C} : PreservesColimit F (coyoneda.obj (op (yoneda.obj X))) := by + suffices IsIso (colimit.post F (coyoneda.obj (op (yoneda.obj X)))) from + preservesColimitOfIsIsoPost _ _ + suffices colimit.post F (coyoneda.obj (op (yoneda.obj X))) = + (colimitObjIsoColimitCompEvaluation _ _).inv ≫ ((yonedaYonedaColimit F).app (op X)).inv from + this ▸ inferInstance + rw [yonedaYonedaColimit_app_inv, Iso.inv_hom_id_assoc] + +end CategoryTheory From 03a37e0be41d95a6f229d3578b5481c50e125cf4 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 20 Oct 2024 17:40:31 +0000 Subject: [PATCH 127/131] chore(LinearAlgebra/TensorProduct): fix duplication (#17912) This was unintentionally made in [#16776](https://github.com/leanprover-community/mathlib4/pull/16776) by myself. --- Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean | 2 +- Mathlib/LinearAlgebra/TensorProduct/Basic.lean | 9 ++------- 2 files changed, 3 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean index 2275dee475858..28e837f52f457 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean @@ -60,7 +60,7 @@ noncomputable instance instMonoidalCategory : MonoidalCategory (AlgebraCat.{u} R (forget₂ (AlgebraCat R) (ModuleCat R)) { μIso := fun _ _ => Iso.refl _ εIso := Iso.refl _ - associator_eq := fun _ _ _ => TensorProduct.ext₃ (fun _ _ _ => rfl) + associator_eq := fun _ _ _ => TensorProduct.ext_threefold (fun _ _ _ => rfl) leftUnitor_eq := fun _ => TensorProduct.ext' (fun _ _ => rfl) rightUnitor_eq := fun _ => TensorProduct.ext' (fun _ _ => rfl) } diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index 8225c569e1d95..f08d4d96958bb 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -526,13 +526,6 @@ theorem ext' {g h : M ⊗[R] N →ₗ[R] P} (H : ∀ x y, g (x ⊗ₜ y) = h (x TensorProduct.induction_on z (by simp_rw [LinearMap.map_zero]) H fun x y ihx ihy => by rw [g.map_add, h.map_add, ihx, ihy] -theorem ext₃ {g h : (M ⊗[R] N) ⊗[R] P →ₗ[R] Q} - (H : ∀ x y z, g (x ⊗ₜ y ⊗ₜ z) = h (x ⊗ₜ y ⊗ₜ z)) : g = h := - ext' fun x => TensorProduct.induction_on x - (fun x => by simp only [zero_tmul, map_zero]) - (fun x y => H x y) - (fun x y ihx ihy z => by rw [add_tmul, g.map_add, h.map_add, ihx, ihy]) - theorem lift.unique {g : M ⊗[R] N →ₗ[R] P} (H : ∀ x y, g (x ⊗ₜ y) = f x y) : g = lift f := ext' fun m n => by rw [H, lift.tmul] @@ -621,6 +614,8 @@ theorem ext_threefold {g h : (M ⊗[R] N) ⊗[R] P →ₗ[R] Q} ext x y z exact H x y z +@[deprecated (since := "2024-10-18")] alias ext₃ := ext_threefold + -- We'll need this one for checking the pentagon identity! theorem ext_fourfold {g h : ((M ⊗[R] N) ⊗[R] P) ⊗[R] Q →ₗ[R] S} (H : ∀ w x y z, g (w ⊗ₜ x ⊗ₜ y ⊗ₜ z) = h (w ⊗ₜ x ⊗ₜ y ⊗ₜ z)) : g = h := by From 15fb8246a353ad60012ade7c341d315e827d89d2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 20 Oct 2024 18:58:06 +0000 Subject: [PATCH 128/131] chore(Topology/Exterior): golf (#17955) Use `exterior_mono` to golf 3 lemmas --- Mathlib/Topology/Exterior.lean | 24 +++++++++--------------- 1 file changed, 9 insertions(+), 15 deletions(-) diff --git a/Mathlib/Topology/Exterior.lean b/Mathlib/Topology/Exterior.lean index 511879082540a..8b46e6e2b00e1 100644 --- a/Mathlib/Topology/Exterior.lean +++ b/Mathlib/Topology/Exterior.lean @@ -60,21 +60,6 @@ theorem exterior_union (s t : Set X) : exterior (s ∪ t) = exterior s ∪ exter theorem exterior_sUnion (S : Set (Set X)) : exterior (⋃₀ S) = ⋃ s ∈ S, exterior s := by simp only [sUnion_eq_biUnion, exterior_iUnion] -theorem exterior_iInter_subset {s : ι → Set X} : exterior (⋂ i, s i) ⊆ ⋂ i, exterior (s i) := by - simp_rw [exterior] - refine ker_mono (nhdsSet_iInter_le _) |>.trans_eq ?_ - simp_rw [ker_iInf] - -theorem exterior_inter_subset {s t : Set X} : exterior (s ∩ t) ⊆ exterior s ∩ exterior t := by - simp_rw [exterior] - refine ker_mono (nhdsSet_inter_le _ _) |>.trans_eq ?_ - rw [ker_inf _ _] - -theorem exterior_sInter_subset {s : Set (Set X)} : exterior (⋂₀ s) ⊆ ⋂ x ∈ s, exterior x := by - simp_rw [exterior] - refine ker_mono (nhdsSet_sInter_le _) |>.trans_eq ?_ - simp_rw [ker_iInf] - theorem mem_exterior_iff_specializes : x ∈ exterior s ↔ ∃ y ∈ s, x ⤳ y := calc x ∈ exterior s ↔ x ∈ exterior (⋃ y ∈ s, {y}) := by simp _ ↔ ∃ y ∈ s, x ⤳ y := by @@ -98,6 +83,15 @@ theorem exterior_eq_exterior_iff_nhdsSet : exterior s = exterior t ↔ 𝓝ˢ s lemma specializes_iff_exterior_subset : x ⤳ y ↔ exterior {x} ⊆ exterior {y} := by simp [Specializes] +theorem exterior_iInter_subset {s : ι → Set X} : exterior (⋂ i, s i) ⊆ ⋂ i, exterior (s i) := + exterior_mono.map_iInf_le + +theorem exterior_inter_subset {s t : Set X} : exterior (s ∩ t) ⊆ exterior s ∩ exterior t := + exterior_mono.map_inf_le _ _ + +theorem exterior_sInter_subset {s : Set (Set X)} : exterior (⋂₀ s) ⊆ ⋂ x ∈ s, exterior x := + exterior_mono.map_sInf_le + @[simp] lemma exterior_empty : exterior (∅ : Set X) = ∅ := isOpen_empty.exterior_eq @[simp] lemma exterior_univ : exterior (univ : Set X) = univ := isOpen_univ.exterior_eq From a3996466391c60195c5b37cdd1d7496f75962049 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Sun, 20 Oct 2024 19:39:56 +0000 Subject: [PATCH 129/131] feat(Analysis/Normed/Field/Ultra): nonarchimedean iff norm of nats le one (#15077) Co-authored-by: Yury G. Kudryashov Co-authored-by: Yakov Pechersky Co-authored-by: Yakov Pechersky --- Mathlib.lean | 1 + Mathlib/Analysis/Normed/Field/Ultra.lean | 131 +++++++++++++++++++++++ 2 files changed, 132 insertions(+) create mode 100644 Mathlib/Analysis/Normed/Field/Ultra.lean diff --git a/Mathlib.lean b/Mathlib.lean index ebba90f308a9c..a64206aa580a9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1239,6 +1239,7 @@ import Mathlib.Analysis.Normed.Field.Basic import Mathlib.Analysis.Normed.Field.InfiniteSum import Mathlib.Analysis.Normed.Field.Lemmas import Mathlib.Analysis.Normed.Field.ProperSpace +import Mathlib.Analysis.Normed.Field.Ultra import Mathlib.Analysis.Normed.Field.UnitBall import Mathlib.Analysis.Normed.Group.AddCircle import Mathlib.Analysis.Normed.Group.AddTorsor diff --git a/Mathlib/Analysis/Normed/Field/Ultra.lean b/Mathlib/Analysis/Normed/Field/Ultra.lean new file mode 100644 index 0000000000000..bbc28a3967fd1 --- /dev/null +++ b/Mathlib/Analysis/Normed/Field/Ultra.lean @@ -0,0 +1,131 @@ +/- +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.Analysis.Normed.Field.Basic +import Mathlib.Analysis.Normed.Group.Ultra + +/-! +## Sufficient conditions to have an ultrametric norm on a division ring + +This file provides ways of constructing an instance of `IsUltrametricDist` based on +facts about the existing norm. + +## Main results + +* `isUltrametricDist_of_forall_norm_natCast_le_one`: a norm in a division ring is ultrametric + if the norm of the image of a natural is less than or equal to one + +## Implementation details + +The proof relies on a bounded-from-above argument. The main result has a longer proof +to be able to be applied in noncommutative division rings. + +## Tags + +ultrametric, nonarchimedean +-/ +open Metric NNReal + +namespace IsUltrametricDist + +section sufficient + +variable {R : Type*} [NormedDivisionRing R] + +lemma isUltrametricDist_of_forall_norm_add_one_le_max_norm_one + (h : ∀ x : R, ‖x + 1‖ ≤ max ‖x‖ 1) : IsUltrametricDist R := by + refine isUltrametricDist_of_forall_norm_add_le_max_norm (fun x y ↦ ?_) + rcases eq_or_ne y 0 with rfl | hy + · simpa only [add_zero] using le_max_left _ _ + · have p : 0 < ‖y‖ := norm_pos_iff.mpr hy + simpa only [div_add_one hy, norm_div, div_le_iff₀ p, max_mul_of_nonneg _ _ p.le, one_mul, + div_mul_cancel₀ _ p.ne'] using h (x / y) + +lemma isUltrametricDist_of_forall_norm_add_one_of_norm_le_one + (h : ∀ x : R, ‖x‖ ≤ 1 → ‖x + 1‖ ≤ 1) : IsUltrametricDist R := by + refine isUltrametricDist_of_forall_norm_add_one_le_max_norm_one fun x ↦ ?_ + rcases le_or_lt ‖x‖ 1 with H|H + · exact (h _ H).trans (le_max_right _ _) + · suffices ‖x + 1‖ ≤ ‖x‖ from this.trans (le_max_left _ _) + rw [← div_le_div_right (c := ‖x‖) (H.trans' zero_lt_one), div_self (H.trans' zero_lt_one).ne', + ← norm_div, add_div, div_self (by simpa using (H.trans' zero_lt_one)), add_comm] + apply h + simp [inv_le_one_iff₀, H.le] + +lemma isUltrametricDist_of_forall_norm_sub_one_of_norm_le_one + (h : ∀ x : R, ‖x‖ ≤ 1 → ‖x - 1‖ ≤ 1) : IsUltrametricDist R := by + have (x : R) (hx : ‖x‖ ≤ 1) : ‖x + 1‖ ≤ 1 := by + simpa only [← neg_add', norm_neg] using h (-x) (norm_neg x ▸ hx) + exact isUltrametricDist_of_forall_norm_add_one_of_norm_le_one this + +/-- This technical lemma is used in the proof of +`isUltrametricDist_of_forall_norm_natCast_le_one`. -/ +lemma isUltrametricDist_of_forall_pow_norm_le_nsmul_pow_max_one_norm + (h : ∀ (x : R) (m : ℕ), ‖x + 1‖ ^ m ≤ (m + 1) • max 1 (‖x‖ ^ m)) : + IsUltrametricDist R := by + -- it will suffice to prove that `‖x + 1‖ ≤ max 1 ‖x‖` + refine isUltrametricDist_of_forall_norm_add_one_le_max_norm_one fun x ↦ ?_ + -- Morally, we want to deduce this from the hypothesis `h` by taking an `m`-th root and showing + -- that `(m + 1) ^ (1 / m)` gets arbitrarily close to 1, although we will formalise this in a way + -- that avoids explicitly mentioning `m`-th roots. + -- First note it suffices to show that `‖x + 1‖ ≤ a` for all `a : ℝ` with `max ‖x‖ 1 < a`. + rw [max_comm] + refine le_of_forall_le_of_dense fun a ha ↦ ?_ + have ha' : 1 < a := (max_lt_iff.mp ha).left + -- `max 1 ‖x‖ < a`, so there must be some `m : ℕ` such that `m + 1 < (a / max 1 ‖x‖) ^ m` + -- by the virtue of exponential growth being faster than linear growth + obtain ⟨m, hm⟩ : ∃ m : ℕ, ((m + 1) : ℕ) < (a / (max 1 ‖x‖)) ^ m := by + apply_mod_cast Real.exists_natCast_add_one_lt_pow_of_one_lt + rwa [one_lt_div (by positivity)] + -- and we rearrange again to get `(m + 1) • max 1 ‖x‖ ^ m < a ^ m` + rw [div_pow, lt_div_iff₀ (by positivity), ← nsmul_eq_mul] at hm + -- which squeezes down to get our `‖x + 1‖ ≤ a` using our to-be-proven hypothesis of + -- `‖x + 1‖ ^ m ≤ (m + 1) • max 1 ‖x‖ ^ m`, so we're done + -- we can distribute powers into the right term of `max` + have hp : max 1 ‖x‖ ^ m = max 1 (‖x‖ ^ m) := by + have : MonotoneOn (fun a : ℝ ↦ a ^ m) (Set.Ici _) := fun a h b _ h' ↦ pow_le_pow_left h h' _ + rw [this.map_max zero_le_one (norm_nonneg x), one_pow] + rw [hp] at hm + refine le_of_pow_le_pow_left (fun h ↦ ?_) (zero_lt_one.trans ha').le ((h _ _).trans hm.le) + simp only [h, zero_add, pow_zero, max_self, one_smul, lt_self_iff_false] at hm + +/-- To prove that a normed division ring is nonarchimedean, it suffices to prove that the norm +of the image of any natural is less than or equal to one. -/ +lemma isUltrametricDist_of_forall_norm_natCast_le_one + (h : ∀ n : ℕ, ‖(n : R)‖ ≤ 1) : IsUltrametricDist R := by + -- from a previous lemma, suffices to prove that for all `m`, we have + -- `‖x + 1‖ ^ m ≤ (m + 1) • max 1 ‖x‖ ^ m` + refine isUltrametricDist_of_forall_pow_norm_le_nsmul_pow_max_one_norm (fun x m ↦ ?_) + -- we first use our hypothesis about the norm of naturals to have that multiplication by + -- naturals keeps the norm small + replace h (x : R) (n : ℕ) : ‖n • x‖ ≤ ‖x‖ := by + rw [nsmul_eq_mul, norm_mul] + rcases (norm_nonneg x).eq_or_lt with hx | hx + · simp only [← hx, mul_zero, le_refl] + · simpa only [mul_le_iff_le_one_left hx] using h _ + -- we expand the LHS using the binomial theorem, and apply the hypothesis to bound each term by + -- a power of ‖x‖ + transitivity ∑ k ∈ Finset.range (m + 1), ‖x‖ ^ k + · simpa only [← norm_pow, (Commute.one_right x).add_pow, one_pow, mul_one, nsmul_eq_mul, + Nat.cast_comm] using (norm_sum_le _ _).trans (Finset.sum_le_sum fun _ _ ↦ h _ _) + -- the nature of the norm means that one of `1` and `‖x‖ ^ m` is the largest of the two, so the + -- other terms in the binomial expansion are bounded by the max of these, and the number of terms + -- in the sum is precisely `m + 1` + rw [← Finset.card_range (m + 1), ← Finset.sum_const, Finset.card_range] + rcases max_cases 1 (‖x‖ ^ m) with (⟨hm, hx⟩|⟨hm, hx⟩) <;> rw [hm] <;> + -- which we show by comparing the terms in the sum one by one + refine Finset.sum_le_sum fun i hi ↦ ?_ + · rcases eq_or_ne m 0 with rfl | hm + · simp only [pow_zero, le_refl, + show i = 0 by simpa only [zero_add, Finset.range_one, Finset.mem_singleton] using hi] + · rw [pow_le_one_iff_of_nonneg (norm_nonneg _) hm] at hx + exact pow_le_one₀ (norm_nonneg _) hx + · refine pow_le_pow_right₀ ?_ (by simpa only [Finset.mem_range, Nat.lt_succ] using hi) + contrapose! hx + exact pow_le_one₀ (norm_nonneg _) hx.le + +end sufficient + +end IsUltrametricDist From cf0049ec812ebc91c7295af712a53f415308a142 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Sun, 20 Oct 2024 19:54:50 +0000 Subject: [PATCH 130/131] feat(CategoryTheory): natural version of `FullyFaithful.homEquiv` (#17450) --- Mathlib/CategoryTheory/Yoneda.lean | 22 +++++++++++++++++++++- 1 file changed, 21 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 61b37dc22ad22..5116a9b934fd8 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -24,7 +24,7 @@ namespace CategoryTheory open Opposite -universe v v₁ u₁ u₂ +universe v v₁ v₂ u₁ u₂ -- morphism levels before object levels. See note [CategoryTheory universes]. variable {C : Type u₁} [Category.{v₁} C] @@ -767,4 +767,24 @@ lemma yonedaMap_app_apply {Y : C} {X : Cᵒᵖ} (f : X.unop ⟶ Y) : end +namespace Functor.FullyFaithful + +variable {C : Type u₁} [Category.{v₁} C] + +/-- `FullyFaithful.homEquiv` as a natural isomorphism. -/ +def homNatIso {D : Type u₂} [Category.{v₂} D] {F : C ⥤ D} (hF : F.FullyFaithful) (X : C) : + F.op ⋙ yoneda.obj (F.obj X) ⋙ uliftFunctor.{v₁} ≅ yoneda.obj X ⋙ uliftFunctor.{v₂} := + NatIso.ofComponents + (fun Y => Equiv.toIso (Equiv.ulift.trans <| hF.homEquiv.symm.trans Equiv.ulift.symm)) + (fun f => by ext; exact Equiv.ulift.injective (hF.map_injective (by simp))) + +/-- `FullyFaithful.homEquiv` as a natural isomorphism. -/ +def homNatIsoMaxRight {D : Type u₂} [Category.{max v₁ v₂} D] {F : C ⥤ D} (hF : F.FullyFaithful) + (X : C) : F.op ⋙ yoneda.obj (F.obj X) ≅ yoneda.obj X ⋙ uliftFunctor.{v₂} := + NatIso.ofComponents + (fun Y => Equiv.toIso (hF.homEquiv.symm.trans Equiv.ulift.symm)) + (fun f => by ext; exact Equiv.ulift.injective (hF.map_injective (by simp))) + +end Functor.FullyFaithful + end CategoryTheory From 1b6c758c8ee0b23cbafb1c5add718cd58333c143 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 20 Oct 2024 21:29:44 +0000 Subject: [PATCH 131/131] =?UTF-8?q?feat:=20`negOnePow=20n=20=3D=20(-1=20:?= =?UTF-8?q?=20=E2=84=A4)=20^=20n`=20(#17967)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and rename `Int.coe_negOnePow` to `Int.cast_negOnePow` for disambiguation --- Mathlib/Algebra/Ring/NegOnePow.lean | 12 +++++++++--- .../SpecialFunctions/Trigonometric/Basic.lean | 12 ++++++------ 2 files changed, 15 insertions(+), 9 deletions(-) diff --git a/Mathlib/Algebra/Ring/NegOnePow.lean b/Mathlib/Algebra/Ring/NegOnePow.lean index 6696881dcbf2e..9f52047d0c32c 100644 --- a/Mathlib/Algebra/Ring/NegOnePow.lean +++ b/Mathlib/Algebra/Ring/NegOnePow.lean @@ -101,11 +101,17 @@ lemma negOnePow_eq_iff (n₁ n₂ : ℤ) : lemma negOnePow_mul_self (n : ℤ) : (n * n).negOnePow = n.negOnePow := by simpa [mul_sub, negOnePow_eq_iff] using n.even_mul_pred_self -lemma coe_negOnePow (K : Type*) (n : ℤ) [Field K] : n.negOnePow = (-1 : K) ^ n := by +lemma cast_negOnePow (K : Type*) (n : ℤ) [Field K] : n.negOnePow = (-1 : K) ^ n := by rcases even_or_odd' n with ⟨k, rfl | rfl⟩ - · rw [zpow_mul, zpow_ofNat] - simp + · simp [zpow_mul, zpow_ofNat] · rw [zpow_add_one₀ (by norm_num), zpow_mul, zpow_ofNat] simp +@[deprecated (since := "2024-10-20")] alias coe_negOnePow := cast_negOnePow + +lemma cast_negOnePow_natCast (R : Type*) [Ring R] (n : ℕ) : negOnePow n = (-1 : R) ^ n := by + obtain ⟨k, rfl | rfl⟩ := Nat.even_or_odd' n <;> simp [pow_succ, pow_mul] + +lemma coe_negOnePow_natCast (n : ℕ) : negOnePow n = (-1 : ℤ) ^ n := cast_negOnePow_natCast .. + end Int diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index 8e6add75779d2..b09fb6620f62e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -280,19 +280,19 @@ theorem sin_int_mul_two_pi_sub (x : ℝ) (n : ℤ) : sin (n * (2 * π) - x) = -s sin_neg x ▸ sin_periodic.int_mul_sub_eq n theorem sin_add_int_mul_pi (x : ℝ) (n : ℤ) : sin (x + n * π) = (-1) ^ n * sin x := - n.coe_negOnePow ℝ ▸ sin_antiperiodic.add_int_mul_eq n + n.cast_negOnePow ℝ ▸ sin_antiperiodic.add_int_mul_eq n theorem sin_add_nat_mul_pi (x : ℝ) (n : ℕ) : sin (x + n * π) = (-1) ^ n * sin x := sin_antiperiodic.add_nat_mul_eq n theorem sin_sub_int_mul_pi (x : ℝ) (n : ℤ) : sin (x - n * π) = (-1) ^ n * sin x := - n.coe_negOnePow ℝ ▸ sin_antiperiodic.sub_int_mul_eq n + n.cast_negOnePow ℝ ▸ sin_antiperiodic.sub_int_mul_eq n theorem sin_sub_nat_mul_pi (x : ℝ) (n : ℕ) : sin (x - n * π) = (-1) ^ n * sin x := sin_antiperiodic.sub_nat_mul_eq n theorem sin_int_mul_pi_sub (x : ℝ) (n : ℤ) : sin (n * π - x) = -((-1) ^ n * sin x) := by - simpa only [sin_neg, mul_neg, Int.coe_negOnePow] using sin_antiperiodic.int_mul_sub_eq n + simpa only [sin_neg, mul_neg, Int.cast_negOnePow] using sin_antiperiodic.int_mul_sub_eq n theorem sin_nat_mul_pi_sub (x : ℝ) (n : ℕ) : sin (n * π - x) = -((-1) ^ n * sin x) := by simpa only [sin_neg, mul_neg] using sin_antiperiodic.nat_mul_sub_eq n @@ -363,19 +363,19 @@ theorem cos_int_mul_two_pi_sub (x : ℝ) (n : ℤ) : cos (n * (2 * π) - x) = co cos_neg x ▸ cos_periodic.int_mul_sub_eq n theorem cos_add_int_mul_pi (x : ℝ) (n : ℤ) : cos (x + n * π) = (-1) ^ n * cos x := - n.coe_negOnePow ℝ ▸ cos_antiperiodic.add_int_mul_eq n + n.cast_negOnePow ℝ ▸ cos_antiperiodic.add_int_mul_eq n theorem cos_add_nat_mul_pi (x : ℝ) (n : ℕ) : cos (x + n * π) = (-1) ^ n * cos x := cos_antiperiodic.add_nat_mul_eq n theorem cos_sub_int_mul_pi (x : ℝ) (n : ℤ) : cos (x - n * π) = (-1) ^ n * cos x := - n.coe_negOnePow ℝ ▸ cos_antiperiodic.sub_int_mul_eq n + n.cast_negOnePow ℝ ▸ cos_antiperiodic.sub_int_mul_eq n theorem cos_sub_nat_mul_pi (x : ℝ) (n : ℕ) : cos (x - n * π) = (-1) ^ n * cos x := cos_antiperiodic.sub_nat_mul_eq n theorem cos_int_mul_pi_sub (x : ℝ) (n : ℤ) : cos (n * π - x) = (-1) ^ n * cos x := - n.coe_negOnePow ℝ ▸ cos_neg x ▸ cos_antiperiodic.int_mul_sub_eq n + n.cast_negOnePow ℝ ▸ cos_neg x ▸ cos_antiperiodic.int_mul_sub_eq n theorem cos_nat_mul_pi_sub (x : ℝ) (n : ℕ) : cos (n * π - x) = (-1) ^ n * cos x := cos_neg x ▸ cos_antiperiodic.nat_mul_sub_eq n